From b178344d6d24ef2e4d9d4ae5f86fc55809a0b373 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 09:32:07 -0400 Subject: [PATCH 001/100] fix adaptive composition --- .../ZeroConcentrated/AdaptiveComposition.lean | 42 ++----------------- .../ZeroConcentrated/DP.lean | 35 +++++++++------- .../Mechanism/Properties.lean | 3 ++ 3 files changed, 27 insertions(+), 53 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index ac102d93..73417a5e 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -131,50 +131,14 @@ theorem privComposeAdaptive_zCDPBound {nq1 : List T → PMF U} {nq2 : U -> List zCDPBound (privComposeAdaptive nq1 nq2) (ε₁ + ε₂) := by rw [zCDPBound] intro α Hα l₁ l₂ Hneighbours - -- This step is loose - apply (@LE.le.trans _ _ _ (ENNReal.ofReal (1/2 * (ε₁)^2 * α + 1/2 * (ε₂)^2 * α : ℝ)) _ _ ?case_sq) - case case_sq => - apply ofReal_le_ofReal - -- 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 - refine mul_nonneg ?bc.bc.ha Hε₂ - refine mul_nonneg ?G Hε₁ - simp -- 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₂) α ≤ ENNReal.ofReal (1 / 2 * ε₂ ^ 2 * α)) := - ciSup_le fun x => h2 x α Hα l₁ l₂ Hneighbours + have conditional_ub : (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) ≤ ENNReal.ofReal ε₂ := -- ENNReal.ofReal (1 / 2 * ε₂ ^ 2 * α)) := + iSup_le fun i => h2 i α 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 => - rw [ENNReal.ofReal_add ?G1 ?G2] - case G1 => - simp - apply mul_nonneg - · apply mul_nonneg - · simp - · exact sq_nonneg ε₁ - · linarith - case G2 => - simp - apply mul_nonneg - · apply mul_nonneg - · simp - · exact sq_nonneg ε₂ - · linarith + rw [ENNReal.ofReal_add Hε₁ Hε₂] exact _root_.add_le_add (h1 α Hα l₁ l₂ Hneighbours) conditional_ub exact privComposeAdaptive_renyi_bound Hα Hneighbours HAC1 HAC2 diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean index f51ad1aa..b0a96b60 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean @@ -44,7 +44,7 @@ satisfying this bound are ``ε``-DP). -/ 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 * α) + RenyiDivergence (q l₁) (q l₂) α ≤ ENNReal.ofReal ε /-- All neighbouring queries are absolutely continuous @@ -63,18 +63,19 @@ lemma zCDP_mono {m : List T -> PMF U} {ε₁ ε₂ : NNReal} (H : ε₁ ≤ ε · assumption · rw [zCDPBound] at * intro α Hα l₁ l₂ N - apply (@le_trans _ _ _ (ENNReal.ofReal (1 / 2 * ↑ε₁ ^ 2 * α)) _ (Hε α Hα l₁ l₂ N)) - apply ENNReal.coe_mono - refine (Real.toNNReal_le_toNNReal_iff ?a.hp).mpr ?a.a - · apply mul_nonneg - · apply mul_nonneg - · simp - · simp - · linarith - · repeat rw [mul_assoc] - apply (mul_le_mul_iff_of_pos_left (by simp)).mpr - apply (mul_le_mul_iff_of_pos_right (by linarith)).mpr - apply pow_le_pow_left' H (OfNat.ofNat 2) + sorry + -- apply (@le_trans _ _ _ (ENNReal.ofReal (1 / 2 * ↑ε₁ ^ 2 * α)) _ (Hε α Hα l₁ l₂ N)) + -- apply ENNReal.coe_mono + -- refine (Real.toNNReal_le_toNNReal_iff ?a.hp).mpr ?a.a + -- · apply mul_nonneg + -- · apply mul_nonneg + -- · simp + -- · simp + -- · linarith + -- · repeat rw [mul_assoc] + -- apply (mul_le_mul_iff_of_pos_left (by simp)).mpr + -- apply (mul_le_mul_iff_of_pos_right (by linarith)).mpr + -- apply pow_le_pow_left' H (OfNat.ofNat 2) /-- Obtain an approximate DP bound from a zCDP bound, when ε > 0 and δ < 1 @@ -82,6 +83,8 @@ Obtain an approximate DP bound from a zCDP bound, when ε > 0 and δ < 1 lemma ApproximateDP_of_zCDP_pos_lt_one [Countable U] (m : Mechanism T U) (ε : ℝ) (Hε_pos : 0 < ε) (h : zCDPBound m ε) (Hm : ACNeighbour m) : ∀ δ : NNReal, (0 < (δ : ℝ)) -> ((δ : ℝ) < 1) -> DP' m (ε^2/2 + ε * (2*Real.log (1/δ))^(1/2 : ℝ)) δ := by + sorry + /- have Hε : 0 ≤ ε := by exact le_of_lt Hε_pos intro δ Hδ0 Hδ1 generalize Dε' : (ε^2/2 + ε * (2*Real.log (1/δ))^(1/2 : ℝ)) = ε' @@ -627,6 +630,7 @@ lemma ApproximateDP_of_zCDP_pos_lt_one [Countable U] (m : Mechanism T U) -- Conclude by simplification simp [add_comm] + -/ @@ -1873,6 +1877,8 @@ Convert ε-DP bound to `(1/2)ε²`-zCDP bound Note that `zCDPBound _ ε` corresponds to `(1/2)ε²`-zCDP (not `ε`-zCDP). -/ lemma ofDP_bound (ε : NNReal) (q' : List T -> PMF U) (H : SLang.PureDP q' ε) : zCDPBound q' ε := by + sorry + /- rw [zCDPBound] intro α Hα l₁ l₂ HN -- Special case: (εα/2 > 1) @@ -2292,6 +2298,7 @@ lemma ofDP_bound (ε : NNReal) (q' : List T -> PMF U) (H : SLang.PureDP q' ε) : apply Eq.le congr 1 linarith + -/ /- Convert ε-DP to `(1/2)ε²`-zCDP. @@ -2301,4 +2308,4 @@ Note that `zCDPBound _ ε` corresponds to `(1/2)ε²`-zCDP (not `ε`-zCDP). lemma ofDP (ε : NNReal) (q : List T -> PMF U) (H : SLang.PureDP q ε) : zCDP q ε := by constructor · exact ACNeighbour_of_DP ε q H - · exact ofDP_bound ε q H \ No newline at end of file + · exact ofDP_bound ε q H diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean index d3653614..72137a87 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean @@ -29,6 +29,8 @@ theorem privNoisedQuery_zCDPBound (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ have A := @discrete_GaussianGenSample_ZeroConcentrated α h1 (Δ * ε₂) ε₁ (query l₁) (query l₂) apply le_trans A clear A + sorry + /- -- Turn it into an equality ASAP rw [sensitivity] at bounded_sensitivity @@ -186,6 +188,7 @@ theorem privNoisedQuery_zCDPBound (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ case G1 => exact NNReal.zero_le_coe congr simp only [Real.toNNReal_coe] + -/ lemma discrete_gaussian_shift {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) (τ x : ℤ) : discrete_gaussian σ μ (x - τ) = discrete_gaussian σ (μ + τ) (x) := by From d0b4650ae2fd22fa1b0dd84a40961754bb49df16 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 09:39:34 -0400 Subject: [PATCH 002/100] fix definition --- .../ZeroConcentrated/AdaptiveComposition.lean | 13 ++++++++++--- .../DifferentialPrivacy/ZeroConcentrated/DP.lean | 2 +- 2 files changed, 11 insertions(+), 4 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean index 73417a5e..7358ef78 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/AdaptiveComposition.lean @@ -134,11 +134,18 @@ theorem privComposeAdaptive_zCDPBound {nq1 : List T → PMF U} {nq2 : U -> List -- 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₂) α) ≤ ENNReal.ofReal ε₂ := -- ENNReal.ofReal (1 / 2 * ε₂ ^ 2 * α)) := - iSup_le fun i => h2 i α Hα l₁ l₂ Hneighbours + have conditional_ub : (⨆ (u : U), RenyiDivergence (nq2 u l₁) (nq2 u l₂) α) ≤ ENNReal.ofReal (ε₂ * α) := + by exact iSup_le fun i => h2 i α 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 => - rw [ENNReal.ofReal_add Hε₁ Hε₂] + rw [add_mul] + rw [ENNReal.ofReal_add ?G1 ?G2] + case G1 => + apply Right.mul_nonneg Hε₁ + linarith + case G2 => + apply Right.mul_nonneg Hε₂ + linarith exact _root_.add_le_add (h1 α Hα l₁ l₂ Hneighbours) conditional_ub exact privComposeAdaptive_renyi_bound Hα Hneighbours HAC1 HAC2 diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean index b0a96b60..09828b5a 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean @@ -44,7 +44,7 @@ satisfying this bound are ``ε``-DP). -/ def zCDPBound (q : List T → PMF U) (ε : ℝ) : Prop := ∀ α : ℝ, 1 < α → ∀ l₁ l₂ : List T, Neighbour l₁ l₂ → - RenyiDivergence (q l₁) (q l₂) α ≤ ENNReal.ofReal ε + RenyiDivergence (q l₁) (q l₂) α ≤ ENNReal.ofReal (ε * α) /-- All neighbouring queries are absolutely continuous From 7c1eac8ca6e42f0aa6ce3b71c6a68f269a83f4ea Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 09:47:58 -0400 Subject: [PATCH 003/100] update base proof for gaussian mechanism property --- .../ZeroConcentrated/Mechanism/Properties.lean | 10 ++++------ 1 file changed, 4 insertions(+), 6 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean index 72137a87..0b4c1ecb 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean @@ -23,14 +23,12 @@ namespace SLang 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 Δ ε₁ ε₂) ((ε₁ : NNReal) / ε₂) := by + zCDPBound (privNoisedQuery query Δ ε₁ ε₂) ((1/2) * ((ε₁ : NNReal) / ε₂) ^ 2) := 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 - sorry - /- -- Turn it into an equality ASAP rw [sensitivity] at bounded_sensitivity @@ -188,7 +186,6 @@ theorem privNoisedQuery_zCDPBound (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ case G1 => exact NNReal.zero_le_coe congr simp only [Real.toNNReal_coe] - -/ lemma discrete_gaussian_shift {σ : ℝ} (h : σ ≠ 0) (μ : ℝ) (τ x : ℤ) : discrete_gaussian σ μ (x - τ) = discrete_gaussian σ (μ + τ) (x) := by @@ -228,8 +225,9 @@ theorem privNoisedQuery_zCDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (b simp [zCDP] apply And.intro · exact privNoisedQuery_AC query Δ ε₁ ε₂ - · apply privNoisedQuery_zCDPBound - exact bounded_sensitivity + · sorry + -- apply privNoisedQuery_zCDPBound + -- exact bounded_sensitivity end SLang From 0b0370b819cf1ad2bce2d14689521eb04fb7f51f Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 09:54:36 -0400 Subject: [PATCH 004/100] fix mono --- .../ZeroConcentrated/DP.lean | 23 ++++++++----------- 1 file changed, 10 insertions(+), 13 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean index 09828b5a..9f3b6f76 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean @@ -63,19 +63,16 @@ lemma zCDP_mono {m : List T -> PMF U} {ε₁ ε₂ : NNReal} (H : ε₁ ≤ ε · assumption · rw [zCDPBound] at * intro α Hα l₁ l₂ N - sorry - -- apply (@le_trans _ _ _ (ENNReal.ofReal (1 / 2 * ↑ε₁ ^ 2 * α)) _ (Hε α Hα l₁ l₂ N)) - -- apply ENNReal.coe_mono - -- refine (Real.toNNReal_le_toNNReal_iff ?a.hp).mpr ?a.a - -- · apply mul_nonneg - -- · apply mul_nonneg - -- · simp - -- · simp - -- · linarith - -- · repeat rw [mul_assoc] - -- apply (mul_le_mul_iff_of_pos_left (by simp)).mpr - -- apply (mul_le_mul_iff_of_pos_right (by linarith)).mpr - -- apply pow_le_pow_left' H (OfNat.ofNat 2) + apply (@le_trans _ _ _ (ENNReal.ofReal (ε₁ * α)) _ ?G1) + case G1 => apply Hε <;> trivial + apply ENNReal.coe_mono + refine (Real.toNNReal_le_toNNReal_iff ?a.hp).mpr ?a.a + · apply mul_nonneg + · exact NNReal.zero_le_coe + · linarith + · apply mul_le_mul_of_nonneg_right + · exact H + · linarith /-- Obtain an approximate DP bound from a zCDP bound, when ε > 0 and δ < 1 From 5f6ef79a5b53dcaac41674fcba39b525b4772146 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 10:02:32 -0400 Subject: [PATCH 005/100] fix base proof for approxDP --- .../ZeroConcentrated/DP.lean | 24 +++++++++++-------- 1 file changed, 14 insertions(+), 10 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean index 9f3b6f76..44977265 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean @@ -78,10 +78,8 @@ lemma zCDP_mono {m : List T -> PMF U} {ε₁ ε₂ : NNReal} (H : ε₁ ≤ ε Obtain an approximate DP bound from a zCDP bound, when ε > 0 and δ < 1 -/ lemma ApproximateDP_of_zCDP_pos_lt_one [Countable U] (m : Mechanism T U) - (ε : ℝ) (Hε_pos : 0 < ε) (h : zCDPBound m ε) (Hm : ACNeighbour m) : + (ε : ℝ) (Hε_pos : 0 < ε) (h : zCDPBound m ((1/2) * ε^2)) (Hm : ACNeighbour m) : ∀ δ : NNReal, (0 < (δ : ℝ)) -> ((δ : ℝ) < 1) -> DP' m (ε^2/2 + ε * (2*Real.log (1/δ))^(1/2 : ℝ)) δ := by - sorry - /- have Hε : 0 ≤ ε := by exact le_of_lt Hε_pos intro δ Hδ0 Hδ1 generalize Dε' : (ε^2/2 + ε * (2*Real.log (1/δ))^(1/2 : ℝ)) = ε' @@ -627,7 +625,6 @@ lemma ApproximateDP_of_zCDP_pos_lt_one [Countable U] (m : Mechanism T U) -- Conclude by simplification simp [add_comm] - -/ @@ -635,12 +632,11 @@ lemma ApproximateDP_of_zCDP_pos_lt_one [Countable U] (m : Mechanism T U) Obtain an approximate DP bound from a zCDP bound, when ε > 0 -/ lemma ApproximateDP_of_zCDP_pos [Countable U] (m : Mechanism T U) - (ε : ℝ) (Hε_pos : 0 < ε) (h : zCDPBound m ε) (Hm : ACNeighbour m) : + (ε : ℝ) (Hε_pos : 0 < ε) (h : zCDPBound m ((1/2) * ε^2)) (Hm : ACNeighbour m) : ∀ δ : NNReal, (0 < (δ : ℝ)) -> DP' m (ε^2/2 + ε * (2*Real.log (1/δ))^(1/2 : ℝ)) δ := by intro δ Hδ0 cases (Classical.em (δ < 1)) - · intro Hδ1 - apply ApproximateDP_of_zCDP_pos_lt_one m ε Hε_pos h Hm δ Hδ0 + · apply ApproximateDP_of_zCDP_pos_lt_one m ε Hε_pos h Hm δ Hδ0 trivial · apply ApproximateDP_gt1 apply le_of_not_lt @@ -650,7 +646,7 @@ lemma ApproximateDP_of_zCDP_pos [Countable U] (m : Mechanism T U) Obtain an approximate DP bound from a zCDP bound -/ theorem ApproximateDP_of_zCDP [Countable U] (m : Mechanism T U) - (ε : ℝ) (Hε : 0 ≤ ε) (h : zCDPBound m ε) (Hm : ACNeighbour m) : + (ε : ℝ) (Hε : 0 ≤ ε) (h : zCDPBound m ((1/2) * ε^2)) (Hm : ACNeighbour m) : ∀ δ : NNReal, (0 < (δ : ℝ)) -> DP' m (ε^2/2 + ε * (2*Real.log (1/δ))^(1/2 : ℝ)) δ := by cases LE.le.lt_or_eq Hε · rename_i Hε @@ -674,7 +670,7 @@ zCDP is no weaker than approximate DP, up to a loss of parameters. -/ lemma zCDP_ApproximateDP [Countable U] {m : Mechanism T U} : ∃ (degrade : (δ : NNReal) -> (ε' : NNReal) -> NNReal), ∀ (δ : NNReal) (_ : 0 < δ) (ε' : NNReal), - (zCDP m (degrade δ ε') -> ApproximateDP m ε' δ) := by + (zCDP m (degrade δ ((1/2) * ε'^2)) -> ApproximateDP m ε' δ) := by let degrade (δ : NNReal) (ε' : NNReal) : NNReal := (√(2 * Real.log (1/δ) + 2 * ε') - √(2 * Real.log (1/δ))).toNNReal have HDdegrade δ ε' : degrade δ ε' = (√(2 * Real.log (1/δ) + 2 * ε') - √(2 * Real.log (1/δ))).toNNReal := by rfl @@ -687,8 +683,15 @@ lemma zCDP_ApproximateDP [Countable U] {m : Mechanism T U} : rename_i Hδ1 rw [ApproximateDP] - have R := ApproximateDP_of_zCDP m (degrade δ ε') (by simp) HB HN δ Hδ + + have R := ApproximateDP_of_zCDP m (degrade δ ε') (by simp) ?G1 HN δ Hδ + case G1 => + -- this proof has to be redone + sorry + sorry + + /- have Hdegrade : ((degrade δ ε') ^ 2) / 2 + (degrade δ ε') * (2 * Real.log (1 / δ))^(1/2 : ℝ) = ε' := by rw [HDdegrade] generalize HD : Real.log (1 / δ) = D @@ -733,6 +736,7 @@ lemma zCDP_ApproximateDP [Countable U] {m : Mechanism T U} : linarith rw [Hdegrade] at R trivial + -/ /-- From 95be45cb014a53353a594a25169b7573a2204d46 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 10:07:57 -0400 Subject: [PATCH 006/100] new zCDP proofs wrapped --- SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean | 7 ++----- .../ZeroConcentrated/Mechanism/Properties.lean | 10 ++++------ .../DifferentialPrivacy/ZeroConcentrated/System.lean | 4 ++-- 3 files changed, 8 insertions(+), 13 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean index 44977265..1bbb753d 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean @@ -1877,9 +1877,7 @@ Convert ε-DP bound to `(1/2)ε²`-zCDP bound Note that `zCDPBound _ ε` corresponds to `(1/2)ε²`-zCDP (not `ε`-zCDP). -/ -lemma ofDP_bound (ε : NNReal) (q' : List T -> PMF U) (H : SLang.PureDP q' ε) : zCDPBound q' ε := by - sorry - /- +lemma ofDP_bound (ε : NNReal) (q' : List T -> PMF U) (H : SLang.PureDP q' ε) : zCDPBound q' ((1/2) * ε^2) := by rw [zCDPBound] intro α Hα l₁ l₂ HN -- Special case: (εα/2 > 1) @@ -2299,14 +2297,13 @@ lemma ofDP_bound (ε : NNReal) (q' : List T -> PMF U) (H : SLang.PureDP q' ε) : apply Eq.le congr 1 linarith - -/ /- Convert ε-DP to `(1/2)ε²`-zCDP. Note that `zCDPBound _ ε` corresponds to `(1/2)ε²`-zCDP (not `ε`-zCDP). -/ -lemma ofDP (ε : NNReal) (q : List T -> PMF U) (H : SLang.PureDP q ε) : zCDP q ε := by +lemma ofDP (ε : NNReal) (q : List T -> PMF U) (H : SLang.PureDP q ε) : zCDP q ((1/2) * ε^2) := by constructor · exact ACNeighbour_of_DP ε q H · exact ofDP_bound ε q H diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean index 0b4c1ecb..8838f8db 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Mechanism/Properties.lean @@ -221,13 +221,11 @@ def privNoisedQuery_AC (query : List T -> ℤ) (Δ ε₁ ε₂ : ℕ+) : ACNeigh The zCDP mechanism is ``(Δε₂/ε₁)^2``-zCDP. -/ theorem privNoisedQuery_zCDP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (bounded_sensitivity : sensitivity query Δ) : - zCDP (privNoisedQuery query Δ ε₁ ε₂) ((ε₁ : NNReal) / ε₂) := by - simp [zCDP] + zCDP (privNoisedQuery query Δ ε₁ ε₂) ((1/2) * ((ε₁ : NNReal) / ε₂) ^ 2) := by + simp only [zCDP] apply And.intro · exact privNoisedQuery_AC query Δ ε₁ ε₂ - · sorry - -- apply privNoisedQuery_zCDPBound - -- exact bounded_sensitivity - + · apply privNoisedQuery_zCDPBound + exact bounded_sensitivity end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean index de28f44d..5381e875 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean @@ -25,10 +25,10 @@ Instance of a DP system for zCDP, using the discrete Gaussian as a noising mecha -/ noncomputable instance gaussian_zCDPSystem : DPSystem T where prop := zCDP - prop_adp := zCDP_ApproximateDP + prop_adp := sorry -- zCDP_ApproximateDP prop_mono := zCDP_mono noise := privNoisedQuery - noise_prop := privNoisedQuery_zCDP + noise_prop := sorry -- privNoisedQuery_zCDP adaptive_compose_prop := privComposeAdaptive_zCDP postprocess_prop := privPostProcess_zCDP const_prop := privConst_zCDP From 569d44e07f8bb90f98cf439170c93cb5f2f6a680 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 11:15:08 -0400 Subject: [PATCH 007/100] stub changes to system --- SampCert.lean | 4 +- SampCert/DifferentialPrivacy/Abstract.lean | 42 ++++++++++++------- SampCert/DifferentialPrivacy/Pure/System.lean | 15 +++++-- .../Queries/BoundedMean/Code.lean | 1 + .../Queries/BoundedMean/Properties.lean | 1 + .../Queries/BoundedSum/Code.lean | 3 +- .../Queries/BoundedSum/Properties.lean | 6 ++- .../Queries/Count/Code.lean | 3 +- .../Queries/Count/Properties.lean | 9 +++- .../Queries/Histogram/Code.lean | 3 +- .../Queries/Histogram/Properties.lean | 6 ++- .../Queries/HistogramMean/Code.lean | 3 +- .../Queries/HistogramMean/Properties.lean | 3 +- .../ZeroConcentrated/System.lean | 26 ++++++++++-- 14 files changed, 88 insertions(+), 37 deletions(-) diff --git a/SampCert.lean b/SampCert.lean index 55d9f9b0..1d7fc16e 100644 --- a/SampCert.lean +++ b/SampCert.lean @@ -14,7 +14,7 @@ import Init.Data.UInt.Lemmas open SLang PMF -def combineConcentrated := @privNoisedBoundedMean_DP gaussian_zCDPSystem +def combineConcentrated := @privNoisedBoundedMean_DP _ gaussian_zCDPSystem def combinePure := @privNoisedBoundedMean_DP PureDPSystem /- @@ -41,7 +41,7 @@ Return an upper bound on the bin value, clipped to 2^(1 + numBins) def unbin (n : Fin numBins) : ℕ+ := 2 ^ (1 + n.val) noncomputable def combineMeanHistogram : Mechanism ℕ (Option ℚ) := - privMeanHistogram PureDPSystem numBins { bin } unbin 1 20 2 1 20 + privMeanHistogram PureDPSystem laplace_pureDPSystem numBins { bin } unbin 1 20 2 1 20 end histogramMeanExample diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index 84d7fb13..ee4813f6 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -15,8 +15,6 @@ import SampCert.DifferentialPrivacy.Approximate.DP This file defines an abstract system of differentially private operators. -/ -noncomputable section - open Classical Nat Int Real ENNReal namespace SLang @@ -31,25 +29,18 @@ class DPSystem (T : Type) where prop : Mechanism T Z → NNReal → Prop /-- - For any δ, prop implies ``ApproximateDP δ ε`` up to a sufficient degradation - of the privacy parameter. + Definition of DP is well-formed: Privacy parameter required to obtain (ε', δ)-approximate DP -/ - prop_adp [Countable Z] {m : Mechanism T Z} : - ∃ (degrade : (δ : NNReal) -> (ε' : NNReal) -> NNReal), ∀ (δ : NNReal) (_ : 0 < δ) (ε' : NNReal), - (prop m (degrade δ ε') -> ApproximateDP m ε' δ) + of_adp : (δ : NNReal) -> (ε' : NNReal) -> NNReal /-- - DP is monotonic + For any ε', this definition of DP implies (ε', δ)-approximate-DP for all δ -/ - prop_mono {m : Mechanism T Z} {ε₁ ε₂: NNReal} (Hε : ε₁ ≤ ε₂) (H : prop m ε₁) : prop m ε₂ + prop_adp [Countable Z] {m : Mechanism T Z} : ∀ (δ : NNReal) (_ : 0 < δ) (ε' : NNReal), + (prop m (of_adp δ ε') -> ApproximateDP m ε' δ) /-- - A noise mechanism (eg. Laplace, Discrete Gaussian, etc) - Paramaterized by a query, sensitivity, and a (rational) security paramater. - -/ - noise : Query T ℤ → (sensitivity : ℕ+) → (num : ℕ+) → (den : ℕ+) → Mechanism T ℤ - /-- - Adding noise to a query makes it private. + DP is monotonic -/ - noise_prop : ∀ q : List T → ℤ, ∀ Δ εn εd : ℕ+, sensitivity q Δ → prop (noise q Δ εn εd) (εn / εd) + prop_mono {m : Mechanism T Z} {ε₁ ε₂: NNReal} (Hε : ε₁ ≤ ε₂) (H : prop m ε₁) : prop m ε₂ /-- Privacy adaptively composes by addition. -/ @@ -96,4 +87,23 @@ lemma bind_bind_indep (p : Mechanism T U) (q : Mechanism T V) (h : U → V → P ext l x simp [privCompose, privComposeAdaptive, tsum_prod'] + +/-- +A noise function for a differential privacy system +-/ +class DPNoise (dps : DPSystem T) where + /-- + A noise mechanism (eg. Laplace, Discrete Gaussian, etc) + Paramaterized by a query, sensitivity, and a (rational) security paramater. + -/ + noise : Query T ℤ → (sensitivity : ℕ+) → (num : ℕ+) → (den : ℕ+) → Mechanism T ℤ + /-- + How much privacy one obtains by adding Δ εn εd noise to a Δ-sensitive query + -/ + noise_priv : NNReal → NNReal + /-- + Adding noise to a query makes it private. + -/ + noise_prop : ∀ q : List T → ℤ, ∀ Δ εn εd : ℕ+, sensitivity q Δ → dps.prop (noise q Δ εn εd) (noise_priv (εn / εd)) + end SLang diff --git a/SampCert/DifferentialPrivacy/Pure/System.lean b/SampCert/DifferentialPrivacy/Pure/System.lean index 00e16cd0..a2b26545 100644 --- a/SampCert/DifferentialPrivacy/Pure/System.lean +++ b/SampCert/DifferentialPrivacy/Pure/System.lean @@ -21,14 +21,21 @@ variable { T : Type } /-- Pure ε-DP with noise drawn from the discrete Laplace distribution. -/ -noncomputable instance PureDPSystem : DPSystem T where +instance PureDPSystem : DPSystem T where prop := PureDP - prop_adp := pure_ApproximateDP + of_adp := sorry + prop_adp := sorry -- pure_ApproximateDP prop_mono := PureDP_mono - noise := privNoisedQueryPure - noise_prop := privNoisedQueryPure_DP adaptive_compose_prop := PureDP_ComposeAdaptive' postprocess_prop := PureDP_PostProcess const_prop := PureDP_privConst + +instance laplace_pureDPSystem : DPNoise (@PureDPSystem T) where + noise := privNoisedQueryPure + noise_priv := sorry + noise_prop := sorry -- privNoisedQueryPure_DP + + + end SLang diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Code.lean b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Code.lean index 2cf226d8..79234d71 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Code.lean @@ -18,6 +18,7 @@ noncomputable section namespace SLang variable [dps : DPSystem ℕ] +variable [dpn : DPNoise dps] /-- Compute a noised mean using a noised sum and noised count. diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean index 775a6435..e157ce3c 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean @@ -21,6 +21,7 @@ noncomputable section namespace SLang variable [dps : DPSystem ℕ] +variable [dpn : DPNoise dps] lemma budget_split (ε₁ ε₂ : ℕ+) : (ε₁ : NNReal) / (ε₂ : NNReal) = (ε₁ : NNReal) / ((2 * ε₂) : ℕ+) + (ε₁ : NNReal) / ((2 * ε₂) : ℕ+) := by diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedSum/Code.lean b/SampCert/DifferentialPrivacy/Queries/BoundedSum/Code.lean index 12a37d7a..ffdc6f98 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedSum/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedSum/Code.lean @@ -19,6 +19,7 @@ noncomputable section namespace SLang variable [dps : DPSystem ℕ] +variable [dpn : DPNoise dps] /-- Bounded sum query: computes a sum and truncates it at an upper bound. @@ -30,6 +31,6 @@ 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 ℕ) : PMF ℤ := do - dps.noise (exactBoundedSum U) U ε₁ ε₂ l + dpn.noise (exactBoundedSum U) U ε₁ ε₂ l end SLang diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedSum/Properties.lean b/SampCert/DifferentialPrivacy/Queries/BoundedSum/Properties.lean index 70e72bfd..4f21348b 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedSum/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedSum/Properties.lean @@ -22,6 +22,7 @@ noncomputable section namespace SLang variable [dps : DPSystem ℕ] +variable [dpn : DPNoise dps] /-- Sensitivity of the bounded sum is equal to the bound. @@ -86,7 +87,8 @@ The noised bounded sum satisfies the DP property of the DP system. @[simp] theorem privNoisedBoundedSum_DP (U : ℕ+) (ε₁ ε₂ : ℕ+) : dps.prop (privNoisedBoundedSum U ε₁ ε₂) ((ε₁ : NNReal) / ε₂) := by - apply dps.noise_prop - apply exactBoundedSum_sensitivity + sorry + -- apply dpn.noise_prop + -- apply exactBoundedSum_sensitivity end SLang diff --git a/SampCert/DifferentialPrivacy/Queries/Count/Code.lean b/SampCert/DifferentialPrivacy/Queries/Count/Code.lean index a94db31e..b38929ef 100644 --- a/SampCert/DifferentialPrivacy/Queries/Count/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/Count/Code.lean @@ -17,6 +17,7 @@ namespace SLang variable {T : Type} variable [dps : DPSystem T] +variable [dpn : DPNoise dps] /-- Query to count the size of the input list. @@ -27,6 +28,6 @@ def exactCount (l : List T) : ℤ := List.length l A noised counting mechanism. -/ def privNoisedCount (ε₁ ε₂ : ℕ+) (l : List T) : PMF ℤ := do - dps.noise exactCount 1 ε₁ ε₂ l + dpn.noise exactCount 1 ε₁ ε₂ l end SLang diff --git a/SampCert/DifferentialPrivacy/Queries/Count/Properties.lean b/SampCert/DifferentialPrivacy/Queries/Count/Properties.lean index 8257d763..3dfc5377 100644 --- a/SampCert/DifferentialPrivacy/Queries/Count/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/Count/Properties.lean @@ -21,6 +21,7 @@ namespace SLang variable {T : Type} variable [dps : DPSystem T] +variable [dpn : DPNoise dps] /-- The counting query is 1-sensitive @@ -46,7 +47,11 @@ The noised counting query satisfies DP property @[simp] theorem privNoisedCount_DP (ε₁ ε₂ : ℕ+) : dps.prop (privNoisedCount ε₁ ε₂) ((ε₁ : NNReal) / ε₂) := by - apply dps.noise_prop - apply exactCount_1_sensitive + apply DPSystem_prop_ext + case H => + sorry + all_goals sorry + -- apply dpn.noise_prop + -- apply exactCount_1_sensitive end SLang diff --git a/SampCert/DifferentialPrivacy/Queries/Histogram/Code.lean b/SampCert/DifferentialPrivacy/Queries/Histogram/Code.lean index 553cdedc..9342a6c9 100644 --- a/SampCert/DifferentialPrivacy/Queries/Histogram/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/Histogram/Code.lean @@ -87,6 +87,7 @@ instance : DiscreteMeasurableSpace (Histogram T numBins B) where namespace SLang variable [dps : DPSystem T] +variable [dpn : DPNoise dps] /-- Compute the exact number of elements inside a histogram bin @@ -98,7 +99,7 @@ def exactBinCount (b : Fin numBins) (l : List T) : ℤ := Compute a noised count of the number of list elements inside a particular histogram bin -/ def privNoisedBinCount (ε₁ ε₂ : ℕ+) (b : Fin numBins) : Mechanism T ℤ := - (dps.noise (exactBinCount numBins B b) 1 ε₁ (ε₂ * numBins)) + (dpn.noise (exactBinCount numBins B b) 1 ε₁ (ε₂ * numBins)) /-- Modify a count inside a Histogram diff --git a/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean b/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean index 155c66b8..531cafd7 100644 --- a/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean @@ -21,6 +21,7 @@ namespace SLang variable {T : Type} variable [dps : DPSystem T] +variable [dpn : DPNoise dps] variable [HT : Inhabited T] variable (numBins : ℕ+) @@ -61,8 +62,9 @@ DP bound for a noised bin count lemma privNoisedBinCount_DP (ε₁ ε₂ : ℕ+) (b : Fin numBins) : dps.prop (privNoisedBinCount numBins B ε₁ ε₂ b) (ε₁ / ((ε₂ * numBins : PNat))) := by unfold privNoisedBinCount - apply dps.noise_prop - apply exactBinCount_sensitivity + sorry + -- apply dpn.noise_prop + -- apply exactBinCount_sensitivity -- MARKUSDE: Looking at this proof it is clear that we need better tactic support for the abstract DP operators -- MARKUSDE: - Lemmas with equality side conditions for the privacy cost diff --git a/SampCert/DifferentialPrivacy/Queries/HistogramMean/Code.lean b/SampCert/DifferentialPrivacy/Queries/HistogramMean/Code.lean index 4a91528c..6bb35172 100644 --- a/SampCert/DifferentialPrivacy/Queries/HistogramMean/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/HistogramMean/Code.lean @@ -23,6 +23,7 @@ Implementation for noncomputable section variable (dps : SLang.DPSystem ℕ) +variable (dpn : SLang.DPNoise dps) variable (numBins : ℕ+) variable (B : Bins ℕ numBins) @@ -44,7 +45,7 @@ and the bounded mean with (ε₃/ε₄)-DP. def privMeanHistogram (ε₁ ε₂ : ℕ+) (τ : ℤ) (ε₃ ε₄ : ℕ+) : Mechanism ℕ (Option ℚ) := privPostProcess (privComposeAdaptive - (@privMaxBinAboveThreshold numBins _ B dps ε₁ ε₂ τ) + (@privMaxBinAboveThreshold numBins _ B dps dpn ε₁ ε₂ τ) (fun opt_max => match opt_max with | none => privConst none diff --git a/SampCert/DifferentialPrivacy/Queries/HistogramMean/Properties.lean b/SampCert/DifferentialPrivacy/Queries/HistogramMean/Properties.lean index 22892f3f..05573c9e 100644 --- a/SampCert/DifferentialPrivacy/Queries/HistogramMean/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/HistogramMean/Properties.lean @@ -22,6 +22,7 @@ noncomputable section namespace SLang variable (dps : DPSystem ℕ) +variable (dpn : DPNoise dps) variable (numBins : ℕ+) variable (B : Bins ℕ numBins) variable (unbin : Fin numBins -> ℕ+) @@ -60,7 +61,7 @@ instance : DiscreteMeasurableSpace (Option (Fin ↑numBins)) where DP bound for the adaptive mean -/ lemma privMeanHistogram_DP (ε₁ ε₂ : ℕ+) (τ : ℤ) (ε₃ ε₄ : ℕ+) : - dps.prop (privMeanHistogram dps numBins B unbin ε₁ ε₂ τ ε₃ ε₄) (ε₁/ε₂ + ε₃/ε₄) := by + dps.prop (privMeanHistogram dps dpn numBins B unbin ε₁ ε₂ τ ε₃ ε₄) (ε₁/ε₂ + ε₃/ε₄) := by rw [privMeanHistogram] apply dps.postprocess_prop apply dps.adaptive_compose_prop diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean index 5381e875..22b26b58 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean @@ -21,16 +21,34 @@ namespace SLang variable { T : Type } /-- -Instance of a DP system for zCDP, using the discrete Gaussian as a noising mechanism. +Instance of a DP system for zCDP -/ -noncomputable instance gaussian_zCDPSystem : DPSystem T where +instance zCDPSystem : DPSystem T where prop := zCDP + of_adp := sorry prop_adp := sorry -- zCDP_ApproximateDP prop_mono := zCDP_mono - noise := privNoisedQuery - noise_prop := sorry -- privNoisedQuery_zCDP + -- noise := privNoisedQuery + -- noise_prop := sorry -- privNoisedQuery_zCDP adaptive_compose_prop := privComposeAdaptive_zCDP postprocess_prop := privPostProcess_zCDP const_prop := privConst_zCDP +/-- +Gaussian noise for zCDP system +-/ +instance gaussian_zCDPSystem : DPNoise zCDPsystem where + noise := privNoisedQuery + noise_priv := sorry + noise_prop := sorry -- privNoisedQuery_zCDP + +/-- +Laplace noise for zCDP system +-/ +instance laplace_zCDPSystem : DPNoise zCDPsystem where + noise := sorry + noise_priv := sorry + noise_prop := sorry -- privNoisedQuery_zCDP + + end SLang From 28efc75dde032cc7e53020ab46dd1198119b8fb0 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 11:59:53 -0400 Subject: [PATCH 008/100] stub changes to noise --- SampCert/DifferentialPrivacy/Abstract.lean | 39 ++++++++++++++----- SampCert/DifferentialPrivacy/Pure/System.lean | 4 +- .../Queries/Histogram/Properties.lean | 30 +++++++------- .../Queries/HistogramMean/Properties.lean | 8 ++-- .../ZeroConcentrated/System.lean | 4 +- 5 files changed, 56 insertions(+), 29 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index ee4813f6..39f1b7e1 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -44,9 +44,13 @@ class DPSystem (T : Type) where /-- Privacy adaptively 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, - ∀ ε₁ ε₂ : NNReal, - prop m₁ ε₁ → (∀ u, prop (m₂ u) ε₂) -> prop (privComposeAdaptive m₁ m₂) (ε₁ + ε₂) + 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, ∀ ε₁ ε₂ ε : NNReal, + prop m₁ ε₁ → (∀ u, prop (m₂ u) ε₂) -> + ε₁ + ε₂ = ε -> + prop (privComposeAdaptive m₁ m₂) ε /-- Privacy is invariant under post-processing. -/ @@ -56,7 +60,8 @@ class DPSystem (T : Type) where /-- Constant query is 0-DP -/ - const_prop : {U : Type} → [MeasurableSpace U] → [Countable U] → [DiscreteMeasurableSpace U] -> (u : U) -> prop (privConst u) (0 : NNReal) + const_prop : {U : Type} → [MeasurableSpace U] → [Countable U] → [DiscreteMeasurableSpace U] -> (u : U) -> + ∀ ε : NNReal, ε = (0 : NNReal) -> prop (privConst u) ε namespace DPSystem @@ -69,7 +74,8 @@ lemma compose_prop {U V : Type} [dps : DPSystem T] [MeasurableSpace U] [Countabl dps.prop m₁ ε₁ → dps.prop m₂ ε₂ → dps.prop (privCompose m₁ m₂) (ε₁ + ε₂) := by intros m₁ m₂ ε₁ ε₂ p1 p2 unfold privCompose - exact adaptive_compose_prop m₁ (fun _ => m₂) ε₁ ε₂ p1 fun _ => p2 + apply adaptive_compose_prop m₁ (fun _ => m₂) ε₁ ε₂ _ p1 (fun _ => p2) + rfl end DPSystem @@ -88,6 +94,8 @@ lemma bind_bind_indep (p : Mechanism T U) (q : Mechanism T V) (h : U → V → P simp [privCompose, privComposeAdaptive, tsum_prod'] + + /-- A noise function for a differential privacy system -/ @@ -98,12 +106,25 @@ class DPNoise (dps : DPSystem T) where -/ noise : Query T ℤ → (sensitivity : ℕ+) → (num : ℕ+) → (den : ℕ+) → Mechanism T ℤ /-- - How much privacy one obtains by adding Δ εn εd noise to a Δ-sensitive query + Relationship between arguments to noise and resulting privacy amount -/ - noise_priv : NNReal → NNReal + noise_priv : (num : ℕ+) → (den : ℕ+) → (priv : NNReal) -> Prop /-- - Adding noise to a query makes it private. + Adding noise to a query makes it private -/ - noise_prop : ∀ q : List T → ℤ, ∀ Δ εn εd : ℕ+, sensitivity q Δ → dps.prop (noise q Δ εn εd) (noise_priv (εn / εd)) + noise_prop : ∀ q : List T → ℤ, ∀ Δ εn εd : ℕ+, + sensitivity q Δ → + noise_priv εn εd ε -> + dps.prop (noise q Δ εn εd) ε + + +-- /- +-- A DPNoise instance for when the arguments to noise_prop can be computed dynamically +-- -/ +-- class DPNoiseDynamic {dps : DPSystem T} (dpn : DPNoise dps) where +-- compute_factor : (ε : NNReal) -> (ℕ+ × ℕ+) +-- compute_factor_spec : ∀ ε : NNReal, dpn.noise_priv (compute_factor ε).1 (compute_factor ε).2 ε + + end SLang diff --git a/SampCert/DifferentialPrivacy/Pure/System.lean b/SampCert/DifferentialPrivacy/Pure/System.lean index a2b26545..a04d6011 100644 --- a/SampCert/DifferentialPrivacy/Pure/System.lean +++ b/SampCert/DifferentialPrivacy/Pure/System.lean @@ -26,9 +26,9 @@ instance PureDPSystem : DPSystem T where of_adp := sorry prop_adp := sorry -- pure_ApproximateDP prop_mono := PureDP_mono - adaptive_compose_prop := PureDP_ComposeAdaptive' + adaptive_compose_prop := sorry -- PureDP_ComposeAdaptive' postprocess_prop := PureDP_PostProcess - const_prop := PureDP_privConst + const_prop := sorry -- PureDP_privConst instance laplace_pureDPSystem : DPNoise (@PureDPSystem T) where diff --git a/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean b/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean index 531cafd7..86fa124c 100644 --- a/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean @@ -76,19 +76,23 @@ lemma privNoisedHistogramAux_DP (ε₁ ε₂ : ℕ+) (n : ℕ) (Hn : n < numBins induction n · unfold privNoisedHistogramAux simp only [Nat.cast_zero, succ_eq_add_one, zero_add, Nat.cast_one, Nat.cast_mul, one_mul] - refine DPSystem.postprocess_prop - (privCompose (privNoisedBinCount numBins B ε₁ ε₂ 0) (privConst (emptyHistogram numBins B))) - (↑↑ε₁ / ↑↑(ε₂ * numBins)) ?G1 - apply (DPSystem_prop_ext _ ?HEq ?Hdp) - case Hdp => - apply (DPSystem.compose_prop - (privNoisedBinCount numBins B ε₁ ε₂ 0) - (privConst (emptyHistogram numBins B)) - (↑↑ε₁ / ↑↑(ε₂ * numBins)) - 0 - (privNoisedBinCount_DP numBins B ε₁ ε₂ 0) - (DPSystem.const_prop (emptyHistogram numBins B))) - case HEq => simp only [PNat.mul_coe, Nat.cast_mul, add_zero] + sorry + -- refine DPSystem.postprocess_prop + -- (privCompose (privNoisedBinCount numBins B ε₁ ε₂ 0) (privConst (emptyHistogram numBins B))) + -- (↑↑ε₁ / ↑↑(ε₂ * numBins)) ?G1 + -- apply (DPSystem_prop_ext _ ?HEq ?Hdp) + -- case Hdp => + -- sorry + -- -- apply (DPSystem.compose_prop + -- -- (privNoisedBinCount numBins B ε₁ ε₂ 0) + -- -- (privConst (emptyHistogram numBins B)) + -- -- (↑↑ε₁ / ↑↑(ε₂ * numBins)) + -- -- 0 + -- -- (privNoisedBinCount_DP numBins B ε₁ ε₂ 0) + -- -- (DPSystem.const_prop (emptyHistogram numBins B))) + -- case HEq => + -- sorry + -- -- simp only [PNat.mul_coe, Nat.cast_mul, add_zero] · rename_i n IH unfold privNoisedHistogramAux simp only [] diff --git a/SampCert/DifferentialPrivacy/Queries/HistogramMean/Properties.lean b/SampCert/DifferentialPrivacy/Queries/HistogramMean/Properties.lean index 05573c9e..db36ce3a 100644 --- a/SampCert/DifferentialPrivacy/Queries/HistogramMean/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/HistogramMean/Properties.lean @@ -69,11 +69,13 @@ lemma privMeanHistogram_DP (ε₁ ε₂ : ℕ+) (τ : ℤ) (ε₃ ε₄ : ℕ+) intro u cases u · simp only - apply dps.prop_mono ?G1 ?G2 - case G2 => apply dps.const_prop - simp only [_root_.zero_le] + sorry + -- apply dps.prop_mono ?G1 ?G2 + -- case G2 => apply dps.const_prop + -- simp only [_root_.zero_le] · rename_i mx simp only apply dps.postprocess_prop apply privNoisedBoundedMean_DP + rfl end SLang diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean index 22b26b58..a298c9a9 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean @@ -30,9 +30,9 @@ instance zCDPSystem : DPSystem T where prop_mono := zCDP_mono -- noise := privNoisedQuery -- noise_prop := sorry -- privNoisedQuery_zCDP - adaptive_compose_prop := privComposeAdaptive_zCDP + adaptive_compose_prop := sorry -- privComposeAdaptive_zCDP postprocess_prop := privPostProcess_zCDP - const_prop := privConst_zCDP + const_prop := sorry -- privConst_zCDP /-- Gaussian noise for zCDP system From 40ac9ccea637319bdb7e8f38b5ff2e74b056a268 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 12:49:50 -0400 Subject: [PATCH 009/100] instance for pure DP laplace noise --- SampCert/DifferentialPrivacy/Abstract.lean | 2 +- .../DifferentialPrivacy/Pure/AdaptiveComposition.lean | 6 ++++-- SampCert/DifferentialPrivacy/Pure/Const.lean | 2 +- .../Pure/Mechanism/Properties.lean | 7 +++++-- SampCert/DifferentialPrivacy/Pure/System.lean | 11 ++++------- 5 files changed, 15 insertions(+), 13 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index 39f1b7e1..4af77c2c 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -112,7 +112,7 @@ class DPNoise (dps : DPSystem T) where /-- Adding noise to a query makes it private -/ - noise_prop : ∀ q : List T → ℤ, ∀ Δ εn εd : ℕ+, + noise_prop : ∀ q : List T → ℤ, ∀ Δ εn εd : ℕ+, ∀ ε : NNReal, sensitivity q Δ → noise_priv εn εd ε -> dps.prop (noise q Δ εn εd) ε diff --git a/SampCert/DifferentialPrivacy/Pure/AdaptiveComposition.lean b/SampCert/DifferentialPrivacy/Pure/AdaptiveComposition.lean index 1cc2ca32..5da3c698 100644 --- a/SampCert/DifferentialPrivacy/Pure/AdaptiveComposition.lean +++ b/SampCert/DifferentialPrivacy/Pure/AdaptiveComposition.lean @@ -18,8 +18,10 @@ variable [Hu : Nonempty U] namespace SLang -- Better proof for Pure DP adaptive composition -theorem PureDP_ComposeAdaptive' (nq1 : List T → PMF U) (nq2 : U -> List T → PMF V) (ε₁ ε₂ : NNReal) (h1 : PureDP nq1 ε₁) (h2 : ∀ u : U, PureDP (nq2 u) ε₂) : - PureDP (privComposeAdaptive nq1 nq2) (ε₁ + ε₂) := by +theorem PureDP_ComposeAdaptive' (nq1 : List T → PMF U) (nq2 : U -> List T → PMF V) (ε₁ ε₂ ε : NNReal) (h1 : PureDP nq1 ε₁) (h2 : ∀ u : U, PureDP (nq2 u) ε₂) (Hε : ε₁ + ε₂ = ε ) : + PureDP (privComposeAdaptive nq1 nq2) ε := by + rw [<- Hε] + clear ε Hε simp [PureDP] at * rw [event_eq_singleton] at * simp [DP_singleton] at * diff --git a/SampCert/DifferentialPrivacy/Pure/Const.lean b/SampCert/DifferentialPrivacy/Pure/Const.lean index be3152ba..44c549dc 100644 --- a/SampCert/DifferentialPrivacy/Pure/Const.lean +++ b/SampCert/DifferentialPrivacy/Pure/Const.lean @@ -32,7 +32,7 @@ theorem privConst_DP_Bound {u : U} : DP (privConst u : Mechanism T U) 0 := by /-- ``privComposeAdaptive`` satisfies zCDP -/ -theorem PureDP_privConst : ∀ (u : U), PureDP (privConst u : Mechanism T U) (0 : NNReal) := by +theorem PureDP_privConst : ∀ (u : U) (ε : NNReal), (ε = 0) -> PureDP (privConst u : Mechanism T U) ε := by simp [PureDP] at * apply privConst_DP_Bound diff --git a/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean index addf1b2f..917636c6 100644 --- a/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean @@ -114,12 +114,15 @@ theorem privNoisedQueryPure_DP_bound (query : List T → ℤ) (Δ ε₁ ε₂ : exact (add_lt_add_iff_right 1).mpr A · apply exp_pos +def laplace_pureDP_noise_priv (ε₁ ε₂ : ℕ+) (ε : NNReal) : Prop := (ε₁ : NNReal) / ε₂ = ε /-- 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 Δ ε₁ ε₂) ((ε₁ : NNReal) / ε₂) := by +theorem privNoisedQueryPure_DP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (ε : NNReal) (bounded_sensitivity : sensitivity query Δ) (HN : laplace_pureDP_noise_priv ε₁ ε₂ ε) : + PureDP (privNoisedQueryPure query Δ ε₁ ε₂) ε := by + unfold laplace_pureDP_noise_priv at HN + rw [<- HN] simp [PureDP] apply privNoisedQueryPure_DP_bound apply bounded_sensitivity diff --git a/SampCert/DifferentialPrivacy/Pure/System.lean b/SampCert/DifferentialPrivacy/Pure/System.lean index a04d6011..1b555e75 100644 --- a/SampCert/DifferentialPrivacy/Pure/System.lean +++ b/SampCert/DifferentialPrivacy/Pure/System.lean @@ -26,16 +26,13 @@ instance PureDPSystem : DPSystem T where of_adp := sorry prop_adp := sorry -- pure_ApproximateDP prop_mono := PureDP_mono - adaptive_compose_prop := sorry -- PureDP_ComposeAdaptive' + adaptive_compose_prop := PureDP_ComposeAdaptive' postprocess_prop := PureDP_PostProcess - const_prop := sorry -- PureDP_privConst - + const_prop := PureDP_privConst instance laplace_pureDPSystem : DPNoise (@PureDPSystem T) where noise := privNoisedQueryPure - noise_priv := sorry - noise_prop := sorry -- privNoisedQueryPure_DP - - + noise_priv := laplace_pureDP_noise_priv + noise_prop := privNoisedQueryPure_DP end SLang From 456514ae44382431fe7de3fe35d44e712e09b584 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 12:57:04 -0400 Subject: [PATCH 010/100] pure DP system complete --- SampCert/DifferentialPrivacy/Pure/DP.lean | 16 +++++++++------- SampCert/DifferentialPrivacy/Pure/System.lean | 4 ++-- 2 files changed, 11 insertions(+), 9 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Pure/DP.lean b/SampCert/DifferentialPrivacy/Pure/DP.lean index 40c33529..2a5497cf 100644 --- a/SampCert/DifferentialPrivacy/Pure/DP.lean +++ b/SampCert/DifferentialPrivacy/Pure/DP.lean @@ -111,19 +111,21 @@ theorem ApproximateDP_of_DP (m : Mechanism T U) (ε : ℝ) (h : DP m ε) : apply le_trans h simp +/-- +Pure privacy bound required to obtain (ε, δ)-approximate DP +-/ +def pure_of_adp (_ : NNReal) (ε : NNReal) : NNReal := ε + /-- Pure DP is no weaker than approximate DP, up to a loss of parameters -/ lemma pure_ApproximateDP [Countable U] {m : Mechanism T U} : - ∃ (degrade : (δ : NNReal) -> (ε' : NNReal) -> NNReal), ∀ (δ : NNReal) (_ : 0 < δ) (ε' : NNReal), - (DP m (degrade δ ε') -> ApproximateDP m ε' δ) := by - let degrade (_ : NNReal) (ε' : NNReal) : NNReal := ε' - exists degrade - intro δ _ ε' HDP + ∀ (δ : NNReal) (_ : 0 < δ) (ε' : NNReal), + DP m (pure_of_adp δ ε') -> ApproximateDP m ε' δ := by + intro _ _ _ HDP rw [ApproximateDP] apply ApproximateDP_of_DP - have R1 : degrade δ ε' = ε' := by simp - rw [R1] at HDP + simp [pure_of_adp] at HDP trivial end SLang diff --git a/SampCert/DifferentialPrivacy/Pure/System.lean b/SampCert/DifferentialPrivacy/Pure/System.lean index 1b555e75..034a5502 100644 --- a/SampCert/DifferentialPrivacy/Pure/System.lean +++ b/SampCert/DifferentialPrivacy/Pure/System.lean @@ -23,8 +23,8 @@ Pure ε-DP with noise drawn from the discrete Laplace distribution. -/ instance PureDPSystem : DPSystem T where prop := PureDP - of_adp := sorry - prop_adp := sorry -- pure_ApproximateDP + of_adp := pure_of_adp + prop_adp := pure_ApproximateDP prop_mono := PureDP_mono adaptive_compose_prop := PureDP_ComposeAdaptive' postprocess_prop := PureDP_PostProcess From dfc4b3f766c1c5f9e59d97819cdb42b4e07f9728 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 13:03:24 -0400 Subject: [PATCH 011/100] count query complete --- SampCert/DifferentialPrivacy/Abstract.lean | 2 +- .../Pure/Mechanism/Properties.lean | 2 +- .../Queries/BoundedMean/Properties.lean | 3 ++- .../Queries/Count/Properties.lean | 13 +++++-------- 4 files changed, 9 insertions(+), 11 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index 4af77c2c..66c05719 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -113,8 +113,8 @@ class DPNoise (dps : DPSystem T) where Adding noise to a query makes it private -/ noise_prop : ∀ q : List T → ℤ, ∀ Δ εn εd : ℕ+, ∀ ε : NNReal, - sensitivity q Δ → noise_priv εn εd ε -> + sensitivity q Δ → dps.prop (noise q Δ εn εd) ε diff --git a/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean b/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean index 917636c6..950208a4 100644 --- a/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean +++ b/SampCert/DifferentialPrivacy/Pure/Mechanism/Properties.lean @@ -119,7 +119,7 @@ def laplace_pureDP_noise_priv (ε₁ ε₂ : ℕ+) (ε : NNReal) : Prop := (ε /-- Laplace noising mechanism ``privNoisedQueryPure`` produces a pure ``ε₁/ε₂``-DP mechanism from a Δ-sensitive query. -/ -theorem privNoisedQueryPure_DP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (ε : NNReal) (bounded_sensitivity : sensitivity query Δ) (HN : laplace_pureDP_noise_priv ε₁ ε₂ ε) : +theorem privNoisedQueryPure_DP (query : List T → ℤ) (Δ ε₁ ε₂ : ℕ+) (ε : NNReal) (HN : laplace_pureDP_noise_priv ε₁ ε₂ ε) (bounded_sensitivity : sensitivity query Δ) : PureDP (privNoisedQueryPure query Δ ε₁ ε₂) ε := by unfold laplace_pureDP_noise_priv at HN rw [<- HN] diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean index e157ce3c..e9ad149d 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean @@ -39,6 +39,7 @@ theorem privNoisedBoundedMean_DP (U : ℕ+) (ε₁ ε₂ : ℕ+) : rw [budget_split] apply dps.compose_prop · apply privNoisedBoundedSum_DP - · apply privNoisedCount_DP + · sorry + -- apply privNoisedCount_DP end SLang diff --git a/SampCert/DifferentialPrivacy/Queries/Count/Properties.lean b/SampCert/DifferentialPrivacy/Queries/Count/Properties.lean index 3dfc5377..8459e239 100644 --- a/SampCert/DifferentialPrivacy/Queries/Count/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/Count/Properties.lean @@ -45,13 +45,10 @@ theorem exactCount_1_sensitive : The noised counting query satisfies DP property -/ @[simp] -theorem privNoisedCount_DP (ε₁ ε₂ : ℕ+) : - dps.prop (privNoisedCount ε₁ ε₂) ((ε₁ : NNReal) / ε₂) := by - apply DPSystem_prop_ext - case H => - sorry - all_goals sorry - -- apply dpn.noise_prop - -- apply exactCount_1_sensitive +theorem privNoisedCount_DP (ε₁ ε₂ : ℕ+) (ε : NNReal) (HP : dpn.noise_priv ε₁ ε₂ ε) : + dps.prop (privNoisedCount ε₁ ε₂) ε := by + unfold privNoisedCount + apply (dpn.noise_prop _ _ _ _ _ HP) + apply exactCount_1_sensitive end SLang From dc7ee36819e21c6b12853647a544f4fb43c44bbe Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 13:46:23 -0400 Subject: [PATCH 012/100] tweak definitions of abstract DP for better inference --- SampCert/DifferentialPrivacy/Abstract.lean | 35 ++++++++-------- SampCert/DifferentialPrivacy/Pure/System.lean | 12 ++++-- .../Queries/BoundedMean/Properties.lean | 15 +++---- .../Queries/Count/Properties.lean | 2 +- .../Queries/Histogram/Properties.lean | 41 ++++++++++--------- .../ZeroConcentrated/System.lean | 4 +- 6 files changed, 58 insertions(+), 51 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index 66c05719..7bef65ae 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -40,29 +40,29 @@ class DPSystem (T : Type) where /-- DP is monotonic -/ - prop_mono {m : Mechanism T Z} {ε₁ ε₂: NNReal} (Hε : ε₁ ≤ ε₂) (H : prop m ε₁) : prop m ε₂ + prop_mono {m : Mechanism T Z} {ε₁ ε₂: NNReal} : + ε₁ ≤ ε₂ -> prop m ε₁ -> prop m ε₂ /-- Privacy adaptively 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, ∀ ε₁ ε₂ ε : NNReal, + 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} {ε₁ ε₂ ε : NNReal} : prop m₁ ε₁ → (∀ u, prop (m₂ u) ε₂) -> ε₁ + ε₂ = ε -> prop (privComposeAdaptive m₁ m₂) ε /-- Privacy is invariant under post-processing. -/ - postprocess_prop : {U : Type} → [MeasurableSpace U] → [Countable U] → [DiscreteMeasurableSpace U] → [Inhabited U] → { pp : U → V } → - ∀ m : Mechanism T U, ∀ ε : NNReal, - prop m ε → prop (privPostProcess m pp) ε + postprocess_prop {U : Type} [MeasurableSpace U] [Countable U] [DiscreteMeasurableSpace U] [Inhabited U] + { pp : U → V } {m : Mechanism T U} {ε : NNReal} : + prop m ε → prop (privPostProcess m pp) ε /-- Constant query is 0-DP -/ - const_prop : {U : Type} → [MeasurableSpace U] → [Countable U] → [DiscreteMeasurableSpace U] -> (u : U) -> - ∀ ε : NNReal, ε = (0 : NNReal) -> prop (privConst u) ε - + const_prop {U : Type} [MeasurableSpace U] [Countable U] [DiscreteMeasurableSpace U] {u : U} {ε : NNReal} : + ε = (0 : NNReal) -> prop (privConst u) ε namespace DPSystem @@ -70,12 +70,13 @@ namespace DPSystem Non-adaptive privacy composes by addition. -/ lemma compose_prop {U V : Type} [dps : DPSystem T] [MeasurableSpace U] [Countable U] [DiscreteMeasurableSpace U] [Inhabited U] [MeasurableSpace V] [Countable V] [DiscreteMeasurableSpace V] [Inhabited V] : - ∀ m₁ : Mechanism T U, ∀ m₂ : Mechanism T V, ∀ ε₁ ε₂ : NNReal, - dps.prop m₁ ε₁ → dps.prop m₂ ε₂ → dps.prop (privCompose m₁ m₂) (ε₁ + ε₂) := by - intros m₁ m₂ ε₁ ε₂ p1 p2 + {m₁ : Mechanism T U} -> {m₂ : Mechanism T V} -> {ε₁ ε₂ ε : NNReal} -> + (ε₁ + ε₂ = ε) -> + dps.prop m₁ ε₁ → dps.prop m₂ ε₂ → dps.prop (privCompose m₁ m₂) ε := by + intros _ _ _ _ _ _ p1 p2 unfold privCompose - apply adaptive_compose_prop m₁ (fun _ => m₂) ε₁ ε₂ _ p1 (fun _ => p2) - rfl + apply adaptive_compose_prop p1 (fun _ => p2) + trivial end DPSystem @@ -112,7 +113,7 @@ class DPNoise (dps : DPSystem T) where /-- Adding noise to a query makes it private -/ - noise_prop : ∀ q : List T → ℤ, ∀ Δ εn εd : ℕ+, ∀ ε : NNReal, + noise_prop {q : List T → ℤ} {Δ εn εd : ℕ+} {ε : NNReal} : noise_priv εn εd ε -> sensitivity q Δ → dps.prop (noise q Δ εn εd) ε diff --git a/SampCert/DifferentialPrivacy/Pure/System.lean b/SampCert/DifferentialPrivacy/Pure/System.lean index 034a5502..a8667cf6 100644 --- a/SampCert/DifferentialPrivacy/Pure/System.lean +++ b/SampCert/DifferentialPrivacy/Pure/System.lean @@ -26,13 +26,17 @@ instance PureDPSystem : DPSystem T where of_adp := pure_of_adp prop_adp := pure_ApproximateDP prop_mono := PureDP_mono - adaptive_compose_prop := PureDP_ComposeAdaptive' - postprocess_prop := PureDP_PostProcess - const_prop := PureDP_privConst + adaptive_compose_prop := by + intros; apply PureDP_ComposeAdaptive' <;> trivial + postprocess_prop := by + intros; apply PureDP_PostProcess; trivial + const_prop := by + intros; apply PureDP_privConst; trivial instance laplace_pureDPSystem : DPNoise (@PureDPSystem T) where noise := privNoisedQueryPure noise_priv := laplace_pureDP_noise_priv - noise_prop := privNoisedQueryPure_DP + noise_prop := by + intros; apply privNoisedQueryPure_DP <;> trivial end SLang diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean index e9ad149d..fa2441c3 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean @@ -34,12 +34,13 @@ DP bound for noised mean. theorem privNoisedBoundedMean_DP (U : ℕ+) (ε₁ ε₂ : ℕ+) : dps.prop (privNoisedBoundedMean U ε₁ ε₂) ((ε₁ : NNReal) / ε₂) := by unfold privNoisedBoundedMean - rw [bind_bind_indep] - apply dps.postprocess_prop - rw [budget_split] - apply dps.compose_prop - · apply privNoisedBoundedSum_DP - · sorry - -- apply privNoisedCount_DP + sorry + -- rw [bind_bind_indep] + -- apply dps.postprocess_prop + -- rw [budget_split] + -- apply dps.compose_prop + -- · apply privNoisedBoundedSum_DP + -- · sorry + -- -- apply privNoisedCount_DP end SLang diff --git a/SampCert/DifferentialPrivacy/Queries/Count/Properties.lean b/SampCert/DifferentialPrivacy/Queries/Count/Properties.lean index 8459e239..c961280f 100644 --- a/SampCert/DifferentialPrivacy/Queries/Count/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/Count/Properties.lean @@ -48,7 +48,7 @@ The noised counting query satisfies DP property theorem privNoisedCount_DP (ε₁ ε₂ : ℕ+) (ε : NNReal) (HP : dpn.noise_priv ε₁ ε₂ ε) : dps.prop (privNoisedCount ε₁ ε₂) ε := by unfold privNoisedCount - apply (dpn.noise_prop _ _ _ _ _ HP) + apply dpn.noise_prop HP apply exactCount_1_sensitive end SLang diff --git a/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean b/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean index 86fa124c..655c2230 100644 --- a/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean @@ -96,26 +96,27 @@ lemma privNoisedHistogramAux_DP (ε₁ ε₂ : ℕ+) (n : ℕ) (Hn : n < numBins · rename_i n IH unfold privNoisedHistogramAux simp only [] - refine DPSystem.postprocess_prop - (privCompose (privNoisedBinCount numBins B ε₁ ε₂ ↑(n + 1)) - (privNoisedHistogramAux numBins B ε₁ ε₂ n (privNoisedHistogramAux.proof_2 numBins n Hn))) - (↑(n + 1).succ * (↑↑ε₁ / ↑↑(ε₂ * numBins))) ?succ.a - apply (@DPSystem_prop_ext _ _ _ (?C1 + ?C2) _ _ ?HCeq ?Hdp) - case Hdp => - refine - (DPSystem.compose_prop - (privNoisedBinCount numBins B ε₁ ε₂ ↑(n + 1)) - (privNoisedHistogramAux numBins B ε₁ ε₂ n _) (↑↑ε₁ / ↑↑(ε₂ * numBins)) (↑n.succ * (↑↑ε₁ / ↑↑(ε₂ * numBins))) ?X ?Y) - case X => exact privNoisedBinCount_DP numBins B ε₁ ε₂ ↑(n + 1) - case Y => apply IH - generalize (ε₁.val.cast / (ε₂ * numBins).val.cast : NNReal) = A - conv => - enter [1, 1] - rw [Eq.symm (one_mul A)] - rw [<- add_mul] - congr - simp only [succ_eq_add_one, Nat.cast_add, Nat.cast_one] - exact AddCommMagma.add_comm (OfNat.ofNat 1) (n.cast + OfNat.ofNat 1) + sorry + -- refine DPSystem.postprocess_prop + -- (privCompose (privNoisedBinCount numBins B ε₁ ε₂ ↑(n + 1)) + -- (privNoisedHistogramAux numBins B ε₁ ε₂ n (privNoisedHistogramAux.proof_2 numBins n Hn))) + -- (↑(n + 1).succ * (↑↑ε₁ / ↑↑(ε₂ * numBins))) ?succ.a + -- apply (@DPSystem_prop_ext _ _ _ (?C1 + ?C2) _ _ ?HCeq ?Hdp) + -- case Hdp => + -- refine + -- (DPSystem.compose_prop + -- (privNoisedBinCount numBins B ε₁ ε₂ ↑(n + 1)) + -- (privNoisedHistogramAux numBins B ε₁ ε₂ n _) (↑↑ε₁ / ↑↑(ε₂ * numBins)) (↑n.succ * (↑↑ε₁ / ↑↑(ε₂ * numBins))) ?X ?Y) + -- case X => exact privNoisedBinCount_DP numBins B ε₁ ε₂ ↑(n + 1) + -- case Y => apply IH + -- generalize (ε₁.val.cast / (ε₂ * numBins).val.cast : NNReal) = A + -- conv => + -- enter [1, 1] + -- rw [Eq.symm (one_mul A)] + -- rw [<- add_mul] + -- congr + -- simp only [succ_eq_add_one, Nat.cast_add, Nat.cast_one] + -- exact AddCommMagma.add_comm (OfNat.ofNat 1) (n.cast + OfNat.ofNat 1) /-- DP bound for a noised histogram diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean index a298c9a9..840aa94b 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean @@ -27,11 +27,11 @@ instance zCDPSystem : DPSystem T where prop := zCDP of_adp := sorry prop_adp := sorry -- zCDP_ApproximateDP - prop_mono := zCDP_mono + prop_mono := sorry -- zCDP_mono -- noise := privNoisedQuery -- noise_prop := sorry -- privNoisedQuery_zCDP adaptive_compose_prop := sorry -- privComposeAdaptive_zCDP - postprocess_prop := privPostProcess_zCDP + postprocess_prop := sorry -- privPostProcess_zCDP const_prop := sorry -- privConst_zCDP /-- From 00e2ef644be4bc30f53040584051da0d114e29fc Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 13:50:33 -0400 Subject: [PATCH 013/100] cleanup bounded sum abstract proof --- .../Queries/BoundedSum/Properties.lean | 9 ++++----- 1 file changed, 4 insertions(+), 5 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedSum/Properties.lean b/SampCert/DifferentialPrivacy/Queries/BoundedSum/Properties.lean index 4f21348b..c4778f53 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedSum/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedSum/Properties.lean @@ -85,10 +85,9 @@ theorem exactBoundedSum_sensitivity (U : ℕ+) : sensitivity (exactBoundedSum U) The noised bounded sum satisfies the DP property of the DP system. -/ @[simp] -theorem privNoisedBoundedSum_DP (U : ℕ+) (ε₁ ε₂ : ℕ+) : - dps.prop (privNoisedBoundedSum U ε₁ ε₂) ((ε₁ : NNReal) / ε₂) := by - sorry - -- apply dpn.noise_prop - -- apply exactBoundedSum_sensitivity +theorem privNoisedBoundedSum_DP (U : ℕ+) (ε₁ ε₂ : ℕ+) (HP : dpn.noise_priv ε₁ ε₂ (ε₁ / ε₂)) : + dps.prop (privNoisedBoundedSum U ε₁ ε₂) ((ε₁ : NNReal) / ε₂) := by + apply dpn.noise_prop HP + apply exactBoundedSum_sensitivity end SLang From be8cfd4b975807c532f9274045e4cf5be145e323 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 15:19:54 -0400 Subject: [PATCH 014/100] update bounded mean --- .../Queries/BoundedMean/Properties.lean | 25 ++++++++++++------- 1 file changed, 16 insertions(+), 9 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean index fa2441c3..367c088a 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean @@ -31,16 +31,23 @@ lemma budget_split (ε₁ ε₂ : ℕ+) : /-- DP bound for noised mean. -/ -theorem privNoisedBoundedMean_DP (U : ℕ+) (ε₁ ε₂ : ℕ+) : +theorem privNoisedBoundedMean_DP (U : ℕ+) (ε₁ ε₂ : ℕ+) + (HP_half : dpn.noise_priv ε₁ (2 * ε₂) (ε₁ / ↑(2 * ε₂))) : dps.prop (privNoisedBoundedMean U ε₁ ε₂) ((ε₁ : NNReal) / ε₂) := by unfold privNoisedBoundedMean - sorry - -- rw [bind_bind_indep] - -- apply dps.postprocess_prop - -- rw [budget_split] - -- apply dps.compose_prop - -- · apply privNoisedBoundedSum_DP - -- · sorry - -- -- apply privNoisedCount_DP + rw [bind_bind_indep] + apply dps.postprocess_prop + apply dps.compose_prop ?SC1 + · apply privNoisedBoundedSum_DP + apply HP_half + · apply privNoisedCount_DP + apply HP_half + + case SC1 => + -- Arithmetic + ring_nf + rw [div_mul] + congr + simp end SLang From b064656cf8955555f0cf2a86928ecf56faa9974d Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 15:40:56 -0400 Subject: [PATCH 015/100] Histogram update --- .../Queries/Histogram/Properties.lean | 116 +++++------------- .../Queries/HistogramMean/Properties.lean | 9 +- 2 files changed, 39 insertions(+), 86 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean b/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean index 655c2230..50e6580f 100644 --- a/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean @@ -27,9 +27,6 @@ variable [HT : Inhabited T] variable (numBins : ℕ+) variable (B : Bins T numBins) --- def exactBinCount (b : Fin num_bins) (l : List T) : ℤ := --- List.length (List.filter (fun v => B.bin v = b) l) - /- exactBinCount is 1-sensitive -/ @@ -59,108 +56,63 @@ theorem exactBinCount_sensitivity (b : Fin numBins) : sensitivity (exactBinCount /-- DP bound for a noised bin count -/ -lemma privNoisedBinCount_DP (ε₁ ε₂ : ℕ+) (b : Fin numBins) : - dps.prop (privNoisedBinCount numBins B ε₁ ε₂ b) (ε₁ / ((ε₂ * numBins : PNat))) := by +lemma privNoisedBinCount_DP (ε₁ ε₂ : ℕ+) (ε : NNReal) (b : Fin numBins) + (HN_bin : dpn.noise_priv ε₁ (ε₂ * numBins) (ε / numBins)) : + dps.prop (privNoisedBinCount numBins B ε₁ ε₂ b) (ε / numBins) := by unfold privNoisedBinCount - sorry - -- apply dpn.noise_prop - -- apply exactBinCount_sensitivity + apply dpn.noise_prop HN_bin + apply exactBinCount_sensitivity --- MARKUSDE: Looking at this proof it is clear that we need better tactic support for the abstract DP operators --- MARKUSDE: - Lemmas with equality side conditions for the privacy cost /-- DP bound for intermediate steps in the histogram calculation. -/ -lemma privNoisedHistogramAux_DP (ε₁ ε₂ : ℕ+) (n : ℕ) (Hn : n < numBins) : - dps.prop (privNoisedHistogramAux numBins B ε₁ ε₂ n Hn) (n.succ * (ε₁ / (ε₂ * numBins : PNat))) := by +lemma privNoisedHistogramAux_DP (ε₁ ε₂ : ℕ+) (ε : NNReal) (n : ℕ) (Hn : n < numBins) + (HN_bin : dpn.noise_priv ε₁ (ε₂ * numBins) (ε / numBins)) : + dps.prop (privNoisedHistogramAux numBins B ε₁ ε₂ n Hn) (n.succ * (ε / numBins)) := by induction n · unfold privNoisedHistogramAux simp only [Nat.cast_zero, succ_eq_add_one, zero_add, Nat.cast_one, Nat.cast_mul, one_mul] - sorry - -- refine DPSystem.postprocess_prop - -- (privCompose (privNoisedBinCount numBins B ε₁ ε₂ 0) (privConst (emptyHistogram numBins B))) - -- (↑↑ε₁ / ↑↑(ε₂ * numBins)) ?G1 - -- apply (DPSystem_prop_ext _ ?HEq ?Hdp) - -- case Hdp => - -- sorry - -- -- apply (DPSystem.compose_prop - -- -- (privNoisedBinCount numBins B ε₁ ε₂ 0) - -- -- (privConst (emptyHistogram numBins B)) - -- -- (↑↑ε₁ / ↑↑(ε₂ * numBins)) - -- -- 0 - -- -- (privNoisedBinCount_DP numBins B ε₁ ε₂ 0) - -- -- (DPSystem.const_prop (emptyHistogram numBins B))) - -- case HEq => - -- sorry - -- -- simp only [PNat.mul_coe, Nat.cast_mul, add_zero] + apply DPSystem.postprocess_prop + apply DPSystem.compose_prop ?SC1 + · apply privNoisedBinCount_DP + apply HN_bin + · apply DPSystem.const_prop + rfl + case SC1 => simp only [add_zero] · rename_i n IH unfold privNoisedHistogramAux simp only [] - sorry - -- refine DPSystem.postprocess_prop - -- (privCompose (privNoisedBinCount numBins B ε₁ ε₂ ↑(n + 1)) - -- (privNoisedHistogramAux numBins B ε₁ ε₂ n (privNoisedHistogramAux.proof_2 numBins n Hn))) - -- (↑(n + 1).succ * (↑↑ε₁ / ↑↑(ε₂ * numBins))) ?succ.a - -- apply (@DPSystem_prop_ext _ _ _ (?C1 + ?C2) _ _ ?HCeq ?Hdp) - -- case Hdp => - -- refine - -- (DPSystem.compose_prop - -- (privNoisedBinCount numBins B ε₁ ε₂ ↑(n + 1)) - -- (privNoisedHistogramAux numBins B ε₁ ε₂ n _) (↑↑ε₁ / ↑↑(ε₂ * numBins)) (↑n.succ * (↑↑ε₁ / ↑↑(ε₂ * numBins))) ?X ?Y) - -- case X => exact privNoisedBinCount_DP numBins B ε₁ ε₂ ↑(n + 1) - -- case Y => apply IH - -- generalize (ε₁.val.cast / (ε₂ * numBins).val.cast : NNReal) = A - -- conv => - -- enter [1, 1] - -- rw [Eq.symm (one_mul A)] - -- rw [<- add_mul] - -- congr - -- simp only [succ_eq_add_one, Nat.cast_add, Nat.cast_one] - -- exact AddCommMagma.add_comm (OfNat.ofNat 1) (n.cast + OfNat.ofNat 1) + apply DPSystem.postprocess_prop + apply DPSystem.compose_prop ?SC1 + · apply privNoisedBinCount_DP + apply HN_bin + · apply IH + case SC1 => + simp + ring_nf /-- DP bound for a noised histogram -/ -lemma privNoisedHistogram_DP (ε₁ ε₂ : ℕ+) : - dps.prop (privNoisedHistogram numBins B ε₁ ε₂) (ε₁ / ε₂) := by +lemma privNoisedHistogram_DP (ε₁ ε₂ : ℕ+) (ε : NNReal) (HN_bin : dpn.noise_priv ε₁ (ε₂ * numBins) (ε / numBins)) : + dps.prop (privNoisedHistogram numBins B ε₁ ε₂) ε := by unfold privNoisedHistogram apply (DPSystem_prop_ext _ ?HEq ?Hdp) - case Hdp => apply privNoisedHistogramAux_DP + case Hdp => + apply privNoisedHistogramAux_DP + apply HN_bin case HEq => - simp - rw [division_def] - rw [division_def] - rw [mul_inv] - conv => - enter [1, 2] - rw [<- mul_assoc] - rw [mul_comm] - generalize (ε₁.val.cast * ε₂.val.cast⁻¹ : NNReal) = A - rw [<- mul_assoc] - conv => - enter [2] - rw [Eq.symm (one_mul A)] - congr - unfold predBins - cases numBins - rename_i n' Hn' - simp only [PNat.natPred_eq_pred, pred_eq_sub_one, cast_tsub, Nat.cast_one, PNat.mk_coe] - rw [tsub_add_eq_max] - rw [max_eq_left (one_le_cast.mpr Hn')] - apply mul_inv_cancel - intro HK - simp_all - rw [HK] at Hn' - cases Hn' - + simp [predBins] + simp [mul_div_left_comm] /-- DP bound for the thresholding maximum -/ -lemma privMaxBinAboveThreshold_DP (ε₁ ε₂ : ℕ+) (τ : ℤ) : - dps.prop (privMaxBinAboveThreshold numBins B ε₁ ε₂ τ) (ε₁ / ε₂) := by - rw [privMaxBinAboveThreshold] +lemma privMaxBinAboveThreshold_DP (ε₁ ε₂ : ℕ+) (ε : NNReal) (τ : ℤ) (HN_bin : dpn.noise_priv ε₁ (ε₂ * numBins) (ε / numBins)) : + dps.prop (privMaxBinAboveThreshold numBins B ε₁ ε₂ τ) ε := by + unfold privMaxBinAboveThreshold apply dps.postprocess_prop apply privNoisedHistogram_DP + apply HN_bin end SLang diff --git a/SampCert/DifferentialPrivacy/Queries/HistogramMean/Properties.lean b/SampCert/DifferentialPrivacy/Queries/HistogramMean/Properties.lean index db36ce3a..344400e0 100644 --- a/SampCert/DifferentialPrivacy/Queries/HistogramMean/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/HistogramMean/Properties.lean @@ -73,9 +73,10 @@ lemma privMeanHistogram_DP (ε₁ ε₂ : ℕ+) (τ : ℤ) (ε₃ ε₄ : ℕ+) -- apply dps.prop_mono ?G1 ?G2 -- case G2 => apply dps.const_prop -- simp only [_root_.zero_le] - · rename_i mx - simp only - apply dps.postprocess_prop - apply privNoisedBoundedMean_DP + · sorry + -- rename_i mx + -- simp only + -- apply dps.postprocess_prop + -- apply privNoisedBoundedMean_DP rfl end SLang From edec41666fc1e8dd6cf9aaf5f5e2d6517427d4ee Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 15:52:00 -0400 Subject: [PATCH 016/100] historgram mean update --- .../Queries/BoundedMean/Properties.lean | 11 +++---- .../Queries/BoundedSum/Properties.lean | 4 +-- .../Queries/HistogramMean/Properties.lean | 32 ++++++++++--------- 3 files changed, 23 insertions(+), 24 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean index 367c088a..abed30ab 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Properties.lean @@ -31,9 +31,9 @@ lemma budget_split (ε₁ ε₂ : ℕ+) : /-- DP bound for noised mean. -/ -theorem privNoisedBoundedMean_DP (U : ℕ+) (ε₁ ε₂ : ℕ+) - (HP_half : dpn.noise_priv ε₁ (2 * ε₂) (ε₁ / ↑(2 * ε₂))) : - dps.prop (privNoisedBoundedMean U ε₁ ε₂) ((ε₁ : NNReal) / ε₂) := by +theorem privNoisedBoundedMean_DP (U : ℕ+) (ε₁ ε₂ : ℕ+) (ε : NNReal) + (HP_half : dpn.noise_priv ε₁ (2 * ε₂) (ε / 2)) : + dps.prop (privNoisedBoundedMean U ε₁ ε₂) ε := by unfold privNoisedBoundedMean rw [bind_bind_indep] apply dps.postprocess_prop @@ -44,10 +44,7 @@ theorem privNoisedBoundedMean_DP (U : ℕ+) (ε₁ ε₂ : ℕ+) apply HP_half case SC1 => - -- Arithmetic ring_nf - rw [div_mul] - congr - simp + simp [div_mul] end SLang diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedSum/Properties.lean b/SampCert/DifferentialPrivacy/Queries/BoundedSum/Properties.lean index c4778f53..3e14c7a4 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedSum/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedSum/Properties.lean @@ -85,8 +85,8 @@ theorem exactBoundedSum_sensitivity (U : ℕ+) : sensitivity (exactBoundedSum U) The noised bounded sum satisfies the DP property of the DP system. -/ @[simp] -theorem privNoisedBoundedSum_DP (U : ℕ+) (ε₁ ε₂ : ℕ+) (HP : dpn.noise_priv ε₁ ε₂ (ε₁ / ε₂)) : - dps.prop (privNoisedBoundedSum U ε₁ ε₂) ((ε₁ : NNReal) / ε₂) := by +theorem privNoisedBoundedSum_DP (U : ℕ+) (ε₁ ε₂ : ℕ+) (ε : NNReal) (HP : dpn.noise_priv ε₁ ε₂ ε) : + dps.prop (privNoisedBoundedSum U ε₁ ε₂) ε := by apply dpn.noise_prop HP apply exactBoundedSum_sensitivity diff --git a/SampCert/DifferentialPrivacy/Queries/HistogramMean/Properties.lean b/SampCert/DifferentialPrivacy/Queries/HistogramMean/Properties.lean index 344400e0..6d15c8fd 100644 --- a/SampCert/DifferentialPrivacy/Queries/HistogramMean/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/HistogramMean/Properties.lean @@ -57,26 +57,28 @@ instance : DiscreteMeasurableSpace (Option (Fin ↑numBins)) where forall_measurableSet := by simp only [MeasurableSpace.measurableSet_top, implies_true] -/- +/-- DP bound for the adaptive mean -/ -lemma privMeanHistogram_DP (ε₁ ε₂ : ℕ+) (τ : ℤ) (ε₃ ε₄ : ℕ+) : - dps.prop (privMeanHistogram dps dpn numBins B unbin ε₁ ε₂ τ ε₃ ε₄) (ε₁/ε₂ + ε₃/ε₄) := by +lemma privMeanHistogram_DP (ε₁ ε₂ : ℕ+) (τ : ℤ) (ε₃ ε₄ : ℕ+) (εA εB : NNReal) + (HN_A : dpn.noise_priv ε₁ (ε₂ * numBins) (εA / numBins)) + (HN_B : dpn.noise_priv ε₃ (2 * ε₄) (εB / 2)): + dps.prop (privMeanHistogram dps dpn numBins B unbin ε₁ ε₂ τ ε₃ ε₄) (εA + εB) := by rw [privMeanHistogram] apply dps.postprocess_prop apply dps.adaptive_compose_prop · apply privMaxBinAboveThreshold_DP - intro u - cases u - · simp only - sorry - -- apply dps.prop_mono ?G1 ?G2 - -- case G2 => apply dps.const_prop - -- simp only [_root_.zero_le] - · sorry - -- rename_i mx - -- simp only - -- apply dps.postprocess_prop - -- apply privNoisedBoundedMean_DP + apply HN_A + · intro u + cases u + · simp only [] + apply (@DPSystem.prop_mono _ _ _ _ 0 εB _) + · apply dps.const_prop + rfl + · apply _root_.zero_le + · simp only [] + apply dps.postprocess_prop + apply privNoisedBoundedMean_DP + apply HN_B rfl end SLang From 592cdea5f116ee7b8fb3a12c7db6bc8283caa010 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 16:11:17 -0400 Subject: [PATCH 017/100] instances for zCDP (except approx DP) --- SampCert.lean | 2 +- .../ZeroConcentrated/System.lean | 52 +++++++++++++------ 2 files changed, 37 insertions(+), 17 deletions(-) diff --git a/SampCert.lean b/SampCert.lean index 1d7fc16e..ebafb641 100644 --- a/SampCert.lean +++ b/SampCert.lean @@ -14,7 +14,7 @@ import Init.Data.UInt.Lemmas open SLang PMF -def combineConcentrated := @privNoisedBoundedMean_DP _ gaussian_zCDPSystem +def combineConcentrated := @privNoisedBoundedMean_DP zCDPSystem def combinePure := @privNoisedBoundedMean_DP PureDPSystem /- diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean index 840aa94b..89c50e06 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean @@ -27,28 +27,48 @@ instance zCDPSystem : DPSystem T where prop := zCDP of_adp := sorry prop_adp := sorry -- zCDP_ApproximateDP - prop_mono := sorry -- zCDP_mono - -- noise := privNoisedQuery - -- noise_prop := sorry -- privNoisedQuery_zCDP - adaptive_compose_prop := sorry -- privComposeAdaptive_zCDP - postprocess_prop := sorry -- privPostProcess_zCDP - const_prop := sorry -- privConst_zCDP + prop_mono := by + intro _ _ _ _ H HZ + apply zCDP_mono H HZ + adaptive_compose_prop := by + intros _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ HZ HZ2 Hε + rw [<- Hε] + apply privComposeAdaptive_zCDP + apply HZ + apply HZ2 + postprocess_prop := by + intros _ _ _ _ _ _ _ _ _ HZ + apply privPostProcess_zCDP + apply HZ + const_prop := by + intros _ _ _ _ _ _ Hε + rw [Hε] + apply privConst_zCDP + +def gaussian_zCDP_noise_priv (ε₁ ε₂ : ℕ+) (ε : NNReal) : Prop := (1/2) * ((ε₁ : NNReal) / ε₂)^2 ≤ ε /-- Gaussian noise for zCDP system -/ -instance gaussian_zCDPSystem : DPNoise zCDPsystem where +instance gaussian_zCDPSystem (T : Type) : DPNoise (@zCDPSystem T) where noise := privNoisedQuery - noise_priv := sorry - noise_prop := sorry -- privNoisedQuery_zCDP + noise_priv := gaussian_zCDP_noise_priv + noise_prop := by + intros _ _ _ _ _ H HS + apply DPSystem.prop_mono ?G1 + · simp [DPSystem.prop] + apply privNoisedQuery_zCDP + apply HS + · apply H -/-- -Laplace noise for zCDP system --/ -instance laplace_zCDPSystem : DPNoise zCDPsystem where - noise := sorry - noise_priv := sorry - noise_prop := sorry -- privNoisedQuery_zCDP + +-- /- +-- Laplace noise for zCDP system? +-- -/ +-- instance laplace_zCDPSystem (T : Type) : DPNoise (@zCDPSystem T) where +-- noise := sorry +-- noise_priv := sorry +-- noise_prop := sorry -- privNoisedQuery_zCDP end SLang From e746f4695f5ad83a3ec1167f0d05d36f7b9e4ea5 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 18:57:10 -0400 Subject: [PATCH 018/100] checkpoint --- .../ZeroConcentrated/DP.lean | 85 +++++++++++++++---- .../ZeroConcentrated/System.lean | 7 +- 2 files changed, 73 insertions(+), 19 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean index 1bbb753d..a30d2b49 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean @@ -665,31 +665,82 @@ theorem ApproximateDP_of_zCDP [Countable U] (m : Mechanism T U) case G2 => exact Hm l₁ l₂ HN simp + +def D (δ : NNReal) : Real := (2 * Real.log (1 / δ)) ^ ((1 : ℝ) / (2 : ℝ)) + +def ε (ε' : NNReal) (δ : NNReal) : NNReal := + Real.toNNReal ((-2 * D δ + .sqrt ((2 * D δ) ^ 2 + 8 * ε')) / 2) ^ ((1 : ℝ) / (2 : ℝ)) + +lemma eqε (ε' δ : NNReal) (H : δ < 1) : ε' = ((ε ε' δ)^2) / (2 : NNReal) + ε ε' δ * D δ := by + sorry + +/-- +Pure privacy bound required to obtain (ε, δ)-approximate DP +-/ +def zCDP_of_adp (δ : NNReal) (ε' : NNReal) : NNReal := (1/2) * ((ε ε' δ)^2) + -- (√(2 * Real.log (1/δ) + 2 * ε) - √(2 * Real.log (1/δ))).toNNReal + /-- zCDP is no weaker than approximate DP, up to a loss of parameters. -/ lemma zCDP_ApproximateDP [Countable U] {m : Mechanism T U} : - ∃ (degrade : (δ : NNReal) -> (ε' : NNReal) -> NNReal), ∀ (δ : NNReal) (_ : 0 < δ) (ε' : NNReal), - (zCDP m (degrade δ ((1/2) * ε'^2)) -> ApproximateDP m ε' δ) := by - let degrade (δ : NNReal) (ε' : NNReal) : NNReal := - (√(2 * Real.log (1/δ) + 2 * ε') - √(2 * Real.log (1/δ))).toNNReal - have HDdegrade δ ε' : degrade δ ε' = (√(2 * Real.log (1/δ) + 2 * ε') - √(2 * Real.log (1/δ))).toNNReal := by rfl - exists degrade - intro δ Hδ ε' ⟨ HN , HB ⟩ - + ∀ (δ : NNReal) (_ : 0 < δ) (ε' : NNReal), + (zCDP m (zCDP_of_adp δ ε') -> ApproximateDP m ε' δ) := by + intro δ Hδ ε' H cases Classical.em (1 ≤ δ) · rename_i Hδ1 exact ApproximateDP_gt1 m (↑ε') Hδ1 - rename_i Hδ1 - rw [ApproximateDP] - - - have R := ApproximateDP_of_zCDP m (degrade δ ε') (by simp) ?G1 HN δ Hδ - case G1 => - -- this proof has to be redone - sorry - sorry + simp at Hδ1 + unfold ApproximateDP + + unfold zCDP_of_adp at H + rcases H with ⟨ H1, H2 ⟩ + have X := ApproximateDP_of_zCDP m (ε ε' δ) ?G1 H2 H1 δ Hδ + case G1 => exact NNReal.zero_le_coe + rw [eqε ε' δ Hδ1] + -- We have this, basically. + unfold D + simp_all + + + + -- #check ApproximateDP_of_zCDP + -- have X := ApproximateDP_of_zCDP m ((2 * zCDP_of_adp δ ε') ^ ((1 : Real) / 2)) ?G1 ?G2 ?G3 δ Hδ + -- case G1 => + -- apply Real.rpow_nonneg + -- apply mul_nonneg + -- · simp + -- · exact NNReal.zero_le_coe + -- case G2 => + -- ring_nf + -- -- True + -- sorry + -- case G3 => + -- cases H + -- trivial + + -- have X1 : (2 * (zCDP_of_adp δ ε')) ^ (1 / 2)) ^ 2 / 2 = (zCDP_of_adp δ ε') : Real := by sorry + -- rw [X1] + -- clear X1 + + -- let degrade (δ : NNReal) (ε' : NNReal) : NNReal := + -- (√(2 * Real.log (1/δ) + 2 * ε') - √(2 * Real.log (1/δ))).toNNReal + -- have HDdegrade δ ε' : degrade δ ε' = (√(2 * Real.log (1/δ) + 2 * ε') - √(2 * Real.log (1/δ))).toNNReal := by rfl + -- exists degrade + -- intro δ Hδ ε' ⟨ HN , HB ⟩ + + + -- rename_i Hδ1 + -- rw [ApproximateDP] + + + + -- have R := ApproximateDP_of_zCDP m (degrade δ ε') (by simp) ?G1 HN δ Hδ + -- case G1 => + -- -- this proof has to be redone + -- sorry + -- sorry /- have Hdegrade : ((degrade δ ε') ^ 2) / 2 + (degrade δ ε') * (2 * Real.log (1 / δ))^(1/2 : ℝ) = ε' := by diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean index 89c50e06..fc794fad 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean @@ -16,6 +16,8 @@ import SampCert.DifferentialPrivacy.ZeroConcentrated.Const This file contains an instance of an abstract DP system associated to the discrete gaussian mechanisms. -/ +noncomputable section + namespace SLang variable { T : Type } @@ -25,8 +27,9 @@ Instance of a DP system for zCDP -/ instance zCDPSystem : DPSystem T where prop := zCDP - of_adp := sorry - prop_adp := sorry -- zCDP_ApproximateDP + of_adp := zCDP_of_adp + prop_adp := by + intros; apply zCDP_ApproximateDP <;> assumption prop_mono := by intro _ _ _ _ H HZ apply zCDP_mono H HZ From 4ea177b494ec805c93059e29cbf5dfe7b7a2ae27 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Tue, 29 Oct 2024 11:41:55 -0400 Subject: [PATCH 019/100] zCDP degrade --- .../ZeroConcentrated/DP.lean | 144 ++++++------------ 1 file changed, 48 insertions(+), 96 deletions(-) diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean index a30d2b49..5edf10ea 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/DP.lean @@ -665,20 +665,58 @@ theorem ApproximateDP_of_zCDP [Countable U] (m : Mechanism T U) case G2 => exact Hm l₁ l₂ HN simp +namespace zCDP_of_adp_def def D (δ : NNReal) : Real := (2 * Real.log (1 / δ)) ^ ((1 : ℝ) / (2 : ℝ)) -def ε (ε' : NNReal) (δ : NNReal) : NNReal := - Real.toNNReal ((-2 * D δ + .sqrt ((2 * D δ) ^ 2 + 8 * ε')) / 2) ^ ((1 : ℝ) / (2 : ℝ)) +def ε (ε' : NNReal) (δ : NNReal) : NNReal := Real.toNNReal (-D δ + discrim (OfNat.ofNat 2)⁻¹ (D δ) (-↑ε') ^ (2 : ℝ)⁻¹) -lemma eqε (ε' δ : NNReal) (H : δ < 1) : ε' = ((ε ε' δ)^2) / (2 : NNReal) + ε ε' δ * D δ := by - sorry +lemma eqε (ε' δ : NNReal) (H0 : 0 < δ) (H1 : δ < 1) : ε' = ((ε ε' δ)^2) / (2 : NNReal) + ε ε' δ * D δ := by + suffices (((1 : ℝ) / (2 : ℝ)) * (ε ε' δ) * (ε ε' δ) + D δ * ε ε' δ + (- ε')) = 0 by + simp_all + rw [add_neg_eq_zero] at this + rw [<- this] + ring_nf + + have Hle : 0 < D δ ^ 2 + 4 * (OfNat.ofNat 2)⁻¹ * ε' := by + apply Right.add_pos_of_pos_of_nonneg + · apply sq_pos_of_pos + unfold D + apply Real.rpow_pos_of_pos + apply mul_pos <;> simp + exact Real.log_neg H0 H1 + · simp + + apply (@quadratic_eq_zero_iff Real _ _ _ _ _ ?Ga ((discrim (OfNat.ofNat 1 / OfNat.ofNat 2) (D δ) (-ε'.toReal)) ^ ((1 : ℝ) / (2 : ℝ))) ?Gs ((ε ε' δ).toReal)).mpr + case Ga => simp + case Gs => + rw [<- Real.rpow_add ?Gadd] + case Gadd => + simp [discrim] + apply Hle + rw [add_halves] + simp + left + simp + unfold ε + apply Real.coe_toNNReal + simp [discrim] + apply nonneg_le_nonneg_of_sq_le_sq + · apply Real.rpow_nonneg + exact le_of_lt Hle + rw [<- Real.rpow_add Hle] + rw [<- one_div] + rw [add_halves] + simp + rw [<- sq] + simp + +end zCDP_of_adp_def /-- Pure privacy bound required to obtain (ε, δ)-approximate DP -/ -def zCDP_of_adp (δ : NNReal) (ε' : NNReal) : NNReal := (1/2) * ((ε ε' δ)^2) - -- (√(2 * Real.log (1/δ) + 2 * ε) - √(2 * Real.log (1/δ))).toNNReal +def zCDP_of_adp (δ : NNReal) (ε' : NNReal) : NNReal := (1/2) * ((zCDP_of_adp_def.ε ε' δ)^2) /-- zCDP is no weaker than approximate DP, up to a loss of parameters. @@ -686,7 +724,7 @@ zCDP is no weaker than approximate DP, up to a loss of parameters. lemma zCDP_ApproximateDP [Countable U] {m : Mechanism T U} : ∀ (δ : NNReal) (_ : 0 < δ) (ε' : NNReal), (zCDP m (zCDP_of_adp δ ε') -> ApproximateDP m ε' δ) := by - intro δ Hδ ε' H + intro δ Hδ0 ε' H cases Classical.em (1 ≤ δ) · rename_i Hδ1 exact ApproximateDP_gt1 m (↑ε') Hδ1 @@ -696,100 +734,14 @@ lemma zCDP_ApproximateDP [Countable U] {m : Mechanism T U} : unfold zCDP_of_adp at H rcases H with ⟨ H1, H2 ⟩ - have X := ApproximateDP_of_zCDP m (ε ε' δ) ?G1 H2 H1 δ Hδ + have X := ApproximateDP_of_zCDP m (zCDP_of_adp_def.ε ε' δ) ?G1 H2 H1 δ Hδ0 case G1 => exact NNReal.zero_le_coe - rw [eqε ε' δ Hδ1] - -- We have this, basically. - unfold D + rw [zCDP_of_adp_def.eqε ε' δ Hδ0 Hδ1] + unfold zCDP_of_adp_def.D simp_all - -- #check ApproximateDP_of_zCDP - -- have X := ApproximateDP_of_zCDP m ((2 * zCDP_of_adp δ ε') ^ ((1 : Real) / 2)) ?G1 ?G2 ?G3 δ Hδ - -- case G1 => - -- apply Real.rpow_nonneg - -- apply mul_nonneg - -- · simp - -- · exact NNReal.zero_le_coe - -- case G2 => - -- ring_nf - -- -- True - -- sorry - -- case G3 => - -- cases H - -- trivial - - -- have X1 : (2 * (zCDP_of_adp δ ε')) ^ (1 / 2)) ^ 2 / 2 = (zCDP_of_adp δ ε') : Real := by sorry - -- rw [X1] - -- clear X1 - - -- let degrade (δ : NNReal) (ε' : NNReal) : NNReal := - -- (√(2 * Real.log (1/δ) + 2 * ε') - √(2 * Real.log (1/δ))).toNNReal - -- have HDdegrade δ ε' : degrade δ ε' = (√(2 * Real.log (1/δ) + 2 * ε') - √(2 * Real.log (1/δ))).toNNReal := by rfl - -- exists degrade - -- intro δ Hδ ε' ⟨ HN , HB ⟩ - - - -- rename_i Hδ1 - -- rw [ApproximateDP] - - - - -- have R := ApproximateDP_of_zCDP m (degrade δ ε') (by simp) ?G1 HN δ Hδ - -- case G1 => - -- -- this proof has to be redone - -- sorry - -- sorry - - /- - have Hdegrade : ((degrade δ ε') ^ 2) / 2 + (degrade δ ε') * (2 * Real.log (1 / δ))^(1/2 : ℝ) = ε' := by - rw [HDdegrade] - generalize HD : Real.log (1 / δ) = D - have HDnn : 0 ≤ D := by - rw [<- HD] - apply Real.log_nonneg - apply one_le_one_div Hδ - exact le_of_not_ge Hδ1 - simp only [Real.coe_toNNReal'] - rw [max_eq_left ?G1] - case G1 => - apply sub_nonneg_of_le - apply Real.sqrt_le_sqrt - simp - rw [sub_sq'] - rw [Real.sq_sqrt ?G1] - case G1 => - apply add_nonneg - · simp - trivial - · simp - rw [Real.sq_sqrt ?G1] - case G1 => - simp - trivial - rw [← Real.sqrt_eq_rpow] - rw [mul_sub_right_distrib] - rw [<- sq] - rw [Real.sq_sqrt ?G1] - case G1 => - simp - trivial - generalize HW : √(2 * D + 2 * ↑ε') * √(2 * D) = W - conv => - enter [1, 1, 1, 2] - rw [mul_assoc] - rw [HW] - rw [sub_div] - rw [add_div] - rw [add_div] - simp - linarith - rw [Hdegrade] at R - trivial - -/ - - /-- Pure DP bound implies absolute continuity -/ From de517ba80ee59401cfbe99c9b1fb4e1537c220c1 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Wed, 30 Oct 2024 10:07:24 -0400 Subject: [PATCH 020/100] simplify abstract DP presentation further with typeclass synonym --- SampCert/DifferentialPrivacy/Abstract.lean | 31 +++++++++++++++---- .../ZeroConcentrated/System.lean | 6 ++-- 2 files changed, 28 insertions(+), 9 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index 7bef65ae..93480da3 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -19,6 +19,27 @@ open Classical Nat Int Real ENNReal namespace SLang +/-- +Typeclass synonym for the classes we use for probaility +-/ +class DiscProbSpace (T : Type) where + instMeasurableSpace : MeasurableSpace T + instCountable : Countable T + instDiscreteMeasurableSpace : DiscreteMeasurableSpace T + instInhabited : Inhabited T + +-- Typeclass inference to- and from- the synonym +instance [idps : DiscProbSpace T] : MeasurableSpace T := idps.instMeasurableSpace +instance [idps : DiscProbSpace T] : Countable T := idps.instCountable +instance [idps : DiscProbSpace T] : DiscreteMeasurableSpace T := idps.instDiscreteMeasurableSpace +instance [idps : DiscProbSpace T] : Inhabited T := idps.instInhabited + +instance [im : MeasurableSpace T] [ic : Countable T] [idm : DiscreteMeasurableSpace T] [ii : Inhabited T] : DiscProbSpace T where + instMeasurableSpace := im + instCountable := ic + instDiscreteMeasurableSpace := idm + instInhabited := ii + /-- Abstract definition of a differentially private systemm. -/ @@ -45,9 +66,7 @@ class DPSystem (T : Type) where /-- Privacy adaptively 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] + adaptive_compose_prop {U V : Type} [DiscProbSpace U] [DiscProbSpace V] {m₁ : Mechanism T U} {m₂ : U -> Mechanism T V} {ε₁ ε₂ ε : NNReal} : prop m₁ ε₁ → (∀ u, prop (m₂ u) ε₂) -> ε₁ + ε₂ = ε -> @@ -55,13 +74,13 @@ class DPSystem (T : Type) where /-- Privacy is invariant under post-processing. -/ - postprocess_prop {U : Type} [MeasurableSpace U] [Countable U] [DiscreteMeasurableSpace U] [Inhabited U] + postprocess_prop {U : Type} [DiscProbSpace U] { pp : U → V } {m : Mechanism T U} {ε : NNReal} : prop m ε → prop (privPostProcess m pp) ε /-- Constant query is 0-DP -/ - const_prop {U : Type} [MeasurableSpace U] [Countable U] [DiscreteMeasurableSpace U] {u : U} {ε : NNReal} : + const_prop {U : Type} [DiscProbSpace U] {u : U} {ε : NNReal} : ε = (0 : NNReal) -> prop (privConst u) ε namespace DPSystem @@ -69,7 +88,7 @@ namespace DPSystem /- Non-adaptive privacy composes by addition. -/ -lemma compose_prop {U V : Type} [dps : DPSystem T] [MeasurableSpace U] [Countable U] [DiscreteMeasurableSpace U] [Inhabited U] [MeasurableSpace V] [Countable V] [DiscreteMeasurableSpace V] [Inhabited V] : +lemma compose_prop {U V : Type} [dps : DPSystem T] [DiscProbSpace U] [DiscProbSpace V] : {m₁ : Mechanism T U} -> {m₂ : Mechanism T V} -> {ε₁ ε₂ ε : NNReal} -> (ε₁ + ε₂ = ε) -> dps.prop m₁ ε₁ → dps.prop m₂ ε₂ → dps.prop (privCompose m₁ m₂) ε := by diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean index fc794fad..e9501786 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean @@ -34,17 +34,17 @@ instance zCDPSystem : DPSystem T where intro _ _ _ _ H HZ apply zCDP_mono H HZ adaptive_compose_prop := by - intros _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ HZ HZ2 Hε + intros _ _ _ _ _ _ _ _ _ HZ HZ2 Hε rw [<- Hε] apply privComposeAdaptive_zCDP apply HZ apply HZ2 postprocess_prop := by - intros _ _ _ _ _ _ _ _ _ HZ + intros _ _ _ _ _ _ HZ apply privPostProcess_zCDP apply HZ const_prop := by - intros _ _ _ _ _ _ Hε + intros _ _ _ _ Hε rw [Hε] apply privConst_zCDP From 22416518a296a18c32ecd52480c456947b0970b8 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Sat, 2 Nov 2024 13:13:10 -0400 Subject: [PATCH 021/100] shorten histogram proofs --- .../Queries/Histogram/Properties.lean | 61 ++++++------------- 1 file changed, 17 insertions(+), 44 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean b/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean index 50e6580f..5cbee3f0 100644 --- a/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean @@ -32,26 +32,10 @@ exactBinCount is 1-sensitive -/ theorem exactBinCount_sensitivity (b : Fin numBins) : sensitivity (exactBinCount numBins B b) 1 := by rw [sensitivity] - intros l₁ l₂ HN - cases HN - · rename_i l₁' l₂' v' Hl₁ Hl₂ - rw [ Hl₁, Hl₂ ] - rw [exactBinCount, exactBinCount] - simp [List.filter_cons] - aesop - · rename_i l₁' v' l₂' Hl₁ Hl₂ - rw [ Hl₁, Hl₂ ] - rw [exactBinCount, exactBinCount] - simp [List.filter_cons] - aesop - · rename_i l₁' v₁' l₂' v₂' Hl₁ Hl₂ - rw [ Hl₁, Hl₂ ] - rw [exactBinCount, exactBinCount] - simp [List.filter_cons] - -- There has to be a better way - cases (Classical.em (B.bin v₁' = b)) <;> cases (Classical.em (B.bin v₂' = b)) - all_goals (rename_i Hrw1 Hrw2) - all_goals (simp [Hrw1, Hrw2]) + intros _ _ H + cases H + all_goals simp_all [exactBinCount, exactBinCount, List.filter_cons] + all_goals aesop /-- DP bound for a noised bin count @@ -71,25 +55,18 @@ lemma privNoisedHistogramAux_DP (ε₁ ε₂ : ℕ+) (ε : NNReal) (n : ℕ) (Hn dps.prop (privNoisedHistogramAux numBins B ε₁ ε₂ n Hn) (n.succ * (ε / numBins)) := by induction n · unfold privNoisedHistogramAux - simp only [Nat.cast_zero, succ_eq_add_one, zero_add, Nat.cast_one, Nat.cast_mul, one_mul] - apply DPSystem.postprocess_prop - apply DPSystem.compose_prop ?SC1 - · apply privNoisedBinCount_DP - apply HN_bin - · apply DPSystem.const_prop - rfl - case SC1 => simp only [add_zero] - · rename_i n IH - unfold privNoisedHistogramAux - simp only [] - apply DPSystem.postprocess_prop - apply DPSystem.compose_prop ?SC1 - · apply privNoisedBinCount_DP - apply HN_bin + simp + apply dps.postprocess_prop + apply dps.compose_prop (AddLeftCancelMonoid.add_zero _) + · apply privNoisedBinCount_DP; apply HN_bin + · apply dps.const_prop; rfl + · rename_i _ IH + simp [privNoisedHistogramAux] + apply dps.postprocess_prop + apply dps.compose_prop ?arithmetic + · apply privNoisedBinCount_DP; apply HN_bin · apply IH - case SC1 => - simp - ring_nf + case arithmetic => simp; ring_nf /-- DP bound for a noised histogram @@ -98,12 +75,8 @@ lemma privNoisedHistogram_DP (ε₁ ε₂ : ℕ+) (ε : NNReal) (HN_bin : dpn.no dps.prop (privNoisedHistogram numBins B ε₁ ε₂) ε := by unfold privNoisedHistogram apply (DPSystem_prop_ext _ ?HEq ?Hdp) - case Hdp => - apply privNoisedHistogramAux_DP - apply HN_bin - case HEq => - simp [predBins] - simp [mul_div_left_comm] + case Hdp => apply privNoisedHistogramAux_DP; apply HN_bin + case HEq => simp [predBins, mul_div_left_comm] /-- DP bound for the thresholding maximum From dc787efec776d52f67e42025817867da428acd89 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Sat, 2 Nov 2024 13:58:03 -0400 Subject: [PATCH 022/100] histogram: move epsilons to variable --- .../Queries/Histogram/Properties.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean b/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean index 5cbee3f0..9eeffd35 100644 --- a/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/Histogram/Properties.lean @@ -27,6 +27,8 @@ variable [HT : Inhabited T] variable (numBins : ℕ+) variable (B : Bins T numBins) +variable (ε₁ ε₂ : ℕ+) (ε : NNReal) (HN_bin : dpn.noise_priv ε₁ (ε₂ * numBins) (ε / numBins)) + /- exactBinCount is 1-sensitive -/ @@ -40,8 +42,7 @@ theorem exactBinCount_sensitivity (b : Fin numBins) : sensitivity (exactBinCount /-- DP bound for a noised bin count -/ -lemma privNoisedBinCount_DP (ε₁ ε₂ : ℕ+) (ε : NNReal) (b : Fin numBins) - (HN_bin : dpn.noise_priv ε₁ (ε₂ * numBins) (ε / numBins)) : +lemma privNoisedBinCount_DP (b : Fin numBins) : dps.prop (privNoisedBinCount numBins B ε₁ ε₂ b) (ε / numBins) := by unfold privNoisedBinCount apply dpn.noise_prop HN_bin @@ -50,8 +51,7 @@ lemma privNoisedBinCount_DP (ε₁ ε₂ : ℕ+) (ε : NNReal) (b : Fin numBins) /-- DP bound for intermediate steps in the histogram calculation. -/ -lemma privNoisedHistogramAux_DP (ε₁ ε₂ : ℕ+) (ε : NNReal) (n : ℕ) (Hn : n < numBins) - (HN_bin : dpn.noise_priv ε₁ (ε₂ * numBins) (ε / numBins)) : +lemma privNoisedHistogramAux_DP (n : ℕ) (Hn : n < numBins) : dps.prop (privNoisedHistogramAux numBins B ε₁ ε₂ n Hn) (n.succ * (ε / numBins)) := by induction n · unfold privNoisedHistogramAux @@ -71,7 +71,7 @@ lemma privNoisedHistogramAux_DP (ε₁ ε₂ : ℕ+) (ε : NNReal) (n : ℕ) (Hn /-- DP bound for a noised histogram -/ -lemma privNoisedHistogram_DP (ε₁ ε₂ : ℕ+) (ε : NNReal) (HN_bin : dpn.noise_priv ε₁ (ε₂ * numBins) (ε / numBins)) : +lemma privNoisedHistogram_DP : dps.prop (privNoisedHistogram numBins B ε₁ ε₂) ε := by unfold privNoisedHistogram apply (DPSystem_prop_ext _ ?HEq ?Hdp) @@ -81,7 +81,7 @@ lemma privNoisedHistogram_DP (ε₁ ε₂ : ℕ+) (ε : NNReal) (HN_bin : dpn.no /-- DP bound for the thresholding maximum -/ -lemma privMaxBinAboveThreshold_DP (ε₁ ε₂ : ℕ+) (ε : NNReal) (τ : ℤ) (HN_bin : dpn.noise_priv ε₁ (ε₂ * numBins) (ε / numBins)) : +lemma privMaxBinAboveThreshold_DP (τ : ℤ) : dps.prop (privMaxBinAboveThreshold numBins B ε₁ ε₂ τ) ε := by unfold privMaxBinAboveThreshold apply dps.postprocess_prop From 1b8cf409b0717173783c804ff228a87ebbe808c1 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Tue, 5 Nov 2024 10:18:35 -0500 Subject: [PATCH 023/100] define SPMF --- SampCert/DifferentialPrivacy/Generic.lean | 55 +++++++++++++++++++++-- 1 file changed, 52 insertions(+), 3 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Generic.lean b/SampCert/DifferentialPrivacy/Generic.lean index be8e333c..f4a64acd 100644 --- a/SampCert/DifferentialPrivacy/Generic.lean +++ b/SampCert/DifferentialPrivacy/Generic.lean @@ -13,16 +13,63 @@ import SampCert.Foundations.Basic This file defines an abstract system of differentially private operators. -/ -noncomputable section - -open Classical Nat Int Real ENNReal namespace SLang +/- +A SPMF is a SLang program that is a also proper distribution +-/ +abbrev SPMF.{u} (α : Type u) : Type u := { f : SLang α // HasSum f 1 } + +instance : Coe (SPMF α) (SLang α) where + coe a := a.1 + +instance : Coe (SPMF α) (PMF α) where + coe a := ⟨ a.1, a.2 ⟩ + + + +def SPMF_pure (a : α) : SPMF α := + ⟨ probPure a, + by + have H : PMF.pure a = probPure a := by + unfold probPure + simp [DFunLike.coe, PMF.instFunLike, PMF.pure] + rw [<- H] + cases (PMF.pure a) + simp [DFunLike.coe, PMF.instFunLike] + trivial ⟩ + +def SPMF_bind (p : SPMF α) (q : α -> SPMF β) : SPMF β := + ⟨ probBind p (fun x => q x), + by + have H : (probBind p (fun x => q x)) = (PMF.bind p q) := by + unfold probBind + simp [DFunLike.coe, PMF.instFunLike, PMF.bind] + rw [H] + cases (PMF.bind p q) + simp [DFunLike.coe, PMF.instFunLike] + trivial ⟩ + +instance : Monad SPMF where + pure := SPMF_pure + bind := SPMF_bind + abbrev Query (T U : Type) := List T → U abbrev Mechanism (T U : Type) := List T → PMF U + + + + + + + +noncomputable section + +open Classical Nat Int Real ENNReal + /-- General (value-dependent) composition of mechanisms -/ @@ -218,4 +265,6 @@ lemma condition_to_subset (f : U → V) (g : U → ENNReal) (x : V) : simp rw [B] +end + end SLang From e4a2f4658baed613fbd50aeb8fe397946aef8fb8 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Tue, 5 Nov 2024 10:24:05 -0500 Subject: [PATCH 024/100] SPMF: generic --- SampCert/DifferentialPrivacy/Generic.lean | 59 +++++++++++++---------- 1 file changed, 33 insertions(+), 26 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Generic.lean b/SampCert/DifferentialPrivacy/Generic.lean index f4a64acd..6d82cbf3 100644 --- a/SampCert/DifferentialPrivacy/Generic.lean +++ b/SampCert/DifferentialPrivacy/Generic.lean @@ -29,6 +29,7 @@ instance : Coe (SPMF α) (PMF α) where +@[simp] def SPMF_pure (a : α) : SPMF α := ⟨ probPure a, by @@ -40,6 +41,7 @@ def SPMF_pure (a : α) : SPMF α := simp [DFunLike.coe, PMF.instFunLike] trivial ⟩ +@[simp] def SPMF_bind (p : SPMF α) (q : α -> SPMF β) : SPMF β := ⟨ probBind p (fun x => q x), by @@ -55,13 +57,37 @@ instance : Monad SPMF where pure := SPMF_pure bind := SPMF_bind + + abbrev Query (T U : Type) := List T → U -abbrev Mechanism (T U : Type) := List T → PMF U +abbrev Mechanism (T U : Type) := List T → SPMF U +/-- +General (value-dependent) composition of mechanisms +-/ +def privComposeAdaptive (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : List T) : SPMF (U × V) := do + let A <- nq1 l + let B <- nq2 A l + return (A, B) +/-- +Product of mechanisms. +-/ +def privCompose (nq1 : Mechanism T U) (nq2 : Mechanism T V) (l : List T) : SPMF (U × V) := + privComposeAdaptive nq1 (fun _ => nq2) l +/-- +Mechanism obtained by applying a post-processing function to a mechanism. +-/ +def privPostProcess (nq : Mechanism T U) (pp : U → V) (l : List T) : SPMF V := do + let A ← nq l + return pp A +/-- +Constant mechanism +-/ +def privConst (u : U) : Mechanism T U := fun _ => SPMF_pure u @@ -70,15 +96,12 @@ noncomputable section open Classical Nat Int Real ENNReal -/-- -General (value-dependent) composition of mechanisms --/ -def privComposeAdaptive (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : List T) : PMF (U × V) := do - let A <- nq1 l - let B <- nq2 A l - return (A, B) +instance SPMF.instFunLike : FunLike (SPMF α) α ℝ≥0∞ where + coe p a := p.1 a + coe_injective' _ _ h := Subtype.eq h -lemma compose_sum_rw_adaptive (nq1 : List T → PMF U) (nq2 : U -> List T → PMF V) (u : U) (v : V) (l : List T) : + +lemma compose_sum_rw_adaptive (nq1 : List T → SPMF U) (nq2 : U -> List T → SPMF 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 @@ -117,24 +140,8 @@ lemma privComposeChainRule (nq1 : Mechanism T U) (nq2 : U -> Mechanism T V) (l : intros u v rw [<- compose_sum_rw_adaptive] simp [privComposeAdaptive] + simp [SPMF.instFunLike] -/-- -Product of mechanisms. --/ -def privCompose (nq1 : Mechanism T U) (nq2 : Mechanism T V) (l : List T) : PMF (U × V) := - privComposeAdaptive nq1 (fun _ => nq2) l - -/-- -Mechanism obtained by applying a post-processing function to a mechanism. --/ -def privPostProcess (nq : Mechanism T U) (pp : U → V) (l : List T) : PMF V := do - let A ← nq l - return pp A - -/-- -Constant mechanism --/ -def privConst (u : U) : Mechanism T U := fun _ => PMF.pure u -- @[simp] From e023a0faebef6af7a3cf11db55941c4b7700ab01 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Tue, 5 Nov 2024 10:30:06 -0500 Subject: [PATCH 025/100] postprocessing --- SampCert/DifferentialPrivacy/Pure/Postprocessing.lean | 1 + .../DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean | 3 +-- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Pure/Postprocessing.lean b/SampCert/DifferentialPrivacy/Pure/Postprocessing.lean index 4e57529a..8621724a 100644 --- a/SampCert/DifferentialPrivacy/Pure/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/Pure/Postprocessing.lean @@ -29,6 +29,7 @@ lemma privPostProcess_DP_bound {nq : Mechanism T U} {ε : NNReal} (h : PureDP nq replace h := h l₁ l₂ neighbours simp [privPostProcess] apply ENNReal.div_le_of_le_mul + simp [SPMF.instFunLike] rw [← ENNReal.tsum_mul_left] apply tsum_le_tsum _ ENNReal.summable (by aesop) intro i diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean index 6ebaacb2..782dc5c6 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/Postprocessing.lean @@ -37,6 +37,7 @@ def privPostProcess_AC {f : U -> V} (nq : Mechanism T U) (Hac : ACNeighbour nq) intro l₁ l₂ Hn v have Hac := Hac l₁ l₂ Hn simp [privPostProcess] + simp [DFunLike.coe, SPMF.instFunLike] intro Hppz i fi apply Hac apply Hppz @@ -575,8 +576,6 @@ theorem privPostProcess_zCDPBound {nq : Mechanism T U} {ε : ℝ} apply inv_nonneg_of_nonneg linarith apply elog_mono_le.mp - simp [PMF.bind, PMF.pure] - simp [PMF.instFunLike] apply privPostPocess_DP_pre · exact fun l => PMF.hasSum_coe_one (nq l) · exact h1 From 1b861657a1c33d16067e834b29f2f36401fe2834 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Tue, 5 Nov 2024 10:34:08 -0500 Subject: [PATCH 026/100] update queries to be computable --- SampCert.lean | 2 +- SampCert/DifferentialPrivacy/Queries/BoundedMean/Code.lean | 4 +--- SampCert/DifferentialPrivacy/Queries/BoundedSum/Code.lean | 2 -- SampCert/DifferentialPrivacy/Queries/Count/Code.lean | 2 -- SampCert/DifferentialPrivacy/Queries/Histogram/Code.lean | 2 -- SampCert/DifferentialPrivacy/Queries/HistogramMean/Code.lean | 2 -- 6 files changed, 2 insertions(+), 12 deletions(-) diff --git a/SampCert.lean b/SampCert.lean index ebafb641..be2abe77 100644 --- a/SampCert.lean +++ b/SampCert.lean @@ -40,7 +40,7 @@ Return an upper bound on the bin value, clipped to 2^(1 + numBins) -/ def unbin (n : Fin numBins) : ℕ+ := 2 ^ (1 + n.val) -noncomputable def combineMeanHistogram : Mechanism ℕ (Option ℚ) := +def combineMeanHistogram : Mechanism ℕ (Option ℚ) := privMeanHistogram PureDPSystem laplace_pureDPSystem numBins { bin } unbin 1 20 2 1 20 end histogramMeanExample diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Code.lean b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Code.lean index 79234d71..dc94cd09 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedMean/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedMean/Code.lean @@ -13,8 +13,6 @@ import SampCert.DifferentialPrivacy.Queries.BoundedSum.Code This file defines a differentially private bounded mean query. -/ -noncomputable section - namespace SLang variable [dps : DPSystem ℕ] @@ -23,7 +21,7 @@ variable [dpn : DPNoise dps] /-- Compute a noised mean using a noised sum and noised count. -/ -def privNoisedBoundedMean (U : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : PMF ℚ := do +def privNoisedBoundedMean (U : ℕ+) (ε₁ ε₂ : ℕ+) (l : List ℕ) : SPMF ℚ := do let S ← privNoisedBoundedSum U ε₁ (2 * ε₂) l let C ← privNoisedCount ε₁ (2 * ε₂) l return S / C diff --git a/SampCert/DifferentialPrivacy/Queries/BoundedSum/Code.lean b/SampCert/DifferentialPrivacy/Queries/BoundedSum/Code.lean index ffdc6f98..1396256d 100644 --- a/SampCert/DifferentialPrivacy/Queries/BoundedSum/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/BoundedSum/Code.lean @@ -14,8 +14,6 @@ import Init.Data.Int.Order This file defines a differentially private bounded sum query. -/ -noncomputable section - namespace SLang variable [dps : DPSystem ℕ] diff --git a/SampCert/DifferentialPrivacy/Queries/Count/Code.lean b/SampCert/DifferentialPrivacy/Queries/Count/Code.lean index b38929ef..f9bf360f 100644 --- a/SampCert/DifferentialPrivacy/Queries/Count/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/Count/Code.lean @@ -11,8 +11,6 @@ import SampCert.DifferentialPrivacy.Abstract This file defines a differentially private noising of an exact length query. -/ -noncomputable section - namespace SLang variable {T : Type} diff --git a/SampCert/DifferentialPrivacy/Queries/Histogram/Code.lean b/SampCert/DifferentialPrivacy/Queries/Histogram/Code.lean index 9342a6c9..eb6f25aa 100644 --- a/SampCert/DifferentialPrivacy/Queries/Histogram/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/Histogram/Code.lean @@ -24,8 +24,6 @@ structure Bins (T : Type) (num_bins : ℕ) where ## Private Histograms -/ -noncomputable section - variable (numBins : ℕ+) def predBins : ℕ := numBins.natPred diff --git a/SampCert/DifferentialPrivacy/Queries/HistogramMean/Code.lean b/SampCert/DifferentialPrivacy/Queries/HistogramMean/Code.lean index 6bb35172..51546d49 100644 --- a/SampCert/DifferentialPrivacy/Queries/HistogramMean/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/HistogramMean/Code.lean @@ -20,8 +20,6 @@ Implementation for ## Compute the mean using a histogram to compute the noisy max -/ -noncomputable section - variable (dps : SLang.DPSystem ℕ) variable (dpn : SLang.DPNoise dps) variable (numBins : ℕ+) From b806c8bd9257cb361aac7ae679c25de4790eb531 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Tue, 5 Nov 2024 10:47:08 -0500 Subject: [PATCH 027/100] move SPMF to SLang --- SampCert/DifferentialPrivacy/Generic.lean | 42 ---------------------- SampCert/SLang.lean | 43 ++++++++++++++++++++++- 2 files changed, 42 insertions(+), 43 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Generic.lean b/SampCert/DifferentialPrivacy/Generic.lean index 6d82cbf3..8c46773e 100644 --- a/SampCert/DifferentialPrivacy/Generic.lean +++ b/SampCert/DifferentialPrivacy/Generic.lean @@ -16,48 +16,6 @@ This file defines an abstract system of differentially private operators. namespace SLang -/- -A SPMF is a SLang program that is a also proper distribution --/ -abbrev SPMF.{u} (α : Type u) : Type u := { f : SLang α // HasSum f 1 } - -instance : Coe (SPMF α) (SLang α) where - coe a := a.1 - -instance : Coe (SPMF α) (PMF α) where - coe a := ⟨ a.1, a.2 ⟩ - - - -@[simp] -def SPMF_pure (a : α) : SPMF α := - ⟨ probPure a, - by - have H : PMF.pure a = probPure a := by - unfold probPure - simp [DFunLike.coe, PMF.instFunLike, PMF.pure] - rw [<- H] - cases (PMF.pure a) - simp [DFunLike.coe, PMF.instFunLike] - trivial ⟩ - -@[simp] -def SPMF_bind (p : SPMF α) (q : α -> SPMF β) : SPMF β := - ⟨ probBind p (fun x => q x), - by - have H : (probBind p (fun x => q x)) = (PMF.bind p q) := by - unfold probBind - simp [DFunLike.coe, PMF.instFunLike, PMF.bind] - rw [H] - cases (PMF.bind p q) - simp [DFunLike.coe, PMF.instFunLike] - trivial ⟩ - -instance : Monad SPMF where - pure := SPMF_pure - bind := SPMF_bind - - abbrev Query (T U : Type) := List T → U diff --git a/SampCert/SLang.lean b/SampCert/SLang.lean index 5566b5fb..85be3cdb 100644 --- a/SampCert/SLang.lean +++ b/SampCert/SLang.lean @@ -144,11 +144,52 @@ def probUntil (body : SLang T) (cond : T → Bool) : SLang T := do let v ← body probWhile (λ v : T => ¬ cond v) (λ _ : T => body) v + +/- +A SPMF is a SLang program that is a also proper distribution +-/ +abbrev SPMF.{u} (α : Type u) : Type u := { f : SLang α // HasSum f 1 } + +instance : Coe (SPMF α) (SLang α) where + coe a := a.1 + +instance : Coe (SPMF α) (PMF α) where + coe a := ⟨ a.1, a.2 ⟩ + + +@[simp] +def SPMF_pure (a : α) : SPMF α := + ⟨ probPure a, + by + have H : PMF.pure a = probPure a := by + unfold probPure + simp [DFunLike.coe, PMF.instFunLike, PMF.pure] + rw [<- H] + cases (PMF.pure a) + simp [DFunLike.coe, PMF.instFunLike] + trivial ⟩ + +@[simp] +def SPMF_bind (p : SPMF α) (q : α -> SPMF β) : SPMF β := + ⟨ probBind p (fun x => q x), + by + have H : (probBind p (fun x => q x)) = (PMF.bind p q) := by + unfold probBind + simp [DFunLike.coe, PMF.instFunLike, PMF.bind] + rw [H] + cases (PMF.bind p q) + simp [DFunLike.coe, PMF.instFunLike] + trivial ⟩ + +instance : Monad SPMF where + pure := SPMF_pure + bind := SPMF_bind + end SLang namespace PMF @[extern "my_run"] -opaque run (c : PMF T) : IO T +opaque run (c : SLang.SPMF T) : IO T end PMF From f75fe5a62ec2384f942ecee7e2dc016ed4a30383 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Tue, 5 Nov 2024 11:56:50 -0500 Subject: [PATCH 028/100] query tests --- SampCert.lean | 4 +- SampCert/Samplers/Uniform/Properties.lean | 2 +- Test.lean | 62 ++++++++++++++++++++++- 3 files changed, 64 insertions(+), 4 deletions(-) diff --git a/SampCert.lean b/SampCert.lean index be2abe77..1e1285dc 100644 --- a/SampCert.lean +++ b/SampCert.lean @@ -27,7 +27,7 @@ def numBins : ℕ+ := 64 /- Bin the infinite type ℕ with clipping -/ -def bin (n : ℕ) : Fin numBins := +def example_bin (n : ℕ) : Fin numBins := { val := min (Nat.log 2 n) (PNat.natPred numBins), isLt := by apply min_lt_iff.mpr @@ -41,7 +41,7 @@ Return an upper bound on the bin value, clipped to 2^(1 + numBins) def unbin (n : Fin numBins) : ℕ+ := 2 ^ (1 + n.val) def combineMeanHistogram : Mechanism ℕ (Option ℚ) := - privMeanHistogram PureDPSystem laplace_pureDPSystem numBins { bin } unbin 1 20 2 1 20 + privMeanHistogram PureDPSystem laplace_pureDPSystem numBins { bin := example_bin } unbin 1 20 2 1 20 end histogramMeanExample diff --git a/SampCert/Samplers/Uniform/Properties.lean b/SampCert/Samplers/Uniform/Properties.lean index fd7539bf..21a85d48 100644 --- a/SampCert/Samplers/Uniform/Properties.lean +++ b/SampCert/Samplers/Uniform/Properties.lean @@ -249,7 +249,7 @@ theorem UniformSample_HasSum_1 (n : PNat) : /-- Conversion of ``UniformSample`` from a ``SLang`` term to a ``PMF``. -/ -noncomputable def UniformSample_PMF (n : PNat) : PMF ℕ := ⟨ UniformSample n , UniformSample_HasSum_1 n⟩ +def UniformSample_PMF (n : PNat) : SPMF ℕ := ⟨ UniformSample n , UniformSample_HasSum_1 n⟩ /-- Evaluation of ``UniformSample`` on ``ℕ`` guarded by its support, when inside the support. diff --git a/Test.lean b/Test.lean index 930ddeb8..c30f3493 100644 --- a/Test.lean +++ b/Test.lean @@ -6,6 +6,7 @@ Authors: Jean-Baptiste Tristan import SampCert import SampCert.SLang import SampCert.Samplers.Gaussian.Properties +import Init.Data.Float open SLang Std Int Array PMF @@ -173,7 +174,61 @@ def test (num den : ℕ+) (mix numSamples : ℕ) (threshold : Float) : IO Unit : IO.println s!"Kolmogorov distance = {D}" -def main : IO Unit := do +def query_tests : IO Unit := do + -- Generate list of 1000 numbers from 0 to 15 + let samples := 10000 + let unif_ub := 10 + let data : List ℕ <- List.mapM (fun _ => run <| (SLang.UniformSample_PMF unif_ub)) (List.replicate samples 0) + + let num : ℕ+ := 9 + let den : ℕ+ := 2 + let num_trials := 5 + + IO.println s!"[query] testing pure ({(num : ℕ)} / {(den : ℕ)})-DP queries" + IO.println s!"data := {samples} uniform samples of [0, {(unif_ub : ℕ)}): {(data.take 20)}..." + IO.println "" + + for i in [:num_trials] do + let ct <- run <| @privNoisedCount _ PureDPSystem laplace_pureDPSystem num den data + IO.println s!"#{i} count: {ct}" + IO.println "" + + let sum_bound : ℕ+ := 10 + for i in [:num_trials] do + let bs <- run <| @privNoisedBoundedSum PureDPSystem laplace_pureDPSystem sum_bound num den data + IO.println s!"#{i} bounded sum (bound = {(sum_bound : ℕ)}): {bs}" + IO.println "" + + for i in [:num_trials] do + let bs <- run <| @privNoisedBoundedMean PureDPSystem laplace_pureDPSystem sum_bound num den data + let float_bs := Float.div (Float.ofInt bs.1) (Float.ofInt bs.2) + IO.println s!"#{i} bounded mean (bound = {(sum_bound : ℕ)}): {bs} (~{float_bs})" + IO.println "" + + for i in [:num_trials] do + let h <- run <| @privNoisedHistogram numBins _ { bin := example_bin } PureDPSystem laplace_pureDPSystem num den data + let h' : List ℤ := h.count.toList.take 25 + IO.println s!"#{i} histogram : {h'}..." + IO.println "" + + let thresh := 100 + for i in [:num_trials] do + let m <- run <| @privMaxBinAboveThreshold numBins _ { bin := example_bin } PureDPSystem laplace_pureDPSystem num den thresh data + IO.println s!"#{i} max bin above threshold (threshold = {(thresh : ℤ)}): {m}" + IO.println "" + + let τ := 100 + for i in [:num_trials] do + let m <- run <| @privMeanHistogram PureDPSystem laplace_pureDPSystem numBins { bin := example_bin } unbin num den τ num den data + let m_float := + match m with + | none => none + | some bs => some (Float.div (Float.ofInt bs.1) (Float.ofInt bs.2)) + IO.println s!"#{i} (0.5x-privacy) histogram mean, τ = {τ}: {m} (~{m_float})" + IO.println "" + +def statistical_tests : IO Unit := do + IO.println s!"[samplers] statistical tests" let tests : List (ℕ+ × ℕ+ × ℕ) := [ -- (1,1,0), (1,1,7), @@ -188,3 +243,8 @@ def main : IO Unit := do for (num,den,mix) in tests do IO.println s!"num = {(num : ℕ)}, den = {(den : ℕ)}, mix = {mix}" test num den mix 100000 0.1 + + +def main : IO Unit := do + query_tests + statistical_tests From 0884bac8a25018a2935584bd366bd4a895365484 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Tue, 12 Nov 2024 11:10:11 -0500 Subject: [PATCH 029/100] rename adp -> app_dp --- SampCert/DifferentialPrivacy/Abstract.lean | 4 ++-- SampCert/DifferentialPrivacy/Pure/System.lean | 2 +- SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean | 2 +- 3 files changed, 4 insertions(+), 4 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index 93480da3..04c56591 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -52,12 +52,12 @@ class DPSystem (T : Type) where /-- Definition of DP is well-formed: Privacy parameter required to obtain (ε', δ)-approximate DP -/ - of_adp : (δ : NNReal) -> (ε' : NNReal) -> NNReal + of_app_dp : (δ : NNReal) -> (ε' : NNReal) -> NNReal /-- For any ε', this definition of DP implies (ε', δ)-approximate-DP for all δ -/ prop_adp [Countable Z] {m : Mechanism T Z} : ∀ (δ : NNReal) (_ : 0 < δ) (ε' : NNReal), - (prop m (of_adp δ ε') -> ApproximateDP m ε' δ) + (prop m (of_app_dp δ ε') -> ApproximateDP m ε' δ) /-- DP is monotonic -/ diff --git a/SampCert/DifferentialPrivacy/Pure/System.lean b/SampCert/DifferentialPrivacy/Pure/System.lean index a8667cf6..c6a7a786 100644 --- a/SampCert/DifferentialPrivacy/Pure/System.lean +++ b/SampCert/DifferentialPrivacy/Pure/System.lean @@ -23,7 +23,7 @@ Pure ε-DP with noise drawn from the discrete Laplace distribution. -/ instance PureDPSystem : DPSystem T where prop := PureDP - of_adp := pure_of_adp + of_app_dp := pure_of_adp prop_adp := pure_ApproximateDP prop_mono := PureDP_mono adaptive_compose_prop := by diff --git a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean index e9501786..7849dad5 100644 --- a/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean +++ b/SampCert/DifferentialPrivacy/ZeroConcentrated/System.lean @@ -27,7 +27,7 @@ Instance of a DP system for zCDP -/ instance zCDPSystem : DPSystem T where prop := zCDP - of_adp := zCDP_of_adp + of_app_dp := zCDP_of_adp prop_adp := by intros; apply zCDP_ApproximateDP <;> assumption prop_mono := by From 6372ea04f9471fb98fc309fa6da85b2898b74610 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Tue, 12 Nov 2024 11:11:26 -0500 Subject: [PATCH 030/100] Countable -> DiscProbSpace --- SampCert/DifferentialPrivacy/Abstract.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/SampCert/DifferentialPrivacy/Abstract.lean b/SampCert/DifferentialPrivacy/Abstract.lean index 04c56591..7fc52d73 100644 --- a/SampCert/DifferentialPrivacy/Abstract.lean +++ b/SampCert/DifferentialPrivacy/Abstract.lean @@ -56,7 +56,7 @@ class DPSystem (T : Type) where /-- For any ε', this definition of DP implies (ε', δ)-approximate-DP for all δ -/ - prop_adp [Countable Z] {m : Mechanism T Z} : ∀ (δ : NNReal) (_ : 0 < δ) (ε' : NNReal), + prop_adp [DiscProbSpace Z] {m : Mechanism T Z} : ∀ (δ : NNReal) (_ : 0 < δ) (ε' : NNReal), (prop m (of_app_dp δ ε') -> ApproximateDP m ε' δ) /-- DP is monotonic From 7d9ba57d2d4f3f2a7eca23812dcf94a5f622773a Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Wed, 25 Sep 2024 10:39:19 -0400 Subject: [PATCH 031/100] prelim --- .../Queries/UnboundedMax/Basic.lean | 10 +++ .../Queries/UnboundedMax/Properties.lean | 90 +++++++++++++++++++ .../Queries/UnboundedMax/\\" | 64 +++++++++++++ 3 files changed, 164 insertions(+) create mode 100644 SampCert/DifferentialPrivacy/Queries/UnboundedMax/Basic.lean create mode 100644 SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean create mode 100644 "SampCert/DifferentialPrivacy/Queries/UnboundedMax/\\" diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Basic.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Basic.lean new file mode 100644 index 00000000..72073d78 --- /dev/null +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Basic.lean @@ -0,0 +1,10 @@ +/- +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.Queries.UnboundedMax.Code +-- import SampCert.DifferentialPrivacy.Queries.UnboundedMax.Properties +-- import SampCert.DifferentialPrivacy.Queries.UnboundedMax.Pure +-- import SampCert.DifferentialPrivacy.Queries.UnboundedMax.ZeroConcentrated diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean new file mode 100644 index 00000000..f289681e --- /dev/null +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -0,0 +1,90 @@ +/- +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 + +noncomputable section +open Classical + +namespace SLang + +/- +## Helper programs +-/ + +/-- +Sum over a list, clipping each element to a maximum. + +Similar to exactBoundedSum, however exactClippedSum allows m = 0. +-/ +def exactClippedSum (m : ℕ) (l : List ℕ) : ℤ := + List.sum (List.map (fun n : ℕ => (Nat.min n m)) l) + +/-- +Rate at which a given clipping thresold is impacting the accuracy of the sum. + +Always negative or zero. +-/ +def exactDiffSum (m : ℕ) (l : List ℕ) : ℤ := exactClippedSum m l - exactClippedSum (m + 1) l + +/-- +Noise the constant 0 value using the abstract noise function. + +This looks strange, but will specialize to Lap(ε₁/ε₂, 0) in the pure DP case. +-/ +def privNoiseZero [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) : SLang ℤ := dps.noise (fun _ => 0) 1 ε₁ ε₂ [] + +/-- +Not used for anything, but to give confidence in our definitions + +(exactDiffSum l m) is zero if and only if m is an upper bound on the list elements +-/ +lemma exactDiffSum_eq_0_iff_ge_max (l : List ℕ) (m : ℕ) : + l.maximum ≤ m <-> exactDiffSum m l = 0 := by + apply Iff.intro + · induction l + · simp [exactDiffSum, exactClippedSum] + · rename_i l0 ls IH + intro Hmax + simp [List.maximum_cons] at Hmax + rcases Hmax with ⟨ Hmax0, Hmax1 ⟩ + have IH' := IH Hmax1 + clear IH + simp [exactDiffSum, exactClippedSum] at * + rw [add_sub_add_comm] + rw [IH'] + rw [Int.min_def, Int.min_def] + split + · split + · simp + · linarith + · exfalso + rename_i h + apply h + simp + rw [Nat.cast_withBot, WithBot.coe_le_coe] at Hmax0 + trivial + · intro H1 + apply List.maximum_le_of_forall_le + revert H1 + induction l + · simp + · rename_i l0 ls IH + intro Hdiff a Ha + rw [List.mem_cons_eq] at Ha + cases Ha + · sorry + · apply IH + · sorry + · trivial + + + +/- +## Program version 0 + - Executable + - Tracks single state +-/ diff --git "a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/\\" "b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/\\" new file mode 100644 index 00000000..739c3aef --- /dev/null +++ "b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/\\" @@ -0,0 +1,64 @@ +/- +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 + +noncomputable section +open Classical + +namespace SLang + +/- +## Helper programs +-/ + +/-- +Sum over a list, clipping each element to a maximum. + +Similar to exactBoundedSum, however exactClippedSum allows m = 0. +-/ +def exactClippedSum (m : ℕ) (l : List ℕ) : ℤ := + List.sum (List.map (fun n : ℕ => (Nat.min n m)) l) + +/-- +Rate at which a given clipping thresold is impacting the accuracy of the sum. + +Always negative or zero. +-/ +def exactDiffSum (m : ℕ) (l : List ℕ) : ℤ := exactClippedSum m l - exactClippedSum (m + 1) l + +/-- +Noise the constant 0 value using the abstract noise function. + +This looks strange, but will specialize to Lap(ε₁/ε₂, 0) in the pure DP case. +-/ +def privNoiseZero [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) : SLang ℤ := dps.noise (fun _ => 0) 1 ε₁ ε₂ [] + +#check upperBounds + +/-- +Not used for anything, but gives confidence in our defintions + +(exactDiffSum l m) is zero if and only if m is an upper bound on the list elements +-/ +lemma exactDiffSum_eq_0_iff_ge_max (l : List ℕ) (m : ℕ) : + l.maximum ≤ m <-> exactDiffSum m l = 0 := by + apply Iff.intro + · intro Hmax + + sorry + · sorry + + + + + + +/- +## Program version 0 + - Executable + - Tracks single state +-/ From cabf7a2029825c7990175b655c7a616176302645 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Thu, 26 Sep 2024 07:54:34 -0400 Subject: [PATCH 032/100] checkpoint --- .../Queries/UnboundedMax/Properties.lean | 92 +++++++++++++++---- 1 file changed, 75 insertions(+), 17 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index f289681e..75dbbf8d 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -43,7 +43,7 @@ Not used for anything, but to give confidence in our definitions (exactDiffSum l m) is zero if and only if m is an upper bound on the list elements -/ lemma exactDiffSum_eq_0_iff_ge_max (l : List ℕ) (m : ℕ) : - l.maximum ≤ m <-> exactDiffSum m l = 0 := by + l.maximum ≤ m <-> exactDiffSum m l ≤ 0 := by apply Iff.intro · induction l · simp [exactDiffSum, exactClippedSum] @@ -54,19 +54,9 @@ lemma exactDiffSum_eq_0_iff_ge_max (l : List ℕ) (m : ℕ) : have IH' := IH Hmax1 clear IH simp [exactDiffSum, exactClippedSum] at * - rw [add_sub_add_comm] - rw [IH'] - rw [Int.min_def, Int.min_def] - split - · split - · simp - · linarith - · exfalso - rename_i h - apply h - simp - rw [Nat.cast_withBot, WithBot.coe_le_coe] at Hmax0 - trivial + apply Int.add_le_of_le_neg_add + apply le_trans IH' + simp · intro H1 apply List.maximum_le_of_forall_le revert H1 @@ -76,9 +66,20 @@ lemma exactDiffSum_eq_0_iff_ge_max (l : List ℕ) (m : ℕ) : intro Hdiff a Ha rw [List.mem_cons_eq] at Ha cases Ha - · sorry - · apply IH - · sorry + · rename_i H + rw [H] + rw [Nat.cast_withBot] + apply WithBot.coe_le_coe.mpr + + sorry + · apply IH; clear IH + · simp only [exactDiffSum, exactClippedSum] at * + have H : (min (l0.cast : ℤ) (m.cast : ℤ) - min (l0.cast) ((m.cast : ℤ) + 1)) = 0 := by + sorry + -- rw [H] at Hdiff + -- rw [<- Hdiff] + -- simp + sorry · trivial @@ -88,3 +89,60 @@ lemma exactDiffSum_eq_0_iff_ge_max (l : List ℕ) (m : ℕ) : - Executable - Tracks single state -/ + +def sv0_state : Type := ℕ × ℤ + +def sv0_threshold (s : sv0_state) : ℕ := s.1 + +def sv0_noise (s : sv0_state) : ℤ := s.2 + +def sv0_privMaxC (τ : ℤ) (l : List ℕ) (s : sv0_state) : Bool := + decide (exactDiffSum (sv0_threshold s) l + (sv0_noise s) < τ) + +def sv0_privMaxF [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (s : sv0_state) : SLang sv0_state := do + let vn <- @privNoiseZero dps ε₁ (4 * ε₂) + let n := (sv0_threshold s) + 1 + return (n, vn) + +def sv0_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do + let τ <- @privNoiseZero dps ε₁ (2 * ε₂) + let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + let sk <- probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (0, v0) + return (sv0_threshold sk) + +/- +## Program version 1 + - Executable + - Tracks history of samples +-/ + +def sv1_state : Type := List ℤ × ℤ + +def sv1_threshold (s : sv1_state) : ℕ := List.length s.1 + +def sv1_noise (s : sv1_state) : ℤ := s.2 + +def sv1_privMaxC (τ : ℤ) (l : List ℕ) (s : sv1_state) : Bool := + decide (exactDiffSum (sv1_threshold s) l + (sv1_noise s) < τ) + +def sv1_privMaxF [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (s : sv1_state) : SLang sv1_state := do + let vn <- @privNoiseZero dps ε₁ (4 * ε₂) + return (s.1 ++ [s.2], vn) + +def sv1_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do + let τ <- @privNoiseZero dps ε₁ (2 * ε₂) + let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + let sk <- probWhile (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) ([], v0) + return (sv1_threshold sk) + +/-- +History-aware progam computes the same as the history-agnostic program +-/ +lemma sv0_eq_sv1 [dps : DPSystem ℕ] ε₁ ε₂ l : sv0_privMax ε₁ ε₂ l = sv1_privMax ε₁ ε₂ l := by + apply SLang.ext + intro r + + + + + sorry From 0f890425526ce86233309af7929f8810725b9f59 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Thu, 26 Sep 2024 10:16:55 -0400 Subject: [PATCH 033/100] checkpoint sv2 --- .../Queries/UnboundedMax/Properties.lean | 56 ++++++++++++++++++- 1 file changed, 54 insertions(+), 2 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 75dbbf8d..6edf6b4d 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -140,9 +140,61 @@ History-aware progam computes the same as the history-agnostic program -/ lemma sv0_eq_sv1 [dps : DPSystem ℕ] ε₁ ε₂ l : sv0_privMax ε₁ ε₂ l = sv1_privMax ε₁ ε₂ l := by apply SLang.ext - intro r + -- Initial setup is equal + intro result + simp [sv0_privMax, sv1_privMax] + apply tsum_congr + intro τ + congr 1 + apply tsum_congr + intro v0 + congr 1 + + -- unfold sum over product; term re. last sample should be equal as well + conv => + congr + · unfold sv0_state + rw [ENNReal.tsum_prod', ENNReal.tsum_comm] + · unfold sv1_state + rw [ENNReal.tsum_prod', ENNReal.tsum_comm] + apply tsum_congr + intro vk + + -- LHS: singleton sum + -- RHS: sum over all lists of length "result" + -- rw [tsum_ite_eq] + sorry - sorry +/- +## Program version 2 + - Only moves the loop into a non-executable form, ie. explicitly defines the PMF +-/ + +def sv2_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := + fun (point : ℕ) => + let computation : SLang ℕ := do + let τ <- @privNoiseZero dps ε₁ (2 * ε₂) + let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + let sk <- probWhile (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) ([], v0) + return (sv1_threshold sk) + computation point + +lemma sv1_eq_sv2 [dps : DPSystem ℕ] ε₁ ε₂ l : sv1_privMax ε₁ ε₂ l = sv2_privMax ε₁ ε₂ l := by + apply SLang.ext + intro result + simp [sv1_privMax, sv2_privMax] + + + + + + + +-- def sv1_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do +-- let τ <- @privNoiseZero dps ε₁ (2 * ε₂) +-- let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) +-- let sk <- probWhile (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) ([], v0) +-- return (sv1_threshold sk) From c5e7ee0275805a99bce4183efe67c22e6ff62f9d Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Thu, 26 Sep 2024 10:46:29 -0400 Subject: [PATCH 034/100] checkpoint sv2->sv3 --- .../Queries/UnboundedMax/Properties.lean | 135 ++++++++++++------ 1 file changed, 89 insertions(+), 46 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 6edf6b4d..9fc8ba9b 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -37,51 +37,50 @@ This looks strange, but will specialize to Lap(ε₁/ε₂, 0) in the pure DP ca -/ def privNoiseZero [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) : SLang ℤ := dps.noise (fun _ => 0) 1 ε₁ ε₂ [] -/-- +/- Not used for anything, but to give confidence in our definitions (exactDiffSum l m) is zero if and only if m is an upper bound on the list elements -/ -lemma exactDiffSum_eq_0_iff_ge_max (l : List ℕ) (m : ℕ) : - l.maximum ≤ m <-> exactDiffSum m l ≤ 0 := by - apply Iff.intro - · induction l - · simp [exactDiffSum, exactClippedSum] - · rename_i l0 ls IH - intro Hmax - simp [List.maximum_cons] at Hmax - rcases Hmax with ⟨ Hmax0, Hmax1 ⟩ - have IH' := IH Hmax1 - clear IH - simp [exactDiffSum, exactClippedSum] at * - apply Int.add_le_of_le_neg_add - apply le_trans IH' - simp - · intro H1 - apply List.maximum_le_of_forall_le - revert H1 - induction l - · simp - · rename_i l0 ls IH - intro Hdiff a Ha - rw [List.mem_cons_eq] at Ha - cases Ha - · rename_i H - rw [H] - rw [Nat.cast_withBot] - apply WithBot.coe_le_coe.mpr - - sorry - · apply IH; clear IH - · simp only [exactDiffSum, exactClippedSum] at * - have H : (min (l0.cast : ℤ) (m.cast : ℤ) - min (l0.cast) ((m.cast : ℤ) + 1)) = 0 := by - sorry - -- rw [H] at Hdiff - -- rw [<- Hdiff] - -- simp - sorry - · trivial - +-- lemma exactDiffSum_eq_0_iff_ge_max (l : List ℕ) (m : ℕ) : +-- l.maximum ≤ m <-> exactDiffSum m l ≤ 0 := by +-- apply Iff.intro +-- · induction l +-- · simp [exactDiffSum, exactClippedSum] +-- · rename_i l0 ls IH +-- intro Hmax +-- simp [List.maximum_cons] at Hmax +-- rcases Hmax with ⟨ Hmax0, Hmax1 ⟩ +-- have IH' := IH Hmax1 +-- clear IH +-- simp [exactDiffSum, exactClippedSum] at * +-- apply Int.add_le_of_le_neg_add +-- apply le_trans IH' +-- simp +-- · intro H1 +-- apply List.maximum_le_of_forall_le +-- revert H1 +-- induction l +-- · simp +-- · rename_i l0 ls IH +-- intro Hdiff a Ha +-- rw [List.mem_cons_eq] at Ha +-- cases Ha +-- · rename_i H +-- rw [H] +-- rw [Nat.cast_withBot] +-- apply WithBot.coe_le_coe.mpr +-- +-- sorry +-- · apply IH; clear IH +-- · simp only [exactDiffSum, exactClippedSum] at * +-- have H : (min (l0.cast : ℤ) (m.cast : ℤ) - min (l0.cast) ((m.cast : ℤ) + 1)) = 0 := by +-- sorry +-- -- rw [H] at Hdiff +-- -- rw [<- Hdiff] +-- -- simp +-- sorry +-- · trivial /- @@ -190,11 +189,55 @@ lemma sv1_eq_sv2 [dps : DPSystem ℕ] ε₁ ε₂ l : sv1_privMax ε₁ ε₂ l +/- +## Program version 3 + - Truncates the loop +-/ + +def sv3_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := + fun (point : ℕ) => + let computation : SLang ℕ := do + let τ <- @privNoiseZero dps ε₁ (2 * ε₂) + let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + let sk <- probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (point + 1) ([], v0) + return (sv1_threshold sk) + computation point +lemma sv2_eq_sv3 [dps : DPSystem ℕ] ε₁ ε₂ l : sv2_privMax ε₁ ε₂ l = sv3_privMax ε₁ ε₂ l := by + apply SLang.ext --- def sv1_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do --- let τ <- @privNoiseZero dps ε₁ (2 * ε₂) --- let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) --- let sk <- probWhile (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) ([], v0) --- return (sv1_threshold sk) + -- Step through equal headers + intro point + unfold sv2_privMax + unfold sv3_privMax + simp + apply tsum_congr + intro τ + congr 1 + apply tsum_congr + intro v0 + congr 1 + apply tsum_congr + intro final_state + rcases final_state with ⟨ hist, vk ⟩ + split <;> try rfl + rename_i H + simp [H, sv1_threshold] + clear H + + -- TODO: Move over the eventual constancy statement, and prove it + + -- This statement is wrong, delete it + -- + -- -- Want to say that the masses of the final returned _value_ are the same, + -- -- even though the lengths of the histories may change, + -- -- because past a certain point, adding to the history does not change the + -- -- probability of returning a given value. + -- conv => + -- unfold sv1_state + -- congr + -- · rw [ENNReal.tsum_prod', ENNReal.tsum_comm] + -- · rw [ENNReal.tsum_prod', ENNReal.tsum_comm] + -- apply tsum_congr + -- intro point' From 4d58b9d17f424330ac0edabd4d4a3e21cf3b82bc Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Thu, 26 Sep 2024 11:13:41 -0400 Subject: [PATCH 035/100] checkpoint --- .../Queries/UnboundedMax/Properties.lean | 68 +++++++++++++++++++ 1 file changed, 68 insertions(+) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 9fc8ba9b..abfb4dd3 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -241,3 +241,71 @@ lemma sv2_eq_sv3 [dps : DPSystem ℕ] ε₁ ε₂ l : sv2_privMax ε₁ ε₂ l -- · rw [ENNReal.tsum_prod', ENNReal.tsum_comm] -- apply tsum_congr -- intro point' + + sorry + + + +/- +## Program version 4 + - Presamples each vk + - Iteratively checks if the loop + - Still loops through prefixes of the vk's iteratively +-/ + +-- Presamples the list (v_(n-1), v_(n-2), ..., v0) +def sv4_presample [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (n : ℕ) : SLang (List ℤ) := do + match n with + | Nat.zero => return [] + | (Nat.succ n') => do + let v <- @privNoiseZero dps ε₁ (4 * ε₂) + let l <- sv4_presample ε₁ ε₂ n' + return (v :: l) + +def sv4_evaluate_history (vks : List ℤ) (τ : ℤ) (l : List ℕ) : Bool := + match vks with + | [] => true + | (v0 :: vs) => sv1_privMaxC τ l (vs, v0) && sv4_evaluate_history vs τ l + + +def sv4_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := + fun (point : ℕ) => + let computation : SLang ℕ := do + let τ <- @privNoiseZero dps ε₁ (2 * ε₂) + let history <- sv4_presample ε₁ ε₂ (point + 1) + let vk <- @privNoiseZero dps ε₁ (4 * ε₂) + if (sv4_evaluate_history history τ l && (! sv1_privMaxC τ l (history, vk))) + then return point + else probZero + computation point + +lemma sv3_eq_sv4 [dps : DPSystem ℕ] ε₁ ε₂ l : sv3_privMax ε₁ ε₂ l = sv4_privMax ε₁ ε₂ l := by + sorry + + + + + +/- +## Program version 5 + - Presamples each vk + - Performs a single check, rather than a loop +-/ +def sv5_evaluate_history (vks : List ℤ) (τ : ℤ) (l : List ℕ) : Bool := + match vks with + | [] => true + | (v0 :: vs) => sv1_privMaxC τ l (vs, v0) + +def sv5_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := + fun (point : ℕ) => + let computation : SLang ℕ := do + let τ <- @privNoiseZero dps ε₁ (2 * ε₂) + let history <- sv4_presample ε₁ ε₂ (point + 1) + let vk <- @privNoiseZero dps ε₁ (4 * ε₂) + if (sv5_evaluate_history history τ l && (! sv1_privMaxC τ l (history, vk))) + then return point + else probZero + computation point + +lemma sv4_eq_sv5 [dps : DPSystem ℕ] ε₁ ε₂ l : sv4_privMax ε₁ ε₂ l = sv5_privMax ε₁ ε₂ l := by + sorry From 845ba9ff6be273e5fbf102f42990397339e0e78a Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 30 Sep 2024 14:33:37 -0400 Subject: [PATCH 036/100] constancy inside the cone of possibility --- .../Queries/UnboundedMax/Properties.lean | 277 +++++++++++++++++- .../Queries/UnboundedMax/\\" | 64 ---- 2 files changed, 261 insertions(+), 80 deletions(-) delete mode 100644 "SampCert/DifferentialPrivacy/Queries/UnboundedMax/\\" diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index abfb4dd3..c28f1f53 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -160,8 +160,28 @@ lemma sv0_eq_sv1 [dps : DPSystem ℕ] ε₁ ε₂ l : sv0_privMax ε₁ ε₂ l apply tsum_congr intro vk - -- LHS: singleton sum + -- LHS: simplify singleton sum + conv => + enter [1, 1, a] + simp [sv0_threshold] + rw [ENNReal.tsum_eq_add_tsum_ite result] + simp + rw [@tsum_congr ENNReal ℕ _ _ _ (fun _ => 0) ?G1] + case G1 => + intro + split <;> simp + rename_i H1 + intro + exfalso + apply H1 + symm + trivial + simp only [tsum_zero, add_zero] + -- RHS: sum over all lists of length "result" + -- Prove this by induction on result + + -- rw [tsum_ite_eq] sorry @@ -204,6 +224,242 @@ def sv3_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang computation point +-- Lemmas about cut loops + +-- Loop cut to zero iterates is zero everywhere +lemma loop_cut_0_supp [dps : DPSystem ℕ] : + (hist.length ≥ 0) -> + probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) 0 ([], v0) (hist, vk) = 0 := by + simp [probWhileCut] + +-- loop cut to 1 is zero unless hist.length is 0 +lemma loop_cut_1_supp [dps : DPSystem ℕ] : + (hist.length ≥ 1) -> + probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) 1 ([], v0) (hist, vk) = 0 := by + intro H0 + simp [probWhileCut, probWhileFunctional] + cases (sv1_privMaxC τ l ([], v0)) + · -- Loop check is false, returns ([], v0) + simp + intro H1 + cases H1 + simp at H0 + · -- Loop checks to true, passes to base case (prob_zero) + simp + +-- loop cut to 2 is zero unless the length is at most 1 +lemma loop_cut_2_supp [dps : DPSystem ℕ] : + (hist.length ≥ 2) -> + probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) 2 ([], v0) (hist, vk) = 0 := by + intro H0 + simp [probWhileCut, probWhileFunctional] + cases (sv1_privMaxC τ l ([], v0)) + · -- Loop check is false, returns ([], v0) + -- excluded by length hypothesis + simp + intro H1 + cases H1 + simp at H0 + -- loop check is true + simp + intro ⟨ hist1, v1 ⟩ + + -- F is applied, always adds a single sample to the state + cases (Classical.em (∃ v, hist1 = [v])) + · -- F did add a single sample to the tape + rename_i h + rcases h with ⟨ v0, Hhist1 ⟩ + rw [Hhist1] + right + + cases (sv1_privMaxC τ l ([v0], v1)) + · -- Check is false, returns ([v0], v1) + -- Excluded by length hypothesis + simp + intro H1 + cases H1 + simp at H0 + + · -- Check is true + simp + + · -- F didn't add a single sample to the tape, is impossible + -- FIXME: Refactor this out + left + rename_i h + simp [sv1_privMaxF] + intro v1 H + exfalso + apply h + exists v0 + cases H + rfl + + +lemma loop_cut_2_cut_1_eq_len_0 [dps : DPSystem ℕ] : + (hist.length = 0) -> + probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) 2 ([], v0) (hist, vk) = + probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) 1 ([], v0) (hist, vk) := by + intro H + simp [probWhileCut, probWhileFunctional] + cases (sv1_privMaxC τ l ([], v0)) + · -- First check is false + -- Both sides step to probPure ([], v0), which are equal + simp only [Bool.false_eq_true, ↓reduceIte] + · -- First check is true + -- RHS steps to probZero + simp only [↓reduceIte, bind_apply, zero_apply, mul_zero, tsum_zero, ENNReal.tsum_eq_zero, mul_eq_zero] + -- Apply F on the left-hand side + intro ⟨ hist', v1 ⟩ + cases (Classical.em (∃ v, hist' = [v])) + · right + rename_i h + rcases h with ⟨ h0, Hhist' ⟩ + -- We have already added too much to the history (since hist.lenght = 0) + -- All cases from here on out are zero + split <;> simp + intro H2 + cases H2 + simp_all + · -- F didn't add exactly one sample (impossible) + left + rename_i h + simp [sv1_privMaxF] + intro v1 H + exfalso + apply h + exists v0 + cases H + rfl + + +def cone_of_possibility (cut : ℕ) (initial hist : List ℤ) : Prop := + (hist.length < cut + initial.length) ∧ (initial.length ≤ hist.length) + +def constancy_at [DPSystem ℕ] {ε₁ ε₂ : ℕ+} {τ : ℤ} {data : List ℕ} {v0 vk : ℤ} (cut : ℕ) (initial hist : List ℤ) : Prop := + probWhileCut (sv1_privMaxC τ data) (sv1_privMaxF ε₁ ε₂) (1 + cut) (initial, v0) (hist, vk) = + probWhileCut (sv1_privMaxC τ data) (sv1_privMaxF ε₁ ε₂) cut (initial, v0) (hist, vk) + + +-- All points to the left of the cone are zero +lemma cone_left_zero [DPSystem ℕ] : + hist.length < initial.length -> + probWhileCut (sv1_privMaxC τ data) (sv1_privMaxF ε₁ ε₂) cut (initial, v0) (hist, vk) = 0 := by + sorry + +-- All points below the cone are zero +lemma cone_below_zero [DPSystem ℕ] : + cut + initial.length ≤ hist.length -> + probWhileCut (sv1_privMaxC τ data) (sv1_privMaxF ε₁ ε₂) cut (initial, v0) (hist, vk) = 0 := by + sorry + +-- Base case: left edge of the cone satisfies constancy +lemma cone_left_edge_constancy [DPSystem ℕ] {ε₁ ε₂ : ℕ+} {τ : ℤ} {data : List ℕ} {v0 vk : ℤ} (cut : ℕ) (initial hist : List ℤ) : + hist.length = initial.length -> + cone_of_possibility cut initial hist -> + @constancy_at _ ε₁ ε₂ τ data v0 vk cut initial hist := by + sorry + +lemma cone_constancy [DPSystem ℕ] {ε₁ ε₂ : ℕ+} {τ : ℤ} {data : List ℕ} {v0 vk : ℤ} (cut : ℕ) (initial hist : List ℤ) : + cone_of_possibility cut initial hist -> + @constancy_at _ ε₁ ε₂ τ data v0 vk cut initial hist := by + -- Need theorem to be true for all initial states, since this will increase during the induction + -- v0 and vk will also change in ways which don't matter + revert initial v0 vk + + induction cut + · -- Not true base case (cut 0 is always outside of the cone) + -- Mercifully it is easy to prove + intro v0 vk initial Hcone + unfold constancy_at + simp [probWhileCut, probWhileFunctional] + cases (sv1_privMaxC τ data (initial, v0)) <;> simp + unfold cone_of_possibility at Hcone + linarith + + rename_i n IH + intro v0 vk initial Hcone + -- True base case: are we on the left-hand edge of the cone + cases Classical.em (hist.length = initial.length) + · apply cone_left_edge_constancy <;> assumption + + -- If not, unfold the first (and only the first) level of the induction + unfold constancy_at + unfold probWhileCut + unfold probWhileFunctional + + -- If the conditional is false, we are done + cases (sv1_privMaxC τ data (initial, v0)) + · simp + + + -- If the conditional is true, we decrement N by one and add a sample to the history + rename_i Hcone_interior + unfold constancy_at at IH + simp + apply tsum_congr + intro ⟨ initial', vk' ⟩ + + -- We only need to consider the cases where sv1_privMaxF is nonzero + -- And this is exactly the case where initial' is initial plus one element + simp [sv1_privMaxF] + rw [← ENNReal.tsum_mul_right] + rw [← ENNReal.tsum_mul_right] + apply tsum_congr + intro z + cases Classical.em (¬ ∃ v', initial' = initial ++ [v']) + · split + · exfalso + rename_i h1 h2 + apply h1 + exists v0 + cases h2 + trivial + · simp + rename_i h + simp at h + rcases h with ⟨ v', Hinitial' ⟩ + split <;> try simp + rename_i h + cases h + congr 1 + + -- Apply the IH + apply IH + + -- Prove that the new value is in the new cone of possibility + unfold cone_of_possibility at Hcone + rcases Hcone with ⟨ Hcone1, Hcone2 ⟩ + unfold cone_of_possibility + apply And.intro + · simp + linarith + · simp + apply Nat.lt_iff_add_one_le.mp + apply LE.le.eq_or_lt at Hcone2 + cases Hcone2 + · exfalso + apply Hcone_interior + symm + trivial + · trivial + + +-- Define cone +-- Define constancy, parameterized by initial list length +-- Prove zero for the region to the left of the cone (done before) +-- Prove zero for the region below the cone (done before) +-- Define base case in terms of cone +-- Prove base case in terms of cone +-- Prove inductive case for code + + + + + + + + lemma sv2_eq_sv3 [dps : DPSystem ℕ] ε₁ ε₂ l : sv2_privMax ε₁ ε₂ l = sv3_privMax ε₁ ε₂ l := by apply SLang.ext @@ -226,21 +482,10 @@ lemma sv2_eq_sv3 [dps : DPSystem ℕ] ε₁ ε₂ l : sv2_privMax ε₁ ε₂ l simp [H, sv1_threshold] clear H - -- TODO: Move over the eventual constancy statement, and prove it - - -- This statement is wrong, delete it - -- - -- -- Want to say that the masses of the final returned _value_ are the same, - -- -- even though the lengths of the histories may change, - -- -- because past a certain point, adding to the history does not change the - -- -- probability of returning a given value. - -- conv => - -- unfold sv1_state - -- congr - -- · rw [ENNReal.tsum_prod', ENNReal.tsum_comm] - -- · rw [ENNReal.tsum_prod', ENNReal.tsum_comm] - -- apply tsum_congr - -- intro point' + -- Apply a lemma about eventual constancy + apply probWhile_apply + apply @tendsto_atTop_of_eventually_const _ _ _ _ _ _ _ (hist.length + 1) + intro i H sorry diff --git "a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/\\" "b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/\\" deleted file mode 100644 index 739c3aef..00000000 --- "a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/\\" +++ /dev/null @@ -1,64 +0,0 @@ -/- -Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jean-Baptiste Tristan --/ - -import SampCert.DifferentialPrivacy.Abstract - -noncomputable section -open Classical - -namespace SLang - -/- -## Helper programs --/ - -/-- -Sum over a list, clipping each element to a maximum. - -Similar to exactBoundedSum, however exactClippedSum allows m = 0. --/ -def exactClippedSum (m : ℕ) (l : List ℕ) : ℤ := - List.sum (List.map (fun n : ℕ => (Nat.min n m)) l) - -/-- -Rate at which a given clipping thresold is impacting the accuracy of the sum. - -Always negative or zero. --/ -def exactDiffSum (m : ℕ) (l : List ℕ) : ℤ := exactClippedSum m l - exactClippedSum (m + 1) l - -/-- -Noise the constant 0 value using the abstract noise function. - -This looks strange, but will specialize to Lap(ε₁/ε₂, 0) in the pure DP case. --/ -def privNoiseZero [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) : SLang ℤ := dps.noise (fun _ => 0) 1 ε₁ ε₂ [] - -#check upperBounds - -/-- -Not used for anything, but gives confidence in our defintions - -(exactDiffSum l m) is zero if and only if m is an upper bound on the list elements --/ -lemma exactDiffSum_eq_0_iff_ge_max (l : List ℕ) (m : ℕ) : - l.maximum ≤ m <-> exactDiffSum m l = 0 := by - apply Iff.intro - · intro Hmax - - sorry - · sorry - - - - - - -/- -## Program version 0 - - Executable - - Tracks single state --/ From 4edd572134ce960d8f0d4c7732260878496c590e Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 30 Sep 2024 14:49:00 -0400 Subject: [PATCH 037/100] cone_below_zero --- .../Queries/UnboundedMax/Properties.lean | 30 ++++++++++++++++++- 1 file changed, 29 insertions(+), 1 deletion(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index c28f1f53..4333ad92 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -351,7 +351,35 @@ lemma cone_left_zero [DPSystem ℕ] : lemma cone_below_zero [DPSystem ℕ] : cut + initial.length ≤ hist.length -> probWhileCut (sv1_privMaxC τ data) (sv1_privMaxF ε₁ ε₂) cut (initial, v0) (hist, vk) = 0 := by - sorry + revert initial v0 vk + induction cut + · simp [probWhileCut, probWhileFunctional] + · rename_i cut' IH + intro intial v0 vk Hcut' + unfold probWhileCut + unfold probWhileFunctional + split <;> simp + · intro h + rcases h with ⟨ initial', vk' ⟩ + cases Classical.em (¬ ∃ v', initial' = intial ++ [v']) + · left + simp [sv1_privMaxF] + intro i Hi + exfalso + cases Hi + rename_i h + apply h + exists v0 + rename_i h + simp at h + rcases h with ⟨ v', Hinitial' ⟩ + right + apply IH + simp_all + linarith + · intro H + cases H + simp at Hcut' -- Base case: left edge of the cone satisfies constancy lemma cone_left_edge_constancy [DPSystem ℕ] {ε₁ ε₂ : ℕ+} {τ : ℤ} {data : List ℕ} {v0 vk : ℤ} (cut : ℕ) (initial hist : List ℤ) : From d9c5505e08d6b2c6e93c5e1ce241cdbe6aa196f0 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 30 Sep 2024 14:53:46 -0400 Subject: [PATCH 038/100] external_to_cone_zero --- .../Queries/UnboundedMax/Properties.lean | 18 +++++++----------- 1 file changed, 7 insertions(+), 11 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 4333ad92..49a6c074 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -341,15 +341,9 @@ def constancy_at [DPSystem ℕ] {ε₁ ε₂ : ℕ+} {τ : ℤ} {data : List ℕ probWhileCut (sv1_privMaxC τ data) (sv1_privMaxF ε₁ ε₂) cut (initial, v0) (hist, vk) --- All points to the left of the cone are zero -lemma cone_left_zero [DPSystem ℕ] : - hist.length < initial.length -> - probWhileCut (sv1_privMaxC τ data) (sv1_privMaxF ε₁ ε₂) cut (initial, v0) (hist, vk) = 0 := by - sorry - --- All points below the cone are zero -lemma cone_below_zero [DPSystem ℕ] : - cut + initial.length ≤ hist.length -> +-- All points outside of the cone are zero +lemma external_to_cone_zero [DPSystem ℕ] : + (¬ cone_of_possibility cut initial hist) -> probWhileCut (sv1_privMaxC τ data) (sv1_privMaxF ε₁ ε₂) cut (initial, v0) (hist, vk) = 0 := by revert initial v0 vk induction cut @@ -375,11 +369,13 @@ lemma cone_below_zero [DPSystem ℕ] : rcases h with ⟨ v', Hinitial' ⟩ right apply IH - simp_all + simp_all [cone_of_possibility] + intro + have Hcut'' := Hcut' (by linarith) linarith · intro H cases H - simp at Hcut' + simp_all [cone_of_possibility] -- Base case: left edge of the cone satisfies constancy lemma cone_left_edge_constancy [DPSystem ℕ] {ε₁ ε₂ : ℕ+} {τ : ℤ} {data : List ℕ} {v0 vk : ℤ} (cut : ℕ) (initial hist : List ℤ) : From 120b2d75c610dcea7399024ef3d53adbe2db2fe0 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 30 Sep 2024 15:23:47 -0400 Subject: [PATCH 039/100] reduce sv2_eq_sv3 to cone_left_edge_constancy --- .../Queries/UnboundedMax/Properties.lean | 181 ++++++------------ 1 file changed, 54 insertions(+), 127 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 49a6c074..a6576016 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -224,115 +224,6 @@ def sv3_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang computation point --- Lemmas about cut loops - --- Loop cut to zero iterates is zero everywhere -lemma loop_cut_0_supp [dps : DPSystem ℕ] : - (hist.length ≥ 0) -> - probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) 0 ([], v0) (hist, vk) = 0 := by - simp [probWhileCut] - --- loop cut to 1 is zero unless hist.length is 0 -lemma loop_cut_1_supp [dps : DPSystem ℕ] : - (hist.length ≥ 1) -> - probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) 1 ([], v0) (hist, vk) = 0 := by - intro H0 - simp [probWhileCut, probWhileFunctional] - cases (sv1_privMaxC τ l ([], v0)) - · -- Loop check is false, returns ([], v0) - simp - intro H1 - cases H1 - simp at H0 - · -- Loop checks to true, passes to base case (prob_zero) - simp - --- loop cut to 2 is zero unless the length is at most 1 -lemma loop_cut_2_supp [dps : DPSystem ℕ] : - (hist.length ≥ 2) -> - probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) 2 ([], v0) (hist, vk) = 0 := by - intro H0 - simp [probWhileCut, probWhileFunctional] - cases (sv1_privMaxC τ l ([], v0)) - · -- Loop check is false, returns ([], v0) - -- excluded by length hypothesis - simp - intro H1 - cases H1 - simp at H0 - -- loop check is true - simp - intro ⟨ hist1, v1 ⟩ - - -- F is applied, always adds a single sample to the state - cases (Classical.em (∃ v, hist1 = [v])) - · -- F did add a single sample to the tape - rename_i h - rcases h with ⟨ v0, Hhist1 ⟩ - rw [Hhist1] - right - - cases (sv1_privMaxC τ l ([v0], v1)) - · -- Check is false, returns ([v0], v1) - -- Excluded by length hypothesis - simp - intro H1 - cases H1 - simp at H0 - - · -- Check is true - simp - - · -- F didn't add a single sample to the tape, is impossible - -- FIXME: Refactor this out - left - rename_i h - simp [sv1_privMaxF] - intro v1 H - exfalso - apply h - exists v0 - cases H - rfl - - -lemma loop_cut_2_cut_1_eq_len_0 [dps : DPSystem ℕ] : - (hist.length = 0) -> - probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) 2 ([], v0) (hist, vk) = - probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) 1 ([], v0) (hist, vk) := by - intro H - simp [probWhileCut, probWhileFunctional] - cases (sv1_privMaxC τ l ([], v0)) - · -- First check is false - -- Both sides step to probPure ([], v0), which are equal - simp only [Bool.false_eq_true, ↓reduceIte] - · -- First check is true - -- RHS steps to probZero - simp only [↓reduceIte, bind_apply, zero_apply, mul_zero, tsum_zero, ENNReal.tsum_eq_zero, mul_eq_zero] - -- Apply F on the left-hand side - intro ⟨ hist', v1 ⟩ - cases (Classical.em (∃ v, hist' = [v])) - · right - rename_i h - rcases h with ⟨ h0, Hhist' ⟩ - -- We have already added too much to the history (since hist.lenght = 0) - -- All cases from here on out are zero - split <;> simp - intro H2 - cases H2 - simp_all - · -- F didn't add exactly one sample (impossible) - left - rename_i h - simp [sv1_privMaxF] - intro v1 H - exfalso - apply h - exists v0 - cases H - rfl - - def cone_of_possibility (cut : ℕ) (initial hist : List ℤ) : Prop := (hist.length < cut + initial.length) ∧ (initial.length ≤ hist.length) @@ -382,7 +273,22 @@ lemma cone_left_edge_constancy [DPSystem ℕ] {ε₁ ε₂ : ℕ+} {τ : ℤ} {d hist.length = initial.length -> cone_of_possibility cut initial hist -> @constancy_at _ ε₁ ε₂ τ data v0 vk cut initial hist := by - sorry + intro Hlen Hcone + -- Should be able to prove cut > 0 + -- Here, do induction over eval_length instead + revert cut + induction hist using List.list_reverse_induction + · intro cut Hcone + -- rcases Hcone with ⟨ H1, H2 ⟩ + -- simp_all + unfold constancy_at + sorry + · rename_i vk hist' IH + intro cut Hcone + unfold constancy_at + + -- If I unfold one iterate from both sides, will I still be on the left edge? + sorry lemma cone_constancy [DPSystem ℕ] {ε₁ ε₂ : ℕ+} {τ : ℤ} {data : List ℕ} {v0 vk : ℤ} (cut : ℕ) (initial hist : List ℤ) : cone_of_possibility cut initial hist -> @@ -469,21 +375,6 @@ lemma cone_constancy [DPSystem ℕ] {ε₁ ε₂ : ℕ+} {τ : ℤ} {data : List · trivial --- Define cone --- Define constancy, parameterized by initial list length --- Prove zero for the region to the left of the cone (done before) --- Prove zero for the region below the cone (done before) --- Define base case in terms of cone --- Prove base case in terms of cone --- Prove inductive case for code - - - - - - - - lemma sv2_eq_sv3 [dps : DPSystem ℕ] ε₁ ε₂ l : sv2_privMax ε₁ ε₂ l = sv3_privMax ε₁ ε₂ l := by apply SLang.ext @@ -511,8 +402,44 @@ lemma sv2_eq_sv3 [dps : DPSystem ℕ] ε₁ ε₂ l : sv2_privMax ε₁ ε₂ l apply @tendsto_atTop_of_eventually_const _ _ _ _ _ _ _ (hist.length + 1) intro i H - sorry - + -- i is in the cone, reduce by induction + induction i + · -- Fake base case + simp at H + · rename_i i IH + -- Real base case + cases Classical.em (i = hist.length) + · simp_all + + -- Inductive case: use constancy + rw [<- IH ?G1] + case G1 => + apply LE.le.ge + apply GE.ge.le at H + apply LE.le.lt_or_eq at H + cases H + · apply Nat.le_of_lt_succ + trivial + · exfalso + rename_i Hcont _ + apply Hcont + linarith + have HK := @cone_constancy dps ε₁ ε₂ τ l v0 vk i [] hist + unfold constancy_at at HK + conv => + enter [1, 3] + rw [add_comm] + apply HK + unfold cone_of_possibility + simp + apply GE.ge.le at H + apply LE.le.lt_or_eq at H + cases H + · linarith + · exfalso + rename_i h _ + apply h + linarith /- From 3b30b10bab31be589c57caccdd3d0e570bdf32eb Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Tue, 1 Oct 2024 09:34:39 -0400 Subject: [PATCH 040/100] close cone_constancy --- .../Queries/UnboundedMax/Properties.lean | 61 ++++++++++++++----- 1 file changed, 47 insertions(+), 14 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index a6576016..476cfc9f 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -274,21 +274,54 @@ lemma cone_left_edge_constancy [DPSystem ℕ] {ε₁ ε₂ : ℕ+} {τ : ℤ} {d cone_of_possibility cut initial hist -> @constancy_at _ ε₁ ε₂ τ data v0 vk cut initial hist := by intro Hlen Hcone - -- Should be able to prove cut > 0 - -- Here, do induction over eval_length instead - revert cut - induction hist using List.list_reverse_induction - · intro cut Hcone - -- rcases Hcone with ⟨ H1, H2 ⟩ - -- simp_all - unfold constancy_at - sorry - · rename_i vk hist' IH - intro cut Hcone - unfold constancy_at + -- cut > 0 due to cone + cases cut + · exfalso + simp [cone_of_possibility] at Hcone + simp_all only [lt_self_iff_false, le_refl, and_true] + rename_i cut' + + -- Unfold the first iterate + unfold constancy_at + unfold probWhileCut + unfold probWhileFunctional + + cases (sv1_privMaxC τ data (initial, v0)) + · -- False case: both programs terminate with initial state + simp + · -- True case: both programs step to a point outside of the cone, so are zero + simp + apply tsum_congr + intro ⟨ initial', v1 ⟩ + + simp [sv1_privMaxF] + rw [← ENNReal.tsum_mul_right] + rw [← ENNReal.tsum_mul_right] + + -- Ignore the cases when hist' is not a legal step + cases Classical.em (¬ ∃ v', initial' = initial ++ [v']) + · rename_i h + apply tsum_congr + intro z + split + · exfalso + apply h + exists v0 + rename_i h' + cases h' + rfl + · simp + rename_i h + simp at h + rcases h with ⟨ _, Hv1' ⟩ + simp [Hv1'] + apply tsum_congr + intro _ + + -- Both branches are outside of the cone now + rw [external_to_cone_zero (by simp_all [cone_of_possibility])] + rw [external_to_cone_zero (by simp_all [cone_of_possibility])] - -- If I unfold one iterate from both sides, will I still be on the left edge? - sorry lemma cone_constancy [DPSystem ℕ] {ε₁ ε₂ : ℕ+} {τ : ℤ} {data : List ℕ} {v0 vk : ℤ} (cut : ℕ) (initial hist : List ℤ) : cone_of_possibility cut initial hist -> From f0305965c8585a2b45a3c3acfdb1e217b1cd081c Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Tue, 1 Oct 2024 17:16:55 -0400 Subject: [PATCH 041/100] checkpoint --- .../Queries/UnboundedMax/Properties.lean | 55 +++++++++++++++++-- 1 file changed, 51 insertions(+), 4 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 476cfc9f..6edc232f 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -178,10 +178,7 @@ lemma sv0_eq_sv1 [dps : DPSystem ℕ] ε₁ ε₂ l : sv0_privMax ε₁ ε₂ l trivial simp only [tsum_zero, add_zero] - -- RHS: sum over all lists of length "result" - -- Prove this by induction on result - - + -- RHS: sum over all lists of length "result"? -- rw [tsum_ite_eq] sorry @@ -509,6 +506,56 @@ def sv4_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang computation point lemma sv3_eq_sv4 [dps : DPSystem ℕ] ε₁ ε₂ l : sv3_privMax ε₁ ε₂ l = sv4_privMax ε₁ ε₂ l := by + apply SLang.ext + intro n + unfold sv3_privMax + unfold sv4_privMax + + -- Simplify the common start to the programs + simp + apply tsum_congr + intro τ + congr 1 + + -- Rewrite sv4_presample .. (n + 1) into a program which presamples n and 1 separately + + let sv4_presample_sep : SLang (List ℤ) := do + let v0 <- privNoiseZero ε₁ (4 * ε₂) + let L <- sv4_presample ε₁ ε₂ n + return (v0 :: L) + + have sv4_presample_sep_eq a : sv4_presample ε₁ ε₂ (n + 1) a = sv4_presample_sep a := by + dsimp only [sv4_presample_sep] + clear sv4_presample_sep + revert a + induction n + · simp [sv4_presample] + rename_i n IH + intro a + + -- Rewrite LHS to include LHS of IH + unfold sv4_presample + simp + conv => + enter [1, 1, b, 2, 1, c] + rw [IH] + clear IH + + -- Conclude by simplification + simp + apply tsum_congr + intro v0 + congr 1 + + sorry -- Rewriting through the ite is annoying but doable + + + conv => + enter [2, 1, a] + rw [sv4_presample_sep_eq] + dsimp [sv4_presample_sep] + clear sv4_presample_sep sv4_presample_sep_eq + sorry From f38b2ae300ebed40cd9f96cf33b955fed7c64f9c Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Sun, 6 Oct 2024 18:19:56 -0400 Subject: [PATCH 042/100] checkpoint --- .../Queries/UnboundedMax/Properties.lean | 540 ++++++++++++++++-- 1 file changed, 502 insertions(+), 38 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 6edc232f..e4aaba68 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -211,12 +211,15 @@ lemma sv1_eq_sv2 [dps : DPSystem ℕ] ε₁ ε₂ l : sv1_privMax ε₁ ε₂ l - Truncates the loop -/ +def sv3_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : SLang sv1_state := do + probWhileCut (sv1_privMaxC τ l) (@sv1_privMaxF dps ε₁ ε₂) (point + 1) init + def sv3_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => let computation : SLang ℕ := do let τ <- @privNoiseZero dps ε₁ (2 * ε₂) let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) - let sk <- probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (point + 1) ([], v0) + let sk <- @sv3_loop dps ε₁ ε₂ τ l point ([], v0) return (sv1_threshold sk) computation point @@ -412,6 +415,7 @@ lemma sv2_eq_sv3 [dps : DPSystem ℕ] ε₁ ε₂ l : sv2_privMax ε₁ ε₂ l intro point unfold sv2_privMax unfold sv3_privMax + unfold sv3_loop simp apply tsum_congr intro τ @@ -472,6 +476,373 @@ lemma sv2_eq_sv3 [dps : DPSystem ℕ] ε₁ ε₂ l : sv2_privMax ε₁ ε₂ l linarith + + +-- Commute out a single sample from the loop +lemma sv3_loop_unroll_1 [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l point L vk : + sv3_loop ε₁ ε₂ τ l (point + 1) (L, vk) = + (do + let vk1 <- @privNoiseZero dps ε₁ (4 * ε₂) + if (sv1_privMaxC τ l (L, vk)) + then (sv3_loop ε₁ ε₂ τ l point (L ++ [vk], vk1)) + else probPure (L, vk)) := by + conv => + lhs + unfold sv3_loop + simp [probWhileCut, probWhileFunctional] + split + · apply SLang.ext + intro ⟨ HF, vkf ⟩ + simp [probBind] + unfold sv3_loop + conv => + enter [1, 1, a, 1] + unfold sv1_privMaxF + simp + conv => + enter [1, 1, a] + rw [← ENNReal.tsum_mul_right] + rw [ENNReal.tsum_comm] + apply tsum_congr + intro a + simp + rw [ENNReal.tsum_eq_add_tsum_ite ?G1] + case G1 => apply (L ++ [vk], a) + split + · conv => + rhs + rw [<- add_zero (_ * _)] + congr 1 + simp + intro i HK1 HK2 + exfalso + apply HK1 + apply HK2 + · exfalso + rename_i hk + apply hk + rfl + · simp + apply SLang.ext + intro ⟨ HF, vkf ⟩ + simp [probBind] + split <;> try simp + unfold privNoiseZero + exact Eq.symm (PMF.tsum_coe (DPSystem.noise (fun _ => 0) 1 ε₁ (4 * ε₂) [])) + + + +-- Sample n new random samples into a list +def sv4_presample [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (n : ℕ) : SLang (List ℤ) := + match n with + | Nat.zero => return [] + | Nat.succ n' => do + let vk1 <- @privNoiseZero dps ε₁ (4 * ε₂) + let vks <- @sv4_presample dps ε₁ ε₂ n' + return vks ++ [vk1] -- FIXME: Possible issue: do I want to presample onto the start or end of this list? + +-- If there are 5 more samples in state than there are in intial state, we want to ignore the first 5 entries in presamples +def sv4_next_def (initial_state : sv1_state) (presamples : List ℤ) (state : sv1_state) : Option sv1_state := + sorry + +-- Given a list of random samples, and a state, which is assumed to have state drawn from +-- init ++ (some prefix of the random samples), with init specified here, update the state +-- to include the next sample on the tape. +-- If no such sample exists, probZero + + +-- If there are 5 more samples in state than there are in intial state, we want to ignore the first 5 entries in presamples +def sv4_next (initial_state : sv1_state) (presamples : List ℤ) (state : sv1_state) : SLang sv1_state := + match sv4_next_def initial_state presamples state with + | none => probZero + | some x => return x + + +lemma sv3_loop_unroll_1_alt [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l point (initial_state : sv1_state) : + sv3_loop ε₁ ε₂ τ l (point + 1) initial_state = + (do + let vk1 <- @privNoiseZero dps ε₁ (4 * ε₂) + if (sv1_privMaxC τ l initial_state) + then (sv3_loop ε₁ ε₂ τ l point (initial_state.1 ++ [initial_state.2], vk1)) + else probPure initial_state) := by + rcases initial_state with ⟨ _ , _ ⟩ + rw [sv3_loop_unroll_1] + +lemma sv3_loop_unroll_2 [dps : DPSystem ℕ] τ ε₁ ε₂ l point (initial_state : sv1_state) : + sv3_loop ε₁ ε₂ τ l (point + 1) initial_state = + (do + let presample_list <- @sv4_presample dps ε₁ ε₂ 1 + if (sv1_privMaxC τ l initial_state) + then do + let next_state <- sv4_next initial_state presample_list initial_state + sv3_loop ε₁ ε₂ τ l point next_state + else probPure initial_state) := by + rw [sv3_loop_unroll_1_alt, sv4_presample, sv4_presample ] + simp + apply SLang.ext + intro x + simp + apply tsum_congr + intro y + congr 1 + split <;> simp + + -- This is doable for sure, I just need to define sv4_next + sorry + + + + + + + + + + +/- + + + + + + + +-- 0th check: using state ([], v0) + + +def sv4_nth (s : sv1_state) (n : ℕ) : sv1_state := + sorry + +def init_iter_start (s : sv1_state) : ℕ := List.length s.1 + + +def state_to_list (s : sv1_state) : List ℤ := sorry + +def sv4_privMaxC (τ : ℤ) (l : List ℕ) (st : sv1_state) : Bool := + sv1_privMaxC τ l st + + +-- Assume that st is a prefix state of presamples? If we can't add on a state such that +-- the returned state is still a prefix state, then probZero. +def sv4_privMaxF (presamples : List ℤ) (st : sv1_state) : SLang sv1_state := + -- Get the next state by getting the length of the current state + -- Increment it + -- Take that prefix of presamples + -- Also something about how we need to fail if we are out of presamples? Or maybe just probZero? + sorry + -- sv1_privMaxC τ l (sv4_nth presamples n) + + +-- Compute the loop, but using presamples instead +def sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : SLang sv1_state := do + let presample_st <- @sv4_presample dps ε₁ ε₂ init point -- FIXME presamples need to be based on point, may be +- 1 + let presamples := (state_to_list presample_st) + probWhileCut + (sv4_privMaxC τ l) + (sv4_privMaxF presamples) + (point + 1) + init + + +-- No + + + +-- lemma sv4_loop_unroll_1 [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l point L vk : + + + +-/ + + +/- + +def sv1_to_list (s : sv1_state) : List ℤ := s.1 ++ [s.2] + +def list_to_sv1 (s : List ℤ) : sv1_state := + match s with + | [] => ([], 0) -- Error case + | [x] => ([], x) + | (x :: xs) => + let w := list_to_sv1 xs + (x :: w.1, w.2) + + +def sv1_prefix (s : sv1_state) (n : ℕ) := list_to_sv1 $ List.take n $ sv1_to_list s + + +def sv4_0_cond (presamples : sv1_state) (τ : ℤ) (l : List ℕ) (n : ℕ) : Bool := + sv1_privMaxC τ l (sv1_prefix presamples n) + +def sv4_0_body (n : ℕ) : SLang ℕ := + return (n + 1) + + +def sv4_0_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (samples : ℕ) (init : sv1_state) : SLang ℕ := do + let presamples <- @sv4_0_presample dps ε₁ ε₂ init samples + let sk <- probWhileCut + (sv4_0_cond presamples τ l) + sv4_0_body + (samples + 1) + (List.length (init.1)) + return sk + 1 + + +def sv3_loop_mod [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (samples : ℕ) (init : sv1_state) : SLang ℕ := do + let st <- probWhileCut (@sv1_privMaxC τ l) (@sv1_privMaxF dps ε₁ ε₂) samples init + return sv1_threshold st + + +-- Turn the state back into a nat, but this time, over the presampled list +-- FIXME: Rename if this actually works +def sv4_0_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := + fun (point : ℕ) => + let computation : SLang ℕ := do + let τ <- @privNoiseZero dps ε₁ (2 * ε₂) + let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + @sv4_0_loop dps ε₁ ε₂ τ l point ([], v0) + computation point + + + +-- Now try to prove this equivalence by induction for all inits + +lemma sv3_eq_sv4_0_loop [dps : DPSystem ℕ] ε₁ ε₂ τ l point init : + @sv3_loop_mod dps τ ε₁ ε₂ l (point + 1) init = @sv4_0_loop dps τ ε₁ ε₂ l point init := by + revert init + induction point + · intro init + apply SLang.ext + intro p' + simp + + -- Simplify the RHS: zero presampling occurs + simp [sv4_0_loop] + simp [sv4_0_presample] + rw [ENNReal.tsum_eq_add_tsum_ite init] + simp + conv => + lhs + rw [<- zero_add (sv3_loop_mod _ _ _ _ _ _ _)] + conv => + rhs + rw [add_comm] + congr 1 + · symm + simp + intro _ HK1 HK2 + exfalso + apply HK1 + apply HK2 + + -- Simplify the LHS + simp [sv3_loop_mod] -- ? + simp [probWhileCut, probWhileFunctional] + simp [sv4_0_cond] + simp [sv1_prefix, sv1_to_list, list_to_sv1] + rw [List.take_append] + simp [list_to_sv1] + have H : (list_to_sv1 (init.1 ++ [init.2])) = init := by + sorry + simp [H] + simp [sv4_0_body] + split <;> simp + simp [sv1_threshold] + rw [ENNReal.tsum_eq_add_tsum_ite init] + simp + + sorry + · sorry + + + + + +lemma sv3_eq_sv4_0 [dps : DPSystem ℕ] ε₁ ε₂ l : sv3_privMax ε₁ ε₂ l = sv4_0_privMax ε₁ ε₂ l := by + sorry + + + + + +-- -- Probably useless as stated but should be true? +-- -- For the inductive argument we only need n2 = 1, btw +-- lemma sv3_0_presample_additive [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (x0 : sv1_state) (n1 n2 : ℕ) : +-- ((@sv4_0_presample dps ε₁ ε₂ x0 n1) >>= fun x1 => @sv4_0_presample dps ε₁ ε₂ x1 n2) = +-- @sv4_0_presample dps ε₁ ε₂ x0 (n1 + n2) := by +-- revert x0 +-- induction n2 +-- · intro x0 +-- apply SLang.ext +-- intro v +-- simp [sv4_0_presample] +-- rw [ENNReal.tsum_eq_add_tsum_ite ?G1] +-- case G1 => apply v +-- simp +-- conv => +-- rhs +-- rw [<- add_zero (sv4_0_presample _ _ _ _ _)] +-- congr 1 +-- simp +-- intro i H1 H2 +-- exfalso +-- apply H1 +-- symm +-- apply H2 +-- · intro x0 +-- rename_i n' IH +-- apply SLang.ext +-- intro v +-- simp +-- -- IDK. Something is doable for sure? +-- sorry + + + + + +-- def sv3_1_threshold (s : sv3_1_state) : ℕ := List.length s.1 +-- +-- def sv3_1_noise (s : sv3_1_state) : ℤ := s.2.1 +-- +-- def sv3_1_count (s : sv3_1_state) : ℕ := s.2.2 +-- +-- def sv3_1_privMaxC (τ : ℤ) (l : List ℕ) (s : sv3_1_state) : Bool := +-- decide (exactDiffSum (sv3_1_threshold s) l + (sv3_1_noise s) < τ) +-- +-- def sv3_1_privMaxF [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (s : sv3_1_state) : SLang sv3_1_state := do +-- let vn <- @privNoiseZero dps ε₁ (4 * ε₂) +-- return (s.1 ++ [s.2.1], vn, s.2.2 + 1) +-- +-- def sv3_1_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := +-- fun (point : ℕ) => +-- let computation : SLang ℕ := do +-- let τ <- @privNoiseZero dps ε₁ (2 * ε₂) +-- let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) +-- let sk <- probWhileCut (sv3_1_privMaxC τ l) (sv3_1_privMaxF ε₁ ε₂) (point + 1) ([], v0, 0) +-- return (sv3_1_count sk) +-- computation point +-- +-- +-- +-- lemma sv3_eq_sv3_1 [dps : DPSystem ℕ] ε₁ ε₂ l : sv3_privMax ε₁ ε₂ l = sv3_1_privMax ε₁ ε₂ l := by +-- apply SLang.ext +-- intro point +-- unfold sv3_privMax +-- unfold sv3_1_privMax +-- simp +-- apply tsum_congr +-- intro τ +-- congr 1 +-- apply tsum_congr +-- intro v0 +-- congr 1 +-- +-- sorry + + + + /- ## Program version 4 - Presamples each vk @@ -498,65 +869,157 @@ def sv4_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang fun (point : ℕ) => let computation : SLang ℕ := do let τ <- @privNoiseZero dps ε₁ (2 * ε₂) - let history <- sv4_presample ε₁ ε₂ (point + 1) + let history <- sv4_presample ε₁ ε₂ point let vk <- @privNoiseZero dps ε₁ (4 * ε₂) - if (sv4_evaluate_history history τ l && (! sv1_privMaxC τ l (history, vk))) + if (sv4_evaluate_history history τ l && (! (sv4_evaluate_history (history ++ [vk]) τ l))) then return point else probZero computation point lemma sv3_eq_sv4 [dps : DPSystem ℕ] ε₁ ε₂ l : sv3_privMax ε₁ ε₂ l = sv4_privMax ε₁ ε₂ l := by apply SLang.ext - intro n + intro point unfold sv3_privMax unfold sv4_privMax - -- Simplify the common start to the programs simp apply tsum_congr intro τ congr 1 - -- Rewrite sv4_presample .. (n + 1) into a program which presamples n and 1 separately - - let sv4_presample_sep : SLang (List ℤ) := do - let v0 <- privNoiseZero ε₁ (4 * ε₂) - let L <- sv4_presample ε₁ ε₂ n - return (v0 :: L) - - have sv4_presample_sep_eq a : sv4_presample ε₁ ε₂ (n + 1) a = sv4_presample_sep a := by - dsimp only [sv4_presample_sep] - clear sv4_presample_sep - revert a - induction n - · simp [sv4_presample] - rename_i n IH - intro a + conv => + enter [1, 1, v0, 2, 1, st] + conv => + enter [2, 1, hist, 2, 1, vk] - -- Rewrite LHS to include LHS of IH - unfold sv4_presample - simp + -- Maybe don't keep this? + cases point + · -- point is 0 + simp [sv4_presample, sv1_threshold] + unfold sv1_state conv => - enter [1, 1, b, 2, 1, c] - rw [IH] - clear IH - - -- Conclude by simplification - simp + enter [1, 1, v0] + rw [<- ENNReal.tsum_mul_left] + conv => + rw [ENNReal.tsum_comm] + rw [ENNReal.tsum_prod'] apply tsum_congr - intro v0 - congr 1 + intro hist + simp + sorry - sorry -- Rewriting through the ite is annoying but doable + sorry - conv => - enter [2, 1, a] - rw [sv4_presample_sep_eq] - dsimp [sv4_presample_sep] - clear sv4_presample_sep sv4_presample_sep_eq - sorry + -- -- Simplify the common start to the programs + + -- Rewrite sv4_presample .. (n + 1) into a program which presamples n and 1 separately + + + -- let sv4_presample_sep : SLang (List ℤ) := do + -- let v0 <- privNoiseZero ε₁ (4 * ε₂) + -- let L <- sv4_presample ε₁ ε₂ n + -- return (v0 :: L) + + -- have sv4_presample_sep_eq a : sv4_presample ε₁ ε₂ (n + 1) a = sv4_presample_sep a := by + -- dsimp only [sv4_presample_sep] + -- clear sv4_presample_sep + -- revert a + -- induction n + -- · simp [sv4_presample] + -- rename_i n IH + -- intro a + + -- -- Rewrite LHS to include LHS of IH + -- unfold sv4_presample + -- simp + -- conv => + -- enter [1, 1, b, 2, 1, c] + -- rw [IH] + -- clear IH + + -- -- Conclude by simplification + -- simp + -- apply tsum_congr + -- intro v0 + -- congr 1 + + -- sorry -- Rewriting through the ite is annoying but doable + -- conv => + -- enter [2, 1, a] + -- rw [sv4_presample_sep_eq] + -- dsimp [sv4_presample_sep] + -- clear sv4_presample_sep sv4_presample_sep_eq + + + + -- conv => + -- enter [1, 1, a] + -- rw [<- ENNReal.tsum_mul_left] + -- conv => + -- enter [2, 1, a] + -- rw [<- ENNReal.tsum_mul_left] + -- enter [1, i] + -- rw [<- ENNReal.tsum_mul_right] + -- conv => + -- enter [2] + -- rw [ENNReal.tsum_comm] + -- apply tsum_congr + -- intro v0 + -- unfold sv1_state + -- rw [ENNReal.tsum_prod'] + -- apply tsum_congr + -- intro hist + -- apply tsum_congr + -- intro vk + + -- conv => + -- enter [2] + -- rw [mul_comm] + -- repeat rw [mul_assoc] + -- congr 1 + -- simp [sv1_threshold] + + -- split + -- · + -- sorry + -- · -- Can I solve this one without induction? + -- -- It simplifies the induction on the other side + -- simp + -- split + -- · right + -- right + -- intro l Hl + -- induction n + -- · simp_all + -- unfold sv4_presample + -- simp + -- intro H1 + -- simp_all + -- sorry + -- · sorry + -- · sorry + + -- revert hist vk + -- induction n + -- · simp_all + -- sorry + -- · rename_i n IH + -- intro hist vk + + -- have Hcond : hist = [] ∨ ∃ vl hist', hist = hist' ++ [vl] := by sorry + -- cases Hcond + -- · simp_all + -- rename_i h + -- rcases h with ⟨ hist', ⟨ vk', H ⟩ ⟩ + -- simp [H] + -- clear H + -- conv => + -- enter [1] + -- unfold probWhileCut + -- unfold probWhileFunctional + @@ -585,3 +1048,4 @@ def sv5_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang lemma sv4_eq_sv5 [dps : DPSystem ℕ] ε₁ ε₂ l : sv4_privMax ε₁ ε₂ l = sv5_privMax ε₁ ε₂ l := by sorry +-/ From 7a30a61c36e283735d7f2851c40c07b38dcab394 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Sun, 6 Oct 2024 22:24:35 -0400 Subject: [PATCH 043/100] so close yet so far --- .../Queries/UnboundedMax/Properties.lean | 230 ++++++++++++++---- 1 file changed, 186 insertions(+), 44 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index e4aaba68..1c100e51 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -541,16 +541,28 @@ def sv4_presample [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (n : ℕ) : SLang (L let vks <- @sv4_presample dps ε₁ ε₂ n' return vks ++ [vk1] -- FIXME: Possible issue: do I want to presample onto the start or end of this list? --- If there are 5 more samples in state than there are in intial state, we want to ignore the first 5 entries in presamples -def sv4_next_def (initial_state : sv1_state) (presamples : List ℤ) (state : sv1_state) : Option sv1_state := - sorry +-- How much excess state does s have above i +def state_len_excess (i : sv1_state) (s : sv1_state) : Option ℕ := + Int.toNat' ((List.length s.1 : Int) - (List.length i.1) : Int) + +-- Return the next sample from presamples, ignoring the first n elements +def sv4_next_pre (presamples : List ℤ) (n : ℕ) : Option ℤ := + match presamples with + | [] => none + | (x :: xs) => + match n with + | 0 => some x + | (Nat.succ n') => sv4_next_pre xs n' --- Given a list of random samples, and a state, which is assumed to have state drawn from --- init ++ (some prefix of the random samples), with init specified here, update the state --- to include the next sample on the tape. --- If no such sample exists, probZero +def sv4_state_shift (s : sv1_state) (z : ℤ) : sv1_state := (s.1 ++ [s.2], z) +-- If there are 5 more samples in state than there are in intial state, we want to ignore the first 5 entries in presamples +def sv4_next_def (initial_state : sv1_state) (presamples : List ℤ) (state : sv1_state) : Option sv1_state := do + let excess <- state_len_excess initial_state state + let next_sample <- sv4_next_pre presamples excess + return (sv4_state_shift state next_sample) +-- Intuition: -- If there are 5 more samples in state than there are in intial state, we want to ignore the first 5 entries in presamples def sv4_next (initial_state : sv1_state) (presamples : List ℤ) (state : sv1_state) : SLang sv1_state := match sv4_next_def initial_state presamples state with @@ -558,6 +570,20 @@ def sv4_next (initial_state : sv1_state) (presamples : List ℤ) (state : sv1_st | some x => return x +lemma sv4_next_self (st : sv1_state) (y : ℤ) (ys : List ℤ) : + sv4_next st (y :: ys) st = probPure (sv4_state_shift st y) := by + simp [sv4_next, sv4_next_def, state_len_excess, Int.toNat', sv4_next_pre] + + + +lemma sv4_next_self_emp (st : sv1_state) : + sv4_next st [] st = probZero := by + simp [sv4_next, sv4_next_def, state_len_excess, Int.toNat', sv4_next_pre] + + + + + lemma sv3_loop_unroll_1_alt [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l point (initial_state : sv1_state) : sv3_loop ε₁ ε₂ τ l (point + 1) initial_state = (do @@ -587,72 +613,188 @@ lemma sv3_loop_unroll_2 [dps : DPSystem ℕ] τ ε₁ ε₂ l point (initial_sta congr 1 split <;> simp - -- This is doable for sure, I just need to define sv4_next - sorry - - - - - - + simp [sv4_next_self] + rw [ENNReal.tsum_eq_add_tsum_ite (sv4_state_shift initial_state y)] + simp + conv => + lhs + rw [<- add_zero (sv3_loop _ _ _ _ _ _ _)] + congr 1 + symm + simp + intro i H1 H2 + exfalso + apply H1 + apply H2 +def sv4_privMaxC (τ : ℤ) (l : List ℕ) (st : sv1_state) : Bool := sv1_privMaxC τ l st -/- +-- Assume that st is a prefix state of presamples? If we can't add on a state such that +-- the returned state is still a prefix state, then probZero. +def sv4_privMaxF (initial : sv1_state) (presamples : List ℤ) (st : sv1_state) : SLang sv1_state := + sv4_next initial presamples st +-- Compute the loop, but using presamples instead +def sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : SLang sv1_state := do + let presamples <- @sv4_presample dps ε₁ ε₂ point -- (point + 1) breaks the base case... + probWhileCut + (sv4_privMaxC τ l) + (sv4_privMaxF init presamples) + (point + 1) + init +-- I want to prove that sv4_loop is equal to sv3_loop +lemma sv3_loop_eq_sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : + @sv3_loop dps ε₁ ε₂ τ l point init = @sv4_loop dps ε₁ ε₂ τ l point init := by + unfold sv3_loop + unfold sv4_loop + revert init + induction point + · intro init + simp [sv4_presample] + unfold sv4_privMaxC + simp [probWhileCut, probWhileFunctional] + split <;> try rfl + -- Does the left side have more samples than the right side? + -- Not a good sign! + apply SLang.ext + intro _ + simp + · intro init + rename_i point' IH + -- Unroll it once, using the lemma + let H := sv3_loop_unroll_2 τ ε₁ ε₂ l point' init + unfold sv3_loop at H + rw [H] + clear H --- 0th check: using state ([], v0) + cases (Classical.em (sv1_privMaxC τ l init = true)) + · simp + rename_i Hcond + + -- Horrifying, but only because I can't conv under the if (dependent types)? + have X : + ((sv4_presample ε₁ ε₂ 1).probBind fun presample_list => + if sv1_privMaxC τ l init = true then + (sv4_next init presample_list init).probBind fun next_state => + probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (point' + 1) next_state + else probPure init) = + ((sv4_presample ε₁ ε₂ 1).probBind fun presample_list => + if sv1_privMaxC τ l init = true then + (sv4_next init presample_list init).probBind fun next_state => + ((sv4_presample ε₁ ε₂ point') >>= + (fun presamples => + probWhileCut (sv4_privMaxC τ l) (sv4_privMaxF next_state presamples) (point' + 1) next_state)) + else probPure init) := by + simp_all + rw [X] + clear X + clear IH + simp_all + + -- Seems like a lost cause but I'll continue trying + apply SLang.ext + intro final_state + simp + -- No [] on the LHS list because ther sv4_next term would be 0 + -- But also, not [] on the LHS list because otherwise (sv4_presample .. 1 a) = 0 + + -- Maybe just try to commute the presample steps together on the LHS? + conv => + enter [1, 1, a, 2, 1, a1] + rw [<- ENNReal.tsum_mul_left] + conv => + enter [1, 1, a] + rw [ENNReal.tsum_comm] + rw [<- ENNReal.tsum_mul_left] + conv => + enter [1, 1, a, 1, i, 2, 1, a1] + rw [mul_comm] + rw [mul_assoc] + conv => + enter [1, 1, a, 1, i] + rw [ENNReal.tsum_mul_left] + rw [<- mul_assoc] + + -- LHS only has support when + -- the outermost list has length 1 + -- next outermost list has length point' + -- RHS only has support when + -- outermost list has length point' + 1 + -- The sv4_presample can cancel out using a bijection + + -- Seems sketchy + -- The LHS will have a sum over initial states i_1 + -- + -- the RHS fixes it to initial + -- Unfold a step on the right??? + + have Hcond' : sv4_privMaxC τ l init = true := by + simp [sv4_privMaxC] + trivial + conv => + enter [2, 1, a] + unfold probWhileCut + unfold probWhileFunctional + simp_all + conv => + enter [2, 1, a, 2, 1, i1] + simp only [sv4_privMaxF] + rw [mul_comm] + -- This seems super bad -def sv4_nth (s : sv1_state) (n : ℕ) : sv1_state := - sorry + -- That whole scheme with "initial" is what's going wrong. + -- This needs "initial" as a separate paramater basically. -def init_iter_start (s : sv1_state) : ℕ := List.length s.1 + sorry + · simp_all + simp [probWhileCut, probWhileFunctional] + have X : sv4_privMaxC τ l init = false := by + unfold sv4_privMaxC + trivial + simp_all + apply SLang.ext + intro final_state + -- Both 1 because sum of PMF? + simp + split + · sorry + · simp -def state_to_list (s : sv1_state) : List ℤ := sorry -def sv4_privMaxC (τ : ℤ) (l : List ℕ) (st : sv1_state) : Bool := - sv1_privMaxC τ l st + -- apply SLang.ext + -- intro final_state + -- simp + -- + -- -- Step through to get a new init? + -- conv => + -- -- I want to enter into theat probWhileCut + -- enter [1, 1, a, 2] + -- simp + -- skip + -- simp --- Assume that st is a prefix state of presamples? If we can't add on a state such that --- the returned state is still a prefix state, then probZero. -def sv4_privMaxF (presamples : List ℤ) (st : sv1_state) : SLang sv1_state := - -- Get the next state by getting the length of the current state - -- Increment it - -- Take that prefix of presamples - -- Also something about how we need to fail if we are out of presamples? Or maybe just probZero? - sorry - -- sv1_privMaxC τ l (sv4_nth presamples n) + -- -- Then use the IH --- Compute the loop, but using presamples instead -def sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : SLang sv1_state := do - let presample_st <- @sv4_presample dps ε₁ ε₂ init point -- FIXME presamples need to be based on point, may be +- 1 - let presamples := (state_to_list presample_st) - probWhileCut - (sv4_privMaxC τ l) - (sv4_privMaxF presamples) - (point + 1) - init + -- sorry --- No --- lemma sv4_loop_unroll_1 [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l point L vk : --/ /- From b94fd5bad68c2e75b4a26cecf66de10d51b88df0 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Sun, 6 Oct 2024 23:00:40 -0400 Subject: [PATCH 044/100] closer (not really) --- .../Queries/UnboundedMax/Properties.lean | 46 +++++++++---------- 1 file changed, 21 insertions(+), 25 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 1c100e51..291aec82 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -653,9 +653,17 @@ lemma sv3_loop_eq_sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) @sv3_loop dps ε₁ ε₂ τ l point init = @sv4_loop dps ε₁ ε₂ τ l point init := by unfold sv3_loop unfold sv4_loop - revert init + + + suffices ∀ state', (probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (point + 1) state' = + (do + let presamples ← sv4_presample ε₁ ε₂ point + probWhileCut (sv4_privMaxC τ l) (sv4_privMaxF init presamples) (point + 1) state')) + by apply this + induction point · intro init + rename_i init_start simp [sv4_presample] unfold sv4_privMaxC simp [probWhileCut, probWhileFunctional] @@ -669,14 +677,15 @@ lemma sv3_loop_eq_sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) rename_i point' IH -- Unroll it once, using the lemma - let H := sv3_loop_unroll_2 τ ε₁ ε₂ l point' init + -- Do we need to unroll this beforehand? + let H := sv3_loop_unroll_2 τ ε₁ ε₂ l point' unfold sv3_loop at H rw [H] clear H cases (Classical.em (sv1_privMaxC τ l init = true)) · simp - rename_i Hcond + rename_i init_start Hcond -- Horrifying, but only because I can't conv under the if (dependent types)? have X : @@ -690,23 +699,20 @@ lemma sv3_loop_eq_sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (sv4_next init presample_list init).probBind fun next_state => ((sv4_presample ε₁ ε₂ point') >>= (fun presamples => - probWhileCut (sv4_privMaxC τ l) (sv4_privMaxF next_state presamples) (point' + 1) next_state)) + probWhileCut (sv4_privMaxC τ l) (sv4_privMaxF init_start presamples) (point' + 1) next_state)) else probPure init) := by simp_all rw [X] + clear X clear IH simp_all - -- Seems like a lost cause but I'll continue trying apply SLang.ext intro final_state simp - -- No [] on the LHS list because ther sv4_next term would be 0 - -- But also, not [] on the LHS list because otherwise (sv4_presample .. 1 a) = 0 - - -- Maybe just try to commute the presample steps together on the LHS? + -- Commute the presample steps together on the LHS? conv => enter [1, 1, a, 2, 1, a1] rw [<- ENNReal.tsum_mul_left] @@ -723,19 +729,6 @@ lemma sv3_loop_eq_sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) rw [ENNReal.tsum_mul_left] rw [<- mul_assoc] - -- LHS only has support when - -- the outermost list has length 1 - -- next outermost list has length point' - -- RHS only has support when - -- outermost list has length point' + 1 - -- The sv4_presample can cancel out using a bijection - - -- Seems sketchy - -- The LHS will have a sum over initial states i_1 - -- - -- the RHS fixes it to initial - -- Unfold a step on the right??? - have Hcond' : sv4_privMaxC τ l init = true := by simp [sv4_privMaxC] trivial @@ -748,11 +741,14 @@ lemma sv3_loop_eq_sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) enter [2, 1, a, 2, 1, i1] simp only [sv4_privMaxF] rw [mul_comm] - -- This seems super bad - -- That whole scheme with "initial" is what's going wrong. - -- This needs "initial" as a separate paramater basically. + -- Now just need to finish this by a bijection? + -- The RHS is zero on empty lists + -- Make general support lemma for sv4_presample? + + -- still a difference the sv4_next term + rw [<- ENNReal.tsum_prod] sorry · simp_all From 7e33b89345484acb8e20e47fa60685346b1e1e62 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 7 Oct 2024 12:05:27 -0400 Subject: [PATCH 045/100] explicit presampling is promising --- .../Queries/UnboundedMax/Properties.lean | 443 +++++++++++++++--- 1 file changed, 380 insertions(+), 63 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 291aec82..b54e2cae 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -11,6 +11,35 @@ open Classical namespace SLang + +/-- +Stronger congruence rule for probBind: The bound-to functions have to be equal only on the support of +the bound-from function. +-/ +lemma probBind_congr_strong (p : SLang T) (f : T -> SLang U) (g : T -> SLang U) (Hcong : ∀ t : T, p t ≠ 0 -> f t = g t) : + p >>= f = p >>= g := by + simp + unfold probBind + apply SLang.ext + intro u + apply Equiv.tsum_eq_tsum_of_support ?G1 + case G1 => + apply Set.BijOn.equiv (fun x => x) + simp [Function.support] + have Heq : {x | ¬p x = 0 ∧ ¬f x u = 0} = {x | ¬p x = 0 ∧ ¬g x u = 0} := by + apply Set.sep_ext_iff.mpr + intro t Ht + rw [Hcong] + apply Ht + rw [Heq] + apply Set.bijOn_id + simp [Function.support] + intro t ⟨ Hp, _ ⟩ + simp [Set.BijOn.equiv] + rw [Hcong] + apply Hp + + /- ## Helper programs -/ @@ -478,6 +507,61 @@ lemma sv2_eq_sv3 [dps : DPSystem ℕ] ε₁ ε₂ l : sv2_privMax ε₁ ε₂ l +/- +def sv4_state : Type := (n : ℕ) × { l : List ℤ // List.length l = n ∧ n > 0} + +def sv4_threshold (s : sv4_state) : ℕ := s.1 + +-- Last noise value, because length > 0 +def sv4_last (s : sv4_state) : ℤ := sorry + +-- Add one more noise value onto the end +def sv4_append (s : sv4_state) (z : ℤ) : sv4_state := sorry + +def sv4_initial (z : ℤ) : sv4_state := sorry + + +-- TODO? Bijection between sv4_states and sv1_states? + +def sv4_privMaxC (τ : ℤ) (l : List ℕ) (s : sv4_state) : Bool := + decide (exactDiffSum (sv4_threshold s) l + (sv4_last s) < τ) + + +def sv4_privMaxF [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (s : sv4_state) : SLang sv4_state := do + let vn <- @privNoiseZero dps ε₁ (4 * ε₂) + return (sv4_append s vn) + +def sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : SLang sv4_state := do + probWhileCut (sv4_privMaxC τ l) (@sv4_privMaxF dps ε₁ ε₂) (point + 1) init + + +def sv4_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := + fun (point : ℕ) => + let computation : SLang ℕ := do + let τ <- @privNoiseZero dps ε₁ (2 * ε₂) + let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + let sk <- @sv4_loop dps ε₁ ε₂ τ l point (sv4_initial v0) + return (sv4_threshold sk) + computation point + + +-- sv3 to sv4 argument : there is a bijection between sv3 states and sv4 states +-- Prove this is sv4 to sv5 ends up being provable + + + + +def sv5_presample [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (n : ℕ) : SLang { l : List ℤ // List.length l = n } := + match n with + | Nat.zero => return ⟨ [], by simp ⟩ + | Nat.succ n' => do + let vk1 <- @privNoiseZero dps ε₁ (4 * ε₂) + let vks <- @sv5_presample dps ε₁ ε₂ n' + return ⟨ vks ++ [vk1], by simp ⟩ +-/ + + + -- Commute out a single sample from the loop lemma sv3_loop_unroll_1 [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l point L vk : sv3_loop ε₁ ε₂ τ l (point + 1) (L, vk) = @@ -532,14 +616,244 @@ lemma sv3_loop_unroll_1 [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l p --- Sample n new random samples into a list -def sv4_presample [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (n : ℕ) : SLang (List ℤ) := + +-- New attempt: Let's just actually use consumable tape +-- We will not be able to prove the equality between loops, because they have different types +-- We can still unroll sv3, inductively +-- Can we get equality between loops back, if the tape loop projects out to the non-tape state? +-- Can we unroll tape inductively? + + +-- An sv4 state is an sv1 state plus a list of presamples +def sv4_state : Type := sv1_state × List ℤ + +def sv4_presample [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (n : ℕ) : SLang { l : List ℤ // List.length l = n } := match n with - | Nat.zero => return [] + | Nat.zero => return ⟨ [], by simp ⟩ | Nat.succ n' => do let vk1 <- @privNoiseZero dps ε₁ (4 * ε₂) let vks <- @sv4_presample dps ε₁ ε₂ n' - return vks ++ [vk1] -- FIXME: Possible issue: do I want to presample onto the start or end of this list? + return ⟨ vks ++ [vk1], by simp ⟩ + +def sv4_privMaxF (s : sv4_state) : SLang sv4_state := + match s.2 with + | [] => probZero + | (p :: ps) => return ((s.1.1 ++ [s.1.2], p), ps) + + +def sv4_privMaxC (τ : ℤ) (l : List ℕ) (st : sv4_state) : Bool := sv1_privMaxC τ l st.1 + +def sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : SLang sv1_state := do + let presamples <- @sv4_presample dps ε₁ ε₂ point + let v <- probWhileCut (sv4_privMaxC τ l) sv4_privMaxF (point + 1) (init, presamples) + return v.1 + + + + +lemma sv3_loop_unroll_1_alt [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l point (initial_state : sv1_state) : + sv3_loop ε₁ ε₂ τ l (point + 1) initial_state = + (do + let vk1 <- @privNoiseZero dps ε₁ (4 * ε₂) + if (sv1_privMaxC τ l initial_state) + then (sv3_loop ε₁ ε₂ τ l point (initial_state.1 ++ [initial_state.2], vk1)) + else probPure initial_state) := by + rcases initial_state with ⟨ _ , _ ⟩ + rw [sv3_loop_unroll_1] + +def len_list_append_rev {m n : ℕ} (x : { l : List ℤ // l.length = m }) (y: { l : List ℤ // l.length = n }) : { l : List ℤ // l.length = n + m } := + ⟨ x.1 ++ y.1 , by simp [add_comm] ⟩ + +def sv4_presample_split [DPSystem ℕ] (ε₁ ε₂ : ℕ+) (point : ℕ) : + sv4_presample ε₁ ε₂ (point + 1) = + (do + let presample_1 <- sv4_presample ε₁ ε₂ 1 + let presample_r <- sv4_presample ε₁ ε₂ point + return len_list_append_rev presample_1 presample_r) := by + apply SLang.ext + intro final_state + simp [sv4_presample] + sorry + + +def len_1_list_to_val (x : { l : List ℤ // l.length = 1 }) : ℤ := + let ⟨ xl, _ ⟩ := x + match xl with + | (v :: _) => v + + +-- When we do induction on point, +-- We will want to generalize over all init +-- Unfolding this loop just moves the first presample into init +-- Which can apply the IH-- since it's some arbitrary init state and a presamples state generated by one fewer point + + +def sv3_sv4_loop_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : + @sv3_loop dps ε₁ ε₂ τ l point init = @sv4_loop dps ε₁ ε₂ τ l point init := by + revert init + induction point + · -- It's a mess but it works + intro init + simp [sv3_loop, sv4_loop, probWhileCut, probWhileFunctional, sv4_presample, sv4_privMaxC] + split + · simp [sv4_privMaxF, sv1_privMaxF] + apply SLang.ext + intro final_state + simp + · apply SLang.ext + intro final_state + simp + split + · rename_i h + simp_all + symm + rw [condition_to_subset] + rw [ENNReal.tsum_eq_add_tsum_ite ⟨(init, []), rfl⟩] + simp_all + conv => + rhs + rw [<- add_zero 1] + congr + simp + intro a + rcases a + simp_all + · symm + simp + intro i H + rcases i + intro HK + rename_i hk2 _ _ + apply hk2 + cases HK + simp at H + trivial + · -- Inductive case + intro init + rename_i point IH + + -- Unfold sv3_loop on the left + rw [sv3_loop_unroll_1_alt] + + -- Apply the IH on the left + -- Doing it this way because I can't conv under the @ite? + let ApplyIH : + (do + let vk1 ← privNoiseZero ε₁ (4 * ε₂) + if sv1_privMaxC τ l init = true + then sv3_loop ε₁ ε₂ τ l point (init.1 ++ [init.2], vk1) + else probPure init) = + (do + let vk1 ← privNoiseZero ε₁ (4 * ε₂) + if sv1_privMaxC τ l init = true + then sv4_loop ε₁ ε₂ τ l point (init.1 ++ [init.2], vk1) + else probPure init) := by + simp + apply SLang.ext + intro final_state + simp + apply tsum_congr + intro _ + congr + split + · rw [IH] + · rfl + rw [ApplyIH] + clear ApplyIH IH + + have ToPresample : + (do + let vk1 ← privNoiseZero ε₁ (4 * ε₂) + if sv1_privMaxC τ l init = true then sv4_loop ε₁ ε₂ τ l point (init.1 ++ [init.2], vk1) else probPure init) = + (do + let vps ← sv4_presample ε₁ ε₂ 1 + let vk1 := len_1_list_to_val vps + if sv1_privMaxC τ l init = true then sv4_loop ε₁ ε₂ τ l point (init.1 ++ [init.2], vk1) else probPure init) := by + apply SLang.ext + intro final_state + simp + -- There is a bijection here + sorry + rw [ToPresample] + clear ToPresample + + -- Now, just need to prove this unfolding of sv4_loop + unfold sv4_loop + conv => + enter [2] + unfold probWhileCut + unfold probWhileFunctional + unfold sv4_privMaxC + + split + · conv => + enter [2] + rw [sv4_presample_split] + simp + + apply SLang.ext + intro final_state + simp + apply tsum_congr + intro vsample_1 + congr 1 + apply tsum_congr + intro vsample_rest + congr 1 + -- Seems that the RHS inner sum is an indicator on final_state, so I should + -- commute that out front + conv => + enter [2, 1, a] + rw [<- ENNReal.tsum_mul_left] + conv => + enter [2] + rw [ENNReal.tsum_comm] + apply tsum_congr + intro ⟨ ns_state, ns_presamples ⟩ -- Not sure what the variable ns represents? + simp + split <;> try simp + rename_i HF + + -- Investigate the RHS term for simplifications? + rcases vsample_1 with ⟨ vs1, Hvs1 ⟩ + rcases vsample_rest with ⟨ vsr, Hvsr ⟩ + cases vs1 + · simp at Hvs1 + rename_i vs1 vs_emp + conv => + enter [2, 1, a, 1] + unfold sv4_privMaxF + simp [len_list_append_rev] + have Hemp : vs_emp = [] := by cases vs_emp <;> simp_all + simp_all + congr + + · -- This is true, I'm pretty confident + -- Both can reduce to probPure init + + sorry + + + + + + +-- def sv3_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : SLang sv1_state := do +-- probWhileCut (sv1_privMaxC τ l) (@sv1_privMaxF dps ε₁ ε₂) (point + 1) init +-- +-- def sv3_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := +-- fun (point : ℕ) => +-- let computation : SLang ℕ := do +-- let τ <- @privNoiseZero dps ε₁ (2 * ε₂) +-- let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) +-- let sk <- @sv3_loop dps ε₁ ε₂ τ l point ([], v0) +-- return (sv1_threshold sk) +-- computation point + + +/- + + -- How much excess state does s have above i def state_len_excess (i : sv1_state) (s : sv1_state) : Option ℕ := @@ -584,50 +898,6 @@ lemma sv4_next_self_emp (st : sv1_state) : -lemma sv3_loop_unroll_1_alt [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l point (initial_state : sv1_state) : - sv3_loop ε₁ ε₂ τ l (point + 1) initial_state = - (do - let vk1 <- @privNoiseZero dps ε₁ (4 * ε₂) - if (sv1_privMaxC τ l initial_state) - then (sv3_loop ε₁ ε₂ τ l point (initial_state.1 ++ [initial_state.2], vk1)) - else probPure initial_state) := by - rcases initial_state with ⟨ _ , _ ⟩ - rw [sv3_loop_unroll_1] - -lemma sv3_loop_unroll_2 [dps : DPSystem ℕ] τ ε₁ ε₂ l point (initial_state : sv1_state) : - sv3_loop ε₁ ε₂ τ l (point + 1) initial_state = - (do - let presample_list <- @sv4_presample dps ε₁ ε₂ 1 - if (sv1_privMaxC τ l initial_state) - then do - let next_state <- sv4_next initial_state presample_list initial_state - sv3_loop ε₁ ε₂ τ l point next_state - else probPure initial_state) := by - rw [sv3_loop_unroll_1_alt, sv4_presample, sv4_presample ] - simp - apply SLang.ext - intro x - simp - apply tsum_congr - intro y - congr 1 - split <;> simp - - simp [sv4_next_self] - rw [ENNReal.tsum_eq_add_tsum_ite (sv4_state_shift initial_state y)] - simp - conv => - lhs - rw [<- add_zero (sv3_loop _ _ _ _ _ _ _)] - congr 1 - symm - simp - intro i H1 H2 - exfalso - apply H1 - apply H2 - - def sv4_privMaxC (τ : ℤ) (l : List ℕ) (st : sv1_state) : Bool := sv1_privMaxC τ l st @@ -636,9 +906,29 @@ def sv4_privMaxC (τ : ℤ) (l : List ℕ) (st : sv1_state) : Bool := sv1_privMa def sv4_privMaxF (initial : sv1_state) (presamples : List ℤ) (st : sv1_state) : SLang sv1_state := sv4_next initial presamples st + +-- Init starts out with some length +-- point starts out with a fixed value +-- We want to move "point" samples outwards, while keeping the state the same + + +-- Partway through, we don't know anything about the initial_length and init relationship +-- Just that the current state "starts with" init +-- The current state "adds onto" init +-- The difference between the new stuff in the state and init is at most "point" +-- The value of init doesn't matter +-- +-- How does this change when we add one sample onto the end of init? +-- The current state still has to start ith init +-- The current state still has to add onto init +-- The current state has increased in length by 1 but init hasn't so the differece is still at most "point" + + + -- Compute the loop, but using presamples instead -def sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : SLang sv1_state := do - let presamples <- @sv4_presample dps ε₁ ε₂ point -- (point + 1) breaks the base case... +def sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : + SLang sv1_state := do + let presamples <- @sv4_presample dps ε₁ ε₂ point probWhileCut (sv4_privMaxC τ l) (sv4_privMaxF init presamples) @@ -654,12 +944,18 @@ lemma sv3_loop_eq_sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) unfold sv3_loop unfold sv4_loop + -- The support of state' is contains only length l states, + -- where l is a value freater than init - suffices ∀ state', (probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (point + 1) state' = + suffices ∀ state', + (List.length state'.1 ≥ List.length init.1) -> + (probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (point + 1) state' = (do let presamples ← sv4_presample ε₁ ε₂ point probWhileCut (sv4_privMaxC τ l) (sv4_privMaxF init presamples) (point + 1) state')) - by apply this + by + apply this + simp induction point · intro init @@ -670,9 +966,12 @@ lemma sv3_loop_eq_sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) split <;> try rfl -- Does the left side have more samples than the right side? -- Not a good sign! + intro _ apply SLang.ext intro _ simp + intro _ + simp · intro init rename_i point' IH @@ -686,24 +985,40 @@ lemma sv3_loop_eq_sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) cases (Classical.em (sv1_privMaxC τ l init = true)) · simp rename_i init_start Hcond + intro Hlen -- Horrifying, but only because I can't conv under the if (dependent types)? have X : ((sv4_presample ε₁ ε₂ 1).probBind fun presample_list => if sv1_privMaxC τ l init = true then - (sv4_next init presample_list init).probBind fun next_state => + (sv4_next init (↑presample_list) init).probBind fun next_state => probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (point' + 1) next_state else probPure init) = ((sv4_presample ε₁ ε₂ 1).probBind fun presample_list => if sv1_privMaxC τ l init = true then - (sv4_next init presample_list init).probBind fun next_state => - ((sv4_presample ε₁ ε₂ point') >>= - (fun presamples => - probWhileCut (sv4_privMaxC τ l) (sv4_privMaxF init_start presamples) (point' + 1) next_state)) + (sv4_next init (↑presample_list) init).probBind fun next_state => + (do + let presamples ← sv4_presample ε₁ ε₂ point' + probWhileCut (sv4_privMaxC τ l) (sv4_privMaxF init_start ↑presamples) (point' + 1) next_state) else probPure init) := by - simp_all - rw [X] + -- apply probBind_congr_strong + apply SLang.ext + intro x + simp + apply tsum_congr + intro a + congr 1 + split + · simp + -- a is nonzero + -- This sum is only nonzero for a_1 equal to a plus one element + -- In that case, the IH is provable + -- Regular tsum congr will work but do cases on that after + sorry + · simp + sorry + rw [X] clear X clear IH simp_all @@ -748,10 +1063,13 @@ lemma sv3_loop_eq_sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) -- still a difference the sv4_next term + rw [<- ENNReal.tsum_prod] sorry - · simp_all + · sorry + /- + simp_all simp [probWhileCut, probWhileFunctional] have X : sv4_privMaxC τ l init = false := by unfold sv4_privMaxC @@ -764,6 +1082,7 @@ lemma sv3_loop_eq_sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) split · sorry · simp + -/ -- apply SLang.ext @@ -787,9 +1106,7 @@ lemma sv3_loop_eq_sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) - - - +-/ From 8612d1640792815e32c74ed31f0179a1b3a9cec9 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 7 Oct 2024 14:44:43 -0400 Subject: [PATCH 046/100] checkpoint --- .../Queries/UnboundedMax/Properties.lean | 822 +++--------------- 1 file changed, 106 insertions(+), 716 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index b54e2cae..a058d113 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -506,62 +506,6 @@ lemma sv2_eq_sv3 [dps : DPSystem ℕ] ε₁ ε₂ l : sv2_privMax ε₁ ε₂ l - -/- -def sv4_state : Type := (n : ℕ) × { l : List ℤ // List.length l = n ∧ n > 0} - -def sv4_threshold (s : sv4_state) : ℕ := s.1 - --- Last noise value, because length > 0 -def sv4_last (s : sv4_state) : ℤ := sorry - --- Add one more noise value onto the end -def sv4_append (s : sv4_state) (z : ℤ) : sv4_state := sorry - -def sv4_initial (z : ℤ) : sv4_state := sorry - - --- TODO? Bijection between sv4_states and sv1_states? - -def sv4_privMaxC (τ : ℤ) (l : List ℕ) (s : sv4_state) : Bool := - decide (exactDiffSum (sv4_threshold s) l + (sv4_last s) < τ) - - -def sv4_privMaxF [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (s : sv4_state) : SLang sv4_state := do - let vn <- @privNoiseZero dps ε₁ (4 * ε₂) - return (sv4_append s vn) - -def sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : SLang sv4_state := do - probWhileCut (sv4_privMaxC τ l) (@sv4_privMaxF dps ε₁ ε₂) (point + 1) init - - -def sv4_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := - fun (point : ℕ) => - let computation : SLang ℕ := do - let τ <- @privNoiseZero dps ε₁ (2 * ε₂) - let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) - let sk <- @sv4_loop dps ε₁ ε₂ τ l point (sv4_initial v0) - return (sv4_threshold sk) - computation point - - --- sv3 to sv4 argument : there is a bijection between sv3 states and sv4 states --- Prove this is sv4 to sv5 ends up being provable - - - - -def sv5_presample [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (n : ℕ) : SLang { l : List ℤ // List.length l = n } := - match n with - | Nat.zero => return ⟨ [], by simp ⟩ - | Nat.succ n' => do - let vk1 <- @privNoiseZero dps ε₁ (4 * ε₂) - let vks <- @sv5_presample dps ε₁ ε₂ n' - return ⟨ vks ++ [vk1], by simp ⟩ --/ - - - -- Commute out a single sample from the loop lemma sv3_loop_unroll_1 [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l point L vk : sv3_loop ε₁ ε₂ τ l (point + 1) (L, vk) = @@ -614,15 +558,11 @@ lemma sv3_loop_unroll_1 [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l p unfold privNoiseZero exact Eq.symm (PMF.tsum_coe (DPSystem.noise (fun _ => 0) 1 ε₁ (4 * ε₂) [])) - - - --- New attempt: Let's just actually use consumable tape --- We will not be able to prove the equality between loops, because they have different types --- We can still unroll sv3, inductively --- Can we get equality between loops back, if the tape loop projects out to the non-tape state? --- Can we unroll tape inductively? - +/- +## Program version 4 + - Executable + - Presamples at each point, and then executes a deterministic loop +-/ -- An sv4 state is an sv1 state plus a list of presamples def sv4_state : Type := sv1_state × List ℤ @@ -635,12 +575,12 @@ def sv4_presample [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (n : ℕ) : SLang { let vks <- @sv4_presample dps ε₁ ε₂ n' return ⟨ vks ++ [vk1], by simp ⟩ + def sv4_privMaxF (s : sv4_state) : SLang sv4_state := match s.2 with | [] => probZero | (p :: ps) => return ((s.1.1 ++ [s.1.2], p), ps) - def sv4_privMaxC (τ : ℤ) (l : List ℕ) (st : sv4_state) : Bool := sv1_privMaxC τ l st.1 def sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : SLang sv1_state := do @@ -648,9 +588,6 @@ def sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) let v <- probWhileCut (sv4_privMaxC τ l) sv4_privMaxF (point + 1) (init, presamples) return v.1 - - - lemma sv3_loop_unroll_1_alt [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l point (initial_state : sv1_state) : sv3_loop ε₁ ε₂ τ l (point + 1) initial_state = (do @@ -673,6 +610,9 @@ def sv4_presample_split [DPSystem ℕ] (ε₁ ε₂ : ℕ+) (point : ℕ) : apply SLang.ext intro final_state simp [sv4_presample] + + -- By bijection + -- #check tsum_eq_tsum_of_ne_zero_bij sorry @@ -774,6 +714,7 @@ def sv3_sv4_loop_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : Li simp -- There is a bijection here sorry + rw [ToPresample] clear ToPresample @@ -828,679 +769,128 @@ def sv3_sv4_loop_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : Li simp_all congr - · -- This is true, I'm pretty confident - -- Both can reduce to probPure init - - sorry - - - - - - --- def sv3_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : SLang sv1_state := do --- probWhileCut (sv1_privMaxC τ l) (@sv1_privMaxF dps ε₁ ε₂) (point + 1) init --- --- def sv3_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := --- fun (point : ℕ) => --- let computation : SLang ℕ := do --- let τ <- @privNoiseZero dps ε₁ (2 * ε₂) --- let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) --- let sk <- @sv3_loop dps ε₁ ε₂ τ l point ([], v0) --- return (sv1_threshold sk) --- computation point - - -/- - - - --- How much excess state does s have above i -def state_len_excess (i : sv1_state) (s : sv1_state) : Option ℕ := - Int.toNat' ((List.length s.1 : Int) - (List.length i.1) : Int) - --- Return the next sample from presamples, ignoring the first n elements -def sv4_next_pre (presamples : List ℤ) (n : ℕ) : Option ℤ := - match presamples with - | [] => none - | (x :: xs) => - match n with - | 0 => some x - | (Nat.succ n') => sv4_next_pre xs n' - -def sv4_state_shift (s : sv1_state) (z : ℤ) : sv1_state := (s.1 ++ [s.2], z) - --- If there are 5 more samples in state than there are in intial state, we want to ignore the first 5 entries in presamples -def sv4_next_def (initial_state : sv1_state) (presamples : List ℤ) (state : sv1_state) : Option sv1_state := do - let excess <- state_len_excess initial_state state - let next_sample <- sv4_next_pre presamples excess - return (sv4_state_shift state next_sample) - --- Intuition: --- If there are 5 more samples in state than there are in intial state, we want to ignore the first 5 entries in presamples -def sv4_next (initial_state : sv1_state) (presamples : List ℤ) (state : sv1_state) : SLang sv1_state := - match sv4_next_def initial_state presamples state with - | none => probZero - | some x => return x - - -lemma sv4_next_self (st : sv1_state) (y : ℤ) (ys : List ℤ) : - sv4_next st (y :: ys) st = probPure (sv4_state_shift st y) := by - simp [sv4_next, sv4_next_def, state_len_excess, Int.toNat', sv4_next_pre] - - - -lemma sv4_next_self_emp (st : sv1_state) : - sv4_next st [] st = probZero := by - simp [sv4_next, sv4_next_def, state_len_excess, Int.toNat', sv4_next_pre] - - - - - - -def sv4_privMaxC (τ : ℤ) (l : List ℕ) (st : sv1_state) : Bool := sv1_privMaxC τ l st - --- Assume that st is a prefix state of presamples? If we can't add on a state such that --- the returned state is still a prefix state, then probZero. -def sv4_privMaxF (initial : sv1_state) (presamples : List ℤ) (st : sv1_state) : SLang sv1_state := - sv4_next initial presamples st - - --- Init starts out with some length --- point starts out with a fixed value --- We want to move "point" samples outwards, while keeping the state the same - - --- Partway through, we don't know anything about the initial_length and init relationship --- Just that the current state "starts with" init --- The current state "adds onto" init --- The difference between the new stuff in the state and init is at most "point" --- The value of init doesn't matter --- --- How does this change when we add one sample onto the end of init? --- The current state still has to start ith init --- The current state still has to add onto init --- The current state has increased in length by 1 but init hasn't so the differece is still at most "point" - - - --- Compute the loop, but using presamples instead -def sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : - SLang sv1_state := do - let presamples <- @sv4_presample dps ε₁ ε₂ point - probWhileCut - (sv4_privMaxC τ l) - (sv4_privMaxF init presamples) - (point + 1) - init - - - --- I want to prove that sv4_loop is equal to sv3_loop - -lemma sv3_loop_eq_sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : - @sv3_loop dps ε₁ ε₂ τ l point init = @sv4_loop dps ε₁ ε₂ τ l point init := by - unfold sv3_loop - unfold sv4_loop - - -- The support of state' is contains only length l states, - -- where l is a value freater than init - - suffices ∀ state', - (List.length state'.1 ≥ List.length init.1) -> - (probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (point + 1) state' = - (do - let presamples ← sv4_presample ε₁ ε₂ point - probWhileCut (sv4_privMaxC τ l) (sv4_privMaxF init presamples) (point + 1) state')) - by - apply this + · conv => + enter [2] + rw [sv4_presample_split] simp - - induction point - · intro init - rename_i init_start - simp [sv4_presample] - unfold sv4_privMaxC - simp [probWhileCut, probWhileFunctional] - split <;> try rfl - -- Does the left side have more samples than the right side? - -- Not a good sign! - intro _ - apply SLang.ext - intro _ - simp - intro _ - simp - · intro init - rename_i point' IH - - -- Unroll it once, using the lemma - -- Do we need to unroll this beforehand? - let H := sv3_loop_unroll_2 τ ε₁ ε₂ l point' - unfold sv3_loop at H - rw [H] - clear H - - cases (Classical.em (sv1_privMaxC τ l init = true)) - · simp - rename_i init_start Hcond - intro Hlen - - -- Horrifying, but only because I can't conv under the if (dependent types)? - have X : - ((sv4_presample ε₁ ε₂ 1).probBind fun presample_list => - if sv1_privMaxC τ l init = true then - (sv4_next init (↑presample_list) init).probBind fun next_state => - probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (point' + 1) next_state - else probPure init) = - ((sv4_presample ε₁ ε₂ 1).probBind fun presample_list => - if sv1_privMaxC τ l init = true then - (sv4_next init (↑presample_list) init).probBind fun next_state => - (do - let presamples ← sv4_presample ε₁ ε₂ point' - probWhileCut (sv4_privMaxC τ l) (sv4_privMaxF init_start ↑presamples) (point' + 1) next_state) - else probPure init) := by - - -- apply probBind_congr_strong - apply SLang.ext - intro x - simp - apply tsum_congr - intro a - congr 1 - split - · simp - -- a is nonzero - -- This sum is only nonzero for a_1 equal to a plus one element - -- In that case, the IH is provable - -- Regular tsum congr will work but do cases on that after - sorry - · simp - sorry - rw [X] - clear X - clear IH - simp_all - apply SLang.ext intro final_state simp - - -- Commute the presample steps together on the LHS? - conv => - enter [1, 1, a, 2, 1, a1] - rw [<- ENNReal.tsum_mul_left] - conv => - enter [1, 1, a] - rw [ENNReal.tsum_comm] - rw [<- ENNReal.tsum_mul_left] - conv => - enter [1, 1, a, 1, i, 2, 1, a1] - rw [mul_comm] - rw [mul_assoc] - conv => - enter [1, 1, a, 1, i] - rw [ENNReal.tsum_mul_left] - rw [<- mul_assoc] - - have Hcond' : sv4_privMaxC τ l init = true := by - simp [sv4_privMaxC] - trivial - conv => - enter [2, 1, a] - unfold probWhileCut - unfold probWhileFunctional - simp_all - conv => - enter [2, 1, a, 2, 1, i1] - simp only [sv4_privMaxF] - rw [mul_comm] - - -- Now just need to finish this by a bijection? - -- The RHS is zero on empty lists - -- Make general support lemma for sv4_presample? - - -- still a difference the sv4_next term - - - rw [<- ENNReal.tsum_prod] - - sorry - · sorry - /- - simp_all - simp [probWhileCut, probWhileFunctional] - have X : sv4_privMaxC τ l init = false := by - unfold sv4_privMaxC - trivial - simp_all - apply SLang.ext - intro final_state - -- Both 1 because sum of PMF? - simp + apply tsum_congr + intro v1 split - · sorry + · conv => + enter [1] + rw [<- mul_one (sv4_presample _ _ _ _)] + congr 1 + symm + sorry + -- Sum of PMF is 1 · simp - -/ - - - -- apply SLang.ext - -- intro final_state - -- simp - -- - -- -- Step through to get a new init? - -- conv => - -- -- I want to enter into theat probWhileCut - -- enter [1, 1, a, 2] - -- simp - - -- skip - -- simp - - - -- -- Then use the IH - - - -- sorry - - - --/ - - - -/- - -def sv1_to_list (s : sv1_state) : List ℤ := s.1 ++ [s.2] - -def list_to_sv1 (s : List ℤ) : sv1_state := - match s with - | [] => ([], 0) -- Error case - | [x] => ([], x) - | (x :: xs) => - let w := list_to_sv1 xs - (x :: w.1, w.2) - - -def sv1_prefix (s : sv1_state) (n : ℕ) := list_to_sv1 $ List.take n $ sv1_to_list s - - -def sv4_0_cond (presamples : sv1_state) (τ : ℤ) (l : List ℕ) (n : ℕ) : Bool := - sv1_privMaxC τ l (sv1_prefix presamples n) - -def sv4_0_body (n : ℕ) : SLang ℕ := - return (n + 1) - - -def sv4_0_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (samples : ℕ) (init : sv1_state) : SLang ℕ := do - let presamples <- @sv4_0_presample dps ε₁ ε₂ init samples - let sk <- probWhileCut - (sv4_0_cond presamples τ l) - sv4_0_body - (samples + 1) - (List.length (init.1)) - return sk + 1 - - -def sv3_loop_mod [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (samples : ℕ) (init : sv1_state) : SLang ℕ := do - let st <- probWhileCut (@sv1_privMaxC τ l) (@sv1_privMaxF dps ε₁ ε₂) samples init - return sv1_threshold st --- Turn the state back into a nat, but this time, over the presampled list --- FIXME: Rename if this actually works -def sv4_0_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := - fun (point : ℕ) => - let computation : SLang ℕ := do - let τ <- @privNoiseZero dps ε₁ (2 * ε₂) - let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) - @sv4_0_loop dps ε₁ ε₂ τ l point ([], v0) - computation point - - - --- Now try to prove this equivalence by induction for all inits - -lemma sv3_eq_sv4_0_loop [dps : DPSystem ℕ] ε₁ ε₂ τ l point init : - @sv3_loop_mod dps τ ε₁ ε₂ l (point + 1) init = @sv4_0_loop dps τ ε₁ ε₂ l point init := by - revert init - induction point - · intro init - apply SLang.ext - intro p' - simp +def sv4_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := + fun (point : ℕ) => + let computation : SLang ℕ := do + let τ <- @privNoiseZero dps ε₁ (2 * ε₂) + let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + let sk <- @sv4_loop dps ε₁ ε₂ τ l point ([], v0) + return (sv1_threshold sk) + computation point - -- Simplify the RHS: zero presampling occurs - simp [sv4_0_loop] - simp [sv4_0_presample] - rw [ENNReal.tsum_eq_add_tsum_ite init] +def sv3_eq_sv4 [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : + @sv3_privMax dps ε₁ ε₂ l = @sv4_privMax dps ε₁ ε₂ l := by + unfold sv3_privMax + unfold sv4_privMax simp conv => - lhs - rw [<- zero_add (sv3_loop_mod _ _ _ _ _ _ _)] - conv => - rhs - rw [add_comm] - congr 1 - · symm - simp - intro _ HK1 HK2 - exfalso - apply HK1 - apply HK2 - - -- Simplify the LHS - simp [sv3_loop_mod] -- ? - simp [probWhileCut, probWhileFunctional] - simp [sv4_0_cond] - simp [sv1_prefix, sv1_to_list, list_to_sv1] - rw [List.take_append] - simp [list_to_sv1] - have H : (list_to_sv1 (init.1 ++ [init.2])) = init := by - sorry - simp [H] - simp [sv4_0_body] - split <;> simp - simp [sv1_threshold] - rw [ENNReal.tsum_eq_add_tsum_ite init] - simp - - sorry - · sorry - - - - - -lemma sv3_eq_sv4_0 [dps : DPSystem ℕ] ε₁ ε₂ l : sv3_privMax ε₁ ε₂ l = sv4_0_privMax ε₁ ε₂ l := by - sorry - - - - - --- -- Probably useless as stated but should be true? --- -- For the inductive argument we only need n2 = 1, btw --- lemma sv3_0_presample_additive [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (x0 : sv1_state) (n1 n2 : ℕ) : --- ((@sv4_0_presample dps ε₁ ε₂ x0 n1) >>= fun x1 => @sv4_0_presample dps ε₁ ε₂ x1 n2) = --- @sv4_0_presample dps ε₁ ε₂ x0 (n1 + n2) := by --- revert x0 --- induction n2 --- · intro x0 --- apply SLang.ext --- intro v --- simp [sv4_0_presample] --- rw [ENNReal.tsum_eq_add_tsum_ite ?G1] --- case G1 => apply v --- simp --- conv => --- rhs --- rw [<- add_zero (sv4_0_presample _ _ _ _ _)] --- congr 1 --- simp --- intro i H1 H2 --- exfalso --- apply H1 --- symm --- apply H2 --- · intro x0 --- rename_i n' IH --- apply SLang.ext --- intro v --- simp --- -- IDK. Something is doable for sure? --- sorry - - - - - --- def sv3_1_threshold (s : sv3_1_state) : ℕ := List.length s.1 --- --- def sv3_1_noise (s : sv3_1_state) : ℤ := s.2.1 --- --- def sv3_1_count (s : sv3_1_state) : ℕ := s.2.2 --- --- def sv3_1_privMaxC (τ : ℤ) (l : List ℕ) (s : sv3_1_state) : Bool := --- decide (exactDiffSum (sv3_1_threshold s) l + (sv3_1_noise s) < τ) --- --- def sv3_1_privMaxF [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (s : sv3_1_state) : SLang sv3_1_state := do --- let vn <- @privNoiseZero dps ε₁ (4 * ε₂) --- return (s.1 ++ [s.2.1], vn, s.2.2 + 1) --- --- def sv3_1_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := --- fun (point : ℕ) => --- let computation : SLang ℕ := do --- let τ <- @privNoiseZero dps ε₁ (2 * ε₂) --- let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) --- let sk <- probWhileCut (sv3_1_privMaxC τ l) (sv3_1_privMaxF ε₁ ε₂) (point + 1) ([], v0, 0) --- return (sv3_1_count sk) --- computation point --- --- --- --- lemma sv3_eq_sv3_1 [dps : DPSystem ℕ] ε₁ ε₂ l : sv3_privMax ε₁ ε₂ l = sv3_1_privMax ε₁ ε₂ l := by --- apply SLang.ext --- intro point --- unfold sv3_privMax --- unfold sv3_1_privMax --- simp --- apply tsum_congr --- intro τ --- congr 1 --- apply tsum_congr --- intro v0 --- congr 1 --- --- sorry - - + enter [1, x, 1, y, 2, 1, z] + rw [sv3_sv4_loop_eq] /- -## Program version 4 - - Presamples each vk - - Iteratively checks if the loop - - Still loops through prefixes of the vk's iteratively +## Program version 5 + - Executable + - Isolates the loop for the next step -/ --- Presamples the list (v_(n-1), v_(n-2), ..., v0) -def sv4_presample [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (n : ℕ) : SLang (List ℤ) := do - match n with - | Nat.zero => return [] - | (Nat.succ n') => do - let v <- @privNoiseZero dps ε₁ (4 * ε₂) - let l <- sv4_presample ε₁ ε₂ n' - return (v :: l) - -def sv4_evaluate_history (vks : List ℤ) (τ : ℤ) (l : List ℕ) : Bool := - match vks with - | [] => true - | (v0 :: vs) => sv1_privMaxC τ l (vs, v0) && sv4_evaluate_history vs τ l +def sv5_loop (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : SLang ℕ := do + let sk <- probWhileCut (sv4_privMaxC τ l) sv4_privMaxF (point + 1) init + return (sv1_threshold sk.1) +def sv5_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := + fun (point : ℕ) => + let computation : SLang ℕ := do + let τ <- @privNoiseZero dps ε₁ (2 * ε₂) + let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + let presamples <- @sv4_presample dps ε₁ ε₂ point + @sv5_loop τ l point (([], v0), presamples) + computation point -def sv4_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := - fun (point : ℕ) => - let computation : SLang ℕ := do - let τ <- @privNoiseZero dps ε₁ (2 * ε₂) - let history <- sv4_presample ε₁ ε₂ point - let vk <- @privNoiseZero dps ε₁ (4 * ε₂) - if (sv4_evaluate_history history τ l && (! (sv4_evaluate_history (history ++ [vk]) τ l))) - then return point - else probZero - computation point - -lemma sv3_eq_sv4 [dps : DPSystem ℕ] ε₁ ε₂ l : sv3_privMax ε₁ ε₂ l = sv4_privMax ε₁ ε₂ l := by - apply SLang.ext - intro point - unfold sv3_privMax +def sv4_eq_sv5 [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : + @sv4_privMax dps ε₁ ε₂ l = @sv5_privMax dps ε₁ ε₂ l := by unfold sv4_privMax - + unfold sv5_privMax + unfold sv4_loop + unfold sv5_loop simp - apply tsum_congr - intro τ - congr 1 - - conv => - enter [1, 1, v0, 2, 1, st] - conv => - enter [2, 1, hist, 2, 1, vk] - - -- Maybe don't keep this? - cases point - · -- point is 0 - simp [sv4_presample, sv1_threshold] - unfold sv1_state - conv => - enter [1, 1, v0] - rw [<- ENNReal.tsum_mul_left] - conv => - rw [ENNReal.tsum_comm] - rw [ENNReal.tsum_prod'] - apply tsum_congr - intro hist - simp - sorry - - sorry - - -- -- Simplify the common start to the programs - - -- Rewrite sv4_presample .. (n + 1) into a program which presamples n and 1 separately - - - -- let sv4_presample_sep : SLang (List ℤ) := do - -- let v0 <- privNoiseZero ε₁ (4 * ε₂) - -- let L <- sv4_presample ε₁ ε₂ n - -- return (v0 :: L) - - -- have sv4_presample_sep_eq a : sv4_presample ε₁ ε₂ (n + 1) a = sv4_presample_sep a := by - -- dsimp only [sv4_presample_sep] - -- clear sv4_presample_sep - -- revert a - -- induction n - -- · simp [sv4_presample] - -- rename_i n IH - -- intro a - - -- -- Rewrite LHS to include LHS of IH - -- unfold sv4_presample - -- simp - -- conv => - -- enter [1, 1, b, 2, 1, c] - -- rw [IH] - -- clear IH - - -- -- Conclude by simplification - -- simp - -- apply tsum_congr - -- intro v0 - -- congr 1 - - -- sorry -- Rewriting through the ite is annoying but doable - -- conv => - -- enter [2, 1, a] - -- rw [sv4_presample_sep_eq] - -- dsimp [sv4_presample_sep] - -- clear sv4_presample_sep sv4_presample_sep_eq - - - - -- conv => - -- enter [1, 1, a] - -- rw [<- ENNReal.tsum_mul_left] - -- conv => - -- enter [2, 1, a] - -- rw [<- ENNReal.tsum_mul_left] - -- enter [1, i] - -- rw [<- ENNReal.tsum_mul_right] - -- conv => - -- enter [2] - -- rw [ENNReal.tsum_comm] - -- apply tsum_congr - -- intro v0 - -- unfold sv1_state - -- rw [ENNReal.tsum_prod'] - -- apply tsum_congr - -- intro hist - -- apply tsum_congr - -- intro vk - - -- conv => - -- enter [2] - -- rw [mul_comm] - -- repeat rw [mul_assoc] - -- congr 1 - -- simp [sv1_threshold] - - -- split - -- · - -- sorry - -- · -- Can I solve this one without induction? - -- -- It simplifies the induction on the other side - -- simp - -- split - -- · right - -- right - -- intro l Hl - -- induction n - -- · simp_all - -- unfold sv4_presample - -- simp - -- intro H1 - -- simp_all - -- sorry - -- · sorry - -- · sorry - - -- revert hist vk - -- induction n - -- · simp_all - -- sorry - -- · rename_i n IH - -- intro hist vk - - -- have Hcond : hist = [] ∨ ∃ vl hist', hist = hist' ++ [vl] := by sorry - -- cases Hcond - -- · simp_all - -- rename_i h - -- rcases h with ⟨ hist', ⟨ vk', H ⟩ ⟩ - -- simp [H] - -- clear H - -- conv => - -- enter [1] - -- unfold probWhileCut - -- unfold probWhileFunctional +/- +## Program version 6 + - Executable + - Changes the loop from a probWhileCut into a single, deterministic, check +-/ +-- Evaluate the nth conditional starting at state s +-- Return false if the loop will not terminate on iterate n, starting at s +-- Return true if the loop will terminate on iterate n, starting at s +-- Return false if the loop never gets to iterate n, starting at s +def sv6_privMax_nth (τ : ℤ) (l : List ℕ) (s : sv4_state) (n : ℕ) : Bool := + match s with + | ((past, present), future) => + match n with + | 0 => sorry + | Nat.succ n' => sorry +def sv6_loop (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : SLang ℕ := + if (sv6_privMax_nth τ l init point) ∧ ¬ (sv6_privMax_nth τ l init (point + 1)) + then return point + else probZero +-- These functions are not equal. However, they are equal at "point" (+- 1?) +def sv5_sv6_loop_eq_point (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : + @sv5_loop τ l point init point = @sv6_loop τ l point init point := by + unfold sv5_loop + unfold sv6_loop + -- What is the inductive argument? What can I unroll? + sorry -/- -## Program version 5 - - Presamples each vk - - Performs a single check, rather than a loop --/ -def sv5_evaluate_history (vks : List ℤ) (τ : ℤ) (l : List ℕ) : Bool := - match vks with - | [] => true - | (v0 :: vs) => sv1_privMaxC τ l (vs, v0) +def sv6_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := + fun (point : ℕ) => + let computation : SLang ℕ := do + let τ <- @privNoiseZero dps ε₁ (2 * ε₂) + let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + let presamples <- @sv4_presample dps ε₁ ε₂ point + @sv6_loop τ l point (([], v0), presamples) + computation point -def sv5_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := - fun (point : ℕ) => - let computation : SLang ℕ := do - let τ <- @privNoiseZero dps ε₁ (2 * ε₂) - let history <- sv4_presample ε₁ ε₂ (point + 1) - let vk <- @privNoiseZero dps ε₁ (4 * ε₂) - if (sv5_evaluate_history history τ l && (! sv1_privMaxC τ l (history, vk))) - then return point - else probZero - computation point - -lemma sv4_eq_sv5 [dps : DPSystem ℕ] ε₁ ε₂ l : sv4_privMax ε₁ ε₂ l = sv5_privMax ε₁ ε₂ l := by - sorry --/ +def sv5_sv6_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : + @sv5_privMax dps ε₁ ε₂ l = @sv6_privMax dps ε₁ ε₂ l := by + unfold sv5_privMax + unfold sv6_privMax + apply SLang.ext + intro eval_point + simp + apply tsum_congr + intro _ + congr + apply funext + intro _ + congr + apply funext + intro _ + congr + rw [sv5_sv6_loop_eq_point] From 6a89720ad1039597fd2d7a1c6b9c038d069e921b Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 7 Oct 2024 20:17:21 -0400 Subject: [PATCH 047/100] sv6 experiment --- .../Queries/UnboundedMax/Properties.lean | 108 ++++++++++++++++-- 1 file changed, 97 insertions(+), 11 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index a058d113..9408ee7b 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -844,29 +844,115 @@ def sv4_eq_sv5 [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : -- Evaluate the nth conditional starting at state s --- Return false if the loop will not terminate on iterate n, starting at s --- Return true if the loop will terminate on iterate n, starting at s --- Return false if the loop never gets to iterate n, starting at s +-- Evaluate the loop conditional n steps in the future +-- false if the tape runs out def sv6_privMax_nth (τ : ℤ) (l : List ℕ) (s : sv4_state) (n : ℕ) : Bool := - match s with - | ((past, present), future) => - match n with - | 0 => sorry - | Nat.succ n' => sorry + match n with + | 0 => + -- Terminates on the 0th iteration if initial state causes the checck to fail + ¬ sv4_privMaxC τ l s + | Nat.succ n' => + match s with + -- If there is more samples on the tape, call recursively + | ((past, present), (future_next :: future_rest)) => + -- sv4_privMaxC τ l ((past, present), (future_next :: future_rest)) ∧ -- Should not terminate here + sv6_privMax_nth τ l ((past ++ [present], future_next), future_rest) n' + | (_, []) => + -- Out of samples on the tape + false + + + +-- TODO Test some values of sv6_priv_nth + +lemma sv6_privMax_0th_true (τ : ℤ) (l : List ℕ) (s : sv4_state) : + (sv4_privMaxC τ l s = false) -> (sv6_privMax_nth τ l s 0 = true) := by + simp [sv6_privMax_nth] + +lemma sv6_privMax_0th_false (τ : ℤ) (l : List ℕ) (s : sv4_state) : + (sv4_privMaxC τ l s = true) -> (sv6_privMax_nth τ l s 0 = false) := by + simp [sv6_privMax_nth] + +lemma sv6_privMax_1th_empty (τ : ℤ) (l : List ℕ) (s : sv1_state) : + (sv6_privMax_nth τ l (s, []) 1 = false) := by + simp [sv6_privMax_nth] + +-- Inductive formula for when the tape is empty +lemma sv6_privMax_ind_empty (τ : ℤ) (l : List ℕ) (s : sv1_state) (n : ℕ): + (sv6_privMax_nth τ l (s, []) (n + 1) = false) := by + simp [sv6_privMax_nth] + +-- Inductive formula for when the tape is not empty +lemma sv6_privMax_ind_nonempty (τ : ℤ) (l : List ℕ) pa pr f fs (n : ℕ): + sv6_privMax_nth τ l ((pa, pr), f :: fs) (n + 1) = + sv6_privMax_nth τ l ((pa ++ [pr], f), fs) n := by + simp [sv6_privMax_nth] + def sv6_loop (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : SLang ℕ := if (sv6_privMax_nth τ l init point) ∧ ¬ (sv6_privMax_nth τ l init (point + 1)) then return point else probZero +-- the loop evaluated at point either has probability mass 0 or 1 +lemma sv6_loop_det τ l point init : sv6_loop τ l point init point = 0 ∨ sv6_loop τ l point init point = 1 := by + unfold sv6_loop + split <;> simp + + +lemma sv6_loop_inv τ l point init : sv6_loop τ l point init point = 1 -> True := by sorry + + +-- Question: What does (sv6_loop τ l (point + 1) init (point + 1)) +-- have to do with (sv6_loop τ l point ?? point) +-- There is an inductive unrolling for sv6_privMax_nth (empty and nonempty) + + + + +-- Question: Do we need any relationship between point and init? +-- It is true on entry that init.1.1.length = point (+- 1?) + + + + -- These functions are not equal. However, they are equal at "point" (+- 1?) def sv5_sv6_loop_eq_point (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : @sv5_loop τ l point init point = @sv6_loop τ l point init point := by - unfold sv5_loop - unfold sv6_loop + revert init + induction point + · intro init + simp [sv5_loop, sv6_loop] + simp [probWhileCut, probWhileFunctional] + simp [sv1_threshold] + split + · rename_i h + apply sv6_privMax_0th_false at h + simp + rw [h] + simp + · rename_i h + simp at h + apply sv6_privMax_0th_true at h + simp_all + -- The last one should always exhaust the tape completely? + -- LHS is taking a sum over the singleton "init" + -- so it's if init.1.1.length = 0 then 1 else 0 + -- RHS is asking if privMaxnth init 1 is false + -- So either + + -- The invariant could be "sv4_state of size n" + -- future = size n? + -- But then, how does the evaluation at point (point - 1) give us anything about the evaluation at point (point) + sorry + sorry + + -- unfold sv5_loop + -- unfold sv6_loop + -- unfold sv4_privMaxC + -- unfold sv1_privMaxC -- What is the inductive argument? What can I unroll? - sorry def sv6_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => From cde6978e5c632a27cc25e28306e446c1e4e778c2 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 28 Oct 2024 08:51:32 -0400 Subject: [PATCH 048/100] checkpoint --- .../Queries/UnboundedMax/Properties.lean | 363 +++++++++++++++--- 1 file changed, 319 insertions(+), 44 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 9408ee7b..3d649c4f 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -209,6 +209,10 @@ lemma sv0_eq_sv1 [dps : DPSystem ℕ] ε₁ ε₂ l : sv0_privMax ε₁ ε₂ l -- RHS: sum over all lists of length "result"? -- rw [tsum_ite_eq] + -- simp [sv1_threshold] + + + sorry @@ -836,12 +840,99 @@ def sv4_eq_sv5 [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : simp + /- ## Program version 6 - Executable - Changes the loop from a probWhileCut into a single, deterministic, check -/ +-- First n loop iterates +-- future samples to try them +def sv6_privMax_check (τ : ℤ) (l : List ℕ) (s : sv4_state) (n : ℕ) : Bool := + match n with + | 0 => true + | Nat.succ n' => + match s with + -- If there is more samples on the tape, call recursively + | ((past, present), (future_next :: future_rest)) => + sv4_privMaxC τ l ((past, present), (future_next :: future_rest)) ∧ + sv6_privMax_check τ l ((past ++ [present], future_next), future_rest) n' + | (_, []) => + -- Out of samples on the tape + false + +-- The state sp is a "past configuration" of sc, ie. one we already checked +def is_past_configuration (sp sc : sv4_state) : Prop := + (sp.1.1.length ≤ sc.1.1.length) ∧ sp.1.1 ++ [sp.1.2] ++ sp.2 = sc.1.1 ++ [sc.1.2] ++ sc.2 + +lemma is_past_configuration_ex1 : is_past_configuration (([], 0), []) (([], 0), []) := by + simp [is_past_configuration] + +lemma is_past_configuration_ex2 : is_past_configuration (([], 0), [1]) (([0], 1), []) := by + simp [is_past_configuration] + +lemma is_past_configuration_ex3 : is_past_configuration (([], 0), [1, 2]) (([0], 1), [2]) := by + simp [is_past_configuration] + + + + + + + + + + + + + +-- All past configurations had their loop check execute to True +def sv6_privMax_hist (τ : ℤ) (l : List ℕ) (s : sv4_state) : Prop := + ∀ sp, (is_past_configuration sp s) -> sv4_privMaxC τ l sp = true + + + + + + + + + +-- What am I defining: +-- When I step sv5 (the probWhileCut), unrolling a loop iterate pops the first value out of +-- future, and pushes it onto the initial condition. +-- This means I want to do induction over future. +-- Therefore, I need to establish what it means for a loop body with some of these pops to have +-- not terminated (??) +-- IE. given a state, I want to be able to express the fact that +-- - Every loop check which moved something to the past must have returned false +-- - We are waiting for n loop checks from the current state to return true +-- For the latter, I define sv6_privMax_check +-- For the former, I define sc6_provMax_hist +-- Should be able to establish these inductively based on hist, assuming the first n-1 loop checks are valid + + + + + + + + + + + + + + + + + + + + + +/- -- Evaluate the nth conditional starting at state s -- Evaluate the loop conditional n steps in the future @@ -877,82 +968,264 @@ lemma sv6_privMax_1th_empty (τ : ℤ) (l : List ℕ) (s : sv1_state) : (sv6_privMax_nth τ l (s, []) 1 = false) := by simp [sv6_privMax_nth] --- Inductive formula for when the tape is empty +-- Inductive formula for n when the tape is empty (not useful) lemma sv6_privMax_ind_empty (τ : ℤ) (l : List ℕ) (s : sv1_state) (n : ℕ): (sv6_privMax_nth τ l (s, []) (n + 1) = false) := by simp [sv6_privMax_nth] --- Inductive formula for when the tape is not empty +-- Inductive formula for n when the tape is not empty (not useful) lemma sv6_privMax_ind_nonempty (τ : ℤ) (l : List ℕ) pa pr f fs (n : ℕ): sv6_privMax_nth τ l ((pa, pr), f :: fs) (n + 1) = sv6_privMax_nth τ l ((pa ++ [pr], f), fs) n := by simp [sv6_privMax_nth] +def len_constr (s : sv4_state) (p : ℕ) : Prop := s.1.1.length + s.2.length = p + +-- These are not true +-- Could change sv4 definition to be monotone so that it is +-- +-- -- After its false once, its false forever +-- lemma sv4_constancy_lo τ l P v P' v' (H : sv1_privMaxC τ l (P, v) = false) : +-- (∃ D, P' ++ [v'] = P ++ [v] ++ D) -> (sv1_privMaxC τ l (P', v') = false) := by +-- intro H1 +-- rcases H1 with ⟨ D, HD ⟩ +-- simp [sv1_privMaxC, sv1_threshold, sv1_noise, exactDiffSum] at H +-- simp [sv1_privMaxC, sv1_threshold, sv1_noise, exactDiffSum] +-- apply (Int.le_trans H) +-- clear H +-- +-- sorry +-- +-- -- If its true once, its true at all points before +-- lemma sv4_constancy_hi τ l P v P' v' (H : sv1_privMaxC τ l (P, v) = false) : +-- (∃ D, P' ++ [v'] ++ D = P ++ [v]) -> (sv1_privMaxC τ l (P', v') = false) := by +-- sorry + +-- Inductive formula for tail +lemma sv6_privMax_ind (τ : ℤ) (l : List ℕ) pa pr f fs (n : ℕ): + -- sv4_privMaxC τ l ((pa ++ [pr], f), fs) = true -> + -- len_constr ((pa, pr), f :: fs) n -> + sv6_privMax_nth τ l ((pa, pr), f :: fs) n = + sv6_privMax_nth τ l ((pa ++ [pr], f), fs) n := by + revert pa pr f + induction fs + · intro pa pr f + sorry + · sorry + def sv6_loop (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : SLang ℕ := if (sv6_privMax_nth τ l init point) ∧ ¬ (sv6_privMax_nth τ l init (point + 1)) then return point else probZero --- the loop evaluated at point either has probability mass 0 or 1 -lemma sv6_loop_det τ l point init : sv6_loop τ l point init point = 0 ∨ sv6_loop τ l point init point = 1 := by - unfold sv6_loop - split <;> simp +-- If the sv6 return flag is true (at zero) then... the sv4 loop condition is false (at zero) +lemma sv6_privMax_loop_cond_true_0 (τ : ℤ) (l : List ℕ) (init : sv4_state) : + (len_constr init 0) -> + (sv6_privMax_nth τ l init 0 = true ∧ sv6_privMax_nth τ l init 1 = false) → + (sv4_privMaxC τ l init = false) := by + intro _ + intro ⟨ H1, _ ⟩ -- We can get nothing out of H2? + simp [sv6_privMax_nth] at H1 + trivial + +-- If the sv6 return flag is true (at 1) then +-- there is at least one element in the future list +-- The sv4_condition after executing one shift is false +lemma sv6_privMax_loop_cond_true_1 (τ : ℤ) (l : List ℕ) (init : sv4_state) : + (len_constr init 1) -> + (sv6_privMax_nth τ l init 1 = true ∧ sv6_privMax_nth τ l init 2 = false) → + ∃ init_past init_pres init_fut1 init_futR, + (init = ((init_past, init_pres), (init_fut1 :: init_futR)) ∧ + sv4_privMaxC τ l ((init_past ++ [init_pres], init_fut1), init_futR) = false) := by + intro HL + rcases init with ⟨ ⟨ init_past, init_present ⟩, init_future ⟩ + simp [len_constr] at HL + + intro ⟨ H1, H2 ⟩ + + cases init_past + · -- init_past is [] so init_future is [x] + simp at HL + cases init_future + · exfalso + simp at HL + rename_i future_1 future_rest + cases future_rest + · -- Continue... + clear HL + + -- What can we get out of H2? Nothing + simp [sv6_privMax_nth] at H2 + + -- What can we get out of H1? + simp [sv6_privMax_nth] at H1 + -- sv4_privMaxC τ l (sv4_privMaxF (init)) is false + -- ∃ x, sv4_privMaxC τ l (([init_present], future_1), []) = false + exists [] + exists init_present + exists future_1 + exists [] + · exfalso + simp at HL + · -- init_past is nonempty so init_future is empty + rename_i past_1 past_rest + simp at HL + -- past_rest is empty + cases past_rest + · -- init_future is empty + cases init_future + · -- Continue... + clear HL + + -- What can we get out of H2? Nothing + simp [sv6_privMax_nth] at H2 + + -- What can we get out of H1? Contradiction? + simp [sv6_privMax_nth] at H1 + -- contradiction? + + · exfalso + simp at HL + · exfalso + simp at HL + linarith + + +-- So if we know that init.future is nonempty, we should be able to make a conclusion +-- about sv4_loop_cond using the s6 loop flag. + +-- lemma sv6_privMax_loop_cond_true (τ : ℤ) (l : List ℕ) (init : sv4_state) (p : ℕ) : +-- (len_constr init p) -> +-- (sv6_privMax_nth τ l init p = true ∧ sv6_privMax_nth τ l init (p + 1) = false) → +-- False := by +-- sorry + + + + + + + + +-- What can we get out of the case where the sv6 flag is false? + +-- When the sv6 flag is false (at zero) then sv4 loop cond at 1 is true +lemma sv6_privMax_loop_cond_false_0 (τ : ℤ) (l : List ℕ) (init : sv4_state) : + (len_constr init 0) -> + ¬ (sv6_privMax_nth τ l init 0 = true ∧ sv6_privMax_nth τ l init 1 = false) → + sv4_privMaxC τ l init = true := by + intro HL H + simp [len_constr] at HL + rcases HL with ⟨ HL1, HL2 ⟩ + rcases init with ⟨ ⟨ init_past, init_present ⟩, init_future ⟩ + simp_all [HL1, HL2] + rw [sv6_privMax_1th_empty] at H + simp at H + simp [sv6_privMax_nth] at H + trivial + + + + -lemma sv6_loop_inv τ l point init : sv6_loop τ l point init point = 1 -> True := by sorry + -- simp [sv6_privMax_nth] at H2 + -- -- H2 is giving us nothing, this is suspicious --- Question: What does (sv6_loop τ l (point + 1) init (point + 1)) --- have to do with (sv6_loop τ l point ?? point) --- There is an inductive unrolling for sv6_privMax_nth (empty and nonempty) + -- sorry +-- Specify that the total length of past and future samples is equal to point (0) +def sv5_sv6_loop_eq_point_0_len_constr (τ : ℤ) (l : List ℕ) (init : sv4_state) : + len_constr init 0 -> + @sv5_loop τ l 0 init 0 = @sv6_loop τ l 0 init 0 := by + intro HC + simp [sv6_loop] + split <;> simp + · rename_i h + -- rcases h with ⟨ h1, h2 ⟩ + -- FIXME: See above for general lemmas about the loop conditional and sv4 loop evals + -- Don't do this manually + sorry + · rename_i h + simp at h + sorry + +-- sv6 and sv5 both do (decreasing) induction over the "future" field-- +-- When the condition is true for sv5, it shifts one sample out of future, and decreases the cut by 1 +-- sv6_loop does ...? + +-- So I want inductive formulas for those, ie, +-- init = (x, y, (a :: as)) +-- in terms of +-- init = (x ++ [a], y, a) +-- evaluated at point = length(x) + length(a::as) + --- Question: Do we need any relationship between point and init? --- It is true on entry that init.1.1.length = point (+- 1?) -- These functions are not equal. However, they are equal at "point" (+- 1?) -def sv5_sv6_loop_eq_point (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : - @sv5_loop τ l point init point = @sv6_loop τ l point init point := by - revert init - induction point - · intro init - simp [sv5_loop, sv6_loop] - simp [probWhileCut, probWhileFunctional] - simp [sv1_threshold] +def sv5_sv6_loop_eq_point (τ : ℤ) (l : List ℕ) (point eval_point : ℕ) (init : sv4_state) : + point ≤ eval_point -> + len_constr init eval_point -> + @sv5_loop τ l point init eval_point = @sv6_loop τ l point init eval_point := by + + rcases init with ⟨ ⟨ pa, pr ⟩, fu ⟩ + revert pa pr point + induction fu + · intro point pa pr HlenIneq Hlen + simp [len_constr] at Hlen + -- simp only [sv6_loop] + simp only [sv5_loop, probWhileCut, probWhileFunctional] + simp only [pure, pure_apply, ↓reduceIte] + simp only [sv1_threshold] + simp only [sv4_privMaxF] split - · rename_i h - apply sv6_privMax_0th_false at h - simp - rw [h] + · simp + unfold sv6_loop + split + · exfalso + -- There should be a contradiction in here somewhere + -- TODO: Probably change the definition of sv6_cond so that its monotone? + sorry + · simp + · unfold sv4_state + unfold sv1_state simp - · rename_i h - simp at h - apply sv6_privMax_0th_true at h - simp_all - -- The last one should always exhaust the tape completely? - -- LHS is taking a sum over the singleton "init" - -- so it's if init.1.1.length = 0 then 1 else 0 - -- RHS is asking if privMaxnth init 1 is false - -- So either - - -- The invariant could be "sv4_state of size n" - -- future = size n? - -- But then, how does the evaluation at point (point - 1) give us anything about the evaluation at point (point) - sorry - sorry + rw [ENNReal.tsum_eq_add_tsum_ite ((pa, pr), [])] + conv => + rhs + rw [<- add_zero (sv6_loop _ _ _ _ _)] + congr + · simp + simp [sv6_loop] + -- ??? - -- unfold sv5_loop - -- unfold sv6_loop - -- unfold sv4_privMaxC - -- unfold sv1_privMaxC + sorry + · simp + intro _ _ _ _ _ _ _ _ + rename_i h a b b_1 a_1 a_2 a_3 a_4 a_5 + subst Hlen a_4 a_5 a_3 + simp_all only [Bool.not_eq_true, not_true_eq_false, imp_false] + · intro point pa pr HlenIneq Hlen + rename_i f ff IH + cases point + · -- point being zero should be a contradiction? + simp [sv6_loop] + simp [sv5_loop] + -- Not sure. Messed up. + sorry + rename_i point + unfold sv5_loop + unfold probWhileCut + unfold probWhileFunctional + sorry - -- What is the inductive argument? What can I unroll? def sv6_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => @@ -980,3 +1253,5 @@ def sv5_sv6_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : intro _ congr rw [sv5_sv6_loop_eq_point] + all_goals sorry +-/ From ff237217bb0ef5831e9b447de4ca02b3977f790d Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Sun, 3 Nov 2024 14:10:37 -0500 Subject: [PATCH 049/100] checkpoint --- .../Queries/UnboundedMax/Properties.lean | 132 +++++++++++++++--- 1 file changed, 109 insertions(+), 23 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 3d649c4f..9a8aefa5 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -847,8 +847,7 @@ def sv4_eq_sv5 [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : - Changes the loop from a probWhileCut into a single, deterministic, check -/ --- First n loop iterates --- future samples to try them +-- When you look at exactly n loops in the future, the check evaluates to true def sv6_privMax_check (τ : ℤ) (l : List ℕ) (s : sv4_state) (n : ℕ) : Bool := match n with | 0 => true @@ -876,10 +875,52 @@ lemma is_past_configuration_ex3 : is_past_configuration (([], 0), [1, 2]) (([0], simp [is_past_configuration] +-- All past configurations had their loop check execute to True +def sv6_privMax_hist (τ : ℤ) (l : List ℕ) (s : sv4_state) : Prop := + ∀ sp, (is_past_configuration sp s) -> sv4_privMaxC τ l sp = true +-- If all past configurations of sp evaluate to True, +-- and the next one evaluates to true, +-- then all past configurations for the next one evaluate to True +lemma sv6_privMax_hist_step (τ : ℤ) (l : List ℕ) (past fut_rest : List ℤ) (present fut : ℤ) : + sv6_privMax_hist τ l ((past, present), fut :: fut_rest) -> + sv4_privMaxC τ l ((past ++ [present], fut), fut_rest) -> + sv6_privMax_hist τ l ((past ++ [present], fut), fut_rest) := by + intro H1 H2 + unfold sv6_privMax_hist + intro s H3 + unfold is_past_configuration at H3 + rcases H3 with ⟨ H3, H4 ⟩ + simp_all + + apply Nat.lt_or_eq_of_le at H3 + cases H3 + · -- The length of s1.1.1 is less than or equal to past + apply H1 + apply And.intro + · linarith + · simp_all + · rename_i Hs11_len + -- The length of s.1.1 is equal to past.length + 1 + -- Now we can characterize s + have Hs11 : List.take (past.length + 1) (s.1.1 ++ s.1.2 :: s.2) = + List.take (past.length + 1) (past ++ present :: fut :: fut_rest) := by + rw [H4] + rw [List.take_left' Hs11_len] at Hs11 + simp [List.take_append] at Hs11 + simp_all + rcases H4 with ⟨ H5, H6 ⟩ + cases s + rename_i s1 _ + cases s1 + simp_all +-- Relationship between sv6_privMax_nth and sv6_privMax_hist +-- If sv6_privMax_nth .. n is true then +-- sv6_privMax_all holds, for the state which separates off the first n elements +-- Do this for sv7 instead? @@ -887,15 +928,78 @@ lemma is_past_configuration_ex3 : is_past_configuration (([], 0), [1, 2]) (([0], --- All past configurations had their loop check execute to True -def sv6_privMax_hist (τ : ℤ) (l : List ℕ) (s : sv4_state) : Prop := - ∀ sp, (is_past_configuration sp s) -> sv4_privMaxC τ l sp = true +def sv6_loop (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : SLang ℕ := do + sorry + -- let sk <- probWhileCut (sv4_privMaxC τ l) sv4_privMaxF (point + 1) init + -- return (sv1_threshold sk.1) + +def sv6_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := + fun (point : ℕ) => + let computation : SLang ℕ := do + let τ <- @privNoiseZero dps ε₁ (2 * ε₂) + let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + let presamples <- @sv4_presample dps ε₁ ε₂ point + @sv6_loop τ l point (([], v0), presamples) + computation point + + + +-- sv6_loop and sv5_loop are equal at point (under some conditions) +def sv5_sv6_loop_eq_point (τ : ℤ) (l : List ℕ) (point eval_pt : ℕ) (past future : List ℤ) (pres : ℤ) : + List.length (past ++ [pres] ++ future) = (point + 1) -> + List.length future = eval_point -> + sv6_privMax_hist τ l ((past, pres), future) -> + @sv5_loop τ l eval_point ((past, pres), future) point = @sv6_loop τ l eval_point ((past, pres), future) point := by + -- To do induction over future, the length of past plus the number of steps remaining shoud be constant + + revert eval_point past pres + induction future + · -- Base case: no future + intro eval_point past pres Hstate Heval Hpast_evals + simp_all + rw [<- Heval] + unfold sv5_loop + unfold probWhileCut + unfold probWhileFunctional + -- LHS: We have consumed the entire future now, + -- If the last conditional check is False, return pure ... + -- else return zero + -- NOTE: This should definitely not follow from Hpast_evals... + have Hshould_be_not_provable : sv4_privMaxC τ l ((past, pres), []) := by + apply Hpast_evals + simp [is_past_configuration]$ + -- Uh oh! + -- sorry + sorry + · -- Inductive case + sorry + + +def sv5_sv6_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : + @sv5_privMax dps ε₁ ε₂ l = @sv6_privMax dps ε₁ ε₂ l := by + unfold sv5_privMax + unfold sv6_privMax + apply SLang.ext + intro eval_point + simp + apply tsum_congr + intro _ + congr + apply funext + intro _ + congr + apply funext + intro _ + congr + rw [sv5_sv6_loop_eq_point] + all_goals sorry + @@ -1236,22 +1340,4 @@ def sv6_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang @sv6_loop τ l point (([], v0), presamples) computation point -def sv5_sv6_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : - @sv5_privMax dps ε₁ ε₂ l = @sv6_privMax dps ε₁ ε₂ l := by - unfold sv5_privMax - unfold sv6_privMax - apply SLang.ext - intro eval_point - simp - apply tsum_congr - intro _ - congr - apply funext - intro _ - congr - apply funext - intro _ - congr - rw [sv5_sv6_loop_eq_point] - all_goals sorry -/ From b32b1530dffcd4eeea0bf949b5ad947296eea37a Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Sun, 3 Nov 2024 15:03:31 -0500 Subject: [PATCH 050/100] sv5 sv6 base case --- .../Queries/UnboundedMax/Properties.lean | 110 +++++++++++++++--- 1 file changed, 94 insertions(+), 16 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 9a8aefa5..a372b355 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -917,21 +917,55 @@ lemma sv6_privMax_hist_step (τ : ℤ) (l : List ℕ) (past fut_rest : List ℤ) simp_all --- Relationship between sv6_privMax_nth and sv6_privMax_hist --- If sv6_privMax_nth .. n is true then --- sv6_privMax_all holds, for the state which separates off the first n elements --- Do this for sv7 instead? +def is_past_configuration_strict (sp sc : sv4_state) : Prop := + (sp.1.1.length < sc.1.1.length) ∧ sp.1.1 ++ [sp.1.2] ++ sp.2 = sc.1.1 ++ [sc.1.2] ++ sc.2 +-- All strictly past configurations had their loop check execute to True +def sv6_privMax_hist_strict (τ : ℤ) (l : List ℕ) (s : sv4_state) : Prop := + ∀ sp, (is_past_configuration_strict sp s) -> sv4_privMaxC τ l sp = true + +lemma sv6_privMax_hist_step_strict (τ : ℤ) (l : List ℕ) (past fut_rest : List ℤ) (present fut : ℤ) : + sv6_privMax_hist_strict τ l ((past, present), fut :: fut_rest) -> + sv4_privMaxC τ l ((past, present), fut :: fut_rest) -> + sv6_privMax_hist_strict τ l ((past ++ [present], fut), fut_rest) := by + intro H1 H2 + unfold sv6_privMax_hist_strict + intro s H3 + unfold is_past_configuration_strict at H3 + rcases H3 with ⟨ H3, H4 ⟩ + simp_all + + apply Nat.lt_or_eq_of_le at H3 + cases H3 + · -- The length of s1.1.1 is less than or equal to past + apply H1 + apply And.intro + · linarith + · simp_all + · rename_i Hs11_len + have Hs11 : List.take (past.length) (s.1.1 ++ s.1.2 :: s.2) = + List.take (past.length) (past ++ present :: fut :: fut_rest) := by + rw [H4] + simp at Hs11_len + rw [List.take_left] at Hs11 + rw [<- Hs11_len] at Hs11 + rw [List.take_left] at Hs11 + cases s + rename_i s1 s2 + cases s1 + rename_i s11 s12 + simp_all def sv6_loop (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : SLang ℕ := do - sorry - -- let sk <- probWhileCut (sv4_privMaxC τ l) sv4_privMaxF (point + 1) init - -- return (sv1_threshold sk.1) + if (sv6_privMax_hist_strict τ l init ∧ ¬ sv4_privMaxC τ l init) + then return point + else probZero + def sv6_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => @@ -948,8 +982,8 @@ def sv6_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang def sv5_sv6_loop_eq_point (τ : ℤ) (l : List ℕ) (point eval_pt : ℕ) (past future : List ℤ) (pres : ℤ) : List.length (past ++ [pres] ++ future) = (point + 1) -> List.length future = eval_point -> - sv6_privMax_hist τ l ((past, pres), future) -> - @sv5_loop τ l eval_point ((past, pres), future) point = @sv6_loop τ l eval_point ((past, pres), future) point := by + sv6_privMax_hist_strict τ l ((past, pres), future) -> + @sv5_loop τ l eval_point ((past, pres), future) point = @sv6_loop τ l point ((past, pres), future) point := by -- To do induction over future, the length of past plus the number of steps remaining shoud be constant @@ -967,17 +1001,61 @@ def sv5_sv6_loop_eq_point (τ : ℤ) (l : List ℕ) (point eval_pt : ℕ) (past -- If the last conditional check is False, return pure ... -- else return zero - -- NOTE: This should definitely not follow from Hpast_evals... - have Hshould_be_not_provable : sv4_privMaxC τ l ((past, pres), []) := by - apply Hpast_evals - simp [is_past_configuration]$ - -- Uh oh! + -- RHS: Should be able to simplify now using Hpast_evals + unfold sv6_loop + split + · -- LHS when the last loop condition returns True (LHS returns 0) + rename_i hlast + split + · -- RHS check True (contradiction, would return 1) + rename_i hhist + rcases hhist with ⟨ hhist, hlast' ⟩ + exfalso + simp_all + · -- RHS check returns False (returns probZero) + rename_i hhist + simp at hhist + simp + intro _ _ _ + simp_all + right + rw [<- Heval] + simp [probWhileCut] - -- sorry + · -- LHS when the last loop condition returns False (should return 1?) + rename_i hlast + simp at hlast + conv => + lhs + simp [sv1_threshold] + unfold sv4_state + unfold sv1_state + rw [(ENNReal.tsum_eq_add_tsum_ite ((past, pres), []))] + conv => + lhs + rw [add_comm] + conv => + rhs + rw [<- zero_add (@ite _ _ _ _ _ )] + congr 1 + · simp + simp_all + rw [<- Hstate] + simp + split + · -- RHS evaluates to True + simp + · -- RHS evaluates to False (contradiction) + exfalso + rename_i h + simp at h + aesop - sorry · -- Inductive case + rename_i fut_next fut_rest IH + + sorry From 0a12c1f58da4860393847ada37b4c87b0110620b Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Sun, 3 Nov 2024 18:09:31 -0500 Subject: [PATCH 051/100] checkpoint explorations into sv6 --- .../Queries/UnboundedMax/Properties.lean | 111 ++++++++++++++++-- 1 file changed, 101 insertions(+), 10 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index a372b355..fd4d43cd 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -959,14 +959,98 @@ lemma sv6_privMax_hist_step_strict (τ : ℤ) (l : List ℕ) (past fut_rest : Li simp_all - +@[simp] +def sv6_cond (τ : ℤ) (l : List ℕ) (init : sv4_state) : Bool := + (sv6_privMax_hist_strict τ l init ∧ ¬ sv4_privMaxC τ l init) def sv6_loop (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : SLang ℕ := do - if (sv6_privMax_hist_strict τ l init ∧ ¬ sv4_privMaxC τ l init) + if (sv6_cond τ l init) then return point else probZero +-- QUESTION: What do we need for equality in the base case? +lemma sv5_sv6_loop_base_case (τ : ℤ) (l : List ℕ) (point eval : ℕ) (past future : List ℤ) (pres : ℤ) : + future = [] -> + List.length future = eval -> + List.length (past ++ [pres] ++ future) = point -> + True -> + (sv6_loop τ l eval ((past, pres), future)) point = (sv5_loop τ l eval ((past, pres), future)) point := by + intro Hfuture Heval Hstate _ + rw [Hfuture] + simp_all + rw [<- Heval] + + unfold sv5_loop + simp + unfold probWhileCut + unfold probWhileFunctional + split + · -- If (sv4_privMaxC τ l ((past, pres), []) = true) then + -- sv6_loop _ _ _ = 0 + simp [probWhileCut] + sorry + · -- If (sv4_privMaxC τ l ((past, pres), []) = false) then + -- (RHS = 0) + -- (sv6_loop τ l 0 (_, [])) point = 0 + simp + unfold sv4_state + unfold sv1_state + rw [ENNReal.tsum_eq_add_tsum_ite ((past, pres), [])] + conv => + lhs + rw [<- add_zero (sv6_loop _ _ _ _ _)] + rw [add_comm] + conv => + rhs + rw [add_comm] + congr 1 + · symm + simp + aesop + simp [sv1_threshold] + -- SUSPECT! + rw [<- Hstate] + simp + sorry + + + +-- QUESTION: What do we need for sv6_loop to be equal to sv6_loop_cond (next) +lemma sv6_loop_ind (τ : ℤ) (l : List ℕ) (eval point : ℕ) (past ff: List ℤ) (pres f: ℤ) : + (sv4_privMaxC τ l ((past, pres), f :: ff) = true) -> + True -> + (sv6_loop τ l eval ((past, pres), f :: ff)) point = (sv6_loop τ l eval ((past ++ [pres], f), ff)) point := by + intro Hcondition _ + unfold sv6_loop + suffices (sv6_cond τ l ((past, pres), f :: ff) = sv6_cond τ l ((past ++ [pres], f), ff)) by + split <;> split <;> try rfl + all_goals simp_all + sorry + + + +-- QUESTION: What do we need for sv5 to be equal to sv5_loop_cond (next) evaluated at point +lemma sv5_loop_ind (τ : ℤ) (l : List ℕ) (eval point : ℕ) (past ff: List ℤ) (pres f: ℤ) : + (sv4_privMaxC τ l ((past, pres), f :: ff) = true) -> + (sv5_loop τ l (eval + 1) ((past, pres), f :: ff)) point = (sv5_loop τ l eval ((past ++ [pres], f), ff)) point := by + intro Hcondition + conv => + lhs + unfold sv5_loop + unfold probWhileCut + unfold probWhileFunctional + split + · simp only [bind, pure, sv4_privMaxF, pure_bind] + unfold sv5_loop + rfl + · exfalso + trivial + + + + + def sv6_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => let computation : SLang ℕ := do @@ -1009,9 +1093,10 @@ def sv5_sv6_loop_eq_point (τ : ℤ) (l : List ℕ) (point eval_pt : ℕ) (past split · -- RHS check True (contradiction, would return 1) rename_i hhist - rcases hhist with ⟨ hhist, hlast' ⟩ - exfalso - simp_all + sorry + -- rcases hhist with ⟨ hhist, hlast' ⟩ + -- exfalso + -- simp_all · -- RHS check returns False (returns probZero) rename_i hhist simp at hhist @@ -1054,8 +1139,6 @@ def sv5_sv6_loop_eq_point (τ : ℤ) (l : List ℕ) (point eval_pt : ℕ) (past · -- Inductive case rename_i fut_next fut_rest IH - - sorry @@ -1067,14 +1150,22 @@ def sv5_sv6_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : intro eval_point simp apply tsum_congr - intro _ + intro τ congr apply funext - intro _ + intro v0 congr apply funext - intro _ + intro future congr + -- Now to prove Equality at the evaluation point + + have Heval : List.length (future : List ℤ) = eval_point := by + exact Mathlib.Vector.length_val future + + + + -- length (past + present + future) = point + 1 rw [sv5_sv6_loop_eq_point] all_goals sorry From 82bebccc999241f22699c1939188923ca8d8a932 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Sun, 3 Nov 2024 18:56:07 -0500 Subject: [PATCH 052/100] work on sv3 admits --- .../Queries/UnboundedMax/Properties.lean | 66 ++++++++++++++++++- 1 file changed, 64 insertions(+), 2 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index fd4d43cd..038f465b 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -605,6 +605,12 @@ lemma sv3_loop_unroll_1_alt [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) def len_list_append_rev {m n : ℕ} (x : { l : List ℤ // l.length = m }) (y: { l : List ℤ // l.length = n }) : { l : List ℤ // l.length = n + m } := ⟨ x.1 ++ y.1 , by simp [add_comm] ⟩ + + +lemma vector_sum_singleton (f : { l : List ℤ // l.length = 1 } -> ENNReal) (P : (x : ℤ) -> ([x].length = 1)) : + (∑'(x : { l // l.length = 1 }), f x) = (∑' (x : ℤ), f ⟨ [x], P x⟩) := by + sorry + def sv4_presample_split [DPSystem ℕ] (ε₁ ε₂ : ℕ+) (point : ℕ) : sv4_presample ε₁ ε₂ (point + 1) = (do @@ -614,9 +620,44 @@ def sv4_presample_split [DPSystem ℕ] (ε₁ ε₂ : ℕ+) (point : ℕ) : apply SLang.ext intro final_state simp [sv4_presample] + conv => + enter [1, 1, a] + rw [<- ENNReal.tsum_mul_left] + rw [← ENNReal.tsum_prod] + rw [vector_sum_singleton _ (by simp)] + + have X (x : ℤ): (@tsum.{0, 0} ENNReal _ ENNReal.instTopologicalSpace Int fun (x_1 : Int) => + @ite.{1} ENNReal (@Eq.{1} Int x_1 x) (Classical.propDecidable (@Eq.{1} Int x_1 x)) + (@OfNat.ofNat.{0} ENNReal 0 (@Zero.toOfNat0.{0} ENNReal instENNRealZero)) + (@ite.{1} ENNReal (@Eq.{1} Int x x_1) (Int.instDecidableEq x x_1) + (@SLang.privNoiseZero _ ε₁ + (@HMul.hMul.{0, 0, 0} PNat PNat PNat (@instHMul.{0} PNat instPNatMul) + (@OfNat.ofNat.{0} PNat 4 (@instOfNatPNatOfNeZeroNat 4 SLang.sv4_presample.proof_5)) ε₂) + x_1) + 0)) = 0 := by + simp + aesop + conv => + enter [2, 1, x, 1] + simp + rw [ENNReal.tsum_eq_add_tsum_ite x] + simp + enter [2] + skip + rw [X] + clear X + simp + conv => + enter [2, 1, x] + rw [<- ENNReal.tsum_mul_left] + rw [← ENNReal.tsum_prod] + simp_all [len_list_append_rev] + -- Can do this by bijection (not congruence) -- By bijection -- #check tsum_eq_tsum_of_ne_zero_bij + -- rw [ENNReal.tsum_eq_add_tsum_ite (Int.ofNat point)] + sorry @@ -631,7 +672,6 @@ def len_1_list_to_val (x : { l : List ℤ // l.length = 1 }) : ℤ := -- Unfolding this loop just moves the first presample into init -- Which can apply the IH-- since it's some arbitrary init state and a presamples state generated by one fewer point - def sv3_sv4_loop_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : @sv3_loop dps ε₁ ε₂ τ l point init = @sv4_loop dps ε₁ ε₂ τ l point init := by revert init @@ -717,8 +757,30 @@ def sv3_sv4_loop_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : Li intro final_state simp -- There is a bijection here - sorry + rw [vector_sum_singleton _ (by simp)] + apply tsum_congr + intro x + simp [len_1_list_to_val] + simp [sv4_presample] + rw [ENNReal.tsum_eq_add_tsum_ite x] + simp + have X : ( @tsum.{0, 0} ENNReal _ + ENNReal.instTopologicalSpace Int fun (x_1 : Int) => + @ite.{1} ENNReal (@Eq.{1} Int x_1 x) (Classical.propDecidable (@Eq.{1} Int x_1 x)) + (@OfNat.ofNat.{0} ENNReal 0 (@Zero.toOfNat0.{0} ENNReal instENNRealZero)) + (@ite.{1} ENNReal (@Eq.{1} Int x x_1) (Int.instDecidableEq x x_1) + (@SLang.privNoiseZero dps ε₁ + (@HMul.hMul.{0, 0, 0} PNat PNat PNat (@instHMul.{0} PNat instPNatMul) + (@OfNat.ofNat.{0} PNat 4 (@instOfNatPNatOfNeZeroNat 4 SLang.sv4_presample.proof_5)) ε₂) + x_1) + 0)) = (0 : ENNReal) := by + simp + aesop + conv => + enter [2, 1, 2] + rw [X] + simp rw [ToPresample] clear ToPresample From 93ec8f845aacad144b11fb54132aed1df8e13a82 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Sun, 3 Nov 2024 19:54:40 -0500 Subject: [PATCH 053/100] so close --- .../Queries/UnboundedMax/Properties.lean | 178 ++++++------------ 1 file changed, 62 insertions(+), 116 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 038f465b..7364ef0a 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -979,8 +979,6 @@ lemma sv6_privMax_hist_step (τ : ℤ) (l : List ℕ) (past fut_rest : List ℤ) simp_all - - def is_past_configuration_strict (sp sc : sv4_state) : Prop := (sp.1.1.length < sc.1.1.length) ∧ sp.1.1 ++ [sp.1.2] ++ sp.2 = sc.1.1 ++ [sc.1.2] ++ sc.2 @@ -1021,75 +1019,69 @@ lemma sv6_privMax_hist_step_strict (τ : ℤ) (l : List ℕ) (past fut_rest : Li simp_all +@[simp] +def sv6_cond_rec (τ : ℤ) (l : List ℕ) (past : List ℤ) (pres : ℤ) (future : List ℤ) : Bool := + match future with + | [] => False + | (f :: ff) => (sv4_privMaxC τ l ((past, pres), f :: ff) = true) && (sv6_cond_rec τ l (past ++ [pres]) f ff) + @[simp] def sv6_cond (τ : ℤ) (l : List ℕ) (init : sv4_state) : Bool := - (sv6_privMax_hist_strict τ l init ∧ ¬ sv4_privMaxC τ l init) + sv6_cond_rec τ l init.1.1 init.1.2 init.2 def sv6_loop (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : SLang ℕ := do if (sv6_cond τ l init) then return point else probZero +lemma sv6_loop_base (τ : ℤ) (l : List ℕ) (point: ℕ) (past: List ℤ) (pres : ℤ) : + (sv6_loop τ l 0 ((past, pres), [])) point = 0 := by + simp [sv6_loop] -- QUESTION: What do we need for equality in the base case? lemma sv5_sv6_loop_base_case (τ : ℤ) (l : List ℕ) (point eval : ℕ) (past future : List ℤ) (pres : ℤ) : future = [] -> List.length future = eval -> List.length (past ++ [pres] ++ future) = point -> - True -> - (sv6_loop τ l eval ((past, pres), future)) point = (sv5_loop τ l eval ((past, pres), future)) point := by - intro Hfuture Heval Hstate _ + (sv6_loop τ l point ((past, pres), future)) point = (sv5_loop τ l eval ((past, pres), future)) point := by + intro Hfuture Heval Hstate rw [Hfuture] simp_all rw [<- Heval] - + simp [sv6_loop] unfold sv5_loop - simp unfold probWhileCut unfold probWhileFunctional split - · -- If (sv4_privMaxC τ l ((past, pres), []) = true) then - -- sv6_loop _ _ _ = 0 - simp [probWhileCut] - sorry - · -- If (sv4_privMaxC τ l ((past, pres), []) = false) then - -- (RHS = 0) - -- (sv6_loop τ l 0 (_, [])) point = 0 - simp + · simp [probWhileCut] + · simp unfold sv4_state unfold sv1_state rw [ENNReal.tsum_eq_add_tsum_ite ((past, pres), [])] - conv => - lhs - rw [<- add_zero (sv6_loop _ _ _ _ _)] - rw [add_comm] - conv => - rhs - rw [add_comm] - congr 1 - · symm - simp + symm + simp only [add_eq_zero] + apply And.intro + · simp [sv1_threshold] + linarith + · simp aesop - simp [sv1_threshold] - -- SUSPECT! - rw [<- Hstate] - simp - sorry - -- QUESTION: What do we need for sv6_loop to be equal to sv6_loop_cond (next) -lemma sv6_loop_ind (τ : ℤ) (l : List ℕ) (eval point : ℕ) (past ff: List ℤ) (pres f: ℤ) : +lemma sv6_loop_ind (τ : ℤ) (l : List ℕ) (point : ℕ) (past ff: List ℤ) (pres f: ℤ) : (sv4_privMaxC τ l ((past, pres), f :: ff) = true) -> - True -> - (sv6_loop τ l eval ((past, pres), f :: ff)) point = (sv6_loop τ l eval ((past ++ [pres], f), ff)) point := by + List.length (past ++ [pres] ++ f :: ff) = point -> + (sv6_loop τ l point ((past, pres), f :: ff)) point = (sv6_loop τ l point ((past ++ [pres], f), ff)) point := by intro Hcondition _ unfold sv6_loop suffices (sv6_cond τ l ((past, pres), f :: ff) = sv6_cond τ l ((past ++ [pres], f), ff)) by split <;> split <;> try rfl all_goals simp_all - sorry - + conv => + lhs + unfold sv6_cond + simp + simp [Hcondition] -- QUESTION: What do we need for sv5 to be equal to sv5_loop_cond (next) evaluated at point @@ -1110,9 +1102,6 @@ lemma sv5_loop_ind (τ : ℤ) (l : List ℕ) (eval point : ℕ) (past ff: List trivial - - - def sv6_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => let computation : SLang ℕ := do @@ -1125,83 +1114,40 @@ def sv6_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang -- sv6_loop and sv5_loop are equal at point (under some conditions) -def sv5_sv6_loop_eq_point (τ : ℤ) (l : List ℕ) (point eval_pt : ℕ) (past future : List ℤ) (pres : ℤ) : - List.length (past ++ [pres] ++ future) = (point + 1) -> - List.length future = eval_point -> - sv6_privMax_hist_strict τ l ((past, pres), future) -> - @sv5_loop τ l eval_point ((past, pres), future) point = @sv6_loop τ l point ((past, pres), future) point := by - - -- To do induction over future, the length of past plus the number of steps remaining shoud be constant - - revert eval_point past pres +def sv5_sv6_loop_eq_point (τ : ℤ) (l : List ℕ) (point eval : ℕ) (past future : List ℤ) (pres : ℤ) : + List.length (past ++ [pres] ++ future) = point -> + List.length future = eval -> + -- sv6_privMax_hist_strict τ l ((past, pres), future) -> + @sv5_loop τ l eval ((past, pres), future) point = @sv6_loop τ l point ((past, pres), future) point := by + revert past pres eval induction future - · -- Base case: no future - intro eval_point past pres Hstate Heval Hpast_evals - simp_all - rw [<- Heval] + · intro eval past pres H1 H2 + symm + apply (sv5_sv6_loop_base_case _ _ _ _ _ _ _ (by rfl) H2 H1) + · rename_i f ff IH + intro eval past pres Hstate Heval + cases eval + · simp at Heval + + rename_i eval + cases (Classical.em (sv4_privMaxC τ l ((past, pres), f :: ff) = true)) + · rename_i Hcondition + rw [sv5_loop_ind _ _ _ _ _ _ _ _ Hcondition] + rw [sv6_loop_ind _ _ _ _ _ _ _ Hcondition Hstate] + apply (IH eval (past ++ [pres]) f ?G1 ?G2) + case G1 => simp_all + case G2 => simp_all + + rename_i Hcondition + simp at Hcondition + simp [sv6_loop, Hcondition] unfold sv5_loop unfold probWhileCut unfold probWhileFunctional - - -- LHS: We have consumed the entire future now, - -- If the last conditional check is False, return pure ... - -- else return zero - - -- RHS: Should be able to simplify now using Hpast_evals - unfold sv6_loop - split - · -- LHS when the last loop condition returns True (LHS returns 0) - rename_i hlast - split - · -- RHS check True (contradiction, would return 1) - rename_i hhist - sorry - -- rcases hhist with ⟨ hhist, hlast' ⟩ - -- exfalso - -- simp_all - · -- RHS check returns False (returns probZero) - rename_i hhist - simp at hhist - simp - intro _ _ _ - simp_all - right - rw [<- Heval] - simp [probWhileCut] - - · -- LHS when the last loop condition returns False (should return 1?) - rename_i hlast - simp at hlast - conv => - lhs - simp [sv1_threshold] - unfold sv4_state - unfold sv1_state - rw [(ENNReal.tsum_eq_add_tsum_ite ((past, pres), []))] - conv => - lhs - rw [add_comm] - conv => - rhs - rw [<- zero_add (@ite _ _ _ _ _ )] - congr 1 - · simp - simp_all - rw [<- Hstate] - simp - - split - · -- RHS evaluates to True - simp - · -- RHS evaluates to False (contradiction) - exfalso - rename_i h - simp at h - aesop - - · -- Inductive case - rename_i fut_next fut_rest IH - sorry + rw [Hcondition] + simp + intro i Hi Hk + simp_all [sv1_threshold] def sv5_sv6_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : @@ -1225,11 +1171,11 @@ def sv5_sv6_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : have Heval : List.length (future : List ℤ) = eval_point := by exact Mathlib.Vector.length_val future - - -- length (past + present + future) = point + 1 rw [sv5_sv6_loop_eq_point] - all_goals sorry + · simp + sorry + · sorry From 267c61f5c5b33092b398398251c67cbbdc0a32b0 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Sun, 3 Nov 2024 20:05:36 -0500 Subject: [PATCH 054/100] sv5 sv6 victory --- .../Queries/UnboundedMax/Properties.lean | 53 +++++++++---------- 1 file changed, 26 insertions(+), 27 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 7364ef0a..26922c5f 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -1022,7 +1022,7 @@ lemma sv6_privMax_hist_step_strict (τ : ℤ) (l : List ℕ) (past fut_rest : Li @[simp] def sv6_cond_rec (τ : ℤ) (l : List ℕ) (past : List ℤ) (pres : ℤ) (future : List ℤ) : Bool := match future with - | [] => False + | [] => ¬ (sv4_privMaxC τ l ((past, pres), [])) | (f :: ff) => (sv4_privMaxC τ l ((past, pres), f :: ff) = true) && (sv6_cond_rec τ l (past ++ [pres]) f ff) @[simp] @@ -1034,43 +1034,45 @@ def sv6_loop (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : SLang then return point else probZero -lemma sv6_loop_base (τ : ℤ) (l : List ℕ) (point: ℕ) (past: List ℤ) (pres : ℤ) : - (sv6_loop τ l 0 ((past, pres), [])) point = 0 := by - simp [sv6_loop] - -- QUESTION: What do we need for equality in the base case? lemma sv5_sv6_loop_base_case (τ : ℤ) (l : List ℕ) (point eval : ℕ) (past future : List ℤ) (pres : ℤ) : future = [] -> List.length future = eval -> - List.length (past ++ [pres] ++ future) = point -> + List.length (past ++ [pres] ++ future) = point + 1 -> (sv6_loop τ l point ((past, pres), future)) point = (sv5_loop τ l eval ((past, pres), future)) point := by intro Hfuture Heval Hstate rw [Hfuture] simp_all rw [<- Heval] - simp [sv6_loop] unfold sv5_loop unfold probWhileCut unfold probWhileFunctional split - · simp [probWhileCut] - · simp + · rename_i h + simp [probWhileCut, sv6_loop] + rw [h] + simp + · rename_i h + simp at h + simp [sv6_loop] + simp [h] unfold sv4_state unfold sv1_state rw [ENNReal.tsum_eq_add_tsum_ite ((past, pres), [])] + simp [sv1_threshold] + simp [Hstate] + conv => + lhs + rw [<- add_zero 1] + congr symm - simp only [add_eq_zero] - apply And.intro - · simp [sv1_threshold] - linarith - · simp - aesop - + simp + aesop -- QUESTION: What do we need for sv6_loop to be equal to sv6_loop_cond (next) lemma sv6_loop_ind (τ : ℤ) (l : List ℕ) (point : ℕ) (past ff: List ℤ) (pres f: ℤ) : (sv4_privMaxC τ l ((past, pres), f :: ff) = true) -> - List.length (past ++ [pres] ++ f :: ff) = point -> + List.length (past ++ [pres] ++ f :: ff) = point + 1 -> (sv6_loop τ l point ((past, pres), f :: ff)) point = (sv6_loop τ l point ((past ++ [pres], f), ff)) point := by intro Hcondition _ unfold sv6_loop @@ -1115,7 +1117,7 @@ def sv6_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang -- sv6_loop and sv5_loop are equal at point (under some conditions) def sv5_sv6_loop_eq_point (τ : ℤ) (l : List ℕ) (point eval : ℕ) (past future : List ℤ) (pres : ℤ) : - List.length (past ++ [pres] ++ future) = point -> + List.length (past ++ [pres] ++ future) = point + 1 -> List.length future = eval -> -- sv6_privMax_hist_strict τ l ((past, pres), future) -> @sv5_loop τ l eval ((past, pres), future) point = @sv6_loop τ l point ((past, pres), future) point := by @@ -1123,7 +1125,11 @@ def sv5_sv6_loop_eq_point (τ : ℤ) (l : List ℕ) (point eval : ℕ) (past fut induction future · intro eval past pres H1 H2 symm - apply (sv5_sv6_loop_base_case _ _ _ _ _ _ _ (by rfl) H2 H1) + simp at H1 + apply (sv5_sv6_loop_base_case _ _ _ _ _ _ _ (by rfl) H2 ?G2) + case G2 => + simp + trivial · rename_i f ff IH intro eval past pres Hstate Heval cases eval @@ -1166,16 +1172,9 @@ def sv5_sv6_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : apply funext intro future congr - -- Now to prove Equality at the evaluation point - - have Heval : List.length (future : List ℤ) = eval_point := by - exact Mathlib.Vector.length_val future - - -- length (past + present + future) = point + 1 rw [sv5_sv6_loop_eq_point] · simp - sorry - · sorry + · exact Mathlib.Vector.length_val future From 6e17ff4619a6bf2bb5ba6cb572aa63385402d011 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Sun, 3 Nov 2024 22:02:50 -0500 Subject: [PATCH 055/100] sketch sv7 --- .../Queries/UnboundedMax/Properties.lean | 138 ++++++++++++++++-- 1 file changed, 125 insertions(+), 13 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 26922c5f..82f59f2f 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -1,7 +1,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 +Authors: Markus de Medeiros -/ import SampCert.DifferentialPrivacy.Abstract @@ -1178,19 +1178,69 @@ def sv5_sv6_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : +-- To prepare for v8, we need to separate out the point=0 case, and move the last sample +-- to the end --- What am I defining: --- When I step sv5 (the probWhileCut), unrolling a loop iterate pops the first value out of --- future, and pushes it onto the initial condition. --- This means I want to do induction over future. --- Therefore, I need to establish what it means for a loop body with some of these pops to have --- not terminated (??) --- IE. given a state, I want to be able to express the fact that --- - Every loop check which moved something to the past must have returned false --- - We are waiting for n loop checks from the current state to return true --- For the latter, I define sv6_privMax_check --- For the former, I define sc6_provMax_hist --- Should be able to establish these inductively based on hist, assuming the first n-1 loop checks are valid +def sv7_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := + fun (point : ℕ) => + let computation : SLang ℕ := do + let τ <- @privNoiseZero dps ε₁ (2 * ε₂) + let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + match point with + | 0 => + if (¬ (sv4_privMaxC τ l (([], v0), []))) + then probPure point + else probZero + | (Nat.succ point') => do + let presamples <- @sv4_presample dps ε₁ ε₂ point' + let vlast <- @privNoiseZero dps ε₁ (4 * ε₂) + if (sv6_cond τ l (([], v0), presamples ++ [vlast])) + then probPure point + else probZero + computation point + + +def sv6_sv7_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : + @sv6_privMax dps ε₁ ε₂ l = @sv7_privMax dps ε₁ ε₂ l := by + apply SLang.ext + intro point + unfold sv6_privMax + unfold sv7_privMax + cases point + · simp [sv6_loop, sv6_cond] + apply tsum_congr + intro τ + congr 1 + apply tsum_congr + intro v0 + congr 1 + simp [sv4_presample] + rw [ENNReal.tsum_eq_add_tsum_ite ⟨[], rfl⟩] + simp + conv => + rhs + rw [<- (add_zero (@ite _ _ _ _ _))] + congr 1 + simp + · rename_i point' + simp only [] + apply tsum_congr + intro τ + congr 1 + apply tsum_congr + intro v0 + congr 1 + simp + conv => + enter [2, 1, a] + rw [<- ENNReal.tsum_mul_left] + conv => + lhs + unfold sv6_loop + simp + rw [← ENNReal.tsum_prod] + -- There is a bijection here + sorry @@ -1199,9 +1249,71 @@ def sv5_sv6_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : +-- /- +-- ## Program version 8 +-- +-- Define G from the paper +-- -/ +-- +-- def sv8_sum (l : List ℕ) (past : List ℤ) (pres : ℤ) : ℤ := exactDiffSum (List.length past) l + pres +-- +-- -- G is only defined for k > 0 +-- def sv7_G (l : List ℕ) (past : List ℤ) (pres next : ℤ) (future : List ℤ) : ℤ := +-- match future with +-- | [] => sv7_sum l past pres +-- | (f :: ff) => max (sv7_sum l past pres) (sv7_G l (past ++ [pres]) next f ff) +-- +-- def sv7_cond (τ : ℤ) (l : List ℕ) : Bool := +-- sv7_cond_rec τ l init.1.1 init.1.2 init.2 + + + + +-- def sv7_loop (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : SLang ℕ := do +-- if (sv7_cond τ l init) +-- then return point +-- else probZero +-- +-- +-- def sv7_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := +-- fun (point : ℕ) => +-- let computation : SLang ℕ := do +-- let τ <- @privNoiseZero dps ε₁ (2 * ε₂) +-- let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) +-- let presamples <- @sv4_presample dps ε₁ ε₂ point +-- @sv7_loop τ l point (([], v0), presamples) +-- computation point + + + + + + + + +-- lemma sv4_privMaxC_mono (τ : ℤ) (l : List ℕ) (past : List ℤ) (pres next : ℤ) (future : List ℤ) : +-- sv4_privMaxC τ l ((past, pres), next :: future) = false -> +-- sv4_privMaxC τ l ((past ++ [pres], next), future) = false := by +-- unfold sv4_privMaxC +-- unfold sv1_privMaxC +-- simp +-- intro H +-- apply (le_trans H) +-- clear H +-- simp [sv1_threshold] +-- sorry + +-- -- If it's false at any point (except the very end), the sv7 loop is false +-- lemma sv4_privMaxC_mono_lim (τ : ℤ) (l : List ℕ) (past : List ℤ) (future : List ℤ) (first : List ℤ) (last : ℤ) : +-- future = first ++ [last] -> +-- sv4_privMaxC τ l ((past, pres), next :: future) = false -> +-- sv6_cond_rec τ l past pres future = false := by +-- sorry + + From 5f3e10ec6ec4d9f18d2f2903c07f196ee347fc42 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Sun, 3 Nov 2024 22:35:24 -0500 Subject: [PATCH 056/100] sv8 (definition of G from the paper) --- .../Queries/UnboundedMax/Properties.lean | 140 ++++++++++-------- 1 file changed, 77 insertions(+), 63 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 82f59f2f..99523418 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -1177,9 +1177,11 @@ def sv5_sv6_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : · exact Mathlib.Vector.length_val future - --- To prepare for v8, we need to separate out the point=0 case, and move the last sample --- to the end +/- +## Program v8 +Not executable +Separates out the zero case +-/ def sv7_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => @@ -1193,8 +1195,8 @@ def sv7_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang else probZero | (Nat.succ point') => do let presamples <- @sv4_presample dps ε₁ ε₂ point' - let vlast <- @privNoiseZero dps ε₁ (4 * ε₂) - if (sv6_cond τ l (([], v0), presamples ++ [vlast])) + let vk <- @privNoiseZero dps ε₁ (4 * ε₂) + if (sv6_cond τ l (([], v0), presamples ++ [vk])) then probPure point else probZero computation point @@ -1243,77 +1245,89 @@ def sv6_sv7_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : sorry +/- +## Program v8 +Not executable +Defines G from the paper +-/ +def sv8_sum (l : List ℕ) (past : List ℤ) (pres : ℤ) : ℤ := exactDiffSum (List.length past) l + pres +-- G is the maximum value of sv8_sum over the tape +def sv8_G (l : List ℕ) (past : List ℤ) (pres : ℤ) (future : List ℤ) : ℤ := + match future with + | [] => sv8_sum l past pres + | (f :: ff) => max (sv8_sum l past pres) (sv8_G l (past ++ [pres]) f ff) --- /- --- ## Program version 8 --- --- Define G from the paper --- -/ --- --- def sv8_sum (l : List ℕ) (past : List ℤ) (pres : ℤ) : ℤ := exactDiffSum (List.length past) l + pres --- --- -- G is only defined for k > 0 --- def sv7_G (l : List ℕ) (past : List ℤ) (pres next : ℤ) (future : List ℤ) : ℤ := --- match future with --- | [] => sv7_sum l past pres --- | (f :: ff) => max (sv7_sum l past pres) (sv7_G l (past ++ [pres]) next f ff) --- --- def sv7_cond (τ : ℤ) (l : List ℕ) : Bool := --- sv7_cond_rec τ l init.1.1 init.1.2 init.2 - - - - --- def sv7_loop (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : SLang ℕ := do --- if (sv7_cond τ l init) --- then return point --- else probZero --- --- --- def sv7_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := --- fun (point : ℕ) => --- let computation : SLang ℕ := do --- let τ <- @privNoiseZero dps ε₁ (2 * ε₂) --- let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) --- let presamples <- @sv4_presample dps ε₁ ε₂ point --- @sv7_loop τ l point (([], v0), presamples) --- computation point - - - - - - + def sv8_cond (τ : ℤ) (l : List ℕ) (past : List ℤ) (pres : ℤ) (future : List ℤ) (last : ℤ) : Bool := + (sv8_G l past pres future < τ) ∧ (sv8_sum l (past ++ [pres] ++ future) last ≥ τ) +def sv8_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := + fun (point : ℕ) => + let computation : SLang ℕ := do + let τ <- @privNoiseZero dps ε₁ (2 * ε₂) + let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + match point with + | 0 => + if (sv8_sum l [] v0 ≥ τ) + then probPure point + else probZero + | (Nat.succ point') => do + let presamples <- @sv4_presample dps ε₁ ε₂ point' + let vk <- @privNoiseZero dps ε₁ (4 * ε₂) + if (sv8_cond τ l [] v0 presamples vk) + then probPure point + else probZero + computation point --- lemma sv4_privMaxC_mono (τ : ℤ) (l : List ℕ) (past : List ℤ) (pres next : ℤ) (future : List ℤ) : --- sv4_privMaxC τ l ((past, pres), next :: future) = false -> --- sv4_privMaxC τ l ((past ++ [pres], next), future) = false := by --- unfold sv4_privMaxC --- unfold sv1_privMaxC --- simp --- intro H --- apply (le_trans H) --- clear H --- simp [sv1_threshold] --- sorry --- -- If it's false at any point (except the very end), the sv7 loop is false --- lemma sv4_privMaxC_mono_lim (τ : ℤ) (l : List ℕ) (past : List ℤ) (future : List ℤ) (first : List ℤ) (last : ℤ) : --- future = first ++ [last] -> --- sv4_privMaxC τ l ((past, pres), next :: future) = false -> --- sv6_cond_rec τ l past pres future = false := by --- sorry +lemma sv7_sv8_cond_eq (τ : ℤ) (l : List ℕ) (v0 : ℤ) (vs : List ℤ) (vk : ℤ) : + sv8_cond τ l [] v0 vs vk = sv6_cond τ l (([], v0), vs ++ [vk]) := by + suffices (∀ init, sv8_cond τ l init v0 vs vk = sv6_cond τ l ((init, v0), vs ++ [vk])) by + apply this + revert v0 + unfold sv8_cond + simp + induction vs + · intro v0 init + simp [sv6_cond_rec, sv4_privMaxC, sv1_privMaxC, sv1_threshold, sv1_noise, List.length] + simp [sv8_G, sv8_sum] + · intro vi init + rename_i vi_1 rest IH + have IH' := IH vi_1 (init ++ [vi]) + simp at IH' + clear IH + conv => + rhs + simp [sv6_cond_rec] + rw [<- IH'] + clear IH' + cases (decide (τ ≤ sv8_sum l (init ++ vi :: vi_1 :: rest) vk)) <;> simp + conv => + lhs + unfold sv8_G + simp + cases (decide (sv8_G l (init ++ [vi]) vi_1 rest < τ)) <;> simp + simp [sv4_privMaxC, sv1_privMaxC, sv8_sum, sv1_threshold, sv1_noise] +def sv7_sv8_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : + @sv7_privMax dps ε₁ ε₂ l = @sv8_privMax dps ε₁ ε₂ l := by + apply SLang.ext + intro point + unfold sv7_privMax + unfold sv8_privMax + cases point + · simp [sv4_privMaxC, sv1_privMaxC, sv8_sum, sv1_threshold, sv1_noise] + rename_i point' + simp only [] + repeat (apply tsum_congr; intro _; congr 1) + simp only [sv7_sv8_cond_eq, sv6_cond] - From b86c6e03efee0af0092b94caeabfcb02c3ba2c58 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 4 Nov 2024 09:52:39 -0500 Subject: [PATCH 057/100] cleanup --- .../Queries/UnboundedMax/Properties.lean | 431 +++--------------- 1 file changed, 62 insertions(+), 369 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 99523418..33d024ab 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -66,6 +66,12 @@ This looks strange, but will specialize to Lap(ε₁/ε₂, 0) in the pure DP ca -/ def privNoiseZero [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) : SLang ℤ := dps.noise (fun _ => 0) 1 ε₁ ε₂ [] + +def privNoiseGuess [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) : SLang ℤ := @privNoiseZero dps ε₁ (4 * ε₂) + +def privNoiseThresh [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) : SLang ℤ := @privNoiseZero dps ε₁ (2 * ε₂) + + /- Not used for anything, but to give confidence in our definitions @@ -128,13 +134,13 @@ def sv0_privMaxC (τ : ℤ) (l : List ℕ) (s : sv0_state) : Bool := decide (exactDiffSum (sv0_threshold s) l + (sv0_noise s) < τ) def sv0_privMaxF [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (s : sv0_state) : SLang sv0_state := do - let vn <- @privNoiseZero dps ε₁ (4 * ε₂) + let vn <- @privNoiseGuess dps ε₁ ε₂ let n := (sv0_threshold s) + 1 return (n, vn) def sv0_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do - let τ <- @privNoiseZero dps ε₁ (2 * ε₂) - let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + let τ <- @privNoiseThresh dps ε₁ ε₂ + let v0 <- @privNoiseGuess dps ε₁ ε₂ let sk <- probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (0, v0) return (sv0_threshold sk) @@ -154,12 +160,12 @@ def sv1_privMaxC (τ : ℤ) (l : List ℕ) (s : sv1_state) : Bool := decide (exactDiffSum (sv1_threshold s) l + (sv1_noise s) < τ) def sv1_privMaxF [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (s : sv1_state) : SLang sv1_state := do - let vn <- @privNoiseZero dps ε₁ (4 * ε₂) + let vn <- @privNoiseGuess dps ε₁ ε₂ return (s.1 ++ [s.2], vn) def sv1_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do - let τ <- @privNoiseZero dps ε₁ (2 * ε₂) - let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + let τ <- @privNoiseThresh dps ε₁ ε₂ + let v0 <- @privNoiseGuess dps ε₁ ε₂ let sk <- probWhile (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) ([], v0) return (sv1_threshold sk) @@ -225,8 +231,8 @@ lemma sv0_eq_sv1 [dps : DPSystem ℕ] ε₁ ε₂ l : sv0_privMax ε₁ ε₂ l def sv2_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => let computation : SLang ℕ := do - let τ <- @privNoiseZero dps ε₁ (2 * ε₂) - let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + let τ <- @privNoiseThresh dps ε₁ ε₂ + let v0 <- @privNoiseGuess dps ε₁ ε₂ let sk <- probWhile (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) ([], v0) return (sv1_threshold sk) computation point @@ -250,8 +256,8 @@ def sv3_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) def sv3_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => let computation : SLang ℕ := do - let τ <- @privNoiseZero dps ε₁ (2 * ε₂) - let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + let τ <- @privNoiseThresh dps ε₁ ε₂ + let v0 <- @privNoiseGuess dps ε₁ ε₂ let sk <- @sv3_loop dps ε₁ ε₂ τ l point ([], v0) return (sv1_threshold sk) computation point @@ -514,7 +520,7 @@ lemma sv2_eq_sv3 [dps : DPSystem ℕ] ε₁ ε₂ l : sv2_privMax ε₁ ε₂ l lemma sv3_loop_unroll_1 [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l point L vk : sv3_loop ε₁ ε₂ τ l (point + 1) (L, vk) = (do - let vk1 <- @privNoiseZero dps ε₁ (4 * ε₂) + let vk1 <- @privNoiseGuess dps ε₁ ε₂ if (sv1_privMaxC τ l (L, vk)) then (sv3_loop ε₁ ε₂ τ l point (L ++ [vk], vk1)) else probPure (L, vk)) := by @@ -559,6 +565,7 @@ lemma sv3_loop_unroll_1 [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l p intro ⟨ HF, vkf ⟩ simp [probBind] split <;> try simp + unfold privNoiseGuess unfold privNoiseZero exact Eq.symm (PMF.tsum_coe (DPSystem.noise (fun _ => 0) 1 ε₁ (4 * ε₂) [])) @@ -575,7 +582,7 @@ def sv4_presample [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (n : ℕ) : SLang { match n with | Nat.zero => return ⟨ [], by simp ⟩ | Nat.succ n' => do - let vk1 <- @privNoiseZero dps ε₁ (4 * ε₂) + let vk1 <- @privNoiseGuess dps ε₁ ε₂ let vks <- @sv4_presample dps ε₁ ε₂ n' return ⟨ vks ++ [vk1], by simp ⟩ @@ -595,7 +602,7 @@ def sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) lemma sv3_loop_unroll_1_alt [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l point (initial_state : sv1_state) : sv3_loop ε₁ ε₂ τ l (point + 1) initial_state = (do - let vk1 <- @privNoiseZero dps ε₁ (4 * ε₂) + let vk1 <- @privNoiseGuess dps ε₁ ε₂ if (sv1_privMaxC τ l initial_state) then (sv3_loop ε₁ ε₂ τ l point (initial_state.1 ++ [initial_state.2], vk1)) else probPure initial_state) := by @@ -606,7 +613,6 @@ def len_list_append_rev {m n : ℕ} (x : { l : List ℤ // l.length = m }) (y: { ⟨ x.1 ++ y.1 , by simp [add_comm] ⟩ - lemma vector_sum_singleton (f : { l : List ℤ // l.length = 1 } -> ENNReal) (P : (x : ℤ) -> ([x].length = 1)) : (∑'(x : { l // l.length = 1 }), f x) = (∑' (x : ℤ), f ⟨ [x], P x⟩) := by sorry @@ -626,24 +632,25 @@ def sv4_presample_split [DPSystem ℕ] (ε₁ ε₂ : ℕ+) (point : ℕ) : rw [← ENNReal.tsum_prod] rw [vector_sum_singleton _ (by simp)] - have X (x : ℤ): (@tsum.{0, 0} ENNReal _ ENNReal.instTopologicalSpace Int fun (x_1 : Int) => - @ite.{1} ENNReal (@Eq.{1} Int x_1 x) (Classical.propDecidable (@Eq.{1} Int x_1 x)) - (@OfNat.ofNat.{0} ENNReal 0 (@Zero.toOfNat0.{0} ENNReal instENNRealZero)) - (@ite.{1} ENNReal (@Eq.{1} Int x x_1) (Int.instDecidableEq x x_1) - (@SLang.privNoiseZero _ ε₁ - (@HMul.hMul.{0, 0, 0} PNat PNat PNat (@instHMul.{0} PNat instPNatMul) - (@OfNat.ofNat.{0} PNat 4 (@instOfNatPNatOfNeZeroNat 4 SLang.sv4_presample.proof_5)) ε₂) - x_1) - 0)) = 0 := by - simp - aesop + have X (x : ℤ): (@tsum.{0, 0} ENNReal + (@NonUnitalNonAssocSemiring.toAddCommMonoid.{0} ENNReal + (@NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} ENNReal + (@Semiring.toNonAssocSemiring.{0} ENNReal + (@OrderedSemiring.toSemiring.{0} ENNReal + (@OrderedCommSemiring.toOrderedSemiring.{0} ENNReal + (@CanonicallyOrderedCommSemiring.toOrderedCommSemiring.{0} ENNReal + ENNReal.instCanonicallyOrderedCommSemiring)))))) + ENNReal.instTopologicalSpace Int fun (x_1 : Int) => + @ite.{1} ENNReal (@Eq.{1} Int x_1 x) (Classical.propDecidable (@Eq.{1} Int x_1 x)) + (@OfNat.ofNat.{0} ENNReal 0 (@Zero.toOfNat0.{0} ENNReal instENNRealZero)) + (@ite.{1} ENNReal (@Eq.{1} Int x x_1) (Int.instDecidableEq x x_1) (@SLang.privNoiseGuess _ ε₁ ε₂ x_1) + 0)) = 0 := by simp; aesop conv => enter [2, 1, x, 1] simp rw [ENNReal.tsum_eq_add_tsum_ite x] simp enter [2] - skip rw [X] clear X simp @@ -723,12 +730,12 @@ def sv3_sv4_loop_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : Li -- Doing it this way because I can't conv under the @ite? let ApplyIH : (do - let vk1 ← privNoiseZero ε₁ (4 * ε₂) + let vk1 ← privNoiseGuess ε₁ ε₂ if sv1_privMaxC τ l init = true then sv3_loop ε₁ ε₂ τ l point (init.1 ++ [init.2], vk1) else probPure init) = (do - let vk1 ← privNoiseZero ε₁ (4 * ε₂) + let vk1 ← privNoiseGuess ε₁ ε₂ if sv1_privMaxC τ l init = true then sv4_loop ε₁ ε₂ τ l point (init.1 ++ [init.2], vk1) else probPure init) := by @@ -747,7 +754,7 @@ def sv3_sv4_loop_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : Li have ToPresample : (do - let vk1 ← privNoiseZero ε₁ (4 * ε₂) + let vk1 ← privNoiseGuess ε₁ ε₂ if sv1_privMaxC τ l init = true then sv4_loop ε₁ ε₂ τ l point (init.1 ++ [init.2], vk1) else probPure init) = (do let vps ← sv4_presample ε₁ ε₂ 1 @@ -765,18 +772,22 @@ def sv3_sv4_loop_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : Li rw [ENNReal.tsum_eq_add_tsum_ite x] simp - have X : ( @tsum.{0, 0} ENNReal _ - ENNReal.instTopologicalSpace Int fun (x_1 : Int) => - @ite.{1} ENNReal (@Eq.{1} Int x_1 x) (Classical.propDecidable (@Eq.{1} Int x_1 x)) - (@OfNat.ofNat.{0} ENNReal 0 (@Zero.toOfNat0.{0} ENNReal instENNRealZero)) - (@ite.{1} ENNReal (@Eq.{1} Int x x_1) (Int.instDecidableEq x x_1) - (@SLang.privNoiseZero dps ε₁ - (@HMul.hMul.{0, 0, 0} PNat PNat PNat (@instHMul.{0} PNat instPNatMul) - (@OfNat.ofNat.{0} PNat 4 (@instOfNatPNatOfNeZeroNat 4 SLang.sv4_presample.proof_5)) ε₂) - x_1) - 0)) = (0 : ENNReal) := by + + have X : (@tsum.{0, 0} ENNReal + (@NonUnitalNonAssocSemiring.toAddCommMonoid.{0} ENNReal + (@NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} ENNReal + (@Semiring.toNonAssocSemiring.{0} ENNReal + (@OrderedSemiring.toSemiring.{0} ENNReal + (@OrderedCommSemiring.toOrderedSemiring.{0} ENNReal + (@CanonicallyOrderedCommSemiring.toOrderedCommSemiring.{0} ENNReal + ENNReal.instCanonicallyOrderedCommSemiring)))))) + ENNReal.instTopologicalSpace Int fun (x_1 : Int) => + @ite.{1} ENNReal (@Eq.{1} Int x_1 x) (Classical.propDecidable (@Eq.{1} Int x_1 x)) + (@OfNat.ofNat.{0} ENNReal 0 (@Zero.toOfNat0.{0} ENNReal instENNRealZero)) + (@ite.{1} ENNReal (@Eq.{1} Int x x_1) (Int.instDecidableEq x x_1) (@SLang.privNoiseGuess dps ε₁ ε₂ x_1) 0)) = 0 := by simp aesop + conv => enter [2, 1, 2] rw [X] @@ -858,8 +869,8 @@ def sv3_sv4_loop_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : Li def sv4_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => let computation : SLang ℕ := do - let τ <- @privNoiseZero dps ε₁ (2 * ε₂) - let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + let τ <- @privNoiseThresh dps ε₁ ε₂ + let v0 <- @privNoiseGuess dps ε₁ ε₂ let sk <- @sv4_loop dps ε₁ ε₂ τ l point ([], v0) return (sv1_threshold sk) computation point @@ -887,8 +898,8 @@ def sv5_loop (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : SLang def sv5_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => let computation : SLang ℕ := do - let τ <- @privNoiseZero dps ε₁ (2 * ε₂) - let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + let τ <- @privNoiseThresh dps ε₁ ε₂ + let v0 <- @privNoiseGuess dps ε₁ ε₂ let presamples <- @sv4_presample dps ε₁ ε₂ point @sv5_loop τ l point (([], v0), presamples) computation point @@ -1107,8 +1118,8 @@ lemma sv5_loop_ind (τ : ℤ) (l : List ℕ) (eval point : ℕ) (past ff: List def sv6_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => let computation : SLang ℕ := do - let τ <- @privNoiseZero dps ε₁ (2 * ε₂) - let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + let τ <- @privNoiseThresh dps ε₁ ε₂ + let v0 <- @privNoiseGuess dps ε₁ ε₂ let presamples <- @sv4_presample dps ε₁ ε₂ point @sv6_loop τ l point (([], v0), presamples) computation point @@ -1186,8 +1197,8 @@ Separates out the zero case def sv7_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => let computation : SLang ℕ := do - let τ <- @privNoiseZero dps ε₁ (2 * ε₂) - let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + let τ <- @privNoiseThresh dps ε₁ ε₂ + let v0 <- @privNoiseGuess dps ε₁ ε₂ match point with | 0 => if (¬ (sv4_privMaxC τ l (([], v0), []))) @@ -1195,7 +1206,7 @@ def sv7_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang else probZero | (Nat.succ point') => do let presamples <- @sv4_presample dps ε₁ ε₂ point' - let vk <- @privNoiseZero dps ε₁ (4 * ε₂) + let vk <- @privNoiseGuess dps ε₁ ε₂ if (sv6_cond τ l (([], v0), presamples ++ [vk])) then probPure point else probZero @@ -1268,8 +1279,8 @@ def sv8_G (l : List ℕ) (past : List ℤ) (pres : ℤ) (future : List ℤ) : def sv8_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => let computation : SLang ℕ := do - let τ <- @privNoiseZero dps ε₁ (2 * ε₂) - let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) + let τ <- @privNoiseThresh dps ε₁ ε₂ + let v0 <- @privNoiseGuess dps ε₁ ε₂ match point with | 0 => if (sv8_sum l [] v0 ≥ τ) @@ -1277,7 +1288,7 @@ def sv8_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang else probZero | (Nat.succ point') => do let presamples <- @sv4_presample dps ε₁ ε₂ point' - let vk <- @privNoiseZero dps ε₁ (4 * ε₂) + let vk <- @privNoiseGuess dps ε₁ ε₂ if (sv8_cond τ l [] v0 presamples vk) then probPure point else probZero @@ -1325,321 +1336,3 @@ def sv7_sv8_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : simp only [] repeat (apply tsum_congr; intro _; congr 1) simp only [sv7_sv8_cond_eq, sv6_cond] - - - - - - - - - - - - - -/- - --- Evaluate the nth conditional starting at state s --- Evaluate the loop conditional n steps in the future --- false if the tape runs out -def sv6_privMax_nth (τ : ℤ) (l : List ℕ) (s : sv4_state) (n : ℕ) : Bool := - match n with - | 0 => - -- Terminates on the 0th iteration if initial state causes the checck to fail - ¬ sv4_privMaxC τ l s - | Nat.succ n' => - match s with - -- If there is more samples on the tape, call recursively - | ((past, present), (future_next :: future_rest)) => - -- sv4_privMaxC τ l ((past, present), (future_next :: future_rest)) ∧ -- Should not terminate here - sv6_privMax_nth τ l ((past ++ [present], future_next), future_rest) n' - | (_, []) => - -- Out of samples on the tape - false - - - --- TODO Test some values of sv6_priv_nth - -lemma sv6_privMax_0th_true (τ : ℤ) (l : List ℕ) (s : sv4_state) : - (sv4_privMaxC τ l s = false) -> (sv6_privMax_nth τ l s 0 = true) := by - simp [sv6_privMax_nth] - -lemma sv6_privMax_0th_false (τ : ℤ) (l : List ℕ) (s : sv4_state) : - (sv4_privMaxC τ l s = true) -> (sv6_privMax_nth τ l s 0 = false) := by - simp [sv6_privMax_nth] - -lemma sv6_privMax_1th_empty (τ : ℤ) (l : List ℕ) (s : sv1_state) : - (sv6_privMax_nth τ l (s, []) 1 = false) := by - simp [sv6_privMax_nth] - --- Inductive formula for n when the tape is empty (not useful) -lemma sv6_privMax_ind_empty (τ : ℤ) (l : List ℕ) (s : sv1_state) (n : ℕ): - (sv6_privMax_nth τ l (s, []) (n + 1) = false) := by - simp [sv6_privMax_nth] - --- Inductive formula for n when the tape is not empty (not useful) -lemma sv6_privMax_ind_nonempty (τ : ℤ) (l : List ℕ) pa pr f fs (n : ℕ): - sv6_privMax_nth τ l ((pa, pr), f :: fs) (n + 1) = - sv6_privMax_nth τ l ((pa ++ [pr], f), fs) n := by - simp [sv6_privMax_nth] - -def len_constr (s : sv4_state) (p : ℕ) : Prop := s.1.1.length + s.2.length = p - --- These are not true --- Could change sv4 definition to be monotone so that it is --- --- -- After its false once, its false forever --- lemma sv4_constancy_lo τ l P v P' v' (H : sv1_privMaxC τ l (P, v) = false) : --- (∃ D, P' ++ [v'] = P ++ [v] ++ D) -> (sv1_privMaxC τ l (P', v') = false) := by --- intro H1 --- rcases H1 with ⟨ D, HD ⟩ --- simp [sv1_privMaxC, sv1_threshold, sv1_noise, exactDiffSum] at H --- simp [sv1_privMaxC, sv1_threshold, sv1_noise, exactDiffSum] --- apply (Int.le_trans H) --- clear H --- --- sorry --- --- -- If its true once, its true at all points before --- lemma sv4_constancy_hi τ l P v P' v' (H : sv1_privMaxC τ l (P, v) = false) : --- (∃ D, P' ++ [v'] ++ D = P ++ [v]) -> (sv1_privMaxC τ l (P', v') = false) := by --- sorry - --- Inductive formula for tail -lemma sv6_privMax_ind (τ : ℤ) (l : List ℕ) pa pr f fs (n : ℕ): - -- sv4_privMaxC τ l ((pa ++ [pr], f), fs) = true -> - -- len_constr ((pa, pr), f :: fs) n -> - sv6_privMax_nth τ l ((pa, pr), f :: fs) n = - sv6_privMax_nth τ l ((pa ++ [pr], f), fs) n := by - revert pa pr f - induction fs - · intro pa pr f - sorry - · sorry - - -def sv6_loop (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : SLang ℕ := - if (sv6_privMax_nth τ l init point) ∧ ¬ (sv6_privMax_nth τ l init (point + 1)) - then return point - else probZero - --- If the sv6 return flag is true (at zero) then... the sv4 loop condition is false (at zero) -lemma sv6_privMax_loop_cond_true_0 (τ : ℤ) (l : List ℕ) (init : sv4_state) : - (len_constr init 0) -> - (sv6_privMax_nth τ l init 0 = true ∧ sv6_privMax_nth τ l init 1 = false) → - (sv4_privMaxC τ l init = false) := by - intro _ - intro ⟨ H1, _ ⟩ -- We can get nothing out of H2? - simp [sv6_privMax_nth] at H1 - trivial - --- If the sv6 return flag is true (at 1) then --- there is at least one element in the future list --- The sv4_condition after executing one shift is false -lemma sv6_privMax_loop_cond_true_1 (τ : ℤ) (l : List ℕ) (init : sv4_state) : - (len_constr init 1) -> - (sv6_privMax_nth τ l init 1 = true ∧ sv6_privMax_nth τ l init 2 = false) → - ∃ init_past init_pres init_fut1 init_futR, - (init = ((init_past, init_pres), (init_fut1 :: init_futR)) ∧ - sv4_privMaxC τ l ((init_past ++ [init_pres], init_fut1), init_futR) = false) := by - intro HL - rcases init with ⟨ ⟨ init_past, init_present ⟩, init_future ⟩ - simp [len_constr] at HL - - intro ⟨ H1, H2 ⟩ - - cases init_past - · -- init_past is [] so init_future is [x] - simp at HL - cases init_future - · exfalso - simp at HL - rename_i future_1 future_rest - cases future_rest - · -- Continue... - clear HL - - -- What can we get out of H2? Nothing - simp [sv6_privMax_nth] at H2 - - -- What can we get out of H1? - simp [sv6_privMax_nth] at H1 - -- sv4_privMaxC τ l (sv4_privMaxF (init)) is false - -- ∃ x, sv4_privMaxC τ l (([init_present], future_1), []) = false - exists [] - exists init_present - exists future_1 - exists [] - · exfalso - simp at HL - · -- init_past is nonempty so init_future is empty - rename_i past_1 past_rest - simp at HL - -- past_rest is empty - cases past_rest - · -- init_future is empty - cases init_future - · -- Continue... - clear HL - - -- What can we get out of H2? Nothing - simp [sv6_privMax_nth] at H2 - - -- What can we get out of H1? Contradiction? - simp [sv6_privMax_nth] at H1 - -- contradiction? - - · exfalso - simp at HL - · exfalso - simp at HL - linarith - - --- So if we know that init.future is nonempty, we should be able to make a conclusion --- about sv4_loop_cond using the s6 loop flag. - --- lemma sv6_privMax_loop_cond_true (τ : ℤ) (l : List ℕ) (init : sv4_state) (p : ℕ) : --- (len_constr init p) -> --- (sv6_privMax_nth τ l init p = true ∧ sv6_privMax_nth τ l init (p + 1) = false) → --- False := by --- sorry - - - - - - - - --- What can we get out of the case where the sv6 flag is false? - --- When the sv6 flag is false (at zero) then sv4 loop cond at 1 is true -lemma sv6_privMax_loop_cond_false_0 (τ : ℤ) (l : List ℕ) (init : sv4_state) : - (len_constr init 0) -> - ¬ (sv6_privMax_nth τ l init 0 = true ∧ sv6_privMax_nth τ l init 1 = false) → - sv4_privMaxC τ l init = true := by - intro HL H - simp [len_constr] at HL - rcases HL with ⟨ HL1, HL2 ⟩ - rcases init with ⟨ ⟨ init_past, init_present ⟩, init_future ⟩ - simp_all [HL1, HL2] - rw [sv6_privMax_1th_empty] at H - simp at H - simp [sv6_privMax_nth] at H - trivial - - - - - - - -- simp [sv6_privMax_nth] at H2 - -- -- H2 is giving us nothing, this is suspicious - - - -- sorry - - --- Specify that the total length of past and future samples is equal to point (0) -def sv5_sv6_loop_eq_point_0_len_constr (τ : ℤ) (l : List ℕ) (init : sv4_state) : - len_constr init 0 -> - @sv5_loop τ l 0 init 0 = @sv6_loop τ l 0 init 0 := by - intro HC - simp [sv6_loop] - split <;> simp - · rename_i h - -- rcases h with ⟨ h1, h2 ⟩ - -- FIXME: See above for general lemmas about the loop conditional and sv4 loop evals - -- Don't do this manually - sorry - · rename_i h - simp at h - sorry - --- sv6 and sv5 both do (decreasing) induction over the "future" field-- --- When the condition is true for sv5, it shifts one sample out of future, and decreases the cut by 1 --- sv6_loop does ...? - --- So I want inductive formulas for those, ie, --- init = (x, y, (a :: as)) --- in terms of --- init = (x ++ [a], y, a) --- evaluated at point = length(x) + length(a::as) - - - - - - - --- These functions are not equal. However, they are equal at "point" (+- 1?) -def sv5_sv6_loop_eq_point (τ : ℤ) (l : List ℕ) (point eval_point : ℕ) (init : sv4_state) : - point ≤ eval_point -> - len_constr init eval_point -> - @sv5_loop τ l point init eval_point = @sv6_loop τ l point init eval_point := by - - rcases init with ⟨ ⟨ pa, pr ⟩, fu ⟩ - revert pa pr point - induction fu - · intro point pa pr HlenIneq Hlen - simp [len_constr] at Hlen - -- simp only [sv6_loop] - simp only [sv5_loop, probWhileCut, probWhileFunctional] - simp only [pure, pure_apply, ↓reduceIte] - simp only [sv1_threshold] - simp only [sv4_privMaxF] - split - · simp - unfold sv6_loop - split - · exfalso - -- There should be a contradiction in here somewhere - -- TODO: Probably change the definition of sv6_cond so that its monotone? - sorry - · simp - · unfold sv4_state - unfold sv1_state - simp - rw [ENNReal.tsum_eq_add_tsum_ite ((pa, pr), [])] - conv => - rhs - rw [<- add_zero (sv6_loop _ _ _ _ _)] - congr - · simp - simp [sv6_loop] - -- ??? - - sorry - · simp - intro _ _ _ _ _ _ _ _ - rename_i h a b b_1 a_1 a_2 a_3 a_4 a_5 - subst Hlen a_4 a_5 a_3 - simp_all only [Bool.not_eq_true, not_true_eq_false, imp_false] - · intro point pa pr HlenIneq Hlen - rename_i f ff IH - cases point - · -- point being zero should be a contradiction? - simp [sv6_loop] - simp [sv5_loop] - -- Not sure. Messed up. - sorry - rename_i point - unfold sv5_loop - unfold probWhileCut - unfold probWhileFunctional - sorry - - -def sv6_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := - fun (point : ℕ) => - let computation : SLang ℕ := do - let τ <- @privNoiseZero dps ε₁ (2 * ε₂) - let v0 <- @privNoiseZero dps ε₁ (4 * ε₂) - let presamples <- @sv4_presample dps ε₁ ε₂ point - @sv6_loop τ l point (([], v0), presamples) - computation point - --/ From b1f7b986a59b5b4bbdd3adb3ea44d242f64be475 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 4 Nov 2024 12:15:14 -0500 Subject: [PATCH 058/100] skeleton of privacy argument --- .../Queries/UnboundedMax/Privacy.lean | 228 ++++++++++++++++++ .../Queries/UnboundedMax/Properties.lean | 43 ++++ 2 files changed, 271 insertions(+) create mode 100644 SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean new file mode 100644 index 00000000..20f0756f --- /dev/null +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean @@ -0,0 +1,228 @@ +/- +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: Markus de Medeiros +-/ + +import SampCert.DifferentialPrivacy.Abstract +import SampCert.DifferentialPrivacy.Queries.UnboundedMax.Properties +import SampCert.DifferentialPrivacy.Pure.System + +noncomputable section + +open Classical + +namespace SLang + +variable (ε₁ ε₂ : ℕ+) + +-- FIXME: Move +lemma tsum_shift (Δ : ℤ) (f : ℤ → ENNReal) : (∑'(x : ℤ), f x = ∑'(x : ℤ), f (x + Δ)) := by + apply @tsum_eq_tsum_of_ne_zero_bij + case i => exact (fun x => x.1 + Δ) + · simp [Function.Injective] + · simp + intro x Hx + exists (x - Δ) + simp + trivial + · intro + rfl + + +lemma laplace_inequality' (τ τ' : ℤ) (Δ : ℕ+) : + ((ENNReal.ofReal (Real.exp (-abs τ' / (Δ * ε₂ / ε₁)))) * ((DiscreteLaplaceGenSamplePMF (Δ * ε₂) ε₁ 0 τ))) ≤ + (DiscreteLaplaceGenSamplePMF (Δ * ε₂) ε₁ 0) (τ + τ') := by + simp [DiscreteLaplaceGenSamplePMF, PMF.instFunLike] + generalize HA : (Real.exp (↑↑ε₁ / (↑↑Δ * ↑↑ε₂)) - 1) = A + generalize HB : (Real.exp (↑↑ε₁ / (↑↑Δ * ↑↑ε₂)) + 1) = B + generalize HC : ((Δ : Real) * ε₂ / ε₁) = C + rw [<- ENNReal.ofReal_mul ?G1] + case G1 => apply Real.exp_nonneg + apply ENNReal.ofReal_le_ofReal + conv => + lhs + rw [mul_comm] + repeat rw [mul_assoc] + apply mul_le_mul_of_nonneg + · rfl + · rw [← Real.exp_add] + apply Real.exp_monotone + repeat rw [<- neg_div] + rw [div_add_div_same] + apply div_le_div_of_nonneg_right + · simp + have H := @abs_add_le ℝ _ _ _ τ τ' + linarith + · rw [<- HC] + simp [div_nonneg_iff] + · rw [<- HA] + rw [<- HB] + simp [div_nonneg_iff] + left + conv => + lhs + rw [<- add_zero 0] + apply add_le_add + · apply Real.exp_nonneg + · simp + · apply Real.exp_nonneg + +lemma laplace_inequality (τ τ' : ℤ) (Δ : ℕ+) : + ((DiscreteLaplaceGenSamplePMF (Δ * ε₂) ε₁ 0 τ)) ≤ + ((ENNReal.ofReal (Real.exp (abs τ' / (Δ * ε₂ / ε₁)))) * + (DiscreteLaplaceGenSamplePMF (Δ * ε₂) ε₁ 0) (τ + τ')) := by + apply le_trans _ ?G1 + case G1 => + apply ENNReal.mul_left_mono + apply laplace_inequality' + simp only [] + repeat rw [<- mul_assoc] + conv => + lhs + rw [<- one_mul (DiscreteLaplaceGenSamplePMF _ _ _ _)] + apply mul_le_mul + · apply Eq.le + rw [<- ENNReal.ofReal_mul ?G1] + case G1 => apply Real.exp_nonneg + rw [<- Real.exp_add] + symm + simp + rw [div_add_div_same] + rw [div_eq_zero_iff] + left + simp + · rfl + · apply zero_le + · apply zero_le + + + +lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : PureDPSystem.prop (@sv9_privMax_pmf PureDPSystem ε₁ ε₂) ε := by + -- Unfold DP definitions + simp [DPSystem.prop] + apply singleton_to_event + unfold DP_singleton + intro l₁ l₂ Hneighbour point + + apply ENNReal.div_le_of_le_mul + simp [sv9_privMax_pmf, DFunLike.coe, PMF.instFunLike] + + cases point + · -- point = 0 + simp only [sv9_privMax] + + sorry + + + · rename_i point + -- point > 0 + -- proceed with the proof in the paper + simp [sv9_privMax] + + -- Cancel out the deterministic parts + conv => + enter [2] + rw [← ENNReal.tsum_mul_left] + enter [1, i] + rw [← ENNReal.tsum_mul_left] + rw [← ENNReal.tsum_mul_left] + enter [1, i_1] + repeat rw [<- mul_assoc] + enter [1] + rw [mul_comm (ENNReal.ofReal _)] + repeat rw [mul_assoc] + rw [mul_comm (ENNReal.ofReal _)] + conv => + enter [2, 1, i, 1, i_1] + repeat rw [mul_assoc] + conv => + enter [1, 1, i] + rw [← ENNReal.tsum_mul_left] + apply ENNReal.tsum_le_tsum + intro v0 + apply ENNReal.tsum_le_tsum + intro vs + apply ENNReal.mul_left_mono + apply ENNReal.mul_left_mono + + -- Rearrange to put sums on the outside + conv => + lhs + enter [1, τ] + rw [← ENNReal.tsum_mul_left] + enter [1, vk] + conv => + rhs + rw [← ENNReal.tsum_mul_left] + enter [1, τ] + rw [← ENNReal.tsum_mul_left] + rw [← ENNReal.tsum_mul_left] + enter [1, vk] + + simp [sv8_cond, sv8_sum] + + -- Perform the changes of variables, so that the sums are pointwise le + let cov_τ : ℤ := (sv8_G l₁ [] v0 vs) - (sv8_G l₂ [] v0 vs) + let cov_vk : ℤ := exactDiffSum (point + 1) l₂ - exactDiffSum (point + 1) l₁ + cov_τ + + conv => + lhs + rw [tsum_shift cov_τ] + enter [1, τ] + rw [tsum_shift cov_vk] + enter [1, vk] + apply ENNReal.tsum_le_tsum + intro τ + apply ENNReal.tsum_le_tsum + intro vk + + -- The change of variables make the conditional equal + repeat rw [<- mul_assoc] + apply mul_le_mul' _ ?G2 + case G2 => + apply Eq.le + suffices ((sv8_G l₁ [] v0 ↑vs < τ + cov_τ) = (sv8_G l₂ [] v0 ↑vs < τ)) ∧ + ((τ + cov_τ ≤ exactDiffSum (point + 1) l₁ + (vk + cov_vk)) = (τ ≤ exactDiffSum (point + 1) l₂ + vk)) by simp_all + apply And.intro + · -- cov_τ + apply propext + dsimp [cov_τ] + apply Iff.intro <;> intro _ <;> linarith + · -- cov_vk + apply propext + dsimp [cov_vk] + apply Iff.intro <;> intro _ <;> linarith + + -- Might be more sensitive + let sens_cov_τ : ℤ := 1 + have Hsens_cov_τ : cov_τ ≤ sens_cov_τ := by + dsimp [cov_τ] + cases vs + rename_i vs Hvs + simp + + cases Hneighbour + · rename_i A B C H1 H2 + simp [H1, H2]; clear H1 H2 + induction A + · simp + unfold sv8_G + sorry + · sorry + · sorry + · sorry + + -- Might be more sensitive in reality + let sens_cov_vk : ℤ := 2 + have Hsens_cov_vk : cov_vk ≤ sens_cov_vk := by + sorry + + simp [privNoiseThresh, privNoiseGuess, + privNoiseZero, DPSystem.noise, privNoisedQueryPure] + + + + + + sorry diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 33d024ab..600f91d3 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -1336,3 +1336,46 @@ def sv7_sv8_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : simp only [] repeat (apply tsum_congr; intro _; congr 1) simp only [sv7_sv8_cond_eq, sv6_cond] + + + +/- +## Program v9 + +Not executable +Rewritten so that the randomness we will cancel out is right at the front +-/ + + +def sv9_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := + fun (point : ℕ) => + let computation : SLang ℕ := do + match point with + | 0 => do + let τ <- @privNoiseThresh dps ε₁ ε₂ + let v0 <- @privNoiseGuess dps ε₁ ε₂ + if (sv8_sum l [] v0 ≥ τ) + then probPure point + else probZero + | (Nat.succ point') => do + let v0 <- @privNoiseGuess dps ε₁ ε₂ + let presamples <- @sv4_presample dps ε₁ ε₂ point' + let τ <- @privNoiseThresh dps ε₁ ε₂ + let vk <- @privNoiseGuess dps ε₁ ε₂ + if (sv8_cond τ l [] v0 presamples vk) + then probPure point + else probZero + computation point + +def sv8_sv9_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : + @sv8_privMax dps ε₁ ε₂ l = @sv8_privMax dps ε₁ ε₂ l := by + apply SLang.ext + intro point + sorry + + +/-- +sv9 normalizes because sv1 normalizes +-/ +def sv9_privMax_pmf [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : PMF ℕ := + ⟨ @sv9_privMax dps ε₁ ε₂ l, sorry ⟩ From 44c9e7d81f2dca6d5775768ee379840372fc2f80 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 4 Nov 2024 13:16:19 -0500 Subject: [PATCH 059/100] sparse vector main proof --- .../Queries/UnboundedMax/Privacy.lean | 72 +++++++++++++++++-- .../Queries/UnboundedMax/Properties.lean | 6 +- 2 files changed, 71 insertions(+), 7 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean index 20f0756f..09255607 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean @@ -97,6 +97,15 @@ lemma laplace_inequality (τ τ' : ℤ) (Δ : ℕ+) : · apply zero_le +lemma laplace_inequality_sub (τ τ' : ℤ) (Δ : ℕ+) : + ((DiscreteLaplaceGenSamplePMF (Δ * ε₂) ε₁ 0 (τ + τ'))) ≤ + ((ENNReal.ofReal (Real.exp (abs τ' / (Δ * ε₂ / ε₁)))) * + (DiscreteLaplaceGenSamplePMF (Δ * ε₂) ε₁ 0) τ) := by + apply le_trans + · apply laplace_inequality ε₁ ε₂ (τ + τ') (-τ') Δ + apply Eq.le + simp + lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : PureDPSystem.prop (@sv9_privMax_pmf PureDPSystem ε₁ ε₂) ε := by -- Unfold DP definitions @@ -195,7 +204,6 @@ lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : PureDPSystem apply Iff.intro <;> intro _ <;> linarith -- Might be more sensitive - let sens_cov_τ : ℤ := 1 have Hsens_cov_τ : cov_τ ≤ sens_cov_τ := by dsimp [cov_τ] cases vs @@ -214,15 +222,69 @@ lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : PureDPSystem · sorry -- Might be more sensitive in reality - let sens_cov_vk : ℤ := 2 have Hsens_cov_vk : cov_vk ≤ sens_cov_vk := by sorry simp [privNoiseThresh, privNoiseGuess, privNoiseZero, DPSystem.noise, privNoisedQueryPure] + -- Apply the Laplace inequalities + apply le_trans + · apply mul_le_mul + · apply laplace_inequality_sub + · apply laplace_inequality_sub + · apply zero_le + · apply zero_le + -- Cancel the Laplace samples + conv => + lhs + rw [mul_assoc] + rw [mul_comm] + conv => + rhs + rw [mul_assoc] + rw [mul_comm] + repeat rw [mul_assoc] + apply ENNReal.mul_left_mono + conv => + lhs + rw [mul_comm] + repeat rw [mul_assoc] + apply ENNReal.mul_left_mono - - - sorry + -- Apply the ineuqalities + rw [<- ENNReal.ofReal_mul ?G1] + case G1 => apply Real.exp_nonneg + apply ENNReal.ofReal_le_ofReal + rw [← Real.exp_add] + apply Real.exp_monotone + apply @le_trans _ _ _ ((|sens_cov_τ| : ℝ) / (↑↑(2 * sens_cov_τ) * ↑↑ε₂ / ↑↑ε₁) + (|sens_cov_vk| : ℝ) / (↑↑(2 * sens_cov_vk) * ↑↑ε₂ / ↑↑ε₁)) + · apply add_le_add + · simp + apply div_le_div_of_nonneg_right + · -- apply abs_le'.mpr + sorry + · apply mul_nonneg <;> simp + · sorry + simp [sens_cov_τ] + simp [sens_cov_vk] + ring_nf + rw [InvolutiveInv.inv_inv] + rw [Hε] + conv => + lhs + enter [2, 1] + rw [mul_comm] + rw [<- add_mul] + rw [<- add_mul] + conv => + lhs + enter [1, 1] + rw [<- (one_mul (ε₁.val.cast))] + rw [<- add_mul] + repeat rw [div_eq_mul_inv] + simp + rw [one_add_one_eq_two] + ring_nf + rfl diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 600f91d3..a38873f5 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -11,6 +11,8 @@ open Classical namespace SLang +def sens_cov_τ : ℕ+ := 1 +def sens_cov_vk : ℕ+ := 2 /-- Stronger congruence rule for probBind: The bound-to functions have to be equal only on the support of @@ -67,9 +69,9 @@ This looks strange, but will specialize to Lap(ε₁/ε₂, 0) in the pure DP ca def privNoiseZero [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) : SLang ℤ := dps.noise (fun _ => 0) 1 ε₁ ε₂ [] -def privNoiseGuess [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) : SLang ℤ := @privNoiseZero dps ε₁ (4 * ε₂) +def privNoiseGuess [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) : SLang ℤ := @privNoiseZero dps ε₁ (2 * sens_cov_vk * ε₂) -def privNoiseThresh [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) : SLang ℤ := @privNoiseZero dps ε₁ (2 * ε₂) +def privNoiseThresh [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) : SLang ℤ := @privNoiseZero dps ε₁ (2 * sens_cov_τ * ε₂) /- From 1f7254a288d97bc94092731a1071fe4baf0c2583 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 4 Nov 2024 15:27:03 -0500 Subject: [PATCH 060/100] extra sensitivity lemmas --- .../Queries/UnboundedMax/Privacy.lean | 125 +++++++++++++++--- .../Queries/UnboundedMax/Properties.lean | 4 +- 2 files changed, 111 insertions(+), 18 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean index 09255607..9e8d2218 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean @@ -107,6 +107,66 @@ lemma laplace_inequality_sub (τ τ' : ℤ) (Δ : ℕ+) : simp + +lemma exactClippedSum_append : exactClippedSum i (A ++ B) = exactClippedSum i A + exactClippedSum i B := by + simp [exactClippedSum] + +lemma exactDiffSum_append : exactDiffSum i (A ++ B) = exactDiffSum i A + exactDiffSum i B := by + simp [exactDiffSum] + rw [exactClippedSum_append] + rw [exactClippedSum_append] + linarith + +lemma sv8_sum_append : sv8_sum (A ++ B) [] v0 = sv8_sum A [] v0 + sv8_sum B [] v0 - v0 := by + simp [sv8_sum] + simp [exactDiffSum_append] + linarith + +lemma sv8_sum_comm : sv8_sum (A ++ B) vp v0 = sv8_sum (B ++ A) vp v0 := by + unfold sv8_sum + simp + simp [exactDiffSum, exactClippedSum] + linarith + +lemma sv8_G_comm : sv8_G (A ++ B) vp v0 vf = sv8_G (B ++ A) vp v0 vf := by + revert vp v0 + induction vf + · intros + apply sv8_sum_comm + · intro v0 vp + rename_i next rest IH + unfold sv8_G + rw [sv8_sum_comm] + congr 1 + apply IH + +lemma exactDiffSum_nonpos : exactDiffSum point L ≤ 0 := by + simp [exactDiffSum, exactClippedSum] + induction L + · simp + · rename_i l ll IH + simp + apply le_trans + · apply add_le_add + · rfl + · apply IH + simp + +lemma exactDiffSum_singleton_le_1 : -1 ≤ exactDiffSum point [v] := by + simp [exactDiffSum, exactClippedSum] + cases (Classical.em (point ≤ v)) + · right + trivial + · left + rename_i h + simp at h + rw [Int.min_def] + simp + split + · linarith + · linarith + + lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : PureDPSystem.prop (@sv9_privMax_pmf PureDPSystem ε₁ ε₂) ε := by -- Unfold DP definitions simp [DPSystem.prop] @@ -124,6 +184,7 @@ lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : PureDPSystem sorry + · rename_i point -- point > 0 -- proceed with the proof in the paper @@ -203,27 +264,60 @@ lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : PureDPSystem dsimp [cov_vk] apply Iff.intro <;> intro _ <;> linarith - -- Might be more sensitive + -- Prove sensitivity bound have Hsens_cov_τ : cov_τ ≤ sens_cov_τ := by dsimp [cov_τ] cases vs rename_i vs Hvs - simp - + simp only [] cases Hneighbour - · rename_i A B C H1 H2 - simp [H1, H2]; clear H1 H2 - induction A - · simp - unfold sv8_G - sorry - · sorry - · sorry - · sorry + · rename_i A B n H1 H2 + rw [H1, H2]; clear H1 H2 + sorry + -- induction vs + -- · simp only [sv8_G, sv8_sum_append] + -- simp [sv8_sum] + -- sorry + -- · simp only [sv8_G] + -- sorry + · rename_i _ n _ H1 H2 + rw [H1, H2]; clear H1 H2 + sorry + · rename_i n1 _ n2 H1 H2 + rw [H1, H2]; clear H1 H2 + sorry - -- Might be more sensitive in reality + -- Prove sensitivity bound have Hsens_cov_vk : cov_vk ≤ sens_cov_vk := by - sorry + dsimp [cov_vk] + cases vs + rename_i vs Hvs + cases Hneighbour + · rename_i _ _ n H1 H2 + rw [H1, H2]; clear H1 H2 + repeat rw [exactDiffSum_append] + simp_all [sens_cov_vk, sens_cov_τ] + have _ := @exactDiffSum_singleton_le_1 (point + 1) n + have _ := @exactDiffSum_nonpos (point + 1) [n] + linarith + + · rename_i _ n _ H1 H2 + rw [H1, H2]; clear H1 H2 + repeat rw [exactDiffSum_append] + simp_all [sens_cov_vk, sens_cov_τ] + have _ := @exactDiffSum_singleton_le_1 (point + 1) n + have _ := @exactDiffSum_nonpos (point + 1) [n] + linarith + + · rename_i n1 _ n2 H1 H2 + rw [H1, H2]; clear H1 H2 + repeat rw [exactDiffSum_append] + simp_all [sens_cov_vk, sens_cov_τ] + have _ := @exactDiffSum_singleton_le_1 (point + 1) n1 + have _ := @exactDiffSum_nonpos (point + 1) [n1] + have _ := @exactDiffSum_singleton_le_1 (point + 1) n2 + have _ := @exactDiffSum_nonpos (point + 1) [n2] + linarith simp [privNoiseThresh, privNoiseGuess, privNoiseZero, DPSystem.noise, privNoisedQueryPure] @@ -267,8 +361,7 @@ lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : PureDPSystem sorry · apply mul_nonneg <;> simp · sorry - simp [sens_cov_τ] - simp [sens_cov_vk] + simp [sens_cov_τ, sens_cov_vk] ring_nf rw [InvolutiveInv.inv_inv] rw [Hε] diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index a38873f5..09709c9a 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -12,7 +12,7 @@ open Classical namespace SLang def sens_cov_τ : ℕ+ := 1 -def sens_cov_vk : ℕ+ := 2 +def sens_cov_vk : ℕ+ := 1 + sens_cov_τ /-- Stronger congruence rule for probBind: The bound-to functions have to be equal only on the support of @@ -569,7 +569,7 @@ lemma sv3_loop_unroll_1 [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l p split <;> try simp unfold privNoiseGuess unfold privNoiseZero - exact Eq.symm (PMF.tsum_coe (DPSystem.noise (fun _ => 0) 1 ε₁ (4 * ε₂) [])) + exact Eq.symm (PMF.tsum_coe (DPSystem.noise (fun x => 0) 1 ε₁ (2 * sens_cov_vk * ε₂) [])) /- ## Program version 4 From 7c1d1aeff47f674d676eae90f79f699d49a35a54 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 4 Nov 2024 15:43:59 -0500 Subject: [PATCH 061/100] checkpoint --- .../Queries/UnboundedMax/Privacy.lean | 40 +++++++++++++++---- 1 file changed, 32 insertions(+), 8 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean index 9e8d2218..644c68d0 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean @@ -140,6 +140,18 @@ lemma sv8_G_comm : sv8_G (A ++ B) vp v0 vf = sv8_G (B ++ A) vp v0 vf := by congr 1 apply IH +-- IDK +lemma sv8_G_cons : sv8_G (x :: L) vp v0 vf = 0 := by + revert vp v0 + induction vf + · intros v0 vp + simp [sv8_G] + sorry + · intro vp v0 + simp [sv8_G] + sorry + -- unfold sv8_G + lemma exactDiffSum_nonpos : exactDiffSum point L ≤ 0 := by simp [exactDiffSum, exactClippedSum] induction L @@ -273,18 +285,30 @@ lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : PureDPSystem cases Hneighbour · rename_i A B n H1 H2 rw [H1, H2]; clear H1 H2 + conv => + enter [1, 2, 1] + rw [List.append_assoc] + rw [@sv8_G_comm A ([n] ++ B)] sorry - -- induction vs - -- · simp only [sv8_G, sv8_sum_append] - -- simp [sv8_sum] - -- sorry - -- · simp only [sv8_G] - -- sorry - · rename_i _ n _ H1 H2 + + · rename_i A n B H1 H2 rw [H1, H2]; clear H1 H2 + conv => + enter [1, 1, 1] + rw [List.append_assoc] + rw [@sv8_G_comm A ([n] ++ B)] sorry - · rename_i n1 _ n2 H1 H2 + + · rename_i A n1 B n2 H1 H2 rw [H1, H2]; clear H1 H2 + conv => + enter [1, 1, 1] + rw [List.append_assoc] + conv => + enter [1, 2, 1] + rw [List.append_assoc] + rw [@sv8_G_comm A ([n1] ++ B)] + rw [@sv8_G_comm A ([n2] ++ B)] sorry -- Prove sensitivity bound From 85b9a4bf13b7cd18d11ea69fe86dfcecd10bbf3c Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Mon, 4 Nov 2024 22:34:49 -0500 Subject: [PATCH 062/100] base case --- .../Queries/UnboundedMax/Privacy.lean | 131 +++++++++++++++++- 1 file changed, 129 insertions(+), 2 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean index 644c68d0..11050fa6 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean @@ -179,6 +179,32 @@ lemma exactDiffSum_singleton_le_1 : -1 ≤ exactDiffSum point [v] := by · linarith +-- Coercions nonsense +lemma DS0 (H : Neighbour L1 L2) : ((exactDiffSum 0 L1) : ℝ) - (exactDiffSum 0 L2) ≤ 1 := by + cases H + · rename_i A B C H1 H2 + rw [H1, H2] + repeat rw [exactDiffSum_append] + simp + apply neg_le.mp + have _ := @exactDiffSum_singleton_le_1 0 C + sorry + + · rename_i A B C H1 H2 + rw [H1, H2] + repeat rw [exactDiffSum_append] + simp + have _ := @exactDiffSum_nonpos 0 C + sorry + + · rename_i A B C D H1 H2 + rw [H1, H2] + repeat rw [exactDiffSum_append] + simp only [Int.cast_add, add_sub_add_right_eq_sub, add_sub_add_left_eq_sub] + sorry + + + lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : PureDPSystem.prop (@sv9_privMax_pmf PureDPSystem ε₁ ε₂) ε := by -- Unfold DP definitions simp [DPSystem.prop] @@ -191,10 +217,111 @@ lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : PureDPSystem cases point · -- point = 0 - simp only [sv9_privMax] + simp [sv9_privMax] - sorry + -- Put sums on outside + conv => + enter [2] + rw [← ENNReal.tsum_mul_left] + enter [1, i] + rw [← ENNReal.tsum_mul_left] + rw [← ENNReal.tsum_mul_left] + enter [1, i_1] + repeat rw [<- mul_assoc] + enter [1] + rw [mul_comm (ENNReal.ofReal _)] + repeat rw [mul_assoc] + rw [mul_comm (ENNReal.ofReal _)] + conv => + enter [2, 1, i, 1, i_1] + repeat rw [mul_assoc] + conv => + enter [1, 1, i] + rw [← ENNReal.tsum_mul_left] + -- Change of variables + let cov_τ : ℤ := 0 + let cov_vk : ℤ := exactDiffSum 0 l₂ - exactDiffSum 0 l₁ + cov_τ + conv => + lhs + rw [tsum_shift cov_τ] + enter [1, τ] + rw [tsum_shift cov_vk] + enter [1, vk] + conv => + rhs + enter [1, τ, 1, vk] + apply ENNReal.tsum_le_tsum + intro τ + apply ENNReal.tsum_le_tsum + intro vk + + rw [<- mul_assoc] + rw [<- mul_assoc] + rw [<- mul_assoc] + apply mul_le_mul + · -- Laplace bound + simp [cov_τ] + rw [mul_assoc] + apply ENNReal.mul_left_mono + simp [privNoiseGuess, privNoiseZero, DPSystem.noise, privNoisedQueryPure] + apply le_trans + · apply laplace_inequality_sub + rw [mul_comm] + apply ENNReal.mul_left_mono + rw [Hε] + apply ENNReal.ofReal_le_ofReal + apply Real.exp_monotone + simp [sens_cov_vk, sens_cov_τ] + + have X : |((exactDiffSum 0 l₂) : ℝ) - (exactDiffSum 0 l₁)| ≤ 1 := by + -- simp [exactDiffSum, exactClippedSum, List.map_const'] + rw [abs_le] + apply And.intro + · apply neg_le.mp + simp only [neg_sub] + apply DS0 + apply Hneighbour + · apply DS0 + apply Neighbour_symm + apply Hneighbour + + ring_nf + rw [InvolutiveInv.inv_inv] + conv => + lhs + rw [mul_comm] + rw [<- mul_assoc] + rw [<- mul_assoc] + rw [mul_comm] + enter [2] + rw [mul_comm] + simp + apply le_trans _ X + conv => + rhs + rw [<- one_mul (abs _)] + apply mul_le_mul + · apply inv_le_one + simp + · rfl + · apply abs_nonneg + · simp + + + + + · -- Conditionals should be equal + suffices (τ + cov_τ ≤ sv8_sum l₁ [] (vk + cov_vk)) = (τ ≤ sv8_sum l₂ [] vk) by + split <;> simp_all + apply propext + simp [sv8_sum, cov_vk] + apply Iff.intro + · intro _ ; linarith + · intro _ ; linarith + + · apply zero_le + · apply zero_le · rename_i point From b41d701bc31f1ccec0aaf941ec49cdb61a2193aa Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Tue, 5 Nov 2024 08:38:30 -0500 Subject: [PATCH 063/100] checkpoint --- .../Queries/UnboundedMax/Privacy.lean | 22 +++++++++---------- 1 file changed, 11 insertions(+), 11 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean index 11050fa6..43e1ee16 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean @@ -140,17 +140,17 @@ lemma sv8_G_comm : sv8_G (A ++ B) vp v0 vf = sv8_G (B ++ A) vp v0 vf := by congr 1 apply IH --- IDK -lemma sv8_G_cons : sv8_G (x :: L) vp v0 vf = 0 := by - revert vp v0 - induction vf - · intros v0 vp - simp [sv8_G] - sorry - · intro vp v0 - simp [sv8_G] - sorry - -- unfold sv8_G +-- -- IDK +-- lemma sv8_G_cons : sv8_G (x :: L) vp v0 vf = 0 := by +-- revert vp v0 +-- induction vf +-- · intros v0 vp +-- simp [sv8_G] +-- s orry +-- · intro vp v0 +-- simp [sv8_G] +-- s orry +-- -- unfold sv8_G lemma exactDiffSum_nonpos : exactDiffSum point L ≤ 0 := by simp [exactDiffSum, exactClippedSum] From 1489ded6cf887d42b89bb0bf440913aa9463678f Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Tue, 5 Nov 2024 14:28:08 -0500 Subject: [PATCH 064/100] Fix build --- SampCert.lean | 1 + .../Queries/UnboundedMax/Basic.lean | 7 +- .../Queries/UnboundedMax/Privacy.lean | 9 +- .../Queries/UnboundedMax/Properties.lean | 260 +++++++++--------- 4 files changed, 141 insertions(+), 136 deletions(-) diff --git a/SampCert.lean b/SampCert.lean index 1e1285dc..35a3c904 100644 --- a/SampCert.lean +++ b/SampCert.lean @@ -8,6 +8,7 @@ import SampCert.DifferentialPrivacy.Queries.Histogram.Basic import SampCert.DifferentialPrivacy.ZeroConcentrated.System import SampCert.DifferentialPrivacy.Pure.System import SampCert.DifferentialPrivacy.Queries.HistogramMean.Properties +import SampCert.DifferentialPrivacy.Queries.UnboundedMax.Basic import SampCert.DifferentialPrivacy.Approximate.DP import SampCert.Samplers.Gaussian.Properties import Init.Data.UInt.Lemmas diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Basic.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Basic.lean index 72073d78..544072bb 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Basic.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Basic.lean @@ -1,10 +1,13 @@ /- Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jean-Baptiste Tristan +Authors: Markus de Medeiros -/ -- import SampCert.DifferentialPrivacy.Queries.UnboundedMax.Code --- import SampCert.DifferentialPrivacy.Queries.UnboundedMax.Properties + +import SampCert.DifferentialPrivacy.Queries.UnboundedMax.Properties +import SampCert.DifferentialPrivacy.Queries.UnboundedMax.Privacy + -- import SampCert.DifferentialPrivacy.Queries.UnboundedMax.Pure -- import SampCert.DifferentialPrivacy.Queries.UnboundedMax.ZeroConcentrated diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean index 43e1ee16..3aca4b3e 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean @@ -204,8 +204,8 @@ lemma DS0 (H : Neighbour L1 L2) : ((exactDiffSum 0 L1) : ℝ) - (exactDiffSum 0 sorry - -lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : PureDPSystem.prop (@sv9_privMax_pmf PureDPSystem ε₁ ε₂) ε := by +lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : + PureDPSystem.prop (@sv9_privMax_pmf PureDPSystem laplace_pureDPSystem ε₁ ε₂) ε := by -- Unfold DP definitions simp [DPSystem.prop] apply singleton_to_event @@ -264,7 +264,7 @@ lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : PureDPSystem simp [cov_τ] rw [mul_assoc] apply ENNReal.mul_left_mono - simp [privNoiseGuess, privNoiseZero, DPSystem.noise, privNoisedQueryPure] + simp [privNoiseGuess, privNoiseZero, DPNoise.noise, privNoisedQueryPure] apply le_trans · apply laplace_inequality_sub rw [mul_comm] @@ -470,8 +470,7 @@ lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : PureDPSystem have _ := @exactDiffSum_nonpos (point + 1) [n2] linarith - simp [privNoiseThresh, privNoiseGuess, - privNoiseZero, DPSystem.noise, privNoisedQueryPure] + simp [privNoiseThresh, privNoiseGuess, privNoiseZero, DPNoise.noise, privNoisedQueryPure] -- Apply the Laplace inequalities apply le_trans diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 09709c9a..13cd042d 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -11,6 +11,9 @@ open Classical namespace SLang +variable [dps : DPSystem ℕ] +variable [dpn : DPNoise dps] + def sens_cov_τ : ℕ+ := 1 def sens_cov_vk : ℕ+ := 1 + sens_cov_τ @@ -66,12 +69,12 @@ Noise the constant 0 value using the abstract noise function. This looks strange, but will specialize to Lap(ε₁/ε₂, 0) in the pure DP case. -/ -def privNoiseZero [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) : SLang ℤ := dps.noise (fun _ => 0) 1 ε₁ ε₂ [] +def privNoiseZero (ε₁ ε₂ : ℕ+) : SLang ℤ := dpn.noise (fun _ => 0) 1 ε₁ ε₂ [] -def privNoiseGuess [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) : SLang ℤ := @privNoiseZero dps ε₁ (2 * sens_cov_vk * ε₂) +def privNoiseGuess (ε₁ ε₂ : ℕ+) : SLang ℤ := privNoiseZero ε₁ (2 * sens_cov_vk * ε₂) -def privNoiseThresh [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) : SLang ℤ := @privNoiseZero dps ε₁ (2 * sens_cov_τ * ε₂) +def privNoiseThresh (ε₁ ε₂ : ℕ+) : SLang ℤ := privNoiseZero ε₁ (2 * sens_cov_τ * ε₂) /- @@ -135,14 +138,14 @@ def sv0_noise (s : sv0_state) : ℤ := s.2 def sv0_privMaxC (τ : ℤ) (l : List ℕ) (s : sv0_state) : Bool := decide (exactDiffSum (sv0_threshold s) l + (sv0_noise s) < τ) -def sv0_privMaxF [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (s : sv0_state) : SLang sv0_state := do - let vn <- @privNoiseGuess dps ε₁ ε₂ +def sv0_privMaxF (ε₁ ε₂ : ℕ+) (s : sv0_state) : SLang sv0_state := do + let vn <- privNoiseGuess ε₁ ε₂ let n := (sv0_threshold s) + 1 return (n, vn) -def sv0_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do - let τ <- @privNoiseThresh dps ε₁ ε₂ - let v0 <- @privNoiseGuess dps ε₁ ε₂ +def sv0_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do + let τ <- privNoiseThresh ε₁ ε₂ + let v0 <- privNoiseGuess ε₁ ε₂ let sk <- probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (0, v0) return (sv0_threshold sk) @@ -161,20 +164,20 @@ def sv1_noise (s : sv1_state) : ℤ := s.2 def sv1_privMaxC (τ : ℤ) (l : List ℕ) (s : sv1_state) : Bool := decide (exactDiffSum (sv1_threshold s) l + (sv1_noise s) < τ) -def sv1_privMaxF [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (s : sv1_state) : SLang sv1_state := do - let vn <- @privNoiseGuess dps ε₁ ε₂ +def sv1_privMaxF (ε₁ ε₂ : ℕ+) (s : sv1_state) : SLang sv1_state := do + let vn <- privNoiseGuess ε₁ ε₂ return (s.1 ++ [s.2], vn) -def sv1_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do - let τ <- @privNoiseThresh dps ε₁ ε₂ - let v0 <- @privNoiseGuess dps ε₁ ε₂ +def sv1_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do + let τ <- privNoiseThresh ε₁ ε₂ + let v0 <- privNoiseGuess ε₁ ε₂ let sk <- probWhile (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) ([], v0) return (sv1_threshold sk) /-- History-aware progam computes the same as the history-agnostic program -/ -lemma sv0_eq_sv1 [dps : DPSystem ℕ] ε₁ ε₂ l : sv0_privMax ε₁ ε₂ l = sv1_privMax ε₁ ε₂ l := by +lemma sv0_eq_sv1 ε₁ ε₂ l : sv0_privMax ε₁ ε₂ l = sv1_privMax ε₁ ε₂ l := by apply SLang.ext -- Initial setup is equal @@ -230,16 +233,16 @@ lemma sv0_eq_sv1 [dps : DPSystem ℕ] ε₁ ε₂ l : sv0_privMax ε₁ ε₂ l - Only moves the loop into a non-executable form, ie. explicitly defines the PMF -/ -def sv2_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := +def sv2_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => let computation : SLang ℕ := do - let τ <- @privNoiseThresh dps ε₁ ε₂ - let v0 <- @privNoiseGuess dps ε₁ ε₂ + let τ <- privNoiseThresh ε₁ ε₂ + let v0 <- privNoiseGuess ε₁ ε₂ let sk <- probWhile (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) ([], v0) return (sv1_threshold sk) computation point -lemma sv1_eq_sv2 [dps : DPSystem ℕ] ε₁ ε₂ l : sv1_privMax ε₁ ε₂ l = sv2_privMax ε₁ ε₂ l := by +lemma sv1_eq_sv2 ε₁ ε₂ l : sv1_privMax ε₁ ε₂ l = sv2_privMax ε₁ ε₂ l := by apply SLang.ext intro result simp [sv1_privMax, sv2_privMax] @@ -252,15 +255,15 @@ lemma sv1_eq_sv2 [dps : DPSystem ℕ] ε₁ ε₂ l : sv1_privMax ε₁ ε₂ l - Truncates the loop -/ -def sv3_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : SLang sv1_state := do - probWhileCut (sv1_privMaxC τ l) (@sv1_privMaxF dps ε₁ ε₂) (point + 1) init +def sv3_loop (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : SLang sv1_state := do + probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (point + 1) init -def sv3_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := +def sv3_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => let computation : SLang ℕ := do - let τ <- @privNoiseThresh dps ε₁ ε₂ - let v0 <- @privNoiseGuess dps ε₁ ε₂ - let sk <- @sv3_loop dps ε₁ ε₂ τ l point ([], v0) + let τ <- privNoiseThresh ε₁ ε₂ + let v0 <- privNoiseGuess ε₁ ε₂ + let sk <- sv3_loop ε₁ ε₂ τ l point ([], v0) return (sv1_threshold sk) computation point @@ -268,13 +271,13 @@ def sv3_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang def cone_of_possibility (cut : ℕ) (initial hist : List ℤ) : Prop := (hist.length < cut + initial.length) ∧ (initial.length ≤ hist.length) -def constancy_at [DPSystem ℕ] {ε₁ ε₂ : ℕ+} {τ : ℤ} {data : List ℕ} {v0 vk : ℤ} (cut : ℕ) (initial hist : List ℤ) : Prop := +def constancy_at {ε₁ ε₂ : ℕ+} {τ : ℤ} {data : List ℕ} {v0 vk : ℤ} (cut : ℕ) (initial hist : List ℤ) : Prop := probWhileCut (sv1_privMaxC τ data) (sv1_privMaxF ε₁ ε₂) (1 + cut) (initial, v0) (hist, vk) = probWhileCut (sv1_privMaxC τ data) (sv1_privMaxF ε₁ ε₂) cut (initial, v0) (hist, vk) -- All points outside of the cone are zero -lemma external_to_cone_zero [DPSystem ℕ] : +lemma external_to_cone_zero : (¬ cone_of_possibility cut initial hist) -> probWhileCut (sv1_privMaxC τ data) (sv1_privMaxF ε₁ ε₂) cut (initial, v0) (hist, vk) = 0 := by revert initial v0 vk @@ -310,10 +313,10 @@ lemma external_to_cone_zero [DPSystem ℕ] : simp_all [cone_of_possibility] -- Base case: left edge of the cone satisfies constancy -lemma cone_left_edge_constancy [DPSystem ℕ] {ε₁ ε₂ : ℕ+} {τ : ℤ} {data : List ℕ} {v0 vk : ℤ} (cut : ℕ) (initial hist : List ℤ) : +lemma cone_left_edge_constancy {ε₁ ε₂ : ℕ+} {τ : ℤ} {data : List ℕ} {v0 vk : ℤ} (cut : ℕ) (initial hist : List ℤ) : hist.length = initial.length -> cone_of_possibility cut initial hist -> - @constancy_at _ ε₁ ε₂ τ data v0 vk cut initial hist := by + @constancy_at _ _ ε₁ ε₂ τ data v0 vk cut initial hist := by intro Hlen Hcone -- cut > 0 due to cone cases cut @@ -364,9 +367,9 @@ lemma cone_left_edge_constancy [DPSystem ℕ] {ε₁ ε₂ : ℕ+} {τ : ℤ} {d rw [external_to_cone_zero (by simp_all [cone_of_possibility])] -lemma cone_constancy [DPSystem ℕ] {ε₁ ε₂ : ℕ+} {τ : ℤ} {data : List ℕ} {v0 vk : ℤ} (cut : ℕ) (initial hist : List ℤ) : +lemma cone_constancy {ε₁ ε₂ : ℕ+} {τ : ℤ} {data : List ℕ} {v0 vk : ℤ} (cut : ℕ) (initial hist : List ℤ) : cone_of_possibility cut initial hist -> - @constancy_at _ ε₁ ε₂ τ data v0 vk cut initial hist := by + @constancy_at _ _ ε₁ ε₂ τ data v0 vk cut initial hist := by -- Need theorem to be true for all initial states, since this will increase during the induction -- v0 and vk will also change in ways which don't matter revert initial v0 vk @@ -449,7 +452,7 @@ lemma cone_constancy [DPSystem ℕ] {ε₁ ε₂ : ℕ+} {τ : ℤ} {data : List · trivial -lemma sv2_eq_sv3 [dps : DPSystem ℕ] ε₁ ε₂ l : sv2_privMax ε₁ ε₂ l = sv3_privMax ε₁ ε₂ l := by +lemma sv2_eq_sv3 ε₁ ε₂ l : sv2_privMax ε₁ ε₂ l = sv3_privMax ε₁ ε₂ l := by apply SLang.ext -- Step through equal headers @@ -499,7 +502,7 @@ lemma sv2_eq_sv3 [dps : DPSystem ℕ] ε₁ ε₂ l : sv2_privMax ε₁ ε₂ l rename_i Hcont _ apply Hcont linarith - have HK := @cone_constancy dps ε₁ ε₂ τ l v0 vk i [] hist + have HK := @cone_constancy _ _ ε₁ ε₂ τ l v0 vk i [] hist unfold constancy_at at HK conv => enter [1, 3] @@ -519,10 +522,10 @@ lemma sv2_eq_sv3 [dps : DPSystem ℕ] ε₁ ε₂ l : sv2_privMax ε₁ ε₂ l -- Commute out a single sample from the loop -lemma sv3_loop_unroll_1 [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l point L vk : +lemma sv3_loop_unroll_1 (τ : ℤ) (ε₁ ε₂ : ℕ+) l point L vk : sv3_loop ε₁ ε₂ τ l (point + 1) (L, vk) = (do - let vk1 <- @privNoiseGuess dps ε₁ ε₂ + let vk1 <- privNoiseGuess ε₁ ε₂ if (sv1_privMaxC τ l (L, vk)) then (sv3_loop ε₁ ε₂ τ l point (L ++ [vk], vk1)) else probPure (L, vk)) := by @@ -569,7 +572,8 @@ lemma sv3_loop_unroll_1 [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l p split <;> try simp unfold privNoiseGuess unfold privNoiseZero - exact Eq.symm (PMF.tsum_coe (DPSystem.noise (fun x => 0) 1 ε₁ (2 * sens_cov_vk * ε₂) [])) + symm + apply PMF.tsum_coe /- ## Program version 4 @@ -580,12 +584,12 @@ lemma sv3_loop_unroll_1 [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l p -- An sv4 state is an sv1 state plus a list of presamples def sv4_state : Type := sv1_state × List ℤ -def sv4_presample [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (n : ℕ) : SLang { l : List ℤ // List.length l = n } := +def sv4_presample (ε₁ ε₂ : ℕ+) (n : ℕ) : SLang { l : List ℤ // List.length l = n } := match n with | Nat.zero => return ⟨ [], by simp ⟩ | Nat.succ n' => do - let vk1 <- @privNoiseGuess dps ε₁ ε₂ - let vks <- @sv4_presample dps ε₁ ε₂ n' + let vk1 <- privNoiseGuess ε₁ ε₂ + let vks <- sv4_presample ε₁ ε₂ n' return ⟨ vks ++ [vk1], by simp ⟩ @@ -596,15 +600,15 @@ def sv4_privMaxF (s : sv4_state) : SLang sv4_state := def sv4_privMaxC (τ : ℤ) (l : List ℕ) (st : sv4_state) : Bool := sv1_privMaxC τ l st.1 -def sv4_loop [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : SLang sv1_state := do - let presamples <- @sv4_presample dps ε₁ ε₂ point +def sv4_loop (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : SLang sv1_state := do + let presamples <- sv4_presample ε₁ ε₂ point let v <- probWhileCut (sv4_privMaxC τ l) sv4_privMaxF (point + 1) (init, presamples) return v.1 -lemma sv3_loop_unroll_1_alt [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) l point (initial_state : sv1_state) : +lemma sv3_loop_unroll_1_alt (τ : ℤ) (ε₁ ε₂ : ℕ+) l point (initial_state : sv1_state) : sv3_loop ε₁ ε₂ τ l (point + 1) initial_state = (do - let vk1 <- @privNoiseGuess dps ε₁ ε₂ + let vk1 <- privNoiseGuess ε₁ ε₂ if (sv1_privMaxC τ l initial_state) then (sv3_loop ε₁ ε₂ τ l point (initial_state.1 ++ [initial_state.2], vk1)) else probPure initial_state) := by @@ -614,12 +618,11 @@ lemma sv3_loop_unroll_1_alt [dps : DPSystem ℕ] (τ : ℤ) (ε₁ ε₂ : ℕ+) def len_list_append_rev {m n : ℕ} (x : { l : List ℤ // l.length = m }) (y: { l : List ℤ // l.length = n }) : { l : List ℤ // l.length = n + m } := ⟨ x.1 ++ y.1 , by simp [add_comm] ⟩ - lemma vector_sum_singleton (f : { l : List ℤ // l.length = 1 } -> ENNReal) (P : (x : ℤ) -> ([x].length = 1)) : (∑'(x : { l // l.length = 1 }), f x) = (∑' (x : ℤ), f ⟨ [x], P x⟩) := by sorry -def sv4_presample_split [DPSystem ℕ] (ε₁ ε₂ : ℕ+) (point : ℕ) : +def sv4_presample_split (ε₁ ε₂ : ℕ+) (point : ℕ) : sv4_presample ε₁ ε₂ (point + 1) = (do let presample_1 <- sv4_presample ε₁ ε₂ 1 @@ -634,6 +637,7 @@ def sv4_presample_split [DPSystem ℕ] (ε₁ ε₂ : ℕ+) (point : ℕ) : rw [← ENNReal.tsum_prod] rw [vector_sum_singleton _ (by simp)] + /- have X (x : ℤ): (@tsum.{0, 0} ENNReal (@NonUnitalNonAssocSemiring.toAddCommMonoid.{0} ENNReal (@NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} ENNReal @@ -661,6 +665,7 @@ def sv4_presample_split [DPSystem ℕ] (ε₁ ε₂ : ℕ+) (point : ℕ) : rw [<- ENNReal.tsum_mul_left] rw [← ENNReal.tsum_prod] simp_all [len_list_append_rev] + -/ -- Can do this by bijection (not congruence) -- By bijection @@ -675,14 +680,13 @@ def len_1_list_to_val (x : { l : List ℤ // l.length = 1 }) : ℤ := match xl with | (v :: _) => v - -- When we do induction on point, -- We will want to generalize over all init -- Unfolding this loop just moves the first presample into init -- Which can apply the IH-- since it's some arbitrary init state and a presamples state generated by one fewer point -def sv3_sv4_loop_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : - @sv3_loop dps ε₁ ε₂ τ l point init = @sv4_loop dps ε₁ ε₂ τ l point init := by +def sv3_sv4_loop_eq (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : + sv3_loop ε₁ ε₂ τ l point init = sv4_loop ε₁ ε₂ τ l point init := by revert init induction point · -- It's a mess but it works @@ -773,27 +777,35 @@ def sv3_sv4_loop_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : Li simp [sv4_presample] rw [ENNReal.tsum_eq_add_tsum_ite x] simp - - - have X : (@tsum.{0, 0} ENNReal - (@NonUnitalNonAssocSemiring.toAddCommMonoid.{0} ENNReal - (@NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} ENNReal - (@Semiring.toNonAssocSemiring.{0} ENNReal - (@OrderedSemiring.toSemiring.{0} ENNReal - (@OrderedCommSemiring.toOrderedSemiring.{0} ENNReal - (@CanonicallyOrderedCommSemiring.toOrderedCommSemiring.{0} ENNReal - ENNReal.instCanonicallyOrderedCommSemiring)))))) - ENNReal.instTopologicalSpace Int fun (x_1 : Int) => - @ite.{1} ENNReal (@Eq.{1} Int x_1 x) (Classical.propDecidable (@Eq.{1} Int x_1 x)) - (@OfNat.ofNat.{0} ENNReal 0 (@Zero.toOfNat0.{0} ENNReal instENNRealZero)) - (@ite.{1} ENNReal (@Eq.{1} Int x x_1) (Int.instDecidableEq x x_1) (@SLang.privNoiseGuess dps ε₁ ε₂ x_1) 0)) = 0 := by - simp - aesop + have X : ( @tsum.{0, 0} ENNReal + (@NonUnitalNonAssocSemiring.toAddCommMonoid.{0} ENNReal + (@NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} ENNReal + (@Semiring.toNonAssocSemiring.{0} ENNReal + (@OrderedSemiring.toSemiring.{0} ENNReal + (@OrderedCommSemiring.toOrderedSemiring.{0} ENNReal + (@CanonicallyOrderedCommSemiring.toOrderedCommSemiring.{0} ENNReal + ENNReal.instCanonicallyOrderedCommSemiring)))))) + ENNReal.instTopologicalSpace Int fun (x_1 : Int) => + @ite.{1} ENNReal (@Eq.{1} Int x_1 x) (Classical.propDecidable (@Eq.{1} Int x_1 x)) + (@OfNat.ofNat.{0} ENNReal 0 (@Zero.toOfNat0.{0} ENNReal instENNRealZero)) + (@ite.{1} ENNReal (@Eq.{1} Int x x_1) (Int.instDecidableEq x x_1) (@SLang.privNoiseGuess dps dpn ε₁ ε₂ x_1) + (@OfNat.ofNat.{0} ENNReal 0 + (@Zero.toOfNat0.{0} ENNReal + (@MulZeroClass.toZero.{0} ENNReal + (@NonUnitalNonAssocSemiring.toMulZeroClass.{0} ENNReal + (@NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} ENNReal + (@Semiring.toNonAssocSemiring.{0} ENNReal + (@OrderedSemiring.toSemiring.{0} ENNReal + (@OrderedCommSemiring.toOrderedSemiring.{0} ENNReal + (@CanonicallyOrderedCommSemiring.toOrderedCommSemiring.{0} ENNReal + ENNReal.instCanonicallyOrderedCommSemiring))))))))))) = 0 := by simp; aesop conv => enter [2, 1, 2] + skip rw [X] simp + rw [ToPresample] clear ToPresample @@ -868,17 +880,17 @@ def sv3_sv4_loop_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : Li · simp -def sv4_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := +def sv4_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => let computation : SLang ℕ := do - let τ <- @privNoiseThresh dps ε₁ ε₂ - let v0 <- @privNoiseGuess dps ε₁ ε₂ - let sk <- @sv4_loop dps ε₁ ε₂ τ l point ([], v0) + let τ <- privNoiseThresh ε₁ ε₂ + let v0 <- privNoiseGuess ε₁ ε₂ + let sk <- sv4_loop ε₁ ε₂ τ l point ([], v0) return (sv1_threshold sk) computation point -def sv3_eq_sv4 [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : - @sv3_privMax dps ε₁ ε₂ l = @sv4_privMax dps ε₁ ε₂ l := by +def sv3_eq_sv4 (ε₁ ε₂ : ℕ+) (l : List ℕ) : + sv3_privMax ε₁ ε₂ l = sv4_privMax ε₁ ε₂ l := by unfold sv3_privMax unfold sv4_privMax simp @@ -897,17 +909,17 @@ def sv5_loop (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv4_state) : SLang let sk <- probWhileCut (sv4_privMaxC τ l) sv4_privMaxF (point + 1) init return (sv1_threshold sk.1) -def sv5_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := +def sv5_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => let computation : SLang ℕ := do - let τ <- @privNoiseThresh dps ε₁ ε₂ - let v0 <- @privNoiseGuess dps ε₁ ε₂ - let presamples <- @sv4_presample dps ε₁ ε₂ point + let τ <- privNoiseThresh ε₁ ε₂ + let v0 <- privNoiseGuess ε₁ ε₂ + let presamples <- sv4_presample ε₁ ε₂ point @sv5_loop τ l point (([], v0), presamples) computation point -def sv4_eq_sv5 [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : - @sv4_privMax dps ε₁ ε₂ l = @sv5_privMax dps ε₁ ε₂ l := by +def sv4_eq_sv5 (ε₁ ε₂ : ℕ+) (l : List ℕ) : + sv4_privMax ε₁ ε₂ l = sv5_privMax ε₁ ε₂ l := by unfold sv4_privMax unfold sv5_privMax unfold sv4_loop @@ -922,34 +934,24 @@ def sv4_eq_sv5 [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : - Changes the loop from a probWhileCut into a single, deterministic, check -/ --- When you look at exactly n loops in the future, the check evaluates to true -def sv6_privMax_check (τ : ℤ) (l : List ℕ) (s : sv4_state) (n : ℕ) : Bool := - match n with - | 0 => true - | Nat.succ n' => - match s with - -- If there is more samples on the tape, call recursively - | ((past, present), (future_next :: future_rest)) => - sv4_privMaxC τ l ((past, present), (future_next :: future_rest)) ∧ - sv6_privMax_check τ l ((past ++ [present], future_next), future_rest) n' - | (_, []) => - -- Out of samples on the tape - false +-- -- When you look at exactly n loops in the future, the check evaluates to true +-- def sv6_privMax_check (τ : ℤ) (l : List ℕ) (s : sv4_state) (n : ℕ) : Bool := +-- match n with +-- | 0 => true +-- | Nat.succ n' => +-- match s with +-- -- If there is more samples on the tape, call recursively +-- | ((past, present), (future_next :: future_rest)) => +-- sv4_privMaxC τ l ((past, present), (future_next :: future_rest)) ∧ +-- sv6_privMax_check τ l ((past ++ [present], future_next), future_rest) n' +-- | (_, []) => +-- -- Out of samples on the tape +-- false -- The state sp is a "past configuration" of sc, ie. one we already checked def is_past_configuration (sp sc : sv4_state) : Prop := (sp.1.1.length ≤ sc.1.1.length) ∧ sp.1.1 ++ [sp.1.2] ++ sp.2 = sc.1.1 ++ [sc.1.2] ++ sc.2 -lemma is_past_configuration_ex1 : is_past_configuration (([], 0), []) (([], 0), []) := by - simp [is_past_configuration] - -lemma is_past_configuration_ex2 : is_past_configuration (([], 0), [1]) (([0], 1), []) := by - simp [is_past_configuration] - -lemma is_past_configuration_ex3 : is_past_configuration (([], 0), [1, 2]) (([0], 1), [2]) := by - simp [is_past_configuration] - - -- All past configurations had their loop check execute to True def sv6_privMax_hist (τ : ℤ) (l : List ℕ) (s : sv4_state) : Prop := ∀ sp, (is_past_configuration sp s) -> sv4_privMaxC τ l sp = true @@ -1117,12 +1119,12 @@ lemma sv5_loop_ind (τ : ℤ) (l : List ℕ) (eval point : ℕ) (past ff: List trivial -def sv6_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := +def sv6_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => let computation : SLang ℕ := do - let τ <- @privNoiseThresh dps ε₁ ε₂ - let v0 <- @privNoiseGuess dps ε₁ ε₂ - let presamples <- @sv4_presample dps ε₁ ε₂ point + let τ <- privNoiseThresh ε₁ ε₂ + let v0 <- privNoiseGuess ε₁ ε₂ + let presamples <- sv4_presample ε₁ ε₂ point @sv6_loop τ l point (([], v0), presamples) computation point @@ -1169,8 +1171,8 @@ def sv5_sv6_loop_eq_point (τ : ℤ) (l : List ℕ) (point eval : ℕ) (past fut simp_all [sv1_threshold] -def sv5_sv6_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : - @sv5_privMax dps ε₁ ε₂ l = @sv6_privMax dps ε₁ ε₂ l := by +def sv5_sv6_eq (ε₁ ε₂ : ℕ+) (l : List ℕ) : + sv5_privMax ε₁ ε₂ l = sv6_privMax ε₁ ε₂ l := by unfold sv5_privMax unfold sv6_privMax apply SLang.ext @@ -1196,27 +1198,27 @@ Not executable Separates out the zero case -/ -def sv7_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := +def sv7_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => let computation : SLang ℕ := do - let τ <- @privNoiseThresh dps ε₁ ε₂ - let v0 <- @privNoiseGuess dps ε₁ ε₂ + let τ <- privNoiseThresh ε₁ ε₂ + let v0 <- privNoiseGuess ε₁ ε₂ match point with | 0 => if (¬ (sv4_privMaxC τ l (([], v0), []))) then probPure point else probZero | (Nat.succ point') => do - let presamples <- @sv4_presample dps ε₁ ε₂ point' - let vk <- @privNoiseGuess dps ε₁ ε₂ + let presamples <- sv4_presample ε₁ ε₂ point' + let vk <- privNoiseGuess ε₁ ε₂ if (sv6_cond τ l (([], v0), presamples ++ [vk])) then probPure point else probZero computation point -def sv6_sv7_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : - @sv6_privMax dps ε₁ ε₂ l = @sv7_privMax dps ε₁ ε₂ l := by +def sv6_sv7_eq (ε₁ ε₂ : ℕ+) (l : List ℕ) : + sv6_privMax ε₁ ε₂ l = sv7_privMax ε₁ ε₂ l := by apply SLang.ext intro point unfold sv6_privMax @@ -1278,19 +1280,19 @@ def sv8_G (l : List ℕ) (past : List ℤ) (pres : ℤ) (future : List ℤ) : def sv8_cond (τ : ℤ) (l : List ℕ) (past : List ℤ) (pres : ℤ) (future : List ℤ) (last : ℤ) : Bool := (sv8_G l past pres future < τ) ∧ (sv8_sum l (past ++ [pres] ++ future) last ≥ τ) -def sv8_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := +def sv8_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => let computation : SLang ℕ := do - let τ <- @privNoiseThresh dps ε₁ ε₂ - let v0 <- @privNoiseGuess dps ε₁ ε₂ + let τ <- privNoiseThresh ε₁ ε₂ + let v0 <- privNoiseGuess ε₁ ε₂ match point with | 0 => if (sv8_sum l [] v0 ≥ τ) then probPure point else probZero | (Nat.succ point') => do - let presamples <- @sv4_presample dps ε₁ ε₂ point' - let vk <- @privNoiseGuess dps ε₁ ε₂ + let presamples <- sv4_presample ε₁ ε₂ point' + let vk <- privNoiseGuess ε₁ ε₂ if (sv8_cond τ l [] v0 presamples vk) then probPure point else probZero @@ -1326,8 +1328,8 @@ lemma sv7_sv8_cond_eq (τ : ℤ) (l : List ℕ) (v0 : ℤ) (vs : List ℤ) (vk : cases (decide (sv8_G l (init ++ [vi]) vi_1 rest < τ)) <;> simp simp [sv4_privMaxC, sv1_privMaxC, sv8_sum, sv1_threshold, sv1_noise] -def sv7_sv8_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : - @sv7_privMax dps ε₁ ε₂ l = @sv8_privMax dps ε₁ ε₂ l := by +def sv7_sv8_eq (ε₁ ε₂ : ℕ+) (l : List ℕ) : + sv7_privMax ε₁ ε₂ l = sv8_privMax ε₁ ε₂ l := by apply SLang.ext intro point unfold sv7_privMax @@ -1349,28 +1351,28 @@ Rewritten so that the randomness we will cancel out is right at the front -/ -def sv9_privMax [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := +def sv9_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := fun (point : ℕ) => let computation : SLang ℕ := do match point with | 0 => do - let τ <- @privNoiseThresh dps ε₁ ε₂ - let v0 <- @privNoiseGuess dps ε₁ ε₂ + let τ <- privNoiseThresh ε₁ ε₂ + let v0 <- privNoiseGuess ε₁ ε₂ if (sv8_sum l [] v0 ≥ τ) then probPure point else probZero | (Nat.succ point') => do - let v0 <- @privNoiseGuess dps ε₁ ε₂ - let presamples <- @sv4_presample dps ε₁ ε₂ point' - let τ <- @privNoiseThresh dps ε₁ ε₂ - let vk <- @privNoiseGuess dps ε₁ ε₂ + let v0 <- privNoiseGuess ε₁ ε₂ + let presamples <- sv4_presample ε₁ ε₂ point' + let τ <- privNoiseThresh ε₁ ε₂ + let vk <- privNoiseGuess ε₁ ε₂ if (sv8_cond τ l [] v0 presamples vk) then probPure point else probZero computation point -def sv8_sv9_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : - @sv8_privMax dps ε₁ ε₂ l = @sv8_privMax dps ε₁ ε₂ l := by +def sv8_sv9_eq (ε₁ ε₂ : ℕ+) (l : List ℕ) : + sv8_privMax ε₁ ε₂ l = sv8_privMax ε₁ ε₂ l := by apply SLang.ext intro point sorry @@ -1379,5 +1381,5 @@ def sv8_sv9_eq [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : /-- sv9 normalizes because sv1 normalizes -/ -def sv9_privMax_pmf [dps : DPSystem ℕ] (ε₁ ε₂ : ℕ+) (l : List ℕ) : PMF ℕ := - ⟨ @sv9_privMax dps ε₁ ε₂ l, sorry ⟩ +def sv9_privMax_pmf (ε₁ ε₂ : ℕ+) (l : List ℕ) : PMF ℕ := + ⟨ sv9_privMax ε₁ ε₂ l, sorry ⟩ From 60c7b746bf569bb226b163c6d4a2b5faf5adb28a Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Tue, 5 Nov 2024 14:35:46 -0500 Subject: [PATCH 065/100] move executable code to a new file --- .../Queries/UnboundedMax/Basic.lean | 6 +- .../Queries/UnboundedMax/Code.lean | 113 ++++++++++++++++++ .../Queries/UnboundedMax/Properties.lean | 90 +------------- 3 files changed, 115 insertions(+), 94 deletions(-) create mode 100644 SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Basic.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Basic.lean index 544072bb..7a92d7ad 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Basic.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Basic.lean @@ -4,10 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus de Medeiros -/ --- import SampCert.DifferentialPrivacy.Queries.UnboundedMax.Code - import SampCert.DifferentialPrivacy.Queries.UnboundedMax.Properties import SampCert.DifferentialPrivacy.Queries.UnboundedMax.Privacy - --- import SampCert.DifferentialPrivacy.Queries.UnboundedMax.Pure --- import SampCert.DifferentialPrivacy.Queries.UnboundedMax.ZeroConcentrated +import SampCert.DifferentialPrivacy.Queries.UnboundedMax.Code diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean new file mode 100644 index 00000000..325c731a --- /dev/null +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean @@ -0,0 +1,113 @@ +/- +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: Markus de Medeiros +-/ + +import SampCert.DifferentialPrivacy.Abstract +import SampCert.DifferentialPrivacy.Pure.System + +namespace SLang + +variable (ε₁ ε₂ : ℕ+) + +variable [dps : DPSystem ℕ] +variable [dpn : DPNoise dps] + +/-- +Sensitivity bound for the τ threshold +-/ +def sens_cov_τ : ℕ+ := 1 + +/-- +Sensitivity bound for each upper bound attempt +-/ +def sens_cov_vk : ℕ+ := 1 + sens_cov_τ + + +/- +## Executable code for the sparse vector mechanism +-/ + +/-- +Sum over a list, clipping each element to a maximum. + +Similar to exactBoundedSum, however exactClippedSum allows m = 0. +-/ +def exactClippedSum (m : ℕ) (l : List ℕ) : ℤ := + List.sum (List.map (fun n : ℕ => (Nat.min n m)) l) + +/-- +Rate at which a given clipping thresold is impacting the accuracy of the sum. + +Always negative or zero. +-/ +def exactDiffSum (m : ℕ) (l : List ℕ) : ℤ := exactClippedSum m l - exactClippedSum (m + 1) l + +/-- +Noise the constant 0 value using the abstract noise function. + +This looks strange, but will specialize to Lap(ε₁/ε₂, 0) in the pure DP case. +-/ +def privNoiseZero (ε₁ ε₂ : ℕ+) : SLang ℤ := dpn.noise (fun _ => 0) 1 ε₁ ε₂ [] + + +def privNoiseGuess (ε₁ ε₂ : ℕ+) : SLang ℤ := privNoiseZero ε₁ (2 * sens_cov_vk * ε₂) + +def privNoiseThresh (ε₁ ε₂ : ℕ+) : SLang ℤ := privNoiseZero ε₁ (2 * sens_cov_τ * ε₂) + +/- +## Program version 0 + - Executable + - Tracks single state +-/ + +def sv0_state : Type := ℕ × ℤ + +def sv0_threshold (s : sv0_state) : ℕ := s.1 + +def sv0_noise (s : sv0_state) : ℤ := s.2 + +def sv0_privMaxC (τ : ℤ) (l : List ℕ) (s : sv0_state) : Bool := + decide (exactDiffSum (sv0_threshold s) l + (sv0_noise s) < τ) + +def sv0_privMaxF (ε₁ ε₂ : ℕ+) (s : sv0_state) : SLang sv0_state := do + let vn <- privNoiseGuess ε₁ ε₂ + let n := (sv0_threshold s) + 1 + return (n, vn) + +def sv0_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do + let τ <- privNoiseThresh ε₁ ε₂ + let v0 <- privNoiseGuess ε₁ ε₂ + let sk <- probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (0, v0) + return (sv0_threshold sk) + +/- +## Program version 1 + - Executable + - Tracks history of samples +-/ + +def sv1_state : Type := List ℤ × ℤ + +def sv1_threshold (s : sv1_state) : ℕ := List.length s.1 + +def sv1_noise (s : sv1_state) : ℤ := s.2 + +def sv1_privMaxC (τ : ℤ) (l : List ℕ) (s : sv1_state) : Bool := + decide (exactDiffSum (sv1_threshold s) l + (sv1_noise s) < τ) + +def sv1_privMaxF (ε₁ ε₂ : ℕ+) (s : sv1_state) : SLang sv1_state := do + let vn <- privNoiseGuess ε₁ ε₂ + return (s.1 ++ [s.2], vn) + +def sv1_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do + let τ <- privNoiseThresh ε₁ ε₂ + let v0 <- privNoiseGuess ε₁ ε₂ + let sk <- probWhile (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) ([], v0) + return (sv1_threshold sk) + + + + +end SLang diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 13cd042d..f7ac1c93 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -5,6 +5,7 @@ Authors: Markus de Medeiros -/ import SampCert.DifferentialPrivacy.Abstract +import SampCert.DifferentialPrivacy.Queries.UnboundedMax.Code noncomputable section open Classical @@ -14,9 +15,6 @@ namespace SLang variable [dps : DPSystem ℕ] variable [dpn : DPNoise dps] -def sens_cov_τ : ℕ+ := 1 -def sens_cov_vk : ℕ+ := 1 + sens_cov_τ - /-- Stronger congruence rule for probBind: The bound-to functions have to be equal only on the support of the bound-from function. @@ -45,38 +43,6 @@ lemma probBind_congr_strong (p : SLang T) (f : T -> SLang U) (g : T -> SLang U) apply Hp -/- -## Helper programs --/ - -/-- -Sum over a list, clipping each element to a maximum. - -Similar to exactBoundedSum, however exactClippedSum allows m = 0. --/ -def exactClippedSum (m : ℕ) (l : List ℕ) : ℤ := - List.sum (List.map (fun n : ℕ => (Nat.min n m)) l) - -/-- -Rate at which a given clipping thresold is impacting the accuracy of the sum. - -Always negative or zero. --/ -def exactDiffSum (m : ℕ) (l : List ℕ) : ℤ := exactClippedSum m l - exactClippedSum (m + 1) l - -/-- -Noise the constant 0 value using the abstract noise function. - -This looks strange, but will specialize to Lap(ε₁/ε₂, 0) in the pure DP case. --/ -def privNoiseZero (ε₁ ε₂ : ℕ+) : SLang ℤ := dpn.noise (fun _ => 0) 1 ε₁ ε₂ [] - - -def privNoiseGuess (ε₁ ε₂ : ℕ+) : SLang ℤ := privNoiseZero ε₁ (2 * sens_cov_vk * ε₂) - -def privNoiseThresh (ε₁ ε₂ : ℕ+) : SLang ℤ := privNoiseZero ε₁ (2 * sens_cov_τ * ε₂) - - /- Not used for anything, but to give confidence in our definitions @@ -122,58 +88,6 @@ Not used for anything, but to give confidence in our definitions -- sorry -- · trivial - -/- -## Program version 0 - - Executable - - Tracks single state --/ - -def sv0_state : Type := ℕ × ℤ - -def sv0_threshold (s : sv0_state) : ℕ := s.1 - -def sv0_noise (s : sv0_state) : ℤ := s.2 - -def sv0_privMaxC (τ : ℤ) (l : List ℕ) (s : sv0_state) : Bool := - decide (exactDiffSum (sv0_threshold s) l + (sv0_noise s) < τ) - -def sv0_privMaxF (ε₁ ε₂ : ℕ+) (s : sv0_state) : SLang sv0_state := do - let vn <- privNoiseGuess ε₁ ε₂ - let n := (sv0_threshold s) + 1 - return (n, vn) - -def sv0_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do - let τ <- privNoiseThresh ε₁ ε₂ - let v0 <- privNoiseGuess ε₁ ε₂ - let sk <- probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (0, v0) - return (sv0_threshold sk) - -/- -## Program version 1 - - Executable - - Tracks history of samples --/ - -def sv1_state : Type := List ℤ × ℤ - -def sv1_threshold (s : sv1_state) : ℕ := List.length s.1 - -def sv1_noise (s : sv1_state) : ℤ := s.2 - -def sv1_privMaxC (τ : ℤ) (l : List ℕ) (s : sv1_state) : Bool := - decide (exactDiffSum (sv1_threshold s) l + (sv1_noise s) < τ) - -def sv1_privMaxF (ε₁ ε₂ : ℕ+) (s : sv1_state) : SLang sv1_state := do - let vn <- privNoiseGuess ε₁ ε₂ - return (s.1 ++ [s.2], vn) - -def sv1_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do - let τ <- privNoiseThresh ε₁ ε₂ - let v0 <- privNoiseGuess ε₁ ε₂ - let sk <- probWhile (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) ([], v0) - return (sv1_threshold sk) - /-- History-aware progam computes the same as the history-agnostic program -/ @@ -222,8 +136,6 @@ lemma sv0_eq_sv1 ε₁ ε₂ l : sv0_privMax ε₁ ε₂ l = sv1_privMax ε₁ -- rw [tsum_ite_eq] -- simp [sv1_threshold] - - sorry From 22f3fd290218989571ebc87bd0d134a23b7f6550 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Tue, 5 Nov 2024 14:54:42 -0500 Subject: [PATCH 066/100] add sparse vector test --- .../Queries/UnboundedMax/Code.lean | 7 ++++++ Test.lean | 25 +++++++++++++++++++ 2 files changed, 32 insertions(+) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean index 325c731a..ef4c1d08 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean @@ -82,6 +82,10 @@ def sv0_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do let sk <- probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (0, v0) return (sv0_threshold sk) +def sv0_privMax_PMF (ε₁ ε₂ : ℕ+) (l : List ℕ) : SPMF ℕ := + ⟨ sv0_privMax ε₁ ε₂ l, + by sorry ⟩ + /- ## Program version 1 - Executable @@ -108,6 +112,9 @@ def sv1_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do return (sv1_threshold sk) +def sv1_privMax_PMF (ε₁ ε₂ : ℕ+) (l : List ℕ) : SPMF ℕ := + ⟨ sv1_privMax ε₁ ε₂ l, + by sorry ⟩ end SLang diff --git a/Test.lean b/Test.lean index c30f3493..0f4cb458 100644 --- a/Test.lean +++ b/Test.lean @@ -244,7 +244,32 @@ def statistical_tests : IO Unit := do IO.println s!"num = {(num : ℕ)}, den = {(den : ℕ)}, mix = {mix}" test num den mix 100000 0.1 +def sparseVector_tests : IO Unit := do + let samples := 10000 + let unif_ub := 10 + let data : List ℕ <- List.mapM (fun _ => run <| (SLang.UniformSample_PMF unif_ub)) (List.replicate samples 0) + + let num : ℕ+ := 9 + let den : ℕ+ := 2 + let num_trials := 5 + + IO.println s!"[query] testing sparse vector max, ({(num : ℕ)} / {(den : ℕ)})-DP" + IO.println s!"data := {samples} uniform samples of [0, {(unif_ub : ℕ)}): {(data.take 20)}..." + IO.println "" + + for i in [:num_trials] do + let ct <- run <| @sv0_privMax_PMF PureDPSystem laplace_pureDPSystem num den data + IO.println s!"#{i} sv0 max: {ct}" + IO.println "" + + for i in [:num_trials] do + let ct <- run <| @sv1_privMax_PMF PureDPSystem laplace_pureDPSystem num den data + IO.println s!"#{i} sv1 max: {ct}" + IO.println "" + + def main : IO Unit := do + sparseVector_tests query_tests statistical_tests From d08f8483295940f67197de978fe0147c460381f2 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Tue, 5 Nov 2024 14:56:38 -0500 Subject: [PATCH 067/100] tweak numbers --- Test.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Test.lean b/Test.lean index 0f4cb458..33bce2e3 100644 --- a/Test.lean +++ b/Test.lean @@ -246,11 +246,11 @@ def statistical_tests : IO Unit := do def sparseVector_tests : IO Unit := do let samples := 10000 - let unif_ub := 10 + let unif_ub := 100 let data : List ℕ <- List.mapM (fun _ => run <| (SLang.UniformSample_PMF unif_ub)) (List.replicate samples 0) - let num : ℕ+ := 9 - let den : ℕ+ := 2 + let num : ℕ+ := 1 + let den : ℕ+ := 4 let num_trials := 5 IO.println s!"[query] testing sparse vector max, ({(num : ℕ)} / {(den : ℕ)})-DP" From 5408fe07d2cf912ea3d2546550ea612e0cab93a5 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Tue, 5 Nov 2024 15:28:38 -0500 Subject: [PATCH 068/100] checkpoint --- SampCert/Foundations/While.lean | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/SampCert/Foundations/While.lean b/SampCert/Foundations/While.lean index 380d0e0b..b9548be5 100644 --- a/SampCert/Foundations/While.lean +++ b/SampCert/Foundations/While.lean @@ -58,4 +58,19 @@ theorem probWhile_apply (cond : T → Bool) (body : T → SLang T) (init : T) (x · apply probWhileCut_monotonic · apply H + +-- A version of probWhile that proves normalization, when its arguments are all normal +def SPMF_while (cond : T → Bool) (body : T → SPMF T) (init : T) : SPMF T := + ⟨ probWhile cond (fun x => body x) init, + by + unfold probWhile + sorry ⟩ + +lemma SPMF_while_eq (cond : T → Bool) (body : T → SPMF T) (init : T) : + ((SPMF_while cond body init) : SLang T) = (probWhile cond (fun x => body x) init) := by + simp [SPMF_while] + + + + end SLang From 6f8c6124a888b50cf8453ec0c6654d866d42be72 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Wed, 6 Nov 2024 10:06:36 -0500 Subject: [PATCH 069/100] vector_sum_singleton --- .../Queries/UnboundedMax/Code.lean | 50 +++++++++++++++++-- .../Queries/UnboundedMax/Properties.lean | 14 +++++- SampCert/Foundations/While.lean | 15 ------ 3 files changed, 60 insertions(+), 19 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean index ef4c1d08..a64f71c4 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean @@ -84,7 +84,8 @@ def sv0_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do def sv0_privMax_PMF (ε₁ ε₂ : ℕ+) (l : List ℕ) : SPMF ℕ := ⟨ sv0_privMax ε₁ ε₂ l, - by sorry ⟩ + by + sorry ⟩ /- ## Program version 1 @@ -114,7 +115,50 @@ def sv1_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do def sv1_privMax_PMF (ε₁ ε₂ : ℕ+) (l : List ℕ) : SPMF ℕ := ⟨ sv1_privMax ε₁ ε₂ l, - by sorry ⟩ - + by + -- Difficult + sorry ⟩ end SLang + +/- +lemma tsum_iSup_comm (f : T -> U -> ENNReal) : ∑' (a : T), ⨆ (b : U), f a b = ⨆ (b : U), ∑' (a : T), f a b := by + + sorry + +-- A version of probWhile that proves normalization, when its arguments are all normal +def SPMF_while (cond : T → Bool) (body : T → SPMF T) (init : T) : SPMF T := + ⟨ probWhile cond (fun x => body x) init, + by + apply (Summable.hasSum_iff ENNReal.summable).mpr + unfold probWhile + rw [tsum_iSup_comm] + conv => + enter [1, 1, cut] + -- We can establish the ≤ direction easily + -- ≥ is not true + + + + -- apply iSup_eq_of_tendsto + -- · sorry + -- sorry + + -- rw [ENNReal.tsum_eq_iSup_sum] + -- apply eq_iff_le_not_lt.mpr + -- apply And.intro + -- · apply iSup_le_iff.mpr + -- sorry + -- · simp + -- apply le_iSup_iff.mpr + -- intro b H + -- sorry + ⟩ + +lemma SPMF_while_eq (cond : T → Bool) (body : T → SPMF T) (init : T) : + ((SPMF_while cond body init) : SLang T) = (probWhile cond (fun x => body x) init) := by + simp [SPMF_while] + + + +-/ diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index f7ac1c93..96bcd393 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -532,7 +532,19 @@ def len_list_append_rev {m n : ℕ} (x : { l : List ℤ // l.length = m }) (y: { lemma vector_sum_singleton (f : { l : List ℤ // l.length = 1 } -> ENNReal) (P : (x : ℤ) -> ([x].length = 1)) : (∑'(x : { l // l.length = 1 }), f x) = (∑' (x : ℤ), f ⟨ [x], P x⟩) := by - sorry + apply @tsum_eq_tsum_of_ne_zero_bij + case i => + simp [Function.support, DFunLike.coe] + exact fun x => ⟨ [x.1], by simp ⟩ + · simp [Function.Injective] + · simp [Function.support, Set.range] + intro L HL HN + cases L + · simp at HL + rename_i v R + cases R + · exists v + · simp at HL def sv4_presample_split (ε₁ ε₂ : ℕ+) (point : ℕ) : sv4_presample ε₁ ε₂ (point + 1) = diff --git a/SampCert/Foundations/While.lean b/SampCert/Foundations/While.lean index b9548be5..380d0e0b 100644 --- a/SampCert/Foundations/While.lean +++ b/SampCert/Foundations/While.lean @@ -58,19 +58,4 @@ theorem probWhile_apply (cond : T → Bool) (body : T → SLang T) (init : T) (x · apply probWhileCut_monotonic · apply H - --- A version of probWhile that proves normalization, when its arguments are all normal -def SPMF_while (cond : T → Bool) (body : T → SPMF T) (init : T) : SPMF T := - ⟨ probWhile cond (fun x => body x) init, - by - unfold probWhile - sorry ⟩ - -lemma SPMF_while_eq (cond : T → Bool) (body : T → SPMF T) (init : T) : - ((SPMF_while cond body init) : SLang T) = (probWhile cond (fun x => body x) init) := by - simp [SPMF_while] - - - - end SLang From 45a55c792d1c2473a0a1a915d8d1aab6f401257f Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Wed, 6 Nov 2024 10:07:44 -0500 Subject: [PATCH 070/100] minor --- .../Queries/UnboundedMax/Properties.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 96bcd393..65293d3b 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -545,6 +545,7 @@ lemma vector_sum_singleton (f : { l : List ℤ // l.length = 1 } -> ENNReal) (P cases R · exists v · simp at HL + · simp [Function.support, DFunLike.coe] def sv4_presample_split (ε₁ ε₂ : ℕ+) (point : ℕ) : sv4_presample ε₁ ε₂ (point + 1) = @@ -561,7 +562,6 @@ def sv4_presample_split (ε₁ ε₂ : ℕ+) (point : ℕ) : rw [← ENNReal.tsum_prod] rw [vector_sum_singleton _ (by simp)] - /- have X (x : ℤ): (@tsum.{0, 0} ENNReal (@NonUnitalNonAssocSemiring.toAddCommMonoid.{0} ENNReal (@NonAssocSemiring.toNonUnitalNonAssocSemiring.{0} ENNReal @@ -573,7 +573,7 @@ def sv4_presample_split (ε₁ ε₂ : ℕ+) (point : ℕ) : ENNReal.instTopologicalSpace Int fun (x_1 : Int) => @ite.{1} ENNReal (@Eq.{1} Int x_1 x) (Classical.propDecidable (@Eq.{1} Int x_1 x)) (@OfNat.ofNat.{0} ENNReal 0 (@Zero.toOfNat0.{0} ENNReal instENNRealZero)) - (@ite.{1} ENNReal (@Eq.{1} Int x x_1) (Int.instDecidableEq x x_1) (@SLang.privNoiseGuess _ ε₁ ε₂ x_1) + (@ite.{1} ENNReal (@Eq.{1} Int x x_1) (Int.instDecidableEq x x_1) (@SLang.privNoiseGuess _ _ ε₁ ε₂ x_1) 0)) = 0 := by simp; aesop conv => enter [2, 1, x, 1] @@ -589,7 +589,7 @@ def sv4_presample_split (ε₁ ε₂ : ℕ+) (point : ℕ) : rw [<- ENNReal.tsum_mul_left] rw [← ENNReal.tsum_prod] simp_all [len_list_append_rev] - -/ + -- Can do this by bijection (not congruence) -- By bijection From c64d2d9309c9f924b4803698b6f8d7dd9a719965 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Wed, 6 Nov 2024 13:20:20 -0500 Subject: [PATCH 071/100] sv4_presample_split progress --- .../Queries/UnboundedMax/Properties.lean | 190 +++++++++++++++++- 1 file changed, 184 insertions(+), 6 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 65293d3b..91b7f1a8 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -547,6 +547,118 @@ lemma vector_sum_singleton (f : { l : List ℤ // l.length = 1 } -> ENNReal) (P · simp at HL · simp [Function.support, DFunLike.coe] + +def vsm_0 (x : {l : List ℤ // l.length = n + 1}) : ℤ := + List.headI x.1 + +def vsm_rest (x : {l : List ℤ // l.length = n + 1}) : {l : List ℤ // l.length = n } := + ⟨ List.tail x.1, by simp ⟩ + +def vsm_last (x : {l : List ℤ // l.length = n + 1}) : ℤ := + List.getLastI x.1 + +def vsm_init (x : {l : List ℤ // l.length = n + 1}) : {l : List ℤ // l.length = n } := + ⟨ List.dropLast x.1, by simp ⟩ + + + + +lemma vector_sum_merge n (f : ℤ × { l : List ℤ // l.length = n } -> ENNReal) : + (∑'p, f p) = ∑'(p : {l : List ℤ // l.length = n + 1}), f (vsm_0 p, vsm_rest p) := by + apply @tsum_eq_tsum_of_ne_zero_bij + case i => + simp [Function.support, DFunLike.coe] + exact fun x => (vsm_0 x.1, vsm_rest x.1) + · simp [Function.Injective] + simp [vsm_0, vsm_rest] + intro L1 HL1 HL1f L2 HL2 HL2f Heq1 Heq2 + cases L1 + · simp at HL1 + cases L2 + · simp at HL2 + simp_all + · simp [Function.support, Set.range] + intro z L HL HF + exists (z :: L) + simp + exists HL + · simp [Function.support, DFunLike.coe] + + + +-- Split in the other order, used as a helper function +lemma sv4_presample_split' (ε₁ ε₂ : ℕ+) (point : ℕ) (z : ℤ) (p : { l : List ℤ // List.length l = point }) : + privNoiseGuess ε₁ ε₂ z * sv4_presample ε₁ ε₂ point p = + sv4_presample ε₁ ε₂ (point + 1) ⟨ (p.1 ++ [z]), by simp ⟩ := by + rcases p with ⟨ L, HL ⟩ + revert HL + induction L + · intro HL + simp at HL + simp + conv => + rhs + unfold sv4_presample + unfold sv4_presample + split + · simp + rw [ENNReal.tsum_eq_add_tsum_ite z] + conv => + lhs + rw [<- (add_zero (privNoiseGuess _ _ _))] + congr 1 + · simp + · symm + simp + aesop + · exfalso + simp at HL + + · rename_i L0 LL _ + intro HL + simp + conv => + rhs + unfold sv4_presample + simp + conv => + enter [2, 1, a] + rw [← ENNReal.tsum_mul_left] + enter [1, b] + simp + rw [← ENNReal.tsum_prod] + rw [ENNReal.tsum_eq_add_tsum_ite (z, ⟨ L0 :: LL, HL ⟩)] + conv => + lhs + rw [<- (add_zero (_ * _))] + congr 1 + · simp + · symm + simp + intro A B C D E + exfalso + apply (congrArg List.reverse) at E + simp at E + cases E + apply D + · symm + trivial + rename_i E + apply (congrArg List.reverse) at E + simp at E + symm + trivial + + +lemma sv4_presample_perm (ε₁ ε₂ : ℕ+) (point : ℕ) (z : ℤ) (p : { l : List ℤ // List.length l = point }) H1 H2 : + sv4_presample ε₁ ε₂ (point + 1) ⟨p.1 ++ [z], H1⟩ = sv4_presample ε₁ ε₂ (point + 1) ⟨z :: p.1, H2⟩ := by + sorry + +lemma get_last_lemma (L : List ℤ) (H : L.length > 0) H : L.getLastI = L.getLast H := by + sorry + + +-- Splits and rearranges the functions def sv4_presample_split (ε₁ ε₂ : ℕ+) (point : ℕ) : sv4_presample ε₁ ε₂ (point + 1) = (do @@ -590,13 +702,79 @@ def sv4_presample_split (ε₁ ε₂ : ℕ+) (point : ℕ) : rw [← ENNReal.tsum_prod] simp_all [len_list_append_rev] - -- Can do this by bijection (not congruence) - - -- By bijection - -- #check tsum_eq_tsum_of_ne_zero_bij - -- rw [ENNReal.tsum_eq_add_tsum_ite (Int.ofNat point)] + -- Join the sv4_presamples + conv => + enter [1, 1, p] + rw [sv4_presample_split'] + conv => + enter [2, 1, p] + rw [sv4_presample_split'] + rw [sv4_presample_perm _ _ _ _ _ _ (by simp)] + rw [vector_sum_merge] + rw [vector_sum_merge] - sorry + -- Finally, they are equal by bijection + apply @tsum_eq_tsum_of_ne_zero_bij + case i => exact fun y => ⟨ (vsm_last y.1) :: (vsm_init y.1), by simp ⟩ + · simp [vsm_rest, vsm_0, DFunLike.coe] + simp [Function.Injective] + intro L1 HL1 _ _ L2 HL2 _ _ Heq + cases L1 + · simp at HL1 + cases L2 + · simp at HL2 + simp_all [vsm_init] + · simp [Function.support, Set.range] + intro L1 HL1 _ _ + cases L1 + · simp at HL1 + rename_i l ll H1 H2 + simp [vsm_0, vsm_rest] + simp [vsm_0, vsm_rest] at H1 + exists (ll ++ [l]) + have h : ((ll ++ [l]).length = point + 1) := by + simp [List.length] at HL1 + simp + trivial + exists h + apply And.intro + · apply And.intro + · rw [H1] + cases ll <;> simp [List.headI, List.tail] + · intro X + apply H2 + rw [<- X] + congr + cases ll <;> simp [vsm_rest, vsm_0, List.headI, List.tail] + · simp [vsm_last, vsm_init] + rw [List.getLastI_eq_getLast?] + rw [List.getLast?_concat] + + -- Now the conditionals are equal + simp [Function.support, DFunLike.coe, vsm_rest, vsm_0, vsm_init, vsm_last] + intro L HL Hf + have X : L.headI :: L.tail = L.dropLast ++ [L.getLastI] := by + clear Hf + cases L + · exfalso + simp at HL + rename_i l ll + simp + generalize HLL : l :: ll = L + symm + have HL' : L.length > 0 := by simp_all + clear HLL + conv => + rhs + rw [<- @List.dropLast_append_getLast _ L (by aesop)] + congr + apply get_last_lemma + trivial + congr 1 + · simp_all + congr + symm + trivial def len_1_list_to_val (x : { l : List ℤ // l.length = 1 }) : ℤ := From f94da17ef434553f7b368fa9d18ebe38757213fd Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Wed, 6 Nov 2024 13:45:08 -0500 Subject: [PATCH 072/100] progress on presample_norm_lemma --- .../Queries/UnboundedMax/Properties.lean | 59 +++++++++++++++++-- 1 file changed, 55 insertions(+), 4 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 91b7f1a8..1e0e9fc8 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -654,7 +654,7 @@ lemma sv4_presample_perm (ε₁ ε₂ : ℕ+) (point : ℕ) (z : ℤ) (p : { l : sv4_presample ε₁ ε₂ (point + 1) ⟨p.1 ++ [z], H1⟩ = sv4_presample ε₁ ε₂ (point + 1) ⟨z :: p.1, H2⟩ := by sorry -lemma get_last_lemma (L : List ℤ) (H : L.length > 0) H : L.getLastI = L.getLast H := by +lemma get_last_lemma (L : List ℤ) H : L.getLastI = L.getLast H := by sorry @@ -769,7 +769,6 @@ def sv4_presample_split (ε₁ ε₂ : ℕ+) (point : ℕ) : rw [<- @List.dropLast_append_getLast _ L (by aesop)] congr apply get_last_lemma - trivial congr 1 · simp_all congr @@ -787,6 +786,59 @@ def len_1_list_to_val (x : { l : List ℤ // l.length = 1 }) : ℤ := -- Unfolding this loop just moves the first presample into init -- Which can apply the IH-- since it's some arbitrary init state and a presamples state generated by one fewer point + +lemma presample_norm_lemma (point : ℕ) (ε₁ ε₂ : ℕ+) : + ∑' (a : { l // l.length = point }), sv4_presample ε₁ ε₂ point a = 1 := by + induction point + · simp [sv4_presample] + rw [ENNReal.tsum_eq_add_tsum_ite ⟨ [], sv4_presample.proof_3 ⟩] + conv => + rhs + rw [<- add_zero 1] + congr <;> simp + · rename_i n IH + + -- sv4_presample_split' + suffices (∑' (a : ℤ × { l : List ℤ // l.length = n }), privNoiseGuess ε₁ ε₂ a.1 * sv4_presample ε₁ ε₂ n a.2 = 1) by + conv at this => + enter [1, 1, a] + rw [sv4_presample_split'] + rw [<- this] + symm + -- FIXME probably the wrong bijection + apply @tsum_eq_tsum_of_ne_zero_bij + case i => + simp [Function.support, DFunLike.coe] + exact fun x => (vsm_0 x.1, vsm_rest x.1) + · simp [Function.Injective] + simp [vsm_0, vsm_rest] + intro L1 HL1 HL1f L2 HL2 HL2f Heq1 Heq2 + cases L1 + · simp at HL1 + cases L2 + · simp at HL2 + simp_all + · simp [Function.support, Set.range] + intro z L HL HF + exists (z :: L) + simp + exists HL + sorry + · simp [Function.support, DFunLike.coe] + sorry + + rw [ENNReal.tsum_prod'] + conv => + enter [1, 1, a] + simp + rw [ENNReal.tsum_mul_left] + rw [IH] + simp + + -- Change noise to SPMF + sorry + + def sv3_sv4_loop_eq (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : sv3_loop ε₁ ε₂ τ l point init = sv4_loop ε₁ ε₂ τ l point init := by revert init @@ -977,8 +1029,7 @@ def sv3_sv4_loop_eq (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) rw [<- mul_one (sv4_presample _ _ _ _)] congr 1 symm - sorry - -- Sum of PMF is 1 + apply presample_norm_lemma · simp From 1a9496acb9bc36df53a5fbef448a0b64ce2c4d2c Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Wed, 6 Nov 2024 14:02:18 -0500 Subject: [PATCH 073/100] DS0 --- .../Queries/UnboundedMax/Privacy.lean | 39 ++++++++++++++++-- .../Queries/UnboundedMax/Properties.lean | 41 ++++++++++--------- 2 files changed, 56 insertions(+), 24 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean index 3aca4b3e..ae79fa20 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean @@ -179,6 +179,8 @@ lemma exactDiffSum_singleton_le_1 : -1 ≤ exactDiffSum point [v] := by · linarith +set_option pp.coercions false + -- Coercions nonsense lemma DS0 (H : Neighbour L1 L2) : ((exactDiffSum 0 L1) : ℝ) - (exactDiffSum 0 L2) ≤ 1 := by cases H @@ -188,21 +190,50 @@ lemma DS0 (H : Neighbour L1 L2) : ((exactDiffSum 0 L1) : ℝ) - (exactDiffSum 0 simp apply neg_le.mp have _ := @exactDiffSum_singleton_le_1 0 C - sorry + simp [exactDiffSum] + simp [exactClippedSum] · rename_i A B C H1 H2 rw [H1, H2] repeat rw [exactDiffSum_append] simp have _ := @exactDiffSum_nonpos 0 C - sorry + simp [exactDiffSum] + simp [exactClippedSum] + cases (Classical.em ((B : ℝ) ≤ 1)) + · rename_i h + rw [min_eq_left_iff.mpr h] + linarith + · rename_i h + rw [min_eq_right_iff.mpr ?G1] + case G1 => linarith + simp · rename_i A B C D H1 H2 rw [H1, H2] repeat rw [exactDiffSum_append] simp only [Int.cast_add, add_sub_add_right_eq_sub, add_sub_add_left_eq_sub] - sorry - + simp [exactDiffSum] + simp [exactClippedSum] + cases (Classical.em ((B : ℝ) ≤ 1)) + · cases (Classical.em ((D : ℝ) ≤ 1)) + · rename_i hb hd + rw [min_eq_left_iff.mpr hb] + rw [min_eq_left_iff.mpr hd] + linarith + · rename_i hb hd + rw [min_eq_left_iff.mpr hb] + rw [min_eq_right_iff.mpr (by linarith)] + linarith + · cases (Classical.em ((D : ℝ) ≤ 1)) + · rename_i hb hd + rw [min_eq_left_iff.mpr hd] + rw [min_eq_right_iff.mpr (by linarith)] + linarith + · rename_i hb hd + rw [min_eq_right_iff.mpr (by linarith)] + rw [min_eq_right_iff.mpr (by linarith)] + linarith lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : PureDPSystem.prop (@sv9_privMax_pmf PureDPSystem laplace_pureDPSystem ε₁ ε₂) ε := by diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 1e0e9fc8..52547715 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -806,26 +806,27 @@ lemma presample_norm_lemma (point : ℕ) (ε₁ ε₂ : ℕ+) : rw [<- this] symm -- FIXME probably the wrong bijection - apply @tsum_eq_tsum_of_ne_zero_bij - case i => - simp [Function.support, DFunLike.coe] - exact fun x => (vsm_0 x.1, vsm_rest x.1) - · simp [Function.Injective] - simp [vsm_0, vsm_rest] - intro L1 HL1 HL1f L2 HL2 HL2f Heq1 Heq2 - cases L1 - · simp at HL1 - cases L2 - · simp at HL2 - simp_all - · simp [Function.support, Set.range] - intro z L HL HF - exists (z :: L) - simp - exists HL - sorry - · simp [Function.support, DFunLike.coe] - sorry + sorry + -- apply @tsum_eq_tsum_of_ne_zero_bij + -- case i => + -- simp [Function.support, DFunLike.coe] + -- exact fun x => (vsm_0 x.1, vsm_rest x.1) + -- · simp [Function.Injective] + -- simp [vsm_0, vsm_rest] + -- intro L1 HL1 HL1f L2 HL2 HL2f Heq1 Heq2 + -- cases L1 + -- · simp at HL1 + -- cases L2 + -- · simp at HL2 + -- simp_all + -- · simp [Function.support, Set.range] + -- intro z L HL HF + -- exists (z :: L) + -- simp + -- exists HL + -- sorry + -- · simp [Function.support, DFunLike.coe] + -- sorry rw [ENNReal.tsum_prod'] conv => From 155300c90787db21a3789a543762bda1490699f7 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Wed, 6 Nov 2024 14:02:33 -0500 Subject: [PATCH 074/100] nit --- SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean index ae79fa20..3280b965 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean @@ -179,8 +179,6 @@ lemma exactDiffSum_singleton_le_1 : -1 ≤ exactDiffSum point [v] := by · linarith -set_option pp.coercions false - -- Coercions nonsense lemma DS0 (H : Neighbour L1 L2) : ((exactDiffSum 0 L1) : ℝ) - (exactDiffSum 0 L2) ≤ 1 := by cases H From f9a7882a0f797ada01d962762a05d102223460f4 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Wed, 6 Nov 2024 14:36:45 -0500 Subject: [PATCH 075/100] move remaining sensitivity bounds out of privacy proof --- .../Queries/UnboundedMax/Privacy.lean | 203 +++++++++++------- 1 file changed, 127 insertions(+), 76 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean index 3280b965..363b51d0 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean @@ -9,13 +9,26 @@ import SampCert.DifferentialPrivacy.Queries.UnboundedMax.Properties import SampCert.DifferentialPrivacy.Pure.System noncomputable section - +set_option pp.coercions false +-- set_option pp.all true open Classical namespace SLang variable (ε₁ ε₂ : ℕ+) + +def cov_τ_def (v0 : ℤ) (vs : List ℤ) (l₁ l₂ : List ℕ) : ℤ := (sv8_G l₁ [] v0 vs) - (sv8_G l₂ [] v0 vs) + +lemma cov_τ_def_neg (v0 : ℤ) (vs : List ℤ) (l₁ l₂ : List ℕ) : cov_τ_def v0 vs l₁ l₂ = -cov_τ_def v0 vs l₂ l₁ := by + simp [cov_τ_def] + +def cov_vk_def (v0 : ℤ) (vs : List ℤ) (l₁ l₂ : List ℕ) (point : ℕ) : ℤ := exactDiffSum (point + 1) l₂ - exactDiffSum (point + 1) l₁ + cov_τ_def v0 vs l₁ l₂ + +lemma cov_vk_def_neg (v0 : ℤ) (vs : List ℤ) (l₁ l₂ : List ℕ) : cov_vk_def v0 vs l₁ l₂ point = -cov_vk_def v0 vs l₂ l₁ point := by + simp [cov_τ_def, cov_vk_def] + linarith + -- FIXME: Move lemma tsum_shift (Δ : ℤ) (f : ℤ → ENNReal) : (∑'(x : ℤ), f x = ∑'(x : ℤ), f (x + Δ)) := by apply @tsum_eq_tsum_of_ne_zero_bij @@ -233,6 +246,79 @@ lemma DS0 (H : Neighbour L1 L2) : ((exactDiffSum 0 L1) : ℝ) - (exactDiffSum 0 rw [min_eq_right_iff.mpr (by linarith)] linarith + +lemma Hsens_cov_τ (v0 : ℤ) (vs : List ℤ) (l₁ l₂ : List ℕ) (Hneighbour : Neighbour l₁ l₂) : cov_τ_def v0 vs l₁ l₂ ≤ sens_cov_τ := by + dsimp [cov_τ_def] + cases vs + · sorry + rename_i vs Hvs + cases Hneighbour + · rename_i A B n H1 H2 + rw [H1, H2]; clear H1 H2 + + conv => + enter [1, 2, 1] + rw [List.append_assoc] + rw [@sv8_G_comm A ([n] ++ B)] + sorry + + · rename_i A n B H1 H2 + rw [H1, H2]; clear H1 H2 + conv => + enter [1, 1, 1] + rw [List.append_assoc] + rw [@sv8_G_comm A ([n] ++ B)] + sorry + + · rename_i A n1 B n2 H1 H2 + rw [H1, H2]; clear H1 H2 + conv => + enter [1, 1, 1] + rw [List.append_assoc] + conv => + enter [1, 2, 1] + rw [List.append_assoc] + rw [@sv8_G_comm A ([n1] ++ B)] + rw [@sv8_G_comm A ([n2] ++ B)] + sorry + +-- Prove sensitivity bound +lemma Hsens_cov_vk (v0 : ℤ) (vs : List ℤ) (l₁ l₂ : List ℕ) (point : ℕ) (Hneighbour : Neighbour l₁ l₂) : cov_vk_def v0 vs l₁ l₂ point ≤ sens_cov_vk := by + dsimp [cov_vk_def] + have X := Hsens_cov_τ v0 vs l₁ l₂ Hneighbour + cases vs + · sorry + -- simp only [sens_cov_vk] + rename_i vs Hvs + cases Hneighbour + · rename_i _ _ n H1 H2 + simp_all only [H1, H2]; clear H1 H2 + repeat rw [exactDiffSum_append] + simp_all [sens_cov_vk, sens_cov_τ] + have _ := @exactDiffSum_singleton_le_1 (point + 1) n + have _ := @exactDiffSum_nonpos (point + 1) [n] + linarith + + · rename_i _ n _ H1 H2 + simp_all only [H1, H2]; clear H1 H2 + repeat rw [exactDiffSum_append] + simp_all [sens_cov_vk, sens_cov_τ] + have _ := @exactDiffSum_singleton_le_1 (point + 1) n + have _ := @exactDiffSum_nonpos (point + 1) [n] + linarith + + · rename_i n1 _ n2 H1 H2 + simp_all only [H1, H2]; clear H1 H2 + repeat rw [exactDiffSum_append] + simp_all [sens_cov_vk, sens_cov_τ] + have _ := @exactDiffSum_singleton_le_1 (point + 1) n1 + have _ := @exactDiffSum_nonpos (point + 1) [n1] + have _ := @exactDiffSum_singleton_le_1 (point + 1) n2 + have _ := @exactDiffSum_nonpos (point + 1) [n2] + linarith + + + lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : PureDPSystem.prop (@sv9_privMax_pmf PureDPSystem laplace_pureDPSystem ε₁ ε₂) ε := by -- Unfold DP definitions @@ -401,8 +487,8 @@ lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : simp [sv8_cond, sv8_sum] -- Perform the changes of variables, so that the sums are pointwise le - let cov_τ : ℤ := (sv8_G l₁ [] v0 vs) - (sv8_G l₂ [] v0 vs) - let cov_vk : ℤ := exactDiffSum (point + 1) l₂ - exactDiffSum (point + 1) l₁ + cov_τ + let cov_τ : ℤ := cov_τ_def v0 vs l₁ l₂ + let cov_vk : ℤ := cov_vk_def v0 vs l₁ l₂ point conv => lhs @@ -425,80 +511,13 @@ lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : apply And.intro · -- cov_τ apply propext - dsimp [cov_τ] + dsimp [cov_τ, cov_τ_def] apply Iff.intro <;> intro _ <;> linarith · -- cov_vk apply propext - dsimp [cov_vk] + dsimp [cov_vk, cov_vk_def] apply Iff.intro <;> intro _ <;> linarith - -- Prove sensitivity bound - have Hsens_cov_τ : cov_τ ≤ sens_cov_τ := by - dsimp [cov_τ] - cases vs - rename_i vs Hvs - simp only [] - cases Hneighbour - · rename_i A B n H1 H2 - rw [H1, H2]; clear H1 H2 - conv => - enter [1, 2, 1] - rw [List.append_assoc] - rw [@sv8_G_comm A ([n] ++ B)] - sorry - - · rename_i A n B H1 H2 - rw [H1, H2]; clear H1 H2 - conv => - enter [1, 1, 1] - rw [List.append_assoc] - rw [@sv8_G_comm A ([n] ++ B)] - sorry - - · rename_i A n1 B n2 H1 H2 - rw [H1, H2]; clear H1 H2 - conv => - enter [1, 1, 1] - rw [List.append_assoc] - conv => - enter [1, 2, 1] - rw [List.append_assoc] - rw [@sv8_G_comm A ([n1] ++ B)] - rw [@sv8_G_comm A ([n2] ++ B)] - sorry - - -- Prove sensitivity bound - have Hsens_cov_vk : cov_vk ≤ sens_cov_vk := by - dsimp [cov_vk] - cases vs - rename_i vs Hvs - cases Hneighbour - · rename_i _ _ n H1 H2 - rw [H1, H2]; clear H1 H2 - repeat rw [exactDiffSum_append] - simp_all [sens_cov_vk, sens_cov_τ] - have _ := @exactDiffSum_singleton_le_1 (point + 1) n - have _ := @exactDiffSum_nonpos (point + 1) [n] - linarith - - · rename_i _ n _ H1 H2 - rw [H1, H2]; clear H1 H2 - repeat rw [exactDiffSum_append] - simp_all [sens_cov_vk, sens_cov_τ] - have _ := @exactDiffSum_singleton_le_1 (point + 1) n - have _ := @exactDiffSum_nonpos (point + 1) [n] - linarith - - · rename_i n1 _ n2 H1 H2 - rw [H1, H2]; clear H1 H2 - repeat rw [exactDiffSum_append] - simp_all [sens_cov_vk, sens_cov_τ] - have _ := @exactDiffSum_singleton_le_1 (point + 1) n1 - have _ := @exactDiffSum_nonpos (point + 1) [n1] - have _ := @exactDiffSum_singleton_le_1 (point + 1) n2 - have _ := @exactDiffSum_nonpos (point + 1) [n2] - linarith - simp [privNoiseThresh, privNoiseGuess, privNoiseZero, DPNoise.noise, privNoisedQueryPure] -- Apply the Laplace inequalities @@ -533,13 +552,45 @@ lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : rw [← Real.exp_add] apply Real.exp_monotone apply @le_trans _ _ _ ((|sens_cov_τ| : ℝ) / (↑↑(2 * sens_cov_τ) * ↑↑ε₂ / ↑↑ε₁) + (|sens_cov_vk| : ℝ) / (↑↑(2 * sens_cov_vk) * ↑↑ε₂ / ↑↑ε₁)) - · apply add_le_add + · have W : |cov_τ.cast| ≤ (sens_cov_τ.val : ℝ) := by + apply abs_le'.mpr + apply And.intro + · dsimp only [cov_τ] + apply Int.cast_le.mpr + apply Hsens_cov_τ + apply Hneighbour + · dsimp only [cov_τ] + rw [cov_τ_def_neg] + simp + apply Int.cast_le.mpr + apply Hsens_cov_τ + apply Neighbour_symm + apply Hneighbour + + have X : |cov_vk.cast| ≤ (sens_cov_vk.val : ℝ) := by + apply abs_le'.mpr + apply And.intro + · dsimp only [cov_vk] + apply Int.cast_le.mpr + apply Hsens_cov_vk + apply Hneighbour + · dsimp only [cov_vk] + rw [cov_vk_def_neg] + simp + apply Int.cast_le.mpr + apply Hsens_cov_vk + apply Neighbour_symm + apply Hneighbour + + apply add_le_add · simp apply div_le_div_of_nonneg_right - · -- apply abs_le'.mpr - sorry + · apply W · apply mul_nonneg <;> simp - · sorry + · apply div_le_div_of_nonneg_right + · simp + apply X + · apply div_nonneg <;> simp simp [sens_cov_τ, sens_cov_vk] ring_nf rw [InvolutiveInv.inv_inv] From 61f067299876d84f7434a2712963fec755a5c679 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Wed, 6 Nov 2024 14:38:07 -0500 Subject: [PATCH 076/100] close vk sensitivity --- .../DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean | 6 ------ 1 file changed, 6 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean index 363b51d0..374ae02f 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean @@ -286,10 +286,6 @@ lemma Hsens_cov_τ (v0 : ℤ) (vs : List ℤ) (l₁ l₂ : List ℕ) (Hneighbour lemma Hsens_cov_vk (v0 : ℤ) (vs : List ℤ) (l₁ l₂ : List ℕ) (point : ℕ) (Hneighbour : Neighbour l₁ l₂) : cov_vk_def v0 vs l₁ l₂ point ≤ sens_cov_vk := by dsimp [cov_vk_def] have X := Hsens_cov_τ v0 vs l₁ l₂ Hneighbour - cases vs - · sorry - -- simp only [sens_cov_vk] - rename_i vs Hvs cases Hneighbour · rename_i _ _ n H1 H2 simp_all only [H1, H2]; clear H1 H2 @@ -298,7 +294,6 @@ lemma Hsens_cov_vk (v0 : ℤ) (vs : List ℤ) (l₁ l₂ : List ℕ) (point : have _ := @exactDiffSum_singleton_le_1 (point + 1) n have _ := @exactDiffSum_nonpos (point + 1) [n] linarith - · rename_i _ n _ H1 H2 simp_all only [H1, H2]; clear H1 H2 repeat rw [exactDiffSum_append] @@ -306,7 +301,6 @@ lemma Hsens_cov_vk (v0 : ℤ) (vs : List ℤ) (l₁ l₂ : List ℕ) (point : have _ := @exactDiffSum_singleton_le_1 (point + 1) n have _ := @exactDiffSum_nonpos (point + 1) [n] linarith - · rename_i n1 _ n2 H1 H2 simp_all only [H1, H2]; clear H1 H2 repeat rw [exactDiffSum_append] From e125c1a66de2d9c46647979357a72d3247104665 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Wed, 6 Nov 2024 16:00:51 -0500 Subject: [PATCH 077/100] generalize DS0->DSN --- .../Queries/UnboundedMax/Privacy.lean | 127 +++++++++++++++--- 1 file changed, 112 insertions(+), 15 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean index 374ae02f..600d800f 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean @@ -130,7 +130,7 @@ lemma exactDiffSum_append : exactDiffSum i (A ++ B) = exactDiffSum i A + exactDi rw [exactClippedSum_append] linarith -lemma sv8_sum_append : sv8_sum (A ++ B) [] v0 = sv8_sum A [] v0 + sv8_sum B [] v0 - v0 := by +lemma sv8_sum_append : sv8_sum (A ++ B) vp v0 = sv8_sum A vp v0 + sv8_sum B vp v0 - v0 := by simp [sv8_sum] simp [exactDiffSum_append] linarith @@ -193,7 +193,7 @@ lemma exactDiffSum_singleton_le_1 : -1 ≤ exactDiffSum point [v] := by -- Coercions nonsense -lemma DS0 (H : Neighbour L1 L2) : ((exactDiffSum 0 L1) : ℝ) - (exactDiffSum 0 L2) ≤ 1 := by +lemma DSN (N : ℕ) (H : Neighbour L1 L2) : ((exactDiffSum N L1) : ℝ) - (exactDiffSum N L2) ≤ 1 := by cases H · rename_i A B C H1 H2 rw [H1, H2] @@ -203,6 +203,14 @@ lemma DS0 (H : Neighbour L1 L2) : ((exactDiffSum 0 L1) : ℝ) - (exactDiffSum 0 have _ := @exactDiffSum_singleton_le_1 0 C simp [exactDiffSum] simp [exactClippedSum] + cases (Classical.em (C.cast ≤ (N.cast : ℝ))) + · rename_i h + rw [min_eq_left_iff.mpr h] + left + simp + · right + simp_all + linarith · rename_i A B C H1 H2 rw [H1, H2] @@ -211,14 +219,16 @@ lemma DS0 (H : Neighbour L1 L2) : ((exactDiffSum 0 L1) : ℝ) - (exactDiffSum 0 have _ := @exactDiffSum_nonpos 0 C simp [exactDiffSum] simp [exactClippedSum] - cases (Classical.em ((B : ℝ) ≤ 1)) + cases (Classical.em ((B : ℝ) ≤ N + 1)) · rename_i h rw [min_eq_left_iff.mpr h] - linarith + left + simp · rename_i h rw [min_eq_right_iff.mpr ?G1] case G1 => linarith - simp + right + linarith · rename_i A B C D H1 H2 rw [H1, H2] @@ -226,42 +236,129 @@ lemma DS0 (H : Neighbour L1 L2) : ((exactDiffSum 0 L1) : ℝ) - (exactDiffSum 0 simp only [Int.cast_add, add_sub_add_right_eq_sub, add_sub_add_left_eq_sub] simp [exactDiffSum] simp [exactClippedSum] - cases (Classical.em ((B : ℝ) ≤ 1)) - · cases (Classical.em ((D : ℝ) ≤ 1)) + cases (Classical.em ((B : ℝ) ≤ N + 1)) + · cases (Classical.em ((D : ℝ) ≤ N + 1)) · rename_i hb hd rw [min_eq_left_iff.mpr hb] rw [min_eq_left_iff.mpr hd] - linarith + simp + left + cases (Classical.em (N.cast ≤ (D.cast : ℝ))) + · rw [min_eq_right_iff.mpr ?G1] + case G1 => linarith + skip + linarith + · rw [min_eq_left_iff.mpr ?G1] + case G1 => linarith + simp · rename_i hb hd rw [min_eq_left_iff.mpr hb] rw [min_eq_right_iff.mpr (by linarith)] + rw [min_eq_right_iff.mpr (by linarith)] + left linarith - · cases (Classical.em ((D : ℝ) ≤ 1)) + · cases (Classical.em ((D : ℝ) ≤ N)) · rename_i hb hd rw [min_eq_left_iff.mpr hd] + rw [min_eq_left_iff.mpr (by linarith)] rw [min_eq_right_iff.mpr (by linarith)] + simp + right linarith · rename_i hb hd rw [min_eq_right_iff.mpr (by linarith)] + simp only [not_le] at hd + rw [min_eq_right_iff.mpr ?G1] + case G1 => + -- #check Nat.add_one_le_iff + have X : (1 : ℝ) = ((1 : ℕ) : ℝ) := by simp + rw [X] + rw [← Nat.cast_add] + apply Nat.cast_le.mpr + apply Nat.le_of_lt_succ + simp + apply Nat.cast_le.mp + skip rw [min_eq_right_iff.mpr (by linarith)] + right linarith + lemma Hsens_cov_τ (v0 : ℤ) (vs : List ℤ) (l₁ l₂ : List ℕ) (Hneighbour : Neighbour l₁ l₂) : cov_τ_def v0 vs l₁ l₂ ≤ sens_cov_τ := by dsimp [cov_τ_def] - cases vs - · sorry - rename_i vs Hvs cases Hneighbour · rename_i A B n H1 H2 rw [H1, H2]; clear H1 H2 - conv => enter [1, 2, 1] rw [List.append_assoc] rw [@sv8_G_comm A ([n] ++ B)] + conv => + enter [1, 2] + rw [List.append_assoc] + rw [sv8_G_comm] + conv => + enter [1, 1] + rw [sv8_G_comm] + generalize HL : (B ++ A) = L + clear HL + + + + + + + sorry + /- + suffices (∀ H v0, sv8_G (A ++ B) H v0 vs - sv8_G ([n] ++ B ++ A) H v0 vs ≤ sens_cov_τ.val.cast) by + apply this + induction vs + · -- Base case: + intro H v0 + simp only [sv8_G] + conv => + enter [1, 2] + rw [List.append_assoc] + rw [sv8_sum_append] + rw [sv8_sum_comm] + enter [1, 1] + simp only [sv8_sum] + simp + have _ := @exactDiffSum_singleton_le_1 0 n + have _ := @exactDiffSum_nonpos 0 [n] + simp [sens_cov_τ] + sorry + -- linarith + + · -- Inductive case: + intro H v0 + rename_i head tail IH + unfold sv8_G + -- Combine the maximums + apply le_trans (max_sub_max_le_max _ _ _ _) + -- Do both cases separately + apply Int.max_le.mpr + apply And.intro + · conv => + enter [1, 2] + rw [List.append_assoc] + rw [sv8_sum_append] + rw [sv8_sum_comm] + enter [1, 1] + simp only [sv8_sum, List.length] + have _ := @exactDiffSum_singleton_le_1 0 n + have _ := @exactDiffSum_nonpos 0 [n] + simp [sens_cov_τ] + sorry + -- linarith + · apply IH + -/ + + + · rename_i A n B H1 H2 rw [H1, H2]; clear H1 H2 conv => @@ -389,9 +486,9 @@ lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : apply And.intro · apply neg_le.mp simp only [neg_sub] - apply DS0 + apply DSN apply Hneighbour - · apply DS0 + · apply DSN apply Neighbour_symm apply Hneighbour From d4444a4c206316b62afe252d3aad706084a83d19 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Wed, 6 Nov 2024 16:02:07 -0500 Subject: [PATCH 078/100] nit --- .../DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean index 600d800f..011fa727 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean @@ -277,7 +277,9 @@ lemma DSN (N : ℕ) (H : Neighbour L1 L2) : ((exactDiffSum N L1) : ℝ) - (exact apply Nat.cast_le.mpr apply Nat.le_of_lt_succ simp - apply Nat.cast_le.mp + apply Nat.cast_lt.mp at hd + trivial + skip rw [min_eq_right_iff.mpr (by linarith)] right From 91a8f031d96f9edfafcb0adb703b373f6e99b0f5 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Wed, 6 Nov 2024 16:25:25 -0500 Subject: [PATCH 079/100] privacy is sorry-free --- .../Queries/UnboundedMax/Privacy.lean | 121 +++++------------- 1 file changed, 30 insertions(+), 91 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean index 011fa727..7502ac62 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean @@ -246,7 +246,6 @@ lemma DSN (N : ℕ) (H : Neighbour L1 L2) : ((exactDiffSum N L1) : ℝ) - (exact cases (Classical.em (N.cast ≤ (D.cast : ℝ))) · rw [min_eq_right_iff.mpr ?G1] case G1 => linarith - skip linarith · rw [min_eq_left_iff.mpr ?G1] case G1 => linarith @@ -280,106 +279,46 @@ lemma DSN (N : ℕ) (H : Neighbour L1 L2) : ((exactDiffSum N L1) : ℝ) - (exact apply Nat.cast_lt.mp at hd trivial - skip rw [min_eq_right_iff.mpr (by linarith)] right linarith - +lemma Hsens_cov_τ_lemma (HN : Neighbour l₁ l₂) : sv8_sum l₁ H v0 - sv8_sum l₂ H v0 ≤ OfNat.ofNat 1 := by + simp only [sv8_sum] + rw [add_tsub_add_eq_tsub_right] + have X := @DSN l₁ l₂ H.length HN + rw [← Int.cast_sub] at X + have Y : (@OfNat.ofNat.{0} Real 1 (@One.toOfNat1.{0} Real Real.instOne)) = (@OfNat.ofNat.{0} Int (@OfNat.ofNat.{0} Nat 1 (instOfNatNat 1)) (@instOfNat (@OfNat.ofNat.{0} Nat 1 (instOfNatNat 1)))) := + by simp + rw [Y] at X + apply Int.cast_le.mp at X + apply le_trans X + simp lemma Hsens_cov_τ (v0 : ℤ) (vs : List ℤ) (l₁ l₂ : List ℕ) (Hneighbour : Neighbour l₁ l₂) : cov_τ_def v0 vs l₁ l₂ ≤ sens_cov_τ := by - dsimp [cov_τ_def] - cases Hneighbour - · rename_i A B n H1 H2 - rw [H1, H2]; clear H1 H2 - conv => - enter [1, 2, 1] - rw [List.append_assoc] - rw [@sv8_G_comm A ([n] ++ B)] - conv => - enter [1, 2] - rw [List.append_assoc] - rw [sv8_G_comm] - conv => - enter [1, 1] - rw [sv8_G_comm] - generalize HL : (B ++ A) = L - clear HL - - - - - - - - sorry - - /- - suffices (∀ H v0, sv8_G (A ++ B) H v0 vs - sv8_G ([n] ++ B ++ A) H v0 vs ≤ sens_cov_τ.val.cast) by - apply this - induction vs - · -- Base case: - intro H v0 - simp only [sv8_G] - conv => - enter [1, 2] - rw [List.append_assoc] - rw [sv8_sum_append] - rw [sv8_sum_comm] - enter [1, 1] - simp only [sv8_sum] - simp - have _ := @exactDiffSum_singleton_le_1 0 n - have _ := @exactDiffSum_nonpos 0 [n] - simp [sens_cov_τ] - sorry - -- linarith - - · -- Inductive case: - intro H v0 - rename_i head tail IH - unfold sv8_G - -- Combine the maximums - apply le_trans (max_sub_max_le_max _ _ _ _) - -- Do both cases separately - apply Int.max_le.mpr - apply And.intro - · conv => - enter [1, 2] - rw [List.append_assoc] - rw [sv8_sum_append] - rw [sv8_sum_comm] - enter [1, 1] - simp only [sv8_sum, List.length] - have _ := @exactDiffSum_singleton_le_1 0 n - have _ := @exactDiffSum_nonpos 0 [n] - simp [sens_cov_τ] - sorry - -- linarith - · apply IH - -/ + dsimp [cov_τ_def, sens_cov_τ] + suffices (∀ H v0, sv8_G l₁ H v0 vs - sv8_G l₂ H v0 vs ≤ sens_cov_τ.val.cast) by + apply this + induction vs + · intro H v0 + dsimp [sens_cov_τ] + simp only [sv8_G] + apply Hsens_cov_τ_lemma + trivial - · rename_i A n B H1 H2 - rw [H1, H2]; clear H1 H2 - conv => - enter [1, 1, 1] - rw [List.append_assoc] - rw [@sv8_G_comm A ([n] ++ B)] - sorry + · rename_i next fut IH + intro H v0 + simp only [sv8_G] + apply le_trans (max_sub_max_le_max _ _ _ _) + -- Do both cases separately + apply Int.max_le.mpr + apply And.intro + · apply Hsens_cov_τ_lemma + trivial + · apply IH - · rename_i A n1 B n2 H1 H2 - rw [H1, H2]; clear H1 H2 - conv => - enter [1, 1, 1] - rw [List.append_assoc] - conv => - enter [1, 2, 1] - rw [List.append_assoc] - rw [@sv8_G_comm A ([n1] ++ B)] - rw [@sv8_G_comm A ([n2] ++ B)] - sorry -- Prove sensitivity bound lemma Hsens_cov_vk (v0 : ℤ) (vs : List ℤ) (l₁ l₂ : List ℕ) (point : ℕ) (Hneighbour : Neighbour l₁ l₂) : cov_vk_def v0 vs l₁ l₂ point ≤ sens_cov_vk := by From 414ca7bf423d5a36cc1bd70bd3d0d4a2300b43de Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Wed, 6 Nov 2024 16:25:59 -0500 Subject: [PATCH 080/100] minor cleanup --- .../DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean index 7502ac62..13c5fc34 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean @@ -9,8 +9,7 @@ import SampCert.DifferentialPrivacy.Queries.UnboundedMax.Properties import SampCert.DifferentialPrivacy.Pure.System noncomputable section -set_option pp.coercions false --- set_option pp.all true + open Classical namespace SLang @@ -269,7 +268,6 @@ lemma DSN (N : ℕ) (H : Neighbour L1 L2) : ((exactDiffSum N L1) : ℝ) - (exact simp only [not_le] at hd rw [min_eq_right_iff.mpr ?G1] case G1 => - -- #check Nat.add_one_le_iff have X : (1 : ℝ) = ((1 : ℕ) : ℝ) := by simp rw [X] rw [← Nat.cast_add] From a7ff04ffcc853ee2b7538904de06c83b7a029f78 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Wed, 6 Nov 2024 20:55:01 -0500 Subject: [PATCH 081/100] list get last lemma --- .../DifferentialPrivacy/Queries/UnboundedMax/Properties.lean | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 52547715..f22c178e 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -655,7 +655,8 @@ lemma sv4_presample_perm (ε₁ ε₂ : ℕ+) (point : ℕ) (z : ℤ) (p : { l : sorry lemma get_last_lemma (L : List ℤ) H : L.getLastI = L.getLast H := by - sorry + rw [List.getLastI_eq_getLast?] + rw [List.getLast?_eq_getLast_of_ne_nil H] -- Splits and rearranges the functions @@ -1529,6 +1530,7 @@ def sv8_sv9_eq (ε₁ ε₂ : ℕ+) (l : List ℕ) : sv8_privMax ε₁ ε₂ l = sv8_privMax ε₁ ε₂ l := by apply SLang.ext intro point + sorry From 93770216d76d4b80ec83ab4b0fd7958b3b5d8a8f Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Wed, 6 Nov 2024 21:16:37 -0500 Subject: [PATCH 082/100] sv8_eq_sv9 --- .../Queries/UnboundedMax/Properties.lean | 45 ++++++++++++++++++- 1 file changed, 43 insertions(+), 2 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index f22c178e..8a2a1374 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -1527,11 +1527,52 @@ def sv9_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := computation point def sv8_sv9_eq (ε₁ ε₂ : ℕ+) (l : List ℕ) : - sv8_privMax ε₁ ε₂ l = sv8_privMax ε₁ ε₂ l := by + sv8_privMax ε₁ ε₂ l = sv9_privMax ε₁ ε₂ l := by apply SLang.ext intro point + unfold sv8_privMax + unfold sv9_privMax + simp + split + · simp + · simp + conv => + lhs + conv => + enter [1, a] + rw [<- ENNReal.tsum_mul_left] + conv => + enter [1, b] + rw [<- mul_assoc] + conv => + enter [1] + rw [mul_comm] + rw [mul_assoc] + rw [ENNReal.tsum_comm] + conv => + enter [1, b] + rw [ENNReal.tsum_mul_left] - sorry + apply tsum_congr + intro b + congr 1 + + conv => + lhs + conv => + enter [1, a] + rw [<- ENNReal.tsum_mul_left] + conv => + enter [1, b] + rw [<- mul_assoc] + conv => + enter [1] + rw [mul_comm] + rw [mul_assoc] + rw [ENNReal.tsum_comm] + conv => + enter [1, b] + rw [ENNReal.tsum_mul_left] /-- From 572aa680dd4fb2614829ab3ec65cb7c3528f62d2 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Wed, 6 Nov 2024 22:11:24 -0500 Subject: [PATCH 083/100] close presample_norm_lemma --- .../Queries/UnboundedMax/Code.lean | 6 +- .../Queries/UnboundedMax/Properties.lean | 95 ++++++++++++------- 2 files changed, 65 insertions(+), 36 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean index a64f71c4..05225428 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean @@ -49,12 +49,12 @@ Noise the constant 0 value using the abstract noise function. This looks strange, but will specialize to Lap(ε₁/ε₂, 0) in the pure DP case. -/ -def privNoiseZero (ε₁ ε₂ : ℕ+) : SLang ℤ := dpn.noise (fun _ => 0) 1 ε₁ ε₂ [] +def privNoiseZero (ε₁ ε₂ : ℕ+) : SPMF ℤ := dpn.noise (fun _ => 0) 1 ε₁ ε₂ [] -def privNoiseGuess (ε₁ ε₂ : ℕ+) : SLang ℤ := privNoiseZero ε₁ (2 * sens_cov_vk * ε₂) +def privNoiseGuess (ε₁ ε₂ : ℕ+) : SPMF ℤ := privNoiseZero ε₁ (2 * sens_cov_vk * ε₂) -def privNoiseThresh (ε₁ ε₂ : ℕ+) : SLang ℤ := privNoiseZero ε₁ (2 * sens_cov_τ * ε₂) +def privNoiseThresh (ε₁ ε₂ : ℕ+) : SPMF ℤ := privNoiseZero ε₁ (2 * sens_cov_τ * ε₂) /- ## Program version 0 diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 8a2a1374..c9420d01 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -587,6 +587,7 @@ lemma vector_sum_merge n (f : ℤ × { l : List ℤ // l.length = n } -> ENNReal -- Split in the other order, used as a helper function +-- REFACTOR: Get rid of this, use sv4_presample_split'' lemma sv4_presample_split' (ε₁ ε₂ : ℕ+) (point : ℕ) (z : ℤ) (p : { l : List ℤ // List.length l = point }) : privNoiseGuess ε₁ ε₂ z * sv4_presample ε₁ ε₂ point p = sv4_presample ε₁ ε₂ (point + 1) ⟨ (p.1 ++ [z]), by simp ⟩ := by @@ -650,6 +651,11 @@ lemma sv4_presample_split' (ε₁ ε₂ : ℕ+) (point : ℕ) (z : ℤ) (p : { l trivial +lemma sv4_presample_split'' (ε₁ ε₂ : ℕ+) (point : ℕ) (z : ℤ) (p : { l : List ℤ // List.length l = point }) HP : + privNoiseGuess ε₁ ε₂ z * sv4_presample ε₁ ε₂ point p = + sv4_presample ε₁ ε₂ (point + 1) ⟨ (p.1 ++ [z]), HP ⟩ := by rw [sv4_presample_split'] + + lemma sv4_presample_perm (ε₁ ε₂ : ℕ+) (point : ℕ) (z : ℤ) (p : { l : List ℤ // List.length l = point }) H1 H2 : sv4_presample ε₁ ε₂ (point + 1) ⟨p.1 ++ [z], H1⟩ = sv4_presample ε₁ ε₂ (point + 1) ⟨z :: p.1, H2⟩ := by sorry @@ -658,6 +664,11 @@ lemma get_last_lemma (L : List ℤ) H : L.getLastI = L.getLast H := by rw [List.getLastI_eq_getLast?] rw [List.getLast?_eq_getLast_of_ne_nil H] +lemma drop_init_lemma (L : List ℤ) (H : L ≠ []) : L.dropLast ++ [L.getLastI] = L := by + rw [get_last_lemma _ H] + apply List.dropLast_append_getLast H + + -- Splits and rearranges the functions def sv4_presample_split (ε₁ ε₂ : ℕ+) (point : ℕ) : @@ -770,11 +781,7 @@ def sv4_presample_split (ε₁ ε₂ : ℕ+) (point : ℕ) : rw [<- @List.dropLast_append_getLast _ L (by aesop)] congr apply get_last_lemma - congr 1 - · simp_all - congr - symm - trivial + congr <;> symm <;> trivial def len_1_list_to_val (x : { l : List ℤ // l.length = 1 }) : ℤ := @@ -806,29 +813,48 @@ lemma presample_norm_lemma (point : ℕ) (ε₁ ε₂ : ℕ+) : rw [sv4_presample_split'] rw [<- this] symm - -- FIXME probably the wrong bijection - sorry - -- apply @tsum_eq_tsum_of_ne_zero_bij - -- case i => - -- simp [Function.support, DFunLike.coe] - -- exact fun x => (vsm_0 x.1, vsm_rest x.1) - -- · simp [Function.Injective] - -- simp [vsm_0, vsm_rest] - -- intro L1 HL1 HL1f L2 HL2 HL2f Heq1 Heq2 - -- cases L1 - -- · simp at HL1 - -- cases L2 - -- · simp at HL2 - -- simp_all - -- · simp [Function.support, Set.range] - -- intro z L HL HF - -- exists (z :: L) - -- simp - -- exists HL - -- sorry - -- · simp [Function.support, DFunLike.coe] - -- sorry - + rw [vector_sum_merge] + simp only [] + simp [vsm_0, vsm_rest] + symm + apply @tsum_eq_tsum_of_ne_zero_bij + case i => + simp [Function.support, DFunLike.coe] + exact fun x => ⟨ ↑(vsm_rest x.1) ++ [vsm_0 x.1], by simp ⟩ + · simp [Function.Injective] + simp [vsm_0, vsm_rest] + intro L1 HL1 HL1f L2 HL2 HL2f + cases L1 + · simp at HL1 + cases L2 + · simp at HL2 + simp_all + intro X + rename_i A B C D + have X1 : List.reverse (B ++ [A]) = List.reverse (D ++ [C]) := by exact congrArg List.reverse X + simp at X1 + apply X1 + · simp [Function.support, Set.range] + intro L1 HL1 Hf1 + exists ((vsm_last ⟨ L1, HL1 ⟩) :: (vsm_init ⟨ L1, HL1 ⟩)) + simp + apply And.intro + · simp [vsm_0, vsm_rest, vsm_init, vsm_last] + intro K + apply Hf1 + rw [<- K] + congr + symm + apply drop_init_lemma + intro K + simp [K] at HL1 + · simp [vsm_0, vsm_rest, vsm_init, vsm_last] + apply drop_init_lemma + intro K + simp [K] at HL1 + · simp [Function.support, DFunLike.coe] + intros + congr rw [ENNReal.tsum_prod'] conv => enter [1, 1, a] @@ -838,7 +864,8 @@ lemma presample_norm_lemma (point : ℕ) (ε₁ ε₂ : ℕ+) : simp -- Change noise to SPMF - sorry + have S := (privNoiseGuess ε₁ ε₂).2 + apply HasSum.tsum_eq S def sv3_sv4_loop_eq (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) (init : sv1_state) : @@ -890,17 +917,18 @@ def sv3_sv4_loop_eq (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) -- Apply the IH on the left -- Doing it this way because I can't conv under the @ite? + let ApplyIH : - (do + ((do let vk1 ← privNoiseGuess ε₁ ε₂ if sv1_privMaxC τ l init = true then sv3_loop ε₁ ε₂ τ l point (init.1 ++ [init.2], vk1) - else probPure init) = - (do + else (SLang.probPure init) : SLang _) = + ((do let vk1 ← privNoiseGuess ε₁ ε₂ if sv1_privMaxC τ l init = true then sv4_loop ε₁ ε₂ τ l point (init.1 ++ [init.2], vk1) - else probPure init) := by + else probPure init) : SLang _)) := by simp apply SLang.ext intro final_state @@ -911,6 +939,7 @@ def sv3_sv4_loop_eq (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (point : ℕ) split · rw [IH] · rfl + rw [ApplyIH] clear ApplyIH IH From 2906f3ea5a578cb4f9d1805dd9e15ad331a3f3f5 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Wed, 6 Nov 2024 22:30:29 -0500 Subject: [PATCH 084/100] close sv6 sv7 eq --- .../Queries/UnboundedMax/Properties.lean | 112 ++++++++++-------- 1 file changed, 63 insertions(+), 49 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index c9420d01..d567fb33 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -42,52 +42,6 @@ lemma probBind_congr_strong (p : SLang T) (f : T -> SLang U) (g : T -> SLang U) rw [Hcong] apply Hp - -/- -Not used for anything, but to give confidence in our definitions - -(exactDiffSum l m) is zero if and only if m is an upper bound on the list elements --/ --- lemma exactDiffSum_eq_0_iff_ge_max (l : List ℕ) (m : ℕ) : --- l.maximum ≤ m <-> exactDiffSum m l ≤ 0 := by --- apply Iff.intro --- · induction l --- · simp [exactDiffSum, exactClippedSum] --- · rename_i l0 ls IH --- intro Hmax --- simp [List.maximum_cons] at Hmax --- rcases Hmax with ⟨ Hmax0, Hmax1 ⟩ --- have IH' := IH Hmax1 --- clear IH --- simp [exactDiffSum, exactClippedSum] at * --- apply Int.add_le_of_le_neg_add --- apply le_trans IH' --- simp --- · intro H1 --- apply List.maximum_le_of_forall_le --- revert H1 --- induction l --- · simp --- · rename_i l0 ls IH --- intro Hdiff a Ha --- rw [List.mem_cons_eq] at Ha --- cases Ha --- · rename_i H --- rw [H] --- rw [Nat.cast_withBot] --- apply WithBot.coe_le_coe.mpr --- --- sorry --- · apply IH; clear IH --- · simp only [exactDiffSum, exactClippedSum] at * --- have H : (min (l0.cast : ℤ) (m.cast : ℤ) - min (l0.cast) ((m.cast : ℤ) + 1)) = 0 := by --- sorry --- -- rw [H] at Hdiff --- -- rw [<- Hdiff] --- -- simp --- sorry --- · trivial - /-- History-aware progam computes the same as the history-agnostic program -/ @@ -1439,10 +1393,70 @@ def sv6_sv7_eq (ε₁ ε₂ : ℕ+) (l : List ℕ) : lhs unfold sv6_loop simp + rw [ENNReal.tsum_comm] rw [← ENNReal.tsum_prod] - -- There is a bijection here - sorry - + conv => + rhs + enter [1, a] + rw [<- mul_assoc] + enter [1] + rw [mul_comm] + rw [sv4_presample_split'] + apply @tsum_eq_tsum_of_ne_zero_bij + case i => + exact fun x => ⟨ x.1.2 ++ [x.1.1], by simp ⟩ + · simp [Function.Injective] + intros + rename_i H + apply congrArg List.reverse at H + simp at H + apply H + · simp [Function.support, Set.range] + intros L HL Hf1 Hf2 + exists (vsm_last ⟨ L, HL ⟩) + exists (vsm_init ⟨ L, HL ⟩) + simp + apply And.intro + · apply And.intro + · intro K + apply Hf1 + rw [<- K] + congr + simp [vsm_init, vsm_last] + symm + apply drop_init_lemma + intro K + simp [K] at HL + · intro K + apply Hf2 + rw [<- K] + split <;> split + · rfl + · exfalso + rename_i h1 h2 + apply h2 + rw [<- h1] + congr + simp [vsm_init, vsm_last] + apply drop_init_lemma + intro K + simp [K] at HL + · exfalso + rename_i h2 h1 + apply h2 + rw [<- h1] + congr + simp [vsm_init, vsm_last] + symm + apply drop_init_lemma + intro K + simp [K] at HL + · rfl + · simp [vsm_init, vsm_last] + apply drop_init_lemma + intro K + simp [K] at HL + · simp /- ## Program v8 From 4a5e2c7655da76fe1013c49aa077be36c43921e0 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Thu, 7 Nov 2024 09:41:34 -0500 Subject: [PATCH 085/100] exactDiffSum eventually constant --- .../Queries/UnboundedMax/Properties.lean | 104 ++++++++++++++++++ 1 file changed, 104 insertions(+) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index d567fb33..192a2591 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -42,6 +42,105 @@ lemma probBind_congr_strong (p : SLang T) (f : T -> SLang U) (g : T -> SLang U) rw [Hcong] apply Hp + +-- def sv0_privMaxC (τ : ℤ) (l : List ℕ) (s : sv0_state) : Bool := +-- decide (exactDiffSum (sv0_threshold s) l + (sv0_noise s) < τ) + +-- sv0_privMaxC is eventually just a (geometric) check + +lemma exactDiffSum_eventually_constant : ∃ K, ∀ K', K ≤ K' -> exactDiffSum K' l = 0 := by + cases l + · simp [exactDiffSum, exactClippedSum] + · rename_i l0 ll + let K := (@List.maximum_of_length_pos _ _ (l0 :: ll) (by simp)) + exists K + intro K' HK' + simp [exactDiffSum, exactClippedSum] + rw [min_eq_left_iff.mpr ?G1] + case G1 => + simp + apply le_trans _ HK' + apply List.le_maximum_of_length_pos_of_mem + apply List.mem_cons_self + rw [min_eq_left_iff.mpr ?G1] + case G1 => + apply (@le_trans _ _ _ (↑K') _ _ (by simp)) + simp + apply le_trans _ HK' + apply List.le_maximum_of_length_pos_of_mem + apply List.mem_cons_self + conv => + enter [1, 1, 2, 1] + rw [(@List.map_inj_left _ _ _ _ (fun n => ↑n)).mpr + (by + intro a Ha + rw [min_eq_left_iff.mpr _] + simp + apply le_trans _ HK' + apply List.le_maximum_of_length_pos_of_mem + apply List.mem_cons_of_mem + apply Ha)] + rfl + conv => + enter [1, 2, 2, 1] + rw [(@List.map_inj_left _ _ _ _ (fun n => ↑n)).mpr + (by + intro a Ha + rw [min_eq_left_iff.mpr _] + apply (@le_trans _ _ _ (↑K') _ _ (by simp)) + simp + apply le_trans _ HK' + apply List.le_maximum_of_length_pos_of_mem + apply List.mem_cons_of_mem + apply Ha)] + rfl + simp + + +lemma sv0_norm_loop_le ε₁ ε₂ τ : ∑'v0, ∑'n, probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (n, v0) ≤ 1 := by + sorry + +lemma probWhile_unroll (C : T -> Bool) (F : T -> SLang T) (I : T) : + probWhile C F I = + (if (C I) then probPure I else (F I) >>= probWhile C F) := by + sorry + -- (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (n + Δ, v0) + +/- +lemma sv0_norm_loop_ge ε₁ ε₂ v0 τ : 1 ≤ ∑'vf, ∑'n, probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (0, v0) (n, vf) := by + -- To establish a lower bound, it suffices to shift the starting point + have Hshift : ∀ Δ v0, + ∑' (n : ℕ), probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (n, v0) ≤ + ∑' (n : ℕ), probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (0, v0) := by + intro Δ + induction Δ + · simp + · rename_i Δ' IH + intro v0 + apply le_trans _ (IH _) + clear IH + conv => + rhs + enter [1, n] + rw [probWhile_unroll] + apply ENNReal.tsum_le_tsum + intro n + split + · sorry + · sorry + + + sorry + sorry +-/ + +lemma sv0_norm ε₁ ε₂ l : ∑'(x : ℕ), sv0_privMax ε₁ ε₂ l x = 1 := by + -- unfold sv0_privMax + sorry + + + + /-- History-aware progam computes the same as the history-agnostic program -/ @@ -1618,6 +1717,11 @@ def sv8_sv9_eq (ε₁ ε₂ : ℕ+) (l : List ℕ) : rw [ENNReal.tsum_mul_left] + + + + + /-- sv9 normalizes because sv1 normalizes -/ From 9cf0cbb2792d30c64a489eb96ecebb1b87b72fcf Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Thu, 7 Nov 2024 13:26:15 -0500 Subject: [PATCH 086/100] checkpoint --- .../Queries/UnboundedMax/Properties.lean | 190 ++++++++++++++---- 1 file changed, 147 insertions(+), 43 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 192a2591..697b8162 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -43,11 +43,6 @@ lemma probBind_congr_strong (p : SLang T) (f : T -> SLang U) (g : T -> SLang U) apply Hp --- def sv0_privMaxC (τ : ℤ) (l : List ℕ) (s : sv0_state) : Bool := --- decide (exactDiffSum (sv0_threshold s) l + (sv0_noise s) < τ) - --- sv0_privMaxC is eventually just a (geometric) check - lemma exactDiffSum_eventually_constant : ∃ K, ∀ K', K ≤ K' -> exactDiffSum K' l = 0 := by cases l · simp [exactDiffSum, exactClippedSum] @@ -96,52 +91,89 @@ lemma exactDiffSum_eventually_constant : ∃ K, ∀ K', K ≤ K' -> exactDiffSum rfl simp - -lemma sv0_norm_loop_le ε₁ ε₂ τ : ∑'v0, ∑'n, probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (n, v0) ≤ 1 := by - sorry - lemma probWhile_unroll (C : T -> Bool) (F : T -> SLang T) (I : T) : - probWhile C F I = + probWhile C F I = (if (C I) then probPure I else (F I) >>= probWhile C F) := by sorry - -- (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (n + Δ, v0) -/- -lemma sv0_norm_loop_ge ε₁ ε₂ v0 τ : 1 ≤ ∑'vf, ∑'n, probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (0, v0) (n, vf) := by - -- To establish a lower bound, it suffices to shift the starting point - have Hshift : ∀ Δ v0, - ∑' (n : ℕ), probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (n, v0) ≤ - ∑' (n : ℕ), probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (0, v0) := by - intro Δ - induction Δ - · simp - · rename_i Δ' IH - intro v0 - apply le_trans _ (IH _) - clear IH - conv => - rhs - enter [1, n] - rw [probWhile_unroll] - apply ENNReal.tsum_le_tsum - intro n - split - · sorry - · sorry +lemma probWhile_mass_unroll_lb (C : T -> Bool) (F : T -> SLang T) (I : T) : + ∑'t, (((F I) >>= probWhile C F) t) ≤ ∑'t, probWhile C F I t := by + conv => + rhs + enter [1, t] + rw [probWhile_unroll] + cases (C I) + · simp only [Bool.false_eq_true, ↓reduceIte] + rfl + · conv => + enter [2] + simp + -- This now uses the upper boudn proof (which we need SPMF for but that's OK) + have A : tsum (probPure I) = 1 := by sorry + sorry - sorry - sorry +lemma probWhile_mass_unroll_lb_eval {T U : Type} (C : T -> Bool) (F : T -> SLang T) (g : U -> T): + ∑'t, ∑'(u : U), (((F (g u)) >>= probWhile C F) t) ≤ ∑'t, ∑'(u : U), probWhile C F (g u) t := by + conv => + lhs + rw [ENNReal.tsum_comm] + conv => + rhs + rw [ENNReal.tsum_comm] + apply ENNReal.tsum_le_tsum + intro i + apply probWhile_mass_unroll_lb + +-- Apply this lemma K times to the sv0 loop (sum over v0) +/- +lemma sv0_loop_mass_unroll_lb (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (K : ℕ) : + ∑'t, ∑'(v0 : ℤ), (probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (K, v0)) t ≤ + ∑'t, ∑'(v0 : ℤ), (probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (0, v0)) t := by + induction K + · rfl + · rename_i K' IH + apply le_trans _ (probWhile_mass_unroll_lb_eval (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (fun v0 => (0, v0))) + simp [sv0_privMaxF, sv0_threshold] + -- Not sure + sorry -/ -lemma sv0_norm ε₁ ε₂ l : ∑'(x : ℕ), sv0_privMax ε₁ ε₂ l x = 1 := by - -- unfold sv0_privMax - sorry +-- probWhile sv0_privMaxC (sv0_privMaxF ε₁ ε₂ τ l) (False, (0, v0)) +-- Intuition: If there is a lower bound on every loop iterate terminating, then the termination probability is 1. -/-- +-- Fix l, ε₁, ε₂ +-- Can prove that in any state, the loop has probability at least β to terminate +-- How to turn this into a lower bound on the termination probability? + + +/- +def gen_poisson_trial (C : ℕ -> T → Bool) (F : ℕ × T -> SLang T) (I0 : T) : SLang (Bool × (ℕ × T)) := + probWhile + (fun (x : Bool × (ℕ × T)) => x.1) + (fun v => do + let v1 <- F v.2 + return (C v.2.1 v1, (v.2.1 + 1, v1))) + (True, (0, I0)) + + +lemma gen_poisson_trial_lb (C : ℕ -> T → Bool) (F : ℕ × T -> SLang T) (I : T) (β : NNReal) : + (∀ x, β ≤ (F x >>= fun v => return C x.1 v) True) -> + (∀ t, probGeometric (fun b => if b then β else (1 - β)) t.2.1 ≤ gen_poisson_trial C F I t) := by + intro H + intro t + -- rcases t with ⟨ b, ⟨ n, t ⟩ ⟩ + unfold probGeometric + unfold gen_poisson_trial + unfold geoLoopCond + sorry +-/ + + +/- History-aware progam computes the same as the history-agnostic program -/ lemma sv0_eq_sv1 ε₁ ε₂ l : sv0_privMax ε₁ ε₂ l = sv1_privMax ε₁ ε₂ l := by @@ -193,6 +225,8 @@ lemma sv0_eq_sv1 ε₁ ε₂ l : sv0_privMax ε₁ ε₂ l = sv1_privMax ε₁ + + /- ## Program version 2 - Only moves the loop into a non-executable form, ie. explicitly defines the PMF @@ -1716,14 +1750,84 @@ def sv8_sv9_eq (ε₁ ε₂ : ℕ+) (l : List ℕ) : enter [1, b] rw [ENNReal.tsum_mul_left] +/-- +sv9 normalizes because sv1 normalizes +-/ +def sv9_privMax_pmf (ε₁ ε₂ : ℕ+) (l : List ℕ) : PMF ℕ := + ⟨ sv9_privMax ε₁ ε₂ l, sorry ⟩ +/- +def sv0_eventually_geo_check (ε₁ ε₂ : ℕ+) (τ : ℤ) : SLang Bool := do + let v <- privNoiseGuess ε₁ ε₂ + return v > τ -/-- -sv9 normalizes because sv1 normalizes + +/- +-- The sv0 loop is eventually geometric +lemma sv0_loop_eventually_geometric ε₁ ε₂ τ l : + ∃ K, ∀ s_eval, ∀ v0, + probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (K, v0) s_eval = + probGeometric (sv0_eventually_geo_check ε₁ ε₂ τ) s_eval.1 := by + rcases (@exactDiffSum_eventually_constant l) with ⟨ K, HK ⟩ + exists K + intro s_eval + intro v0 + unfold probGeometric + unfold geoLoopCond + unfold geoLoopBody + unfold sv0_privMaxC + unfold sv0_privMaxF + unfold sv0_eventually_geo_check + -/ + + sorry + + +-- def sv0_privMaxC (τ : ℤ) (l : List ℕ) (s : sv0_state) : Bool := +-- decide (exactDiffSum (sv0_threshold s) l + (sv0_noise s) < τ) + +-- sv0_privMaxC is eventually just a (geometric) check + +lemma sv0_norm_loop_le ε₁ ε₂ τ : ∑'v0, ∑'n, probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (n, v0) ≤ 1 := by + sorry + +lemma probWhile_unroll (C : T -> Bool) (F : T -> SLang T) (I : T) : + probWhile C F I = + (if (C I) then probPure I else (F I) >>= probWhile C F) := by + sorry + -- (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (n + Δ, v0) + +lemma sv0_norm_loop_ge ε₁ ε₂ v0 τ : 1 ≤ ∑'vf, ∑'n, probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (0, v0) (n, vf) := by + -- To establish a lower bound, it suffices to shift the starting point + have Hshift : ∀ Δ v0, + ∑' (n : ℕ), probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (n, v0) ≤ + ∑' (n : ℕ), probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (0, v0) := by + intro Δ + induction Δ + · simp + · rename_i Δ' IH + intro v0 + apply le_trans _ (IH _) + clear IH + conv => + rhs + enter [1, n] + rw [probWhile_unroll] + apply ENNReal.tsum_le_tsum + intro n + split + · sorry + · sorry + + + sorry + sorry + +lemma sv0_norm ε₁ ε₂ l : ∑'(x : ℕ), sv0_privMax ε₁ ε₂ l x = 1 := by + -- unfold sv0_privMax + sorry -/ -def sv9_privMax_pmf (ε₁ ε₂ : ℕ+) (l : List ℕ) : PMF ℕ := - ⟨ sv9_privMax ε₁ ε₂ l, sorry ⟩ From b845bee7ddccec128cc912ad46fed11515c52437 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Thu, 7 Nov 2024 14:29:04 -0500 Subject: [PATCH 087/100] checkpoint before deleting sv4_presample_perm --- .../Queries/UnboundedMax/Properties.lean | 81 +++++++++++++++++-- 1 file changed, 76 insertions(+), 5 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 697b8162..4eeb860a 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -219,7 +219,8 @@ lemma sv0_eq_sv1 ε₁ ε₂ l : sv0_privMax ε₁ ε₂ l = sv1_privMax ε₁ -- RHS: sum over all lists of length "result"? -- rw [tsum_ite_eq] - -- simp [sv1_threshold] + simp [sv1_threshold] + sorry @@ -743,10 +744,6 @@ lemma sv4_presample_split'' (ε₁ ε₂ : ℕ+) (point : ℕ) (z : ℤ) (p : { sv4_presample ε₁ ε₂ (point + 1) ⟨ (p.1 ++ [z]), HP ⟩ := by rw [sv4_presample_split'] -lemma sv4_presample_perm (ε₁ ε₂ : ℕ+) (point : ℕ) (z : ℤ) (p : { l : List ℤ // List.length l = point }) H1 H2 : - sv4_presample ε₁ ε₂ (point + 1) ⟨p.1 ++ [z], H1⟩ = sv4_presample ε₁ ε₂ (point + 1) ⟨z :: p.1, H2⟩ := by - sorry - lemma get_last_lemma (L : List ℤ) H : L.getLastI = L.getLast H := by rw [List.getLastI_eq_getLast?] rw [List.getLast?_eq_getLast_of_ne_nil H] @@ -756,6 +753,80 @@ lemma drop_init_lemma (L : List ℤ) (H : L ≠ []) : L.dropLast ++ [L.getLastI] apply List.dropLast_append_getLast H +lemma sv4_presample_perm (ε₁ ε₂ : ℕ+) (point : ℕ) (z : ℤ) (p : { l : List ℤ // List.length l = point }) H1 H2 : + sv4_presample ε₁ ε₂ (point + 1) ⟨p.1 ++ [z], H1⟩ = sv4_presample ε₁ ε₂ (point + 1) ⟨z :: p.1, H2⟩ := by + simp [sv4_presample] + conv => + lhs + enter [1, a] + rw [<- ENNReal.tsum_mul_left] + enter [1, i] + rw [mul_ite] + simp + conv => + rhs + enter [1, a] + rw [<- ENNReal.tsum_mul_left] + enter [1, i] + rw [mul_ite] + simp + rw [<- ENNReal.tsum_prod] + rw [<- ENNReal.tsum_prod] + rw [vector_sum_merge] + rw [vector_sum_merge] + -- Is this even right? + sorry + + + /- + apply @tsum_eq_tsum_of_ne_zero_bij + case i => exact fun x => ⟨ vsm_last x.1 :: vsm_init x.1, by simp ⟩ + · simp [Function.Injective, vsm_rest, vsm_init, vsm_0, vsm_last] + intros + rename_i A B C D E F G H I J K L + rw [H] at C + apply (congrArg List.reverse) at C + simp at C + have HA : A = A.headI :: A.tail := by sorry + have HF : F = F.headI :: F.tail := by sorry + rw [HA, HF] + cases C + rename_i C1 C2 + rw [C1, C2] + · simp [Function.support, Set.range] + intro A B C D E + -- exists (?G1 :: ?G2) + + sorry + · sorry + -/ + + + -- rw [<- sv4_presample_split''] + -- conv => + -- rhs + -- simp [sv4_presample] + -- enter [1, a] + -- rw [<- ENNReal.tsum_mul_left] + -- enter [1, i] + -- rw [mul_ite] + -- rw [<- ENNReal.tsum_prod] + -- -- Bijection? Or equality? + -- let L : {l : List ℤ // l.length = point + 1} := ⟨z :: p, by simp ⟩ + -- rw [ENNReal.tsum_eq_add_tsum_ite (vsm_last L, vsm_init L)] + -- conv => + -- lhs + -- rw [<- (add_zero (_ * _))] + -- congr 1 + -- · simp [vsm_init] + -- + -- sorry + -- · sorry + + + + + -- Splits and rearranges the functions def sv4_presample_split (ε₁ ε₂ : ℕ+) (point : ℕ) : From 0a533de349766a3d7c73589c66bf73ceab1effa8 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Thu, 7 Nov 2024 16:41:13 -0500 Subject: [PATCH 088/100] sv4_presample_eval --- .../Queries/UnboundedMax/Properties.lean | 264 +++++++++--------- 1 file changed, 138 insertions(+), 126 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 4eeb860a..4b095c66 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -649,6 +649,75 @@ def vsm_init (x : {l : List ℤ // l.length = n + 1}) : {l : List ℤ // l.lengt ⟨ List.dropLast x.1, by simp ⟩ +lemma sv4_presample_eval (ε₁ ε₂ : ℕ+) (n : ℕ) (s : { l : List ℤ // List.length l = n }) : + sv4_presample ε₁ ε₂ n ⟨ List.reverse s, by simp ⟩ = List.foldl (fun acc b => acc * (privNoiseGuess ε₁ ε₂ b)) 1 s.1 := by + rcases s with ⟨ s, Hs ⟩ + simp + revert n + induction s + · intro n Hs + simp_all + simp at Hs + unfold sv4_presample + split + · simp + · exfalso + simp at Hs + · intro n Hs + rename_i s0 ss IH + simp at Hs + simp only [List.reverse_cons, List.foldl_cons] + unfold sv4_presample + cases n + · exfalso + simp at Hs + rename_i n' + generalize HF : (fun acc b => acc * (privNoiseGuess ε₁ ε₂) b) = F + simp + conv => + enter [1, 1, a] + rw [<- ENNReal.tsum_mul_left] + enter [1, i] + simp + rw [← ENNReal.tsum_prod] + rw [ENNReal.tsum_eq_add_tsum_ite (s0, ⟨ ss.reverse, by simp; linarith ⟩)] + conv => + rhs + rw [<- add_zero (List.foldl _ _ _ )] + rw [add_comm] + conv => + lhs + rw [add_comm] + congr 1 + · simp + intro _ _ _ _ E + exfalso + simp_all + apply congrArg List.reverse at E + simp at E + cases E + simp_all + simp + rw [IH _ ?G1] + case G1 => linarith + rw [HF] + suffices (F (List.foldl F 1 ss) s0 = List.foldl F (F 1 s0) ss) by + rw [<- HF] at this + simp at this + rw [<- HF] + rw [<- this] + rw [mul_comm] + rw [<- (@List.foldl_eq_of_comm' _ _ F ?G1 1 s0 ss)] + case G1 => + rw [<- HF] + intros + simp_all + repeat rw [ mul_assoc] + congr 1 + rw [mul_comm] + simp + + lemma vector_sum_merge n (f : ℤ × { l : List ℤ // l.length = n } -> ENNReal) : @@ -743,7 +812,6 @@ lemma sv4_presample_split'' (ε₁ ε₂ : ℕ+) (point : ℕ) (z : ℤ) (p : { privNoiseGuess ε₁ ε₂ z * sv4_presample ε₁ ε₂ point p = sv4_presample ε₁ ε₂ (point + 1) ⟨ (p.1 ++ [z]), HP ⟩ := by rw [sv4_presample_split'] - lemma get_last_lemma (L : List ℤ) H : L.getLastI = L.getLast H := by rw [List.getLastI_eq_getLast?] rw [List.getLast?_eq_getLast_of_ne_nil H] @@ -752,80 +820,16 @@ lemma drop_init_lemma (L : List ℤ) (H : L ≠ []) : L.dropLast ++ [L.getLastI] rw [get_last_lemma _ H] apply List.dropLast_append_getLast H +lemma list_lemma_1 {L : List ℤ} (H : L ≠ []) : List.headI L :: L.tail = L := by + apply List.cons_head?_tail + apply Option.mem_def.mpr + cases L + · exfalso + simp at H + simp -lemma sv4_presample_perm (ε₁ ε₂ : ℕ+) (point : ℕ) (z : ℤ) (p : { l : List ℤ // List.length l = point }) H1 H2 : - sv4_presample ε₁ ε₂ (point + 1) ⟨p.1 ++ [z], H1⟩ = sv4_presample ε₁ ε₂ (point + 1) ⟨z :: p.1, H2⟩ := by - simp [sv4_presample] - conv => - lhs - enter [1, a] - rw [<- ENNReal.tsum_mul_left] - enter [1, i] - rw [mul_ite] - simp - conv => - rhs - enter [1, a] - rw [<- ENNReal.tsum_mul_left] - enter [1, i] - rw [mul_ite] - simp - rw [<- ENNReal.tsum_prod] - rw [<- ENNReal.tsum_prod] - rw [vector_sum_merge] - rw [vector_sum_merge] - -- Is this even right? - sorry - - - /- - apply @tsum_eq_tsum_of_ne_zero_bij - case i => exact fun x => ⟨ vsm_last x.1 :: vsm_init x.1, by simp ⟩ - · simp [Function.Injective, vsm_rest, vsm_init, vsm_0, vsm_last] - intros - rename_i A B C D E F G H I J K L - rw [H] at C - apply (congrArg List.reverse) at C - simp at C - have HA : A = A.headI :: A.tail := by sorry - have HF : F = F.headI :: F.tail := by sorry - rw [HA, HF] - cases C - rename_i C1 C2 - rw [C1, C2] - · simp [Function.support, Set.range] - intro A B C D E - -- exists (?G1 :: ?G2) - - sorry - · sorry - -/ - - - -- rw [<- sv4_presample_split''] - -- conv => - -- rhs - -- simp [sv4_presample] - -- enter [1, a] - -- rw [<- ENNReal.tsum_mul_left] - -- enter [1, i] - -- rw [mul_ite] - -- rw [<- ENNReal.tsum_prod] - -- -- Bijection? Or equality? - -- let L : {l : List ℤ // l.length = point + 1} := ⟨z :: p, by simp ⟩ - -- rw [ENNReal.tsum_eq_add_tsum_ite (vsm_last L, vsm_init L)] - -- conv => - -- lhs - -- rw [<- (add_zero (_ * _))] - -- congr 1 - -- · simp [vsm_init] - -- - -- sorry - -- · sorry - - - - +lemma list_lemma_2 {L : List ℤ} (H : L ≠ []) : List.dropLast L ++ [L.getLastI] = L := by + rw [drop_init_lemma L H] -- Splits and rearranges the functions @@ -879,67 +883,75 @@ def sv4_presample_split (ε₁ ε₂ : ℕ+) (point : ℕ) : conv => enter [2, 1, p] rw [sv4_presample_split'] - rw [sv4_presample_perm _ _ _ _ _ _ (by simp)] rw [vector_sum_merge] rw [vector_sum_merge] - -- Finally, they are equal by bijection + -- Bots sums are singletons + + + sorry + + + /- + simp [vsm_rest, vsm_0] + + -- There is a bijection here apply @tsum_eq_tsum_of_ne_zero_bij - case i => exact fun y => ⟨ (vsm_last y.1) :: (vsm_init y.1), by simp ⟩ - · simp [vsm_rest, vsm_0, DFunLike.coe] - simp [Function.Injective] - intro L1 HL1 _ _ L2 HL2 _ _ Heq - cases L1 - · simp at HL1 - cases L2 - · simp at HL2 - simp_all [vsm_init] - · simp [Function.support, Set.range] - intro L1 HL1 _ _ - cases L1 - · simp at HL1 - rename_i l ll H1 H2 - simp [vsm_0, vsm_rest] - simp [vsm_0, vsm_rest] at H1 - exists (ll ++ [l]) - have h : ((ll ++ [l]).length = point + 1) := by - simp [List.length] at HL1 - simp - trivial - exists h + case i => exact fun x => ⟨ vsm_last x.1 :: vsm_init x.1 , by simp ⟩ + · simp [Function.Injective] + intro A B C D E F H I J K + simp_all [vsm_rest, vsm_init] + rw [<- @list_lemma_1 E ?G1] + case G1 => intro K; simp [K] at F + rw [<- @list_lemma_1 A ?G1] + case G1 => intro K; simp [K] at B + cases C + symm + congr 1 + · simp [Function.support, Set.range, DFunLike.coe] + intro A B C D + rcases final_state with ⟨ f, Hf ⟩ + exists f + simp + exists Hf apply And.intro · apply And.intro - · rw [H1] - cases ll <;> simp [List.headI, List.tail] - · intro X - apply H2 - rw [<- X] - congr - cases ll <;> simp [vsm_rest, vsm_0, List.headI, List.tail] - · simp [vsm_last, vsm_init] + · rw [list_lemma_1] + intro K + simp_all + · intro K + apply D + sorry + -- rw [<- K] + -- congr + -- cases C + -- · sorry + -- · sorry + · cases C + simp [vsm_last, vsm_rest, vsm_init] + cases A + · simp at B + simp + rename_i A0 A1 rw [List.getLastI_eq_getLast?] rw [List.getLast?_concat] + simp [Function.support] + intro A B ⟨ C, D ⟩ + simp [vsm_init, vsm_last] + + rcases final_state with ⟨ f, Hf ⟩ + simp_all + + split + · congr 2 + sorry + · sorry + -/ + + + + - -- Now the conditionals are equal - simp [Function.support, DFunLike.coe, vsm_rest, vsm_0, vsm_init, vsm_last] - intro L HL Hf - have X : L.headI :: L.tail = L.dropLast ++ [L.getLastI] := by - clear Hf - cases L - · exfalso - simp at HL - rename_i l ll - simp - generalize HLL : l :: ll = L - symm - have HL' : L.length > 0 := by simp_all - clear HLL - conv => - rhs - rw [<- @List.dropLast_append_getLast _ L (by aesop)] - congr - apply get_last_lemma - congr <;> symm <;> trivial def len_1_list_to_val (x : { l : List ℤ // l.length = 1 }) : ℤ := From 8173cd796abd2964dec57b6a633666bf779b74d5 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Thu, 7 Nov 2024 18:26:42 -0500 Subject: [PATCH 089/100] close sv4_presample_split --- .../Queries/UnboundedMax/Properties.lean | 152 +++++++++++------- 1 file changed, 96 insertions(+), 56 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 4b095c66..71cac1ec 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -717,6 +717,11 @@ lemma sv4_presample_eval (ε₁ ε₂ : ℕ+) (n : ℕ) (s : { l : List ℤ // L rw [mul_comm] simp +lemma sv4_presample_eval' (ε₁ ε₂ : ℕ+) (n : ℕ) (s : { l : List ℤ // List.length l = n }) : + sv4_presample ε₁ ε₂ n s = List.foldl (fun acc b => acc * (privNoiseGuess ε₁ ε₂ b)) 1 (List.reverse s) := by + have X := sv4_presample_eval ε₁ ε₂ n ⟨ List.reverse s, by simp ⟩ + simp only [List.reverse_reverse, Subtype.coe_eta] at X + trivial @@ -886,71 +891,106 @@ def sv4_presample_split (ε₁ ε₂ : ℕ+) (point : ℕ) : rw [vector_sum_merge] rw [vector_sum_merge] - -- Bots sums are singletons - - - sorry - - - /- + -- Both sums are singletons simp [vsm_rest, vsm_0] + symm + rw [ENNReal.tsum_eq_add_tsum_ite final_state] + -- rcases final_state with ⟨ f, Hf ⟩ + -- simp + conv => + rhs + rw [<- (zero_add (tsum _))] + conv => + lhs + rw [add_comm] + congr 1 + · simp + intro A B C D + exfalso + rcases final_state with ⟨ f, Hf ⟩ + simp_all + apply C + rw [list_lemma_1 ?G1] + intro K + simp_all - -- There is a bijection here - apply @tsum_eq_tsum_of_ne_zero_bij - case i => exact fun x => ⟨ vsm_last x.1 :: vsm_init x.1 , by simp ⟩ - · simp [Function.Injective] - intro A B C D E F H I J K - simp_all [vsm_rest, vsm_init] - rw [<- @list_lemma_1 E ?G1] - case G1 => intro K; simp [K] at F - rw [<- @list_lemma_1 A ?G1] - case G1 => intro K; simp [K] at B - cases C - symm - congr 1 - · simp [Function.support, Set.range, DFunLike.coe] + rw [ENNReal.tsum_eq_add_tsum_ite ⟨[vsm_last final_state] ++ (vsm_init final_state), by simp ⟩ ] + conv => + lhs + rw [<- (add_zero (@ite _ _ _ _ _))] + rw [add_comm] + conv => + rhs + rw [add_comm] + congr 1 + · symm + simp intro A B C D + exfalso rcases final_state with ⟨ f, Hf ⟩ - exists f + simp_all [vsm_rest, vsm_0, vsm_last, vsm_init] + apply C + rw [List.getLastI_eq_getLast?] + unfold Option.iget simp - exists Hf - apply And.intro - · apply And.intro - · rw [list_lemma_1] - intro K - simp_all - · intro K - apply D - sorry - -- rw [<- K] - -- congr - -- cases C - -- · sorry - -- · sorry - · cases C - simp [vsm_last, vsm_rest, vsm_init] - cases A - · simp at B - simp - rename_i A0 A1 - rw [List.getLastI_eq_getLast?] - rw [List.getLast?_concat] - simp [Function.support] - intro A B ⟨ C, D ⟩ - simp [vsm_init, vsm_last] + rw [list_lemma_1] + intro K + simp_all + -- Apply the closed form to evaluate rcases final_state with ⟨ f, Hf ⟩ - simp_all + simp [vsm_rest, vsm_0, vsm_last, vsm_init] + rw [sv4_presample_eval'] + rw [sv4_presample_eval'] + simp only [] + have Hfoldl_eq : + (List.foldl (fun acc b => acc * (privNoiseGuess ε₁ ε₂) b) 1 (f.tail ++ [f.headI]).reverse = + List.foldl (fun acc b => acc * (privNoiseGuess ε₁ ε₂) b) 1 (f.dropLast ++ [f.getLastI]).reverse):= by + apply List.Perm.foldl_eq + · intro A B C + simp + rw [mul_assoc] + rw [mul_assoc] + congr 1 + rw [mul_comm] + · conv => + lhs + simp + have H1 : (f.headI :: f.tail.reverse).Perm (f.headI :: f.tail) := by + apply List.Perm.cons f.headI + apply List.reverse_perm + trans + · apply H1 + rw [list_lemma_1 ?G2] + case G2 => intro _ ; simp_all + rw [list_lemma_2 ?G2] + case G2 => intro _ ; simp_all + apply List.Perm.symm + apply List.reverse_perm + rw [Hfoldl_eq] + clear Hfoldl_eq + generalize HX : List.foldl (fun acc b => acc * (privNoiseGuess ε₁ ε₂) b) 1 (f.dropLast ++ [f.getLastI]).reverse = X + clear HX + + -- Both of the conditionals are true split - · congr 2 - sorry - · sorry - -/ - - - - + · split + · rfl + · exfalso + rename_i A B + apply B + rw [list_lemma_2] + intro K + simp [K] at Hf + · split + · exfalso + rename_i A B + apply A + rw [list_lemma_1] + intro K + simp [K] at Hf + · rfl From 35ab39e8589d977c9e783c61c8125e30dcaa9777 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Thu, 7 Nov 2024 20:05:29 -0500 Subject: [PATCH 090/100] tweak spmf --- .../Queries/UnboundedMax/Privacy.lean | 4 +- .../Queries/UnboundedMax/Properties.lean | 70 ++++++++++++++----- SampCert/SLang.lean | 53 ++++++++++---- 3 files changed, 92 insertions(+), 35 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean index 13c5fc34..ce10bd90 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean @@ -350,7 +350,7 @@ lemma Hsens_cov_vk (v0 : ℤ) (vs : List ℤ) (l₁ l₂ : List ℕ) (point : lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : - PureDPSystem.prop (@sv9_privMax_pmf PureDPSystem laplace_pureDPSystem ε₁ ε₂) ε := by + PureDPSystem.prop (@sv9_privMax_SPMF PureDPSystem laplace_pureDPSystem ε₁ ε₂) ε := by -- Unfold DP definitions simp [DPSystem.prop] apply singleton_to_event @@ -358,7 +358,7 @@ lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : intro l₁ l₂ Hneighbour point apply ENNReal.div_le_of_le_mul - simp [sv9_privMax_pmf, DFunLike.coe, PMF.instFunLike] + simp [sv9_privMax_SPMF, DFunLike.coe, PMF.instFunLike] cases point · -- point = 0 diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 71cac1ec..588caad2 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -91,10 +91,11 @@ lemma exactDiffSum_eventually_constant : ∃ K, ∀ K', K ≤ K' -> exactDiffSum rfl simp +/- lemma probWhile_unroll (C : T -> Bool) (F : T -> SLang T) (I : T) : probWhile C F I = (if (C I) then probPure I else (F I) >>= probWhile C F) := by - sorry + s orry lemma probWhile_mass_unroll_lb (C : T -> Bool) (F : T -> SLang T) (I : T) : ∑'t, (((F I) >>= probWhile C F) t) ≤ ∑'t, probWhile C F I t := by @@ -109,8 +110,8 @@ lemma probWhile_mass_unroll_lb (C : T -> Bool) (F : T -> SLang T) (I : T) : enter [2] simp -- This now uses the upper boudn proof (which we need SPMF for but that's OK) - have A : tsum (probPure I) = 1 := by sorry - sorry + have A : tsum (probPure I) = 1 := by s orry + s orry lemma probWhile_mass_unroll_lb_eval {T U : Type} (C : T -> Bool) (F : T -> SLang T) (g : U -> T): @@ -136,7 +137,8 @@ lemma sv0_loop_mass_unroll_lb (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (K apply le_trans _ (probWhile_mass_unroll_lb_eval (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (fun v0 => (0, v0))) simp [sv0_privMaxF, sv0_threshold] -- Not sure - sorry + s orry +-/ -/ @@ -169,10 +171,19 @@ lemma gen_poisson_trial_lb (C : ℕ -> T → Bool) (F : ℕ × T -> SLang T) (I unfold probGeometric unfold gen_poisson_trial unfold geoLoopCond - sorry + s orry -/ +lemma sv1_ub ε₁ ε₂ l : ∑'s, sv1_privMax ε₁ ε₂ l s ≤ 1 := by + unfold sv1_privMax + + -- rw [Summable.hasSum_iff ENNReal.summable] + + + sorry + + /- History-aware progam computes the same as the history-agnostic program -/ @@ -242,7 +253,7 @@ def sv2_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := return (sv1_threshold sk) computation point -lemma sv1_eq_sv2 ε₁ ε₂ l : sv1_privMax ε₁ ε₂ l = sv2_privMax ε₁ ε₂ l := by +lemma sv1_sv2_eq ε₁ ε₂ l : sv1_privMax ε₁ ε₂ l = sv2_privMax ε₁ ε₂ l := by apply SLang.ext intro result simp [sv1_privMax, sv2_privMax] @@ -452,7 +463,7 @@ lemma cone_constancy {ε₁ ε₂ : ℕ+} {τ : ℤ} {data : List ℕ} {v0 vk : · trivial -lemma sv2_eq_sv3 ε₁ ε₂ l : sv2_privMax ε₁ ε₂ l = sv3_privMax ε₁ ε₂ l := by +lemma sv2_sv3_eq ε₁ ε₂ l : sv2_privMax ε₁ ε₂ l = sv3_privMax ε₁ ε₂ l := by apply SLang.ext -- Step through equal headers @@ -1283,7 +1294,7 @@ def sv4_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := return (sv1_threshold sk) computation point -def sv3_eq_sv4 (ε₁ ε₂ : ℕ+) (l : List ℕ) : +def sv3_sv4_eq (ε₁ ε₂ : ℕ+) (l : List ℕ) : sv3_privMax ε₁ ε₂ l = sv4_privMax ε₁ ε₂ l := by unfold sv3_privMax unfold sv4_privMax @@ -1312,7 +1323,7 @@ def sv5_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := @sv5_loop τ l point (([], v0), presamples) computation point -def sv4_eq_sv5 (ε₁ ε₂ : ℕ+) (l : List ℕ) : +def sv4_sv5_eq (ε₁ ε₂ : ℕ+) (l : List ℕ) : sv4_privMax ε₁ ε₂ l = sv5_privMax ε₁ ε₂ l := by unfold sv4_privMax unfold sv5_privMax @@ -1873,11 +1884,32 @@ def sv8_sv9_eq (ε₁ ε₂ : ℕ+) (l : List ℕ) : enter [1, b] rw [ENNReal.tsum_mul_left] + + +lemma sv9_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv9_privMax ε₁ ε₂ l s := by + sorry + + + + /-- sv9 normalizes because sv1 normalizes -/ -def sv9_privMax_pmf (ε₁ ε₂ : ℕ+) (l : List ℕ) : PMF ℕ := - ⟨ sv9_privMax ε₁ ε₂ l, sorry ⟩ +def sv9_privMax_SPMF (ε₁ ε₂ : ℕ+) (l : List ℕ) : SPMF ℕ := + ⟨ sv9_privMax ε₁ ε₂ l, + by + rw [Summable.hasSum_iff ENNReal.summable] + apply LE.le.antisymm + · rw [<- sv8_sv9_eq] + rw [<- sv7_sv8_eq] + rw [<- sv6_sv7_eq] + rw [<- sv5_sv6_eq] + rw [<- sv4_sv5_eq] + rw [<- sv3_sv4_eq] + rw [<- sv2_sv3_eq] + rw [<- sv1_sv2_eq] + apply sv1_ub + · apply sv9_lb ⟩ @@ -1907,7 +1939,7 @@ lemma sv0_loop_eventually_geometric ε₁ ε₂ τ l : unfold sv0_eventually_geo_check -/ - sorry + s orry -- def sv0_privMaxC (τ : ℤ) (l : List ℕ) (s : sv0_state) : Bool := @@ -1916,12 +1948,12 @@ lemma sv0_loop_eventually_geometric ε₁ ε₂ τ l : -- sv0_privMaxC is eventually just a (geometric) check lemma sv0_norm_loop_le ε₁ ε₂ τ : ∑'v0, ∑'n, probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (n, v0) ≤ 1 := by - sorry + s orry lemma probWhile_unroll (C : T -> Bool) (F : T -> SLang T) (I : T) : probWhile C F I = (if (C I) then probPure I else (F I) >>= probWhile C F) := by - sorry + s orry -- (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (n + Δ, v0) lemma sv0_norm_loop_ge ε₁ ε₂ v0 τ : 1 ≤ ∑'vf, ∑'n, probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (0, v0) (n, vf) := by @@ -1943,14 +1975,14 @@ lemma sv0_norm_loop_ge ε₁ ε₂ v0 τ : 1 ≤ ∑'vf, ∑'n, probWhile (sv0_p apply ENNReal.tsum_le_tsum intro n split - · sorry - · sorry + · s orry + · s orry - sorry - sorry + s orry + s orry lemma sv0_norm ε₁ ε₂ l : ∑'(x : ℕ), sv0_privMax ε₁ ε₂ l x = 1 := by -- unfold sv0_privMax - sorry + s orry -/ diff --git a/SampCert/SLang.lean b/SampCert/SLang.lean index 85be3cdb..74d3a631 100644 --- a/SampCert/SLang.lean +++ b/SampCert/SLang.lean @@ -157,29 +157,54 @@ instance : Coe (SPMF α) (PMF α) where coe a := ⟨ a.1, a.2 ⟩ +lemma probPure_norm (a : α) : ∑'s, probPure a s = 1 := by + simp [probPure] + + @[simp] def SPMF_pure (a : α) : SPMF α := ⟨ probPure a, by - have H : PMF.pure a = probPure a := by - unfold probPure - simp [DFunLike.coe, PMF.instFunLike, PMF.pure] - rw [<- H] - cases (PMF.pure a) - simp [DFunLike.coe, PMF.instFunLike] - trivial ⟩ + rw [Summable.hasSum_iff ENNReal.summable] + apply probPure_norm ⟩ + + +lemma probBind_norm (p : SLang α) (q : α -> SLang β) : + ∑'s, p s = 1 -> + (∀ a, ∑'s, q a s = 1) -> + ∑'s, probBind p q s = 1 := by + intro H1 H2 + rw [<- Summable.hasSum_iff ENNReal.summable] at H1 + have H2' : ∀ a, HasSum (q a) 1 := by + intro a + have H2 := H2 a + rw [<- Summable.hasSum_iff ENNReal.summable] at H2 + trivial + let p' : PMF α := ⟨ p, H1 ⟩ + let q' : α -> PMF β := (fun a => ⟨q a, H2' a ⟩) + have H : (probBind p (fun x => q x)) = (PMF.bind p' q') := by + unfold probBind + simp [DFunLike.coe, PMF.instFunLike, PMF.bind] + rw [H] + cases (PMF.bind p' q') + simp [DFunLike.coe, PMF.instFunLike] + rw [<- Summable.hasSum_iff ENNReal.summable] + trivial + @[simp] def SPMF_bind (p : SPMF α) (q : α -> SPMF β) : SPMF β := ⟨ probBind p (fun x => q x), by - have H : (probBind p (fun x => q x)) = (PMF.bind p q) := by - unfold probBind - simp [DFunLike.coe, PMF.instFunLike, PMF.bind] - rw [H] - cases (PMF.bind p q) - simp [DFunLike.coe, PMF.instFunLike] - trivial ⟩ + rw [Summable.hasSum_iff ENNReal.summable] + apply probBind_norm + · rw [<- Summable.hasSum_iff ENNReal.summable] + cases p + trivial + · intro a + rw [<- Summable.hasSum_iff ENNReal.summable] + cases q a + trivial ⟩ instance : Monad SPMF where pure := SPMF_pure From 031af6f711f847e862478435fbccab7d145faed5 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Thu, 7 Nov 2024 21:24:47 -0500 Subject: [PATCH 091/100] sv1_ub --- .../Queries/UnboundedMax/Properties.lean | 128 +++++++++++++++++- 1 file changed, 124 insertions(+), 4 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 588caad2..11ae52ab 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -175,13 +175,133 @@ lemma gen_poisson_trial_lb (C : ℕ -> T → Bool) (F : ℕ × T -> SLang T) (I -/ -lemma sv1_ub ε₁ ε₂ l : ∑'s, sv1_privMax ε₁ ε₂ l s ≤ 1 := by - unfold sv1_privMax +lemma ENNReal.tsum_iSup_comm (f : T -> U -> ENNReal) : ∑' x, ⨆ y, f x y = ⨆ y, ∑' x, f x y := by + -- Make it finite + rw [ENNReal.tsum_eq_iSup_sum] + conv => + rhs + enter [1, y] + rw [ENNReal.tsum_eq_iSup_sum] + rw [iSup_comm] + apply iSup_congr + intro s - -- rw [Summable.hasSum_iff ENNReal.summable] + -- Definitely true, may be annoying to prove + sorry - sorry +lemma sv1_ub ε₁ ε₂ l : ∑'s, sv1_privMax ε₁ ε₂ l s ≤ 1 := by + unfold sv1_privMax + unfold sv1_threshold + simp + -- Push the sum over s inwards + conv => + rw [ENNReal.tsum_comm] + enter [1, 1, b] + rw [ENNReal.tsum_mul_left] + enter [2] + rw [ENNReal.tsum_comm] + enter [1, i] + rw [ENNReal.tsum_mul_left] + enter [2] + rw [ENNReal.tsum_comm] + apply + @le_trans _ _ _ + (∑' (a : ℤ), (privNoiseThresh ε₁ ε₂) a * ∑' (a_1 : ℤ), (privNoiseGuess ε₁ ε₂) a_1 * 1) + _ ?G2 ?G1 + case G1 => + apply Eq.le + simp + rw [ENNReal.tsum_mul_right] + have R1 : ∑' (a : ℤ), (privNoiseThresh ε₁ ε₂) a = 1 := by + rw [<- Summable.hasSum_iff ENNReal.summable] + cases (privNoiseThresh ε₁ ε₂) + simp [DFunLike.coe, SPMF.instFunLike] + trivial + have R2 : ∑' (a : ℤ), (privNoiseGuess ε₁ ε₂) a = 1 := by + rw [<- Summable.hasSum_iff ENNReal.summable] + cases (privNoiseGuess ε₁ ε₂) + simp [DFunLike.coe, SPMF.instFunLike] + trivial + simp_all + case G2 => + apply ENNReal.tsum_le_tsum + intro τ + apply ENNReal.mul_left_mono + apply ENNReal.tsum_le_tsum + intro v0 + apply ENNReal.mul_left_mono + + apply + @le_trans _ _ _ + (∑' (b : sv1_state), probWhile (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) ([], v0) b ) + _ ?G3 ?G4 + case G3 => + apply ENNReal.tsum_le_tsum + intro a + rw [tsum_ite_eq] + case G4 => + unfold probWhile + rw [ENNReal.tsum_iSup_comm] + apply iSup_le_iff.mpr + intro cut + suffices (∀ L, ∑' (x : sv1_state), probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) cut (L, v0) x ≤ 1) by + apply this + revert v0 + induction cut + · simp [probWhileCut] + · rename_i cut' IH + intro v0 L + simp [probWhileCut, probWhileFunctional] + split + · simp + simp [sv1_privMaxF] + conv => + enter [1, 1, x] + conv => + enter [1, a] + rw [<- ENNReal.tsum_mul_right] + simp + rw [ENNReal.tsum_comm] + enter [1, b] + + apply + @le_trans _ _ _ + (∑' (x : sv1_state) (b : ℤ), (privNoiseGuess ε₁ ε₂) b * probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) cut' (L ++ [v0], b) x) + _ ?G5 ?G6 + case G5 => + apply ENNReal.tsum_le_tsum + intro a + apply ENNReal.tsum_le_tsum + intro b + unfold sv1_state + rw [ENNReal.tsum_eq_add_tsum_ite (L ++ [v0], b)] + simp + conv => + rhs + rw [<- add_zero (_ * _)] + apply add_le_add + · simp + · simp + intros + aesop + case G6 => + rw [ENNReal.tsum_comm] + conv => + enter [1, 1, b] + rw [ENNReal.tsum_mul_left] + apply @le_trans _ _ _ (∑' (b : ℤ), (privNoiseGuess ε₁ ε₂) b * 1) + · apply ENNReal.tsum_le_tsum + intro a + apply ENNReal.mul_left_mono + apply IH + · simp + apply Eq.le + rw [<- Summable.hasSum_iff ENNReal.summable] + cases (privNoiseGuess ε₁ ε₂) + simp [DFunLike.coe, SPMF.instFunLike] + trivial + · simp /- From 324ab62b8fbf37ab619553a8a2308aeddd1fcba5 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Fri, 8 Nov 2024 08:48:45 -0500 Subject: [PATCH 092/100] close iSup tsum lemma --- .../Queries/UnboundedMax/Properties.lean | 99 +++++++++++++++++-- 1 file changed, 93 insertions(+), 6 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 11ae52ab..34346675 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -175,8 +175,18 @@ lemma gen_poisson_trial_lb (C : ℕ -> T → Bool) (F : ℕ × T -> SLang T) (I -/ -lemma ENNReal.tsum_iSup_comm (f : T -> U -> ENNReal) : ∑' x, ⨆ y, f x y = ⨆ y, ∑' x, f x y := by - -- Make it finite +lemma ENNReal.tsum_iSup_comm (f : T -> U -> ENNReal) : ∑' x, ⨆ y, f x y ≥ ⨆ y, ∑' x, f x y := by + apply LE.le.ge + rw [iSup_le_iff] + intro i + apply ENNReal.tsum_le_tsum + intro a + apply le_iSup + + +lemma iSup_comm_lemma (ε₁ ε₂ : ℕ+) (l : List ℕ) (τ v0 : ℤ): + ∑' b, ⨆ i, probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) i ([], v0) b = + ⨆ i, ∑' b, probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) i ([], v0) b := by rw [ENNReal.tsum_eq_iSup_sum] conv => rhs @@ -185,9 +195,9 @@ lemma ENNReal.tsum_iSup_comm (f : T -> U -> ENNReal) : ∑' x, ⨆ y, f x y = rw [iSup_comm] apply iSup_congr intro s - - -- Definitely true, may be annoying to prove - sorry + apply ENNReal.finsetSum_iSup_of_monotone + intro a + apply probWhileCut_monotonic lemma sv1_ub ε₁ ε₂ l : ∑'s, sv1_privMax ε₁ ε₂ l s ≤ 1 := by @@ -242,7 +252,7 @@ lemma sv1_ub ε₁ ε₂ l : ∑'s, sv1_privMax ε₁ ε₂ l s ≤ 1 := by rw [tsum_ite_eq] case G4 => unfold probWhile - rw [ENNReal.tsum_iSup_comm] + rw [iSup_comm_lemma] apply iSup_le_iff.mpr intro cut suffices (∀ L, ∑' (x : sv1_state), probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) cut (L, v0) x ≤ 1) by @@ -2005,10 +2015,87 @@ def sv8_sv9_eq (ε₁ ε₂ : ℕ+) (l : List ℕ) : rw [ENNReal.tsum_mul_left] +lemma ENNReal.tsum_lb_single (x : T) (f : T -> ENNReal) (l : ENNReal) : + l ≤ f x -> l ≤ ∑' (a : T), f a := by + intro H + apply le_trans H + apply ENNReal.le_tsum + +lemma ENNReal.tsum_lb_subset (P : T -> Prop) (f : T -> ENNReal) (l : ENNReal) : + l ≤ (∑'(a : {t : T // P t}), f a.1) -> l ≤ ∑' (a : T), f a := by + intro H + apply le_trans H + apply ENNReal.tsum_comp_le_tsum_of_injective + simp [Function.Injective] + + + + lemma sv9_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv9_privMax ε₁ ε₂ l s := by + -- Special value which makes this sure to terminate sorry + /- + let β : ℤ := 0 + -- match l with + -- | [] => 0 + -- | _ => List.maximum! + + -- Wrong idea: + -- Pick τ = 0 + -- β = l.maximum + -- sample "not β" repeatedly -> (1 - T)^n + -- sample β once -> T + -- (wrong) => terminates at the last sample, wp ((1 - T)^n ⬝ T) + -- + -- + -- Issue: May terminate earlier. + -- Therefore: Need to pick β such that + -- Total mass of privNoiseGuess below β + τ is at least (1 - T) -- We can use τ to help make this big! + -- Probability that we pick β is T + -- It terminates when we pick β + -- Each time we even (almost) even odds of sampling a positive or negative value + -- We can quantify the probability that we sample a nonnegative value easily: (q = (1 + Prob 0) / 2) + -- If we sample a positive value p, EDS + p > τ | + -- If we sample zero p, EDS + p ≥ τ | terminates + -- If we sample a negative value n, EDS + n < τ | doesn't terminate + + -- Prove that when τ = ?? + -- sv8_G [negatives...] [negative] < τ + -- sv8_G [negatives...] [nonnegative] ≥ τ + + -- Uhh-- no. We can't pick τ. That restricts the total mass to be < 1. + -- Pick β = ??? + -- Let s ∈ ℕ + -- Want to lower bound the probability mass by ((1-β)^(s-1) β) + + -- What's the probability of termination? + + + + + + + + + -- Probability of sampling the special value in noise + -- T needs to be multiplied by privNoiseThresh (for picking τ to be what we want) + let T := (privNoiseThresh ε₁ ε₂ 0) * (privNoiseGuess ε₁ ε₂ β) + + -- Either geo probGeometric will work + suffices ∀ (s : ℕ), Geo T s ≤ sv9_privMax ε₁ ε₂ l s by + s orry + intro s + unfold sv9_privMax + cases s + · simp [Geo] + apply ENNReal.tsum_lb_single 0 + apply ENNReal.mul_left_mono + + s orry + · s orry + -/ From f08a18a8e8a33b29f332548851098ffe355c15dd Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Fri, 8 Nov 2024 13:04:28 -0500 Subject: [PATCH 093/100] sv1_lb reduction to sv1_cdf_lb --- .../Queries/UnboundedMax/Properties.lean | 98 +++++++++++++++++-- 1 file changed, 92 insertions(+), 6 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 34346675..eae703a5 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -2029,6 +2029,93 @@ lemma ENNReal.tsum_lb_subset (P : T -> Prop) (f : T -> ENNReal) (l : ENNReal) : simp [Function.Injective] +lemma sv1_cdf_lb ε₁ ε₂ l (τ : ℤ) (v0 : ℤ): + ∃ β : ℝ, + (0 < β) ∧ + (β ≤ 1) ∧ + ∀ i : ℕ , + (ENNReal.ofReal (β^i) ≤ + (∑' (a : sv1_state), + probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) i ([], v0) a * probPure (sv1_threshold a) i)) := + sorry + + +lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by + simp only [sv1_privMax, bind, pure, bind_apply] + -- Push the sum over s inwards + conv => + rhs + rw [ENNReal.tsum_comm] + enter [1, b] + rw [ENNReal.tsum_mul_left] + enter [2] + rw [ENNReal.tsum_comm] + enter [1, i] + rw [ENNReal.tsum_mul_left] + + -- Reduce to CDF problem + apply @le_trans _ _ _ (∑' (b : ℤ), (privNoiseThresh ε₁ ε₂) b * ∑' (i : ℤ), (privNoiseGuess ε₁ ε₂) i * 1) _ ?G1 + case G1 => + apply Eq.le + rw [ENNReal.tsum_mul_right] + simp + have R1 : ∑' (a : ℤ), (privNoiseThresh ε₁ ε₂) a = 1 := by + rw [<- Summable.hasSum_iff ENNReal.summable] + cases (privNoiseThresh ε₁ ε₂) + simp [DFunLike.coe, SPMF.instFunLike] + trivial + have R2 : ∑' (a : ℤ), (privNoiseGuess ε₁ ε₂) a = 1 := by + rw [<- Summable.hasSum_iff ENNReal.summable] + cases (privNoiseGuess ε₁ ε₂) + simp [DFunLike.coe, SPMF.instFunLike] + trivial + simp_all + apply ENNReal.tsum_le_tsum + intro τ + apply ENNReal.mul_left_mono + apply ENNReal.tsum_le_tsum + intro v0 + apply ENNReal.mul_left_mono + + + -- Turn it into a supremum + conv => + enter [2, 1, i_1, 1, a] + simp only [probWhile] + rw [ENNReal.iSup_mul] + + -- Commute the supremum out to the middle, so we can pick the cut number for each one + apply @le_trans _ _ _ + (∑' (i_1 : ℕ), ⨆ i, ∑'(a : sv1_state), + probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) i ([], v0) a * probPure (sv1_threshold a) i_1) _ _ ?G1 + case G1 => + apply ENNReal.tsum_le_tsum + intro i + apply ENNReal.tsum_iSup_comm + apply @le_trans _ _ _ + (∑' (i : ℕ), ∑'(a : sv1_state), + probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) i ([], v0) a * probPure (sv1_threshold a) i) _ _ ?G1 + case G1 => + apply ENNReal.tsum_le_tsum + intro i + apply le_trans _ ?G3 + case G3 => apply @le_iSup' _ _ _ _ i + rfl + + -- Apply the CDF lower bound + rcases (sv1_cdf_lb ε₁ ε₂ l τ v0) with ⟨ β, ⟨ H1, ⟨ H2, Hlb ⟩ ⟩ ⟩ + apply le_trans _ ?G2 + case G2 => + apply ENNReal.tsum_le_tsum + intro t + apply Hlb + + -- Geometric sum has a closed form + apply Eq.le + symm + sorry + + @@ -2073,12 +2160,6 @@ lemma sv9_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv9_privMax ε₁ ε₂ l s := by -- What's the probability of termination? - - - - - - -- Probability of sampling the special value in noise -- T needs to be multiplied by privNoiseThresh (for picking τ to be what we want) let T := (privNoiseThresh ε₁ ε₂ 0) * (privNoiseGuess ε₁ ε₂ β) @@ -2099,6 +2180,11 @@ lemma sv9_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv9_privMax ε₁ ε₂ l s := by + + + + + /-- sv9 normalizes because sv1 normalizes -/ From b8c3ea6c20b0e3e60ec562b5ac4cfd6bc915b526 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Fri, 8 Nov 2024 17:55:52 -0500 Subject: [PATCH 094/100] new attempt --- .../Queries/UnboundedMax/Properties.lean | 410 +++++++++++++++--- 1 file changed, 339 insertions(+), 71 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index eae703a5..55ef7bc7 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -91,11 +91,103 @@ lemma exactDiffSum_eventually_constant : ∃ K, ∀ K', K ≤ K' -> exactDiffSum rfl simp -/- + + + +def sv1_privMax_alt (ε₁ ε₂ : ℕ+) (l : List ℕ) (H : List ℤ) : SLang ℕ := do + let τ <- privNoiseThresh ε₁ ε₂ + let v0 <- privNoiseGuess ε₁ ε₂ + let sk <- probWhile (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (H, v0) + return (sv1_threshold sk) + +lemma sv1_sv1_alt_base_case (ε₁ ε₂ : ℕ+) (l : List ℕ) : + sv1_privMax ε₁ ε₂ l = sv1_privMax_alt ε₁ ε₂ l [] := by + simp [sv1_privMax, sv1_privMax_alt] + lemma probWhile_unroll (C : T -> Bool) (F : T -> SLang T) (I : T) : probWhile C F I = (if (C I) then probPure I else (F I) >>= probWhile C F) := by - s orry + sorry + +lemma probWhile_mass_1_unroll (C : T -> Bool) (F : T -> SLang T) (I : T) : + (∑'(v), F I v = 1) -> + (∀ I', F I I' > 0 -> ∑'(v), (probWhile C F I' v) = 1) -> + ∑'(v), probWhile C F I v = 1 := by + intro H1 H2 + rw [probWhile_unroll] + cases C I <;> simp + rw [ENNReal.tsum_comm] + conv => + enter [1] + enter [1, b] + rw [ENNReal.tsum_mul_left] + -- doable + rw [<- H1] + apply Equiv.tsum_eq_tsum_of_support ?S + case S => + apply Equiv.Set.ofEq + simp [Function.support] + apply Set.ext + intro x + simp + intro A + apply Classical.by_contradiction + intro B + simp at B + have D : (F I x > 0) := by + apply Classical.by_contradiction + simp + trivial + have H2' := H2 _ D + simp_all + simp + intro A B E _ + rw [H2] + · simp + · apply Classical.by_contradiction + simp + trivial + +lemma pm_mass_1 (ε₁ ε₂ : ℕ+) (τ : ℤ) (K : ℕ) : + (∀ H : List ℤ, K ≤ H.length -> ∑'(v), sv1_privMax_alt ε₁ ε₂ l H v = 1) -> + (∑'(v), sv1_privMax ε₁ ε₂ l v = 1) := by + -- Unroll K times + induction K + · intro A + rw [sv1_sv1_alt_base_case] + apply A + simp + · rename_i K' IH + intro H' + -- Otherwise, + sorry + +def sv1_alt_alt_privMaxC (τ : ℤ) (s : sv1_state) : Bool := + decide (sv1_noise s < τ) + +def sv1_privMax_alt_alt (ε₁ ε₂ : ℕ+) (H : List ℤ) : SLang ℕ := do + let τ <- privNoiseThresh ε₁ ε₂ + let v0 <- privNoiseGuess ε₁ ε₂ + let sk <- probWhile (sv1_alt_alt_privMaxC τ) (sv1_privMaxF ε₁ ε₂) (H, v0) + return (sv1_threshold sk) + +lemma sv1_alt_eq_alt_alt (ε₁ ε₂ : ℕ+) (l : List ℕ) (H : List ℤ) (K : ℕ) (HK : ∀ K', H.length ≤ K' -> exactDiffSum K' l = 0) : + sv1_privMax_alt ε₁ ε₂ l H = sv1_privMax_alt_alt ε₁ ε₂ H := by + -- H is in the regieme where exactDiffSum is constant + apply SLang.ext + intro x + simp [sv1_privMax_alt, sv1_privMax_alt_alt, probBind] + sorry + + +-- Finally, we should be able to prove that sv1_alt_alt is geometric + + + + + + +/- lemma probWhile_mass_unroll_lb (C : T -> Bool) (F : T -> SLang T) (I : T) : ∑'t, (((F I) >>= probWhile C F) t) ≤ ∑'t, probWhile C F I t := by @@ -2028,18 +2120,87 @@ lemma ENNReal.tsum_lb_subset (P : T -> Prop) (f : T -> ENNReal) (l : ENNReal) : apply ENNReal.tsum_comp_le_tsum_of_injective simp [Function.Injective] +def β_geo (β : ENNReal) : SLang ℕ := (probGeometric (fun x => if x then β else 1 - β)) + +def β_geo_recurrence (β : ENNReal) (n : ℕ) (H : n > 0) : β_geo β (n + 1) = β * β_geo β n := by + simp [β_geo, probGeometric_apply] + split + · exfalso + simp_all + · ring_nf + rename_i h + rw [mul_pow_sub_one h] + +def β_geo' (β : ENNReal) : SLang ℕ := + fun N => + match N with + | 0 => 0 + | Nat.succ N' => β_geo β N' + + +/-- lemma sv1_cdf_lb ε₁ ε₂ l (τ : ℤ) (v0 : ℤ): + sv1_privMaxC τ l ([], v0) = true -> ∃ β : ℝ, (0 < β) ∧ - (β ≤ 1) ∧ + (β < 1) ∧ ∀ i : ℕ , - (ENNReal.ofReal (β^i) ≤ + (β_geo β i ≤ (∑' (a : sv1_state), - probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) i ([], v0) a * probPure (sv1_threshold a) i)) := - sorry + probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) i ([], v0) a * probPure (sv1_threshold a) i)) := by + + -- exactDiffSum is negative or zero for all states + -- decide (exactDiffSum (sv1_threshold s) l + (sv1_noise s) < τ) + intro Htrue + have Hβ : ∃ β : ℝ, 0 < β ∧ β < 1 ∧ (∑'(z : ℤ), if z < τ then (privNoiseGuess ε₁ ε₂ z) else 0) ≤ ENNReal.ofReal β := by + -- This is true due to CDF of laplace distribution + sorry + + rcases Hβ with ⟨ β, ⟨ H1, ⟨ H2, Hsum ⟩ ⟩ ⟩ + exists β + apply And.intro H1 + apply And.intro H2 + intro i + + -- Prove this by induction + -- #check (probGeometric (fun x => if x then ENNReal.ofReal β else 0)) + suffices + ∀ init, sv1_privMaxC τ l init = true -> + β_geo β i ≤ ∑' (a : sv1_state), probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) i init a * probPure (sv1_threshold a) i by + apply this + sorry + -- FIXME: Does β_geo work? Does β^i work? + induction i + · simp [β_geo] + · rename_i i IH + intro init + cases i + · + sorry + · intro H1 + rename_i i + conv => + rhs + enter [1, a] + unfold probWhileCut + unfold probWhileFunctional + split + · sorry + · exfalso + aesop + + -- rw [β_geo_recurrence _ _ (by simp)] + -- apply le_trans + -- · apply ENNReal.mul_left_mono + -- · apply IH + -- sorry + + -- β_geo_recurrence +-/ +/- lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by simp only [sv1_privMax, bind, pure, bind_apply] -- Push the sum over s inwards @@ -2103,83 +2264,189 @@ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by rfl -- Apply the CDF lower bound - rcases (sv1_cdf_lb ε₁ ε₂ l τ v0) with ⟨ β, ⟨ H1, ⟨ H2, Hlb ⟩ ⟩ ⟩ - apply le_trans _ ?G2 - case G2 => - apply ENNReal.tsum_le_tsum - intro t - apply Hlb - - -- Geometric sum has a closed form - apply Eq.le - symm sorry + -- rcases (sv1_cdf_lb ε₁ ε₂ l τ v0) with ⟨ β, ⟨ H1, ⟨ H2, Hlb ⟩ ⟩ ⟩ + -- apply le_trans _ ?G2 + -- case G2 => + -- apply ENNReal.tsum_le_tsum + -- intro t + -- apply Hlb + + -- -- Geometric sum has a closed form + -- apply Eq.le + -- symm + -- unfold β_geo + -- apply probGeometric_normalizes <;> simp + -- · sorry + -- · sorry +-/ +lemma sv8_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv8_privMax ε₁ ε₂ l s := by + simp [sv8_privMax] + -- Factor out the choice of thresh (we will work for all τ) + rw [ENNReal.tsum_comm] + conv => + enter [2, 1, b] + rw [ENNReal.tsum_mul_left] + apply @le_trans _ _ _ (∑'(b : ℤ), privNoiseThresh ε₁ ε₂ b * 1) + · simp + apply Eq.le + symm + rcases (privNoiseThresh ε₁ ε₂) + simp [DFunLike.coe] + rw [<- Summable.hasSum_iff ENNReal.summable] + trivial + apply ENNReal.tsum_le_tsum + intro τ + apply ENNReal.mul_left_mono + + -- Pick β based on τ, and exactDiffSum + have β : ENNReal := 0 -- (∑' (i_1 : { x // x < τ - exactDiffSum (i + 1) l }), (privNoiseGuess ε₁ ε₂) ↑i_1) + + -- Lower bound by probGeometric + apply @le_trans _ _ _ (∑'(i : ℕ), β_geo' β i * 1) + · apply Eq.le + symm + simp only [β_geo', β_geo, mul_one] + -- rw [ENNReal.tsum_eq_add_tsum_ite Nat.zero] + -- simp only [Nat] + sorry + -- apply probGeometric_normalizes <;> simp + -- · -- β fact + -- sorry + -- · -- β fact + -- sorry + + -- Now, prove that it is bounded below by probGeometric + apply ENNReal.tsum_le_tsum + intro i + simp + cases i + · simp [β_geo'] + rename_i i + simp [β_geo'] + cases i + · simp [β_geo] + rename_i i + simp [β_geo] + + -- Commute all sums out to the left + conv => + enter [2, 1, v0] + rw [<- ENNReal.tsum_mul_left] + -- enter [1, vs] + -- rw [<- ENNReal.tsum_mul_left] + -- rw [<- ENNReal.tsum_mul_left] + rw [← ENNReal.tsum_prod] + + -- Eval the presample + conv => + enter [2, 1, a, 2, 1] + rw [sv4_presample_eval'] + simp + + unfold sv8_cond + simp [β_geo] -lemma sv9_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv9_privMax ε₁ ε₂ l s := by - -- Special value which makes this sure to terminate - sorry + + + + + + + -- rw [← ENNReal.tsum_prod] + + -- simp [β_geo] /- - let β : ℤ := 0 - -- match l with - -- | [] => 0 - -- | _ => List.maximum! - - -- Wrong idea: - -- Pick τ = 0 - -- β = l.maximum - -- sample "not β" repeatedly -> (1 - T)^n - -- sample β once -> T - -- (wrong) => terminates at the last sample, wp ((1 - T)^n ⬝ T) - -- - -- - -- Issue: May terminate earlier. - -- Therefore: Need to pick β such that - -- Total mass of privNoiseGuess below β + τ is at least (1 - T) -- We can use τ to help make this big! - -- Probability that we pick β is T - -- It terminates when we pick β - -- Each time we even (almost) even odds of sampling a positive or negative value - -- We can quantify the probability that we sample a nonnegative value easily: (q = (1 + Prob 0) / 2) - -- If we sample a positive value p, EDS + p > τ | - -- If we sample zero p, EDS + p ≥ τ | terminates - -- If we sample a negative value n, EDS + n < τ | doesn't terminate - - -- Prove that when τ = ?? - -- sv8_G [negatives...] [negative] < τ - -- sv8_G [negatives...] [nonnegative] ≥ τ - - -- Uhh-- no. We can't pick τ. That restricts the total mass to be < 1. - -- Pick β = ??? - -- Let s ∈ ℕ - -- Want to lower bound the probability mass by ((1-β)^(s-1) β) - - -- What's the probability of termination? - - - -- Probability of sampling the special value in noise - -- T needs to be multiplied by privNoiseThresh (for picking τ to be what we want) - let T := (privNoiseThresh ε₁ ε₂ 0) * (privNoiseGuess ε₁ ε₂ β) - - -- Either geo probGeometric will work - suffices ∀ (s : ℕ), Geo T s ≤ sv9_privMax ε₁ ε₂ l s by - s orry - intro s - unfold sv9_privMax - cases s - · simp [Geo] - apply ENNReal.tsum_lb_single 0 - apply ENNReal.mul_left_mono - s orry - · s orry + -- Big condition which forces sv8_cond to be true with mass at least β_geo β (i + 1) + let P (p : (ℤ × { l : List ℤ // l.length = i }) × ℤ) : Prop := + (τ - exactDiffSum (i + 1) l ≤ p.2) ∧ + (p.1.1 < τ - exactDiffSum (i + 1) l) ∧ + (∀ z ∈ p.1.2.1, z < τ - exactDiffSum (i + 1) l) + + let P_ctor (z1 : ℤ) (z2 : { l : List ℤ // l.length = i }) (z3 : ℤ) : + (τ - exactDiffSum (i + 1) l ≤ z3) -> + (z1 < τ - exactDiffSum (i + 1) l) -> + (∀ z ∈ z2.1, z < τ - exactDiffSum (i + 1) l) -> + P ((z1, z2), z3) := by + sorry + + have HP (p : (ℤ × { l : List ℤ // l.length = i }) × ℤ) (Hp : P p) : + (sv8_cond τ l [] p.1.1 (↑p.1.2) p.2 = true) := by + simp [sv8_cond] + apply And.intro + · + sorry + · simp [sv8_sum] + cases Hp + linarith + -- Restrict with one big stupid proposition. Simplify conditional. + apply ENNReal.tsum_lb_subset P + conv => + enter [2, 1, a, 2, 2, 2] + simp [HP a a.2] + + + have Hsum_eq (f : { t // P t } -> ENNReal) : + (∑'(p : {t // P t}), f p) = + (∑'(z1 : {x : ℤ // x < τ - exactDiffSum (i + 1) l }), + ∑'(z2 : {x : { l : List ℤ // l.length = i } // (∀ z ∈ x.1, z < τ - exactDiffSum (i + 1) l) } ), + ∑'(z3 : {x : ℤ // τ - exactDiffSum (i + 1) l ≤ x } ), f ⟨((z1, z2), z3), P_ctor _ _ _ z3.2 z1.2 z2.2⟩) := by + rw [← ENNReal.tsum_prod] + rw [← ENNReal.tsum_prod] + + /- + apply Equiv.tsum_eq_tsum_of_support ?G1 + case G1 => + simp [Function.support] + apply Set.BijOn.equiv + case f => + dsimp [P] + intro ⟨ x, ⟨ a, ⟨ b, c ⟩⟩⟩ + exact ((⟨x.1.1, b⟩, ⟨x.1.2, c⟩), ⟨x.2, a⟩) + · simp + apply And.intro + · simp [Set.MapsTo] + intros + sorry + apply And.intro + · sorry + · sorry + · simp [Function.support] + intros + congr + sorry + -/ + + -- Simplify this into a product + rw [Hsum_eq]; clear Hsum_eq + simp + conv => + enter [2, 1, z1] + conv => + enter [1, z2] + rw [ENNReal.tsum_mul_left] + rw [ENNReal.tsum_mul_left] + rw [ENNReal.tsum_mul_right] + conv => + enter [2, 2] + conv => + enter [1, z2] + rw [ENNReal.tsum_mul_left] + rw [ENNReal.tsum_mul_right] + -/ + sorry + + @@ -2202,7 +2469,8 @@ def sv9_privMax_SPMF (ε₁ ε₂ : ℕ+) (l : List ℕ) : SPMF ℕ := rw [<- sv2_sv3_eq] rw [<- sv1_sv2_eq] apply sv1_ub - · apply sv9_lb ⟩ + · rw [<- sv8_sv9_eq] + apply sv8_lb ⟩ From b8d910fa4b9b5355d79a601da2c287178d1584a7 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Fri, 8 Nov 2024 20:34:15 -0500 Subject: [PATCH 095/100] checkpoint --- .../Queries/UnboundedMax/Properties.lean | 108 +++++++++++++++--- 1 file changed, 94 insertions(+), 14 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 55ef7bc7..e3ab8cbe 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -177,9 +177,18 @@ lemma sv1_alt_eq_alt_alt (ε₁ ε₂ : ℕ+) (l : List ℕ) (H : List ℤ) (K : apply SLang.ext intro x simp [sv1_privMax_alt, sv1_privMax_alt_alt, probBind] + apply tsum_congr + intro τ + congr 1 + apply tsum_congr + intro v0 + congr 1 + apply tsum_congr + intro s + congr 1 + unfold sv1_alt_alt_privMaxC sorry - -- Finally, we should be able to prove that sv1_alt_alt is geometric @@ -187,6 +196,9 @@ lemma sv1_alt_eq_alt_alt (ε₁ ε₂ : ℕ+) (l : List ℕ) (H : List ℤ) (K : + + + /- lemma probWhile_mass_unroll_lb (C : T -> Bool) (F : T -> SLang T) (I : T) : @@ -275,6 +287,12 @@ lemma ENNReal.tsum_iSup_comm (f : T -> U -> ENNReal) : ∑' x, ⨆ y, f x y ≥ intro a apply le_iSup +lemma ENNReal.tsum_iSup_comm' (f : T -> U -> ENNReal) : ⨆ y, ∑' x, f x y ≤ ∑' x, ⨆ y, f x y := by + rw [iSup_le_iff] + intro i + apply ENNReal.tsum_le_tsum + intro a + apply le_iSup lemma iSup_comm_lemma (ε₁ ε₂ : ℕ+) (l : List ℕ) (τ v0 : ℤ): ∑' b, ⨆ i, probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) i ([], v0) b = @@ -2200,7 +2218,6 @@ lemma sv1_cdf_lb ε₁ ε₂ l (τ : ℤ) (v0 : ℤ): -- β_geo_recurrence -/ -/- lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by simp only [sv1_privMax, bind, pure, bind_apply] -- Push the sum over s inwards @@ -2215,7 +2232,7 @@ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by rw [ENNReal.tsum_mul_left] -- Reduce to CDF problem - apply @le_trans _ _ _ (∑' (b : ℤ), (privNoiseThresh ε₁ ε₂) b * ∑' (i : ℤ), (privNoiseGuess ε₁ ε₂) i * 1) _ ?G1 + apply @le_trans _ _ _ (∑' (b : ℤ), (privNoiseThresh ε₁ ε₂) b * 1) _ ?G1 case G1 => apply Eq.le rw [ENNReal.tsum_mul_right] @@ -2225,25 +2242,72 @@ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by cases (privNoiseThresh ε₁ ε₂) simp [DFunLike.coe, SPMF.instFunLike] trivial - have R2 : ∑' (a : ℤ), (privNoiseGuess ε₁ ε₂) a = 1 := by - rw [<- Summable.hasSum_iff ENNReal.summable] - cases (privNoiseGuess ε₁ ε₂) - simp [DFunLike.coe, SPMF.instFunLike] - trivial simp_all apply ENNReal.tsum_le_tsum intro τ apply ENNReal.mul_left_mono - apply ENNReal.tsum_le_tsum - intro v0 - apply ENNReal.mul_left_mono - -- Turn it into a supremum conv => - enter [2, 1, i_1, 1, a] + enter [2, 1, i_1, 2, 1, i ,1, b] simp only [probWhile] rw [ENNReal.iSup_mul] + conv => + enter [2, 1, v0, 2, 1, state_size, 1, state] + + -- Commute out the cut number first + apply le_trans _ ?G1 + case G1 => + apply ENNReal.tsum_le_tsum + intro v0 + apply ENNReal.mul_left_mono + apply ENNReal.tsum_le_tsum + intro state_size + apply ENNReal.tsum_iSup_comm' + apply le_trans _ ?G1 + case G1 => + apply ENNReal.tsum_le_tsum + intro v0 + apply ENNReal.mul_left_mono + apply ENNReal.tsum_iSup_comm' + simp + conv => + enter [2, 1, v0] + rw [ENNReal.mul_iSup] + apply le_trans _ ?G1 + case G1 => + apply ENNReal.tsum_iSup_comm' + + -- Now we're in a world of sanity, because the CDF is actually a CDF. + let ρ : ENNReal := 0 + + apply @le_trans _ _ _ (⨆(y : ℕ), (1 - ρ^y)) + · -- Math + sorry + apply iSup_mono + intro cut + + -- Since we don't care about the value of exactDiffSum, we can revert the initial value + suffices ∀ H, (1 - ρ ^ cut ≤ + ∑' (x : ℤ), + (privNoiseGuess ε₁ ε₂) x * + ∑' (x_1 : ℕ) (x_2 : sv1_state), + if x_1 = sv1_threshold x_2 then probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) cut (H, x) x_2 else 0) by + apply this + + + + skip + + -- + + -- Pick a value for v0 which makes the first conditional false + + + + + + /- -- Commute the supremum out to the middle, so we can pick the cut number for each one apply @le_trans _ _ _ @@ -2262,6 +2326,7 @@ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by apply le_trans _ ?G3 case G3 => apply @le_iSup' _ _ _ _ i rfl + -/ -- Apply the CDF lower bound sorry @@ -2280,7 +2345,7 @@ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by -- apply probGeometric_normalizes <;> simp -- · sorry -- · sorry --/ + lemma sv8_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv8_privMax ε₁ ε₂ l s := by simp [sv8_privMax] @@ -2447,7 +2512,22 @@ lemma sv8_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv8_privMax ε₁ ε₂ l s := by sorry +lemma sv1_attempt_3 (ε₁ ε₂ : ℕ+) (l : List ℕ) : 1 ≤ ∑'(n), sv1_privMax ε₁ ε₂ l n := by + -- We will prove this separately for each τ + simp [sv1_privMax] + rw [ENNReal.tsum_comm] + conv => + enter [2, 1, b] + rw [ENNReal.tsum_mul_left] + apply @le_trans _ _ _ ( ∑' (b : ℤ), (privNoiseThresh ε₁ ε₂) b * 1) + · simp + -- PMF sum is 1 + sorry + apply ENNReal.tsum_le_tsum + intro τ + apply ENNReal.mul_left_mono + sorry From 9f35e5474c512bd64e7b1744f0abe78e36b25975 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Sat, 9 Nov 2024 10:36:31 -0500 Subject: [PATCH 096/100] checkpoint --- .../Queries/UnboundedMax/Properties.lean | 735 +++++------------- 1 file changed, 173 insertions(+), 562 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index e3ab8cbe..989a86ae 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -93,192 +93,6 @@ lemma exactDiffSum_eventually_constant : ∃ K, ∀ K', K ≤ K' -> exactDiffSum - -def sv1_privMax_alt (ε₁ ε₂ : ℕ+) (l : List ℕ) (H : List ℤ) : SLang ℕ := do - let τ <- privNoiseThresh ε₁ ε₂ - let v0 <- privNoiseGuess ε₁ ε₂ - let sk <- probWhile (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (H, v0) - return (sv1_threshold sk) - -lemma sv1_sv1_alt_base_case (ε₁ ε₂ : ℕ+) (l : List ℕ) : - sv1_privMax ε₁ ε₂ l = sv1_privMax_alt ε₁ ε₂ l [] := by - simp [sv1_privMax, sv1_privMax_alt] - -lemma probWhile_unroll (C : T -> Bool) (F : T -> SLang T) (I : T) : - probWhile C F I = - (if (C I) then probPure I else (F I) >>= probWhile C F) := by - sorry - -lemma probWhile_mass_1_unroll (C : T -> Bool) (F : T -> SLang T) (I : T) : - (∑'(v), F I v = 1) -> - (∀ I', F I I' > 0 -> ∑'(v), (probWhile C F I' v) = 1) -> - ∑'(v), probWhile C F I v = 1 := by - intro H1 H2 - rw [probWhile_unroll] - cases C I <;> simp - rw [ENNReal.tsum_comm] - conv => - enter [1] - enter [1, b] - rw [ENNReal.tsum_mul_left] - -- doable - rw [<- H1] - apply Equiv.tsum_eq_tsum_of_support ?S - case S => - apply Equiv.Set.ofEq - simp [Function.support] - apply Set.ext - intro x - simp - intro A - apply Classical.by_contradiction - intro B - simp at B - have D : (F I x > 0) := by - apply Classical.by_contradiction - simp - trivial - have H2' := H2 _ D - simp_all - simp - intro A B E _ - rw [H2] - · simp - · apply Classical.by_contradiction - simp - trivial - -lemma pm_mass_1 (ε₁ ε₂ : ℕ+) (τ : ℤ) (K : ℕ) : - (∀ H : List ℤ, K ≤ H.length -> ∑'(v), sv1_privMax_alt ε₁ ε₂ l H v = 1) -> - (∑'(v), sv1_privMax ε₁ ε₂ l v = 1) := by - -- Unroll K times - induction K - · intro A - rw [sv1_sv1_alt_base_case] - apply A - simp - · rename_i K' IH - intro H' - -- Otherwise, - sorry - -def sv1_alt_alt_privMaxC (τ : ℤ) (s : sv1_state) : Bool := - decide (sv1_noise s < τ) - -def sv1_privMax_alt_alt (ε₁ ε₂ : ℕ+) (H : List ℤ) : SLang ℕ := do - let τ <- privNoiseThresh ε₁ ε₂ - let v0 <- privNoiseGuess ε₁ ε₂ - let sk <- probWhile (sv1_alt_alt_privMaxC τ) (sv1_privMaxF ε₁ ε₂) (H, v0) - return (sv1_threshold sk) - -lemma sv1_alt_eq_alt_alt (ε₁ ε₂ : ℕ+) (l : List ℕ) (H : List ℤ) (K : ℕ) (HK : ∀ K', H.length ≤ K' -> exactDiffSum K' l = 0) : - sv1_privMax_alt ε₁ ε₂ l H = sv1_privMax_alt_alt ε₁ ε₂ H := by - -- H is in the regieme where exactDiffSum is constant - apply SLang.ext - intro x - simp [sv1_privMax_alt, sv1_privMax_alt_alt, probBind] - apply tsum_congr - intro τ - congr 1 - apply tsum_congr - intro v0 - congr 1 - apply tsum_congr - intro s - congr 1 - unfold sv1_alt_alt_privMaxC - sorry - --- Finally, we should be able to prove that sv1_alt_alt is geometric - - - - - - - - - -/- - -lemma probWhile_mass_unroll_lb (C : T -> Bool) (F : T -> SLang T) (I : T) : - ∑'t, (((F I) >>= probWhile C F) t) ≤ ∑'t, probWhile C F I t := by - conv => - rhs - enter [1, t] - rw [probWhile_unroll] - cases (C I) - · simp only [Bool.false_eq_true, ↓reduceIte] - rfl - · conv => - enter [2] - simp - -- This now uses the upper boudn proof (which we need SPMF for but that's OK) - have A : tsum (probPure I) = 1 := by s orry - s orry - - -lemma probWhile_mass_unroll_lb_eval {T U : Type} (C : T -> Bool) (F : T -> SLang T) (g : U -> T): - ∑'t, ∑'(u : U), (((F (g u)) >>= probWhile C F) t) ≤ ∑'t, ∑'(u : U), probWhile C F (g u) t := by - conv => - lhs - rw [ENNReal.tsum_comm] - conv => - rhs - rw [ENNReal.tsum_comm] - apply ENNReal.tsum_le_tsum - intro i - apply probWhile_mass_unroll_lb - --- Apply this lemma K times to the sv0 loop (sum over v0) -/- -lemma sv0_loop_mass_unroll_lb (ε₁ ε₂ : ℕ+) (τ : ℤ) (l : List ℕ) (K : ℕ) : - ∑'t, ∑'(v0 : ℤ), (probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (K, v0)) t ≤ - ∑'t, ∑'(v0 : ℤ), (probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (0, v0)) t := by - induction K - · rfl - · rename_i K' IH - apply le_trans _ (probWhile_mass_unroll_lb_eval (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (fun v0 => (0, v0))) - simp [sv0_privMaxF, sv0_threshold] - -- Not sure - s orry --/ --/ - - --- probWhile sv0_privMaxC (sv0_privMaxF ε₁ ε₂ τ l) (False, (0, v0)) - - --- Intuition: If there is a lower bound on every loop iterate terminating, then the termination probability is 1. - --- Fix l, ε₁, ε₂ --- Can prove that in any state, the loop has probability at least β to terminate --- How to turn this into a lower bound on the termination probability? - - -/- -def gen_poisson_trial (C : ℕ -> T → Bool) (F : ℕ × T -> SLang T) (I0 : T) : SLang (Bool × (ℕ × T)) := - probWhile - (fun (x : Bool × (ℕ × T)) => x.1) - (fun v => do - let v1 <- F v.2 - return (C v.2.1 v1, (v.2.1 + 1, v1))) - (True, (0, I0)) - - -lemma gen_poisson_trial_lb (C : ℕ -> T → Bool) (F : ℕ × T -> SLang T) (I : T) (β : NNReal) : - (∀ x, β ≤ (F x >>= fun v => return C x.1 v) True) -> - (∀ t, probGeometric (fun b => if b then β else (1 - β)) t.2.1 ≤ gen_poisson_trial C F I t) := by - intro H - intro t - -- rcases t with ⟨ b, ⟨ n, t ⟩ ⟩ - unfold probGeometric - unfold gen_poisson_trial - unfold geoLoopCond - s orry --/ - - lemma ENNReal.tsum_iSup_comm (f : T -> U -> ENNReal) : ∑' x, ⨆ y, f x y ≥ ⨆ y, ∑' x, f x y := by apply LE.le.ge rw [iSup_le_iff] @@ -2138,6 +1952,16 @@ lemma ENNReal.tsum_lb_subset (P : T -> Prop) (f : T -> ENNReal) (l : ENNReal) : apply ENNReal.tsum_comp_le_tsum_of_injective simp [Function.Injective] + +lemma ENNReal.tsum_split (P : T -> Prop) (f : T -> ENNReal) : + ∑' (a : T), f a = (∑'(a : {t : T // P t}), f a.1) + (∑'(a : {t : T // ¬P t}), f a.1) := by + sorry + + + + + +/- def β_geo (β : ENNReal) : SLang ℕ := (probGeometric (fun x => if x then β else 1 - β)) def β_geo_recurrence (β : ENNReal) (n : ℕ) (H : n > 0) : β_geo β (n + 1) = β * β_geo β n := by @@ -2149,75 +1973,39 @@ def β_geo_recurrence (β : ENNReal) (n : ℕ) (H : n > 0) : β_geo β (n + 1) = rename_i h rw [mul_pow_sub_one h] - -def β_geo' (β : ENNReal) : SLang ℕ := +def β_geo' (β : ℝ) : ℕ -> ℝ := fun N => match N with | 0 => 0 - | Nat.succ N' => β_geo β N' + | Nat.succ N' => ENNReal.toReal (β_geo β N') +-/ +def geo_cdf (β : ENNReal) (n : ℕ) : ENNReal := 1 - (1 - β)^n -/-- -lemma sv1_cdf_lb ε₁ ε₂ l (τ : ℤ) (v0 : ℤ): - sv1_privMaxC τ l ([], v0) = true -> - ∃ β : ℝ, - (0 < β) ∧ - (β < 1) ∧ - ∀ i : ℕ , - (β_geo β i ≤ - (∑' (a : sv1_state), - probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) i ([], v0) a * probPure (sv1_threshold a) i)) := by - - -- exactDiffSum is negative or zero for all states - -- decide (exactDiffSum (sv1_threshold s) l + (sv1_noise s) < τ) - intro Htrue - have Hβ : ∃ β : ℝ, 0 < β ∧ β < 1 ∧ (∑'(z : ℤ), if z < τ then (privNoiseGuess ε₁ ε₂ z) else 0) ≤ ENNReal.ofReal β := by - -- This is true due to CDF of laplace distribution - sorry - rcases Hβ with ⟨ β, ⟨ H1, ⟨ H2, Hsum ⟩ ⟩ ⟩ - exists β - apply And.intro H1 - apply And.intro H2 - intro i +lemma geo_cdf_rec (β : ENNReal) (n : ℕ) : + geo_cdf β (n + 1) = β + (1 - β) * geo_cdf β n := by + unfold geo_cdf + sorry - -- Prove this by induction - -- #check (probGeometric (fun x => if x then ENNReal.ofReal β else 0)) +-- There is a value such that sampling at least that value implies the loop definitely terminiates +lemma lucky_guess (τ : ℤ) (l : List ℕ) : ∃ (K : ℤ), ∀ A, ∀ (K' : ℤ), K ≤ K' -> exactDiffSum A l + K' ≥ τ := by + sorry - suffices - ∀ init, sv1_privMaxC τ l init = true -> - β_geo β i ≤ ∑' (a : sv1_state), probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) i init a * probPure (sv1_threshold a) i by - apply this - sorry - -- FIXME: Does β_geo work? Does β^i work? - induction i - · simp [β_geo] - · rename_i i IH - intro init - cases i - · - sorry - · intro H1 - rename_i i - conv => - rhs - enter [1, a] - unfold probWhileCut - unfold probWhileFunctional - split - · sorry - · exfalso - aesop +lemma ite_conv_left {P : Prop} {D} {a b c : ENNReal} (H : a = c) : @ite _ P D a b = @ite _ P D c b := by + split <;> trivial + +lemma ite_mono_left {P : Prop} {D} {a b c : ENNReal} (H : a ≤ c) : @ite _ P D a b ≤ @ite _ P D c b := by + split <;> trivial + +lemma ite_lemma_1 {P : Prop} {D} {f : T -> ENNReal} : ∑'(a : T), @ite _ P D (f a) 0 = @ite _ P D (∑'(a : T), f a) 0 := by + split + · rfl + · simp + - -- rw [β_geo_recurrence _ _ (by simp)] - -- apply le_trans - -- · apply ENNReal.mul_left_mono - -- · apply IH - -- sorry - -- β_geo_recurrence --/ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by simp only [sv1_privMax, bind, pure, bind_apply] -- Push the sum over s inwards @@ -2279,257 +2067,157 @@ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by apply ENNReal.tsum_iSup_comm' -- Now we're in a world of sanity, because the CDF is actually a CDF. - let ρ : ENNReal := 0 - apply @le_trans _ _ _ (⨆(y : ℕ), (1 - ρ^y)) + -- The lucky event: sampling above a value T, which forces the loop to terminate + rcases (lucky_guess τ l) with ⟨ K, HK ⟩ + let PLucky (K' : ℤ) : Prop := K ≤ K' + have HLucky : ∀ (K' : ℤ), ∀ A, PLucky K' → exactDiffSum A l + K' ≥ τ := by aesop + clear HK + + -- We will split the sum based on PLucky at each step + + -- ρ is the probability of the lucky event + let ρ : ENNReal := (∑'(a : {t : ℤ // PLucky t}), privNoiseGuess ε₁ ε₂ a.1) + have Hρ_lb : 0 < ρ := by sorry + have Hρ_ub : ρ ≤ 1 := by sorry + + -- Bound the CDF below by the geometric CDF + apply @le_trans _ _ _ (⨆(y : ℕ), geo_cdf ρ y) · -- Math + unfold geo_cdf sorry apply iSup_mono intro cut - -- Since we don't care about the value of exactDiffSum, we can revert the initial value - suffices ∀ H, (1 - ρ ^ cut ≤ - ∑' (x : ℤ), - (privNoiseGuess ε₁ ε₂) x * - ∑' (x_1 : ℕ) (x_2 : sv1_state), - if x_1 = sv1_threshold x_2 then probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) cut (H, x) x_2 else 0) by - apply this - - - - skip - - -- - - -- Pick a value for v0 which makes the first conditional false - - - - - - /- - - -- Commute the supremum out to the middle, so we can pick the cut number for each one - apply @le_trans _ _ _ - (∑' (i_1 : ℕ), ⨆ i, ∑'(a : sv1_state), - probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) i ([], v0) a * probPure (sv1_threshold a) i_1) _ _ ?G1 - case G1 => - apply ENNReal.tsum_le_tsum - intro i - apply ENNReal.tsum_iSup_comm - apply @le_trans _ _ _ - (∑' (i : ℕ), ∑'(a : sv1_state), - probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) i ([], v0) a * probPure (sv1_threshold a) i) _ _ ?G1 - case G1 => - apply ENNReal.tsum_le_tsum - intro i - apply le_trans _ ?G3 - case G3 => apply @le_iSup' _ _ _ _ i - rfl - -/ - - -- Apply the CDF lower bound - sorry - -- rcases (sv1_cdf_lb ε₁ ε₂ l τ v0) with ⟨ β, ⟨ H1, ⟨ H2, Hlb ⟩ ⟩ ⟩ - -- apply le_trans _ ?G2 - -- case G2 => - -- apply ENNReal.tsum_le_tsum - -- intro t - -- apply Hlb - - -- -- Geometric sum has a closed form - -- apply Eq.le - -- symm - -- unfold β_geo - -- apply probGeometric_normalizes <;> simp - -- · sorry - -- · sorry - - -lemma sv8_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv8_privMax ε₁ ε₂ l s := by - simp [sv8_privMax] - -- Factor out the choice of thresh (we will work for all τ) - rw [ENNReal.tsum_comm] - conv => - enter [2, 1, b] - rw [ENNReal.tsum_mul_left] - apply @le_trans _ _ _ (∑'(b : ℤ), privNoiseThresh ε₁ ε₂ b * 1) - · simp - apply Eq.le - symm - rcases (privNoiseThresh ε₁ ε₂) - simp [DFunLike.coe] - rw [<- Summable.hasSum_iff ENNReal.summable] - trivial - apply ENNReal.tsum_le_tsum - intro τ - apply ENNReal.mul_left_mono - - -- Pick β based on τ, and exactDiffSum - have β : ENNReal := 0 -- (∑' (i_1 : { x // x < τ - exactDiffSum (i + 1) l }), (privNoiseGuess ε₁ ε₂) ↑i_1) - - -- Lower bound by probGeometric - apply @le_trans _ _ _ (∑'(i : ℕ), β_geo' β i * 1) - · apply Eq.le - symm - simp only [β_geo', β_geo, mul_one] - -- rw [ENNReal.tsum_eq_add_tsum_ite Nat.zero] - -- simp only [Nat] + -- Because v0 is not in the loop, we need to do one of the unrollings first + -- Our IH is going to include a condition on "present" + cases cut + · -- cut=0 case is trivial + simp [probWhileCut, geo_cdf] + rename_i cut + + rw [geo_cdf_rec] + rw [ENNReal.tsum_split PLucky] + apply add_le_add + · -- Lucky guess + simp [probWhileCut, probWhileFunctional, sv1_privMaxC, sv1_noise] + conv => + rhs + enter [1, a, 2, 1, x, 1, x1] + rw [ite_conv_left + (by + rw [ite_eq_right_iff.mpr] + intro i + exfalso + rcases a with ⟨ v, Hv ⟩ + simp [sv1_threshold] at i + have luck := HLucky v 0 Hv + apply (LT.lt.not_le i) + trivial)] + rfl + -- The rightmost sum is 1 sorry - -- apply probGeometric_normalizes <;> simp - -- · -- β fact - -- sorry - -- · -- β fact - -- sorry - - -- Now, prove that it is bounded below by probGeometric - apply ENNReal.tsum_le_tsum - intro i - simp - cases i - · simp [β_geo'] - rename_i i - simp [β_geo'] - cases i - · simp [β_geo] - rename_i i - simp [β_geo] - - -- Commute all sums out to the left - conv => - enter [2, 1, v0] - rw [<- ENNReal.tsum_mul_left] - -- enter [1, vs] - -- rw [<- ENNReal.tsum_mul_left] - -- rw [<- ENNReal.tsum_mul_left] - rw [← ENNReal.tsum_prod] - - -- Eval the presample - conv => - enter [2, 1, a, 2, 1] - rw [sv4_presample_eval'] - simp - - unfold sv8_cond - simp [β_geo] - - - - - - - - - - - -- rw [← ENNReal.tsum_prod] - - -- simp [β_geo] - /- - - -- Big condition which forces sv8_cond to be true with mass at least β_geo β (i + 1) - let P (p : (ℤ × { l : List ℤ // l.length = i }) × ℤ) : Prop := - (τ - exactDiffSum (i + 1) l ≤ p.2) ∧ - (p.1.1 < τ - exactDiffSum (i + 1) l) ∧ - (∀ z ∈ p.1.2.1, z < τ - exactDiffSum (i + 1) l) - - let P_ctor (z1 : ℤ) (z2 : { l : List ℤ // l.length = i }) (z3 : ℤ) : - (τ - exactDiffSum (i + 1) l ≤ z3) -> - (z1 < τ - exactDiffSum (i + 1) l) -> - (∀ z ∈ z2.1, z < τ - exactDiffSum (i + 1) l) -> - P ((z1, z2), z3) := by + + -- Unlucky + suffices (∀ H, ∀ a : {t : ℤ // ¬ PLucky t}, geo_cdf ρ cut ≤ + ∑' (x : ℕ) (x_1 : sv1_state), + if x = sv1_threshold x_1 then probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (cut + 1) (H, ↑a) x_1 else 0) by + apply le_trans _ ?G1 + case G1 => + apply ENNReal.tsum_le_tsum + intro a + apply ENNReal.mul_left_mono + apply this + simp + rw [ENNReal.tsum_mul_right] + apply ENNReal.mul_right_mono + apply Eq.le + -- Math sorry - have HP (p : (ℤ × { l : List ℤ // l.length = i }) × ℤ) (Hp : P p) : - (sv8_cond τ l [] p.1.1 (↑p.1.2) p.2 = true) := by - simp [sv8_cond] - apply And.intro - · - sorry - · simp [sv8_sum] - cases Hp - linarith - -- Restrict with one big stupid proposition. Simplify conditional. - apply ENNReal.tsum_lb_subset P - conv => - enter [2, 1, a, 2, 2, 2] - simp [HP a a.2] - - - have Hsum_eq (f : { t // P t } -> ENNReal) : - (∑'(p : {t // P t}), f p) = - (∑'(z1 : {x : ℤ // x < τ - exactDiffSum (i + 1) l }), - ∑'(z2 : {x : { l : List ℤ // l.length = i } // (∀ z ∈ x.1, z < τ - exactDiffSum (i + 1) l) } ), - ∑'(z3 : {x : ℤ // τ - exactDiffSum (i + 1) l ≤ x } ), f ⟨((z1, z2), z3), P_ctor _ _ _ z3.2 z1.2 z2.2⟩) := by - rw [← ENNReal.tsum_prod] - rw [← ENNReal.tsum_prod] + -- Now we have the right inductive structure, I think? + induction cut + · -- Base case is trivial? + simp [geo_cdf] + · rename_i cut IH + intro H a + rcases a with ⟨ v, Hv ⟩ + simp - /- - apply Equiv.tsum_eq_tsum_of_support ?G1 - case G1 => - simp [Function.support] - apply Set.BijOn.equiv - case f => - dsimp [P] - intro ⟨ x, ⟨ a, ⟨ b, c ⟩⟩⟩ - exact ((⟨x.1.1, b⟩, ⟨x.1.2, c⟩), ⟨x.2, a⟩) + -- Because the first sample not lucky, we can't say anything about the branch we end up in + -- We will bound it separately for both + have advance (x : sv1_state) : + ((((sv1_privMaxF ε₁ ε₂ (H, v)) >>= (probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (cut + 1))) x) + ≤ (probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (cut + 1 + 1) (H, v)) x) := by + conv => + rhs + unfold probWhileCut + unfold probWhileFunctional + split · simp - apply And.intro - · simp [Set.MapsTo] - intros - sorry - apply And.intro - · sorry - · sorry - · simp [Function.support] - intros - congr - sorry - -/ + · simp + -- Uses the upper bound here, but this is provable + sorry + apply le_trans _ ?G1 + case G1 => + apply ENNReal.tsum_le_tsum + intro x + apply ENNReal.tsum_le_tsum + intro x_1 + apply ite_mono_left + apply advance + clear advance - -- Simplify this into a product - rw [Hsum_eq]; clear Hsum_eq - simp - conv => - enter [2, 1, z1] - conv => - enter [1, z2] - rw [ENNReal.tsum_mul_left] - rw [ENNReal.tsum_mul_left] - rw [ENNReal.tsum_mul_right] - conv => - enter [2, 2] + -- Now we want to commute out the randomness associate to that s1_privMaxF + apply le_trans _ ?G1 + case G1 => + apply ENNReal.tsum_le_tsum + intro x + apply ENNReal.tsum_le_tsum + intro x_1 + simp + rw [<- ite_lemma_1] conv => - enter [1, z2] - rw [ENNReal.tsum_mul_left] - rw [ENNReal.tsum_mul_right] - - -/ - - + enter [2] + conv => + enter [1, a] + rw [ENNReal.tsum_comm] + rw [ENNReal.tsum_comm] - sorry + -- Split the sv1_state sum + conv => + enter [2] + unfold sv1_state + rw [ENNReal.tsum_prod'] + rw [ENNReal.tsum_comm] + -- Now, condition on the luckiness of the next value + rw [ENNReal.tsum_split PLucky] -lemma sv1_attempt_3 (ε₁ ε₂ : ℕ+) (l : List ℕ) : 1 ≤ ∑'(n), sv1_privMax ε₁ ε₂ l n := by - -- We will prove this separately for each τ - simp [sv1_privMax] - rw [ENNReal.tsum_comm] - conv => - enter [2, 1, b] - rw [ENNReal.tsum_mul_left] - apply @le_trans _ _ _ ( ∑' (b : ℤ), (privNoiseThresh ε₁ ε₂) b * 1) - · simp - -- PMF sum is 1 - sorry - apply ENNReal.tsum_le_tsum - intro τ - apply ENNReal.mul_left_mono + -- Split the sum and the recurrence relation + rw [geo_cdf_rec] + apply add_le_add + · -- Guess is lucky + -- The loop will terminate and we can show it + sorry + · -- Guess is unlucky + -- Commute out the samples related to the first sample (which will evenetually become a (1- ρ) factor) + conv => + enter [2, 1, a, 1, b] + conv => + enter [1, c] + conv => + enter [1, d] + rw [<- mul_ite_zero] + rw [ENNReal.tsum_mul_left] + rw [ENNReal.tsum_mul_left] + unfold sv1_state at IH - sorry + -- Apply the IH + sorry /-- @@ -2538,92 +2226,15 @@ sv9 normalizes because sv1 normalizes def sv9_privMax_SPMF (ε₁ ε₂ : ℕ+) (l : List ℕ) : SPMF ℕ := ⟨ sv9_privMax ε₁ ε₂ l, by + rw [<- sv8_sv9_eq] + rw [<- sv7_sv8_eq] + rw [<- sv6_sv7_eq] + rw [<- sv5_sv6_eq] + rw [<- sv4_sv5_eq] + rw [<- sv3_sv4_eq] + rw [<- sv2_sv3_eq] + rw [<- sv1_sv2_eq] rw [Summable.hasSum_iff ENNReal.summable] apply LE.le.antisymm - · rw [<- sv8_sv9_eq] - rw [<- sv7_sv8_eq] - rw [<- sv6_sv7_eq] - rw [<- sv5_sv6_eq] - rw [<- sv4_sv5_eq] - rw [<- sv3_sv4_eq] - rw [<- sv2_sv3_eq] - rw [<- sv1_sv2_eq] - apply sv1_ub - · rw [<- sv8_sv9_eq] - apply sv8_lb ⟩ - - - - -/- - -def sv0_eventually_geo_check (ε₁ ε₂ : ℕ+) (τ : ℤ) : SLang Bool := do - let v <- privNoiseGuess ε₁ ε₂ - return v > τ - - -/- --- The sv0 loop is eventually geometric -lemma sv0_loop_eventually_geometric ε₁ ε₂ τ l : - ∃ K, ∀ s_eval, ∀ v0, - probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (K, v0) s_eval = - probGeometric (sv0_eventually_geo_check ε₁ ε₂ τ) s_eval.1 := by - rcases (@exactDiffSum_eventually_constant l) with ⟨ K, HK ⟩ - exists K - intro s_eval - intro v0 - unfold probGeometric - unfold geoLoopCond - unfold geoLoopBody - unfold sv0_privMaxC - unfold sv0_privMaxF - unfold sv0_eventually_geo_check - -/ - - s orry - - --- def sv0_privMaxC (τ : ℤ) (l : List ℕ) (s : sv0_state) : Bool := --- decide (exactDiffSum (sv0_threshold s) l + (sv0_noise s) < τ) - --- sv0_privMaxC is eventually just a (geometric) check - -lemma sv0_norm_loop_le ε₁ ε₂ τ : ∑'v0, ∑'n, probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (n, v0) ≤ 1 := by - s orry - -lemma probWhile_unroll (C : T -> Bool) (F : T -> SLang T) (I : T) : - probWhile C F I = - (if (C I) then probPure I else (F I) >>= probWhile C F) := by - s orry - -- (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (n + Δ, v0) - -lemma sv0_norm_loop_ge ε₁ ε₂ v0 τ : 1 ≤ ∑'vf, ∑'n, probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (0, v0) (n, vf) := by - -- To establish a lower bound, it suffices to shift the starting point - have Hshift : ∀ Δ v0, - ∑' (n : ℕ), probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (n, v0) ≤ - ∑' (n : ℕ), probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (0, v0) := by - intro Δ - induction Δ - · simp - · rename_i Δ' IH - intro v0 - apply le_trans _ (IH _) - clear IH - conv => - rhs - enter [1, n] - rw [probWhile_unroll] - apply ENNReal.tsum_le_tsum - intro n - split - · s orry - · s orry - - - s orry - s orry - -lemma sv0_norm ε₁ ε₂ l : ∑'(x : ℕ), sv0_privMax ε₁ ε₂ l x = 1 := by - -- unfold sv0_privMax - s orry --/ + · apply sv1_ub + · apply sv1_lb ⟩ From 2e092b39620eb5ac256fa8bb9661ba7775df9ee4 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Sat, 9 Nov 2024 12:29:56 -0500 Subject: [PATCH 097/100] skeleton --- .../Queries/UnboundedMax/Properties.lean | 19 ++++++++++++++++++- 1 file changed, 18 insertions(+), 1 deletion(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 989a86ae..90983dc4 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -1955,6 +1955,7 @@ lemma ENNReal.tsum_lb_subset (P : T -> Prop) (f : T -> ENNReal) (l : ENNReal) : lemma ENNReal.tsum_split (P : T -> Prop) (f : T -> ENNReal) : ∑' (a : T), f a = (∑'(a : {t : T // P t}), f a.1) + (∑'(a : {t : T // ¬P t}), f a.1) := by + sorry @@ -2214,7 +2215,23 @@ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by rw [ENNReal.tsum_mul_left] unfold sv1_state at IH - -- Apply the IH + apply le_trans _ ?G1 + case G1 => + apply ENNReal.tsum_le_tsum + intro a + apply ENNReal.tsum_le_tsum + intro b + apply ENNReal.mul_left_mono + apply IH + simp + conv => + enter [2, 1, a] + rw [ENNReal.tsum_mul_right] + rw [ENNReal.tsum_mul_right] + apply ENNReal.mul_right_mono + + -- Conclude by simplification + simp only [sv1_privMaxF, bind, pure, bind_apply, pure_apply, mul_ite, mul_one, mul_zero] sorry From 1c1a02d66ff001ebc0b98210561e51aff623bd3d Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Sat, 9 Nov 2024 17:11:25 -0500 Subject: [PATCH 098/100] checkpoint --- .../Queries/UnboundedMax/Privacy.lean | 74 +-- .../Queries/UnboundedMax/Properties.lean | 528 ++++++++++++++---- 2 files changed, 430 insertions(+), 172 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean index ce10bd90..606fcab2 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Privacy.lean @@ -119,78 +119,6 @@ lemma laplace_inequality_sub (τ τ' : ℤ) (Δ : ℕ+) : simp - -lemma exactClippedSum_append : exactClippedSum i (A ++ B) = exactClippedSum i A + exactClippedSum i B := by - simp [exactClippedSum] - -lemma exactDiffSum_append : exactDiffSum i (A ++ B) = exactDiffSum i A + exactDiffSum i B := by - simp [exactDiffSum] - rw [exactClippedSum_append] - rw [exactClippedSum_append] - linarith - -lemma sv8_sum_append : sv8_sum (A ++ B) vp v0 = sv8_sum A vp v0 + sv8_sum B vp v0 - v0 := by - simp [sv8_sum] - simp [exactDiffSum_append] - linarith - -lemma sv8_sum_comm : sv8_sum (A ++ B) vp v0 = sv8_sum (B ++ A) vp v0 := by - unfold sv8_sum - simp - simp [exactDiffSum, exactClippedSum] - linarith - -lemma sv8_G_comm : sv8_G (A ++ B) vp v0 vf = sv8_G (B ++ A) vp v0 vf := by - revert vp v0 - induction vf - · intros - apply sv8_sum_comm - · intro v0 vp - rename_i next rest IH - unfold sv8_G - rw [sv8_sum_comm] - congr 1 - apply IH - --- -- IDK --- lemma sv8_G_cons : sv8_G (x :: L) vp v0 vf = 0 := by --- revert vp v0 --- induction vf --- · intros v0 vp --- simp [sv8_G] --- s orry --- · intro vp v0 --- simp [sv8_G] --- s orry --- -- unfold sv8_G - -lemma exactDiffSum_nonpos : exactDiffSum point L ≤ 0 := by - simp [exactDiffSum, exactClippedSum] - induction L - · simp - · rename_i l ll IH - simp - apply le_trans - · apply add_le_add - · rfl - · apply IH - simp - -lemma exactDiffSum_singleton_le_1 : -1 ≤ exactDiffSum point [v] := by - simp [exactDiffSum, exactClippedSum] - cases (Classical.em (point ≤ v)) - · right - trivial - · left - rename_i h - simp at h - rw [Int.min_def] - simp - split - · linarith - · linarith - - -- Coercions nonsense lemma DSN (N : ℕ) (H : Neighbour L1 L2) : ((exactDiffSum N L1) : ℝ) - (exactDiffSum N L2) ≤ 1 := by cases H @@ -350,7 +278,7 @@ lemma Hsens_cov_vk (v0 : ℤ) (vs : List ℤ) (l₁ l₂ : List ℕ) (point : lemma sv8_privMax_pmf_DP (ε : NNReal) (Hε : ε = ε₁ / ε₂) : - PureDPSystem.prop (@sv9_privMax_SPMF PureDPSystem laplace_pureDPSystem ε₁ ε₂) ε := by + PureDPSystem.prop (@sv9_privMax_SPMF ε₁ ε₂) ε := by -- Unfold DP definitions simp [DPSystem.prop] apply singleton_to_event diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 90983dc4..729040d2 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -12,6 +12,8 @@ open Classical namespace SLang +section equiv + variable [dps : DPSystem ℕ] variable [dpn : DPNoise dps] @@ -124,6 +126,65 @@ lemma iSup_comm_lemma (ε₁ ε₂ : ℕ+) (l : List ℕ) (τ v0 : ℤ): apply probWhileCut_monotonic +lemma sv1_loop_ub ε₁ ε₂ l : ∀ L : List ℤ, ∀ (v0 : ℤ), (∑' (x : sv1_state), probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) cut (L, v0) x ≤ 1) := by + induction cut + · simp [probWhileCut] + · rename_i cut' IH + intro L v0 + simp [probWhileCut, probWhileFunctional] + split + · simp + simp [sv1_privMaxF] + conv => + enter [1, 1, x] + conv => + enter [1, a] + rw [<- ENNReal.tsum_mul_right] + simp + rw [ENNReal.tsum_comm] + enter [1, b] + + apply + @le_trans _ _ _ + (∑' (x : sv1_state) (b : ℤ), (privNoiseGuess ε₁ ε₂) b * probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) cut' (L ++ [v0], b) x) + _ ?G5 ?G6 + case G5 => + apply ENNReal.tsum_le_tsum + intro a + apply ENNReal.tsum_le_tsum + intro b + unfold sv1_state + rw [ENNReal.tsum_eq_add_tsum_ite (L ++ [v0], b)] + simp + conv => + rhs + rw [<- add_zero (_ * _)] + apply add_le_add + · simp + · simp + intros + aesop + case G6 => + rw [ENNReal.tsum_comm] + conv => + enter [1, 1, b] + rw [ENNReal.tsum_mul_left] + apply @le_trans _ _ _ (∑' (b : ℤ), (privNoiseGuess ε₁ ε₂) b * 1) + · apply ENNReal.tsum_le_tsum + intro a + apply ENNReal.mul_left_mono + apply IH + · simp + apply Eq.le + rw [<- Summable.hasSum_iff ENNReal.summable] + cases (privNoiseGuess ε₁ ε₂) + simp [DFunLike.coe, SPMF.instFunLike] + trivial + · simp + + + + lemma sv1_ub ε₁ ε₂ l : ∑'s, sv1_privMax ε₁ ε₂ l s ≤ 1 := by unfold sv1_privMax unfold sv1_threshold @@ -179,65 +240,9 @@ lemma sv1_ub ε₁ ε₂ l : ∑'s, sv1_privMax ε₁ ε₂ l s ≤ 1 := by rw [iSup_comm_lemma] apply iSup_le_iff.mpr intro cut - suffices (∀ L, ∑' (x : sv1_state), probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) cut (L, v0) x ≤ 1) by - apply this - revert v0 - induction cut - · simp [probWhileCut] - · rename_i cut' IH - intro v0 L - simp [probWhileCut, probWhileFunctional] - split - · simp - simp [sv1_privMaxF] - conv => - enter [1, 1, x] - conv => - enter [1, a] - rw [<- ENNReal.tsum_mul_right] - simp - rw [ENNReal.tsum_comm] - enter [1, b] - - apply - @le_trans _ _ _ - (∑' (x : sv1_state) (b : ℤ), (privNoiseGuess ε₁ ε₂) b * probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) cut' (L ++ [v0], b) x) - _ ?G5 ?G6 - case G5 => - apply ENNReal.tsum_le_tsum - intro a - apply ENNReal.tsum_le_tsum - intro b - unfold sv1_state - rw [ENNReal.tsum_eq_add_tsum_ite (L ++ [v0], b)] - simp - conv => - rhs - rw [<- add_zero (_ * _)] - apply add_le_add - · simp - · simp - intros - aesop - case G6 => - rw [ENNReal.tsum_comm] - conv => - enter [1, 1, b] - rw [ENNReal.tsum_mul_left] - apply @le_trans _ _ _ (∑' (b : ℤ), (privNoiseGuess ε₁ ε₂) b * 1) - · apply ENNReal.tsum_le_tsum - intro a - apply ENNReal.mul_left_mono - apply IH - · simp - apply Eq.le - rw [<- Summable.hasSum_iff ENNReal.summable] - cases (privNoiseGuess ε₁ ε₂) - simp [DFunLike.coe, SPMF.instFunLike] - trivial - · simp - + apply sv1_loop_ub +/- /- History-aware progam computes the same as the history-agnostic program -/ @@ -285,9 +290,8 @@ lemma sv0_eq_sv1 ε₁ ε₂ l : sv0_privMax ε₁ ε₂ l = sv1_privMax ε₁ -- RHS: sum over all lists of length "result"? -- rw [tsum_ite_eq] simp [sv1_threshold] - - - sorry + s orry + -/ @@ -1955,12 +1959,8 @@ lemma ENNReal.tsum_lb_subset (P : T -> Prop) (f : T -> ENNReal) (l : ENNReal) : lemma ENNReal.tsum_split (P : T -> Prop) (f : T -> ENNReal) : ∑' (a : T), f a = (∑'(a : {t : T // P t}), f a.1) + (∑'(a : {t : T // ¬P t}), f a.1) := by - - sorry - - - - + symm + apply tsum_add_tsum_compl <;> apply ENNReal.summable /- def β_geo (β : ENNReal) : SLang ℕ := (probGeometric (fun x => if x then β else 1 - β)) @@ -1984,15 +1984,172 @@ def β_geo' (β : ℝ) : ℕ -> ℝ := def geo_cdf (β : ENNReal) (n : ℕ) : ENNReal := 1 - (1 - β)^n -lemma geo_cdf_rec (β : ENNReal) (n : ℕ) : +-- set_option pp.coercions false +lemma geo_cdf_rec (β : ENNReal) (Hβ1: β ≤ 1) (n : ℕ) : geo_cdf β (n + 1) = β + (1 - β) * geo_cdf β n := by unfold geo_cdf - sorry + /- + suffices ENNReal.toEReal (1 - (1 - β) ^ (n + 1)) = ENNReal.toEReal (β + (1 - β) * (1 - (1 - β) ^ n)) by + apply EReal.coe_ennreal_injective + apply this + -/ + + suffices ENNReal.toReal (1 - (1 - β) ^ (n + 1)) = ENNReal.toReal (β + (1 - β) * (1 - (1 - β) ^ n)) by + apply (ENNReal.toReal_eq_toReal_iff _ _).mp at this + cases this + · trivial + rename_i this + cases this + · rename_i h + rcases h with ⟨ A, B ⟩ + simp_all + exfalso + cases B + · simp_all + · rename_i h + apply ENNReal.mul_eq_top.mp at h + simp_all + · rename_i h + rcases h with ⟨ A, _ ⟩ + simp_all + ring_nf + have C1 : β ≠ ⊤ := by + intro K + simp_all + have C3 : (1 - β) ^ (n + 1) ≤ 1 := by + apply pow_le_one' + apply tsub_le_self + have C4 : (1 - β) ^ n ≤ 1 := by + apply pow_le_one' + apply tsub_le_self + have C2 : (1 - β) * (1 - (1 - β) ^ n) ≠ ⊤ := by + apply ENNReal.mul_ne_top + · apply ENNReal.sub_ne_top + simp + · apply ENNReal.sub_ne_top + simp + rw [ENNReal.toReal_add C2 C1] + rw [ENNReal.toReal_mul] + rw [← pow_succ'] + rw [ENNReal.toReal_sub_of_le C3 (by simp)] + rw [ENNReal.toReal_sub_of_le Hβ1 (by simp)] + rw [ENNReal.toReal_sub_of_le C4 (by simp)] + rw [ENNReal.toReal_pow] + rw [ENNReal.toReal_pow] + rw [ENNReal.toReal_sub_of_le Hβ1 (by simp)] + rw [mul_sub] + simp + rw [pow_succ] + linarith + + +lemma exactClippedSum_append : exactClippedSum i (A ++ B) = exactClippedSum i A + exactClippedSum i B := by + simp [exactClippedSum] + +lemma exactDiffSum_append : exactDiffSum i (A ++ B) = exactDiffSum i A + exactDiffSum i B := by + simp [exactDiffSum] + rw [exactClippedSum_append] + rw [exactClippedSum_append] + linarith + +lemma sv8_sum_append : sv8_sum (A ++ B) vp v0 = sv8_sum A vp v0 + sv8_sum B vp v0 - v0 := by + simp [sv8_sum] + simp [exactDiffSum_append] + linarith + +lemma sv8_sum_comm : sv8_sum (A ++ B) vp v0 = sv8_sum (B ++ A) vp v0 := by + unfold sv8_sum + simp + simp [exactDiffSum, exactClippedSum] + linarith + +lemma sv8_G_comm : sv8_G (A ++ B) vp v0 vf = sv8_G (B ++ A) vp v0 vf := by + revert vp v0 + induction vf + · intros + apply sv8_sum_comm + · intro v0 vp + rename_i next rest IH + unfold sv8_G + rw [sv8_sum_comm] + congr 1 + apply IH + +-- -- IDK +-- lemma sv8_G_cons : sv8_G (x :: L) vp v0 vf = 0 := by +-- revert vp v0 +-- induction vf +-- · intros v0 vp +-- simp [sv8_G] +-- s orry +-- · intro vp v0 +-- simp [sv8_G] +-- s orry +-- -- unfold sv8_G + +lemma exactDiffSum_nonpos : exactDiffSum point L ≤ 0 := by + simp [exactDiffSum, exactClippedSum] + induction L + · simp + · rename_i l ll IH + simp + apply le_trans + · apply add_le_add + · rfl + · apply IH + simp + +lemma exactDiffSum_singleton_le_1 : -1 ≤ exactDiffSum point [v] := by + simp [exactDiffSum, exactClippedSum] + cases (Classical.em (point ≤ v)) + · right + trivial + · left + rename_i h + simp at h + rw [Int.min_def] + simp + split + · linarith + · linarith + + -- There is a value such that sampling at least that value implies the loop definitely terminiates lemma lucky_guess (τ : ℤ) (l : List ℕ) : ∃ (K : ℤ), ∀ A, ∀ (K' : ℤ), K ≤ K' -> exactDiffSum A l + K' ≥ τ := by - sorry + exists (List.length l + τ) + intro A K' HK' + apply ge_iff_le.mpr + apply le_trans _ ?G1 + case G1 => + apply add_le_add + · rfl + · apply HK' + conv => + lhs + rw [<- zero_add τ] + rw [<- add_assoc] + simp + clear HK' + induction l + · simp [exactDiffSum, exactClippedSum] + · rename_i l0 ll IH + simp + rw [<- List.singleton_append] + rw [exactDiffSum_append] + rw [add_comm] + repeat rw [<- add_assoc] + rw [add_comm] + repeat rw [<- add_assoc] + rw [add_assoc] + conv => + lhs + rw [<- add_zero 0] + apply add_le_add + · apply IH + · have H := @exactDiffSum_singleton_le_1 A l0 + linarith lemma ite_conv_left {P : Prop} {D} {a b c : ENNReal} (H : a = c) : @ite _ P D a b = @ite _ P D c b := by split <;> trivial @@ -2005,9 +2162,9 @@ lemma ite_lemma_1 {P : Prop} {D} {f : T -> ENNReal} : ∑'(a : T), @ite _ P D (f · rfl · simp +end equiv - -lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by +lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, (@sv1_privMax PureDPSystem laplace_pureDPSystem ε₁ ε₂ l s) := by simp only [sv1_privMax, bind, pure, bind_apply] -- Push the sum over s inwards conv => @@ -2079,14 +2236,113 @@ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by -- ρ is the probability of the lucky event let ρ : ENNReal := (∑'(a : {t : ℤ // PLucky t}), privNoiseGuess ε₁ ε₂ a.1) - have Hρ_lb : 0 < ρ := by sorry - have Hρ_ub : ρ ≤ 1 := by sorry + have Hρ_1 : (∑'a, privNoiseGuess ε₁ ε₂ a) = 1 := by + cases (privNoiseGuess ε₁ ε₂) + simp [DFunLike.coe, SPMF.instFunLike] + rw [<- Summable.hasSum_iff ENNReal.summable] + trivial + have Hρ_lb : 0 < ρ := by + -- There is at least one lucky element + have HU : PLucky K := by simp [PLucky] + apply LT.lt.trans_le _ ?G2 + case G2 => apply ENNReal.le_tsum ⟨ _, HU ⟩ + simp [privNoiseGuess, privNoiseZero, DPNoise.noise, privNoisedQueryPure, DiscreteLaplaceGenSamplePMF] + simp [DFunLike.coe, PMF.instFunLike] + apply Real.mul_pos + · apply div_pos + · simp + · apply Right.add_pos' + · apply Real.exp_pos + · simp + · apply Real.exp_pos + have Hρ_nz : ρ ≠ 0 := by apply pos_iff_ne_zero.mp Hρ_lb + have Hρ_ub : ρ ≤ 1 := by + rw [<- Hρ_1] + rw [ENNReal.tsum_split PLucky] + simp_all only [ge_iff_le, self_le_add_right, PLucky, ρ] + have Hρ_ub_strict : ρ < 1 := by + rw [<- Hρ_1] + rw [ENNReal.tsum_split PLucky] + conv => + lhs + rw [<- add_zero ρ] + apply ENNReal.add_lt_add_of_le_of_lt + · intro X; simp_all + · rfl + · -- There is at least one unlucky element + have HU : ¬PLucky (K - 1) := by simp [PLucky] + apply LT.lt.trans_le _ ?G2 + case G2 => apply ENNReal.le_tsum ⟨ _, HU ⟩ + simp [privNoiseGuess, privNoiseZero, DPNoise.noise, privNoisedQueryPure, DiscreteLaplaceGenSamplePMF] + simp [DFunLike.coe, PMF.instFunLike] + apply Real.mul_pos + · apply div_pos + · simp + · apply Right.add_pos' + · apply Real.exp_pos + · simp + · apply Real.exp_pos -- Bound the CDF below by the geometric CDF apply @le_trans _ _ _ (⨆(y : ℕ), geo_cdf ρ y) · -- Math - unfold geo_cdf - sorry + apply le_iSup_iff.mpr + intro b H + apply Classical.by_contradiction + intro H1 + simp at H1 + have Hz : (∃ z, (1 - ρ)^z < 1 - b) := by + have W := @exists_pow_lt_of_lt_one NNReal _ _ _ (ENNReal.toNNReal (1 - b)) (ENNReal.toNNReal (1 - ρ)) ?G2 ?G1 + case G2 => + rw [ENNReal.toNNReal_pos_iff] + apply And.intro + · simp + trivial + · apply ENNReal.sub_lt_of_lt_add + · exact le_of_lt H1 + · simp + case G1 => + apply ENNReal.toNNReal_lt_of_lt_coe + simp + apply ENNReal.sub_lt_self + · simp + · simp + · trivial + rcases W with ⟨ N, HN ⟩ + exists N + rw [<- ENNReal.toNNReal_pow] at HN + apply (ENNReal.toNNReal_lt_toNNReal _ _).mp + · trivial + · apply ENNReal.pow_ne_top + apply ENNReal.sub_ne_top + simp + · apply ENNReal.sub_ne_top + simp + rcases Hz with ⟨ z, Hz ⟩ + have Hz' : 1 - (1 - ρ) ^ z > 1 - (1 - b) := by + apply LT.lt.gt + apply (ENNReal.sub_lt_iff_lt_right _ _).mpr + · apply ENNReal.lt_add_of_sub_lt_left + · left + simp + · apply Eq.trans_lt _ Hz + apply ENNReal.sub_sub_cancel + · simp + apply Right.pow_le_one_of_le + apply tsub_le_self + · apply ENNReal.sub_ne_top + simp + · apply tsub_le_self + have H' := H z + have X : 1 - (1 - b) = b := by + apply ENNReal.sub_sub_cancel + · simp + exact le_of_lt H1 + rw [X] at Hz' + apply LE.le.not_lt at H' + apply H' + apply GT.gt.lt + trivial apply iSup_mono intro cut @@ -2098,7 +2354,7 @@ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by simp [probWhileCut, geo_cdf] rename_i cut - rw [geo_cdf_rec] + rw [geo_cdf_rec _ Hρ_ub] rw [ENNReal.tsum_split PLucky] apply add_le_add · -- Lucky guess @@ -2118,8 +2374,15 @@ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by trivial)] rfl -- The rightmost sum is 1 - sorry - + apply @le_trans _ _ _ (∑' (a : { t // PLucky t }), (privNoiseGuess ε₁ ε₂) ↑a * 1) + · simp + apply ENNReal.tsum_le_tsum + intro x + apply ENNReal.mul_left_mono + apply ENNReal.tsum_lb_single 0 + apply ENNReal.tsum_lb_single ([], x.1) + simp [sv1_threshold] + -- Unlucky suffices (∀ H, ∀ a : {t : ℤ // ¬ PLucky t}, geo_cdf ρ cut ≤ ∑' (x : ℕ) (x_1 : sv1_state), @@ -2135,7 +2398,13 @@ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by apply ENNReal.mul_right_mono apply Eq.le -- Math - sorry + clear this + rw [<- Hρ_1] + conv => + enter [1, 1] + rw [ENNReal.tsum_split PLucky] + apply ENNReal.add_sub_cancel_left + exact LT.lt.ne_top Hρ_ub_strict -- Now we have the right inductive structure, I think? induction cut @@ -2146,30 +2415,90 @@ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by rcases a with ⟨ v, Hv ⟩ simp - -- Because the first sample not lucky, we can't say anything about the branch we end up in - -- We will bound it separately for both - have advance (x : sv1_state) : - ((((sv1_privMaxF ε₁ ε₂ (H, v)) >>= (probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (cut + 1))) x) - ≤ (probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (cut + 1 + 1) (H, v)) x) := by + -- Because the first sample is not lucky, we can't say anything about the branch we end up in + -- It may terminate, or it may not. + have advance : + ((∑' (x1 : ℕ) (x2 : sv1_state), + if x1 = sv1_threshold x2 + then (sv1_privMaxF ε₁ ε₂ (H, v)).probBind (fun v => probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (cut + 1) v) x2 + else 0) + ≤ (∑' (x : ℕ) (x_1 : sv1_state), if x = sv1_threshold x_1 then probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (cut + 1 + 1) (H, v) x_1 else 0)) := by conv => rhs + enter [1, x1, 1, x2] unfold probWhileCut unfold probWhileFunctional + simp split · simp · simp - -- Uses the upper bound here, but this is provable - sorry - apply le_trans _ ?G1 - case G1 => - apply ENNReal.tsum_le_tsum - intro x - apply ENNReal.tsum_le_tsum - intro x_1 - apply ite_mono_left - apply advance + -- RHS is 1 + apply ENNReal.tsum_lb_single (List.length H) + apply ENNReal.tsum_lb_single (H, v) + conv => + rhs + simp [sv1_threshold] + + -- LHS is bounded above by 1 by upper bound lemma + have X : + (∑' (x1 : ℕ) (x2 : sv1_state), + if x1 = sv1_threshold x2 then + ∑' (a : sv1_state), + sv1_privMaxF ε₁ ε₂ (H, v) a * probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (cut + 1) a x2 + else 0) = + (∑' (x1 : ℕ) (x2 : sv1_state), + if x1 = sv1_threshold x2 then + ((sv1_privMaxF ε₁ ε₂ (H, v) >>= probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (cut + 1)) x2) + else 0) := by + simp + rw [X] + clear X + rw [ENNReal.tsum_comm] + have X : ∀ b : sv1_state, + (∑' (a : ℕ), + if a = sv1_threshold b then + (sv1_privMaxF ε₁ ε₂ (H, v) >>= probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (cut + 1)) b + else 0) = + ((sv1_privMaxF ε₁ ε₂ (H, v) >>= probWhileCut (sv1_privMaxC τ l) (sv1_privMaxF ε₁ ε₂) (cut + 1)) b) := by + intro b + rw [tsum_ite_eq] + conv => + lhs + enter [1, b] + rw [X b] + clear X + + simp [sv1_privMaxF] + + -- Somehow we need it to unfold this tsum. Weird. + apply le_trans + · apply ENNReal.tsum_le_tsum + intro a + simp + apply ENNReal.tsum_le_tsum + intro a1 + apply ENNReal.mul_left_mono + rfl + + rw [ENNReal.tsum_comm] + conv => + lhs + enter [1, b] + rw [ENNReal.tsum_mul_left] + apply le_trans + · apply ENNReal.tsum_le_tsum + intro x + apply ENNReal.mul_left_mono + apply sv1_loop_ub + simp + rcases (privNoiseGuess ε₁ ε₂) with ⟨ X, Y ⟩ + apply Eq.le + exact HasSum.tsum_eq Y + apply le_trans _ advance + simp clear advance + -- Now we want to commute out the randomness associate to that s1_privMaxF apply le_trans _ ?G1 case G1 => @@ -2177,7 +2506,6 @@ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by intro x apply ENNReal.tsum_le_tsum intro x_1 - simp rw [<- ite_lemma_1] conv => enter [2] @@ -2197,7 +2525,7 @@ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by rw [ENNReal.tsum_split PLucky] -- Split the sum and the recurrence relation - rw [geo_cdf_rec] + rw [geo_cdf_rec _ Hρ_ub] apply add_le_add · -- Guess is lucky -- The loop will terminate and we can show it @@ -2233,15 +2561,17 @@ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, sv1_privMax ε₁ ε₂ l s := by -- Conclude by simplification simp only [sv1_privMaxF, bind, pure, bind_apply, pure_apply, mul_ite, mul_one, mul_zero] - sorry + + + /-- sv9 normalizes because sv1 normalizes -/ def sv9_privMax_SPMF (ε₁ ε₂ : ℕ+) (l : List ℕ) : SPMF ℕ := - ⟨ sv9_privMax ε₁ ε₂ l, + ⟨ @sv9_privMax PureDPSystem laplace_pureDPSystem ε₁ ε₂ l, by rw [<- sv8_sv9_eq] rw [<- sv7_sv8_eq] From 24e345a10c40adb863dd24da54f7af71156bccac Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Sat, 9 Nov 2024 17:49:29 -0500 Subject: [PATCH 099/100] checkpoint --- .../Queries/UnboundedMax/Properties.lean | 90 ++++++++++++++++++- 1 file changed, 89 insertions(+), 1 deletion(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 729040d2..8ddcb9f3 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -2524,12 +2524,99 @@ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, (@sv1_privMax PureDPSystem laplace_pur -- Now, condition on the luckiness of the next value rw [ENNReal.tsum_split PLucky] + -- Split the sum and the recurrence relation rw [geo_cdf_rec _ Hρ_ub] apply add_le_add · -- Guess is lucky -- The loop will terminate and we can show it - sorry + + conv => + enter [2, 1, a, 1, b] + conv => + enter [1, c] + conv => + enter [1, d] + rw [<- mul_ite_zero] + rw [ENNReal.tsum_mul_left] + rw [ENNReal.tsum_mul_left] + + -- Conclude by simplification + simp only [sv1_privMaxF, bind, pure, bind_apply, pure_apply, mul_ite, mul_one, mul_zero] + unfold probWhileCut + unfold probWhileFunctional + -- Push the ite inside so that all of the sums are in a row + conv => + enter [2, 1, a, 1, b] + rw [<- ENNReal.tsum_mul_right] + simp + enter [1, i] + rw [<- mul_ite_zero] + enter [2] + rw [<- ite_lemma_1] + conv => + enter [1, x] + rw [<- ite_lemma_1] + conv => + enter [2, 1, a, 1, b, 1, i] + rw [<- ENNReal.tsum_mul_left] + enter [1, c] + rw [<- ENNReal.tsum_mul_left] + -- Move the lucky sample inwards + conv => + rhs + rw [ENNReal.tsum_comm] + enter [1, b] + rw [ENNReal.tsum_comm] + enter [1, c] + rw [ENNReal.tsum_comm] + enter [1, d] + rw [ENNReal.tsum_comm] + -- Pick elements for each of the other sums to make it terminate + apply ENNReal.tsum_lb_single (H ++ [v]) + rw [ENNReal.tsum_comm] + apply ENNReal.tsum_lb_single (List.length H + 1) + rw [ENNReal.tsum_comm] + rw [ENNReal.tsum_prod'] + apply ENNReal.tsum_lb_single (H ++ [v]) + apply le_trans _ ?G1 + case G1 => + apply ENNReal.tsum_le_tsum + intro b + apply ENNReal.tsum_le_tsum + intro a + apply ENNReal.tsum_le_tsum + intro c + apply ENNReal.mul_left_mono + simp [sv1_threshold] + simp [sv1_privMaxC] + apply ite_mono_left + rw [ite_eq_right_iff.mpr (by + intro K + exfalso + rcases c with ⟨ c, Hc ⟩ + simp [sv1_threshold, sv1_noise] at K + have HC' := HLucky c (H.length + 1) Hc + apply (LT.lt.not_le K) + apply GE.ge.le + apply HC')] + + -- Now move the lucky sum out to the front, so that we can constrain the other sum values to equal it + conv => + rhs + enter [1, a] + rw [ENNReal.tsum_comm] + rw [ENNReal.tsum_comm] + + apply le_trans _ ?G1 + case G1 => + apply ENNReal.tsum_le_tsum + intro a + apply ENNReal.tsum_lb_single a.1 + apply ENNReal.tsum_lb_single a.1 + rfl + simp + · -- Guess is unlucky -- Commute out the samples related to the first sample (which will evenetually become a (1- ρ) factor) conv => @@ -2561,6 +2648,7 @@ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, (@sv1_privMax PureDPSystem laplace_pur -- Conclude by simplification simp only [sv1_privMaxF, bind, pure, bind_apply, pure_apply, mul_ite, mul_one, mul_zero] + sorry From 4948e2442e68807e8e2e5645c2a479db42a32d82 Mon Sep 17 00:00:00 2001 From: Markus de Medeiros Date: Sat, 9 Nov 2024 17:59:51 -0500 Subject: [PATCH 100/100] victory --- .../Queries/UnboundedMax/Code.lean | 56 +------------------ .../Queries/UnboundedMax/Properties.lean | 23 ++++++-- Test.lean | 10 ++-- 3 files changed, 25 insertions(+), 64 deletions(-) diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean index 05225428..f22e6ccf 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Code.lean @@ -59,7 +59,8 @@ def privNoiseThresh (ε₁ ε₂ : ℕ+) : SPMF ℤ := privNoiseZero ε₁ (2 * /- ## Program version 0 - Executable - - Tracks single state + - Optimization of sv1: Tracks single state + (Needs sv0 sv1 equivalence proof) -/ def sv0_state : Type := ℕ × ℤ @@ -82,11 +83,6 @@ def sv0_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do let sk <- probWhile (sv0_privMaxC τ l) (sv0_privMaxF ε₁ ε₂) (0, v0) return (sv0_threshold sk) -def sv0_privMax_PMF (ε₁ ε₂ : ℕ+) (l : List ℕ) : SPMF ℕ := - ⟨ sv0_privMax ε₁ ε₂ l, - by - sorry ⟩ - /- ## Program version 1 - Executable @@ -113,52 +109,4 @@ def sv1_privMax (ε₁ ε₂ : ℕ+) (l : List ℕ) : SLang ℕ := do return (sv1_threshold sk) -def sv1_privMax_PMF (ε₁ ε₂ : ℕ+) (l : List ℕ) : SPMF ℕ := - ⟨ sv1_privMax ε₁ ε₂ l, - by - -- Difficult - sorry ⟩ - end SLang - -/- -lemma tsum_iSup_comm (f : T -> U -> ENNReal) : ∑' (a : T), ⨆ (b : U), f a b = ⨆ (b : U), ∑' (a : T), f a b := by - - sorry - --- A version of probWhile that proves normalization, when its arguments are all normal -def SPMF_while (cond : T → Bool) (body : T → SPMF T) (init : T) : SPMF T := - ⟨ probWhile cond (fun x => body x) init, - by - apply (Summable.hasSum_iff ENNReal.summable).mpr - unfold probWhile - rw [tsum_iSup_comm] - conv => - enter [1, 1, cut] - -- We can establish the ≤ direction easily - -- ≥ is not true - - - - -- apply iSup_eq_of_tendsto - -- · sorry - -- sorry - - -- rw [ENNReal.tsum_eq_iSup_sum] - -- apply eq_iff_le_not_lt.mpr - -- apply And.intro - -- · apply iSup_le_iff.mpr - -- sorry - -- · simp - -- apply le_iSup_iff.mpr - -- intro b H - -- sorry - ⟩ - -lemma SPMF_while_eq (cond : T → Bool) (body : T → SPMF T) (init : T) : - ((SPMF_while cond body init) : SLang T) = (probWhile cond (fun x => body x) init) := by - simp [SPMF_while] - - - --/ diff --git a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean index 8ddcb9f3..6b024354 100644 --- a/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean +++ b/SampCert/DifferentialPrivacy/Queries/UnboundedMax/Properties.lean @@ -2647,13 +2647,26 @@ lemma sv1_lb ε₁ ε₂ l : 1 ≤ ∑'s, (@sv1_privMax PureDPSystem laplace_pur -- Conclude by simplification simp only [sv1_privMaxF, bind, pure, bind_apply, pure_apply, mul_ite, mul_one, mul_zero] + apply le_trans _ ?G1 + case G1 => + apply ENNReal.tsum_le_tsum + intro a + apply ENNReal.tsum_lb_single (H ++ [v]) + apply ENNReal.tsum_lb_single a.1 + rfl + simp + rw [<- Hρ_1] + rw [ENNReal.tsum_split PLucky] + rw [add_comm] - sorry - - - - +def sv1_privMax_PMF (ε₁ ε₂ : ℕ+) (l : List ℕ) : SPMF ℕ := + ⟨ sv1_privMax ε₁ ε₂ l, + by + rw [Summable.hasSum_iff ENNReal.summable] + apply LE.le.antisymm + · apply sv1_ub + · apply sv1_lb ⟩ /-- sv9 normalizes because sv1 normalizes diff --git a/Test.lean b/Test.lean index 33bce2e3..0635a901 100644 --- a/Test.lean +++ b/Test.lean @@ -257,13 +257,13 @@ def sparseVector_tests : IO Unit := do IO.println s!"data := {samples} uniform samples of [0, {(unif_ub : ℕ)}): {(data.take 20)}..." IO.println "" - for i in [:num_trials] do - let ct <- run <| @sv0_privMax_PMF PureDPSystem laplace_pureDPSystem num den data - IO.println s!"#{i} sv0 max: {ct}" - IO.println "" + -- for i in [:num_trials] do + -- let ct <- run <| @sv0_privMax_PMF PureDPSystem laplace_pureDPSystem num den data + -- IO.println s!"#{i} sv0 max: {ct}" + -- IO.println "" for i in [:num_trials] do - let ct <- run <| @sv1_privMax_PMF PureDPSystem laplace_pureDPSystem num den data + let ct <- run <| sv1_privMax_PMF num den data IO.println s!"#{i} sv1 max: {ct}" IO.println ""