From 351d7e8b2374772f32a6af19eb73aa5c81fcef7e Mon Sep 17 00:00:00 2001 From: Hei Li Date: Tue, 17 Sep 2024 13:21:02 +0200 Subject: [PATCH 01/69] Progres with flip_spec --- theories/coneris/lib/flip_spec.v | 222 +++++++++++++++++++++++++++++++ theories/coneris/lib/hocap.v | 66 ++++----- theories/prelude/stdpp_ext.v | 16 +++ 3 files changed, 271 insertions(+), 33 deletions(-) create mode 100644 theories/coneris/lib/flip_spec.v diff --git a/theories/coneris/lib/flip_spec.v b/theories/coneris/lib/flip_spec.v new file mode 100644 index 00000000..58e2c96d --- /dev/null +++ b/theories/coneris/lib/flip_spec.v @@ -0,0 +1,222 @@ +From clutch.coneris Require Import coneris. +From clutch.coneris Require Import flip hocap. + +Set Default Proof Using "Type*". + +Class flip_spec `{!conerisGS Σ} := FlipSpec +{ + (** * Operations *) + allocate_tape : val; + flip_tape : val; + (** * Ghost state *) + (** The assumptions about [Σ] *) + flipG : gFunctors → Type; + (** [name] is used to associate [locked] with [is_lock] *) + error_name : Type; + tape_name: Type; + (** * Predicates *) + is_flip {L : flipG Σ} (N:namespace) + (γ1: error_name) (γ2: tape_name): iProp Σ; + flip_error_auth {L : flipG Σ} (γ: error_name) (x:R): iProp Σ; + flip_error_frag {L : flipG Σ} (γ: error_name) (x:R): iProp Σ; + flip_tapes_auth {L : flipG Σ} (γ: tape_name) (m:gmap loc (list bool)): iProp Σ; + flip_tapes_frag {L : flipG Σ} (γ: tape_name) (α:loc) (ns:list bool): iProp Σ; + (** * General properties of the predicates *) + #[global] is_flip_persistent {L : flipG Σ} N γ1 γ2 :: + Persistent (is_flip (L:=L) N γ1 γ2); + #[global] flip_error_auth_timeless {L : flipG Σ} γ x :: + Timeless (flip_error_auth (L:=L) γ x); + #[global] flip_error_frag_timeless {L : flipG Σ} γ x :: + Timeless (flip_error_frag (L:=L) γ x); + #[global] flip_tapes_auth_timeless {L : flipG Σ} γ m :: + Timeless (flip_tapes_auth (L:=L) γ m); + #[global] flip_tapes_frag_timeless {L : flipG Σ} γ α ns :: + Timeless (flip_tapes_frag (L:=L) γ α ns); + + flip_error_auth_exclusive {L : flipG Σ} γ x1 x2: + flip_error_auth (L:=L) γ x1 -∗ flip_error_auth (L:=L) γ x2 -∗ False; + flip_error_frag_exclusive {L : flipG Σ} γ x1 x2: + flip_error_frag (L:=L) γ x1 -∗ flip_error_frag (L:=L) γ x2 -∗ False; + flip_error_auth_pos {L : flipG Σ} γ x: + flip_error_auth (L:=L) γ x -∗ ⌜(0<=x)%R⌝; + flip_error_auth_frag {L : flipG Σ} γ x: + flip_error_frag (L:=L) γ x -∗ ⌜(0<=x)%R⌝; + flip_error_agree {L : flipG Σ} γ x1 x2: + flip_error_auth (L:=L) γ x1 -∗ flip_error_frag (L:=L) γ x2 -∗ ⌜x1 = x2 ⌝; + flip_error_update {L : flipG Σ} γ x1 x2 (x3:nonnegreal): + flip_error_auth (L:=L) γ x1 -∗ flip_error_frag (L:=L) γ x2 ==∗ + flip_error_auth (L:=L) γ x3 ∗ flip_error_frag (L:=L) γ x3; + + flip_tapes_auth_exclusive {L : flipG Σ} γ m m': + flip_tapes_auth (L:=L) γ m -∗ flip_tapes_auth (L:=L) γ m' -∗ False; + flip_tapes_frag_exclusive {L : flipG Σ} γ α ns ns': + flip_tapes_frag (L:=L) γ α ns -∗ flip_tapes_frag (L:=L) γ α ns' -∗ False; + flip_tapes_agree {L : flipG Σ} γ α m ns: + flip_tapes_auth (L:=L) γ m -∗ flip_tapes_frag (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; + flip_tapes_update {L : flipG Σ} γ α m ns n: + flip_tapes_auth (L:=L) γ m -∗ flip_tapes_frag (L:=L) γ α ns ==∗ + flip_tapes_auth (L:=L) γ (<[α := (ns ++[n])]> m) ∗ flip_tapes_frag (L:=L) γ α (ns ++ [n]); + + (** * Program specs *) + flip_inv_create_spec {L : flipG Σ} N E ε: + ↯ ε -∗ + wp_update E (∃ γ1 γ2, is_flip (L:=L) N γ1 γ2 ∗ + flip_error_frag (L:=L) γ1 ε); + + flip_allocate_tape_spec {L: flipG Σ} N E γ1 γ2: + ↑N ⊆ E-> + {{{ is_flip (L:=L) N γ1 γ2 }}} + allocate_tape #() @ E + {{{ (v:val), RET v; + ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ flip_tapes_frag (L:=L) γ2 α [] + }}}; + flip_tape_spec_some {L: flipG Σ} N E γ1 γ2 (P: iProp Σ) (Q:iProp Σ) (α:loc) (n:bool) ns: + ↑N⊆E -> + {{{ is_flip (L:=L) N γ1 γ2 ∗ + □ (P ={E∖↑N}=∗ Q) ∗ + P ∗ flip_tapes_frag (L:=L) γ2 α (n::ns) + }}} + flip_tape #lbl:α @ E + {{{ RET #n; Q ∗ flip_tapes_frag (L:=L) γ2 α ns}}}; + flip_presample_spec {L: flipG Σ} NS E ns α + (ε2 : R -> bool -> R) + (P : iProp Σ) T γ1 γ2: + ↑NS ⊆ E -> + (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> + (∀ (ε:R), 0<= ε -> (ε2 ε true + ε2 ε false)/2 <= ε)%R-> + is_flip (L:=L) NS γ1 γ2 -∗ + (□∀ (ε:R) n, (P ∗ flip_error_auth (L:=L) γ1 ε) ={E∖↑NS}=∗ + (⌜(1<=ε2 ε n)%R⌝ ∨ (flip_error_auth (L:=L) γ1 (ε2 ε n) ∗ T (n)))) + -∗ + P -∗ flip_tapes_frag (L:=L) γ2 α ns-∗ + wp_update E (∃ n, T (n) ∗ flip_tapes_frag (L:=L) γ2 α (ns++[n])) +}. + + +Section lemmas. + Context `{rc:flip_spec} {L: flipG Σ}. + + Lemma flip_tape_spec_none N E γ1 γ2 (ε2:R -> bool -> R) (P: iProp Σ) (T: bool -> iProp Σ) (Q: bool -> iProp Σ)(α:loc): + ↑N ⊆ E-> + (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> + (∀ (ε:R), 0<=ε -> ((ε2 ε true) + (ε2 ε false))/2 <= ε)%R → + {{{ is_flip (L:=L) N γ1 γ2 ∗ + □(∀ (ε:R) (n : bool), P ∗ flip_error_auth (L:=L) γ1 ε + ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨ flip_error_auth (L:=L) γ1 (ε2 ε n) ∗ T n) ) ∗ + □ (∀ (n:bool), T n ={E∖↑N}=∗ Q n) ∗ + P ∗ flip_tapes_frag (L:=L) γ2 α [] + }}} + flip_tape #lbl:α @ E + {{{ (b:bool), RET #b; Q b ∗ flip_tapes_frag (L:=L) γ2 α [] }}}. + Proof. + iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP & Hα) HΦ". + iMod (flip_presample_spec with "[//][//][$][$]") as "(%&HT&Hα)"; [done..|]. + iApply (flip_tape_spec_some _ _ _ _ (T n) with "[$Hα $HT]"); try done. + { by iSplit. } + by iNext. + Qed. + +End lemmas. + + +(** Instantiate flip *) +Class flipG1 Σ := FlipG1 { flipG1_error::hocap_errorGS Σ; + flipG1_tapes:: hocap_tapesGS Σ; + }. +Local Definition flip_inv_pred1 `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapesGS Σ} γ1 γ2:= + (∃ (ε:R) (m:gmap loc (nat * list nat)) , + ↯ ε ∗ ●↯ ε @ γ1 ∗ + ([∗ map] α ↦ t ∈ m, α ↪N ( t.1 ; t.2 )) ∗ ●m@γ2 + )%I. + +#[local] Program Definition flip_spec1 `{!conerisGS Σ}: flip_spec := + {| allocate_tape:= (λ: <>, allocB); + flip_tape:= flipL; + flipG := flipG1; + error_name := gname; + tape_name := gname; + is_flip _ N γ1 γ2 := inv N (flip_inv_pred1 γ1 γ2); + flip_error_auth _ γ x := ●↯ x @ γ; + flip_error_frag _ γ x := ◯↯ x @ γ; + flip_tapes_auth _ γ m := (●((λ ns, (1, bool_to_nat<$>ns))<$>m)@γ)%I; + flip_tapes_frag _ γ α ns := (α◯↪N (1%nat;bool_to_nat<$> ns) @ γ)%I; + |}. + +Next Obligation. + simpl. + iIntros (??????) "(%&<-&H1)(%&<-&H2)". + iCombine "H1 H2" gives "%H". by rewrite excl_auth.excl_auth_auth_op_valid in H. +Qed. +Next Obligation. + simpl. + iIntros (??????) "(%&<-&H1)(%&<-&H2)". + iCombine "H1 H2" gives "%H". by rewrite excl_auth.excl_auth_frag_op_valid in H. +Qed. +Next Obligation. + simpl. + iIntros (?????) "H". + iApply (hocap_error_auth_pos with "[$]"). +Qed. +Next Obligation. + simpl. + iIntros (?????) "H". + iApply (hocap_error_frag_pos with "[$]"). +Qed. +Next Obligation. + simpl. + iIntros (??????) "H1 H2". + iApply (hocap_error_agree with "[$][$]"). +Qed. +Next Obligation. + simpl. iIntros (???????) "??". + iApply (hocap_error_update with "[$][$]"). +Qed. +Next Obligation. + simpl. + iIntros (??????) "H1 H2". + by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[%H _]". +Qed. +Next Obligation. + simpl. + iIntros (???????) "H1 H2". + iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. + rewrite dfrac_op_own dfrac_valid_own. by intro. +Qed. +Next Obligation. + simpl. + iIntros. + iDestruct (hocap_tapes_agree with "[$][$]") as "%H". + iPureIntro. + rewrite lookup_fmap_Some in H. destruct H as (?&?&K). + simplify_eq. + rewrite K. + f_equal. + eapply fmap_inj; last done. + intros [][]?; by simplify_eq. +Qed. +Next Obligation. + iIntros. + iMod (hocap_tapes_presample with "[$][$]") as "[??]". + iModIntro. iFrame. + rewrite fmap_insert. + rewrite fmap_app; iFrame. +Qed. +Next Obligation. + simpl. + iIntros (????? ε) "H1". + iApply fupd_wp_update. + iApply wp_update_ret. + iDestruct (ec_valid with "[$]") as "%". + unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "(%γ1 & H2 & H3)"; first naive_solver. + simpl. + iMod (hocap_tapes_alloc ∅) as "(%γ2 & H4 & H5)". + iMod (inv_alloc _ _ (flip_inv_pred1 γ1 γ2) with "[H1 H2 H4]"). + { iFrame. by iNext. } + by iFrame. +Qed. +Next Obligation. + Admitted. +Next Obligation. +Admitted. +Next Obligation. +Admitted. diff --git a/theories/coneris/lib/hocap.v b/theories/coneris/lib/hocap.v index 3b56af31..e77a641c 100644 --- a/theories/coneris/lib/hocap.v +++ b/theories/coneris/lib/hocap.v @@ -195,39 +195,39 @@ Section HOCAP. intros n; inv_fin n. Qed. - (** With tapes *) - Context `{!hocap_tapesGS Σ}. - - Definition tapes_inv (γ :gname):= - inv hocap_tapes_nroot (∃ m, ●m@γ ∗ [∗ map] α ↦ t ∈ m, α ↪N ( t.1 ; t.2 ) ). - Lemma wp_hocap_presample_adv_comp (N : nat) z E ns α - (ε2 : R -> fin (S N) -> R) - (P : iProp Σ) (Q : val-> iProp Σ) T γ γ': - TCEq N (Z.to_nat z) → - ↑hocap_error_nroot ⊆ E -> - ↑hocap_tapes_nroot ⊆ E -> - (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> - (∀ (ε:R), 0<= ε -> SeriesC (λ n, (1 / (S N)) * (ε2 ε n))%R <= ε)%R → - error_inv γ -∗ tapes_inv γ' -∗ - (□∀ (ε:R) m, (⌜m!!α = Some (N, ns)⌝ ∗ P ∗ ●↯ ε @ γ ∗ ●m@γ') - ={E∖↑hocap_error_nroot∖↑hocap_tapes_nroot}=∗ - ∃ n, (⌜(1<=ε2 ε n)%R⌝ ∨(●↯ (ε2 ε n) @ γ ∗ ●(<[α := (N, ns ++ [fin_to_nat n])]>m) @ γ' ∗ T (n)))) - -∗ P -∗ α ◯↪N (N; ns) @ γ' -∗ - wp_update E (∃ n, T (n) ∗ α◯↪N (N; ns++[fin_to_nat n]) @ γ'). - Proof. - iIntros (-> Hval Hsubset Hubset' Hineq) "#Hinv #Hinv' #Hshift HP Htape". - iApply fupd_wp_update. - iInv "Hinv" as ">(%ε' & Hε & Hauth)" "Hclose". - iInv "Hinv'" as ">(%m & Hm & Hmauth)" "Hclose'". - iDestruct (ec_valid with "[$]") as "%". - iApply wp_update_mono_fupd. iApply fupd_frame_l. - iSplitL "Hε". - - iApply (wp_update_presample_exp _ α _ ns ε' (ε2 ε) with "[$Hε]"). - + naive_solver. - + naive_solver. - + admit. - - iMod ("Hclose'" with "[$]"). - Abort. + (* (** With tapes *) *) + (* Context `{!hocap_tapesGS Σ}. *) + + (* Definition tapes_inv (γ :gname):= *) + (* inv hocap_tapes_nroot (∃ m, ●m@γ ∗ [∗ map] α ↦ t ∈ m, α ↪N ( t.1 ; t.2 ) ). *) + (* Lemma wp_hocap_presample_adv_comp (N : nat) z E ns α *) + (* (ε2 : R -> fin (S N) -> R) *) + (* (P : iProp Σ) (Q : val-> iProp Σ) T γ γ': *) + (* TCEq N (Z.to_nat z) → *) + (* ↑hocap_error_nroot ⊆ E -> *) + (* ↑hocap_tapes_nroot ⊆ E -> *) + (* (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> *) + (* (∀ (ε:R), 0<= ε -> SeriesC (λ n, (1 / (S N)) * (ε2 ε n))%R <= ε)%R → *) + (* error_inv γ -∗ tapes_inv γ' -∗ *) + (* (□∀ (ε:R) m, (⌜m!!α = Some (N, ns)⌝ ∗ P ∗ ●↯ ε @ γ ∗ ●m@γ') *) + (* ={E∖↑hocap_error_nroot∖↑hocap_tapes_nroot}=∗ *) + (* ∃ n, (⌜(1<=ε2 ε n)%R⌝ ∨(●↯ (ε2 ε n) @ γ ∗ ●(<[α := (N, ns ++ [fin_to_nat n])]>m) @ γ' ∗ T (n)))) *) + (* -∗ P -∗ α ◯↪N (N; ns) @ γ' -∗ *) + (* wp_update E (∃ n, T (n) ∗ α◯↪N (N; ns++[fin_to_nat n]) @ γ'). *) + (* Proof. *) + (* iIntros (-> Hval Hsubset Hubset' Hineq) "#Hinv #Hinv' #Hshift HP Htape". *) + (* iApply fupd_wp_update. *) + (* iInv "Hinv" as ">(%ε' & Hε & Hauth)" "Hclose". *) + (* iInv "Hinv'" as ">(%m & Hm & Hmauth)" "Hclose'". *) + (* iDestruct (ec_valid with "[$]") as "%". *) + (* iApply wp_update_mono_fupd. iApply fupd_frame_l. *) + (* iSplitL "Hε". *) + (* - iApply (wp_update_presample_exp _ α _ ns ε' (ε2 ε) with "[$Hε]"). *) + (* + naive_solver. *) + (* + naive_solver. *) + (* + admit. *) + (* - iMod ("Hclose'" with "[$]"). *) + (* Abort. *) (* iApply (wp_presample_adv_comp); [done|exact|..]. *) (* iApply fupd_frame_l. *) diff --git a/theories/prelude/stdpp_ext.v b/theories/prelude/stdpp_ext.v index e8cf0184..a11247ae 100644 --- a/theories/prelude/stdpp_ext.v +++ b/theories/prelude/stdpp_ext.v @@ -371,6 +371,22 @@ Section list. a ∈ l1 ∧ b ∈ l2 → (a, b) ∈ list_prod l1 l2. Proof. apply elem_of_list_prod. Qed. + Lemma fmap_inj (f:A -> B) (l1 l2: list A): + Inj (=) (=) f -> f<$>l1=f<$>l2 -> l1 = l2. + Proof. + intros H. + revert l2. + induction l1. + - intros ? K. simpl in *. + by erewrite fmap_nil_inv. + - intros l2 K. + rewrite fmap_cons in K. + destruct l2. + + rewrite fmap_nil in K. simplify_eq. + + rewrite fmap_cons in K. + by simplify_eq. + Qed. + End list. Tactic Notation "case_match" "in" ident(H) "eqn" ":" ident(Hd) := From 4396686b9aa0dd4d7657cf63caf5076047a5d857 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Tue, 17 Sep 2024 13:47:52 +0200 Subject: [PATCH 02/69] Progress with flip_spec --- theories/coneris/lib/flip_spec.v | 48 ++++++++++++++++++++++++++++++-- 1 file changed, 46 insertions(+), 2 deletions(-) diff --git a/theories/coneris/lib/flip_spec.v b/theories/coneris/lib/flip_spec.v index 58e2c96d..d883a5c8 100644 --- a/theories/coneris/lib/flip_spec.v +++ b/theories/coneris/lib/flip_spec.v @@ -215,8 +215,52 @@ Next Obligation. by iFrame. Qed. Next Obligation. - Admitted. + simpl. + iIntros (???????? Φ) "#Hinv HΦ". + wp_pures. + rewrite /allocB. + iInv "Hinv" as "(%&%&?&?&?&?)" "Hclose". + wp_apply (wp_alloc_tape); first done. + iIntros (α) "Hα". + iDestruct (hocap_tapes_notin with "[$][$]") as "%". + iMod (hocap_tapes_new _ _ α 1%nat [] with "[$]") as "[?H]"; first done. + iMod ("Hclose" with "[-H HΦ]"). + { iFrame. iModIntro. + rewrite big_sepM_insert; [iFrame|done]. + } + iApply "HΦ". + by iFrame. +Qed. Next Obligation. -Admitted. + simpl. + iIntros (????????????? Φ) "(#Hinv & #Hvs & HP & Hfrag) HΦ". + rewrite /flipL. + wp_pures. + wp_bind (rand(_) _)%E. + iInv "Hinv" as ">(%&%&?&?&H3&?)" "Hclose". + iDestruct (hocap_tapes_agree with "[$][$]") as "%". + erewrite <-(insert_delete m) at 1; last done. + rewrite big_sepM_insert; last apply lookup_delete. + iDestruct "H3" as "[Htape H3]". + simpl. + wp_apply (wp_rand_tape with "[$]") as "[Htape %]". + iMod ("Hvs" with "[$]") as "HQ". + iMod (hocap_tapes_pop with "[$][$]") as "[H4 Hfrag]". + iMod ("Hclose" with "[-Hfrag HΦ HQ]") as "_". + { iFrame. + rewrite <-(insert_delete m) at 2; last done. + iNext. + rewrite insert_insert. + rewrite big_sepM_insert; last apply lookup_delete. iFrame. + } + iModIntro. + wp_apply conversion.wp_int_to_bool as "_"; first done. + replace (Z_to_bool _) with n. + - iApply "HΦ"; iFrame. + - destruct n; simpl. + + rewrite Z_to_bool_neq_0; [done|lia]. + + by rewrite Z_to_bool_eq_0. +Qed. Next Obligation. + simpl. Admitted. From b13e39896a876c2bb1d39d814a74ca96815e10d3 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Tue, 17 Sep 2024 14:04:32 +0200 Subject: [PATCH 03/69] hocap flip spec --- theories/coneris/lib/flip_spec.v | 51 +++++++++++++++++++++++++++++++- 1 file changed, 50 insertions(+), 1 deletion(-) diff --git a/theories/coneris/lib/flip_spec.v b/theories/coneris/lib/flip_spec.v index d883a5c8..20e6f7eb 100644 --- a/theories/coneris/lib/flip_spec.v +++ b/theories/coneris/lib/flip_spec.v @@ -263,4 +263,53 @@ Next Obligation. Qed. Next Obligation. simpl. -Admitted. + iIntros (???????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". + iApply wp_update_state_step_coupl. + iIntros (σ ε) "((Hheap&Htapes)&Hε)". + iMod (inv_acc with "Hinv") as "[>(% & % & H1 & H2 & H3 & H4 ) Hclose]"; [done|]. + iDestruct (hocap_tapes_agree with "[$][$]") as "%". + erewrite <-(insert_delete m) at 1; last done. + rewrite big_sepM_insert; last apply lookup_delete. + simpl. + iDestruct "H3" as "[Htape H3]". + iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". + iDestruct (ec_supply_bound with "[$][$]") as "%". + iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. + iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". + iApply state_step_coupl_state_adv_comp_con_prob_lang; first done. + unshelve iExists (λ x, mknonnegreal (ε2 ε1' (nat_to_bool (fin_to_nat x))) _). + { apply Hpos. apply cond_nonneg. } + iSplit. + { iPureIntro. + simpl. + unshelve epose proof (Hineq ε1' _) as H'; first apply cond_nonneg. + rewrite SeriesC_finite_foldr/=. + rewrite nat_to_bool_eq_0 nat_to_bool_neq_0; last lia. + simpl in *. lra. + } + iIntros (sample). + destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' (nat_to_bool (fin_to_nat sample))))%R 1%R) as [Hdec|Hdec]; last first. + { apply Rnot_lt_ge, Rge_le in Hdec. + iApply state_step_coupl_ret_err_ge_1. + simpl. simpl in *. lra. + } + iApply state_step_coupl_ret. + unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' (nat_to_bool (fin_to_nat sample))) _) with "Hε_supply") as "[Hε_supply Hε]". + { apply Hpos. apply cond_nonneg. } + { simpl. done. } + simpl. + iMod (tapeN_update_append _ _ _ _ sample with "[$][Htape]") as "[Htapes Htape]". + { by erewrite Heq. } + iMod (hocap_tapes_presample _ _ _ _ _ (fin_to_nat sample) with "[$][$]") as "[H4 Hfrag]". + iMod "Hclose'" as "_". + iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". + { iExFalso. iApply (ec_contradict with "[$]"). exact. } + iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4]") as "_". + { iNext. + rewrite big_sepM_insert_delete Heq/=; iFrame. + } + iApply fupd_mask_intro_subseteq; first set_solver. + iFrame. + rewrite fmap_app/= nat_to_bool_to_nat; first done. + pose proof fin_to_nat_lt sample. lia. +Qed. From eb5024bb2a87a9c1de06292c5cf700da2f6e21d2 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Tue, 17 Sep 2024 14:47:29 +0200 Subject: [PATCH 04/69] Basic template for impl2 with flip --- .../coneris/examples/random_counter/impl2.v | 488 +++++++++--------- theories/coneris/lib/flip_spec.v | 24 +- 2 files changed, 262 insertions(+), 250 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index 86efae3d..20ccfad8 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -1,6 +1,6 @@ From iris.algebra Require Import frac_auth. From iris.base_logic.lib Require Import invariants. -From clutch.coneris Require Import coneris hocap random_counter. +From clutch.coneris Require Import coneris hocap random_counter flip_spec. From clutch Require Import uniform_list. Set Default Proof Using "Type*". @@ -78,29 +78,36 @@ Section tapes_lemmas. End tapes_lemmas. Section impl2. - + Context `{F: flip_spec Σ}. Definition new_counter2 : val:= λ: "_", ref #0. - Definition incr_counter2 : val := λ: "l", let: "n" := rand #1 in - let: "n'" := rand #1 in - let: "x" := #2 * "n" + "n'" in - (FAA "l" "x", "x"). + (* Definition incr_counter2 : val := λ: "l", let: "n" := rand #1 in *) + (* let: "n'" := rand #1 in *) + (* let: "x" := #2 * "n" + "n'" in *) + (* (FAA "l" "x", "x"). *) Definition allocate_tape2 : val := λ: "_", AllocTape #1. - Definition incr_counter_tape2 :val := λ: "l" "α", let: "n" := rand("α") #1 in - let: "n'" := rand("α") #1 in + Definition incr_counter_tape2 :val := λ: "l" "α", let: "n" := + conversion.bool_to_int (flip_allocate_tape "α") + in + let: "n'" := + conversion.bool_to_int (flip_allocate_tape "α") + in let: "x" := #2 * "n" + "n'" in (FAA "l" "x", "x"). Definition read_counter2 : val := λ: "l", !"l". Class counterG2 Σ := CounterG2 { counterG2_error::hocap_errorGS Σ; counterG2_tapes:: hocap_tapesGS' Σ; - counterG2_frac_authR:: inG Σ (frac_authR natR) }. + counterG2_frac_authR:: inG Σ (frac_authR natR); + counterG2_flipG:: flipG Σ + }. - Context `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapesGS' Σ, !inG Σ (frac_authR natR)}. + Context `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapesGS' Σ, !inG Σ (frac_authR natR), L:flipG Σ}. Definition counter_inv_pred2 (c:val) γ1 γ2 γ3:= - (∃ (ε:R) (m:gmap loc (bool*list nat)) (l:loc) (z:nat), - ↯ ε ∗ ●↯ ε @ γ1 ∗ - ([∗ map] α ↦ t ∈ m, α ↪N ( 1%nat ; if t.1:bool then expander t.2 else drop 1%nat (expander t.2)) ) + (∃ (ε:R) (m:gmap loc (bool*list nat)) (l:loc) (z:nat) γ1' γ2' N', + is_flip (L:=L) N' γ1' γ2' ∗ + flip_error_frag (L:=L) γ1' ε ∗ ●↯ ε @ γ1 ∗ + ([∗ map] α ↦ t ∈ m, flip_tapes_frag (L:=L) γ2' α (if t.1:bool then nat_to_bool <$> (expander t.2) else nat_to_bool <$> (drop 1%nat (expander t.2))) ) ∗ ●m@γ2 ∗ ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ3 (●F z) )%I. @@ -113,46 +120,47 @@ Section impl2. ◯↯ε @ γ1 ∗ own γ3 (◯F 0%nat) }}}. Proof. - rewrite /new_counter2. - iIntros (Φ) "Hε HΦ". - wp_pures. - wp_alloc l as "Hl". - iDestruct (ec_valid with "[$]") as "%". - unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "[%γ1 [H1 H2]]". - { lra. } - simpl. - iMod (hocap_tapes_alloc' (∅:gmap _ _)) as "[%γ2 [H3 H4]]". - iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". - { by apply frac_auth_valid. } - replace (#0) with (#0%nat) by done. - iMod (inv_alloc N _ (counter_inv_pred2 (#l) γ1 γ2 γ3) with "[$Hε $Hl $H1 $H3 $H5]") as "#Hinv". - { iSplit; last done. by iApply big_sepM_empty. } - iApply "HΦ". - iExists _, _, _. by iFrame. - Qed. + Admitted. + (* rewrite /new_counter2. *) + (* iIntros (Φ) "Hε HΦ". *) + (* wp_pures. *) + (* wp_alloc l as "Hl". *) + (* iDestruct (ec_valid with "[$]") as "%". *) + (* unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "[%γ1 [H1 H2]]". *) + (* { lra. } *) + (* simpl. *) + (* iMod (hocap_tapes_alloc' (∅:gmap _ _)) as "[%γ2 [H3 H4]]". *) + (* iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". *) + (* { by apply frac_auth_valid. } *) + (* replace (#0) with (#0%nat) by done. *) + (* iMod (inv_alloc N _ (counter_inv_pred2 (#l) γ1 γ2 γ3) with "[$Hε $Hl $H1 $H3 $H5]") as "#Hinv". *) + (* { iSplit; last done. by iApply big_sepM_empty. } *) + (* iApply "HΦ". *) + (* iExists _, _, _. by iFrame. *) + (* Qed. *) (** This lemma is not possible as only one view shift*) - Lemma incr_counter_spec2 E N c γ1 γ2 γ3 (ε2:R -> nat -> R) (P: iProp Σ) (T: nat -> iProp Σ) (Q: nat->nat->iProp Σ): - ↑N ⊆ E-> - (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> - (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → - {{{ inv N (counter_inv_pred2 c γ1 γ2 γ3) ∗ - □(∀ (ε:R) (n : nat), P ∗ ●↯ ε @ γ1 ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨●↯ (ε2 ε n) @ γ1 ∗ T n) ) ∗ - □ (∀ (n z:nat), T n ∗ own γ3 (●F z) ={E∖↑N}=∗ - own γ3 (●F(z+n)%nat)∗ Q z n) ∗ - P - }}} - incr_counter2 c @ E - {{{ (n:nat) (z:nat), RET (#z, #n); Q z n }}}. - Proof. - iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP) HΦ". - rewrite /incr_counter2. - wp_pures. - wp_bind (rand _)%E. - iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". - (** cant do two view shifts! *) - Abort. + (* Lemma incr_counter_spec2 E N c γ1 γ2 γ3 (ε2:R -> nat -> R) (P: iProp Σ) (T: nat -> iProp Σ) (Q: nat->nat->iProp Σ): *) + (* ↑N ⊆ E-> *) + (* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) + (* (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → *) + (* {{{ inv N (counter_inv_pred2 c γ1 γ2 γ3) ∗ *) + (* □(∀ (ε:R) (n : nat), P ∗ ●↯ ε @ γ1 ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨●↯ (ε2 ε n) @ γ1 ∗ T n) ) ∗ *) + (* □ (∀ (n z:nat), T n ∗ own γ3 (●F z) ={E∖↑N}=∗ *) + (* own γ3 (●F(z+n)%nat)∗ Q z n) ∗ *) + (* P *) + (* }}} *) + (* incr_counter2 c @ E *) + (* {{{ (n:nat) (z:nat), RET (#z, #n); Q z n }}}. *) + (* Proof. *) + (* iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP) HΦ". *) + (* rewrite /incr_counter2. *) + (* wp_pures. *) + (* wp_bind (rand _)%E. *) + (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* (** cant do two view shifts! *) *) + (* Abort. *) Lemma allocate_tape_spec2 N E c γ1 γ2 γ3: ↑N ⊆ E-> @@ -162,21 +170,22 @@ Section impl2. ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ α ◯↪N (true, []) @ γ2 }}}. Proof. - iIntros (Hsubset Φ) "#Hinv HΦ". - rewrite /allocate_tape2. - wp_pures. - wp_alloctape α as "Hα". - iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". - iDestruct (hocap_tapes_notin' with "[$][$]") as "%". - iMod (hocap_tapes_new' _ _ _ _ true with "[$]") as "[H4 H7]"; first done. - replace ([]) with (expander []) by done. - iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Hα]") as "_". - { iNext. iSplitL; last done. - rewrite big_sepM_insert; [iFrame|done]. - } - iApply "HΦ". - by iFrame. - Qed. + Admitted. + (* iIntros (Hsubset Φ) "#Hinv HΦ". *) + (* rewrite /allocate_tape2. *) + (* wp_pures. *) + (* wp_alloctape α as "Hα". *) + (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* iDestruct (hocap_tapes_notin' with "[$][$]") as "%". *) + (* iMod (hocap_tapes_new' _ _ _ _ true with "[$]") as "[H4 H7]"; first done. *) + (* replace ([]) with (expander []) by done. *) + (* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Hα]") as "_". *) + (* { iNext. iSplitL; last done. *) + (* rewrite big_sepM_insert; [iFrame|done]. *) + (* } *) + (* iApply "HΦ". *) + (* by iFrame. *) + (* Qed. *) Lemma incr_counter_tape_spec_some2 N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: ↑N⊆E -> @@ -188,62 +197,63 @@ Section impl2. incr_counter_tape2 c #lbl:α @ E {{{ (z:nat), RET (#z, #n); Q z ∗ α ◯↪N (true, ns) @ γ2}}}. Proof. - iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP & Hα) HΦ". - rewrite /incr_counter_tape2. - wp_pures. - wp_bind (rand(_) _)%E. - iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". - iDestruct (hocap_tapes_agree' with "[$][$]") as "%". - erewrite <-(insert_delete m) at 1; last done. - rewrite big_sepM_insert; last apply lookup_delete. - simpl. - iDestruct "H3" as "[Htape H3]". - wp_apply (wp_rand_tape with "[$]"). - iIntros "[Htape %H1]". - iMod (hocap_tapes_pop1' with "[$][$]") as "[H4 Hα]". - iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". - { iSplitL; last done. - erewrite <-(insert_delete m) at 2; last done. - iNext. - rewrite insert_insert. - rewrite big_sepM_insert; last apply lookup_delete. iFrame. - } - iModIntro. - wp_pures. - clear -Hsubset H1. - wp_bind (rand(_) _)%E. - iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". - iDestruct (hocap_tapes_agree' with "[$][$]") as "%". - erewrite <-(insert_delete m) at 1; last done. - rewrite big_sepM_insert; last apply lookup_delete. - simpl. - iDestruct "H3" as "[Htape H3]". - wp_apply (wp_rand_tape with "[$]"). - iIntros "[Htape %H2]". - iMod (hocap_tapes_pop2' with "[$][$]") as "[H4 Hα]". - iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". - { iSplitL; last done. - erewrite <-(insert_delete m) at 2; last done. - iNext. - rewrite insert_insert. - rewrite big_sepM_insert; last apply lookup_delete. iFrame. - } - iModIntro. - wp_pures. - clear -Hsubset H1 H2. - wp_bind (FAA _ _). - iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". - wp_faa. - iMod ("Hvs" with "[$]") as "[H6 HQ]". - replace (#(z+n)) with (#(z+n)%nat); last first. - { by rewrite Nat2Z.inj_add. } - replace 2%Z with (Z.of_nat 2%nat) by done. - rewrite -Nat2Z.inj_mul -Nat2Z.inj_add -Nat.div2_odd -Nat2Z.inj_add. - iMod ("Hclose" with "[$H1 $H2 $H3 $H4 H5 $H6]") as "_"; first by iFrame. - iModIntro. wp_pures. - iApply "HΦ". - by iFrame. - Qed. + Admitted. + (* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP & Hα) HΦ". *) + (* rewrite /incr_counter_tape2. *) + (* wp_pures. *) + (* wp_bind (rand(_) _)%E. *) + (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* iDestruct (hocap_tapes_agree' with "[$][$]") as "%". *) + (* erewrite <-(insert_delete m) at 1; last done. *) + (* rewrite big_sepM_insert; last apply lookup_delete. *) + (* simpl. *) + (* iDestruct "H3" as "[Htape H3]". *) + (* wp_apply (wp_rand_tape with "[$]"). *) + (* iIntros "[Htape %H1]". *) + (* iMod (hocap_tapes_pop1' with "[$][$]") as "[H4 Hα]". *) + (* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". *) + (* { iSplitL; last done. *) + (* erewrite <-(insert_delete m) at 2; last done. *) + (* iNext. *) + (* rewrite insert_insert. *) + (* rewrite big_sepM_insert; last apply lookup_delete. iFrame. *) + (* } *) + (* iModIntro. *) + (* wp_pures. *) + (* clear -Hsubset H1. *) + (* wp_bind (rand(_) _)%E. *) + (* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* iDestruct (hocap_tapes_agree' with "[$][$]") as "%". *) + (* erewrite <-(insert_delete m) at 1; last done. *) + (* rewrite big_sepM_insert; last apply lookup_delete. *) + (* simpl. *) + (* iDestruct "H3" as "[Htape H3]". *) + (* wp_apply (wp_rand_tape with "[$]"). *) + (* iIntros "[Htape %H2]". *) + (* iMod (hocap_tapes_pop2' with "[$][$]") as "[H4 Hα]". *) + (* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". *) + (* { iSplitL; last done. *) + (* erewrite <-(insert_delete m) at 2; last done. *) + (* iNext. *) + (* rewrite insert_insert. *) + (* rewrite big_sepM_insert; last apply lookup_delete. iFrame. *) + (* } *) + (* iModIntro. *) + (* wp_pures. *) + (* clear -Hsubset H1 H2. *) + (* wp_bind (FAA _ _). *) + (* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* wp_faa. *) + (* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) + (* replace (#(z+n)) with (#(z+n)%nat); last first. *) + (* { by rewrite Nat2Z.inj_add. } *) + (* replace 2%Z with (Z.of_nat 2%nat) by done. *) + (* rewrite -Nat2Z.inj_mul -Nat2Z.inj_add -Nat.div2_odd -Nat2Z.inj_add. *) + (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 H5 $H6]") as "_"; first by iFrame. *) + (* iModIntro. wp_pures. *) + (* iApply "HΦ". *) + (* by iFrame. *) + (* Qed. *) Lemma counter_presample_spec2 NS E ns α (ε2 : R -> nat -> R) @@ -258,89 +268,90 @@ Section impl2. P -∗ α ◯↪N (true, ns) @ γ2 -∗ wp_update E (∃ n, T (n) ∗ α◯↪N (true, ns++[n]) @ γ2). Proof. - iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". - iApply wp_update_state_step_coupl. - iIntros (σ ε) "((Hheap&Htapes)&Hε)". - iMod (inv_acc with "Hinv") as "[>(% & % & % & % & H1 & H2 & H3 & H4 & -> & H5 & H6) Hclose]"; [done|]. - iDestruct (hocap_tapes_agree' with "[$][$]") as "%". - erewrite <-(insert_delete m) at 1; last done. - rewrite big_sepM_insert; last apply lookup_delete. - simpl. - iDestruct "H3" as "[Htape H3]". - iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%H1)". - iDestruct (ec_supply_bound with "[$][$]") as "%". - iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. - iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". - iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang 2%nat); first done. - pose (f (l:list (fin 2)) := (match l with - | x::[x'] => 2*fin_to_nat x + fin_to_nat x' - | _ => 0 - end)). - unshelve iExists - (λ l, mknonnegreal (ε2 ε1' (f l) ) _). - { apply Hpos; simpl. auto. } - simpl. - iSplit. - { iPureIntro. - etrans; last apply Hineq; last auto. - pose (k:=(λ n, 1 / ((1 + 1) * ((1 + 1) * 1)) * ε2 ε1' (f n))%R). - erewrite (SeriesC_ext _ (λ x, if bool_decide (x ∈ enum_uniform_list 1%nat 2%nat) - then k x else 0%R))%R; last first. - - intros. case_bool_decide as K. - + rewrite elem_of_enum_uniform_list in K. by rewrite K /k. - + rewrite elem_of_enum_uniform_list in K. - case_match eqn:K'; [by rewrite Nat.eqb_eq in K'|done]. - - rewrite SeriesC_list; last apply NoDup_enum_uniform_list. - rewrite /k /=. rewrite SeriesC_nat_bounded'/=. lra. - } - iIntros (sample ?). - destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' (f sample)))%R 1%R) as [Hdec|Hdec]; last first. - { apply Rnot_lt_ge, Rge_le in Hdec. - iApply state_step_coupl_ret_err_ge_1. - simpl. simpl in *. lra. - } - iApply state_step_coupl_ret. - unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' (f sample)) _) with "Hε_supply") as "[Hε_supply Hε]". - { apply Hpos. apply cond_nonneg. } - { simpl. done. } - assert (Nat.div2 (f sample)<2)%nat as K1. - { rewrite Nat.div2_div. apply Nat.Div0.div_lt_upper_bound. rewrite /f. - simpl. repeat case_match; try lia. pose proof fin_to_nat_lt t. pose proof fin_to_nat_lt t0. lia. - } - rewrite -H1. - iMod (tapeN_update_append _ _ _ _ (nat_to_fin K1) with "[$][$]") as "[Htapes Htape]". - assert (Nat.b2n (Nat.odd (f sample))<2)%nat as K2. - { edestruct (Nat.odd _); simpl; lia. } - rewrite -(list_fmap_singleton (fin_to_nat)) -fmap_app. - iMod (tapeN_update_append _ _ _ _ (nat_to_fin K2) with "[$][$]") as "[Htapes Htape]". - iMod (hocap_tapes_presample' _ _ _ _ _ (f sample) with "[$][$]") as "[H4 Hfrag]". - iMod "Hclose'" as "_". - iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". - { iExFalso. iApply (ec_contradict with "[$]"). exact. } - rewrite insert_insert. - rewrite fmap_app -!app_assoc /=. - assert (([nat_to_fin K1;nat_to_fin K2]) = sample) as K. - { destruct sample as [|x xs]; first done. - destruct xs as [|x' xs]; first done. - destruct xs as [|]; last done. - pose proof fin_to_nat_lt x. pose proof fin_to_nat_lt x'. - repeat f_equal; apply fin_to_nat_inj; rewrite fin_to_nat_to_fin. - - rewrite /f. rewrite Nat.div2_div. - rewrite Nat.mul_comm Nat.div_add_l; last lia. rewrite Nat.div_small; lia. - - rewrite /f. rewrite Nat.add_comm Nat.odd_add_even. - + destruct (fin_to_nat x') as [|[|]]; simpl; lia. - + by econstructor. - } - iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4 $H5 $H6]") as "_". - { iNext. iSplit; last done. - rewrite big_sepM_insert_delete; iFrame. - simpl. rewrite !fin_to_nat_to_fin. - rewrite /expander bind_app -/(expander ns). simpl. by rewrite H1. - } - iApply fupd_mask_intro_subseteq; first set_solver. - rewrite K. - iFrame. - Qed. + Admitted. + (* iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) + (* iApply wp_update_state_step_coupl. *) + (* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) + (* iMod (inv_acc with "Hinv") as "[>(% & % & % & % & H1 & H2 & H3 & H4 & -> & H5 & H6) Hclose]"; [done|]. *) + (* iDestruct (hocap_tapes_agree' with "[$][$]") as "%". *) + (* erewrite <-(insert_delete m) at 1; last done. *) + (* rewrite big_sepM_insert; last apply lookup_delete. *) + (* simpl. *) + (* iDestruct "H3" as "[Htape H3]". *) + (* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%H1)". *) + (* iDestruct (ec_supply_bound with "[$][$]") as "%". *) + (* iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. *) + (* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) + (* iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang 2%nat); first done. *) + (* pose (f (l:list (fin 2)) := (match l with *) + (* | x::[x'] => 2*fin_to_nat x + fin_to_nat x' *) + (* | _ => 0 *) + (* end)). *) + (* unshelve iExists *) + (* (λ l, mknonnegreal (ε2 ε1' (f l) ) _). *) + (* { apply Hpos; simpl. auto. } *) + (* simpl. *) + (* iSplit. *) + (* { iPureIntro. *) + (* etrans; last apply Hineq; last auto. *) + (* pose (k:=(λ n, 1 / ((1 + 1) * ((1 + 1) * 1)) * ε2 ε1' (f n))%R). *) + (* erewrite (SeriesC_ext _ (λ x, if bool_decide (x ∈ enum_uniform_list 1%nat 2%nat) *) + (* then k x else 0%R))%R; last first. *) + (* - intros. case_bool_decide as K. *) + (* + rewrite elem_of_enum_uniform_list in K. by rewrite K /k. *) + (* + rewrite elem_of_enum_uniform_list in K. *) + (* case_match eqn:K'; [by rewrite Nat.eqb_eq in K'|done]. *) + (* - rewrite SeriesC_list; last apply NoDup_enum_uniform_list. *) + (* rewrite /k /=. rewrite SeriesC_nat_bounded'/=. lra. *) + (* } *) + (* iIntros (sample ?). *) + (* destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' (f sample)))%R 1%R) as [Hdec|Hdec]; last first. *) + (* { apply Rnot_lt_ge, Rge_le in Hdec. *) + (* iApply state_step_coupl_ret_err_ge_1. *) + (* simpl. simpl in *. lra. *) + (* } *) + (* iApply state_step_coupl_ret. *) + (* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' (f sample)) _) with "Hε_supply") as "[Hε_supply Hε]". *) + (* { apply Hpos. apply cond_nonneg. } *) + (* { simpl. done. } *) + (* assert (Nat.div2 (f sample)<2)%nat as K1. *) + (* { rewrite Nat.div2_div. apply Nat.Div0.div_lt_upper_bound. rewrite /f. *) + (* simpl. repeat case_match; try lia. pose proof fin_to_nat_lt t. pose proof fin_to_nat_lt t0. lia. *) + (* } *) + (* rewrite -H1. *) + (* iMod (tapeN_update_append _ _ _ _ (nat_to_fin K1) with "[$][$]") as "[Htapes Htape]". *) + (* assert (Nat.b2n (Nat.odd (f sample))<2)%nat as K2. *) + (* { edestruct (Nat.odd _); simpl; lia. } *) + (* rewrite -(list_fmap_singleton (fin_to_nat)) -fmap_app. *) + (* iMod (tapeN_update_append _ _ _ _ (nat_to_fin K2) with "[$][$]") as "[Htapes Htape]". *) + (* iMod (hocap_tapes_presample' _ _ _ _ _ (f sample) with "[$][$]") as "[H4 Hfrag]". *) + (* iMod "Hclose'" as "_". *) + (* iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". *) + (* { iExFalso. iApply (ec_contradict with "[$]"). exact. } *) + (* rewrite insert_insert. *) + (* rewrite fmap_app -!app_assoc /=. *) + (* assert (([nat_to_fin K1;nat_to_fin K2]) = sample) as K. *) + (* { destruct sample as [|x xs]; first done. *) + (* destruct xs as [|x' xs]; first done. *) + (* destruct xs as [|]; last done. *) + (* pose proof fin_to_nat_lt x. pose proof fin_to_nat_lt x'. *) + (* repeat f_equal; apply fin_to_nat_inj; rewrite fin_to_nat_to_fin. *) + (* - rewrite /f. rewrite Nat.div2_div. *) + (* rewrite Nat.mul_comm Nat.div_add_l; last lia. rewrite Nat.div_small; lia. *) + (* - rewrite /f. rewrite Nat.add_comm Nat.odd_add_even. *) + (* + destruct (fin_to_nat x') as [|[|]]; simpl; lia. *) + (* + by econstructor. *) + (* } *) + (* iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4 $H5 $H6]") as "_". *) + (* { iNext. iSplit; last done. *) + (* rewrite big_sepM_insert_delete; iFrame. *) + (* simpl. rewrite !fin_to_nat_to_fin. *) + (* rewrite /expander bind_app -/(expander ns). simpl. by rewrite H1. *) + (* } *) + (* iApply fupd_mask_intro_subseteq; first set_solver. *) + (* rewrite K. *) + (* iFrame. *) + (* Qed. *) Lemma read_counter_spec2 N E c γ1 γ2 γ3 P Q: ↑N ⊆ E -> @@ -353,19 +364,20 @@ Section impl2. {{{ (n':nat), RET #n'; Q n' }}}. Proof. - iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP) HΦ". - rewrite /read_counter2. - wp_pure. - iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". - wp_load. - iMod ("Hvs" with "[$]") as "[H6 HQ]". - iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]"); first done. - iApply ("HΦ" with "[$]"). - Qed. + Admitted. + (* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP) HΦ". *) + (* rewrite /read_counter2. *) + (* wp_pure. *) + (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* wp_load. *) + (* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) + (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]"); first done. *) + (* iApply ("HΦ" with "[$]"). *) + (* Qed. *) End impl2. -Program Definition random_counter2 `{!conerisGS Σ}: random_counter := +Program Definition random_counter2 `{!conerisGS Σ, flip_spec Σ}: random_counter := {| new_counter := new_counter2; allocate_tape:= allocate_tape2; incr_counter_tape := incr_counter_tape2; @@ -374,7 +386,7 @@ Program Definition random_counter2 `{!conerisGS Σ}: random_counter := error_name := gname; tape_name := gname; counter_name :=gname; - is_counter _ N c γ1 γ2 γ3 := inv N (counter_inv_pred2 c γ1 γ2 γ3); + is_counter _ N c γ1 γ2 γ3 := inv N (counter_inv_pred2 (L:=_) c γ1 γ2 γ3); counter_error_auth _ γ x := ●↯ x @ γ; counter_error_frag _ γ x := ◯↯ x @ γ; counter_tapes_auth _ γ m := (●((λ ns, (true, ns))<$>m)@γ)%I; @@ -389,49 +401,49 @@ Program Definition random_counter2 `{!conerisGS Σ}: random_counter := |}. Next Obligation. simpl. - iIntros (??????) "(%&<-&H1)(%&<-&H2)". - iCombine "H1 H2" gives "%H". by rewrite excl_auth.excl_auth_auth_op_valid in H. + iIntros (????????) "(%&<-&H1)(%&<-&H2)". + iCombine "H1 H2" gives "%K". by rewrite excl_auth.excl_auth_auth_op_valid in K. Qed. Next Obligation. simpl. - iIntros (??????) "(%&<-&H1)(%&<-&H2)". - iCombine "H1 H2" gives "%H". by rewrite excl_auth.excl_auth_frag_op_valid in H. + iIntros (????????) "(%&<-&H1)(%&<-&H2)". + iCombine "H1 H2" gives "%K". by rewrite excl_auth.excl_auth_frag_op_valid in K. Qed. Next Obligation. simpl. - iIntros (?????) "H". + iIntros (???????) "H". iApply (hocap_error_auth_pos with "[$]"). Qed. Next Obligation. simpl. - iIntros (?????) "H". + iIntros (???????) "H". iApply (hocap_error_frag_pos with "[$]"). Qed. Next Obligation. simpl. - iIntros (??????) "H1 H2". + iIntros (????????) "H1 H2". iApply (hocap_error_agree with "[$][$]"). Qed. Next Obligation. - simpl. iIntros (???????) "??". + simpl. iIntros (?????????) "??". iApply (hocap_error_update with "[$][$]"). Qed. Next Obligation. simpl. - iIntros (??????) "H1 H2". - by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[%H _]". + iIntros (????????) "H1 H2". + by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[% _]". Qed. Next Obligation. simpl. - iIntros (???????) "H1 H2". + iIntros (?????????) "H1 H2". iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. rewrite dfrac_op_own dfrac_valid_own. by intro. Qed. Next Obligation. simpl. iIntros. - iDestruct (hocap_tapes_agree' with "[$][$]") as "%H". - rewrite lookup_fmap_Some in H. destruct H as (?&?&?). + iDestruct (hocap_tapes_agree' with "[$][$]") as "%K". + rewrite lookup_fmap_Some in K. destruct K as (?&?&?). by simplify_eq. Qed. Next Obligation. @@ -441,28 +453,28 @@ Next Obligation. Qed. Next Obligation. simpl. - iIntros (??????) "H1 H2". - iCombine "H1 H2" gives "%H". by rewrite auth_auth_op_valid in H. + iIntros (????????) "H1 H2". + iCombine "H1 H2" gives "%K". by rewrite auth_auth_op_valid in K. Qed. Next Obligation. - simpl. iIntros (???? z z' ?) "H1 H2". - iCombine "H1 H2" gives "%H". - apply frac_auth_included_total in H. iPureIntro. + simpl. iIntros (?????? z z' ?) "H1 H2". + iCombine "H1 H2" gives "%K". + apply frac_auth_included_total in K. iPureIntro. by apply nat_included. Qed. Next Obligation. simpl. - iIntros (????????). + iIntros (??????????). rewrite frac_auth_frag_op. by rewrite own_op. Qed. Next Obligation. - simpl. iIntros (??????) "H1 H2". - iCombine "H1 H2" gives "%H". + simpl. iIntros (????????) "H1 H2". + iCombine "H1 H2" gives "%K". iPureIntro. - by apply frac_auth_agree_L in H. + by apply frac_auth_agree_L in K. Qed. Next Obligation. - simpl. iIntros (????????) "H1 H2". + simpl. iIntros (??????????) "H1 H2". iMod (own_update_2 _ _ _ (_ ⋅ _) with "[$][$]") as "[$$]"; last done. apply frac_auth_update. apply nat_local_update. lia. diff --git a/theories/coneris/lib/flip_spec.v b/theories/coneris/lib/flip_spec.v index 20e6f7eb..641e44a2 100644 --- a/theories/coneris/lib/flip_spec.v +++ b/theories/coneris/lib/flip_spec.v @@ -6,21 +6,21 @@ Set Default Proof Using "Type*". Class flip_spec `{!conerisGS Σ} := FlipSpec { (** * Operations *) - allocate_tape : val; + flip_allocate_tape : val; flip_tape : val; (** * Ghost state *) (** The assumptions about [Σ] *) flipG : gFunctors → Type; (** [name] is used to associate [locked] with [is_lock] *) - error_name : Type; - tape_name: Type; + flip_error_name : Type; + flip_tape_name: Type; (** * Predicates *) is_flip {L : flipG Σ} (N:namespace) - (γ1: error_name) (γ2: tape_name): iProp Σ; - flip_error_auth {L : flipG Σ} (γ: error_name) (x:R): iProp Σ; - flip_error_frag {L : flipG Σ} (γ: error_name) (x:R): iProp Σ; - flip_tapes_auth {L : flipG Σ} (γ: tape_name) (m:gmap loc (list bool)): iProp Σ; - flip_tapes_frag {L : flipG Σ} (γ: tape_name) (α:loc) (ns:list bool): iProp Σ; + (γ1: flip_error_name) (γ2: flip_tape_name): iProp Σ; + flip_error_auth {L : flipG Σ} (γ: flip_error_name) (x:R): iProp Σ; + flip_error_frag {L : flipG Σ} (γ: flip_error_name) (x:R): iProp Σ; + flip_tapes_auth {L : flipG Σ} (γ: flip_tape_name) (m:gmap loc (list bool)): iProp Σ; + flip_tapes_frag {L : flipG Σ} (γ: flip_tape_name) (α:loc) (ns:list bool): iProp Σ; (** * General properties of the predicates *) #[global] is_flip_persistent {L : flipG Σ} N γ1 γ2 :: Persistent (is_flip (L:=L) N γ1 γ2); @@ -66,7 +66,7 @@ Class flip_spec `{!conerisGS Σ} := FlipSpec flip_allocate_tape_spec {L: flipG Σ} N E γ1 γ2: ↑N ⊆ E-> {{{ is_flip (L:=L) N γ1 γ2 }}} - allocate_tape #() @ E + flip_allocate_tape #() @ E {{{ (v:val), RET v; ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ flip_tapes_frag (L:=L) γ2 α [] }}}; @@ -130,11 +130,11 @@ Local Definition flip_inv_pred1 `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapes )%I. #[local] Program Definition flip_spec1 `{!conerisGS Σ}: flip_spec := - {| allocate_tape:= (λ: <>, allocB); + {| flip_allocate_tape:= (λ: <>, allocB); flip_tape:= flipL; flipG := flipG1; - error_name := gname; - tape_name := gname; + flip_error_name := gname; + flip_tape_name := gname; is_flip _ N γ1 γ2 := inv N (flip_inv_pred1 γ1 γ2); flip_error_auth _ γ x := ●↯ x @ γ; flip_error_frag _ γ x := ◯↯ x @ γ; From f2d886e0f5a90d687df29649e0abadcd12b761ba Mon Sep 17 00:00:00 2001 From: Hei Li Date: Tue, 17 Sep 2024 15:21:11 +0200 Subject: [PATCH 05/69] new_counter_spec2 --- .../coneris/examples/random_counter/impl2.v | 67 ++++++++++--------- 1 file changed, 35 insertions(+), 32 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index 20ccfad8..6e8ec3c1 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -97,15 +97,15 @@ Section impl2. Class counterG2 Σ := CounterG2 { counterG2_error::hocap_errorGS Σ; counterG2_tapes:: hocap_tapesGS' Σ; counterG2_frac_authR:: inG Σ (frac_authR natR); - counterG2_flipG:: flipG Σ + counterG2_flipG: flipG Σ }. - Context `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapesGS' Σ, !inG Σ (frac_authR natR), L:flipG Σ}. + Context `{L:!flipG Σ, !hocap_errorGS Σ, !hocap_tapesGS' Σ, !inG Σ (frac_authR natR)}. - Definition counter_inv_pred2 (c:val) γ1 γ2 γ3:= - (∃ (ε:R) (m:gmap loc (bool*list nat)) (l:loc) (z:nat) γ1' γ2' N', - is_flip (L:=L) N' γ1' γ2' ∗ + Definition counter_inv_pred2 N (c:val) γ1 γ2 γ3:= + (∃ (ε:R) (m:gmap loc (bool*list nat)) (l:loc) (z:nat) γ1' γ2', + is_flip (L:=L) (N.@"flip") γ1' γ2' ∗ flip_error_frag (L:=L) γ1' ε ∗ ●↯ ε @ γ1 ∗ ([∗ map] α ↦ t ∈ m, flip_tapes_frag (L:=L) γ2' α (if t.1:bool then nat_to_bool <$> (expander t.2) else nat_to_bool <$> (drop 1%nat (expander t.2))) ) ∗ ●m@γ2 ∗ @@ -116,28 +116,30 @@ Section impl2. {{{ ↯ ε }}} new_counter2 #() @ E {{{ (c:val), RET c; - ∃ γ1 γ2 γ3, inv N (counter_inv_pred2 c γ1 γ2 γ3) ∗ + ∃ γ1 γ2 γ3, inv (N.@ "counter") (counter_inv_pred2 N c γ1 γ2 γ3) ∗ ◯↯ε @ γ1 ∗ own γ3 (◯F 0%nat) }}}. Proof. - Admitted. - (* rewrite /new_counter2. *) - (* iIntros (Φ) "Hε HΦ". *) - (* wp_pures. *) - (* wp_alloc l as "Hl". *) - (* iDestruct (ec_valid with "[$]") as "%". *) - (* unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "[%γ1 [H1 H2]]". *) - (* { lra. } *) - (* simpl. *) - (* iMod (hocap_tapes_alloc' (∅:gmap _ _)) as "[%γ2 [H3 H4]]". *) - (* iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". *) - (* { by apply frac_auth_valid. } *) - (* replace (#0) with (#0%nat) by done. *) - (* iMod (inv_alloc N _ (counter_inv_pred2 (#l) γ1 γ2 γ3) with "[$Hε $Hl $H1 $H3 $H5]") as "#Hinv". *) - (* { iSplit; last done. by iApply big_sepM_empty. } *) - (* iApply "HΦ". *) - (* iExists _, _, _. by iFrame. *) - (* Qed. *) + rewrite /new_counter2. + iIntros (Φ) "Hε HΦ". + wp_pures. + iDestruct (ec_valid with "[$]") as "%". + iMod (flip_inv_create_spec with "[$]") as "(%γ1'&%γ2'&#?&Hε)". + wp_alloc l as "Hl". + unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "[%γ1 [H1 H2]]". + { lra. } + simpl. + iMod (hocap_tapes_alloc' (∅:gmap _ _)) as "[%γ2 [H3 H4]]". + iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". + { by apply frac_auth_valid. } + replace (#0) with (#0%nat) by done. + iMod (inv_alloc (N.@"counter") _ (counter_inv_pred2 N (#l) γ1 γ2 γ3) with "[$Hε $Hl $H1 $H3 $H5]") as "#Hinv". + { iNext. iExists _. iSplit; first done. + iSplit; last done. + by iApply big_sepM_empty. } + iApply "HΦ". + iExists _, _, _. by iFrame. + Qed. (** This lemma is not possible as only one view shift*) @@ -164,7 +166,7 @@ Section impl2. Lemma allocate_tape_spec2 N E c γ1 γ2 γ3: ↑N ⊆ E-> - {{{ inv N (counter_inv_pred2 c γ1 γ2 γ3) }}} + {{{ inv (N.@ "counter") (counter_inv_pred2 N c γ1 γ2 γ3) }}} allocate_tape2 #() @ E {{{ (v:val), RET v; ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ α ◯↪N (true, []) @ γ2 @@ -189,7 +191,7 @@ Section impl2. Lemma incr_counter_tape_spec_some2 N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: ↑N⊆E -> - {{{ inv N (counter_inv_pred2 c γ1 γ2 γ3) ∗ + {{{ inv (N.@ "counter") (counter_inv_pred2 N c γ1 γ2 γ3) ∗ □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ own γ3 (●F(z+n)%nat)∗ Q z) ∗ P ∗ α ◯↪N (true, n::ns) @ γ2 @@ -255,14 +257,14 @@ Section impl2. (* by iFrame. *) (* Qed. *) - Lemma counter_presample_spec2 NS E ns α + Lemma counter_presample_spec2 N E ns α (ε2 : R -> nat -> R) (P : iProp Σ) T γ1 γ2 γ3 c: - ↑NS ⊆ E -> + ↑N ⊆ E -> (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> (∀ (ε:R), 0<= ε ->SeriesC (λ n, if (bool_decide (n≤3%nat)) then 1 / (S 3%nat) * ε2 ε n else 0%R)%R <= ε)%R-> - inv NS (counter_inv_pred2 c γ1 γ2 γ3) -∗ - (□∀ (ε:R) n, (P ∗ ●↯ ε@ γ1) ={E∖↑NS}=∗ + inv (N.@ "counter") (counter_inv_pred2 N c γ1 γ2 γ3) -∗ + (□∀ (ε:R) n, (P ∗ ●↯ ε@ γ1) ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝ ∨(●↯ (ε2 ε n) @ γ1 ∗ T (n)))) -∗ P -∗ α ◯↪N (true, ns) @ γ2 -∗ @@ -355,7 +357,7 @@ Section impl2. Lemma read_counter_spec2 N E c γ1 γ2 γ3 P Q: ↑N ⊆ E -> - {{{ inv N (counter_inv_pred2 c γ1 γ2 γ3) ∗ + {{{ inv (N.@ "counter") (counter_inv_pred2 N c γ1 γ2 γ3) ∗ □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ own γ3 (●F z)∗ Q z) ∗ P @@ -386,7 +388,7 @@ Program Definition random_counter2 `{!conerisGS Σ, flip_spec Σ}: random_counte error_name := gname; tape_name := gname; counter_name :=gname; - is_counter _ N c γ1 γ2 γ3 := inv N (counter_inv_pred2 (L:=_) c γ1 γ2 γ3); + is_counter _ N c γ1 γ2 γ3 := inv (N.@ "counter") (counter_inv_pred2 (L:=_) N c γ1 γ2 γ3); counter_error_auth _ γ x := ●↯ x @ γ; counter_error_frag _ γ x := ◯↯ x @ γ; counter_tapes_auth _ γ m := (●((λ ns, (true, ns))<$>m)@γ)%I; @@ -480,3 +482,4 @@ Next Obligation. apply nat_local_update. lia. Qed. + From ee83b1bbc2f9d0c7fc7d1553edad10bb0c189091 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Tue, 17 Sep 2024 15:55:32 +0200 Subject: [PATCH 06/69] Progress on impl 2, getting stuck a bit --- theories/coneris/examples/random_counter/impl2.v | 12 +++++++----- theories/coneris/lib/flip_spec.v | 4 ++++ 2 files changed, 11 insertions(+), 5 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index 6e8ec3c1..3a4d7252 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -84,7 +84,7 @@ Section impl2. (* let: "n'" := rand #1 in *) (* let: "x" := #2 * "n" + "n'" in *) (* (FAA "l" "x", "x"). *) - Definition allocate_tape2 : val := λ: "_", AllocTape #1. + Definition allocate_tape2 : val := flip_allocate_tape. Definition incr_counter_tape2 :val := λ: "l" "α", let: "n" := conversion.bool_to_int (flip_allocate_tape "α") in @@ -172,12 +172,14 @@ Section impl2. ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ α ◯↪N (true, []) @ γ2 }}}. Proof. + iIntros (Hsubset Φ) "#Hinv HΦ". + rewrite /allocate_tape2. + iApply fupd_pgl_wp. + iApply (flip_allocate_tape_spec). + (** Spec not strong enough *) Admitted. - (* iIntros (Hsubset Φ) "#Hinv HΦ". *) - (* rewrite /allocate_tape2. *) - (* wp_pures. *) + (* iInv (N.@"counter") as "(%ε & %m & %l & %z & %γ1' & %γ2' & #Hinv' & >H1 & >H2 & >H3 & >H4 & > (-> & H5 & H6))" "Hclose". *) (* wp_alloctape α as "Hα". *) - (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) (* iDestruct (hocap_tapes_notin' with "[$][$]") as "%". *) (* iMod (hocap_tapes_new' _ _ _ _ true with "[$]") as "[H4 H7]"; first done. *) (* replace ([]) with (expander []) by done. *) diff --git a/theories/coneris/lib/flip_spec.v b/theories/coneris/lib/flip_spec.v index 641e44a2..ca455555 100644 --- a/theories/coneris/lib/flip_spec.v +++ b/theories/coneris/lib/flip_spec.v @@ -32,6 +32,10 @@ Class flip_spec `{!conerisGS Σ} := FlipSpec Timeless (flip_tapes_auth (L:=L) γ m); #[global] flip_tapes_frag_timeless {L : flipG Σ} γ α ns :: Timeless (flip_tapes_frag (L:=L) γ α ns); + #[global] flip_error_name_inhabited:: + Inhabited flip_error_name; + #[global] flip_tape_name_inhabited:: + Inhabited flip_tape_name; flip_error_auth_exclusive {L : flipG Σ} γ x1 x2: flip_error_auth (L:=L) γ x1 -∗ flip_error_auth (L:=L) γ x2 -∗ False; From bf4565aa805dc277793fc803136fd239bf6621bd Mon Sep 17 00:00:00 2001 From: Hei Li Date: Tue, 17 Sep 2024 16:16:22 +0200 Subject: [PATCH 07/69] another try? --- .../coneris/examples/random_counter/impl2.v | 30 +++++++++++-------- 1 file changed, 18 insertions(+), 12 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index 3a4d7252..2fc64e87 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -103,9 +103,8 @@ Section impl2. Context `{L:!flipG Σ, !hocap_errorGS Σ, !hocap_tapesGS' Σ, !inG Σ (frac_authR natR)}. - Definition counter_inv_pred2 N (c:val) γ1 γ2 γ3:= - (∃ (ε:R) (m:gmap loc (bool*list nat)) (l:loc) (z:nat) γ1' γ2', - is_flip (L:=L) (N.@"flip") γ1' γ2' ∗ + Definition counter_inv_pred2 (c:val) γ1 γ1' γ2 γ2' γ3:= + (∃ (ε:R) (m:gmap loc (bool*list nat)) (l:loc) (z:nat), flip_error_frag (L:=L) γ1' ε ∗ ●↯ ε @ γ1 ∗ ([∗ map] α ↦ t ∈ m, flip_tapes_frag (L:=L) γ2' α (if t.1:bool then nat_to_bool <$> (expander t.2) else nat_to_bool <$> (drop 1%nat (expander t.2))) ) ∗ ●m@γ2 ∗ @@ -116,7 +115,8 @@ Section impl2. {{{ ↯ ε }}} new_counter2 #() @ E {{{ (c:val), RET c; - ∃ γ1 γ2 γ3, inv (N.@ "counter") (counter_inv_pred2 N c γ1 γ2 γ3) ∗ + ∃ γ1 γ2 γ3, (∃ γ1' γ2', is_flip (L:=L) (N.@"flip") γ1' γ2' ∗ + inv (N.@ "counter") (counter_inv_pred2 c γ1 γ1' γ2 γ2' γ3)) ∗ ◯↯ε @ γ1 ∗ own γ3 (◯F 0%nat) }}}. Proof. @@ -133,12 +133,13 @@ Section impl2. iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". { by apply frac_auth_valid. } replace (#0) with (#0%nat) by done. - iMod (inv_alloc (N.@"counter") _ (counter_inv_pred2 N (#l) γ1 γ2 γ3) with "[$Hε $Hl $H1 $H3 $H5]") as "#Hinv". - { iNext. iExists _. iSplit; first done. + iMod (inv_alloc (N.@"counter") _ (counter_inv_pred2 (#l) γ1 γ1' γ2 γ2' γ3) with "[$Hε $Hl $H1 $H3 $H5]") as "#Hinv". + { iNext. iSplit; last done. by iApply big_sepM_empty. } iApply "HΦ". - iExists _, _, _. by iFrame. + iExists _, _, _. iModIntro. iFrame. + iExists _, _. by iSplit. Qed. @@ -166,7 +167,8 @@ Section impl2. Lemma allocate_tape_spec2 N E c γ1 γ2 γ3: ↑N ⊆ E-> - {{{ inv (N.@ "counter") (counter_inv_pred2 N c γ1 γ2 γ3) }}} + {{{ (∃ γ1' γ2', is_flip (L:=L) (N.@"flip") γ1' γ2' ∗ + inv (N.@ "counter") (counter_inv_pred2 c γ1 γ1' γ2 γ2' γ3)) }}} allocate_tape2 #() @ E {{{ (v:val), RET v; ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ α ◯↪N (true, []) @ γ2 @@ -193,7 +195,8 @@ Section impl2. Lemma incr_counter_tape_spec_some2 N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: ↑N⊆E -> - {{{ inv (N.@ "counter") (counter_inv_pred2 N c γ1 γ2 γ3) ∗ + {{{ (∃ γ1' γ2', is_flip (L:=L) (N.@"flip") γ1' γ2' ∗ + inv (N.@ "counter") (counter_inv_pred2 c γ1 γ1' γ2 γ2' γ3)) ∗ □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ own γ3 (●F(z+n)%nat)∗ Q z) ∗ P ∗ α ◯↪N (true, n::ns) @ γ2 @@ -265,7 +268,8 @@ Section impl2. ↑N ⊆ E -> (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> (∀ (ε:R), 0<= ε ->SeriesC (λ n, if (bool_decide (n≤3%nat)) then 1 / (S 3%nat) * ε2 ε n else 0%R)%R <= ε)%R-> - inv (N.@ "counter") (counter_inv_pred2 N c γ1 γ2 γ3) -∗ + (∃ γ1' γ2', is_flip (L:=L) (N.@"flip") γ1' γ2' ∗ + inv (N.@ "counter") (counter_inv_pred2 c γ1 γ1' γ2 γ2' γ3)) -∗ (□∀ (ε:R) n, (P ∗ ●↯ ε@ γ1) ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝ ∨(●↯ (ε2 ε n) @ γ1 ∗ T (n)))) -∗ @@ -359,7 +363,8 @@ Section impl2. Lemma read_counter_spec2 N E c γ1 γ2 γ3 P Q: ↑N ⊆ E -> - {{{ inv (N.@ "counter") (counter_inv_pred2 N c γ1 γ2 γ3) ∗ + {{{ (∃ γ1' γ2', is_flip (L:=L) (N.@"flip") γ1' γ2' ∗ + inv (N.@ "counter") (counter_inv_pred2 c γ1 γ1' γ2 γ2' γ3)) ∗ □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ own γ3 (●F z)∗ Q z) ∗ P @@ -390,7 +395,8 @@ Program Definition random_counter2 `{!conerisGS Σ, flip_spec Σ}: random_counte error_name := gname; tape_name := gname; counter_name :=gname; - is_counter _ N c γ1 γ2 γ3 := inv (N.@ "counter") (counter_inv_pred2 (L:=_) N c γ1 γ2 γ3); + is_counter _ N c γ1 γ2 γ3 := (∃ γ1' γ2', is_flip (N.@"flip") γ1' γ2' ∗ + inv (N.@ "counter") (counter_inv_pred2 (L:=_) c γ1 γ1' γ2 γ2' γ3))%I; counter_error_auth _ γ x := ●↯ x @ γ; counter_error_frag _ γ x := ◯↯ x @ γ; counter_tapes_auth _ γ m := (●((λ ns, (true, ns))<$>m)@γ)%I; From 0a5480f7a3baf7cfbb728f20659293433a33a489 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Thu, 19 Sep 2024 09:47:10 +0200 Subject: [PATCH 08/69] first rough sketch of proof of impl2 --- .../coneris/examples/random_counter/impl2.v | 128 +++++++++--------- 1 file changed, 62 insertions(+), 66 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index 2fc64e87..bcf09776 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -94,30 +94,28 @@ Section impl2. let: "x" := #2 * "n" + "n'" in (FAA "l" "x", "x"). Definition read_counter2 : val := λ: "l", !"l". - Class counterG2 Σ := CounterG2 { counterG2_error::hocap_errorGS Σ; - counterG2_tapes:: hocap_tapesGS' Σ; + Class counterG2 Σ := CounterG2 { counterG2_tapes:: hocap_tapesGS' Σ; counterG2_frac_authR:: inG Σ (frac_authR natR); counterG2_flipG: flipG Σ }. - Context `{L:!flipG Σ, !hocap_errorGS Σ, !hocap_tapesGS' Σ, !inG Σ (frac_authR natR)}. + Context `{L:!flipG Σ, !hocap_tapesGS' Σ, !inG Σ (frac_authR natR)}. - Definition counter_inv_pred2 (c:val) γ1 γ1' γ2 γ2' γ3:= - (∃ (ε:R) (m:gmap loc (bool*list nat)) (l:loc) (z:nat), - flip_error_frag (L:=L) γ1' ε ∗ ●↯ ε @ γ1 ∗ - ([∗ map] α ↦ t ∈ m, flip_tapes_frag (L:=L) γ2' α (if t.1:bool then nat_to_bool <$> (expander t.2) else nat_to_bool <$> (drop 1%nat (expander t.2))) ) - ∗ ●m@γ2 ∗ - ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ3 (●F z) + Definition counter_inv_pred2 (c:val) γ1 γ1' γ2 := + (∃ (m:gmap loc (bool*list nat)) (l:loc) (z:nat), + ([∗ map] α ↦ t ∈ m, flip_tapes_frag (L:=L) γ1' α (if t.1:bool then nat_to_bool <$> (expander t.2) else nat_to_bool <$> (drop 1%nat (expander t.2))) ) + ∗ ●m@γ1 ∗ + ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ2 (●F z) )%I. Lemma new_counter_spec2 E ε N: {{{ ↯ ε }}} new_counter2 #() @ E {{{ (c:val), RET c; - ∃ γ1 γ2 γ3, (∃ γ1' γ2', is_flip (L:=L) (N.@"flip") γ1' γ2' ∗ - inv (N.@ "counter") (counter_inv_pred2 c γ1 γ1' γ2 γ2' γ3)) ∗ - ◯↯ε @ γ1 ∗ own γ3 (◯F 0%nat) + ∃ γ1 γ2 γ3, (∃ γ2', is_flip (L:=L) (N.@"flip") γ1 γ2' ∗ + inv (N.@ "counter") (counter_inv_pred2 c γ2 γ2' γ3)) ∗ + flip_error_frag (L:=L) γ1 ε ∗ own γ3 (◯F 0%nat) }}}. Proof. rewrite /new_counter2. @@ -126,20 +124,18 @@ Section impl2. iDestruct (ec_valid with "[$]") as "%". iMod (flip_inv_create_spec with "[$]") as "(%γ1'&%γ2'&#?&Hε)". wp_alloc l as "Hl". - unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "[%γ1 [H1 H2]]". - { lra. } simpl. iMod (hocap_tapes_alloc' (∅:gmap _ _)) as "[%γ2 [H3 H4]]". iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". { by apply frac_auth_valid. } replace (#0) with (#0%nat) by done. - iMod (inv_alloc (N.@"counter") _ (counter_inv_pred2 (#l) γ1 γ1' γ2 γ2' γ3) with "[$Hε $Hl $H1 $H3 $H5]") as "#Hinv". - { iNext. + iMod (inv_alloc (N.@"counter") _ (counter_inv_pred2 (#l) γ2 γ2' γ3) with "[$Hl $H3 $H5]") as "#Hinv". + { iNext. iSplit; last done. by iApply big_sepM_empty. } iApply "HΦ". iExists _, _, _. iModIntro. iFrame. - iExists _, _. by iSplit. + iExists _. by iSplit. Qed. @@ -167,8 +163,8 @@ Section impl2. Lemma allocate_tape_spec2 N E c γ1 γ2 γ3: ↑N ⊆ E-> - {{{ (∃ γ1' γ2', is_flip (L:=L) (N.@"flip") γ1' γ2' ∗ - inv (N.@ "counter") (counter_inv_pred2 c γ1 γ1' γ2 γ2' γ3)) }}} + {{{ (∃ γ2', is_flip (L:=L) (N.@"flip") γ1 γ2' ∗ + inv (N.@ "counter") (counter_inv_pred2 c γ2 γ2' γ3)) }}} allocate_tape2 #() @ E {{{ (v:val), RET v; ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ α ◯↪N (true, []) @ γ2 @@ -195,8 +191,8 @@ Section impl2. Lemma incr_counter_tape_spec_some2 N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: ↑N⊆E -> - {{{ (∃ γ1' γ2', is_flip (L:=L) (N.@"flip") γ1' γ2' ∗ - inv (N.@ "counter") (counter_inv_pred2 c γ1 γ1' γ2 γ2' γ3)) ∗ + {{{ (∃ γ2', is_flip (L:=L) (N.@"flip") γ1 γ2' ∗ + inv (N.@ "counter") (counter_inv_pred2 c γ2 γ2' γ3)) ∗ □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ own γ3 (●F(z+n)%nat)∗ Q z) ∗ P ∗ α ◯↪N (true, n::ns) @ γ2 @@ -204,10 +200,10 @@ Section impl2. incr_counter_tape2 c #lbl:α @ E {{{ (z:nat), RET (#z, #n); Q z ∗ α ◯↪N (true, ns) @ γ2}}}. Proof. - Admitted. - (* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP & Hα) HΦ". *) - (* rewrite /incr_counter_tape2. *) - (* wp_pures. *) + iIntros (Hsubset Φ) "((%γ2' & #Hinv) & #Hvs & HP & Hα) HΦ". + rewrite /incr_counter_tape2. + wp_pures. + Admitted. (* wp_bind (rand(_) _)%E. *) (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) (* iDestruct (hocap_tapes_agree' with "[$][$]") as "%". *) @@ -255,7 +251,7 @@ Section impl2. (* replace (#(z+n)) with (#(z+n)%nat); last first. *) (* { by rewrite Nat2Z.inj_add. } *) (* replace 2%Z with (Z.of_nat 2%nat) by done. *) - (* rewrite -Nat2Z.inj_mul -Nat2Z.inj_add -Nat.div2_odd -Nat2Z.inj_add. *) + (* rewrite -Nat2Z.inj_mul -Nat2Z.inj_add -Nat.div2_odd -Nat2Z.inj_add. *) (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 H5 $H6]") as "_"; first by iFrame. *) (* iModIntro. wp_pures. *) (* iApply "HΦ". *) @@ -268,10 +264,10 @@ Section impl2. ↑N ⊆ E -> (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> (∀ (ε:R), 0<= ε ->SeriesC (λ n, if (bool_decide (n≤3%nat)) then 1 / (S 3%nat) * ε2 ε n else 0%R)%R <= ε)%R-> - (∃ γ1' γ2', is_flip (L:=L) (N.@"flip") γ1' γ2' ∗ - inv (N.@ "counter") (counter_inv_pred2 c γ1 γ1' γ2 γ2' γ3)) -∗ - (□∀ (ε:R) n, (P ∗ ●↯ ε@ γ1) ={E∖↑N}=∗ - (⌜(1<=ε2 ε n)%R⌝ ∨(●↯ (ε2 ε n) @ γ1 ∗ T (n)))) + (∃ γ2', is_flip (L:=L) (N.@"flip") γ1 γ2' ∗ + inv (N.@ "counter") (counter_inv_pred2 c γ2 γ2' γ3)) -∗ + (□∀ (ε:R) n, (P ∗ flip_error_auth (L:=L) γ1 ε) ={E∖↑N}=∗ + (⌜(1<=ε2 ε n)%R⌝ ∨(flip_error_auth (L:=L) γ1 (ε2 ε n) ∗ T (n)))) -∗ P -∗ α ◯↪N (true, ns) @ γ2 -∗ wp_update E (∃ n, T (n) ∗ α◯↪N (true, ns++[n]) @ γ2). @@ -363,8 +359,8 @@ Section impl2. Lemma read_counter_spec2 N E c γ1 γ2 γ3 P Q: ↑N ⊆ E -> - {{{ (∃ γ1' γ2', is_flip (L:=L) (N.@"flip") γ1' γ2' ∗ - inv (N.@ "counter") (counter_inv_pred2 c γ1 γ1' γ2 γ2' γ3)) ∗ + {{{ (∃ γ2', is_flip (L:=L) (N.@"flip") γ1 γ2' ∗ + inv (N.@ "counter") (counter_inv_pred2 c γ2 γ2' γ3)) ∗ □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ own γ3 (●F z)∗ Q z) ∗ P @@ -392,13 +388,13 @@ Program Definition random_counter2 `{!conerisGS Σ, flip_spec Σ}: random_counte incr_counter_tape := incr_counter_tape2; read_counter:=read_counter2; counterG := counterG2; - error_name := gname; + error_name := flip_error_name; tape_name := gname; counter_name :=gname; - is_counter _ N c γ1 γ2 γ3 := (∃ γ1' γ2', is_flip (N.@"flip") γ1' γ2' ∗ - inv (N.@ "counter") (counter_inv_pred2 (L:=_) c γ1 γ1' γ2 γ2' γ3))%I; - counter_error_auth _ γ x := ●↯ x @ γ; - counter_error_frag _ γ x := ◯↯ x @ γ; + is_counter _ N c γ1 γ2 γ3 := (∃ γ2', is_flip (N.@"flip") γ1 γ2' ∗ + inv (N.@ "counter") (counter_inv_pred2 (L:=_) c γ2 γ2' γ3))%I; + counter_error_auth _ γ x := flip_error_auth γ x; + counter_error_frag _ γ x := flip_error_frag γ x; counter_tapes_auth _ γ m := (●((λ ns, (true, ns))<$>m)@γ)%I; counter_tapes_frag _ γ α ns := (α◯↪N (true, ns) @ γ)%I; counter_content_auth _ γ z := own γ (●F z); @@ -409,35 +405,35 @@ Program Definition random_counter2 `{!conerisGS Σ, flip_spec Σ}: random_counte counter_presample_spec _ :=counter_presample_spec2; read_counter_spec _ :=read_counter_spec2 |}. -Next Obligation. - simpl. - iIntros (????????) "(%&<-&H1)(%&<-&H2)". - iCombine "H1 H2" gives "%K". by rewrite excl_auth.excl_auth_auth_op_valid in K. -Qed. -Next Obligation. - simpl. - iIntros (????????) "(%&<-&H1)(%&<-&H2)". - iCombine "H1 H2" gives "%K". by rewrite excl_auth.excl_auth_frag_op_valid in K. -Qed. -Next Obligation. - simpl. - iIntros (???????) "H". - iApply (hocap_error_auth_pos with "[$]"). -Qed. -Next Obligation. - simpl. - iIntros (???????) "H". - iApply (hocap_error_frag_pos with "[$]"). -Qed. -Next Obligation. - simpl. - iIntros (????????) "H1 H2". - iApply (hocap_error_agree with "[$][$]"). -Qed. -Next Obligation. - simpl. iIntros (?????????) "??". - iApply (hocap_error_update with "[$][$]"). -Qed. +(* Next Obligation. *) +(* simpl. *) +(* iIntros (????????) "(%&<-&H1)(%&<-&H2)". *) +(* iCombine "H1 H2" gives "%K". by rewrite excl_auth.excl_auth_auth_op_valid in K. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (????????) "(%&<-&H1)(%&<-&H2)". *) +(* iCombine "H1 H2" gives "%K". by rewrite excl_auth.excl_auth_frag_op_valid in K. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (???????) "H". *) +(* iApply (hocap_error_auth_pos with "[$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (???????) "H". *) +(* iApply (hocap_error_frag_pos with "[$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (????????) "H1 H2". *) +(* iApply (hocap_error_agree with "[$][$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. iIntros (?????????) "??". *) +(* iApply (hocap_error_update with "[$][$]"). *) +(* Qed. *) Next Obligation. simpl. iIntros (????????) "H1 H2". From 4f74c4fe7950fd4f7cdea78f4caacee63ad7b80f Mon Sep 17 00:00:00 2001 From: Hei Li Date: Thu, 19 Sep 2024 10:18:50 +0200 Subject: [PATCH 09/69] allocate_tape_spec2 --- .../coneris/examples/random_counter/impl2.v | 41 +++++++++---------- 1 file changed, 19 insertions(+), 22 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index bcf09776..7e0464f4 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -19,7 +19,7 @@ Notation "α ◯↪N ( b , ns ) @ γ":= (α ↪[ γ ] (b, ns))%I Notation "● m @ γ" := (ghost_map_auth γ 1 m) (at level 20) : bi_scope. Section tapes_lemmas. - Context `{!conerisGS Σ, !hocap_tapesGS' Σ}. + Context `{!hocap_tapesGS' Σ}. Lemma hocap_tapes_alloc' m: ⊢ |==>∃ γ, (● m @ γ) ∗ [∗ map] k↦v ∈ m, (k ◯↪N (v.1, v.2) @ γ). @@ -65,14 +65,14 @@ Section tapes_lemmas. iApply (ghost_map_update with "[$][$]"). Qed. - Lemma hocap_tapes_notin' α N ns m (f:(bool*list nat)-> nat) g: - α ↪N (N; ns) -∗ ([∗ map] α0↦t ∈ m, α0 ↪N (f t; g t)) -∗ ⌜m!!α=None ⌝. + Lemma hocap_tapes_notin' `{F: flip_spec Σ} {L:flipG Σ} α ns m γ (f:(bool*list nat)-> (list bool)) : + flip_tapes_frag (L:=L) γ α ns -∗ ([∗ map] α0↦t ∈ m, flip_tapes_frag (L:=L) γ α0 (f t)) -∗ ⌜m!!α=None ⌝. Proof. destruct (m!!α) eqn:Heqn; last by iIntros. iIntros "Hα Hmap". iDestruct (big_sepM_lookup with "[$]") as "?"; first done. iExFalso. - iApply (tapeN_tapeN_contradict with "[$][$]"). + by iDestruct (flip_tapes_frag_exclusive with "[$][$]") as "%". Qed. End tapes_lemmas. @@ -170,25 +170,22 @@ Section impl2. ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ α ◯↪N (true, []) @ γ2 }}}. Proof. - iIntros (Hsubset Φ) "#Hinv HΦ". + iIntros (Hsubset Φ) "(%γ2' & #Hinv & #Hinv') HΦ". rewrite /allocate_tape2. - iApply fupd_pgl_wp. - iApply (flip_allocate_tape_spec). - (** Spec not strong enough *) - Admitted. - (* iInv (N.@"counter") as "(%ε & %m & %l & %z & %γ1' & %γ2' & #Hinv' & >H1 & >H2 & >H3 & >H4 & > (-> & H5 & H6))" "Hclose". *) - (* wp_alloctape α as "Hα". *) - (* iDestruct (hocap_tapes_notin' with "[$][$]") as "%". *) - (* iMod (hocap_tapes_new' _ _ _ _ true with "[$]") as "[H4 H7]"; first done. *) - (* replace ([]) with (expander []) by done. *) - (* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Hα]") as "_". *) - (* { iNext. iSplitL; last done. *) - (* rewrite big_sepM_insert; [iFrame|done]. *) - (* } *) - (* iApply "HΦ". *) - (* by iFrame. *) - (* Qed. *) - + iApply wptac_wp_fupd. + iApply (flip_allocate_tape_spec); [by eapply nclose_subseteq'|done|]. + iNext. + iIntros (?) "(%α & -> & Hfrag)". + iInv "Hinv'" as ">(%m&%l&%z &Hfrags & Hauth & Hrest)" "Hclose". + iDestruct (hocap_tapes_notin' with "[$][$]") as "%". + iMod (hocap_tapes_new' _ _ _ [] true with "[$]") as "[Hauth H]"; first done. + iMod ("Hclose" with "[$Hrest $Hauth Hfrags Hfrag]") as "_". + { iNext. rewrite big_sepM_insert; last done. + iFrame. } + iApply "HΦ". + by iFrame. + Qed. + Lemma incr_counter_tape_spec_some2 N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: ↑N⊆E -> {{{ (∃ γ2', is_flip (L:=L) (N.@"flip") γ1 γ2' ∗ From 2248da2d19fcdd8c3bfe7d7da157de8395a159de Mon Sep 17 00:00:00 2001 From: Hei Li Date: Thu, 19 Sep 2024 10:31:55 +0200 Subject: [PATCH 10/69] read_counter_spec2 --- .../coneris/examples/random_counter/impl2.v | 28 +++++++++++-------- 1 file changed, 16 insertions(+), 12 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index 7e0464f4..8d6f6bae 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -86,10 +86,10 @@ Section impl2. (* (FAA "l" "x", "x"). *) Definition allocate_tape2 : val := flip_allocate_tape. Definition incr_counter_tape2 :val := λ: "l" "α", let: "n" := - conversion.bool_to_int (flip_allocate_tape "α") + conversion.bool_to_int (flip_tape "α") in let: "n'" := - conversion.bool_to_int (flip_allocate_tape "α") + conversion.bool_to_int (flip_tape "α") in let: "x" := #2 * "n" + "n'" in (FAA "l" "x", "x"). @@ -200,6 +200,8 @@ Section impl2. iIntros (Hsubset Φ) "((%γ2' & #Hinv) & #Hvs & HP & Hα) HΦ". rewrite /incr_counter_tape2. wp_pures. + wp_apply (flip_tape_spec_some). + (** spec not strong enough !*) Admitted. (* wp_bind (rand(_) _)%E. *) (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) @@ -366,16 +368,18 @@ Section impl2. {{{ (n':nat), RET #n'; Q n' }}}. Proof. - Admitted. - (* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP) HΦ". *) - (* rewrite /read_counter2. *) - (* wp_pure. *) - (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) - (* wp_load. *) - (* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) - (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]"); first done. *) - (* iApply ("HΦ" with "[$]"). *) - (* Qed. *) + iIntros (Hsubset Φ) "((%γ2' & #Hinv & #Hinv') & #Hvs & HP) HΦ". + rewrite /read_counter2. + wp_pure. + iInv "Hinv'" as ">(%m&%l&%z &Hfrags & Hauth & -> & Hloc & Hcont)" "Hclose". + wp_load. + iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". + { apply difference_mono_l. by apply nclose_subseteq'. } + iMod ("Hvs" with "[$]") as "[Hcont HQ]". + iMod "Hclose'". + iMod ("Hclose" with "[$Hfrags $Hauth $Hloc $Hcont]"); first done. + iApply "HΦ". by iFrame. + Qed. End impl2. From a22771332d8e7f25cd9a57ced545d8ed52237fbc Mon Sep 17 00:00:00 2001 From: Hei Li Date: Thu, 19 Sep 2024 10:36:43 +0200 Subject: [PATCH 11/69] NIT --- theories/coneris/examples/random_counter/impl2.v | 11 ++++++----- 1 file changed, 6 insertions(+), 5 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index 8d6f6bae..b8e340e0 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -271,8 +271,9 @@ Section impl2. P -∗ α ◯↪N (true, ns) @ γ2 -∗ wp_update E (∃ n, T (n) ∗ α◯↪N (true, ns++[n]) @ γ2). Proof. + iIntros (Hsubset Hpos Hineq) "(%γ2' & #Hinv & #Hinv') #Hvs HP Hfrag". Admitted. - (* iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) + (* iApply flip_presample_spec. *) (* iApply wp_update_state_step_coupl. *) (* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) (* iMod (inv_acc with "Hinv") as "[>(% & % & % & % & H1 & H2 & H3 & H4 & -> & H5 & H6) Hclose]"; [done|]. *) @@ -338,22 +339,22 @@ Section impl2. (* destruct xs as [|x' xs]; first done. *) (* destruct xs as [|]; last done. *) (* pose proof fin_to_nat_lt x. pose proof fin_to_nat_lt x'. *) - (* repeat f_equal; apply fin_to_nat_inj; rewrite fin_to_nat_to_fin. *) + (* repeat f_equal; apply fin_to_nat_inj; rewrite fin_to_nat_to_fin. *) (* - rewrite /f. rewrite Nat.div2_div. *) (* rewrite Nat.mul_comm Nat.div_add_l; last lia. rewrite Nat.div_small; lia. *) (* - rewrite /f. rewrite Nat.add_comm Nat.odd_add_even. *) (* + destruct (fin_to_nat x') as [|[|]]; simpl; lia. *) - (* + by econstructor. *) + (* + by econstructor. *) (* } *) (* iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4 $H5 $H6]") as "_". *) (* { iNext. iSplit; last done. *) (* rewrite big_sepM_insert_delete; iFrame. *) (* simpl. rewrite !fin_to_nat_to_fin. *) - (* rewrite /expander bind_app -/(expander ns). simpl. by rewrite H1. *) + (* rewrite /expander bind_app -/(expander ns). simpl. by rewrite H1. *) (* } *) (* iApply fupd_mask_intro_subseteq; first set_solver. *) (* rewrite K. *) - (* iFrame. *) + (* iFrame. *) (* Qed. *) Lemma read_counter_spec2 N E c γ1 γ2 γ3 P Q: From e1027da8f076d27eb14a965848637f160a427b75 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Thu, 19 Sep 2024 12:25:51 +0200 Subject: [PATCH 12/69] strengthen flip_tape_spec_some --- theories/coneris/lib/flip_spec.v | 96 +++++++++++++++----------------- 1 file changed, 46 insertions(+), 50 deletions(-) diff --git a/theories/coneris/lib/flip_spec.v b/theories/coneris/lib/flip_spec.v index ca455555..4854dfb9 100644 --- a/theories/coneris/lib/flip_spec.v +++ b/theories/coneris/lib/flip_spec.v @@ -77,11 +77,13 @@ Class flip_spec `{!conerisGS Σ} := FlipSpec flip_tape_spec_some {L: flipG Σ} N E γ1 γ2 (P: iProp Σ) (Q:iProp Σ) (α:loc) (n:bool) ns: ↑N⊆E -> {{{ is_flip (L:=L) N γ1 γ2 ∗ - □ (P ={E∖↑N}=∗ Q) ∗ - P ∗ flip_tapes_frag (L:=L) γ2 α (n::ns) + □ (∀ m, P ∗ + flip_tapes_auth (L:=L) γ2 m + ={E∖↑N}=∗ ⌜m!!α=Some (n::ns)⌝ ∗ Q ∗ flip_tapes_auth (L:=L) γ2 (<[α:=ns]> m)) ∗ + P }}} flip_tape #lbl:α @ E - {{{ RET #n; Q ∗ flip_tapes_frag (L:=L) γ2 α ns}}}; + {{{ RET #n; Q }}}; flip_presample_spec {L: flipG Σ} NS E ns α (ε2 : R -> bool -> R) (P : iProp Σ) T γ1 γ2: @@ -97,40 +99,15 @@ Class flip_spec `{!conerisGS Σ} := FlipSpec }. -Section lemmas. - Context `{rc:flip_spec} {L: flipG Σ}. - - Lemma flip_tape_spec_none N E γ1 γ2 (ε2:R -> bool -> R) (P: iProp Σ) (T: bool -> iProp Σ) (Q: bool -> iProp Σ)(α:loc): - ↑N ⊆ E-> - (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> - (∀ (ε:R), 0<=ε -> ((ε2 ε true) + (ε2 ε false))/2 <= ε)%R → - {{{ is_flip (L:=L) N γ1 γ2 ∗ - □(∀ (ε:R) (n : bool), P ∗ flip_error_auth (L:=L) γ1 ε - ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨ flip_error_auth (L:=L) γ1 (ε2 ε n) ∗ T n) ) ∗ - □ (∀ (n:bool), T n ={E∖↑N}=∗ Q n) ∗ - P ∗ flip_tapes_frag (L:=L) γ2 α [] - }}} - flip_tape #lbl:α @ E - {{{ (b:bool), RET #b; Q b ∗ flip_tapes_frag (L:=L) γ2 α [] }}}. - Proof. - iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP & Hα) HΦ". - iMod (flip_presample_spec with "[//][//][$][$]") as "(%&HT&Hα)"; [done..|]. - iApply (flip_tape_spec_some _ _ _ _ (T n) with "[$Hα $HT]"); try done. - { by iSplit. } - by iNext. - Qed. - -End lemmas. - - (** Instantiate flip *) Class flipG1 Σ := FlipG1 { flipG1_error::hocap_errorGS Σ; flipG1_tapes:: hocap_tapesGS Σ; }. Local Definition flip_inv_pred1 `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapesGS Σ} γ1 γ2:= - (∃ (ε:R) (m:gmap loc (nat * list nat)) , - ↯ ε ∗ ●↯ ε @ γ1 ∗ - ([∗ map] α ↦ t ∈ m, α ↪N ( t.1 ; t.2 )) ∗ ●m@γ2 + (∃ (ε:R) (m:gmap loc (list bool)) , + ↯ ε ∗ ●↯ ε @ γ1 ∗ + ([∗ map] α ↦ t ∈ ((λ (ns:list bool), (1%nat, bool_to_nat<$>ns))<$>m), α ↪N ( t.1 ; t.2 )) ∗ + ●(((λ (ns:list bool), (1, bool_to_nat<$>ns))<$>m):gmap _ _)@γ2 )%I. #[local] Program Definition flip_spec1 `{!conerisGS Σ}: flip_spec := @@ -215,7 +192,10 @@ Next Obligation. simpl. iMod (hocap_tapes_alloc ∅) as "(%γ2 & H4 & H5)". iMod (inv_alloc _ _ (flip_inv_pred1 γ1 γ2) with "[H1 H2 H4]"). - { iFrame. by iNext. } + { iFrame. iNext. iExists ∅. + iFrame. + by rewrite fmap_empty. + } by iFrame. Qed. Next Obligation. @@ -229,7 +209,10 @@ Next Obligation. iDestruct (hocap_tapes_notin with "[$][$]") as "%". iMod (hocap_tapes_new _ _ α 1%nat [] with "[$]") as "[?H]"; first done. iMod ("Hclose" with "[-H HΦ]"). - { iFrame. iModIntro. + { iModIntro. + iExists _, (<[α:=([])]> m). + iFrame. + rewrite fmap_insert. rewrite big_sepM_insert; [iFrame|done]. } iApply "HΦ". @@ -237,25 +220,27 @@ Next Obligation. Qed. Next Obligation. simpl. - iIntros (????????????? Φ) "(#Hinv & #Hvs & HP & Hfrag) HΦ". + iIntros (????????????? Φ) "(#Hinv & #Hvs & HP) HΦ". rewrite /flipL. wp_pures. wp_bind (rand(_) _)%E. iInv "Hinv" as ">(%&%&?&?&H3&?)" "Hclose". - iDestruct (hocap_tapes_agree with "[$][$]") as "%". + iMod ("Hvs" with "[$]") as "(%&HQ&Hauth)". erewrite <-(insert_delete m) at 1; last done. - rewrite big_sepM_insert; last apply lookup_delete. + rewrite fmap_insert. + rewrite big_sepM_insert; last first. + { rewrite fmap_delete. apply lookup_delete. } iDestruct "H3" as "[Htape H3]". simpl. wp_apply (wp_rand_tape with "[$]") as "[Htape %]". - iMod ("Hvs" with "[$]") as "HQ". - iMod (hocap_tapes_pop with "[$][$]") as "[H4 Hfrag]". - iMod ("Hclose" with "[-Hfrag HΦ HQ]") as "_". - { iFrame. + iMod ("Hclose" with "[- HΦ HQ]") as "_". + { iExists _, (<[α:=ns]> m). + iFrame. rewrite <-(insert_delete m) at 2; last done. - iNext. - rewrite insert_insert. - rewrite big_sepM_insert; last apply lookup_delete. iFrame. + rewrite insert_insert fmap_insert. + rewrite big_sepM_insert; first iFrame. + rewrite fmap_delete. + apply lookup_delete. } iModIntro. wp_apply conversion.wp_int_to_bool as "_"; first done. @@ -264,16 +249,22 @@ Next Obligation. - destruct n; simpl. + rewrite Z_to_bool_neq_0; [done|lia]. + by rewrite Z_to_bool_eq_0. -Qed. +Qed. Next Obligation. simpl. iIntros (???????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". iApply wp_update_state_step_coupl. iIntros (σ ε) "((Hheap&Htapes)&Hε)". iMod (inv_acc with "Hinv") as "[>(% & % & H1 & H2 & H3 & H4 ) Hclose]"; [done|]. - iDestruct (hocap_tapes_agree with "[$][$]") as "%". + iDestruct (hocap_tapes_agree with "[$][$]") as "%K". + rewrite lookup_fmap_Some in K. destruct K as (?&M&?). + simplify_eq. + unshelve epose proof fmap_inj _ _ _ _ M as ->. + { intros [][]?; by simplify_eq. } erewrite <-(insert_delete m) at 1; last done. - rewrite big_sepM_insert; last apply lookup_delete. + rewrite fmap_insert. + rewrite big_sepM_insert; last first. + { rewrite fmap_delete. apply lookup_delete. } simpl. iDestruct "H3" as "[Htape H3]". iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". @@ -308,9 +299,14 @@ Next Obligation. iMod "Hclose'" as "_". iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". { iExFalso. iApply (ec_contradict with "[$]"). exact. } - iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4]") as "_". - { iNext. - rewrite big_sepM_insert_delete Heq/=; iFrame. + iMod ("Hclose" with "[$Hε $H2 Htape H3 H4]") as "_". + { iNext. + iExists (<[α:=(ns ++ [nat_to_bool sample])]>m). + rewrite fmap_insert. + rewrite big_sepM_insert_delete Heq/=. + rewrite fmap_delete. iFrame. + rewrite fmap_app/= nat_to_bool_to_nat; first iFrame. + pose proof fin_to_nat_lt sample. lia. } iApply fupd_mask_intro_subseteq; first set_solver. iFrame. From c618b3b07ec93116e105fbd8816488d5deb285ac Mon Sep 17 00:00:00 2001 From: Hei Li Date: Thu, 19 Sep 2024 15:38:43 +0200 Subject: [PATCH 13/69] Progress with incr counter tape spec some 2 --- .../coneris/examples/random_counter/impl2.v | 43 ++++++++++++------- theories/coneris/lib/flip_spec.v | 10 +++++ 2 files changed, 38 insertions(+), 15 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index b8e340e0..06a3516a 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -197,22 +197,34 @@ Section impl2. incr_counter_tape2 c #lbl:α @ E {{{ (z:nat), RET (#z, #n); Q z ∗ α ◯↪N (true, ns) @ γ2}}}. Proof. - iIntros (Hsubset Φ) "((%γ2' & #Hinv) & #Hvs & HP & Hα) HΦ". + iIntros (Hsubset Φ) "((%γ2' & #Hinv & #Hinv') & #Hvs & HP & Hα) HΦ". rewrite /incr_counter_tape2. wp_pures. - wp_apply (flip_tape_spec_some). - (** spec not strong enough !*) - Admitted. - (* wp_bind (rand(_) _)%E. *) - (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) - (* iDestruct (hocap_tapes_agree' with "[$][$]") as "%". *) - (* erewrite <-(insert_delete m) at 1; last done. *) - (* rewrite big_sepM_insert; last apply lookup_delete. *) - (* simpl. *) - (* iDestruct "H3" as "[Htape H3]". *) - (* wp_apply (wp_rand_tape with "[$]"). *) - (* iIntros "[Htape %H1]". *) - (* iMod (hocap_tapes_pop1' with "[$][$]") as "[H4 Hα]". *) + wp_apply (flip_tape_spec_some _ _ _ _ (α ◯↪N (true, n :: ns)@ γ2) (α ◯↪N (false, n :: ns)@ γ2)with "[$Hα]"); [|iSplit; first done|]. + { by apply nclose_subseteq'. } + { iModIntro. iIntros (m) "[Hfrag Hauth]". + iInv "Hinv'" as ">(%m'&%&%&H1&H2&?&?&?)" "Hclose". + iDestruct (hocap_tapes_agree' with "[$][$]") as "%". + iDestruct (flip_tapes_agree with "[$][H1]") as "%". + { iDestruct (big_sepM_lookup with "[$]") as "?"; first done. + iFrame. + } + erewrite <-(insert_delete m') at 1; last done. + rewrite big_sepM_insert; last apply lookup_delete. + simpl. + iDestruct "H1" as "[Htape H1]". + iMod (hocap_tapes_pop1' with "[$][$]") as "[H4 Hα]". + iMod (flip_tapes_pop with "[$][$]") as "[Hauth Htape]". + iMod ("Hclose" with "[-Hα Hauth]"). + - iExists (<[_:=_]> _). iFrame. + iNext. + erewrite <-(insert_delete m') at 2; last done. + rewrite insert_insert. + rewrite big_sepM_insert; [iFrame|apply lookup_delete]. + - iModIntro. simpl in *. iSplit; first done. iFrame. + } + iIntros "Hα". + Admitted. (* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". *) (* { iSplitL; last done. *) (* erewrite <-(insert_delete m) at 2; last done. *) @@ -255,7 +267,7 @@ Section impl2. (* iModIntro. wp_pures. *) (* iApply "HΦ". *) (* by iFrame. *) - (* Qed. *) + (* Qed. *) Lemma counter_presample_spec2 N E ns α (ε2 : R -> nat -> R) @@ -272,6 +284,7 @@ Section impl2. wp_update E (∃ n, T (n) ∗ α◯↪N (true, ns++[n]) @ γ2). Proof. iIntros (Hsubset Hpos Hineq) "(%γ2' & #Hinv & #Hinv') #Hvs HP Hfrag". + iApply fupd_wp_update. Admitted. (* iApply flip_presample_spec. *) (* iApply wp_update_state_step_coupl. *) diff --git a/theories/coneris/lib/flip_spec.v b/theories/coneris/lib/flip_spec.v index 4854dfb9..608a2403 100644 --- a/theories/coneris/lib/flip_spec.v +++ b/theories/coneris/lib/flip_spec.v @@ -60,6 +60,9 @@ Class flip_spec `{!conerisGS Σ} := FlipSpec flip_tapes_update {L : flipG Σ} γ α m ns n: flip_tapes_auth (L:=L) γ m -∗ flip_tapes_frag (L:=L) γ α ns ==∗ flip_tapes_auth (L:=L) γ (<[α := (ns ++[n])]> m) ∗ flip_tapes_frag (L:=L) γ α (ns ++ [n]); + flip_tapes_pop {L : flipG Σ} γ α m ns n: + flip_tapes_auth (L:=L) γ m -∗ flip_tapes_frag (L:=L) γ α (n::ns) ==∗ + flip_tapes_auth (L:=L) γ (<[α := (ns)]> m) ∗ flip_tapes_frag (L:=L) γ α (ns); (** * Program specs *) flip_inv_create_spec {L : flipG Σ} N E ε: @@ -182,6 +185,13 @@ Next Obligation. rewrite fmap_insert. rewrite fmap_app; iFrame. Qed. +Next Obligation. + iIntros. + iMod (hocap_tapes_pop with "[$][$]") as "[??]". + iModIntro. iFrame. + rewrite fmap_insert. + iFrame. +Qed. Next Obligation. simpl. iIntros (????? ε) "H1". From 0236bcc39b606bbfba31be38972b6edc7a207ec3 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Thu, 19 Sep 2024 15:50:40 +0200 Subject: [PATCH 14/69] progress with incr counter tape spec some 2 --- .../coneris/examples/random_counter/impl2.v | 90 +++++++++---------- 1 file changed, 45 insertions(+), 45 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index 06a3516a..a6601460 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -190,7 +190,7 @@ Section impl2. ↑N⊆E -> {{{ (∃ γ2', is_flip (L:=L) (N.@"flip") γ1 γ2' ∗ inv (N.@ "counter") (counter_inv_pred2 c γ2 γ2' γ3)) ∗ - □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ + □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N.@"counter"}=∗ own γ3 (●F(z+n)%nat)∗ Q z) ∗ P ∗ α ◯↪N (true, n::ns) @ γ2 }}} @@ -224,50 +224,50 @@ Section impl2. - iModIntro. simpl in *. iSplit; first done. iFrame. } iIntros "Hα". - Admitted. - (* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". *) - (* { iSplitL; last done. *) - (* erewrite <-(insert_delete m) at 2; last done. *) - (* iNext. *) - (* rewrite insert_insert. *) - (* rewrite big_sepM_insert; last apply lookup_delete. iFrame. *) - (* } *) - (* iModIntro. *) - (* wp_pures. *) - (* clear -Hsubset H1. *) - (* wp_bind (rand(_) _)%E. *) - (* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) - (* iDestruct (hocap_tapes_agree' with "[$][$]") as "%". *) - (* erewrite <-(insert_delete m) at 1; last done. *) - (* rewrite big_sepM_insert; last apply lookup_delete. *) - (* simpl. *) - (* iDestruct "H3" as "[Htape H3]". *) - (* wp_apply (wp_rand_tape with "[$]"). *) - (* iIntros "[Htape %H2]". *) - (* iMod (hocap_tapes_pop2' with "[$][$]") as "[H4 Hα]". *) - (* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". *) - (* { iSplitL; last done. *) - (* erewrite <-(insert_delete m) at 2; last done. *) - (* iNext. *) - (* rewrite insert_insert. *) - (* rewrite big_sepM_insert; last apply lookup_delete. iFrame. *) - (* } *) - (* iModIntro. *) - (* wp_pures. *) - (* clear -Hsubset H1 H2. *) - (* wp_bind (FAA _ _). *) - (* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) - (* wp_faa. *) - (* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) - (* replace (#(z+n)) with (#(z+n)%nat); last first. *) - (* { by rewrite Nat2Z.inj_add. } *) - (* replace 2%Z with (Z.of_nat 2%nat) by done. *) - (* rewrite -Nat2Z.inj_mul -Nat2Z.inj_add -Nat.div2_odd -Nat2Z.inj_add. *) - (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 H5 $H6]") as "_"; first by iFrame. *) - (* iModIntro. wp_pures. *) - (* iApply "HΦ". *) - (* by iFrame. *) - (* Qed. *) + wp_apply (conversion.wp_bool_to_int) as "_"; first done. + wp_pures. + wp_apply (flip_tape_spec_some _ _ _ _ (α ◯↪N (false, n :: ns)@ γ2) (α ◯↪N (true, ns)@ γ2)with "[$Hα]"); [|iSplit; first done|]. + { by apply nclose_subseteq'. } + { iModIntro. iIntros (m) "[Hfrag Hauth]". + iInv "Hinv'" as ">(%m'&%&%&H1&H2&?&?&?)" "Hclose". + iDestruct (hocap_tapes_agree' with "[$][$]") as "%". + iDestruct (flip_tapes_agree with "[$][H1]") as "%". + { iDestruct (big_sepM_lookup with "[$]") as "?"; first done. + iFrame. + } + erewrite <-(insert_delete m') at 1; last done. + rewrite big_sepM_insert; last apply lookup_delete. + simpl in *. + iDestruct "H1" as "[Htape H1]". + iMod (hocap_tapes_pop2' with "[$][$]") as "[H4 Hα]". + iMod (flip_tapes_pop with "[$][$]") as "[Hauth Htape]". + iMod ("Hclose" with "[-Hα Hauth]"). + - iExists (<[_:=_]> _). iFrame. + iNext. + erewrite <-(insert_delete m') at 2; last done. + rewrite insert_insert. + rewrite big_sepM_insert; [iFrame|apply lookup_delete]. + - iModIntro. simpl in *. iSplit; first done. iFrame. + } + iIntros "Hα". + wp_apply (conversion.wp_bool_to_int) as "_"; first done. + wp_pures. + wp_bind (FAA _ _). + iInv "Hinv'" as ">(%&%&%&?&?&->&?&?)" "Hclose". + wp_faa. simpl. + iMod ("Hvs" with "[$]") as "[H6 HQ]". + replace (#(z+n)) with (#(z+n)%nat); last first. + { by rewrite Nat2Z.inj_add. } + replace 2%Z with (Z.of_nat 2%nat) by done. + replace (_*_+_)%Z with (Z.of_nat n); last first. + { admit. (* rewrite -Nat2Z.inj_mul -Nat2Z.inj_add -Nat.div2_odd -Nat2Z.inj_add. *) } + iMod ("Hclose" with "[-HΦ HQ Hα]") as "_"; first iFrame. + { admit. } + iModIntro. + wp_pures. + iApply "HΦ". + by iFrame. + Admitted. Lemma counter_presample_spec2 N E ns α (ε2 : R -> nat -> R) From b602e83098b3ba38b8af797f0cbc229ed5697c1d Mon Sep 17 00:00:00 2001 From: Hei Li Date: Fri, 20 Sep 2024 10:20:41 +0200 Subject: [PATCH 15/69] impl2 refaactor --- .../coneris/examples/random_counter/impl2.v | 345 +++++++----------- 1 file changed, 142 insertions(+), 203 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index a6601460..21eb9cff 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -5,77 +5,84 @@ From clutch Require Import uniform_list. Set Default Proof Using "Type*". + Local Definition expander (l:list nat):= - l ≫= (λ x, [Nat.div2 x; Nat.b2n (Nat.odd x)]). + l ≫= (λ x, [2<=?x; (Nat.odd x)]). +Local Lemma expander_eta x: x<4->(Z.of_nat x= Z.of_nat 2*Z.b2z (2<=? x)%nat + Z.b2z (Nat.odd x))%Z. +Proof. + destruct x as [|[|[|[|]]]]; last lia; intros; simpl; lia. +Qed. +(* Local Definition expander (l:list nat):= *) +(* l ≫= (λ x, [Nat.div2 x; Nat.b2n (Nat.odd x)]). *) -Class hocap_tapesGS' (Σ : gFunctors) := Hocap_tapesGS' { - hocap_tapesGS_inG' :: ghost_mapG Σ loc (bool* list nat) - }. -Definition hocap_tapesΣ' := ghost_mapΣ loc (bool *list nat). +(* Class hocap_tapesGS' (Σ : gFunctors) := Hocap_tapesGS' { *) +(* hocap_tapesGS_inG' :: ghost_mapG Σ loc (bool* list nat) *) +(* }. *) +(* Definition hocap_tapesΣ' := ghost_mapΣ loc (bool *list nat). *) -Notation "α ◯↪N ( b , ns ) @ γ":= (α ↪[ γ ] (b, ns))%I - (at level 20) : bi_scope. +(* Notation "α ◯↪N ( b , ns ) @ γ":= (α ↪[ γ ] (b, ns))%I *) +(* (at level 20) : bi_scope. *) -Notation "● m @ γ" := (ghost_map_auth γ 1 m) (at level 20) : bi_scope. +(* Notation "● m @ γ" := (ghost_map_auth γ 1 m) (at level 20) : bi_scope. *) -Section tapes_lemmas. - Context `{!hocap_tapesGS' Σ}. +(* Section tapes_lemmas. *) +(* Context `{!hocap_tapesGS' Σ}. *) - Lemma hocap_tapes_alloc' m: - ⊢ |==>∃ γ, (● m @ γ) ∗ [∗ map] k↦v ∈ m, (k ◯↪N (v.1, v.2) @ γ). - Proof. - iMod ghost_map_alloc as (γ) "[??]". - iFrame. iModIntro. - iApply big_sepM_mono; last done. - by iIntros (?[??]). - Qed. +(* Lemma hocap_tapes_alloc' m: *) +(* ⊢ |==>∃ γ, (● m @ γ) ∗ [∗ map] k↦v ∈ m, (k ◯↪N (v.1, v.2) @ γ). *) +(* Proof. *) +(* iMod ghost_map_alloc as (γ) "[??]". *) +(* iFrame. iModIntro. *) +(* iApply big_sepM_mono; last done. *) +(* by iIntros (?[??]). *) +(* Qed. *) - Lemma hocap_tapes_agree' m b γ k ns: - (● m @ γ) -∗ (k ◯↪N (b, ns) @ γ) -∗ ⌜ m!!k = Some (b, ns) ⌝. - Proof. - iIntros "H1 H2". - by iCombine "H1 H2" gives "%". - Qed. +(* Lemma hocap_tapes_agree' m b γ k ns: *) +(* (● m @ γ) -∗ (k ◯↪N (b, ns) @ γ) -∗ ⌜ m!!k = Some (b, ns) ⌝. *) +(* Proof. *) +(* iIntros "H1 H2". *) +(* by iCombine "H1 H2" gives "%". *) +(* Qed. *) - Lemma hocap_tapes_new' γ m k ns b: - m!!k=None -> ⊢ (● m @ γ) ==∗ (● (<[k:=(b, ns)]>m) @ γ) ∗ (k ◯↪N (b, ns) @ γ). - Proof. - iIntros (Hlookup) "H". - by iApply ghost_map_insert. - Qed. +(* Lemma hocap_tapes_new' γ m k ns b: *) +(* m!!k=None -> ⊢ (● m @ γ) ==∗ (● (<[k:=(b, ns)]>m) @ γ) ∗ (k ◯↪N (b, ns) @ γ). *) +(* Proof. *) +(* iIntros (Hlookup) "H". *) +(* by iApply ghost_map_insert. *) +(* Qed. *) - Lemma hocap_tapes_presample' b γ m k ns n: - (● m @ γ) -∗ (k ◯↪N (b, ns) @ γ) ==∗ (● (<[k:=(b, ns++[n])]>m) @ γ) ∗ (k ◯↪N (b, ns++[n]) @ γ). - Proof. - iIntros "H1 H2". - iApply (ghost_map_update with "[$][$]"). - Qed. +(* Lemma hocap_tapes_presample' b γ m k ns n: *) +(* (● m @ γ) -∗ (k ◯↪N (b, ns) @ γ) ==∗ (● (<[k:=(b, ns++[n])]>m) @ γ) ∗ (k ◯↪N (b, ns++[n]) @ γ). *) +(* Proof. *) +(* iIntros "H1 H2". *) +(* iApply (ghost_map_update with "[$][$]"). *) +(* Qed. *) - Lemma hocap_tapes_pop1' γ m k ns: - (● m @ γ) -∗ (k ◯↪N (true, ns) @ γ) ==∗ (● (<[k:=(false, ns)]>m) @ γ) ∗ (k ◯↪N (false, ns) @ γ). - Proof. - iIntros "H1 H2". - iApply (ghost_map_update with "[$][$]"). - Qed. +(* Lemma hocap_tapes_pop1' γ m k ns: *) +(* (● m @ γ) -∗ (k ◯↪N (true, ns) @ γ) ==∗ (● (<[k:=(false, ns)]>m) @ γ) ∗ (k ◯↪N (false, ns) @ γ). *) +(* Proof. *) +(* iIntros "H1 H2". *) +(* iApply (ghost_map_update with "[$][$]"). *) +(* Qed. *) - Lemma hocap_tapes_pop2' γ m k ns n: - (● m @ γ) -∗ (k ◯↪N (false, n::ns) @ γ) ==∗ (● (<[k:=(true, ns)]>m) @ γ) ∗ (k ◯↪N (true, ns) @ γ). - Proof. - iIntros "H1 H2". - iApply (ghost_map_update with "[$][$]"). - Qed. +(* Lemma hocap_tapes_pop2' γ m k ns n: *) +(* (● m @ γ) -∗ (k ◯↪N (false, n::ns) @ γ) ==∗ (● (<[k:=(true, ns)]>m) @ γ) ∗ (k ◯↪N (true, ns) @ γ). *) +(* Proof. *) +(* iIntros "H1 H2". *) +(* iApply (ghost_map_update with "[$][$]"). *) +(* Qed. *) - Lemma hocap_tapes_notin' `{F: flip_spec Σ} {L:flipG Σ} α ns m γ (f:(bool*list nat)-> (list bool)) : - flip_tapes_frag (L:=L) γ α ns -∗ ([∗ map] α0↦t ∈ m, flip_tapes_frag (L:=L) γ α0 (f t)) -∗ ⌜m!!α=None ⌝. - Proof. - destruct (m!!α) eqn:Heqn; last by iIntros. - iIntros "Hα Hmap". - iDestruct (big_sepM_lookup with "[$]") as "?"; first done. - iExFalso. - by iDestruct (flip_tapes_frag_exclusive with "[$][$]") as "%". - Qed. +(* Lemma hocap_tapes_notin' `{F: flip_spec Σ} {L:flipG Σ} α ns m γ (f:(bool*list nat)-> (list bool)) : *) +(* flip_tapes_frag (L:=L) γ α ns -∗ ([∗ map] α0↦t ∈ m, flip_tapes_frag (L:=L) γ α0 (f t)) -∗ ⌜m!!α=None ⌝. *) +(* Proof. *) +(* destruct (m!!α) eqn:Heqn; last by iIntros. *) +(* iIntros "Hα Hmap". *) +(* iDestruct (big_sepM_lookup with "[$]") as "?"; first done. *) +(* iExFalso. *) +(* by iDestruct (flip_tapes_frag_exclusive with "[$][$]") as "%". *) +(* Qed. *) -End tapes_lemmas. +(* End tapes_lemmas. *) Section impl2. Context `{F: flip_spec Σ}. @@ -94,27 +101,27 @@ Section impl2. let: "x" := #2 * "n" + "n'" in (FAA "l" "x", "x"). Definition read_counter2 : val := λ: "l", !"l". - Class counterG2 Σ := CounterG2 { counterG2_tapes:: hocap_tapesGS' Σ; + Class counterG2 Σ := CounterG2 { (* counterG2_tapes:: hocap_tapesGS' Σ; *) counterG2_frac_authR:: inG Σ (frac_authR natR); counterG2_flipG: flipG Σ }. - Context `{L:!flipG Σ, !hocap_tapesGS' Σ, !inG Σ (frac_authR natR)}. + Context `{L:!flipG Σ, !inG Σ (frac_authR natR)}. - Definition counter_inv_pred2 (c:val) γ1 γ1' γ2 := - (∃ (m:gmap loc (bool*list nat)) (l:loc) (z:nat), - ([∗ map] α ↦ t ∈ m, flip_tapes_frag (L:=L) γ1' α (if t.1:bool then nat_to_bool <$> (expander t.2) else nat_to_bool <$> (drop 1%nat (expander t.2))) ) - ∗ ●m@γ1 ∗ - ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ2 (●F z) + Definition counter_inv_pred2 (c:val) γ := + (∃ (* (m:gmap loc (bool*list nat)) *) (l:loc) (z:nat), + (* ([∗ map] α ↦ t ∈ m, flip_tapes_frag (L:=L) γ1' α (if t.1:bool then nat_to_bool <$> (expander t.2) else nat_to_bool <$> (drop 1%nat (expander t.2))) ) *) + (* ∗ ●m@γ1 ∗ *) + ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ (●F z) )%I. Lemma new_counter_spec2 E ε N: {{{ ↯ ε }}} new_counter2 #() @ E {{{ (c:val), RET c; - ∃ γ1 γ2 γ3, (∃ γ2', is_flip (L:=L) (N.@"flip") γ1 γ2' ∗ - inv (N.@ "counter") (counter_inv_pred2 c γ2 γ2' γ3)) ∗ + ∃ γ1 γ2 γ3, (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ + inv (N.@ "counter") (counter_inv_pred2 c γ3)) ∗ flip_error_frag (L:=L) γ1 ε ∗ own γ3 (◯F 0%nat) }}}. Proof. @@ -125,17 +132,14 @@ Section impl2. iMod (flip_inv_create_spec with "[$]") as "(%γ1'&%γ2'&#?&Hε)". wp_alloc l as "Hl". simpl. - iMod (hocap_tapes_alloc' (∅:gmap _ _)) as "[%γ2 [H3 H4]]". + (* iMod (hocap_tapes_alloc' (∅:gmap _ _)) as "[%γ2 [H3 H4]]". *) iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". { by apply frac_auth_valid. } replace (#0) with (#0%nat) by done. - iMod (inv_alloc (N.@"counter") _ (counter_inv_pred2 (#l) γ2 γ2' γ3) with "[$Hl $H3 $H5]") as "#Hinv". - { iNext. - iSplit; last done. - by iApply big_sepM_empty. } + iMod (inv_alloc (N.@"counter") _ (counter_inv_pred2 (#l) γ3) with "[$Hl $H5]") as "#Hinv"; first done. iApply "HΦ". iExists _, _, _. iModIntro. iFrame. - iExists _. by iSplit. + by iSplit. Qed. @@ -163,111 +167,89 @@ Section impl2. Lemma allocate_tape_spec2 N E c γ1 γ2 γ3: ↑N ⊆ E-> - {{{ (∃ γ2', is_flip (L:=L) (N.@"flip") γ1 γ2' ∗ - inv (N.@ "counter") (counter_inv_pred2 c γ2 γ2' γ3)) }}} + {{{ (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ + inv (N.@ "counter") (counter_inv_pred2 c γ3)) }}} allocate_tape2 #() @ E {{{ (v:val), RET v; - ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ α ◯↪N (true, []) @ γ2 + ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ (flip_tapes_frag (L:=L) γ2 α (expander []) ∗ ⌜Forall (λ x, x<4) []⌝) }}}. Proof. - iIntros (Hsubset Φ) "(%γ2' & #Hinv & #Hinv') HΦ". + iIntros (Hsubset Φ) "(#Hinv & #Hinv') HΦ". rewrite /allocate_tape2. iApply wptac_wp_fupd. iApply (flip_allocate_tape_spec); [by eapply nclose_subseteq'|done|]. iNext. iIntros (?) "(%α & -> & Hfrag)". - iInv "Hinv'" as ">(%m&%l&%z &Hfrags & Hauth & Hrest)" "Hclose". - iDestruct (hocap_tapes_notin' with "[$][$]") as "%". - iMod (hocap_tapes_new' _ _ _ [] true with "[$]") as "[Hauth H]"; first done. - iMod ("Hclose" with "[$Hrest $Hauth Hfrags Hfrag]") as "_". - { iNext. rewrite big_sepM_insert; last done. - iFrame. } iApply "HΦ". - by iFrame. + iFrame. + iPureIntro. + split; first done. + by apply Forall_nil. Qed. Lemma incr_counter_tape_spec_some2 N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: ↑N⊆E -> - {{{ (∃ γ2', is_flip (L:=L) (N.@"flip") γ1 γ2' ∗ - inv (N.@ "counter") (counter_inv_pred2 c γ2 γ2' γ3)) ∗ - □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N.@"counter"}=∗ + {{{ (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ + inv (N.@ "counter") (counter_inv_pred2 c γ3)) ∗ + □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ own γ3 (●F(z+n)%nat)∗ Q z) ∗ - P ∗ α ◯↪N (true, n::ns) @ γ2 + P ∗ (flip_tapes_frag (L:=L) γ2 α (expander (n::ns))∗ ⌜Forall (λ x, x<4) (n::ns)⌝) }}} incr_counter_tape2 c #lbl:α @ E - {{{ (z:nat), RET (#z, #n); Q z ∗ α ◯↪N (true, ns) @ γ2}}}. + {{{ (z:nat), RET (#z, #n); Q z ∗ (flip_tapes_frag (L:=L) γ2 α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝)}}}. Proof. - iIntros (Hsubset Φ) "((%γ2' & #Hinv & #Hinv') & #Hvs & HP & Hα) HΦ". + iIntros (Hsubset Φ) "((#Hinv & #Hinv') & #Hvs & HP & [Hα %]) HΦ". rewrite /incr_counter_tape2. wp_pures. - wp_apply (flip_tape_spec_some _ _ _ _ (α ◯↪N (true, n :: ns)@ γ2) (α ◯↪N (false, n :: ns)@ γ2)with "[$Hα]"); [|iSplit; first done|]. + wp_apply (flip_tape_spec_some _ _ _ _ (flip_tapes_frag γ2 α (expander (n::ns)) ) (flip_tapes_frag γ2 α (drop 1%nat (expander (n::ns))))with "[$Hα]"); [|iSplit; first done|]. { by apply nclose_subseteq'. } { iModIntro. iIntros (m) "[Hfrag Hauth]". - iInv "Hinv'" as ">(%m'&%&%&H1&H2&?&?&?)" "Hclose". - iDestruct (hocap_tapes_agree' with "[$][$]") as "%". - iDestruct (flip_tapes_agree with "[$][H1]") as "%". - { iDestruct (big_sepM_lookup with "[$]") as "?"; first done. - iFrame. - } - erewrite <-(insert_delete m') at 1; last done. - rewrite big_sepM_insert; last apply lookup_delete. + iDestruct (flip_tapes_agree with "[$][$]") as "%". simpl. - iDestruct "H1" as "[Htape H1]". - iMod (hocap_tapes_pop1' with "[$][$]") as "[H4 Hα]". iMod (flip_tapes_pop with "[$][$]") as "[Hauth Htape]". - iMod ("Hclose" with "[-Hα Hauth]"). - - iExists (<[_:=_]> _). iFrame. - iNext. - erewrite <-(insert_delete m') at 2; last done. - rewrite insert_insert. - rewrite big_sepM_insert; [iFrame|apply lookup_delete]. - - iModIntro. simpl in *. iSplit; first done. iFrame. + iModIntro. + simpl in *. + by iFrame. } iIntros "Hα". wp_apply (conversion.wp_bool_to_int) as "_"; first done. wp_pures. - wp_apply (flip_tape_spec_some _ _ _ _ (α ◯↪N (false, n :: ns)@ γ2) (α ◯↪N (true, ns)@ γ2)with "[$Hα]"); [|iSplit; first done|]. + wp_apply (flip_tape_spec_some _ _ _ _ (flip_tapes_frag γ2 α (drop 1%nat (expander (n::ns)))) (flip_tapes_frag γ2 α (expander ns))with "[$Hα]"); [|iSplit; first done|]. { by apply nclose_subseteq'. } { iModIntro. iIntros (m) "[Hfrag Hauth]". - iInv "Hinv'" as ">(%m'&%&%&H1&H2&?&?&?)" "Hclose". - iDestruct (hocap_tapes_agree' with "[$][$]") as "%". - iDestruct (flip_tapes_agree with "[$][H1]") as "%". - { iDestruct (big_sepM_lookup with "[$]") as "?"; first done. - iFrame. - } - erewrite <-(insert_delete m') at 1; last done. - rewrite big_sepM_insert; last apply lookup_delete. - simpl in *. - iDestruct "H1" as "[Htape H1]". - iMod (hocap_tapes_pop2' with "[$][$]") as "[H4 Hα]". + iDestruct (flip_tapes_agree with "[$][$]") as "%". + simpl. iMod (flip_tapes_pop with "[$][$]") as "[Hauth Htape]". - iMod ("Hclose" with "[-Hα Hauth]"). - - iExists (<[_:=_]> _). iFrame. - iNext. - erewrite <-(insert_delete m') at 2; last done. - rewrite insert_insert. - rewrite big_sepM_insert; [iFrame|apply lookup_delete]. - - iModIntro. simpl in *. iSplit; first done. iFrame. + iModIntro. + simpl in *. + by iFrame. } iIntros "Hα". wp_apply (conversion.wp_bool_to_int) as "_"; first done. wp_pures. wp_bind (FAA _ _). - iInv "Hinv'" as ">(%&%&%&?&?&->&?&?)" "Hclose". + iInv "Hinv'" as ">(%&%&->&?&?)" "Hclose". wp_faa. simpl. + iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". + { apply difference_mono; [done|by apply nclose_subseteq']. } iMod ("Hvs" with "[$]") as "[H6 HQ]". - replace (#(z+n)) with (#(z+n)%nat); last first. - { by rewrite Nat2Z.inj_add. } replace 2%Z with (Z.of_nat 2%nat) by done. replace (_*_+_)%Z with (Z.of_nat n); last first. - { admit. (* rewrite -Nat2Z.inj_mul -Nat2Z.inj_add -Nat.div2_odd -Nat2Z.inj_add. *) } - iMod ("Hclose" with "[-HΦ HQ Hα]") as "_"; first iFrame. - { admit. } + { assert (n<4). + - by eapply (@Forall_inv _ (λ x, x<4)). + - by apply expander_eta. + } + replace (#(z+n)) with (#(z+n)%nat); last first. + { by rewrite Nat2Z.inj_add. } + iMod "Hclose'" as "_". + iMod ("Hclose" with "[-HΦ HQ Hα]") as "_"; first by iFrame. iModIntro. wp_pures. iApply "HΦ". - by iFrame. - Admitted. + iFrame. + iPureIntro. + by eapply Forall_inv_tail. + Qed. Lemma counter_presample_spec2 N E ns α (ε2 : R -> nat -> R) @@ -275,15 +257,15 @@ Section impl2. ↑N ⊆ E -> (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> (∀ (ε:R), 0<= ε ->SeriesC (λ n, if (bool_decide (n≤3%nat)) then 1 / (S 3%nat) * ε2 ε n else 0%R)%R <= ε)%R-> - (∃ γ2', is_flip (L:=L) (N.@"flip") γ1 γ2' ∗ - inv (N.@ "counter") (counter_inv_pred2 c γ2 γ2' γ3)) -∗ + (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ + inv (N.@ "counter") (counter_inv_pred2 c γ3)) -∗ (□∀ (ε:R) n, (P ∗ flip_error_auth (L:=L) γ1 ε) ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝ ∨(flip_error_auth (L:=L) γ1 (ε2 ε n) ∗ T (n)))) -∗ - P -∗ α ◯↪N (true, ns) @ γ2 -∗ - wp_update E (∃ n, T (n) ∗ α◯↪N (true, ns++[n]) @ γ2). + P -∗ (flip_tapes_frag (L:=L) γ2 α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝) -∗ + wp_update E (∃ n, T (n) ∗ (flip_tapes_frag (L:=L) γ2 α (expander (ns++[n])) ∗ ⌜Forall (λ x, x<4) (ns++[n])⌝)). Proof. - iIntros (Hsubset Hpos Hineq) "(%γ2' & #Hinv & #Hinv') #Hvs HP Hfrag". + iIntros (Hsubset Hpos Hineq) "(#Hinv & #Hinv') #Hvs HP Hfrag". iApply fupd_wp_update. Admitted. (* iApply flip_presample_spec. *) @@ -372,8 +354,8 @@ Section impl2. Lemma read_counter_spec2 N E c γ1 γ2 γ3 P Q: ↑N ⊆ E -> - {{{ (∃ γ2', is_flip (L:=L) (N.@"flip") γ1 γ2' ∗ - inv (N.@ "counter") (counter_inv_pred2 c γ2 γ2' γ3)) ∗ + {{{ (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ + inv (N.@ "counter") (counter_inv_pred2 c γ3)) ∗ □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ own γ3 (●F z)∗ Q z) ∗ P @@ -382,36 +364,36 @@ Section impl2. {{{ (n':nat), RET #n'; Q n' }}}. Proof. - iIntros (Hsubset Φ) "((%γ2' & #Hinv & #Hinv') & #Hvs & HP) HΦ". + iIntros (Hsubset Φ) "((#Hinv & #Hinv') & #Hvs & HP) HΦ". rewrite /read_counter2. wp_pure. - iInv "Hinv'" as ">(%m&%l&%z &Hfrags & Hauth & -> & Hloc & Hcont)" "Hclose". + iInv "Hinv'" as ">(%l&%z & -> & Hloc & Hcont)" "Hclose". wp_load. iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". { apply difference_mono_l. by apply nclose_subseteq'. } iMod ("Hvs" with "[$]") as "[Hcont HQ]". iMod "Hclose'". - iMod ("Hclose" with "[$Hfrags $Hauth $Hloc $Hcont]"); first done. + iMod ("Hclose" with "[ $Hloc $Hcont]"); first done. iApply "HΦ". by iFrame. Qed. End impl2. -Program Definition random_counter2 `{!conerisGS Σ, flip_spec Σ}: random_counter := +Program Definition random_counter2 `{flip_spec Σ}: random_counter := {| new_counter := new_counter2; allocate_tape:= allocate_tape2; incr_counter_tape := incr_counter_tape2; read_counter:=read_counter2; counterG := counterG2; error_name := flip_error_name; - tape_name := gname; + tape_name := flip_tape_name; counter_name :=gname; - is_counter _ N c γ1 γ2 γ3 := (∃ γ2', is_flip (N.@"flip") γ1 γ2' ∗ - inv (N.@ "counter") (counter_inv_pred2 (L:=_) c γ2 γ2' γ3))%I; + is_counter _ N c γ1 γ2 γ3 := (is_flip (N.@"flip") γ1 γ2 ∗ + inv (N.@ "counter") (counter_inv_pred2 c γ3))%I; counter_error_auth _ γ x := flip_error_auth γ x; counter_error_frag _ γ x := flip_error_frag γ x; - counter_tapes_auth _ γ m := (●((λ ns, (true, ns))<$>m)@γ)%I; - counter_tapes_frag _ γ α ns := (α◯↪N (true, ns) @ γ)%I; + counter_tapes_auth _ γ m := flip_tapes_auth (L:=_) γ ((λ ns, expander ns)<$>m); + counter_tapes_frag _ γ α ns := (flip_tapes_frag (L:=_) γ α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝)%I; counter_content_auth _ γ z := own γ (●F z); counter_content_frag _ γ f z := own γ (◯F{f} z); new_counter_spec _ := new_counter_spec2; @@ -420,82 +402,39 @@ Program Definition random_counter2 `{!conerisGS Σ, flip_spec Σ}: random_counte counter_presample_spec _ :=counter_presample_spec2; read_counter_spec _ :=read_counter_spec2 |}. -(* Next Obligation. *) -(* simpl. *) -(* iIntros (????????) "(%&<-&H1)(%&<-&H2)". *) -(* iCombine "H1 H2" gives "%K". by rewrite excl_auth.excl_auth_auth_op_valid in K. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (????????) "(%&<-&H1)(%&<-&H2)". *) -(* iCombine "H1 H2" gives "%K". by rewrite excl_auth.excl_auth_frag_op_valid in K. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (???????) "H". *) -(* iApply (hocap_error_auth_pos with "[$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (???????) "H". *) -(* iApply (hocap_error_frag_pos with "[$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (????????) "H1 H2". *) -(* iApply (hocap_error_agree with "[$][$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. iIntros (?????????) "??". *) -(* iApply (hocap_error_update with "[$][$]"). *) -(* Qed. *) -Next Obligation. - simpl. - iIntros (????????) "H1 H2". - by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[% _]". -Qed. Next Obligation. - simpl. - iIntros (?????????) "H1 H2". - iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. - rewrite dfrac_op_own dfrac_valid_own. by intro. + intros ?? H ???. + apply counterG2_flipG. Qed. Next Obligation. simpl. - iIntros. - iDestruct (hocap_tapes_agree' with "[$][$]") as "%K". - rewrite lookup_fmap_Some in K. destruct K as (?&?&?). - by simplify_eq. -Qed. -Next Obligation. - iIntros. - iDestruct (hocap_tapes_presample' with "[$][$]") as ">[? $]". - by rewrite fmap_insert. + iIntros (???????) "H1 H2". + iApply (flip_tapes_auth_exclusive with "[$][$]"). Qed. Next Obligation. simpl. - iIntros (????????) "H1 H2". + iIntros (???????) "H1 H2". iCombine "H1 H2" gives "%K". by rewrite auth_auth_op_valid in K. Qed. Next Obligation. - simpl. iIntros (?????? z z' ?) "H1 H2". + simpl. iIntros (????? z z' ?) "H1 H2". iCombine "H1 H2" gives "%K". apply frac_auth_included_total in K. iPureIntro. by apply nat_included. Qed. Next Obligation. simpl. - iIntros (??????????). + iIntros (?????????). rewrite frac_auth_frag_op. by rewrite own_op. Qed. Next Obligation. - simpl. iIntros (????????) "H1 H2". + simpl. iIntros (???????) "H1 H2". iCombine "H1 H2" gives "%K". iPureIntro. by apply frac_auth_agree_L in K. Qed. Next Obligation. - simpl. iIntros (??????????) "H1 H2". + simpl. iIntros (?????????) "H1 H2". iMod (own_update_2 _ _ _ (_ ⋅ _) with "[$][$]") as "[$$]"; last done. apply frac_auth_update. apply nat_local_update. lia. From 45b4474d070d9ad879a16d78d908866e8d15f00a Mon Sep 17 00:00:00 2001 From: Hei Li Date: Fri, 20 Sep 2024 10:49:00 +0200 Subject: [PATCH 16/69] flip presample spec simple --- theories/coneris/lib/flip_spec.v | 181 ++++++++++++++++++++----------- 1 file changed, 117 insertions(+), 64 deletions(-) diff --git a/theories/coneris/lib/flip_spec.v b/theories/coneris/lib/flip_spec.v index 608a2403..76fea537 100644 --- a/theories/coneris/lib/flip_spec.v +++ b/theories/coneris/lib/flip_spec.v @@ -88,6 +88,22 @@ Class flip_spec `{!conerisGS Σ} := FlipSpec flip_tape #lbl:α @ E {{{ RET #n; Q }}}; flip_presample_spec {L: flipG Σ} NS E ns α + (ε2 : R -> list bool -> R) + (P : iProp Σ) num T γ1 γ2: + ↑NS ⊆ E -> + (∀ ε l, 0<= ε -> length l = num -> 0<=ε2 ε l)%R -> + (∀ (ε:R), 0<= ε -> SeriesC (λ l, if length l =? num then ε2 ε l else 0) /(2^num) <= ε)%R-> + is_flip (L:=L) NS γ1 γ2 -∗ + (□∀ (ε:R) n, (P ∗ flip_error_auth (L:=L) γ1 ε) ={E∖↑NS}=∗ + (⌜(1<=ε2 ε n)%R⌝ ∨ (flip_error_auth (L:=L) γ1 (ε2 ε n) ∗ T (n)))) + -∗ + P -∗ flip_tapes_frag (L:=L) γ2 α ns-∗ + wp_update E (∃ n, T (n) ∗ flip_tapes_frag (L:=L) γ2 α (ns++n)) +}. + +Section test. + Context `{F:flip_spec}. + Lemma flip_presample_spec_simple {L: flipG Σ} NS E ns α (ε2 : R -> bool -> R) (P : iProp Σ) T γ1 γ2: ↑NS ⊆ E -> @@ -98,8 +114,44 @@ Class flip_spec `{!conerisGS Σ} := FlipSpec (⌜(1<=ε2 ε n)%R⌝ ∨ (flip_error_auth (L:=L) γ1 (ε2 ε n) ∗ T (n)))) -∗ P -∗ flip_tapes_frag (L:=L) γ2 α ns-∗ - wp_update E (∃ n, T (n) ∗ flip_tapes_frag (L:=L) γ2 α (ns++[n])) -}. + wp_update E (∃ n, T (n) ∗ flip_tapes_frag (L:=L) γ2 α (ns++[n])). + Proof. + iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". + pose (ε2' ε l:= match l with + | [b]=> ε2 ε b + | _ => 1%R + end + ). + iMod (flip_presample_spec _ _ _ _ ε2' P 1%nat (λ l, match l with | [b] => T b | _ => False end )%I with "[//][][$][$]") as "(%l & [??])"; first done. + - rewrite /ε2'. + intros. repeat case_match; try done. naive_solver. + - intros. + etrans; last apply Hineq; try done. + erewrite (SeriesC_subset (λ x, x ∈ [[true]; [false]])); last first. + + intros ? H'. + case_match eqn:K; last done. + rewrite Nat.eqb_eq in K. + exfalso. + apply H'. + destruct a as [|[|] [|]]; try (simpl in *; done). + all: set_solver. + + rewrite SeriesC_list; last first. + { rewrite !NoDup_cons; repeat split; last apply NoDup_nil; set_solver. } + simpl. lra. + - iModIntro. iIntros (ε n) "?". + destruct n as [?|b [|]]. + + iLeft. iPureIntro. + by rewrite /ε2'. + + iMod ("Hvs" $! ε b with "[$]") as "[% | ?]". + * iLeft. iPureIntro. by rewrite /ε2'. + * iRight. by rewrite /ε2'. + + iLeft. iPureIntro. + by rewrite /ε2'. + - repeat case_match; try done. + iModIntro. + iFrame. + Qed. +End test. (** Instantiate flip *) @@ -261,65 +313,66 @@ Next Obligation. + by rewrite Z_to_bool_eq_0. Qed. Next Obligation. - simpl. - iIntros (???????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". - iApply wp_update_state_step_coupl. - iIntros (σ ε) "((Hheap&Htapes)&Hε)". - iMod (inv_acc with "Hinv") as "[>(% & % & H1 & H2 & H3 & H4 ) Hclose]"; [done|]. - iDestruct (hocap_tapes_agree with "[$][$]") as "%K". - rewrite lookup_fmap_Some in K. destruct K as (?&M&?). - simplify_eq. - unshelve epose proof fmap_inj _ _ _ _ M as ->. - { intros [][]?; by simplify_eq. } - erewrite <-(insert_delete m) at 1; last done. - rewrite fmap_insert. - rewrite big_sepM_insert; last first. - { rewrite fmap_delete. apply lookup_delete. } - simpl. - iDestruct "H3" as "[Htape H3]". - iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". - iDestruct (ec_supply_bound with "[$][$]") as "%". - iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. - iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". - iApply state_step_coupl_state_adv_comp_con_prob_lang; first done. - unshelve iExists (λ x, mknonnegreal (ε2 ε1' (nat_to_bool (fin_to_nat x))) _). - { apply Hpos. apply cond_nonneg. } - iSplit. - { iPureIntro. - simpl. - unshelve epose proof (Hineq ε1' _) as H'; first apply cond_nonneg. - rewrite SeriesC_finite_foldr/=. - rewrite nat_to_bool_eq_0 nat_to_bool_neq_0; last lia. - simpl in *. lra. - } - iIntros (sample). - destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' (nat_to_bool (fin_to_nat sample))))%R 1%R) as [Hdec|Hdec]; last first. - { apply Rnot_lt_ge, Rge_le in Hdec. - iApply state_step_coupl_ret_err_ge_1. - simpl. simpl in *. lra. - } - iApply state_step_coupl_ret. - unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' (nat_to_bool (fin_to_nat sample))) _) with "Hε_supply") as "[Hε_supply Hε]". - { apply Hpos. apply cond_nonneg. } - { simpl. done. } - simpl. - iMod (tapeN_update_append _ _ _ _ sample with "[$][Htape]") as "[Htapes Htape]". - { by erewrite Heq. } - iMod (hocap_tapes_presample _ _ _ _ _ (fin_to_nat sample) with "[$][$]") as "[H4 Hfrag]". - iMod "Hclose'" as "_". - iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". - { iExFalso. iApply (ec_contradict with "[$]"). exact. } - iMod ("Hclose" with "[$Hε $H2 Htape H3 H4]") as "_". - { iNext. - iExists (<[α:=(ns ++ [nat_to_bool sample])]>m). - rewrite fmap_insert. - rewrite big_sepM_insert_delete Heq/=. - rewrite fmap_delete. iFrame. - rewrite fmap_app/= nat_to_bool_to_nat; first iFrame. - pose proof fin_to_nat_lt sample. lia. - } - iApply fupd_mask_intro_subseteq; first set_solver. - iFrame. - rewrite fmap_app/= nat_to_bool_to_nat; first done. - pose proof fin_to_nat_lt sample. lia. -Qed. +Admitted. +(* simpl. *) +(* iIntros (???????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) +(* iApply wp_update_state_step_coupl. *) +(* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) +(* iMod (inv_acc with "Hinv") as "[>(% & % & H1 & H2 & H3 & H4 ) Hclose]"; [done|]. *) +(* iDestruct (hocap_tapes_agree with "[$][$]") as "%K". *) +(* rewrite lookup_fmap_Some in K. destruct K as (?&M&?). *) +(* simplify_eq. *) +(* unshelve epose proof fmap_inj _ _ _ _ M as ->. *) +(* { intros [][]?; by simplify_eq. } *) +(* erewrite <-(insert_delete m) at 1; last done. *) +(* rewrite fmap_insert. *) +(* rewrite big_sepM_insert; last first. *) +(* { rewrite fmap_delete. apply lookup_delete. } *) +(* simpl. *) +(* iDestruct "H3" as "[Htape H3]". *) +(* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". *) +(* iDestruct (ec_supply_bound with "[$][$]") as "%". *) +(* iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. *) +(* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) +(* iApply state_step_coupl_state_adv_comp_con_prob_lang; first done. *) +(* unshelve iExists (λ x, mknonnegreal (ε2 ε1' (nat_to_bool (fin_to_nat x))) _). *) +(* { apply Hpos. apply cond_nonneg. } *) +(* iSplit. *) +(* { iPureIntro. *) +(* simpl. *) +(* unshelve epose proof (Hineq ε1' _) as H'; first apply cond_nonneg. *) +(* rewrite SeriesC_finite_foldr/=. *) +(* rewrite nat_to_bool_eq_0 nat_to_bool_neq_0; last lia. *) +(* simpl in *. lra. *) +(* } *) +(* iIntros (sample). *) +(* destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' (nat_to_bool (fin_to_nat sample))))%R 1%R) as [Hdec|Hdec]; last first. *) +(* { apply Rnot_lt_ge, Rge_le in Hdec. *) +(* iApply state_step_coupl_ret_err_ge_1. *) +(* simpl. simpl in *. lra. *) +(* } *) +(* iApply state_step_coupl_ret. *) +(* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' (nat_to_bool (fin_to_nat sample))) _) with "Hε_supply") as "[Hε_supply Hε]". *) +(* { apply Hpos. apply cond_nonneg. } *) +(* { simpl. done. } *) +(* simpl. *) +(* iMod (tapeN_update_append _ _ _ _ sample with "[$][Htape]") as "[Htapes Htape]". *) +(* { by erewrite Heq. } *) +(* iMod (hocap_tapes_presample _ _ _ _ _ (fin_to_nat sample) with "[$][$]") as "[H4 Hfrag]". *) +(* iMod "Hclose'" as "_". *) +(* iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". *) +(* { iExFalso. iApply (ec_contradict with "[$]"). exact. } *) +(* iMod ("Hclose" with "[$Hε $H2 Htape H3 H4]") as "_". *) +(* { iNext. *) +(* iExists (<[α:=(ns ++ [nat_to_bool sample])]>m). *) +(* rewrite fmap_insert. *) +(* rewrite big_sepM_insert_delete Heq/=. *) +(* rewrite fmap_delete. iFrame. *) +(* rewrite fmap_app/= nat_to_bool_to_nat; first iFrame. *) +(* pose proof fin_to_nat_lt sample. lia. *) +(* } *) +(* iApply fupd_mask_intro_subseteq; first set_solver. *) +(* iFrame. *) +(* rewrite fmap_app/= nat_to_bool_to_nat; first done. *) +(* pose proof fin_to_nat_lt sample. lia. *) +(* Qed. *) From c0be184e82bf6174964538fa7ac98df7ed73a04e Mon Sep 17 00:00:00 2001 From: Hei Li Date: Fri, 20 Sep 2024 11:25:18 +0200 Subject: [PATCH 17/69] Progress with flip presample spec --- theories/coneris/lib/flip_spec.v | 135 ++++++++++++++++-------------- theories/coneris/lib/hocap.v | 7 ++ theories/coneris/primitive_laws.v | 9 ++ 3 files changed, 89 insertions(+), 62 deletions(-) diff --git a/theories/coneris/lib/flip_spec.v b/theories/coneris/lib/flip_spec.v index 76fea537..decc4a45 100644 --- a/theories/coneris/lib/flip_spec.v +++ b/theories/coneris/lib/flip_spec.v @@ -313,66 +313,77 @@ Next Obligation. + by rewrite Z_to_bool_eq_0. Qed. Next Obligation. + simpl. + iIntros (????????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". + iApply wp_update_state_step_coupl. + iIntros (σ ε) "((Hheap&Htapes)&Hε)". + iMod (inv_acc with "Hinv") as "[>(% & % & H1 & H2 & H3 & H4 ) Hclose]"; [done|]. + iDestruct (hocap_tapes_agree with "[$][$]") as "%K". + rewrite lookup_fmap_Some in K. destruct K as (?&M&?). + simplify_eq. + unshelve epose proof fmap_inj _ _ _ _ M as ->. + { intros [][]?; by simplify_eq. } + erewrite <-(insert_delete m) at 1; last done. + rewrite fmap_insert. + rewrite big_sepM_insert; last first. + { rewrite fmap_delete. apply lookup_delete. } + simpl. + iDestruct "H3" as "[Htape H3]". + iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". + iDestruct (ec_supply_bound with "[$][$]") as "%". + iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. + iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". + iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. + unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 ε1' (nat_to_bool <$> (fin_to_nat <$> x)) else 1) _). + { case_match; last (simpl;lra). apply Hpos; first apply cond_nonneg. + rewrite !fmap_length. + by apply Nat.eqb_eq. + } + iSplit. + { iPureIntro. + simpl. + unshelve epose proof (Hineq ε1' _) as H'; first apply cond_nonneg. + admit. + (* rewrite SeriesC_finite_foldr/=. *) + (* rewrite nat_to_bool_eq_0 nat_to_bool_neq_0; last lia. *) + (* simpl in *. lra. *) + } + iIntros (sample) "<-". + destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' (nat_to_bool <$> (fin_to_nat <$> sample))))%R 1%R) as [Hdec|Hdec]; last first. + { apply Rnot_lt_ge, Rge_le in Hdec. + iApply state_step_coupl_ret_err_ge_1. + simpl. simpl in *. rewrite Nat.eqb_refl. lra. + } + iApply state_step_coupl_ret. + unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' (nat_to_bool <$> (fin_to_nat <$> sample))) _) with "Hε_supply") as "[Hε_supply Hε]". + { apply Hpos; first apply cond_nonneg. by rewrite !fmap_length. } + { simpl. done. } + simpl. + iMod (tapeN_update_append' _ _ _ _ sample with "[$][Htape]") as "[Htapes Htape]". + { by erewrite Heq. } + iMod (hocap_tapes_presample' _ _ _ _ _ (fin_to_nat <$> sample) with "[$][$]") as "[H4 Hfrag]". + iMod "Hclose'" as "_". + iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". + { iExFalso. iApply (ec_contradict with "[$]"). exact. } + assert (∀ (x: list (fin (2))), bool_to_nat <$> (nat_to_bool <$> (fin_to_nat <$> x)) = fin_to_nat <$> x) as Hrewrite. + { intros l. + rewrite -list_fmap_compose. + rewrite <-list_fmap_id. + apply list_fmap_ext. + intros ?? K. simpl. rewrite nat_to_bool_to_nat; first done. + apply list_lookup_fmap_inv in K as (x'&->&?). + pose proof fin_to_nat_lt x'. lia. } + iMod ("Hclose" with "[$Hε $H2 Htape H3 H4]") as "_". + { iNext. + iExists (<[α:=(ns ++ (nat_to_bool<$>(fin_to_nat <$> sample)))]>m). + rewrite fmap_insert. + rewrite big_sepM_insert_delete Heq/=. + rewrite fmap_delete. iFrame. + rewrite fmap_app/=. + rewrite !Hrewrite. iFrame. + } + iApply fupd_mask_intro_subseteq; first set_solver. + iFrame. rewrite fmap_app/=Hrewrite. iFrame. + erewrite (nnreal_ext _ _); first done. + simpl. by rewrite Nat.eqb_refl. Admitted. -(* simpl. *) -(* iIntros (???????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) -(* iApply wp_update_state_step_coupl. *) -(* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) -(* iMod (inv_acc with "Hinv") as "[>(% & % & H1 & H2 & H3 & H4 ) Hclose]"; [done|]. *) -(* iDestruct (hocap_tapes_agree with "[$][$]") as "%K". *) -(* rewrite lookup_fmap_Some in K. destruct K as (?&M&?). *) -(* simplify_eq. *) -(* unshelve epose proof fmap_inj _ _ _ _ M as ->. *) -(* { intros [][]?; by simplify_eq. } *) -(* erewrite <-(insert_delete m) at 1; last done. *) -(* rewrite fmap_insert. *) -(* rewrite big_sepM_insert; last first. *) -(* { rewrite fmap_delete. apply lookup_delete. } *) -(* simpl. *) -(* iDestruct "H3" as "[Htape H3]". *) -(* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". *) -(* iDestruct (ec_supply_bound with "[$][$]") as "%". *) -(* iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. *) -(* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) -(* iApply state_step_coupl_state_adv_comp_con_prob_lang; first done. *) -(* unshelve iExists (λ x, mknonnegreal (ε2 ε1' (nat_to_bool (fin_to_nat x))) _). *) -(* { apply Hpos. apply cond_nonneg. } *) -(* iSplit. *) -(* { iPureIntro. *) -(* simpl. *) -(* unshelve epose proof (Hineq ε1' _) as H'; first apply cond_nonneg. *) -(* rewrite SeriesC_finite_foldr/=. *) -(* rewrite nat_to_bool_eq_0 nat_to_bool_neq_0; last lia. *) -(* simpl in *. lra. *) -(* } *) -(* iIntros (sample). *) -(* destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' (nat_to_bool (fin_to_nat sample))))%R 1%R) as [Hdec|Hdec]; last first. *) -(* { apply Rnot_lt_ge, Rge_le in Hdec. *) -(* iApply state_step_coupl_ret_err_ge_1. *) -(* simpl. simpl in *. lra. *) -(* } *) -(* iApply state_step_coupl_ret. *) -(* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' (nat_to_bool (fin_to_nat sample))) _) with "Hε_supply") as "[Hε_supply Hε]". *) -(* { apply Hpos. apply cond_nonneg. } *) -(* { simpl. done. } *) -(* simpl. *) -(* iMod (tapeN_update_append _ _ _ _ sample with "[$][Htape]") as "[Htapes Htape]". *) -(* { by erewrite Heq. } *) -(* iMod (hocap_tapes_presample _ _ _ _ _ (fin_to_nat sample) with "[$][$]") as "[H4 Hfrag]". *) -(* iMod "Hclose'" as "_". *) -(* iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". *) -(* { iExFalso. iApply (ec_contradict with "[$]"). exact. } *) -(* iMod ("Hclose" with "[$Hε $H2 Htape H3 H4]") as "_". *) -(* { iNext. *) -(* iExists (<[α:=(ns ++ [nat_to_bool sample])]>m). *) -(* rewrite fmap_insert. *) -(* rewrite big_sepM_insert_delete Heq/=. *) -(* rewrite fmap_delete. iFrame. *) -(* rewrite fmap_app/= nat_to_bool_to_nat; first iFrame. *) -(* pose proof fin_to_nat_lt sample. lia. *) -(* } *) -(* iApply fupd_mask_intro_subseteq; first set_solver. *) -(* iFrame. *) -(* rewrite fmap_app/= nat_to_bool_to_nat; first done. *) -(* pose proof fin_to_nat_lt sample. lia. *) -(* Qed. *) diff --git a/theories/coneris/lib/hocap.v b/theories/coneris/lib/hocap.v index e77a641c..4bca030e 100644 --- a/theories/coneris/lib/hocap.v +++ b/theories/coneris/lib/hocap.v @@ -111,6 +111,13 @@ Section tapes_lemmas. iApply (ghost_map_update with "[$][$]"). Qed. + Lemma hocap_tapes_presample' γ m k N ns ns': + (● m @ γ) -∗ (k ◯↪N (N; ns) @ γ) ==∗ (● (<[k:=(N,ns++ns')]>m) @ γ) ∗ (k ◯↪N (N; ns++ns') @ γ). + Proof. + iIntros "H1 H2". + iApply (ghost_map_update with "[$][$]"). + Qed. + Lemma hocap_tapes_pop γ m k N ns n: (● m @ γ) -∗ (k ◯↪N (N; n::ns) @ γ) ==∗ (● (<[k:=(N,ns)]>m) @ γ) ∗ (k ◯↪N (N; ns) @ γ). Proof. diff --git a/theories/coneris/primitive_laws.v b/theories/coneris/primitive_laws.v index c2d4bc20..cf715c48 100644 --- a/theories/coneris/primitive_laws.v +++ b/theories/coneris/primitive_laws.v @@ -127,6 +127,15 @@ Section tape_interface. iFrame. by rewrite fmap_app. Qed. + + Lemma tapeN_update_append' α N m (ns ns':list (fin (S N))): + tapes_auth 1 m -∗ α ↪N (N; fin_to_nat <$> ns) ==∗ tapes_auth 1 (<[α:=(N; ns ++ ns')]> m) ∗ α ↪N (N; (fin_to_nat <$> ns) ++ (fin_to_nat <$> ns')). + Proof. + iIntros "? (%&%&?)". + iMod (ghost_map_update with "[$][$]") as "[??]". + iFrame. + by rewrite fmap_app. + Qed. (* Lemma spec_tapeN_to_empty l M : From 4dc3f5059bd6226080e4a8c094c41f3afb5671a0 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Fri, 20 Sep 2024 16:02:38 +0200 Subject: [PATCH 18/69] Progress with changing hocap RA to authR, TODO: fix rand counter interface and its implementations --- theories/coneris/error_rules.v | 42 +++--- .../coneris/examples/random_counter/impl1.v | 1 + .../coneris/examples/random_counter/impl2.v | 35 ++++- .../examples/random_counter/random_counter.v | 8 +- theories/coneris/lib/flip_spec.v | 112 ++++++++++----- theories/coneris/lib/hocap.v | 128 +++++++++++++----- theories/prelude/uniform_list.v | 57 ++++++-- theories/prob/countable_sum.v | 7 + theories/prob/distribution.v | 20 +-- 9 files changed, 298 insertions(+), 112 deletions(-) diff --git a/theories/coneris/error_rules.v b/theories/coneris/error_rules.v index defd65e9..3ff7b576 100644 --- a/theories/coneris/error_rules.v +++ b/theories/coneris/error_rules.v @@ -836,7 +836,7 @@ Lemma pgl_iterM_state N p σ α ns: pgl (iterM p (λ σ, state_step σ α) σ) (λ σ', ∃ ns' : list (fin (S N)), - ns' ∈ enum_uniform_list N p ∧ σ' = state_upd_tapes <[α:=(N; ns ++ ns')]> σ) 0. + ns' ∈ enum_uniform_fin_list N p ∧ σ' = state_upd_tapes <[α:=(N; ns ++ ns')]> σ) 0. Proof. intros H. rewrite /pgl. @@ -855,7 +855,7 @@ Proof. rewrite -dunifv_pos in H1. apply H0. exists x. - split; [by apply elem_of_enum_uniform_list|done]. + split; [by apply elem_of_enum_uniform_fin_list|done]. Qed. Lemma state_step_coupl_iterM_state_adv_comp_con_prob_lang (p:nat) α σ1 Z (ε ε_rem: nonnegreal) N ns: @@ -871,11 +871,11 @@ Proof. by apply elem_of_list_In, elem_of_list_In, elem_of_elements, elem_of_dom. } assert (0<=1 / S N ^ p)%R as Hineq. { apply Rcomplements.Rdiv_le_0_compat; first lra. apply pow_lt. apply pos_INR_S. } - (* R: predicate should hold iff tapes σ' at α is ns ++ [nx] where ns is in enum_uniform_list N p *) + (* R: predicate should hold iff tapes σ' at α is ns ++ [nx] where ns is in enum_uniform_fin_list N p *) unshelve iExists - (fun σ' : state => exists ns', ns' ∈ enum_uniform_list N p /\ σ' = (state_upd_tapes <[α:=(_; ns ++ ns') : tape]> σ1)), nnreal_zero, + (fun σ' : state => exists ns', ns' ∈ enum_uniform_fin_list N p /\ σ' = (state_upd_tapes <[α:=(_; ns ++ ns') : tape]> σ1)), nnreal_zero, (fun ρ => (ε_rem + - match ClassicalEpsilon.excluded_middle_informative (exists ns', ns' ∈ enum_uniform_list N p /\ ρ = (state_upd_tapes <[α:=(_; ns ++ ns') : tape]> σ1)) with + match ClassicalEpsilon.excluded_middle_informative (exists ns', ns' ∈ enum_uniform_fin_list N p /\ ρ = (state_upd_tapes <[α:=(_; ns ++ ns') : tape]> σ1)) with | left p => mknonnegreal (ε2 (epsilon p)) _ | _ => nnreal_zero end))%NNR. @@ -893,14 +893,14 @@ Proof. - apply Rlt_gt. apply pow_lt. pose proof pos_INR N; lra. - rewrite -S_INR; simpl in *; lra. } - rewrite elem_of_enum_uniform_list in H1. + rewrite elem_of_enum_uniform_fin_list in H1. etrans; last exact. etrans; last apply (SeriesC_ge_elem _ (epsilon e)). + rewrite S_INR. rewrite H1. by rewrite Nat.eqb_refl. + intros; case_match; last lra. apply Rmult_le_pos; try done. by simpl. - + apply (ex_seriesC_ext (λ n, if bool_decide (n∈enum_uniform_list N p) then (1 / S N ^ p * ε2 n)%R else 0%R)); last apply ex_seriesC_list. - intros. case_bool_decide as H'; rewrite elem_of_enum_uniform_list in H'. + + apply (ex_seriesC_ext (λ n, if bool_decide (n∈enum_uniform_fin_list N p) then (1 / S N ^ p * ε2 n)%R else 0%R)); last apply ex_seriesC_list. + intros. case_bool_decide as H'; rewrite elem_of_enum_uniform_fin_list in H'. * subst. by rewrite Nat.eqb_refl. * rewrite -Nat.eqb_neq in H'. by rewrite H'. - iPureIntro. @@ -912,13 +912,13 @@ Proof. by rewrite dmap_unfold_pmf -SeriesC_scal_r. } rewrite fubini_pos_seriesC'; last first. - + eapply (ex_seriesC_ext (λ a, if bool_decide (a ∈ enum_uniform_list N p) then _ else 0%R)); last apply ex_seriesC_list. + + eapply (ex_seriesC_ext (λ a, if bool_decide (a ∈ enum_uniform_fin_list N p) then _ else 0%R)); last apply ex_seriesC_list. intros n. case_bool_decide as H; first done. rewrite SeriesC_0; first done. intros x. rewrite dunifv_pmf bool_decide_eq_false_2; first lra. - by rewrite -elem_of_enum_uniform_list. + by rewrite -elem_of_enum_uniform_fin_list. + intros a. rewrite dunifv_pmf. eapply (ex_seriesC_ext (λ b, if bool_decide (b=state_upd_tapes <[α:=(N; ns ++ a)]> σ1) then _ else 0%R)); last apply ex_seriesC_singleton_dependent. @@ -927,11 +927,11 @@ Proof. + intros. repeat apply Rmult_le_pos; repeat case_match; simpl; try lra; try done. all: apply Rplus_le_le_0_compat; by try lra. - + erewrite (SeriesC_ext _ (λ n, (if bool_decide (n∈enum_uniform_list N p) then 1 / S N ^ p * ε2 n else 0) + (if bool_decide (n∈enum_uniform_list N p) then 1 / S N ^ p * ε_rem else 0)))%R. + + erewrite (SeriesC_ext _ (λ n, (if bool_decide (n∈enum_uniform_fin_list N p) then 1 / S N ^ p * ε2 n else 0) + (if bool_decide (n∈enum_uniform_fin_list N p) then 1 / S N ^ p * ε_rem else 0)))%R. * rewrite SeriesC_plus; [|apply ex_seriesC_list..]. - rewrite SeriesC_list_2; last apply NoDup_enum_uniform_list. - rewrite enum_uniform_list_length. - setoid_rewrite elem_of_enum_uniform_list'. + rewrite SeriesC_list_2; last apply NoDup_enum_uniform_fin_list. + rewrite enum_uniform_fin_list_length. + setoid_rewrite elem_of_enum_uniform_fin_list'. rewrite Rplus_0_l. rewrite Rplus_comm. apply Rplus_le_compat; last done. rewrite -pow_INR. simpl. @@ -951,7 +951,7 @@ Proof. case_match; last first. { exfalso. apply n. naive_solver. } rewrite dunifv_pmf. - rewrite bool_decide_eq_true_2; last by apply elem_of_enum_uniform_list. + rewrite bool_decide_eq_true_2; last by apply elem_of_enum_uniform_fin_list. rewrite -pow_INR. simpl. pose proof epsilon_correct _ e as H'. simpl in H'. replace (epsilon e) with l; try lra. @@ -962,7 +962,7 @@ Proof. intros. rewrite dunifv_pmf. rewrite bool_decide_eq_false_2; first lra. - by rewrite -elem_of_enum_uniform_list. + by rewrite -elem_of_enum_uniform_fin_list. - simpl. iPureIntro. eapply pgl_mon_pred; last first. @@ -970,7 +970,7 @@ Proof. + done. - iIntros (σ2 [ns' [Helem ->]]). pose proof Helem as Helem'. - rewrite elem_of_enum_uniform_list in Helem. rewrite <- Helem. + rewrite elem_of_enum_uniform_fin_list in Helem. rewrite <- Helem. iMod ("H" with "[]") as "H"; first done. case_match; last first. + (* contradiction *) @@ -1001,16 +1001,16 @@ Proof. * intros. split; repeat case_match; try rewrite S_INR; simpl; try rewrite Rmult_1_r; try lra. all: apply Rmult_le_pos; last done. all: rewrite -S_INR; apply Rdiv_INR_ge_0. - * eapply (ex_seriesC_ext (λ x, if (bool_decide (x∈enum_uniform_list N 1%nat)) then + * eapply (ex_seriesC_ext (λ x, if (bool_decide (x∈enum_uniform_fin_list N 1%nat)) then match x with |[x'] => (1 / S N * ε2 x')%R | _ => 0 end else 0)); last apply ex_seriesC_list. intros [|n[|]]. -- rewrite bool_decide_eq_false_2; first done. - by rewrite elem_of_enum_uniform_list. + by rewrite elem_of_enum_uniform_fin_list. -- rewrite bool_decide_eq_true_2; first done. - by rewrite elem_of_enum_uniform_list. + by rewrite elem_of_enum_uniform_fin_list. -- rewrite bool_decide_eq_false_2; first done. - by rewrite elem_of_enum_uniform_list. + by rewrite elem_of_enum_uniform_fin_list. + intros; apply Rmult_le_pos; last by simpl. apply Rdiv_INR_ge_0. + intros. repeat case_match; by simplify_eq. diff --git a/theories/coneris/examples/random_counter/impl1.v b/theories/coneris/examples/random_counter/impl1.v index 64261d43..a6713079 100644 --- a/theories/coneris/examples/random_counter/impl1.v +++ b/theories/coneris/examples/random_counter/impl1.v @@ -38,6 +38,7 @@ Section impl1. iDestruct (ec_valid with "[$]") as "%". unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "[%γ1 [H1 H2]]". { lra. } + { simpl. lra. } simpl. iMod (hocap_tapes_alloc (∅:gmap _ _)) as "[%γ2 [H3 H4]]". iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index 21eb9cff..f5fa6344 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -439,5 +439,36 @@ Next Obligation. apply frac_auth_update. apply nat_local_update. lia. Qed. - - +Next Obligation. + intros ?? H ?. + apply counterG2_flipG. +Qed. +Next Obligation. + simpl. + iIntros (????????) "[? _] [? _]". + by iDestruct (flip_tapes_frag_exclusive with "[$][$]") as "%". +Qed. +Next Obligation. + simpl. + iIntros. + iApply (flip_error_update with "[$][$]"). +Qed. +Next Obligation. + iIntros. + iApply (flip_error_agree with "[$][$]"). +Qed. +Next Obligation. + iIntros. + iApply (flip_error_frag_valid with "[$]"). +Admitted. +Next Obligation. +Admitted. +Next Obligation. +Admitted. +Next Obligation. +Admitted. +Next Obligation. + simpl. +Admitted. +Next Obligation. +Admitted. diff --git a/theories/coneris/examples/random_counter/random_counter.v b/theories/coneris/examples/random_counter/random_counter.v index b2a86280..9c0c06b2 100644 --- a/theories/coneris/examples/random_counter/random_counter.v +++ b/theories/coneris/examples/random_counter/random_counter.v @@ -46,10 +46,10 @@ Class random_counter `{!conerisGS Σ} := RandCounter counter_error_auth (L:=L) γ x1 -∗ counter_error_auth (L:=L) γ x2 -∗ False; counter_error_frag_exclusive {L : counterG Σ} γ x1 x2: counter_error_frag (L:=L) γ x1 -∗ counter_error_frag (L:=L) γ x2 -∗ False; - counter_error_auth_pos {L : counterG Σ} γ x: - counter_error_auth (L:=L) γ x -∗ ⌜(0<=x)%R⌝; - counter_error_auth_frag {L : counterG Σ} γ x: - counter_error_frag (L:=L) γ x -∗ ⌜(0<=x)%R⌝; + counter_error_auth_valid {L : counterG Σ} γ x: + counter_error_auth (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; + counter_error_auth_valid {L : counterG Σ} γ x: + counter_error_frag (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; counter_error_agree {L : counterG Σ} γ x1 x2: counter_error_auth (L:=L) γ x1 -∗ counter_error_frag (L:=L) γ x2 -∗ ⌜x1 = x2 ⌝; counter_error_update {L : counterG Σ} γ x1 x2 (x3:nonnegreal): diff --git a/theories/coneris/lib/flip_spec.v b/theories/coneris/lib/flip_spec.v index decc4a45..50296ae0 100644 --- a/theories/coneris/lib/flip_spec.v +++ b/theories/coneris/lib/flip_spec.v @@ -39,17 +39,20 @@ Class flip_spec `{!conerisGS Σ} := FlipSpec flip_error_auth_exclusive {L : flipG Σ} γ x1 x2: flip_error_auth (L:=L) γ x1 -∗ flip_error_auth (L:=L) γ x2 -∗ False; - flip_error_frag_exclusive {L : flipG Σ} γ x1 x2: - flip_error_frag (L:=L) γ x1 -∗ flip_error_frag (L:=L) γ x2 -∗ False; - flip_error_auth_pos {L : flipG Σ} γ x: - flip_error_auth (L:=L) γ x -∗ ⌜(0<=x)%R⌝; - flip_error_auth_frag {L : flipG Σ} γ x: - flip_error_frag (L:=L) γ x -∗ ⌜(0<=x)%R⌝; - flip_error_agree {L : flipG Σ} γ x1 x2: - flip_error_auth (L:=L) γ x1 -∗ flip_error_frag (L:=L) γ x2 -∗ ⌜x1 = x2 ⌝; - flip_error_update {L : flipG Σ} γ x1 x2 (x3:nonnegreal): - flip_error_auth (L:=L) γ x1 -∗ flip_error_frag (L:=L) γ x2 ==∗ - flip_error_auth (L:=L) γ x3 ∗ flip_error_frag (L:=L) γ x3; + flip_error_frag_split {L : flipG Σ} γ (x1 x2:nonnegreal): + flip_error_frag (L:=L) γ x1 ∗ flip_error_frag (L:=L) γ x2 ⊣⊢ + flip_error_frag (L:=L) γ (x1+x2)%R ; + flip_error_auth_valid {L : flipG Σ} γ x: + flip_error_auth (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; + flip_error_frag_valid {L : flipG Σ} γ x: + flip_error_frag (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; + flip_error_ineq {L : flipG Σ} γ x1 x2: + flip_error_auth (L:=L) γ x1 -∗ flip_error_frag (L:=L) γ x2 -∗ ⌜(x2 <= x1)%R ⌝; + flip_error_decrease {L : flipG Σ} γ (x1 x2:nonnegreal): + flip_error_auth (L:=L) γ (x1 + x2)%NNR -∗ flip_error_frag (L:=L) γ x1 ==∗ flip_error_auth (L:=L) γ x2; + flip_error_increase {L : flipG Σ} γ (x1 x2:nonnegreal): + (x1+x2<1)%R -> ⊢ flip_error_auth (L:=L) γ x1 ==∗ + flip_error_auth (L:=L) γ (x1+x2) ∗ flip_error_frag (L:=L) γ x2; flip_tapes_auth_exclusive {L : flipG Σ} γ m m': flip_tapes_auth (L:=L) γ m -∗ flip_tapes_auth (L:=L) γ m' -∗ False; @@ -139,7 +142,7 @@ Section test. { rewrite !NoDup_cons; repeat split; last apply NoDup_nil; set_solver. } simpl. lra. - iModIntro. iIntros (ε n) "?". - destruct n as [?|b [|]]. + destruct n as [|b [|]]. + iLeft. iPureIntro. by rewrite /ε2'. + iMod ("Hvs" $! ε b with "[$]") as "[% | ?]". @@ -180,32 +183,36 @@ Local Definition flip_inv_pred1 `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapes Next Obligation. simpl. - iIntros (??????) "(%&<-&H1)(%&<-&H2)". - iCombine "H1 H2" gives "%H". by rewrite excl_auth.excl_auth_auth_op_valid in H. + iIntros. + iApply (hocap_error_auth_exclusive with "[$][$]"). Qed. Next Obligation. simpl. - iIntros (??????) "(%&<-&H1)(%&<-&H2)". - iCombine "H1 H2" gives "%H". by rewrite excl_auth.excl_auth_frag_op_valid in H. + iIntros. + iApply (hocap_error_frag_split). Qed. Next Obligation. simpl. iIntros (?????) "H". - iApply (hocap_error_auth_pos with "[$]"). + iApply (hocap_error_auth_valid with "[$]"). Qed. Next Obligation. simpl. iIntros (?????) "H". - iApply (hocap_error_frag_pos with "[$]"). + iApply (hocap_error_frag_valid with "[$]"). Qed. Next Obligation. simpl. iIntros (??????) "H1 H2". - iApply (hocap_error_agree with "[$][$]"). + iApply (hocap_error_ineq with "[$][$]"). +Qed. +Next Obligation. + iIntros. + iApply (hocap_error_decrease with "[$][$]"). Qed. Next Obligation. - simpl. iIntros (???????) "??". - iApply (hocap_error_update with "[$][$]"). + iIntros. + by iApply (hocap_error_increase with "[$]"). Qed. Next Obligation. simpl. @@ -250,7 +257,7 @@ Next Obligation. iApply fupd_wp_update. iApply wp_update_ret. iDestruct (ec_valid with "[$]") as "%". - unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "(%γ1 & H2 & H3)"; first naive_solver. + unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "(%γ1 & H2 & H3)"; simpl; [naive_solver..|]. simpl. iMod (hocap_tapes_alloc ∅) as "(%γ2 & H4 & H5)". iMod (inv_alloc _ _ (flip_inv_pred1 γ1 γ2) with "[H1 H2 H4]"). @@ -333,6 +340,14 @@ Next Obligation. iDestruct (ec_supply_bound with "[$][$]") as "%". iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". + assert (∀ (x: list (fin (2))), bool_to_nat <$> (nat_to_bool <$> (fin_to_nat <$> x)) = fin_to_nat <$> x) as Hrewrite. + { intros l. + rewrite -list_fmap_compose. + rewrite <-list_fmap_id. + apply list_fmap_ext. + intros ?? K. simpl. rewrite nat_to_bool_to_nat; first done. + apply list_lookup_fmap_inv in K as (x'&->&?). + pose proof fin_to_nat_lt x'. lia. } iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 ε1' (nat_to_bool <$> (fin_to_nat <$> x)) else 1) _). { case_match; last (simpl;lra). apply Hpos; first apply cond_nonneg. @@ -343,10 +358,45 @@ Next Obligation. { iPureIntro. simpl. unshelve epose proof (Hineq ε1' _) as H'; first apply cond_nonneg. - admit. - (* rewrite SeriesC_finite_foldr/=. *) - (* rewrite nat_to_bool_eq_0 nat_to_bool_neq_0; last lia. *) - (* simpl in *. lra. *) + etrans; last exact. + rewrite (Rdiv_def _ (2^_)) -SeriesC_scal_r. + etrans; last eapply (SeriesC_le_inj _ (λ x, Some (nat_to_bool <$> (fin_to_nat <$> x)))). + - apply SeriesC_le. + + simpl. intros n. + rewrite !fmap_length. + case_match. + * replace (1+1)%R with 2%R by lra. + rewrite -Rdiv_1_l. simpl. + split; last lra. + apply Rmult_le_pos. + -- apply Rcomplements.Rdiv_le_0_compat; first lra. + apply pow_lt; lra. + -- apply Hpos; first done. + rewrite !fmap_length. + by apply Nat.eqb_eq. + * lra. + + simpl. + apply (ex_seriesC_list_length _ num). + intros ?. rewrite !fmap_length. + case_match; last lra. + intros. by apply Nat.eqb_eq. + - intros. case_match; last lra. + apply Rmult_le_pos; first apply Hpos; simpl; auto. + + by apply Nat.eqb_eq. + + rewrite -Rdiv_1_l. + apply Rcomplements.Rdiv_le_0_compat; first lra. + apply pow_lt; lra. + - intros ??? <- K. + simplify_eq. + rewrite -!list_fmap_compose in K. + apply list_fmap_eq_inj in K; try done. + intros x x'. + repeat (inv_fin x; try intros x); + repeat (inv_fin x'; try intros x'); done. + - apply (ex_seriesC_list_length _ num). + intros ?. + case_match; last lra. + intros. by apply Nat.eqb_eq. } iIntros (sample) "<-". destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' (nat_to_bool <$> (fin_to_nat <$> sample))))%R 1%R) as [Hdec|Hdec]; last first. @@ -365,14 +415,6 @@ Next Obligation. iMod "Hclose'" as "_". iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". { iExFalso. iApply (ec_contradict with "[$]"). exact. } - assert (∀ (x: list (fin (2))), bool_to_nat <$> (nat_to_bool <$> (fin_to_nat <$> x)) = fin_to_nat <$> x) as Hrewrite. - { intros l. - rewrite -list_fmap_compose. - rewrite <-list_fmap_id. - apply list_fmap_ext. - intros ?? K. simpl. rewrite nat_to_bool_to_nat; first done. - apply list_lookup_fmap_inv in K as (x'&->&?). - pose proof fin_to_nat_lt x'. lia. } iMod ("Hclose" with "[$Hε $H2 Htape H3 H4]") as "_". { iNext. iExists (<[α:=(ns ++ (nat_to_bool<$>(fin_to_nat <$> sample)))]>m). @@ -386,4 +428,4 @@ Next Obligation. iFrame. rewrite fmap_app/=Hrewrite. iFrame. erewrite (nnreal_ext _ _); first done. simpl. by rewrite Nat.eqb_refl. -Admitted. +Qed. diff --git a/theories/coneris/lib/hocap.v b/theories/coneris/lib/hocap.v index 4bca030e..33da79ec 100644 --- a/theories/coneris/lib/hocap.v +++ b/theories/coneris/lib/hocap.v @@ -6,70 +6,123 @@ From clutch.coneris Require Import coneris flip. Set Default Proof Using "Type*". Definition hocap_error_nroot:= nroot.@ "error". -Definition hocap_error_RA := excl_authR NNRO. +Definition hocap_error_RA := authR nonnegrealR. Class hocap_errorGS (Σ : gFunctors) := Hocap_errorGS { hocap_errorGS_inG :: inG Σ (hocap_error_RA); }. -Definition hocap_errorΣ := #[GFunctor (excl_authR (nonnegrealUR))]. +Definition hocap_errorΣ := #[GFunctor (hocap_error_RA)]. -Notation "'●↯' ε '@' γ" := (∃ (x : nonnegreal), ⌜x.(nonneg) = ε%R⌝ ∗ own γ (●E x))%I +Notation "'●↯' ε '@' γ" := (∃ (x : nonnegreal), ⌜x.(nonneg) = ε%R⌝ ∗ own γ (● x))%I (at level 1). -Notation "'◯↯' ε '@' γ" := (∃ (x : nonnegreal), ⌜x.(nonneg) = ε%R⌝ ∗ own γ (◯E x))%I +Notation "'◯↯' ε '@' γ" := (∃ (x : nonnegreal), ⌜x.(nonneg) = ε%R⌝ ∗ own γ (◯ x))%I (at level 1). -Definition hocap_tapes_nroot:=nroot.@"tapes". -Class hocap_tapesGS (Σ : gFunctors) := Hocap_tapesGS { - hocap_tapesGS_inG :: ghost_mapG Σ loc (nat*list nat) - }. -Definition hocap_tapesΣ := ghost_mapΣ loc (nat*list nat). - -Notation "α ◯↪N ( M ; ns ) @ γ":= (α ↪[ γ ] (M,ns))%I - (at level 20, format "α ◯↪N ( M ; ns ) @ γ") : bi_scope. - -Notation "● m @ γ" := (ghost_map_auth γ 1 m) (at level 20) : bi_scope. Section error_lemmas. Context `{!conerisGS Σ, !hocap_errorGS Σ}. + + Lemma hocap_error_auth_exclusive b b' γ: + ●↯ b @ γ -∗ ●↯ b' @ γ -∗ False. + Proof. + iIntros "[%[% H1]] [%[% H2]]". + iCombine "H1 H2" gives "%H1". + compute in H1. destruct H1. + exfalso. + apply H1. done. + Qed. + + Lemma hocap_error_frag_split (b b':nonnegreal) γ: + ◯↯ b @ γ ∗ ◯↯ b' @ γ ⊣⊢ ◯↯ (b+b') @ γ. + Proof. + iSplit. + - iIntros "[[%x1[% H1]] [%x2[% H2]]]". + iExists (x1 + x2)%NNR. + iSplit; [simpl; simpl in *; iPureIntro; lra|]. + rewrite auth_frag_op own_op. iFrame. + - iIntros "[%x [% H]]". + simpl in *. + replace x with (b+b')%NNR; last (apply nnreal_ext; simpl; lra). + rewrite auth_frag_op own_op. + by iDestruct "H" as "[$ $]". + Qed. + (* Helpful lemmas *) - Lemma hocap_error_auth_pos (b:R) γ: - (●↯ b @ γ) -∗ ⌜(0<=b)%R⌝. + Lemma hocap_error_auth_valid (b:R) γ: + (●↯ b @ γ) -∗ ⌜(0<=b<1)%R⌝. Proof. - by iIntros "[%[<- H]]". + iIntros "[%x[<- H]]". + iDestruct (own_valid with "[$]") as "%H". + iPureIntro. + rewrite auth_auth_valid in H. + destruct x. + compute in H. + split; simpl; lra. Qed. - Lemma hocap_error_frag_pos (b:R) γ: - (◯↯ b @ γ) -∗ ⌜(0<=b)%R⌝. + Lemma hocap_error_frag_valid (b:R) γ: + (◯↯ b @ γ) -∗ ⌜(0<=b<1)%R⌝. Proof. - by iIntros "[%[<- H]]". + iIntros "[%[<- H]]". + iDestruct (own_valid with "[$]") as "%H". + iPureIntro. + rewrite auth_frag_valid in H. + destruct x. + compute in H. + split; simpl; lra. Qed. Lemma hocap_error_alloc (ε:nonnegreal): - ⊢ |==>∃ γ, (●↯ ε @ γ) ∗ (◯↯ ε @ γ). + (ε<1)%R -> ⊢ |==>∃ γ, (●↯ ε @ γ) ∗ (◯↯ ε @ γ). Proof. - iMod (own_alloc (●E ε ⋅ ◯E ε)) as "[% [??]]". - - by apply excl_auth_valid. + intros H. + iMod (own_alloc (● ε ⋅ ◯ ε)) as "[% [??]]". + - apply auth_both_valid_2. + + compute. destruct ε; simpl in H. lra. + + apply nonnegreal_included; lra. - by eauto with iFrame. Qed. - Lemma hocap_error_agree γ (b c:R) : - (●↯ b @ γ) -∗ (◯↯ c @ γ) -∗ ⌜ b = c ⌝. + Lemma hocap_error_ineq γ (b c:R) : + (●↯ b @ γ) -∗ (◯↯ c @ γ) -∗ ⌜ (c<=b)%R ⌝. Proof. iIntros "[% [<- Hγ●]] [% [<-Hγ◯]]". - by iCombine "Hγ● Hγ◯" gives %<-%excl_auth_agree_L. + iCombine "Hγ● Hγ◯" gives "%Hop". + by eapply auth_both_valid_discrete in Hop as [Hlt%nonnegreal_included ?]. Qed. - Lemma hocap_error_update γ (b':nonnegreal) (b c:R) : - (●↯ b @ γ) -∗ (◯↯ c @ γ) ==∗ (●↯ b' @ γ) ∗ (◯↯ b' @ γ). + Lemma hocap_error_decrease γ (b' b:nonnegreal) : + (●↯ (b+b') @ γ) -∗ (◯↯ b @ γ) ==∗ (●↯ b' @ γ). Proof. - iIntros "[% [<- Hγ●]] [% [<-Hγ◯]]". - iMod (own_update_2 _ _ _ (_ ⋅ _) with "Hγ● Hγ◯") as "[$$]". - { by apply excl_auth_update. } - done. + iIntros "H1 H2". + simpl. + iDestruct "H1" as "[% [% H1]]". + iDestruct "H2" as "[% [% H2]]". + iMod (own_update_2 with "H1 H2") as "Hown". + { eapply (auth_update_dealloc _ _ b'), nonnegreal_local_update. + - apply cond_nonneg. + - apply nnreal_ext =>/=. lra. } + iFrame. by iPureIntro. Qed. + + Lemma hocap_error_increase γ (b b':nonnegreal) : + (b+b'<1)%R -> ⊢ (●↯ b @ γ) ==∗ (●↯ (b+b')%NNR @ γ) ∗ (◯↯ b' @ γ). + Proof. + iIntros (Hineq) "[% [% H]]". + iMod (own_update with "H") as "[$ $]"; last (iPureIntro; split; last done). + - apply auth_update_alloc. + apply (local_update_unital_discrete _ _ (b+b')%NNR) => z H1 H2. + split; first done. + apply nnreal_ext. simpl. + rewrite Rplus_comm. + apply Rplus_eq_compat_l. + simpl in *. rewrite -H H2. simpl. lra. + - done. + Qed. + Lemma hocap_error_irrel γ (b c:R) : (b=c)%R -> (●↯ b @ γ) -∗ (●↯ c @ γ). Proof. @@ -78,6 +131,17 @@ Section error_lemmas. End error_lemmas. +Definition hocap_tapes_nroot:=nroot.@"tapes". +Class hocap_tapesGS (Σ : gFunctors) := Hocap_tapesGS { + hocap_tapesGS_inG :: ghost_mapG Σ loc (nat*list nat) + }. +Definition hocap_tapesΣ := ghost_mapΣ loc (nat*list nat). + +Notation "α ◯↪N ( M ; ns ) @ γ":= (α ↪[ γ ] (M,ns))%I + (at level 20, format "α ◯↪N ( M ; ns ) @ γ") : bi_scope. + +Notation "● m @ γ" := (ghost_map_auth γ 1 m) (at level 20) : bi_scope. + Section tapes_lemmas. Context `{!conerisGS Σ, !hocap_tapesGS Σ}. diff --git a/theories/prelude/uniform_list.v b/theories/prelude/uniform_list.v index 5a946068..bec7bc0f 100644 --- a/theories/prelude/uniform_list.v +++ b/theories/prelude/uniform_list.v @@ -4,19 +4,19 @@ Require Import Coq.Program.Equality. Set Default Proof Using "Type*". Section uniform_list. - Variables N:nat. + Context `{Hfinite: Finite A}. Fixpoint enum_uniform_list (p:nat):= match p with | O => [[]] | S p' => - x ← enum (fin (S N)); + x ← enum (A); y ← enum_uniform_list p'; mret (x::y) end. Lemma enum_uniform_list_S (p:nat) : enum_uniform_list (S p) = - x ← enum (fin (S N)); + x ← enum A; y ← enum_uniform_list p; mret (x::y). Proof. @@ -31,7 +31,8 @@ Section uniform_list. + simpl. rewrite elem_of_list_singleton. by intros ->. + rewrite enum_uniform_list_S. rewrite elem_of_list_bind. elim. intros x. rewrite elem_of_list_bind. do 2 elim. intros y. - intros [?%elem_of_list_ret H%IHp]. subst. done. + intros [?%elem_of_list_ret ?]. subst. + simpl. intros. f_equal. naive_solver. - revert l; induction p; intros l. + simpl. set_unfold. intros ?%nil_length_inv. naive_solver. + destruct l as [|t l'] eqn:Heqn; first done. @@ -51,7 +52,7 @@ Section uniform_list. - symmetry; by rewrite Nat.eqb_neq. Qed. - Lemma bind_length1 {A:Type} (l:list (list A)) a: + Lemma bind_length1 (l:list (list A)) a: length (l ≫= λ y, mret (a :: y)) = length l. Proof. induction l. @@ -60,7 +61,7 @@ Section uniform_list. rewrite app_length. simpl. f_equal. done. Qed. - Lemma bind_length2 {A:Type} (l1 : list A) l2: + Lemma bind_length2 (l1 : list A) l2: length (l1 ≫= λ x, l2 ≫= λ y, mret (x :: y)) = length l1 * length l2. Proof. revert l2. @@ -73,14 +74,13 @@ Section uniform_list. Qed. Lemma enum_uniform_list_length p: - length (enum_uniform_list p) = (S N)^p. + length (enum_uniform_list p) = (length (enum A))^p. Proof. induction p. - done. - rewrite enum_uniform_list_S. rewrite bind_length2. rewrite IHp. - rewrite length_enum_fin. rewrite Nat.pow_succ_r'. lia. Qed. @@ -100,3 +100,44 @@ Section uniform_list. Qed. End uniform_list. + + +Section uniform_fin_list. + Variables N:nat. + + Definition enum_uniform_fin_list:= + enum_uniform_list (A:=fin (S N)). + + Lemma enum_uniform_fin_list_S (p:nat) : + enum_uniform_fin_list (S p) = + x ← enum (fin (S N)); + y ← enum_uniform_fin_list p; + mret (x::y). + Proof. + simpl. done. + Qed. + + Lemma elem_of_enum_uniform_fin_list l p: + l ∈ enum_uniform_fin_list p <-> length l = p. + Proof. + apply elem_of_enum_uniform_list. + Qed. + + Lemma elem_of_enum_uniform_fin_list' l p: + bool_decide (l∈enum_uniform_fin_list p) = (length l =? p). + Proof. + apply elem_of_enum_uniform_list'. + Qed. + + Lemma enum_uniform_fin_list_length p: + length (enum_uniform_fin_list p) = (S N)^p. + Proof. + by rewrite enum_uniform_list_length length_enum_fin. + Qed. + + Lemma NoDup_enum_uniform_fin_list p: + NoDup (enum_uniform_fin_list p). + Proof. + apply NoDup_enum_uniform_list. + Qed. +End uniform_fin_list. diff --git a/theories/prob/countable_sum.v b/theories/prob/countable_sum.v index 19a24450..6eaa80f5 100644 --- a/theories/prob/countable_sum.v +++ b/theories/prob/countable_sum.v @@ -710,6 +710,13 @@ Section filter. End filter. +Lemma ex_seriesC_list_length `{Countable A} (f:list A -> R) num: + (forall x, (0 length x = num) -> + ex_seriesC f. +Proof. + intros. +Admitted. + Lemma SeriesC_Series_nat (f : nat → R) : SeriesC f = Series f. Proof. diff --git a/theories/prob/distribution.v b/theories/prob/distribution.v index 193aaa72..50ca8f5d 100644 --- a/theories/prob/distribution.v +++ b/theories/prob/distribution.v @@ -2247,8 +2247,8 @@ Section uniform. End uniform. -(** Uniform lists *) -Section uniform_lists. +(** Uniform fin lists *) +Section uniform_fin_lists. Program Definition dunifv (N p: nat): distr (list (fin (S N))) := MkDistr (λ x, if (bool_decide (length x = p)) then /(S N^p)%nat else 0) _ _ _. Next Obligation. @@ -2264,19 +2264,19 @@ Section uniform_lists. intros. eapply ex_seriesC_ext; last apply ex_seriesC_list. simpl. intros. erewrite bool_decide_ext; first done. - by erewrite elem_of_enum_uniform_list. + by erewrite elem_of_enum_uniform_fin_list. Qed. Next Obligation. intros N p. erewrite SeriesC_ext. - - erewrite SeriesC_list_2; last apply NoDup_enum_uniform_list. - rewrite enum_uniform_list_length. + - erewrite SeriesC_list_2; last apply NoDup_enum_uniform_fin_list. + rewrite enum_uniform_fin_list_length. erewrite Rinv_l; first done. apply not_0_INR. apply Nat.pow_nonzero. lia. - intros. erewrite bool_decide_ext; first done. - by erewrite elem_of_enum_uniform_list. + by erewrite elem_of_enum_uniform_fin_list. Qed. Lemma dunifv_pmf N p v: @@ -2287,15 +2287,15 @@ Section uniform_lists. SeriesC (dunifv N p) = 1. Proof. erewrite SeriesC_ext. - - erewrite SeriesC_list_2; last apply NoDup_enum_uniform_list. - rewrite enum_uniform_list_length. + - erewrite SeriesC_list_2; last apply NoDup_enum_uniform_fin_list. + rewrite enum_uniform_fin_list_length. erewrite Rinv_l; first done. apply not_0_INR. apply Nat.pow_nonzero. lia. - intros. rewrite /dunifv_pmf. erewrite bool_decide_ext; first done. - by erewrite elem_of_enum_uniform_list. + by erewrite elem_of_enum_uniform_fin_list. Qed. Lemma dunifv_pos N p v: @@ -2316,7 +2316,7 @@ Section uniform_lists. done. Qed. -End uniform_lists. +End uniform_fin_lists. Ltac inv_distr := repeat From fe3ce88e25acc2723ad294f00b2053a81dfe575a Mon Sep 17 00:00:00 2001 From: Hei Li Date: Fri, 20 Sep 2024 16:53:40 +0200 Subject: [PATCH 19/69] BIG refactor, a lot of code commented out --- theories/coneris/examples/parallel_add.v | 434 ++++----- .../coneris/examples/random_counter/impl1.v | 664 ++++++------- .../coneris/examples/random_counter/impl2.v | 890 +++++++++--------- .../coneris/examples/random_counter/impl3.v | 778 +++++++-------- .../examples/random_counter/random_counter.v | 408 ++++---- theories/coneris/lib/flip_spec.v | 840 ++++++++--------- theories/coneris/lib/hocap.v | 412 ++++---- 7 files changed, 2215 insertions(+), 2211 deletions(-) diff --git a/theories/coneris/examples/parallel_add.v b/theories/coneris/examples/parallel_add.v index fd2aa1cd..111c754a 100644 --- a/theories/coneris/examples/parallel_add.v +++ b/theories/coneris/examples/parallel_add.v @@ -239,7 +239,7 @@ Section attempt1. own γ1 (●E b1) ∗ own γ2 (●E b2) ∗ l ↦ #z ∗ ↯ err ∗ ⌜ (nonneg err = 1- Rpower 2%R (INR flip_num-2))%R⌝ ∗ - ⌜(flip_num = bool_to_nat (ssrbool.isSome b1) +bool_to_nat (ssrbool.isSome b2))%nat⌝∗ + ⌜(flip_num = bool_to_nat (ssrbool.isSome b1) +bool_to_nat (ssrbool.isSome b2))%nat⌝∗ ⌜(z = bool_to_Z (is_Some_true b1) + bool_to_Z (is_Some_true b2))%Z⌝ . @@ -314,7 +314,7 @@ Section attempt1. iMod (ghost_var_update' _ (Some true) with "[$Hγ1●][$]") as "[Hγ1● Hγ1◯]". iMod ("Hclose" with "[Hγ1● Hγ2● Hl Herr]") as "_"; last done. iNext. iExists _, _, _, (mknonnegreal _ _ ), _. iFrame; simpl. - repeat iSplit; iPureIntro; [done|done|lia]. + repeat iSplit; iPureIntro; [done|done|lia]. } { rewrite /flipL. wp_pures. @@ -356,7 +356,7 @@ Section attempt1. iMod (ghost_var_update' _ (Some true) with "[$Hγ2●][$]") as "[Hγ2● Hγ2◯]". iMod ("Hclose" with "[Hγ1● Hγ2● Hl Herr]") as "_"; last done. iNext. iExists _, _, _, (mknonnegreal _ _ ), _. iFrame; simpl. - repeat iSplit; iPureIntro; [done|done|lia]. + repeat iSplit; iPureIntro; [done|done|lia]. } iIntros (??) "[Hγ1◯ Hγ2◯]". iNext. wp_pures. @@ -370,7 +370,7 @@ Section attempt1. Unshelve. all: try lra. all: simpl; try apply cond_nonneg. - all: apply resource_nonneg. + all: apply resource_nonneg. Qed. (** This time we use the property of flip being logically atomic *) @@ -435,7 +435,7 @@ Section attempt1. replace 1%R with (INR 1) by done. apply le_INR. rewrite /bool_to_nat; case_match; lia. - } + } destruct b; last first. { iIntros. iExFalso. iApply ec_contradict; last done. simpl. lra. } simpl. @@ -497,7 +497,7 @@ Section attempt1. replace 1%R with (INR 1) by done. apply le_INR. rewrite /bool_to_nat; case_match; lia. - } + } destruct b; last first. { iIntros. iExFalso. iApply ec_contradict; last done. simpl. lra. } simpl. @@ -542,220 +542,220 @@ Section attempt1. Qed. End attempt1. -Section attempt2. - Context `{!conerisGS Σ, !spawnG Σ, !hocap_errorGS Σ, !inG Σ (excl_authR ZR), !inG Σ (excl_authR (optionO boolO))}. (* A hocap style spec for half_FAA*) Definition loc_nroot:=nroot.@"loc". Definition both_nroot:=nroot.@"both". - Definition loc_inv (l:loc) γ:= - inv loc_nroot (∃ (z:Z), l↦#z ∗ own γ (●E z)). - - Lemma wp_half_FAA E - (ε2 : R -> bool -> R) - (P Q : iProp Σ) (T : bool -> iProp Σ) γ1 γ2 (l:loc): - ↑hocap_error_nroot ⊆ E -> - ↑loc_nroot ⊆ E -> - (∀ ε b, 0<= ε -> 0<= ε2 ε b)%R-> - (∀ (ε:R), 0<=ε -> (((ε2 ε true) + (ε2 ε false))/2 <= (ε)))%R → - {{{ error_inv γ1 ∗ loc_inv l γ2 ∗ - □(∀ (ε:R) (b : bool), P ∗ ●↯ ε @ γ1 ={E∖↑hocap_error_nroot}=∗ (⌜(1<=ε2 ε b)%R⌝∨●↯ (ε2 ε b) @ γ1 ∗ T b) ) ∗ - □ (∀ (b:bool) (z:Z), T b ∗ own γ2 (●E z) ={E∖↑loc_nroot}=∗ - if b then own γ2 (●E(z+1)%Z)∗ Q else own γ2 (●E(z))∗ Q ) ∗ - P }}} half_FAA l @ E {{{ (v: val), RET v; Q }}}. - Proof. - iIntros (Hsubset1 Hsubset2 ? Hineq Φ) "(#Hinv1 & #Hinv2 & #Hchange1 & #Hchange2 & HP) HΦ". - rewrite /half_FAA. - wp_apply (wp_hocap_flip_adv_comp _ ε2 P with "[-HΦ]"); [done|..]. - - intros. naive_solver. - - simpl. intros; naive_solver. - - by repeat iSplit. - - iIntros. destruct b. - + wp_pures. iInv "Hinv2" as "(%&?&?)" "Hclose". - wp_faa. - iMod ("Hchange2" with "[$]") as "[??]". - iMod ("Hclose" with "[$]"). - by iApply "HΦ". - + iInv "Hinv2" as "(%&?&?)" "Hclose". - wp_pure. - iMod ("Hchange2" with "[$]") as "[??]". - iMod ("Hclose" with "[$]"). - by iApply "HΦ". - Qed. - - Lemma parallel_add_spec'': - {{{ ↯ (3/4) }}} - parallel_add - {{{ (z:Z), RET #z; ⌜(z=2)⌝ }}}. - Proof. - iIntros (Φ) "Herr HΦ". - rewrite /parallel_add. - wp_alloc l as "Hl". - wp_pures. - rewrite -/(half_FAA l). - (** Allocate hocap error*) - unshelve iMod (hocap_error_alloc (mknonnegreal (3/4) _)) as "(%γ1 & Hεauth & Hεfrag)". - { lra. } - simpl. - (** Allocate hocap heap *) - iMod (own_alloc ((●E 0) ⋅ (◯E 0))) as "[%γ2 [Hauth_loc Hfrag_loc]]". - { by apply excl_auth_valid. } - (** Allocate hocap for tracking contributions *) - iMod (own_alloc ((●E None) ⋅ (◯E None))) as "[%γ3 [Hauth_a Hfrag_a]]". - { by apply excl_auth_valid. } - iMod (own_alloc ((●E None) ⋅ (◯E None))) as "[%γ4 [Hauth_b Hfrag_b]]". - { by apply excl_auth_valid. } - (** Allocate error invariants *) - iMod (inv_alloc hocap_error_nroot _ ((∃ (ε:R), ↯ ε ∗ ●↯ ε @ γ1)) with "[Herr Hεauth]") as "#Hεinv". - { iExists (mknonnegreal (3/4) _); iFrame. - Unshelve. lra. } - (** Allocate location inv *) - iMod (inv_alloc loc_nroot _ (∃ (z:Z), l↦#z ∗ own γ2 (●E z) - ) with "[Hl Hauth_loc]") as "#Hlocinv". - { iFrame. } - iMod (inv_alloc both_nroot _ - (∃ (z:Z) (a b:option bool), own γ2 (◯E z) ∗ - own γ3 (●E a) ∗ - own γ4 (●E b) ∗ - ⌜(z=bool_to_Z (is_Some_true a) + bool_to_Z (is_Some_true b))%Z⌝ ∗ - ◯↯ (1-Rpower 2 (IZR (bool_to_Z (bool_decide (is_Some a)))+IZR (bool_to_Z (bool_decide (is_Some b)))-2))%R@ γ1) - with "[Hfrag_loc Hεfrag Hauth_a Hauth_b]") as "#Hinvboth". - { iFrame. simpl. iModIntro. iSplitR; first by iPureIntro. - replace (1-_)%R with (3/4)%R; first done. - replace (Rpower _ _) with (1/4)%R; try lra. - replace (0+0-2)%R with (-2)%R by lra. - rewrite Rpower_Ropp/ Rpower. replace (IPR 2) with (INR 2); last done. - rewrite -ln_pow; last lra. - rewrite exp_ln; lra. - } - wp_apply (wp_par (λ _, own γ3 (◯E (Some true)))(λ _, own γ4 (◯E (Some true))) with "[Hfrag_a][Hfrag_b]"). - - (* first branch *) - wp_apply (wp_half_FAA _ (λ x b, - match ClassicalEpsilon.excluded_middle_informative (1/2<=x)%R - with - | left P => if b then (2*x - 1)%R else 1%R - | _ => x - end ) (own γ3 (◯E None)) - (own γ3 (◯E (Some true))) - (λ b, ⌜b=true⌝ ∗ (own γ3 (◯E (Some false))))%I γ1 γ2 l with "[$Hfrag_a]"); [done|done|..]. - + intros; repeat case_match; lra. - + intros. case_match; simpl; lra. - + repeat iSplit. - * done. - * done. - * iModIntro. iIntros (ε res) "[Hfrag_a Hεauth]". - iInv both_nroot as ">(%z & %a & %b & Hloc_frag & Hauth_a & Hauth_b & -> & Hεfrag)" "Hclose". - iDestruct (hocap_error_agree with "[$][$]") as "%K". - iDestruct (ghost_var_agree' with "[$Hauth_a][$]") as "->". - simpl in *. - case_match; last first. - { exfalso. apply n. - simpl in *. - rewrite K. - assert (Rpower 2 (0+IZR (bool_to_Z (bool_decide (is_Some b))) - 2) <= 1/2)%R; try lra. - replace (_/_)%R with (/2) by lra. - rewrite -{3}(Rpower_1 2); last lra. - rewrite -Rpower_Ropp. - apply Rle_Rpower; first lra. - case_bool_decide; simpl; lra. - } - iMod (ghost_var_update' _ (Some false) with "[$Hauth_a][$]") as "[Hauth_a Hfrag_a]". - iMod (hocap_error_update _ (mknonnegreal _ _)with "[$][$]") as "[Hεauth Hεfrag]". - simpl. - iMod ("Hclose" with "[Hloc_frag Hauth_a Hauth_b Hεfrag]"); iFrame; first done. - simpl. case_match; last first. - { iLeft. done. } - iRight. iFrame. iModIntro; iSplitL; last done. - simpl. rewrite K. - iApply (hocap_error_irrel with "[$]"). - rewrite !Rpower_plus Rpower_O; last lra. - rewrite Rpower_1; lra. - * iModIntro. - iIntros (??) "[[-> Hfrag_a] Hauth_loc]". - iInv both_nroot as ">(% & %a & %b & Hloc_frag & Hauth_a & Hauth_b & -> & Hεfrag)" "Hclose". - iDestruct (ghost_var_agreeZ with "[$][$]") as "->". - iDestruct (ghost_var_agree' with "[$Hauth_a][$]") as "->". - iMod (ghost_var_updateZ with "[$][$]") as "[Hauth_loc Hfrag_loc]". - iMod (ghost_var_update' _ (Some true) with "[$Hauth_a][$]") as "[Hauth_a Hfrag_a]". - simpl. - iMod ("Hclose" with "[$Hfrag_loc $Hauth_a $Hauth_b $Hεfrag]") as "_"; first done. - iFrame. simpl. - iModIntro. - replace (0+_+1)%Z with (1+bool_to_Z (is_Some_true b)); [done|lia]. - + iIntros. done. - - (* first branch *) - wp_apply (wp_half_FAA _ (λ x b, - match ClassicalEpsilon.excluded_middle_informative (1/2<=x)%R - with - | left P => if b then (2*x - 1)%R else 1%R - | _ => x - end ) (own γ4 (◯E None)) - (own γ4 (◯E (Some true))) - (λ b, ⌜b=true⌝ ∗ (own γ4 (◯E (Some false))))%I γ1 γ2 l with "[$Hfrag_b]"); [done|done|..]. - + intros; repeat case_match; lra. - + intros. case_match; simpl; lra. - + repeat iSplit. - * done. - * done. - * iModIntro. iIntros (ε res) "[Hfrag_b Hεauth]". - iInv both_nroot as ">(%z & %a & %b & Hloc_frag & Hauth_a & Hauth_b & -> & Hεfrag)" "Hclose". - iDestruct (hocap_error_agree with "[$][$]") as "%K". - iDestruct (ghost_var_agree' with "[$Hauth_b][$]") as "->". - simpl in *. - case_match; last first. - { exfalso. apply n. - simpl in *. - rewrite K. - assert (Rpower 2 (IZR (bool_to_Z (bool_decide (is_Some a)))+0 - 2) <= 1/2)%R; try lra. - replace (_/_)%R with (/2) by lra. - rewrite -{3}(Rpower_1 2); last lra. - rewrite -Rpower_Ropp. - apply Rle_Rpower; first lra. - case_bool_decide; simpl; lra. - } - iMod (ghost_var_update' _ (Some false) with "[$Hauth_b][$]") as "[Hauth_b Hfrag_b]". - iMod (hocap_error_update _ (mknonnegreal _ _)with "[$][$]") as "[Hεauth Hεfrag]". - simpl. - iMod ("Hclose" with "[Hloc_frag Hauth_a Hauth_b Hεfrag]"); iFrame; first done. - simpl. case_match; last first. - { iLeft. done. } - iRight. iFrame. iModIntro; iSplitL; last done. - simpl. rewrite K. - iApply (hocap_error_irrel with "[$]"). - rewrite !Rpower_plus Rpower_O; last lra. - rewrite Rpower_1; lra. - * iModIntro. - iIntros (??) "[[-> Hfrag_a] Hauth_loc]". - iInv both_nroot as ">(% & %a & %b & Hloc_frag & Hauth_a & Hauth_b & -> & Hεfrag)" "Hclose". - iDestruct (ghost_var_agreeZ with "[$][$]") as "->". - iDestruct (ghost_var_agree' with "[$Hauth_b][$]") as "->". - iMod (ghost_var_updateZ with "[$][$]") as "[Hauth_loc Hfrag_loc]". - iMod (ghost_var_update' _ (Some true) with "[$Hauth_b][$]") as "[Hauth_b Hfrag_b]". - simpl. - iMod ("Hclose" with "[$Hfrag_loc $Hauth_a $Hauth_b $Hεfrag]") as "_"; first done. - iFrame. simpl. - iModIntro. - replace (_+0+1)%Z with (bool_to_Z (is_Some_true a)+1); [done|lia]. - + iIntros. done. - - iIntros (??) "[Hfrag_a Hfrag_b]". iModIntro. wp_pures. - iInv loc_nroot as ">(%&Hl&Hloc_auth)" "Hclose". - iInv both_nroot as ">(%&%&%&Hloc_frag&Hauth_a & Hauth_b & -> & Hεfrag)" "Hclose'". - iDestruct (ghost_var_agreeZ with "[$][$]") as "->". - iDestruct (ghost_var_agree' with "[$Hauth_a][$]") as "->". - iDestruct (ghost_var_agree' with "[$Hauth_b][$]") as "->". - simpl. - wp_load. - iMod ("Hclose'" with "[$Hloc_frag $Hauth_a $Hauth_b $Hεfrag]"); first done. - iModIntro. - iMod ("Hclose" with "[Hl Hloc_auth]"); first by iFrame. - by iApply "HΦ". - Unshelve. - all: simpl; try lra. - + replace (1+_-2)%R with (INR (bool_to_nat (bool_decide (is_Some b)))-1)%R; first apply resource_nonneg. - case_bool_decide; simpl; lra. - + replace (_+1-2)%R with (INR (bool_to_nat (bool_decide (is_Some a)))-1)%R; first apply resource_nonneg. - case_bool_decide; simpl; lra. - Qed. -End attempt2. +(* Section attempt2. *) +(* Context `{!conerisGS Σ, !spawnG Σ, !hocap_errorGS Σ, !inG Σ (excl_authR ZR), !inG Σ (excl_authR (optionO boolO))}. *) + + (* Definition loc_inv (l:loc) γ:= *) + (* inv loc_nroot (∃ (z:Z), l↦#z ∗ own γ (●E z)). *) +(* Lemma wp_half_FAA E *) +(* (ε2 : R -> bool -> R) *) +(* (P Q : iProp Σ) (T : bool -> iProp Σ) γ1 γ2 (l:loc): *) +(* ↑hocap_error_nroot ⊆ E -> *) +(* ↑loc_nroot ⊆ E -> *) +(* (∀ ε b, 0<= ε -> 0<= ε2 ε b)%R-> *) +(* (∀ (ε:R), 0<=ε -> (((ε2 ε true) + (ε2 ε false))/2 <= (ε)))%R → *) +(* {{{ error_inv γ1 ∗ loc_inv l γ2 ∗ *) +(* □(∀ (ε:R) (b : bool), P ∗ ●↯ ε @ γ1 ={E∖↑hocap_error_nroot}=∗ (⌜(1<=ε2 ε b)%R⌝∨●↯ (ε2 ε b) @ γ1 ∗ T b) ) ∗ *) +(* □ (∀ (b:bool) (z:Z), T b ∗ own γ2 (●E z) ={E∖↑loc_nroot}=∗ *) +(* if b then own γ2 (●E(z+1)%Z)∗ Q else own γ2 (●E(z))∗ Q ) ∗ *) +(* P }}} half_FAA l @ E {{{ (v: val), RET v; Q }}}. *) +(* Proof. *) +(* iIntros (Hsubset1 Hsubset2 ? Hineq Φ) "(#Hinv1 & #Hinv2 & #Hchange1 & #Hchange2 & HP) HΦ". *) +(* rewrite /half_FAA. *) +(* wp_apply (wp_hocap_flip_adv_comp _ ε2 P with "[-HΦ]"); [done|..]. *) +(* - intros. naive_solver. *) +(* - simpl. intros; naive_solver. *) +(* - by repeat iSplit. *) +(* - iIntros. destruct b. *) +(* + wp_pures. iInv "Hinv2" as "(%&?&?)" "Hclose". *) +(* wp_faa. *) +(* iMod ("Hchange2" with "[$]") as "[??]". *) +(* iMod ("Hclose" with "[$]"). *) +(* by iApply "HΦ". *) +(* + iInv "Hinv2" as "(%&?&?)" "Hclose". *) +(* wp_pure. *) +(* iMod ("Hchange2" with "[$]") as "[??]". *) +(* iMod ("Hclose" with "[$]"). *) +(* by iApply "HΦ". *) +(* Qed. *) + +(* Lemma parallel_add_spec'': *) +(* {{{ ↯ (3/4) }}} *) +(* parallel_add *) +(* {{{ (z:Z), RET #z; ⌜(z=2)⌝ }}}. *) +(* Proof. *) +(* iIntros (Φ) "Herr HΦ". *) +(* rewrite /parallel_add. *) +(* wp_alloc l as "Hl". *) +(* wp_pures. *) +(* rewrite -/(half_FAA l). *) +(* (** Allocate hocap error*) *) +(* unshelve iMod (hocap_error_alloc (mknonnegreal (3/4) _)) as "(%γ1 & Hεauth & Hεfrag)". *) +(* { lra. } *) +(* simpl. *) +(* (** Allocate hocap heap *) *) +(* iMod (own_alloc ((●E 0) ⋅ (◯E 0))) as "[%γ2 [Hauth_loc Hfrag_loc]]". *) +(* { by apply excl_auth_valid. } *) +(* (** Allocate hocap for tracking contributions *) *) +(* iMod (own_alloc ((●E None) ⋅ (◯E None))) as "[%γ3 [Hauth_a Hfrag_a]]". *) +(* { by apply excl_auth_valid. } *) +(* iMod (own_alloc ((●E None) ⋅ (◯E None))) as "[%γ4 [Hauth_b Hfrag_b]]". *) +(* { by apply excl_auth_valid. } *) +(* (** Allocate error invariants *) *) +(* iMod (inv_alloc hocap_error_nroot _ ((∃ (ε:R), ↯ ε ∗ ●↯ ε @ γ1)) with "[Herr Hεauth]") as "#Hεinv". *) +(* { iExists (mknonnegreal (3/4) _); iFrame. *) +(* Unshelve. lra. } *) +(* (** Allocate location inv *) *) +(* iMod (inv_alloc loc_nroot _ (∃ (z:Z), l↦#z ∗ own γ2 (●E z) *) +(* ) with "[Hl Hauth_loc]") as "#Hlocinv". *) +(* { iFrame. } *) +(* iMod (inv_alloc both_nroot _ *) +(* (∃ (z:Z) (a b:option bool), own γ2 (◯E z) ∗ *) +(* own γ3 (●E a) ∗ *) +(* own γ4 (●E b) ∗ *) +(* ⌜(z=bool_to_Z (is_Some_true a) + bool_to_Z (is_Some_true b))%Z⌝ ∗ *) +(* ◯↯ (1-Rpower 2 (IZR (bool_to_Z (bool_decide (is_Some a)))+IZR (bool_to_Z (bool_decide (is_Some b)))-2))%R@ γ1) *) +(* with "[Hfrag_loc Hεfrag Hauth_a Hauth_b]") as "#Hinvboth". *) +(* { iFrame. simpl. iModIntro. iSplitR; first by iPureIntro. *) +(* replace (1-_)%R with (3/4)%R; first done. *) +(* replace (Rpower _ _) with (1/4)%R; try lra. *) +(* replace (0+0-2)%R with (-2)%R by lra. *) +(* rewrite Rpower_Ropp/ Rpower. replace (IPR 2) with (INR 2); last done. *) +(* rewrite -ln_pow; last lra. *) +(* rewrite exp_ln; lra. *) +(* } *) +(* wp_apply (wp_par (λ _, own γ3 (◯E (Some true)))(λ _, own γ4 (◯E (Some true))) with "[Hfrag_a][Hfrag_b]"). *) +(* - (* first branch *) *) +(* wp_apply (wp_half_FAA _ (λ x b, *) +(* match ClassicalEpsilon.excluded_middle_informative (1/2<=x)%R *) +(* with *) +(* | left P => if b then (2*x - 1)%R else 1%R *) +(* | _ => x *) +(* end ) (own γ3 (◯E None)) *) +(* (own γ3 (◯E (Some true))) *) +(* (λ b, ⌜b=true⌝ ∗ (own γ3 (◯E (Some false))))%I γ1 γ2 l with "[$Hfrag_a]"); [done|done|..]. *) +(* + intros; repeat case_match; lra. *) +(* + intros. case_match; simpl; lra. *) +(* + repeat iSplit. *) +(* * done. *) +(* * done. *) +(* * iModIntro. iIntros (ε res) "[Hfrag_a Hεauth]". *) +(* iInv both_nroot as ">(%z & %a & %b & Hloc_frag & Hauth_a & Hauth_b & -> & Hεfrag)" "Hclose". *) +(* iDestruct (hocap_error_agree with "[$][$]") as "%K". *) +(* iDestruct (ghost_var_agree' with "[$Hauth_a][$]") as "->". *) +(* simpl in *. *) +(* case_match; last first. *) +(* { exfalso. apply n. *) +(* simpl in *. *) +(* rewrite K. *) +(* assert (Rpower 2 (0+IZR (bool_to_Z (bool_decide (is_Some b))) - 2) <= 1/2)%R; try lra. *) +(* replace (_/_)%R with (/2) by lra. *) +(* rewrite -{3}(Rpower_1 2); last lra. *) +(* rewrite -Rpower_Ropp. *) +(* apply Rle_Rpower; first lra. *) +(* case_bool_decide; simpl; lra. *) +(* } *) +(* iMod (ghost_var_update' _ (Some false) with "[$Hauth_a][$]") as "[Hauth_a Hfrag_a]". *) +(* iMod (hocap_error_update _ (mknonnegreal _ _)with "[$][$]") as "[Hεauth Hεfrag]". *) +(* simpl. *) +(* iMod ("Hclose" with "[Hloc_frag Hauth_a Hauth_b Hεfrag]"); iFrame; first done. *) +(* simpl. case_match; last first. *) +(* { iLeft. done. } *) +(* iRight. iFrame. iModIntro; iSplitL; last done. *) +(* simpl. rewrite K. *) +(* iApply (hocap_error_irrel with "[$]"). *) +(* rewrite !Rpower_plus Rpower_O; last lra. *) +(* rewrite Rpower_1; lra. *) +(* * iModIntro. *) +(* iIntros (??) "[[-> Hfrag_a] Hauth_loc]". *) +(* iInv both_nroot as ">(% & %a & %b & Hloc_frag & Hauth_a & Hauth_b & -> & Hεfrag)" "Hclose". *) +(* iDestruct (ghost_var_agreeZ with "[$][$]") as "->". *) +(* iDestruct (ghost_var_agree' with "[$Hauth_a][$]") as "->". *) +(* iMod (ghost_var_updateZ with "[$][$]") as "[Hauth_loc Hfrag_loc]". *) +(* iMod (ghost_var_update' _ (Some true) with "[$Hauth_a][$]") as "[Hauth_a Hfrag_a]". *) +(* simpl. *) +(* iMod ("Hclose" with "[$Hfrag_loc $Hauth_a $Hauth_b $Hεfrag]") as "_"; first done. *) +(* iFrame. simpl. *) +(* iModIntro. *) +(* replace (0+_+1)%Z with (1+bool_to_Z (is_Some_true b)); [done|lia]. *) +(* + iIntros. done. *) +(* - (* first branch *) *) +(* wp_apply (wp_half_FAA _ (λ x b, *) +(* match ClassicalEpsilon.excluded_middle_informative (1/2<=x)%R *) +(* with *) +(* | left P => if b then (2*x - 1)%R else 1%R *) +(* | _ => x *) +(* end ) (own γ4 (◯E None)) *) +(* (own γ4 (◯E (Some true))) *) +(* (λ b, ⌜b=true⌝ ∗ (own γ4 (◯E (Some false))))%I γ1 γ2 l with "[$Hfrag_b]"); [done|done|..]. *) +(* + intros; repeat case_match; lra. *) +(* + intros. case_match; simpl; lra. *) +(* + repeat iSplit. *) +(* * done. *) +(* * done. *) +(* * iModIntro. iIntros (ε res) "[Hfrag_b Hεauth]". *) +(* iInv both_nroot as ">(%z & %a & %b & Hloc_frag & Hauth_a & Hauth_b & -> & Hεfrag)" "Hclose". *) +(* iDestruct (hocap_error_agree with "[$][$]") as "%K". *) +(* iDestruct (ghost_var_agree' with "[$Hauth_b][$]") as "->". *) +(* simpl in *. *) +(* case_match; last first. *) +(* { exfalso. apply n. *) +(* simpl in *. *) +(* rewrite K. *) +(* assert (Rpower 2 (IZR (bool_to_Z (bool_decide (is_Some a)))+0 - 2) <= 1/2)%R; try lra. *) +(* replace (_/_)%R with (/2) by lra. *) +(* rewrite -{3}(Rpower_1 2); last lra. *) +(* rewrite -Rpower_Ropp. *) +(* apply Rle_Rpower; first lra. *) +(* case_bool_decide; simpl; lra. *) +(* } *) +(* iMod (ghost_var_update' _ (Some false) with "[$Hauth_b][$]") as "[Hauth_b Hfrag_b]". *) +(* iMod (hocap_error_update _ (mknonnegreal _ _)with "[$][$]") as "[Hεauth Hεfrag]". *) +(* simpl. *) +(* iMod ("Hclose" with "[Hloc_frag Hauth_a Hauth_b Hεfrag]"); iFrame; first done. *) +(* simpl. case_match; last first. *) +(* { iLeft. done. } *) +(* iRight. iFrame. iModIntro; iSplitL; last done. *) +(* simpl. rewrite K. *) +(* iApply (hocap_error_irrel with "[$]"). *) +(* rewrite !Rpower_plus Rpower_O; last lra. *) +(* rewrite Rpower_1; lra. *) +(* * iModIntro. *) +(* iIntros (??) "[[-> Hfrag_a] Hauth_loc]". *) +(* iInv both_nroot as ">(% & %a & %b & Hloc_frag & Hauth_a & Hauth_b & -> & Hεfrag)" "Hclose". *) +(* iDestruct (ghost_var_agreeZ with "[$][$]") as "->". *) +(* iDestruct (ghost_var_agree' with "[$Hauth_b][$]") as "->". *) +(* iMod (ghost_var_updateZ with "[$][$]") as "[Hauth_loc Hfrag_loc]". *) +(* iMod (ghost_var_update' _ (Some true) with "[$Hauth_b][$]") as "[Hauth_b Hfrag_b]". *) +(* simpl. *) +(* iMod ("Hclose" with "[$Hfrag_loc $Hauth_a $Hauth_b $Hεfrag]") as "_"; first done. *) +(* iFrame. simpl. *) +(* iModIntro. *) +(* replace (_+0+1)%Z with (bool_to_Z (is_Some_true a)+1); [done|lia]. *) +(* + iIntros. done. *) +(* - iIntros (??) "[Hfrag_a Hfrag_b]". iModIntro. wp_pures. *) +(* iInv loc_nroot as ">(%&Hl&Hloc_auth)" "Hclose". *) +(* iInv both_nroot as ">(%&%&%&Hloc_frag&Hauth_a & Hauth_b & -> & Hεfrag)" "Hclose'". *) +(* iDestruct (ghost_var_agreeZ with "[$][$]") as "->". *) +(* iDestruct (ghost_var_agree' with "[$Hauth_a][$]") as "->". *) +(* iDestruct (ghost_var_agree' with "[$Hauth_b][$]") as "->". *) +(* simpl. *) +(* wp_load. *) +(* iMod ("Hclose'" with "[$Hloc_frag $Hauth_a $Hauth_b $Hεfrag]"); first done. *) +(* iModIntro. *) +(* iMod ("Hclose" with "[Hl Hloc_auth]"); first by iFrame. *) +(* by iApply "HΦ". *) +(* Unshelve. *) +(* all: simpl; try lra. *) +(* + replace (1+_-2)%R with (INR (bool_to_nat (bool_decide (is_Some b)))-1)%R; first apply resource_nonneg. *) +(* case_bool_decide; simpl; lra. *) +(* + replace (_+1-2)%R with (INR (bool_to_nat (bool_decide (is_Some a)))-1)%R; first apply resource_nonneg. *) +(* case_bool_decide; simpl; lra. *) +(* Qed. *) +(* End attempt2. *) Definition half_FAA' (l:loc) α:= diff --git a/theories/coneris/examples/random_counter/impl1.v b/theories/coneris/examples/random_counter/impl1.v index a6713079..c6b91daa 100644 --- a/theories/coneris/examples/random_counter/impl1.v +++ b/theories/coneris/examples/random_counter/impl1.v @@ -1,347 +1,347 @@ -From iris.algebra Require Import frac_auth. -From iris.base_logic.lib Require Import invariants. -From clutch.coneris Require Import coneris hocap random_counter. +(* From iris.algebra Require Import frac_auth. *) +(* From iris.base_logic.lib Require Import invariants. *) +(* From clutch.coneris Require Import coneris hocap random_counter. *) -Set Default Proof Using "Type*". +(* Set Default Proof Using "Type*". *) -Section impl1. +(* Section impl1. *) - Definition new_counter1 : val:= λ: "_", ref #0. - Definition incr_counter1 : val := λ: "l", let: "n" := rand #3 in (FAA "l" "n", "n"). - Definition allocate_tape1 : val := λ: "_", AllocTape #3. - Definition incr_counter_tape1 :val := λ: "l" "α", let: "n" := rand("α") #3 in (FAA "l" "n", "n"). - Definition read_counter1 : val := λ: "l", !"l". - Class counterG1 Σ := CounterG1 { counterG1_error::hocap_errorGS Σ; - counterG1_tapes:: hocap_tapesGS Σ; - counterG1_frac_authR:: inG Σ (frac_authR natR) }. +(* Definition new_counter1 : val:= λ: "_", ref #0. *) +(* Definition incr_counter1 : val := λ: "l", let: "n" := rand #3 in (FAA "l" "n", "n"). *) +(* Definition allocate_tape1 : val := λ: "_", AllocTape #3. *) +(* Definition incr_counter_tape1 :val := λ: "l" "α", let: "n" := rand("α") #3 in (FAA "l" "n", "n"). *) +(* Definition read_counter1 : val := λ: "l", !"l". *) +(* Class counterG1 Σ := CounterG1 { counterG1_error::hocap_errorGS Σ; *) +(* counterG1_tapes:: hocap_tapesGS Σ; *) +(* counterG1_frac_authR:: inG Σ (frac_authR natR) }. *) - Context `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapesGS Σ, !inG Σ (frac_authR natR)}. - Definition counter_inv_pred1 (c:val) γ1 γ2 γ3:= - (∃ (ε:R) m (l:loc) (z:nat), - ↯ ε ∗ ●↯ ε @ γ1 ∗ - ([∗ map] α ↦ t ∈ m, α ↪N ( t.1 ; t.2 )) ∗ ●m@γ2 ∗ - ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ3 (●F z) - )%I. +(* Context `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapesGS Σ, !inG Σ (frac_authR natR)}. *) +(* Definition counter_inv_pred1 (c:val) γ1 γ2 γ3:= *) +(* (∃ (ε:R) m (l:loc) (z:nat), *) +(* ↯ ε ∗ ●↯ ε @ γ1 ∗ *) +(* ([∗ map] α ↦ t ∈ m, α ↪N ( t.1 ; t.2 )) ∗ ●m@γ2 ∗ *) +(* ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ3 (●F z) *) +(* )%I. *) - Lemma new_counter_spec1 E ε N: - {{{ ↯ ε }}} - new_counter1 #() @ E - {{{ (c:val), RET c; - ∃ γ1 γ2 γ3, inv N (counter_inv_pred1 c γ1 γ2 γ3) ∗ - ◯↯ε @ γ1 ∗ own γ3 (◯F 0%nat) - }}}. - Proof. - rewrite /new_counter1. - iIntros (Φ) "Hε HΦ". - wp_pures. - wp_alloc l as "Hl". - iDestruct (ec_valid with "[$]") as "%". - unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "[%γ1 [H1 H2]]". - { lra. } - { simpl. lra. } - simpl. - iMod (hocap_tapes_alloc (∅:gmap _ _)) as "[%γ2 [H3 H4]]". - iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". - { by apply frac_auth_valid. } - replace (#0) with (#0%nat) by done. - iMod (inv_alloc N _ (counter_inv_pred1 (#l) γ1 γ2 γ3) with "[$Hε $Hl $H1 $H3 $H5]") as "#Hinv". - { iSplit; last done. by iApply big_sepM_empty. } - iApply "HΦ". - iExists _, _, _. by iFrame. - Qed. +(* Lemma new_counter_spec1 E ε N: *) +(* {{{ ↯ ε }}} *) +(* new_counter1 #() @ E *) +(* {{{ (c:val), RET c; *) +(* ∃ γ1 γ2 γ3, inv N (counter_inv_pred1 c γ1 γ2 γ3) ∗ *) +(* ◯↯ε @ γ1 ∗ own γ3 (◯F 0%nat) *) +(* }}}. *) +(* Proof. *) +(* rewrite /new_counter1. *) +(* iIntros (Φ) "Hε HΦ". *) +(* wp_pures. *) +(* wp_alloc l as "Hl". *) +(* iDestruct (ec_valid with "[$]") as "%". *) +(* unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "[%γ1 [H1 H2]]". *) +(* { lra. } *) +(* { simpl. lra. } *) +(* simpl. *) +(* iMod (hocap_tapes_alloc (∅:gmap _ _)) as "[%γ2 [H3 H4]]". *) +(* iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". *) +(* { by apply frac_auth_valid. } *) +(* replace (#0) with (#0%nat) by done. *) +(* iMod (inv_alloc N _ (counter_inv_pred1 (#l) γ1 γ2 γ3) with "[$Hε $Hl $H1 $H3 $H5]") as "#Hinv". *) +(* { iSplit; last done. by iApply big_sepM_empty. } *) +(* iApply "HΦ". *) +(* iExists _, _, _. by iFrame. *) +(* Qed. *) - (** Not used in instantiating type class*) - Lemma incr_counter_spec1 E N c γ1 γ2 γ3 (ε2:R -> nat -> R) (P: iProp Σ) (T: nat -> iProp Σ) (Q: nat->nat->iProp Σ): - ↑N ⊆ E-> - (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> - (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → - {{{ inv N (counter_inv_pred1 c γ1 γ2 γ3) ∗ - □(∀ (ε:R) (n : nat), P ∗ ●↯ ε @ γ1 ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨●↯ (ε2 ε n) @ γ1 ∗ T n) ) ∗ - □ (∀ (n z:nat), T n ∗ own γ3 (●F z) ={E∖↑N}=∗ - own γ3 (●F(z+n)%nat)∗ Q z n) ∗ - P - }}} - incr_counter1 c @ E - {{{ (n:nat) (z:nat), RET (#z, #n); Q z n }}}. - Proof. - iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP) HΦ". - rewrite /incr_counter1. - wp_pures. - wp_bind (rand _)%E. - iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". - iDestruct (ec_valid with "[$]") as "[%K1 %K2]". - wp_apply (wp_couple_rand_adv_comp1' _ _ _ _ (λ x, ε2 ε (fin_to_nat x)) with "[$]"). - { intros. naive_solver. } - { rewrite SeriesC_finite_foldr; specialize (Hineq ε K1). simpl; lra. } - iIntros (n) "H1". - iMod ("Hvs1" with "[$]") as "[%|[H2 HT]]". - { iExFalso. iApply ec_contradict; last done. done. } - iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. - iModIntro. wp_pures. - clear -Hsubset. - wp_bind (FAA _ _). - iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". - wp_faa. - iMod ("Hvs2" with "[$]") as "[H6 HQ]". - replace (#(z+n)) with (#(z+n)%nat); last first. - { by rewrite Nat2Z.inj_add. } - iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. - iModIntro. - wp_pures. - by iApply "HΦ". - Qed. +(* (** Not used in instantiating type class*) *) +(* Lemma incr_counter_spec1 E N c γ1 γ2 γ3 (ε2:R -> nat -> R) (P: iProp Σ) (T: nat -> iProp Σ) (Q: nat->nat->iProp Σ): *) +(* ↑N ⊆ E-> *) +(* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) +(* (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → *) +(* {{{ inv N (counter_inv_pred1 c γ1 γ2 γ3) ∗ *) +(* □(∀ (ε:R) (n : nat), P ∗ ●↯ ε @ γ1 ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨●↯ (ε2 ε n) @ γ1 ∗ T n) ) ∗ *) +(* □ (∀ (n z:nat), T n ∗ own γ3 (●F z) ={E∖↑N}=∗ *) +(* own γ3 (●F(z+n)%nat)∗ Q z n) ∗ *) +(* P *) +(* }}} *) +(* incr_counter1 c @ E *) +(* {{{ (n:nat) (z:nat), RET (#z, #n); Q z n }}}. *) +(* Proof. *) +(* iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP) HΦ". *) +(* rewrite /incr_counter1. *) +(* wp_pures. *) +(* wp_bind (rand _)%E. *) +(* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) +(* iDestruct (ec_valid with "[$]") as "[%K1 %K2]". *) +(* wp_apply (wp_couple_rand_adv_comp1' _ _ _ _ (λ x, ε2 ε (fin_to_nat x)) with "[$]"). *) +(* { intros. naive_solver. } *) +(* { rewrite SeriesC_finite_foldr; specialize (Hineq ε K1). simpl; lra. } *) +(* iIntros (n) "H1". *) +(* iMod ("Hvs1" with "[$]") as "[%|[H2 HT]]". *) +(* { iExFalso. iApply ec_contradict; last done. done. } *) +(* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) +(* iModIntro. wp_pures. *) +(* clear -Hsubset. *) +(* wp_bind (FAA _ _). *) +(* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) +(* wp_faa. *) +(* iMod ("Hvs2" with "[$]") as "[H6 HQ]". *) +(* replace (#(z+n)) with (#(z+n)%nat); last first. *) +(* { by rewrite Nat2Z.inj_add. } *) +(* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) +(* iModIntro. *) +(* wp_pures. *) +(* by iApply "HΦ". *) +(* Qed. *) - Lemma allocate_tape_spec1 N E c γ1 γ2 γ3: - ↑N ⊆ E-> - {{{ inv N (counter_inv_pred1 c γ1 γ2 γ3) }}} - allocate_tape1 #() @ E - {{{ (v:val), RET v; - ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ α ◯↪N (3%nat; []) @ γ2 - }}}. - Proof. - iIntros (Hsubset Φ) "#Hinv HΦ". - rewrite /allocate_tape1. - wp_pures. - wp_alloctape α as "Hα". - iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". - iDestruct (hocap_tapes_notin with "[$][$]") as "%". - iMod (hocap_tapes_new with "[$]") as "[H4 H7]"; first done. - iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Hα]") as "_". - { iNext. iSplitL; last done. - rewrite big_sepM_insert; [simpl; iFrame|done]. - } - iApply "HΦ". - by iFrame. - Qed. +(* Lemma allocate_tape_spec1 N E c γ1 γ2 γ3: *) +(* ↑N ⊆ E-> *) +(* {{{ inv N (counter_inv_pred1 c γ1 γ2 γ3) }}} *) +(* allocate_tape1 #() @ E *) +(* {{{ (v:val), RET v; *) +(* ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ α ◯↪N (3%nat; []) @ γ2 *) +(* }}}. *) +(* Proof. *) +(* iIntros (Hsubset Φ) "#Hinv HΦ". *) +(* rewrite /allocate_tape1. *) +(* wp_pures. *) +(* wp_alloctape α as "Hα". *) +(* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) +(* iDestruct (hocap_tapes_notin with "[$][$]") as "%". *) +(* iMod (hocap_tapes_new with "[$]") as "[H4 H7]"; first done. *) +(* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Hα]") as "_". *) +(* { iNext. iSplitL; last done. *) +(* rewrite big_sepM_insert; [simpl; iFrame|done]. *) +(* } *) +(* iApply "HΦ". *) +(* by iFrame. *) +(* Qed. *) - Lemma incr_counter_tape_spec_some1 N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: - ↑N⊆E -> - {{{ inv N (counter_inv_pred1 c γ1 γ2 γ3) ∗ - □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ - own γ3 (●F(z+n)%nat)∗ Q z) ∗ - P ∗ α ◯↪N (3%nat; n::ns) @ γ2 - }}} - incr_counter_tape1 c #lbl:α @ E - {{{ (z:nat), RET (#z, #n); Q z ∗ α ◯↪N (3%nat; ns) @ γ2}}}. - Proof. - iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP & Hα) HΦ". - rewrite /incr_counter_tape1. - wp_pures. - wp_bind (rand(_) _)%E. - iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". - iDestruct (hocap_tapes_agree with "[$][$]") as "%". - erewrite <-(insert_delete m) at 1; last done. - rewrite big_sepM_insert; last apply lookup_delete. - simpl. - iDestruct "H3" as "[Htape H3]". - wp_apply (wp_rand_tape with "[$]"). - iIntros "[Htape %]". - iMod (hocap_tapes_pop with "[$][$]") as "[H4 Hα]". - iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". - { iSplitL; last done. - erewrite <-(insert_delete m) at 2; last done. - iNext. - rewrite insert_insert. - rewrite big_sepM_insert; last apply lookup_delete. iFrame. - } - iModIntro. - wp_pures. - clear -Hsubset. - wp_bind (FAA _ _). - iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". - wp_faa. - iMod ("Hvs" with "[$]") as "[H6 HQ]". - replace (#(z+n)) with (#(z+n)%nat); last first. - { by rewrite Nat2Z.inj_add. } - iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. - iModIntro. wp_pures. - iApply "HΦ". - by iFrame. - Qed. +(* Lemma incr_counter_tape_spec_some1 N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: *) +(* ↑N⊆E -> *) +(* {{{ inv N (counter_inv_pred1 c γ1 γ2 γ3) ∗ *) +(* □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ *) +(* own γ3 (●F(z+n)%nat)∗ Q z) ∗ *) +(* P ∗ α ◯↪N (3%nat; n::ns) @ γ2 *) +(* }}} *) +(* incr_counter_tape1 c #lbl:α @ E *) +(* {{{ (z:nat), RET (#z, #n); Q z ∗ α ◯↪N (3%nat; ns) @ γ2}}}. *) +(* Proof. *) +(* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP & Hα) HΦ". *) +(* rewrite /incr_counter_tape1. *) +(* wp_pures. *) +(* wp_bind (rand(_) _)%E. *) +(* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) +(* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) +(* erewrite <-(insert_delete m) at 1; last done. *) +(* rewrite big_sepM_insert; last apply lookup_delete. *) +(* simpl. *) +(* iDestruct "H3" as "[Htape H3]". *) +(* wp_apply (wp_rand_tape with "[$]"). *) +(* iIntros "[Htape %]". *) +(* iMod (hocap_tapes_pop with "[$][$]") as "[H4 Hα]". *) +(* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". *) +(* { iSplitL; last done. *) +(* erewrite <-(insert_delete m) at 2; last done. *) +(* iNext. *) +(* rewrite insert_insert. *) +(* rewrite big_sepM_insert; last apply lookup_delete. iFrame. *) +(* } *) +(* iModIntro. *) +(* wp_pures. *) +(* clear -Hsubset. *) +(* wp_bind (FAA _ _). *) +(* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) +(* wp_faa. *) +(* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) +(* replace (#(z+n)) with (#(z+n)%nat); last first. *) +(* { by rewrite Nat2Z.inj_add. } *) +(* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) +(* iModIntro. wp_pures. *) +(* iApply "HΦ". *) +(* by iFrame. *) +(* Qed. *) - Lemma counter_presample_spec1 NS E ns α - (ε2 : R -> nat -> R) - (P : iProp Σ) T γ1 γ2 γ3 c: - ↑NS ⊆ E -> - (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> - (∀ (ε:R), 0<= ε ->SeriesC (λ n, if (bool_decide (n≤3%nat)) then 1 / (S 3%nat) * ε2 ε n else 0%R)%R <= ε)%R-> - inv NS (counter_inv_pred1 c γ1 γ2 γ3) -∗ - (□∀ (ε:R) n, (P ∗ ●↯ ε@ γ1) ={E∖↑NS}=∗ - (⌜(1<=ε2 ε n)%R⌝ ∨(●↯ (ε2 ε n) @ γ1 ∗ T (n)))) - -∗ - P -∗ α ◯↪N (3%nat; ns) @ γ2 -∗ - wp_update E (∃ n, T (n) ∗ α◯↪N (3%nat; ns++[n]) @ γ2). - Proof. - iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". - iApply wp_update_state_step_coupl. - iIntros (σ ε) "((Hheap&Htapes)&Hε)". - iMod (inv_acc with "Hinv") as "[>(% & % & % & % & H1 & H2 & H3 & H4 & -> & H5 & H6) Hclose]"; [done|]. - iDestruct (hocap_tapes_agree with "[$][$]") as "%". - erewrite <-(insert_delete m) at 1; last done. - rewrite big_sepM_insert; last apply lookup_delete. - simpl. - iDestruct "H3" as "[Htape H3]". - iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%)". - iDestruct (ec_supply_bound with "[$][$]") as "%". - iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. - iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". - iApply state_step_coupl_state_adv_comp_con_prob_lang; first done. - unshelve iExists (λ x, mknonnegreal (ε2 ε1' x) _). - { apply Hpos. apply cond_nonneg. } - iSplit. - { iPureIntro. - unshelve epose proof (Hineq ε1' _) as H1; first apply cond_nonneg. - by rewrite SeriesC_nat_bounded_fin in H1. } - iIntros (sample). +(* Lemma counter_presample_spec1 NS E ns α *) +(* (ε2 : R -> nat -> R) *) +(* (P : iProp Σ) T γ1 γ2 γ3 c: *) +(* ↑NS ⊆ E -> *) +(* (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> *) +(* (∀ (ε:R), 0<= ε ->SeriesC (λ n, if (bool_decide (n≤3%nat)) then 1 / (S 3%nat) * ε2 ε n else 0%R)%R <= ε)%R-> *) +(* inv NS (counter_inv_pred1 c γ1 γ2 γ3) -∗ *) +(* (□∀ (ε:R) n, (P ∗ ●↯ ε@ γ1) ={E∖↑NS}=∗ *) +(* (⌜(1<=ε2 ε n)%R⌝ ∨(●↯ (ε2 ε n) @ γ1 ∗ T (n)))) *) +(* -∗ *) +(* P -∗ α ◯↪N (3%nat; ns) @ γ2 -∗ *) +(* wp_update E (∃ n, T (n) ∗ α◯↪N (3%nat; ns++[n]) @ γ2). *) +(* Proof. *) +(* iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) +(* iApply wp_update_state_step_coupl. *) +(* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) +(* iMod (inv_acc with "Hinv") as "[>(% & % & % & % & H1 & H2 & H3 & H4 & -> & H5 & H6) Hclose]"; [done|]. *) +(* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) +(* erewrite <-(insert_delete m) at 1; last done. *) +(* rewrite big_sepM_insert; last apply lookup_delete. *) +(* simpl. *) +(* iDestruct "H3" as "[Htape H3]". *) +(* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%)". *) +(* iDestruct (ec_supply_bound with "[$][$]") as "%". *) +(* iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. *) +(* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) +(* iApply state_step_coupl_state_adv_comp_con_prob_lang; first done. *) +(* unshelve iExists (λ x, mknonnegreal (ε2 ε1' x) _). *) +(* { apply Hpos. apply cond_nonneg. } *) +(* iSplit. *) +(* { iPureIntro. *) +(* unshelve epose proof (Hineq ε1' _) as H1; first apply cond_nonneg. *) +(* by rewrite SeriesC_nat_bounded_fin in H1. } *) +(* iIntros (sample). *) - destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' sample))%R 1%R) as [Hdec|Hdec]; last first. - { apply Rnot_lt_ge, Rge_le in Hdec. - iApply state_step_coupl_ret_err_ge_1. - simpl. simpl in *. lra. - } - iApply state_step_coupl_ret. - unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' sample) _) with "Hε_supply") as "[Hε_supply Hε]". - { apply Hpos. apply cond_nonneg. } - { simpl. done. } - iMod (tapeN_update_append _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". - iMod (hocap_tapes_presample _ _ _ _ _ (fin_to_nat sample) with "[$][$]") as "[H4 Hfrag]". - iMod "Hclose'" as "_". - iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". - { iExFalso. iApply (ec_contradict with "[$]"). exact. } - iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4 $H5 $H6]") as "_". - { iNext. iSplit; last done. - rewrite big_sepM_insert_delete; iFrame. - } - iApply fupd_mask_intro_subseteq; first set_solver. - iFrame. - Qed. +(* destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' sample))%R 1%R) as [Hdec|Hdec]; last first. *) +(* { apply Rnot_lt_ge, Rge_le in Hdec. *) +(* iApply state_step_coupl_ret_err_ge_1. *) +(* simpl. simpl in *. lra. *) +(* } *) +(* iApply state_step_coupl_ret. *) +(* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' sample) _) with "Hε_supply") as "[Hε_supply Hε]". *) +(* { apply Hpos. apply cond_nonneg. } *) +(* { simpl. done. } *) +(* iMod (tapeN_update_append _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". *) +(* iMod (hocap_tapes_presample _ _ _ _ _ (fin_to_nat sample) with "[$][$]") as "[H4 Hfrag]". *) +(* iMod "Hclose'" as "_". *) +(* iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". *) +(* { iExFalso. iApply (ec_contradict with "[$]"). exact. } *) +(* iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4 $H5 $H6]") as "_". *) +(* { iNext. iSplit; last done. *) +(* rewrite big_sepM_insert_delete; iFrame. *) +(* } *) +(* iApply fupd_mask_intro_subseteq; first set_solver. *) +(* iFrame. *) +(* Qed. *) - Lemma read_counter_spec1 N E c γ1 γ2 γ3 P Q: - ↑N ⊆ E -> - {{{ inv N (counter_inv_pred1 c γ1 γ2 γ3) ∗ - □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ - own γ3 (●F z)∗ Q z) - ∗ P - }}} - read_counter1 c @ E - {{{ (n':nat), RET #n'; Q n' - }}}. - Proof. - iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP) HΦ". - rewrite /read_counter1. - wp_pure. - iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". - wp_load. - iMod ("Hvs" with "[$]") as "[H6 HQ]". - iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]"); first done. - iApply ("HΦ" with "[$]"). - Qed. +(* Lemma read_counter_spec1 N E c γ1 γ2 γ3 P Q: *) +(* ↑N ⊆ E -> *) +(* {{{ inv N (counter_inv_pred1 c γ1 γ2 γ3) ∗ *) +(* □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ *) +(* own γ3 (●F z)∗ Q z) *) +(* ∗ P *) +(* }}} *) +(* read_counter1 c @ E *) +(* {{{ (n':nat), RET #n'; Q n' *) +(* }}}. *) +(* Proof. *) +(* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP) HΦ". *) +(* rewrite /read_counter1. *) +(* wp_pure. *) +(* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) +(* wp_load. *) +(* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) +(* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]"); first done. *) +(* iApply ("HΦ" with "[$]"). *) +(* Qed. *) -End impl1. +(* End impl1. *) -Program Definition random_counter1 `{!conerisGS Σ}: random_counter := - {| new_counter := new_counter1; - allocate_tape:= allocate_tape1; - incr_counter_tape := incr_counter_tape1; - read_counter:=read_counter1; - counterG := counterG1; - error_name := gname; - tape_name := gname; - counter_name :=gname; - is_counter _ N c γ1 γ2 γ3 := inv N (counter_inv_pred1 c γ1 γ2 γ3); - counter_error_auth _ γ x := ●↯ x @ γ; - counter_error_frag _ γ x := ◯↯ x @ γ; - counter_tapes_auth _ γ m := (●((λ ns, (3, ns))<$>m)@γ)%I; - counter_tapes_frag _ γ α ns := (α◯↪N (3%nat;ns) @ γ)%I; - counter_content_auth _ γ z := own γ (●F z); - counter_content_frag _ γ f z := own γ (◯F{f} z); - new_counter_spec _ := new_counter_spec1; - allocate_tape_spec _ :=allocate_tape_spec1; - incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some1; - counter_presample_spec _ :=counter_presample_spec1; - read_counter_spec _ :=read_counter_spec1 - |}. -Next Obligation. - simpl. - iIntros (??????) "(%&<-&H1)(%&<-&H2)". - iCombine "H1 H2" gives "%H". by rewrite excl_auth.excl_auth_auth_op_valid in H. -Qed. -Next Obligation. - simpl. - iIntros (??????) "(%&<-&H1)(%&<-&H2)". - iCombine "H1 H2" gives "%H". by rewrite excl_auth.excl_auth_frag_op_valid in H. -Qed. -Next Obligation. - simpl. - iIntros (?????) "H". - iApply (hocap_error_auth_pos with "[$]"). -Qed. -Next Obligation. - simpl. - iIntros (?????) "H". - iApply (hocap_error_frag_pos with "[$]"). -Qed. -Next Obligation. - simpl. - iIntros (??????) "H1 H2". - iApply (hocap_error_agree with "[$][$]"). -Qed. -Next Obligation. - simpl. iIntros (???????) "??". - iApply (hocap_error_update with "[$][$]"). -Qed. -Next Obligation. - simpl. - iIntros (??????) "H1 H2". - by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[%H _]". -Qed. -Next Obligation. - simpl. - iIntros (???????) "H1 H2". - iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. - rewrite dfrac_op_own dfrac_valid_own. by intro. -Qed. -Next Obligation. - simpl. - iIntros. - iDestruct (hocap_tapes_agree with "[$][$]") as "%H". - iPureIntro. - rewrite lookup_fmap_Some in H. destruct H as (?&?&?). - by simplify_eq. -Qed. -Next Obligation. - iIntros. - iMod (hocap_tapes_presample with "[$][$]") as "[??]". - iModIntro. iFrame. - by rewrite fmap_insert. -Qed. -Next Obligation. - simpl. - iIntros (??????) "H1 H2". - iCombine "H1 H2" gives "%H". by rewrite auth_auth_op_valid in H. -Qed. -Next Obligation. - simpl. iIntros (???? z z' ?) "H1 H2". - iCombine "H1 H2" gives "%H". - apply frac_auth_included_total in H. iPureIntro. - by apply nat_included. -Qed. -Next Obligation. - simpl. - iIntros (????????). - rewrite frac_auth_frag_op. by rewrite own_op. -Qed. -Next Obligation. - simpl. iIntros (??????) "H1 H2". - iCombine "H1 H2" gives "%H". - iPureIntro. - by apply frac_auth_agree_L in H. -Qed. -Next Obligation. - simpl. iIntros (????????) "H1 H2". - iMod (own_update_2 _ _ _ (_ ⋅ _) with "[$][$]") as "[$$]"; last done. - apply frac_auth_update. - apply nat_local_update. lia. -Qed. +(* Program Definition random_counter1 `{!conerisGS Σ}: random_counter := *) +(* {| new_counter := new_counter1; *) +(* allocate_tape:= allocate_tape1; *) +(* incr_counter_tape := incr_counter_tape1; *) +(* read_counter:=read_counter1; *) +(* counterG := counterG1; *) +(* error_name := gname; *) +(* tape_name := gname; *) +(* counter_name :=gname; *) +(* is_counter _ N c γ1 γ2 γ3 := inv N (counter_inv_pred1 c γ1 γ2 γ3); *) +(* counter_error_auth _ γ x := ●↯ x @ γ; *) +(* counter_error_frag _ γ x := ◯↯ x @ γ; *) +(* counter_tapes_auth _ γ m := (●((λ ns, (3, ns))<$>m)@γ)%I; *) +(* counter_tapes_frag _ γ α ns := (α◯↪N (3%nat;ns) @ γ)%I; *) +(* counter_content_auth _ γ z := own γ (●F z); *) +(* counter_content_frag _ γ f z := own γ (◯F{f} z); *) +(* new_counter_spec _ := new_counter_spec1; *) +(* allocate_tape_spec _ :=allocate_tape_spec1; *) +(* incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some1; *) +(* counter_presample_spec _ :=counter_presample_spec1; *) +(* read_counter_spec _ :=read_counter_spec1 *) +(* |}. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (??????) "(%&<-&H1)(%&<-&H2)". *) +(* iCombine "H1 H2" gives "%H". by rewrite excl_auth.excl_auth_auth_op_valid in H. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (??????) "(%&<-&H1)(%&<-&H2)". *) +(* iCombine "H1 H2" gives "%H". by rewrite excl_auth.excl_auth_frag_op_valid in H. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (?????) "H". *) +(* iApply (hocap_error_auth_pos with "[$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (?????) "H". *) +(* iApply (hocap_error_frag_pos with "[$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (??????) "H1 H2". *) +(* iApply (hocap_error_agree with "[$][$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. iIntros (???????) "??". *) +(* iApply (hocap_error_update with "[$][$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (??????) "H1 H2". *) +(* by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[%H _]". *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (???????) "H1 H2". *) +(* iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. *) +(* rewrite dfrac_op_own dfrac_valid_own. by intro. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros. *) +(* iDestruct (hocap_tapes_agree with "[$][$]") as "%H". *) +(* iPureIntro. *) +(* rewrite lookup_fmap_Some in H. destruct H as (?&?&?). *) +(* by simplify_eq. *) +(* Qed. *) +(* Next Obligation. *) +(* iIntros. *) +(* iMod (hocap_tapes_presample with "[$][$]") as "[??]". *) +(* iModIntro. iFrame. *) +(* by rewrite fmap_insert. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (??????) "H1 H2". *) +(* iCombine "H1 H2" gives "%H". by rewrite auth_auth_op_valid in H. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. iIntros (???? z z' ?) "H1 H2". *) +(* iCombine "H1 H2" gives "%H". *) +(* apply frac_auth_included_total in H. iPureIntro. *) +(* by apply nat_included. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (????????). *) +(* rewrite frac_auth_frag_op. by rewrite own_op. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. iIntros (??????) "H1 H2". *) +(* iCombine "H1 H2" gives "%H". *) +(* iPureIntro. *) +(* by apply frac_auth_agree_L in H. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. iIntros (????????) "H1 H2". *) +(* iMod (own_update_2 _ _ _ (_ ⋅ _) with "[$][$]") as "[$$]"; last done. *) +(* apply frac_auth_update. *) +(* apply nat_local_update. lia. *) +(* Qed. *) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index f5fa6344..5554b146 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -1,474 +1,474 @@ -From iris.algebra Require Import frac_auth. -From iris.base_logic.lib Require Import invariants. -From clutch.coneris Require Import coneris hocap random_counter flip_spec. -From clutch Require Import uniform_list. +(* From iris.algebra Require Import frac_auth. *) +(* From iris.base_logic.lib Require Import invariants. *) +(* From clutch.coneris Require Import coneris hocap random_counter flip_spec. *) +(* From clutch Require Import uniform_list. *) -Set Default Proof Using "Type*". +(* Set Default Proof Using "Type*". *) -Local Definition expander (l:list nat):= - l ≫= (λ x, [2<=?x; (Nat.odd x)]). -Local Lemma expander_eta x: x<4->(Z.of_nat x= Z.of_nat 2*Z.b2z (2<=? x)%nat + Z.b2z (Nat.odd x))%Z. -Proof. - destruct x as [|[|[|[|]]]]; last lia; intros; simpl; lia. -Qed. (* Local Definition expander (l:list nat):= *) -(* l ≫= (λ x, [Nat.div2 x; Nat.b2n (Nat.odd x)]). *) +(* l ≫= (λ x, [2<=?x; (Nat.odd x)]). *) +(* Local Lemma expander_eta x: x<4->(Z.of_nat x= Z.of_nat 2*Z.b2z (2<=? x)%nat + Z.b2z (Nat.odd x))%Z. *) +(* Proof. *) +(* destruct x as [|[|[|[|]]]]; last lia; intros; simpl; lia. *) +(* Qed. *) +(* (* Local Definition expander (l:list nat):= *) *) +(* (* l ≫= (λ x, [Nat.div2 x; Nat.b2n (Nat.odd x)]). *) *) -(* Class hocap_tapesGS' (Σ : gFunctors) := Hocap_tapesGS' { *) -(* hocap_tapesGS_inG' :: ghost_mapG Σ loc (bool* list nat) *) -(* }. *) -(* Definition hocap_tapesΣ' := ghost_mapΣ loc (bool *list nat). *) +(* (* Class hocap_tapesGS' (Σ : gFunctors) := Hocap_tapesGS' { *) *) +(* (* hocap_tapesGS_inG' :: ghost_mapG Σ loc (bool* list nat) *) *) +(* (* }. *) *) +(* (* Definition hocap_tapesΣ' := ghost_mapΣ loc (bool *list nat). *) *) -(* Notation "α ◯↪N ( b , ns ) @ γ":= (α ↪[ γ ] (b, ns))%I *) -(* (at level 20) : bi_scope. *) +(* (* Notation "α ◯↪N ( b , ns ) @ γ":= (α ↪[ γ ] (b, ns))%I *) *) +(* (* (at level 20) : bi_scope. *) *) -(* Notation "● m @ γ" := (ghost_map_auth γ 1 m) (at level 20) : bi_scope. *) +(* (* Notation "● m @ γ" := (ghost_map_auth γ 1 m) (at level 20) : bi_scope. *) *) -(* Section tapes_lemmas. *) -(* Context `{!hocap_tapesGS' Σ}. *) +(* (* Section tapes_lemmas. *) *) +(* (* Context `{!hocap_tapesGS' Σ}. *) *) -(* Lemma hocap_tapes_alloc' m: *) -(* ⊢ |==>∃ γ, (● m @ γ) ∗ [∗ map] k↦v ∈ m, (k ◯↪N (v.1, v.2) @ γ). *) -(* Proof. *) -(* iMod ghost_map_alloc as (γ) "[??]". *) -(* iFrame. iModIntro. *) -(* iApply big_sepM_mono; last done. *) -(* by iIntros (?[??]). *) -(* Qed. *) +(* (* Lemma hocap_tapes_alloc' m: *) *) +(* (* ⊢ |==>∃ γ, (● m @ γ) ∗ [∗ map] k↦v ∈ m, (k ◯↪N (v.1, v.2) @ γ). *) *) +(* (* Proof. *) *) +(* (* iMod ghost_map_alloc as (γ) "[??]". *) *) +(* (* iFrame. iModIntro. *) *) +(* (* iApply big_sepM_mono; last done. *) *) +(* (* by iIntros (?[??]). *) *) +(* (* Qed. *) *) -(* Lemma hocap_tapes_agree' m b γ k ns: *) -(* (● m @ γ) -∗ (k ◯↪N (b, ns) @ γ) -∗ ⌜ m!!k = Some (b, ns) ⌝. *) -(* Proof. *) -(* iIntros "H1 H2". *) -(* by iCombine "H1 H2" gives "%". *) -(* Qed. *) +(* (* Lemma hocap_tapes_agree' m b γ k ns: *) *) +(* (* (● m @ γ) -∗ (k ◯↪N (b, ns) @ γ) -∗ ⌜ m!!k = Some (b, ns) ⌝. *) *) +(* (* Proof. *) *) +(* (* iIntros "H1 H2". *) *) +(* (* by iCombine "H1 H2" gives "%". *) *) +(* (* Qed. *) *) -(* Lemma hocap_tapes_new' γ m k ns b: *) -(* m!!k=None -> ⊢ (● m @ γ) ==∗ (● (<[k:=(b, ns)]>m) @ γ) ∗ (k ◯↪N (b, ns) @ γ). *) -(* Proof. *) -(* iIntros (Hlookup) "H". *) -(* by iApply ghost_map_insert. *) -(* Qed. *) +(* (* Lemma hocap_tapes_new' γ m k ns b: *) *) +(* (* m!!k=None -> ⊢ (● m @ γ) ==∗ (● (<[k:=(b, ns)]>m) @ γ) ∗ (k ◯↪N (b, ns) @ γ). *) *) +(* (* Proof. *) *) +(* (* iIntros (Hlookup) "H". *) *) +(* (* by iApply ghost_map_insert. *) *) +(* (* Qed. *) *) -(* Lemma hocap_tapes_presample' b γ m k ns n: *) -(* (● m @ γ) -∗ (k ◯↪N (b, ns) @ γ) ==∗ (● (<[k:=(b, ns++[n])]>m) @ γ) ∗ (k ◯↪N (b, ns++[n]) @ γ). *) -(* Proof. *) -(* iIntros "H1 H2". *) -(* iApply (ghost_map_update with "[$][$]"). *) -(* Qed. *) +(* (* Lemma hocap_tapes_presample' b γ m k ns n: *) *) +(* (* (● m @ γ) -∗ (k ◯↪N (b, ns) @ γ) ==∗ (● (<[k:=(b, ns++[n])]>m) @ γ) ∗ (k ◯↪N (b, ns++[n]) @ γ). *) *) +(* (* Proof. *) *) +(* (* iIntros "H1 H2". *) *) +(* (* iApply (ghost_map_update with "[$][$]"). *) *) +(* (* Qed. *) *) -(* Lemma hocap_tapes_pop1' γ m k ns: *) -(* (● m @ γ) -∗ (k ◯↪N (true, ns) @ γ) ==∗ (● (<[k:=(false, ns)]>m) @ γ) ∗ (k ◯↪N (false, ns) @ γ). *) -(* Proof. *) -(* iIntros "H1 H2". *) -(* iApply (ghost_map_update with "[$][$]"). *) -(* Qed. *) +(* (* Lemma hocap_tapes_pop1' γ m k ns: *) *) +(* (* (● m @ γ) -∗ (k ◯↪N (true, ns) @ γ) ==∗ (● (<[k:=(false, ns)]>m) @ γ) ∗ (k ◯↪N (false, ns) @ γ). *) *) +(* (* Proof. *) *) +(* (* iIntros "H1 H2". *) *) +(* (* iApply (ghost_map_update with "[$][$]"). *) *) +(* (* Qed. *) *) -(* Lemma hocap_tapes_pop2' γ m k ns n: *) -(* (● m @ γ) -∗ (k ◯↪N (false, n::ns) @ γ) ==∗ (● (<[k:=(true, ns)]>m) @ γ) ∗ (k ◯↪N (true, ns) @ γ). *) -(* Proof. *) -(* iIntros "H1 H2". *) -(* iApply (ghost_map_update with "[$][$]"). *) -(* Qed. *) +(* (* Lemma hocap_tapes_pop2' γ m k ns n: *) *) +(* (* (● m @ γ) -∗ (k ◯↪N (false, n::ns) @ γ) ==∗ (● (<[k:=(true, ns)]>m) @ γ) ∗ (k ◯↪N (true, ns) @ γ). *) *) +(* (* Proof. *) *) +(* (* iIntros "H1 H2". *) *) +(* (* iApply (ghost_map_update with "[$][$]"). *) *) +(* (* Qed. *) *) -(* Lemma hocap_tapes_notin' `{F: flip_spec Σ} {L:flipG Σ} α ns m γ (f:(bool*list nat)-> (list bool)) : *) -(* flip_tapes_frag (L:=L) γ α ns -∗ ([∗ map] α0↦t ∈ m, flip_tapes_frag (L:=L) γ α0 (f t)) -∗ ⌜m!!α=None ⌝. *) -(* Proof. *) -(* destruct (m!!α) eqn:Heqn; last by iIntros. *) -(* iIntros "Hα Hmap". *) -(* iDestruct (big_sepM_lookup with "[$]") as "?"; first done. *) -(* iExFalso. *) -(* by iDestruct (flip_tapes_frag_exclusive with "[$][$]") as "%". *) -(* Qed. *) +(* (* Lemma hocap_tapes_notin' `{F: flip_spec Σ} {L:flipG Σ} α ns m γ (f:(bool*list nat)-> (list bool)) : *) *) +(* (* flip_tapes_frag (L:=L) γ α ns -∗ ([∗ map] α0↦t ∈ m, flip_tapes_frag (L:=L) γ α0 (f t)) -∗ ⌜m!!α=None ⌝. *) *) +(* (* Proof. *) *) +(* (* destruct (m!!α) eqn:Heqn; last by iIntros. *) *) +(* (* iIntros "Hα Hmap". *) *) +(* (* iDestruct (big_sepM_lookup with "[$]") as "?"; first done. *) *) +(* (* iExFalso. *) *) +(* (* by iDestruct (flip_tapes_frag_exclusive with "[$][$]") as "%". *) *) +(* (* Qed. *) *) -(* End tapes_lemmas. *) +(* (* End tapes_lemmas. *) *) -Section impl2. - Context `{F: flip_spec Σ}. - Definition new_counter2 : val:= λ: "_", ref #0. - (* Definition incr_counter2 : val := λ: "l", let: "n" := rand #1 in *) - (* let: "n'" := rand #1 in *) - (* let: "x" := #2 * "n" + "n'" in *) - (* (FAA "l" "x", "x"). *) - Definition allocate_tape2 : val := flip_allocate_tape. - Definition incr_counter_tape2 :val := λ: "l" "α", let: "n" := - conversion.bool_to_int (flip_tape "α") - in - let: "n'" := - conversion.bool_to_int (flip_tape "α") - in - let: "x" := #2 * "n" + "n'" in - (FAA "l" "x", "x"). - Definition read_counter2 : val := λ: "l", !"l". - Class counterG2 Σ := CounterG2 { (* counterG2_tapes:: hocap_tapesGS' Σ; *) - counterG2_frac_authR:: inG Σ (frac_authR natR); - counterG2_flipG: flipG Σ - }. +(* Section impl2. *) +(* Context `{F: flip_spec Σ}. *) +(* Definition new_counter2 : val:= λ: "_", ref #0. *) +(* (* Definition incr_counter2 : val := λ: "l", let: "n" := rand #1 in *) *) +(* (* let: "n'" := rand #1 in *) *) +(* (* let: "x" := #2 * "n" + "n'" in *) *) +(* (* (FAA "l" "x", "x"). *) *) +(* Definition allocate_tape2 : val := flip_allocate_tape. *) +(* Definition incr_counter_tape2 :val := λ: "l" "α", let: "n" := *) +(* conversion.bool_to_int (flip_tape "α") *) +(* in *) +(* let: "n'" := *) +(* conversion.bool_to_int (flip_tape "α") *) +(* in *) +(* let: "x" := #2 * "n" + "n'" in *) +(* (FAA "l" "x", "x"). *) +(* Definition read_counter2 : val := λ: "l", !"l". *) +(* Class counterG2 Σ := CounterG2 { (* counterG2_tapes:: hocap_tapesGS' Σ; *) *) +(* counterG2_frac_authR:: inG Σ (frac_authR natR); *) +(* counterG2_flipG: flipG Σ *) +(* }. *) - Context `{L:!flipG Σ, !inG Σ (frac_authR natR)}. +(* Context `{L:!flipG Σ, !inG Σ (frac_authR natR)}. *) - Definition counter_inv_pred2 (c:val) γ := - (∃ (* (m:gmap loc (bool*list nat)) *) (l:loc) (z:nat), - (* ([∗ map] α ↦ t ∈ m, flip_tapes_frag (L:=L) γ1' α (if t.1:bool then nat_to_bool <$> (expander t.2) else nat_to_bool <$> (drop 1%nat (expander t.2))) ) *) - (* ∗ ●m@γ1 ∗ *) - ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ (●F z) - )%I. +(* Definition counter_inv_pred2 (c:val) γ := *) +(* (∃ (* (m:gmap loc (bool*list nat)) *) (l:loc) (z:nat), *) +(* (* ([∗ map] α ↦ t ∈ m, flip_tapes_frag (L:=L) γ1' α (if t.1:bool then nat_to_bool <$> (expander t.2) else nat_to_bool <$> (drop 1%nat (expander t.2))) ) *) *) +(* (* ∗ ●m@γ1 ∗ *) *) +(* ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ (●F z) *) +(* )%I. *) - Lemma new_counter_spec2 E ε N: - {{{ ↯ ε }}} - new_counter2 #() @ E - {{{ (c:val), RET c; - ∃ γ1 γ2 γ3, (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ - inv (N.@ "counter") (counter_inv_pred2 c γ3)) ∗ - flip_error_frag (L:=L) γ1 ε ∗ own γ3 (◯F 0%nat) - }}}. - Proof. - rewrite /new_counter2. - iIntros (Φ) "Hε HΦ". - wp_pures. - iDestruct (ec_valid with "[$]") as "%". - iMod (flip_inv_create_spec with "[$]") as "(%γ1'&%γ2'&#?&Hε)". - wp_alloc l as "Hl". - simpl. - (* iMod (hocap_tapes_alloc' (∅:gmap _ _)) as "[%γ2 [H3 H4]]". *) - iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". - { by apply frac_auth_valid. } - replace (#0) with (#0%nat) by done. - iMod (inv_alloc (N.@"counter") _ (counter_inv_pred2 (#l) γ3) with "[$Hl $H5]") as "#Hinv"; first done. - iApply "HΦ". - iExists _, _, _. iModIntro. iFrame. - by iSplit. - Qed. +(* Lemma new_counter_spec2 E ε N: *) +(* {{{ ↯ ε }}} *) +(* new_counter2 #() @ E *) +(* {{{ (c:val), RET c; *) +(* ∃ γ1 γ2 γ3, (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ *) +(* inv (N.@ "counter") (counter_inv_pred2 c γ3)) ∗ *) +(* flip_error_frag (L:=L) γ1 ε ∗ own γ3 (◯F 0%nat) *) +(* }}}. *) +(* Proof. *) +(* rewrite /new_counter2. *) +(* iIntros (Φ) "Hε HΦ". *) +(* wp_pures. *) +(* iDestruct (ec_valid with "[$]") as "%". *) +(* iMod (flip_inv_create_spec with "[$]") as "(%γ1'&%γ2'&#?&Hε)". *) +(* wp_alloc l as "Hl". *) +(* simpl. *) +(* (* iMod (hocap_tapes_alloc' (∅:gmap _ _)) as "[%γ2 [H3 H4]]". *) *) +(* iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". *) +(* { by apply frac_auth_valid. } *) +(* replace (#0) with (#0%nat) by done. *) +(* iMod (inv_alloc (N.@"counter") _ (counter_inv_pred2 (#l) γ3) with "[$Hl $H5]") as "#Hinv"; first done. *) +(* iApply "HΦ". *) +(* iExists _, _, _. iModIntro. iFrame. *) +(* by iSplit. *) +(* Qed. *) - (** This lemma is not possible as only one view shift*) - (* Lemma incr_counter_spec2 E N c γ1 γ2 γ3 (ε2:R -> nat -> R) (P: iProp Σ) (T: nat -> iProp Σ) (Q: nat->nat->iProp Σ): *) - (* ↑N ⊆ E-> *) - (* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) - (* (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → *) - (* {{{ inv N (counter_inv_pred2 c γ1 γ2 γ3) ∗ *) - (* □(∀ (ε:R) (n : nat), P ∗ ●↯ ε @ γ1 ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨●↯ (ε2 ε n) @ γ1 ∗ T n) ) ∗ *) - (* □ (∀ (n z:nat), T n ∗ own γ3 (●F z) ={E∖↑N}=∗ *) - (* own γ3 (●F(z+n)%nat)∗ Q z n) ∗ *) - (* P *) - (* }}} *) - (* incr_counter2 c @ E *) - (* {{{ (n:nat) (z:nat), RET (#z, #n); Q z n }}}. *) - (* Proof. *) - (* iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP) HΦ". *) - (* rewrite /incr_counter2. *) - (* wp_pures. *) - (* wp_bind (rand _)%E. *) - (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) - (* (** cant do two view shifts! *) *) - (* Abort. *) +(* (** This lemma is not possible as only one view shift*) *) +(* (* Lemma incr_counter_spec2 E N c γ1 γ2 γ3 (ε2:R -> nat -> R) (P: iProp Σ) (T: nat -> iProp Σ) (Q: nat->nat->iProp Σ): *) *) +(* (* ↑N ⊆ E-> *) *) +(* (* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) *) +(* (* (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → *) *) +(* (* {{{ inv N (counter_inv_pred2 c γ1 γ2 γ3) ∗ *) *) +(* (* □(∀ (ε:R) (n : nat), P ∗ ●↯ ε @ γ1 ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨●↯ (ε2 ε n) @ γ1 ∗ T n) ) ∗ *) *) +(* (* □ (∀ (n z:nat), T n ∗ own γ3 (●F z) ={E∖↑N}=∗ *) *) +(* (* own γ3 (●F(z+n)%nat)∗ Q z n) ∗ *) *) +(* (* P *) *) +(* (* }}} *) *) +(* (* incr_counter2 c @ E *) *) +(* (* {{{ (n:nat) (z:nat), RET (#z, #n); Q z n }}}. *) *) +(* (* Proof. *) *) +(* (* iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP) HΦ". *) *) +(* (* rewrite /incr_counter2. *) *) +(* (* wp_pures. *) *) +(* (* wp_bind (rand _)%E. *) *) +(* (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) *) +(* (* (** cant do two view shifts! *) *) *) +(* (* Abort. *) *) - Lemma allocate_tape_spec2 N E c γ1 γ2 γ3: - ↑N ⊆ E-> - {{{ (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ - inv (N.@ "counter") (counter_inv_pred2 c γ3)) }}} - allocate_tape2 #() @ E - {{{ (v:val), RET v; - ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ (flip_tapes_frag (L:=L) γ2 α (expander []) ∗ ⌜Forall (λ x, x<4) []⌝) - }}}. - Proof. - iIntros (Hsubset Φ) "(#Hinv & #Hinv') HΦ". - rewrite /allocate_tape2. - iApply wptac_wp_fupd. - iApply (flip_allocate_tape_spec); [by eapply nclose_subseteq'|done|]. - iNext. - iIntros (?) "(%α & -> & Hfrag)". - iApply "HΦ". - iFrame. - iPureIntro. - split; first done. - by apply Forall_nil. - Qed. +(* Lemma allocate_tape_spec2 N E c γ1 γ2 γ3: *) +(* ↑N ⊆ E-> *) +(* {{{ (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ *) +(* inv (N.@ "counter") (counter_inv_pred2 c γ3)) }}} *) +(* allocate_tape2 #() @ E *) +(* {{{ (v:val), RET v; *) +(* ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ (flip_tapes_frag (L:=L) γ2 α (expander []) ∗ ⌜Forall (λ x, x<4) []⌝) *) +(* }}}. *) +(* Proof. *) +(* iIntros (Hsubset Φ) "(#Hinv & #Hinv') HΦ". *) +(* rewrite /allocate_tape2. *) +(* iApply wptac_wp_fupd. *) +(* iApply (flip_allocate_tape_spec); [by eapply nclose_subseteq'|done|]. *) +(* iNext. *) +(* iIntros (?) "(%α & -> & Hfrag)". *) +(* iApply "HΦ". *) +(* iFrame. *) +(* iPureIntro. *) +(* split; first done. *) +(* by apply Forall_nil. *) +(* Qed. *) - Lemma incr_counter_tape_spec_some2 N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: - ↑N⊆E -> - {{{ (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ - inv (N.@ "counter") (counter_inv_pred2 c γ3)) ∗ - □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ - own γ3 (●F(z+n)%nat)∗ Q z) ∗ - P ∗ (flip_tapes_frag (L:=L) γ2 α (expander (n::ns))∗ ⌜Forall (λ x, x<4) (n::ns)⌝) - }}} - incr_counter_tape2 c #lbl:α @ E - {{{ (z:nat), RET (#z, #n); Q z ∗ (flip_tapes_frag (L:=L) γ2 α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝)}}}. - Proof. - iIntros (Hsubset Φ) "((#Hinv & #Hinv') & #Hvs & HP & [Hα %]) HΦ". - rewrite /incr_counter_tape2. - wp_pures. - wp_apply (flip_tape_spec_some _ _ _ _ (flip_tapes_frag γ2 α (expander (n::ns)) ) (flip_tapes_frag γ2 α (drop 1%nat (expander (n::ns))))with "[$Hα]"); [|iSplit; first done|]. - { by apply nclose_subseteq'. } - { iModIntro. iIntros (m) "[Hfrag Hauth]". - iDestruct (flip_tapes_agree with "[$][$]") as "%". - simpl. - iMod (flip_tapes_pop with "[$][$]") as "[Hauth Htape]". - iModIntro. - simpl in *. - by iFrame. - } - iIntros "Hα". - wp_apply (conversion.wp_bool_to_int) as "_"; first done. - wp_pures. - wp_apply (flip_tape_spec_some _ _ _ _ (flip_tapes_frag γ2 α (drop 1%nat (expander (n::ns)))) (flip_tapes_frag γ2 α (expander ns))with "[$Hα]"); [|iSplit; first done|]. - { by apply nclose_subseteq'. } - { iModIntro. iIntros (m) "[Hfrag Hauth]". - iDestruct (flip_tapes_agree with "[$][$]") as "%". - simpl. - iMod (flip_tapes_pop with "[$][$]") as "[Hauth Htape]". - iModIntro. - simpl in *. - by iFrame. - } - iIntros "Hα". - wp_apply (conversion.wp_bool_to_int) as "_"; first done. - wp_pures. - wp_bind (FAA _ _). - iInv "Hinv'" as ">(%&%&->&?&?)" "Hclose". - wp_faa. simpl. - iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". - { apply difference_mono; [done|by apply nclose_subseteq']. } - iMod ("Hvs" with "[$]") as "[H6 HQ]". - replace 2%Z with (Z.of_nat 2%nat) by done. - replace (_*_+_)%Z with (Z.of_nat n); last first. - { assert (n<4). - - by eapply (@Forall_inv _ (λ x, x<4)). - - by apply expander_eta. - } - replace (#(z+n)) with (#(z+n)%nat); last first. - { by rewrite Nat2Z.inj_add. } - iMod "Hclose'" as "_". - iMod ("Hclose" with "[-HΦ HQ Hα]") as "_"; first by iFrame. - iModIntro. - wp_pures. - iApply "HΦ". - iFrame. - iPureIntro. - by eapply Forall_inv_tail. - Qed. +(* Lemma incr_counter_tape_spec_some2 N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: *) +(* ↑N⊆E -> *) +(* {{{ (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ *) +(* inv (N.@ "counter") (counter_inv_pred2 c γ3)) ∗ *) +(* □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ *) +(* own γ3 (●F(z+n)%nat)∗ Q z) ∗ *) +(* P ∗ (flip_tapes_frag (L:=L) γ2 α (expander (n::ns))∗ ⌜Forall (λ x, x<4) (n::ns)⌝) *) +(* }}} *) +(* incr_counter_tape2 c #lbl:α @ E *) +(* {{{ (z:nat), RET (#z, #n); Q z ∗ (flip_tapes_frag (L:=L) γ2 α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝)}}}. *) +(* Proof. *) +(* iIntros (Hsubset Φ) "((#Hinv & #Hinv') & #Hvs & HP & [Hα %]) HΦ". *) +(* rewrite /incr_counter_tape2. *) +(* wp_pures. *) +(* wp_apply (flip_tape_spec_some _ _ _ _ (flip_tapes_frag γ2 α (expander (n::ns)) ) (flip_tapes_frag γ2 α (drop 1%nat (expander (n::ns))))with "[$Hα]"); [|iSplit; first done|]. *) +(* { by apply nclose_subseteq'. } *) +(* { iModIntro. iIntros (m) "[Hfrag Hauth]". *) +(* iDestruct (flip_tapes_agree with "[$][$]") as "%". *) +(* simpl. *) +(* iMod (flip_tapes_pop with "[$][$]") as "[Hauth Htape]". *) +(* iModIntro. *) +(* simpl in *. *) +(* by iFrame. *) +(* } *) +(* iIntros "Hα". *) +(* wp_apply (conversion.wp_bool_to_int) as "_"; first done. *) +(* wp_pures. *) +(* wp_apply (flip_tape_spec_some _ _ _ _ (flip_tapes_frag γ2 α (drop 1%nat (expander (n::ns)))) (flip_tapes_frag γ2 α (expander ns))with "[$Hα]"); [|iSplit; first done|]. *) +(* { by apply nclose_subseteq'. } *) +(* { iModIntro. iIntros (m) "[Hfrag Hauth]". *) +(* iDestruct (flip_tapes_agree with "[$][$]") as "%". *) +(* simpl. *) +(* iMod (flip_tapes_pop with "[$][$]") as "[Hauth Htape]". *) +(* iModIntro. *) +(* simpl in *. *) +(* by iFrame. *) +(* } *) +(* iIntros "Hα". *) +(* wp_apply (conversion.wp_bool_to_int) as "_"; first done. *) +(* wp_pures. *) +(* wp_bind (FAA _ _). *) +(* iInv "Hinv'" as ">(%&%&->&?&?)" "Hclose". *) +(* wp_faa. simpl. *) +(* iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". *) +(* { apply difference_mono; [done|by apply nclose_subseteq']. } *) +(* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) +(* replace 2%Z with (Z.of_nat 2%nat) by done. *) +(* replace (_*_+_)%Z with (Z.of_nat n); last first. *) +(* { assert (n<4). *) +(* - by eapply (@Forall_inv _ (λ x, x<4)). *) +(* - by apply expander_eta. *) +(* } *) +(* replace (#(z+n)) with (#(z+n)%nat); last first. *) +(* { by rewrite Nat2Z.inj_add. } *) +(* iMod "Hclose'" as "_". *) +(* iMod ("Hclose" with "[-HΦ HQ Hα]") as "_"; first by iFrame. *) +(* iModIntro. *) +(* wp_pures. *) +(* iApply "HΦ". *) +(* iFrame. *) +(* iPureIntro. *) +(* by eapply Forall_inv_tail. *) +(* Qed. *) - Lemma counter_presample_spec2 N E ns α - (ε2 : R -> nat -> R) - (P : iProp Σ) T γ1 γ2 γ3 c: - ↑N ⊆ E -> - (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> - (∀ (ε:R), 0<= ε ->SeriesC (λ n, if (bool_decide (n≤3%nat)) then 1 / (S 3%nat) * ε2 ε n else 0%R)%R <= ε)%R-> - (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ - inv (N.@ "counter") (counter_inv_pred2 c γ3)) -∗ - (□∀ (ε:R) n, (P ∗ flip_error_auth (L:=L) γ1 ε) ={E∖↑N}=∗ - (⌜(1<=ε2 ε n)%R⌝ ∨(flip_error_auth (L:=L) γ1 (ε2 ε n) ∗ T (n)))) - -∗ - P -∗ (flip_tapes_frag (L:=L) γ2 α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝) -∗ - wp_update E (∃ n, T (n) ∗ (flip_tapes_frag (L:=L) γ2 α (expander (ns++[n])) ∗ ⌜Forall (λ x, x<4) (ns++[n])⌝)). - Proof. - iIntros (Hsubset Hpos Hineq) "(#Hinv & #Hinv') #Hvs HP Hfrag". - iApply fupd_wp_update. - Admitted. - (* iApply flip_presample_spec. *) - (* iApply wp_update_state_step_coupl. *) - (* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) - (* iMod (inv_acc with "Hinv") as "[>(% & % & % & % & H1 & H2 & H3 & H4 & -> & H5 & H6) Hclose]"; [done|]. *) - (* iDestruct (hocap_tapes_agree' with "[$][$]") as "%". *) - (* erewrite <-(insert_delete m) at 1; last done. *) - (* rewrite big_sepM_insert; last apply lookup_delete. *) - (* simpl. *) - (* iDestruct "H3" as "[Htape H3]". *) - (* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%H1)". *) - (* iDestruct (ec_supply_bound with "[$][$]") as "%". *) - (* iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. *) - (* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) - (* iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang 2%nat); first done. *) - (* pose (f (l:list (fin 2)) := (match l with *) - (* | x::[x'] => 2*fin_to_nat x + fin_to_nat x' *) - (* | _ => 0 *) - (* end)). *) - (* unshelve iExists *) - (* (λ l, mknonnegreal (ε2 ε1' (f l) ) _). *) - (* { apply Hpos; simpl. auto. } *) - (* simpl. *) - (* iSplit. *) - (* { iPureIntro. *) - (* etrans; last apply Hineq; last auto. *) - (* pose (k:=(λ n, 1 / ((1 + 1) * ((1 + 1) * 1)) * ε2 ε1' (f n))%R). *) - (* erewrite (SeriesC_ext _ (λ x, if bool_decide (x ∈ enum_uniform_list 1%nat 2%nat) *) - (* then k x else 0%R))%R; last first. *) - (* - intros. case_bool_decide as K. *) - (* + rewrite elem_of_enum_uniform_list in K. by rewrite K /k. *) - (* + rewrite elem_of_enum_uniform_list in K. *) - (* case_match eqn:K'; [by rewrite Nat.eqb_eq in K'|done]. *) - (* - rewrite SeriesC_list; last apply NoDup_enum_uniform_list. *) - (* rewrite /k /=. rewrite SeriesC_nat_bounded'/=. lra. *) - (* } *) - (* iIntros (sample ?). *) - (* destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' (f sample)))%R 1%R) as [Hdec|Hdec]; last first. *) - (* { apply Rnot_lt_ge, Rge_le in Hdec. *) - (* iApply state_step_coupl_ret_err_ge_1. *) - (* simpl. simpl in *. lra. *) - (* } *) - (* iApply state_step_coupl_ret. *) - (* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' (f sample)) _) with "Hε_supply") as "[Hε_supply Hε]". *) - (* { apply Hpos. apply cond_nonneg. } *) - (* { simpl. done. } *) - (* assert (Nat.div2 (f sample)<2)%nat as K1. *) - (* { rewrite Nat.div2_div. apply Nat.Div0.div_lt_upper_bound. rewrite /f. *) - (* simpl. repeat case_match; try lia. pose proof fin_to_nat_lt t. pose proof fin_to_nat_lt t0. lia. *) - (* } *) - (* rewrite -H1. *) - (* iMod (tapeN_update_append _ _ _ _ (nat_to_fin K1) with "[$][$]") as "[Htapes Htape]". *) - (* assert (Nat.b2n (Nat.odd (f sample))<2)%nat as K2. *) - (* { edestruct (Nat.odd _); simpl; lia. } *) - (* rewrite -(list_fmap_singleton (fin_to_nat)) -fmap_app. *) - (* iMod (tapeN_update_append _ _ _ _ (nat_to_fin K2) with "[$][$]") as "[Htapes Htape]". *) - (* iMod (hocap_tapes_presample' _ _ _ _ _ (f sample) with "[$][$]") as "[H4 Hfrag]". *) - (* iMod "Hclose'" as "_". *) - (* iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". *) - (* { iExFalso. iApply (ec_contradict with "[$]"). exact. } *) - (* rewrite insert_insert. *) - (* rewrite fmap_app -!app_assoc /=. *) - (* assert (([nat_to_fin K1;nat_to_fin K2]) = sample) as K. *) - (* { destruct sample as [|x xs]; first done. *) - (* destruct xs as [|x' xs]; first done. *) - (* destruct xs as [|]; last done. *) - (* pose proof fin_to_nat_lt x. pose proof fin_to_nat_lt x'. *) - (* repeat f_equal; apply fin_to_nat_inj; rewrite fin_to_nat_to_fin. *) - (* - rewrite /f. rewrite Nat.div2_div. *) - (* rewrite Nat.mul_comm Nat.div_add_l; last lia. rewrite Nat.div_small; lia. *) - (* - rewrite /f. rewrite Nat.add_comm Nat.odd_add_even. *) - (* + destruct (fin_to_nat x') as [|[|]]; simpl; lia. *) - (* + by econstructor. *) - (* } *) - (* iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4 $H5 $H6]") as "_". *) - (* { iNext. iSplit; last done. *) - (* rewrite big_sepM_insert_delete; iFrame. *) - (* simpl. rewrite !fin_to_nat_to_fin. *) - (* rewrite /expander bind_app -/(expander ns). simpl. by rewrite H1. *) - (* } *) - (* iApply fupd_mask_intro_subseteq; first set_solver. *) - (* rewrite K. *) - (* iFrame. *) - (* Qed. *) +(* Lemma counter_presample_spec2 N E ns α *) +(* (ε2 : R -> nat -> R) *) +(* (P : iProp Σ) T γ1 γ2 γ3 c: *) +(* ↑N ⊆ E -> *) +(* (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> *) +(* (∀ (ε:R), 0<= ε ->SeriesC (λ n, if (bool_decide (n≤3%nat)) then 1 / (S 3%nat) * ε2 ε n else 0%R)%R <= ε)%R-> *) +(* (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ *) +(* inv (N.@ "counter") (counter_inv_pred2 c γ3)) -∗ *) +(* (□∀ (ε:R) n, (P ∗ flip_error_auth (L:=L) γ1 ε) ={E∖↑N}=∗ *) +(* (⌜(1<=ε2 ε n)%R⌝ ∨(flip_error_auth (L:=L) γ1 (ε2 ε n) ∗ T (n)))) *) +(* -∗ *) +(* P -∗ (flip_tapes_frag (L:=L) γ2 α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝) -∗ *) +(* wp_update E (∃ n, T (n) ∗ (flip_tapes_frag (L:=L) γ2 α (expander (ns++[n])) ∗ ⌜Forall (λ x, x<4) (ns++[n])⌝)). *) +(* Proof. *) +(* iIntros (Hsubset Hpos Hineq) "(#Hinv & #Hinv') #Hvs HP Hfrag". *) +(* iApply fupd_wp_update. *) +(* Admitted. *) +(* (* iApply flip_presample_spec. *) *) +(* (* iApply wp_update_state_step_coupl. *) *) +(* (* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) *) +(* (* iMod (inv_acc with "Hinv") as "[>(% & % & % & % & H1 & H2 & H3 & H4 & -> & H5 & H6) Hclose]"; [done|]. *) *) +(* (* iDestruct (hocap_tapes_agree' with "[$][$]") as "%". *) *) +(* (* erewrite <-(insert_delete m) at 1; last done. *) *) +(* (* rewrite big_sepM_insert; last apply lookup_delete. *) *) +(* (* simpl. *) *) +(* (* iDestruct "H3" as "[Htape H3]". *) *) +(* (* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%H1)". *) *) +(* (* iDestruct (ec_supply_bound with "[$][$]") as "%". *) *) +(* (* iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. *) *) +(* (* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) *) +(* (* iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang 2%nat); first done. *) *) +(* (* pose (f (l:list (fin 2)) := (match l with *) *) +(* (* | x::[x'] => 2*fin_to_nat x + fin_to_nat x' *) *) +(* (* | _ => 0 *) *) +(* (* end)). *) *) +(* (* unshelve iExists *) *) +(* (* (λ l, mknonnegreal (ε2 ε1' (f l) ) _). *) *) +(* (* { apply Hpos; simpl. auto. } *) *) +(* (* simpl. *) *) +(* (* iSplit. *) *) +(* (* { iPureIntro. *) *) +(* (* etrans; last apply Hineq; last auto. *) *) +(* (* pose (k:=(λ n, 1 / ((1 + 1) * ((1 + 1) * 1)) * ε2 ε1' (f n))%R). *) *) +(* (* erewrite (SeriesC_ext _ (λ x, if bool_decide (x ∈ enum_uniform_list 1%nat 2%nat) *) *) +(* (* then k x else 0%R))%R; last first. *) *) +(* (* - intros. case_bool_decide as K. *) *) +(* (* + rewrite elem_of_enum_uniform_list in K. by rewrite K /k. *) *) +(* (* + rewrite elem_of_enum_uniform_list in K. *) *) +(* (* case_match eqn:K'; [by rewrite Nat.eqb_eq in K'|done]. *) *) +(* (* - rewrite SeriesC_list; last apply NoDup_enum_uniform_list. *) *) +(* (* rewrite /k /=. rewrite SeriesC_nat_bounded'/=. lra. *) *) +(* (* } *) *) +(* (* iIntros (sample ?). *) *) +(* (* destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' (f sample)))%R 1%R) as [Hdec|Hdec]; last first. *) *) +(* (* { apply Rnot_lt_ge, Rge_le in Hdec. *) *) +(* (* iApply state_step_coupl_ret_err_ge_1. *) *) +(* (* simpl. simpl in *. lra. *) *) +(* (* } *) *) +(* (* iApply state_step_coupl_ret. *) *) +(* (* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' (f sample)) _) with "Hε_supply") as "[Hε_supply Hε]". *) *) +(* (* { apply Hpos. apply cond_nonneg. } *) *) +(* (* { simpl. done. } *) *) +(* (* assert (Nat.div2 (f sample)<2)%nat as K1. *) *) +(* (* { rewrite Nat.div2_div. apply Nat.Div0.div_lt_upper_bound. rewrite /f. *) *) +(* (* simpl. repeat case_match; try lia. pose proof fin_to_nat_lt t. pose proof fin_to_nat_lt t0. lia. *) *) +(* (* } *) *) +(* (* rewrite -H1. *) *) +(* (* iMod (tapeN_update_append _ _ _ _ (nat_to_fin K1) with "[$][$]") as "[Htapes Htape]". *) *) +(* (* assert (Nat.b2n (Nat.odd (f sample))<2)%nat as K2. *) *) +(* (* { edestruct (Nat.odd _); simpl; lia. } *) *) +(* (* rewrite -(list_fmap_singleton (fin_to_nat)) -fmap_app. *) *) +(* (* iMod (tapeN_update_append _ _ _ _ (nat_to_fin K2) with "[$][$]") as "[Htapes Htape]". *) *) +(* (* iMod (hocap_tapes_presample' _ _ _ _ _ (f sample) with "[$][$]") as "[H4 Hfrag]". *) *) +(* (* iMod "Hclose'" as "_". *) *) +(* (* iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". *) *) +(* (* { iExFalso. iApply (ec_contradict with "[$]"). exact. } *) *) +(* (* rewrite insert_insert. *) *) +(* (* rewrite fmap_app -!app_assoc /=. *) *) +(* (* assert (([nat_to_fin K1;nat_to_fin K2]) = sample) as K. *) *) +(* (* { destruct sample as [|x xs]; first done. *) *) +(* (* destruct xs as [|x' xs]; first done. *) *) +(* (* destruct xs as [|]; last done. *) *) +(* (* pose proof fin_to_nat_lt x. pose proof fin_to_nat_lt x'. *) *) +(* (* repeat f_equal; apply fin_to_nat_inj; rewrite fin_to_nat_to_fin. *) *) +(* (* - rewrite /f. rewrite Nat.div2_div. *) *) +(* (* rewrite Nat.mul_comm Nat.div_add_l; last lia. rewrite Nat.div_small; lia. *) *) +(* (* - rewrite /f. rewrite Nat.add_comm Nat.odd_add_even. *) *) +(* (* + destruct (fin_to_nat x') as [|[|]]; simpl; lia. *) *) +(* (* + by econstructor. *) *) +(* (* } *) *) +(* (* iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4 $H5 $H6]") as "_". *) *) +(* (* { iNext. iSplit; last done. *) *) +(* (* rewrite big_sepM_insert_delete; iFrame. *) *) +(* (* simpl. rewrite !fin_to_nat_to_fin. *) *) +(* (* rewrite /expander bind_app -/(expander ns). simpl. by rewrite H1. *) *) +(* (* } *) *) +(* (* iApply fupd_mask_intro_subseteq; first set_solver. *) *) +(* (* rewrite K. *) *) +(* (* iFrame. *) *) +(* (* Qed. *) *) - Lemma read_counter_spec2 N E c γ1 γ2 γ3 P Q: - ↑N ⊆ E -> - {{{ (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ - inv (N.@ "counter") (counter_inv_pred2 c γ3)) ∗ - □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ - own γ3 (●F z)∗ Q z) - ∗ P - }}} - read_counter2 c @ E - {{{ (n':nat), RET #n'; Q n' - }}}. - Proof. - iIntros (Hsubset Φ) "((#Hinv & #Hinv') & #Hvs & HP) HΦ". - rewrite /read_counter2. - wp_pure. - iInv "Hinv'" as ">(%l&%z & -> & Hloc & Hcont)" "Hclose". - wp_load. - iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". - { apply difference_mono_l. by apply nclose_subseteq'. } - iMod ("Hvs" with "[$]") as "[Hcont HQ]". - iMod "Hclose'". - iMod ("Hclose" with "[ $Hloc $Hcont]"); first done. - iApply "HΦ". by iFrame. - Qed. +(* Lemma read_counter_spec2 N E c γ1 γ2 γ3 P Q: *) +(* ↑N ⊆ E -> *) +(* {{{ (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ *) +(* inv (N.@ "counter") (counter_inv_pred2 c γ3)) ∗ *) +(* □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ *) +(* own γ3 (●F z)∗ Q z) *) +(* ∗ P *) +(* }}} *) +(* read_counter2 c @ E *) +(* {{{ (n':nat), RET #n'; Q n' *) +(* }}}. *) +(* Proof. *) +(* iIntros (Hsubset Φ) "((#Hinv & #Hinv') & #Hvs & HP) HΦ". *) +(* rewrite /read_counter2. *) +(* wp_pure. *) +(* iInv "Hinv'" as ">(%l&%z & -> & Hloc & Hcont)" "Hclose". *) +(* wp_load. *) +(* iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". *) +(* { apply difference_mono_l. by apply nclose_subseteq'. } *) +(* iMod ("Hvs" with "[$]") as "[Hcont HQ]". *) +(* iMod "Hclose'". *) +(* iMod ("Hclose" with "[ $Hloc $Hcont]"); first done. *) +(* iApply "HΦ". by iFrame. *) +(* Qed. *) -End impl2. +(* End impl2. *) -Program Definition random_counter2 `{flip_spec Σ}: random_counter := - {| new_counter := new_counter2; - allocate_tape:= allocate_tape2; - incr_counter_tape := incr_counter_tape2; - read_counter:=read_counter2; - counterG := counterG2; - error_name := flip_error_name; - tape_name := flip_tape_name; - counter_name :=gname; - is_counter _ N c γ1 γ2 γ3 := (is_flip (N.@"flip") γ1 γ2 ∗ - inv (N.@ "counter") (counter_inv_pred2 c γ3))%I; - counter_error_auth _ γ x := flip_error_auth γ x; - counter_error_frag _ γ x := flip_error_frag γ x; - counter_tapes_auth _ γ m := flip_tapes_auth (L:=_) γ ((λ ns, expander ns)<$>m); - counter_tapes_frag _ γ α ns := (flip_tapes_frag (L:=_) γ α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝)%I; - counter_content_auth _ γ z := own γ (●F z); - counter_content_frag _ γ f z := own γ (◯F{f} z); - new_counter_spec _ := new_counter_spec2; - allocate_tape_spec _ :=allocate_tape_spec2; - incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some2; - counter_presample_spec _ :=counter_presample_spec2; - read_counter_spec _ :=read_counter_spec2 - |}. -Next Obligation. - intros ?? H ???. - apply counterG2_flipG. -Qed. -Next Obligation. - simpl. - iIntros (???????) "H1 H2". - iApply (flip_tapes_auth_exclusive with "[$][$]"). -Qed. -Next Obligation. - simpl. - iIntros (???????) "H1 H2". - iCombine "H1 H2" gives "%K". by rewrite auth_auth_op_valid in K. -Qed. -Next Obligation. - simpl. iIntros (????? z z' ?) "H1 H2". - iCombine "H1 H2" gives "%K". - apply frac_auth_included_total in K. iPureIntro. - by apply nat_included. -Qed. -Next Obligation. - simpl. - iIntros (?????????). - rewrite frac_auth_frag_op. by rewrite own_op. -Qed. -Next Obligation. - simpl. iIntros (???????) "H1 H2". - iCombine "H1 H2" gives "%K". - iPureIntro. - by apply frac_auth_agree_L in K. -Qed. -Next Obligation. - simpl. iIntros (?????????) "H1 H2". - iMod (own_update_2 _ _ _ (_ ⋅ _) with "[$][$]") as "[$$]"; last done. - apply frac_auth_update. - apply nat_local_update. lia. -Qed. -Next Obligation. - intros ?? H ?. - apply counterG2_flipG. -Qed. -Next Obligation. - simpl. - iIntros (????????) "[? _] [? _]". - by iDestruct (flip_tapes_frag_exclusive with "[$][$]") as "%". -Qed. -Next Obligation. - simpl. - iIntros. - iApply (flip_error_update with "[$][$]"). -Qed. -Next Obligation. - iIntros. - iApply (flip_error_agree with "[$][$]"). -Qed. -Next Obligation. - iIntros. - iApply (flip_error_frag_valid with "[$]"). -Admitted. -Next Obligation. -Admitted. -Next Obligation. -Admitted. -Next Obligation. -Admitted. -Next Obligation. - simpl. -Admitted. -Next Obligation. -Admitted. +(* Program Definition random_counter2 `{flip_spec Σ}: random_counter := *) +(* {| new_counter := new_counter2; *) +(* allocate_tape:= allocate_tape2; *) +(* incr_counter_tape := incr_counter_tape2; *) +(* read_counter:=read_counter2; *) +(* counterG := counterG2; *) +(* error_name := flip_error_name; *) +(* tape_name := flip_tape_name; *) +(* counter_name :=gname; *) +(* is_counter _ N c γ1 γ2 γ3 := (is_flip (N.@"flip") γ1 γ2 ∗ *) +(* inv (N.@ "counter") (counter_inv_pred2 c γ3))%I; *) +(* counter_error_auth _ γ x := flip_error_auth γ x; *) +(* counter_error_frag _ γ x := flip_error_frag γ x; *) +(* counter_tapes_auth _ γ m := flip_tapes_auth (L:=_) γ ((λ ns, expander ns)<$>m); *) +(* counter_tapes_frag _ γ α ns := (flip_tapes_frag (L:=_) γ α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝)%I; *) +(* counter_content_auth _ γ z := own γ (●F z); *) +(* counter_content_frag _ γ f z := own γ (◯F{f} z); *) +(* new_counter_spec _ := new_counter_spec2; *) +(* allocate_tape_spec _ :=allocate_tape_spec2; *) +(* incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some2; *) +(* counter_presample_spec _ :=counter_presample_spec2; *) +(* read_counter_spec _ :=read_counter_spec2 *) +(* |}. *) +(* Next Obligation. *) +(* intros ?? H ???. *) +(* apply counterG2_flipG. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (???????) "H1 H2". *) +(* iApply (flip_tapes_auth_exclusive with "[$][$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (???????) "H1 H2". *) +(* iCombine "H1 H2" gives "%K". by rewrite auth_auth_op_valid in K. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. iIntros (????? z z' ?) "H1 H2". *) +(* iCombine "H1 H2" gives "%K". *) +(* apply frac_auth_included_total in K. iPureIntro. *) +(* by apply nat_included. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (?????????). *) +(* rewrite frac_auth_frag_op. by rewrite own_op. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. iIntros (???????) "H1 H2". *) +(* iCombine "H1 H2" gives "%K". *) +(* iPureIntro. *) +(* by apply frac_auth_agree_L in K. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. iIntros (?????????) "H1 H2". *) +(* iMod (own_update_2 _ _ _ (_ ⋅ _) with "[$][$]") as "[$$]"; last done. *) +(* apply frac_auth_update. *) +(* apply nat_local_update. lia. *) +(* Qed. *) +(* Next Obligation. *) +(* intros ?? H ?. *) +(* apply counterG2_flipG. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (????????) "[? _] [? _]". *) +(* by iDestruct (flip_tapes_frag_exclusive with "[$][$]") as "%". *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros. *) +(* iApply (flip_error_update with "[$][$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* iIntros. *) +(* iApply (flip_error_agree with "[$][$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* iIntros. *) +(* iApply (flip_error_frag_valid with "[$]"). *) +(* Admitted. *) +(* Next Obligation. *) +(* Admitted. *) +(* Next Obligation. *) +(* Admitted. *) +(* Next Obligation. *) +(* Admitted. *) +(* Next Obligation. *) +(* simpl. *) +(* Admitted. *) +(* Next Obligation. *) +(* Admitted. *) diff --git a/theories/coneris/examples/random_counter/impl3.v b/theories/coneris/examples/random_counter/impl3.v index 0abb395d..94ca521b 100644 --- a/theories/coneris/examples/random_counter/impl3.v +++ b/theories/coneris/examples/random_counter/impl3.v @@ -1,404 +1,404 @@ -From iris.algebra Require Import frac_auth. -From iris.base_logic.lib Require Import invariants. -From clutch.coneris Require Import coneris hocap random_counter. +(* From iris.algebra Require Import frac_auth. *) +(* From iris.base_logic.lib Require Import invariants. *) +(* From clutch.coneris Require Import coneris hocap random_counter. *) -Set Default Proof Using "Type*". +(* Set Default Proof Using "Type*". *) -Section filter. - Definition filter_f (n:nat):= bool_decide (n<4)%nat. - Definition filtered_list (l:list _) := filter filter_f l. -End filter. +(* Section filter. *) +(* Definition filter_f (n:nat):= bool_decide (n<4)%nat. *) +(* Definition filtered_list (l:list _) := filter filter_f l. *) +(* End filter. *) -Section lemmas. - Context `{!conerisGS Σ}. - Lemma hocap_tapes_notin3 α N ns (m:gmap loc (nat*list nat)) : - α ↪N (N; ns) -∗ ([∗ map] α↦t ∈ m,∃ (ls:list nat), ⌜ (filter filter_f ls) = t.2⌝ ∗ α ↪N ( 4%nat ; ls)) -∗ ⌜m!!α=None ⌝. - Proof. - destruct (m!!α) eqn:Heqn; last by iIntros. - iIntros "Hα Hmap". - iDestruct (big_sepM_lookup with "[$]") as "(%&%&?)"; first done. - iExFalso. - iApply (tapeN_tapeN_contradict with "[$][$]"). - Qed. -End lemmas. +(* Section lemmas. *) +(* Context `{!conerisGS Σ}. *) +(* Lemma hocap_tapes_notin3 α N ns (m:gmap loc (nat*list nat)) : *) +(* α ↪N (N; ns) -∗ ([∗ map] α↦t ∈ m,∃ (ls:list nat), ⌜ (filter filter_f ls) = t.2⌝ ∗ α ↪N ( 4%nat ; ls)) -∗ ⌜m!!α=None ⌝. *) +(* Proof. *) +(* destruct (m!!α) eqn:Heqn; last by iIntros. *) +(* iIntros "Hα Hmap". *) +(* iDestruct (big_sepM_lookup with "[$]") as "(%&%&?)"; first done. *) +(* iExFalso. *) +(* iApply (tapeN_tapeN_contradict with "[$][$]"). *) +(* Qed. *) +(* End lemmas. *) -Section impl3. +(* Section impl3. *) - Definition new_counter3 : val:= λ: "_", ref #0. - Definition allocate_tape3 : val := λ: "_", AllocTape #4. - Definition incr_counter_tape3 :val := rec: "f" "l" "α":= - let: "n" := rand("α") #4 in - if: "n" < #4 - then (FAA "l" "n", "n") - else "f" "l" "α". - Definition read_counter3 : val := λ: "l", !"l". - Class counterG3 Σ := CounterG3 { counterG2_error::hocap_errorGS Σ; - counterG2_tapes:: hocap_tapesGS Σ; - counterG2_frac_authR:: inG Σ (frac_authR natR) }. +(* Definition new_counter3 : val:= λ: "_", ref #0. *) +(* Definition allocate_tape3 : val := λ: "_", AllocTape #4. *) +(* Definition incr_counter_tape3 :val := rec: "f" "l" "α":= *) +(* let: "n" := rand("α") #4 in *) +(* if: "n" < #4 *) +(* then (FAA "l" "n", "n") *) +(* else "f" "l" "α". *) +(* Definition read_counter3 : val := λ: "l", !"l". *) +(* Class counterG3 Σ := CounterG3 { counterG2_error::hocap_errorGS Σ; *) +(* counterG2_tapes:: hocap_tapesGS Σ; *) +(* counterG2_frac_authR:: inG Σ (frac_authR natR) }. *) - Context `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapesGS Σ, !inG Σ (frac_authR natR)}. - Definition counter_inv_pred3 (c:val) γ1 γ2 γ3:= - (∃ (ε:R) (m:gmap loc (nat * list nat)) (l:loc) (z:nat), - ↯ ε ∗ ●↯ ε @ γ1 ∗ - ([∗ map] α ↦ t ∈ m, ∃ (ls:list nat), ⌜ (filter filter_f ls) = t.2⌝ ∗ α ↪N ( 4%nat ; ls) ) - ∗ ●m@γ2 ∗ - ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ3 (●F z) - )%I. +(* Context `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapesGS Σ, !inG Σ (frac_authR natR)}. *) +(* Definition counter_inv_pred3 (c:val) γ1 γ2 γ3:= *) +(* (∃ (ε:R) (m:gmap loc (nat * list nat)) (l:loc) (z:nat), *) +(* ↯ ε ∗ ●↯ ε @ γ1 ∗ *) +(* ([∗ map] α ↦ t ∈ m, ∃ (ls:list nat), ⌜ (filter filter_f ls) = t.2⌝ ∗ α ↪N ( 4%nat ; ls) ) *) +(* ∗ ●m@γ2 ∗ *) +(* ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ3 (●F z) *) +(* )%I. *) - Lemma new_counter_spec3 E ε N: - {{{ ↯ ε }}} - new_counter3 #() @ E - {{{ (c:val), RET c; - ∃ γ1 γ2 γ3, inv N (counter_inv_pred3 c γ1 γ2 γ3) ∗ - ◯↯ε @ γ1 ∗ own γ3 (◯F 0%nat) - }}}. - Proof. - rewrite /new_counter3. - iIntros (Φ) "Hε HΦ". - wp_pures. - wp_alloc l as "Hl". - iDestruct (ec_valid with "[$]") as "%". - unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "[%γ1 [H1 H2]]". - { lra. } - simpl. - iMod (hocap_tapes_alloc (∅:gmap _ _)) as "[%γ2 [H3 H4]]". - iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". - { by apply frac_auth_valid. } - replace (#0) with (#0%nat) by done. - iMod (inv_alloc N _ (counter_inv_pred3 (#l) γ1 γ2 γ3) with "[$Hε $Hl $H1 $H3 $H5]") as "#Hinv". - { iSplit; last done. by iApply big_sepM_empty. } - iApply "HΦ". - iExists _, _, _. by iFrame. - Qed. +(* Lemma new_counter_spec3 E ε N: *) +(* {{{ ↯ ε }}} *) +(* new_counter3 #() @ E *) +(* {{{ (c:val), RET c; *) +(* ∃ γ1 γ2 γ3, inv N (counter_inv_pred3 c γ1 γ2 γ3) ∗ *) +(* ◯↯ε @ γ1 ∗ own γ3 (◯F 0%nat) *) +(* }}}. *) +(* Proof. *) +(* rewrite /new_counter3. *) +(* iIntros (Φ) "Hε HΦ". *) +(* wp_pures. *) +(* wp_alloc l as "Hl". *) +(* iDestruct (ec_valid with "[$]") as "%". *) +(* unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "[%γ1 [H1 H2]]". *) +(* { lra. } *) +(* simpl. *) +(* iMod (hocap_tapes_alloc (∅:gmap _ _)) as "[%γ2 [H3 H4]]". *) +(* iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". *) +(* { by apply frac_auth_valid. } *) +(* replace (#0) with (#0%nat) by done. *) +(* iMod (inv_alloc N _ (counter_inv_pred3 (#l) γ1 γ2 γ3) with "[$Hε $Hl $H1 $H3 $H5]") as "#Hinv". *) +(* { iSplit; last done. by iApply big_sepM_empty. } *) +(* iApply "HΦ". *) +(* iExists _, _, _. by iFrame. *) +(* Qed. *) - Lemma allocate_tape_spec3 N E c γ1 γ2 γ3: - ↑N ⊆ E-> - {{{ inv N (counter_inv_pred3 c γ1 γ2 γ3) }}} - allocate_tape3 #() @ E - {{{ (v:val), RET v; - ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ α ◯↪N (3%nat; []) @ γ2 - }}}. - Proof. - iIntros (Hsubset Φ) "#Hinv HΦ". - rewrite /allocate_tape3. - wp_pures. - wp_alloctape α as "Hα". - iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". - iDestruct (hocap_tapes_notin3 with "[$][$]") as "%". - iMod (hocap_tapes_new with "[$]") as "[H4 H7]"; first done. - iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Hα]") as "_". - { iNext. iSplitL; last done. - rewrite big_sepM_insert; [simpl; iFrame|done]. - by rewrite filter_nil. - } - iApply "HΦ". - by iFrame. - Qed. +(* Lemma allocate_tape_spec3 N E c γ1 γ2 γ3: *) +(* ↑N ⊆ E-> *) +(* {{{ inv N (counter_inv_pred3 c γ1 γ2 γ3) }}} *) +(* allocate_tape3 #() @ E *) +(* {{{ (v:val), RET v; *) +(* ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ α ◯↪N (3%nat; []) @ γ2 *) +(* }}}. *) +(* Proof. *) +(* iIntros (Hsubset Φ) "#Hinv HΦ". *) +(* rewrite /allocate_tape3. *) +(* wp_pures. *) +(* wp_alloctape α as "Hα". *) +(* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) +(* iDestruct (hocap_tapes_notin3 with "[$][$]") as "%". *) +(* iMod (hocap_tapes_new with "[$]") as "[H4 H7]"; first done. *) +(* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Hα]") as "_". *) +(* { iNext. iSplitL; last done. *) +(* rewrite big_sepM_insert; [simpl; iFrame|done]. *) +(* by rewrite filter_nil. *) +(* } *) +(* iApply "HΦ". *) +(* by iFrame. *) +(* Qed. *) - Lemma incr_counter_tape_spec_some3 N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: - ↑N⊆E -> - {{{ inv N (counter_inv_pred3 c γ1 γ2 γ3) ∗ - □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ - own γ3 (●F(z+n)%nat)∗ Q z) ∗ - P ∗ α ◯↪N (3%nat; n::ns) @ γ2 - }}} - incr_counter_tape3 c #lbl:α @ E - {{{ (z:nat), RET (#z, #n); Q z ∗ α ◯↪N (3%nat; ns) @ γ2}}}. - Proof. - iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP & Hα) HΦ". - rewrite /incr_counter_tape3. - iLöb as "IH". - wp_pures. - wp_bind (rand(_) _)%E. - iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". - iDestruct (hocap_tapes_agree with "[$][$]") as "%". - erewrite <-(insert_delete m) at 1; last done. - rewrite big_sepM_insert; last apply lookup_delete. - simpl. - iDestruct "H3" as "[(%ls & %Hls & Htape) H3]". - destruct ls as [|x ls]. - { rewrite filter_nil in Hls. simplify_eq. } - wp_apply (wp_rand_tape with "[$]") as "[Htape %Hineq]". - rewrite Nat.le_lteq in Hineq. - destruct Hineq as [? | ->]. - - (* first value is valid *) - iMod (hocap_tapes_pop with "[$][$]") as "[H4 Hα]". - rewrite filter_cons /filter_f in Hls. - rewrite bool_decide_eq_true_2 in Hls; last done. simpl in *. - simplify_eq. - iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". - { iSplitL; last done. - erewrite <-(insert_delete m) at 2; last done. - iNext. - rewrite insert_insert. - rewrite big_sepM_insert; last apply lookup_delete. by iFrame. - } - iModIntro. - wp_pures. - rewrite bool_decide_eq_true_2; last lia. - clear -Hsubset. - wp_pures. - wp_bind (FAA _ _). - iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". - wp_faa. - iMod ("Hvs" with "[$]") as "[H6 HQ]". - replace (#(z+n)) with (#(z+n)%nat); last first. - { by rewrite Nat2Z.inj_add. } - iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. - iModIntro. wp_pures. - iApply "HΦ". - by iFrame. - - (* we get a 5, do iLöb induction *) - rewrite filter_cons /filter_f in Hls. - rewrite bool_decide_eq_false_2 in Hls; last lia. simpl in *. - iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". - { iSplitL; last done. - erewrite <-(insert_delete m) at 2; last done. - iNext. - rewrite big_sepM_insert; last apply lookup_delete. by iFrame. - } - iModIntro. - do 4 wp_pure. - by iApply ("IH" with "[$][$]"). - Qed. +(* Lemma incr_counter_tape_spec_some3 N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: *) +(* ↑N⊆E -> *) +(* {{{ inv N (counter_inv_pred3 c γ1 γ2 γ3) ∗ *) +(* □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ *) +(* own γ3 (●F(z+n)%nat)∗ Q z) ∗ *) +(* P ∗ α ◯↪N (3%nat; n::ns) @ γ2 *) +(* }}} *) +(* incr_counter_tape3 c #lbl:α @ E *) +(* {{{ (z:nat), RET (#z, #n); Q z ∗ α ◯↪N (3%nat; ns) @ γ2}}}. *) +(* Proof. *) +(* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP & Hα) HΦ". *) +(* rewrite /incr_counter_tape3. *) +(* iLöb as "IH". *) +(* wp_pures. *) +(* wp_bind (rand(_) _)%E. *) +(* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) +(* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) +(* erewrite <-(insert_delete m) at 1; last done. *) +(* rewrite big_sepM_insert; last apply lookup_delete. *) +(* simpl. *) +(* iDestruct "H3" as "[(%ls & %Hls & Htape) H3]". *) +(* destruct ls as [|x ls]. *) +(* { rewrite filter_nil in Hls. simplify_eq. } *) +(* wp_apply (wp_rand_tape with "[$]") as "[Htape %Hineq]". *) +(* rewrite Nat.le_lteq in Hineq. *) +(* destruct Hineq as [? | ->]. *) +(* - (* first value is valid *) *) +(* iMod (hocap_tapes_pop with "[$][$]") as "[H4 Hα]". *) +(* rewrite filter_cons /filter_f in Hls. *) +(* rewrite bool_decide_eq_true_2 in Hls; last done. simpl in *. *) +(* simplify_eq. *) +(* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". *) +(* { iSplitL; last done. *) +(* erewrite <-(insert_delete m) at 2; last done. *) +(* iNext. *) +(* rewrite insert_insert. *) +(* rewrite big_sepM_insert; last apply lookup_delete. by iFrame. *) +(* } *) +(* iModIntro. *) +(* wp_pures. *) +(* rewrite bool_decide_eq_true_2; last lia. *) +(* clear -Hsubset. *) +(* wp_pures. *) +(* wp_bind (FAA _ _). *) +(* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) +(* wp_faa. *) +(* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) +(* replace (#(z+n)) with (#(z+n)%nat); last first. *) +(* { by rewrite Nat2Z.inj_add. } *) +(* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) +(* iModIntro. wp_pures. *) +(* iApply "HΦ". *) +(* by iFrame. *) +(* - (* we get a 5, do iLöb induction *) *) +(* rewrite filter_cons /filter_f in Hls. *) +(* rewrite bool_decide_eq_false_2 in Hls; last lia. simpl in *. *) +(* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". *) +(* { iSplitL; last done. *) +(* erewrite <-(insert_delete m) at 2; last done. *) +(* iNext. *) +(* rewrite big_sepM_insert; last apply lookup_delete. by iFrame. *) +(* } *) +(* iModIntro. *) +(* do 4 wp_pure. *) +(* by iApply ("IH" with "[$][$]"). *) +(* Qed. *) - Lemma counter_presample_spec3 NS E ns α - (ε2 : R -> nat -> R) - (P : iProp Σ) T γ1 γ2 γ3 c: - ↑NS ⊆ E -> - (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> - (∀ (ε:R), 0<= ε ->SeriesC (λ n, if (bool_decide (n≤3%nat)) then 1 / (S 3%nat) * ε2 ε n else 0%R)%R <= ε)%R-> - inv NS (counter_inv_pred3 c γ1 γ2 γ3) -∗ - (□∀ (ε:R) n, (P ∗ ●↯ ε@ γ1) ={E∖↑NS}=∗ - (⌜(1<=ε2 ε n)%R⌝ ∨(●↯ (ε2 ε n) @ γ1 ∗ T (n)))) - -∗ - P -∗ α ◯↪N (3%nat; ns) @ γ2 -∗ - wp_update E (∃ n, T (n) ∗ α◯↪N (3%nat; ns++[n]) @ γ2). - Proof. - iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". - iMod wp_update_epsilon_err as "(%epsilon_err&%H&?)". - iApply wp_update_state_step_coupl. - iRevert "HP Hfrag". - iApply (ec_ind_amp _ 5%R with "[][$]"); [done|lra|]. - iModIntro. - clear epsilon_err H. - iIntros (epsilon_err ?) "#IH Hepsilon_err HP Hfrag". - iIntros (σ ε) "((Hheap&Htapes)&Hε)". - iMod (inv_acc with "Hinv") as "[>(%ε0 & % & % & % & H1 & H2 & H3 & H4 & -> & H5 & H6) Hclose]"; [done|]. - iDestruct (hocap_tapes_agree with "[$][$]") as "%". - erewrite <-(insert_delete m) at 1; last done. - rewrite big_sepM_insert; last apply lookup_delete. - simpl. - iDestruct "H3" as "[(%ls & %Hls & Htape) H3]". - iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%)". - iDestruct (ec_valid with "H1") as "%". - iDestruct (ec_combine with "[$]") as "Htotal". - iDestruct (ec_supply_bound with "[$][$]") as "%". - iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. - iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". - pose (ε2' := λ (x:nat), (if bool_decide (x<=3)%nat then ε2 ε0 x else - if bool_decide (x=4)%nat then ε0 + 5 * epsilon_err - else 0)%R - ). - iApply state_step_coupl_state_adv_comp_con_prob_lang; first done. - assert (forall x, 0<=ε2' x)%R. - { intros. rewrite /ε2'. repeat case_bool_decide; try done; first naive_solver. - apply Rplus_le_le_0_compat; simpl; [naive_solver|lra]. } - unshelve iExists (λ x, mknonnegreal (ε2' x) _); auto. - iSplit. - { iPureIntro. - unshelve epose proof (Hineq ε0 _) as H'; first (by naive_solver). - rewrite SeriesC_nat_bounded_fin in H'. - replace (INR 5%nat) with 5%R by (simpl; lra). - replace (INR 4%nat) with 4%R in H' by (simpl; lra). - rewrite /=/ε2' Hε1'. - rewrite SeriesC_finite_foldr/=. - rewrite SeriesC_finite_foldr/= in H'. - lra. - } - iIntros (sample). - simpl. - destruct (Rlt_decision (nonneg ε_rem + (ε2' sample))%R 1%R) as [Hdec|Hdec]; last first. - { apply Rnot_lt_ge, Rge_le in Hdec. - iApply state_step_coupl_ret_err_ge_1. - simpl. simpl in *. lra. - } - destruct (Nat.lt_total (fin_to_nat sample) 4%nat) as [|[|]]. - - (* we sample a value less than 4*) - iApply state_step_coupl_ret. - unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2' sample) _) with "Hε_supply") as "[Hε_supply Hε]"; auto. - iMod (tapeN_update_append _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". - iMod (hocap_tapes_presample _ _ _ _ _ (fin_to_nat sample) with "[$][$]") as "[H4 Hfrag]". - iMod "Hclose'" as "_". - iMod ("Hvs" $! _ (fin_to_nat sample)%nat with "[$]") as "[%|[H2 HT]]". - { iExFalso. iApply (ec_contradict with "[$]"). - simpl. - rewrite /ε2'. - by rewrite bool_decide_eq_true_2; last lia. - } - simpl. - iMod ("Hclose" with "[$Hε H2 Htape H3 $H4 $H5 $H6]") as "_". - { iNext. - rewrite /ε2'. rewrite bool_decide_eq_true_2; last lia. - iFrame. - rewrite big_sepM_insert_delete; iFrame. simpl. iPureIntro. - split; last done. - rewrite filter_app. f_equal. - rewrite filter_cons. - rewrite decide_True; first done. - rewrite /filter_f. by apply bool_decide_pack. - } - iApply fupd_mask_intro_subseteq; first set_solver. - iFrame. - - (* we sampled a 4, and hence löb *) - unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2' sample) _) with "Hε_supply") as "[Hε_supply Hε]"; auto. - iMod (tapeN_update_append _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". - iMod "Hclose'" as "_". - simpl. - rewrite {2}/ε2'. - rewrite bool_decide_eq_false_2; last lia. - rewrite bool_decide_eq_true_2; last lia. - iDestruct (ec_split with "[$]") as "[Hε Hampl]"; simpl; [naive_solver|lra|]. - simpl. - iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4 $H5 $H6]") as "_". - { iNext. - rewrite (big_sepM_delete _ m); [iFrame|done]. - simpl. iPureIntro; split; last done. - rewrite filter_app filter_cons decide_False; simpl. - - by rewrite filter_nil app_nil_r. - - rewrite /filter_f. case_bool_decide; [lia|auto]. - } - iDestruct ("IH" with "[$][$][$]") as "IH'". - by iMod ("IH'" $! (state_upd_tapes <[_:=_]> _) with "[$]") as "IH'". - - pose proof fin_to_nat_lt sample; lia. - Qed. +(* Lemma counter_presample_spec3 NS E ns α *) +(* (ε2 : R -> nat -> R) *) +(* (P : iProp Σ) T γ1 γ2 γ3 c: *) +(* ↑NS ⊆ E -> *) +(* (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> *) +(* (∀ (ε:R), 0<= ε ->SeriesC (λ n, if (bool_decide (n≤3%nat)) then 1 / (S 3%nat) * ε2 ε n else 0%R)%R <= ε)%R-> *) +(* inv NS (counter_inv_pred3 c γ1 γ2 γ3) -∗ *) +(* (□∀ (ε:R) n, (P ∗ ●↯ ε@ γ1) ={E∖↑NS}=∗ *) +(* (⌜(1<=ε2 ε n)%R⌝ ∨(●↯ (ε2 ε n) @ γ1 ∗ T (n)))) *) +(* -∗ *) +(* P -∗ α ◯↪N (3%nat; ns) @ γ2 -∗ *) +(* wp_update E (∃ n, T (n) ∗ α◯↪N (3%nat; ns++[n]) @ γ2). *) +(* Proof. *) +(* iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) +(* iMod wp_update_epsilon_err as "(%epsilon_err&%H&?)". *) +(* iApply wp_update_state_step_coupl. *) +(* iRevert "HP Hfrag". *) +(* iApply (ec_ind_amp _ 5%R with "[][$]"); [done|lra|]. *) +(* iModIntro. *) +(* clear epsilon_err H. *) +(* iIntros (epsilon_err ?) "#IH Hepsilon_err HP Hfrag". *) +(* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) +(* iMod (inv_acc with "Hinv") as "[>(%ε0 & % & % & % & H1 & H2 & H3 & H4 & -> & H5 & H6) Hclose]"; [done|]. *) +(* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) +(* erewrite <-(insert_delete m) at 1; last done. *) +(* rewrite big_sepM_insert; last apply lookup_delete. *) +(* simpl. *) +(* iDestruct "H3" as "[(%ls & %Hls & Htape) H3]". *) +(* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%)". *) +(* iDestruct (ec_valid with "H1") as "%". *) +(* iDestruct (ec_combine with "[$]") as "Htotal". *) +(* iDestruct (ec_supply_bound with "[$][$]") as "%". *) +(* iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. *) +(* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) +(* pose (ε2' := λ (x:nat), (if bool_decide (x<=3)%nat then ε2 ε0 x else *) +(* if bool_decide (x=4)%nat then ε0 + 5 * epsilon_err *) +(* else 0)%R *) +(* ). *) +(* iApply state_step_coupl_state_adv_comp_con_prob_lang; first done. *) +(* assert (forall x, 0<=ε2' x)%R. *) +(* { intros. rewrite /ε2'. repeat case_bool_decide; try done; first naive_solver. *) +(* apply Rplus_le_le_0_compat; simpl; [naive_solver|lra]. } *) +(* unshelve iExists (λ x, mknonnegreal (ε2' x) _); auto. *) +(* iSplit. *) +(* { iPureIntro. *) +(* unshelve epose proof (Hineq ε0 _) as H'; first (by naive_solver). *) +(* rewrite SeriesC_nat_bounded_fin in H'. *) +(* replace (INR 5%nat) with 5%R by (simpl; lra). *) +(* replace (INR 4%nat) with 4%R in H' by (simpl; lra). *) +(* rewrite /=/ε2' Hε1'. *) +(* rewrite SeriesC_finite_foldr/=. *) +(* rewrite SeriesC_finite_foldr/= in H'. *) +(* lra. *) +(* } *) +(* iIntros (sample). *) +(* simpl. *) +(* destruct (Rlt_decision (nonneg ε_rem + (ε2' sample))%R 1%R) as [Hdec|Hdec]; last first. *) +(* { apply Rnot_lt_ge, Rge_le in Hdec. *) +(* iApply state_step_coupl_ret_err_ge_1. *) +(* simpl. simpl in *. lra. *) +(* } *) +(* destruct (Nat.lt_total (fin_to_nat sample) 4%nat) as [|[|]]. *) +(* - (* we sample a value less than 4*) *) +(* iApply state_step_coupl_ret. *) +(* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2' sample) _) with "Hε_supply") as "[Hε_supply Hε]"; auto. *) +(* iMod (tapeN_update_append _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". *) +(* iMod (hocap_tapes_presample _ _ _ _ _ (fin_to_nat sample) with "[$][$]") as "[H4 Hfrag]". *) +(* iMod "Hclose'" as "_". *) +(* iMod ("Hvs" $! _ (fin_to_nat sample)%nat with "[$]") as "[%|[H2 HT]]". *) +(* { iExFalso. iApply (ec_contradict with "[$]"). *) +(* simpl. *) +(* rewrite /ε2'. *) +(* by rewrite bool_decide_eq_true_2; last lia. *) +(* } *) +(* simpl. *) +(* iMod ("Hclose" with "[$Hε H2 Htape H3 $H4 $H5 $H6]") as "_". *) +(* { iNext. *) +(* rewrite /ε2'. rewrite bool_decide_eq_true_2; last lia. *) +(* iFrame. *) +(* rewrite big_sepM_insert_delete; iFrame. simpl. iPureIntro. *) +(* split; last done. *) +(* rewrite filter_app. f_equal. *) +(* rewrite filter_cons. *) +(* rewrite decide_True; first done. *) +(* rewrite /filter_f. by apply bool_decide_pack. *) +(* } *) +(* iApply fupd_mask_intro_subseteq; first set_solver. *) +(* iFrame. *) +(* - (* we sampled a 4, and hence löb *) *) +(* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2' sample) _) with "Hε_supply") as "[Hε_supply Hε]"; auto. *) +(* iMod (tapeN_update_append _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". *) +(* iMod "Hclose'" as "_". *) +(* simpl. *) +(* rewrite {2}/ε2'. *) +(* rewrite bool_decide_eq_false_2; last lia. *) +(* rewrite bool_decide_eq_true_2; last lia. *) +(* iDestruct (ec_split with "[$]") as "[Hε Hampl]"; simpl; [naive_solver|lra|]. *) +(* simpl. *) +(* iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4 $H5 $H6]") as "_". *) +(* { iNext. *) +(* rewrite (big_sepM_delete _ m); [iFrame|done]. *) +(* simpl. iPureIntro; split; last done. *) +(* rewrite filter_app filter_cons decide_False; simpl. *) +(* - by rewrite filter_nil app_nil_r. *) +(* - rewrite /filter_f. case_bool_decide; [lia|auto]. *) +(* } *) +(* iDestruct ("IH" with "[$][$][$]") as "IH'". *) +(* by iMod ("IH'" $! (state_upd_tapes <[_:=_]> _) with "[$]") as "IH'". *) +(* - pose proof fin_to_nat_lt sample; lia. *) +(* Qed. *) - Lemma read_counter_spec3 N E c γ1 γ2 γ3 P Q: - ↑N ⊆ E -> - {{{ inv N (counter_inv_pred3 c γ1 γ2 γ3) ∗ - □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ - own γ3 (●F z)∗ Q z) - ∗ P - }}} - read_counter3 c @ E - {{{ (n':nat), RET #n'; Q n' - }}}. - Proof. - iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP) HΦ". - rewrite /read_counter3. - wp_pure. - iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". - wp_load. - iMod ("Hvs" with "[$]") as "[H6 HQ]". - iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]"); first done. - iApply ("HΦ" with "[$]"). - Qed. +(* Lemma read_counter_spec3 N E c γ1 γ2 γ3 P Q: *) +(* ↑N ⊆ E -> *) +(* {{{ inv N (counter_inv_pred3 c γ1 γ2 γ3) ∗ *) +(* □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ *) +(* own γ3 (●F z)∗ Q z) *) +(* ∗ P *) +(* }}} *) +(* read_counter3 c @ E *) +(* {{{ (n':nat), RET #n'; Q n' *) +(* }}}. *) +(* Proof. *) +(* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP) HΦ". *) +(* rewrite /read_counter3. *) +(* wp_pure. *) +(* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) +(* wp_load. *) +(* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) +(* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]"); first done. *) +(* iApply ("HΦ" with "[$]"). *) +(* Qed. *) -End impl3. +(* End impl3. *) -Program Definition random_counter3 `{!conerisGS Σ}: random_counter := - {| new_counter := new_counter3; - allocate_tape:= allocate_tape3; - incr_counter_tape := incr_counter_tape3; - read_counter:=read_counter3; - counterG := counterG3; - error_name := gname; - tape_name := gname; - counter_name :=gname; - is_counter _ N c γ1 γ2 γ3 := inv N (counter_inv_pred3 c γ1 γ2 γ3); - counter_error_auth _ γ x := ●↯ x @ γ; - counter_error_frag _ γ x := ◯↯ x @ γ; - counter_tapes_auth _ γ m := (●((λ ns, (3, ns))<$>m)@γ)%I; - counter_tapes_frag _ γ α ns := (α◯↪N (3%nat;ns) @ γ)%I; - counter_content_auth _ γ z := own γ (●F z); - counter_content_frag _ γ f z := own γ (◯F{f} z); - new_counter_spec _ := new_counter_spec3; - allocate_tape_spec _ :=allocate_tape_spec3; - incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some3; - counter_presample_spec _ :=counter_presample_spec3; - read_counter_spec _ :=read_counter_spec3 - |}. -Next Obligation. - simpl. - iIntros (??????) "(%&<-&H1)(%&<-&H2)". - iCombine "H1 H2" gives "%H". by rewrite excl_auth.excl_auth_auth_op_valid in H. -Qed. -Next Obligation. - simpl. - iIntros (??????) "(%&<-&H1)(%&<-&H2)". - iCombine "H1 H2" gives "%H". by rewrite excl_auth.excl_auth_frag_op_valid in H. -Qed. -Next Obligation. - simpl. - iIntros (?????) "H". - iApply (hocap_error_auth_pos with "[$]"). -Qed. -Next Obligation. - simpl. - iIntros (?????) "H". - iApply (hocap_error_frag_pos with "[$]"). -Qed. -Next Obligation. - simpl. - iIntros (??????) "H1 H2". - iApply (hocap_error_agree with "[$][$]"). -Qed. -Next Obligation. - simpl. iIntros (???????) "??". - iApply (hocap_error_update with "[$][$]"). -Qed. -Next Obligation. - simpl. - iIntros (??????) "H1 H2". - by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[%H _]". -Qed. -Next Obligation. - simpl. - iIntros (???????) "H1 H2". - iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. - rewrite dfrac_op_own dfrac_valid_own. by intro. -Qed. -Next Obligation. - simpl. - iIntros. - iDestruct (hocap_tapes_agree with "[$][$]") as "%H". - iPureIntro. - rewrite lookup_fmap_Some in H. destruct H as (?&?&?). - by simplify_eq. -Qed. -Next Obligation. - iIntros. - iMod (hocap_tapes_presample with "[$][$]") as "[??]". - iModIntro. iFrame. - by rewrite fmap_insert. -Qed. -Next Obligation. - simpl. - iIntros (??????) "H1 H2". - iCombine "H1 H2" gives "%H". by rewrite auth_auth_op_valid in H. -Qed. -Next Obligation. - simpl. iIntros (???? z z' ?) "H1 H2". - iCombine "H1 H2" gives "%H". - apply frac_auth_included_total in H. iPureIntro. - by apply nat_included. -Qed. -Next Obligation. - simpl. - iIntros (????????). - rewrite frac_auth_frag_op. by rewrite own_op. -Qed. -Next Obligation. - simpl. iIntros (??????) "H1 H2". - iCombine "H1 H2" gives "%H". - iPureIntro. - by apply frac_auth_agree_L in H. -Qed. -Next Obligation. - simpl. iIntros (????????) "H1 H2". - iMod (own_update_2 _ _ _ (_ ⋅ _) with "[$][$]") as "[$$]"; last done. - apply frac_auth_update. - apply nat_local_update. lia. -Qed. +(* Program Definition random_counter3 `{!conerisGS Σ}: random_counter := *) +(* {| new_counter := new_counter3; *) +(* allocate_tape:= allocate_tape3; *) +(* incr_counter_tape := incr_counter_tape3; *) +(* read_counter:=read_counter3; *) +(* counterG := counterG3; *) +(* error_name := gname; *) +(* tape_name := gname; *) +(* counter_name :=gname; *) +(* is_counter _ N c γ1 γ2 γ3 := inv N (counter_inv_pred3 c γ1 γ2 γ3); *) +(* counter_error_auth _ γ x := ●↯ x @ γ; *) +(* counter_error_frag _ γ x := ◯↯ x @ γ; *) +(* counter_tapes_auth _ γ m := (●((λ ns, (3, ns))<$>m)@γ)%I; *) +(* counter_tapes_frag _ γ α ns := (α◯↪N (3%nat;ns) @ γ)%I; *) +(* counter_content_auth _ γ z := own γ (●F z); *) +(* counter_content_frag _ γ f z := own γ (◯F{f} z); *) +(* new_counter_spec _ := new_counter_spec3; *) +(* allocate_tape_spec _ :=allocate_tape_spec3; *) +(* incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some3; *) +(* counter_presample_spec _ :=counter_presample_spec3; *) +(* read_counter_spec _ :=read_counter_spec3 *) +(* |}. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (??????) "(%&<-&H1)(%&<-&H2)". *) +(* iCombine "H1 H2" gives "%H". by rewrite excl_auth.excl_auth_auth_op_valid in H. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (??????) "(%&<-&H1)(%&<-&H2)". *) +(* iCombine "H1 H2" gives "%H". by rewrite excl_auth.excl_auth_frag_op_valid in H. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (?????) "H". *) +(* iApply (hocap_error_auth_pos with "[$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (?????) "H". *) +(* iApply (hocap_error_frag_pos with "[$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (??????) "H1 H2". *) +(* iApply (hocap_error_agree with "[$][$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. iIntros (???????) "??". *) +(* iApply (hocap_error_update with "[$][$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (??????) "H1 H2". *) +(* by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[%H _]". *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (???????) "H1 H2". *) +(* iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. *) +(* rewrite dfrac_op_own dfrac_valid_own. by intro. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros. *) +(* iDestruct (hocap_tapes_agree with "[$][$]") as "%H". *) +(* iPureIntro. *) +(* rewrite lookup_fmap_Some in H. destruct H as (?&?&?). *) +(* by simplify_eq. *) +(* Qed. *) +(* Next Obligation. *) +(* iIntros. *) +(* iMod (hocap_tapes_presample with "[$][$]") as "[??]". *) +(* iModIntro. iFrame. *) +(* by rewrite fmap_insert. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (??????) "H1 H2". *) +(* iCombine "H1 H2" gives "%H". by rewrite auth_auth_op_valid in H. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. iIntros (???? z z' ?) "H1 H2". *) +(* iCombine "H1 H2" gives "%H". *) +(* apply frac_auth_included_total in H. iPureIntro. *) +(* by apply nat_included. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (????????). *) +(* rewrite frac_auth_frag_op. by rewrite own_op. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. iIntros (??????) "H1 H2". *) +(* iCombine "H1 H2" gives "%H". *) +(* iPureIntro. *) +(* by apply frac_auth_agree_L in H. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. iIntros (????????) "H1 H2". *) +(* iMod (own_update_2 _ _ _ (_ ⋅ _) with "[$][$]") as "[$$]"; last done. *) +(* apply frac_auth_update. *) +(* apply nat_local_update. lia. *) +(* Qed. *) diff --git a/theories/coneris/examples/random_counter/random_counter.v b/theories/coneris/examples/random_counter/random_counter.v index 9c0c06b2..2491e18c 100644 --- a/theories/coneris/examples/random_counter/random_counter.v +++ b/theories/coneris/examples/random_counter/random_counter.v @@ -1,212 +1,216 @@ -From clutch.coneris Require Import coneris. +(* From clutch.coneris Require Import coneris. *) -Set Default Proof Using "Type*". +(* Set Default Proof Using "Type*". *) -Class random_counter `{!conerisGS Σ} := RandCounter -{ - (** * Operations *) - new_counter : val; - (* incr_counter : val; *) - allocate_tape : val; - incr_counter_tape : val; - read_counter : val; - (** * Ghost state *) - (** The assumptions about [Σ] *) - counterG : gFunctors → Type; - (** [name] is used to associate [locked] with [is_lock] *) - error_name : Type; - tape_name: Type; - counter_name: Type; - (** * Predicates *) - is_counter {L : counterG Σ} (N:namespace) (counter: val) - (γ1: error_name) (γ2: tape_name) (γ3: counter_name): iProp Σ; - counter_error_auth {L : counterG Σ} (γ: error_name) (x:R): iProp Σ; - counter_error_frag {L : counterG Σ} (γ: error_name) (x:R): iProp Σ; - counter_tapes_auth {L : counterG Σ} (γ: tape_name) (m:gmap loc (list nat)): iProp Σ; - counter_tapes_frag {L : counterG Σ} (γ: tape_name) (α:loc) (ns:list nat): iProp Σ; - counter_content_auth {L : counterG Σ} (γ: counter_name) (z:nat): iProp Σ; - counter_content_frag {L : counterG Σ} (γ: counter_name) (f:frac) (z:nat): iProp Σ; - (** * General properties of the predicates *) - #[global] is_counter_persistent {L : counterG Σ} N c γ1 γ2 γ3 :: - Persistent (is_counter (L:=L) N c γ1 γ2 γ3); - #[global] counter_error_auth_timeless {L : counterG Σ} γ x :: - Timeless (counter_error_auth (L:=L) γ x); - #[global] counter_error_frag_timeless {L : counterG Σ} γ x :: - Timeless (counter_error_frag (L:=L) γ x); - #[global] counter_tapes_auth_timeless {L : counterG Σ} γ m :: - Timeless (counter_tapes_auth (L:=L) γ m); - #[global] counter_tapes_frag_timeless {L : counterG Σ} γ α ns :: - Timeless (counter_tapes_frag (L:=L) γ α ns); - #[global] counter_content_auth_timeless {L : counterG Σ} γ z :: - Timeless (counter_content_auth (L:=L) γ z); - #[global] counter_content_frag_timeless {L : counterG Σ} γ f z :: - Timeless (counter_content_frag (L:=L) γ f z); +(* Class random_counter `{!conerisGS Σ} := RandCounter *) +(* { *) +(* (** * Operations *) *) +(* new_counter : val; *) +(* (* incr_counter : val; *) *) +(* allocate_tape : val; *) +(* incr_counter_tape : val; *) +(* read_counter : val; *) +(* (** * Ghost state *) *) +(* (** The assumptions about [Σ] *) *) +(* counterG : gFunctors → Type; *) +(* (** [name] is used to associate [locked] with [is_lock] *) *) +(* error_name : Type; *) +(* tape_name: Type; *) +(* counter_name: Type; *) +(* (** * Predicates *) *) +(* is_counter {L : counterG Σ} (N:namespace) (counter: val) *) +(* (γ1: error_name) (γ2: tape_name) (γ3: counter_name): iProp Σ; *) +(* counter_error_auth {L : counterG Σ} (γ: error_name) (x:R): iProp Σ; *) +(* counter_error_frag {L : counterG Σ} (γ: error_name) (x:R): iProp Σ; *) +(* counter_tapes_auth {L : counterG Σ} (γ: tape_name) (m:gmap loc (list nat)): iProp Σ; *) +(* counter_tapes_frag {L : counterG Σ} (γ: tape_name) (α:loc) (ns:list nat): iProp Σ; *) +(* counter_content_auth {L : counterG Σ} (γ: counter_name) (z:nat): iProp Σ; *) +(* counter_content_frag {L : counterG Σ} (γ: counter_name) (f:frac) (z:nat): iProp Σ; *) +(* (** * General properties of the predicates *) *) +(* #[global] is_counter_persistent {L : counterG Σ} N c γ1 γ2 γ3 :: *) +(* Persistent (is_counter (L:=L) N c γ1 γ2 γ3); *) +(* #[global] counter_error_auth_timeless {L : counterG Σ} γ x :: *) +(* Timeless (counter_error_auth (L:=L) γ x); *) +(* #[global] counter_error_frag_timeless {L : counterG Σ} γ x :: *) +(* Timeless (counter_error_frag (L:=L) γ x); *) +(* #[global] counter_tapes_auth_timeless {L : counterG Σ} γ m :: *) +(* Timeless (counter_tapes_auth (L:=L) γ m); *) +(* #[global] counter_tapes_frag_timeless {L : counterG Σ} γ α ns :: *) +(* Timeless (counter_tapes_frag (L:=L) γ α ns); *) +(* #[global] counter_content_auth_timeless {L : counterG Σ} γ z :: *) +(* Timeless (counter_content_auth (L:=L) γ z); *) +(* #[global] counter_content_frag_timeless {L : counterG Σ} γ f z :: *) +(* Timeless (counter_content_frag (L:=L) γ f z); *) - counter_error_auth_exclusive {L : counterG Σ} γ x1 x2: - counter_error_auth (L:=L) γ x1 -∗ counter_error_auth (L:=L) γ x2 -∗ False; - counter_error_frag_exclusive {L : counterG Σ} γ x1 x2: - counter_error_frag (L:=L) γ x1 -∗ counter_error_frag (L:=L) γ x2 -∗ False; - counter_error_auth_valid {L : counterG Σ} γ x: - counter_error_auth (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; - counter_error_auth_valid {L : counterG Σ} γ x: - counter_error_frag (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; - counter_error_agree {L : counterG Σ} γ x1 x2: - counter_error_auth (L:=L) γ x1 -∗ counter_error_frag (L:=L) γ x2 -∗ ⌜x1 = x2 ⌝; - counter_error_update {L : counterG Σ} γ x1 x2 (x3:nonnegreal): - counter_error_auth (L:=L) γ x1 -∗ counter_error_frag (L:=L) γ x2 ==∗ - counter_error_auth (L:=L) γ x3 ∗ counter_error_frag (L:=L) γ x3; - - counter_tapes_auth_exclusive {L : counterG Σ} γ m m': - counter_tapes_auth (L:=L) γ m -∗ counter_tapes_auth (L:=L) γ m' -∗ False; - counter_tapes_frag_exclusive {L : counterG Σ} γ α ns ns': - counter_tapes_frag (L:=L) γ α ns -∗ counter_tapes_frag (L:=L) γ α ns' -∗ False; - counter_tapes_agree {L : counterG Σ} γ α m ns: - counter_tapes_auth (L:=L) γ m -∗ counter_tapes_frag (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; - counter_tapes_update {L : counterG Σ} γ α m ns n: - counter_tapes_auth (L:=L) γ m -∗ counter_tapes_frag (L:=L) γ α ns ==∗ - counter_tapes_auth (L:=L) γ (<[α := (ns ++[n])]> m) ∗ counter_tapes_frag (L:=L) γ α (ns ++ [n]); +(* counter_error_auth_exclusive {L : counterG Σ} γ x1 x2: *) +(* counter_error_auth (L:=L) γ x1 -∗ counter_error_auth (L:=L) γ x2 -∗ False; *) +(* counter_error_frag_split {L : counterG Σ} γ (x1 x2:nonnegreal): *) +(* counter_error_frag (L:=L) γ x1 ∗ counter_error_frag (L:=L) γ x2 ⊣⊢ *) +(* counter_error_frag (L:=L) γ (x1+x2)%R ; *) +(* counter_error_auth_valid {L : counterG Σ} γ x: *) +(* counter_error_auth (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; *) +(* counter_error_frag_valid {L : counterG Σ} γ x: *) +(* counter_error_frag (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; *) +(* counter_error_ineq {L : counterG Σ} γ x1 x2: *) +(* counter_error_auth (L:=L) γ x1 -∗ counter_error_frag (L:=L) γ x2 -∗ ⌜(x2<=x1)%R ⌝; *) +(* counter_error_decrease {L : counterG Σ} γ (x1 x2:nonnegreal): *) +(* counter_error_auth (L:=L) γ (x1 + x2)%NNR -∗ counter_error_frag (L:=L) γ x1 ==∗ counter_error_auth (L:=L) γ x2; *) +(* counter_error_increase {L : counterG Σ} γ (x1 x2:nonnegreal): *) +(* (x1+x2<1)%R -> ⊢ counter_error_auth (L:=L) γ x1 ==∗ *) +(* counter_error_auth (L:=L) γ (x1+x2) ∗ counter_error_frag (L:=L) γ x2; *) + +(* counter_tapes_auth_exclusive {L : counterG Σ} γ m m': *) +(* counter_tapes_auth (L:=L) γ m -∗ counter_tapes_auth (L:=L) γ m' -∗ False; *) +(* counter_tapes_frag_exclusive {L : counterG Σ} γ α ns ns': *) +(* counter_tapes_frag (L:=L) γ α ns -∗ counter_tapes_frag (L:=L) γ α ns' -∗ False; *) +(* counter_tapes_agree {L : counterG Σ} γ α m ns: *) +(* counter_tapes_auth (L:=L) γ m -∗ counter_tapes_frag (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; *) +(* counter_tapes_update {L : counterG Σ} γ α m ns n: *) +(* counter_tapes_auth (L:=L) γ m -∗ counter_tapes_frag (L:=L) γ α ns ==∗ *) +(* counter_tapes_auth (L:=L) γ (<[α := (ns ++[n])]> m) ∗ counter_tapes_frag (L:=L) γ α (ns ++ [n]); *) - counter_content_auth_exclusive {L : counterG Σ} γ z1 z2: - counter_content_auth (L:=L) γ z1 -∗ counter_content_auth (L:=L) γ z2 -∗ False; - counter_content_less_than {L : counterG Σ} γ z z' f: - counter_content_auth (L:=L) γ z -∗ counter_content_frag (L:=L) γ f z' -∗ ⌜(z'<=z)%nat ⌝; - counter_content_frag_combine {L:counterG Σ} γ f f' z z': - (counter_content_frag (L:=L) γ f z ∗ counter_content_frag (L:=L) γ f' z')%I ≡ - counter_content_frag (L:=L) γ (f+f')%Qp (z+z')%nat; - counter_content_agree {L : counterG Σ} γ z z': - counter_content_auth (L:=L) γ z -∗ counter_content_frag (L:=L) γ 1%Qp z' -∗ ⌜(z'=z)%nat ⌝; - counter_content_update {L : counterG Σ} γ f z1 z2 z3: - counter_content_auth (L:=L) γ z1 -∗ counter_content_frag (L:=L) γ f z2 ==∗ - counter_content_auth (L:=L) γ (z1+z3)%nat ∗ counter_content_frag (L:=L) γ f (z2+z3)%nat; +(* counter_content_auth_exclusive {L : counterG Σ} γ z1 z2: *) +(* counter_content_auth (L:=L) γ z1 -∗ counter_content_auth (L:=L) γ z2 -∗ False; *) +(* counter_content_less_than {L : counterG Σ} γ z z' f: *) +(* counter_content_auth (L:=L) γ z -∗ counter_content_frag (L:=L) γ f z' -∗ ⌜(z'<=z)%nat ⌝; *) +(* counter_content_frag_combine {L:counterG Σ} γ f f' z z': *) +(* (counter_content_frag (L:=L) γ f z ∗ counter_content_frag (L:=L) γ f' z')%I ≡ *) +(* counter_content_frag (L:=L) γ (f+f')%Qp (z+z')%nat; *) +(* counter_content_agree {L : counterG Σ} γ z z': *) +(* counter_content_auth (L:=L) γ z -∗ counter_content_frag (L:=L) γ 1%Qp z' -∗ ⌜(z'=z)%nat ⌝; *) +(* counter_content_update {L : counterG Σ} γ f z1 z2 z3: *) +(* counter_content_auth (L:=L) γ z1 -∗ counter_content_frag (L:=L) γ f z2 ==∗ *) +(* counter_content_auth (L:=L) γ (z1+z3)%nat ∗ counter_content_frag (L:=L) γ f (z2+z3)%nat; *) - (** * Program specs *) - new_counter_spec {L : counterG Σ} E ε N: - {{{ ↯ ε }}} new_counter #() @E - {{{ c, RET c; ∃ γ1 γ2 γ3, is_counter (L:=L) N c γ1 γ2 γ3 ∗ - counter_error_frag (L:=L) γ1 ε ∗ - counter_content_frag (L:=L) γ3 1%Qp 0%nat - }}}; - allocate_tape_spec {L: counterG Σ} N E c γ1 γ2 γ3: - ↑N ⊆ E-> - {{{ is_counter (L:=L) N c γ1 γ2 γ3 }}} - allocate_tape #() @ E - {{{ (v:val), RET v; - ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ counter_tapes_frag (L:=L) γ2 α [] - }}}; - incr_counter_tape_spec_some {L: counterG Σ} N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: - ↑N⊆E -> - {{{ is_counter (L:=L) N c γ1 γ2 γ3 ∗ - □ (∀ (z:nat), P ∗ counter_content_auth (L:=L) γ3 z ={E∖↑N}=∗ - counter_content_auth (L:=L) γ3 (z+n)%nat ∗ Q z) ∗ - P ∗ counter_tapes_frag (L:=L) γ2 α (n::ns) - }}} - incr_counter_tape c #lbl:α @ E - {{{ (z:nat), RET (#z, #n); Q z ∗ counter_tapes_frag (L:=L) γ2 α ns}}}; - counter_presample_spec {L: counterG Σ} NS E ns α - (ε2 : R -> nat -> R) - (P : iProp Σ) T γ1 γ2 γ3 c: - ↑NS ⊆ E -> - (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> - (∀ (ε:R), 0<= ε ->SeriesC (λ n, if (bool_decide (n≤3%nat)) then 1 / (S 3%nat) * ε2 ε n else 0%R)%R <= ε)%R-> - is_counter (L:=L) NS c γ1 γ2 γ3 -∗ - (□∀ (ε:R) n, (P ∗ counter_error_auth (L:=L) γ1 ε) ={E∖↑NS}=∗ - (⌜(1<=ε2 ε n)%R⌝ ∨ (counter_error_auth (L:=L) γ1 (ε2 ε n) ∗ T (n)))) - -∗ - P -∗ counter_tapes_frag (L:=L) γ2 α ns-∗ - wp_update E (∃ n, T (n) ∗ counter_tapes_frag (L:=L) γ2 α (ns++[n])); - read_counter_spec {L: counterG Σ} N E c γ1 γ2 γ3 P Q: - ↑N ⊆ E -> - {{{ is_counter (L:=L) N c γ1 γ2 γ3 ∗ - □ (∀ (z:nat), P ∗ counter_content_auth (L:=L) γ3 z ={E∖↑N}=∗ - counter_content_auth (L:=L) γ3 z∗ Q z) - ∗ P - }}} - read_counter c @ E - {{{ (n':nat), RET #n'; Q n' - }}} -}. +(* (** * Program specs *) *) +(* new_counter_spec {L : counterG Σ} E ε N: *) +(* {{{ ↯ ε }}} new_counter #() @E *) +(* {{{ c, RET c; ∃ γ1 γ2 γ3, is_counter (L:=L) N c γ1 γ2 γ3 ∗ *) +(* counter_error_frag (L:=L) γ1 ε ∗ *) +(* counter_content_frag (L:=L) γ3 1%Qp 0%nat *) +(* }}}; *) +(* allocate_tape_spec {L: counterG Σ} N E c γ1 γ2 γ3: *) +(* ↑N ⊆ E-> *) +(* {{{ is_counter (L:=L) N c γ1 γ2 γ3 }}} *) +(* allocate_tape #() @ E *) +(* {{{ (v:val), RET v; *) +(* ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ counter_tapes_frag (L:=L) γ2 α [] *) +(* }}}; *) +(* incr_counter_tape_spec_some {L: counterG Σ} N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: *) +(* ↑N⊆E -> *) +(* {{{ is_counter (L:=L) N c γ1 γ2 γ3 ∗ *) +(* □ (∀ (z:nat), P ∗ counter_content_auth (L:=L) γ3 z ={E∖↑N}=∗ *) +(* counter_content_auth (L:=L) γ3 (z+n)%nat ∗ Q z) ∗ *) +(* P ∗ counter_tapes_frag (L:=L) γ2 α (n::ns) *) +(* }}} *) +(* incr_counter_tape c #lbl:α @ E *) +(* {{{ (z:nat), RET (#z, #n); Q z ∗ counter_tapes_frag (L:=L) γ2 α ns}}}; *) +(* counter_presample_spec {L: counterG Σ} NS E ns α *) +(* (ε2 : R -> nat -> R) *) +(* (P : iProp Σ) T γ1 γ2 γ3 c: *) +(* ↑NS ⊆ E -> *) +(* (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> *) +(* (∀ (ε:R), 0<= ε ->SeriesC (λ n, if (bool_decide (n≤3%nat)) then 1 / (S 3%nat) * ε2 ε n else 0%R)%R <= ε)%R-> *) +(* is_counter (L:=L) NS c γ1 γ2 γ3 -∗ *) +(* (□∀ (ε ε':nonnegreal) n, (P ∗ counter_error_auth (L:=L) γ1 (ε')%R) ={E∖↑NS}=∗ *) +(* (⌜(1<=ε2 ε n)%R⌝ ∨ (counter_error_auth (L:=L) γ1 (ε' + )%R ∗ T (n)))) *) +(* -∗ *) +(* P -∗ counter_tapes_frag (L:=L) γ2 α ns-∗ *) +(* wp_update E (∃ n, T (n) ∗ counter_tapes_frag (L:=L) γ2 α (ns++[n])); *) +(* read_counter_spec {L: counterG Σ} N E c γ1 γ2 γ3 P Q: *) +(* ↑N ⊆ E -> *) +(* {{{ is_counter (L:=L) N c γ1 γ2 γ3 ∗ *) +(* □ (∀ (z:nat), P ∗ counter_content_auth (L:=L) γ3 z ={E∖↑N}=∗ *) +(* counter_content_auth (L:=L) γ3 z∗ Q z) *) +(* ∗ P *) +(* }}} *) +(* read_counter c @ E *) +(* {{{ (n':nat), RET #n'; Q n' *) +(* }}} *) +(* }. *) -Section lemmas. - Context `{rc:random_counter} {L: counterG Σ}. +(* Section lemmas. *) +(* Context `{rc:random_counter} {L: counterG Σ}. *) - Lemma incr_counter_tape_spec_none N E c γ1 γ2 γ3 (ε2:R -> nat -> R) (P: iProp Σ) (T: nat -> iProp Σ) (Q: nat -> nat -> iProp Σ)(α:loc): - ↑N ⊆ E-> - (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> - (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → - {{{ is_counter (L:=L) N c γ1 γ2 γ3 ∗ - □(∀ (ε:R) (n : nat), P ∗ counter_error_auth (L:=L) γ1 ε - ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨ counter_error_auth (L:=L) γ1 (ε2 ε n) ∗ T n) ) ∗ - □ (∀ (n:nat) (z:nat), T n ∗ counter_content_auth (L:=L) γ3 z ={E∖↑N}=∗ - counter_content_auth (L:=L) γ3 (z+n)%nat ∗ Q z n) ∗ - P ∗ counter_tapes_frag (L:=L) γ2 α [] - }}} - incr_counter_tape c #lbl:α @ E - {{{ (z:nat) (n:nat), RET (#z, #n); Q z n ∗ counter_tapes_frag (L:=L) γ2 α [] }}}. - Proof. - iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP & Hα) HΦ". - iMod (counter_presample_spec with "[//][//][$][$]") as "(%&HT&Hα)"; try done. - { intros ε Hε. specialize (Hineq ε Hε). - rewrite SeriesC_nat_bounded_fin SeriesC_finite_foldr /=. lra. - } - iApply (incr_counter_tape_spec_some _ _ _ _ _ _ (T n) (λ x, Q x n) with "[$Hα $HT]"); try done. - { by iSplit. } - iNext. - iIntros. iApply ("HΦ" with "[$]"). - Qed. +(* Lemma incr_counter_tape_spec_none N E c γ1 γ2 γ3 (ε2:R -> nat -> R) (P: iProp Σ) (T: nat -> iProp Σ) (Q: nat -> nat -> iProp Σ)(α:loc): *) +(* ↑N ⊆ E-> *) +(* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) +(* (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → *) +(* {{{ is_counter (L:=L) N c γ1 γ2 γ3 ∗ *) +(* □(∀ (ε ε': nonnegreal) (n : nat), P ∗ counter_error_auth (L:=L) γ1 (ε+ε') *) +(* ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨ counter_error_auth (L:=L) γ1 (ε2 ε n + ε') ∗ T n) ) ∗ *) +(* □ (∀ (n:nat) (z:nat), T n ∗ counter_content_auth (L:=L) γ3 z ={E∖↑N}=∗ *) +(* counter_content_auth (L:=L) γ3 (z+n)%nat ∗ Q z n) ∗ *) +(* P ∗ counter_tapes_frag (L:=L) γ2 α [] *) +(* }}} *) +(* incr_counter_tape c #lbl:α @ E *) +(* {{{ (z:nat) (n:nat), RET (#z, #n); Q z n ∗ counter_tapes_frag (L:=L) γ2 α [] }}}. *) +(* Proof. *) +(* iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP & Hα) HΦ". *) +(* iMod (counter_presample_spec with "[//][//][$][$]") as "(%&HT&Hα)"; try done. *) +(* { intros ε Hε. specialize (Hineq ε Hε). *) +(* rewrite SeriesC_nat_bounded_fin SeriesC_finite_foldr /=. lra. *) +(* } *) +(* iApply (incr_counter_tape_spec_some _ _ _ _ _ _ (T n) (λ x, Q x n) with "[$Hα $HT]"); try done. *) +(* { by iSplit. } *) +(* iNext. *) +(* iIntros. iApply ("HΦ" with "[$]"). *) +(* Qed. *) - Lemma incr_counter_tape_spec_none' N E c γ1 γ2 γ3 ε (ε2:R -> nat -> R)(α:loc) (ns:list nat) (q:Qp) (z:nat): - ↑N ⊆ E-> - (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> - (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → - {{{ is_counter (L:=L) N c γ1 γ2 γ3 ∗ - counter_error_frag (L:=L) γ1 ε ∗ - counter_tapes_frag (L:=L) γ2 α [] ∗ - counter_content_frag (L:=L) γ3 q z - }}} - incr_counter_tape c #lbl:α @ E - {{{ (z':nat) (n:nat), - RET (#z', #n); ⌜(0<=n<4)%nat⌝ ∗ - ⌜(z<=z')%nat⌝ ∗ - counter_error_frag (L:=L) γ1 (ε2 ε n) ∗ - counter_tapes_frag (L:=L) γ2 α [] ∗ - counter_content_frag (L:=L) γ3 q (z+n)%nat }}}. - Proof. - iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & Herr & Htapes & Hcontent) HΦ". - pose (ε2' := (λ ε x, if (x<=?3)%nat then ε2 ε x else 1%R)). - wp_apply (incr_counter_tape_spec_none _ _ _ _ _ _ ε2' - (counter_error_frag γ1 ε ∗ counter_content_frag γ3 q z)%I - (λ n, ⌜0<=n<4⌝ ∗ counter_error_frag γ1 (ε2' ε n) ∗ counter_content_frag γ3 q z)%I - (λ z' n, ⌜0<=n<4⌝ ∗ ⌜z<=z'⌝ ∗ counter_error_frag γ1 (ε2' ε n) ∗ counter_content_frag γ3 q (z+n)%nat)%I - with "[$Herr $Htapes $Hcontent]"). - - done. - - rewrite /ε2'. intros. - case_match; [naive_solver|lra]. - - rewrite /ε2'. simpl. intros. naive_solver. - - repeat iSplit; first done. - + iModIntro. - iIntros (ε' n) "[(Herr &Hcontent) Herr_auth]". - destruct (n<=?3) eqn:H; last first. - { iLeft. iPureIntro. rewrite /ε2'. by rewrite H. } - iRight. - iFrame. - iDestruct (counter_error_agree with "[$][$]") as "->". - iDestruct (counter_error_auth_pos with "[$]") as "%". - unshelve iMod (counter_error_update _ _ _ (mknonnegreal (ε2' ε n) _) with "[$][$]") as "[$$]". - { rewrite /ε2' H. - naive_solver. } - iPureIntro. split; first lia. - apply leb_complete in H. lia. - + iModIntro. - iIntros (??) "[(%&Herr & Hcontent) Hcontent_auth]". - iFrame. - iDestruct (counter_content_less_than with "[$][$]") as "%". - by iMod (counter_content_update with "[$][$]") as "[$ $]". - - iIntros (??) "[(%&%&?&?)?]". - iApply "HΦ". - iFrame. - rewrite /ε2'. case_match eqn: H2; first by iFrame. - apply leb_iff_conv in H2. lia. - Qed. +(* Lemma incr_counter_tape_spec_none' N E c γ1 γ2 γ3 ε (ε2:R -> nat -> R)(α:loc) (ns:list nat) (q:Qp) (z:nat): *) +(* ↑N ⊆ E-> *) +(* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) +(* (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → *) +(* {{{ is_counter (L:=L) N c γ1 γ2 γ3 ∗ *) +(* counter_error_frag (L:=L) γ1 ε ∗ *) +(* counter_tapes_frag (L:=L) γ2 α [] ∗ *) +(* counter_content_frag (L:=L) γ3 q z *) +(* }}} *) +(* incr_counter_tape c #lbl:α @ E *) +(* {{{ (z':nat) (n:nat), *) +(* RET (#z', #n); ⌜(0<=n<4)%nat⌝ ∗ *) +(* ⌜(z<=z')%nat⌝ ∗ *) +(* counter_error_frag (L:=L) γ1 (ε2 ε n) ∗ *) +(* counter_tapes_frag (L:=L) γ2 α [] ∗ *) +(* counter_content_frag (L:=L) γ3 q (z+n)%nat }}}. *) +(* Proof. *) +(* iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & Herr & Htapes & Hcontent) HΦ". *) +(* pose (ε2' := (λ ε x, if (x<=?3)%nat then ε2 ε x else 1%R)). *) +(* wp_apply (incr_counter_tape_spec_none _ _ _ _ _ _ ε2' *) +(* (counter_error_frag γ1 ε ∗ counter_content_frag γ3 q z)%I *) +(* (λ n, ⌜0<=n<4⌝ ∗ counter_error_frag γ1 (ε2' ε n) ∗ counter_content_frag γ3 q z)%I *) +(* (λ z' n, ⌜0<=n<4⌝ ∗ ⌜z<=z'⌝ ∗ counter_error_frag γ1 (ε2' ε n) ∗ counter_content_frag γ3 q (z+n)%nat)%I *) +(* with "[$Herr $Htapes $Hcontent]"). *) +(* - done. *) +(* - rewrite /ε2'. intros. *) +(* case_match; [naive_solver|lra]. *) +(* - rewrite /ε2'. simpl. intros. naive_solver. *) +(* - repeat iSplit; first done. *) +(* + iModIntro. *) +(* iIntros (εa εb n) "[(Herr &Hcontent) Herr_auth]". *) +(* destruct (n<=?3) eqn:H; last first. *) +(* { iLeft. iPureIntro. rewrite /ε2'. by rewrite H. } *) +(* iRight. *) +(* iFrame. *) +(* iDestruct (counter_error_ineq with "[$][$]") as "%". *) +(* iDestruct (counter_error_auth_valid with "[$]") as "%". *) +(* iDestruct (counter_error_frag_valid with "[$]") as "%". *) +(* unshelve iMod (counter_error_update _ _ _ (mknonnegreal (ε2' ε n) _) with "[$][$]") as "[$$]". *) +(* { rewrite /ε2' H. *) +(* naive_solver. } *) +(* iPureIntro. split; first lia. *) +(* apply leb_complete in H. lia. *) +(* + iModIntro. *) +(* iIntros (??) "[(%&Herr & Hcontent) Hcontent_auth]". *) +(* iFrame. *) +(* iDestruct (counter_content_less_than with "[$][$]") as "%". *) +(* by iMod (counter_content_update with "[$][$]") as "[$ $]". *) +(* - iIntros (??) "[(%&%&?&?)?]". *) +(* iApply "HΦ". *) +(* iFrame. *) +(* rewrite /ε2'. case_match eqn: H2; first by iFrame. *) +(* apply leb_iff_conv in H2. lia. *) +(* Qed. *) -End lemmas. +(* End lemmas. *) diff --git a/theories/coneris/lib/flip_spec.v b/theories/coneris/lib/flip_spec.v index 50296ae0..c5f1bd71 100644 --- a/theories/coneris/lib/flip_spec.v +++ b/theories/coneris/lib/flip_spec.v @@ -1,431 +1,431 @@ -From clutch.coneris Require Import coneris. -From clutch.coneris Require Import flip hocap. +(* From clutch.coneris Require Import coneris. *) +(* From clutch.coneris Require Import flip hocap. *) -Set Default Proof Using "Type*". +(* Set Default Proof Using "Type*". *) -Class flip_spec `{!conerisGS Σ} := FlipSpec -{ - (** * Operations *) - flip_allocate_tape : val; - flip_tape : val; - (** * Ghost state *) - (** The assumptions about [Σ] *) - flipG : gFunctors → Type; - (** [name] is used to associate [locked] with [is_lock] *) - flip_error_name : Type; - flip_tape_name: Type; - (** * Predicates *) - is_flip {L : flipG Σ} (N:namespace) - (γ1: flip_error_name) (γ2: flip_tape_name): iProp Σ; - flip_error_auth {L : flipG Σ} (γ: flip_error_name) (x:R): iProp Σ; - flip_error_frag {L : flipG Σ} (γ: flip_error_name) (x:R): iProp Σ; - flip_tapes_auth {L : flipG Σ} (γ: flip_tape_name) (m:gmap loc (list bool)): iProp Σ; - flip_tapes_frag {L : flipG Σ} (γ: flip_tape_name) (α:loc) (ns:list bool): iProp Σ; - (** * General properties of the predicates *) - #[global] is_flip_persistent {L : flipG Σ} N γ1 γ2 :: - Persistent (is_flip (L:=L) N γ1 γ2); - #[global] flip_error_auth_timeless {L : flipG Σ} γ x :: - Timeless (flip_error_auth (L:=L) γ x); - #[global] flip_error_frag_timeless {L : flipG Σ} γ x :: - Timeless (flip_error_frag (L:=L) γ x); - #[global] flip_tapes_auth_timeless {L : flipG Σ} γ m :: - Timeless (flip_tapes_auth (L:=L) γ m); - #[global] flip_tapes_frag_timeless {L : flipG Σ} γ α ns :: - Timeless (flip_tapes_frag (L:=L) γ α ns); - #[global] flip_error_name_inhabited:: - Inhabited flip_error_name; - #[global] flip_tape_name_inhabited:: - Inhabited flip_tape_name; +(* Class flip_spec `{!conerisGS Σ} := FlipSpec *) +(* { *) +(* (** * Operations *) *) +(* flip_allocate_tape : val; *) +(* flip_tape : val; *) +(* (** * Ghost state *) *) +(* (** The assumptions about [Σ] *) *) +(* flipG : gFunctors → Type; *) +(* (** [name] is used to associate [locked] with [is_lock] *) *) +(* flip_error_name : Type; *) +(* flip_tape_name: Type; *) +(* (** * Predicates *) *) +(* is_flip {L : flipG Σ} (N:namespace) *) +(* (γ1: flip_error_name) (γ2: flip_tape_name): iProp Σ; *) +(* flip_error_auth {L : flipG Σ} (γ: flip_error_name) (x:R): iProp Σ; *) +(* flip_error_frag {L : flipG Σ} (γ: flip_error_name) (x:R): iProp Σ; *) +(* flip_tapes_auth {L : flipG Σ} (γ: flip_tape_name) (m:gmap loc (list bool)): iProp Σ; *) +(* flip_tapes_frag {L : flipG Σ} (γ: flip_tape_name) (α:loc) (ns:list bool): iProp Σ; *) +(* (** * General properties of the predicates *) *) +(* #[global] is_flip_persistent {L : flipG Σ} N γ1 γ2 :: *) +(* Persistent (is_flip (L:=L) N γ1 γ2); *) +(* #[global] flip_error_auth_timeless {L : flipG Σ} γ x :: *) +(* Timeless (flip_error_auth (L:=L) γ x); *) +(* #[global] flip_error_frag_timeless {L : flipG Σ} γ x :: *) +(* Timeless (flip_error_frag (L:=L) γ x); *) +(* #[global] flip_tapes_auth_timeless {L : flipG Σ} γ m :: *) +(* Timeless (flip_tapes_auth (L:=L) γ m); *) +(* #[global] flip_tapes_frag_timeless {L : flipG Σ} γ α ns :: *) +(* Timeless (flip_tapes_frag (L:=L) γ α ns); *) +(* #[global] flip_error_name_inhabited:: *) +(* Inhabited flip_error_name; *) +(* #[global] flip_tape_name_inhabited:: *) +(* Inhabited flip_tape_name; *) - flip_error_auth_exclusive {L : flipG Σ} γ x1 x2: - flip_error_auth (L:=L) γ x1 -∗ flip_error_auth (L:=L) γ x2 -∗ False; - flip_error_frag_split {L : flipG Σ} γ (x1 x2:nonnegreal): - flip_error_frag (L:=L) γ x1 ∗ flip_error_frag (L:=L) γ x2 ⊣⊢ - flip_error_frag (L:=L) γ (x1+x2)%R ; - flip_error_auth_valid {L : flipG Σ} γ x: - flip_error_auth (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; - flip_error_frag_valid {L : flipG Σ} γ x: - flip_error_frag (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; - flip_error_ineq {L : flipG Σ} γ x1 x2: - flip_error_auth (L:=L) γ x1 -∗ flip_error_frag (L:=L) γ x2 -∗ ⌜(x2 <= x1)%R ⌝; - flip_error_decrease {L : flipG Σ} γ (x1 x2:nonnegreal): - flip_error_auth (L:=L) γ (x1 + x2)%NNR -∗ flip_error_frag (L:=L) γ x1 ==∗ flip_error_auth (L:=L) γ x2; - flip_error_increase {L : flipG Σ} γ (x1 x2:nonnegreal): - (x1+x2<1)%R -> ⊢ flip_error_auth (L:=L) γ x1 ==∗ - flip_error_auth (L:=L) γ (x1+x2) ∗ flip_error_frag (L:=L) γ x2; +(* flip_error_auth_exclusive {L : flipG Σ} γ x1 x2: *) +(* flip_error_auth (L:=L) γ x1 -∗ flip_error_auth (L:=L) γ x2 -∗ False; *) +(* flip_error_frag_split {L : flipG Σ} γ (x1 x2:nonnegreal): *) +(* flip_error_frag (L:=L) γ x1 ∗ flip_error_frag (L:=L) γ x2 ⊣⊢ *) +(* flip_error_frag (L:=L) γ (x1+x2)%R ; *) +(* flip_error_auth_valid {L : flipG Σ} γ x: *) +(* flip_error_auth (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; *) +(* flip_error_frag_valid {L : flipG Σ} γ x: *) +(* flip_error_frag (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; *) +(* flip_error_ineq {L : flipG Σ} γ x1 x2: *) +(* flip_error_auth (L:=L) γ x1 -∗ flip_error_frag (L:=L) γ x2 -∗ ⌜(x2 <= x1)%R ⌝; *) +(* flip_error_decrease {L : flipG Σ} γ (x1 x2:nonnegreal): *) +(* flip_error_auth (L:=L) γ (x1 + x2)%NNR -∗ flip_error_frag (L:=L) γ x1 ==∗ flip_error_auth (L:=L) γ x2; *) +(* flip_error_increase {L : flipG Σ} γ (x1 x2:nonnegreal): *) +(* (x1+x2<1)%R -> ⊢ flip_error_auth (L:=L) γ x1 ==∗ *) +(* flip_error_auth (L:=L) γ (x1+x2) ∗ flip_error_frag (L:=L) γ x2; *) - flip_tapes_auth_exclusive {L : flipG Σ} γ m m': - flip_tapes_auth (L:=L) γ m -∗ flip_tapes_auth (L:=L) γ m' -∗ False; - flip_tapes_frag_exclusive {L : flipG Σ} γ α ns ns': - flip_tapes_frag (L:=L) γ α ns -∗ flip_tapes_frag (L:=L) γ α ns' -∗ False; - flip_tapes_agree {L : flipG Σ} γ α m ns: - flip_tapes_auth (L:=L) γ m -∗ flip_tapes_frag (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; - flip_tapes_update {L : flipG Σ} γ α m ns n: - flip_tapes_auth (L:=L) γ m -∗ flip_tapes_frag (L:=L) γ α ns ==∗ - flip_tapes_auth (L:=L) γ (<[α := (ns ++[n])]> m) ∗ flip_tapes_frag (L:=L) γ α (ns ++ [n]); - flip_tapes_pop {L : flipG Σ} γ α m ns n: - flip_tapes_auth (L:=L) γ m -∗ flip_tapes_frag (L:=L) γ α (n::ns) ==∗ - flip_tapes_auth (L:=L) γ (<[α := (ns)]> m) ∗ flip_tapes_frag (L:=L) γ α (ns); +(* flip_tapes_auth_exclusive {L : flipG Σ} γ m m': *) +(* flip_tapes_auth (L:=L) γ m -∗ flip_tapes_auth (L:=L) γ m' -∗ False; *) +(* flip_tapes_frag_exclusive {L : flipG Σ} γ α ns ns': *) +(* flip_tapes_frag (L:=L) γ α ns -∗ flip_tapes_frag (L:=L) γ α ns' -∗ False; *) +(* flip_tapes_agree {L : flipG Σ} γ α m ns: *) +(* flip_tapes_auth (L:=L) γ m -∗ flip_tapes_frag (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; *) +(* flip_tapes_update {L : flipG Σ} γ α m ns n: *) +(* flip_tapes_auth (L:=L) γ m -∗ flip_tapes_frag (L:=L) γ α ns ==∗ *) +(* flip_tapes_auth (L:=L) γ (<[α := (ns ++[n])]> m) ∗ flip_tapes_frag (L:=L) γ α (ns ++ [n]); *) +(* flip_tapes_pop {L : flipG Σ} γ α m ns n: *) +(* flip_tapes_auth (L:=L) γ m -∗ flip_tapes_frag (L:=L) γ α (n::ns) ==∗ *) +(* flip_tapes_auth (L:=L) γ (<[α := (ns)]> m) ∗ flip_tapes_frag (L:=L) γ α (ns); *) - (** * Program specs *) - flip_inv_create_spec {L : flipG Σ} N E ε: - ↯ ε -∗ - wp_update E (∃ γ1 γ2, is_flip (L:=L) N γ1 γ2 ∗ - flip_error_frag (L:=L) γ1 ε); +(* (** * Program specs *) *) +(* flip_inv_create_spec {L : flipG Σ} N E ε: *) +(* ↯ ε -∗ *) +(* wp_update E (∃ γ1 γ2, is_flip (L:=L) N γ1 γ2 ∗ *) +(* flip_error_frag (L:=L) γ1 ε); *) - flip_allocate_tape_spec {L: flipG Σ} N E γ1 γ2: - ↑N ⊆ E-> - {{{ is_flip (L:=L) N γ1 γ2 }}} - flip_allocate_tape #() @ E - {{{ (v:val), RET v; - ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ flip_tapes_frag (L:=L) γ2 α [] - }}}; - flip_tape_spec_some {L: flipG Σ} N E γ1 γ2 (P: iProp Σ) (Q:iProp Σ) (α:loc) (n:bool) ns: - ↑N⊆E -> - {{{ is_flip (L:=L) N γ1 γ2 ∗ - □ (∀ m, P ∗ - flip_tapes_auth (L:=L) γ2 m - ={E∖↑N}=∗ ⌜m!!α=Some (n::ns)⌝ ∗ Q ∗ flip_tapes_auth (L:=L) γ2 (<[α:=ns]> m)) ∗ - P - }}} - flip_tape #lbl:α @ E - {{{ RET #n; Q }}}; - flip_presample_spec {L: flipG Σ} NS E ns α - (ε2 : R -> list bool -> R) - (P : iProp Σ) num T γ1 γ2: - ↑NS ⊆ E -> - (∀ ε l, 0<= ε -> length l = num -> 0<=ε2 ε l)%R -> - (∀ (ε:R), 0<= ε -> SeriesC (λ l, if length l =? num then ε2 ε l else 0) /(2^num) <= ε)%R-> - is_flip (L:=L) NS γ1 γ2 -∗ - (□∀ (ε:R) n, (P ∗ flip_error_auth (L:=L) γ1 ε) ={E∖↑NS}=∗ - (⌜(1<=ε2 ε n)%R⌝ ∨ (flip_error_auth (L:=L) γ1 (ε2 ε n) ∗ T (n)))) - -∗ - P -∗ flip_tapes_frag (L:=L) γ2 α ns-∗ - wp_update E (∃ n, T (n) ∗ flip_tapes_frag (L:=L) γ2 α (ns++n)) -}. +(* flip_allocate_tape_spec {L: flipG Σ} N E γ1 γ2: *) +(* ↑N ⊆ E-> *) +(* {{{ is_flip (L:=L) N γ1 γ2 }}} *) +(* flip_allocate_tape #() @ E *) +(* {{{ (v:val), RET v; *) +(* ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ flip_tapes_frag (L:=L) γ2 α [] *) +(* }}}; *) +(* flip_tape_spec_some {L: flipG Σ} N E γ1 γ2 (P: iProp Σ) (Q:iProp Σ) (α:loc) (n:bool) ns: *) +(* ↑N⊆E -> *) +(* {{{ is_flip (L:=L) N γ1 γ2 ∗ *) +(* □ (∀ m, P ∗ *) +(* flip_tapes_auth (L:=L) γ2 m *) +(* ={E∖↑N}=∗ ⌜m!!α=Some (n::ns)⌝ ∗ Q ∗ flip_tapes_auth (L:=L) γ2 (<[α:=ns]> m)) ∗ *) +(* P *) +(* }}} *) +(* flip_tape #lbl:α @ E *) +(* {{{ RET #n; Q }}}; *) +(* flip_presample_spec {L: flipG Σ} NS E ns α *) +(* (ε2 : R -> list bool -> R) *) +(* (P : _ -> iProp Σ) num T γ1 γ2 ε: *) +(* ↑NS ⊆ E -> *) +(* (∀ ε l, 0<= ε -> length l = num -> 0<=ε2 ε l)%R -> *) +(* (∀ (ε:R), 0<= ε -> SeriesC (λ l, if length l =? num then ε2 ε l else 0) /(2^num) <= ε)%R-> *) +(* is_flip (L:=L) NS γ1 γ2 -∗ *) +(* (□∀ (ε ε':nonnegreal) n, (P ε ∗ flip_error_auth (L:=L) γ1 (ε + ε')%R) ={E∖↑NS}=∗ *) +(* (⌜(1<=ε2 ε n)%R⌝ ∨ (flip_error_auth (L:=L) γ1 (ε2 ε n + ε')%R ∗ T ε n))) *) +(* -∗ *) +(* P ε -∗ flip_tapes_frag (L:=L) γ2 α ns-∗ *) +(* wp_update E (∃ n, T ε n ∗ flip_tapes_frag (L:=L) γ2 α (ns++n)) *) +(* }. *) -Section test. - Context `{F:flip_spec}. - Lemma flip_presample_spec_simple {L: flipG Σ} NS E ns α - (ε2 : R -> bool -> R) - (P : iProp Σ) T γ1 γ2: - ↑NS ⊆ E -> - (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> - (∀ (ε:R), 0<= ε -> (ε2 ε true + ε2 ε false)/2 <= ε)%R-> - is_flip (L:=L) NS γ1 γ2 -∗ - (□∀ (ε:R) n, (P ∗ flip_error_auth (L:=L) γ1 ε) ={E∖↑NS}=∗ - (⌜(1<=ε2 ε n)%R⌝ ∨ (flip_error_auth (L:=L) γ1 (ε2 ε n) ∗ T (n)))) - -∗ - P -∗ flip_tapes_frag (L:=L) γ2 α ns-∗ - wp_update E (∃ n, T (n) ∗ flip_tapes_frag (L:=L) γ2 α (ns++[n])). - Proof. - iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". - pose (ε2' ε l:= match l with - | [b]=> ε2 ε b - | _ => 1%R - end - ). - iMod (flip_presample_spec _ _ _ _ ε2' P 1%nat (λ l, match l with | [b] => T b | _ => False end )%I with "[//][][$][$]") as "(%l & [??])"; first done. - - rewrite /ε2'. - intros. repeat case_match; try done. naive_solver. - - intros. - etrans; last apply Hineq; try done. - erewrite (SeriesC_subset (λ x, x ∈ [[true]; [false]])); last first. - + intros ? H'. - case_match eqn:K; last done. - rewrite Nat.eqb_eq in K. - exfalso. - apply H'. - destruct a as [|[|] [|]]; try (simpl in *; done). - all: set_solver. - + rewrite SeriesC_list; last first. - { rewrite !NoDup_cons; repeat split; last apply NoDup_nil; set_solver. } - simpl. lra. - - iModIntro. iIntros (ε n) "?". - destruct n as [|b [|]]. - + iLeft. iPureIntro. - by rewrite /ε2'. - + iMod ("Hvs" $! ε b with "[$]") as "[% | ?]". - * iLeft. iPureIntro. by rewrite /ε2'. - * iRight. by rewrite /ε2'. - + iLeft. iPureIntro. - by rewrite /ε2'. - - repeat case_match; try done. - iModIntro. - iFrame. - Qed. -End test. +(* Section test. *) +(* Context `{F:flip_spec}. *) +(* Lemma flip_presample_spec_simple {L: flipG Σ} NS E ns α *) +(* (ε2 : R -> bool -> R) *) +(* (P : iProp Σ) T γ1 γ2: *) +(* ↑NS ⊆ E -> *) +(* (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> *) +(* (∀ (ε:R), 0<= ε -> (ε2 ε true + ε2 ε false)/2 <= ε)%R-> *) +(* is_flip (L:=L) NS γ1 γ2 -∗ *) +(* (□∀ (ε:R) n, (P ∗ flip_error_auth (L:=L) γ1 ε) ={E∖↑NS}=∗ *) +(* (⌜(1<=ε2 ε n)%R⌝ ∨ (flip_error_auth (L:=L) γ1 (ε2 ε n) ∗ T (n)))) *) +(* -∗ *) +(* P -∗ flip_tapes_frag (L:=L) γ2 α ns-∗ *) +(* wp_update E (∃ n, T (n) ∗ flip_tapes_frag (L:=L) γ2 α (ns++[n])). *) +(* Proof. *) +(* iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) +(* pose (ε2' ε l:= match l with *) +(* | [b]=> ε2 ε b *) +(* | _ => 1%R *) +(* end *) +(* ). *) +(* iMod (flip_presample_spec _ _ _ _ ε2' P 1%nat (λ l, match l with | [b] => T b | _ => False end )%I with "[//][][$][$]") as "(%l & [??])"; first done. *) +(* - rewrite /ε2'. *) +(* intros. repeat case_match; try done. naive_solver. *) +(* - intros. *) +(* etrans; last apply Hineq; try done. *) +(* erewrite (SeriesC_subset (λ x, x ∈ [[true]; [false]])); last first. *) +(* + intros ? H'. *) +(* case_match eqn:K; last done. *) +(* rewrite Nat.eqb_eq in K. *) +(* exfalso. *) +(* apply H'. *) +(* destruct a as [|[|] [|]]; try (simpl in *; done). *) +(* all: set_solver. *) +(* + rewrite SeriesC_list; last first. *) +(* { rewrite !NoDup_cons; repeat split; last apply NoDup_nil; set_solver. } *) +(* simpl. lra. *) +(* - iModIntro. iIntros (ε n) "?". *) +(* destruct n as [|b [|]]. *) +(* + iLeft. iPureIntro. *) +(* by rewrite /ε2'. *) +(* + iMod ("Hvs" $! ε b with "[$]") as "[% | ?]". *) +(* * iLeft. iPureIntro. by rewrite /ε2'. *) +(* * iRight. by rewrite /ε2'. *) +(* + iLeft. iPureIntro. *) +(* by rewrite /ε2'. *) +(* - repeat case_match; try done. *) +(* iModIntro. *) +(* iFrame. *) +(* Qed. *) +(* End test. *) -(** Instantiate flip *) -Class flipG1 Σ := FlipG1 { flipG1_error::hocap_errorGS Σ; - flipG1_tapes:: hocap_tapesGS Σ; - }. -Local Definition flip_inv_pred1 `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapesGS Σ} γ1 γ2:= - (∃ (ε:R) (m:gmap loc (list bool)) , - ↯ ε ∗ ●↯ ε @ γ1 ∗ - ([∗ map] α ↦ t ∈ ((λ (ns:list bool), (1%nat, bool_to_nat<$>ns))<$>m), α ↪N ( t.1 ; t.2 )) ∗ - ●(((λ (ns:list bool), (1, bool_to_nat<$>ns))<$>m):gmap _ _)@γ2 - )%I. +(* (** Instantiate flip *) *) +(* Class flipG1 Σ := FlipG1 { flipG1_error::hocap_errorGS Σ; *) +(* flipG1_tapes:: hocap_tapesGS Σ; *) +(* }. *) +(* Local Definition flip_inv_pred1 `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapesGS Σ} γ1 γ2:= *) +(* (∃ (ε:R) (m:gmap loc (list bool)) , *) +(* ↯ ε ∗ ●↯ ε @ γ1 ∗ *) +(* ([∗ map] α ↦ t ∈ ((λ (ns:list bool), (1%nat, bool_to_nat<$>ns))<$>m), α ↪N ( t.1 ; t.2 )) ∗ *) +(* ●(((λ (ns:list bool), (1, bool_to_nat<$>ns))<$>m):gmap _ _)@γ2 *) +(* )%I. *) -#[local] Program Definition flip_spec1 `{!conerisGS Σ}: flip_spec := - {| flip_allocate_tape:= (λ: <>, allocB); - flip_tape:= flipL; - flipG := flipG1; - flip_error_name := gname; - flip_tape_name := gname; - is_flip _ N γ1 γ2 := inv N (flip_inv_pred1 γ1 γ2); - flip_error_auth _ γ x := ●↯ x @ γ; - flip_error_frag _ γ x := ◯↯ x @ γ; - flip_tapes_auth _ γ m := (●((λ ns, (1, bool_to_nat<$>ns))<$>m)@γ)%I; - flip_tapes_frag _ γ α ns := (α◯↪N (1%nat;bool_to_nat<$> ns) @ γ)%I; - |}. +(* #[local] Program Definition flip_spec1 `{!conerisGS Σ}: flip_spec := *) +(* {| flip_allocate_tape:= (λ: <>, allocB); *) +(* flip_tape:= flipL; *) +(* flipG := flipG1; *) +(* flip_error_name := gname; *) +(* flip_tape_name := gname; *) +(* is_flip _ N γ1 γ2 := inv N (flip_inv_pred1 γ1 γ2); *) +(* flip_error_auth _ γ x := ●↯ x @ γ; *) +(* flip_error_frag _ γ x := ◯↯ x @ γ; *) +(* flip_tapes_auth _ γ m := (●((λ ns, (1, bool_to_nat<$>ns))<$>m)@γ)%I; *) +(* flip_tapes_frag _ γ α ns := (α◯↪N (1%nat;bool_to_nat<$> ns) @ γ)%I; *) +(* |}. *) -Next Obligation. - simpl. - iIntros. - iApply (hocap_error_auth_exclusive with "[$][$]"). -Qed. -Next Obligation. - simpl. - iIntros. - iApply (hocap_error_frag_split). -Qed. -Next Obligation. - simpl. - iIntros (?????) "H". - iApply (hocap_error_auth_valid with "[$]"). -Qed. -Next Obligation. - simpl. - iIntros (?????) "H". - iApply (hocap_error_frag_valid with "[$]"). -Qed. -Next Obligation. - simpl. - iIntros (??????) "H1 H2". - iApply (hocap_error_ineq with "[$][$]"). -Qed. -Next Obligation. - iIntros. - iApply (hocap_error_decrease with "[$][$]"). -Qed. -Next Obligation. - iIntros. - by iApply (hocap_error_increase with "[$]"). -Qed. -Next Obligation. - simpl. - iIntros (??????) "H1 H2". - by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[%H _]". -Qed. -Next Obligation. - simpl. - iIntros (???????) "H1 H2". - iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. - rewrite dfrac_op_own dfrac_valid_own. by intro. -Qed. -Next Obligation. - simpl. - iIntros. - iDestruct (hocap_tapes_agree with "[$][$]") as "%H". - iPureIntro. - rewrite lookup_fmap_Some in H. destruct H as (?&?&K). - simplify_eq. - rewrite K. - f_equal. - eapply fmap_inj; last done. - intros [][]?; by simplify_eq. -Qed. -Next Obligation. - iIntros. - iMod (hocap_tapes_presample with "[$][$]") as "[??]". - iModIntro. iFrame. - rewrite fmap_insert. - rewrite fmap_app; iFrame. -Qed. -Next Obligation. - iIntros. - iMod (hocap_tapes_pop with "[$][$]") as "[??]". - iModIntro. iFrame. - rewrite fmap_insert. - iFrame. -Qed. -Next Obligation. - simpl. - iIntros (????? ε) "H1". - iApply fupd_wp_update. - iApply wp_update_ret. - iDestruct (ec_valid with "[$]") as "%". - unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "(%γ1 & H2 & H3)"; simpl; [naive_solver..|]. - simpl. - iMod (hocap_tapes_alloc ∅) as "(%γ2 & H4 & H5)". - iMod (inv_alloc _ _ (flip_inv_pred1 γ1 γ2) with "[H1 H2 H4]"). - { iFrame. iNext. iExists ∅. - iFrame. - by rewrite fmap_empty. - } - by iFrame. -Qed. -Next Obligation. - simpl. - iIntros (???????? Φ) "#Hinv HΦ". - wp_pures. - rewrite /allocB. - iInv "Hinv" as "(%&%&?&?&?&?)" "Hclose". - wp_apply (wp_alloc_tape); first done. - iIntros (α) "Hα". - iDestruct (hocap_tapes_notin with "[$][$]") as "%". - iMod (hocap_tapes_new _ _ α 1%nat [] with "[$]") as "[?H]"; first done. - iMod ("Hclose" with "[-H HΦ]"). - { iModIntro. - iExists _, (<[α:=([])]> m). - iFrame. - rewrite fmap_insert. - rewrite big_sepM_insert; [iFrame|done]. - } - iApply "HΦ". - by iFrame. -Qed. -Next Obligation. - simpl. - iIntros (????????????? Φ) "(#Hinv & #Hvs & HP) HΦ". - rewrite /flipL. - wp_pures. - wp_bind (rand(_) _)%E. - iInv "Hinv" as ">(%&%&?&?&H3&?)" "Hclose". - iMod ("Hvs" with "[$]") as "(%&HQ&Hauth)". - erewrite <-(insert_delete m) at 1; last done. - rewrite fmap_insert. - rewrite big_sepM_insert; last first. - { rewrite fmap_delete. apply lookup_delete. } - iDestruct "H3" as "[Htape H3]". - simpl. - wp_apply (wp_rand_tape with "[$]") as "[Htape %]". - iMod ("Hclose" with "[- HΦ HQ]") as "_". - { iExists _, (<[α:=ns]> m). - iFrame. - rewrite <-(insert_delete m) at 2; last done. - rewrite insert_insert fmap_insert. - rewrite big_sepM_insert; first iFrame. - rewrite fmap_delete. - apply lookup_delete. - } - iModIntro. - wp_apply conversion.wp_int_to_bool as "_"; first done. - replace (Z_to_bool _) with n. - - iApply "HΦ"; iFrame. - - destruct n; simpl. - + rewrite Z_to_bool_neq_0; [done|lia]. - + by rewrite Z_to_bool_eq_0. -Qed. -Next Obligation. - simpl. - iIntros (????????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". - iApply wp_update_state_step_coupl. - iIntros (σ ε) "((Hheap&Htapes)&Hε)". - iMod (inv_acc with "Hinv") as "[>(% & % & H1 & H2 & H3 & H4 ) Hclose]"; [done|]. - iDestruct (hocap_tapes_agree with "[$][$]") as "%K". - rewrite lookup_fmap_Some in K. destruct K as (?&M&?). - simplify_eq. - unshelve epose proof fmap_inj _ _ _ _ M as ->. - { intros [][]?; by simplify_eq. } - erewrite <-(insert_delete m) at 1; last done. - rewrite fmap_insert. - rewrite big_sepM_insert; last first. - { rewrite fmap_delete. apply lookup_delete. } - simpl. - iDestruct "H3" as "[Htape H3]". - iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". - iDestruct (ec_supply_bound with "[$][$]") as "%". - iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. - iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". - assert (∀ (x: list (fin (2))), bool_to_nat <$> (nat_to_bool <$> (fin_to_nat <$> x)) = fin_to_nat <$> x) as Hrewrite. - { intros l. - rewrite -list_fmap_compose. - rewrite <-list_fmap_id. - apply list_fmap_ext. - intros ?? K. simpl. rewrite nat_to_bool_to_nat; first done. - apply list_lookup_fmap_inv in K as (x'&->&?). - pose proof fin_to_nat_lt x'. lia. } - iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. - unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 ε1' (nat_to_bool <$> (fin_to_nat <$> x)) else 1) _). - { case_match; last (simpl;lra). apply Hpos; first apply cond_nonneg. - rewrite !fmap_length. - by apply Nat.eqb_eq. - } - iSplit. - { iPureIntro. - simpl. - unshelve epose proof (Hineq ε1' _) as H'; first apply cond_nonneg. - etrans; last exact. - rewrite (Rdiv_def _ (2^_)) -SeriesC_scal_r. - etrans; last eapply (SeriesC_le_inj _ (λ x, Some (nat_to_bool <$> (fin_to_nat <$> x)))). - - apply SeriesC_le. - + simpl. intros n. - rewrite !fmap_length. - case_match. - * replace (1+1)%R with 2%R by lra. - rewrite -Rdiv_1_l. simpl. - split; last lra. - apply Rmult_le_pos. - -- apply Rcomplements.Rdiv_le_0_compat; first lra. - apply pow_lt; lra. - -- apply Hpos; first done. - rewrite !fmap_length. - by apply Nat.eqb_eq. - * lra. - + simpl. - apply (ex_seriesC_list_length _ num). - intros ?. rewrite !fmap_length. - case_match; last lra. - intros. by apply Nat.eqb_eq. - - intros. case_match; last lra. - apply Rmult_le_pos; first apply Hpos; simpl; auto. - + by apply Nat.eqb_eq. - + rewrite -Rdiv_1_l. - apply Rcomplements.Rdiv_le_0_compat; first lra. - apply pow_lt; lra. - - intros ??? <- K. - simplify_eq. - rewrite -!list_fmap_compose in K. - apply list_fmap_eq_inj in K; try done. - intros x x'. - repeat (inv_fin x; try intros x); - repeat (inv_fin x'; try intros x'); done. - - apply (ex_seriesC_list_length _ num). - intros ?. - case_match; last lra. - intros. by apply Nat.eqb_eq. - } - iIntros (sample) "<-". - destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' (nat_to_bool <$> (fin_to_nat <$> sample))))%R 1%R) as [Hdec|Hdec]; last first. - { apply Rnot_lt_ge, Rge_le in Hdec. - iApply state_step_coupl_ret_err_ge_1. - simpl. simpl in *. rewrite Nat.eqb_refl. lra. - } - iApply state_step_coupl_ret. - unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' (nat_to_bool <$> (fin_to_nat <$> sample))) _) with "Hε_supply") as "[Hε_supply Hε]". - { apply Hpos; first apply cond_nonneg. by rewrite !fmap_length. } - { simpl. done. } - simpl. - iMod (tapeN_update_append' _ _ _ _ sample with "[$][Htape]") as "[Htapes Htape]". - { by erewrite Heq. } - iMod (hocap_tapes_presample' _ _ _ _ _ (fin_to_nat <$> sample) with "[$][$]") as "[H4 Hfrag]". - iMod "Hclose'" as "_". - iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". - { iExFalso. iApply (ec_contradict with "[$]"). exact. } - iMod ("Hclose" with "[$Hε $H2 Htape H3 H4]") as "_". - { iNext. - iExists (<[α:=(ns ++ (nat_to_bool<$>(fin_to_nat <$> sample)))]>m). - rewrite fmap_insert. - rewrite big_sepM_insert_delete Heq/=. - rewrite fmap_delete. iFrame. - rewrite fmap_app/=. - rewrite !Hrewrite. iFrame. - } - iApply fupd_mask_intro_subseteq; first set_solver. - iFrame. rewrite fmap_app/=Hrewrite. iFrame. - erewrite (nnreal_ext _ _); first done. - simpl. by rewrite Nat.eqb_refl. -Qed. +(* Next Obligation. *) +(* simpl. *) +(* iIntros. *) +(* iApply (hocap_error_auth_exclusive with "[$][$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros. *) +(* iApply (hocap_error_frag_split). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (?????) "H". *) +(* iApply (hocap_error_auth_valid with "[$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (?????) "H". *) +(* iApply (hocap_error_frag_valid with "[$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (??????) "H1 H2". *) +(* iApply (hocap_error_ineq with "[$][$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* iIntros. *) +(* iApply (hocap_error_decrease with "[$][$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* iIntros. *) +(* by iApply (hocap_error_increase with "[$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (??????) "H1 H2". *) +(* by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[%H _]". *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (???????) "H1 H2". *) +(* iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. *) +(* rewrite dfrac_op_own dfrac_valid_own. by intro. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros. *) +(* iDestruct (hocap_tapes_agree with "[$][$]") as "%H". *) +(* iPureIntro. *) +(* rewrite lookup_fmap_Some in H. destruct H as (?&?&K). *) +(* simplify_eq. *) +(* rewrite K. *) +(* f_equal. *) +(* eapply fmap_inj; last done. *) +(* intros [][]?; by simplify_eq. *) +(* Qed. *) +(* Next Obligation. *) +(* iIntros. *) +(* iMod (hocap_tapes_presample with "[$][$]") as "[??]". *) +(* iModIntro. iFrame. *) +(* rewrite fmap_insert. *) +(* rewrite fmap_app; iFrame. *) +(* Qed. *) +(* Next Obligation. *) +(* iIntros. *) +(* iMod (hocap_tapes_pop with "[$][$]") as "[??]". *) +(* iModIntro. iFrame. *) +(* rewrite fmap_insert. *) +(* iFrame. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (????? ε) "H1". *) +(* iApply fupd_wp_update. *) +(* iApply wp_update_ret. *) +(* iDestruct (ec_valid with "[$]") as "%". *) +(* unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "(%γ1 & H2 & H3)"; simpl; [naive_solver..|]. *) +(* simpl. *) +(* iMod (hocap_tapes_alloc ∅) as "(%γ2 & H4 & H5)". *) +(* iMod (inv_alloc _ _ (flip_inv_pred1 γ1 γ2) with "[H1 H2 H4]"). *) +(* { iFrame. iNext. iExists ∅. *) +(* iFrame. *) +(* by rewrite fmap_empty. *) +(* } *) +(* by iFrame. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (???????? Φ) "#Hinv HΦ". *) +(* wp_pures. *) +(* rewrite /allocB. *) +(* iInv "Hinv" as "(%&%&?&?&?&?)" "Hclose". *) +(* wp_apply (wp_alloc_tape); first done. *) +(* iIntros (α) "Hα". *) +(* iDestruct (hocap_tapes_notin with "[$][$]") as "%". *) +(* iMod (hocap_tapes_new _ _ α 1%nat [] with "[$]") as "[?H]"; first done. *) +(* iMod ("Hclose" with "[-H HΦ]"). *) +(* { iModIntro. *) +(* iExists _, (<[α:=([])]> m). *) +(* iFrame. *) +(* rewrite fmap_insert. *) +(* rewrite big_sepM_insert; [iFrame|done]. *) +(* } *) +(* iApply "HΦ". *) +(* by iFrame. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (????????????? Φ) "(#Hinv & #Hvs & HP) HΦ". *) +(* rewrite /flipL. *) +(* wp_pures. *) +(* wp_bind (rand(_) _)%E. *) +(* iInv "Hinv" as ">(%&%&?&?&H3&?)" "Hclose". *) +(* iMod ("Hvs" with "[$]") as "(%&HQ&Hauth)". *) +(* erewrite <-(insert_delete m) at 1; last done. *) +(* rewrite fmap_insert. *) +(* rewrite big_sepM_insert; last first. *) +(* { rewrite fmap_delete. apply lookup_delete. } *) +(* iDestruct "H3" as "[Htape H3]". *) +(* simpl. *) +(* wp_apply (wp_rand_tape with "[$]") as "[Htape %]". *) +(* iMod ("Hclose" with "[- HΦ HQ]") as "_". *) +(* { iExists _, (<[α:=ns]> m). *) +(* iFrame. *) +(* rewrite <-(insert_delete m) at 2; last done. *) +(* rewrite insert_insert fmap_insert. *) +(* rewrite big_sepM_insert; first iFrame. *) +(* rewrite fmap_delete. *) +(* apply lookup_delete. *) +(* } *) +(* iModIntro. *) +(* wp_apply conversion.wp_int_to_bool as "_"; first done. *) +(* replace (Z_to_bool _) with n. *) +(* - iApply "HΦ"; iFrame. *) +(* - destruct n; simpl. *) +(* + rewrite Z_to_bool_neq_0; [done|lia]. *) +(* + by rewrite Z_to_bool_eq_0. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (????????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) +(* iApply wp_update_state_step_coupl. *) +(* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) +(* iMod (inv_acc with "Hinv") as "[>(% & % & H1 & H2 & H3 & H4 ) Hclose]"; [done|]. *) +(* iDestruct (hocap_tapes_agree with "[$][$]") as "%K". *) +(* rewrite lookup_fmap_Some in K. destruct K as (?&M&?). *) +(* simplify_eq. *) +(* unshelve epose proof fmap_inj _ _ _ _ M as ->. *) +(* { intros [][]?; by simplify_eq. } *) +(* erewrite <-(insert_delete m) at 1; last done. *) +(* rewrite fmap_insert. *) +(* rewrite big_sepM_insert; last first. *) +(* { rewrite fmap_delete. apply lookup_delete. } *) +(* simpl. *) +(* iDestruct "H3" as "[Htape H3]". *) +(* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". *) +(* iDestruct (ec_supply_bound with "[$][$]") as "%". *) +(* iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. *) +(* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) +(* assert (∀ (x: list (fin (2))), bool_to_nat <$> (nat_to_bool <$> (fin_to_nat <$> x)) = fin_to_nat <$> x) as Hrewrite. *) +(* { intros l. *) +(* rewrite -list_fmap_compose. *) +(* rewrite <-list_fmap_id. *) +(* apply list_fmap_ext. *) +(* intros ?? K. simpl. rewrite nat_to_bool_to_nat; first done. *) +(* apply list_lookup_fmap_inv in K as (x'&->&?). *) +(* pose proof fin_to_nat_lt x'. lia. } *) +(* iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. *) +(* unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 ε1' (nat_to_bool <$> (fin_to_nat <$> x)) else 1) _). *) +(* { case_match; last (simpl;lra). apply Hpos; first apply cond_nonneg. *) +(* rewrite !fmap_length. *) +(* by apply Nat.eqb_eq. *) +(* } *) +(* iSplit. *) +(* { iPureIntro. *) +(* simpl. *) +(* unshelve epose proof (Hineq ε1' _) as H'; first apply cond_nonneg. *) +(* etrans; last exact. *) +(* rewrite (Rdiv_def _ (2^_)) -SeriesC_scal_r. *) +(* etrans; last eapply (SeriesC_le_inj _ (λ x, Some (nat_to_bool <$> (fin_to_nat <$> x)))). *) +(* - apply SeriesC_le. *) +(* + simpl. intros n. *) +(* rewrite !fmap_length. *) +(* case_match. *) +(* * replace (1+1)%R with 2%R by lra. *) +(* rewrite -Rdiv_1_l. simpl. *) +(* split; last lra. *) +(* apply Rmult_le_pos. *) +(* -- apply Rcomplements.Rdiv_le_0_compat; first lra. *) +(* apply pow_lt; lra. *) +(* -- apply Hpos; first done. *) +(* rewrite !fmap_length. *) +(* by apply Nat.eqb_eq. *) +(* * lra. *) +(* + simpl. *) +(* apply (ex_seriesC_list_length _ num). *) +(* intros ?. rewrite !fmap_length. *) +(* case_match; last lra. *) +(* intros. by apply Nat.eqb_eq. *) +(* - intros. case_match; last lra. *) +(* apply Rmult_le_pos; first apply Hpos; simpl; auto. *) +(* + by apply Nat.eqb_eq. *) +(* + rewrite -Rdiv_1_l. *) +(* apply Rcomplements.Rdiv_le_0_compat; first lra. *) +(* apply pow_lt; lra. *) +(* - intros ??? <- K. *) +(* simplify_eq. *) +(* rewrite -!list_fmap_compose in K. *) +(* apply list_fmap_eq_inj in K; try done. *) +(* intros x x'. *) +(* repeat (inv_fin x; try intros x); *) +(* repeat (inv_fin x'; try intros x'); done. *) +(* - apply (ex_seriesC_list_length _ num). *) +(* intros ?. *) +(* case_match; last lra. *) +(* intros. by apply Nat.eqb_eq. *) +(* } *) +(* iIntros (sample) "<-". *) +(* destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' (nat_to_bool <$> (fin_to_nat <$> sample))))%R 1%R) as [Hdec|Hdec]; last first. *) +(* { apply Rnot_lt_ge, Rge_le in Hdec. *) +(* iApply state_step_coupl_ret_err_ge_1. *) +(* simpl. simpl in *. rewrite Nat.eqb_refl. lra. *) +(* } *) +(* iApply state_step_coupl_ret. *) +(* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' (nat_to_bool <$> (fin_to_nat <$> sample))) _) with "Hε_supply") as "[Hε_supply Hε]". *) +(* { apply Hpos; first apply cond_nonneg. by rewrite !fmap_length. } *) +(* { simpl. done. } *) +(* simpl. *) +(* iMod (tapeN_update_append' _ _ _ _ sample with "[$][Htape]") as "[Htapes Htape]". *) +(* { by erewrite Heq. } *) +(* iMod (hocap_tapes_presample' _ _ _ _ _ (fin_to_nat <$> sample) with "[$][$]") as "[H4 Hfrag]". *) +(* iMod "Hclose'" as "_". *) +(* iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". *) +(* { iExFalso. iApply (ec_contradict with "[$]"). exact. } *) +(* iMod ("Hclose" with "[$Hε $H2 Htape H3 H4]") as "_". *) +(* { iNext. *) +(* iExists (<[α:=(ns ++ (nat_to_bool<$>(fin_to_nat <$> sample)))]>m). *) +(* rewrite fmap_insert. *) +(* rewrite big_sepM_insert_delete Heq/=. *) +(* rewrite fmap_delete. iFrame. *) +(* rewrite fmap_app/=. *) +(* rewrite !Hrewrite. iFrame. *) +(* } *) +(* iApply fupd_mask_intro_subseteq; first set_solver. *) +(* iFrame. rewrite fmap_app/=Hrewrite. iFrame. *) +(* erewrite (nnreal_ext _ _); first done. *) +(* simpl. by rewrite Nat.eqb_refl. *) +(* Qed. *) diff --git a/theories/coneris/lib/hocap.v b/theories/coneris/lib/hocap.v index 33da79ec..971b7bb7 100644 --- a/theories/coneris/lib/hocap.v +++ b/theories/coneris/lib/hocap.v @@ -131,210 +131,210 @@ Section error_lemmas. End error_lemmas. -Definition hocap_tapes_nroot:=nroot.@"tapes". -Class hocap_tapesGS (Σ : gFunctors) := Hocap_tapesGS { - hocap_tapesGS_inG :: ghost_mapG Σ loc (nat*list nat) - }. -Definition hocap_tapesΣ := ghost_mapΣ loc (nat*list nat). - -Notation "α ◯↪N ( M ; ns ) @ γ":= (α ↪[ γ ] (M,ns))%I - (at level 20, format "α ◯↪N ( M ; ns ) @ γ") : bi_scope. - -Notation "● m @ γ" := (ghost_map_auth γ 1 m) (at level 20) : bi_scope. - -Section tapes_lemmas. - Context `{!conerisGS Σ, !hocap_tapesGS Σ}. - - Lemma hocap_tapes_alloc m: - ⊢ |==>∃ γ, (● m @ γ) ∗ [∗ map] k↦v ∈ m, (k ◯↪N (v.1; v.2) @ γ). - Proof. - iMod ghost_map_alloc as (γ) "[??]". - iFrame. iModIntro. - iApply big_sepM_mono; last done. - by iIntros (?[??]). - Qed. - - Lemma hocap_tapes_agree m γ k N ns: - (● m @ γ) -∗ (k ◯↪N (N; ns) @ γ) -∗ ⌜ m!!k = Some (N, ns) ⌝. - Proof. - iIntros "H1 H2". - by iCombine "H1 H2" gives "%". - Qed. - - Lemma hocap_tapes_new γ m k N ns : - m!!k=None -> ⊢ (● m @ γ) ==∗ (● (<[k:=(N,ns)]>m) @ γ) ∗ (k ◯↪N (N; ns) @ γ). - Proof. - iIntros (Hlookup) "H". - by iApply ghost_map_insert. - Qed. - - Lemma hocap_tapes_presample γ m k N ns n: - (● m @ γ) -∗ (k ◯↪N (N; ns) @ γ) ==∗ (● (<[k:=(N,ns++[n])]>m) @ γ) ∗ (k ◯↪N (N; ns++[n]) @ γ). - Proof. - iIntros "H1 H2". - iApply (ghost_map_update with "[$][$]"). - Qed. - - Lemma hocap_tapes_presample' γ m k N ns ns': - (● m @ γ) -∗ (k ◯↪N (N; ns) @ γ) ==∗ (● (<[k:=(N,ns++ns')]>m) @ γ) ∗ (k ◯↪N (N; ns++ns') @ γ). - Proof. - iIntros "H1 H2". - iApply (ghost_map_update with "[$][$]"). - Qed. - - Lemma hocap_tapes_pop γ m k N ns n: - (● m @ γ) -∗ (k ◯↪N (N; n::ns) @ γ) ==∗ (● (<[k:=(N,ns)]>m) @ γ) ∗ (k ◯↪N (N; ns) @ γ). - Proof. - iIntros "H1 H2". - iApply (ghost_map_update with "[$][$]"). - Qed. - - Lemma hocap_tapes_notin α N ns m (f:(nat*list nat)-> nat) g: - α ↪N (N; ns) -∗ ([∗ map] α0↦t ∈ m, α0 ↪N (f t; g t)) -∗ ⌜m!!α=None ⌝. - Proof. - destruct (m!!α) eqn:Heqn; last by iIntros. - iIntros "Hα Hmap". - iDestruct (big_sepM_lookup with "[$]") as "?"; first done. - iExFalso. - iApply (tapeN_tapeN_contradict with "[$][$]"). - Qed. - - (** * TODO add*) -End tapes_lemmas. - -Section HOCAP. - - Context `{!conerisGS Σ, !hocap_errorGS Σ}. - - Definition error_inv (γ :gname):= - inv hocap_error_nroot (∃ (ε:R), ↯ ε ∗ ●↯ ε @ γ). - - Lemma wp_hocap_rand_adv_comp (N : nat) z E - (ε2 : R -> fin (S N) -> R) - (P : iProp Σ) (Q : (fin (S N)) -> iProp Σ) γ: - TCEq N (Z.to_nat z) → - ↑hocap_error_nroot ⊆ E -> - (∀ ε n, (0<=ε -> 0<=ε2 ε n))%R -> - (∀ (ε:R), 0<=ε ->SeriesC (λ n, (1 / (S N)) * (ε2 ε n))%R <= (ε))%R → - {{{ error_inv γ∗ - □(∀ (ε:R) (n : fin (S N)), P ∗ ●↯ ε @ γ ={E∖↑hocap_error_nroot}=∗ - (⌜(1<=ε2 ε n)%R⌝ ∨(●↯ (ε2 ε n) @ γ ∗ Q (n))) ) ∗ - P }}} rand #z @ E {{{ n, RET #n; Q (n)}}}. - Proof. - iIntros (-> Hsubset ? Hineq) "%Φ [#Hinv [#Hchange HP]] HΦ". - iInv "Hinv" as ">(%ε & Hε & Hauth)" "Hclose". - iDestruct (ec_valid with "[$]") as "%". - wp_apply (wp_couple_rand_adv_comp1' _ _ _ _ (ε2 ε) with "[$]"); [naive_solver|naive_solver|]. - iIntros (n) "Hε". - iMod ("Hchange" $! _ n with "[$]") as "[%|[Hauth HQ]]"; last first. - { iMod ("Hclose" with "[Hauth Hε]") as "_". - - unshelve iExists (mknonnegreal _ _); last iFrame. - naive_solver. - - by iApply "HΦ". } - iExFalso. - by iApply (ec_contradict with "[$]"). - Qed. - - Lemma wp_hocap_flip_adv_comp E - (ε2 : R -> bool -> R) - (P : iProp Σ) (Q : bool -> iProp Σ) γ: - ↑hocap_error_nroot ⊆ E -> - (∀ ε b, 0<=ε -> 0<=ε2 ε b)%R -> - (∀ (ε:R), 0<=ε -> (((ε2 ε true) + (ε2 ε false))/2 <= (ε)))%R → - {{{ error_inv γ∗ - □(∀ (ε:R) (b : bool), P ∗ ●↯ ε @ γ ={E∖↑hocap_error_nroot}=∗ (⌜(1<=ε2 ε b)%R⌝ ∨ ●↯ (ε2 ε b) @ γ ∗ Q (b)) ) ∗ - P }}} flip @ E {{{ (b:bool), RET #b; Q (b)}}}. - Proof. - iIntros (Hsubset ? Hineq) "%Φ [#Hinv [#Hchange HP]] HΦ". - rewrite /flip/flipL. - wp_pures. - wp_apply (wp_hocap_rand_adv_comp _ _ _ (λ ε x, if fin_to_nat x =? 1%nat then ε2 ε true else ε2 ε false) P (λ x, Q (fin_to_nat x =? 1%nat)) with "[-HΦ]"); [done|..]. - - intros; case_match; naive_solver. - - intros ε. rewrite SeriesC_finite_foldr; simpl. specialize (Hineq ε). lra. - - iFrame. iSplitR; first iExact "Hinv". - iModIntro. - iIntros (ε n) "H". - iMod ("Hchange" $! _ (fin_to_nat n =? 1%nat)with "[$]") as "[%|[Hε HQ]]". - + iModIntro. iLeft. iPureIntro. by case_match. - + iModIntro. iRight. iFrame. by case_match. - - iIntros. - wp_apply (conversion.wp_int_to_bool); first done. - iIntros "_". - iApply "HΦ". - inv_fin n; simpl; first done. - intros n; inv_fin n; simpl; first done. - intros n; inv_fin n. - Qed. - - (* (** With tapes *) *) - (* Context `{!hocap_tapesGS Σ}. *) - - (* Definition tapes_inv (γ :gname):= *) - (* inv hocap_tapes_nroot (∃ m, ●m@γ ∗ [∗ map] α ↦ t ∈ m, α ↪N ( t.1 ; t.2 ) ). *) - (* Lemma wp_hocap_presample_adv_comp (N : nat) z E ns α *) - (* (ε2 : R -> fin (S N) -> R) *) - (* (P : iProp Σ) (Q : val-> iProp Σ) T γ γ': *) - (* TCEq N (Z.to_nat z) → *) - (* ↑hocap_error_nroot ⊆ E -> *) - (* ↑hocap_tapes_nroot ⊆ E -> *) - (* (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> *) - (* (∀ (ε:R), 0<= ε -> SeriesC (λ n, (1 / (S N)) * (ε2 ε n))%R <= ε)%R → *) - (* error_inv γ -∗ tapes_inv γ' -∗ *) - (* (□∀ (ε:R) m, (⌜m!!α = Some (N, ns)⌝ ∗ P ∗ ●↯ ε @ γ ∗ ●m@γ') *) - (* ={E∖↑hocap_error_nroot∖↑hocap_tapes_nroot}=∗ *) - (* ∃ n, (⌜(1<=ε2 ε n)%R⌝ ∨(●↯ (ε2 ε n) @ γ ∗ ●(<[α := (N, ns ++ [fin_to_nat n])]>m) @ γ' ∗ T (n)))) *) - (* -∗ P -∗ α ◯↪N (N; ns) @ γ' -∗ *) - (* wp_update E (∃ n, T (n) ∗ α◯↪N (N; ns++[fin_to_nat n]) @ γ'). *) - (* Proof. *) - (* iIntros (-> Hval Hsubset Hubset' Hineq) "#Hinv #Hinv' #Hshift HP Htape". *) - (* iApply fupd_wp_update. *) - (* iInv "Hinv" as ">(%ε' & Hε & Hauth)" "Hclose". *) - (* iInv "Hinv'" as ">(%m & Hm & Hmauth)" "Hclose'". *) - (* iDestruct (ec_valid with "[$]") as "%". *) - (* iApply wp_update_mono_fupd. iApply fupd_frame_l. *) - (* iSplitL "Hε". *) - (* - iApply (wp_update_presample_exp _ α _ ns ε' (ε2 ε) with "[$Hε]"). *) - (* + naive_solver. *) - (* + naive_solver. *) - (* + admit. *) - (* - iMod ("Hclose'" with "[$]"). *) - (* Abort. *) +(* Definition hocap_tapes_nroot:=nroot.@"tapes". *) +(* Class hocap_tapesGS (Σ : gFunctors) := Hocap_tapesGS { *) +(* hocap_tapesGS_inG :: ghost_mapG Σ loc (nat*list nat) *) +(* }. *) +(* Definition hocap_tapesΣ := ghost_mapΣ loc (nat*list nat). *) + +(* Notation "α ◯↪N ( M ; ns ) @ γ":= (α ↪[ γ ] (M,ns))%I *) +(* (at level 20, format "α ◯↪N ( M ; ns ) @ γ") : bi_scope. *) + +(* Notation "● m @ γ" := (ghost_map_auth γ 1 m) (at level 20) : bi_scope. *) + +(* Section tapes_lemmas. *) +(* Context `{!conerisGS Σ, !hocap_tapesGS Σ}. *) + +(* Lemma hocap_tapes_alloc m: *) +(* ⊢ |==>∃ γ, (● m @ γ) ∗ [∗ map] k↦v ∈ m, (k ◯↪N (v.1; v.2) @ γ). *) +(* Proof. *) +(* iMod ghost_map_alloc as (γ) "[??]". *) +(* iFrame. iModIntro. *) +(* iApply big_sepM_mono; last done. *) +(* by iIntros (?[??]). *) +(* Qed. *) + +(* Lemma hocap_tapes_agree m γ k N ns: *) +(* (● m @ γ) -∗ (k ◯↪N (N; ns) @ γ) -∗ ⌜ m!!k = Some (N, ns) ⌝. *) +(* Proof. *) +(* iIntros "H1 H2". *) +(* by iCombine "H1 H2" gives "%". *) +(* Qed. *) + +(* Lemma hocap_tapes_new γ m k N ns : *) +(* m!!k=None -> ⊢ (● m @ γ) ==∗ (● (<[k:=(N,ns)]>m) @ γ) ∗ (k ◯↪N (N; ns) @ γ). *) +(* Proof. *) +(* iIntros (Hlookup) "H". *) +(* by iApply ghost_map_insert. *) +(* Qed. *) + +(* Lemma hocap_tapes_presample γ m k N ns n: *) +(* (● m @ γ) -∗ (k ◯↪N (N; ns) @ γ) ==∗ (● (<[k:=(N,ns++[n])]>m) @ γ) ∗ (k ◯↪N (N; ns++[n]) @ γ). *) +(* Proof. *) +(* iIntros "H1 H2". *) +(* iApply (ghost_map_update with "[$][$]"). *) +(* Qed. *) + +(* Lemma hocap_tapes_presample' γ m k N ns ns': *) +(* (● m @ γ) -∗ (k ◯↪N (N; ns) @ γ) ==∗ (● (<[k:=(N,ns++ns')]>m) @ γ) ∗ (k ◯↪N (N; ns++ns') @ γ). *) +(* Proof. *) +(* iIntros "H1 H2". *) +(* iApply (ghost_map_update with "[$][$]"). *) +(* Qed. *) + +(* Lemma hocap_tapes_pop γ m k N ns n: *) +(* (● m @ γ) -∗ (k ◯↪N (N; n::ns) @ γ) ==∗ (● (<[k:=(N,ns)]>m) @ γ) ∗ (k ◯↪N (N; ns) @ γ). *) +(* Proof. *) +(* iIntros "H1 H2". *) +(* iApply (ghost_map_update with "[$][$]"). *) +(* Qed. *) + +(* Lemma hocap_tapes_notin α N ns m (f:(nat*list nat)-> nat) g: *) +(* α ↪N (N; ns) -∗ ([∗ map] α0↦t ∈ m, α0 ↪N (f t; g t)) -∗ ⌜m!!α=None ⌝. *) +(* Proof. *) +(* destruct (m!!α) eqn:Heqn; last by iIntros. *) +(* iIntros "Hα Hmap". *) +(* iDestruct (big_sepM_lookup with "[$]") as "?"; first done. *) +(* iExFalso. *) +(* iApply (tapeN_tapeN_contradict with "[$][$]"). *) +(* Qed. *) + +(* (** * TODO add*) *) +(* End tapes_lemmas. *) + +(* Section HOCAP. *) + +(* Context `{!conerisGS Σ, !hocap_errorGS Σ}. *) + +(* Definition error_inv (γ :gname):= *) +(* inv hocap_error_nroot (∃ (ε:R), ↯ ε ∗ ●↯ ε @ γ). *) + +(* Lemma wp_hocap_rand_adv_comp (N : nat) z E *) +(* (ε2 : R -> fin (S N) -> R) *) +(* (P : iProp Σ) (Q : (fin (S N)) -> iProp Σ) γ: *) +(* TCEq N (Z.to_nat z) → *) +(* ↑hocap_error_nroot ⊆ E -> *) +(* (∀ ε n, (0<=ε -> 0<=ε2 ε n))%R -> *) +(* (∀ (ε:R), 0<=ε ->SeriesC (λ n, (1 / (S N)) * (ε2 ε n))%R <= (ε))%R → *) +(* {{{ error_inv γ∗ *) +(* □(∀ (ε:R) (n : fin (S N)), P ∗ ●↯ ε @ γ ={E∖↑hocap_error_nroot}=∗ *) +(* (⌜(1<=ε2 ε n)%R⌝ ∨(●↯ (ε2 ε n) @ γ ∗ Q (n))) ) ∗ *) +(* P }}} rand #z @ E {{{ n, RET #n; Q (n)}}}. *) +(* Proof. *) +(* iIntros (-> Hsubset ? Hineq) "%Φ [#Hinv [#Hchange HP]] HΦ". *) +(* iInv "Hinv" as ">(%ε & Hε & Hauth)" "Hclose". *) +(* iDestruct (ec_valid with "[$]") as "%". *) +(* wp_apply (wp_couple_rand_adv_comp1' _ _ _ _ (ε2 ε) with "[$]"); [naive_solver|naive_solver|]. *) +(* iIntros (n) "Hε". *) +(* iMod ("Hchange" $! _ n with "[$]") as "[%|[Hauth HQ]]"; last first. *) +(* { iMod ("Hclose" with "[Hauth Hε]") as "_". *) +(* - unshelve iExists (mknonnegreal _ _); last iFrame. *) +(* naive_solver. *) +(* - by iApply "HΦ". } *) +(* iExFalso. *) +(* by iApply (ec_contradict with "[$]"). *) +(* Qed. *) + +(* Lemma wp_hocap_flip_adv_comp E *) +(* (ε2 : R -> bool -> R) *) +(* (P : iProp Σ) (Q : bool -> iProp Σ) γ: *) +(* ↑hocap_error_nroot ⊆ E -> *) +(* (∀ ε b, 0<=ε -> 0<=ε2 ε b)%R -> *) +(* (∀ (ε:R), 0<=ε -> (((ε2 ε true) + (ε2 ε false))/2 <= (ε)))%R → *) +(* {{{ error_inv γ∗ *) +(* □(∀ (ε:R) (b : bool), P ∗ ●↯ ε @ γ ={E∖↑hocap_error_nroot}=∗ (⌜(1<=ε2 ε b)%R⌝ ∨ ●↯ (ε2 ε b) @ γ ∗ Q (b)) ) ∗ *) +(* P }}} flip @ E {{{ (b:bool), RET #b; Q (b)}}}. *) +(* Proof. *) +(* iIntros (Hsubset ? Hineq) "%Φ [#Hinv [#Hchange HP]] HΦ". *) +(* rewrite /flip/flipL. *) +(* wp_pures. *) +(* wp_apply (wp_hocap_rand_adv_comp _ _ _ (λ ε x, if fin_to_nat x =? 1%nat then ε2 ε true else ε2 ε false) P (λ x, Q (fin_to_nat x =? 1%nat)) with "[-HΦ]"); [done|..]. *) +(* - intros; case_match; naive_solver. *) +(* - intros ε. rewrite SeriesC_finite_foldr; simpl. specialize (Hineq ε). lra. *) +(* - iFrame. iSplitR; first iExact "Hinv". *) +(* iModIntro. *) +(* iIntros (ε n) "H". *) +(* iMod ("Hchange" $! _ (fin_to_nat n =? 1%nat)with "[$]") as "[%|[Hε HQ]]". *) +(* + iModIntro. iLeft. iPureIntro. by case_match. *) +(* + iModIntro. iRight. iFrame. by case_match. *) +(* - iIntros. *) +(* wp_apply (conversion.wp_int_to_bool); first done. *) +(* iIntros "_". *) +(* iApply "HΦ". *) +(* inv_fin n; simpl; first done. *) +(* intros n; inv_fin n; simpl; first done. *) +(* intros n; inv_fin n. *) +(* Qed. *) + +(* (* (** With tapes *) *) *) +(* (* Context `{!hocap_tapesGS Σ}. *) *) + +(* (* Definition tapes_inv (γ :gname):= *) *) +(* (* inv hocap_tapes_nroot (∃ m, ●m@γ ∗ [∗ map] α ↦ t ∈ m, α ↪N ( t.1 ; t.2 ) ). *) *) +(* (* Lemma wp_hocap_presample_adv_comp (N : nat) z E ns α *) *) +(* (* (ε2 : R -> fin (S N) -> R) *) *) +(* (* (P : iProp Σ) (Q : val-> iProp Σ) T γ γ': *) *) +(* (* TCEq N (Z.to_nat z) → *) *) +(* (* ↑hocap_error_nroot ⊆ E -> *) *) +(* (* ↑hocap_tapes_nroot ⊆ E -> *) *) +(* (* (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> *) *) +(* (* (∀ (ε:R), 0<= ε -> SeriesC (λ n, (1 / (S N)) * (ε2 ε n))%R <= ε)%R → *) *) +(* (* error_inv γ -∗ tapes_inv γ' -∗ *) *) +(* (* (□∀ (ε:R) m, (⌜m!!α = Some (N, ns)⌝ ∗ P ∗ ●↯ ε @ γ ∗ ●m@γ') *) *) +(* (* ={E∖↑hocap_error_nroot∖↑hocap_tapes_nroot}=∗ *) *) +(* (* ∃ n, (⌜(1<=ε2 ε n)%R⌝ ∨(●↯ (ε2 ε n) @ γ ∗ ●(<[α := (N, ns ++ [fin_to_nat n])]>m) @ γ' ∗ T (n)))) *) *) +(* (* -∗ P -∗ α ◯↪N (N; ns) @ γ' -∗ *) *) +(* (* wp_update E (∃ n, T (n) ∗ α◯↪N (N; ns++[fin_to_nat n]) @ γ'). *) *) +(* (* Proof. *) *) +(* (* iIntros (-> Hval Hsubset Hubset' Hineq) "#Hinv #Hinv' #Hshift HP Htape". *) *) +(* (* iApply fupd_wp_update. *) *) +(* (* iInv "Hinv" as ">(%ε' & Hε & Hauth)" "Hclose". *) *) +(* (* iInv "Hinv'" as ">(%m & Hm & Hmauth)" "Hclose'". *) *) +(* (* iDestruct (ec_valid with "[$]") as "%". *) *) +(* (* iApply wp_update_mono_fupd. iApply fupd_frame_l. *) *) +(* (* iSplitL "Hε". *) *) +(* (* - iApply (wp_update_presample_exp _ α _ ns ε' (ε2 ε) with "[$Hε]"). *) *) +(* (* + naive_solver. *) *) +(* (* + naive_solver. *) *) +(* (* + admit. *) *) +(* (* - iMod ("Hclose'" with "[$]"). *) *) +(* (* Abort. *) *) - (* iApply (wp_presample_adv_comp); [done|exact|..]. *) - (* iApply fupd_frame_l. *) - (* - *) - (* Abort. *) - - - -End HOCAP. - -Section HOCAP_alt. - Context `{!conerisGS Σ}. - - Lemma wp_flip_exp_hocap (P : iProp Σ) (Q : bool → iProp Σ) E1 E2 : - □ (P ={E1, E2}=∗ - ∃ ε1 ε2, ⌜(( (ε2 true) + (ε2 false))/2 <= ε1)%R⌝ ∗ - ⌜(∀ b, 0<=ε2 b)%R⌝ ∗ - ↯ ε1 ∗ (∀ (b : bool), ↯ (ε2 b) ={E2, E1}=∗ Q b)) -∗ - {{{ P }}} flip @ E1 {{{ (b : bool), RET #b; Q b}}}. - Proof. - iIntros "#Hvs". iIntros (Ψ) "!# HP HΨ". - rewrite /flip /flipL. - wp_pures. - wp_bind (rand _)%E. - iMod ("Hvs" with "HP") as (ε1 ε2) "(% & % & Hε1 & HQ)". - set (ε2' := ε2 ∘ fin_to_bool). - iApply (wp_couple_rand_adv_comp1' _ _ _ _ ε2' with "Hε1"). - { intros; rewrite /ε2'. simpl. done. } - { rewrite SeriesC_finite_foldr /ε2' /=. lra. } - iIntros "!>" (n) "Hε2". - assert (ε2' n = ε2 (Z_to_bool n)) as ->. - { inv_fin n; [eauto|]. intros n. inv_fin n; [eauto|]. intros n. inv_fin n. } - iMod ("HQ" with "Hε2") as "HQ". iModIntro. - wp_apply (conversion.wp_int_to_bool); [done|]. - iIntros "_". - by iApply "HΨ". - Qed. - -End HOCAP_alt. +(* (* iApply (wp_presample_adv_comp); [done|exact|..]. *) *) +(* (* iApply fupd_frame_l. *) *) +(* (* - *) *) +(* (* Abort. *) *) + + + +(* End HOCAP. *) + +(* Section HOCAP_alt. *) +(* Context `{!conerisGS Σ}. *) + +(* Lemma wp_flip_exp_hocap (P : iProp Σ) (Q : bool → iProp Σ) E1 E2 : *) +(* □ (P ={E1, E2}=∗ *) +(* ∃ ε1 ε2, ⌜(( (ε2 true) + (ε2 false))/2 <= ε1)%R⌝ ∗ *) +(* ⌜(∀ b, 0<=ε2 b)%R⌝ ∗ *) +(* ↯ ε1 ∗ (∀ (b : bool), ↯ (ε2 b) ={E2, E1}=∗ Q b)) -∗ *) +(* {{{ P }}} flip @ E1 {{{ (b : bool), RET #b; Q b}}}. *) +(* Proof. *) +(* iIntros "#Hvs". iIntros (Ψ) "!# HP HΨ". *) +(* rewrite /flip /flipL. *) +(* wp_pures. *) +(* wp_bind (rand _)%E. *) +(* iMod ("Hvs" with "HP") as (ε1 ε2) "(% & % & Hε1 & HQ)". *) +(* set (ε2' := ε2 ∘ fin_to_bool). *) +(* iApply (wp_couple_rand_adv_comp1' _ _ _ _ ε2' with "Hε1"). *) +(* { intros; rewrite /ε2'. simpl. done. } *) +(* { rewrite SeriesC_finite_foldr /ε2' /=. lra. } *) +(* iIntros "!>" (n) "Hε2". *) +(* assert (ε2' n = ε2 (Z_to_bool n)) as ->. *) +(* { inv_fin n; [eauto|]. intros n. inv_fin n; [eauto|]. intros n. inv_fin n. } *) +(* iMod ("HQ" with "Hε2") as "HQ". iModIntro. *) +(* wp_apply (conversion.wp_int_to_bool); [done|]. *) +(* iIntros "_". *) +(* by iApply "HΨ". *) +(* Qed. *) + +(* End HOCAP_alt. *) From 14b179ece99919ed4745cc56f64654472c39210f Mon Sep 17 00:00:00 2001 From: Hei Li Date: Mon, 23 Sep 2024 10:52:17 +0200 Subject: [PATCH 20/69] first attempt with hocap rand --- theories/coneris/lib/hocap.v | 16 +++-- theories/coneris/lib/hocap_rand.v | 107 ++++++++++++++++++++++++++++++ theories/coneris/primitive_laws.v | 11 +++ theories/prob/countable_sum.v | 16 +++-- 4 files changed, 142 insertions(+), 8 deletions(-) create mode 100644 theories/coneris/lib/hocap_rand.v diff --git a/theories/coneris/lib/hocap.v b/theories/coneris/lib/hocap.v index 971b7bb7..7b210966 100644 --- a/theories/coneris/lib/hocap.v +++ b/theories/coneris/lib/hocap.v @@ -1,7 +1,7 @@ (** * Hocap specs *) From stdpp Require Import namespaces. From iris Require Import excl_auth invariants list. -From clutch.coneris Require Import coneris flip. +From clutch.coneris Require Import coneris. Set Default Proof Using "Type*". @@ -94,14 +94,16 @@ Section error_lemmas. Qed. Lemma hocap_error_decrease γ (b' b:nonnegreal) : - (●↯ (b+b') @ γ) -∗ (◯↯ b @ γ) ==∗ (●↯ b' @ γ). + (●↯ (b) @ γ) -∗ (◯↯ b' @ γ) ==∗ (●↯ (b-b') @ γ). Proof. iIntros "H1 H2". simpl. + iDestruct (hocap_error_ineq with "[$][$]") as "%". iDestruct "H1" as "[% [% H1]]". iDestruct "H2" as "[% [% H2]]". iMod (own_update_2 with "H1 H2") as "Hown". - { eapply (auth_update_dealloc _ _ b'), nonnegreal_local_update. + { unshelve eapply (auth_update_dealloc _ _ ((b-b') _)%NNR), nonnegreal_local_update. + - lra. - apply cond_nonneg. - apply nnreal_ext =>/=. lra. } iFrame. by iPureIntro. @@ -123,11 +125,17 @@ Section error_lemmas. - done. Qed. - Lemma hocap_error_irrel γ (b c:R) : + Lemma hocap_error_auth_irrel γ (b c:R) : (b=c)%R -> (●↯ b @ γ) -∗ (●↯ c @ γ). Proof. iIntros (->) "$". Qed. + + Lemma hocap_error_frag_irrel γ (b c:R) : + (b=c)%R -> (◯↯ b @ γ) -∗ (◯↯ c @ γ). + Proof. + iIntros (->) "$". + Qed. End error_lemmas. diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v new file mode 100644 index 00000000..8fcb7c97 --- /dev/null +++ b/theories/coneris/lib/hocap_rand.v @@ -0,0 +1,107 @@ +(** * Hocap rand specs *) +From clutch.coneris Require Import coneris hocap. + +Set Default Proof Using "Type*". + + +Class rand_spec `{!conerisGS Σ} := RandSpec +{ + (** * Operations *) + rand_allocate_tape : val; + rand_tape : val; + (** * Ghost state *) + (** The assumptions about [Σ] *) + randG : gFunctors → Type; + + rand_error_name : Type; + rand_tape_name: Type; + (** * Predicates *) + is_rand {L : randG Σ} (N:namespace) + (γ1: rand_error_name) (γ2: rand_tape_name): iProp Σ; + rand_error_auth {L : randG Σ} (γ: rand_error_name) (x:R): iProp Σ; + rand_error_frag {L : randG Σ} (γ: rand_error_name) (x:R): iProp Σ; + rand_tapes_auth {L : randG Σ} (γ: rand_tape_name) (m:gmap loc (nat * list nat)): iProp Σ; + rand_tapes_frag {L : randG Σ} (γ: rand_tape_name) (α:loc) (ns: (nat * list nat)): iProp Σ; + (** * General properties of the predicates *) + #[global] is_rand_persistent {L : randG Σ} N γ1 γ2 :: + Persistent (is_rand (L:=L) N γ1 γ2); + #[global] rand_error_auth_timeless {L : randG Σ} γ x :: + Timeless (rand_error_auth (L:=L) γ x); + #[global] rand_error_frag_timeless {L : randG Σ} γ x :: + Timeless (rand_error_frag (L:=L) γ x); + #[global] rand_tapes_auth_timeless {L : randG Σ} γ m :: + Timeless (rand_tapes_auth (L:=L) γ m); + #[global] rand_tapes_frag_timeless {L : randG Σ} γ α ns :: + Timeless (rand_tapes_frag (L:=L) γ α ns); + #[global] rand_error_name_inhabited:: + Inhabited rand_error_name; + #[global] rand_tape_name_inhabited:: + Inhabited rand_tape_name; + + rand_error_auth_exclusive {L : randG Σ} γ x1 x2: + rand_error_auth (L:=L) γ x1 -∗ rand_error_auth (L:=L) γ x2 -∗ False; + rand_error_frag_split {L : randG Σ} γ (x1 x2:nonnegreal): + rand_error_frag (L:=L) γ x1 ∗ rand_error_frag (L:=L) γ x2 ⊣⊢ + rand_error_frag (L:=L) γ (x1+x2)%R ; + rand_error_auth_valid {L : randG Σ} γ x: + rand_error_auth (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; + rand_error_frag_valid {L : randG Σ} γ x: + rand_error_frag (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; + rand_error_ineq {L : randG Σ} γ x1 x2: + rand_error_auth (L:=L) γ x1 -∗ rand_error_frag (L:=L) γ x2 -∗ ⌜(x2 <= x1)%R ⌝; + rand_error_decrease {L : randG Σ} γ (x1 x2:nonnegreal): + rand_error_auth (L:=L) γ x1 -∗ rand_error_frag (L:=L) γ x2 ==∗ rand_error_auth (L:=L) γ (x2-x1)%R; + rand_error_increase {L : randG Σ} γ (x1 x2:nonnegreal): + (x1+x2<1)%R -> ⊢ rand_error_auth (L:=L) γ x1 ==∗ + rand_error_auth (L:=L) γ (x1+x2) ∗ rand_error_frag (L:=L) γ x2; + + rand_tapes_auth_exclusive {L : randG Σ} γ m m': + rand_tapes_auth (L:=L) γ m -∗ rand_tapes_auth (L:=L) γ m' -∗ False; + rand_tapes_frag_exclusive {L : randG Σ} γ α ns ns': + rand_tapes_frag (L:=L) γ α ns -∗ rand_tapes_frag (L:=L) γ α ns' -∗ False; + rand_tapes_agree {L : randG Σ} γ α m ns: + rand_tapes_auth (L:=L) γ m -∗ rand_tapes_frag (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; + rand_tapes_valid {L : randG Σ} γ1 γ2 α N E tb ns: + ⌜↑N⊆E⌝ ∗ is_rand (L:=L) N γ1 γ2 ∗ rand_tapes_frag (L:=L) γ2 α (tb, ns) ={E}=∗ + ⌜Forall (λ n, n<=tb)%nat ns⌝; + rand_tapes_update {L : randG Σ} γ α m ns ns': + rand_tapes_auth (L:=L) γ m -∗ rand_tapes_frag (L:=L) γ α ns ==∗ + rand_tapes_auth (L:=L) γ (<[α := ns']> m) ∗ rand_tapes_frag (L:=L) γ α ns'; + + (** * Program specs *) + rand_inv_create_spec {L : randG Σ} N E ε: + ↯ ε -∗ + wp_update E (∃ γ1 γ2, is_rand (L:=L) N γ1 γ2 ∗ + rand_error_frag (L:=L) γ1 ε); + + rand_allocate_tape_spec {L: randG Σ} N E (tb:nat) γ1 γ2: + ↑N ⊆ E-> + {{{ is_rand (L:=L) N γ1 γ2 }}} + rand_allocate_tape #tb @ E + {{{ (v:val), RET v; + ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ rand_tapes_frag (L:=L) γ2 α (tb, []) + }}}; + rand_tape_spec_some {L: randG Σ} N E γ1 γ2 (P: iProp Σ) (Q:iProp Σ) (α:loc) (n:nat) (tb:nat) ns: + ↑N⊆E -> + {{{ is_rand (L:=L) N γ1 γ2 ∗ + □ (∀ (m:gmap loc (nat * list nat)), P ∗ + rand_tapes_auth (L:=L) γ2 m + ={E∖↑N}=∗ ⌜m!!α=Some (tb, n::ns)⌝ ∗ Q ∗ rand_tapes_auth (L:=L) γ2 (<[α:=(tb, ns)]> m)) ∗ + P + }}} + rand_tape #lbl:α #tb @ E + {{{ RET #n; Q }}}; + rand_presample_spec {L: randG Σ} N E ns α (tb:nat) + (ε2 : list (fin (S tb)) -> R) + (P : iProp Σ) num T γ1 γ2 ε : + ↑N ⊆ E -> + (∀ l, length l = num -> 0<= ε2 l)%R -> + (SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R-> + is_rand (L:=L) N γ1 γ2 -∗ + (□∀ (ε':R) ns', (P ∗ rand_error_auth (L:=L) γ1 (ε')%R ∗ ⌜length ns' = num⌝) ={E∖↑N}=∗ + ∃ diff: R, ⌜(0<=diff)%R⌝ ∗ ⌜(ε' = ε+diff)%R⌝ ∗ + (⌜(1<=ε2 ns' + diff)%R⌝ ∨ (rand_error_auth (L:=L) γ1 (ε2 ns' + diff)%R ∗ T ns'))) + -∗ + P -∗ rand_tapes_frag (L:=L) γ2 α (tb, ns)-∗ + wp_update E (∃ n, T n ∗ rand_tapes_frag (L:=L) γ2 α (tb, ns++(fin_to_nat <$> n))) +}. diff --git a/theories/coneris/primitive_laws.v b/theories/coneris/primitive_laws.v index cf715c48..980f143b 100644 --- a/theories/coneris/primitive_laws.v +++ b/theories/coneris/primitive_laws.v @@ -137,6 +137,17 @@ Section tape_interface. by rewrite fmap_app. Qed. + Lemma tapeN_ineq α N ns: + α↪N (N; ns) -∗ ⌜Forall (λ n, n<=N)%nat ns⌝. + Proof. + iIntros "(% & <- & H)". + iPureIntro. + eapply Forall_impl. + - apply fin.fin_forall_leq. + - simpl. intros. + lia. + Qed. + (* Lemma spec_tapeN_to_empty l M : (l ↪ₛN ( M ; [] ) -∗ l ↪ₛ ( M ; [] )). diff --git a/theories/prob/countable_sum.v b/theories/prob/countable_sum.v index 6eaa80f5..7d872fe7 100644 --- a/theories/prob/countable_sum.v +++ b/theories/prob/countable_sum.v @@ -2,7 +2,7 @@ From Coq Require Import Reals Psatz. From Coquelicot Require Import Series Hierarchy Lim_seq Rbar Lub. From stdpp Require Import option. From stdpp Require Export countable finite gmap. -From clutch.prelude Require Import base Reals_ext Coquelicot_ext Series_ext stdpp_ext classical fin. +From clutch.prelude Require Import base Reals_ext Coquelicot_ext Series_ext stdpp_ext classical fin uniform_list. Set Default Proof Using "Type*". Import Hierarchy. @@ -710,12 +710,20 @@ Section filter. End filter. -Lemma ex_seriesC_list_length `{Countable A} (f:list A -> R) num: - (forall x, (0 length x = num) -> +Lemma ex_seriesC_list_length `{Finite A} (f:list A -> R) num: + (forall x, (0≠f x)%R -> length x = num) -> ex_seriesC f. Proof. intros. -Admitted. + eapply (ex_seriesC_ext (λ x, if bool_decide(x∈enum_uniform_list num) then f x else 0))%R. + - intros n. + case_bool_decide as K; first done. + destruct (Req_dec (f n) 0); first done. + exfalso. + apply K. apply elem_of_enum_uniform_list. + naive_solver. + - apply ex_seriesC_list. +Qed. Lemma SeriesC_Series_nat (f : nat → R) : SeriesC f = Series f. From 758b6d500c53dcb849d987d4a80bde8985f4f96f Mon Sep 17 00:00:00 2001 From: Hei Li Date: Mon, 23 Sep 2024 14:25:41 +0200 Subject: [PATCH 21/69] progress with hocap rand --- theories/coneris/lib/hocap.v | 150 ++++++++++++------------- theories/coneris/lib/hocap_rand.v | 177 +++++++++++++++++++++++++++++- theories/coneris/primitive_laws.v | 10 ++ theories/coneris/wp_update.v | 8 ++ 4 files changed, 265 insertions(+), 80 deletions(-) diff --git a/theories/coneris/lib/hocap.v b/theories/coneris/lib/hocap.v index 7b210966..00cb1ec0 100644 --- a/theories/coneris/lib/hocap.v +++ b/theories/coneris/lib/hocap.v @@ -93,36 +93,38 @@ Section error_lemmas. by eapply auth_both_valid_discrete in Hop as [Hlt%nonnegreal_included ?]. Qed. - Lemma hocap_error_decrease γ (b' b:nonnegreal) : + Lemma hocap_error_decrease γ (b' b:R) : (●↯ (b) @ γ) -∗ (◯↯ b' @ γ) ==∗ (●↯ (b-b') @ γ). Proof. iIntros "H1 H2". simpl. iDestruct (hocap_error_ineq with "[$][$]") as "%". - iDestruct "H1" as "[% [% H1]]". - iDestruct "H2" as "[% [% H2]]". + iDestruct "H1" as "[%x [% H1]]". + iDestruct "H2" as "[%x' [% H2]]". iMod (own_update_2 with "H1 H2") as "Hown". - { unshelve eapply (auth_update_dealloc _ _ ((b-b') _)%NNR), nonnegreal_local_update. + { unshelve eapply (auth_update_dealloc _ _ ((x-x') _)%NNR), nonnegreal_local_update. - lra. - apply cond_nonneg. - apply nnreal_ext =>/=. lra. } - iFrame. by iPureIntro. + iFrame. simpl. iPureIntro. + by subst. Qed. - Lemma hocap_error_increase γ (b b':nonnegreal) : - (b+b'<1)%R -> ⊢ (●↯ b @ γ) ==∗ (●↯ (b+b')%NNR @ γ) ∗ (◯↯ b' @ γ). + Lemma hocap_error_increase γ (b:R) (b':nonnegreal) : + (b+b'<1)%R -> ⊢ (●↯ b @ γ) ==∗ (●↯ (b+b')%R @ γ) ∗ (◯↯ b' @ γ). Proof. - iIntros (Hineq) "[% [% H]]". + iIntros (Hineq) "[%x [% H]]". iMod (own_update with "H") as "[$ $]"; last (iPureIntro; split; last done). - apply auth_update_alloc. - apply (local_update_unital_discrete _ _ (b+b')%NNR) => z H1 H2. - split; first done. - apply nnreal_ext. simpl. - rewrite Rplus_comm. - apply Rplus_eq_compat_l. - simpl in *. rewrite -H H2. simpl. lra. - - done. + apply (local_update_unital_discrete _ _ (x+b')%NNR) => z H1 H2. + split. + + destruct x, b'. compute. simpl in *. lra. + + apply nnreal_ext. simpl. + rewrite Rplus_comm. + apply Rplus_eq_compat_l. + simpl in *. rewrite H2. simpl. lra. + - simpl. lra. Qed. Lemma hocap_error_auth_irrel γ (b c:R) : @@ -139,76 +141,74 @@ Section error_lemmas. End error_lemmas. -(* Definition hocap_tapes_nroot:=nroot.@"tapes". *) -(* Class hocap_tapesGS (Σ : gFunctors) := Hocap_tapesGS { *) -(* hocap_tapesGS_inG :: ghost_mapG Σ loc (nat*list nat) *) -(* }. *) -(* Definition hocap_tapesΣ := ghost_mapΣ loc (nat*list nat). *) +Definition hocap_tapes_nroot:=nroot.@"tapes". +Class hocap_tapesGS (Σ : gFunctors) := Hocap_tapesGS { + hocap_tapesGS_inG :: ghost_mapG Σ loc (nat*list nat) + }. +Definition hocap_tapesΣ := ghost_mapΣ loc (nat*list nat). -(* Notation "α ◯↪N ( M ; ns ) @ γ":= (α ↪[ γ ] (M,ns))%I *) -(* (at level 20, format "α ◯↪N ( M ; ns ) @ γ") : bi_scope. *) +Notation "α ◯↪N ( M ; ns ) @ γ":= (α ↪[ γ ] (M,ns))%I + (at level 20, format "α ◯↪N ( M ; ns ) @ γ") : bi_scope. -(* Notation "● m @ γ" := (ghost_map_auth γ 1 m) (at level 20) : bi_scope. *) +Notation "● m @ γ" := (ghost_map_auth γ 1 m) (at level 20) : bi_scope. -(* Section tapes_lemmas. *) -(* Context `{!conerisGS Σ, !hocap_tapesGS Σ}. *) +Section tapes_lemmas. + Context `{!conerisGS Σ, !hocap_tapesGS Σ}. -(* Lemma hocap_tapes_alloc m: *) -(* ⊢ |==>∃ γ, (● m @ γ) ∗ [∗ map] k↦v ∈ m, (k ◯↪N (v.1; v.2) @ γ). *) -(* Proof. *) -(* iMod ghost_map_alloc as (γ) "[??]". *) -(* iFrame. iModIntro. *) -(* iApply big_sepM_mono; last done. *) -(* by iIntros (?[??]). *) -(* Qed. *) - -(* Lemma hocap_tapes_agree m γ k N ns: *) -(* (● m @ γ) -∗ (k ◯↪N (N; ns) @ γ) -∗ ⌜ m!!k = Some (N, ns) ⌝. *) -(* Proof. *) -(* iIntros "H1 H2". *) -(* by iCombine "H1 H2" gives "%". *) -(* Qed. *) + Lemma hocap_tapes_alloc m: + ⊢ |==>∃ γ, (● m @ γ) ∗ [∗ map] k↦v ∈ m, (k ◯↪N (v.1; v.2) @ γ). + Proof. + iMod ghost_map_alloc as (γ) "[??]". + iFrame. iModIntro. + iApply big_sepM_mono; last done. + by iIntros (?[??]). + Qed. -(* Lemma hocap_tapes_new γ m k N ns : *) -(* m!!k=None -> ⊢ (● m @ γ) ==∗ (● (<[k:=(N,ns)]>m) @ γ) ∗ (k ◯↪N (N; ns) @ γ). *) -(* Proof. *) -(* iIntros (Hlookup) "H". *) -(* by iApply ghost_map_insert. *) -(* Qed. *) + Lemma hocap_tapes_agree m γ k N ns: + (● m @ γ) -∗ (k ◯↪N (N; ns) @ γ) -∗ ⌜ m!!k = Some (N, ns) ⌝. + Proof. + iIntros "H1 H2". + by iCombine "H1 H2" gives "%". + Qed. -(* Lemma hocap_tapes_presample γ m k N ns n: *) -(* (● m @ γ) -∗ (k ◯↪N (N; ns) @ γ) ==∗ (● (<[k:=(N,ns++[n])]>m) @ γ) ∗ (k ◯↪N (N; ns++[n]) @ γ). *) -(* Proof. *) -(* iIntros "H1 H2". *) -(* iApply (ghost_map_update with "[$][$]"). *) -(* Qed. *) + Lemma hocap_tapes_new γ m k N ns : + m!!k=None -> ⊢ (● m @ γ) ==∗ (● (<[k:=(N,ns)]>m) @ γ) ∗ (k ◯↪N (N; ns) @ γ). + Proof. + iIntros (Hlookup) "H". + by iApply ghost_map_insert. + Qed. -(* Lemma hocap_tapes_presample' γ m k N ns ns': *) -(* (● m @ γ) -∗ (k ◯↪N (N; ns) @ γ) ==∗ (● (<[k:=(N,ns++ns')]>m) @ γ) ∗ (k ◯↪N (N; ns++ns') @ γ). *) -(* Proof. *) -(* iIntros "H1 H2". *) -(* iApply (ghost_map_update with "[$][$]"). *) -(* Qed. *) + Lemma hocap_tapes_update γ m k N N' ns ns': + (● m @ γ) -∗ (k ◯↪N (N; ns) @ γ) ==∗ (● (<[k:=(N',ns')]>m) @ γ) ∗ (k ◯↪N (N'; ns') @ γ). + Proof. + iIntros "H1 H2". + iApply (ghost_map_update with "[$][$]"). + Qed. -(* Lemma hocap_tapes_pop γ m k N ns n: *) -(* (● m @ γ) -∗ (k ◯↪N (N; n::ns) @ γ) ==∗ (● (<[k:=(N,ns)]>m) @ γ) ∗ (k ◯↪N (N; ns) @ γ). *) -(* Proof. *) -(* iIntros "H1 H2". *) -(* iApply (ghost_map_update with "[$][$]"). *) -(* Qed. *) + (* Lemma hocap_tapes_presample γ m k N ns n: *) + (* (● m @ γ) -∗ (k ◯↪N (N; ns) @ γ) ==∗ (● (<[k:=(N,ns++[n])]>m) @ γ) ∗ (k ◯↪N (N; ns++[n]) @ γ). *) + (* Proof. *) + (* iIntros "H1 H2". *) + (* iApply (ghost_map_update with "[$][$]"). *) + (* Qed. *) + + (* Lemma hocap_tapes_presample' γ m k N ns ns': *) + (* (● m @ γ) -∗ (k ◯↪N (N; ns) @ γ) ==∗ (● (<[k:=(N,ns++ns')]>m) @ γ) ∗ (k ◯↪N (N; ns++ns') @ γ). *) + (* Proof. *) + (* iIntros "H1 H2". *) + (* iApply (ghost_map_update with "[$][$]"). *) + (* Qed. *) + + (* Lemma hocap_tapes_pop γ m k N ns n: *) + (* (● m @ γ) -∗ (k ◯↪N (N; n::ns) @ γ) ==∗ (● (<[k:=(N,ns)]>m) @ γ) ∗ (k ◯↪N (N; ns) @ γ). *) + (* Proof. *) + (* iIntros "H1 H2". *) + (* iApply (ghost_map_update with "[$][$]"). *) + (* Qed. *) -(* Lemma hocap_tapes_notin α N ns m (f:(nat*list nat)-> nat) g: *) -(* α ↪N (N; ns) -∗ ([∗ map] α0↦t ∈ m, α0 ↪N (f t; g t)) -∗ ⌜m!!α=None ⌝. *) -(* Proof. *) -(* destruct (m!!α) eqn:Heqn; last by iIntros. *) -(* iIntros "Hα Hmap". *) -(* iDestruct (big_sepM_lookup with "[$]") as "?"; first done. *) -(* iExFalso. *) -(* iApply (tapeN_tapeN_contradict with "[$][$]"). *) -(* Qed. *) + -(* (** * TODO add*) *) -(* End tapes_lemmas. *) +End tapes_lemmas. (* Section HOCAP. *) diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index 8fcb7c97..3d578203 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -50,7 +50,7 @@ Class rand_spec `{!conerisGS Σ} := RandSpec rand_error_ineq {L : randG Σ} γ x1 x2: rand_error_auth (L:=L) γ x1 -∗ rand_error_frag (L:=L) γ x2 -∗ ⌜(x2 <= x1)%R ⌝; rand_error_decrease {L : randG Σ} γ (x1 x2:nonnegreal): - rand_error_auth (L:=L) γ x1 -∗ rand_error_frag (L:=L) γ x2 ==∗ rand_error_auth (L:=L) γ (x2-x1)%R; + rand_error_auth (L:=L) γ x1 -∗ rand_error_frag (L:=L) γ x2 ==∗ rand_error_auth (L:=L) γ (x1-x2)%R; rand_error_increase {L : randG Σ} γ (x1 x2:nonnegreal): (x1+x2<1)%R -> ⊢ rand_error_auth (L:=L) γ x1 ==∗ rand_error_auth (L:=L) γ (x1+x2) ∗ rand_error_frag (L:=L) γ x2; @@ -69,11 +69,18 @@ Class rand_spec `{!conerisGS Σ} := RandSpec rand_tapes_auth (L:=L) γ (<[α := ns']> m) ∗ rand_tapes_frag (L:=L) γ α ns'; (** * Program specs *) - rand_inv_create_spec {L : randG Σ} N E ε: + rand_inv_create_spec {L : randG Σ} N E: + ⊢ wp_update E (∃ γ1 γ2, is_rand (L:=L) N γ1 γ2); + rand_err_convert_spec1 {L : randG Σ} N E ε γ1 γ2: + ↑N ⊆ E-> + is_rand (L:=L) N γ1 γ2 -∗ ↯ ε -∗ - wp_update E (∃ γ1 γ2, is_rand (L:=L) N γ1 γ2 ∗ - rand_error_frag (L:=L) γ1 ε); - + wp_update E (rand_error_frag (L:=L) γ1 ε); + rand_err_convert_spec2 {L : randG Σ} N E ε γ1 γ2: + ↑N ⊆ E-> + is_rand (L:=L) N γ1 γ2 -∗ + rand_error_frag (L:=L) γ1 ε -∗ + wp_update E (↯ ε); rand_allocate_tape_spec {L: randG Σ} N E (tb:nat) γ1 γ2: ↑N ⊆ E-> {{{ is_rand (L:=L) N γ1 γ2 }}} @@ -105,3 +112,163 @@ Class rand_spec `{!conerisGS Σ} := RandSpec P -∗ rand_tapes_frag (L:=L) γ2 α (tb, ns)-∗ wp_update E (∃ n, T n ∗ rand_tapes_frag (L:=L) γ2 α (tb, ns++(fin_to_nat <$> n))) }. + +Section impl. + Local Opaque INR. + (* (** Instantiate rand *) *) + Class randG1 Σ := RandG1 { flipG1_error::hocap_errorGS Σ; + flipG1_tapes:: hocap_tapesGS Σ; + }. + Local Definition rand_inv_pred1 `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapesGS Σ} γ1 γ2:= + (∃ (ε:R) (m:gmap loc (nat * list nat)) , + ↯ ε ∗ ●↯ ε @ γ1 ∗ + ([∗ map] α ↦ t ∈ m, α ↪N ( t.1 ; t.2 )) ∗ + ●m@γ2 + )%I. + + #[local] Program Definition rand_spec1 `{!conerisGS Σ}: rand_spec := + {| rand_allocate_tape:= (λ: "N", alloc "N"); + rand_tape:= (λ: "α" "N", rand("α") "N"); + randG := randG1; + rand_error_name := gname; + rand_tape_name := gname; + is_rand _ N γ1 γ2 := inv N (rand_inv_pred1 γ1 γ2); + rand_error_auth _ γ x := ●↯ x @ γ; + rand_error_frag _ γ x := ◯↯ x @ γ; + rand_tapes_auth _ γ m := (●m@γ)%I; + rand_tapes_frag _ γ α ns := (α ◯↪N (ns.1; ns.2) @ γ)%I; + |}. + Next Obligation. + simpl. + iIntros. + iApply (hocap_error_auth_exclusive with "[$][$]"). + Qed. + Next Obligation. + simpl. + iIntros. + iApply (hocap_error_frag_split). + Qed. + Next Obligation. + simpl. + iIntros (?????) "H". + iApply (hocap_error_auth_valid with "[$]"). + Qed. + Next Obligation. + simpl. + iIntros (?????) "H". + iApply (hocap_error_frag_valid with "[$]"). + Qed. + Next Obligation. + simpl. + iIntros (??????) "H1 H2". + iApply (hocap_error_ineq with "[$][$]"). + Qed. + Next Obligation. + iIntros. + iApply (hocap_error_decrease with "[$][$]"). + Qed. + Next Obligation. + iIntros. + by iApply (hocap_error_increase with "[$]"). + Qed. + Next Obligation. + simpl. + iIntros (??????) "H1 H2". + by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[%H _]". + Qed. + Next Obligation. + simpl. + iIntros (???????) "H1 H2". + iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. + rewrite dfrac_op_own dfrac_valid_own. by intro. + Qed. + Next Obligation. + simpl. + iIntros (??????[]) "??". + by iDestruct (hocap_tapes_agree with "[$][$]") as "%H". + Qed. + Next Obligation. + simpl. + iIntros (??????????) "(%&#Hinv&?)". + iInv "Hinv" as ">(%&%&?&?&?&?)" "Hclose". + iDestruct (hocap_tapes_agree with "[$][$]") as "%". + iDestruct (big_sepM_lookup_acc with "[$]") as "[H H']"; first done. + iDestruct (tapeN_ineq with "[$]") as "%". + iMod ("Hclose" with "[-]"); last done. + iFrame. + by iApply "H'". + Qed. + Next Obligation. + iIntros (???????[]) "??". + iMod (hocap_tapes_update with "[$][$]") as "[??]". + by iFrame. + Qed. + Next Obligation. + simpl. + iIntros (?????). + iApply fupd_wp_update_ret. + unshelve iMod (hocap_error_alloc (nnreal_zero)) as "(%γ1 & ? & ?)"; simpl; [rewrite INR_0; lra|]. + iMod ec_zero as "?". + iMod (hocap_tapes_alloc ∅) as "(%γ2 & H4 & H5)". + iMod (inv_alloc _ _ (rand_inv_pred1 γ1 γ2) with "[-]"). + { iFrame. by iNext. } + by iFrame. + Qed. + Next Obligation. + simpl. + iIntros (????? ε ???) "#Hinv Herr". + iApply fupd_wp_update_ret. + iInv "Hinv" as ">(%ε'&%&H1&?&?&?)" "Hclose". + iDestruct (ec_valid with "[$Herr]") as "%". + iCombine "Herr H1" as "?". + iDestruct (ec_valid with "[$]") as "[% %]". + simpl in *. + unshelve iMod (hocap_error_increase _ _ (mknonnegreal ε _) with "[$]") as "[? Hfrag]"; [naive_solver|simpl;lra|]. + simpl. + iMod ("Hclose" with "[-Hfrag]") as "_"; last done. + iFrame. iApply (ec_eq with "[$]"). lra. + Qed. + Next Obligation. + simpl. + iIntros (????? ε ???) "#Hinv Herr". + iApply fupd_wp_update_ret. + iInv "Hinv" as ">(%ε'&%&H1&?&?&?)" "Hclose". + iDestruct (hocap_error_frag_valid with "[$]") as "%". + iDestruct (ec_valid with "[$]") as "%". + iDestruct (hocap_error_ineq with "[$][$]") as "%". + unshelve iMod (hocap_error_decrease with "[$][$]") as "?". + iDestruct (ec_eq with "[$]") as "?"; last first. + { iDestruct (ec_split with "[$]") as "[? H]"; last first. + - iMod ("Hclose" with "[-H]") as "_"; last done. + iFrame. + - naive_solver. + - lra. + } + lra. + Qed. + Next Obligation. + simpl. + iIntros (????????? Φ) "#Hinv HΦ". + wp_pures. + iInv "Hinv" as "(%&%&?&?&?&?)" "Hclose". + wp_apply (wp_alloc_tape); first done. + iIntros (α) "Hα". + iDestruct (hocap_tapes_notin with "[$][$]") as "%". + iMod (hocap_tapes_new _ _ α tb [] with "[$]") as "[?H]"; first done. + iMod ("Hclose" with "[-H HΦ]"). + { iModIntro. iFrame. + rewrite big_sepM_insert; [iFrame|done]. + } + iApply "HΦ". + by iFrame. + Qed. + Next Obligation. + simpl. + Admitted. + Next Obligation. + simpl. + Admitted. + + +End impl. + diff --git a/theories/coneris/primitive_laws.v b/theories/coneris/primitive_laws.v index 980f143b..c1eae4b0 100644 --- a/theories/coneris/primitive_laws.v +++ b/theories/coneris/primitive_laws.v @@ -147,6 +147,16 @@ Section tape_interface. - simpl. intros. lia. Qed. + + Lemma hocap_tapes_notin α N ns m (f:(nat*list nat)-> nat) g: + α ↪N (N; ns) -∗ ([∗ map] α0↦t ∈ m, α0 ↪N (f t; g t)) -∗ ⌜m!!α=None ⌝. + Proof. + destruct (m!!α) eqn:Heqn; last by iIntros. + iIntros "Hα Hmap". + iDestruct (big_sepM_lookup with "[$]") as "?"; first done. + iExFalso. + iApply (tapeN_tapeN_contradict with "[$][$]"). + Qed. (* Lemma spec_tapeN_to_empty l M : diff --git a/theories/coneris/wp_update.v b/theories/coneris/wp_update.v index e3472ce7..fde19e70 100644 --- a/theories/coneris/wp_update.v +++ b/theories/coneris/wp_update.v @@ -71,6 +71,14 @@ Section wp_update. iMod "H". by iApply "H". Qed. + Lemma fupd_wp_update_ret E P: + (|={E}=> P) ⊢ wp_update E P. + Proof. + iIntros "H". + iApply fupd_wp_update. + by iApply wp_update_ret. + Qed. + Lemma wp_update_frame_l R E P : R ∗ wp_update E P ⊢ wp_update E (P ∗ R). Proof. From c6361a9e23c3f48e6caf03c1bf4cb89a6cabe9a6 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Mon, 23 Sep 2024 15:05:11 +0200 Subject: [PATCH 22/69] NIT --- theories/coneris/lib/hocap_rand.v | 116 ++++++++++++++++++++++++++++-- 1 file changed, 110 insertions(+), 6 deletions(-) diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index 3d578203..3d839ffc 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -263,12 +263,116 @@ Section impl. by iFrame. Qed. Next Obligation. - simpl. - Admitted. - Next Obligation. - simpl. - Admitted. - + simpl. + iIntros (?????????????? Φ) "(#Hinv & #Hvs & HP) HΦ". + wp_pures. + wp_bind (rand(_) _)%E. + iInv "Hinv" as ">(%&%&H1&H2&H3&H4)" "Hclose". + iMod ("Hvs" with "[$]") as "(%&HQ&Hauth)". + iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. + simpl. + wp_apply (wp_rand_tape with "[$]") as "[Htape %]". + iMod ("Hclose" with "[- HΦ HQ]") as "_". + { iExists _, (<[α:=_]> m). + iFrame. + iApply "H3". by iNext. + } + by iApply "HΦ". +Qed. +Next Obligation. + simpl. + iIntros (??????????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". + iApply wp_update_state_step_coupl. + iIntros (σ1 ε_supply) "((Hheap&Htapes)&Hε)". + iMod (inv_acc with "Hinv") as "[>(%ε' & % & H1 & H2 & H3 & H4 ) Hclose]"; [done|]. + iDestruct (hocap_tapes_agree with "[$][$]") as "%K". + iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. + simpl. + iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". + iDestruct (ec_supply_bound with "[$][$]") as "%". + iMod (ec_supply_decrease with "[$][$]") as (ε'' ε_rem -> Hε') "Hε_supply". subst. +Admitted. +(* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) +(* iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. *) +(* pose (diff := (ε'' - ε)%R). *) +(* unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 x + (diff/ ((S tb) ^ num))else 1)%R _). *) +(* { case_match; last (simpl;lra). apply Hpos. by apply Nat.eqb_eq. } *) +(* iSplit. *) +(* { iPureIntro. *) +(* simpl. *) +(* unshelve epose proof (Hineq ε1' _) as H'; first apply cond_nonneg. *) +(* etrans; last exact. *) +(* rewrite (Rdiv_def _ (2^_)) -SeriesC_scal_r. *) +(* etrans; last eapply (SeriesC_le_inj _ (λ x, Some (nat_to_bool <$> (fin_to_nat <$> x)))). *) +(* - apply SeriesC_le. *) +(* + simpl. intros n. *) +(* rewrite !fmap_length. *) +(* case_match. *) +(* * replace (1+1)%R with 2%R by lra. *) +(* rewrite -Rdiv_1_l. simpl. *) +(* split; last lra. *) +(* apply Rmult_le_pos. *) +(* -- apply Rcomplements.Rdiv_le_0_compat; first lra. *) +(* apply pow_lt; lra. *) +(* -- apply Hpos; first done. *) +(* rewrite !fmap_length. *) +(* by apply Nat.eqb_eq. *) +(* * lra. *) +(* + simpl. *) +(* apply (ex_seriesC_list_length _ num). *) +(* intros ?. rewrite !fmap_length. *) +(* case_match; last lra. *) +(* intros. by apply Nat.eqb_eq. *) +(* - intros. case_match; last lra. *) +(* apply Rmult_le_pos; first apply Hpos; simpl; auto. *) +(* + by apply Nat.eqb_eq. *) +(* + rewrite -Rdiv_1_l. *) +(* apply Rcomplements.Rdiv_le_0_compat; first lra. *) +(* apply pow_lt; lra. *) +(* - intros ??? <- K. *) +(* simplify_eq. *) +(* rewrite -!list_fmap_compose in K. *) +(* apply list_fmap_eq_inj in K; try done. *) +(* intros x x'. *) +(* repeat (inv_fin x; try intros x); *) +(* repeat (inv_fin x'; try intros x'); done. *) +(* - apply (ex_seriesC_list_length _ num). *) +(* intros ?. *) +(* case_match; last lra. *) +(* intros. by apply Nat.eqb_eq. *) +(* } *) +(* iIntros (sample) "<-". *) +(* destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' (nat_to_bool <$> (fin_to_nat <$> sample))))%R 1%R) as [Hdec|Hdec]; last first. *) +(* { apply Rnot_lt_ge, Rge_le in Hdec. *) +(* iApply state_step_coupl_ret_err_ge_1. *) +(* simpl. simpl in *. rewrite Nat.eqb_refl. lra. *) +(* } *) +(* iApply state_step_coupl_ret. *) +(* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' (nat_to_bool <$> (fin_to_nat <$> sample))) _) with "Hε_supply") as "[Hε_supply Hε]". *) +(* { apply Hpos; first apply cond_nonneg. by rewrite !fmap_length. } *) +(* { simpl. done. } *) +(* simpl. *) +(* iMod (tapeN_update_append' _ _ _ _ sample with "[$][Htape]") as "[Htapes Htape]". *) +(* { by erewrite Heq. } *) +(* iMod (hocap_tapes_presample' _ _ _ _ _ (fin_to_nat <$> sample) with "[$][$]") as "[H4 Hfrag]". *) +(* iMod "Hclose'" as "_". *) +(* iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". *) +(* { iExFalso. iApply (ec_contradict with "[$]"). exact. } *) +(* iMod ("Hclose" with "[$Hε $H2 Htape H3 H4]") as "_". *) +(* { iNext. *) +(* iExists (<[α:=(ns ++ (nat_to_bool<$>(fin_to_nat <$> sample)))]>m). *) +(* rewrite fmap_insert. *) +(* rewrite big_sepM_insert_delete Heq/=. *) +(* rewrite fmap_delete. iFrame. *) +(* rewrite fmap_app/=. *) +(* rewrite !Hrewrite. iFrame. *) +(* } *) +(* iApply fupd_mask_intro_subseteq; first set_solver. *) +(* iFrame. rewrite fmap_app/=Hrewrite. iFrame. *) +(* erewrite (nnreal_ext _ _); first done. *) +(* simpl. by rewrite Nat.eqb_refl. *) +(* Qed. *) + End impl. From 30584833a5871312471b1a3a13a168ed1ca8628c Mon Sep 17 00:00:00 2001 From: Hei Li Date: Tue, 24 Sep 2024 10:51:58 +0200 Subject: [PATCH 23/69] hocap rand progress --- theories/coneris/lib/hocap_rand.v | 443 ++++++++++++++++++++++++------ 1 file changed, 355 insertions(+), 88 deletions(-) diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index 3d839ffc..7e95f85d 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -105,12 +105,18 @@ Class rand_spec `{!conerisGS Σ} := RandSpec (∀ l, length l = num -> 0<= ε2 l)%R -> (SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R-> is_rand (L:=L) N γ1 γ2 -∗ - (□∀ (ε':R) ns', (P ∗ rand_error_auth (L:=L) γ1 (ε')%R ∗ ⌜length ns' = num⌝) ={E∖↑N}=∗ - ∃ diff: R, ⌜(0<=diff)%R⌝ ∗ ⌜(ε' = ε+diff)%R⌝ ∗ - (⌜(1<=ε2 ns' + diff)%R⌝ ∨ (rand_error_auth (L:=L) γ1 (ε2 ns' + diff)%R ∗ T ns'))) + □(∀ (ε': R) m, P ∗rand_error_auth (L:=L) γ1 (ε')%R ∗ rand_tapes_auth (L:=L) γ2 m ={E∖↑N}=∗ + ⌜(ε<=ε')%R⌝ ∗ ⌜(m!!α = Some (tb, ns))⌝ ∗ + P ∗rand_error_auth (L:=L) γ1 (ε')%R ∗ rand_tapes_auth (L:=L) γ2 m + ) -∗ + (□∀ (ε':R) ns' m, (P ∗ rand_error_auth (L:=L) γ1 (ε')%R ∗ ⌜length ns' = num⌝ ∗ + rand_tapes_auth (L:=L) γ2 m + ) ={E∖↑N}=∗ + (⌜(1<=ε2 ns' + (ε'-ε))%R⌝ ∨ (rand_error_auth (L:=L) γ1 (ε2 ns' + (ε'-ε))%R ∗ + rand_tapes_auth (L:=L) γ2 (<[α:= (tb, ns++(fin_to_nat <$> ns'))]> m) ∗ T ns'))) -∗ - P -∗ rand_tapes_frag (L:=L) γ2 α (tb, ns)-∗ - wp_update E (∃ n, T n ∗ rand_tapes_frag (L:=L) γ2 α (tb, ns++(fin_to_nat <$> n))) + P -∗ + wp_update E (∃ n, T n) }. Section impl. @@ -281,98 +287,359 @@ Section impl. Qed. Next Obligation. simpl. - iIntros (??????????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". + iIntros (??????????????? Hsubset Hpos Hineq) "#Hinv #Hchange #Hvs HP". iApply wp_update_state_step_coupl. iIntros (σ1 ε_supply) "((Hheap&Htapes)&Hε)". iMod (inv_acc with "Hinv") as "[>(%ε' & % & H1 & H2 & H3 & H4 ) Hclose]"; [done|]. - iDestruct (hocap_tapes_agree with "[$][$]") as "%K". + iMod ("Hchange" with "[$]") as "(%&%&HP & H2 &H4)". + (* iDestruct (hocap_tapes_agree with "[$][$]") as "%K". *) iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. simpl. iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". iDestruct (ec_supply_bound with "[$][$]") as "%". iMod (ec_supply_decrease with "[$][$]") as (ε'' ε_rem -> Hε') "Hε_supply". subst. + iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". + iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. + unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 x + (ε''-ε) else 1)%R _). + { case_match; last (simpl;lra). admit. } + iSplit. + { iPureIntro. + simpl. admit. + (* unshelve epose proof (Hineq ε1' _) as H'; first apply cond_nonneg. *) + (* etrans; last exact. *) + (* rewrite (Rdiv_def _ (2^_)) -SeriesC_scal_r. *) + (* etrans; last eapply (SeriesC_le_inj _ (λ x, Some (nat_to_bool <$> (fin_to_nat <$> x)))). *) + (* - apply SeriesC_le. *) + (* + simpl. intros n. *) + (* rewrite !fmap_length. *) + (* case_match. *) + (* * replace (1+1)%R with 2%R by lra. *) + (* rewrite -Rdiv_1_l. simpl. *) + (* split; last lra. *) + (* apply Rmult_le_pos. *) + (* -- apply Rcomplements.Rdiv_le_0_compat; first lra. *) + (* apply pow_lt; lra. *) + (* -- apply Hpos; first done. *) + (* rewrite !fmap_length. *) + (* by apply Nat.eqb_eq. *) + (* * lra. *) + (* + simpl. *) + (* apply (ex_seriesC_list_length _ num). *) + (* intros ?. rewrite !fmap_length. *) + (* case_match; last lra. *) + (* intros. by apply Nat.eqb_eq. *) + (* - intros. case_match; last lra. *) + (* apply Rmult_le_pos; first apply Hpos; simpl; auto. *) + (* + by apply Nat.eqb_eq. *) + (* + rewrite -Rdiv_1_l. *) + (* apply Rcomplements.Rdiv_le_0_compat; first lra. *) + (* apply pow_lt; lra. *) + (* - intros ??? <- K. *) + (* simplify_eq. *) + (* rewrite -!list_fmap_compose in K. *) + (* apply list_fmap_eq_inj in K; try done. *) + (* intros x x'. *) + (* repeat (inv_fin x; try intros x); *) + (* repeat (inv_fin x'; try intros x'); done. *) + (* - apply (ex_seriesC_list_length _ num). *) + (* intros ?. *) + (* case_match; last lra. *) + (* intros. by apply Nat.eqb_eq. *) + } + iIntros (sample) "<-". + destruct (Rlt_decision (nonneg ε_rem + (ε2 sample + (ε''-ε))%R) 1%R) as [Hdec|Hdec]; last first. + { apply Rnot_lt_ge, Rge_le in Hdec. + iApply state_step_coupl_ret_err_ge_1. + simpl. simpl in *. rewrite Nat.eqb_refl. lra. + } + iApply state_step_coupl_ret. + unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 sample + (ε''-ε))%R _) with "Hε_supply") as "[Hε_supply Hε]". + { admit. } + { simpl. done. } + simpl. + iMod (tapeN_update_append' _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". + (* iMod (hocap_tapes_update _ _ _ _ _ _ ((fin_to_nat <$> ns')++(fin_to_nat <$> sample)) with "[$][$]") as "[H4 Hfrag]". *) + iMod "Hclose'" as "_". + iMod ("Hvs" $! ε'' sample with "[$HP $H2 $H4]") as "[%|(Hauth & H2 & HT)]"; first done. + { iExFalso. iApply (ec_contradict with "[$]"). exact. } + iMod ("Hclose" with "[$Hε $Hauth Htape H3 $H2]") as "_". + { iNext. + by iApply "H3". + } + iApply fupd_mask_intro_subseteq; first set_solver. + iFrame. + erewrite (nnreal_ext _ _); first done. + simpl. by rewrite Nat.eqb_refl. Admitted. -(* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) -(* iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. *) -(* pose (diff := (ε'' - ε)%R). *) -(* unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 x + (diff/ ((S tb) ^ num))else 1)%R _). *) -(* { case_match; last (simpl;lra). apply Hpos. by apply Nat.eqb_eq. } *) -(* iSplit. *) -(* { iPureIntro. *) -(* simpl. *) -(* unshelve epose proof (Hineq ε1' _) as H'; first apply cond_nonneg. *) -(* etrans; last exact. *) -(* rewrite (Rdiv_def _ (2^_)) -SeriesC_scal_r. *) -(* etrans; last eapply (SeriesC_le_inj _ (λ x, Some (nat_to_bool <$> (fin_to_nat <$> x)))). *) -(* - apply SeriesC_le. *) -(* + simpl. intros n. *) -(* rewrite !fmap_length. *) -(* case_match. *) -(* * replace (1+1)%R with 2%R by lra. *) -(* rewrite -Rdiv_1_l. simpl. *) -(* split; last lra. *) -(* apply Rmult_le_pos. *) -(* -- apply Rcomplements.Rdiv_le_0_compat; first lra. *) -(* apply pow_lt; lra. *) -(* -- apply Hpos; first done. *) -(* rewrite !fmap_length. *) -(* by apply Nat.eqb_eq. *) -(* * lra. *) -(* + simpl. *) -(* apply (ex_seriesC_list_length _ num). *) -(* intros ?. rewrite !fmap_length. *) -(* case_match; last lra. *) -(* intros. by apply Nat.eqb_eq. *) -(* - intros. case_match; last lra. *) -(* apply Rmult_le_pos; first apply Hpos; simpl; auto. *) -(* + by apply Nat.eqb_eq. *) -(* + rewrite -Rdiv_1_l. *) -(* apply Rcomplements.Rdiv_le_0_compat; first lra. *) -(* apply pow_lt; lra. *) -(* - intros ??? <- K. *) -(* simplify_eq. *) -(* rewrite -!list_fmap_compose in K. *) -(* apply list_fmap_eq_inj in K; try done. *) -(* intros x x'. *) -(* repeat (inv_fin x; try intros x); *) -(* repeat (inv_fin x'; try intros x'); done. *) -(* - apply (ex_seriesC_list_length _ num). *) -(* intros ?. *) -(* case_match; last lra. *) -(* intros. by apply Nat.eqb_eq. *) -(* } *) -(* iIntros (sample) "<-". *) -(* destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' (nat_to_bool <$> (fin_to_nat <$> sample))))%R 1%R) as [Hdec|Hdec]; last first. *) -(* { apply Rnot_lt_ge, Rge_le in Hdec. *) -(* iApply state_step_coupl_ret_err_ge_1. *) -(* simpl. simpl in *. rewrite Nat.eqb_refl. lra. *) -(* } *) -(* iApply state_step_coupl_ret. *) -(* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' (nat_to_bool <$> (fin_to_nat <$> sample))) _) with "Hε_supply") as "[Hε_supply Hε]". *) -(* { apply Hpos; first apply cond_nonneg. by rewrite !fmap_length. } *) -(* { simpl. done. } *) -(* simpl. *) -(* iMod (tapeN_update_append' _ _ _ _ sample with "[$][Htape]") as "[Htapes Htape]". *) -(* { by erewrite Heq. } *) -(* iMod (hocap_tapes_presample' _ _ _ _ _ (fin_to_nat <$> sample) with "[$][$]") as "[H4 Hfrag]". *) -(* iMod "Hclose'" as "_". *) -(* iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". *) -(* { iExFalso. iApply (ec_contradict with "[$]"). exact. } *) -(* iMod ("Hclose" with "[$Hε $H2 Htape H3 H4]") as "_". *) -(* { iNext. *) -(* iExists (<[α:=(ns ++ (nat_to_bool<$>(fin_to_nat <$> sample)))]>m). *) -(* rewrite fmap_insert. *) -(* rewrite big_sepM_insert_delete Heq/=. *) -(* rewrite fmap_delete. iFrame. *) -(* rewrite fmap_app/=. *) -(* rewrite !Hrewrite. iFrame. *) -(* } *) -(* iApply fupd_mask_intro_subseteq; first set_solver. *) -(* iFrame. rewrite fmap_app/=Hrewrite. iFrame. *) -(* erewrite (nnreal_ext _ _); first done. *) -(* simpl. by rewrite Nat.eqb_refl. *) -(* Qed. *) - End impl. + +(** * EXPERIMENT *) + +Class rand_spec' `{!conerisGS Σ} := RandSpec' +{ + (** * Operations *) + rand_allocate_tape' : val; + rand_tape' : val; + (** * Ghost state *) + (** The assumptions about [Σ] *) + randG' : gFunctors → Type; + + rand_tape_name': Type; + (** * Predicates *) + is_rand' {L : randG' Σ} (N:namespace) + (γ1: rand_tape_name') : iProp Σ; + rand_tapes_auth' {L : randG' Σ} (γ: rand_tape_name') (m:gmap loc (nat * list nat)): iProp Σ; + rand_tapes_frag' {L : randG' Σ} (γ: rand_tape_name') (α:loc) (ns: (nat * list nat)): iProp Σ; + (** * General properties of the predicates *) + #[global] is_rand_persistent' {L : randG' Σ} N γ1 :: + Persistent (is_rand' (L:=L) N γ1); + #[global] rand_tapes_auth_timeless' {L : randG' Σ} γ m :: + Timeless (rand_tapes_auth' (L:=L) γ m); + #[global] rand_tapes_frag_timeless' {L : randG' Σ} γ α ns :: + Timeless (rand_tapes_frag' (L:=L) γ α ns); + #[global] rand_tape_name_inhabited' :: + Inhabited rand_tape_name'; + + rand_tapes_auth_exclusive' {L : randG' Σ} γ m m': + rand_tapes_auth' (L:=L) γ m -∗ rand_tapes_auth' (L:=L) γ m' -∗ False; + rand_tapes_frag_exclusive' {L : randG' Σ} γ α ns ns': + rand_tapes_frag' (L:=L) γ α ns -∗ rand_tapes_frag' (L:=L) γ α ns' -∗ False; + rand_tapes_agree' {L : randG' Σ} γ α m ns: + rand_tapes_auth' (L:=L) γ m -∗ rand_tapes_frag' (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; + rand_tapes_valid' {L : randG' Σ} γ2 α N E tb ns: + ⌜↑N⊆E⌝ ∗ is_rand' (L:=L) N γ2 ∗ rand_tapes_frag' (L:=L) γ2 α (tb, ns) ={E}=∗ + ⌜Forall (λ n, n<=tb)%nat ns⌝; + rand_tapes_update' {L : randG' Σ} γ α m ns ns': + rand_tapes_auth' (L:=L) γ m -∗ rand_tapes_frag' (L:=L) γ α ns ==∗ + rand_tapes_auth' (L:=L) γ (<[α := ns']> m) ∗ rand_tapes_frag' (L:=L) γ α ns'; + + (** * Program specs *) + rand_inv_create_spec' {L : randG' Σ} N E: + ⊢ wp_update E (∃ γ1, is_rand' (L:=L) N γ1); + rand_allocate_tape_spec' {L: randG' Σ} N E (tb:nat) γ2: + ↑N ⊆ E-> + {{{ is_rand' (L:=L) N γ2 }}} + rand_allocate_tape' #tb @ E + {{{ (v:val), RET v; + ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ rand_tapes_frag' (L:=L) γ2 α (tb, []) + }}}; + rand_tape_spec_some' {L: randG' Σ} N E γ2 (P: iProp Σ) (Q:iProp Σ) (α:loc) (n:nat) (tb:nat) ns: + ↑N⊆E -> + {{{ is_rand' (L:=L) N γ2 ∗ + □ (∀ (m:gmap loc (nat * list nat)), P ∗ + rand_tapes_auth' (L:=L) γ2 m + ={E∖↑N}=∗ ⌜m!!α=Some (tb, n::ns)⌝ ∗ Q ∗ rand_tapes_auth' (L:=L) γ2 (<[α:=(tb, ns)]> m)) ∗ + P + }}} + rand_tape' #lbl:α #tb @ E + {{{ RET #n; Q }}}; + rand_presample_spec' {L: randG' Σ} N E ns α (tb:nat) + (ε2 : list (fin (S tb)) -> R) + (P : iProp Σ) num T γ2 (ε:nonnegreal) : + ↑N ⊆ E -> + (∀ l, length l = num -> 0<= ε2 l)%R -> + (SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R-> + is_rand' (L:=L) N γ2 -∗ + □(∀ (ε': nonnegreal) m, P ∗ err_interp ε' ∗ rand_tapes_auth' (L:=L) γ2 m ={E∖↑N}=∗ + ⌜(ε<=ε')%R⌝ ∗ ⌜(m!!α = Some (tb, ns))⌝ ∗ + P ∗err_interp ε' ∗ rand_tapes_auth' (L:=L) γ2 m + ) -∗ + (□∀ (ε': nonnegreal) ns' m, (P ∗ err_interp ε' ∗ ⌜length ns' = num⌝ ∗ + rand_tapes_auth' (L:=L) γ2 m + ) ={E∖↑N}=∗ + (⌜(1<=ε2 ns' + (ε'-ε))%R⌝ ∨ ∃ (ε_final:nonnegreal), + ⌜(nonneg ε_final = ε2 ns' + (ε'-ε))%R ⌝ ∗ + (err_interp ε_final ∗ + rand_tapes_auth' (L:=L) γ2 (<[α:= (tb, ns++(fin_to_nat <$> ns'))]> m) ∗ T ns'))) + -∗ + P -∗ + wp_update E (∃ n, T n) +}. + +Section impl'. + Local Opaque INR. + (* (** Instantiate rand *) *) + Class randG1' Σ := RandG1' { + flipG1_tapes':: hocap_tapesGS Σ; + }. + Local Definition rand_inv_pred1' `{!conerisGS Σ, !hocap_tapesGS Σ} γ2:= + (∃ (m:gmap loc (nat * list nat)) , + ([∗ map] α ↦ t ∈ m, α ↪N ( t.1 ; t.2 )) ∗ + ●m@γ2 + )%I. + + #[local] Program Definition rand_spec1' `{!conerisGS Σ}: rand_spec' := + {| rand_allocate_tape':= (λ: "N", alloc "N"); + rand_tape':= (λ: "α" "N", rand("α") "N"); + randG' := randG1'; + rand_tape_name' := gname; + is_rand' _ N γ2 := inv N (rand_inv_pred1' γ2); + rand_tapes_auth' _ γ m := (●m@γ)%I; + rand_tapes_frag' _ γ α ns := (α ◯↪N (ns.1; ns.2) @ γ)%I; + |}. + Next Obligation. + simpl. + iIntros (??????) "H1 H2". + by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[%H _]". + Qed. + Next Obligation. + simpl. + iIntros (???????) "H1 H2". + iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. + rewrite dfrac_op_own dfrac_valid_own. by intro. + Qed. + Next Obligation. + simpl. + iIntros (??????[]) "??". + by iDestruct (hocap_tapes_agree with "[$][$]") as "%H". + Qed. + Next Obligation. + simpl. + iIntros (?????????) "(%&#Hinv&?)". + iInv "Hinv" as ">(%&?&?)" "Hclose". + iDestruct (hocap_tapes_agree with "[$][$]") as "%". + iDestruct (big_sepM_lookup_acc with "[$]") as "[H H']"; first done. + iDestruct (tapeN_ineq with "[$]") as "%". + iMod ("Hclose" with "[-]"); last done. + iFrame. + by iApply "H'". + Qed. + Next Obligation. + iIntros (???????[]) "??". + iMod (hocap_tapes_update with "[$][$]") as "[??]". + by iFrame. + Qed. + Next Obligation. + simpl. + iIntros (?????). + iApply fupd_wp_update_ret. + iMod (hocap_tapes_alloc ∅) as "(%γ2 & H4 & H5)". + iMod (inv_alloc _ _ (rand_inv_pred1' γ2) with "[-]"). + { iFrame. by iNext. } + by iFrame. + Qed. + Next Obligation. + simpl. + iIntros (???????? Φ) "#Hinv HΦ". + wp_pures. + iInv "Hinv" as "(%&?&?)" "Hclose". + wp_apply (wp_alloc_tape); first done. + iIntros (α) "Hα". + iDestruct (hocap_tapes_notin with "[$][$]") as "%". + iMod (hocap_tapes_new _ _ α tb [] with "[$]") as "[?H]"; first done. + iMod ("Hclose" with "[-H HΦ]"). + { iModIntro. iFrame. + rewrite big_sepM_insert; [iFrame|done]. + } + iApply "HΦ". + by iFrame. + Qed. + Next Obligation. + simpl. + iIntros (????????????? Φ) "(#Hinv & #Hvs & HP) HΦ". + wp_pures. + wp_bind (rand(_) _)%E. + iInv "Hinv" as ">(%&H3&H4)" "Hclose". + iMod ("Hvs" with "[$]") as "(%&HQ&Hauth)". + iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. + simpl. + wp_apply (wp_rand_tape with "[$]") as "[Htape %]". + iMod ("Hclose" with "[- HΦ HQ]") as "_". + { iExists (<[α:=_]> m). + iFrame. + iApply "H3". by iNext. + } + by iApply "HΦ". +Qed. +Next Obligation. + simpl. + iIntros (?????????????? Hsubset Hpos Hineq) "#Hinv #Hcheck #Hvs HP ". + iApply wp_update_state_step_coupl. + iIntros (σ1 ε_supply) "((Hheap&Htapes)&Hε)". + iMod (inv_acc with "Hinv") as "[>(% & H3 & H4 ) Hclose]"; [done|]. + iMod ("Hcheck" with "[$]") as "(%&%&HP&Hε& H4)". + iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. + simpl. + iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". + (* iDestruct (ec_supply_bound with "[$][$]") as "%". *) + (* iMod (ec_supply_decrease with "[$][$]") as (ε'' ε_rem -> Hε') "Hε_supply". subst. *) + iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". + unshelve epose (ε_rem := mknonnegreal (ε_supply-ε) _). + { simpl in *. lra. } + replace (ε_supply) with (ε_rem + ε)%NNR; last first. + { rewrite /ε_rem. apply nnreal_ext. simpl. lra. } + iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. + unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 x else 1)%R _). + { case_match; last (simpl;lra). admit. } + iSplit. + { iPureIntro. + simpl. admit. + (* unshelve epose proof (Hineq ε1' _) as H'; first apply cond_nonneg. *) + (* etrans; last exact. *) + (* rewrite (Rdiv_def _ (2^_)) -SeriesC_scal_r. *) + (* etrans; last eapply (SeriesC_le_inj _ (λ x, Some (nat_to_bool <$> (fin_to_nat <$> x)))). *) + (* - apply SeriesC_le. *) + (* + simpl. intros n. *) + (* rewrite !fmap_length. *) + (* case_match. *) + (* * replace (1+1)%R with 2%R by lra. *) + (* rewrite -Rdiv_1_l. simpl. *) + (* split; last lra. *) + (* apply Rmult_le_pos. *) + (* -- apply Rcomplements.Rdiv_le_0_compat; first lra. *) + (* apply pow_lt; lra. *) + (* -- apply Hpos; first done. *) + (* rewrite !fmap_length. *) + (* by apply Nat.eqb_eq. *) + (* * lra. *) + (* + simpl. *) + (* apply (ex_seriesC_list_length _ num). *) + (* intros ?. rewrite !fmap_length. *) + (* case_match; last lra. *) + (* intros. by apply Nat.eqb_eq. *) + (* - intros. case_match; last lra. *) + (* apply Rmult_le_pos; first apply Hpos; simpl; auto. *) + (* + by apply Nat.eqb_eq. *) + (* + rewrite -Rdiv_1_l. *) + (* apply Rcomplements.Rdiv_le_0_compat; first lra. *) + (* apply pow_lt; lra. *) + (* - intros ??? <- K. *) + (* simplify_eq. *) + (* rewrite -!list_fmap_compose in K. *) + (* apply list_fmap_eq_inj in K; try done. *) + (* intros x x'. *) + (* repeat (inv_fin x; try intros x); *) + (* repeat (inv_fin x'; try intros x'); done. *) + (* - apply (ex_seriesC_list_length _ num). *) + (* intros ?. *) + (* case_match; last lra. *) + (* intros. by apply Nat.eqb_eq. *) + } + iIntros (sample) "<-". + destruct (Rlt_decision (nonneg ε_rem + (ε2 sample)%R) 1%R) as [Hdec|Hdec]; last first. + { apply Rnot_lt_ge, Rge_le in Hdec. + iApply state_step_coupl_ret_err_ge_1. + simpl. simpl in *. rewrite Nat.eqb_refl. lra. + } + iApply state_step_coupl_ret. + simpl. + rewrite -Heq. + iMod (tapeN_update_append' _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". + iMod "Hclose'" as "_". + iMod ("Hvs" $! _ sample with "[$HP $H4 $Hε]") as "Hor"; first done. + iDestruct "Hor" as "[%|(%ε_final & %Heq' & Hε & Hauth & HT)]". + { exfalso. simpl in *. lra. } + iMod ("Hclose" with "[$Hauth Htape H3 ]") as "_". + { iNext. + by iApply "H3". + } + iApply fupd_mask_intro_subseteq; first set_solver. + iFrame. + erewrite (nnreal_ext (_+_)%NNR _); first done. + simpl. rewrite Nat.eqb_refl. simpl in *. lra. +Admitted. + +End impl'. + From 8ab40ccbd79abe5c6695ea0e66f466765a9c2da9 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Tue, 24 Sep 2024 12:41:08 +0200 Subject: [PATCH 24/69] finish hocap rand --- theories/coneris/lib/hocap_rand.v | 124 +++++++++--------------------- 1 file changed, 36 insertions(+), 88 deletions(-) diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index 7e95f85d..ff53dfe0 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -301,50 +301,31 @@ Next Obligation. iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 x + (ε''-ε) else 1)%R _). - { case_match; last (simpl;lra). admit. } + { case_match; last (simpl;lra). + apply Rplus_le_le_0_compat; last (simpl in *; lra). + apply Hpos. by apply Nat.eqb_eq. + } iSplit. { iPureIntro. - simpl. admit. - (* unshelve epose proof (Hineq ε1' _) as H'; first apply cond_nonneg. *) - (* etrans; last exact. *) - (* rewrite (Rdiv_def _ (2^_)) -SeriesC_scal_r. *) - (* etrans; last eapply (SeriesC_le_inj _ (λ x, Some (nat_to_bool <$> (fin_to_nat <$> x)))). *) - (* - apply SeriesC_le. *) - (* + simpl. intros n. *) - (* rewrite !fmap_length. *) - (* case_match. *) - (* * replace (1+1)%R with 2%R by lra. *) - (* rewrite -Rdiv_1_l. simpl. *) - (* split; last lra. *) - (* apply Rmult_le_pos. *) - (* -- apply Rcomplements.Rdiv_le_0_compat; first lra. *) - (* apply pow_lt; lra. *) - (* -- apply Hpos; first done. *) - (* rewrite !fmap_length. *) - (* by apply Nat.eqb_eq. *) - (* * lra. *) - (* + simpl. *) - (* apply (ex_seriesC_list_length _ num). *) - (* intros ?. rewrite !fmap_length. *) - (* case_match; last lra. *) - (* intros. by apply Nat.eqb_eq. *) - (* - intros. case_match; last lra. *) - (* apply Rmult_le_pos; first apply Hpos; simpl; auto. *) - (* + by apply Nat.eqb_eq. *) - (* + rewrite -Rdiv_1_l. *) - (* apply Rcomplements.Rdiv_le_0_compat; first lra. *) - (* apply pow_lt; lra. *) - (* - intros ??? <- K. *) - (* simplify_eq. *) - (* rewrite -!list_fmap_compose in K. *) - (* apply list_fmap_eq_inj in K; try done. *) - (* intros x x'. *) - (* repeat (inv_fin x; try intros x); *) - (* repeat (inv_fin x'; try intros x'); done. *) - (* - apply (ex_seriesC_list_length _ num). *) - (* intros ?. *) - (* case_match; last lra. *) - (* intros. by apply Nat.eqb_eq. *) + simpl. + rewrite (SeriesC_subset (λ x, (x∈enum_uniform_fin_list tb num))); last first. + - intros a. intros K. + rewrite elem_of_enum_uniform_fin_list in K. + case_match eqn :H4; last done. + by rewrite Nat.eqb_eq in H4. + - erewrite (SeriesC_ext _ (λ a, 1/S tb ^ num * (if bool_decide (a ∈ enum_uniform_fin_list tb num) then ε2 a else 0) + (if bool_decide (a ∈ enum_uniform_fin_list tb num) then 1/S tb ^ num * (ε''-ε) else 0)))%R; last first. + + intros n. + case_bool_decide; last lra. + replace (_ =? _) with true; last first. + * symmetry. rewrite Nat.eqb_eq. by rewrite -elem_of_enum_uniform_fin_list. + * simpl. lra. + + rewrite SeriesC_plus; [|apply ex_seriesC_scal_l, ex_seriesC_list|apply ex_seriesC_list..]. + rewrite SeriesC_list_2; last apply NoDup_enum_uniform_fin_list. + rewrite SeriesC_scal_l enum_uniform_fin_list_length pow_INR. + replace (1/_*_*_)%R with (ε''-ε)%R; first (simpl; lra). + rewrite Rmult_comm -Rmult_assoc Rdiv_1_l Rmult_inv_r; first lra. + apply pow_nonzero. + pose proof pos_INR_S tb; lra. } iIntros (sample) "<-". destruct (Rlt_decision (nonneg ε_rem + (ε2 sample + (ε''-ε))%R) 1%R) as [Hdec|Hdec]; last first. @@ -354,7 +335,7 @@ Next Obligation. } iApply state_step_coupl_ret. unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 sample + (ε''-ε))%R _) with "Hε_supply") as "[Hε_supply Hε]". - { admit. } + { apply Rplus_le_le_0_compat; simpl in *; [naive_solver|lra]. } { simpl. done. } simpl. iMod (tapeN_update_append' _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". @@ -370,8 +351,8 @@ Next Obligation. iFrame. erewrite (nnreal_ext _ _); first done. simpl. by rewrite Nat.eqb_refl. -Admitted. - +Qed. + End impl. @@ -572,50 +553,17 @@ Next Obligation. { rewrite /ε_rem. apply nnreal_ext. simpl. lra. } iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 x else 1)%R _). - { case_match; last (simpl;lra). admit. } + { case_match; last (simpl;lra). apply Hpos. by rewrite -Nat.eqb_eq. } iSplit. { iPureIntro. - simpl. admit. - (* unshelve epose proof (Hineq ε1' _) as H'; first apply cond_nonneg. *) - (* etrans; last exact. *) - (* rewrite (Rdiv_def _ (2^_)) -SeriesC_scal_r. *) - (* etrans; last eapply (SeriesC_le_inj _ (λ x, Some (nat_to_bool <$> (fin_to_nat <$> x)))). *) - (* - apply SeriesC_le. *) - (* + simpl. intros n. *) - (* rewrite !fmap_length. *) - (* case_match. *) - (* * replace (1+1)%R with 2%R by lra. *) - (* rewrite -Rdiv_1_l. simpl. *) - (* split; last lra. *) - (* apply Rmult_le_pos. *) - (* -- apply Rcomplements.Rdiv_le_0_compat; first lra. *) - (* apply pow_lt; lra. *) - (* -- apply Hpos; first done. *) - (* rewrite !fmap_length. *) - (* by apply Nat.eqb_eq. *) - (* * lra. *) - (* + simpl. *) - (* apply (ex_seriesC_list_length _ num). *) - (* intros ?. rewrite !fmap_length. *) - (* case_match; last lra. *) - (* intros. by apply Nat.eqb_eq. *) - (* - intros. case_match; last lra. *) - (* apply Rmult_le_pos; first apply Hpos; simpl; auto. *) - (* + by apply Nat.eqb_eq. *) - (* + rewrite -Rdiv_1_l. *) - (* apply Rcomplements.Rdiv_le_0_compat; first lra. *) - (* apply pow_lt; lra. *) - (* - intros ??? <- K. *) - (* simplify_eq. *) - (* rewrite -!list_fmap_compose in K. *) - (* apply list_fmap_eq_inj in K; try done. *) - (* intros x x'. *) - (* repeat (inv_fin x; try intros x); *) - (* repeat (inv_fin x'; try intros x'); done. *) - (* - apply (ex_seriesC_list_length _ num). *) - (* intros ?. *) - (* case_match; last lra. *) - (* intros. by apply Nat.eqb_eq. *) + simpl. etrans; last exact. + rewrite (Rdiv_def (SeriesC _)(S tb ^num)%R) -SeriesC_scal_r. + apply SeriesC_le; last apply ex_seriesC_scal_r, ex_seriesC_list. + intros. rewrite elem_of_enum_uniform_fin_list'. + case_match; last lra. + split; last lra. + apply Rmult_le_pos; last (apply Hpos; by rewrite -Nat.eqb_eq). + rewrite -pow_INR. apply Rdiv_INR_ge_0. } iIntros (sample) "<-". destruct (Rlt_decision (nonneg ε_rem + (ε2 sample)%R) 1%R) as [Hdec|Hdec]; last first. @@ -639,7 +587,7 @@ Next Obligation. iFrame. erewrite (nnreal_ext (_+_)%NNR _); first done. simpl. rewrite Nat.eqb_refl. simpl in *. lra. -Admitted. +Qed. End impl'. From e611498ca42d80f12c1a49cec6692592bf58c50b Mon Sep 17 00:00:00 2001 From: Hei Li Date: Tue, 24 Sep 2024 14:38:12 +0200 Subject: [PATCH 25/69] hocap rand --- theories/coneris/lib/hocap_rand.v | 259 ++++++++++++++++++++++++++---- theories/coneris/wp_update.v | 5 + 2 files changed, 229 insertions(+), 35 deletions(-) diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index ff53dfe0..9c9d8bc1 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -49,9 +49,9 @@ Class rand_spec `{!conerisGS Σ} := RandSpec rand_error_frag (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; rand_error_ineq {L : randG Σ} γ x1 x2: rand_error_auth (L:=L) γ x1 -∗ rand_error_frag (L:=L) γ x2 -∗ ⌜(x2 <= x1)%R ⌝; - rand_error_decrease {L : randG Σ} γ (x1 x2:nonnegreal): + rand_error_decrease {L : randG Σ} γ (x1 x2:R): rand_error_auth (L:=L) γ x1 -∗ rand_error_frag (L:=L) γ x2 ==∗ rand_error_auth (L:=L) γ (x1-x2)%R; - rand_error_increase {L : randG Σ} γ (x1 x2:nonnegreal): + rand_error_increase {L : randG Σ} γ (x1:R) (x2:nonnegreal): (x1+x2<1)%R -> ⊢ rand_error_auth (L:=L) γ x1 ==∗ rand_error_auth (L:=L) γ (x1+x2) ∗ rand_error_frag (L:=L) γ x2; @@ -61,10 +61,11 @@ Class rand_spec `{!conerisGS Σ} := RandSpec rand_tapes_frag (L:=L) γ α ns -∗ rand_tapes_frag (L:=L) γ α ns' -∗ False; rand_tapes_agree {L : randG Σ} γ α m ns: rand_tapes_auth (L:=L) γ m -∗ rand_tapes_frag (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; - rand_tapes_valid {L : randG Σ} γ1 γ2 α N E tb ns: - ⌜↑N⊆E⌝ ∗ is_rand (L:=L) N γ1 γ2 ∗ rand_tapes_frag (L:=L) γ2 α (tb, ns) ={E}=∗ + rand_tapes_valid {L : randG Σ} γ2 α tb ns: + rand_tapes_frag (L:=L) γ2 α (tb, ns) -∗ ⌜Forall (λ n, n<=tb)%nat ns⌝; - rand_tapes_update {L : randG Σ} γ α m ns ns': + rand_tapes_update {L : randG Σ} γ α m ns ns': + Forall (λ n, n<=ns'.1)%nat ns'.2 -> rand_tapes_auth (L:=L) γ m -∗ rand_tapes_frag (L:=L) γ α ns ==∗ rand_tapes_auth (L:=L) γ (<[α := ns']> m) ∗ rand_tapes_frag (L:=L) γ α ns'; @@ -94,7 +95,7 @@ Class rand_spec `{!conerisGS Σ} := RandSpec □ (∀ (m:gmap loc (nat * list nat)), P ∗ rand_tapes_auth (L:=L) γ2 m ={E∖↑N}=∗ ⌜m!!α=Some (tb, n::ns)⌝ ∗ Q ∗ rand_tapes_auth (L:=L) γ2 (<[α:=(tb, ns)]> m)) ∗ - P + ▷P }}} rand_tape #lbl:α #tb @ E {{{ RET #n; Q }}}; @@ -142,7 +143,7 @@ Section impl. rand_error_auth _ γ x := ●↯ x @ γ; rand_error_frag _ γ x := ◯↯ x @ γ; rand_tapes_auth _ γ m := (●m@γ)%I; - rand_tapes_frag _ γ α ns := (α ◯↪N (ns.1; ns.2) @ γ)%I; + rand_tapes_frag _ γ α ns := ((α ◯↪N (ns.1; ns.2) @ γ) ∗ ⌜Forall (λ x, x<=ns.1) ns.2⌝)%I; |}. Next Obligation. simpl. @@ -184,28 +185,21 @@ Section impl. Qed. Next Obligation. simpl. - iIntros (???????) "H1 H2". + iIntros (???????) "[H1 %] [H2 %]". iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. rewrite dfrac_op_own dfrac_valid_own. by intro. Qed. Next Obligation. simpl. - iIntros (??????[]) "??". + iIntros (??????[]) "? [??]". by iDestruct (hocap_tapes_agree with "[$][$]") as "%H". Qed. Next Obligation. simpl. - iIntros (??????????) "(%&#Hinv&?)". - iInv "Hinv" as ">(%&%&?&?&?&?)" "Hclose". - iDestruct (hocap_tapes_agree with "[$][$]") as "%". - iDestruct (big_sepM_lookup_acc with "[$]") as "[H H']"; first done. - iDestruct (tapeN_ineq with "[$]") as "%". - iMod ("Hclose" with "[-]"); last done. - iFrame. - by iApply "H'". + by iIntros (???????) "[? $]". Qed. Next Obligation. - iIntros (???????[]) "??". + iIntros (??????[??][??]?) "?[? %]". iMod (hocap_tapes_update with "[$][$]") as "[??]". by iFrame. Qed. @@ -389,10 +383,11 @@ Class rand_spec' `{!conerisGS Σ} := RandSpec' rand_tapes_frag' (L:=L) γ α ns -∗ rand_tapes_frag' (L:=L) γ α ns' -∗ False; rand_tapes_agree' {L : randG' Σ} γ α m ns: rand_tapes_auth' (L:=L) γ m -∗ rand_tapes_frag' (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; - rand_tapes_valid' {L : randG' Σ} γ2 α N E tb ns: - ⌜↑N⊆E⌝ ∗ is_rand' (L:=L) N γ2 ∗ rand_tapes_frag' (L:=L) γ2 α (tb, ns) ={E}=∗ - ⌜Forall (λ n, n<=tb)%nat ns⌝; + rand_tapes_valid' {L : randG' Σ} γ2 α tb ns: + rand_tapes_frag' (L:=L) γ2 α (tb, ns) -∗ + ⌜Forall (λ n, n<=tb)%nat ns⌝ ; rand_tapes_update' {L : randG' Σ} γ α m ns ns': + Forall (λ x, x<=ns'.1) ns'.2 -> rand_tapes_auth' (L:=L) γ m -∗ rand_tapes_frag' (L:=L) γ α ns ==∗ rand_tapes_auth' (L:=L) γ (<[α := ns']> m) ∗ rand_tapes_frag' (L:=L) γ α ns'; @@ -412,7 +407,7 @@ Class rand_spec' `{!conerisGS Σ} := RandSpec' □ (∀ (m:gmap loc (nat * list nat)), P ∗ rand_tapes_auth' (L:=L) γ2 m ={E∖↑N}=∗ ⌜m!!α=Some (tb, n::ns)⌝ ∗ Q ∗ rand_tapes_auth' (L:=L) γ2 (<[α:=(tb, ns)]> m)) ∗ - P + ▷ P }}} rand_tape' #lbl:α #tb @ E {{{ RET #n; Q }}}; @@ -458,7 +453,7 @@ Section impl'. rand_tape_name' := gname; is_rand' _ N γ2 := inv N (rand_inv_pred1' γ2); rand_tapes_auth' _ γ m := (●m@γ)%I; - rand_tapes_frag' _ γ α ns := (α ◯↪N (ns.1; ns.2) @ γ)%I; + rand_tapes_frag' _ γ α ns := (α ◯↪N (ns.1; ns.2) @ γ ∗ ⌜Forall (λ x, x<=ns.1) ns.2⌝)%I; |}. Next Obligation. simpl. @@ -467,28 +462,21 @@ Section impl'. Qed. Next Obligation. simpl. - iIntros (???????) "H1 H2". + iIntros (???????) "[H1 %] [H2 %]". iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. rewrite dfrac_op_own dfrac_valid_own. by intro. Qed. Next Obligation. simpl. - iIntros (??????[]) "??". - by iDestruct (hocap_tapes_agree with "[$][$]") as "%H". + iIntros (??????[]) "?[? %]". + by iDestruct (hocap_tapes_agree with "[$][$]") as "%". Qed. Next Obligation. simpl. - iIntros (?????????) "(%&#Hinv&?)". - iInv "Hinv" as ">(%&?&?)" "Hclose". - iDestruct (hocap_tapes_agree with "[$][$]") as "%". - iDestruct (big_sepM_lookup_acc with "[$]") as "[H H']"; first done. - iDestruct (tapeN_ineq with "[$]") as "%". - iMod ("Hclose" with "[-]"); last done. - iFrame. - by iApply "H'". + by iIntros (???????) "[? $]". Qed. Next Obligation. - iIntros (???????[]) "??". + iIntros (??????[][]?) "?[? %]". iMod (hocap_tapes_update with "[$][$]") as "[??]". by iFrame. Qed. @@ -591,3 +579,204 @@ Qed. End impl'. +Section checks. + Context `{H: conerisGS Σ, r1:@rand_spec Σ H, r2:@rand_spec' Σ H, L:randG Σ, L': randG' Σ}. + + Lemma wp_rand_alloc_tape NS (N:nat) E γ1 γ2 : + ↑NS ⊆ E -> + {{{ is_rand (L:=L) NS γ1 γ2 }}} rand_allocate_tape #N @ E {{{ α, RET #lbl:α; rand_tapes_frag (L:=L) γ2 α (N,[]) }}}. + Proof. + iIntros (Hsubset Φ) "#Hinv HΦ". + wp_apply (rand_allocate_tape_spec with "[//]"); first exact. + iIntros (?) "(%&->&?)". + by iApply "HΦ". + Qed. + + Lemma wp_rand_alloc_tape' NS (N:nat) E γ2: + ↑NS ⊆ E -> + {{{ is_rand' (L:=L') NS γ2 }}} rand_allocate_tape' #N @ E {{{ α, RET #lbl:α; rand_tapes_frag' (L:=L') γ2 α (N,[]) }}}. + Proof. + iIntros (Hsubset Φ) "#Hinv HΦ". + wp_apply (rand_allocate_tape_spec' with "[//]"); first exact. + iIntros (?) "(%&->&?)". + by iApply "HΦ". + Qed. + + Lemma wp_rand_tape_1 NS (N:nat) E γ1 γ2 n ns α: + ↑NS ⊆ E -> + {{{ is_rand (L:=L) NS γ1 γ2 ∗ ▷ rand_tapes_frag (L:=L) γ2 α (N, n :: ns) }}} + rand_tape #lbl:α #N@ E + {{{ RET #(LitInt n); rand_tapes_frag (L:=L) γ2 α (N, ns) ∗ ⌜n <= N⌝ }}}. + Proof. + iIntros (Hsubset Φ) "(#Hinv & Hfrag) HΦ". + wp_apply (rand_tape_spec_some _ _ _ _ _ (⌜n<=N⌝ ∗ rand_tapes_frag _ _ _)%I with "[Hfrag]"); first exact. + - iSplit; first done. iSplit; last iApply "Hfrag". + iModIntro. + iIntros (?) "[Hfrag Hauth]". + iDestruct (rand_tapes_agree with "[$][$]") as "%". + iDestruct (rand_tapes_valid with "[$Hfrag]") as "%K". + inversion K; subst. + iMod (rand_tapes_update _ _ _ _ (N, ns) with "[$][$]") as "[Hauth Hfrag]"; first done. + iFrame "Hauth". iModIntro. + iFrame. + by iPureIntro. + - iIntros "[??]". iApply "HΦ". + iFrame. + Qed. + + Lemma wp_rand_tape_2 NS (N:nat) E γ2 n ns α: + ↑NS ⊆ E -> + {{{ is_rand' (L:=L') NS γ2 ∗ ▷ rand_tapes_frag' (L:=L') γ2 α (N, n :: ns) }}} + rand_tape' #lbl:α #N@ E + {{{ RET #(LitInt n); rand_tapes_frag' (L:=L') γ2 α (N, ns) ∗ ⌜n <= N⌝ }}}. + Proof. + iIntros (Hsubset Φ) "(#Hinv &>Hfrag) HΦ". + wp_apply (rand_tape_spec_some' _ _ _ _ (⌜n<=N⌝ ∗ rand_tapes_frag' _ _ _)%I with "[Hfrag]"); first exact. + - iSplit; first done. iSplit; last iApply "Hfrag". + iModIntro. + iIntros (?) "[Hfrag Hauth]". + iDestruct (rand_tapes_agree' with "[$][$]") as "%". + iDestruct (rand_tapes_valid' with "[$Hfrag]") as "%K". + inversion K; subst. + iMod (rand_tapes_update' _ _ _ _ (N, ns) with "[$][$]") as "[Hauth Hfrag]"; first done. + iFrame "Hauth". iModIntro. + iFrame. + by iPureIntro. + - iIntros "[??]". iApply "HΦ". + iFrame. + Qed. + + Lemma wp_presample_adv_comp_rand_tape (N : nat) NS E α ns (ε1 : R) (ε2 : fin (S N) -> R) γ1 γ2: + ↑NS ⊆ E -> + (∀ n, 0<=ε2 n)%R -> + (SeriesC (λ n, (1 / (S N)) * ε2 n)%R <= ε1)%R → + is_rand (L:=L) NS γ1 γ2 -∗ + ▷ rand_tapes_frag (L:=L) γ2 α (N, ns) -∗ + rand_error_frag (L:=L) γ1 ε1 -∗ + wp_update E (∃ n, rand_error_frag (L:=L) γ1 (ε2 n) ∗ rand_tapes_frag (L:=L) γ2 α (N, ns ++[fin_to_nat n]))%I. + Proof. + iIntros (Hsubset Hpos Hineq) "#Hinv >Htape Herr". + iMod (rand_presample_spec _ _ _ _ _ (λ l, match l with + | [x] => ε2 x + | _ => 1%R + end + ) + (rand_tapes_frag (L:=L) γ2 α (N, ns) ∗rand_error_frag (L:=L) γ1 ε1) 1%nat + (λ ls, ∃ n, ⌜ls=[n]⌝ ∗ rand_error_frag γ1 (ε2 n) ∗ rand_tapes_frag γ2 α (N, ns ++ [fin_to_nat n]))%I + with "[//][][][$Htape $Herr]") as "(%l&%sample & -> & Herr &Htape)". + - done. + - by intros [|?[|]]. + - etrans; last exact. + etrans; last eapply (SeriesC_le_inj _ (λ l, match l with |[x] => Some x | _ => None end)). + + rewrite Rdiv_def -SeriesC_scal_r. apply Req_le_sym. + apply SeriesC_ext. + intros. case_match; subst. + * rewrite bool_decide_eq_false_2; first (simpl;lra). + by rewrite elem_of_enum_uniform_fin_list. + * case_match; last first. + { rewrite bool_decide_eq_false_2; first (simpl;lra). + by rewrite elem_of_enum_uniform_fin_list. } + rewrite bool_decide_eq_true_2; last by apply elem_of_enum_uniform_fin_list. + simpl. + rewrite Rdiv_def Rmult_1_r; lra. + + intros. apply Rmult_le_pos; last done. + apply Rdiv_INR_ge_0. + + intros. repeat case_match; by simplify_eq. + + apply ex_seriesC_finite. + - iModIntro. + iIntros (??) "([??]&?&?)". + iDestruct (rand_error_ineq with "[$][$]") as "%". + iDestruct (rand_tapes_agree with "[$][$]") as "%". + iModIntro. iFrame. + by iPureIntro. + - iModIntro. + iIntros (?[|sample[|]]?) "([??]&?&%&?)"; try done. + destruct (Rlt_decision (ε2 sample + (ε' - ε1))%R 1%R). + + iRight. + iMod (rand_error_decrease with "[$][$]") as "?". + unshelve iMod (rand_error_increase _ _ (mknonnegreal (ε2 sample) _) with "[$]") as "[? ?]"; first done. + * simpl. lra. + * simpl. iFrame. + iDestruct (rand_tapes_valid with "[$]") as "%". + iMod (rand_tapes_update with "[$][$]") as "[$?]". + -- rewrite Forall_app; split; first done. + rewrite Forall_singleton. + pose proof fin_to_nat_lt sample; simpl. lia. + -- iFrame. + iModIntro. iSplit; last done. + replace (_-_+_)%R with (ε2 sample + (ε' - ε1))%R; first done. + lra. + + iLeft. + iPureIntro. lra. + - iModIntro. iFrame. + Qed. + + Lemma wp_presample_adv_comp_rand_tape' (N : nat) NS E α ns (ε1 : R) (ε2 : fin (S N) -> R) γ2: + ↑NS ⊆ E -> + (∀ n, 0<=ε2 n)%R -> + (SeriesC (λ n, (1 / (S N)) * ε2 n)%R <= ε1)%R → + is_rand' (L:=L') NS γ2 -∗ + ▷ rand_tapes_frag' (L:=L') γ2 α (N, ns) -∗ + ↯ ε1 -∗ + wp_update E (∃ n, ↯ (ε2 n) ∗ rand_tapes_frag' (L:=L') γ2 α (N, ns ++[fin_to_nat n]))%I. + Proof. + iIntros (Hsubset Hpos Hineq) "#Hinv >Htape Herr". + iDestruct (ec_valid with "[$]") as "%Hpos'". + destruct Hpos' as [Hpos' ?]. + iMod (rand_presample_spec' _ _ _ _ _ (λ l, match l with + | [x] => ε2 x + | _ => 1%R + end + ) + (rand_tapes_frag' (L:=L') γ2 α (N, ns) ∗ ↯ ε1) 1%nat + (λ ls, ∃ n, ⌜ls=[n]⌝ ∗ ↯ (ε2 n) ∗ rand_tapes_frag' γ2 α (N, ns ++ [fin_to_nat n]))%I + _ (mknonnegreal ε1 Hpos') + with "[//][][][$Htape $Herr]") as "(%l&%sample & -> & Herr &Htape)". + - done. + - by intros [|?[|]]. + - etrans; last eapply Hineq. + etrans; last eapply (SeriesC_le_inj _ (λ l, match l with |[x] => Some x | _ => None end)). + + rewrite Rdiv_def -SeriesC_scal_r. apply Req_le_sym. + apply SeriesC_ext. + intros. case_match; subst. + * rewrite bool_decide_eq_false_2; first (simpl;lra). + by rewrite elem_of_enum_uniform_fin_list. + * case_match; last first. + { rewrite bool_decide_eq_false_2; first (simpl;lra). + by rewrite elem_of_enum_uniform_fin_list. } + rewrite bool_decide_eq_true_2; last by apply elem_of_enum_uniform_fin_list. + simpl. + rewrite Rdiv_def Rmult_1_r; lra. + + intros. apply Rmult_le_pos; last done. + apply Rdiv_INR_ge_0. + + intros. repeat case_match; by simplify_eq. + + apply ex_seriesC_finite. + - iModIntro. + iIntros (??) "([??]&?&?)". + iDestruct (ec_supply_bound with "[$][$]") as "%". + iDestruct (rand_tapes_agree' with "[$][$]") as "%". + iModIntro. iFrame. + by iPureIntro. + - iModIntro. + iIntros (?[|sample[|]]?) "([??]&?&%&?)"; try done. + destruct (Rlt_decision (ε2 sample + (ε' - ε1))%R 1%R). + + iRight. + iMod (ec_supply_decrease with "[$][$]") as "(%&%&->&<-&?)". + unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 sample) _) with "[$]") as "[? ?]"; first done. + * simpl. simpl in *. lra. + * simpl. iFrame. + iDestruct (rand_tapes_valid' with "[$]") as "%". + iMod (rand_tapes_update' with "[$][$]") as "[$?]". + -- rewrite Forall_app; split; first done. + rewrite Forall_singleton. + pose proof fin_to_nat_lt sample; simpl. lia. + -- iFrame. + iModIntro. iSplit; last done. + simpl. + iPureIntro. lra. + + iLeft. + iPureIntro. simpl; simpl in *. lra. + - iModIntro. iFrame. + Qed. + +End checks. diff --git a/theories/coneris/wp_update.v b/theories/coneris/wp_update.v index fde19e70..a6a80753 100644 --- a/theories/coneris/wp_update.v +++ b/theories/coneris/wp_update.v @@ -201,6 +201,11 @@ Section wp_update. iApply HP. iFrame. Qed. + Global Instance is_except_0_pgl_wp E Q : IsExcept0 (wp_update E Q). + Proof. + by rewrite /IsExcept0 -{2}fupd_wp_update -except_0_fupd -fupd_intro. + Qed. + Global Instance elim_modal_fupd_wp p E P Q : ElimModal True p false (|={E}=> P) P (wp_update E Q) (wp_update E Q). Proof. From fc7ad54099d5937a41b15aa8bb4941cb9d82eb7b Mon Sep 17 00:00:00 2001 From: Hei Li Date: Wed, 25 Sep 2024 09:45:14 +0200 Subject: [PATCH 26/69] Slight improvement on spec --- theories/coneris/lib/hocap_rand.v | 40 +++++++++++++------------------ 1 file changed, 16 insertions(+), 24 deletions(-) diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index 9c9d8bc1..c0175e7b 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -106,16 +106,13 @@ Class rand_spec `{!conerisGS Σ} := RandSpec (∀ l, length l = num -> 0<= ε2 l)%R -> (SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R-> is_rand (L:=L) N γ1 γ2 -∗ - □(∀ (ε': R) m, P ∗rand_error_auth (L:=L) γ1 (ε')%R ∗ rand_tapes_auth (L:=L) γ2 m ={E∖↑N}=∗ - ⌜(ε<=ε')%R⌝ ∗ ⌜(m!!α = Some (tb, ns))⌝ ∗ - P ∗rand_error_auth (L:=L) γ1 (ε')%R ∗ rand_tapes_auth (L:=L) γ2 m - ) -∗ - (□∀ (ε':R) ns' m, (P ∗ rand_error_auth (L:=L) γ1 (ε')%R ∗ ⌜length ns' = num⌝ ∗ - rand_tapes_auth (L:=L) γ2 m - ) ={E∖↑N}=∗ - (⌜(1<=ε2 ns' + (ε'-ε))%R⌝ ∨ (rand_error_auth (L:=L) γ1 (ε2 ns' + (ε'-ε))%R ∗ - rand_tapes_auth (L:=L) γ2 (<[α:= (tb, ns++(fin_to_nat <$> ns'))]> m) ∗ T ns'))) - -∗ + (□∀ (ε':R) m, (P ∗ rand_error_auth (L:=L) γ1 (ε')%R ∗ rand_tapes_auth (L:=L) γ2 m ) ={E∖↑N}=∗ + ⌜(ε<=ε')%R⌝ ∗ ⌜(m!!α = Some (tb, ns))⌝ ∗ + (∀ ns', ⌜length ns' = num⌝ ={E∖↑N}=∗ + ⌜(1<=ε2 ns' + (ε'-ε))%R⌝ ∨ + (rand_error_auth (L:=L) γ1 (ε2 ns' + (ε'-ε))%R ∗ + rand_tapes_auth (L:=L) γ2 (<[α:= (tb, ns++(fin_to_nat <$> ns'))]> m) ∗ + T ns')))-∗ P -∗ wp_update E (∃ n, T n) }. @@ -281,11 +278,11 @@ Section impl. Qed. Next Obligation. simpl. - iIntros (??????????????? Hsubset Hpos Hineq) "#Hinv #Hchange #Hvs HP". + iIntros (??????????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP". iApply wp_update_state_step_coupl. iIntros (σ1 ε_supply) "((Hheap&Htapes)&Hε)". iMod (inv_acc with "Hinv") as "[>(%ε' & % & H1 & H2 & H3 & H4 ) Hclose]"; [done|]. - iMod ("Hchange" with "[$]") as "(%&%&HP & H2 &H4)". + iMod ("Hvs" with "[$]") as "(%&%&Hchange)". (* iDestruct (hocap_tapes_agree with "[$][$]") as "%K". *) iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. simpl. @@ -335,7 +332,7 @@ Next Obligation. iMod (tapeN_update_append' _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". (* iMod (hocap_tapes_update _ _ _ _ _ _ ((fin_to_nat <$> ns')++(fin_to_nat <$> sample)) with "[$][$]") as "[H4 Hfrag]". *) iMod "Hclose'" as "_". - iMod ("Hvs" $! ε'' sample with "[$HP $H2 $H4]") as "[%|(Hauth & H2 & HT)]"; first done. + iMod ("Hchange" $! sample with "[//]") as "[%|(Hauth & H2 & HT)]". { iExFalso. iApply (ec_contradict with "[$]"). exact. } iMod ("Hclose" with "[$Hε $Hauth Htape H3 $H2]") as "_". { iNext. @@ -418,14 +415,9 @@ Class rand_spec' `{!conerisGS Σ} := RandSpec' (∀ l, length l = num -> 0<= ε2 l)%R -> (SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R-> is_rand' (L:=L) N γ2 -∗ - □(∀ (ε': nonnegreal) m, P ∗ err_interp ε' ∗ rand_tapes_auth' (L:=L) γ2 m ={E∖↑N}=∗ - ⌜(ε<=ε')%R⌝ ∗ ⌜(m!!α = Some (tb, ns))⌝ ∗ - P ∗err_interp ε' ∗ rand_tapes_auth' (L:=L) γ2 m - ) -∗ - (□∀ (ε': nonnegreal) ns' m, (P ∗ err_interp ε' ∗ ⌜length ns' = num⌝ ∗ - rand_tapes_auth' (L:=L) γ2 m - ) ={E∖↑N}=∗ - (⌜(1<=ε2 ns' + (ε'-ε))%R⌝ ∨ ∃ (ε_final:nonnegreal), + (□∀ (ε': nonnegreal) m, (P ∗ err_interp ε' ∗ rand_tapes_auth' (L:=L) γ2 m + ) ={E∖↑N}=∗ ⌜(ε<=ε')%R⌝ ∗ ⌜(m!!α = Some (tb, ns))⌝ ∗ + (∀ ns', ⌜length ns' = num⌝ ={E∖↑N}=∗ ⌜(1<=ε2 ns' + (ε'-ε))%R⌝ ∨ ∃ (ε_final:nonnegreal), ⌜(nonneg ε_final = ε2 ns' + (ε'-ε))%R ⌝ ∗ (err_interp ε_final ∗ rand_tapes_auth' (L:=L) γ2 (<[α:= (tb, ns++(fin_to_nat <$> ns'))]> m) ∗ T ns'))) @@ -524,11 +516,11 @@ Section impl'. Qed. Next Obligation. simpl. - iIntros (?????????????? Hsubset Hpos Hineq) "#Hinv #Hcheck #Hvs HP ". + iIntros (?????????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP ". iApply wp_update_state_step_coupl. iIntros (σ1 ε_supply) "((Hheap&Htapes)&Hε)". iMod (inv_acc with "Hinv") as "[>(% & H3 & H4 ) Hclose]"; [done|]. - iMod ("Hcheck" with "[$]") as "(%&%&HP&Hε& H4)". + iMod ("Hvs" with "[$]") as "(%&%&Hchange)". iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. simpl. iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". @@ -564,7 +556,7 @@ Next Obligation. rewrite -Heq. iMod (tapeN_update_append' _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". iMod "Hclose'" as "_". - iMod ("Hvs" $! _ sample with "[$HP $H4 $Hε]") as "Hor"; first done. + iMod ("Hchange" $! sample with "[//]") as "Hor". iDestruct "Hor" as "[%|(%ε_final & %Heq' & Hε & Hauth & HT)]". { exfalso. simpl in *. lra. } iMod ("Hclose" with "[$Hauth Htape H3 ]") as "_". From 1e707a70854302567cb0fa97b30fc8c507a131c7 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Wed, 25 Sep 2024 09:48:59 +0200 Subject: [PATCH 27/69] fix build --- theories/coneris/lib/hocap_rand.v | 20 ++++++++++---------- 1 file changed, 10 insertions(+), 10 deletions(-) diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index c0175e7b..ab33b702 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -655,7 +655,7 @@ Section checks. ) (rand_tapes_frag (L:=L) γ2 α (N, ns) ∗rand_error_frag (L:=L) γ1 ε1) 1%nat (λ ls, ∃ n, ⌜ls=[n]⌝ ∗ rand_error_frag γ1 (ε2 n) ∗ rand_tapes_frag γ2 α (N, ns ++ [fin_to_nat n]))%I - with "[//][][][$Htape $Herr]") as "(%l&%sample & -> & Herr &Htape)". + with "[//][][$Htape $Herr]") as "(%l&%sample & -> & Herr &Htape)". - done. - by intros [|?[|]]. - etrans; last exact. @@ -679,10 +679,10 @@ Section checks. iIntros (??) "([??]&?&?)". iDestruct (rand_error_ineq with "[$][$]") as "%". iDestruct (rand_tapes_agree with "[$][$]") as "%". - iModIntro. iFrame. - by iPureIntro. - - iModIntro. - iIntros (?[|sample[|]]?) "([??]&?&%&?)"; try done. + iModIntro. + iSplit; first (iPureIntro; lra). + iSplit; first done. + iIntros ([|sample[|]]?); try done. destruct (Rlt_decision (ε2 sample + (ε' - ε1))%R 1%R). + iRight. iMod (rand_error_decrease with "[$][$]") as "?". @@ -723,7 +723,7 @@ Section checks. (rand_tapes_frag' (L:=L') γ2 α (N, ns) ∗ ↯ ε1) 1%nat (λ ls, ∃ n, ⌜ls=[n]⌝ ∗ ↯ (ε2 n) ∗ rand_tapes_frag' γ2 α (N, ns ++ [fin_to_nat n]))%I _ (mknonnegreal ε1 Hpos') - with "[//][][][$Htape $Herr]") as "(%l&%sample & -> & Herr &Htape)". + with "[//][][$Htape $Herr]") as "(%l&%sample & -> & Herr &Htape)". - done. - by intros [|?[|]]. - etrans; last eapply Hineq. @@ -747,10 +747,10 @@ Section checks. iIntros (??) "([??]&?&?)". iDestruct (ec_supply_bound with "[$][$]") as "%". iDestruct (rand_tapes_agree' with "[$][$]") as "%". - iModIntro. iFrame. - by iPureIntro. - - iModIntro. - iIntros (?[|sample[|]]?) "([??]&?&%&?)"; try done. + iModIntro. + iSplit; first (iPureIntro; simpl; lra). + iSplit; first done. + iIntros ([|sample[|]]?); try done. destruct (Rlt_decision (ε2 sample + (ε' - ε1))%R 1%R). + iRight. iMod (ec_supply_decrease with "[$][$]") as "(%&%&->&<-&?)". From acfcf798f6f8700c66ccb9d4a5b47625b4fb0557 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Wed, 25 Sep 2024 10:09:45 +0200 Subject: [PATCH 28/69] NIT --- theories/coneris/lib/hocap_rand.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index ab33b702..cb3a437d 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -770,5 +770,5 @@ Section checks. iPureIntro. simpl; simpl in *. lra. - iModIntro. iFrame. Qed. - + End checks. From 48ada792b4758a46e803e51448dbde82bc0e78bb Mon Sep 17 00:00:00 2001 From: Hei Li Date: Wed, 25 Sep 2024 12:40:25 +0200 Subject: [PATCH 29/69] Nicer --- theories/coneris/lib/hocap_rand.v | 1403 +++++++++++++++++------------ 1 file changed, 813 insertions(+), 590 deletions(-) diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index cb3a437d..552b25f9 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -3,7 +3,6 @@ From clutch.coneris Require Import coneris hocap. Set Default Proof Using "Type*". - Class rand_spec `{!conerisGS Σ} := RandSpec { (** * Operations *) @@ -13,48 +12,22 @@ Class rand_spec `{!conerisGS Σ} := RandSpec (** The assumptions about [Σ] *) randG : gFunctors → Type; - rand_error_name : Type; rand_tape_name: Type; (** * Predicates *) is_rand {L : randG Σ} (N:namespace) - (γ1: rand_error_name) (γ2: rand_tape_name): iProp Σ; - rand_error_auth {L : randG Σ} (γ: rand_error_name) (x:R): iProp Σ; - rand_error_frag {L : randG Σ} (γ: rand_error_name) (x:R): iProp Σ; + (γ1: rand_tape_name) : iProp Σ; rand_tapes_auth {L : randG Σ} (γ: rand_tape_name) (m:gmap loc (nat * list nat)): iProp Σ; rand_tapes_frag {L : randG Σ} (γ: rand_tape_name) (α:loc) (ns: (nat * list nat)): iProp Σ; (** * General properties of the predicates *) - #[global] is_rand_persistent {L : randG Σ} N γ1 γ2 :: - Persistent (is_rand (L:=L) N γ1 γ2); - #[global] rand_error_auth_timeless {L : randG Σ} γ x :: - Timeless (rand_error_auth (L:=L) γ x); - #[global] rand_error_frag_timeless {L : randG Σ} γ x :: - Timeless (rand_error_frag (L:=L) γ x); + #[global] is_rand_persistent {L : randG Σ} N γ1 :: + Persistent (is_rand (L:=L) N γ1); #[global] rand_tapes_auth_timeless {L : randG Σ} γ m :: Timeless (rand_tapes_auth (L:=L) γ m); #[global] rand_tapes_frag_timeless {L : randG Σ} γ α ns :: Timeless (rand_tapes_frag (L:=L) γ α ns); - #[global] rand_error_name_inhabited:: - Inhabited rand_error_name; - #[global] rand_tape_name_inhabited:: + #[global] rand_tape_name_inhabited :: Inhabited rand_tape_name; - rand_error_auth_exclusive {L : randG Σ} γ x1 x2: - rand_error_auth (L:=L) γ x1 -∗ rand_error_auth (L:=L) γ x2 -∗ False; - rand_error_frag_split {L : randG Σ} γ (x1 x2:nonnegreal): - rand_error_frag (L:=L) γ x1 ∗ rand_error_frag (L:=L) γ x2 ⊣⊢ - rand_error_frag (L:=L) γ (x1+x2)%R ; - rand_error_auth_valid {L : randG Σ} γ x: - rand_error_auth (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; - rand_error_frag_valid {L : randG Σ} γ x: - rand_error_frag (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; - rand_error_ineq {L : randG Σ} γ x1 x2: - rand_error_auth (L:=L) γ x1 -∗ rand_error_frag (L:=L) γ x2 -∗ ⌜(x2 <= x1)%R ⌝; - rand_error_decrease {L : randG Σ} γ (x1 x2:R): - rand_error_auth (L:=L) γ x1 -∗ rand_error_frag (L:=L) γ x2 ==∗ rand_error_auth (L:=L) γ (x1-x2)%R; - rand_error_increase {L : randG Σ} γ (x1:R) (x2:nonnegreal): - (x1+x2<1)%R -> ⊢ rand_error_auth (L:=L) γ x1 ==∗ - rand_error_auth (L:=L) γ (x1+x2) ∗ rand_error_frag (L:=L) γ x2; - rand_tapes_auth_exclusive {L : randG Σ} γ m m': rand_tapes_auth (L:=L) γ m -∗ rand_tapes_auth (L:=L) γ m' -∗ False; rand_tapes_frag_exclusive {L : randG Σ} γ α ns ns': @@ -63,69 +36,55 @@ Class rand_spec `{!conerisGS Σ} := RandSpec rand_tapes_auth (L:=L) γ m -∗ rand_tapes_frag (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; rand_tapes_valid {L : randG Σ} γ2 α tb ns: rand_tapes_frag (L:=L) γ2 α (tb, ns) -∗ - ⌜Forall (λ n, n<=tb)%nat ns⌝; - rand_tapes_update {L : randG Σ} γ α m ns ns': - Forall (λ n, n<=ns'.1)%nat ns'.2 -> + ⌜Forall (λ n, n<=tb)%nat ns⌝ ; + rand_tapes_update {L : randG Σ} γ α m ns ns': + Forall (λ x, x<=ns'.1) ns'.2 -> rand_tapes_auth (L:=L) γ m -∗ rand_tapes_frag (L:=L) γ α ns ==∗ rand_tapes_auth (L:=L) γ (<[α := ns']> m) ∗ rand_tapes_frag (L:=L) γ α ns'; (** * Program specs *) rand_inv_create_spec {L : randG Σ} N E: - ⊢ wp_update E (∃ γ1 γ2, is_rand (L:=L) N γ1 γ2); - rand_err_convert_spec1 {L : randG Σ} N E ε γ1 γ2: - ↑N ⊆ E-> - is_rand (L:=L) N γ1 γ2 -∗ - ↯ ε -∗ - wp_update E (rand_error_frag (L:=L) γ1 ε); - rand_err_convert_spec2 {L : randG Σ} N E ε γ1 γ2: - ↑N ⊆ E-> - is_rand (L:=L) N γ1 γ2 -∗ - rand_error_frag (L:=L) γ1 ε -∗ - wp_update E (↯ ε); - rand_allocate_tape_spec {L: randG Σ} N E (tb:nat) γ1 γ2: + ⊢ wp_update E (∃ γ1, is_rand (L:=L) N γ1); + rand_allocate_tape_spec {L: randG Σ} N E (tb:nat) γ2: ↑N ⊆ E-> - {{{ is_rand (L:=L) N γ1 γ2 }}} + {{{ is_rand (L:=L) N γ2 }}} rand_allocate_tape #tb @ E {{{ (v:val), RET v; ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ rand_tapes_frag (L:=L) γ2 α (tb, []) }}}; - rand_tape_spec_some {L: randG Σ} N E γ1 γ2 (P: iProp Σ) (Q:iProp Σ) (α:loc) (n:nat) (tb:nat) ns: + rand_tape_spec_some {L: randG Σ} N E γ2 (P: iProp Σ) (Q:iProp Σ) (α:loc) (n:nat) (tb:nat) ns: ↑N⊆E -> - {{{ is_rand (L:=L) N γ1 γ2 ∗ + {{{ is_rand (L:=L) N γ2 ∗ □ (∀ (m:gmap loc (nat * list nat)), P ∗ rand_tapes_auth (L:=L) γ2 m ={E∖↑N}=∗ ⌜m!!α=Some (tb, n::ns)⌝ ∗ Q ∗ rand_tapes_auth (L:=L) γ2 (<[α:=(tb, ns)]> m)) ∗ - ▷P + ▷ P }}} rand_tape #lbl:α #tb @ E - {{{ RET #n; Q }}}; + {{{ RET #n; Q }}}; rand_presample_spec {L: randG Σ} N E ns α (tb:nat) (ε2 : list (fin (S tb)) -> R) - (P : iProp Σ) num T γ1 γ2 ε : + (P : iProp Σ) num T γ2 (ε:nonnegreal) : ↑N ⊆ E -> (∀ l, length l = num -> 0<= ε2 l)%R -> (SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R-> - is_rand (L:=L) N γ1 γ2 -∗ - (□∀ (ε':R) m, (P ∗ rand_error_auth (L:=L) γ1 (ε')%R ∗ rand_tapes_auth (L:=L) γ2 m ) ={E∖↑N}=∗ - ⌜(ε<=ε')%R⌝ ∗ ⌜(m!!α = Some (tb, ns))⌝ ∗ - (∀ ns', ⌜length ns' = num⌝ ={E∖↑N}=∗ - ⌜(1<=ε2 ns' + (ε'-ε))%R⌝ ∨ - (rand_error_auth (L:=L) γ1 (ε2 ns' + (ε'-ε))%R ∗ - rand_tapes_auth (L:=L) γ2 (<[α:= (tb, ns++(fin_to_nat <$> ns'))]> m) ∗ - T ns')))-∗ - P -∗ + is_rand (L:=L) N γ2 -∗ + (□ (P ={E∖↑N, ∅}=∗ ↯ ε ∗ rand_tapes_frag (L:=L) γ2 α (tb, ns) ∗ + (∀ ns', ⌜length ns' = num⌝ ∗ ↯ (ε2 ns') ∗ rand_tapes_frag (L:=L) γ2 α (tb, ns ++ (fin_to_nat <$> ns')) + ={∅, E∖↑N}=∗ T ns'))) + -∗ + P -∗ wp_update E (∃ n, T n) }. Section impl. Local Opaque INR. (* (** Instantiate rand *) *) - Class randG1 Σ := RandG1 { flipG1_error::hocap_errorGS Σ; + Class randG1 Σ := RandG1 { flipG1_tapes:: hocap_tapesGS Σ; }. - Local Definition rand_inv_pred1 `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapesGS Σ} γ1 γ2:= - (∃ (ε:R) (m:gmap loc (nat * list nat)) , - ↯ ε ∗ ●↯ ε @ γ1 ∗ + Local Definition rand_inv_pred1 `{!conerisGS Σ, !hocap_tapesGS Σ} γ2:= + (∃ (m:gmap loc (nat * list nat)) , ([∗ map] α ↦ t ∈ m, α ↪N ( t.1 ; t.2 )) ∗ ●m@γ2 )%I. @@ -134,318 +93,10 @@ Section impl. {| rand_allocate_tape:= (λ: "N", alloc "N"); rand_tape:= (λ: "α" "N", rand("α") "N"); randG := randG1; - rand_error_name := gname; rand_tape_name := gname; - is_rand _ N γ1 γ2 := inv N (rand_inv_pred1 γ1 γ2); - rand_error_auth _ γ x := ●↯ x @ γ; - rand_error_frag _ γ x := ◯↯ x @ γ; + is_rand _ N γ2 := inv N (rand_inv_pred1 γ2); rand_tapes_auth _ γ m := (●m@γ)%I; - rand_tapes_frag _ γ α ns := ((α ◯↪N (ns.1; ns.2) @ γ) ∗ ⌜Forall (λ x, x<=ns.1) ns.2⌝)%I; - |}. - Next Obligation. - simpl. - iIntros. - iApply (hocap_error_auth_exclusive with "[$][$]"). - Qed. - Next Obligation. - simpl. - iIntros. - iApply (hocap_error_frag_split). - Qed. - Next Obligation. - simpl. - iIntros (?????) "H". - iApply (hocap_error_auth_valid with "[$]"). - Qed. - Next Obligation. - simpl. - iIntros (?????) "H". - iApply (hocap_error_frag_valid with "[$]"). - Qed. - Next Obligation. - simpl. - iIntros (??????) "H1 H2". - iApply (hocap_error_ineq with "[$][$]"). - Qed. - Next Obligation. - iIntros. - iApply (hocap_error_decrease with "[$][$]"). - Qed. - Next Obligation. - iIntros. - by iApply (hocap_error_increase with "[$]"). - Qed. - Next Obligation. - simpl. - iIntros (??????) "H1 H2". - by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[%H _]". - Qed. - Next Obligation. - simpl. - iIntros (???????) "[H1 %] [H2 %]". - iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. - rewrite dfrac_op_own dfrac_valid_own. by intro. - Qed. - Next Obligation. - simpl. - iIntros (??????[]) "? [??]". - by iDestruct (hocap_tapes_agree with "[$][$]") as "%H". - Qed. - Next Obligation. - simpl. - by iIntros (???????) "[? $]". - Qed. - Next Obligation. - iIntros (??????[??][??]?) "?[? %]". - iMod (hocap_tapes_update with "[$][$]") as "[??]". - by iFrame. - Qed. - Next Obligation. - simpl. - iIntros (?????). - iApply fupd_wp_update_ret. - unshelve iMod (hocap_error_alloc (nnreal_zero)) as "(%γ1 & ? & ?)"; simpl; [rewrite INR_0; lra|]. - iMod ec_zero as "?". - iMod (hocap_tapes_alloc ∅) as "(%γ2 & H4 & H5)". - iMod (inv_alloc _ _ (rand_inv_pred1 γ1 γ2) with "[-]"). - { iFrame. by iNext. } - by iFrame. - Qed. - Next Obligation. - simpl. - iIntros (????? ε ???) "#Hinv Herr". - iApply fupd_wp_update_ret. - iInv "Hinv" as ">(%ε'&%&H1&?&?&?)" "Hclose". - iDestruct (ec_valid with "[$Herr]") as "%". - iCombine "Herr H1" as "?". - iDestruct (ec_valid with "[$]") as "[% %]". - simpl in *. - unshelve iMod (hocap_error_increase _ _ (mknonnegreal ε _) with "[$]") as "[? Hfrag]"; [naive_solver|simpl;lra|]. - simpl. - iMod ("Hclose" with "[-Hfrag]") as "_"; last done. - iFrame. iApply (ec_eq with "[$]"). lra. - Qed. - Next Obligation. - simpl. - iIntros (????? ε ???) "#Hinv Herr". - iApply fupd_wp_update_ret. - iInv "Hinv" as ">(%ε'&%&H1&?&?&?)" "Hclose". - iDestruct (hocap_error_frag_valid with "[$]") as "%". - iDestruct (ec_valid with "[$]") as "%". - iDestruct (hocap_error_ineq with "[$][$]") as "%". - unshelve iMod (hocap_error_decrease with "[$][$]") as "?". - iDestruct (ec_eq with "[$]") as "?"; last first. - { iDestruct (ec_split with "[$]") as "[? H]"; last first. - - iMod ("Hclose" with "[-H]") as "_"; last done. - iFrame. - - naive_solver. - - lra. - } - lra. - Qed. - Next Obligation. - simpl. - iIntros (????????? Φ) "#Hinv HΦ". - wp_pures. - iInv "Hinv" as "(%&%&?&?&?&?)" "Hclose". - wp_apply (wp_alloc_tape); first done. - iIntros (α) "Hα". - iDestruct (hocap_tapes_notin with "[$][$]") as "%". - iMod (hocap_tapes_new _ _ α tb [] with "[$]") as "[?H]"; first done. - iMod ("Hclose" with "[-H HΦ]"). - { iModIntro. iFrame. - rewrite big_sepM_insert; [iFrame|done]. - } - iApply "HΦ". - by iFrame. - Qed. - Next Obligation. - simpl. - iIntros (?????????????? Φ) "(#Hinv & #Hvs & HP) HΦ". - wp_pures. - wp_bind (rand(_) _)%E. - iInv "Hinv" as ">(%&%&H1&H2&H3&H4)" "Hclose". - iMod ("Hvs" with "[$]") as "(%&HQ&Hauth)". - iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. - simpl. - wp_apply (wp_rand_tape with "[$]") as "[Htape %]". - iMod ("Hclose" with "[- HΦ HQ]") as "_". - { iExists _, (<[α:=_]> m). - iFrame. - iApply "H3". by iNext. - } - by iApply "HΦ". -Qed. -Next Obligation. - simpl. - iIntros (??????????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP". - iApply wp_update_state_step_coupl. - iIntros (σ1 ε_supply) "((Hheap&Htapes)&Hε)". - iMod (inv_acc with "Hinv") as "[>(%ε' & % & H1 & H2 & H3 & H4 ) Hclose]"; [done|]. - iMod ("Hvs" with "[$]") as "(%&%&Hchange)". - (* iDestruct (hocap_tapes_agree with "[$][$]") as "%K". *) - iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. - simpl. - iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". - iDestruct (ec_supply_bound with "[$][$]") as "%". - iMod (ec_supply_decrease with "[$][$]") as (ε'' ε_rem -> Hε') "Hε_supply". subst. - iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". - iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. - unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 x + (ε''-ε) else 1)%R _). - { case_match; last (simpl;lra). - apply Rplus_le_le_0_compat; last (simpl in *; lra). - apply Hpos. by apply Nat.eqb_eq. - } - iSplit. - { iPureIntro. - simpl. - rewrite (SeriesC_subset (λ x, (x∈enum_uniform_fin_list tb num))); last first. - - intros a. intros K. - rewrite elem_of_enum_uniform_fin_list in K. - case_match eqn :H4; last done. - by rewrite Nat.eqb_eq in H4. - - erewrite (SeriesC_ext _ (λ a, 1/S tb ^ num * (if bool_decide (a ∈ enum_uniform_fin_list tb num) then ε2 a else 0) + (if bool_decide (a ∈ enum_uniform_fin_list tb num) then 1/S tb ^ num * (ε''-ε) else 0)))%R; last first. - + intros n. - case_bool_decide; last lra. - replace (_ =? _) with true; last first. - * symmetry. rewrite Nat.eqb_eq. by rewrite -elem_of_enum_uniform_fin_list. - * simpl. lra. - + rewrite SeriesC_plus; [|apply ex_seriesC_scal_l, ex_seriesC_list|apply ex_seriesC_list..]. - rewrite SeriesC_list_2; last apply NoDup_enum_uniform_fin_list. - rewrite SeriesC_scal_l enum_uniform_fin_list_length pow_INR. - replace (1/_*_*_)%R with (ε''-ε)%R; first (simpl; lra). - rewrite Rmult_comm -Rmult_assoc Rdiv_1_l Rmult_inv_r; first lra. - apply pow_nonzero. - pose proof pos_INR_S tb; lra. - } - iIntros (sample) "<-". - destruct (Rlt_decision (nonneg ε_rem + (ε2 sample + (ε''-ε))%R) 1%R) as [Hdec|Hdec]; last first. - { apply Rnot_lt_ge, Rge_le in Hdec. - iApply state_step_coupl_ret_err_ge_1. - simpl. simpl in *. rewrite Nat.eqb_refl. lra. - } - iApply state_step_coupl_ret. - unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 sample + (ε''-ε))%R _) with "Hε_supply") as "[Hε_supply Hε]". - { apply Rplus_le_le_0_compat; simpl in *; [naive_solver|lra]. } - { simpl. done. } - simpl. - iMod (tapeN_update_append' _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". - (* iMod (hocap_tapes_update _ _ _ _ _ _ ((fin_to_nat <$> ns')++(fin_to_nat <$> sample)) with "[$][$]") as "[H4 Hfrag]". *) - iMod "Hclose'" as "_". - iMod ("Hchange" $! sample with "[//]") as "[%|(Hauth & H2 & HT)]". - { iExFalso. iApply (ec_contradict with "[$]"). exact. } - iMod ("Hclose" with "[$Hε $Hauth Htape H3 $H2]") as "_". - { iNext. - by iApply "H3". - } - iApply fupd_mask_intro_subseteq; first set_solver. - iFrame. - erewrite (nnreal_ext _ _); first done. - simpl. by rewrite Nat.eqb_refl. -Qed. - -End impl. - - -(** * EXPERIMENT *) - -Class rand_spec' `{!conerisGS Σ} := RandSpec' -{ - (** * Operations *) - rand_allocate_tape' : val; - rand_tape' : val; - (** * Ghost state *) - (** The assumptions about [Σ] *) - randG' : gFunctors → Type; - - rand_tape_name': Type; - (** * Predicates *) - is_rand' {L : randG' Σ} (N:namespace) - (γ1: rand_tape_name') : iProp Σ; - rand_tapes_auth' {L : randG' Σ} (γ: rand_tape_name') (m:gmap loc (nat * list nat)): iProp Σ; - rand_tapes_frag' {L : randG' Σ} (γ: rand_tape_name') (α:loc) (ns: (nat * list nat)): iProp Σ; - (** * General properties of the predicates *) - #[global] is_rand_persistent' {L : randG' Σ} N γ1 :: - Persistent (is_rand' (L:=L) N γ1); - #[global] rand_tapes_auth_timeless' {L : randG' Σ} γ m :: - Timeless (rand_tapes_auth' (L:=L) γ m); - #[global] rand_tapes_frag_timeless' {L : randG' Σ} γ α ns :: - Timeless (rand_tapes_frag' (L:=L) γ α ns); - #[global] rand_tape_name_inhabited' :: - Inhabited rand_tape_name'; - - rand_tapes_auth_exclusive' {L : randG' Σ} γ m m': - rand_tapes_auth' (L:=L) γ m -∗ rand_tapes_auth' (L:=L) γ m' -∗ False; - rand_tapes_frag_exclusive' {L : randG' Σ} γ α ns ns': - rand_tapes_frag' (L:=L) γ α ns -∗ rand_tapes_frag' (L:=L) γ α ns' -∗ False; - rand_tapes_agree' {L : randG' Σ} γ α m ns: - rand_tapes_auth' (L:=L) γ m -∗ rand_tapes_frag' (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; - rand_tapes_valid' {L : randG' Σ} γ2 α tb ns: - rand_tapes_frag' (L:=L) γ2 α (tb, ns) -∗ - ⌜Forall (λ n, n<=tb)%nat ns⌝ ; - rand_tapes_update' {L : randG' Σ} γ α m ns ns': - Forall (λ x, x<=ns'.1) ns'.2 -> - rand_tapes_auth' (L:=L) γ m -∗ rand_tapes_frag' (L:=L) γ α ns ==∗ - rand_tapes_auth' (L:=L) γ (<[α := ns']> m) ∗ rand_tapes_frag' (L:=L) γ α ns'; - - (** * Program specs *) - rand_inv_create_spec' {L : randG' Σ} N E: - ⊢ wp_update E (∃ γ1, is_rand' (L:=L) N γ1); - rand_allocate_tape_spec' {L: randG' Σ} N E (tb:nat) γ2: - ↑N ⊆ E-> - {{{ is_rand' (L:=L) N γ2 }}} - rand_allocate_tape' #tb @ E - {{{ (v:val), RET v; - ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ rand_tapes_frag' (L:=L) γ2 α (tb, []) - }}}; - rand_tape_spec_some' {L: randG' Σ} N E γ2 (P: iProp Σ) (Q:iProp Σ) (α:loc) (n:nat) (tb:nat) ns: - ↑N⊆E -> - {{{ is_rand' (L:=L) N γ2 ∗ - □ (∀ (m:gmap loc (nat * list nat)), P ∗ - rand_tapes_auth' (L:=L) γ2 m - ={E∖↑N}=∗ ⌜m!!α=Some (tb, n::ns)⌝ ∗ Q ∗ rand_tapes_auth' (L:=L) γ2 (<[α:=(tb, ns)]> m)) ∗ - ▷ P - }}} - rand_tape' #lbl:α #tb @ E - {{{ RET #n; Q }}}; - rand_presample_spec' {L: randG' Σ} N E ns α (tb:nat) - (ε2 : list (fin (S tb)) -> R) - (P : iProp Σ) num T γ2 (ε:nonnegreal) : - ↑N ⊆ E -> - (∀ l, length l = num -> 0<= ε2 l)%R -> - (SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R-> - is_rand' (L:=L) N γ2 -∗ - (□∀ (ε': nonnegreal) m, (P ∗ err_interp ε' ∗ rand_tapes_auth' (L:=L) γ2 m - ) ={E∖↑N}=∗ ⌜(ε<=ε')%R⌝ ∗ ⌜(m!!α = Some (tb, ns))⌝ ∗ - (∀ ns', ⌜length ns' = num⌝ ={E∖↑N}=∗ ⌜(1<=ε2 ns' + (ε'-ε))%R⌝ ∨ ∃ (ε_final:nonnegreal), - ⌜(nonneg ε_final = ε2 ns' + (ε'-ε))%R ⌝ ∗ - (err_interp ε_final ∗ - rand_tapes_auth' (L:=L) γ2 (<[α:= (tb, ns++(fin_to_nat <$> ns'))]> m) ∗ T ns'))) - -∗ - P -∗ - wp_update E (∃ n, T n) -}. - -Section impl'. - Local Opaque INR. - (* (** Instantiate rand *) *) - Class randG1' Σ := RandG1' { - flipG1_tapes':: hocap_tapesGS Σ; - }. - Local Definition rand_inv_pred1' `{!conerisGS Σ, !hocap_tapesGS Σ} γ2:= - (∃ (m:gmap loc (nat * list nat)) , - ([∗ map] α ↦ t ∈ m, α ↪N ( t.1 ; t.2 )) ∗ - ●m@γ2 - )%I. - - #[local] Program Definition rand_spec1' `{!conerisGS Σ}: rand_spec' := - {| rand_allocate_tape':= (λ: "N", alloc "N"); - rand_tape':= (λ: "α" "N", rand("α") "N"); - randG' := randG1'; - rand_tape_name' := gname; - is_rand' _ N γ2 := inv N (rand_inv_pred1' γ2); - rand_tapes_auth' _ γ m := (●m@γ)%I; - rand_tapes_frag' _ γ α ns := (α ◯↪N (ns.1; ns.2) @ γ ∗ ⌜Forall (λ x, x<=ns.1) ns.2⌝)%I; + rand_tapes_frag _ γ α ns := (α ◯↪N (ns.1; ns.2) @ γ ∗ ⌜Forall (λ x, x<=ns.1) ns.2⌝)%I; |}. Next Obligation. simpl. @@ -477,7 +128,7 @@ Section impl'. iIntros (?????). iApply fupd_wp_update_ret. iMod (hocap_tapes_alloc ∅) as "(%γ2 & H4 & H5)". - iMod (inv_alloc _ _ (rand_inv_pred1' γ2) with "[-]"). + iMod (inv_alloc _ _ (rand_inv_pred1 γ2) with "[-]"). { iFrame. by iNext. } by iFrame. Qed. @@ -520,23 +171,23 @@ Next Obligation. iApply wp_update_state_step_coupl. iIntros (σ1 ε_supply) "((Hheap&Htapes)&Hε)". iMod (inv_acc with "Hinv") as "[>(% & H3 & H4 ) Hclose]"; [done|]. - iMod ("Hvs" with "[$]") as "(%&%&Hchange)". + iMod ("Hvs" with "[$]") as "(Herr & [Hfrag %] & Hrest)". + iDestruct (hocap_tapes_agree with "[$][$]") as "%". iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. - simpl. iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". - (* iDestruct (ec_supply_bound with "[$][$]") as "%". *) - (* iMod (ec_supply_decrease with "[$][$]") as (ε'' ε_rem -> Hε') "Hε_supply". subst. *) + iDestruct (ec_supply_bound with "[$][$]") as "%". + iMod (ec_supply_decrease with "[$][$]") as (ε'' ε_rem -> Heq') "Hε_supply". iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". - unshelve epose (ε_rem := mknonnegreal (ε_supply-ε) _). - { simpl in *. lra. } - replace (ε_supply) with (ε_rem + ε)%NNR; last first. - { rewrite /ε_rem. apply nnreal_ext. simpl. lra. } + (* unshelve epose (ε_rem := mknonnegreal (ε_supply-ε) _). *) + (* { simpl in *. lra. } *) + (* replace (ε_supply) with (ε_rem + ε)%NNR; last first. *) + (* { rewrite /ε_rem. apply nnreal_ext. simpl. lra. } *) iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 x else 1)%R _). { case_match; last (simpl;lra). apply Hpos. by rewrite -Nat.eqb_eq. } iSplit. { iPureIntro. - simpl. etrans; last exact. + simpl. rewrite Heq'. etrans; last exact. rewrite (Rdiv_def (SeriesC _)(S tb ^num)%R) -SeriesC_scal_r. apply SeriesC_le; last apply ex_seriesC_scal_r, ex_seriesC_list. intros. rewrite elem_of_enum_uniform_fin_list'. @@ -552,223 +203,795 @@ Next Obligation. simpl. simpl in *. rewrite Nat.eqb_refl. lra. } iApply state_step_coupl_ret. - simpl. + simpl. simpl in *. rewrite -Heq. iMod (tapeN_update_append' _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". iMod "Hclose'" as "_". - iMod ("Hchange" $! sample with "[//]") as "Hor". - iDestruct "Hor" as "[%|(%ε_final & %Heq' & Hε & Hauth & HT)]". - { exfalso. simpl in *. lra. } - iMod ("Hclose" with "[$Hauth Htape H3 ]") as "_". - { iNext. - by iApply "H3". + unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 sample) _) with "[$]") as "[Hsupply Herr]". + { naive_solver. } + { simpl; lra. } + iMod (hocap_tapes_update with "[$][$]") as "[Hauth Hfrag]". + iMod ("Hrest" $! sample with "[$Herr $Hfrag]") as "HT". + { iPureIntro; split; first done. + rewrite Forall_app; split; subst; first done. + eapply Forall_impl; first apply fin.fin_forall_leq. + simpl; intros; lia. } + iFrame. + iMod ("Hclose" with "[-Hsupply]") as "_". + { iNext. iFrame. by iApply "H3". } iApply fupd_mask_intro_subseteq; first set_solver. - iFrame. - erewrite (nnreal_ext (_+_)%NNR _); first done. - simpl. rewrite Nat.eqb_refl. simpl in *. lra. + iApply ec_supply_eq; last done. + simpl. rewrite Nat.eqb_refl. lra. Qed. -End impl'. +End impl. -Section checks. - Context `{H: conerisGS Σ, r1:@rand_spec Σ H, r2:@rand_spec' Σ H, L:randG Σ, L': randG' Σ}. +(* Class rand_spec `{!conerisGS Σ} := RandSpec *) +(* { *) +(* (** * Operations *) *) +(* rand_allocate_tape : val; *) +(* rand_tape : val; *) +(* (** * Ghost state *) *) +(* (** The assumptions about [Σ] *) *) +(* randG : gFunctors → Type; *) + +(* rand_error_name : Type; *) +(* rand_tape_name: Type; *) +(* (** * Predicates *) *) +(* is_rand {L : randG Σ} (N:namespace) *) +(* (γ1: rand_error_name) (γ2: rand_tape_name): iProp Σ; *) +(* rand_error_auth {L : randG Σ} (γ: rand_error_name) (x:R): iProp Σ; *) +(* rand_error_frag {L : randG Σ} (γ: rand_error_name) (x:R): iProp Σ; *) +(* rand_tapes_auth {L : randG Σ} (γ: rand_tape_name) (m:gmap loc (nat * list nat)): iProp Σ; *) +(* rand_tapes_frag {L : randG Σ} (γ: rand_tape_name) (α:loc) (ns: (nat * list nat)): iProp Σ; *) +(* (** * General properties of the predicates *) *) +(* #[global] is_rand_persistent {L : randG Σ} N γ1 γ2 :: *) +(* Persistent (is_rand (L:=L) N γ1 γ2); *) +(* #[global] rand_error_auth_timeless {L : randG Σ} γ x :: *) +(* Timeless (rand_error_auth (L:=L) γ x); *) +(* #[global] rand_error_frag_timeless {L : randG Σ} γ x :: *) +(* Timeless (rand_error_frag (L:=L) γ x); *) +(* #[global] rand_tapes_auth_timeless {L : randG Σ} γ m :: *) +(* Timeless (rand_tapes_auth (L:=L) γ m); *) +(* #[global] rand_tapes_frag_timeless {L : randG Σ} γ α ns :: *) +(* Timeless (rand_tapes_frag (L:=L) γ α ns); *) +(* #[global] rand_error_name_inhabited:: *) +(* Inhabited rand_error_name; *) +(* #[global] rand_tape_name_inhabited:: *) +(* Inhabited rand_tape_name; *) - Lemma wp_rand_alloc_tape NS (N:nat) E γ1 γ2 : - ↑NS ⊆ E -> - {{{ is_rand (L:=L) NS γ1 γ2 }}} rand_allocate_tape #N @ E {{{ α, RET #lbl:α; rand_tapes_frag (L:=L) γ2 α (N,[]) }}}. - Proof. - iIntros (Hsubset Φ) "#Hinv HΦ". - wp_apply (rand_allocate_tape_spec with "[//]"); first exact. - iIntros (?) "(%&->&?)". - by iApply "HΦ". - Qed. +(* rand_error_auth_exclusive {L : randG Σ} γ x1 x2: *) +(* rand_error_auth (L:=L) γ x1 -∗ rand_error_auth (L:=L) γ x2 -∗ False; *) +(* rand_error_frag_split {L : randG Σ} γ (x1 x2:nonnegreal): *) +(* rand_error_frag (L:=L) γ x1 ∗ rand_error_frag (L:=L) γ x2 ⊣⊢ *) +(* rand_error_frag (L:=L) γ (x1+x2)%R ; *) +(* rand_error_auth_valid {L : randG Σ} γ x: *) +(* rand_error_auth (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; *) +(* rand_error_frag_valid {L : randG Σ} γ x: *) +(* rand_error_frag (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; *) +(* rand_error_ineq {L : randG Σ} γ x1 x2: *) +(* rand_error_auth (L:=L) γ x1 -∗ rand_error_frag (L:=L) γ x2 -∗ ⌜(x2 <= x1)%R ⌝; *) +(* rand_error_decrease {L : randG Σ} γ (x1 x2:R): *) +(* rand_error_auth (L:=L) γ x1 -∗ rand_error_frag (L:=L) γ x2 ==∗ rand_error_auth (L:=L) γ (x1-x2)%R; *) +(* rand_error_increase {L : randG Σ} γ (x1:R) (x2:nonnegreal): *) +(* (x1+x2<1)%R -> ⊢ rand_error_auth (L:=L) γ x1 ==∗ *) +(* rand_error_auth (L:=L) γ (x1+x2) ∗ rand_error_frag (L:=L) γ x2; *) - Lemma wp_rand_alloc_tape' NS (N:nat) E γ2: - ↑NS ⊆ E -> - {{{ is_rand' (L:=L') NS γ2 }}} rand_allocate_tape' #N @ E {{{ α, RET #lbl:α; rand_tapes_frag' (L:=L') γ2 α (N,[]) }}}. - Proof. - iIntros (Hsubset Φ) "#Hinv HΦ". - wp_apply (rand_allocate_tape_spec' with "[//]"); first exact. - iIntros (?) "(%&->&?)". - by iApply "HΦ". - Qed. +(* rand_tapes_auth_exclusive {L : randG Σ} γ m m': *) +(* rand_tapes_auth (L:=L) γ m -∗ rand_tapes_auth (L:=L) γ m' -∗ False; *) +(* rand_tapes_frag_exclusive {L : randG Σ} γ α ns ns': *) +(* rand_tapes_frag (L:=L) γ α ns -∗ rand_tapes_frag (L:=L) γ α ns' -∗ False; *) +(* rand_tapes_agree {L : randG Σ} γ α m ns: *) +(* rand_tapes_auth (L:=L) γ m -∗ rand_tapes_frag (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; *) +(* rand_tapes_valid {L : randG Σ} γ2 α tb ns: *) +(* rand_tapes_frag (L:=L) γ2 α (tb, ns) -∗ *) +(* ⌜Forall (λ n, n<=tb)%nat ns⌝; *) +(* rand_tapes_update {L : randG Σ} γ α m ns ns': *) +(* Forall (λ n, n<=ns'.1)%nat ns'.2 -> *) +(* rand_tapes_auth (L:=L) γ m -∗ rand_tapes_frag (L:=L) γ α ns ==∗ *) +(* rand_tapes_auth (L:=L) γ (<[α := ns']> m) ∗ rand_tapes_frag (L:=L) γ α ns'; *) - Lemma wp_rand_tape_1 NS (N:nat) E γ1 γ2 n ns α: - ↑NS ⊆ E -> - {{{ is_rand (L:=L) NS γ1 γ2 ∗ ▷ rand_tapes_frag (L:=L) γ2 α (N, n :: ns) }}} - rand_tape #lbl:α #N@ E - {{{ RET #(LitInt n); rand_tapes_frag (L:=L) γ2 α (N, ns) ∗ ⌜n <= N⌝ }}}. - Proof. - iIntros (Hsubset Φ) "(#Hinv & Hfrag) HΦ". - wp_apply (rand_tape_spec_some _ _ _ _ _ (⌜n<=N⌝ ∗ rand_tapes_frag _ _ _)%I with "[Hfrag]"); first exact. - - iSplit; first done. iSplit; last iApply "Hfrag". - iModIntro. - iIntros (?) "[Hfrag Hauth]". - iDestruct (rand_tapes_agree with "[$][$]") as "%". - iDestruct (rand_tapes_valid with "[$Hfrag]") as "%K". - inversion K; subst. - iMod (rand_tapes_update _ _ _ _ (N, ns) with "[$][$]") as "[Hauth Hfrag]"; first done. - iFrame "Hauth". iModIntro. - iFrame. - by iPureIntro. - - iIntros "[??]". iApply "HΦ". - iFrame. - Qed. +(* (** * Program specs *) *) +(* rand_inv_create_spec {L : randG Σ} N E: *) +(* ⊢ wp_update E (∃ γ1 γ2, is_rand (L:=L) N γ1 γ2); *) +(* rand_err_convert_spec1 {L : randG Σ} N E ε γ1 γ2: *) +(* ↑N ⊆ E-> *) +(* is_rand (L:=L) N γ1 γ2 -∗ *) +(* ↯ ε -∗ *) +(* wp_update E (rand_error_frag (L:=L) γ1 ε); *) +(* rand_err_convert_spec2 {L : randG Σ} N E ε γ1 γ2: *) +(* ↑N ⊆ E-> *) +(* is_rand (L:=L) N γ1 γ2 -∗ *) +(* rand_error_frag (L:=L) γ1 ε -∗ *) +(* wp_update E (↯ ε); *) +(* rand_allocate_tape_spec {L: randG Σ} N E (tb:nat) γ1 γ2: *) +(* ↑N ⊆ E-> *) +(* {{{ is_rand (L:=L) N γ1 γ2 }}} *) +(* rand_allocate_tape #tb @ E *) +(* {{{ (v:val), RET v; *) +(* ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ rand_tapes_frag (L:=L) γ2 α (tb, []) *) +(* }}}; *) +(* rand_tape_spec_some {L: randG Σ} N E γ1 γ2 (P: iProp Σ) (Q:iProp Σ) (α:loc) (n:nat) (tb:nat) ns: *) +(* ↑N⊆E -> *) +(* {{{ is_rand (L:=L) N γ1 γ2 ∗ *) +(* □ (∀ (m:gmap loc (nat * list nat)), P ∗ *) +(* rand_tapes_auth (L:=L) γ2 m *) +(* ={E∖↑N}=∗ ⌜m!!α=Some (tb, n::ns)⌝ ∗ Q ∗ rand_tapes_auth (L:=L) γ2 (<[α:=(tb, ns)]> m)) ∗ *) +(* ▷P *) +(* }}} *) +(* rand_tape #lbl:α #tb @ E *) +(* {{{ RET #n; Q }}}; *) +(* rand_presample_spec {L: randG Σ} N E ns α (tb:nat) *) +(* (ε2 : list (fin (S tb)) -> R) *) +(* (P : iProp Σ) num T γ1 γ2 ε : *) +(* ↑N ⊆ E -> *) +(* (∀ l, length l = num -> 0<= ε2 l)%R -> *) +(* (SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R-> *) +(* is_rand (L:=L) N γ1 γ2 -∗ *) +(* (□∀ (ε':R) m, (P ∗ rand_error_auth (L:=L) γ1 (ε')%R ∗ rand_tapes_auth (L:=L) γ2 m ) ={E∖↑N}=∗ *) +(* ⌜(ε<=ε')%R⌝ ∗ ⌜(m!!α = Some (tb, ns))⌝ ∗ *) +(* (∀ ns', ⌜length ns' = num⌝ ={E∖↑N}=∗ *) +(* ⌜(1<=ε2 ns' + (ε'-ε))%R⌝ ∨ *) +(* (rand_error_auth (L:=L) γ1 (ε2 ns' + (ε'-ε))%R ∗ *) +(* rand_tapes_auth (L:=L) γ2 (<[α:= (tb, ns++(fin_to_nat <$> ns'))]> m) ∗ *) +(* T ns')))-∗ *) +(* P -∗ *) +(* wp_update E (∃ n, T n) *) +(* }. *) - Lemma wp_rand_tape_2 NS (N:nat) E γ2 n ns α: - ↑NS ⊆ E -> - {{{ is_rand' (L:=L') NS γ2 ∗ ▷ rand_tapes_frag' (L:=L') γ2 α (N, n :: ns) }}} - rand_tape' #lbl:α #N@ E - {{{ RET #(LitInt n); rand_tapes_frag' (L:=L') γ2 α (N, ns) ∗ ⌜n <= N⌝ }}}. - Proof. - iIntros (Hsubset Φ) "(#Hinv &>Hfrag) HΦ". - wp_apply (rand_tape_spec_some' _ _ _ _ (⌜n<=N⌝ ∗ rand_tapes_frag' _ _ _)%I with "[Hfrag]"); first exact. - - iSplit; first done. iSplit; last iApply "Hfrag". - iModIntro. - iIntros (?) "[Hfrag Hauth]". - iDestruct (rand_tapes_agree' with "[$][$]") as "%". - iDestruct (rand_tapes_valid' with "[$Hfrag]") as "%K". - inversion K; subst. - iMod (rand_tapes_update' _ _ _ _ (N, ns) with "[$][$]") as "[Hauth Hfrag]"; first done. - iFrame "Hauth". iModIntro. - iFrame. - by iPureIntro. - - iIntros "[??]". iApply "HΦ". - iFrame. - Qed. +(* Section impl. *) +(* Local Opaque INR. *) +(* (* (** Instantiate rand *) *) *) +(* Class randG1 Σ := RandG1 { flipG1_error::hocap_errorGS Σ; *) +(* flipG1_tapes:: hocap_tapesGS Σ; *) +(* }. *) +(* Local Definition rand_inv_pred1 `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapesGS Σ} γ1 γ2:= *) +(* (∃ (ε:R) (m:gmap loc (nat * list nat)) , *) +(* ↯ ε ∗ ●↯ ε @ γ1 ∗ *) +(* ([∗ map] α ↦ t ∈ m, α ↪N ( t.1 ; t.2 )) ∗ *) +(* ●m@γ2 *) +(* )%I. *) - Lemma wp_presample_adv_comp_rand_tape (N : nat) NS E α ns (ε1 : R) (ε2 : fin (S N) -> R) γ1 γ2: - ↑NS ⊆ E -> - (∀ n, 0<=ε2 n)%R -> - (SeriesC (λ n, (1 / (S N)) * ε2 n)%R <= ε1)%R → - is_rand (L:=L) NS γ1 γ2 -∗ - ▷ rand_tapes_frag (L:=L) γ2 α (N, ns) -∗ - rand_error_frag (L:=L) γ1 ε1 -∗ - wp_update E (∃ n, rand_error_frag (L:=L) γ1 (ε2 n) ∗ rand_tapes_frag (L:=L) γ2 α (N, ns ++[fin_to_nat n]))%I. - Proof. - iIntros (Hsubset Hpos Hineq) "#Hinv >Htape Herr". - iMod (rand_presample_spec _ _ _ _ _ (λ l, match l with - | [x] => ε2 x - | _ => 1%R - end - ) - (rand_tapes_frag (L:=L) γ2 α (N, ns) ∗rand_error_frag (L:=L) γ1 ε1) 1%nat - (λ ls, ∃ n, ⌜ls=[n]⌝ ∗ rand_error_frag γ1 (ε2 n) ∗ rand_tapes_frag γ2 α (N, ns ++ [fin_to_nat n]))%I - with "[//][][$Htape $Herr]") as "(%l&%sample & -> & Herr &Htape)". - - done. - - by intros [|?[|]]. - - etrans; last exact. - etrans; last eapply (SeriesC_le_inj _ (λ l, match l with |[x] => Some x | _ => None end)). - + rewrite Rdiv_def -SeriesC_scal_r. apply Req_le_sym. - apply SeriesC_ext. - intros. case_match; subst. - * rewrite bool_decide_eq_false_2; first (simpl;lra). - by rewrite elem_of_enum_uniform_fin_list. - * case_match; last first. - { rewrite bool_decide_eq_false_2; first (simpl;lra). - by rewrite elem_of_enum_uniform_fin_list. } - rewrite bool_decide_eq_true_2; last by apply elem_of_enum_uniform_fin_list. - simpl. - rewrite Rdiv_def Rmult_1_r; lra. - + intros. apply Rmult_le_pos; last done. - apply Rdiv_INR_ge_0. - + intros. repeat case_match; by simplify_eq. - + apply ex_seriesC_finite. - - iModIntro. - iIntros (??) "([??]&?&?)". - iDestruct (rand_error_ineq with "[$][$]") as "%". - iDestruct (rand_tapes_agree with "[$][$]") as "%". - iModIntro. - iSplit; first (iPureIntro; lra). - iSplit; first done. - iIntros ([|sample[|]]?); try done. - destruct (Rlt_decision (ε2 sample + (ε' - ε1))%R 1%R). - + iRight. - iMod (rand_error_decrease with "[$][$]") as "?". - unshelve iMod (rand_error_increase _ _ (mknonnegreal (ε2 sample) _) with "[$]") as "[? ?]"; first done. - * simpl. lra. - * simpl. iFrame. - iDestruct (rand_tapes_valid with "[$]") as "%". - iMod (rand_tapes_update with "[$][$]") as "[$?]". - -- rewrite Forall_app; split; first done. - rewrite Forall_singleton. - pose proof fin_to_nat_lt sample; simpl. lia. - -- iFrame. - iModIntro. iSplit; last done. - replace (_-_+_)%R with (ε2 sample + (ε' - ε1))%R; first done. - lra. - + iLeft. - iPureIntro. lra. - - iModIntro. iFrame. - Qed. +(* #[local] Program Definition rand_spec1 `{!conerisGS Σ}: rand_spec := *) +(* {| rand_allocate_tape:= (λ: "N", alloc "N"); *) +(* rand_tape:= (λ: "α" "N", rand("α") "N"); *) +(* randG := randG1; *) +(* rand_error_name := gname; *) +(* rand_tape_name := gname; *) +(* is_rand _ N γ1 γ2 := inv N (rand_inv_pred1 γ1 γ2); *) +(* rand_error_auth _ γ x := ●↯ x @ γ; *) +(* rand_error_frag _ γ x := ◯↯ x @ γ; *) +(* rand_tapes_auth _ γ m := (●m@γ)%I; *) +(* rand_tapes_frag _ γ α ns := ((α ◯↪N (ns.1; ns.2) @ γ) ∗ ⌜Forall (λ x, x<=ns.1) ns.2⌝)%I; *) +(* |}. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros. *) +(* iApply (hocap_error_auth_exclusive with "[$][$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros. *) +(* iApply (hocap_error_frag_split). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (?????) "H". *) +(* iApply (hocap_error_auth_valid with "[$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (?????) "H". *) +(* iApply (hocap_error_frag_valid with "[$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (??????) "H1 H2". *) +(* iApply (hocap_error_ineq with "[$][$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* iIntros. *) +(* iApply (hocap_error_decrease with "[$][$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* iIntros. *) +(* by iApply (hocap_error_increase with "[$]"). *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (??????) "H1 H2". *) +(* by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[%H _]". *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (???????) "[H1 %] [H2 %]". *) +(* iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. *) +(* rewrite dfrac_op_own dfrac_valid_own. by intro. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (??????[]) "? [??]". *) +(* by iDestruct (hocap_tapes_agree with "[$][$]") as "%H". *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* by iIntros (???????) "[? $]". *) +(* Qed. *) +(* Next Obligation. *) +(* iIntros (??????[??][??]?) "?[? %]". *) +(* iMod (hocap_tapes_update with "[$][$]") as "[??]". *) +(* by iFrame. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (?????). *) +(* iApply fupd_wp_update_ret. *) +(* unshelve iMod (hocap_error_alloc (nnreal_zero)) as "(%γ1 & ? & ?)"; simpl; [rewrite INR_0; lra|]. *) +(* iMod ec_zero as "?". *) +(* iMod (hocap_tapes_alloc ∅) as "(%γ2 & H4 & H5)". *) +(* iMod (inv_alloc _ _ (rand_inv_pred1 γ1 γ2) with "[-]"). *) +(* { iFrame. by iNext. } *) +(* by iFrame. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (????? ε ???) "#Hinv Herr". *) +(* iApply fupd_wp_update_ret. *) +(* iInv "Hinv" as ">(%ε'&%&H1&?&?&?)" "Hclose". *) +(* iDestruct (ec_valid with "[$Herr]") as "%". *) +(* iCombine "Herr H1" as "?". *) +(* iDestruct (ec_valid with "[$]") as "[% %]". *) +(* simpl in *. *) +(* unshelve iMod (hocap_error_increase _ _ (mknonnegreal ε _) with "[$]") as "[? Hfrag]"; [naive_solver|simpl;lra|]. *) +(* simpl. *) +(* iMod ("Hclose" with "[-Hfrag]") as "_"; last done. *) +(* iFrame. iApply (ec_eq with "[$]"). lra. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (????? ε ???) "#Hinv Herr". *) +(* iApply fupd_wp_update_ret. *) +(* iInv "Hinv" as ">(%ε'&%&H1&?&?&?)" "Hclose". *) +(* iDestruct (hocap_error_frag_valid with "[$]") as "%". *) +(* iDestruct (ec_valid with "[$]") as "%". *) +(* iDestruct (hocap_error_ineq with "[$][$]") as "%". *) +(* unshelve iMod (hocap_error_decrease with "[$][$]") as "?". *) +(* iDestruct (ec_eq with "[$]") as "?"; last first. *) +(* { iDestruct (ec_split with "[$]") as "[? H]"; last first. *) +(* - iMod ("Hclose" with "[-H]") as "_"; last done. *) +(* iFrame. *) +(* - naive_solver. *) +(* - lra. *) +(* } *) +(* lra. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (????????? Φ) "#Hinv HΦ". *) +(* wp_pures. *) +(* iInv "Hinv" as "(%&%&?&?&?&?)" "Hclose". *) +(* wp_apply (wp_alloc_tape); first done. *) +(* iIntros (α) "Hα". *) +(* iDestruct (hocap_tapes_notin with "[$][$]") as "%". *) +(* iMod (hocap_tapes_new _ _ α tb [] with "[$]") as "[?H]"; first done. *) +(* iMod ("Hclose" with "[-H HΦ]"). *) +(* { iModIntro. iFrame. *) +(* rewrite big_sepM_insert; [iFrame|done]. *) +(* } *) +(* iApply "HΦ". *) +(* by iFrame. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (?????????????? Φ) "(#Hinv & #Hvs & HP) HΦ". *) +(* wp_pures. *) +(* wp_bind (rand(_) _)%E. *) +(* iInv "Hinv" as ">(%&%&H1&H2&H3&H4)" "Hclose". *) +(* iMod ("Hvs" with "[$]") as "(%&HQ&Hauth)". *) +(* iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. *) +(* simpl. *) +(* wp_apply (wp_rand_tape with "[$]") as "[Htape %]". *) +(* iMod ("Hclose" with "[- HΦ HQ]") as "_". *) +(* { iExists _, (<[α:=_]> m). *) +(* iFrame. *) +(* iApply "H3". by iNext. *) +(* } *) +(* by iApply "HΦ". *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (??????????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP". *) +(* iApply wp_update_state_step_coupl. *) +(* iIntros (σ1 ε_supply) "((Hheap&Htapes)&Hε)". *) +(* iMod (inv_acc with "Hinv") as "[>(%ε' & % & H1 & H2 & H3 & H4 ) Hclose]"; [done|]. *) +(* iMod ("Hvs" with "[$]") as "(%&%&Hchange)". *) +(* (* iDestruct (hocap_tapes_agree with "[$][$]") as "%K". *) *) +(* iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. *) +(* simpl. *) +(* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". *) +(* iDestruct (ec_supply_bound with "[$][$]") as "%". *) +(* iMod (ec_supply_decrease with "[$][$]") as (ε'' ε_rem -> Hε') "Hε_supply". subst. *) +(* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) +(* iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. *) +(* unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 x + (ε''-ε) else 1)%R _). *) +(* { case_match; last (simpl;lra). *) +(* apply Rplus_le_le_0_compat; last (simpl in *; lra). *) +(* apply Hpos. by apply Nat.eqb_eq. *) +(* } *) +(* iSplit. *) +(* { iPureIntro. *) +(* simpl. *) +(* rewrite (SeriesC_subset (λ x, (x∈enum_uniform_fin_list tb num))); last first. *) +(* - intros a. intros K. *) +(* rewrite elem_of_enum_uniform_fin_list in K. *) +(* case_match eqn :H4; last done. *) +(* by rewrite Nat.eqb_eq in H4. *) +(* - erewrite (SeriesC_ext _ (λ a, 1/S tb ^ num * (if bool_decide (a ∈ enum_uniform_fin_list tb num) then ε2 a else 0) + (if bool_decide (a ∈ enum_uniform_fin_list tb num) then 1/S tb ^ num * (ε''-ε) else 0)))%R; last first. *) +(* + intros n. *) +(* case_bool_decide; last lra. *) +(* replace (_ =? _) with true; last first. *) +(* * symmetry. rewrite Nat.eqb_eq. by rewrite -elem_of_enum_uniform_fin_list. *) +(* * simpl. lra. *) +(* + rewrite SeriesC_plus; [|apply ex_seriesC_scal_l, ex_seriesC_list|apply ex_seriesC_list..]. *) +(* rewrite SeriesC_list_2; last apply NoDup_enum_uniform_fin_list. *) +(* rewrite SeriesC_scal_l enum_uniform_fin_list_length pow_INR. *) +(* replace (1/_*_*_)%R with (ε''-ε)%R; first (simpl; lra). *) +(* rewrite Rmult_comm -Rmult_assoc Rdiv_1_l Rmult_inv_r; first lra. *) +(* apply pow_nonzero. *) +(* pose proof pos_INR_S tb; lra. *) +(* } *) +(* iIntros (sample) "<-". *) +(* destruct (Rlt_decision (nonneg ε_rem + (ε2 sample + (ε''-ε))%R) 1%R) as [Hdec|Hdec]; last first. *) +(* { apply Rnot_lt_ge, Rge_le in Hdec. *) +(* iApply state_step_coupl_ret_err_ge_1. *) +(* simpl. simpl in *. rewrite Nat.eqb_refl. lra. *) +(* } *) +(* iApply state_step_coupl_ret. *) +(* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 sample + (ε''-ε))%R _) with "Hε_supply") as "[Hε_supply Hε]". *) +(* { apply Rplus_le_le_0_compat; simpl in *; [naive_solver|lra]. } *) +(* { simpl. done. } *) +(* simpl. *) +(* iMod (tapeN_update_append' _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". *) +(* (* iMod (hocap_tapes_update _ _ _ _ _ _ ((fin_to_nat <$> ns')++(fin_to_nat <$> sample)) with "[$][$]") as "[H4 Hfrag]". *) *) +(* iMod "Hclose'" as "_". *) +(* iMod ("Hchange" $! sample with "[//]") as "[%|(Hauth & H2 & HT)]". *) +(* { iExFalso. iApply (ec_contradict with "[$]"). exact. } *) +(* iMod ("Hclose" with "[$Hε $Hauth Htape H3 $H2]") as "_". *) +(* { iNext. *) +(* by iApply "H3". *) +(* } *) +(* iApply fupd_mask_intro_subseteq; first set_solver. *) +(* iFrame. *) +(* erewrite (nnreal_ext _ _); first done. *) +(* simpl. by rewrite Nat.eqb_refl. *) +(* Qed. *) - Lemma wp_presample_adv_comp_rand_tape' (N : nat) NS E α ns (ε1 : R) (ε2 : fin (S N) -> R) γ2: - ↑NS ⊆ E -> - (∀ n, 0<=ε2 n)%R -> - (SeriesC (λ n, (1 / (S N)) * ε2 n)%R <= ε1)%R → - is_rand' (L:=L') NS γ2 -∗ - ▷ rand_tapes_frag' (L:=L') γ2 α (N, ns) -∗ - ↯ ε1 -∗ - wp_update E (∃ n, ↯ (ε2 n) ∗ rand_tapes_frag' (L:=L') γ2 α (N, ns ++[fin_to_nat n]))%I. - Proof. - iIntros (Hsubset Hpos Hineq) "#Hinv >Htape Herr". - iDestruct (ec_valid with "[$]") as "%Hpos'". - destruct Hpos' as [Hpos' ?]. - iMod (rand_presample_spec' _ _ _ _ _ (λ l, match l with - | [x] => ε2 x - | _ => 1%R - end - ) - (rand_tapes_frag' (L:=L') γ2 α (N, ns) ∗ ↯ ε1) 1%nat - (λ ls, ∃ n, ⌜ls=[n]⌝ ∗ ↯ (ε2 n) ∗ rand_tapes_frag' γ2 α (N, ns ++ [fin_to_nat n]))%I - _ (mknonnegreal ε1 Hpos') - with "[//][][$Htape $Herr]") as "(%l&%sample & -> & Herr &Htape)". - - done. - - by intros [|?[|]]. - - etrans; last eapply Hineq. - etrans; last eapply (SeriesC_le_inj _ (λ l, match l with |[x] => Some x | _ => None end)). - + rewrite Rdiv_def -SeriesC_scal_r. apply Req_le_sym. - apply SeriesC_ext. - intros. case_match; subst. - * rewrite bool_decide_eq_false_2; first (simpl;lra). - by rewrite elem_of_enum_uniform_fin_list. - * case_match; last first. - { rewrite bool_decide_eq_false_2; first (simpl;lra). - by rewrite elem_of_enum_uniform_fin_list. } - rewrite bool_decide_eq_true_2; last by apply elem_of_enum_uniform_fin_list. - simpl. - rewrite Rdiv_def Rmult_1_r; lra. - + intros. apply Rmult_le_pos; last done. - apply Rdiv_INR_ge_0. - + intros. repeat case_match; by simplify_eq. - + apply ex_seriesC_finite. - - iModIntro. - iIntros (??) "([??]&?&?)". - iDestruct (ec_supply_bound with "[$][$]") as "%". - iDestruct (rand_tapes_agree' with "[$][$]") as "%". - iModIntro. - iSplit; first (iPureIntro; simpl; lra). - iSplit; first done. - iIntros ([|sample[|]]?); try done. - destruct (Rlt_decision (ε2 sample + (ε' - ε1))%R 1%R). - + iRight. - iMod (ec_supply_decrease with "[$][$]") as "(%&%&->&<-&?)". - unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 sample) _) with "[$]") as "[? ?]"; first done. - * simpl. simpl in *. lra. - * simpl. iFrame. - iDestruct (rand_tapes_valid' with "[$]") as "%". - iMod (rand_tapes_update' with "[$][$]") as "[$?]". - -- rewrite Forall_app; split; first done. - rewrite Forall_singleton. - pose proof fin_to_nat_lt sample; simpl. lia. - -- iFrame. - iModIntro. iSplit; last done. - simpl. - iPureIntro. lra. - + iLeft. - iPureIntro. simpl; simpl in *. lra. - - iModIntro. iFrame. - Qed. +(* End impl. *) + + +(* (** * EXPERIMENT *) *) + +(* Class rand_spec' `{!conerisGS Σ} := RandSpec' *) +(* { *) +(* (** * Operations *) *) +(* rand_allocate_tape' : val; *) +(* rand_tape' : val; *) +(* (** * Ghost state *) *) +(* (** The assumptions about [Σ] *) *) +(* randG' : gFunctors → Type; *) + +(* rand_tape_name': Type; *) +(* (** * Predicates *) *) +(* is_rand' {L : randG' Σ} (N:namespace) *) +(* (γ1: rand_tape_name') : iProp Σ; *) +(* rand_tapes_auth' {L : randG' Σ} (γ: rand_tape_name') (m:gmap loc (nat * list nat)): iProp Σ; *) +(* rand_tapes_frag' {L : randG' Σ} (γ: rand_tape_name') (α:loc) (ns: (nat * list nat)): iProp Σ; *) +(* (** * General properties of the predicates *) *) +(* #[global] is_rand_persistent' {L : randG' Σ} N γ1 :: *) +(* Persistent (is_rand' (L:=L) N γ1); *) +(* #[global] rand_tapes_auth_timeless' {L : randG' Σ} γ m :: *) +(* Timeless (rand_tapes_auth' (L:=L) γ m); *) +(* #[global] rand_tapes_frag_timeless' {L : randG' Σ} γ α ns :: *) +(* Timeless (rand_tapes_frag' (L:=L) γ α ns); *) +(* #[global] rand_tape_name_inhabited' :: *) +(* Inhabited rand_tape_name'; *) + +(* rand_tapes_auth_exclusive' {L : randG' Σ} γ m m': *) +(* rand_tapes_auth' (L:=L) γ m -∗ rand_tapes_auth' (L:=L) γ m' -∗ False; *) +(* rand_tapes_frag_exclusive' {L : randG' Σ} γ α ns ns': *) +(* rand_tapes_frag' (L:=L) γ α ns -∗ rand_tapes_frag' (L:=L) γ α ns' -∗ False; *) +(* rand_tapes_agree' {L : randG' Σ} γ α m ns: *) +(* rand_tapes_auth' (L:=L) γ m -∗ rand_tapes_frag' (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; *) +(* rand_tapes_valid' {L : randG' Σ} γ2 α tb ns: *) +(* rand_tapes_frag' (L:=L) γ2 α (tb, ns) -∗ *) +(* ⌜Forall (λ n, n<=tb)%nat ns⌝ ; *) +(* rand_tapes_update' {L : randG' Σ} γ α m ns ns': *) +(* Forall (λ x, x<=ns'.1) ns'.2 -> *) +(* rand_tapes_auth' (L:=L) γ m -∗ rand_tapes_frag' (L:=L) γ α ns ==∗ *) +(* rand_tapes_auth' (L:=L) γ (<[α := ns']> m) ∗ rand_tapes_frag' (L:=L) γ α ns'; *) + +(* (** * Program specs *) *) +(* rand_inv_create_spec' {L : randG' Σ} N E: *) +(* ⊢ wp_update E (∃ γ1, is_rand' (L:=L) N γ1); *) +(* rand_allocate_tape_spec' {L: randG' Σ} N E (tb:nat) γ2: *) +(* ↑N ⊆ E-> *) +(* {{{ is_rand' (L:=L) N γ2 }}} *) +(* rand_allocate_tape' #tb @ E *) +(* {{{ (v:val), RET v; *) +(* ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ rand_tapes_frag' (L:=L) γ2 α (tb, []) *) +(* }}}; *) +(* rand_tape_spec_some' {L: randG' Σ} N E γ2 (P: iProp Σ) (Q:iProp Σ) (α:loc) (n:nat) (tb:nat) ns: *) +(* ↑N⊆E -> *) +(* {{{ is_rand' (L:=L) N γ2 ∗ *) +(* □ (∀ (m:gmap loc (nat * list nat)), P ∗ *) +(* rand_tapes_auth' (L:=L) γ2 m *) +(* ={E∖↑N}=∗ ⌜m!!α=Some (tb, n::ns)⌝ ∗ Q ∗ rand_tapes_auth' (L:=L) γ2 (<[α:=(tb, ns)]> m)) ∗ *) +(* ▷ P *) +(* }}} *) +(* rand_tape' #lbl:α #tb @ E *) +(* {{{ RET #n; Q }}}; *) +(* rand_presample_spec' {L: randG' Σ} N E ns α (tb:nat) *) +(* (ε2 : list (fin (S tb)) -> R) *) +(* (P : iProp Σ) num T γ2 (ε:nonnegreal) : *) +(* ↑N ⊆ E -> *) +(* (∀ l, length l = num -> 0<= ε2 l)%R -> *) +(* (SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R-> *) +(* is_rand' (L:=L) N γ2 -∗ *) +(* (□∀ (ε': nonnegreal) m, (P ∗ err_interp ε' ∗ rand_tapes_auth' (L:=L) γ2 m *) +(* ) ={E∖↑N}=∗ ⌜(ε<=ε')%R⌝ ∗ ⌜(m!!α = Some (tb, ns))⌝ ∗ *) +(* (∀ ns', ⌜length ns' = num⌝ ={E∖↑N}=∗ ⌜(1<=ε2 ns' + (ε'-ε))%R⌝ ∨ ∃ (ε_final:nonnegreal), *) +(* ⌜(nonneg ε_final = ε2 ns' + (ε'-ε))%R ⌝ ∗ *) +(* (err_interp ε_final ∗ *) +(* rand_tapes_auth' (L:=L) γ2 (<[α:= (tb, ns++(fin_to_nat <$> ns'))]> m) ∗ T ns'))) *) +(* -∗ *) +(* P -∗ *) +(* wp_update E (∃ n, T n) *) +(* }. *) + +(* Section impl'. *) +(* Local Opaque INR. *) +(* (* (** Instantiate rand *) *) *) +(* Class randG1' Σ := RandG1' { *) +(* flipG1_tapes':: hocap_tapesGS Σ; *) +(* }. *) +(* Local Definition rand_inv_pred1' `{!conerisGS Σ, !hocap_tapesGS Σ} γ2:= *) +(* (∃ (m:gmap loc (nat * list nat)) , *) +(* ([∗ map] α ↦ t ∈ m, α ↪N ( t.1 ; t.2 )) ∗ *) +(* ●m@γ2 *) +(* )%I. *) + +(* #[local] Program Definition rand_spec1' `{!conerisGS Σ}: rand_spec' := *) +(* {| rand_allocate_tape':= (λ: "N", alloc "N"); *) +(* rand_tape':= (λ: "α" "N", rand("α") "N"); *) +(* randG' := randG1'; *) +(* rand_tape_name' := gname; *) +(* is_rand' _ N γ2 := inv N (rand_inv_pred1' γ2); *) +(* rand_tapes_auth' _ γ m := (●m@γ)%I; *) +(* rand_tapes_frag' _ γ α ns := (α ◯↪N (ns.1; ns.2) @ γ ∗ ⌜Forall (λ x, x<=ns.1) ns.2⌝)%I; *) +(* |}. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (??????) "H1 H2". *) +(* by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[%H _]". *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (???????) "[H1 %] [H2 %]". *) +(* iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. *) +(* rewrite dfrac_op_own dfrac_valid_own. by intro. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (??????[]) "?[? %]". *) +(* by iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* by iIntros (???????) "[? $]". *) +(* Qed. *) +(* Next Obligation. *) +(* iIntros (??????[][]?) "?[? %]". *) +(* iMod (hocap_tapes_update with "[$][$]") as "[??]". *) +(* by iFrame. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (?????). *) +(* iApply fupd_wp_update_ret. *) +(* iMod (hocap_tapes_alloc ∅) as "(%γ2 & H4 & H5)". *) +(* iMod (inv_alloc _ _ (rand_inv_pred1' γ2) with "[-]"). *) +(* { iFrame. by iNext. } *) +(* by iFrame. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (???????? Φ) "#Hinv HΦ". *) +(* wp_pures. *) +(* iInv "Hinv" as "(%&?&?)" "Hclose". *) +(* wp_apply (wp_alloc_tape); first done. *) +(* iIntros (α) "Hα". *) +(* iDestruct (hocap_tapes_notin with "[$][$]") as "%". *) +(* iMod (hocap_tapes_new _ _ α tb [] with "[$]") as "[?H]"; first done. *) +(* iMod ("Hclose" with "[-H HΦ]"). *) +(* { iModIntro. iFrame. *) +(* rewrite big_sepM_insert; [iFrame|done]. *) +(* } *) +(* iApply "HΦ". *) +(* by iFrame. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (????????????? Φ) "(#Hinv & #Hvs & HP) HΦ". *) +(* wp_pures. *) +(* wp_bind (rand(_) _)%E. *) +(* iInv "Hinv" as ">(%&H3&H4)" "Hclose". *) +(* iMod ("Hvs" with "[$]") as "(%&HQ&Hauth)". *) +(* iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. *) +(* simpl. *) +(* wp_apply (wp_rand_tape with "[$]") as "[Htape %]". *) +(* iMod ("Hclose" with "[- HΦ HQ]") as "_". *) +(* { iExists (<[α:=_]> m). *) +(* iFrame. *) +(* iApply "H3". by iNext. *) +(* } *) +(* by iApply "HΦ". *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (?????????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP ". *) +(* iApply wp_update_state_step_coupl. *) +(* iIntros (σ1 ε_supply) "((Hheap&Htapes)&Hε)". *) +(* iMod (inv_acc with "Hinv") as "[>(% & H3 & H4 ) Hclose]"; [done|]. *) +(* iMod ("Hvs" with "[$]") as "(%&%&Hchange)". *) +(* iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. *) +(* simpl. *) +(* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". *) +(* (* iDestruct (ec_supply_bound with "[$][$]") as "%". *) *) +(* (* iMod (ec_supply_decrease with "[$][$]") as (ε'' ε_rem -> Hε') "Hε_supply". subst. *) *) +(* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) +(* unshelve epose (ε_rem := mknonnegreal (ε_supply-ε) _). *) +(* { simpl in *. lra. } *) +(* replace (ε_supply) with (ε_rem + ε)%NNR; last first. *) +(* { rewrite /ε_rem. apply nnreal_ext. simpl. lra. } *) +(* iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. *) +(* unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 x else 1)%R _). *) +(* { case_match; last (simpl;lra). apply Hpos. by rewrite -Nat.eqb_eq. } *) +(* iSplit. *) +(* { iPureIntro. *) +(* simpl. etrans; last exact. *) +(* rewrite (Rdiv_def (SeriesC _)(S tb ^num)%R) -SeriesC_scal_r. *) +(* apply SeriesC_le; last apply ex_seriesC_scal_r, ex_seriesC_list. *) +(* intros. rewrite elem_of_enum_uniform_fin_list'. *) +(* case_match; last lra. *) +(* split; last lra. *) +(* apply Rmult_le_pos; last (apply Hpos; by rewrite -Nat.eqb_eq). *) +(* rewrite -pow_INR. apply Rdiv_INR_ge_0. *) +(* } *) +(* iIntros (sample) "<-". *) +(* destruct (Rlt_decision (nonneg ε_rem + (ε2 sample)%R) 1%R) as [Hdec|Hdec]; last first. *) +(* { apply Rnot_lt_ge, Rge_le in Hdec. *) +(* iApply state_step_coupl_ret_err_ge_1. *) +(* simpl. simpl in *. rewrite Nat.eqb_refl. lra. *) +(* } *) +(* iApply state_step_coupl_ret. *) +(* simpl. *) +(* rewrite -Heq. *) +(* iMod (tapeN_update_append' _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". *) +(* iMod "Hclose'" as "_". *) +(* iMod ("Hchange" $! sample with "[//]") as "Hor". *) +(* iDestruct "Hor" as "[%|(%ε_final & %Heq' & Hε & Hauth & HT)]". *) +(* { exfalso. simpl in *. lra. } *) +(* iMod ("Hclose" with "[$Hauth Htape H3 ]") as "_". *) +(* { iNext. *) +(* by iApply "H3". *) +(* } *) +(* iApply fupd_mask_intro_subseteq; first set_solver. *) +(* iFrame. *) +(* erewrite (nnreal_ext (_+_)%NNR _); first done. *) +(* simpl. rewrite Nat.eqb_refl. simpl in *. lra. *) +(* Qed. *) + +(* End impl'. *) + +(* Section checks. *) +(* Context `{H: conerisGS Σ, r1:@rand_spec Σ H, r2:@rand_spec' Σ H, L:randG Σ, L': randG' Σ}. *) + +(* Lemma wp_rand_alloc_tape NS (N:nat) E γ1 γ2 : *) +(* ↑NS ⊆ E -> *) +(* {{{ is_rand (L:=L) NS γ1 γ2 }}} rand_allocate_tape #N @ E {{{ α, RET #lbl:α; rand_tapes_frag (L:=L) γ2 α (N,[]) }}}. *) +(* Proof. *) +(* iIntros (Hsubset Φ) "#Hinv HΦ". *) +(* wp_apply (rand_allocate_tape_spec with "[//]"); first exact. *) +(* iIntros (?) "(%&->&?)". *) +(* by iApply "HΦ". *) +(* Qed. *) + +(* Lemma wp_rand_alloc_tape' NS (N:nat) E γ2: *) +(* ↑NS ⊆ E -> *) +(* {{{ is_rand' (L:=L') NS γ2 }}} rand_allocate_tape' #N @ E {{{ α, RET #lbl:α; rand_tapes_frag' (L:=L') γ2 α (N,[]) }}}. *) +(* Proof. *) +(* iIntros (Hsubset Φ) "#Hinv HΦ". *) +(* wp_apply (rand_allocate_tape_spec' with "[//]"); first exact. *) +(* iIntros (?) "(%&->&?)". *) +(* by iApply "HΦ". *) +(* Qed. *) + +(* Lemma wp_rand_tape_1 NS (N:nat) E γ1 γ2 n ns α: *) +(* ↑NS ⊆ E -> *) +(* {{{ is_rand (L:=L) NS γ1 γ2 ∗ ▷ rand_tapes_frag (L:=L) γ2 α (N, n :: ns) }}} *) +(* rand_tape #lbl:α #N@ E *) +(* {{{ RET #(LitInt n); rand_tapes_frag (L:=L) γ2 α (N, ns) ∗ ⌜n <= N⌝ }}}. *) +(* Proof. *) +(* iIntros (Hsubset Φ) "(#Hinv & Hfrag) HΦ". *) +(* wp_apply (rand_tape_spec_some _ _ _ _ _ (⌜n<=N⌝ ∗ rand_tapes_frag _ _ _)%I with "[Hfrag]"); first exact. *) +(* - iSplit; first done. iSplit; last iApply "Hfrag". *) +(* iModIntro. *) +(* iIntros (?) "[Hfrag Hauth]". *) +(* iDestruct (rand_tapes_agree with "[$][$]") as "%". *) +(* iDestruct (rand_tapes_valid with "[$Hfrag]") as "%K". *) +(* inversion K; subst. *) +(* iMod (rand_tapes_update _ _ _ _ (N, ns) with "[$][$]") as "[Hauth Hfrag]"; first done. *) +(* iFrame "Hauth". iModIntro. *) +(* iFrame. *) +(* by iPureIntro. *) +(* - iIntros "[??]". iApply "HΦ". *) +(* iFrame. *) +(* Qed. *) + +(* Lemma wp_rand_tape_2 NS (N:nat) E γ2 n ns α: *) +(* ↑NS ⊆ E -> *) +(* {{{ is_rand' (L:=L') NS γ2 ∗ ▷ rand_tapes_frag' (L:=L') γ2 α (N, n :: ns) }}} *) +(* rand_tape' #lbl:α #N@ E *) +(* {{{ RET #(LitInt n); rand_tapes_frag' (L:=L') γ2 α (N, ns) ∗ ⌜n <= N⌝ }}}. *) +(* Proof. *) +(* iIntros (Hsubset Φ) "(#Hinv &>Hfrag) HΦ". *) +(* wp_apply (rand_tape_spec_some' _ _ _ _ (⌜n<=N⌝ ∗ rand_tapes_frag' _ _ _)%I with "[Hfrag]"); first exact. *) +(* - iSplit; first done. iSplit; last iApply "Hfrag". *) +(* iModIntro. *) +(* iIntros (?) "[Hfrag Hauth]". *) +(* iDestruct (rand_tapes_agree' with "[$][$]") as "%". *) +(* iDestruct (rand_tapes_valid' with "[$Hfrag]") as "%K". *) +(* inversion K; subst. *) +(* iMod (rand_tapes_update' _ _ _ _ (N, ns) with "[$][$]") as "[Hauth Hfrag]"; first done. *) +(* iFrame "Hauth". iModIntro. *) +(* iFrame. *) +(* by iPureIntro. *) +(* - iIntros "[??]". iApply "HΦ". *) +(* iFrame. *) +(* Qed. *) + +(* Lemma wp_presample_adv_comp_rand_tape (N : nat) NS E α ns (ε1 : R) (ε2 : fin (S N) -> R) γ1 γ2: *) +(* ↑NS ⊆ E -> *) +(* (∀ n, 0<=ε2 n)%R -> *) +(* (SeriesC (λ n, (1 / (S N)) * ε2 n)%R <= ε1)%R → *) +(* is_rand (L:=L) NS γ1 γ2 -∗ *) +(* ▷ rand_tapes_frag (L:=L) γ2 α (N, ns) -∗ *) +(* rand_error_frag (L:=L) γ1 ε1 -∗ *) +(* wp_update E (∃ n, rand_error_frag (L:=L) γ1 (ε2 n) ∗ rand_tapes_frag (L:=L) γ2 α (N, ns ++[fin_to_nat n]))%I. *) +(* Proof. *) +(* iIntros (Hsubset Hpos Hineq) "#Hinv >Htape Herr". *) +(* iMod (rand_presample_spec _ _ _ _ _ (λ l, match l with *) +(* | [x] => ε2 x *) +(* | _ => 1%R *) +(* end *) +(* ) *) +(* (rand_tapes_frag (L:=L) γ2 α (N, ns) ∗rand_error_frag (L:=L) γ1 ε1) 1%nat *) +(* (λ ls, ∃ n, ⌜ls=[n]⌝ ∗ rand_error_frag γ1 (ε2 n) ∗ rand_tapes_frag γ2 α (N, ns ++ [fin_to_nat n]))%I *) +(* with "[//][][$Htape $Herr]") as "(%l&%sample & -> & Herr &Htape)". *) +(* - done. *) +(* - by intros [|?[|]]. *) +(* - etrans; last exact. *) +(* etrans; last eapply (SeriesC_le_inj _ (λ l, match l with |[x] => Some x | _ => None end)). *) +(* + rewrite Rdiv_def -SeriesC_scal_r. apply Req_le_sym. *) +(* apply SeriesC_ext. *) +(* intros. case_match; subst. *) +(* * rewrite bool_decide_eq_false_2; first (simpl;lra). *) +(* by rewrite elem_of_enum_uniform_fin_list. *) +(* * case_match; last first. *) +(* { rewrite bool_decide_eq_false_2; first (simpl;lra). *) +(* by rewrite elem_of_enum_uniform_fin_list. } *) +(* rewrite bool_decide_eq_true_2; last by apply elem_of_enum_uniform_fin_list. *) +(* simpl. *) +(* rewrite Rdiv_def Rmult_1_r; lra. *) +(* + intros. apply Rmult_le_pos; last done. *) +(* apply Rdiv_INR_ge_0. *) +(* + intros. repeat case_match; by simplify_eq. *) +(* + apply ex_seriesC_finite. *) +(* - iModIntro. *) +(* iIntros (??) "([??]&?&?)". *) +(* iDestruct (rand_error_ineq with "[$][$]") as "%". *) +(* iDestruct (rand_tapes_agree with "[$][$]") as "%". *) +(* iModIntro. *) +(* iSplit; first (iPureIntro; lra). *) +(* iSplit; first done. *) +(* iIntros ([|sample[|]]?); try done. *) +(* destruct (Rlt_decision (ε2 sample + (ε' - ε1))%R 1%R). *) +(* + iRight. *) +(* iMod (rand_error_decrease with "[$][$]") as "?". *) +(* unshelve iMod (rand_error_increase _ _ (mknonnegreal (ε2 sample) _) with "[$]") as "[? ?]"; first done. *) +(* * simpl. lra. *) +(* * simpl. iFrame. *) +(* iDestruct (rand_tapes_valid with "[$]") as "%". *) +(* iMod (rand_tapes_update with "[$][$]") as "[$?]". *) +(* -- rewrite Forall_app; split; first done. *) +(* rewrite Forall_singleton. *) +(* pose proof fin_to_nat_lt sample; simpl. lia. *) +(* -- iFrame. *) +(* iModIntro. iSplit; last done. *) +(* replace (_-_+_)%R with (ε2 sample + (ε' - ε1))%R; first done. *) +(* lra. *) +(* + iLeft. *) +(* iPureIntro. lra. *) +(* - iModIntro. iFrame. *) +(* Qed. *) + +(* Lemma wp_presample_adv_comp_rand_tape' (N : nat) NS E α ns (ε1 : R) (ε2 : fin (S N) -> R) γ2: *) +(* ↑NS ⊆ E -> *) +(* (∀ n, 0<=ε2 n)%R -> *) +(* (SeriesC (λ n, (1 / (S N)) * ε2 n)%R <= ε1)%R → *) +(* is_rand' (L:=L') NS γ2 -∗ *) +(* ▷ rand_tapes_frag' (L:=L') γ2 α (N, ns) -∗ *) +(* ↯ ε1 -∗ *) +(* wp_update E (∃ n, ↯ (ε2 n) ∗ rand_tapes_frag' (L:=L') γ2 α (N, ns ++[fin_to_nat n]))%I. *) +(* Proof. *) +(* iIntros (Hsubset Hpos Hineq) "#Hinv >Htape Herr". *) +(* iDestruct (ec_valid with "[$]") as "%Hpos'". *) +(* destruct Hpos' as [Hpos' ?]. *) +(* iMod (rand_presample_spec' _ _ _ _ _ (λ l, match l with *) +(* | [x] => ε2 x *) +(* | _ => 1%R *) +(* end *) +(* ) *) +(* (rand_tapes_frag' (L:=L') γ2 α (N, ns) ∗ ↯ ε1) 1%nat *) +(* (λ ls, ∃ n, ⌜ls=[n]⌝ ∗ ↯ (ε2 n) ∗ rand_tapes_frag' γ2 α (N, ns ++ [fin_to_nat n]))%I *) +(* _ (mknonnegreal ε1 Hpos') *) +(* with "[//][][$Htape $Herr]") as "(%l&%sample & -> & Herr &Htape)". *) +(* - done. *) +(* - by intros [|?[|]]. *) +(* - etrans; last eapply Hineq. *) +(* etrans; last eapply (SeriesC_le_inj _ (λ l, match l with |[x] => Some x | _ => None end)). *) +(* + rewrite Rdiv_def -SeriesC_scal_r. apply Req_le_sym. *) +(* apply SeriesC_ext. *) +(* intros. case_match; subst. *) +(* * rewrite bool_decide_eq_false_2; first (simpl;lra). *) +(* by rewrite elem_of_enum_uniform_fin_list. *) +(* * case_match; last first. *) +(* { rewrite bool_decide_eq_false_2; first (simpl;lra). *) +(* by rewrite elem_of_enum_uniform_fin_list. } *) +(* rewrite bool_decide_eq_true_2; last by apply elem_of_enum_uniform_fin_list. *) +(* simpl. *) +(* rewrite Rdiv_def Rmult_1_r; lra. *) +(* + intros. apply Rmult_le_pos; last done. *) +(* apply Rdiv_INR_ge_0. *) +(* + intros. repeat case_match; by simplify_eq. *) +(* + apply ex_seriesC_finite. *) +(* - iModIntro. *) +(* iIntros (??) "([??]&?&?)". *) +(* iDestruct (ec_supply_bound with "[$][$]") as "%". *) +(* iDestruct (rand_tapes_agree' with "[$][$]") as "%". *) +(* iModIntro. *) +(* iSplit; first (iPureIntro; simpl; lra). *) +(* iSplit; first done. *) +(* iIntros ([|sample[|]]?); try done. *) +(* destruct (Rlt_decision (ε2 sample + (ε' - ε1))%R 1%R). *) +(* + iRight. *) +(* iMod (ec_supply_decrease with "[$][$]") as "(%&%&->&<-&?)". *) +(* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 sample) _) with "[$]") as "[? ?]"; first done. *) +(* * simpl. simpl in *. lra. *) +(* * simpl. iFrame. *) +(* iDestruct (rand_tapes_valid' with "[$]") as "%". *) +(* iMod (rand_tapes_update' with "[$][$]") as "[$?]". *) +(* -- rewrite Forall_app; split; first done. *) +(* rewrite Forall_singleton. *) +(* pose proof fin_to_nat_lt sample; simpl. lia. *) +(* -- iFrame. *) +(* iModIntro. iSplit; last done. *) +(* simpl. *) +(* iPureIntro. lra. *) +(* + iLeft. *) +(* iPureIntro. simpl; simpl in *. lra. *) +(* - iModIntro. iFrame. *) +(* Qed. *) -End checks. +(* End checks. *) From 66f966ac322a7163f2650d12165046eb4c82e3c0 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Wed, 25 Sep 2024 13:11:19 +0200 Subject: [PATCH 30/69] hocap rand with simons suggestion --- install_au_printer.sh.1 | 242 ++++++++++++++++++++++++++++++ theories/coneris/lib/hocap_rand.v | 89 +++++++++++ 2 files changed, 331 insertions(+) create mode 100644 install_au_printer.sh.1 diff --git a/install_au_printer.sh.1 b/install_au_printer.sh.1 new file mode 100644 index 00000000..0d78be18 --- /dev/null +++ b/install_au_printer.sh.1 @@ -0,0 +1,242 @@ +#!/bin/bash + +# +# Script collected by daleif (@math.au.dk), 2023 with ideas from mato@au.dk +# Long install phrase provided by AUIT +# + + +############################################################ +# Help # +############################################################ +Help() +{ + # Display Help + cat </dev/null; then + cat </dev/null; then + cat </dev/null; then + cat < 1 || "$PRINTER_CHECK" != "$PRINTER" ]] + then + echo "'$LIST' results for '$PRINTER', please be more specific" + BUILDING=$(echo "$PRINTER" | awk -F "-" '{print $1}') + echo "Available printers in $BUILDING:" + $SMB_CLIENT_CMD | grep "$BUILDING" + exit + else + echo "Printer ($PRINTER) is known to the print server, proceeding" + fi + +fi + + +# cat < + {{{ is_rand (L:=L) NS γ1 }}} rand_allocate_tape #N @ E {{{ α, RET #lbl:α; rand_tapes_frag (L:=L) γ1 α (N,[]) }}}. + Proof. + iIntros (Hsubset Φ) "#Hinv HΦ". + wp_apply (rand_allocate_tape_spec with "[//]"); first exact. + iIntros (?) "(%&->&?)". + by iApply "HΦ". + Qed. + + Lemma wp_rand_tape_1 NS (N:nat) E γ1 n ns α: + ↑NS ⊆ E -> + {{{ is_rand (L:=L) NS γ1 ∗ ▷ rand_tapes_frag (L:=L) γ1 α (N, n :: ns) }}} + rand_tape #lbl:α #N@ E + {{{ RET #(LitInt n); rand_tapes_frag (L:=L) γ1 α (N, ns) ∗ ⌜n <= N⌝ }}}. + Proof. + iIntros (Hsubset Φ) "(#Hinv &>Hfrag) HΦ". + wp_apply (rand_tape_spec_some _ _ _ _ (⌜n<=N⌝ ∗ rand_tapes_frag _ _ _)%I with "[Hfrag]"); first exact. + - iSplit; first done. iSplit; last iApply "Hfrag". + iModIntro. + iIntros (?) "[Hfrag Hauth]". + iDestruct (rand_tapes_agree with "[$][$]") as "%". + iDestruct (rand_tapes_valid with "[$Hfrag]") as "%K". + inversion K; subst. + iMod (rand_tapes_update _ _ _ _ (N, ns) with "[$][$]") as "[Hauth Hfrag]"; first done. + iFrame "Hauth". iModIntro. + iFrame. + by iPureIntro. + - iIntros "[??]". iApply "HΦ". + iFrame. + Qed. + + Lemma wp_presample_adv_comp_rand_tape (N : nat) NS E α ns (ε1 : R) (ε2 : fin (S N) -> R) γ1: + ↑NS ⊆ E -> + (∀ n, 0<=ε2 n)%R -> + (SeriesC (λ n, (1 / (S N)) * ε2 n)%R <= ε1)%R → + is_rand (L:=L) NS γ1 -∗ + ▷ rand_tapes_frag (L:=L) γ1 α (N, ns) -∗ + ↯ ε1 -∗ + wp_update E (∃ n, ↯ (ε2 n) ∗ rand_tapes_frag (L:=L) γ1 α (N, ns ++[fin_to_nat n]))%I. + Proof. + iIntros (Hsubset Hpos Hineq) "#Hinv >Htape Herr". + iDestruct (ec_valid with "[$]") as "%Hpos'". + destruct Hpos' as [Hpos' ?]. + iMod (rand_presample_spec _ _ _ _ _ (λ l, match l with + | [x] => ε2 x + | _ => 1%R + end + ) + (rand_tapes_frag (L:=L) γ1 α (N, ns) ∗ ↯ ε1) 1%nat + (λ ls, ∃ n, ⌜ls=[n]⌝ ∗ ↯ (ε2 n) ∗ rand_tapes_frag γ1 α (N, ns ++ [fin_to_nat n]))%I + _ (mknonnegreal ε1 Hpos') + with "[//][][$Htape $Herr]") as "(%l&%sample & -> & Herr &Htape)". + - done. + - by intros [|?[|]]. + - etrans; last eapply Hineq. + etrans; last eapply (SeriesC_le_inj _ (λ l, match l with |[x] => Some x | _ => None end)). + + rewrite Rdiv_def -SeriesC_scal_r. apply Req_le_sym. + apply SeriesC_ext. + intros. case_match; subst. + * rewrite bool_decide_eq_false_2; first (simpl;lra). + by rewrite elem_of_enum_uniform_fin_list. + * case_match; last first. + { rewrite bool_decide_eq_false_2; first (simpl;lra). + by rewrite elem_of_enum_uniform_fin_list. } + rewrite bool_decide_eq_true_2; last by apply elem_of_enum_uniform_fin_list. + simpl. + rewrite Rdiv_def Rmult_1_r; lra. + + intros. apply Rmult_le_pos; last done. + apply Rdiv_INR_ge_0. + + intros. repeat case_match; by simplify_eq. + + apply ex_seriesC_finite. + - iModIntro. + iIntros "(?&?)". + simpl. iFrame. + iApply fupd_mask_intro; first set_solver. + iIntros "Hclose". + iIntros ([|?[|]]) "(%&?&?)"; try done. + iMod "Hclose". + iFrame. + by iModIntro. + - iModIntro. iFrame. + Qed. + +End checks. + (* Class rand_spec `{!conerisGS Σ} := RandSpec *) (* { *) (* (** * Operations *) *) From dfceb3f097cf53d96c056b9cf88ee4987e565050 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Wed, 25 Sep 2024 13:12:26 +0200 Subject: [PATCH 31/69] remove file --- install_au_printer.sh.1 | 242 ---------------------------------------- 1 file changed, 242 deletions(-) delete mode 100644 install_au_printer.sh.1 diff --git a/install_au_printer.sh.1 b/install_au_printer.sh.1 deleted file mode 100644 index 0d78be18..00000000 --- a/install_au_printer.sh.1 +++ /dev/null @@ -1,242 +0,0 @@ -#!/bin/bash - -# -# Script collected by daleif (@math.au.dk), 2023 with ideas from mato@au.dk -# Long install phrase provided by AUIT -# - - -############################################################ -# Help # -############################################################ -Help() -{ - # Display Help - cat </dev/null; then - cat </dev/null; then - cat </dev/null; then - cat < 1 || "$PRINTER_CHECK" != "$PRINTER" ]] - then - echo "'$LIST' results for '$PRINTER', please be more specific" - BUILDING=$(echo "$PRINTER" | awk -F "-" '{print $1}') - echo "Available printers in $BUILDING:" - $SMB_CLIENT_CMD | grep "$BUILDING" - exit - else - echo "Printer ($PRINTER) is known to the print server, proceeding" - fi - -fi - - -# cat < Date: Wed, 25 Sep 2024 15:27:54 +0200 Subject: [PATCH 32/69] FLIP --- theories/coneris/lib/flip_spec.v | 431 ------------------------------ theories/coneris/lib/hocap_flip.v | 254 ++++++++++++++++++ theories/coneris/lib/hocap_rand.v | 53 ++-- 3 files changed, 278 insertions(+), 460 deletions(-) delete mode 100644 theories/coneris/lib/flip_spec.v create mode 100644 theories/coneris/lib/hocap_flip.v diff --git a/theories/coneris/lib/flip_spec.v b/theories/coneris/lib/flip_spec.v deleted file mode 100644 index c5f1bd71..00000000 --- a/theories/coneris/lib/flip_spec.v +++ /dev/null @@ -1,431 +0,0 @@ -(* From clutch.coneris Require Import coneris. *) -(* From clutch.coneris Require Import flip hocap. *) - -(* Set Default Proof Using "Type*". *) - -(* Class flip_spec `{!conerisGS Σ} := FlipSpec *) -(* { *) -(* (** * Operations *) *) -(* flip_allocate_tape : val; *) -(* flip_tape : val; *) -(* (** * Ghost state *) *) -(* (** The assumptions about [Σ] *) *) -(* flipG : gFunctors → Type; *) -(* (** [name] is used to associate [locked] with [is_lock] *) *) -(* flip_error_name : Type; *) -(* flip_tape_name: Type; *) -(* (** * Predicates *) *) -(* is_flip {L : flipG Σ} (N:namespace) *) -(* (γ1: flip_error_name) (γ2: flip_tape_name): iProp Σ; *) -(* flip_error_auth {L : flipG Σ} (γ: flip_error_name) (x:R): iProp Σ; *) -(* flip_error_frag {L : flipG Σ} (γ: flip_error_name) (x:R): iProp Σ; *) -(* flip_tapes_auth {L : flipG Σ} (γ: flip_tape_name) (m:gmap loc (list bool)): iProp Σ; *) -(* flip_tapes_frag {L : flipG Σ} (γ: flip_tape_name) (α:loc) (ns:list bool): iProp Σ; *) -(* (** * General properties of the predicates *) *) -(* #[global] is_flip_persistent {L : flipG Σ} N γ1 γ2 :: *) -(* Persistent (is_flip (L:=L) N γ1 γ2); *) -(* #[global] flip_error_auth_timeless {L : flipG Σ} γ x :: *) -(* Timeless (flip_error_auth (L:=L) γ x); *) -(* #[global] flip_error_frag_timeless {L : flipG Σ} γ x :: *) -(* Timeless (flip_error_frag (L:=L) γ x); *) -(* #[global] flip_tapes_auth_timeless {L : flipG Σ} γ m :: *) -(* Timeless (flip_tapes_auth (L:=L) γ m); *) -(* #[global] flip_tapes_frag_timeless {L : flipG Σ} γ α ns :: *) -(* Timeless (flip_tapes_frag (L:=L) γ α ns); *) -(* #[global] flip_error_name_inhabited:: *) -(* Inhabited flip_error_name; *) -(* #[global] flip_tape_name_inhabited:: *) -(* Inhabited flip_tape_name; *) - -(* flip_error_auth_exclusive {L : flipG Σ} γ x1 x2: *) -(* flip_error_auth (L:=L) γ x1 -∗ flip_error_auth (L:=L) γ x2 -∗ False; *) -(* flip_error_frag_split {L : flipG Σ} γ (x1 x2:nonnegreal): *) -(* flip_error_frag (L:=L) γ x1 ∗ flip_error_frag (L:=L) γ x2 ⊣⊢ *) -(* flip_error_frag (L:=L) γ (x1+x2)%R ; *) -(* flip_error_auth_valid {L : flipG Σ} γ x: *) -(* flip_error_auth (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; *) -(* flip_error_frag_valid {L : flipG Σ} γ x: *) -(* flip_error_frag (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; *) -(* flip_error_ineq {L : flipG Σ} γ x1 x2: *) -(* flip_error_auth (L:=L) γ x1 -∗ flip_error_frag (L:=L) γ x2 -∗ ⌜(x2 <= x1)%R ⌝; *) -(* flip_error_decrease {L : flipG Σ} γ (x1 x2:nonnegreal): *) -(* flip_error_auth (L:=L) γ (x1 + x2)%NNR -∗ flip_error_frag (L:=L) γ x1 ==∗ flip_error_auth (L:=L) γ x2; *) -(* flip_error_increase {L : flipG Σ} γ (x1 x2:nonnegreal): *) -(* (x1+x2<1)%R -> ⊢ flip_error_auth (L:=L) γ x1 ==∗ *) -(* flip_error_auth (L:=L) γ (x1+x2) ∗ flip_error_frag (L:=L) γ x2; *) - -(* flip_tapes_auth_exclusive {L : flipG Σ} γ m m': *) -(* flip_tapes_auth (L:=L) γ m -∗ flip_tapes_auth (L:=L) γ m' -∗ False; *) -(* flip_tapes_frag_exclusive {L : flipG Σ} γ α ns ns': *) -(* flip_tapes_frag (L:=L) γ α ns -∗ flip_tapes_frag (L:=L) γ α ns' -∗ False; *) -(* flip_tapes_agree {L : flipG Σ} γ α m ns: *) -(* flip_tapes_auth (L:=L) γ m -∗ flip_tapes_frag (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; *) -(* flip_tapes_update {L : flipG Σ} γ α m ns n: *) -(* flip_tapes_auth (L:=L) γ m -∗ flip_tapes_frag (L:=L) γ α ns ==∗ *) -(* flip_tapes_auth (L:=L) γ (<[α := (ns ++[n])]> m) ∗ flip_tapes_frag (L:=L) γ α (ns ++ [n]); *) -(* flip_tapes_pop {L : flipG Σ} γ α m ns n: *) -(* flip_tapes_auth (L:=L) γ m -∗ flip_tapes_frag (L:=L) γ α (n::ns) ==∗ *) -(* flip_tapes_auth (L:=L) γ (<[α := (ns)]> m) ∗ flip_tapes_frag (L:=L) γ α (ns); *) - -(* (** * Program specs *) *) -(* flip_inv_create_spec {L : flipG Σ} N E ε: *) -(* ↯ ε -∗ *) -(* wp_update E (∃ γ1 γ2, is_flip (L:=L) N γ1 γ2 ∗ *) -(* flip_error_frag (L:=L) γ1 ε); *) - -(* flip_allocate_tape_spec {L: flipG Σ} N E γ1 γ2: *) -(* ↑N ⊆ E-> *) -(* {{{ is_flip (L:=L) N γ1 γ2 }}} *) -(* flip_allocate_tape #() @ E *) -(* {{{ (v:val), RET v; *) -(* ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ flip_tapes_frag (L:=L) γ2 α [] *) -(* }}}; *) -(* flip_tape_spec_some {L: flipG Σ} N E γ1 γ2 (P: iProp Σ) (Q:iProp Σ) (α:loc) (n:bool) ns: *) -(* ↑N⊆E -> *) -(* {{{ is_flip (L:=L) N γ1 γ2 ∗ *) -(* □ (∀ m, P ∗ *) -(* flip_tapes_auth (L:=L) γ2 m *) -(* ={E∖↑N}=∗ ⌜m!!α=Some (n::ns)⌝ ∗ Q ∗ flip_tapes_auth (L:=L) γ2 (<[α:=ns]> m)) ∗ *) -(* P *) -(* }}} *) -(* flip_tape #lbl:α @ E *) -(* {{{ RET #n; Q }}}; *) -(* flip_presample_spec {L: flipG Σ} NS E ns α *) -(* (ε2 : R -> list bool -> R) *) -(* (P : _ -> iProp Σ) num T γ1 γ2 ε: *) -(* ↑NS ⊆ E -> *) -(* (∀ ε l, 0<= ε -> length l = num -> 0<=ε2 ε l)%R -> *) -(* (∀ (ε:R), 0<= ε -> SeriesC (λ l, if length l =? num then ε2 ε l else 0) /(2^num) <= ε)%R-> *) -(* is_flip (L:=L) NS γ1 γ2 -∗ *) -(* (□∀ (ε ε':nonnegreal) n, (P ε ∗ flip_error_auth (L:=L) γ1 (ε + ε')%R) ={E∖↑NS}=∗ *) -(* (⌜(1<=ε2 ε n)%R⌝ ∨ (flip_error_auth (L:=L) γ1 (ε2 ε n + ε')%R ∗ T ε n))) *) -(* -∗ *) -(* P ε -∗ flip_tapes_frag (L:=L) γ2 α ns-∗ *) -(* wp_update E (∃ n, T ε n ∗ flip_tapes_frag (L:=L) γ2 α (ns++n)) *) -(* }. *) - -(* Section test. *) -(* Context `{F:flip_spec}. *) -(* Lemma flip_presample_spec_simple {L: flipG Σ} NS E ns α *) -(* (ε2 : R -> bool -> R) *) -(* (P : iProp Σ) T γ1 γ2: *) -(* ↑NS ⊆ E -> *) -(* (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> *) -(* (∀ (ε:R), 0<= ε -> (ε2 ε true + ε2 ε false)/2 <= ε)%R-> *) -(* is_flip (L:=L) NS γ1 γ2 -∗ *) -(* (□∀ (ε:R) n, (P ∗ flip_error_auth (L:=L) γ1 ε) ={E∖↑NS}=∗ *) -(* (⌜(1<=ε2 ε n)%R⌝ ∨ (flip_error_auth (L:=L) γ1 (ε2 ε n) ∗ T (n)))) *) -(* -∗ *) -(* P -∗ flip_tapes_frag (L:=L) γ2 α ns-∗ *) -(* wp_update E (∃ n, T (n) ∗ flip_tapes_frag (L:=L) γ2 α (ns++[n])). *) -(* Proof. *) -(* iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) -(* pose (ε2' ε l:= match l with *) -(* | [b]=> ε2 ε b *) -(* | _ => 1%R *) -(* end *) -(* ). *) -(* iMod (flip_presample_spec _ _ _ _ ε2' P 1%nat (λ l, match l with | [b] => T b | _ => False end )%I with "[//][][$][$]") as "(%l & [??])"; first done. *) -(* - rewrite /ε2'. *) -(* intros. repeat case_match; try done. naive_solver. *) -(* - intros. *) -(* etrans; last apply Hineq; try done. *) -(* erewrite (SeriesC_subset (λ x, x ∈ [[true]; [false]])); last first. *) -(* + intros ? H'. *) -(* case_match eqn:K; last done. *) -(* rewrite Nat.eqb_eq in K. *) -(* exfalso. *) -(* apply H'. *) -(* destruct a as [|[|] [|]]; try (simpl in *; done). *) -(* all: set_solver. *) -(* + rewrite SeriesC_list; last first. *) -(* { rewrite !NoDup_cons; repeat split; last apply NoDup_nil; set_solver. } *) -(* simpl. lra. *) -(* - iModIntro. iIntros (ε n) "?". *) -(* destruct n as [|b [|]]. *) -(* + iLeft. iPureIntro. *) -(* by rewrite /ε2'. *) -(* + iMod ("Hvs" $! ε b with "[$]") as "[% | ?]". *) -(* * iLeft. iPureIntro. by rewrite /ε2'. *) -(* * iRight. by rewrite /ε2'. *) -(* + iLeft. iPureIntro. *) -(* by rewrite /ε2'. *) -(* - repeat case_match; try done. *) -(* iModIntro. *) -(* iFrame. *) -(* Qed. *) -(* End test. *) - - -(* (** Instantiate flip *) *) -(* Class flipG1 Σ := FlipG1 { flipG1_error::hocap_errorGS Σ; *) -(* flipG1_tapes:: hocap_tapesGS Σ; *) -(* }. *) -(* Local Definition flip_inv_pred1 `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapesGS Σ} γ1 γ2:= *) -(* (∃ (ε:R) (m:gmap loc (list bool)) , *) -(* ↯ ε ∗ ●↯ ε @ γ1 ∗ *) -(* ([∗ map] α ↦ t ∈ ((λ (ns:list bool), (1%nat, bool_to_nat<$>ns))<$>m), α ↪N ( t.1 ; t.2 )) ∗ *) -(* ●(((λ (ns:list bool), (1, bool_to_nat<$>ns))<$>m):gmap _ _)@γ2 *) -(* )%I. *) - -(* #[local] Program Definition flip_spec1 `{!conerisGS Σ}: flip_spec := *) -(* {| flip_allocate_tape:= (λ: <>, allocB); *) -(* flip_tape:= flipL; *) -(* flipG := flipG1; *) -(* flip_error_name := gname; *) -(* flip_tape_name := gname; *) -(* is_flip _ N γ1 γ2 := inv N (flip_inv_pred1 γ1 γ2); *) -(* flip_error_auth _ γ x := ●↯ x @ γ; *) -(* flip_error_frag _ γ x := ◯↯ x @ γ; *) -(* flip_tapes_auth _ γ m := (●((λ ns, (1, bool_to_nat<$>ns))<$>m)@γ)%I; *) -(* flip_tapes_frag _ γ α ns := (α◯↪N (1%nat;bool_to_nat<$> ns) @ γ)%I; *) -(* |}. *) - -(* Next Obligation. *) -(* simpl. *) -(* iIntros. *) -(* iApply (hocap_error_auth_exclusive with "[$][$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros. *) -(* iApply (hocap_error_frag_split). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (?????) "H". *) -(* iApply (hocap_error_auth_valid with "[$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (?????) "H". *) -(* iApply (hocap_error_frag_valid with "[$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (??????) "H1 H2". *) -(* iApply (hocap_error_ineq with "[$][$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* iIntros. *) -(* iApply (hocap_error_decrease with "[$][$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* iIntros. *) -(* by iApply (hocap_error_increase with "[$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (??????) "H1 H2". *) -(* by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[%H _]". *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (???????) "H1 H2". *) -(* iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. *) -(* rewrite dfrac_op_own dfrac_valid_own. by intro. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros. *) -(* iDestruct (hocap_tapes_agree with "[$][$]") as "%H". *) -(* iPureIntro. *) -(* rewrite lookup_fmap_Some in H. destruct H as (?&?&K). *) -(* simplify_eq. *) -(* rewrite K. *) -(* f_equal. *) -(* eapply fmap_inj; last done. *) -(* intros [][]?; by simplify_eq. *) -(* Qed. *) -(* Next Obligation. *) -(* iIntros. *) -(* iMod (hocap_tapes_presample with "[$][$]") as "[??]". *) -(* iModIntro. iFrame. *) -(* rewrite fmap_insert. *) -(* rewrite fmap_app; iFrame. *) -(* Qed. *) -(* Next Obligation. *) -(* iIntros. *) -(* iMod (hocap_tapes_pop with "[$][$]") as "[??]". *) -(* iModIntro. iFrame. *) -(* rewrite fmap_insert. *) -(* iFrame. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (????? ε) "H1". *) -(* iApply fupd_wp_update. *) -(* iApply wp_update_ret. *) -(* iDestruct (ec_valid with "[$]") as "%". *) -(* unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "(%γ1 & H2 & H3)"; simpl; [naive_solver..|]. *) -(* simpl. *) -(* iMod (hocap_tapes_alloc ∅) as "(%γ2 & H4 & H5)". *) -(* iMod (inv_alloc _ _ (flip_inv_pred1 γ1 γ2) with "[H1 H2 H4]"). *) -(* { iFrame. iNext. iExists ∅. *) -(* iFrame. *) -(* by rewrite fmap_empty. *) -(* } *) -(* by iFrame. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (???????? Φ) "#Hinv HΦ". *) -(* wp_pures. *) -(* rewrite /allocB. *) -(* iInv "Hinv" as "(%&%&?&?&?&?)" "Hclose". *) -(* wp_apply (wp_alloc_tape); first done. *) -(* iIntros (α) "Hα". *) -(* iDestruct (hocap_tapes_notin with "[$][$]") as "%". *) -(* iMod (hocap_tapes_new _ _ α 1%nat [] with "[$]") as "[?H]"; first done. *) -(* iMod ("Hclose" with "[-H HΦ]"). *) -(* { iModIntro. *) -(* iExists _, (<[α:=([])]> m). *) -(* iFrame. *) -(* rewrite fmap_insert. *) -(* rewrite big_sepM_insert; [iFrame|done]. *) -(* } *) -(* iApply "HΦ". *) -(* by iFrame. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (????????????? Φ) "(#Hinv & #Hvs & HP) HΦ". *) -(* rewrite /flipL. *) -(* wp_pures. *) -(* wp_bind (rand(_) _)%E. *) -(* iInv "Hinv" as ">(%&%&?&?&H3&?)" "Hclose". *) -(* iMod ("Hvs" with "[$]") as "(%&HQ&Hauth)". *) -(* erewrite <-(insert_delete m) at 1; last done. *) -(* rewrite fmap_insert. *) -(* rewrite big_sepM_insert; last first. *) -(* { rewrite fmap_delete. apply lookup_delete. } *) -(* iDestruct "H3" as "[Htape H3]". *) -(* simpl. *) -(* wp_apply (wp_rand_tape with "[$]") as "[Htape %]". *) -(* iMod ("Hclose" with "[- HΦ HQ]") as "_". *) -(* { iExists _, (<[α:=ns]> m). *) -(* iFrame. *) -(* rewrite <-(insert_delete m) at 2; last done. *) -(* rewrite insert_insert fmap_insert. *) -(* rewrite big_sepM_insert; first iFrame. *) -(* rewrite fmap_delete. *) -(* apply lookup_delete. *) -(* } *) -(* iModIntro. *) -(* wp_apply conversion.wp_int_to_bool as "_"; first done. *) -(* replace (Z_to_bool _) with n. *) -(* - iApply "HΦ"; iFrame. *) -(* - destruct n; simpl. *) -(* + rewrite Z_to_bool_neq_0; [done|lia]. *) -(* + by rewrite Z_to_bool_eq_0. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (????????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) -(* iApply wp_update_state_step_coupl. *) -(* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) -(* iMod (inv_acc with "Hinv") as "[>(% & % & H1 & H2 & H3 & H4 ) Hclose]"; [done|]. *) -(* iDestruct (hocap_tapes_agree with "[$][$]") as "%K". *) -(* rewrite lookup_fmap_Some in K. destruct K as (?&M&?). *) -(* simplify_eq. *) -(* unshelve epose proof fmap_inj _ _ _ _ M as ->. *) -(* { intros [][]?; by simplify_eq. } *) -(* erewrite <-(insert_delete m) at 1; last done. *) -(* rewrite fmap_insert. *) -(* rewrite big_sepM_insert; last first. *) -(* { rewrite fmap_delete. apply lookup_delete. } *) -(* simpl. *) -(* iDestruct "H3" as "[Htape H3]". *) -(* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". *) -(* iDestruct (ec_supply_bound with "[$][$]") as "%". *) -(* iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. *) -(* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) -(* assert (∀ (x: list (fin (2))), bool_to_nat <$> (nat_to_bool <$> (fin_to_nat <$> x)) = fin_to_nat <$> x) as Hrewrite. *) -(* { intros l. *) -(* rewrite -list_fmap_compose. *) -(* rewrite <-list_fmap_id. *) -(* apply list_fmap_ext. *) -(* intros ?? K. simpl. rewrite nat_to_bool_to_nat; first done. *) -(* apply list_lookup_fmap_inv in K as (x'&->&?). *) -(* pose proof fin_to_nat_lt x'. lia. } *) -(* iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. *) -(* unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 ε1' (nat_to_bool <$> (fin_to_nat <$> x)) else 1) _). *) -(* { case_match; last (simpl;lra). apply Hpos; first apply cond_nonneg. *) -(* rewrite !fmap_length. *) -(* by apply Nat.eqb_eq. *) -(* } *) -(* iSplit. *) -(* { iPureIntro. *) -(* simpl. *) -(* unshelve epose proof (Hineq ε1' _) as H'; first apply cond_nonneg. *) -(* etrans; last exact. *) -(* rewrite (Rdiv_def _ (2^_)) -SeriesC_scal_r. *) -(* etrans; last eapply (SeriesC_le_inj _ (λ x, Some (nat_to_bool <$> (fin_to_nat <$> x)))). *) -(* - apply SeriesC_le. *) -(* + simpl. intros n. *) -(* rewrite !fmap_length. *) -(* case_match. *) -(* * replace (1+1)%R with 2%R by lra. *) -(* rewrite -Rdiv_1_l. simpl. *) -(* split; last lra. *) -(* apply Rmult_le_pos. *) -(* -- apply Rcomplements.Rdiv_le_0_compat; first lra. *) -(* apply pow_lt; lra. *) -(* -- apply Hpos; first done. *) -(* rewrite !fmap_length. *) -(* by apply Nat.eqb_eq. *) -(* * lra. *) -(* + simpl. *) -(* apply (ex_seriesC_list_length _ num). *) -(* intros ?. rewrite !fmap_length. *) -(* case_match; last lra. *) -(* intros. by apply Nat.eqb_eq. *) -(* - intros. case_match; last lra. *) -(* apply Rmult_le_pos; first apply Hpos; simpl; auto. *) -(* + by apply Nat.eqb_eq. *) -(* + rewrite -Rdiv_1_l. *) -(* apply Rcomplements.Rdiv_le_0_compat; first lra. *) -(* apply pow_lt; lra. *) -(* - intros ??? <- K. *) -(* simplify_eq. *) -(* rewrite -!list_fmap_compose in K. *) -(* apply list_fmap_eq_inj in K; try done. *) -(* intros x x'. *) -(* repeat (inv_fin x; try intros x); *) -(* repeat (inv_fin x'; try intros x'); done. *) -(* - apply (ex_seriesC_list_length _ num). *) -(* intros ?. *) -(* case_match; last lra. *) -(* intros. by apply Nat.eqb_eq. *) -(* } *) -(* iIntros (sample) "<-". *) -(* destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' (nat_to_bool <$> (fin_to_nat <$> sample))))%R 1%R) as [Hdec|Hdec]; last first. *) -(* { apply Rnot_lt_ge, Rge_le in Hdec. *) -(* iApply state_step_coupl_ret_err_ge_1. *) -(* simpl. simpl in *. rewrite Nat.eqb_refl. lra. *) -(* } *) -(* iApply state_step_coupl_ret. *) -(* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' (nat_to_bool <$> (fin_to_nat <$> sample))) _) with "Hε_supply") as "[Hε_supply Hε]". *) -(* { apply Hpos; first apply cond_nonneg. by rewrite !fmap_length. } *) -(* { simpl. done. } *) -(* simpl. *) -(* iMod (tapeN_update_append' _ _ _ _ sample with "[$][Htape]") as "[Htapes Htape]". *) -(* { by erewrite Heq. } *) -(* iMod (hocap_tapes_presample' _ _ _ _ _ (fin_to_nat <$> sample) with "[$][$]") as "[H4 Hfrag]". *) -(* iMod "Hclose'" as "_". *) -(* iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". *) -(* { iExFalso. iApply (ec_contradict with "[$]"). exact. } *) -(* iMod ("Hclose" with "[$Hε $H2 Htape H3 H4]") as "_". *) -(* { iNext. *) -(* iExists (<[α:=(ns ++ (nat_to_bool<$>(fin_to_nat <$> sample)))]>m). *) -(* rewrite fmap_insert. *) -(* rewrite big_sepM_insert_delete Heq/=. *) -(* rewrite fmap_delete. iFrame. *) -(* rewrite fmap_app/=. *) -(* rewrite !Hrewrite. iFrame. *) -(* } *) -(* iApply fupd_mask_intro_subseteq; first set_solver. *) -(* iFrame. rewrite fmap_app/=Hrewrite. iFrame. *) -(* erewrite (nnreal_ext _ _); first done. *) -(* simpl. by rewrite Nat.eqb_refl. *) -(* Qed. *) diff --git a/theories/coneris/lib/hocap_flip.v b/theories/coneris/lib/hocap_flip.v new file mode 100644 index 00000000..68b4db4b --- /dev/null +++ b/theories/coneris/lib/hocap_flip.v @@ -0,0 +1,254 @@ +From clutch.coneris Require Import coneris. +From clutch.coneris Require Import flip hocap hocap_rand. + +Set Default Proof Using "Type*". + +Class flip_spec `{!conerisGS Σ} := FlipSpec +{ + (** * Operations *) + flip_allocate_tape : val; + flip_tape : val; + (** * Ghost state *) + (** The assumptions about [Σ] *) + flipG : gFunctors → Type; + (** [name] is used to associate [locked] with [is_lock] *) + flip_tape_name: Type; + (** * Predicates *) + is_flip {L : flipG Σ} (N:namespace) (γ2: flip_tape_name): iProp Σ; + flip_tapes_auth {L : flipG Σ} (γ: flip_tape_name) (m:gmap loc (list bool)): iProp Σ; + flip_tapes_frag {L : flipG Σ} (γ: flip_tape_name) (α:loc) (ns:list bool): iProp Σ; + (** * General properties of the predicates *) + #[global] is_flip_persistent {L : flipG Σ} N γ1 :: + Persistent (is_flip (L:=L) N γ1); + #[global] flip_tapes_auth_timeless {L : flipG Σ} γ m :: + Timeless (flip_tapes_auth (L:=L) γ m); + #[global] flip_tapes_frag_timeless {L : flipG Σ} γ α ns :: + Timeless (flip_tapes_frag (L:=L) γ α ns); + #[global] flip_tape_name_inhabited:: + Inhabited flip_tape_name; + + flip_tapes_auth_exclusive {L : flipG Σ} γ m m': + flip_tapes_auth (L:=L) γ m -∗ flip_tapes_auth (L:=L) γ m' -∗ False; + flip_tapes_frag_exclusive {L : flipG Σ} γ α ns ns': + flip_tapes_frag (L:=L) γ α ns -∗ flip_tapes_frag (L:=L) γ α ns' -∗ False; + flip_tapes_agree {L : flipG Σ} γ α m ns: + flip_tapes_auth (L:=L) γ m -∗ flip_tapes_frag (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; + flip_tapes_update {L : flipG Σ} γ α m ns ns': + flip_tapes_auth (L:=L) γ m -∗ flip_tapes_frag (L:=L) γ α ns ==∗ + flip_tapes_auth (L:=L) γ (<[α := ns']> m) ∗ flip_tapes_frag (L:=L) γ α ns'; + + (** * Program specs *) + flip_inv_create_spec {L : flipG Σ} N E: + ⊢ wp_update E (∃ γ1, is_flip (L:=L) N γ1); + + flip_allocate_tape_spec {L: flipG Σ} N E γ1: + ↑N ⊆ E-> + {{{ is_flip (L:=L) N γ1 }}} + flip_allocate_tape #() @ E + {{{ (v:val), RET v; + ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ flip_tapes_frag (L:=L) γ1 α [] + }}}; + + flip_tape_spec_some {L: flipG Σ} N E γ1 (P: iProp Σ) (Q:iProp Σ) (α:loc) (n:bool) ns: + ↑N⊆E -> + {{{ is_flip (L:=L) N γ1 ∗ + (□ (P ={E∖↑N, ∅}=∗ flip_tapes_frag (L:=L) γ1 α (n::ns) ∗ + (flip_tapes_frag (L:=L) γ1 α ns ={∅, E∖↑N}=∗ Q))) ∗ + ▷ P + }}} + flip_tape #lbl:α @ E + {{{ RET #n; Q }}}; + + flip_presample_spec {L: flipG Σ} NS E ns α + (ε2 : list bool -> R) + (P : iProp Σ) num T γ1 ε: + ↑NS ⊆ E -> + (∀ l, length l = num -> 0<=ε2 l)%R -> + (SeriesC (λ l, if length l =? num then ε2 l else 0) /(2^num) <= ε)%R-> + is_flip (L:=L) NS γ1 -∗ + (□ (P ={E∖↑NS, ∅}=∗ ↯ ε ∗ flip_tapes_frag (L:=L) γ1 α ns ∗ + (∀ ns', ⌜length ns' = num⌝ ∗ ↯ (ε2 ns') ∗ flip_tapes_frag (L:=L) γ1 α (ns ++ ns') + ={∅, E∖↑NS}=∗ T ns'))) + -∗ + P -∗ + wp_update E (∃ n, T n) +}. + + +(** Instantiate flip *) +Section instantiate_flip. + Context `{H: conerisGS Σ, r1:@rand_spec Σ H}. + (* Class flipG1 Σ := FlipG1 { flipG1_tapes:: hocap_tapesGS Σ}. *) + (* Local Definition flip_inv_pred1 `{!hocap_tapesGS Σ} γ1:= *) + (* (∃ (m:gmap loc (list bool)) , *) + (* ([∗ map] α ↦ t ∈ ((λ (ns:list bool), (1%nat, bool_to_nat<$>ns))<$>m), rand_tapes_auth γ1 α ( t.1 ; t.2 )) ∗ *) + (* ●(((λ (ns:list bool), (1, bool_to_nat<$>ns))<$>m):gmap _ _)@γ2 *) + (* )%I. *) + + #[local] Program Definition flip_spec1 : flip_spec := + {| flip_allocate_tape:= (λ: <>, rand_allocate_tape #1%nat); + flip_tape:= (λ: "α", conversion.int_to_bool (rand_tape "α" #1%nat)); + flipG := randG; + flip_tape_name := rand_tape_name; + is_flip L N γ1 := is_rand (L:=L) N γ1; + flip_tapes_auth L γ m := rand_tapes_auth (L:=L) γ ((λ ns, (1, bool_to_nat<$>ns))<$>m); + flip_tapes_frag L γ α ns := rand_tapes_frag (L:=L) γ α (1, (fmap (FMap:=list_fmap) bool_to_nat ns)); + |}. + Next Obligation. + simpl. + iIntros (????) "H1 H2". + by iDestruct (rand_tapes_auth_exclusive with "[$][$]") as "?". + Qed. + Next Obligation. + simpl. + iIntros (?????) "H1 H2". + by iDestruct (rand_tapes_frag_exclusive with "[$][$]") as "?". + Qed. + Next Obligation. + simpl. + iIntros. + iDestruct (rand_tapes_agree with "[$][$]") as "%H'". + rewrite lookup_fmap_Some in H'. destruct H' as (?&?&K). + simplify_eq. + rewrite K. + iPureIntro. + f_equal. + eapply fmap_inj; last done. + intros [][]?; by simplify_eq. + Qed. + Next Obligation. + iIntros. + iMod (rand_tapes_update with "[$][$]") as "[??]"; last iFrame. + - simpl. + rewrite Forall_fmap. + apply Forall_true. + simpl. + intros []; simpl; lia. + - iModIntro. + rewrite fmap_insert; iFrame. + Qed. + Next Obligation. + simpl. + iIntros (???). + iApply rand_inv_create_spec. + Qed. + Next Obligation. + simpl. + iIntros. + wp_pures. + wp_apply rand_allocate_tape_spec; [exact|done..]. + Qed. + Next Obligation. + simpl. + iIntros (??????? n ? ? Φ) "(#Hinv & #Hvs & HP) HΦ". + wp_pures. + wp_apply (rand_tape_spec_some _ _ _ P Q with "[-HΦ]"); [done|..]. + - iSplit; first done. + iSplitR; last iFrame. + iModIntro. iIntros "HP". + iMod ("Hvs" with "[$]") as "[Hfrag Hvs']". + by iFrame. + - iIntros "HQ". + wp_apply conversion.wp_int_to_bool as "_"; first done. + replace (Z_to_bool _) with n; first by iApply "HΦ". + destruct n; simpl. + + rewrite Z_to_bool_neq_0; lia. + + rewrite Z_to_bool_eq_0; lia. + Qed. + Next Obligation. + simpl. + iIntros (??????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP". + iMod (rand_presample_spec _ _ _ _ 1%nat (λ ls, ε2 (nat_to_bool <$>(fin_to_nat <$> ls))) P + _ (λ ls, T (nat_to_bool <$> (fin_to_nat <$> ls))) with "[][][$]") as "(%&?)"; [exact|..]. + - intros. apply Hpos. + by rewrite !fmap_length. + - etrans; last exact. + replace 2%R with (INR 2); last (simpl; lra). + rewrite !Rdiv_def. + apply Rmult_le_compat_r. + { rewrite -Rdiv_1_l -pow_INR. apply Rdiv_INR_ge_0. } + etrans; last eapply (SeriesC_le_inj _ (λ l, Some ((λ x, nat_to_bool (fin_to_nat x)) <$> l))). + + apply Req_le_sym. + apply SeriesC_ext. + intros. simpl. rewrite fmap_length. + rewrite elem_of_enum_uniform_fin_list'. case_match; last lra. + rewrite -list_fmap_compose. + f_equal. + + intros. case_match; last lra. + apply Hpos. + by rewrite -Nat.eqb_eq. + + intros ??? H1 H2. + simplify_eq. + apply fmap_inj in H1; first done. + intros x y ?. + repeat (inv_fin x; simpl; try intros x); + by repeat (inv_fin y; simpl; try intros y). + + eapply ex_seriesC_list_length. intros. + case_match; last done. + by rewrite -Nat.eqb_eq. + - done. + - iModIntro. iIntros "HP". iMod ("Hvs" with "[$]") as "($&$&Hvs')". + iModIntro. + iIntros (ls) "(H1&H2&H3)". + iMod ("Hvs'" with "[H1 $H2 H3]") as "?"; last done. + iSplitL "H1". + + by rewrite !fmap_length. + + rewrite fmap_app -!list_fmap_compose. + erewrite (Forall_fmap_ext_1 (_∘_)); first done. + apply Forall_true. + intros x; by repeat (inv_fin x; simpl; try intros x). + - iModIntro. iFrame. + Qed. +End instantiate_flip. + +(* Section test. *) +(* Context `{F:flip_spec}. *) +(* Lemma flip_presample_spec_simple {L: flipG Σ} NS E ns α *) +(* (ε2 : R -> bool -> R) *) +(* (P : iProp Σ) T γ1 γ2: *) +(* ↑NS ⊆ E -> *) +(* (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> *) +(* (∀ (ε:R), 0<= ε -> (ε2 ε true + ε2 ε false)/2 <= ε)%R-> *) +(* is_flip (L:=L) NS γ1 γ2 -∗ *) +(* (□∀ (ε:R) n, (P ∗ flip_error_auth (L:=L) γ1 ε) ={E∖↑NS}=∗ *) +(* (⌜(1<=ε2 ε n)%R⌝ ∨ (flip_error_auth (L:=L) γ1 (ε2 ε n) ∗ T (n)))) *) +(* -∗ *) +(* P -∗ flip_tapes_frag (L:=L) γ2 α ns-∗ *) +(* wp_update E (∃ n, T (n) ∗ flip_tapes_frag (L:=L) γ2 α (ns++[n])). *) +(* Proof. *) +(* iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) +(* pose (ε2' ε l:= match l with *) +(* | [b]=> ε2 ε b *) +(* | _ => 1%R *) +(* end *) +(* ). *) +(* iMod (flip_presample_spec _ _ _ _ ε2' P 1%nat (λ l, match l with | [b] => T b | _ => False end )%I with "[//][][$][$]") as "(%l & [??])"; first done. *) +(* - rewrite /ε2'. *) +(* intros. repeat case_match; try done. naive_solver. *) +(* - intros. *) +(* etrans; last apply Hineq; try done. *) +(* erewrite (SeriesC_subset (λ x, x ∈ [[true]; [false]])); last first. *) +(* + intros ? H'. *) +(* case_match eqn:K; last done. *) +(* rewrite Nat.eqb_eq in K. *) +(* exfalso. *) +(* apply H'. *) +(* destruct a as [|[|] [|]]; try (simpl in *; done). *) +(* all: set_solver. *) +(* + rewrite SeriesC_list; last first. *) +(* { rewrite !NoDup_cons; repeat split; last apply NoDup_nil; set_solver. } *) +(* simpl. lra. *) +(* - iModIntro. iIntros (ε n) "?". *) +(* destruct n as [|b [|]]. *) +(* + iLeft. iPureIntro. *) +(* by rewrite /ε2'. *) +(* + iMod ("Hvs" $! ε b with "[$]") as "[% | ?]". *) +(* * iLeft. iPureIntro. by rewrite /ε2'. *) +(* * iRight. by rewrite /ε2'. *) +(* + iLeft. iPureIntro. *) +(* by rewrite /ε2'. *) +(* - repeat case_match; try done. *) +(* iModIntro. *) +(* iFrame. *) +(* Qed. *) +(* End test. *) diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index f290bf9c..8b985c73 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -55,16 +55,15 @@ Class rand_spec `{!conerisGS Σ} := RandSpec rand_tape_spec_some {L: randG Σ} N E γ2 (P: iProp Σ) (Q:iProp Σ) (α:loc) (n:nat) (tb:nat) ns: ↑N⊆E -> {{{ is_rand (L:=L) N γ2 ∗ - □ (∀ (m:gmap loc (nat * list nat)), P ∗ - rand_tapes_auth (L:=L) γ2 m - ={E∖↑N}=∗ ⌜m!!α=Some (tb, n::ns)⌝ ∗ Q ∗ rand_tapes_auth (L:=L) γ2 (<[α:=(tb, ns)]> m)) ∗ + (□ (P ={E∖↑N, ∅}=∗ rand_tapes_frag (L:=L) γ2 α (tb, n::ns) ∗ + (rand_tapes_frag (L:=L) γ2 α (tb, ns) ={∅, E∖↑N}=∗ Q))) ∗ ▷ P }}} rand_tape #lbl:α #tb @ E {{{ RET #n; Q }}}; rand_presample_spec {L: randG Σ} N E ns α (tb:nat) (ε2 : list (fin (S tb)) -> R) - (P : iProp Σ) num T γ2 (ε:nonnegreal) : + (P : iProp Σ) num T γ2 (ε:R) : ↑N ⊆ E -> (∀ l, length l = num -> 0<= ε2 l)%R -> (SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R-> @@ -154,16 +153,21 @@ Section impl. wp_pures. wp_bind (rand(_) _)%E. iInv "Hinv" as ">(%&H3&H4)" "Hclose". - iMod ("Hvs" with "[$]") as "(%&HQ&Hauth)". + iMod ("Hvs" with "[$]") as "([Hfrag %] & Hrest)". + iDestruct (hocap_tapes_agree with "[$][$]") as "%". iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. simpl. wp_apply (wp_rand_tape with "[$]") as "[Htape %]". - iMod ("Hclose" with "[- HΦ HQ]") as "_". - { iExists (<[α:=_]> m). - iFrame. - iApply "H3". by iNext. - } - by iApply "HΦ". + iMod (hocap_tapes_update with "[$][$]") as "[? Hfrag]". + iMod ("Hrest" with "[$Hfrag]") as "HQ". + - iPureIntro. by eapply Forall_inv_tail. + - iModIntro. + iMod ("Hclose" with "[- HQ HΦ]") as "_". + { iExists (<[α:=_]> m). + iFrame. + iApply "H3". by iNext. + } + by iApply "HΦ". Qed. Next Obligation. simpl. @@ -230,16 +234,6 @@ End impl. Section checks. Context `{H: conerisGS Σ, r1:@rand_spec Σ H, L:randG Σ}. - Lemma wp_rand_alloc_tape NS (N:nat) E γ1 : - ↑NS ⊆ E -> - {{{ is_rand (L:=L) NS γ1 }}} rand_allocate_tape #N @ E {{{ α, RET #lbl:α; rand_tapes_frag (L:=L) γ1 α (N,[]) }}}. - Proof. - iIntros (Hsubset Φ) "#Hinv HΦ". - wp_apply (rand_allocate_tape_spec with "[//]"); first exact. - iIntros (?) "(%&->&?)". - by iApply "HΦ". - Qed. - Lemma wp_rand_tape_1 NS (N:nat) E γ1 n ns α: ↑NS ⊆ E -> {{{ is_rand (L:=L) NS γ1 ∗ ▷ rand_tapes_frag (L:=L) γ1 α (N, n :: ns) }}} @@ -249,15 +243,16 @@ Section checks. iIntros (Hsubset Φ) "(#Hinv &>Hfrag) HΦ". wp_apply (rand_tape_spec_some _ _ _ _ (⌜n<=N⌝ ∗ rand_tapes_frag _ _ _)%I with "[Hfrag]"); first exact. - iSplit; first done. iSplit; last iApply "Hfrag". - iModIntro. - iIntros (?) "[Hfrag Hauth]". - iDestruct (rand_tapes_agree with "[$][$]") as "%". - iDestruct (rand_tapes_valid with "[$Hfrag]") as "%K". - inversion K; subst. - iMod (rand_tapes_update _ _ _ _ (N, ns) with "[$][$]") as "[Hauth Hfrag]"; first done. - iFrame "Hauth". iModIntro. + iModIntro. iIntros "Hfrag". + iDestruct (rand_tapes_valid with "[$]") as "%H'". iFrame. - by iPureIntro. + iApply fupd_mask_intro; first set_solver. + iIntros "Hclose Hfrag". + iFrame. + iMod "Hclose". + iModIntro. + iPureIntro. + rewrite Forall_cons in H'. naive_solver. - iIntros "[??]". iApply "HΦ". iFrame. Qed. From ef8d71f1f0ee4796a1eadaa54aadcd780aa480b8 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Wed, 25 Sep 2024 15:59:16 +0200 Subject: [PATCH 33/69] NIT --- theories/coneris/lib/hocap_rand.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index 8b985c73..0fca693b 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -233,7 +233,7 @@ End impl. Section checks. Context `{H: conerisGS Σ, r1:@rand_spec Σ H, L:randG Σ}. - + Lemma wp_rand_tape_1 NS (N:nat) E γ1 n ns α: ↑NS ⊆ E -> {{{ is_rand (L:=L) NS γ1 ∗ ▷ rand_tapes_frag (L:=L) γ1 α (N, n :: ns) }}} From 4dbc843d23c24d7f34c70acecc13cd63071888c6 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Thu, 26 Sep 2024 10:21:02 +0200 Subject: [PATCH 34/69] interface spec of random counter --- .../examples/random_counter/random_counter.v | 235 ++++++++---------- theories/coneris/lib/hocap_rand.v | 3 +- 2 files changed, 109 insertions(+), 129 deletions(-) diff --git a/theories/coneris/examples/random_counter/random_counter.v b/theories/coneris/examples/random_counter/random_counter.v index 2491e18c..e4c28cbd 100644 --- a/theories/coneris/examples/random_counter/random_counter.v +++ b/theories/coneris/examples/random_counter/random_counter.v @@ -1,133 +1,114 @@ -(* From clutch.coneris Require Import coneris. *) +From clutch.coneris Require Import coneris. -(* Set Default Proof Using "Type*". *) +Set Default Proof Using "Type*". -(* Class random_counter `{!conerisGS Σ} := RandCounter *) -(* { *) -(* (** * Operations *) *) -(* new_counter : val; *) -(* (* incr_counter : val; *) *) -(* allocate_tape : val; *) -(* incr_counter_tape : val; *) -(* read_counter : val; *) -(* (** * Ghost state *) *) -(* (** The assumptions about [Σ] *) *) -(* counterG : gFunctors → Type; *) -(* (** [name] is used to associate [locked] with [is_lock] *) *) -(* error_name : Type; *) -(* tape_name: Type; *) -(* counter_name: Type; *) -(* (** * Predicates *) *) -(* is_counter {L : counterG Σ} (N:namespace) (counter: val) *) -(* (γ1: error_name) (γ2: tape_name) (γ3: counter_name): iProp Σ; *) -(* counter_error_auth {L : counterG Σ} (γ: error_name) (x:R): iProp Σ; *) -(* counter_error_frag {L : counterG Σ} (γ: error_name) (x:R): iProp Σ; *) -(* counter_tapes_auth {L : counterG Σ} (γ: tape_name) (m:gmap loc (list nat)): iProp Σ; *) -(* counter_tapes_frag {L : counterG Σ} (γ: tape_name) (α:loc) (ns:list nat): iProp Σ; *) -(* counter_content_auth {L : counterG Σ} (γ: counter_name) (z:nat): iProp Σ; *) -(* counter_content_frag {L : counterG Σ} (γ: counter_name) (f:frac) (z:nat): iProp Σ; *) -(* (** * General properties of the predicates *) *) -(* #[global] is_counter_persistent {L : counterG Σ} N c γ1 γ2 γ3 :: *) -(* Persistent (is_counter (L:=L) N c γ1 γ2 γ3); *) -(* #[global] counter_error_auth_timeless {L : counterG Σ} γ x :: *) -(* Timeless (counter_error_auth (L:=L) γ x); *) -(* #[global] counter_error_frag_timeless {L : counterG Σ} γ x :: *) -(* Timeless (counter_error_frag (L:=L) γ x); *) -(* #[global] counter_tapes_auth_timeless {L : counterG Σ} γ m :: *) -(* Timeless (counter_tapes_auth (L:=L) γ m); *) -(* #[global] counter_tapes_frag_timeless {L : counterG Σ} γ α ns :: *) -(* Timeless (counter_tapes_frag (L:=L) γ α ns); *) -(* #[global] counter_content_auth_timeless {L : counterG Σ} γ z :: *) -(* Timeless (counter_content_auth (L:=L) γ z); *) -(* #[global] counter_content_frag_timeless {L : counterG Σ} γ f z :: *) -(* Timeless (counter_content_frag (L:=L) γ f z); *) - -(* counter_error_auth_exclusive {L : counterG Σ} γ x1 x2: *) -(* counter_error_auth (L:=L) γ x1 -∗ counter_error_auth (L:=L) γ x2 -∗ False; *) -(* counter_error_frag_split {L : counterG Σ} γ (x1 x2:nonnegreal): *) -(* counter_error_frag (L:=L) γ x1 ∗ counter_error_frag (L:=L) γ x2 ⊣⊢ *) -(* counter_error_frag (L:=L) γ (x1+x2)%R ; *) -(* counter_error_auth_valid {L : counterG Σ} γ x: *) -(* counter_error_auth (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; *) -(* counter_error_frag_valid {L : counterG Σ} γ x: *) -(* counter_error_frag (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; *) -(* counter_error_ineq {L : counterG Σ} γ x1 x2: *) -(* counter_error_auth (L:=L) γ x1 -∗ counter_error_frag (L:=L) γ x2 -∗ ⌜(x2<=x1)%R ⌝; *) -(* counter_error_decrease {L : counterG Σ} γ (x1 x2:nonnegreal): *) -(* counter_error_auth (L:=L) γ (x1 + x2)%NNR -∗ counter_error_frag (L:=L) γ x1 ==∗ counter_error_auth (L:=L) γ x2; *) -(* counter_error_increase {L : counterG Σ} γ (x1 x2:nonnegreal): *) -(* (x1+x2<1)%R -> ⊢ counter_error_auth (L:=L) γ x1 ==∗ *) -(* counter_error_auth (L:=L) γ (x1+x2) ∗ counter_error_frag (L:=L) γ x2; *) +Class random_counter `{!conerisGS Σ} := RandCounter +{ + (** * Operations *) + new_counter : val; + (* incr_counter : val; *) + allocate_tape : val; + incr_counter_tape : val; + read_counter : val; + (** * Ghost state *) + (** The assumptions about [Σ] *) + counterG : gFunctors → Type; + (** [name] is used to associate [locked] with [is_lock] *) + tape_name: Type; + counter_name: Type; + (** * Predicates *) + is_counter {L : counterG Σ} (N:namespace) (counter: val) + (γ1: tape_name) (γ2: counter_name): iProp Σ; + counter_tapes_auth {L : counterG Σ} (γ: tape_name) (m:gmap loc (list nat)): iProp Σ; + counter_tapes_frag {L : counterG Σ} (γ: tape_name) (α:loc) (ns:list nat): iProp Σ; + counter_content_auth {L : counterG Σ} (γ: counter_name) (z:nat): iProp Σ; + counter_content_frag {L : counterG Σ} (γ: counter_name) (f:frac) (z:nat): iProp Σ; + (** * General properties of the predicates *) + #[global] is_counter_persistent {L : counterG Σ} N c γ1 γ2 :: + Persistent (is_counter (L:=L) N c γ1 γ2); + #[global] counter_tapes_auth_timeless {L : counterG Σ} γ m :: + Timeless (counter_tapes_auth (L:=L) γ m); + #[global] counter_tapes_frag_timeless {L : counterG Σ} γ α ns :: + Timeless (counter_tapes_frag (L:=L) γ α ns); + #[global] counter_content_auth_timeless {L : counterG Σ} γ z :: + Timeless (counter_content_auth (L:=L) γ z); + #[global] counter_content_frag_timeless {L : counterG Σ} γ f z :: + Timeless (counter_content_frag (L:=L) γ f z); -(* counter_tapes_auth_exclusive {L : counterG Σ} γ m m': *) -(* counter_tapes_auth (L:=L) γ m -∗ counter_tapes_auth (L:=L) γ m' -∗ False; *) -(* counter_tapes_frag_exclusive {L : counterG Σ} γ α ns ns': *) -(* counter_tapes_frag (L:=L) γ α ns -∗ counter_tapes_frag (L:=L) γ α ns' -∗ False; *) -(* counter_tapes_agree {L : counterG Σ} γ α m ns: *) -(* counter_tapes_auth (L:=L) γ m -∗ counter_tapes_frag (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; *) -(* counter_tapes_update {L : counterG Σ} γ α m ns n: *) -(* counter_tapes_auth (L:=L) γ m -∗ counter_tapes_frag (L:=L) γ α ns ==∗ *) -(* counter_tapes_auth (L:=L) γ (<[α := (ns ++[n])]> m) ∗ counter_tapes_frag (L:=L) γ α (ns ++ [n]); *) + counter_tapes_auth_exclusive {L : counterG Σ} γ m m': + counter_tapes_auth (L:=L) γ m -∗ counter_tapes_auth (L:=L) γ m' -∗ False; + counter_tapes_frag_exclusive {L : counterG Σ} γ α ns ns': + counter_tapes_frag (L:=L) γ α ns -∗ counter_tapes_frag (L:=L) γ α ns' -∗ False; + counter_tapes_agree {L : counterG Σ} γ α m ns: + counter_tapes_auth (L:=L) γ m -∗ counter_tapes_frag (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; + counter_tapes_valid {L : counterG Σ} γ α ns: + counter_tapes_frag (L:=L) γ α ns -∗ ⌜Forall (λ n, n<=3)%nat ns⌝; + counter_tapes_update {L : counterG Σ} γ α m ns ns': + Forall (λ x, x<=3%nat) ns'-> + counter_tapes_auth (L:=L) γ m -∗ counter_tapes_frag (L:=L) γ α ns ==∗ + counter_tapes_auth (L:=L) γ (<[α := ns']> m) ∗ counter_tapes_frag (L:=L) γ α (ns'); -(* counter_content_auth_exclusive {L : counterG Σ} γ z1 z2: *) -(* counter_content_auth (L:=L) γ z1 -∗ counter_content_auth (L:=L) γ z2 -∗ False; *) -(* counter_content_less_than {L : counterG Σ} γ z z' f: *) -(* counter_content_auth (L:=L) γ z -∗ counter_content_frag (L:=L) γ f z' -∗ ⌜(z'<=z)%nat ⌝; *) -(* counter_content_frag_combine {L:counterG Σ} γ f f' z z': *) -(* (counter_content_frag (L:=L) γ f z ∗ counter_content_frag (L:=L) γ f' z')%I ≡ *) -(* counter_content_frag (L:=L) γ (f+f')%Qp (z+z')%nat; *) -(* counter_content_agree {L : counterG Σ} γ z z': *) -(* counter_content_auth (L:=L) γ z -∗ counter_content_frag (L:=L) γ 1%Qp z' -∗ ⌜(z'=z)%nat ⌝; *) -(* counter_content_update {L : counterG Σ} γ f z1 z2 z3: *) -(* counter_content_auth (L:=L) γ z1 -∗ counter_content_frag (L:=L) γ f z2 ==∗ *) -(* counter_content_auth (L:=L) γ (z1+z3)%nat ∗ counter_content_frag (L:=L) γ f (z2+z3)%nat; *) + counter_content_auth_exclusive {L : counterG Σ} γ z1 z2: + counter_content_auth (L:=L) γ z1 -∗ counter_content_auth (L:=L) γ z2 -∗ False; + counter_content_less_than {L : counterG Σ} γ z z' f: + counter_content_auth (L:=L) γ z -∗ counter_content_frag (L:=L) γ f z' -∗ ⌜(z'<=z)%nat ⌝; + counter_content_frag_combine {L:counterG Σ} γ f f' z z': + (counter_content_frag (L:=L) γ f z ∗ counter_content_frag (L:=L) γ f' z')%I ≡ + counter_content_frag (L:=L) γ (f+f')%Qp (z+z')%nat; + counter_content_agree {L : counterG Σ} γ z z': + counter_content_auth (L:=L) γ z -∗ counter_content_frag (L:=L) γ 1%Qp z' -∗ ⌜(z'=z)%nat ⌝; + counter_content_update {L : counterG Σ} γ f z1 z2 z3: + counter_content_auth (L:=L) γ z1 -∗ counter_content_frag (L:=L) γ f z2 ==∗ + counter_content_auth (L:=L) γ (z1+z3)%nat ∗ counter_content_frag (L:=L) γ f (z2+z3)%nat; -(* (** * Program specs *) *) -(* new_counter_spec {L : counterG Σ} E ε N: *) -(* {{{ ↯ ε }}} new_counter #() @E *) -(* {{{ c, RET c; ∃ γ1 γ2 γ3, is_counter (L:=L) N c γ1 γ2 γ3 ∗ *) -(* counter_error_frag (L:=L) γ1 ε ∗ *) -(* counter_content_frag (L:=L) γ3 1%Qp 0%nat *) -(* }}}; *) -(* allocate_tape_spec {L: counterG Σ} N E c γ1 γ2 γ3: *) -(* ↑N ⊆ E-> *) -(* {{{ is_counter (L:=L) N c γ1 γ2 γ3 }}} *) -(* allocate_tape #() @ E *) -(* {{{ (v:val), RET v; *) -(* ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ counter_tapes_frag (L:=L) γ2 α [] *) -(* }}}; *) -(* incr_counter_tape_spec_some {L: counterG Σ} N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: *) -(* ↑N⊆E -> *) -(* {{{ is_counter (L:=L) N c γ1 γ2 γ3 ∗ *) -(* □ (∀ (z:nat), P ∗ counter_content_auth (L:=L) γ3 z ={E∖↑N}=∗ *) -(* counter_content_auth (L:=L) γ3 (z+n)%nat ∗ Q z) ∗ *) -(* P ∗ counter_tapes_frag (L:=L) γ2 α (n::ns) *) -(* }}} *) -(* incr_counter_tape c #lbl:α @ E *) -(* {{{ (z:nat), RET (#z, #n); Q z ∗ counter_tapes_frag (L:=L) γ2 α ns}}}; *) -(* counter_presample_spec {L: counterG Σ} NS E ns α *) -(* (ε2 : R -> nat -> R) *) -(* (P : iProp Σ) T γ1 γ2 γ3 c: *) -(* ↑NS ⊆ E -> *) -(* (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> *) -(* (∀ (ε:R), 0<= ε ->SeriesC (λ n, if (bool_decide (n≤3%nat)) then 1 / (S 3%nat) * ε2 ε n else 0%R)%R <= ε)%R-> *) -(* is_counter (L:=L) NS c γ1 γ2 γ3 -∗ *) -(* (□∀ (ε ε':nonnegreal) n, (P ∗ counter_error_auth (L:=L) γ1 (ε')%R) ={E∖↑NS}=∗ *) -(* (⌜(1<=ε2 ε n)%R⌝ ∨ (counter_error_auth (L:=L) γ1 (ε' + )%R ∗ T (n)))) *) -(* -∗ *) -(* P -∗ counter_tapes_frag (L:=L) γ2 α ns-∗ *) -(* wp_update E (∃ n, T (n) ∗ counter_tapes_frag (L:=L) γ2 α (ns++[n])); *) -(* read_counter_spec {L: counterG Σ} N E c γ1 γ2 γ3 P Q: *) -(* ↑N ⊆ E -> *) -(* {{{ is_counter (L:=L) N c γ1 γ2 γ3 ∗ *) -(* □ (∀ (z:nat), P ∗ counter_content_auth (L:=L) γ3 z ={E∖↑N}=∗ *) -(* counter_content_auth (L:=L) γ3 z∗ Q z) *) -(* ∗ P *) -(* }}} *) -(* read_counter c @ E *) -(* {{{ (n':nat), RET #n'; Q n' *) -(* }}} *) -(* }. *) + (** * Program specs *) + new_counter_spec {L : counterG Σ} E N: + {{{ True }}} new_counter #() @E + {{{ c, RET c; ∃ γ1 γ2, is_counter (L:=L) N c γ1 γ2 ∗ + counter_content_frag (L:=L) γ2 1%Qp 0%nat + }}}; + allocate_tape_spec {L: counterG Σ} N E c γ1 γ2: + ↑N ⊆ E-> + {{{ is_counter (L:=L) N c γ1 γ2 }}} + allocate_tape #() @ E + {{{ (v:val), RET v; + ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ counter_tapes_frag (L:=L) γ1 α [] + }}}; + incr_counter_tape_spec_some {L: counterG Σ} N E c γ1 γ2 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: + ↑N⊆E -> + {{{ is_counter (L:=L) N c γ1 γ2 ∗ + □ (∀ (z:nat), counter_content_auth (L:=L) γ2 z ∗ P ={E∖ ↑N, ∅}=∗ + counter_tapes_frag (L:=L) γ1 α (n::ns) ∗ + (counter_tapes_frag (L:=L) γ1 α ns ={∅, E∖↑N}=∗ + counter_content_auth (L:=L) γ2 (z+n) ∗ Q z) + ) ∗ P + }}} + incr_counter_tape c #lbl:α @ E + {{{ (z:nat), RET (#z, #n); Q z}}}; + counter_presample_spec {L: counterG Σ} NS E ns α + (ε2 : list nat -> R) + (P : iProp Σ) T γ1 γ2 c num ε: + ↑NS ⊆ E -> + (∀ n, 0<=ε2 n)%R -> + (SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R-> + is_counter (L:=L) NS c γ1 γ2 -∗ + □(P ={E∖↑NS,∅}=∗ ↯ ε ∗ counter_tapes_frag (L:=L) γ1 α ns ∗ + (∀ ns', ⌜length ns' = num⌝ ∗ ↯ (ε2 ns') ∗ counter_tapes_frag (L:=L) γ1 α (ns++ns')={∅,E∖↑NS}=∗ + T ns' + ))-∗ + P -∗ + wp_update E (∃ n, T n); + read_counter_spec {L: counterG Σ} N E c γ1 γ2 P Q: + ↑N ⊆ E -> + {{{ is_counter (L:=L) N c γ1 γ2 ∗ + □ (∀ (z:nat), P ∗ counter_content_auth (L:=L) γ2 z ={E∖↑N}=∗ + counter_content_auth (L:=L) γ2 z∗ Q z) + ∗ P + }}} + read_counter c @ E + {{{ (n':nat), RET #n'; Q n' + }}} +}. (* Section lemmas. *) @@ -165,7 +146,7 @@ (* {{{ is_counter (L:=L) N c γ1 γ2 γ3 ∗ *) (* counter_error_frag (L:=L) γ1 ε ∗ *) (* counter_tapes_frag (L:=L) γ2 α [] ∗ *) -(* counter_content_frag (L:=L) γ3 q z *) +(* counter_content_frag (L:=L) γ3 q z *) (* }}} *) (* incr_counter_tape c #lbl:α @ E *) (* {{{ (z':nat) (n:nat), *) @@ -197,7 +178,7 @@ (* iDestruct (counter_error_auth_valid with "[$]") as "%". *) (* iDestruct (counter_error_frag_valid with "[$]") as "%". *) (* unshelve iMod (counter_error_update _ _ _ (mknonnegreal (ε2' ε n) _) with "[$][$]") as "[$$]". *) -(* { rewrite /ε2' H. *) +(* { rewrite /ε2' H. *) (* naive_solver. } *) (* iPureIntro. split; first lia. *) (* apply leb_complete in H. lia. *) diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index 0fca693b..83211ce5 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -35,8 +35,7 @@ Class rand_spec `{!conerisGS Σ} := RandSpec rand_tapes_agree {L : randG Σ} γ α m ns: rand_tapes_auth (L:=L) γ m -∗ rand_tapes_frag (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; rand_tapes_valid {L : randG Σ} γ2 α tb ns: - rand_tapes_frag (L:=L) γ2 α (tb, ns) -∗ - ⌜Forall (λ n, n<=tb)%nat ns⌝ ; + rand_tapes_frag (L:=L) γ2 α (tb, ns) -∗ ⌜Forall (λ n, n<=tb)%nat ns⌝ ; rand_tapes_update {L : randG Σ} γ α m ns ns': Forall (λ x, x<=ns'.1) ns'.2 -> rand_tapes_auth (L:=L) γ m -∗ rand_tapes_frag (L:=L) γ α ns ==∗ From cbdb96984851e630d27d75ec0de4dfd5829a6c73 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Thu, 26 Sep 2024 13:42:27 +0200 Subject: [PATCH 35/69] Better spec --- theories/coneris/lib/hocap_flip.v | 104 +++++++++++++------- theories/coneris/lib/hocap_rand.v | 154 ++++++++++++++++-------------- 2 files changed, 150 insertions(+), 108 deletions(-) diff --git a/theories/coneris/lib/hocap_flip.v b/theories/coneris/lib/hocap_flip.v index 68b4db4b..cca4352f 100644 --- a/theories/coneris/lib/hocap_flip.v +++ b/theories/coneris/lib/hocap_flip.v @@ -49,29 +49,29 @@ Class flip_spec `{!conerisGS Σ} := FlipSpec ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ flip_tapes_frag (L:=L) γ1 α [] }}}; - flip_tape_spec_some {L: flipG Σ} N E γ1 (P: iProp Σ) (Q:iProp Σ) (α:loc) (n:bool) ns: + flip_tape_spec_some {L: flipG Σ} N E γ1 (P: iProp Σ) (Q:bool -> list bool -> iProp Σ) (α:loc) : ↑N⊆E -> {{{ is_flip (L:=L) N γ1 ∗ - (□ (P ={E∖↑N, ∅}=∗ flip_tapes_frag (L:=L) γ1 α (n::ns) ∗ - (flip_tapes_frag (L:=L) γ1 α ns ={∅, E∖↑N}=∗ Q))) ∗ + (□ (P ={E∖↑N, ∅}=∗ + ∃ n ns, flip_tapes_frag (L:=L) γ1 α (n::ns) ∗ + (flip_tapes_frag (L:=L) γ1 α ns ={∅, E∖↑N}=∗ Q n ns))) ∗ ▷ P }}} flip_tape #lbl:α @ E - {{{ RET #n; Q }}}; + {{{ (n:bool), RET #n; ∃ ns, Q n ns}}}; - flip_presample_spec {L: flipG Σ} NS E ns α - (ε2 : list bool -> R) - (P : iProp Σ) num T γ1 ε: + flip_presample_spec {L: flipG Σ} NS E (P : iProp Σ) T γ1: ↑NS ⊆ E -> - (∀ l, length l = num -> 0<=ε2 l)%R -> - (SeriesC (λ l, if length l =? num then ε2 l else 0) /(2^num) <= ε)%R-> is_flip (L:=L) NS γ1 -∗ - (□ (P ={E∖↑NS, ∅}=∗ ↯ ε ∗ flip_tapes_frag (L:=L) γ1 α ns ∗ + (□ (P ={E∖↑NS, ∅}=∗ + ∃ ε num ε2 α ns, ↯ ε ∗ flip_tapes_frag (L:=L) γ1 α ns ∗ + ⌜(∀ l, length l = num -> 0<=ε2 l)%R⌝ ∗ + ⌜(SeriesC (λ l, if length l =? num then ε2 l else 0) /(2^num) <= ε)%R⌝∗ (∀ ns', ⌜length ns' = num⌝ ∗ ↯ (ε2 ns') ∗ flip_tapes_frag (L:=L) γ1 α (ns ++ ns') - ={∅, E∖↑NS}=∗ T ns'))) + ={∅, E∖↑NS}=∗ T ε num ε2 α ns ns'))) -∗ P -∗ - wp_update E (∃ n, T n) + wp_update E (∃ ε num ε2 α ns ns', T ε num ε2 α ns ns') }. @@ -140,26 +140,48 @@ Section instantiate_flip. Qed. Next Obligation. simpl. - iIntros (??????? n ? ? Φ) "(#Hinv & #Hvs & HP) HΦ". + iIntros (????? Q ? ? Φ) "(#Hinv & #Hvs & HP) HΦ". wp_pures. - wp_apply (rand_tape_spec_some _ _ _ P Q with "[-HΦ]"); [done|..]. - - iSplit; first done. - iSplitR; last iFrame. + wp_apply (rand_tape_spec_some _ _ _ P (λ n ns, ∃ b bs, ⌜n= bool_to_nat b⌝ ∗ + ⌜ns = bool_to_nat <$> bs⌝ ∗ + Q b bs + + )%I with "[-HΦ]"); [done|..]. + - iFrame. iSplit; first done. iModIntro. iIntros "HP". - iMod ("Hvs" with "[$]") as "[Hfrag Hvs']". - by iFrame. - - iIntros "HQ". + iMod ("Hvs" with "[$]") as "(%b & %bs & Hfrag & Hrest)". + iFrame. + iModIntro. + iIntros "?". + iMod ("Hrest" with "[$]"). + by iFrame. + - iIntros (n) "(%&%&%&%&%&HQ)". wp_apply conversion.wp_int_to_bool as "_"; first done. - replace (Z_to_bool _) with n; first by iApply "HΦ". - destruct n; simpl. + iApply "HΦ". + subst. + iFrame. + replace (Z_to_bool _) with b; first iFrame. + destruct b; simpl. + rewrite Z_to_bool_neq_0; lia. + rewrite Z_to_bool_eq_0; lia. Qed. Next Obligation. simpl. - iIntros (??????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP". - iMod (rand_presample_spec _ _ _ _ 1%nat (λ ls, ε2 (nat_to_bool <$>(fin_to_nat <$> ls))) P - _ (λ ls, T (nat_to_bool <$> (fin_to_nat <$> ls))) with "[][][$]") as "(%&?)"; [exact|..]. + iIntros (???? T ??) "#Hinv #Hvs HP". + iMod (rand_presample_spec _ _ P + (λ ε num tb ε2 α ns ns', + ∃ bs bs' ε2', ⌜tb = 1%nat⌝ ∗ ⌜fmap (FMap:= list_fmap) bool_to_nat bs = ns⌝ ∗ ⌜fmap (FMap:= list_fmap) bool_to_nat bs'=fin_to_nat <$> ns'⌝ ∗ + ⌜∀ xs ys, fmap (FMap:= list_fmap) bool_to_nat xs = fmap (FMap:= list_fmap)fin_to_nat ys -> ε2' xs = ε2 ys⌝ ∗ + T ε num ε2' α bs bs' + )%I + with "[][][$]") as "(%&%&%&%&%&%&%&%&%&%&->&<-&%&%&HT)"; [exact|exact| |iModIntro; iFrame]. + iModIntro. + iIntros "HP". + iMod ("Hvs" with "[$]") as "(%&%&%&%&%&Herr&Hfrag&%Hpos&%Hineq&Hvs')". + iFrame. + iModIntro. + iExists num, (λ ls, ε2 (nat_to_bool<$>(fin_to_nat <$> ls))). + repeat iSplit; try iPureIntro. - intros. apply Hpos. by rewrite !fmap_length. - etrans; last exact. @@ -186,18 +208,28 @@ Section instantiate_flip. + eapply ex_seriesC_list_length. intros. case_match; last done. by rewrite -Nat.eqb_eq. - - done. - - iModIntro. iIntros "HP". iMod ("Hvs" with "[$]") as "($&$&Hvs')". - iModIntro. - iIntros (ls) "(H1&H2&H3)". - iMod ("Hvs'" with "[H1 $H2 H3]") as "?"; last done. - iSplitL "H1". - + by rewrite !fmap_length. - + rewrite fmap_app -!list_fmap_compose. - erewrite (Forall_fmap_ext_1 (_∘_)); first done. - apply Forall_true. - intros x; by repeat (inv_fin x; simpl; try intros x). - - iModIntro. iFrame. + - iIntros (ls) "(H1&H2&H3)". + iMod ("Hvs'" with "[H1 $H2 H3]") as "?". + + rewrite !fmap_length !fmap_app. + iSplitL "H1"; first done. + rewrite -!list_fmap_compose. + erewrite (Forall_fmap_ext_1 (_∘_)); first done. + apply Forall_true. + intros x; by repeat (inv_fin x; simpl; try intros x). + + iFrame. iPureIntro; repeat split; try done. + * rewrite -!list_fmap_compose. + erewrite (Forall_fmap_ext_1 (_∘_)); first done. + apply Forall_true. + intros x; by repeat (inv_fin x; simpl; try intros x). + * intros ?? <-. + f_equal. + rewrite -!list_fmap_compose. + rewrite -{1}(list_fmap_id xs). + erewrite (Forall_fmap_ext_1 (_∘_)); first done. + apply Forall_true. + intros []; simpl. + -- by rewrite nat_to_bool_neq_0. + -- by rewrite nat_to_bool_eq_0. Qed. End instantiate_flip. diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index 83211ce5..42a7ed6f 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -51,28 +51,28 @@ Class rand_spec `{!conerisGS Σ} := RandSpec {{{ (v:val), RET v; ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ rand_tapes_frag (L:=L) γ2 α (tb, []) }}}; - rand_tape_spec_some {L: randG Σ} N E γ2 (P: iProp Σ) (Q:iProp Σ) (α:loc) (n:nat) (tb:nat) ns: + rand_tape_spec_some {L: randG Σ} N E γ2 (P: iProp Σ) (Q: nat ->list nat -> iProp Σ) (α:loc) (tb:nat) : ↑N⊆E -> {{{ is_rand (L:=L) N γ2 ∗ - (□ (P ={E∖↑N, ∅}=∗ rand_tapes_frag (L:=L) γ2 α (tb, n::ns) ∗ - (rand_tapes_frag (L:=L) γ2 α (tb, ns) ={∅, E∖↑N}=∗ Q))) ∗ + (□ (P ={E∖↑N, ∅}=∗ ∃ n ns, rand_tapes_frag (L:=L) γ2 α (tb, n::ns) ∗ + (rand_tapes_frag (L:=L) γ2 α (tb, ns) ={∅, E∖↑N}=∗ Q n ns))) ∗ ▷ P }}} rand_tape #lbl:α #tb @ E - {{{ RET #n; Q }}}; - rand_presample_spec {L: randG Σ} N E ns α (tb:nat) - (ε2 : list (fin (S tb)) -> R) - (P : iProp Σ) num T γ2 (ε:R) : + {{{ (n:nat), RET #n; ∃ ns, Q n ns}}}; + rand_presample_spec {L: randG Σ} N E (P : iProp Σ) T γ2 : ↑N ⊆ E -> - (∀ l, length l = num -> 0<= ε2 l)%R -> - (SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R-> is_rand (L:=L) N γ2 -∗ - (□ (P ={E∖↑N, ∅}=∗ ↯ ε ∗ rand_tapes_frag (L:=L) γ2 α (tb, ns) ∗ + (□ (P ={E∖↑N, ∅}=∗ + ∃ ε num tb (ε2 : list (fin (S tb)) -> R) α ns, + ↯ ε ∗ rand_tapes_frag (L:=L) γ2 α (tb, ns) ∗ + ⌜(∀ l, length l = num -> 0<= ε2 l)%R⌝ ∗ + ⌜(SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R⌝ ∗ (∀ ns', ⌜length ns' = num⌝ ∗ ↯ (ε2 ns') ∗ rand_tapes_frag (L:=L) γ2 α (tb, ns ++ (fin_to_nat <$> ns')) - ={∅, E∖↑N}=∗ T ns'))) + ={∅, E∖↑N}=∗ T ε num tb ε2 α ns ns'))) -∗ P -∗ - wp_update E (∃ n, T n) + wp_update E (∃ ε num tb ε2 α ns ns', T ε num tb ε2 α ns ns') }. Section impl. @@ -147,34 +147,34 @@ Section impl. by iFrame. Qed. Next Obligation. - simpl. - iIntros (????????????? Φ) "(#Hinv & #Hvs & HP) HΦ". - wp_pures. - wp_bind (rand(_) _)%E. - iInv "Hinv" as ">(%&H3&H4)" "Hclose". - iMod ("Hvs" with "[$]") as "([Hfrag %] & Hrest)". - iDestruct (hocap_tapes_agree with "[$][$]") as "%". - iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. - simpl. - wp_apply (wp_rand_tape with "[$]") as "[Htape %]". - iMod (hocap_tapes_update with "[$][$]") as "[? Hfrag]". - iMod ("Hrest" with "[$Hfrag]") as "HQ". - - iPureIntro. by eapply Forall_inv_tail. - - iModIntro. - iMod ("Hclose" with "[- HQ HΦ]") as "_". - { iExists (<[α:=_]> m). - iFrame. - iApply "H3". by iNext. + simpl. + iIntros (??????????? Φ) "(#Hinv & #Hvs & HP) HΦ". + wp_pures. + wp_bind (rand(_) _)%E. + iInv "Hinv" as ">(%&H3&H4)" "Hclose". + iMod ("Hvs" with "[$]") as "(%n & %ns & [Hfrag %] & Hrest)". + iDestruct (hocap_tapes_agree with "[$][$]") as "%". + iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. + simpl. + wp_apply (wp_rand_tape with "[$]") as "[Htape %]". + iMod (hocap_tapes_update with "[$][$]") as "[? Hfrag]". + iMod ("Hrest" with "[$Hfrag]") as "HQ". + - iPureIntro. by eapply Forall_inv_tail. + - iModIntro. + iMod ("Hclose" with "[- HQ HΦ]") as "_". + { iExists (<[α:=_]> m). + iFrame. + iApply "H3". by iNext. } - by iApply "HΦ". -Qed. -Next Obligation. - simpl. - iIntros (?????????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP ". + iApply "HΦ". by iFrame. + Qed. + Next Obligation. + simpl. + iIntros (?????????) "#Hinv #Hvs HP ". iApply wp_update_state_step_coupl. iIntros (σ1 ε_supply) "((Hheap&Htapes)&Hε)". iMod (inv_acc with "Hinv") as "[>(% & H3 & H4 ) Hclose]"; [done|]. - iMod ("Hvs" with "[$]") as "(Herr & [Hfrag %] & Hrest)". + iMod ("Hvs" with "[$]") as "(%err & %num & %tb & %ε2 & %α & %ns & Herr & [Hfrag %] & %Hpos & %Hineq & Hrest)". iDestruct (hocap_tapes_agree with "[$][$]") as "%". iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". @@ -240,7 +240,7 @@ Section checks. {{{ RET #(LitInt n); rand_tapes_frag (L:=L) γ1 α (N, ns) ∗ ⌜n <= N⌝ }}}. Proof. iIntros (Hsubset Φ) "(#Hinv &>Hfrag) HΦ". - wp_apply (rand_tape_spec_some _ _ _ _ (⌜n<=N⌝ ∗ rand_tapes_frag _ _ _)%I with "[Hfrag]"); first exact. + wp_apply (rand_tape_spec_some _ _ _ _ (λ n' ns', ⌜n=n'/\ns=ns'⌝ ∗ ⌜n<=N⌝ ∗ rand_tapes_frag _ _ _)%I with "[Hfrag]"); first exact. - iSplit; first done. iSplit; last iApply "Hfrag". iModIntro. iIntros "Hfrag". iDestruct (rand_tapes_valid with "[$]") as "%H'". @@ -252,10 +252,11 @@ Section checks. iModIntro. iPureIntro. rewrite Forall_cons in H'. naive_solver. - - iIntros "[??]". iApply "HΦ". - iFrame. + - iIntros (?) "(%&[-> ->]&%&?)". iApply "HΦ". + by iFrame. Qed. + Local Opaque enum_uniform_fin_list. Lemma wp_presample_adv_comp_rand_tape (N : nat) NS E α ns (ε1 : R) (ε2 : fin (S N) -> R) γ1: ↑NS ⊆ E -> (∀ n, 0<=ε2 n)%R -> @@ -268,43 +269,52 @@ Section checks. iIntros (Hsubset Hpos Hineq) "#Hinv >Htape Herr". iDestruct (ec_valid with "[$]") as "%Hpos'". destruct Hpos' as [Hpos' ?]. - iMod (rand_presample_spec _ _ _ _ _ (λ l, match l with - | [x] => ε2 x - | _ => 1%R - end - ) - (rand_tapes_frag (L:=L) γ1 α (N, ns) ∗ ↯ ε1) 1%nat - (λ ls, ∃ n, ⌜ls=[n]⌝ ∗ ↯ (ε2 n) ∗ rand_tapes_frag γ1 α (N, ns ++ [fin_to_nat n]))%I - _ (mknonnegreal ε1 Hpos') - with "[//][][$Htape $Herr]") as "(%l&%sample & -> & Herr &Htape)". + iMod (rand_presample_spec _ _ + (rand_tapes_frag (L:=L) γ1 α (N, ns) ∗ ↯ ε1) + (λ ε' (num':nat) tb' ε2' α' ns1 ns2, + ⌜ε1=ε'⌝ ∗ ⌜(1=num')%nat⌝ ∗ ⌜N=tb'⌝ ∗ + ⌜∀ x y, fin_to_nat <$> x = fin_to_nat <$> y -> + ε2' x = (λ l, match l with |[x] => ε2 x | _ => 1%R end) y⌝∗ ⌜α=α'⌝ ∗ ⌜ns1=ns⌝ ∗ + ∃ (n:fin (S N)), ⌜fin_to_nat <$> ns2=[fin_to_nat n]⌝ ∗ ↯ (ε2 n) ∗ rand_tapes_frag γ1 α (N, ns ++ [fin_to_nat n]))%I + with "[//][][$Htape $Herr]") as "(%&%&%&%&%&%&%&->&<-&->&%&->&->&%&%&?&?)". - done. - - by intros [|?[|]]. - - etrans; last eapply Hineq. - etrans; last eapply (SeriesC_le_inj _ (λ l, match l with |[x] => Some x | _ => None end)). - + rewrite Rdiv_def -SeriesC_scal_r. apply Req_le_sym. - apply SeriesC_ext. - intros. case_match; subst. - * rewrite bool_decide_eq_false_2; first (simpl;lra). - by rewrite elem_of_enum_uniform_fin_list. - * case_match; last first. - { rewrite bool_decide_eq_false_2; first (simpl;lra). - by rewrite elem_of_enum_uniform_fin_list. } - rewrite bool_decide_eq_true_2; last by apply elem_of_enum_uniform_fin_list. - simpl. - rewrite Rdiv_def Rmult_1_r; lra. - + intros. apply Rmult_le_pos; last done. - apply Rdiv_INR_ge_0. - + intros. repeat case_match; by simplify_eq. - + apply ex_seriesC_finite. - iModIntro. - iIntros "(?&?)". - simpl. iFrame. + iIntros "[Hfrag Herr]". + iFrame. iApply fupd_mask_intro; first set_solver. iIntros "Hclose". - iIntros ([|?[|]]) "(%&?&?)"; try done. - iMod "Hclose". - iFrame. - by iModIntro. + iExists 1, (λ l, match l with |[x] => ε2 x | _ => 1%R end). simpl. + repeat iSplit. + + iPureIntro. by intros [|?[|]]. + + iPureIntro. etrans; last eapply Hineq. + etrans; last eapply (SeriesC_le_inj _ (λ l, match l with |[x] => Some x | _ => None end)). + * rewrite Rdiv_def -SeriesC_scal_r. apply Req_le_sym. + apply SeriesC_ext. + intros. case_match; subst. + -- rewrite bool_decide_eq_false_2; first (simpl;lra). + by rewrite elem_of_enum_uniform_fin_list. + -- case_match; last first. + { rewrite bool_decide_eq_false_2; first (simpl;lra). + subst. + by rewrite elem_of_enum_uniform_fin_list. + } + rewrite bool_decide_eq_true_2; last by apply elem_of_enum_uniform_fin_list. + simpl. + rewrite Rdiv_def Rmult_1_r; lra. + * intros. apply Rmult_le_pos; last done. + apply Rdiv_INR_ge_0. + * intros. repeat case_match; by simplify_eq. + * apply ex_seriesC_finite. + + + iIntros ([|?[|]]) "(%&?&?)"; try done. + iFrame. + iMod "Hclose". + iModIntro. + repeat iSplit; try done. + iPureIntro. + intros x y H'. + apply list_fmap_eq_inj in H'; first by simplify_eq. + apply fin_to_nat_inj. - iModIntro. iFrame. Qed. From 690707ae2ee49800df6d664d328949815ffc1892 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Thu, 26 Sep 2024 15:17:34 +0200 Subject: [PATCH 36/69] Simplify spec by removing P --- .../examples/random_counter/random_counter.v | 20 +++++------ theories/coneris/lib/hocap_flip.v | 31 +++++++--------- theories/coneris/lib/hocap_rand.v | 35 ++++++++----------- 3 files changed, 36 insertions(+), 50 deletions(-) diff --git a/theories/coneris/examples/random_counter/random_counter.v b/theories/coneris/examples/random_counter/random_counter.v index e4c28cbd..60d5430f 100644 --- a/theories/coneris/examples/random_counter/random_counter.v +++ b/theories/coneris/examples/random_counter/random_counter.v @@ -74,36 +74,34 @@ Class random_counter `{!conerisGS Σ} := RandCounter {{{ (v:val), RET v; ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ counter_tapes_frag (L:=L) γ1 α [] }}}; - incr_counter_tape_spec_some {L: counterG Σ} N E c γ1 γ2 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: + incr_counter_tape_spec_some {L: counterG Σ} N E c γ1 γ2 (Q:nat->iProp Σ) (α:loc) (n:nat) ns: ↑N⊆E -> {{{ is_counter (L:=L) N c γ1 γ2 ∗ - □ (∀ (z:nat), counter_content_auth (L:=L) γ2 z ∗ P ={E∖ ↑N, ∅}=∗ + (∀ (z:nat), counter_content_auth (L:=L) γ2 z ={E∖ ↑N, ∅}=∗ counter_tapes_frag (L:=L) γ1 α (n::ns) ∗ (counter_tapes_frag (L:=L) γ1 α ns ={∅, E∖↑N}=∗ counter_content_auth (L:=L) γ2 (z+n) ∗ Q z) - ) ∗ P + ) }}} incr_counter_tape c #lbl:α @ E {{{ (z:nat), RET (#z, #n); Q z}}}; counter_presample_spec {L: counterG Σ} NS E ns α - (ε2 : list nat -> R) - (P : iProp Σ) T γ1 γ2 c num ε: + (ε2 : list nat -> R) T γ1 γ2 c num ε: ↑NS ⊆ E -> (∀ n, 0<=ε2 n)%R -> (SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R-> is_counter (L:=L) NS c γ1 γ2 -∗ - □(P ={E∖↑NS,∅}=∗ ↯ ε ∗ counter_tapes_frag (L:=L) γ1 α ns ∗ + ( |={E∖↑NS,∅}=> ↯ ε ∗ counter_tapes_frag (L:=L) γ1 α ns ∗ (∀ ns', ⌜length ns' = num⌝ ∗ ↯ (ε2 ns') ∗ counter_tapes_frag (L:=L) γ1 α (ns++ns')={∅,E∖↑NS}=∗ T ns' - ))-∗ - P -∗ + ))-∗ wp_update E (∃ n, T n); - read_counter_spec {L: counterG Σ} N E c γ1 γ2 P Q: + read_counter_spec {L: counterG Σ} N E c γ1 γ2 Q: ↑N ⊆ E -> {{{ is_counter (L:=L) N c γ1 γ2 ∗ - □ (∀ (z:nat), P ∗ counter_content_auth (L:=L) γ2 z ={E∖↑N}=∗ + (∀ (z:nat), counter_content_auth (L:=L) γ2 z ={E∖↑N}=∗ counter_content_auth (L:=L) γ2 z∗ Q z) - ∗ P + }}} read_counter c @ E {{{ (n':nat), RET #n'; Q n' diff --git a/theories/coneris/lib/hocap_flip.v b/theories/coneris/lib/hocap_flip.v index cca4352f..7ea8657a 100644 --- a/theories/coneris/lib/hocap_flip.v +++ b/theories/coneris/lib/hocap_flip.v @@ -49,28 +49,26 @@ Class flip_spec `{!conerisGS Σ} := FlipSpec ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ flip_tapes_frag (L:=L) γ1 α [] }}}; - flip_tape_spec_some {L: flipG Σ} N E γ1 (P: iProp Σ) (Q:bool -> list bool -> iProp Σ) (α:loc) : + flip_tape_spec_some {L: flipG Σ} N E γ1 (Q:bool -> list bool -> iProp Σ) (α:loc) : ↑N⊆E -> {{{ is_flip (L:=L) N γ1 ∗ - (□ (P ={E∖↑N, ∅}=∗ + (|={E∖↑N, ∅}=> ∃ n ns, flip_tapes_frag (L:=L) γ1 α (n::ns) ∗ - (flip_tapes_frag (L:=L) γ1 α ns ={∅, E∖↑N}=∗ Q n ns))) ∗ - ▷ P + (flip_tapes_frag (L:=L) γ1 α ns ={∅, E∖↑N}=∗ Q n ns)) }}} flip_tape #lbl:α @ E {{{ (n:bool), RET #n; ∃ ns, Q n ns}}}; - flip_presample_spec {L: flipG Σ} NS E (P : iProp Σ) T γ1: + flip_presample_spec {L: flipG Σ} NS E T γ1: ↑NS ⊆ E -> is_flip (L:=L) NS γ1 -∗ - (□ (P ={E∖↑NS, ∅}=∗ + (|={E∖↑NS, ∅}=> ∃ ε num ε2 α ns, ↯ ε ∗ flip_tapes_frag (L:=L) γ1 α ns ∗ ⌜(∀ l, length l = num -> 0<=ε2 l)%R⌝ ∗ ⌜(SeriesC (λ l, if length l =? num then ε2 l else 0) /(2^num) <= ε)%R⌝∗ (∀ ns', ⌜length ns' = num⌝ ∗ ↯ (ε2 ns') ∗ flip_tapes_frag (L:=L) γ1 α (ns ++ ns') - ={∅, E∖↑NS}=∗ T ε num ε2 α ns ns'))) + ={∅, E∖↑NS}=∗ T ε num ε2 α ns ns')) -∗ - P -∗ wp_update E (∃ ε num ε2 α ns ns', T ε num ε2 α ns ns') }. @@ -140,16 +138,15 @@ Section instantiate_flip. Qed. Next Obligation. simpl. - iIntros (????? Q ? ? Φ) "(#Hinv & #Hvs & HP) HΦ". + iIntros (???? Q ? ? Φ) "(#Hinv & Hvs) HΦ". wp_pures. - wp_apply (rand_tape_spec_some _ _ _ P (λ n ns, ∃ b bs, ⌜n= bool_to_nat b⌝ ∗ + wp_apply (rand_tape_spec_some _ _ _ (λ n ns, ∃ b bs, ⌜n= bool_to_nat b⌝ ∗ ⌜ns = bool_to_nat <$> bs⌝ ∗ Q b bs )%I with "[-HΦ]"); [done|..]. - iFrame. iSplit; first done. - iModIntro. iIntros "HP". - iMod ("Hvs" with "[$]") as "(%b & %bs & Hfrag & Hrest)". + iMod ("Hvs") as "(%b & %bs & Hfrag & Hrest)". iFrame. iModIntro. iIntros "?". @@ -167,17 +164,15 @@ Section instantiate_flip. Qed. Next Obligation. simpl. - iIntros (???? T ??) "#Hinv #Hvs HP". - iMod (rand_presample_spec _ _ P + iIntros (??? T ??) "#Hinv Hvs". + iMod (rand_presample_spec _ _ (λ ε num tb ε2 α ns ns', ∃ bs bs' ε2', ⌜tb = 1%nat⌝ ∗ ⌜fmap (FMap:= list_fmap) bool_to_nat bs = ns⌝ ∗ ⌜fmap (FMap:= list_fmap) bool_to_nat bs'=fin_to_nat <$> ns'⌝ ∗ ⌜∀ xs ys, fmap (FMap:= list_fmap) bool_to_nat xs = fmap (FMap:= list_fmap)fin_to_nat ys -> ε2' xs = ε2 ys⌝ ∗ T ε num ε2' α bs bs' )%I - with "[][][$]") as "(%&%&%&%&%&%&%&%&%&%&->&<-&%&%&HT)"; [exact|exact| |iModIntro; iFrame]. - iModIntro. - iIntros "HP". - iMod ("Hvs" with "[$]") as "(%&%&%&%&%&Herr&Hfrag&%Hpos&%Hineq&Hvs')". + with "[][-]") as "(%&%&%&%&%&%&%&%&%&%&->&<-&%&%&HT)"; [exact|exact| |iModIntro; iFrame]. + iMod "Hvs" as "(%&%&%&%&%&Herr&Hfrag&%Hpos&%Hineq&Hvs')". iFrame. iModIntro. iExists num, (λ ls, ε2 (nat_to_bool<$>(fin_to_nat <$> ls))). diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index 42a7ed6f..c29c8d81 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -51,27 +51,24 @@ Class rand_spec `{!conerisGS Σ} := RandSpec {{{ (v:val), RET v; ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ rand_tapes_frag (L:=L) γ2 α (tb, []) }}}; - rand_tape_spec_some {L: randG Σ} N E γ2 (P: iProp Σ) (Q: nat ->list nat -> iProp Σ) (α:loc) (tb:nat) : + rand_tape_spec_some {L: randG Σ} N E γ2 (Q: nat ->list nat -> iProp Σ) (α:loc) (tb:nat) : ↑N⊆E -> {{{ is_rand (L:=L) N γ2 ∗ - (□ (P ={E∖↑N, ∅}=∗ ∃ n ns, rand_tapes_frag (L:=L) γ2 α (tb, n::ns) ∗ - (rand_tapes_frag (L:=L) γ2 α (tb, ns) ={∅, E∖↑N}=∗ Q n ns))) ∗ - ▷ P + (|={E∖↑N, ∅}=> ∃ n ns, rand_tapes_frag (L:=L) γ2 α (tb, n::ns) ∗ + (rand_tapes_frag (L:=L) γ2 α (tb, ns) ={∅, E∖↑N}=∗ Q n ns)) }}} rand_tape #lbl:α #tb @ E {{{ (n:nat), RET #n; ∃ ns, Q n ns}}}; - rand_presample_spec {L: randG Σ} N E (P : iProp Σ) T γ2 : + rand_presample_spec {L: randG Σ} N E T γ2 : ↑N ⊆ E -> is_rand (L:=L) N γ2 -∗ - (□ (P ={E∖↑N, ∅}=∗ + (|={E∖↑N, ∅}=> ∃ ε num tb (ε2 : list (fin (S tb)) -> R) α ns, ↯ ε ∗ rand_tapes_frag (L:=L) γ2 α (tb, ns) ∗ ⌜(∀ l, length l = num -> 0<= ε2 l)%R⌝ ∗ ⌜(SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R⌝ ∗ (∀ ns', ⌜length ns' = num⌝ ∗ ↯ (ε2 ns') ∗ rand_tapes_frag (L:=L) γ2 α (tb, ns ++ (fin_to_nat <$> ns')) - ={∅, E∖↑N}=∗ T ε num tb ε2 α ns ns'))) - -∗ - P -∗ + ={∅, E∖↑N}=∗ T ε num tb ε2 α ns ns')) -∗ wp_update E (∃ ε num tb ε2 α ns ns', T ε num tb ε2 α ns ns') }. @@ -148,11 +145,11 @@ Section impl. Qed. Next Obligation. simpl. - iIntros (??????????? Φ) "(#Hinv & #Hvs & HP) HΦ". + iIntros (?????????? Φ) "(#Hinv & Hvs) HΦ". wp_pures. wp_bind (rand(_) _)%E. iInv "Hinv" as ">(%&H3&H4)" "Hclose". - iMod ("Hvs" with "[$]") as "(%n & %ns & [Hfrag %] & Hrest)". + iMod ("Hvs") as "(%n & %ns & [Hfrag %] & Hrest)". iDestruct (hocap_tapes_agree with "[$][$]") as "%". iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. simpl. @@ -170,11 +167,11 @@ Section impl. Qed. Next Obligation. simpl. - iIntros (?????????) "#Hinv #Hvs HP ". + iIntros (????????) "#Hinv Hvs". iApply wp_update_state_step_coupl. iIntros (σ1 ε_supply) "((Hheap&Htapes)&Hε)". iMod (inv_acc with "Hinv") as "[>(% & H3 & H4 ) Hclose]"; [done|]. - iMod ("Hvs" with "[$]") as "(%err & %num & %tb & %ε2 & %α & %ns & Herr & [Hfrag %] & %Hpos & %Hineq & Hrest)". + iMod "Hvs" as "(%err & %num & %tb & %ε2 & %α & %ns & Herr & [Hfrag %] & %Hpos & %Hineq & Hrest)". iDestruct (hocap_tapes_agree with "[$][$]") as "%". iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". @@ -240,9 +237,8 @@ Section checks. {{{ RET #(LitInt n); rand_tapes_frag (L:=L) γ1 α (N, ns) ∗ ⌜n <= N⌝ }}}. Proof. iIntros (Hsubset Φ) "(#Hinv &>Hfrag) HΦ". - wp_apply (rand_tape_spec_some _ _ _ _ (λ n' ns', ⌜n=n'/\ns=ns'⌝ ∗ ⌜n<=N⌝ ∗ rand_tapes_frag _ _ _)%I with "[Hfrag]"); first exact. - - iSplit; first done. iSplit; last iApply "Hfrag". - iModIntro. iIntros "Hfrag". + wp_apply (rand_tape_spec_some _ _ _ (λ n' ns', ⌜n=n'/\ns=ns'⌝ ∗ ⌜n<=N⌝ ∗ rand_tapes_frag _ _ _)%I with "[Hfrag]"); first exact. + - iSplit; first done. iDestruct (rand_tapes_valid with "[$]") as "%H'". iFrame. iApply fupd_mask_intro; first set_solver. @@ -270,17 +266,14 @@ Section checks. iDestruct (ec_valid with "[$]") as "%Hpos'". destruct Hpos' as [Hpos' ?]. iMod (rand_presample_spec _ _ - (rand_tapes_frag (L:=L) γ1 α (N, ns) ∗ ↯ ε1) (λ ε' (num':nat) tb' ε2' α' ns1 ns2, ⌜ε1=ε'⌝ ∗ ⌜(1=num')%nat⌝ ∗ ⌜N=tb'⌝ ∗ ⌜∀ x y, fin_to_nat <$> x = fin_to_nat <$> y -> ε2' x = (λ l, match l with |[x] => ε2 x | _ => 1%R end) y⌝∗ ⌜α=α'⌝ ∗ ⌜ns1=ns⌝ ∗ ∃ (n:fin (S N)), ⌜fin_to_nat <$> ns2=[fin_to_nat n]⌝ ∗ ↯ (ε2 n) ∗ rand_tapes_frag γ1 α (N, ns ++ [fin_to_nat n]))%I - with "[//][][$Htape $Herr]") as "(%&%&%&%&%&%&%&->&<-&->&%&->&->&%&%&?&?)". + with "[//][- ]") as "(%&%&%&%&%&%&%&->&<-&->&%&->&->&%&%&?&?)". - done. - - iModIntro. - iIntros "[Hfrag Herr]". - iFrame. + - iFrame. iApply fupd_mask_intro; first set_solver. iIntros "Hclose". iExists 1, (λ l, match l with |[x] => ε2 x | _ => 1%R end). simpl. From fb3f8f75eae451d7f93bdd97632584725a228fa9 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Thu, 26 Sep 2024 15:23:03 +0200 Subject: [PATCH 37/69] NIT --- .../examples/random_counter/random_counter.v | 23 ++++++++++--------- 1 file changed, 12 insertions(+), 11 deletions(-) diff --git a/theories/coneris/examples/random_counter/random_counter.v b/theories/coneris/examples/random_counter/random_counter.v index 60d5430f..f111adb6 100644 --- a/theories/coneris/examples/random_counter/random_counter.v +++ b/theories/coneris/examples/random_counter/random_counter.v @@ -74,28 +74,29 @@ Class random_counter `{!conerisGS Σ} := RandCounter {{{ (v:val), RET v; ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ counter_tapes_frag (L:=L) γ1 α [] }}}; - incr_counter_tape_spec_some {L: counterG Σ} N E c γ1 γ2 (Q:nat->iProp Σ) (α:loc) (n:nat) ns: + incr_counter_tape_spec_some {L: counterG Σ} N E c γ1 γ2 (Q:nat->nat->list nat->iProp Σ) (α:loc) : ↑N⊆E -> {{{ is_counter (L:=L) N c γ1 γ2 ∗ (∀ (z:nat), counter_content_auth (L:=L) γ2 z ={E∖ ↑N, ∅}=∗ - counter_tapes_frag (L:=L) γ1 α (n::ns) ∗ + ∃ n ns, counter_tapes_frag (L:=L) γ1 α (n::ns) ∗ (counter_tapes_frag (L:=L) γ1 α ns ={∅, E∖↑N}=∗ - counter_content_auth (L:=L) γ2 (z+n) ∗ Q z) + counter_content_auth (L:=L) γ2 (z+n) ∗ Q z n ns) ) }}} incr_counter_tape c #lbl:α @ E - {{{ (z:nat), RET (#z, #n); Q z}}}; - counter_presample_spec {L: counterG Σ} NS E ns α - (ε2 : list nat -> R) T γ1 γ2 c num ε: + {{{ (z n:nat), RET (#z, #n); ∃ ns, Q z n ns}}}; + counter_presample_spec {L: counterG Σ} NS E T γ1 γ2 c: ↑NS ⊆ E -> - (∀ n, 0<=ε2 n)%R -> - (SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R-> is_counter (L:=L) NS c γ1 γ2 -∗ - ( |={E∖↑NS,∅}=> ↯ ε ∗ counter_tapes_frag (L:=L) γ1 α ns ∗ + ( |={E∖↑NS,∅}=> + ∃ ε α ε2 num ns, + ↯ ε ∗ counter_tapes_frag (L:=L) γ1 α ns ∗ + ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ + ⌜(SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R⌝ ∗ (∀ ns', ⌜length ns' = num⌝ ∗ ↯ (ε2 ns') ∗ counter_tapes_frag (L:=L) γ1 α (ns++ns')={∅,E∖↑NS}=∗ - T ns' + T ε α ε2 num ns ns' ))-∗ - wp_update E (∃ n, T n); + wp_update E (∃ ε α ε2 num ns ns', T ε α ε2 num ns ns'); read_counter_spec {L: counterG Σ} N E c γ1 γ2 Q: ↑N ⊆ E -> {{{ is_counter (L:=L) N c γ1 γ2 ∗ From 92001283bd87770e4e2b7ff442d14bd2c029b3ef Mon Sep 17 00:00:00 2001 From: Hei Li Date: Thu, 26 Sep 2024 15:58:03 +0200 Subject: [PATCH 38/69] NIT --- .../examples/random_counter/random_counter.v | 176 ++++++++++-------- theories/coneris/lib/hocap_flip.v | 10 +- theories/coneris/lib/hocap_rand.v | 7 +- 3 files changed, 101 insertions(+), 92 deletions(-) diff --git a/theories/coneris/examples/random_counter/random_counter.v b/theories/coneris/examples/random_counter/random_counter.v index f111adb6..1e94342e 100644 --- a/theories/coneris/examples/random_counter/random_counter.v +++ b/theories/coneris/examples/random_counter/random_counter.v @@ -93,7 +93,7 @@ Class random_counter `{!conerisGS Σ} := RandCounter ↯ ε ∗ counter_tapes_frag (L:=L) γ1 α ns ∗ ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ ⌜(SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R⌝ ∗ - (∀ ns', ⌜length ns' = num⌝ ∗ ↯ (ε2 ns') ∗ counter_tapes_frag (L:=L) γ1 α (ns++ns')={∅,E∖↑NS}=∗ + (∀ ns', ↯ (ε2 ns') ∗ counter_tapes_frag (L:=L) γ1 α (ns++ns')={∅,E∖↑NS}=∗ T ε α ε2 num ns ns' ))-∗ wp_update E (∃ ε α ε2 num ns ns', T ε α ε2 num ns ns'); @@ -110,87 +110,99 @@ Class random_counter `{!conerisGS Σ} := RandCounter }. -(* Section lemmas. *) -(* Context `{rc:random_counter} {L: counterG Σ}. *) +Section lemmas. + Context `{rc:random_counter} {L: counterG Σ}. -(* Lemma incr_counter_tape_spec_none N E c γ1 γ2 γ3 (ε2:R -> nat -> R) (P: iProp Σ) (T: nat -> iProp Σ) (Q: nat -> nat -> iProp Σ)(α:loc): *) -(* ↑N ⊆ E-> *) -(* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) -(* (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → *) -(* {{{ is_counter (L:=L) N c γ1 γ2 γ3 ∗ *) -(* □(∀ (ε ε': nonnegreal) (n : nat), P ∗ counter_error_auth (L:=L) γ1 (ε+ε') *) -(* ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨ counter_error_auth (L:=L) γ1 (ε2 ε n + ε') ∗ T n) ) ∗ *) -(* □ (∀ (n:nat) (z:nat), T n ∗ counter_content_auth (L:=L) γ3 z ={E∖↑N}=∗ *) -(* counter_content_auth (L:=L) γ3 (z+n)%nat ∗ Q z n) ∗ *) -(* P ∗ counter_tapes_frag (L:=L) γ2 α [] *) -(* }}} *) -(* incr_counter_tape c #lbl:α @ E *) -(* {{{ (z:nat) (n:nat), RET (#z, #n); Q z n ∗ counter_tapes_frag (L:=L) γ2 α [] }}}. *) -(* Proof. *) -(* iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP & Hα) HΦ". *) -(* iMod (counter_presample_spec with "[//][//][$][$]") as "(%&HT&Hα)"; try done. *) -(* { intros ε Hε. specialize (Hineq ε Hε). *) -(* rewrite SeriesC_nat_bounded_fin SeriesC_finite_foldr /=. lra. *) -(* } *) -(* iApply (incr_counter_tape_spec_some _ _ _ _ _ _ (T n) (λ x, Q x n) with "[$Hα $HT]"); try done. *) -(* { by iSplit. } *) -(* iNext. *) -(* iIntros. iApply ("HΦ" with "[$]"). *) -(* Qed. *) + Lemma incr_counter_tape_spec_none N E c γ1 γ2 Q α: + ↑N ⊆ E-> + {{{ is_counter (L:=L) N c γ1 γ2 ∗ + ( |={E∖↑N, ∅}=> + ∃ ε ε2, + ↯ ε ∗ counter_tapes_frag (L:=L) γ1 α [] ∗ + ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ + ⌜(((ε2 0%nat) + (ε2 1%nat)+ (ε2 2%nat)+ (ε2 3%nat))/4 <= ε)%R⌝ ∗ + (∀ n, |={∅,E∖↑N}=> + ∀ (z:nat), counter_content_auth (L:=L) γ2 z ={E∖↑N, ∅}=∗ + ↯ (ε2 n) ∗ counter_tapes_frag (L:=L) γ1 α [] ={∅, E∖↑N}=∗ + counter_content_auth (L:=L) γ2 (z+n) ∗ Q z n ε ε2 + ) + ) + }}} + incr_counter_tape c #lbl:α @ E + {{{ (z n:nat), RET (#z, #n); ∃ ε ε2, Q z n ε ε2 }}}. + Proof. + iIntros (Hsubset Φ) "(#Hinv & Hvs) HΦ". + iMod (counter_presample_spec _ _ + (λ ε α' ε2 num ns ns', + ∃ n z, + ⌜num=1%nat⌝ ∗ ⌜α'=α⌝ ∗ ⌜ns=[]⌝ ∗ ⌜ns'=[n]⌝ ∗ + Q z n ε (λ x, ε2 [x]) + )%I with "[//][-]") as "?"; try done. + Admitted. + (* { iMod "Hvs" as "(%ε & %ε2 & Herr &Hfrag & %Hpos & %Hineq & Hrest)". *) + (* iFrame. *) + (* intros ε Hε. specialize (Hineq ε Hε). *) + (* rewrite SeriesC_nat_bounded_fin SeriesC_finite_foldr /=. lra. *) + (* } *) + (* iApply (incr_counter_tape_spec_some _ _ _ _ _ _ (T n) (λ x, Q x n) with "[$Hα $HT]"); try done. *) + (* { by iSplit. } *) + (* iNext. *) + (* iIntros. iApply ("HΦ" with "[$]"). *) + (* Qed. *) -(* Lemma incr_counter_tape_spec_none' N E c γ1 γ2 γ3 ε (ε2:R -> nat -> R)(α:loc) (ns:list nat) (q:Qp) (z:nat): *) -(* ↑N ⊆ E-> *) -(* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) -(* (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → *) -(* {{{ is_counter (L:=L) N c γ1 γ2 γ3 ∗ *) -(* counter_error_frag (L:=L) γ1 ε ∗ *) -(* counter_tapes_frag (L:=L) γ2 α [] ∗ *) -(* counter_content_frag (L:=L) γ3 q z *) -(* }}} *) -(* incr_counter_tape c #lbl:α @ E *) -(* {{{ (z':nat) (n:nat), *) -(* RET (#z', #n); ⌜(0<=n<4)%nat⌝ ∗ *) -(* ⌜(z<=z')%nat⌝ ∗ *) -(* counter_error_frag (L:=L) γ1 (ε2 ε n) ∗ *) -(* counter_tapes_frag (L:=L) γ2 α [] ∗ *) -(* counter_content_frag (L:=L) γ3 q (z+n)%nat }}}. *) -(* Proof. *) -(* iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & Herr & Htapes & Hcontent) HΦ". *) -(* pose (ε2' := (λ ε x, if (x<=?3)%nat then ε2 ε x else 1%R)). *) -(* wp_apply (incr_counter_tape_spec_none _ _ _ _ _ _ ε2' *) -(* (counter_error_frag γ1 ε ∗ counter_content_frag γ3 q z)%I *) -(* (λ n, ⌜0<=n<4⌝ ∗ counter_error_frag γ1 (ε2' ε n) ∗ counter_content_frag γ3 q z)%I *) -(* (λ z' n, ⌜0<=n<4⌝ ∗ ⌜z<=z'⌝ ∗ counter_error_frag γ1 (ε2' ε n) ∗ counter_content_frag γ3 q (z+n)%nat)%I *) -(* with "[$Herr $Htapes $Hcontent]"). *) -(* - done. *) -(* - rewrite /ε2'. intros. *) -(* case_match; [naive_solver|lra]. *) -(* - rewrite /ε2'. simpl. intros. naive_solver. *) -(* - repeat iSplit; first done. *) -(* + iModIntro. *) -(* iIntros (εa εb n) "[(Herr &Hcontent) Herr_auth]". *) -(* destruct (n<=?3) eqn:H; last first. *) -(* { iLeft. iPureIntro. rewrite /ε2'. by rewrite H. } *) -(* iRight. *) -(* iFrame. *) -(* iDestruct (counter_error_ineq with "[$][$]") as "%". *) -(* iDestruct (counter_error_auth_valid with "[$]") as "%". *) -(* iDestruct (counter_error_frag_valid with "[$]") as "%". *) -(* unshelve iMod (counter_error_update _ _ _ (mknonnegreal (ε2' ε n) _) with "[$][$]") as "[$$]". *) -(* { rewrite /ε2' H. *) -(* naive_solver. } *) -(* iPureIntro. split; first lia. *) -(* apply leb_complete in H. lia. *) -(* + iModIntro. *) -(* iIntros (??) "[(%&Herr & Hcontent) Hcontent_auth]". *) -(* iFrame. *) -(* iDestruct (counter_content_less_than with "[$][$]") as "%". *) -(* by iMod (counter_content_update with "[$][$]") as "[$ $]". *) -(* - iIntros (??) "[(%&%&?&?)?]". *) -(* iApply "HΦ". *) -(* iFrame. *) -(* rewrite /ε2'. case_match eqn: H2; first by iFrame. *) -(* apply leb_iff_conv in H2. lia. *) -(* Qed. *) + (* Lemma incr_counter_tape_spec_none' N E c γ1 γ2 γ3 ε (ε2:R -> nat -> R)(α:loc) (ns:list nat) (q:Qp) (z:nat): *) + (* ↑N ⊆ E-> *) + (* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) + (* (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → *) + (* {{{ is_counter (L:=L) N c γ1 γ2 γ3 ∗ *) + (* counter_error_frag (L:=L) γ1 ε ∗ *) + (* counter_tapes_frag (L:=L) γ2 α [] ∗ *) + (* counter_content_frag (L:=L) γ3 q z *) + (* }}} *) + (* incr_counter_tape c #lbl:α @ E *) + (* {{{ (z':nat) (n:nat), *) + (* RET (#z', #n); ⌜(0<=n<4)%nat⌝ ∗ *) + (* ⌜(z<=z')%nat⌝ ∗ *) + (* counter_error_frag (L:=L) γ1 (ε2 ε n) ∗ *) + (* counter_tapes_frag (L:=L) γ2 α [] ∗ *) + (* counter_content_frag (L:=L) γ3 q (z+n)%nat }}}. *) + (* Proof. *) + (* iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & Herr & Htapes & Hcontent) HΦ". *) + (* pose (ε2' := (λ ε x, if (x<=?3)%nat then ε2 ε x else 1%R)). *) + (* wp_apply (incr_counter_tape_spec_none _ _ _ _ _ _ ε2' *) + (* (counter_error_frag γ1 ε ∗ counter_content_frag γ3 q z)%I *) + (* (λ n, ⌜0<=n<4⌝ ∗ counter_error_frag γ1 (ε2' ε n) ∗ counter_content_frag γ3 q z)%I *) + (* (λ z' n, ⌜0<=n<4⌝ ∗ ⌜z<=z'⌝ ∗ counter_error_frag γ1 (ε2' ε n) ∗ counter_content_frag γ3 q (z+n)%nat)%I *) + (* with "[$Herr $Htapes $Hcontent]"). *) + (* - done. *) + (* - rewrite /ε2'. intros. *) + (* case_match; [naive_solver|lra]. *) + (* - rewrite /ε2'. simpl. intros. naive_solver. *) + (* - repeat iSplit; first done. *) + (* + iModIntro. *) + (* iIntros (εa εb n) "[(Herr &Hcontent) Herr_auth]". *) + (* destruct (n<=?3) eqn:H; last first. *) + (* { iLeft. iPureIntro. rewrite /ε2'. by rewrite H. } *) + (* iRight. *) + (* iFrame. *) + (* iDestruct (counter_error_ineq with "[$][$]") as "%". *) + (* iDestruct (counter_error_auth_valid with "[$]") as "%". *) + (* iDestruct (counter_error_frag_valid with "[$]") as "%". *) + (* unshelve iMod (counter_error_update _ _ _ (mknonnegreal (ε2' ε n) _) with "[$][$]") as "[$$]". *) + (* { rewrite /ε2' H. *) + (* naive_solver. } *) + (* iPureIntro. split; first lia. *) + (* apply leb_complete in H. lia. *) + (* + iModIntro. *) + (* iIntros (??) "[(%&Herr & Hcontent) Hcontent_auth]". *) + (* iFrame. *) + (* iDestruct (counter_content_less_than with "[$][$]") as "%". *) + (* by iMod (counter_content_update with "[$][$]") as "[$ $]". *) + (* - iIntros (??) "[(%&%&?&?)?]". *) + (* iApply "HΦ". *) + (* iFrame. *) + (* rewrite /ε2'. case_match eqn: H2; first by iFrame. *) + (* apply leb_iff_conv in H2. lia. *) + (* Qed. *) -(* End lemmas. *) +End lemmas. diff --git a/theories/coneris/lib/hocap_flip.v b/theories/coneris/lib/hocap_flip.v index 7ea8657a..dab64a01 100644 --- a/theories/coneris/lib/hocap_flip.v +++ b/theories/coneris/lib/hocap_flip.v @@ -66,7 +66,7 @@ Class flip_spec `{!conerisGS Σ} := FlipSpec ∃ ε num ε2 α ns, ↯ ε ∗ flip_tapes_frag (L:=L) γ1 α ns ∗ ⌜(∀ l, length l = num -> 0<=ε2 l)%R⌝ ∗ ⌜(SeriesC (λ l, if length l =? num then ε2 l else 0) /(2^num) <= ε)%R⌝∗ - (∀ ns', ⌜length ns' = num⌝ ∗ ↯ (ε2 ns') ∗ flip_tapes_frag (L:=L) γ1 α (ns ++ ns') + (∀ ns', ↯ (ε2 ns') ∗ flip_tapes_frag (L:=L) γ1 α (ns ++ ns') ={∅, E∖↑NS}=∗ T ε num ε2 α ns ns')) -∗ wp_update E (∃ ε num ε2 α ns ns', T ε num ε2 α ns ns') @@ -203,11 +203,9 @@ Section instantiate_flip. + eapply ex_seriesC_list_length. intros. case_match; last done. by rewrite -Nat.eqb_eq. - - iIntros (ls) "(H1&H2&H3)". - iMod ("Hvs'" with "[H1 $H2 H3]") as "?". - + rewrite !fmap_length !fmap_app. - iSplitL "H1"; first done. - rewrite -!list_fmap_compose. + - iIntros (ls) "(H2&H3)". + iMod ("Hvs'" with "[$H2 H3]") as "?". + + rewrite fmap_app. rewrite -!list_fmap_compose. erewrite (Forall_fmap_ext_1 (_∘_)); first done. apply Forall_true. intros x; by repeat (inv_fin x; simpl; try intros x). diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index c29c8d81..aee65851 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -67,7 +67,7 @@ Class rand_spec `{!conerisGS Σ} := RandSpec ↯ ε ∗ rand_tapes_frag (L:=L) γ2 α (tb, ns) ∗ ⌜(∀ l, length l = num -> 0<= ε2 l)%R⌝ ∗ ⌜(SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R⌝ ∗ - (∀ ns', ⌜length ns' = num⌝ ∗ ↯ (ε2 ns') ∗ rand_tapes_frag (L:=L) γ2 α (tb, ns ++ (fin_to_nat <$> ns')) + (∀ ns', ↯ (ε2 ns') ∗ rand_tapes_frag (L:=L) γ2 α (tb, ns ++ (fin_to_nat <$> ns')) ={∅, E∖↑N}=∗ T ε num tb ε2 α ns ns')) -∗ wp_update E (∃ ε num tb ε2 α ns ns', T ε num tb ε2 α ns ns') }. @@ -212,7 +212,7 @@ Section impl. { simpl; lra. } iMod (hocap_tapes_update with "[$][$]") as "[Hauth Hfrag]". iMod ("Hrest" $! sample with "[$Herr $Hfrag]") as "HT". - { iPureIntro; split; first done. + { iPureIntro. rewrite Forall_app; split; subst; first done. eapply Forall_impl; first apply fin.fin_forall_leq. simpl; intros; lia. @@ -298,8 +298,7 @@ Section checks. apply Rdiv_INR_ge_0. * intros. repeat case_match; by simplify_eq. * apply ex_seriesC_finite. - + - iIntros ([|?[|]]) "(%&?&?)"; try done. + + iIntros ([|?[|]]) "(?&?)"; [by iDestruct (ec_contradict with "[$]") as "%"| |by iDestruct (ec_contradict with "[$]") as "%"]. iFrame. iMod "Hclose". iModIntro. From 95f38e62bb4856c8f3f313e7b4e0250c78dfb09d Mon Sep 17 00:00:00 2001 From: Hei Li Date: Thu, 26 Sep 2024 16:01:38 +0200 Subject: [PATCH 39/69] NIT --- theories/coneris/examples/random_counter/random_counter.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/coneris/examples/random_counter/random_counter.v b/theories/coneris/examples/random_counter/random_counter.v index 1e94342e..6a9c9a28 100644 --- a/theories/coneris/examples/random_counter/random_counter.v +++ b/theories/coneris/examples/random_counter/random_counter.v @@ -136,7 +136,7 @@ Section lemmas. (λ ε α' ε2 num ns ns', ∃ n z, ⌜num=1%nat⌝ ∗ ⌜α'=α⌝ ∗ ⌜ns=[]⌝ ∗ ⌜ns'=[n]⌝ ∗ - Q z n ε (λ x, ε2 [x]) + _ )%I with "[//][-]") as "?"; try done. Admitted. (* { iMod "Hvs" as "(%ε & %ε2 & Herr &Hfrag & %Hpos & %Hineq & Hrest)". *) From 357f87c66304211fe4d5c6f351dba45f4935623f Mon Sep 17 00:00:00 2001 From: Hei Li Date: Fri, 27 Sep 2024 10:27:54 +0200 Subject: [PATCH 40/69] lemma incr_counter_tape_spec_none --- .../examples/random_counter/random_counter.v | 62 +++++++++++++------ 1 file changed, 43 insertions(+), 19 deletions(-) diff --git a/theories/coneris/examples/random_counter/random_counter.v b/theories/coneris/examples/random_counter/random_counter.v index 6a9c9a28..34218b76 100644 --- a/theories/coneris/examples/random_counter/random_counter.v +++ b/theories/coneris/examples/random_counter/random_counter.v @@ -121,12 +121,14 @@ Section lemmas. ↯ ε ∗ counter_tapes_frag (L:=L) γ1 α [] ∗ ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ ⌜(((ε2 0%nat) + (ε2 1%nat)+ (ε2 2%nat)+ (ε2 3%nat))/4 <= ε)%R⌝ ∗ - (∀ n, |={∅,E∖↑N}=> - ∀ (z:nat), counter_content_auth (L:=L) γ2 z ={E∖↑N, ∅}=∗ - ↯ (ε2 n) ∗ counter_tapes_frag (L:=L) γ1 α [] ={∅, E∖↑N}=∗ + (∀ n, ↯ (ε2 n) ∗ counter_tapes_frag (L:=L) γ1 α [n] ={∅,E∖↑N}=∗ + ∀ (z:nat), counter_content_auth (L:=L) γ2 z ={E∖↑N, ∅}=∗ + counter_tapes_frag (L:=L) γ1 α [n]∗ + (counter_tapes_frag (L:=L) γ1 α []={∅, E∖↑N}=∗ counter_content_auth (L:=L) γ2 (z+n) ∗ Q z n ε ε2 ) - ) + ) + ) }}} incr_counter_tape c #lbl:α @ E {{{ (z n:nat), RET (#z, #n); ∃ ε ε2, Q z n ε ε2 }}}. @@ -134,22 +136,44 @@ Section lemmas. iIntros (Hsubset Φ) "(#Hinv & Hvs) HΦ". iMod (counter_presample_spec _ _ (λ ε α' ε2 num ns ns', - ∃ n z, + ∃ n ε2', ⌜num=1%nat⌝ ∗ ⌜α'=α⌝ ∗ ⌜ns=[]⌝ ∗ ⌜ns'=[n]⌝ ∗ - _ - )%I with "[//][-]") as "?"; try done. - Admitted. - (* { iMod "Hvs" as "(%ε & %ε2 & Herr &Hfrag & %Hpos & %Hineq & Hrest)". *) - (* iFrame. *) - (* intros ε Hε. specialize (Hineq ε Hε). *) - (* rewrite SeriesC_nat_bounded_fin SeriesC_finite_foldr /=. lra. *) - (* } *) - (* iApply (incr_counter_tape_spec_some _ _ _ _ _ _ (T n) (λ x, Q x n) with "[$Hα $HT]"); try done. *) - (* { by iSplit. } *) - (* iNext. *) - (* iIntros. iApply ("HΦ" with "[$]"). *) - (* Qed. *) - + ⌜(ε2' = λ x, ε2 [x])%R⌝ ∗ + (∀ z : nat, + counter_content_auth γ2 z ={E ∖ ↑N,∅}=∗ + counter_tapes_frag γ1 α [n] ∗ + (counter_tapes_frag γ1 α [] ={∅,E ∖ ↑N}=∗ + counter_content_auth γ2 (z + n) ∗ Q z n ε ε2')) + )%I with "[//][-HΦ]") as "(%ε & % & %ε2 & %num & %ns & %ns' & %n & %ε2' & -> & -> & -> & -> & -> & Hrest)"; try done. + - iMod "Hvs" as "(%ε & %ε2 & Herr & Hfrag & %Hpos & %Hineq & Hrest)". + iFrame. iModIntro. + iExists (λ l, match l with |[x] => ε2 x | _ => 1 end)%R, 1%nat. + repeat iSplit. + + iPureIntro. intros; repeat case_match; try lra. naive_solver. + + iPureIntro. etrans; last exact. + rewrite SeriesC_list. + * Local Transparent enum_uniform_fin_list. + simpl. lra. + * apply NoDup_fmap_2; last apply NoDup_enum_uniform_fin_list. + apply list_fmap_eq_inj, fin_to_nat_inj. + + iIntros (ns') "[Herr Hfrag]". destruct (ns') as [|n[|??]]; [by iDestruct (ec_contradict with "[$]") as "%"| |by iDestruct (ec_contradict with "[$]") as "%"]. + iMod ("Hrest" $! n with "[$]"). + iFrame. + by iPureIntro. + - wp_apply (incr_counter_tape_spec_some _ _ _ _ _ (λ z n ns, ∃ ε ε2, Q z n ε ε2)%I with "[Hrest]"); first exact. + + iSplit; first done. + iIntros (z) "Hauth". + iMod ("Hrest" with "[$]") as "[Hfrag Hrest]". + iFrame. + iModIntro. + iIntros "Hfrag". + iMod ("Hrest" with "[$]") as "[Hauth HQ]". + by iFrame. + + iIntros (??) "(%&%&%&HQ)". + iApply "HΦ". + iFrame. + Qed. + (* Lemma incr_counter_tape_spec_none' N E c γ1 γ2 γ3 ε (ε2:R -> nat -> R)(α:loc) (ns:list nat) (q:Qp) (z:nat): *) (* ↑N ⊆ E-> *) (* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) From d0d1865f314df9ed138e7ee35f60dc203b0bbac0 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Fri, 27 Sep 2024 10:58:42 +0200 Subject: [PATCH 41/69] sequential spec incr_counter_tape_spec_none' --- .../examples/random_counter/random_counter.v | 117 ++++++++++-------- 1 file changed, 63 insertions(+), 54 deletions(-) diff --git a/theories/coneris/examples/random_counter/random_counter.v b/theories/coneris/examples/random_counter/random_counter.v index 34218b76..54051f64 100644 --- a/theories/coneris/examples/random_counter/random_counter.v +++ b/theories/coneris/examples/random_counter/random_counter.v @@ -174,59 +174,68 @@ Section lemmas. iFrame. Qed. - (* Lemma incr_counter_tape_spec_none' N E c γ1 γ2 γ3 ε (ε2:R -> nat -> R)(α:loc) (ns:list nat) (q:Qp) (z:nat): *) - (* ↑N ⊆ E-> *) - (* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) - (* (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → *) - (* {{{ is_counter (L:=L) N c γ1 γ2 γ3 ∗ *) - (* counter_error_frag (L:=L) γ1 ε ∗ *) - (* counter_tapes_frag (L:=L) γ2 α [] ∗ *) - (* counter_content_frag (L:=L) γ3 q z *) - (* }}} *) - (* incr_counter_tape c #lbl:α @ E *) - (* {{{ (z':nat) (n:nat), *) - (* RET (#z', #n); ⌜(0<=n<4)%nat⌝ ∗ *) - (* ⌜(z<=z')%nat⌝ ∗ *) - (* counter_error_frag (L:=L) γ1 (ε2 ε n) ∗ *) - (* counter_tapes_frag (L:=L) γ2 α [] ∗ *) - (* counter_content_frag (L:=L) γ3 q (z+n)%nat }}}. *) - (* Proof. *) - (* iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & Herr & Htapes & Hcontent) HΦ". *) - (* pose (ε2' := (λ ε x, if (x<=?3)%nat then ε2 ε x else 1%R)). *) - (* wp_apply (incr_counter_tape_spec_none _ _ _ _ _ _ ε2' *) - (* (counter_error_frag γ1 ε ∗ counter_content_frag γ3 q z)%I *) - (* (λ n, ⌜0<=n<4⌝ ∗ counter_error_frag γ1 (ε2' ε n) ∗ counter_content_frag γ3 q z)%I *) - (* (λ z' n, ⌜0<=n<4⌝ ∗ ⌜z<=z'⌝ ∗ counter_error_frag γ1 (ε2' ε n) ∗ counter_content_frag γ3 q (z+n)%nat)%I *) - (* with "[$Herr $Htapes $Hcontent]"). *) - (* - done. *) - (* - rewrite /ε2'. intros. *) - (* case_match; [naive_solver|lra]. *) - (* - rewrite /ε2'. simpl. intros. naive_solver. *) - (* - repeat iSplit; first done. *) - (* + iModIntro. *) - (* iIntros (εa εb n) "[(Herr &Hcontent) Herr_auth]". *) - (* destruct (n<=?3) eqn:H; last first. *) - (* { iLeft. iPureIntro. rewrite /ε2'. by rewrite H. } *) - (* iRight. *) - (* iFrame. *) - (* iDestruct (counter_error_ineq with "[$][$]") as "%". *) - (* iDestruct (counter_error_auth_valid with "[$]") as "%". *) - (* iDestruct (counter_error_frag_valid with "[$]") as "%". *) - (* unshelve iMod (counter_error_update _ _ _ (mknonnegreal (ε2' ε n) _) with "[$][$]") as "[$$]". *) - (* { rewrite /ε2' H. *) - (* naive_solver. } *) - (* iPureIntro. split; first lia. *) - (* apply leb_complete in H. lia. *) - (* + iModIntro. *) - (* iIntros (??) "[(%&Herr & Hcontent) Hcontent_auth]". *) - (* iFrame. *) - (* iDestruct (counter_content_less_than with "[$][$]") as "%". *) - (* by iMod (counter_content_update with "[$][$]") as "[$ $]". *) - (* - iIntros (??) "[(%&%&?&?)?]". *) - (* iApply "HΦ". *) - (* iFrame. *) - (* rewrite /ε2'. case_match eqn: H2; first by iFrame. *) - (* apply leb_iff_conv in H2. lia. *) - (* Qed. *) + Lemma incr_counter_tape_spec_none' N E c γ1 γ2 ε (ε2:nat -> R)(α:loc) (ns:list nat) (q:Qp) (z:nat): + ↑N ⊆ E-> + (∀ n, 0<= ε2 n)%R-> + (((ε2 0%nat) + (ε2 1%nat)+ (ε2 2%nat)+ (ε2 3%nat))/4 <= ε)%R → + {{{ is_counter (L:=L) N c γ1 γ2 ∗ + ↯ ε ∗ + counter_tapes_frag (L:=L) γ1 α [] ∗ + counter_content_frag (L:=L) γ2 q z + }}} + incr_counter_tape c #lbl:α @ E + {{{ (z':nat) (n:nat), + RET (#z', #n); ⌜(0<=n<4)%nat⌝ ∗ + ⌜(z<=z')%nat⌝ ∗ + ↯(ε2 n) ∗ + counter_tapes_frag (L:=L) γ1 α [] ∗ + counter_content_frag (L:=L) γ2 q (z+n)%nat }}}. + Proof. + iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & Herr & Htapes & Hcontent) HΦ". + pose (ε2' := (λ x, if (x<=?3)%nat then ε2 x else 1%R)). + wp_apply (incr_counter_tape_spec_none _ _ _ _ _ + (λ z' n ε' ε2'', ⌜0<=n<4⌝ ∗ ⌜z<=z'⌝ ∗ ⌜ε=ε'⌝ ∗ ⌜ε2''=ε2'⌝ ∗ + ↯ (ε2' n) ∗ counter_tapes_frag (L:=L) γ1 α [] ∗ + counter_content_frag (L:=L) γ2 q (z+n)%nat + )%I + with "[-HΦ]"). + - done. + - iSplit; first done. + iApply fupd_mask_intro; first set_solver. + iIntros "Hclose". + iFrame. + iExists ε2'. + repeat iSplit. + + iPureIntro. rewrite /ε2'. intros; case_match; [naive_solver|lra]. + + done. + + iIntros (n) "[Herr Hfrag]". + iMod "Hclose" as "_". + iModIntro. + iIntros (z') "Hauth". + iApply fupd_mask_intro; first set_solver. + iIntros "Hclose". + rewrite /ε2'. + case_match eqn:H; last by (iDestruct (ec_contradict with "[$]") as "%"). + iFrame. + iIntros "Hfrag'". + iDestruct (counter_content_less_than with "[$][$]") as "%". + iMod (counter_content_update with "[$][$]") as "[??]". + iMod "Hclose" as "_". + iFrame. + apply leb_complete in H. + iPureIntro; repeat split; try done; lia. + - iIntros (z' n) "(%ε' & %ε2'' & % & % &-> & -> & Herr & Hfrag & Hfrag')". + iApply "HΦ". + iFrame. + iSplit; first done. + iSplit; first done. + iApply ec_eq; last done. + rewrite /ε2'. + case_match eqn :K; first done. + exfalso. + apply leb_complete_conv in K. + lia. + Qed. End lemmas. + From a75cb862a49d1819559a707708b8dd18c295a1e0 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Fri, 27 Sep 2024 12:51:40 +0200 Subject: [PATCH 42/69] Progress with impl1 --- .../coneris/examples/random_counter/impl1.v | 636 +++++++++--------- 1 file changed, 303 insertions(+), 333 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl1.v b/theories/coneris/examples/random_counter/impl1.v index c6b91daa..15bee9d8 100644 --- a/theories/coneris/examples/random_counter/impl1.v +++ b/theories/coneris/examples/random_counter/impl1.v @@ -1,347 +1,317 @@ -(* From iris.algebra Require Import frac_auth. *) -(* From iris.base_logic.lib Require Import invariants. *) -(* From clutch.coneris Require Import coneris hocap random_counter. *) +From iris.algebra Require Import frac_auth. +From iris.base_logic.lib Require Import invariants. +From clutch.coneris Require Import coneris hocap hocap_rand random_counter. -(* Set Default Proof Using "Type*". *) +Set Default Proof Using "Type*". -(* Section impl1. *) +Section impl1. + Context `{H:conerisGS Σ, r1:@rand_spec Σ H, L:randG Σ, !inG Σ (frac_authR natR)}. -(* Definition new_counter1 : val:= λ: "_", ref #0. *) -(* Definition incr_counter1 : val := λ: "l", let: "n" := rand #3 in (FAA "l" "n", "n"). *) -(* Definition allocate_tape1 : val := λ: "_", AllocTape #3. *) -(* Definition incr_counter_tape1 :val := λ: "l" "α", let: "n" := rand("α") #3 in (FAA "l" "n", "n"). *) -(* Definition read_counter1 : val := λ: "l", !"l". *) -(* Class counterG1 Σ := CounterG1 { counterG1_error::hocap_errorGS Σ; *) -(* counterG1_tapes:: hocap_tapesGS Σ; *) -(* counterG1_frac_authR:: inG Σ (frac_authR natR) }. *) - -(* Context `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapesGS Σ, !inG Σ (frac_authR natR)}. *) -(* Definition counter_inv_pred1 (c:val) γ1 γ2 γ3:= *) -(* (∃ (ε:R) m (l:loc) (z:nat), *) -(* ↯ ε ∗ ●↯ ε @ γ1 ∗ *) -(* ([∗ map] α ↦ t ∈ m, α ↪N ( t.1 ; t.2 )) ∗ ●m@γ2 ∗ *) -(* ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ3 (●F z) *) -(* )%I. *) + Definition new_counter1 : val:= λ: "_", ref #0. + Definition allocate_tape1 : val := λ: "_", rand_allocate_tape #3. + Definition incr_counter_tape1 :val := λ: "l" "α", let: "n" := rand_tape "α" #3 in (FAA "l" "n", "n"). + Definition read_counter1 : val := λ: "l", !"l". + Class counterG1 Σ := CounterG1 { counterG1_randG : randG Σ; + counterG1_frac_authR:: inG Σ (frac_authR natR) }. + + Definition counter_inv_pred1 c γ2 := (∃ (l:loc) (z:nat), ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ2 (●F z) )%I. + Definition is_counter1 N (c:val) γ1 γ2:= + (is_rand (L:=L) (N.@"rand") γ1 ∗ + inv (N.@"counter") (counter_inv_pred1 c γ2) + )%I. -(* Lemma new_counter_spec1 E ε N: *) -(* {{{ ↯ ε }}} *) -(* new_counter1 #() @ E *) -(* {{{ (c:val), RET c; *) -(* ∃ γ1 γ2 γ3, inv N (counter_inv_pred1 c γ1 γ2 γ3) ∗ *) -(* ◯↯ε @ γ1 ∗ own γ3 (◯F 0%nat) *) -(* }}}. *) -(* Proof. *) -(* rewrite /new_counter1. *) -(* iIntros (Φ) "Hε HΦ". *) -(* wp_pures. *) -(* wp_alloc l as "Hl". *) -(* iDestruct (ec_valid with "[$]") as "%". *) -(* unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "[%γ1 [H1 H2]]". *) -(* { lra. } *) -(* { simpl. lra. } *) -(* simpl. *) -(* iMod (hocap_tapes_alloc (∅:gmap _ _)) as "[%γ2 [H3 H4]]". *) -(* iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". *) -(* { by apply frac_auth_valid. } *) -(* replace (#0) with (#0%nat) by done. *) -(* iMod (inv_alloc N _ (counter_inv_pred1 (#l) γ1 γ2 γ3) with "[$Hε $Hl $H1 $H3 $H5]") as "#Hinv". *) -(* { iSplit; last done. by iApply big_sepM_empty. } *) -(* iApply "HΦ". *) -(* iExists _, _, _. by iFrame. *) -(* Qed. *) + Lemma new_counter_spec1 E N: + {{{ True }}} + new_counter1 #() @ E + {{{ (c:val), RET c; + ∃ γ1 γ2, is_counter1 N c γ1 γ2 ∗ own γ2 (◯F 0%nat) + }}}. + Proof. + rewrite /new_counter1. + iIntros (Φ) "_ HΦ". + wp_pures. + iMod (rand_inv_create_spec) as "(%γ1 & #Hinv)". + wp_alloc l as "Hl". + iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ2[H5 H6]]". + { by apply frac_auth_valid. } + replace (#0) with (#0%nat) by done. + iMod (inv_alloc _ _ (counter_inv_pred1 (#l) γ2) with "[$Hl $H5]") as "#Hinv'"; first done. + iApply "HΦ". + iFrame. + iModIntro. + iExists _. by iSplit. + Qed. -(* (** Not used in instantiating type class*) *) -(* Lemma incr_counter_spec1 E N c γ1 γ2 γ3 (ε2:R -> nat -> R) (P: iProp Σ) (T: nat -> iProp Σ) (Q: nat->nat->iProp Σ): *) -(* ↑N ⊆ E-> *) -(* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) -(* (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → *) -(* {{{ inv N (counter_inv_pred1 c γ1 γ2 γ3) ∗ *) -(* □(∀ (ε:R) (n : nat), P ∗ ●↯ ε @ γ1 ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨●↯ (ε2 ε n) @ γ1 ∗ T n) ) ∗ *) -(* □ (∀ (n z:nat), T n ∗ own γ3 (●F z) ={E∖↑N}=∗ *) -(* own γ3 (●F(z+n)%nat)∗ Q z n) ∗ *) -(* P *) -(* }}} *) -(* incr_counter1 c @ E *) -(* {{{ (n:nat) (z:nat), RET (#z, #n); Q z n }}}. *) -(* Proof. *) -(* iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP) HΦ". *) -(* rewrite /incr_counter1. *) -(* wp_pures. *) -(* wp_bind (rand _)%E. *) -(* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) -(* iDestruct (ec_valid with "[$]") as "[%K1 %K2]". *) -(* wp_apply (wp_couple_rand_adv_comp1' _ _ _ _ (λ x, ε2 ε (fin_to_nat x)) with "[$]"). *) -(* { intros. naive_solver. } *) -(* { rewrite SeriesC_finite_foldr; specialize (Hineq ε K1). simpl; lra. } *) -(* iIntros (n) "H1". *) -(* iMod ("Hvs1" with "[$]") as "[%|[H2 HT]]". *) -(* { iExFalso. iApply ec_contradict; last done. done. } *) -(* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) -(* iModIntro. wp_pures. *) -(* clear -Hsubset. *) -(* wp_bind (FAA _ _). *) -(* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) -(* wp_faa. *) -(* iMod ("Hvs2" with "[$]") as "[H6 HQ]". *) -(* replace (#(z+n)) with (#(z+n)%nat); last first. *) -(* { by rewrite Nat2Z.inj_add. } *) -(* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) -(* iModIntro. *) -(* wp_pures. *) -(* by iApply "HΦ". *) -(* Qed. *) + (* (** Not used in instantiating type class*) *) + (* Lemma incr_counter_spec1 E N c γ1 γ2 γ3 (ε2:R -> nat -> R) (P: iProp Σ) (T: nat -> iProp Σ) (Q: nat->nat->iProp Σ): *) + (* ↑N ⊆ E-> *) + (* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) + (* (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → *) + (* {{{ inv N (counter_inv_pred1 c γ1 γ2 γ3) ∗ *) + (* □(∀ (ε:R) (n : nat), P ∗ ●↯ ε @ γ1 ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨●↯ (ε2 ε n) @ γ1 ∗ T n) ) ∗ *) + (* □ (∀ (n z:nat), T n ∗ own γ3 (●F z) ={E∖↑N}=∗ *) + (* own γ3 (●F(z+n)%nat)∗ Q z n) ∗ *) + (* P *) + (* }}} *) + (* incr_counter1 c @ E *) + (* {{{ (n:nat) (z:nat), RET (#z, #n); Q z n }}}. *) + (* Proof. *) + (* iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP) HΦ". *) + (* rewrite /incr_counter1. *) + (* wp_pures. *) + (* wp_bind (rand _)%E. *) + (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* iDestruct (ec_valid with "[$]") as "[%K1 %K2]". *) + (* wp_apply (wp_couple_rand_adv_comp1' _ _ _ _ (λ x, ε2 ε (fin_to_nat x)) with "[$]"). *) + (* { intros. naive_solver. } *) + (* { rewrite SeriesC_finite_foldr; specialize (Hineq ε K1). simpl; lra. } *) + (* iIntros (n) "H1". *) + (* iMod ("Hvs1" with "[$]") as "[%|[H2 HT]]". *) + (* { iExFalso. iApply ec_contradict; last done. done. } *) + (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) + (* iModIntro. wp_pures. *) + (* clear -Hsubset. *) + (* wp_bind (FAA _ _). *) + (* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* wp_faa. *) + (* iMod ("Hvs2" with "[$]") as "[H6 HQ]". *) + (* replace (#(z+n)) with (#(z+n)%nat); last first. *) + (* { by rewrite Nat2Z.inj_add. } *) + (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) + (* iModIntro. *) + (* wp_pures. *) + (* by iApply "HΦ". *) + (* Qed. *) -(* Lemma allocate_tape_spec1 N E c γ1 γ2 γ3: *) -(* ↑N ⊆ E-> *) -(* {{{ inv N (counter_inv_pred1 c γ1 γ2 γ3) }}} *) -(* allocate_tape1 #() @ E *) -(* {{{ (v:val), RET v; *) -(* ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ α ◯↪N (3%nat; []) @ γ2 *) -(* }}}. *) -(* Proof. *) -(* iIntros (Hsubset Φ) "#Hinv HΦ". *) -(* rewrite /allocate_tape1. *) -(* wp_pures. *) -(* wp_alloctape α as "Hα". *) -(* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) -(* iDestruct (hocap_tapes_notin with "[$][$]") as "%". *) -(* iMod (hocap_tapes_new with "[$]") as "[H4 H7]"; first done. *) -(* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Hα]") as "_". *) -(* { iNext. iSplitL; last done. *) -(* rewrite big_sepM_insert; [simpl; iFrame|done]. *) -(* } *) -(* iApply "HΦ". *) -(* by iFrame. *) -(* Qed. *) + Lemma allocate_tape_spec1 N E c γ1 γ2: + ↑N ⊆ E-> + {{{ is_counter1 N c γ1 γ2 }}} + allocate_tape1 #() @ E + {{{ (v:val), RET v; + ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ rand_tapes_frag (L:=L) γ1 α (3%nat, []) + }}}. + Proof. + Admitted. + (* iIntros (Hsubset Φ) "#Hinv HΦ". *) + (* rewrite /allocate_tape1. *) + (* wp_pures. *) + (* wp_alloctape α as "Hα". *) + (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* iDestruct (hocap_tapes_notin with "[$][$]") as "%". *) + (* iMod (hocap_tapes_new with "[$]") as "[H4 H7]"; first done. *) + (* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Hα]") as "_". *) + (* { iNext. iSplitL; last done. *) + (* rewrite big_sepM_insert; [simpl; iFrame|done]. *) + (* } *) + (* iApply "HΦ". *) + (* by iFrame. *) + (* Qed. *) -(* Lemma incr_counter_tape_spec_some1 N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: *) -(* ↑N⊆E -> *) -(* {{{ inv N (counter_inv_pred1 c γ1 γ2 γ3) ∗ *) -(* □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ *) -(* own γ3 (●F(z+n)%nat)∗ Q z) ∗ *) -(* P ∗ α ◯↪N (3%nat; n::ns) @ γ2 *) -(* }}} *) -(* incr_counter_tape1 c #lbl:α @ E *) -(* {{{ (z:nat), RET (#z, #n); Q z ∗ α ◯↪N (3%nat; ns) @ γ2}}}. *) -(* Proof. *) -(* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP & Hα) HΦ". *) -(* rewrite /incr_counter_tape1. *) -(* wp_pures. *) -(* wp_bind (rand(_) _)%E. *) -(* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) -(* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) -(* erewrite <-(insert_delete m) at 1; last done. *) -(* rewrite big_sepM_insert; last apply lookup_delete. *) -(* simpl. *) -(* iDestruct "H3" as "[Htape H3]". *) -(* wp_apply (wp_rand_tape with "[$]"). *) -(* iIntros "[Htape %]". *) -(* iMod (hocap_tapes_pop with "[$][$]") as "[H4 Hα]". *) -(* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". *) -(* { iSplitL; last done. *) -(* erewrite <-(insert_delete m) at 2; last done. *) -(* iNext. *) -(* rewrite insert_insert. *) -(* rewrite big_sepM_insert; last apply lookup_delete. iFrame. *) -(* } *) -(* iModIntro. *) -(* wp_pures. *) -(* clear -Hsubset. *) -(* wp_bind (FAA _ _). *) -(* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) -(* wp_faa. *) -(* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) -(* replace (#(z+n)) with (#(z+n)%nat); last first. *) -(* { by rewrite Nat2Z.inj_add. } *) -(* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) -(* iModIntro. wp_pures. *) -(* iApply "HΦ". *) -(* by iFrame. *) -(* Qed. *) + Lemma incr_counter_tape_spec_some1 N E c γ1 γ2 (Q:nat->nat -> list nat -> iProp Σ) (α:loc) : + ↑N⊆E -> + {{{ is_counter1 N c γ1 γ2 ∗ + (∀ (z:nat), own γ2 (●F z) ={E∖ ↑N, ∅}=∗ + ∃ n ns, rand_tapes_frag (L:=L) γ1 α (3%nat, n::ns) ∗ + (rand_tapes_frag (L:=L) γ1 α (3%nat, ns) ={∅, E∖↑N}=∗ + own γ2 (●F (z+n)%nat) ∗ Q z n ns) + ) + }}} + incr_counter_tape1 c #lbl:α @ E + {{{ (z n:nat), RET (#z, #n); ∃ ns, Q z n ns}}}. + Proof. + Admitted. + (* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP & Hα) HΦ". *) + (* rewrite /incr_counter_tape1. *) + (* wp_pures. *) + (* wp_bind (rand(_) _)%E. *) + (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) + (* erewrite <-(insert_delete m) at 1; last done. *) + (* rewrite big_sepM_insert; last apply lookup_delete. *) + (* simpl. *) + (* iDestruct "H3" as "[Htape H3]". *) + (* wp_apply (wp_rand_tape with "[$]"). *) + (* iIntros "[Htape %]". *) + (* iMod (hocap_tapes_pop with "[$][$]") as "[H4 Hα]". *) + (* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". *) + (* { iSplitL; last done. *) + (* erewrite <-(insert_delete m) at 2; last done. *) + (* iNext. *) + (* rewrite insert_insert. *) + (* rewrite big_sepM_insert; last apply lookup_delete. iFrame. *) + (* } *) + (* iModIntro. *) + (* wp_pures. *) + (* clear -Hsubset. *) + (* wp_bind (FAA _ _). *) + (* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* wp_faa. *) + (* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) + (* replace (#(z+n)) with (#(z+n)%nat); last first. *) + (* { by rewrite Nat2Z.inj_add. } *) + (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) + (* iModIntro. wp_pures. *) + (* iApply "HΦ". *) + (* by iFrame. *) + (* Qed. *) -(* Lemma counter_presample_spec1 NS E ns α *) -(* (ε2 : R -> nat -> R) *) -(* (P : iProp Σ) T γ1 γ2 γ3 c: *) -(* ↑NS ⊆ E -> *) -(* (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> *) -(* (∀ (ε:R), 0<= ε ->SeriesC (λ n, if (bool_decide (n≤3%nat)) then 1 / (S 3%nat) * ε2 ε n else 0%R)%R <= ε)%R-> *) -(* inv NS (counter_inv_pred1 c γ1 γ2 γ3) -∗ *) -(* (□∀ (ε:R) n, (P ∗ ●↯ ε@ γ1) ={E∖↑NS}=∗ *) -(* (⌜(1<=ε2 ε n)%R⌝ ∨(●↯ (ε2 ε n) @ γ1 ∗ T (n)))) *) -(* -∗ *) -(* P -∗ α ◯↪N (3%nat; ns) @ γ2 -∗ *) -(* wp_update E (∃ n, T (n) ∗ α◯↪N (3%nat; ns++[n]) @ γ2). *) -(* Proof. *) -(* iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) -(* iApply wp_update_state_step_coupl. *) -(* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) -(* iMod (inv_acc with "Hinv") as "[>(% & % & % & % & H1 & H2 & H3 & H4 & -> & H5 & H6) Hclose]"; [done|]. *) -(* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) -(* erewrite <-(insert_delete m) at 1; last done. *) -(* rewrite big_sepM_insert; last apply lookup_delete. *) -(* simpl. *) -(* iDestruct "H3" as "[Htape H3]". *) -(* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%)". *) -(* iDestruct (ec_supply_bound with "[$][$]") as "%". *) -(* iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. *) -(* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) -(* iApply state_step_coupl_state_adv_comp_con_prob_lang; first done. *) -(* unshelve iExists (λ x, mknonnegreal (ε2 ε1' x) _). *) -(* { apply Hpos. apply cond_nonneg. } *) -(* iSplit. *) -(* { iPureIntro. *) -(* unshelve epose proof (Hineq ε1' _) as H1; first apply cond_nonneg. *) -(* by rewrite SeriesC_nat_bounded_fin in H1. } *) -(* iIntros (sample). *) + Lemma counter_presample_spec1 NS E T γ1 γ2 c: + ↑NS ⊆ E -> + is_counter1 NS c γ1 γ2 -∗ + ( |={E∖↑NS,∅}=> + ∃ ε α ε2 num ns, + ↯ ε ∗ rand_tapes_frag (L:=L) γ1 α (3%nat, ns) ∗ + ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ + ⌜(SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R⌝ ∗ + (∀ ns', ↯ (ε2 ns') ∗ rand_tapes_frag (L:=L) γ1 α (3%nat, ns++ns')={∅,E∖↑NS}=∗ + T ε α ε2 num ns ns' + ))-∗ + wp_update E (∃ ε α ε2 num ns ns', T ε α ε2 num ns ns'). + Proof. + Admitted. + (* iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) + (* iApply wp_update_state_step_coupl. *) + (* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) + (* iMod (inv_acc with "Hinv") as "[>(% & % & % & % & H1 & H2 & H3 & H4 & -> & H5 & H6) Hclose]"; [done|]. *) + (* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) + (* erewrite <-(insert_delete m) at 1; last done. *) + (* rewrite big_sepM_insert; last apply lookup_delete. *) + (* simpl. *) + (* iDestruct "H3" as "[Htape H3]". *) + (* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%)". *) + (* iDestruct (ec_supply_bound with "[$][$]") as "%". *) + (* iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. *) + (* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) + (* iApply state_step_coupl_state_adv_comp_con_prob_lang; first done. *) + (* unshelve iExists (λ x, mknonnegreal (ε2 ε1' x) _). *) + (* { apply Hpos. apply cond_nonneg. } *) + (* iSplit. *) + (* { iPureIntro. *) + (* unshelve epose proof (Hineq ε1' _) as H1; first apply cond_nonneg. *) + (* by rewrite SeriesC_nat_bounded_fin in H1. } *) + (* iIntros (sample). *) -(* destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' sample))%R 1%R) as [Hdec|Hdec]; last first. *) -(* { apply Rnot_lt_ge, Rge_le in Hdec. *) -(* iApply state_step_coupl_ret_err_ge_1. *) -(* simpl. simpl in *. lra. *) -(* } *) -(* iApply state_step_coupl_ret. *) -(* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' sample) _) with "Hε_supply") as "[Hε_supply Hε]". *) -(* { apply Hpos. apply cond_nonneg. } *) -(* { simpl. done. } *) -(* iMod (tapeN_update_append _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". *) -(* iMod (hocap_tapes_presample _ _ _ _ _ (fin_to_nat sample) with "[$][$]") as "[H4 Hfrag]". *) -(* iMod "Hclose'" as "_". *) -(* iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". *) -(* { iExFalso. iApply (ec_contradict with "[$]"). exact. } *) -(* iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4 $H5 $H6]") as "_". *) -(* { iNext. iSplit; last done. *) -(* rewrite big_sepM_insert_delete; iFrame. *) -(* } *) -(* iApply fupd_mask_intro_subseteq; first set_solver. *) -(* iFrame. *) -(* Qed. *) + (* destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' sample))%R 1%R) as [Hdec|Hdec]; last first. *) + (* { apply Rnot_lt_ge, Rge_le in Hdec. *) + (* iApply state_step_coupl_ret_err_ge_1. *) + (* simpl. simpl in *. lra. *) + (* } *) + (* iApply state_step_coupl_ret. *) + (* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' sample) _) with "Hε_supply") as "[Hε_supply Hε]". *) + (* { apply Hpos. apply cond_nonneg. } *) + (* { simpl. done. } *) + (* iMod (tapeN_update_append _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". *) + (* iMod (hocap_tapes_presample _ _ _ _ _ (fin_to_nat sample) with "[$][$]") as "[H4 Hfrag]". *) + (* iMod "Hclose'" as "_". *) + (* iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". *) + (* { iExFalso. iApply (ec_contradict with "[$]"). exact. } *) + (* iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4 $H5 $H6]") as "_". *) + (* { iNext. iSplit; last done. *) + (* rewrite big_sepM_insert_delete; iFrame. *) + (* } *) + (* iApply fupd_mask_intro_subseteq; first set_solver. *) + (* iFrame. *) + (* Qed. *) -(* Lemma read_counter_spec1 N E c γ1 γ2 γ3 P Q: *) -(* ↑N ⊆ E -> *) -(* {{{ inv N (counter_inv_pred1 c γ1 γ2 γ3) ∗ *) -(* □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ *) -(* own γ3 (●F z)∗ Q z) *) -(* ∗ P *) -(* }}} *) -(* read_counter1 c @ E *) -(* {{{ (n':nat), RET #n'; Q n' *) -(* }}}. *) -(* Proof. *) -(* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP) HΦ". *) -(* rewrite /read_counter1. *) -(* wp_pure. *) -(* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) -(* wp_load. *) -(* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) -(* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]"); first done. *) -(* iApply ("HΦ" with "[$]"). *) -(* Qed. *) + Lemma read_counter_spec1 N E c γ1 γ2 Q: + ↑N ⊆ E -> + {{{ is_counter1 N c γ1 γ2 ∗ + (∀ (z:nat), own γ2 (●F z) ={E∖↑N}=∗ + own γ2 (●F z) ∗ Q z) + + }}} + read_counter1 c @ E + {{{ (n':nat), RET #n'; Q n' + }}}. + Proof. + Admitted. + (* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP) HΦ". *) + (* rewrite /read_counter1. *) + (* wp_pure. *) + (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* wp_load. *) + (* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) + (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]"); first done. *) + (* iApply ("HΦ" with "[$]"). *) + (* Qed. *) -(* End impl1. *) +End impl1. -(* Program Definition random_counter1 `{!conerisGS Σ}: random_counter := *) -(* {| new_counter := new_counter1; *) -(* allocate_tape:= allocate_tape1; *) -(* incr_counter_tape := incr_counter_tape1; *) -(* read_counter:=read_counter1; *) -(* counterG := counterG1; *) -(* error_name := gname; *) -(* tape_name := gname; *) -(* counter_name :=gname; *) -(* is_counter _ N c γ1 γ2 γ3 := inv N (counter_inv_pred1 c γ1 γ2 γ3); *) -(* counter_error_auth _ γ x := ●↯ x @ γ; *) -(* counter_error_frag _ γ x := ◯↯ x @ γ; *) -(* counter_tapes_auth _ γ m := (●((λ ns, (3, ns))<$>m)@γ)%I; *) -(* counter_tapes_frag _ γ α ns := (α◯↪N (3%nat;ns) @ γ)%I; *) -(* counter_content_auth _ γ z := own γ (●F z); *) -(* counter_content_frag _ γ f z := own γ (◯F{f} z); *) -(* new_counter_spec _ := new_counter_spec1; *) -(* allocate_tape_spec _ :=allocate_tape_spec1; *) -(* incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some1; *) -(* counter_presample_spec _ :=counter_presample_spec1; *) -(* read_counter_spec _ :=read_counter_spec1 *) -(* |}. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (??????) "(%&<-&H1)(%&<-&H2)". *) -(* iCombine "H1 H2" gives "%H". by rewrite excl_auth.excl_auth_auth_op_valid in H. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (??????) "(%&<-&H1)(%&<-&H2)". *) -(* iCombine "H1 H2" gives "%H". by rewrite excl_auth.excl_auth_frag_op_valid in H. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (?????) "H". *) -(* iApply (hocap_error_auth_pos with "[$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (?????) "H". *) -(* iApply (hocap_error_frag_pos with "[$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (??????) "H1 H2". *) -(* iApply (hocap_error_agree with "[$][$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. iIntros (???????) "??". *) -(* iApply (hocap_error_update with "[$][$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (??????) "H1 H2". *) -(* by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[%H _]". *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (???????) "H1 H2". *) -(* iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. *) -(* rewrite dfrac_op_own dfrac_valid_own. by intro. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros. *) -(* iDestruct (hocap_tapes_agree with "[$][$]") as "%H". *) -(* iPureIntro. *) -(* rewrite lookup_fmap_Some in H. destruct H as (?&?&?). *) -(* by simplify_eq. *) -(* Qed. *) -(* Next Obligation. *) -(* iIntros. *) -(* iMod (hocap_tapes_presample with "[$][$]") as "[??]". *) -(* iModIntro. iFrame. *) -(* by rewrite fmap_insert. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (??????) "H1 H2". *) -(* iCombine "H1 H2" gives "%H". by rewrite auth_auth_op_valid in H. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. iIntros (???? z z' ?) "H1 H2". *) -(* iCombine "H1 H2" gives "%H". *) -(* apply frac_auth_included_total in H. iPureIntro. *) -(* by apply nat_included. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (????????). *) -(* rewrite frac_auth_frag_op. by rewrite own_op. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. iIntros (??????) "H1 H2". *) -(* iCombine "H1 H2" gives "%H". *) -(* iPureIntro. *) -(* by apply frac_auth_agree_L in H. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. iIntros (????????) "H1 H2". *) -(* iMod (own_update_2 _ _ _ (_ ⋅ _) with "[$][$]") as "[$$]"; last done. *) -(* apply frac_auth_update. *) -(* apply nat_local_update. lia. *) -(* Qed. *) +Program Definition random_counter1 `{!conerisGS Σ, F:rand_spec}: random_counter := + {| new_counter := new_counter1; + allocate_tape:= allocate_tape1; + incr_counter_tape := incr_counter_tape1; + read_counter:=read_counter1; + counterG := counterG1; + tape_name := rand_tape_name; + counter_name :=gname; + is_counter _ N c γ1 γ2 := is_counter1 (L:=counterG1_randG) N c γ1 γ2; + counter_tapes_auth _ γ m := rand_tapes_auth (L:=counterG1_randG) γ ((λ ns, (3, ns))<$>m); + counter_tapes_frag _ γ α ns := rand_tapes_frag (L:= counterG1_randG)γ α (3%nat, ns); + counter_content_auth _ γ z := own γ (●F z); + counter_content_frag _ γ f z := own γ (◯F{f} z); + new_counter_spec _ := new_counter_spec1; + allocate_tape_spec _ :=allocate_tape_spec1; + incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some1; + counter_presample_spec _ :=counter_presample_spec1; + read_counter_spec _ :=read_counter_spec1 + |}. +Next Obligation. + simpl. + iIntros (?????????) "H1 H2". + iDestruct (rand_tapes_auth_exclusive with "[$][$]") as "[]". +Qed. +Next Obligation. + simpl. + iIntros. + iDestruct (rand_tapes_frag_exclusive with "[$][$]") as "[]". +Qed. +Next Obligation. + simpl. + iIntros. + iDestruct (rand_tapes_agree with "[$][$]") as "%H". + iPureIntro. + apply lookup_fmap_Some in H as (?&?&?). + by simplify_eq. +Qed. +Next Obligation. + iIntros. + iDestruct (rand_tapes_valid with "[$]") as "$". +Qed. +Next Obligation. + simpl. + iIntros. + iMod (rand_tapes_update with "[$][$]") as "[??]"; last iFrame; first done. + by rewrite fmap_insert. +Qed. +Next Obligation. + simpl. + iIntros (?????????) "H1 H2". + iCombine "H1 H2" gives "%H". by rewrite auth_auth_op_valid in H. +Qed. +Next Obligation. + simpl. iIntros (??????? z z' ?) "H1 H2". + iCombine "H1 H2" gives "%H". + apply frac_auth_included_total in H. iPureIntro. + by apply nat_included. +Qed. +Next Obligation. + simpl. + iIntros. + rewrite frac_auth_frag_op. rewrite own_op. + iSplit; iIntros; iFrame. +Qed. +Next Obligation. + simpl. iIntros (?????????) "H1 H2". + iCombine "H1 H2" gives "%H". + iPureIntro. + by apply frac_auth_agree_L in H. +Qed. +Next Obligation. + simpl. iIntros (???????????) "H1 H2". + iMod (own_update_2 _ _ _ (_ ⋅ _) with "[$][$]") as "[$$]"; last done. + apply frac_auth_update. + apply nat_local_update. lia. +Qed. From 22b341b000813adc84c42e0e879290a906e8fc4c Mon Sep 17 00:00:00 2001 From: Hei Li Date: Fri, 27 Sep 2024 13:39:36 +0200 Subject: [PATCH 43/69] fix spec of incr_counter_tape_spec_some --- .../coneris/examples/random_counter/impl1.v | 46 ++++++-------- .../examples/random_counter/random_counter.v | 61 +++++++++++-------- 2 files changed, 55 insertions(+), 52 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl1.v b/theories/coneris/examples/random_counter/impl1.v index 15bee9d8..2414dfa3 100644 --- a/theories/coneris/examples/random_counter/impl1.v +++ b/theories/coneris/examples/random_counter/impl1.v @@ -8,8 +8,8 @@ Section impl1. Context `{H:conerisGS Σ, r1:@rand_spec Σ H, L:randG Σ, !inG Σ (frac_authR natR)}. Definition new_counter1 : val:= λ: "_", ref #0. - Definition allocate_tape1 : val := λ: "_", rand_allocate_tape #3. - Definition incr_counter_tape1 :val := λ: "l" "α", let: "n" := rand_tape "α" #3 in (FAA "l" "n", "n"). + Definition allocate_tape1 : val := λ: "_", rand_allocate_tape #3%nat. + Definition incr_counter_tape1 :val := λ: "l" "α", let: "n" := rand_tape "α" #3%nat in (FAA "l" "n", "n"). Definition read_counter1 : val := λ: "l", !"l". Class counterG1 Σ := CounterG1 { counterG1_randG : randG Σ; counterG1_frac_authR:: inG Σ (frac_authR natR) }. @@ -91,38 +91,31 @@ Section impl1. ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ rand_tapes_frag (L:=L) γ1 α (3%nat, []) }}}. Proof. - Admitted. - (* iIntros (Hsubset Φ) "#Hinv HΦ". *) - (* rewrite /allocate_tape1. *) - (* wp_pures. *) - (* wp_alloctape α as "Hα". *) - (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) - (* iDestruct (hocap_tapes_notin with "[$][$]") as "%". *) - (* iMod (hocap_tapes_new with "[$]") as "[H4 H7]"; first done. *) - (* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Hα]") as "_". *) - (* { iNext. iSplitL; last done. *) - (* rewrite big_sepM_insert; [simpl; iFrame|done]. *) - (* } *) - (* iApply "HΦ". *) - (* by iFrame. *) - (* Qed. *) + iIntros (Hsubset Φ) "[#Hinv #Hinv'] HΦ". + rewrite /allocate_tape1. + wp_pures. + wp_apply rand_allocate_tape_spec; [by eapply nclose_subseteq'|done..]. + Qed. Lemma incr_counter_tape_spec_some1 N E c γ1 γ2 (Q:nat->nat -> list nat -> iProp Σ) (α:loc) : ↑N⊆E -> {{{ is_counter1 N c γ1 γ2 ∗ - (∀ (z:nat), own γ2 (●F z) ={E∖ ↑N, ∅}=∗ - ∃ n ns, rand_tapes_frag (L:=L) γ1 α (3%nat, n::ns) ∗ - (rand_tapes_frag (L:=L) γ1 α (3%nat, ns) ={∅, E∖↑N}=∗ - own γ2 (●F (z+n)%nat) ∗ Q z n ns) + ( |={E∖ ↑N, ∅}=> + ∃ n ns, rand_tapes_frag (L:=L) γ1 α (3, n::ns) ∗ + (rand_tapes_frag (L:=L) γ1 α (3, ns) ={∅, E∖↑N}=∗ + ∀ (z:nat), own γ2 (●F z) + ={E∖↑N}=∗ + own γ2 (●F (z+n)) ∗ Q z n ns) ) }}} incr_counter_tape1 c #lbl:α @ E {{{ (z n:nat), RET (#z, #n); ∃ ns, Q z n ns}}}. Proof. - Admitted. - (* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP & Hα) HΦ". *) - (* rewrite /incr_counter_tape1. *) - (* wp_pures. *) + iIntros (Hsubset Φ) "([#Hinv #Hinv'] & Hvs) HΦ". + rewrite /incr_counter_tape1. + wp_pures. + wp_apply (rand_tape_spec_some with "[Hvs]"); [by eapply nclose_subseteq'|iSplit; first done|..]. + (* - iSplit *) (* wp_bind (rand(_) _)%E. *) (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) (* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) @@ -153,7 +146,8 @@ Section impl1. (* iModIntro. wp_pures. *) (* iApply "HΦ". *) (* by iFrame. *) - (* Qed. *) + (* Qed. *) + Admitted. Lemma counter_presample_spec1 NS E T γ1 γ2 c: ↑NS ⊆ E -> diff --git a/theories/coneris/examples/random_counter/random_counter.v b/theories/coneris/examples/random_counter/random_counter.v index 54051f64..d7885d8f 100644 --- a/theories/coneris/examples/random_counter/random_counter.v +++ b/theories/coneris/examples/random_counter/random_counter.v @@ -77,10 +77,12 @@ Class random_counter `{!conerisGS Σ} := RandCounter incr_counter_tape_spec_some {L: counterG Σ} N E c γ1 γ2 (Q:nat->nat->list nat->iProp Σ) (α:loc) : ↑N⊆E -> {{{ is_counter (L:=L) N c γ1 γ2 ∗ - (∀ (z:nat), counter_content_auth (L:=L) γ2 z ={E∖ ↑N, ∅}=∗ + ( |={E∖ ↑N, ∅}=> ∃ n ns, counter_tapes_frag (L:=L) γ1 α (n::ns) ∗ (counter_tapes_frag (L:=L) γ1 α ns ={∅, E∖↑N}=∗ - counter_content_auth (L:=L) γ2 (z+n) ∗ Q z n ns) + ∀ (z:nat), counter_content_auth (L:=L) γ2 z + ={E∖↑N}=∗ + counter_content_auth (L:=L) γ2 (z+n) ∗ Q z n ns) ) }}} incr_counter_tape c #lbl:α @ E @@ -121,14 +123,15 @@ Section lemmas. ↯ ε ∗ counter_tapes_frag (L:=L) γ1 α [] ∗ ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ ⌜(((ε2 0%nat) + (ε2 1%nat)+ (ε2 2%nat)+ (ε2 3%nat))/4 <= ε)%R⌝ ∗ - (∀ n, ↯ (ε2 n) ∗ counter_tapes_frag (L:=L) γ1 α [n] ={∅,E∖↑N}=∗ - ∀ (z:nat), counter_content_auth (L:=L) γ2 z ={E∖↑N, ∅}=∗ - counter_tapes_frag (L:=L) γ1 α [n]∗ - (counter_tapes_frag (L:=L) γ1 α []={∅, E∖↑N}=∗ - counter_content_auth (L:=L) γ2 (z+n) ∗ Q z n ε ε2 + (∀ n, ↯ (ε2 n) ∗ counter_tapes_frag (L:=L) γ1 α [n] ={∅, E∖↑N}=∗ + |={E∖↑N, ∅}=> + counter_tapes_frag (L:=L) γ1 α [n] ∗ + (counter_tapes_frag (L:=L) γ1 α [] ={∅,E∖↑N}=∗ + ∀ (z:nat), counter_content_auth (L:=L) γ2 z ={E∖↑N}=∗ + counter_content_auth (L:=L) γ2 (z+n) ∗ Q z n ε ε2 + ) ) - ) - ) + ) }}} incr_counter_tape c #lbl:α @ E {{{ (z n:nat), RET (#z, #n); ∃ ε ε2, Q z n ε ε2 }}}. @@ -138,13 +141,15 @@ Section lemmas. (λ ε α' ε2 num ns ns', ∃ n ε2', ⌜num=1%nat⌝ ∗ ⌜α'=α⌝ ∗ ⌜ns=[]⌝ ∗ ⌜ns'=[n]⌝ ∗ - ⌜(ε2' = λ x, ε2 [x])%R⌝ ∗ - (∀ z : nat, - counter_content_auth γ2 z ={E ∖ ↑N,∅}=∗ - counter_tapes_frag γ1 α [n] ∗ - (counter_tapes_frag γ1 α [] ={∅,E ∖ ↑N}=∗ - counter_content_auth γ2 (z + n) ∗ Q z n ε ε2')) - )%I with "[//][-HΦ]") as "(%ε & % & %ε2 & %num & %ns & %ns' & %n & %ε2' & -> & -> & -> & -> & -> & Hrest)"; try done. + ⌜(ε2' = λ x, ε2 [x])%R⌝ ∗ + ( |={E∖ ↑N, ∅}=> + counter_tapes_frag (L:=L) γ1 α [n] ∗ + (counter_tapes_frag (L:=L) γ1 α [] ={∅, E∖↑N}=∗ + ∀ (z:nat), counter_content_auth (L:=L) γ2 z + ={E∖↑N}=∗ + counter_content_auth (L:=L) γ2 (z+n) ∗ Q z n ε ε2') + ) + )%I with "[//][-HΦ]") as "(%ε & % & %ε2 & %num & %ns & %ns' & %n & %ε2' & -> & -> & -> & -> & -> & Hrest)"; first done. - iMod "Hvs" as "(%ε & %ε2 & Herr & Hfrag & %Hpos & %Hineq & Hrest)". iFrame. iModIntro. iExists (λ l, match l with |[x] => ε2 x | _ => 1 end)%R, 1%nat. @@ -157,17 +162,19 @@ Section lemmas. * apply NoDup_fmap_2; last apply NoDup_enum_uniform_fin_list. apply list_fmap_eq_inj, fin_to_nat_inj. + iIntros (ns') "[Herr Hfrag]". destruct (ns') as [|n[|??]]; [by iDestruct (ec_contradict with "[$]") as "%"| |by iDestruct (ec_contradict with "[$]") as "%"]. - iMod ("Hrest" $! n with "[$]"). - iFrame. - by iPureIntro. + iMod ("Hrest" $! n with "[$]") as "?". + iModIntro. + by iFrame. - wp_apply (incr_counter_tape_spec_some _ _ _ _ _ (λ z n ns, ∃ ε ε2, Q z n ε ε2)%I with "[Hrest]"); first exact. + iSplit; first done. - iIntros (z) "Hauth". - iMod ("Hrest" with "[$]") as "[Hfrag Hrest]". + iMod "Hrest" as "[? Hrest]". iFrame. iModIntro. - iIntros "Hfrag". - iMod ("Hrest" with "[$]") as "[Hauth HQ]". + iIntros. + iMod ("Hrest" with "[$]") as "Hrest". + iModIntro. + iIntros (z) "Hauth". + iMod ("Hrest" with "[$]") as "[??]". by iFrame. + iIntros (??) "(%&%&%&HQ)". iApply "HΦ". @@ -211,18 +218,20 @@ Section lemmas. + iIntros (n) "[Herr Hfrag]". iMod "Hclose" as "_". iModIntro. - iIntros (z') "Hauth". iApply fupd_mask_intro; first set_solver. iIntros "Hclose". rewrite /ε2'. case_match eqn:H; last by (iDestruct (ec_contradict with "[$]") as "%"). iFrame. iIntros "Hfrag'". - iDestruct (counter_content_less_than with "[$][$]") as "%". - iMod (counter_content_update with "[$][$]") as "[??]". iMod "Hclose" as "_". iFrame. apply leb_complete in H. + iModIntro. + iIntros (z') "Hauth". + iDestruct (counter_content_less_than with "[$][$]") as "%". + iMod (counter_content_update with "[$][$]") as "[??]". + iFrame. iPureIntro; repeat split; try done; lia. - iIntros (z' n) "(%ε' & %ε2'' & % & % &-> & -> & Herr & Hfrag & Hfrag')". iApply "HΦ". From a3a21511d739560ba939015d9cad1d5e62b180b8 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Fri, 27 Sep 2024 13:54:33 +0200 Subject: [PATCH 44/69] incr_counter_tape_spec_some1 --- .../coneris/examples/random_counter/impl1.v | 62 +++++++++---------- 1 file changed, 28 insertions(+), 34 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl1.v b/theories/coneris/examples/random_counter/impl1.v index 2414dfa3..8324b4da 100644 --- a/theories/coneris/examples/random_counter/impl1.v +++ b/theories/coneris/examples/random_counter/impl1.v @@ -114,40 +114,34 @@ Section impl1. iIntros (Hsubset Φ) "([#Hinv #Hinv'] & Hvs) HΦ". rewrite /incr_counter_tape1. wp_pures. - wp_apply (rand_tape_spec_some with "[Hvs]"); [by eapply nclose_subseteq'|iSplit; first done|..]. - (* - iSplit *) - (* wp_bind (rand(_) _)%E. *) - (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) - (* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) - (* erewrite <-(insert_delete m) at 1; last done. *) - (* rewrite big_sepM_insert; last apply lookup_delete. *) - (* simpl. *) - (* iDestruct "H3" as "[Htape H3]". *) - (* wp_apply (wp_rand_tape with "[$]"). *) - (* iIntros "[Htape %]". *) - (* iMod (hocap_tapes_pop with "[$][$]") as "[H4 Hα]". *) - (* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". *) - (* { iSplitL; last done. *) - (* erewrite <-(insert_delete m) at 2; last done. *) - (* iNext. *) - (* rewrite insert_insert. *) - (* rewrite big_sepM_insert; last apply lookup_delete. iFrame. *) - (* } *) - (* iModIntro. *) - (* wp_pures. *) - (* clear -Hsubset. *) - (* wp_bind (FAA _ _). *) - (* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) - (* wp_faa. *) - (* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) - (* replace (#(z+n)) with (#(z+n)%nat); last first. *) - (* { by rewrite Nat2Z.inj_add. } *) - (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) - (* iModIntro. wp_pures. *) - (* iApply "HΦ". *) - (* by iFrame. *) - (* Qed. *) - Admitted. + wp_apply (rand_tape_spec_some _ _ _ + (λ n ns, ∀ (z:nat), own γ2 (●F z)={E∖↑N}=∗ + own γ2 (●F (z+n)) ∗ Q z n ns)%I with "[Hvs]"); [by eapply nclose_subseteq'|iSplit; first done|..]. + - iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose". + + apply difference_mono_l. + by apply nclose_subseteq'. + + iMod "Hvs" as "(%n & %ns & Hfrag & Hrest)". + iModIntro. + iFrame. + iIntros "Hfrag". + iMod ("Hrest" with "[$]") as "Hrest". + by iMod "Hclose" as "_". + - iIntros (n) "(%ns & Hvs)". + wp_pures. + wp_bind (FAA _ _). + iInv "Hinv'" as ">(%&%&-> & ?&?)" "Hclose". + wp_faa. + iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". + + apply difference_mono_l. + by apply nclose_subseteq'. + + iMod ("Hvs" with "[$]") as "[? HQ]". + iMod "Hclose'" as "_". + rewrite -Nat2Z.inj_add. + iMod ("Hclose" with "[-HΦ HQ]") as "_"; first by iFrame. + iModIntro. + wp_pures. + iApply "HΦ". by iFrame. + Qed. Lemma counter_presample_spec1 NS E T γ1 γ2 c: ↑NS ⊆ E -> From 140171c9f99588c851be775c5f7e1ec5eee12167 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Fri, 27 Sep 2024 14:12:09 +0200 Subject: [PATCH 45/69] counter_presample_spec1 --- .../coneris/examples/random_counter/impl1.v | 79 ++++++++----------- 1 file changed, 35 insertions(+), 44 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl1.v b/theories/coneris/examples/random_counter/impl1.v index 8324b4da..bdddfc04 100644 --- a/theories/coneris/examples/random_counter/impl1.v +++ b/theories/coneris/examples/random_counter/impl1.v @@ -156,51 +156,42 @@ Section impl1. ))-∗ wp_update E (∃ ε α ε2 num ns ns', T ε α ε2 num ns ns'). Proof. - Admitted. - (* iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) - (* iApply wp_update_state_step_coupl. *) - (* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) - (* iMod (inv_acc with "Hinv") as "[>(% & % & % & % & H1 & H2 & H3 & H4 & -> & H5 & H6) Hclose]"; [done|]. *) - (* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) - (* erewrite <-(insert_delete m) at 1; last done. *) - (* rewrite big_sepM_insert; last apply lookup_delete. *) - (* simpl. *) - (* iDestruct "H3" as "[Htape H3]". *) - (* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%)". *) - (* iDestruct (ec_supply_bound with "[$][$]") as "%". *) - (* iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. *) - (* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) - (* iApply state_step_coupl_state_adv_comp_con_prob_lang; first done. *) - (* unshelve iExists (λ x, mknonnegreal (ε2 ε1' x) _). *) - (* { apply Hpos. apply cond_nonneg. } *) - (* iSplit. *) - (* { iPureIntro. *) - (* unshelve epose proof (Hineq ε1' _) as H1; first apply cond_nonneg. *) - (* by rewrite SeriesC_nat_bounded_fin in H1. } *) - (* iIntros (sample). *) + iIntros (Hsubset) "[#Hinv #Hinv'] Hvs". + iMod (rand_presample_spec _ _ (λ ε num tb ε2 α ns ns', + ∃ ε2', ⌜∀ xs ys, fin_to_nat <$> xs = ys -> ε2 xs = ε2' ys⌝ ∗ + ⌜tb = 3⌝ ∗ T ε α ε2' num ns (fin_to_nat <$> ns'))%I with "[//][Hvs]") as "H"; first by apply nclose_subseteq'. + + iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose". + { apply difference_mono_l. + by apply nclose_subseteq'. } + iMod "Hvs" as "(%ε & %α & %ε2 & %num & %ns & Herr & Hfrag & %Hpos & %Hineq & Hrest)". + iFrame. + iModIntro. + iExists num, (λ ls, ε2 (fin_to_nat <$> ls)). + repeat iSplit. + * done. + * iPureIntro. + etrans; last exact. + apply Req_le. + replace (INR 4) with 4%R; last (simpl; lra). + f_equal. + rewrite !SeriesC_list. + -- by rewrite foldr_fmap. + -- apply NoDup_fmap. + ++ apply list_fmap_eq_inj. + apply fin_to_nat_inj. + ++ apply NoDup_enum_uniform_fin_list. + -- apply NoDup_enum_uniform_fin_list. + * iIntros (ns') "[Herr Hfrag]". + iMod ("Hrest" with "[$]") as "HT". + iMod "Hclose" as "_". + iModIntro. + iFrame. + repeat iSplit; iPureIntro; last done. + by intros ??<-. + + iDestruct "H" as "(%ε & %num & %tb & %ε2 & %α & %ns & %ns' & %ε2' & % & -> & HT)". + iModIntro. iFrame. + Qed. - (* destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' sample))%R 1%R) as [Hdec|Hdec]; last first. *) - (* { apply Rnot_lt_ge, Rge_le in Hdec. *) - (* iApply state_step_coupl_ret_err_ge_1. *) - (* simpl. simpl in *. lra. *) - (* } *) - (* iApply state_step_coupl_ret. *) - (* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' sample) _) with "Hε_supply") as "[Hε_supply Hε]". *) - (* { apply Hpos. apply cond_nonneg. } *) - (* { simpl. done. } *) - (* iMod (tapeN_update_append _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". *) - (* iMod (hocap_tapes_presample _ _ _ _ _ (fin_to_nat sample) with "[$][$]") as "[H4 Hfrag]". *) - (* iMod "Hclose'" as "_". *) - (* iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". *) - (* { iExFalso. iApply (ec_contradict with "[$]"). exact. } *) - (* iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4 $H5 $H6]") as "_". *) - (* { iNext. iSplit; last done. *) - (* rewrite big_sepM_insert_delete; iFrame. *) - (* } *) - (* iApply fupd_mask_intro_subseteq; first set_solver. *) - (* iFrame. *) - (* Qed. *) - Lemma read_counter_spec1 N E c γ1 γ2 Q: ↑N ⊆ E -> {{{ is_counter1 N c γ1 γ2 ∗ From 7d7d463a601b575ef43ecd56ec0bc90ccc00e2fa Mon Sep 17 00:00:00 2001 From: Hei Li Date: Fri, 27 Sep 2024 14:14:34 +0200 Subject: [PATCH 46/69] complete impl1 --- .../coneris/examples/random_counter/impl1.v | 24 ++++++++++--------- 1 file changed, 13 insertions(+), 11 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl1.v b/theories/coneris/examples/random_counter/impl1.v index bdddfc04..03ab881e 100644 --- a/theories/coneris/examples/random_counter/impl1.v +++ b/theories/coneris/examples/random_counter/impl1.v @@ -203,17 +203,19 @@ Section impl1. {{{ (n':nat), RET #n'; Q n' }}}. Proof. - Admitted. - (* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP) HΦ". *) - (* rewrite /read_counter1. *) - (* wp_pure. *) - (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) - (* wp_load. *) - (* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) - (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]"); first done. *) - (* iApply ("HΦ" with "[$]"). *) - (* Qed. *) - + iIntros (Hsubset Φ) "([#Hinv #Hinv'] & Hvs) HΦ". + rewrite /read_counter1. + wp_pure. + iInv "Hinv'" as ">( %l & %z & -> & H5 & H6)" "Hclose". + wp_load. + iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". + { apply difference_mono_l. + by apply nclose_subseteq'. } + iMod ("Hvs" with "[$]") as "[? HQ]". + iMod "Hclose'" as "_". + iMod ("Hclose" with "[-HQ HΦ]"); first by iFrame. + iApply ("HΦ" with "[$]"). + Qed. End impl1. From 5dc07b5aba6357f2996fa0193a9f7f6dfbf066dc Mon Sep 17 00:00:00 2001 From: Hei Li Date: Fri, 27 Sep 2024 14:36:18 +0200 Subject: [PATCH 47/69] NIT --- .../coneris/examples/random_counter/impl2.v | 707 +++++++++--------- 1 file changed, 354 insertions(+), 353 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index 5554b146..456430d2 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -1,383 +1,384 @@ -(* From iris.algebra Require Import frac_auth. *) -(* From iris.base_logic.lib Require Import invariants. *) -(* From clutch.coneris Require Import coneris hocap random_counter flip_spec. *) -(* From clutch Require Import uniform_list. *) +From iris.algebra Require Import frac_auth. +From iris.base_logic.lib Require Import invariants. +From clutch.coneris Require Import coneris hocap random_counter hocap_flip. -(* Set Default Proof Using "Type*". *) +Set Default Proof Using "Type*". +Local Definition expander (l:list nat):= + l ≫= (λ x, [2<=?x; (Nat.odd x)]). +Local Lemma expander_eta x: x<4->(Z.of_nat x= Z.of_nat 2*Z.b2z (2<=? x)%nat + Z.b2z (Nat.odd x))%Z. +Proof. + destruct x as [|[|[|[|]]]]; last lia; intros; simpl; lia. +Qed. (* Local Definition expander (l:list nat):= *) -(* l ≫= (λ x, [2<=?x; (Nat.odd x)]). *) -(* Local Lemma expander_eta x: x<4->(Z.of_nat x= Z.of_nat 2*Z.b2z (2<=? x)%nat + Z.b2z (Nat.odd x))%Z. *) -(* Proof. *) -(* destruct x as [|[|[|[|]]]]; last lia; intros; simpl; lia. *) -(* Qed. *) -(* (* Local Definition expander (l:list nat):= *) *) -(* (* l ≫= (λ x, [Nat.div2 x; Nat.b2n (Nat.odd x)]). *) *) +(* l ≫= (λ x, [Nat.div2 x; Nat.b2n (Nat.odd x)]). *) -(* (* Class hocap_tapesGS' (Σ : gFunctors) := Hocap_tapesGS' { *) *) -(* (* hocap_tapesGS_inG' :: ghost_mapG Σ loc (bool* list nat) *) *) -(* (* }. *) *) -(* (* Definition hocap_tapesΣ' := ghost_mapΣ loc (bool *list nat). *) *) +(* Class hocap_tapesGS' (Σ : gFunctors) := Hocap_tapesGS' { *) +(* hocap_tapesGS_inG' :: ghost_mapG Σ loc (bool* list nat) *) +(* }. *) +(* Definition hocap_tapesΣ' := ghost_mapΣ loc (bool *list nat). *) -(* (* Notation "α ◯↪N ( b , ns ) @ γ":= (α ↪[ γ ] (b, ns))%I *) *) -(* (* (at level 20) : bi_scope. *) *) +(* Notation "α ◯↪N ( b , ns ) @ γ":= (α ↪[ γ ] (b, ns))%I *) +(* (at level 20) : bi_scope. *) -(* (* Notation "● m @ γ" := (ghost_map_auth γ 1 m) (at level 20) : bi_scope. *) *) +(* Notation "● m @ γ" := (ghost_map_auth γ 1 m) (at level 20) : bi_scope. *) -(* (* Section tapes_lemmas. *) *) -(* (* Context `{!hocap_tapesGS' Σ}. *) *) +(* Section tapes_lemmas. *) +(* Context `{!hocap_tapesGS' Σ}. *) -(* (* Lemma hocap_tapes_alloc' m: *) *) -(* (* ⊢ |==>∃ γ, (● m @ γ) ∗ [∗ map] k↦v ∈ m, (k ◯↪N (v.1, v.2) @ γ). *) *) -(* (* Proof. *) *) -(* (* iMod ghost_map_alloc as (γ) "[??]". *) *) -(* (* iFrame. iModIntro. *) *) -(* (* iApply big_sepM_mono; last done. *) *) -(* (* by iIntros (?[??]). *) *) -(* (* Qed. *) *) +(* Lemma hocap_tapes_alloc' m: *) +(* ⊢ |==>∃ γ, (● m @ γ) ∗ [∗ map] k↦v ∈ m, (k ◯↪N (v.1, v.2) @ γ). *) +(* Proof. *) +(* iMod ghost_map_alloc as (γ) "[??]". *) +(* iFrame. iModIntro. *) +(* iApply big_sepM_mono; last done. *) +(* by iIntros (?[??]). *) +(* Qed. *) -(* (* Lemma hocap_tapes_agree' m b γ k ns: *) *) -(* (* (● m @ γ) -∗ (k ◯↪N (b, ns) @ γ) -∗ ⌜ m!!k = Some (b, ns) ⌝. *) *) -(* (* Proof. *) *) -(* (* iIntros "H1 H2". *) *) -(* (* by iCombine "H1 H2" gives "%". *) *) -(* (* Qed. *) *) +(* Lemma hocap_tapes_agree' m b γ k ns: *) +(* (● m @ γ) -∗ (k ◯↪N (b, ns) @ γ) -∗ ⌜ m!!k = Some (b, ns) ⌝. *) +(* Proof. *) +(* iIntros "H1 H2". *) +(* by iCombine "H1 H2" gives "%". *) +(* Qed. *) -(* (* Lemma hocap_tapes_new' γ m k ns b: *) *) -(* (* m!!k=None -> ⊢ (● m @ γ) ==∗ (● (<[k:=(b, ns)]>m) @ γ) ∗ (k ◯↪N (b, ns) @ γ). *) *) -(* (* Proof. *) *) -(* (* iIntros (Hlookup) "H". *) *) -(* (* by iApply ghost_map_insert. *) *) -(* (* Qed. *) *) +(* Lemma hocap_tapes_new' γ m k ns b: *) +(* m!!k=None -> ⊢ (● m @ γ) ==∗ (● (<[k:=(b, ns)]>m) @ γ) ∗ (k ◯↪N (b, ns) @ γ). *) +(* Proof. *) +(* iIntros (Hlookup) "H". *) +(* by iApply ghost_map_insert. *) +(* Qed. *) -(* (* Lemma hocap_tapes_presample' b γ m k ns n: *) *) -(* (* (● m @ γ) -∗ (k ◯↪N (b, ns) @ γ) ==∗ (● (<[k:=(b, ns++[n])]>m) @ γ) ∗ (k ◯↪N (b, ns++[n]) @ γ). *) *) -(* (* Proof. *) *) -(* (* iIntros "H1 H2". *) *) -(* (* iApply (ghost_map_update with "[$][$]"). *) *) -(* (* Qed. *) *) +(* Lemma hocap_tapes_presample' b γ m k ns n: *) +(* (● m @ γ) -∗ (k ◯↪N (b, ns) @ γ) ==∗ (● (<[k:=(b, ns++[n])]>m) @ γ) ∗ (k ◯↪N (b, ns++[n]) @ γ). *) +(* Proof. *) +(* iIntros "H1 H2". *) +(* iApply (ghost_map_update with "[$][$]"). *) +(* Qed. *) -(* (* Lemma hocap_tapes_pop1' γ m k ns: *) *) -(* (* (● m @ γ) -∗ (k ◯↪N (true, ns) @ γ) ==∗ (● (<[k:=(false, ns)]>m) @ γ) ∗ (k ◯↪N (false, ns) @ γ). *) *) -(* (* Proof. *) *) -(* (* iIntros "H1 H2". *) *) -(* (* iApply (ghost_map_update with "[$][$]"). *) *) -(* (* Qed. *) *) +(* Lemma hocap_tapes_pop1' γ m k ns: *) +(* (● m @ γ) -∗ (k ◯↪N (true, ns) @ γ) ==∗ (● (<[k:=(false, ns)]>m) @ γ) ∗ (k ◯↪N (false, ns) @ γ). *) +(* Proof. *) +(* iIntros "H1 H2". *) +(* iApply (ghost_map_update with "[$][$]"). *) +(* Qed. *) -(* (* Lemma hocap_tapes_pop2' γ m k ns n: *) *) -(* (* (● m @ γ) -∗ (k ◯↪N (false, n::ns) @ γ) ==∗ (● (<[k:=(true, ns)]>m) @ γ) ∗ (k ◯↪N (true, ns) @ γ). *) *) -(* (* Proof. *) *) -(* (* iIntros "H1 H2". *) *) -(* (* iApply (ghost_map_update with "[$][$]"). *) *) -(* (* Qed. *) *) +(* Lemma hocap_tapes_pop2' γ m k ns n: *) +(* (● m @ γ) -∗ (k ◯↪N (false, n::ns) @ γ) ==∗ (● (<[k:=(true, ns)]>m) @ γ) ∗ (k ◯↪N (true, ns) @ γ). *) +(* Proof. *) +(* iIntros "H1 H2". *) +(* iApply (ghost_map_update with "[$][$]"). *) +(* Qed. *) -(* (* Lemma hocap_tapes_notin' `{F: flip_spec Σ} {L:flipG Σ} α ns m γ (f:(bool*list nat)-> (list bool)) : *) *) -(* (* flip_tapes_frag (L:=L) γ α ns -∗ ([∗ map] α0↦t ∈ m, flip_tapes_frag (L:=L) γ α0 (f t)) -∗ ⌜m!!α=None ⌝. *) *) -(* (* Proof. *) *) -(* (* destruct (m!!α) eqn:Heqn; last by iIntros. *) *) -(* (* iIntros "Hα Hmap". *) *) -(* (* iDestruct (big_sepM_lookup with "[$]") as "?"; first done. *) *) -(* (* iExFalso. *) *) -(* (* by iDestruct (flip_tapes_frag_exclusive with "[$][$]") as "%". *) *) -(* (* Qed. *) *) +(* Lemma hocap_tapes_notin' `{F: flip_spec Σ} {L:flipG Σ} α ns m γ (f:(bool*list nat)-> (list bool)) : *) +(* flip_tapes_frag (L:=L) γ α ns -∗ ([∗ map] α0↦t ∈ m, flip_tapes_frag (L:=L) γ α0 (f t)) -∗ ⌜m!!α=None ⌝. *) +(* Proof. *) +(* destruct (m!!α) eqn:Heqn; last by iIntros. *) +(* iIntros "Hα Hmap". *) +(* iDestruct (big_sepM_lookup with "[$]") as "?"; first done. *) +(* iExFalso. *) +(* by iDestruct (flip_tapes_frag_exclusive with "[$][$]") as "%". *) +(* Qed. *) -(* (* End tapes_lemmas. *) *) +(* End tapes_lemmas. *) -(* Section impl2. *) -(* Context `{F: flip_spec Σ}. *) -(* Definition new_counter2 : val:= λ: "_", ref #0. *) -(* (* Definition incr_counter2 : val := λ: "l", let: "n" := rand #1 in *) *) -(* (* let: "n'" := rand #1 in *) *) -(* (* let: "x" := #2 * "n" + "n'" in *) *) -(* (* (FAA "l" "x", "x"). *) *) -(* Definition allocate_tape2 : val := flip_allocate_tape. *) -(* Definition incr_counter_tape2 :val := λ: "l" "α", let: "n" := *) -(* conversion.bool_to_int (flip_tape "α") *) -(* in *) -(* let: "n'" := *) -(* conversion.bool_to_int (flip_tape "α") *) -(* in *) -(* let: "x" := #2 * "n" + "n'" in *) -(* (FAA "l" "x", "x"). *) -(* Definition read_counter2 : val := λ: "l", !"l". *) -(* Class counterG2 Σ := CounterG2 { (* counterG2_tapes:: hocap_tapesGS' Σ; *) *) -(* counterG2_frac_authR:: inG Σ (frac_authR natR); *) -(* counterG2_flipG: flipG Σ *) -(* }. *) +Section impl2. + Context `{F: flip_spec Σ}. + Definition new_counter2 : val:= λ: "_", ref #0. + (* Definition incr_counter2 : val := λ: "l", let: "n" := rand #1 in *) + (* let: "n'" := rand #1 in *) + (* let: "x" := #2 * "n" + "n'" in *) + (* (FAA "l" "x", "x"). *) + Definition allocate_tape2 : val := flip_allocate_tape. + Definition incr_counter_tape2 :val := λ: "l" "α", let: "n" := + conversion.bool_to_int (flip_tape "α") + in + let: "n'" := + conversion.bool_to_int (flip_tape "α") + in + let: "x" := #2 * "n" + "n'" in + (FAA "l" "x", "x"). + Definition read_counter2 : val := λ: "l", !"l". + Class counterG2 Σ := CounterG2 { (* counterG2_tapes:: hocap_tapesGS' Σ; *) + counterG2_frac_authR:: inG Σ (frac_authR natR); + counterG2_flipG: flipG Σ + }. -(* Context `{L:!flipG Σ, !inG Σ (frac_authR natR)}. *) + Context `{L:!flipG Σ, !inG Σ (frac_authR natR)}. -(* Definition counter_inv_pred2 (c:val) γ := *) -(* (∃ (* (m:gmap loc (bool*list nat)) *) (l:loc) (z:nat), *) -(* (* ([∗ map] α ↦ t ∈ m, flip_tapes_frag (L:=L) γ1' α (if t.1:bool then nat_to_bool <$> (expander t.2) else nat_to_bool <$> (drop 1%nat (expander t.2))) ) *) *) -(* (* ∗ ●m@γ1 ∗ *) *) -(* ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ (●F z) *) -(* )%I. *) + Definition counter_inv_pred2 (c:val) γ := + (∃ (l:loc) (z:nat), + ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ (●F z) + )%I. -(* Lemma new_counter_spec2 E ε N: *) -(* {{{ ↯ ε }}} *) -(* new_counter2 #() @ E *) -(* {{{ (c:val), RET c; *) -(* ∃ γ1 γ2 γ3, (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ *) -(* inv (N.@ "counter") (counter_inv_pred2 c γ3)) ∗ *) -(* flip_error_frag (L:=L) γ1 ε ∗ own γ3 (◯F 0%nat) *) -(* }}}. *) -(* Proof. *) -(* rewrite /new_counter2. *) -(* iIntros (Φ) "Hε HΦ". *) -(* wp_pures. *) -(* iDestruct (ec_valid with "[$]") as "%". *) -(* iMod (flip_inv_create_spec with "[$]") as "(%γ1'&%γ2'&#?&Hε)". *) -(* wp_alloc l as "Hl". *) -(* simpl. *) -(* (* iMod (hocap_tapes_alloc' (∅:gmap _ _)) as "[%γ2 [H3 H4]]". *) *) -(* iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". *) -(* { by apply frac_auth_valid. } *) -(* replace (#0) with (#0%nat) by done. *) -(* iMod (inv_alloc (N.@"counter") _ (counter_inv_pred2 (#l) γ3) with "[$Hl $H5]") as "#Hinv"; first done. *) -(* iApply "HΦ". *) -(* iExists _, _, _. iModIntro. iFrame. *) -(* by iSplit. *) -(* Qed. *) + Definition is_counter2 N (c:val) γ1 γ2:= + (is_flip (L:=L) (N.@"flip") γ1 ∗ + inv (N.@"counter") (counter_inv_pred2 c γ2) + )%I. + Lemma new_counter_spec2 E ε N: + {{{ ↯ ε }}} + new_counter2 #() @ E + {{{ (c:val), RET c; + ∃ γ1 γ2, is_counter2 N c γ1 γ2 ∗ own γ2 (◯F 0%nat) + }}}. + Proof. + Admitted. + (* rewrite /new_counter2. *) + (* iIntros (Φ) "Hε HΦ". *) + (* wp_pures. *) + (* iDestruct (ec_valid with "[$]") as "%". *) + (* iMod (flip_inv_create_spec with "[$]") as "(%γ1'&%γ2'&#?&Hε)". *) + (* wp_alloc l as "Hl". *) + (* simpl. *) + (* (* iMod (hocap_tapes_alloc' (∅:gmap _ _)) as "[%γ2 [H3 H4]]". *) *) + (* iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". *) + (* { by apply frac_auth_valid. } *) + (* replace (#0) with (#0%nat) by done. *) + (* iMod (inv_alloc (N.@"counter") _ (counter_inv_pred2 (#l) γ3) with "[$Hl $H5]") as "#Hinv"; first done. *) + (* iApply "HΦ". *) + (* iExists _, _, _. iModIntro. iFrame. *) + (* by iSplit. *) + (* Qed. *) -(* (** This lemma is not possible as only one view shift*) *) -(* (* Lemma incr_counter_spec2 E N c γ1 γ2 γ3 (ε2:R -> nat -> R) (P: iProp Σ) (T: nat -> iProp Σ) (Q: nat->nat->iProp Σ): *) *) -(* (* ↑N ⊆ E-> *) *) -(* (* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) *) -(* (* (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → *) *) -(* (* {{{ inv N (counter_inv_pred2 c γ1 γ2 γ3) ∗ *) *) -(* (* □(∀ (ε:R) (n : nat), P ∗ ●↯ ε @ γ1 ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨●↯ (ε2 ε n) @ γ1 ∗ T n) ) ∗ *) *) -(* (* □ (∀ (n z:nat), T n ∗ own γ3 (●F z) ={E∖↑N}=∗ *) *) -(* (* own γ3 (●F(z+n)%nat)∗ Q z n) ∗ *) *) -(* (* P *) *) -(* (* }}} *) *) -(* (* incr_counter2 c @ E *) *) -(* (* {{{ (n:nat) (z:nat), RET (#z, #n); Q z n }}}. *) *) -(* (* Proof. *) *) -(* (* iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP) HΦ". *) *) -(* (* rewrite /incr_counter2. *) *) -(* (* wp_pures. *) *) -(* (* wp_bind (rand _)%E. *) *) -(* (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) *) -(* (* (** cant do two view shifts! *) *) *) -(* (* Abort. *) *) -(* Lemma allocate_tape_spec2 N E c γ1 γ2 γ3: *) -(* ↑N ⊆ E-> *) -(* {{{ (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ *) -(* inv (N.@ "counter") (counter_inv_pred2 c γ3)) }}} *) -(* allocate_tape2 #() @ E *) -(* {{{ (v:val), RET v; *) -(* ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ (flip_tapes_frag (L:=L) γ2 α (expander []) ∗ ⌜Forall (λ x, x<4) []⌝) *) -(* }}}. *) -(* Proof. *) -(* iIntros (Hsubset Φ) "(#Hinv & #Hinv') HΦ". *) -(* rewrite /allocate_tape2. *) -(* iApply wptac_wp_fupd. *) -(* iApply (flip_allocate_tape_spec); [by eapply nclose_subseteq'|done|]. *) -(* iNext. *) -(* iIntros (?) "(%α & -> & Hfrag)". *) -(* iApply "HΦ". *) -(* iFrame. *) -(* iPureIntro. *) -(* split; first done. *) -(* by apply Forall_nil. *) -(* Qed. *) + (** This lemma is not possible as only one view shift*) + (* Lemma incr_counter_spec2 E N c γ1 γ2 γ3 (ε2:R -> nat -> R) (P: iProp Σ) (T: nat -> iProp Σ) (Q: nat->nat->iProp Σ): *) + (* ↑N ⊆ E-> *) + (* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) + (* (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → *) + (* {{{ inv N (counter_inv_pred2 c γ1 γ2 γ3) ∗ *) + (* □(∀ (ε:R) (n : nat), P ∗ ●↯ ε @ γ1 ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨●↯ (ε2 ε n) @ γ1 ∗ T n) ) ∗ *) + (* □ (∀ (n z:nat), T n ∗ own γ3 (●F z) ={E∖↑N}=∗ *) + (* own γ3 (●F(z+n)%nat)∗ Q z n) ∗ *) + (* P *) + (* }}} *) + (* incr_counter2 c @ E *) + (* {{{ (n:nat) (z:nat), RET (#z, #n); Q z n }}}. *) + (* Proof. *) + (* iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP) HΦ". *) + (* rewrite /incr_counter2. *) + (* wp_pures. *) + (* wp_bind (rand _)%E. *) + (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* (** cant do two view shifts! *) *) + (* Abort. *) + + Lemma allocate_tape_spec2 N E c γ1 γ2: + ↑N ⊆ E-> + {{{ is_counter2 N c γ1 γ2 }}} + allocate_tape2 #() @ E + {{{ (v:val), RET v; + ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ (flip_tapes_frag (L:=L) γ1 α (expander []) ∗ ⌜Forall (λ x, x<4) []⌝) + }}}. + Proof. + Admitted. + (* iIntros (Hsubset Φ) "(#Hinv & #Hinv') HΦ". *) + (* rewrite /allocate_tape2. *) + (* iApply wptac_wp_fupd. *) + (* iApply (flip_allocate_tape_spec); [by eapply nclose_subseteq'|done|]. *) + (* iNext. *) + (* iIntros (?) "(%α & -> & Hfrag)". *) + (* iApply "HΦ". *) + (* iFrame. *) + (* iPureIntro. *) + (* split; first done. *) + (* by apply Forall_nil. *) + (* Qed. *) -(* Lemma incr_counter_tape_spec_some2 N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: *) -(* ↑N⊆E -> *) -(* {{{ (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ *) -(* inv (N.@ "counter") (counter_inv_pred2 c γ3)) ∗ *) -(* □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ *) -(* own γ3 (●F(z+n)%nat)∗ Q z) ∗ *) -(* P ∗ (flip_tapes_frag (L:=L) γ2 α (expander (n::ns))∗ ⌜Forall (λ x, x<4) (n::ns)⌝) *) -(* }}} *) -(* incr_counter_tape2 c #lbl:α @ E *) -(* {{{ (z:nat), RET (#z, #n); Q z ∗ (flip_tapes_frag (L:=L) γ2 α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝)}}}. *) -(* Proof. *) -(* iIntros (Hsubset Φ) "((#Hinv & #Hinv') & #Hvs & HP & [Hα %]) HΦ". *) -(* rewrite /incr_counter_tape2. *) -(* wp_pures. *) -(* wp_apply (flip_tape_spec_some _ _ _ _ (flip_tapes_frag γ2 α (expander (n::ns)) ) (flip_tapes_frag γ2 α (drop 1%nat (expander (n::ns))))with "[$Hα]"); [|iSplit; first done|]. *) -(* { by apply nclose_subseteq'. } *) -(* { iModIntro. iIntros (m) "[Hfrag Hauth]". *) -(* iDestruct (flip_tapes_agree with "[$][$]") as "%". *) -(* simpl. *) -(* iMod (flip_tapes_pop with "[$][$]") as "[Hauth Htape]". *) -(* iModIntro. *) -(* simpl in *. *) -(* by iFrame. *) -(* } *) -(* iIntros "Hα". *) -(* wp_apply (conversion.wp_bool_to_int) as "_"; first done. *) -(* wp_pures. *) -(* wp_apply (flip_tape_spec_some _ _ _ _ (flip_tapes_frag γ2 α (drop 1%nat (expander (n::ns)))) (flip_tapes_frag γ2 α (expander ns))with "[$Hα]"); [|iSplit; first done|]. *) -(* { by apply nclose_subseteq'. } *) -(* { iModIntro. iIntros (m) "[Hfrag Hauth]". *) -(* iDestruct (flip_tapes_agree with "[$][$]") as "%". *) -(* simpl. *) -(* iMod (flip_tapes_pop with "[$][$]") as "[Hauth Htape]". *) -(* iModIntro. *) -(* simpl in *. *) -(* by iFrame. *) -(* } *) -(* iIntros "Hα". *) -(* wp_apply (conversion.wp_bool_to_int) as "_"; first done. *) -(* wp_pures. *) -(* wp_bind (FAA _ _). *) -(* iInv "Hinv'" as ">(%&%&->&?&?)" "Hclose". *) -(* wp_faa. simpl. *) -(* iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". *) -(* { apply difference_mono; [done|by apply nclose_subseteq']. } *) -(* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) -(* replace 2%Z with (Z.of_nat 2%nat) by done. *) -(* replace (_*_+_)%Z with (Z.of_nat n); last first. *) -(* { assert (n<4). *) -(* - by eapply (@Forall_inv _ (λ x, x<4)). *) -(* - by apply expander_eta. *) -(* } *) -(* replace (#(z+n)) with (#(z+n)%nat); last first. *) -(* { by rewrite Nat2Z.inj_add. } *) -(* iMod "Hclose'" as "_". *) -(* iMod ("Hclose" with "[-HΦ HQ Hα]") as "_"; first by iFrame. *) -(* iModIntro. *) -(* wp_pures. *) -(* iApply "HΦ". *) -(* iFrame. *) -(* iPureIntro. *) -(* by eapply Forall_inv_tail. *) -(* Qed. *) + (* Lemma incr_counter_tape_spec_some2 N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: *) + (* ↑N⊆E -> *) + (* {{{ (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ *) + (* inv (N.@ "counter") (counter_inv_pred2 c γ3)) ∗ *) + (* □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ *) + (* own γ3 (●F(z+n)%nat)∗ Q z) ∗ *) + (* P ∗ (flip_tapes_frag (L:=L) γ2 α (expander (n::ns))∗ ⌜Forall (λ x, x<4) (n::ns)⌝) *) + (* }}} *) + (* incr_counter_tape2 c #lbl:α @ E *) + (* {{{ (z:nat), RET (#z, #n); Q z ∗ (flip_tapes_frag (L:=L) γ2 α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝)}}}. *) + (* Proof. *) + (* iIntros (Hsubset Φ) "((#Hinv & #Hinv') & #Hvs & HP & [Hα %]) HΦ". *) + (* rewrite /incr_counter_tape2. *) + (* wp_pures. *) + (* wp_apply (flip_tape_spec_some _ _ _ _ (flip_tapes_frag γ2 α (expander (n::ns)) ) (flip_tapes_frag γ2 α (drop 1%nat (expander (n::ns))))with "[$Hα]"); [|iSplit; first done|]. *) + (* { by apply nclose_subseteq'. } *) + (* { iModIntro. iIntros (m) "[Hfrag Hauth]". *) + (* iDestruct (flip_tapes_agree with "[$][$]") as "%". *) + (* simpl. *) + (* iMod (flip_tapes_pop with "[$][$]") as "[Hauth Htape]". *) + (* iModIntro. *) + (* simpl in *. *) + (* by iFrame. *) + (* } *) + (* iIntros "Hα". *) + (* wp_apply (conversion.wp_bool_to_int) as "_"; first done. *) + (* wp_pures. *) + (* wp_apply (flip_tape_spec_some _ _ _ _ (flip_tapes_frag γ2 α (drop 1%nat (expander (n::ns)))) (flip_tapes_frag γ2 α (expander ns))with "[$Hα]"); [|iSplit; first done|]. *) + (* { by apply nclose_subseteq'. } *) + (* { iModIntro. iIntros (m) "[Hfrag Hauth]". *) + (* iDestruct (flip_tapes_agree with "[$][$]") as "%". *) + (* simpl. *) + (* iMod (flip_tapes_pop with "[$][$]") as "[Hauth Htape]". *) + (* iModIntro. *) + (* simpl in *. *) + (* by iFrame. *) + (* } *) + (* iIntros "Hα". *) + (* wp_apply (conversion.wp_bool_to_int) as "_"; first done. *) + (* wp_pures. *) + (* wp_bind (FAA _ _). *) + (* iInv "Hinv'" as ">(%&%&->&?&?)" "Hclose". *) + (* wp_faa. simpl. *) + (* iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". *) + (* { apply difference_mono; [done|by apply nclose_subseteq']. } *) + (* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) + (* replace 2%Z with (Z.of_nat 2%nat) by done. *) + (* replace (_*_+_)%Z with (Z.of_nat n); last first. *) + (* { assert (n<4). *) + (* - by eapply (@Forall_inv _ (λ x, x<4)). *) + (* - by apply expander_eta. *) + (* } *) + (* replace (#(z+n)) with (#(z+n)%nat); last first. *) + (* { by rewrite Nat2Z.inj_add. } *) + (* iMod "Hclose'" as "_". *) + (* iMod ("Hclose" with "[-HΦ HQ Hα]") as "_"; first by iFrame. *) + (* iModIntro. *) + (* wp_pures. *) + (* iApply "HΦ". *) + (* iFrame. *) + (* iPureIntro. *) + (* by eapply Forall_inv_tail. *) + (* Qed. *) -(* Lemma counter_presample_spec2 N E ns α *) -(* (ε2 : R -> nat -> R) *) -(* (P : iProp Σ) T γ1 γ2 γ3 c: *) -(* ↑N ⊆ E -> *) -(* (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> *) -(* (∀ (ε:R), 0<= ε ->SeriesC (λ n, if (bool_decide (n≤3%nat)) then 1 / (S 3%nat) * ε2 ε n else 0%R)%R <= ε)%R-> *) -(* (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ *) -(* inv (N.@ "counter") (counter_inv_pred2 c γ3)) -∗ *) -(* (□∀ (ε:R) n, (P ∗ flip_error_auth (L:=L) γ1 ε) ={E∖↑N}=∗ *) -(* (⌜(1<=ε2 ε n)%R⌝ ∨(flip_error_auth (L:=L) γ1 (ε2 ε n) ∗ T (n)))) *) -(* -∗ *) -(* P -∗ (flip_tapes_frag (L:=L) γ2 α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝) -∗ *) -(* wp_update E (∃ n, T (n) ∗ (flip_tapes_frag (L:=L) γ2 α (expander (ns++[n])) ∗ ⌜Forall (λ x, x<4) (ns++[n])⌝)). *) -(* Proof. *) -(* iIntros (Hsubset Hpos Hineq) "(#Hinv & #Hinv') #Hvs HP Hfrag". *) -(* iApply fupd_wp_update. *) -(* Admitted. *) -(* (* iApply flip_presample_spec. *) *) -(* (* iApply wp_update_state_step_coupl. *) *) -(* (* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) *) -(* (* iMod (inv_acc with "Hinv") as "[>(% & % & % & % & H1 & H2 & H3 & H4 & -> & H5 & H6) Hclose]"; [done|]. *) *) -(* (* iDestruct (hocap_tapes_agree' with "[$][$]") as "%". *) *) -(* (* erewrite <-(insert_delete m) at 1; last done. *) *) -(* (* rewrite big_sepM_insert; last apply lookup_delete. *) *) -(* (* simpl. *) *) -(* (* iDestruct "H3" as "[Htape H3]". *) *) -(* (* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%H1)". *) *) -(* (* iDestruct (ec_supply_bound with "[$][$]") as "%". *) *) -(* (* iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. *) *) -(* (* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) *) -(* (* iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang 2%nat); first done. *) *) -(* (* pose (f (l:list (fin 2)) := (match l with *) *) -(* (* | x::[x'] => 2*fin_to_nat x + fin_to_nat x' *) *) -(* (* | _ => 0 *) *) -(* (* end)). *) *) -(* (* unshelve iExists *) *) -(* (* (λ l, mknonnegreal (ε2 ε1' (f l) ) _). *) *) -(* (* { apply Hpos; simpl. auto. } *) *) -(* (* simpl. *) *) -(* (* iSplit. *) *) -(* (* { iPureIntro. *) *) -(* (* etrans; last apply Hineq; last auto. *) *) -(* (* pose (k:=(λ n, 1 / ((1 + 1) * ((1 + 1) * 1)) * ε2 ε1' (f n))%R). *) *) -(* (* erewrite (SeriesC_ext _ (λ x, if bool_decide (x ∈ enum_uniform_list 1%nat 2%nat) *) *) -(* (* then k x else 0%R))%R; last first. *) *) -(* (* - intros. case_bool_decide as K. *) *) -(* (* + rewrite elem_of_enum_uniform_list in K. by rewrite K /k. *) *) -(* (* + rewrite elem_of_enum_uniform_list in K. *) *) -(* (* case_match eqn:K'; [by rewrite Nat.eqb_eq in K'|done]. *) *) -(* (* - rewrite SeriesC_list; last apply NoDup_enum_uniform_list. *) *) -(* (* rewrite /k /=. rewrite SeriesC_nat_bounded'/=. lra. *) *) -(* (* } *) *) -(* (* iIntros (sample ?). *) *) -(* (* destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' (f sample)))%R 1%R) as [Hdec|Hdec]; last first. *) *) -(* (* { apply Rnot_lt_ge, Rge_le in Hdec. *) *) -(* (* iApply state_step_coupl_ret_err_ge_1. *) *) -(* (* simpl. simpl in *. lra. *) *) -(* (* } *) *) -(* (* iApply state_step_coupl_ret. *) *) -(* (* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' (f sample)) _) with "Hε_supply") as "[Hε_supply Hε]". *) *) -(* (* { apply Hpos. apply cond_nonneg. } *) *) -(* (* { simpl. done. } *) *) -(* (* assert (Nat.div2 (f sample)<2)%nat as K1. *) *) -(* (* { rewrite Nat.div2_div. apply Nat.Div0.div_lt_upper_bound. rewrite /f. *) *) -(* (* simpl. repeat case_match; try lia. pose proof fin_to_nat_lt t. pose proof fin_to_nat_lt t0. lia. *) *) -(* (* } *) *) -(* (* rewrite -H1. *) *) -(* (* iMod (tapeN_update_append _ _ _ _ (nat_to_fin K1) with "[$][$]") as "[Htapes Htape]". *) *) -(* (* assert (Nat.b2n (Nat.odd (f sample))<2)%nat as K2. *) *) -(* (* { edestruct (Nat.odd _); simpl; lia. } *) *) -(* (* rewrite -(list_fmap_singleton (fin_to_nat)) -fmap_app. *) *) -(* (* iMod (tapeN_update_append _ _ _ _ (nat_to_fin K2) with "[$][$]") as "[Htapes Htape]". *) *) -(* (* iMod (hocap_tapes_presample' _ _ _ _ _ (f sample) with "[$][$]") as "[H4 Hfrag]". *) *) -(* (* iMod "Hclose'" as "_". *) *) -(* (* iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". *) *) -(* (* { iExFalso. iApply (ec_contradict with "[$]"). exact. } *) *) -(* (* rewrite insert_insert. *) *) -(* (* rewrite fmap_app -!app_assoc /=. *) *) -(* (* assert (([nat_to_fin K1;nat_to_fin K2]) = sample) as K. *) *) -(* (* { destruct sample as [|x xs]; first done. *) *) -(* (* destruct xs as [|x' xs]; first done. *) *) -(* (* destruct xs as [|]; last done. *) *) -(* (* pose proof fin_to_nat_lt x. pose proof fin_to_nat_lt x'. *) *) -(* (* repeat f_equal; apply fin_to_nat_inj; rewrite fin_to_nat_to_fin. *) *) -(* (* - rewrite /f. rewrite Nat.div2_div. *) *) -(* (* rewrite Nat.mul_comm Nat.div_add_l; last lia. rewrite Nat.div_small; lia. *) *) -(* (* - rewrite /f. rewrite Nat.add_comm Nat.odd_add_even. *) *) -(* (* + destruct (fin_to_nat x') as [|[|]]; simpl; lia. *) *) -(* (* + by econstructor. *) *) -(* (* } *) *) -(* (* iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4 $H5 $H6]") as "_". *) *) -(* (* { iNext. iSplit; last done. *) *) -(* (* rewrite big_sepM_insert_delete; iFrame. *) *) -(* (* simpl. rewrite !fin_to_nat_to_fin. *) *) -(* (* rewrite /expander bind_app -/(expander ns). simpl. by rewrite H1. *) *) -(* (* } *) *) -(* (* iApply fupd_mask_intro_subseteq; first set_solver. *) *) -(* (* rewrite K. *) *) -(* (* iFrame. *) *) -(* (* Qed. *) *) + (* Lemma counter_presample_spec2 N E ns α *) + (* (ε2 : R -> nat -> R) *) + (* (P : iProp Σ) T γ1 γ2 γ3 c: *) + (* ↑N ⊆ E -> *) + (* (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> *) + (* (∀ (ε:R), 0<= ε ->SeriesC (λ n, if (bool_decide (n≤3%nat)) then 1 / (S 3%nat) * ε2 ε n else 0%R)%R <= ε)%R-> *) + (* (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ *) + (* inv (N.@ "counter") (counter_inv_pred2 c γ3)) -∗ *) + (* (□∀ (ε:R) n, (P ∗ flip_error_auth (L:=L) γ1 ε) ={E∖↑N}=∗ *) + (* (⌜(1<=ε2 ε n)%R⌝ ∨(flip_error_auth (L:=L) γ1 (ε2 ε n) ∗ T (n)))) *) + (* -∗ *) + (* P -∗ (flip_tapes_frag (L:=L) γ2 α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝) -∗ *) + (* wp_update E (∃ n, T (n) ∗ (flip_tapes_frag (L:=L) γ2 α (expander (ns++[n])) ∗ ⌜Forall (λ x, x<4) (ns++[n])⌝)). *) + (* Proof. *) + (* iIntros (Hsubset Hpos Hineq) "(#Hinv & #Hinv') #Hvs HP Hfrag". *) + (* iApply fupd_wp_update. *) + (* Admitted. *) + (* iApply flip_presample_spec. *) + (* iApply wp_update_state_step_coupl. *) + (* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) + (* iMod (inv_acc with "Hinv") as "[>(% & % & % & % & H1 & H2 & H3 & H4 & -> & H5 & H6) Hclose]"; [done|]. *) + (* iDestruct (hocap_tapes_agree' with "[$][$]") as "%". *) + (* erewrite <-(insert_delete m) at 1; last done. *) + (* rewrite big_sepM_insert; last apply lookup_delete. *) + (* simpl. *) + (* iDestruct "H3" as "[Htape H3]". *) + (* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%H1)". *) + (* iDestruct (ec_supply_bound with "[$][$]") as "%". *) + (* iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. *) + (* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) + (* iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang 2%nat); first done. *) + (* pose (f (l:list (fin 2)) := (match l with *) + (* | x::[x'] => 2*fin_to_nat x + fin_to_nat x' *) + (* | _ => 0 *) + (* end)). *) + (* unshelve iExists *) + (* (λ l, mknonnegreal (ε2 ε1' (f l) ) _). *) + (* { apply Hpos; simpl. auto. } *) + (* simpl. *) + (* iSplit. *) + (* { iPureIntro. *) + (* etrans; last apply Hineq; last auto. *) + (* pose (k:=(λ n, 1 / ((1 + 1) * ((1 + 1) * 1)) * ε2 ε1' (f n))%R). *) + (* erewrite (SeriesC_ext _ (λ x, if bool_decide (x ∈ enum_uniform_list 1%nat 2%nat) *) + (* then k x else 0%R))%R; last first. *) + (* - intros. case_bool_decide as K. *) + (* + rewrite elem_of_enum_uniform_list in K. by rewrite K /k. *) + (* + rewrite elem_of_enum_uniform_list in K. *) + (* case_match eqn:K'; [by rewrite Nat.eqb_eq in K'|done]. *) + (* - rewrite SeriesC_list; last apply NoDup_enum_uniform_list. *) + (* rewrite /k /=. rewrite SeriesC_nat_bounded'/=. lra. *) + (* } *) + (* iIntros (sample ?). *) + (* destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' (f sample)))%R 1%R) as [Hdec|Hdec]; last first. *) + (* { apply Rnot_lt_ge, Rge_le in Hdec. *) + (* iApply state_step_coupl_ret_err_ge_1. *) + (* simpl. simpl in *. lra. *) + (* } *) + (* iApply state_step_coupl_ret. *) + (* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' (f sample)) _) with "Hε_supply") as "[Hε_supply Hε]". *) + (* { apply Hpos. apply cond_nonneg. } *) + (* { simpl. done. } *) + (* assert (Nat.div2 (f sample)<2)%nat as K1. *) + (* { rewrite Nat.div2_div. apply Nat.Div0.div_lt_upper_bound. rewrite /f. *) + (* simpl. repeat case_match; try lia. pose proof fin_to_nat_lt t. pose proof fin_to_nat_lt t0. lia. *) + (* } *) + (* rewrite -H1. *) + (* iMod (tapeN_update_append _ _ _ _ (nat_to_fin K1) with "[$][$]") as "[Htapes Htape]". *) + (* assert (Nat.b2n (Nat.odd (f sample))<2)%nat as K2. *) + (* { edestruct (Nat.odd _); simpl; lia. } *) + (* rewrite -(list_fmap_singleton (fin_to_nat)) -fmap_app. *) + (* iMod (tapeN_update_append _ _ _ _ (nat_to_fin K2) with "[$][$]") as "[Htapes Htape]". *) + (* iMod (hocap_tapes_presample' _ _ _ _ _ (f sample) with "[$][$]") as "[H4 Hfrag]". *) + (* iMod "Hclose'" as "_". *) + (* iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". *) + (* { iExFalso. iApply (ec_contradict with "[$]"). exact. } *) + (* rewrite insert_insert. *) + (* rewrite fmap_app -!app_assoc /=. *) + (* assert (([nat_to_fin K1;nat_to_fin K2]) = sample) as K. *) + (* { destruct sample as [|x xs]; first done. *) + (* destruct xs as [|x' xs]; first done. *) + (* destruct xs as [|]; last done. *) + (* pose proof fin_to_nat_lt x. pose proof fin_to_nat_lt x'. *) + (* repeat f_equal; apply fin_to_nat_inj; rewrite fin_to_nat_to_fin. *) + (* - rewrite /f. rewrite Nat.div2_div. *) + (* rewrite Nat.mul_comm Nat.div_add_l; last lia. rewrite Nat.div_small; lia. *) + (* - rewrite /f. rewrite Nat.add_comm Nat.odd_add_even. *) + (* + destruct (fin_to_nat x') as [|[|]]; simpl; lia. *) + (* + by econstructor. *) + (* } *) + (* iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4 $H5 $H6]") as "_". *) + (* { iNext. iSplit; last done. *) + (* rewrite big_sepM_insert_delete; iFrame. *) + (* simpl. rewrite !fin_to_nat_to_fin. *) + (* rewrite /expander bind_app -/(expander ns). simpl. by rewrite H1. *) + (* } *) + (* iApply fupd_mask_intro_subseteq; first set_solver. *) + (* rewrite K. *) + (* iFrame. *) + (* Qed. *) -(* Lemma read_counter_spec2 N E c γ1 γ2 γ3 P Q: *) -(* ↑N ⊆ E -> *) -(* {{{ (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ *) -(* inv (N.@ "counter") (counter_inv_pred2 c γ3)) ∗ *) -(* □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ *) -(* own γ3 (●F z)∗ Q z) *) -(* ∗ P *) -(* }}} *) -(* read_counter2 c @ E *) -(* {{{ (n':nat), RET #n'; Q n' *) -(* }}}. *) -(* Proof. *) -(* iIntros (Hsubset Φ) "((#Hinv & #Hinv') & #Hvs & HP) HΦ". *) -(* rewrite /read_counter2. *) -(* wp_pure. *) -(* iInv "Hinv'" as ">(%l&%z & -> & Hloc & Hcont)" "Hclose". *) -(* wp_load. *) -(* iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". *) -(* { apply difference_mono_l. by apply nclose_subseteq'. } *) -(* iMod ("Hvs" with "[$]") as "[Hcont HQ]". *) -(* iMod "Hclose'". *) -(* iMod ("Hclose" with "[ $Hloc $Hcont]"); first done. *) -(* iApply "HΦ". by iFrame. *) -(* Qed. *) + (* Lemma read_counter_spec2 N E c γ1 γ2 γ3 P Q: *) + (* ↑N ⊆ E -> *) + (* {{{ (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ *) + (* inv (N.@ "counter") (counter_inv_pred2 c γ3)) ∗ *) + (* □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ *) + (* own γ3 (●F z)∗ Q z) *) + (* ∗ P *) + (* }}} *) + (* read_counter2 c @ E *) + (* {{{ (n':nat), RET #n'; Q n' *) + (* }}}. *) + (* Proof. *) + (* iIntros (Hsubset Φ) "((#Hinv & #Hinv') & #Hvs & HP) HΦ". *) + (* rewrite /read_counter2. *) + (* wp_pure. *) + (* iInv "Hinv'" as ">(%l&%z & -> & Hloc & Hcont)" "Hclose". *) + (* wp_load. *) + (* iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". *) + (* { apply difference_mono_l. by apply nclose_subseteq'. } *) + (* iMod ("Hvs" with "[$]") as "[Hcont HQ]". *) + (* iMod "Hclose'". *) + (* iMod ("Hclose" with "[ $Hloc $Hcont]"); first done. *) + (* iApply "HΦ". by iFrame. *) + (* Qed. *) -(* End impl2. *) +End impl2. (* Program Definition random_counter2 `{flip_spec Σ}: random_counter := *) (* {| new_counter := new_counter2; *) @@ -439,7 +440,7 @@ (* apply frac_auth_update. *) (* apply nat_local_update. lia. *) (* Qed. *) -(* Next Obligation. *) +(* Next Obligation. *) (* intros ?? H ?. *) (* apply counterG2_flipG. *) (* Qed. *) From 28ec2c86c508f4182dc3e5cd1b49eddfbf4fee56 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Fri, 27 Sep 2024 15:23:43 +0200 Subject: [PATCH 48/69] CHECKPOINT. START OF REDOING SPECS TO MAKE TAPES LOCAL --- .../coneris/examples/random_counter/impl2.v | 80 +++++++++---------- 1 file changed, 39 insertions(+), 41 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index 456430d2..471725dd 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -125,23 +125,20 @@ Section impl2. ∃ γ1 γ2, is_counter2 N c γ1 γ2 ∗ own γ2 (◯F 0%nat) }}}. Proof. - Admitted. - (* rewrite /new_counter2. *) - (* iIntros (Φ) "Hε HΦ". *) - (* wp_pures. *) - (* iDestruct (ec_valid with "[$]") as "%". *) - (* iMod (flip_inv_create_spec with "[$]") as "(%γ1'&%γ2'&#?&Hε)". *) - (* wp_alloc l as "Hl". *) - (* simpl. *) - (* (* iMod (hocap_tapes_alloc' (∅:gmap _ _)) as "[%γ2 [H3 H4]]". *) *) - (* iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". *) - (* { by apply frac_auth_valid. } *) - (* replace (#0) with (#0%nat) by done. *) - (* iMod (inv_alloc (N.@"counter") _ (counter_inv_pred2 (#l) γ3) with "[$Hl $H5]") as "#Hinv"; first done. *) - (* iApply "HΦ". *) - (* iExists _, _, _. iModIntro. iFrame. *) - (* by iSplit. *) - (* Qed. *) + rewrite /new_counter2. + iIntros (Φ) "_ HΦ". + wp_pures. + iMod (flip_inv_create_spec) as "(%γ1 & #Hinv)". + wp_alloc l as "Hl". + iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ2[H5 H6]]". + { by apply frac_auth_valid. } + replace (#0) with (#0%nat) by done. + iMod (inv_alloc _ _ (counter_inv_pred2 (#l) γ2) with "[$Hl $H5]") as "#Hinv'"; first done. + iApply "HΦ". + iFrame. + iModIntro. + iExists _. by iSplit. + Qed. (** This lemma is not possible as only one view shift*) @@ -173,32 +170,33 @@ Section impl2. {{{ (v:val), RET v; ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ (flip_tapes_frag (L:=L) γ1 α (expander []) ∗ ⌜Forall (λ x, x<4) []⌝) }}}. + Proof. + iIntros (Hsubset Φ) "[#Hinv #Hinv'] HΦ". + rewrite /allocate_tape2. + wp_pures. + wp_apply flip_allocate_tape_spec; [by eapply nclose_subseteq'|done|..]. + iIntros (?) "(%&->&?)". + iApply "HΦ". + iFrame. + iPureIntro. split; first done. + by apply Forall_nil. + Qed. + + Lemma incr_counter_tape_spec_some2 N E c γ1 γ2 (Q:nat->nat -> list nat -> iProp Σ) (α:loc) : + ↑N⊆E -> + {{{ is_counter2 N c γ1 γ2 ∗ + ( |={E∖ ↑N, ∅}=> + ∃ n ns, (flip_tapes_frag (L:=L) γ1 α (expander (n::ns)) ∗ ⌜Forall (λ x, x<4) (n::ns)⌝) ∗ + (flip_tapes_frag (L:=L) γ1 α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝) ={∅, E∖↑N}=∗ + ∀ (z:nat), own γ2 (●F z) + ={E∖↑N}=∗ + own γ2 (●F (z+n)) ∗ Q z n ns) + + }}} + incr_counter_tape2 c #lbl:α @ E + {{{ (z n:nat), RET (#z, #n); ∃ ns, Q z n ns}}}. Proof. Admitted. - (* iIntros (Hsubset Φ) "(#Hinv & #Hinv') HΦ". *) - (* rewrite /allocate_tape2. *) - (* iApply wptac_wp_fupd. *) - (* iApply (flip_allocate_tape_spec); [by eapply nclose_subseteq'|done|]. *) - (* iNext. *) - (* iIntros (?) "(%α & -> & Hfrag)". *) - (* iApply "HΦ". *) - (* iFrame. *) - (* iPureIntro. *) - (* split; first done. *) - (* by apply Forall_nil. *) - (* Qed. *) - - (* Lemma incr_counter_tape_spec_some2 N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: *) - (* ↑N⊆E -> *) - (* {{{ (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ *) - (* inv (N.@ "counter") (counter_inv_pred2 c γ3)) ∗ *) - (* □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ *) - (* own γ3 (●F(z+n)%nat)∗ Q z) ∗ *) - (* P ∗ (flip_tapes_frag (L:=L) γ2 α (expander (n::ns))∗ ⌜Forall (λ x, x<4) (n::ns)⌝) *) - (* }}} *) - (* incr_counter_tape2 c #lbl:α @ E *) - (* {{{ (z:nat), RET (#z, #n); Q z ∗ (flip_tapes_frag (L:=L) γ2 α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝)}}}. *) - (* Proof. *) (* iIntros (Hsubset Φ) "((#Hinv & #Hinv') & #Hvs & HP & [Hα %]) HΦ". *) (* rewrite /incr_counter_tape2. *) (* wp_pures. *) From ef09370560cf28fc07bcf24828545192bbab8ce0 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Fri, 27 Sep 2024 16:12:47 +0200 Subject: [PATCH 49/69] Updated hocap rand and hocap flip. TODO: rand counter and its implementations --- .../coneris/examples/random_counter/impl1.v | 568 +++++++++--------- theories/coneris/lib/hocap_flip.v | 101 ++-- theories/coneris/lib/hocap_rand.v | 191 +++--- 3 files changed, 419 insertions(+), 441 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl1.v b/theories/coneris/examples/random_counter/impl1.v index 03ab881e..22519878 100644 --- a/theories/coneris/examples/random_counter/impl1.v +++ b/theories/coneris/examples/random_counter/impl1.v @@ -1,298 +1,298 @@ -From iris.algebra Require Import frac_auth. -From iris.base_logic.lib Require Import invariants. -From clutch.coneris Require Import coneris hocap hocap_rand random_counter. +(* From iris.algebra Require Import frac_auth. *) +(* From iris.base_logic.lib Require Import invariants. *) +(* From clutch.coneris Require Import coneris hocap hocap_rand random_counter. *) -Set Default Proof Using "Type*". +(* Set Default Proof Using "Type*". *) -Section impl1. - Context `{H:conerisGS Σ, r1:@rand_spec Σ H, L:randG Σ, !inG Σ (frac_authR natR)}. +(* Section impl1. *) +(* Context `{H:conerisGS Σ, r1:@rand_spec Σ H, L:randG Σ, !inG Σ (frac_authR natR)}. *) - Definition new_counter1 : val:= λ: "_", ref #0. - Definition allocate_tape1 : val := λ: "_", rand_allocate_tape #3%nat. - Definition incr_counter_tape1 :val := λ: "l" "α", let: "n" := rand_tape "α" #3%nat in (FAA "l" "n", "n"). - Definition read_counter1 : val := λ: "l", !"l". - Class counterG1 Σ := CounterG1 { counterG1_randG : randG Σ; - counterG1_frac_authR:: inG Σ (frac_authR natR) }. +(* Definition new_counter1 : val:= λ: "_", ref #0. *) +(* Definition allocate_tape1 : val := λ: "_", rand_allocate_tape #3%nat. *) +(* Definition incr_counter_tape1 :val := λ: "l" "α", let: "n" := rand_tape "α" #3%nat in (FAA "l" "n", "n"). *) +(* Definition read_counter1 : val := λ: "l", !"l". *) +(* Class counterG1 Σ := CounterG1 { counterG1_randG : randG Σ; *) +(* counterG1_frac_authR:: inG Σ (frac_authR natR) }. *) - Definition counter_inv_pred1 c γ2 := (∃ (l:loc) (z:nat), ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ2 (●F z) )%I. - Definition is_counter1 N (c:val) γ1 γ2:= - (is_rand (L:=L) (N.@"rand") γ1 ∗ - inv (N.@"counter") (counter_inv_pred1 c γ2) - )%I. +(* Definition counter_inv_pred1 c γ2 := (∃ (l:loc) (z:nat), ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ2 (●F z) )%I. *) +(* Definition is_counter1 N (c:val) γ1 γ2:= *) +(* (is_rand (L:=L) (N.@"rand") γ1 ∗ *) +(* inv (N.@"counter") (counter_inv_pred1 c γ2) *) +(* )%I. *) - Lemma new_counter_spec1 E N: - {{{ True }}} - new_counter1 #() @ E - {{{ (c:val), RET c; - ∃ γ1 γ2, is_counter1 N c γ1 γ2 ∗ own γ2 (◯F 0%nat) - }}}. - Proof. - rewrite /new_counter1. - iIntros (Φ) "_ HΦ". - wp_pures. - iMod (rand_inv_create_spec) as "(%γ1 & #Hinv)". - wp_alloc l as "Hl". - iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ2[H5 H6]]". - { by apply frac_auth_valid. } - replace (#0) with (#0%nat) by done. - iMod (inv_alloc _ _ (counter_inv_pred1 (#l) γ2) with "[$Hl $H5]") as "#Hinv'"; first done. - iApply "HΦ". - iFrame. - iModIntro. - iExists _. by iSplit. - Qed. +(* Lemma new_counter_spec1 E N: *) +(* {{{ True }}} *) +(* new_counter1 #() @ E *) +(* {{{ (c:val), RET c; *) +(* ∃ γ1 γ2, is_counter1 N c γ1 γ2 ∗ own γ2 (◯F 0%nat) *) +(* }}}. *) +(* Proof. *) +(* rewrite /new_counter1. *) +(* iIntros (Φ) "_ HΦ". *) +(* wp_pures. *) +(* iMod (rand_inv_create_spec) as "(%γ1 & #Hinv)". *) +(* wp_alloc l as "Hl". *) +(* iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ2[H5 H6]]". *) +(* { by apply frac_auth_valid. } *) +(* replace (#0) with (#0%nat) by done. *) +(* iMod (inv_alloc _ _ (counter_inv_pred1 (#l) γ2) with "[$Hl $H5]") as "#Hinv'"; first done. *) +(* iApply "HΦ". *) +(* iFrame. *) +(* iModIntro. *) +(* iExists _. by iSplit. *) +(* Qed. *) - (* (** Not used in instantiating type class*) *) - (* Lemma incr_counter_spec1 E N c γ1 γ2 γ3 (ε2:R -> nat -> R) (P: iProp Σ) (T: nat -> iProp Σ) (Q: nat->nat->iProp Σ): *) - (* ↑N ⊆ E-> *) - (* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) - (* (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → *) - (* {{{ inv N (counter_inv_pred1 c γ1 γ2 γ3) ∗ *) - (* □(∀ (ε:R) (n : nat), P ∗ ●↯ ε @ γ1 ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨●↯ (ε2 ε n) @ γ1 ∗ T n) ) ∗ *) - (* □ (∀ (n z:nat), T n ∗ own γ3 (●F z) ={E∖↑N}=∗ *) - (* own γ3 (●F(z+n)%nat)∗ Q z n) ∗ *) - (* P *) - (* }}} *) - (* incr_counter1 c @ E *) - (* {{{ (n:nat) (z:nat), RET (#z, #n); Q z n }}}. *) - (* Proof. *) - (* iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP) HΦ". *) - (* rewrite /incr_counter1. *) - (* wp_pures. *) - (* wp_bind (rand _)%E. *) - (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) - (* iDestruct (ec_valid with "[$]") as "[%K1 %K2]". *) - (* wp_apply (wp_couple_rand_adv_comp1' _ _ _ _ (λ x, ε2 ε (fin_to_nat x)) with "[$]"). *) - (* { intros. naive_solver. } *) - (* { rewrite SeriesC_finite_foldr; specialize (Hineq ε K1). simpl; lra. } *) - (* iIntros (n) "H1". *) - (* iMod ("Hvs1" with "[$]") as "[%|[H2 HT]]". *) - (* { iExFalso. iApply ec_contradict; last done. done. } *) - (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) - (* iModIntro. wp_pures. *) - (* clear -Hsubset. *) - (* wp_bind (FAA _ _). *) - (* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) - (* wp_faa. *) - (* iMod ("Hvs2" with "[$]") as "[H6 HQ]". *) - (* replace (#(z+n)) with (#(z+n)%nat); last first. *) - (* { by rewrite Nat2Z.inj_add. } *) - (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) - (* iModIntro. *) - (* wp_pures. *) - (* by iApply "HΦ". *) - (* Qed. *) +(* (* (** Not used in instantiating type class*) *) *) +(* (* Lemma incr_counter_spec1 E N c γ1 γ2 γ3 (ε2:R -> nat -> R) (P: iProp Σ) (T: nat -> iProp Σ) (Q: nat->nat->iProp Σ): *) *) +(* (* ↑N ⊆ E-> *) *) +(* (* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) *) +(* (* (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → *) *) +(* (* {{{ inv N (counter_inv_pred1 c γ1 γ2 γ3) ∗ *) *) +(* (* □(∀ (ε:R) (n : nat), P ∗ ●↯ ε @ γ1 ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨●↯ (ε2 ε n) @ γ1 ∗ T n) ) ∗ *) *) +(* (* □ (∀ (n z:nat), T n ∗ own γ3 (●F z) ={E∖↑N}=∗ *) *) +(* (* own γ3 (●F(z+n)%nat)∗ Q z n) ∗ *) *) +(* (* P *) *) +(* (* }}} *) *) +(* (* incr_counter1 c @ E *) *) +(* (* {{{ (n:nat) (z:nat), RET (#z, #n); Q z n }}}. *) *) +(* (* Proof. *) *) +(* (* iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP) HΦ". *) *) +(* (* rewrite /incr_counter1. *) *) +(* (* wp_pures. *) *) +(* (* wp_bind (rand _)%E. *) *) +(* (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) *) +(* (* iDestruct (ec_valid with "[$]") as "[%K1 %K2]". *) *) +(* (* wp_apply (wp_couple_rand_adv_comp1' _ _ _ _ (λ x, ε2 ε (fin_to_nat x)) with "[$]"). *) *) +(* (* { intros. naive_solver. } *) *) +(* (* { rewrite SeriesC_finite_foldr; specialize (Hineq ε K1). simpl; lra. } *) *) +(* (* iIntros (n) "H1". *) *) +(* (* iMod ("Hvs1" with "[$]") as "[%|[H2 HT]]". *) *) +(* (* { iExFalso. iApply ec_contradict; last done. done. } *) *) +(* (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) *) +(* (* iModIntro. wp_pures. *) *) +(* (* clear -Hsubset. *) *) +(* (* wp_bind (FAA _ _). *) *) +(* (* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) *) +(* (* wp_faa. *) *) +(* (* iMod ("Hvs2" with "[$]") as "[H6 HQ]". *) *) +(* (* replace (#(z+n)) with (#(z+n)%nat); last first. *) *) +(* (* { by rewrite Nat2Z.inj_add. } *) *) +(* (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) *) +(* (* iModIntro. *) *) +(* (* wp_pures. *) *) +(* (* by iApply "HΦ". *) *) +(* (* Qed. *) *) - Lemma allocate_tape_spec1 N E c γ1 γ2: - ↑N ⊆ E-> - {{{ is_counter1 N c γ1 γ2 }}} - allocate_tape1 #() @ E - {{{ (v:val), RET v; - ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ rand_tapes_frag (L:=L) γ1 α (3%nat, []) - }}}. - Proof. - iIntros (Hsubset Φ) "[#Hinv #Hinv'] HΦ". - rewrite /allocate_tape1. - wp_pures. - wp_apply rand_allocate_tape_spec; [by eapply nclose_subseteq'|done..]. - Qed. +(* Lemma allocate_tape_spec1 N E c γ1 γ2: *) +(* ↑N ⊆ E-> *) +(* {{{ is_counter1 N c γ1 γ2 }}} *) +(* allocate_tape1 #() @ E *) +(* {{{ (v:val), RET v; *) +(* ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ rand_tapes_frag (L:=L) γ1 α (3%nat, []) *) +(* }}}. *) +(* Proof. *) +(* iIntros (Hsubset Φ) "[#Hinv #Hinv'] HΦ". *) +(* rewrite /allocate_tape1. *) +(* wp_pures. *) +(* wp_apply rand_allocate_tape_spec; [by eapply nclose_subseteq'|done..]. *) +(* Qed. *) - Lemma incr_counter_tape_spec_some1 N E c γ1 γ2 (Q:nat->nat -> list nat -> iProp Σ) (α:loc) : - ↑N⊆E -> - {{{ is_counter1 N c γ1 γ2 ∗ - ( |={E∖ ↑N, ∅}=> - ∃ n ns, rand_tapes_frag (L:=L) γ1 α (3, n::ns) ∗ - (rand_tapes_frag (L:=L) γ1 α (3, ns) ={∅, E∖↑N}=∗ - ∀ (z:nat), own γ2 (●F z) - ={E∖↑N}=∗ - own γ2 (●F (z+n)) ∗ Q z n ns) - ) - }}} - incr_counter_tape1 c #lbl:α @ E - {{{ (z n:nat), RET (#z, #n); ∃ ns, Q z n ns}}}. - Proof. - iIntros (Hsubset Φ) "([#Hinv #Hinv'] & Hvs) HΦ". - rewrite /incr_counter_tape1. - wp_pures. - wp_apply (rand_tape_spec_some _ _ _ - (λ n ns, ∀ (z:nat), own γ2 (●F z)={E∖↑N}=∗ - own γ2 (●F (z+n)) ∗ Q z n ns)%I with "[Hvs]"); [by eapply nclose_subseteq'|iSplit; first done|..]. - - iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose". - + apply difference_mono_l. - by apply nclose_subseteq'. - + iMod "Hvs" as "(%n & %ns & Hfrag & Hrest)". - iModIntro. - iFrame. - iIntros "Hfrag". - iMod ("Hrest" with "[$]") as "Hrest". - by iMod "Hclose" as "_". - - iIntros (n) "(%ns & Hvs)". - wp_pures. - wp_bind (FAA _ _). - iInv "Hinv'" as ">(%&%&-> & ?&?)" "Hclose". - wp_faa. - iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". - + apply difference_mono_l. - by apply nclose_subseteq'. - + iMod ("Hvs" with "[$]") as "[? HQ]". - iMod "Hclose'" as "_". - rewrite -Nat2Z.inj_add. - iMod ("Hclose" with "[-HΦ HQ]") as "_"; first by iFrame. - iModIntro. - wp_pures. - iApply "HΦ". by iFrame. - Qed. +(* Lemma incr_counter_tape_spec_some1 N E c γ1 γ2 (Q:nat->nat -> list nat -> iProp Σ) (α:loc) : *) +(* ↑N⊆E -> *) +(* {{{ is_counter1 N c γ1 γ2 ∗ *) +(* ( |={E∖ ↑N, ∅}=> *) +(* ∃ n ns, rand_tapes_frag (L:=L) γ1 α (3, n::ns) ∗ *) +(* (rand_tapes_frag (L:=L) γ1 α (3, ns) ={∅, E∖↑N}=∗ *) +(* ∀ (z:nat), own γ2 (●F z) *) +(* ={E∖↑N}=∗ *) +(* own γ2 (●F (z+n)) ∗ Q z n ns) *) +(* ) *) +(* }}} *) +(* incr_counter_tape1 c #lbl:α @ E *) +(* {{{ (z n:nat), RET (#z, #n); ∃ ns, Q z n ns}}}. *) +(* Proof. *) +(* iIntros (Hsubset Φ) "([#Hinv #Hinv'] & Hvs) HΦ". *) +(* rewrite /incr_counter_tape1. *) +(* wp_pures. *) +(* wp_apply (rand_tape_spec_some _ _ _ *) +(* (λ n ns, ∀ (z:nat), own γ2 (●F z)={E∖↑N}=∗ *) +(* own γ2 (●F (z+n)) ∗ Q z n ns)%I with "[Hvs]"); [by eapply nclose_subseteq'|iSplit; first done|..]. *) +(* - iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose". *) +(* + apply difference_mono_l. *) +(* by apply nclose_subseteq'. *) +(* + iMod "Hvs" as "(%n & %ns & Hfrag & Hrest)". *) +(* iModIntro. *) +(* iFrame. *) +(* iIntros "Hfrag". *) +(* iMod ("Hrest" with "[$]") as "Hrest". *) +(* by iMod "Hclose" as "_". *) +(* - iIntros (n) "(%ns & Hvs)". *) +(* wp_pures. *) +(* wp_bind (FAA _ _). *) +(* iInv "Hinv'" as ">(%&%&-> & ?&?)" "Hclose". *) +(* wp_faa. *) +(* iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". *) +(* + apply difference_mono_l. *) +(* by apply nclose_subseteq'. *) +(* + iMod ("Hvs" with "[$]") as "[? HQ]". *) +(* iMod "Hclose'" as "_". *) +(* rewrite -Nat2Z.inj_add. *) +(* iMod ("Hclose" with "[-HΦ HQ]") as "_"; first by iFrame. *) +(* iModIntro. *) +(* wp_pures. *) +(* iApply "HΦ". by iFrame. *) +(* Qed. *) - Lemma counter_presample_spec1 NS E T γ1 γ2 c: - ↑NS ⊆ E -> - is_counter1 NS c γ1 γ2 -∗ - ( |={E∖↑NS,∅}=> - ∃ ε α ε2 num ns, - ↯ ε ∗ rand_tapes_frag (L:=L) γ1 α (3%nat, ns) ∗ - ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ - ⌜(SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R⌝ ∗ - (∀ ns', ↯ (ε2 ns') ∗ rand_tapes_frag (L:=L) γ1 α (3%nat, ns++ns')={∅,E∖↑NS}=∗ - T ε α ε2 num ns ns' - ))-∗ - wp_update E (∃ ε α ε2 num ns ns', T ε α ε2 num ns ns'). - Proof. - iIntros (Hsubset) "[#Hinv #Hinv'] Hvs". - iMod (rand_presample_spec _ _ (λ ε num tb ε2 α ns ns', - ∃ ε2', ⌜∀ xs ys, fin_to_nat <$> xs = ys -> ε2 xs = ε2' ys⌝ ∗ - ⌜tb = 3⌝ ∗ T ε α ε2' num ns (fin_to_nat <$> ns'))%I with "[//][Hvs]") as "H"; first by apply nclose_subseteq'. - + iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose". - { apply difference_mono_l. - by apply nclose_subseteq'. } - iMod "Hvs" as "(%ε & %α & %ε2 & %num & %ns & Herr & Hfrag & %Hpos & %Hineq & Hrest)". - iFrame. - iModIntro. - iExists num, (λ ls, ε2 (fin_to_nat <$> ls)). - repeat iSplit. - * done. - * iPureIntro. - etrans; last exact. - apply Req_le. - replace (INR 4) with 4%R; last (simpl; lra). - f_equal. - rewrite !SeriesC_list. - -- by rewrite foldr_fmap. - -- apply NoDup_fmap. - ++ apply list_fmap_eq_inj. - apply fin_to_nat_inj. - ++ apply NoDup_enum_uniform_fin_list. - -- apply NoDup_enum_uniform_fin_list. - * iIntros (ns') "[Herr Hfrag]". - iMod ("Hrest" with "[$]") as "HT". - iMod "Hclose" as "_". - iModIntro. - iFrame. - repeat iSplit; iPureIntro; last done. - by intros ??<-. - + iDestruct "H" as "(%ε & %num & %tb & %ε2 & %α & %ns & %ns' & %ε2' & % & -> & HT)". - iModIntro. iFrame. - Qed. +(* Lemma counter_presample_spec1 NS E T γ1 γ2 c: *) +(* ↑NS ⊆ E -> *) +(* is_counter1 NS c γ1 γ2 -∗ *) +(* ( |={E∖↑NS,∅}=> *) +(* ∃ ε α ε2 num ns, *) +(* ↯ ε ∗ rand_tapes_frag (L:=L) γ1 α (3%nat, ns) ∗ *) +(* ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ *) +(* ⌜(SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R⌝ ∗ *) +(* (∀ ns', ↯ (ε2 ns') ∗ rand_tapes_frag (L:=L) γ1 α (3%nat, ns++ns')={∅,E∖↑NS}=∗ *) +(* T ε α ε2 num ns ns' *) +(* ))-∗ *) +(* wp_update E (∃ ε α ε2 num ns ns', T ε α ε2 num ns ns'). *) +(* Proof. *) +(* iIntros (Hsubset) "[#Hinv #Hinv'] Hvs". *) +(* iMod (rand_presample_spec _ _ (λ ε num tb ε2 α ns ns', *) +(* ∃ ε2', ⌜∀ xs ys, fin_to_nat <$> xs = ys -> ε2 xs = ε2' ys⌝ ∗ *) +(* ⌜tb = 3⌝ ∗ T ε α ε2' num ns (fin_to_nat <$> ns'))%I with "[//][Hvs]") as "H"; first by apply nclose_subseteq'. *) +(* + iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose". *) +(* { apply difference_mono_l. *) +(* by apply nclose_subseteq'. } *) +(* iMod "Hvs" as "(%ε & %α & %ε2 & %num & %ns & Herr & Hfrag & %Hpos & %Hineq & Hrest)". *) +(* iFrame. *) +(* iModIntro. *) +(* iExists num, (λ ls, ε2 (fin_to_nat <$> ls)). *) +(* repeat iSplit. *) +(* * done. *) +(* * iPureIntro. *) +(* etrans; last exact. *) +(* apply Req_le. *) +(* replace (INR 4) with 4%R; last (simpl; lra). *) +(* f_equal. *) +(* rewrite !SeriesC_list. *) +(* -- by rewrite foldr_fmap. *) +(* -- apply NoDup_fmap. *) +(* ++ apply list_fmap_eq_inj. *) +(* apply fin_to_nat_inj. *) +(* ++ apply NoDup_enum_uniform_fin_list. *) +(* -- apply NoDup_enum_uniform_fin_list. *) +(* * iIntros (ns') "[Herr Hfrag]". *) +(* iMod ("Hrest" with "[$]") as "HT". *) +(* iMod "Hclose" as "_". *) +(* iModIntro. *) +(* iFrame. *) +(* repeat iSplit; iPureIntro; last done. *) +(* by intros ??<-. *) +(* + iDestruct "H" as "(%ε & %num & %tb & %ε2 & %α & %ns & %ns' & %ε2' & % & -> & HT)". *) +(* iModIntro. iFrame. *) +(* Qed. *) - Lemma read_counter_spec1 N E c γ1 γ2 Q: - ↑N ⊆ E -> - {{{ is_counter1 N c γ1 γ2 ∗ - (∀ (z:nat), own γ2 (●F z) ={E∖↑N}=∗ - own γ2 (●F z) ∗ Q z) +(* Lemma read_counter_spec1 N E c γ1 γ2 Q: *) +(* ↑N ⊆ E -> *) +(* {{{ is_counter1 N c γ1 γ2 ∗ *) +(* (∀ (z:nat), own γ2 (●F z) ={E∖↑N}=∗ *) +(* own γ2 (●F z) ∗ Q z) *) - }}} - read_counter1 c @ E - {{{ (n':nat), RET #n'; Q n' - }}}. - Proof. - iIntros (Hsubset Φ) "([#Hinv #Hinv'] & Hvs) HΦ". - rewrite /read_counter1. - wp_pure. - iInv "Hinv'" as ">( %l & %z & -> & H5 & H6)" "Hclose". - wp_load. - iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". - { apply difference_mono_l. - by apply nclose_subseteq'. } - iMod ("Hvs" with "[$]") as "[? HQ]". - iMod "Hclose'" as "_". - iMod ("Hclose" with "[-HQ HΦ]"); first by iFrame. - iApply ("HΦ" with "[$]"). - Qed. +(* }}} *) +(* read_counter1 c @ E *) +(* {{{ (n':nat), RET #n'; Q n' *) +(* }}}. *) +(* Proof. *) +(* iIntros (Hsubset Φ) "([#Hinv #Hinv'] & Hvs) HΦ". *) +(* rewrite /read_counter1. *) +(* wp_pure. *) +(* iInv "Hinv'" as ">( %l & %z & -> & H5 & H6)" "Hclose". *) +(* wp_load. *) +(* iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". *) +(* { apply difference_mono_l. *) +(* by apply nclose_subseteq'. } *) +(* iMod ("Hvs" with "[$]") as "[? HQ]". *) +(* iMod "Hclose'" as "_". *) +(* iMod ("Hclose" with "[-HQ HΦ]"); first by iFrame. *) +(* iApply ("HΦ" with "[$]"). *) +(* Qed. *) -End impl1. +(* End impl1. *) -Program Definition random_counter1 `{!conerisGS Σ, F:rand_spec}: random_counter := - {| new_counter := new_counter1; - allocate_tape:= allocate_tape1; - incr_counter_tape := incr_counter_tape1; - read_counter:=read_counter1; - counterG := counterG1; - tape_name := rand_tape_name; - counter_name :=gname; - is_counter _ N c γ1 γ2 := is_counter1 (L:=counterG1_randG) N c γ1 γ2; - counter_tapes_auth _ γ m := rand_tapes_auth (L:=counterG1_randG) γ ((λ ns, (3, ns))<$>m); - counter_tapes_frag _ γ α ns := rand_tapes_frag (L:= counterG1_randG)γ α (3%nat, ns); - counter_content_auth _ γ z := own γ (●F z); - counter_content_frag _ γ f z := own γ (◯F{f} z); - new_counter_spec _ := new_counter_spec1; - allocate_tape_spec _ :=allocate_tape_spec1; - incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some1; - counter_presample_spec _ :=counter_presample_spec1; - read_counter_spec _ :=read_counter_spec1 - |}. -Next Obligation. - simpl. - iIntros (?????????) "H1 H2". - iDestruct (rand_tapes_auth_exclusive with "[$][$]") as "[]". -Qed. -Next Obligation. - simpl. - iIntros. - iDestruct (rand_tapes_frag_exclusive with "[$][$]") as "[]". -Qed. -Next Obligation. - simpl. - iIntros. - iDestruct (rand_tapes_agree with "[$][$]") as "%H". - iPureIntro. - apply lookup_fmap_Some in H as (?&?&?). - by simplify_eq. -Qed. -Next Obligation. - iIntros. - iDestruct (rand_tapes_valid with "[$]") as "$". -Qed. -Next Obligation. - simpl. - iIntros. - iMod (rand_tapes_update with "[$][$]") as "[??]"; last iFrame; first done. - by rewrite fmap_insert. -Qed. -Next Obligation. - simpl. - iIntros (?????????) "H1 H2". - iCombine "H1 H2" gives "%H". by rewrite auth_auth_op_valid in H. -Qed. -Next Obligation. - simpl. iIntros (??????? z z' ?) "H1 H2". - iCombine "H1 H2" gives "%H". - apply frac_auth_included_total in H. iPureIntro. - by apply nat_included. -Qed. -Next Obligation. - simpl. - iIntros. - rewrite frac_auth_frag_op. rewrite own_op. - iSplit; iIntros; iFrame. -Qed. -Next Obligation. - simpl. iIntros (?????????) "H1 H2". - iCombine "H1 H2" gives "%H". - iPureIntro. - by apply frac_auth_agree_L in H. -Qed. -Next Obligation. - simpl. iIntros (???????????) "H1 H2". - iMod (own_update_2 _ _ _ (_ ⋅ _) with "[$][$]") as "[$$]"; last done. - apply frac_auth_update. - apply nat_local_update. lia. -Qed. +(* Program Definition random_counter1 `{!conerisGS Σ, F:rand_spec}: random_counter := *) +(* {| new_counter := new_counter1; *) +(* allocate_tape:= allocate_tape1; *) +(* incr_counter_tape := incr_counter_tape1; *) +(* read_counter:=read_counter1; *) +(* counterG := counterG1; *) +(* tape_name := rand_tape_name; *) +(* counter_name :=gname; *) +(* is_counter _ N c γ1 γ2 := is_counter1 (L:=counterG1_randG) N c γ1 γ2; *) +(* counter_tapes_auth _ γ m := rand_tapes_auth (L:=counterG1_randG) γ ((λ ns, (3, ns))<$>m); *) +(* counter_tapes_frag _ γ α ns := rand_tapes_frag (L:= counterG1_randG)γ α (3%nat, ns); *) +(* counter_content_auth _ γ z := own γ (●F z); *) +(* counter_content_frag _ γ f z := own γ (◯F{f} z); *) +(* new_counter_spec _ := new_counter_spec1; *) +(* allocate_tape_spec _ :=allocate_tape_spec1; *) +(* incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some1; *) +(* counter_presample_spec _ :=counter_presample_spec1; *) +(* read_counter_spec _ :=read_counter_spec1 *) +(* |}. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (?????????) "H1 H2". *) +(* iDestruct (rand_tapes_auth_exclusive with "[$][$]") as "[]". *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros. *) +(* iDestruct (rand_tapes_frag_exclusive with "[$][$]") as "[]". *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros. *) +(* iDestruct (rand_tapes_agree with "[$][$]") as "%H". *) +(* iPureIntro. *) +(* apply lookup_fmap_Some in H as (?&?&?). *) +(* by simplify_eq. *) +(* Qed. *) +(* Next Obligation. *) +(* iIntros. *) +(* iDestruct (rand_tapes_valid with "[$]") as "$". *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros. *) +(* iMod (rand_tapes_update with "[$][$]") as "[??]"; last iFrame; first done. *) +(* by rewrite fmap_insert. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros (?????????) "H1 H2". *) +(* iCombine "H1 H2" gives "%H". by rewrite auth_auth_op_valid in H. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. iIntros (??????? z z' ?) "H1 H2". *) +(* iCombine "H1 H2" gives "%H". *) +(* apply frac_auth_included_total in H. iPureIntro. *) +(* by apply nat_included. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. *) +(* iIntros. *) +(* rewrite frac_auth_frag_op. rewrite own_op. *) +(* iSplit; iIntros; iFrame. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. iIntros (?????????) "H1 H2". *) +(* iCombine "H1 H2" gives "%H". *) +(* iPureIntro. *) +(* by apply frac_auth_agree_L in H. *) +(* Qed. *) +(* Next Obligation. *) +(* simpl. iIntros (???????????) "H1 H2". *) +(* iMod (own_update_2 _ _ _ (_ ⋅ _) with "[$][$]") as "[$$]"; last done. *) +(* apply frac_auth_update. *) +(* apply nat_local_update. lia. *) +(* Qed. *) diff --git a/theories/coneris/lib/hocap_flip.v b/theories/coneris/lib/hocap_flip.v index dab64a01..c31d5a29 100644 --- a/theories/coneris/lib/hocap_flip.v +++ b/theories/coneris/lib/hocap_flip.v @@ -49,27 +49,26 @@ Class flip_spec `{!conerisGS Σ} := FlipSpec ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ flip_tapes_frag (L:=L) γ1 α [] }}}; - flip_tape_spec_some {L: flipG Σ} N E γ1 (Q:bool -> list bool -> iProp Σ) (α:loc) : + flip_tape_spec_some {L: flipG Σ} N E γ1 (Q:bool -> list bool -> iProp Σ) (α:loc) n ns: ↑N⊆E -> - {{{ is_flip (L:=L) N γ1 ∗ - (|={E∖↑N, ∅}=> - ∃ n ns, flip_tapes_frag (L:=L) γ1 α (n::ns) ∗ - (flip_tapes_frag (L:=L) γ1 α ns ={∅, E∖↑N}=∗ Q n ns)) + {{{ is_flip (L:=L) N γ1 ∗ flip_tapes_frag (L:=L) γ1 α (n::ns) }}} flip_tape #lbl:α @ E - {{{ (n:bool), RET #n; ∃ ns, Q n ns}}}; + {{{ RET #n; flip_tapes_frag (L:=L) γ1 α ns }}}; - flip_presample_spec {L: flipG Σ} NS E T γ1: + flip_presample_spec {L: flipG Σ} NS E γ1 α ns T: ↑NS ⊆ E -> is_flip (L:=L) NS γ1 -∗ + flip_tapes_frag (L:=L) γ1 α ns -∗ (|={E∖↑NS, ∅}=> - ∃ ε num ε2 α ns, ↯ ε ∗ flip_tapes_frag (L:=L) γ1 α ns ∗ - ⌜(∀ l, length l = num -> 0<=ε2 l)%R⌝ ∗ - ⌜(SeriesC (λ l, if length l =? num then ε2 l else 0) /(2^num) <= ε)%R⌝∗ - (∀ ns', ↯ (ε2 ns') ∗ flip_tapes_frag (L:=L) γ1 α (ns ++ ns') - ={∅, E∖↑NS}=∗ T ε num ε2 α ns ns')) + ∃ ε num ε2, ↯ ε ∗ + ⌜(∀ l, length l = num -> 0<=ε2 l)%R⌝ ∗ + ⌜(SeriesC (λ l, if length l =? num then ε2 l else 0) /(2^num) <= ε)%R⌝∗ + (∀ ns', ↯ (ε2 ns') + ={∅, E∖↑NS}=∗ T ε num ε2 ns')) -∗ - wp_update E (∃ ε num ε2 α ns ns', T ε num ε2 α ns ns') + wp_update E (∃ ε num ε2 ns', T ε num ε2 ns' ∗ + flip_tapes_frag (L:=L) γ1 α (ns ++ ns')) }. @@ -138,45 +137,33 @@ Section instantiate_flip. Qed. Next Obligation. simpl. - iIntros (???? Q ? ? Φ) "(#Hinv & Hvs) HΦ". + iIntros (???? Q ???? Φ) "(#Hinv & Hfrag) HΦ". wp_pures. - wp_apply (rand_tape_spec_some _ _ _ (λ n ns, ∃ b bs, ⌜n= bool_to_nat b⌝ ∗ - ⌜ns = bool_to_nat <$> bs⌝ ∗ - Q b bs - - )%I with "[-HΦ]"); [done|..]. - - iFrame. iSplit; first done. - iMod ("Hvs") as "(%b & %bs & Hfrag & Hrest)". - iFrame. - iModIntro. - iIntros "?". - iMod ("Hrest" with "[$]"). - by iFrame. - - iIntros (n) "(%&%&%&%&%&HQ)". + wp_apply (rand_tape_spec_some with "[-HΦ]"); [done|..]. + - by iFrame. + - iIntros "Hfrag". wp_apply conversion.wp_int_to_bool as "_"; first done. - iApply "HΦ". - subst. - iFrame. - replace (Z_to_bool _) with b; first iFrame. - destruct b; simpl. + replace (Z_to_bool _) with n; first by iApply "HΦ". + destruct n; simpl. + rewrite Z_to_bool_neq_0; lia. + rewrite Z_to_bool_eq_0; lia. Qed. Next Obligation. simpl. - iIntros (??? T ??) "#Hinv Hvs". - iMod (rand_presample_spec _ _ - (λ ε num tb ε2 α ns ns', - ∃ bs bs' ε2', ⌜tb = 1%nat⌝ ∗ ⌜fmap (FMap:= list_fmap) bool_to_nat bs = ns⌝ ∗ ⌜fmap (FMap:= list_fmap) bool_to_nat bs'=fin_to_nat <$> ns'⌝ ∗ + iIntros (?????? T ?) "#Hinv Hfrag Hvs". + iMod (rand_presample_spec _ _ _ _ _ _ + (λ ε num ε2 ns', + ∃ bs' ε2', ⌜fmap (FMap:= list_fmap) bool_to_nat bs'=fin_to_nat <$> ns'⌝ ∗ ⌜∀ xs ys, fmap (FMap:= list_fmap) bool_to_nat xs = fmap (FMap:= list_fmap)fin_to_nat ys -> ε2' xs = ε2 ys⌝ ∗ - T ε num ε2' α bs bs' + T ε num ε2' bs' )%I - with "[][-]") as "(%&%&%&%&%&%&%&%&%&%&->&<-&%&%&HT)"; [exact|exact| |iModIntro; iFrame]. - iMod "Hvs" as "(%&%&%&%&%&Herr&Hfrag&%Hpos&%Hineq&Hvs')". + with "[//][$][-]") as "(%&%&%&%&[(%&%&%K&%&?)?])"; [exact| |iModIntro; iFrame]; last first. + { by rewrite fmap_app K. } + iMod "Hvs" as "(%&%&%&Herr &%Hpos&%Hineq &Hrest)". iFrame. - iModIntro. iExists num, (λ ls, ε2 (nat_to_bool<$>(fin_to_nat <$> ls))). - repeat iSplit; try iPureIntro. + iModIntro. + repeat iSplit; try iPureIntro. - intros. apply Hpos. by rewrite !fmap_length. - etrans; last exact. @@ -203,26 +190,24 @@ Section instantiate_flip. + eapply ex_seriesC_list_length. intros. case_match; last done. by rewrite -Nat.eqb_eq. - - iIntros (ls) "(H2&H3)". - iMod ("Hvs'" with "[$H2 H3]") as "?". - + rewrite fmap_app. rewrite -!list_fmap_compose. + - iIntros (ls) "H2". + iMod ("Hrest" with "[$H2 ]") as "?". + iFrame. + iModIntro. + iSplit; iPureIntro. + + rewrite -!list_fmap_compose. erewrite (Forall_fmap_ext_1 (_∘_)); first done. apply Forall_true. intros x; by repeat (inv_fin x; simpl; try intros x). - + iFrame. iPureIntro; repeat split; try done. - * rewrite -!list_fmap_compose. - erewrite (Forall_fmap_ext_1 (_∘_)); first done. - apply Forall_true. - intros x; by repeat (inv_fin x; simpl; try intros x). - * intros ?? <-. - f_equal. - rewrite -!list_fmap_compose. - rewrite -{1}(list_fmap_id xs). - erewrite (Forall_fmap_ext_1 (_∘_)); first done. - apply Forall_true. - intros []; simpl. - -- by rewrite nat_to_bool_neq_0. - -- by rewrite nat_to_bool_eq_0. + + intros ?? <-. + f_equal. + rewrite -!list_fmap_compose. + rewrite -{1}(list_fmap_id xs). + erewrite (Forall_fmap_ext_1 (_∘_)); first done. + apply Forall_true. + intros []; simpl. + ** by rewrite nat_to_bool_neq_0. + ** by rewrite nat_to_bool_eq_0. Qed. End instantiate_flip. diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index aee65851..bf7ab59b 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -50,26 +50,26 @@ Class rand_spec `{!conerisGS Σ} := RandSpec rand_allocate_tape #tb @ E {{{ (v:val), RET v; ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ rand_tapes_frag (L:=L) γ2 α (tb, []) - }}}; - rand_tape_spec_some {L: randG Σ} N E γ2 (Q: nat ->list nat -> iProp Σ) (α:loc) (tb:nat) : + }}}; + rand_tape_spec_some {L: randG Σ} N E γ2 (α:loc) (tb:nat) n ns: ↑N⊆E -> {{{ is_rand (L:=L) N γ2 ∗ - (|={E∖↑N, ∅}=> ∃ n ns, rand_tapes_frag (L:=L) γ2 α (tb, n::ns) ∗ - (rand_tapes_frag (L:=L) γ2 α (tb, ns) ={∅, E∖↑N}=∗ Q n ns)) + rand_tapes_frag (L:=L) γ2 α (tb, n::ns) }}} rand_tape #lbl:α #tb @ E - {{{ (n:nat), RET #n; ∃ ns, Q n ns}}}; - rand_presample_spec {L: randG Σ} N E T γ2 : + {{{ RET #n; rand_tapes_frag (L:=L) γ2 α (tb, ns) }}}; + rand_presample_spec {L: randG Σ} N E γ2 α (tb:nat) ns T: ↑N ⊆ E -> is_rand (L:=L) N γ2 -∗ + rand_tapes_frag (L:=L) γ2 α (tb, ns) -∗ (|={E∖↑N, ∅}=> - ∃ ε num tb (ε2 : list (fin (S tb)) -> R) α ns, - ↯ ε ∗ rand_tapes_frag (L:=L) γ2 α (tb, ns) ∗ + ∃ ε num (ε2 : list (fin (S tb)) -> R), + ↯ ε ∗ ⌜(∀ l, length l = num -> 0<= ε2 l)%R⌝ ∗ ⌜(SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R⌝ ∗ - (∀ ns', ↯ (ε2 ns') ∗ rand_tapes_frag (L:=L) γ2 α (tb, ns ++ (fin_to_nat <$> ns')) - ={∅, E∖↑N}=∗ T ε num tb ε2 α ns ns')) -∗ - wp_update E (∃ ε num tb ε2 α ns ns', T ε num tb ε2 α ns ns') + (∀ (ns':list (fin (S tb))), ↯ (ε2 ns') ={∅, E∖↑N}=∗ T ε num ε2 ns')) -∗ + wp_update E (∃ ε num ε2 ns', T ε num ε2 ns' ∗ + rand_tapes_frag (L:=L) γ2 α (tb, ns ++ (fin_to_nat <$> ns'))) }. Section impl. @@ -145,85 +145,81 @@ Section impl. Qed. Next Obligation. simpl. - iIntros (?????????? Φ) "(#Hinv & Hvs) HΦ". + iIntros (??????????? Φ) "(#Hinv & [Hvs %]) HΦ". wp_pures. - wp_bind (rand(_) _)%E. iInv "Hinv" as ">(%&H3&H4)" "Hclose". - iMod ("Hvs") as "(%n & %ns & [Hfrag %] & Hrest)". iDestruct (hocap_tapes_agree with "[$][$]") as "%". iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. simpl. wp_apply (wp_rand_tape with "[$]") as "[Htape %]". iMod (hocap_tapes_update with "[$][$]") as "[? Hfrag]". - iMod ("Hrest" with "[$Hfrag]") as "HQ". - - iPureIntro. by eapply Forall_inv_tail. - - iModIntro. - iMod ("Hclose" with "[- HQ HΦ]") as "_". - { iExists (<[α:=_]> m). - iFrame. - iApply "H3". by iNext. + iMod ("Hclose" with "[- Hfrag HΦ]") as "_". + { iExists (<[α:=_]> m). + iFrame. + iApply "H3". by iNext. } - iApply "HΦ". by iFrame. + iApply "HΦ". iFrame. + iPureIntro. by eapply Forall_inv_tail. Qed. Next Obligation. simpl. - iIntros (????????) "#Hinv Hvs". - iApply wp_update_state_step_coupl. - iIntros (σ1 ε_supply) "((Hheap&Htapes)&Hε)". - iMod (inv_acc with "Hinv") as "[>(% & H3 & H4 ) Hclose]"; [done|]. - iMod "Hvs" as "(%err & %num & %tb & %ε2 & %α & %ns & Herr & [Hfrag %] & %Hpos & %Hineq & Hrest)". - iDestruct (hocap_tapes_agree with "[$][$]") as "%". - iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. - iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". - iDestruct (ec_supply_bound with "[$][$]") as "%". - iMod (ec_supply_decrease with "[$][$]") as (ε'' ε_rem -> Heq') "Hε_supply". - iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". - (* unshelve epose (ε_rem := mknonnegreal (ε_supply-ε) _). *) - (* { simpl in *. lra. } *) - (* replace (ε_supply) with (ε_rem + ε)%NNR; last first. *) - (* { rewrite /ε_rem. apply nnreal_ext. simpl. lra. } *) - iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. - unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 x else 1)%R _). - { case_match; last (simpl;lra). apply Hpos. by rewrite -Nat.eqb_eq. } - iSplit. - { iPureIntro. - simpl. rewrite Heq'. etrans; last exact. - rewrite (Rdiv_def (SeriesC _)(S tb ^num)%R) -SeriesC_scal_r. - apply SeriesC_le; last apply ex_seriesC_scal_r, ex_seriesC_list. - intros. rewrite elem_of_enum_uniform_fin_list'. - case_match; last lra. - split; last lra. - apply Rmult_le_pos; last (apply Hpos; by rewrite -Nat.eqb_eq). - rewrite -pow_INR. apply Rdiv_INR_ge_0. - } - iIntros (sample) "<-". - destruct (Rlt_decision (nonneg ε_rem + (ε2 sample)%R) 1%R) as [Hdec|Hdec]; last first. - { apply Rnot_lt_ge, Rge_le in Hdec. - iApply state_step_coupl_ret_err_ge_1. - simpl. simpl in *. rewrite Nat.eqb_refl. lra. - } - iApply state_step_coupl_ret. - simpl. simpl in *. - rewrite -Heq. - iMod (tapeN_update_append' _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". - iMod "Hclose'" as "_". - unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 sample) _) with "[$]") as "[Hsupply Herr]". - { naive_solver. } - { simpl; lra. } - iMod (hocap_tapes_update with "[$][$]") as "[Hauth Hfrag]". - iMod ("Hrest" $! sample with "[$Herr $Hfrag]") as "HT". - { iPureIntro. - rewrite Forall_app; split; subst; first done. - eapply Forall_impl; first apply fin.fin_forall_leq. - simpl; intros; lia. - } - iFrame. - iMod ("Hclose" with "[-Hsupply]") as "_". - { iNext. iFrame. by iApply "H3". } - iApply fupd_mask_intro_subseteq; first set_solver. - iApply ec_supply_eq; last done. - simpl. rewrite Nat.eqb_refl. lra. -Qed. + iIntros (???????????) "#Hinv [Hfrag %] Hvs". + iApply wp_update_state_step_coupl. + iIntros (σ1 ε_supply) "((Hheap&Htapes)&Hε)". + iMod (inv_acc with "Hinv") as "[>(% & H3 & H4 ) Hclose]"; [done|]. + iMod "Hvs" as "(%err & %num & %ε2 & Herr & %Hpos & %Hineq & Hrest)". + iDestruct (hocap_tapes_agree with "[$][$]") as "%". + iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. + iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". + iDestruct (ec_supply_bound with "[$][$]") as "%". + iMod (ec_supply_decrease with "[$][$]") as (ε'' ε_rem -> Heq') "Hε_supply". + iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". + (* unshelve epose (ε_rem := mknonnegreal (ε_supply-ε) _). *) + (* { simpl in *. lra. } *) + (* replace (ε_supply) with (ε_rem + ε)%NNR; last first. *) + (* { rewrite /ε_rem. apply nnreal_ext. simpl. lra. } *) + iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. + unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 x else 1)%R _). + { case_match; last (simpl;lra). apply Hpos. by rewrite -Nat.eqb_eq. } + iSplit. + { iPureIntro. + simpl. rewrite Heq'. etrans; last exact. + rewrite (Rdiv_def (SeriesC _)(S tb ^num)%R) -SeriesC_scal_r. + apply SeriesC_le; last apply ex_seriesC_scal_r, ex_seriesC_list. + intros. rewrite elem_of_enum_uniform_fin_list'. + case_match; last lra. + split; last lra. + apply Rmult_le_pos; last (apply Hpos; by rewrite -Nat.eqb_eq). + rewrite -pow_INR. apply Rdiv_INR_ge_0. + } + iIntros (sample) "<-". + destruct (Rlt_decision (nonneg ε_rem + (ε2 sample)%R) 1%R) as [Hdec|Hdec]; last first. + { apply Rnot_lt_ge, Rge_le in Hdec. + iApply state_step_coupl_ret_err_ge_1. + simpl. simpl in *. rewrite Nat.eqb_refl. lra. + } + iApply state_step_coupl_ret. + simpl. simpl in *. + rewrite -Heq. + iMod (tapeN_update_append' _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". + iMod "Hclose'" as "_". + unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 sample) _) with "[$]") as "[Hsupply Herr]". + { naive_solver. } + { simpl; lra. } + iMod (hocap_tapes_update with "[$][$]") as "[Hauth Hfrag]". + iMod ("Hrest" $! sample with "[$Herr]") as "HT". + iFrame. + iMod ("Hclose" with "[-Hsupply]") as "_". + { iNext. iFrame. by iApply "H3". } + iApply fupd_mask_intro_subseteq; first set_solver. + iSplit. + - iApply ec_supply_eq; last done. + simpl. rewrite Nat.eqb_refl. lra. + - iPureIntro. + rewrite Forall_app; split; subst; first done. + eapply Forall_impl; first apply fin.fin_forall_leq. + simpl; intros; lia. + Qed. End impl. @@ -237,19 +233,14 @@ Section checks. {{{ RET #(LitInt n); rand_tapes_frag (L:=L) γ1 α (N, ns) ∗ ⌜n <= N⌝ }}}. Proof. iIntros (Hsubset Φ) "(#Hinv &>Hfrag) HΦ". - wp_apply (rand_tape_spec_some _ _ _ (λ n' ns', ⌜n=n'/\ns=ns'⌝ ∗ ⌜n<=N⌝ ∗ rand_tapes_frag _ _ _)%I with "[Hfrag]"); first exact. - - iSplit; first done. - iDestruct (rand_tapes_valid with "[$]") as "%H'". - iFrame. - iApply fupd_mask_intro; first set_solver. - iIntros "Hclose Hfrag". + iDestruct (rand_tapes_valid with "[$]") as "%H'". + wp_apply (rand_tape_spec_some with "[Hfrag]"); first exact. + - by iFrame. + - iIntros. + iApply "HΦ". iFrame. - iMod "Hclose". - iModIntro. iPureIntro. rewrite Forall_cons in H'. naive_solver. - - iIntros (?) "(%&[-> ->]&%&?)". iApply "HΦ". - by iFrame. Qed. Local Opaque enum_uniform_fin_list. @@ -265,17 +256,17 @@ Section checks. iIntros (Hsubset Hpos Hineq) "#Hinv >Htape Herr". iDestruct (ec_valid with "[$]") as "%Hpos'". destruct Hpos' as [Hpos' ?]. - iMod (rand_presample_spec _ _ - (λ ε' (num':nat) tb' ε2' α' ns1 ns2, - ⌜ε1=ε'⌝ ∗ ⌜(1=num')%nat⌝ ∗ ⌜N=tb'⌝ ∗ + iMod (rand_presample_spec _ _ _ _ _ _ + (λ ε' (num':nat) ε2' ns2, + ⌜ε1=ε'⌝ ∗ ⌜(1=num')%nat⌝ ∗ ⌜∀ x y, fin_to_nat <$> x = fin_to_nat <$> y -> - ε2' x = (λ l, match l with |[x] => ε2 x | _ => 1%R end) y⌝∗ ⌜α=α'⌝ ∗ ⌜ns1=ns⌝ ∗ - ∃ (n:fin (S N)), ⌜fin_to_nat <$> ns2=[fin_to_nat n]⌝ ∗ ↯ (ε2 n) ∗ rand_tapes_frag γ1 α (N, ns ++ [fin_to_nat n]))%I - with "[//][- ]") as "(%&%&%&%&%&%&%&->&<-&->&%&->&->&%&%&?&?)". + ε2' x = (λ l, match l with |[x] => ε2 x | _ => 1%R end) y⌝ ∗ + ∃ (n:fin (S N)), ⌜fin_to_nat <$> ns2=[fin_to_nat n]⌝ ∗ ↯ (ε2 n))%I + with "[//][$ ][-]") as "H". - done. - - iFrame. - iApply fupd_mask_intro; first set_solver. + - iApply fupd_mask_intro; first set_solver. iIntros "Hclose". + iFrame. iExists 1, (λ l, match l with |[x] => ε2 x | _ => 1%R end). simpl. repeat iSplit. + iPureIntro. by intros [|?[|]]. @@ -298,7 +289,7 @@ Section checks. apply Rdiv_INR_ge_0. * intros. repeat case_match; by simplify_eq. * apply ex_seriesC_finite. - + iIntros ([|?[|]]) "(?&?)"; [by iDestruct (ec_contradict with "[$]") as "%"| |by iDestruct (ec_contradict with "[$]") as "%"]. + + iIntros ([|?[|]]) "?"; [by iDestruct (ec_contradict with "[$]") as "%"| |by iDestruct (ec_contradict with "[$]") as "%"]. iFrame. iMod "Hclose". iModIntro. @@ -307,7 +298,9 @@ Section checks. intros x y H'. apply list_fmap_eq_inj in H'; first by simplify_eq. apply fin_to_nat_inj. - - iModIntro. iFrame. + - iModIntro. + iDestruct "H" as "(%&%&%&%&[(%&%&%&%&->&?) ?])". + iFrame. Qed. End checks. From b2857832e849ec94689fbe185f5a0a0ce636ff90 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Mon, 30 Sep 2024 10:11:22 +0200 Subject: [PATCH 50/69] random counter spec fix --- .../examples/random_counter/random_counter.v | 92 ++++++++----------- 1 file changed, 38 insertions(+), 54 deletions(-) diff --git a/theories/coneris/examples/random_counter/random_counter.v b/theories/coneris/examples/random_counter/random_counter.v index d7885d8f..b6e22366 100644 --- a/theories/coneris/examples/random_counter/random_counter.v +++ b/theories/coneris/examples/random_counter/random_counter.v @@ -74,31 +74,30 @@ Class random_counter `{!conerisGS Σ} := RandCounter {{{ (v:val), RET v; ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ counter_tapes_frag (L:=L) γ1 α [] }}}; - incr_counter_tape_spec_some {L: counterG Σ} N E c γ1 γ2 (Q:nat->nat->list nat->iProp Σ) (α:loc) : + incr_counter_tape_spec_some {L: counterG Σ} N E c γ1 γ2 (Q:nat->iProp Σ) (α:loc) n ns: ↑N⊆E -> {{{ is_counter (L:=L) N c γ1 γ2 ∗ - ( |={E∖ ↑N, ∅}=> - ∃ n ns, counter_tapes_frag (L:=L) γ1 α (n::ns) ∗ - (counter_tapes_frag (L:=L) γ1 α ns ={∅, E∖↑N}=∗ - ∀ (z:nat), counter_content_auth (L:=L) γ2 z - ={E∖↑N}=∗ - counter_content_auth (L:=L) γ2 (z+n) ∗ Q z n ns) - ) + counter_tapes_frag (L:=L) γ1 α (n::ns) ∗ + ( ∀ (z:nat), counter_content_auth (L:=L) γ2 z ={E∖↑N}=∗ + counter_content_auth (L:=L) γ2 (z+n) ∗ Q z) + }}} incr_counter_tape c #lbl:α @ E - {{{ (z n:nat), RET (#z, #n); ∃ ns, Q z n ns}}}; - counter_presample_spec {L: counterG Σ} NS E T γ1 γ2 c: + {{{ (z:nat), RET (#z, #n); counter_tapes_frag (L:=L) γ1 α ns ∗ + Q z }}}; + counter_presample_spec {L: counterG Σ} NS E T γ1 γ2 c α ns: ↑NS ⊆ E -> is_counter (L:=L) NS c γ1 γ2 -∗ + counter_tapes_frag (L:=L) γ1 α ns -∗ ( |={E∖↑NS,∅}=> - ∃ ε α ε2 num ns, - ↯ ε ∗ counter_tapes_frag (L:=L) γ1 α ns ∗ + ∃ ε ε2 num, + ↯ ε ∗ ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ ⌜(SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R⌝ ∗ - (∀ ns', ↯ (ε2 ns') ∗ counter_tapes_frag (L:=L) γ1 α (ns++ns')={∅,E∖↑NS}=∗ - T ε α ε2 num ns ns' + (∀ ns', ↯ (ε2 ns') ={∅,E∖↑NS}=∗ + T ε ε2 num ns' ))-∗ - wp_update E (∃ ε α ε2 num ns ns', T ε α ε2 num ns ns'); + wp_update E (∃ ε ε2 num ns', counter_tapes_frag (L:=L) γ1 α (ns++ns') ∗ T ε ε2 num ns'); read_counter_spec {L: counterG Σ} N E c γ1 γ2 Q: ↑N ⊆ E -> {{{ is_counter (L:=L) N c γ1 γ2 ∗ @@ -118,39 +117,34 @@ Section lemmas. Lemma incr_counter_tape_spec_none N E c γ1 γ2 Q α: ↑N ⊆ E-> {{{ is_counter (L:=L) N c γ1 γ2 ∗ + counter_tapes_frag (L:=L) γ1 α [] ∗ ( |={E∖↑N, ∅}=> ∃ ε ε2, - ↯ ε ∗ counter_tapes_frag (L:=L) γ1 α [] ∗ + ↯ ε ∗ ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ ⌜(((ε2 0%nat) + (ε2 1%nat)+ (ε2 2%nat)+ (ε2 3%nat))/4 <= ε)%R⌝ ∗ - (∀ n, ↯ (ε2 n) ∗ counter_tapes_frag (L:=L) γ1 α [n] ={∅, E∖↑N}=∗ - |={E∖↑N, ∅}=> - counter_tapes_frag (L:=L) γ1 α [n] ∗ - (counter_tapes_frag (L:=L) γ1 α [] ={∅,E∖↑N}=∗ + (∀ n, ↯ (ε2 n) ={∅, E∖↑N}=∗ ∀ (z:nat), counter_content_auth (L:=L) γ2 z ={E∖↑N}=∗ counter_content_auth (L:=L) γ2 (z+n) ∗ Q z n ε ε2 ) ) - ) + }}} incr_counter_tape c #lbl:α @ E - {{{ (z n:nat), RET (#z, #n); ∃ ε ε2, Q z n ε ε2 }}}. + {{{ (z n:nat), RET (#z, #n); counter_tapes_frag (L:=L) γ1 α [] ∗ + ∃ ε ε2, Q z n ε ε2 }}}. Proof. - iIntros (Hsubset Φ) "(#Hinv & Hvs) HΦ". + iIntros (Hsubset Φ) "(#Hinv & Hfrag & Hvs) HΦ". iMod (counter_presample_spec _ _ - (λ ε α' ε2 num ns ns', + (λ ε ε2 num ns', ∃ n ε2', - ⌜num=1%nat⌝ ∗ ⌜α'=α⌝ ∗ ⌜ns=[]⌝ ∗ ⌜ns'=[n]⌝ ∗ + ⌜num=1%nat⌝ ∗ ⌜ns'=[n]⌝ ∗ ⌜(ε2' = λ x, ε2 [x])%R⌝ ∗ - ( |={E∖ ↑N, ∅}=> - counter_tapes_frag (L:=L) γ1 α [n] ∗ - (counter_tapes_frag (L:=L) γ1 α [] ={∅, E∖↑N}=∗ - ∀ (z:nat), counter_content_auth (L:=L) γ2 z + ( ∀ (z:nat), counter_content_auth (L:=L) γ2 z ={E∖↑N}=∗ counter_content_auth (L:=L) γ2 (z+n) ∗ Q z n ε ε2') - ) - )%I with "[//][-HΦ]") as "(%ε & % & %ε2 & %num & %ns & %ns' & %n & %ε2' & -> & -> & -> & -> & -> & Hrest)"; first done. - - iMod "Hvs" as "(%ε & %ε2 & Herr & Hfrag & %Hpos & %Hineq & Hrest)". + ) %I with "[//][$][-HΦ]") as "(%ε & %ε2 & %num & %ns' & Hfrag & %n & %ε2' & -> & -> & -> & Hrest)"; first done. + - iMod "Hvs" as "(%ε & %ε2 & Herr & %Hpos & %Hineq & Hrest)". iFrame. iModIntro. iExists (λ l, match l with |[x] => ε2 x | _ => 1 end)%R, 1%nat. repeat iSplit. @@ -161,22 +155,17 @@ Section lemmas. simpl. lra. * apply NoDup_fmap_2; last apply NoDup_enum_uniform_fin_list. apply list_fmap_eq_inj, fin_to_nat_inj. - + iIntros (ns') "[Herr Hfrag]". destruct (ns') as [|n[|??]]; [by iDestruct (ec_contradict with "[$]") as "%"| |by iDestruct (ec_contradict with "[$]") as "%"]. + + iIntros (ns') "Herr". destruct (ns') as [|n[|??]]; [by iDestruct (ec_contradict with "[$]") as "%"| |by iDestruct (ec_contradict with "[$]") as "%"]. iMod ("Hrest" $! n with "[$]") as "?". iModIntro. by iFrame. - - wp_apply (incr_counter_tape_spec_some _ _ _ _ _ (λ z n ns, ∃ ε ε2, Q z n ε ε2)%I with "[Hrest]"); first exact. + - wp_apply (incr_counter_tape_spec_some _ _ _ _ _ (λ z, ∃ ε ε2, Q z n ε ε2)%I with "[$Hfrag Hrest]"); first exact. + iSplit; first done. - iMod "Hrest" as "[? Hrest]". - iFrame. + iIntros (?) "?". + iMod ("Hrest" with "[$]") as "[$ ?]". iModIntro. - iIntros. - iMod ("Hrest" with "[$]") as "Hrest". - iModIntro. - iIntros (z) "Hauth". - iMod ("Hrest" with "[$]") as "[??]". - by iFrame. - + iIntros (??) "(%&%&%&HQ)". + iFrame. + + iIntros (?) "[?(%&%&?)]". iApply "HΦ". iFrame. Qed. @@ -202,12 +191,13 @@ Section lemmas. pose (ε2' := (λ x, if (x<=?3)%nat then ε2 x else 1%R)). wp_apply (incr_counter_tape_spec_none _ _ _ _ _ (λ z' n ε' ε2'', ⌜0<=n<4⌝ ∗ ⌜z<=z'⌝ ∗ ⌜ε=ε'⌝ ∗ ⌜ε2''=ε2'⌝ ∗ - ↯ (ε2' n) ∗ counter_tapes_frag (L:=L) γ1 α [] ∗ + ↯ (ε2' n) ∗ counter_content_frag (L:=L) γ2 q (z+n)%nat )%I with "[-HΦ]"). - done. - iSplit; first done. + iFrame. iApply fupd_mask_intro; first set_solver. iIntros "Hclose". iFrame. @@ -215,25 +205,19 @@ Section lemmas. repeat iSplit. + iPureIntro. rewrite /ε2'. intros; case_match; [naive_solver|lra]. + done. - + iIntros (n) "[Herr Hfrag]". - iMod "Hclose" as "_". + + iIntros (n) "Hfrag". + iMod "Hclose". iModIntro. - iApply fupd_mask_intro; first set_solver. - iIntros "Hclose". + iIntros (?) "?". rewrite /ε2'. case_match eqn:H; last by (iDestruct (ec_contradict with "[$]") as "%"). iFrame. - iIntros "Hfrag'". - iMod "Hclose" as "_". - iFrame. apply leb_complete in H. - iModIntro. - iIntros (z') "Hauth". iDestruct (counter_content_less_than with "[$][$]") as "%". iMod (counter_content_update with "[$][$]") as "[??]". iFrame. iPureIntro; repeat split; try done; lia. - - iIntros (z' n) "(%ε' & %ε2'' & % & % &-> & -> & Herr & Hfrag & Hfrag')". + - iIntros (z' n) "[Hfrag (%ε' & %ε2'' & % & % &-> & -> & Herr & Hfrag' )]". iApply "HΦ". iFrame. iSplit; first done. From ea850778af177ddcefc1113e7f4af12d1f654853 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Mon, 30 Sep 2024 10:54:19 +0200 Subject: [PATCH 51/69] impl1 --- .../coneris/examples/random_counter/impl1.v | 555 +++++++++--------- theories/coneris/lib/hocap_flip.v | 5 +- theories/coneris/lib/hocap_rand.v | 6 +- 3 files changed, 276 insertions(+), 290 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl1.v b/theories/coneris/examples/random_counter/impl1.v index 22519878..e3945c42 100644 --- a/theories/coneris/examples/random_counter/impl1.v +++ b/theories/coneris/examples/random_counter/impl1.v @@ -1,298 +1,285 @@ -(* From iris.algebra Require Import frac_auth. *) -(* From iris.base_logic.lib Require Import invariants. *) -(* From clutch.coneris Require Import coneris hocap hocap_rand random_counter. *) +From iris.algebra Require Import frac_auth. +From iris.base_logic.lib Require Import invariants. +From clutch.coneris Require Import coneris hocap hocap_rand random_counter. -(* Set Default Proof Using "Type*". *) +Set Default Proof Using "Type*". -(* Section impl1. *) -(* Context `{H:conerisGS Σ, r1:@rand_spec Σ H, L:randG Σ, !inG Σ (frac_authR natR)}. *) +Section impl1. + Context `{H:conerisGS Σ, r1:@rand_spec Σ H, L:randG Σ, !inG Σ (frac_authR natR)}. -(* Definition new_counter1 : val:= λ: "_", ref #0. *) -(* Definition allocate_tape1 : val := λ: "_", rand_allocate_tape #3%nat. *) -(* Definition incr_counter_tape1 :val := λ: "l" "α", let: "n" := rand_tape "α" #3%nat in (FAA "l" "n", "n"). *) -(* Definition read_counter1 : val := λ: "l", !"l". *) -(* Class counterG1 Σ := CounterG1 { counterG1_randG : randG Σ; *) -(* counterG1_frac_authR:: inG Σ (frac_authR natR) }. *) + Definition new_counter1 : val:= λ: "_", ref #0. + Definition allocate_tape1 : val := λ: "_", rand_allocate_tape #3%nat. + Definition incr_counter_tape1 :val := λ: "l" "α", let: "n" := rand_tape "α" #3%nat in (FAA "l" "n", "n"). + Definition read_counter1 : val := λ: "l", !"l". + Class counterG1 Σ := CounterG1 { counterG1_randG : randG Σ; + counterG1_frac_authR:: inG Σ (frac_authR natR) }. -(* Definition counter_inv_pred1 c γ2 := (∃ (l:loc) (z:nat), ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ2 (●F z) )%I. *) -(* Definition is_counter1 N (c:val) γ1 γ2:= *) -(* (is_rand (L:=L) (N.@"rand") γ1 ∗ *) -(* inv (N.@"counter") (counter_inv_pred1 c γ2) *) -(* )%I. *) + Definition counter_inv_pred1 c γ2 := (∃ (l:loc) (z:nat), ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ2 (●F z) )%I. + Definition is_counter1 N (c:val) γ1 γ2:= + (is_rand (L:=L) (N.@"rand") γ1 ∗ + inv (N.@"counter") (counter_inv_pred1 c γ2) + )%I. -(* Lemma new_counter_spec1 E N: *) -(* {{{ True }}} *) -(* new_counter1 #() @ E *) -(* {{{ (c:val), RET c; *) -(* ∃ γ1 γ2, is_counter1 N c γ1 γ2 ∗ own γ2 (◯F 0%nat) *) -(* }}}. *) -(* Proof. *) -(* rewrite /new_counter1. *) -(* iIntros (Φ) "_ HΦ". *) -(* wp_pures. *) -(* iMod (rand_inv_create_spec) as "(%γ1 & #Hinv)". *) -(* wp_alloc l as "Hl". *) -(* iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ2[H5 H6]]". *) -(* { by apply frac_auth_valid. } *) -(* replace (#0) with (#0%nat) by done. *) -(* iMod (inv_alloc _ _ (counter_inv_pred1 (#l) γ2) with "[$Hl $H5]") as "#Hinv'"; first done. *) -(* iApply "HΦ". *) -(* iFrame. *) -(* iModIntro. *) -(* iExists _. by iSplit. *) -(* Qed. *) + Lemma new_counter_spec1 E N: + {{{ True }}} + new_counter1 #() @ E + {{{ (c:val), RET c; + ∃ γ1 γ2, is_counter1 N c γ1 γ2 ∗ own γ2 (◯F 0%nat) + }}}. + Proof. + rewrite /new_counter1. + iIntros (Φ) "_ HΦ". + wp_pures. + iMod (rand_inv_create_spec) as "(%γ1 & #Hinv)". + wp_alloc l as "Hl". + iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ2[H5 H6]]". + { by apply frac_auth_valid. } + replace (#0) with (#0%nat) by done. + iMod (inv_alloc _ _ (counter_inv_pred1 (#l) γ2) with "[$Hl $H5]") as "#Hinv'"; first done. + iApply "HΦ". + iFrame. + iModIntro. + iExists _. by iSplit. + Qed. -(* (* (** Not used in instantiating type class*) *) *) -(* (* Lemma incr_counter_spec1 E N c γ1 γ2 γ3 (ε2:R -> nat -> R) (P: iProp Σ) (T: nat -> iProp Σ) (Q: nat->nat->iProp Σ): *) *) -(* (* ↑N ⊆ E-> *) *) -(* (* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) *) -(* (* (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → *) *) -(* (* {{{ inv N (counter_inv_pred1 c γ1 γ2 γ3) ∗ *) *) -(* (* □(∀ (ε:R) (n : nat), P ∗ ●↯ ε @ γ1 ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨●↯ (ε2 ε n) @ γ1 ∗ T n) ) ∗ *) *) -(* (* □ (∀ (n z:nat), T n ∗ own γ3 (●F z) ={E∖↑N}=∗ *) *) -(* (* own γ3 (●F(z+n)%nat)∗ Q z n) ∗ *) *) -(* (* P *) *) -(* (* }}} *) *) -(* (* incr_counter1 c @ E *) *) -(* (* {{{ (n:nat) (z:nat), RET (#z, #n); Q z n }}}. *) *) -(* (* Proof. *) *) -(* (* iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP) HΦ". *) *) -(* (* rewrite /incr_counter1. *) *) -(* (* wp_pures. *) *) -(* (* wp_bind (rand _)%E. *) *) -(* (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) *) -(* (* iDestruct (ec_valid with "[$]") as "[%K1 %K2]". *) *) -(* (* wp_apply (wp_couple_rand_adv_comp1' _ _ _ _ (λ x, ε2 ε (fin_to_nat x)) with "[$]"). *) *) -(* (* { intros. naive_solver. } *) *) -(* (* { rewrite SeriesC_finite_foldr; specialize (Hineq ε K1). simpl; lra. } *) *) -(* (* iIntros (n) "H1". *) *) -(* (* iMod ("Hvs1" with "[$]") as "[%|[H2 HT]]". *) *) -(* (* { iExFalso. iApply ec_contradict; last done. done. } *) *) -(* (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) *) -(* (* iModIntro. wp_pures. *) *) -(* (* clear -Hsubset. *) *) -(* (* wp_bind (FAA _ _). *) *) -(* (* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) *) -(* (* wp_faa. *) *) -(* (* iMod ("Hvs2" with "[$]") as "[H6 HQ]". *) *) -(* (* replace (#(z+n)) with (#(z+n)%nat); last first. *) *) -(* (* { by rewrite Nat2Z.inj_add. } *) *) -(* (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) *) -(* (* iModIntro. *) *) -(* (* wp_pures. *) *) -(* (* by iApply "HΦ". *) *) -(* (* Qed. *) *) + (* (** Not used in instantiating type class*) *) + (* Lemma incr_counter_spec1 E N c γ1 γ2 γ3 (ε2:R -> nat -> R) (P: iProp Σ) (T: nat -> iProp Σ) (Q: nat->nat->iProp Σ): *) + (* ↑N ⊆ E-> *) + (* (∀ ε n, 0<= ε -> 0<= ε2 ε n)%R-> *) + (* (∀ (ε:R), 0<=ε -> ((ε2 ε 0%nat) + (ε2 ε 1%nat)+ (ε2 ε 2%nat)+ (ε2 ε 3%nat))/4 <= ε)%R → *) + (* {{{ inv N (counter_inv_pred1 c γ1 γ2 γ3) ∗ *) + (* □(∀ (ε:R) (n : nat), P ∗ ●↯ ε @ γ1 ={E∖↑N}=∗ (⌜(1<=ε2 ε n)%R⌝∨●↯ (ε2 ε n) @ γ1 ∗ T n) ) ∗ *) + (* □ (∀ (n z:nat), T n ∗ own γ3 (●F z) ={E∖↑N}=∗ *) + (* own γ3 (●F(z+n)%nat)∗ Q z n) ∗ *) + (* P *) + (* }}} *) + (* incr_counter1 c @ E *) + (* {{{ (n:nat) (z:nat), RET (#z, #n); Q z n }}}. *) + (* Proof. *) + (* iIntros (Hsubset Hpos Hineq Φ) "(#Hinv & #Hvs1 & #Hvs2 & HP) HΦ". *) + (* rewrite /incr_counter1. *) + (* wp_pures. *) + (* wp_bind (rand _)%E. *) + (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* iDestruct (ec_valid with "[$]") as "[%K1 %K2]". *) + (* wp_apply (wp_couple_rand_adv_comp1' _ _ _ _ (λ x, ε2 ε (fin_to_nat x)) with "[$]"). *) + (* { intros. naive_solver. } *) + (* { rewrite SeriesC_finite_foldr; specialize (Hineq ε K1). simpl; lra. } *) + (* iIntros (n) "H1". *) + (* iMod ("Hvs1" with "[$]") as "[%|[H2 HT]]". *) + (* { iExFalso. iApply ec_contradict; last done. done. } *) + (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) + (* iModIntro. wp_pures. *) + (* clear -Hsubset. *) + (* wp_bind (FAA _ _). *) + (* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* wp_faa. *) + (* iMod ("Hvs2" with "[$]") as "[H6 HQ]". *) + (* replace (#(z+n)) with (#(z+n)%nat); last first. *) + (* { by rewrite Nat2Z.inj_add. } *) + (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) + (* iModIntro. *) + (* wp_pures. *) + (* by iApply "HΦ". *) + (* Qed. *) -(* Lemma allocate_tape_spec1 N E c γ1 γ2: *) -(* ↑N ⊆ E-> *) -(* {{{ is_counter1 N c γ1 γ2 }}} *) -(* allocate_tape1 #() @ E *) -(* {{{ (v:val), RET v; *) -(* ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ rand_tapes_frag (L:=L) γ1 α (3%nat, []) *) -(* }}}. *) -(* Proof. *) -(* iIntros (Hsubset Φ) "[#Hinv #Hinv'] HΦ". *) -(* rewrite /allocate_tape1. *) -(* wp_pures. *) -(* wp_apply rand_allocate_tape_spec; [by eapply nclose_subseteq'|done..]. *) -(* Qed. *) + Lemma allocate_tape_spec1 N E c γ1 γ2: + ↑N ⊆ E-> + {{{ is_counter1 N c γ1 γ2 }}} + allocate_tape1 #() @ E + {{{ (v:val), RET v; + ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ rand_tapes_frag (L:=L) γ1 α (3%nat, []) + }}}. + Proof. + iIntros (Hsubset Φ) "[#Hinv #Hinv'] HΦ". + rewrite /allocate_tape1. + wp_pures. + wp_apply rand_allocate_tape_spec; [by eapply nclose_subseteq'|done..]. + Qed. -(* Lemma incr_counter_tape_spec_some1 N E c γ1 γ2 (Q:nat->nat -> list nat -> iProp Σ) (α:loc) : *) -(* ↑N⊆E -> *) -(* {{{ is_counter1 N c γ1 γ2 ∗ *) -(* ( |={E∖ ↑N, ∅}=> *) -(* ∃ n ns, rand_tapes_frag (L:=L) γ1 α (3, n::ns) ∗ *) -(* (rand_tapes_frag (L:=L) γ1 α (3, ns) ={∅, E∖↑N}=∗ *) -(* ∀ (z:nat), own γ2 (●F z) *) -(* ={E∖↑N}=∗ *) -(* own γ2 (●F (z+n)) ∗ Q z n ns) *) -(* ) *) -(* }}} *) -(* incr_counter_tape1 c #lbl:α @ E *) -(* {{{ (z n:nat), RET (#z, #n); ∃ ns, Q z n ns}}}. *) -(* Proof. *) -(* iIntros (Hsubset Φ) "([#Hinv #Hinv'] & Hvs) HΦ". *) -(* rewrite /incr_counter_tape1. *) -(* wp_pures. *) -(* wp_apply (rand_tape_spec_some _ _ _ *) -(* (λ n ns, ∀ (z:nat), own γ2 (●F z)={E∖↑N}=∗ *) -(* own γ2 (●F (z+n)) ∗ Q z n ns)%I with "[Hvs]"); [by eapply nclose_subseteq'|iSplit; first done|..]. *) -(* - iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose". *) -(* + apply difference_mono_l. *) -(* by apply nclose_subseteq'. *) -(* + iMod "Hvs" as "(%n & %ns & Hfrag & Hrest)". *) -(* iModIntro. *) -(* iFrame. *) -(* iIntros "Hfrag". *) -(* iMod ("Hrest" with "[$]") as "Hrest". *) -(* by iMod "Hclose" as "_". *) -(* - iIntros (n) "(%ns & Hvs)". *) -(* wp_pures. *) -(* wp_bind (FAA _ _). *) -(* iInv "Hinv'" as ">(%&%&-> & ?&?)" "Hclose". *) -(* wp_faa. *) -(* iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". *) -(* + apply difference_mono_l. *) -(* by apply nclose_subseteq'. *) -(* + iMod ("Hvs" with "[$]") as "[? HQ]". *) -(* iMod "Hclose'" as "_". *) -(* rewrite -Nat2Z.inj_add. *) -(* iMod ("Hclose" with "[-HΦ HQ]") as "_"; first by iFrame. *) -(* iModIntro. *) -(* wp_pures. *) -(* iApply "HΦ". by iFrame. *) -(* Qed. *) + Lemma incr_counter_tape_spec_some1 N E c γ1 γ2 (Q:nat->iProp Σ) (α:loc) n ns: + ↑N⊆E -> + {{{ is_counter1 N c γ1 γ2 ∗ + rand_tapes_frag (L:=L) γ1 α (3, n::ns) ∗ + ( ∀ (z:nat), own γ2 (●F z) ={E∖↑N}=∗ + own γ2 (●F (z+n)) ∗ Q z) + + }}} + incr_counter_tape1 c #lbl:α @ E + {{{ (z:nat), RET (#z, #n); rand_tapes_frag (L:=L) γ1 α (3, ns) ∗ + Q z }}}. + Proof. + iIntros (Hsubset Φ) "([#Hinv #Hinv'] & Hfrag & Hvs) HΦ". + rewrite /incr_counter_tape1. + wp_pures. + wp_apply (rand_tape_spec_some with "[$]") as "Hfrag"; first by eapply nclose_subseteq'. + wp_pures. + wp_bind (FAA _ _). + iInv "Hinv'" as ">(%&%&-> & ?&?)" "Hclose". + wp_faa. + iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". + + apply difference_mono_l. + by apply nclose_subseteq'. + + iMod ("Hvs" with "[$]") as "[? HQ]". + iMod "Hclose'" as "_". + rewrite -Nat2Z.inj_add. + iMod ("Hclose" with "[-HQ Hfrag HΦ]") as "_"; first by iFrame. + iModIntro. + wp_pures. + iApply "HΦ". by iFrame. + Qed. -(* Lemma counter_presample_spec1 NS E T γ1 γ2 c: *) -(* ↑NS ⊆ E -> *) -(* is_counter1 NS c γ1 γ2 -∗ *) -(* ( |={E∖↑NS,∅}=> *) -(* ∃ ε α ε2 num ns, *) -(* ↯ ε ∗ rand_tapes_frag (L:=L) γ1 α (3%nat, ns) ∗ *) -(* ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ *) -(* ⌜(SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R⌝ ∗ *) -(* (∀ ns', ↯ (ε2 ns') ∗ rand_tapes_frag (L:=L) γ1 α (3%nat, ns++ns')={∅,E∖↑NS}=∗ *) -(* T ε α ε2 num ns ns' *) -(* ))-∗ *) -(* wp_update E (∃ ε α ε2 num ns ns', T ε α ε2 num ns ns'). *) -(* Proof. *) -(* iIntros (Hsubset) "[#Hinv #Hinv'] Hvs". *) -(* iMod (rand_presample_spec _ _ (λ ε num tb ε2 α ns ns', *) -(* ∃ ε2', ⌜∀ xs ys, fin_to_nat <$> xs = ys -> ε2 xs = ε2' ys⌝ ∗ *) -(* ⌜tb = 3⌝ ∗ T ε α ε2' num ns (fin_to_nat <$> ns'))%I with "[//][Hvs]") as "H"; first by apply nclose_subseteq'. *) -(* + iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose". *) -(* { apply difference_mono_l. *) -(* by apply nclose_subseteq'. } *) -(* iMod "Hvs" as "(%ε & %α & %ε2 & %num & %ns & Herr & Hfrag & %Hpos & %Hineq & Hrest)". *) -(* iFrame. *) -(* iModIntro. *) -(* iExists num, (λ ls, ε2 (fin_to_nat <$> ls)). *) -(* repeat iSplit. *) -(* * done. *) -(* * iPureIntro. *) -(* etrans; last exact. *) -(* apply Req_le. *) -(* replace (INR 4) with 4%R; last (simpl; lra). *) -(* f_equal. *) -(* rewrite !SeriesC_list. *) -(* -- by rewrite foldr_fmap. *) -(* -- apply NoDup_fmap. *) -(* ++ apply list_fmap_eq_inj. *) -(* apply fin_to_nat_inj. *) -(* ++ apply NoDup_enum_uniform_fin_list. *) -(* -- apply NoDup_enum_uniform_fin_list. *) -(* * iIntros (ns') "[Herr Hfrag]". *) -(* iMod ("Hrest" with "[$]") as "HT". *) -(* iMod "Hclose" as "_". *) -(* iModIntro. *) -(* iFrame. *) -(* repeat iSplit; iPureIntro; last done. *) -(* by intros ??<-. *) -(* + iDestruct "H" as "(%ε & %num & %tb & %ε2 & %α & %ns & %ns' & %ε2' & % & -> & HT)". *) -(* iModIntro. iFrame. *) -(* Qed. *) + Lemma counter_presample_spec1 NS E T γ1 γ2 c α ns: + ↑NS ⊆ E -> + is_counter1 NS c γ1 γ2 -∗ + rand_tapes_frag (L:=L) γ1 α (3%nat, ns) -∗ + ( |={E∖↑NS,∅}=> + ∃ ε ε2 num, + ↯ ε ∗ + ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ + ⌜(SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R⌝ ∗ + (∀ ns', ↯ (ε2 ns') ={∅,E∖↑NS}=∗ + T ε ε2 num ns' + ))-∗ + wp_update E (∃ ε ε2 num ns', rand_tapes_frag (L:=L) γ1 α (3%nat, ns++ns') ∗ T ε ε2 num ns'). + Proof. + iIntros (Hsubset) "[#Hinv #Hinv'] Hfrag Hvs". + iMod (rand_presample_spec _ _ _ _ _ _ (λ ε num ε2 ns', + ∃ ε2', ⌜∀ xs ys, fin_to_nat <$> xs = ys -> ε2 xs = ε2' ys⌝ ∗ + T ε ε2' num (fin_to_nat <$> ns'))%I with "[//][$][Hvs]") as "H"; first by apply nclose_subseteq'. + - iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose". + { apply difference_mono_l. + by apply nclose_subseteq'. } + iMod "Hvs" as "(%ε & %ε2 & %num & Herr & %Hpos & %Hineq & Hrest)". + iFrame. + iModIntro. + iExists num, (λ ls, ε2 (fin_to_nat <$> ls)). + repeat iSplit. + + done. + + iPureIntro. + etrans; last exact. + apply Req_le. + replace (INR 4) with 4%R; last (simpl; lra). + f_equal. + rewrite !SeriesC_list. + * by rewrite foldr_fmap. + * apply NoDup_fmap. + -- apply list_fmap_eq_inj. + apply fin_to_nat_inj. + -- apply NoDup_enum_uniform_fin_list. + * apply NoDup_enum_uniform_fin_list. + + iIntros (ns') "Herr". + iMod ("Hrest" with "[$]") as "HT". + iMod "Hclose" as "_". + iModIntro. + iFrame. + iPureIntro. + by intros ??<-. + - iDestruct "H" as "(%ε & %num & %ε2 & %ns' & Hfrag & %ε2' & % & HT)". + iModIntro. iFrame. + Qed. -(* Lemma read_counter_spec1 N E c γ1 γ2 Q: *) -(* ↑N ⊆ E -> *) -(* {{{ is_counter1 N c γ1 γ2 ∗ *) -(* (∀ (z:nat), own γ2 (●F z) ={E∖↑N}=∗ *) -(* own γ2 (●F z) ∗ Q z) *) + Lemma read_counter_spec1 N E c γ1 γ2 Q: + ↑N ⊆ E -> + {{{ is_counter1 N c γ1 γ2 ∗ + (∀ (z:nat), own γ2 (●F z) ={E∖↑N}=∗ + own γ2 (●F z) ∗ Q z) -(* }}} *) -(* read_counter1 c @ E *) -(* {{{ (n':nat), RET #n'; Q n' *) -(* }}}. *) -(* Proof. *) -(* iIntros (Hsubset Φ) "([#Hinv #Hinv'] & Hvs) HΦ". *) -(* rewrite /read_counter1. *) -(* wp_pure. *) -(* iInv "Hinv'" as ">( %l & %z & -> & H5 & H6)" "Hclose". *) -(* wp_load. *) -(* iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". *) -(* { apply difference_mono_l. *) -(* by apply nclose_subseteq'. } *) -(* iMod ("Hvs" with "[$]") as "[? HQ]". *) -(* iMod "Hclose'" as "_". *) -(* iMod ("Hclose" with "[-HQ HΦ]"); first by iFrame. *) -(* iApply ("HΦ" with "[$]"). *) -(* Qed. *) + }}} + read_counter1 c @ E + {{{ (n':nat), RET #n'; Q n' + }}}. + Proof. + iIntros (Hsubset Φ) "([#Hinv #Hinv'] & Hvs) HΦ". + rewrite /read_counter1. + wp_pure. + iInv "Hinv'" as ">( %l & %z & -> & H5 & H6)" "Hclose". + wp_load. + iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". + { apply difference_mono_l. + by apply nclose_subseteq'. } + iMod ("Hvs" with "[$]") as "[? HQ]". + iMod "Hclose'" as "_". + iMod ("Hclose" with "[-HQ HΦ]"); first by iFrame. + iApply ("HΦ" with "[$]"). + Qed. -(* End impl1. *) +End impl1. -(* Program Definition random_counter1 `{!conerisGS Σ, F:rand_spec}: random_counter := *) -(* {| new_counter := new_counter1; *) -(* allocate_tape:= allocate_tape1; *) -(* incr_counter_tape := incr_counter_tape1; *) -(* read_counter:=read_counter1; *) -(* counterG := counterG1; *) -(* tape_name := rand_tape_name; *) -(* counter_name :=gname; *) -(* is_counter _ N c γ1 γ2 := is_counter1 (L:=counterG1_randG) N c γ1 γ2; *) -(* counter_tapes_auth _ γ m := rand_tapes_auth (L:=counterG1_randG) γ ((λ ns, (3, ns))<$>m); *) -(* counter_tapes_frag _ γ α ns := rand_tapes_frag (L:= counterG1_randG)γ α (3%nat, ns); *) -(* counter_content_auth _ γ z := own γ (●F z); *) -(* counter_content_frag _ γ f z := own γ (◯F{f} z); *) -(* new_counter_spec _ := new_counter_spec1; *) -(* allocate_tape_spec _ :=allocate_tape_spec1; *) -(* incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some1; *) -(* counter_presample_spec _ :=counter_presample_spec1; *) -(* read_counter_spec _ :=read_counter_spec1 *) -(* |}. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (?????????) "H1 H2". *) -(* iDestruct (rand_tapes_auth_exclusive with "[$][$]") as "[]". *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros. *) -(* iDestruct (rand_tapes_frag_exclusive with "[$][$]") as "[]". *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros. *) -(* iDestruct (rand_tapes_agree with "[$][$]") as "%H". *) -(* iPureIntro. *) -(* apply lookup_fmap_Some in H as (?&?&?). *) -(* by simplify_eq. *) -(* Qed. *) -(* Next Obligation. *) -(* iIntros. *) -(* iDestruct (rand_tapes_valid with "[$]") as "$". *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros. *) -(* iMod (rand_tapes_update with "[$][$]") as "[??]"; last iFrame; first done. *) -(* by rewrite fmap_insert. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (?????????) "H1 H2". *) -(* iCombine "H1 H2" gives "%H". by rewrite auth_auth_op_valid in H. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. iIntros (??????? z z' ?) "H1 H2". *) -(* iCombine "H1 H2" gives "%H". *) -(* apply frac_auth_included_total in H. iPureIntro. *) -(* by apply nat_included. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros. *) -(* rewrite frac_auth_frag_op. rewrite own_op. *) -(* iSplit; iIntros; iFrame. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. iIntros (?????????) "H1 H2". *) -(* iCombine "H1 H2" gives "%H". *) -(* iPureIntro. *) -(* by apply frac_auth_agree_L in H. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. iIntros (???????????) "H1 H2". *) -(* iMod (own_update_2 _ _ _ (_ ⋅ _) with "[$][$]") as "[$$]"; last done. *) -(* apply frac_auth_update. *) -(* apply nat_local_update. lia. *) -(* Qed. *) +Program Definition random_counter1 `{!conerisGS Σ, F:rand_spec}: random_counter := + {| new_counter := new_counter1; + allocate_tape:= allocate_tape1; + incr_counter_tape := incr_counter_tape1; + read_counter:=read_counter1; + counterG := counterG1; + tape_name := rand_tape_name; + counter_name :=gname; + is_counter _ N c γ1 γ2 := is_counter1 (L:=counterG1_randG) N c γ1 γ2; + counter_tapes_auth _ γ m := rand_tapes_auth (L:=counterG1_randG) γ ((λ ns, (3, ns))<$>m); + counter_tapes_frag _ γ α ns := rand_tapes_frag (L:= counterG1_randG)γ α (3%nat, ns); + counter_content_auth _ γ z := own γ (●F z); + counter_content_frag _ γ f z := own γ (◯F{f} z); + new_counter_spec _ := new_counter_spec1; + allocate_tape_spec _ :=allocate_tape_spec1; + incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some1; + counter_presample_spec _ :=counter_presample_spec1; + read_counter_spec _ :=read_counter_spec1 + |}. +Next Obligation. + simpl. + iIntros (?????????) "H1 H2". + iDestruct (rand_tapes_auth_exclusive with "[$][$]") as "[]". +Qed. +Next Obligation. + simpl. + iIntros. + iDestruct (rand_tapes_frag_exclusive with "[$][$]") as "[]". +Qed. +Next Obligation. + simpl. + iIntros. + iDestruct (rand_tapes_agree with "[$][$]") as "%H". + iPureIntro. + apply lookup_fmap_Some in H as (?&?&?). + by simplify_eq. +Qed. +Next Obligation. + iIntros. + iDestruct (rand_tapes_valid with "[$]") as "$". +Qed. +Next Obligation. + simpl. + iIntros. + iMod (rand_tapes_update with "[$][$]") as "[??]"; last iFrame; first done. + by rewrite fmap_insert. +Qed. +Next Obligation. + simpl. + iIntros (?????????) "H1 H2". + iCombine "H1 H2" gives "%H". by rewrite auth_auth_op_valid in H. +Qed. +Next Obligation. + simpl. iIntros (??????? z z' ?) "H1 H2". + iCombine "H1 H2" gives "%H". + apply frac_auth_included_total in H. iPureIntro. + by apply nat_included. +Qed. +Next Obligation. + simpl. + iIntros. + rewrite frac_auth_frag_op. rewrite own_op. + iSplit; iIntros; iFrame. +Qed. +Next Obligation. + simpl. iIntros (?????????) "H1 H2". + iCombine "H1 H2" gives "%H". + iPureIntro. + by apply frac_auth_agree_L in H. +Qed. +Next Obligation. + simpl. iIntros (???????????) "H1 H2". + iMod (own_update_2 _ _ _ (_ ⋅ _) with "[$][$]") as "[$$]"; last done. + apply frac_auth_update. + apply nat_local_update. lia. +Qed. diff --git a/theories/coneris/lib/hocap_flip.v b/theories/coneris/lib/hocap_flip.v index c31d5a29..1a880185 100644 --- a/theories/coneris/lib/hocap_flip.v +++ b/theories/coneris/lib/hocap_flip.v @@ -67,8 +67,7 @@ Class flip_spec `{!conerisGS Σ} := FlipSpec (∀ ns', ↯ (ε2 ns') ={∅, E∖↑NS}=∗ T ε num ε2 ns')) -∗ - wp_update E (∃ ε num ε2 ns', T ε num ε2 ns' ∗ - flip_tapes_frag (L:=L) γ1 α (ns ++ ns')) + wp_update E (∃ ε num ε2 ns', flip_tapes_frag (L:=L) γ1 α (ns ++ ns') ∗ T ε num ε2 ns') }. @@ -157,7 +156,7 @@ Section instantiate_flip. ⌜∀ xs ys, fmap (FMap:= list_fmap) bool_to_nat xs = fmap (FMap:= list_fmap)fin_to_nat ys -> ε2' xs = ε2 ys⌝ ∗ T ε num ε2' bs' )%I - with "[//][$][-]") as "(%&%&%&%&[(%&%&%K&%&?)?])"; [exact| |iModIntro; iFrame]; last first. + with "[//][$][-]") as "(%&%&%&%&?&%&%&%K&%&?)"; [exact| |iModIntro; iFrame]; last first. { by rewrite fmap_app K. } iMod "Hvs" as "(%&%&%&Herr &%Hpos&%Hineq &Hrest)". iFrame. diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index bf7ab59b..e29f8e24 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -68,8 +68,8 @@ Class rand_spec `{!conerisGS Σ} := RandSpec ⌜(∀ l, length l = num -> 0<= ε2 l)%R⌝ ∗ ⌜(SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R⌝ ∗ (∀ (ns':list (fin (S tb))), ↯ (ε2 ns') ={∅, E∖↑N}=∗ T ε num ε2 ns')) -∗ - wp_update E (∃ ε num ε2 ns', T ε num ε2 ns' ∗ - rand_tapes_frag (L:=L) γ2 α (tb, ns ++ (fin_to_nat <$> ns'))) + wp_update E (∃ ε num ε2 ns', rand_tapes_frag (L:=L) γ2 α (tb, ns ++ (fin_to_nat <$> ns')) ∗ + T ε num ε2 ns') }. Section impl. @@ -299,7 +299,7 @@ Section checks. apply list_fmap_eq_inj in H'; first by simplify_eq. apply fin_to_nat_inj. - iModIntro. - iDestruct "H" as "(%&%&%&%&[(%&%&%&%&->&?) ?])". + iDestruct "H" as "(%&%&%&%&?&->&<-&%&%&->&?)". iFrame. Qed. From 33577e0724583e2abe02df2a8830aa928a914d12 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Mon, 30 Sep 2024 11:22:55 +0200 Subject: [PATCH 52/69] NIT --- .../coneris/examples/random_counter/impl2.v | 248 +++++++++--------- 1 file changed, 118 insertions(+), 130 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index 471725dd..a66fb394 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -118,8 +118,8 @@ Section impl2. inv (N.@"counter") (counter_inv_pred2 c γ2) )%I. - Lemma new_counter_spec2 E ε N: - {{{ ↯ ε }}} + Lemma new_counter_spec2 E N: + {{{ True }}} new_counter2 #() @ E {{{ (c:val), RET c; ∃ γ1 γ2, is_counter2 N c γ1 γ2 ∗ own γ2 (◯F 0%nat) @@ -182,19 +182,17 @@ Section impl2. by apply Forall_nil. Qed. - Lemma incr_counter_tape_spec_some2 N E c γ1 γ2 (Q:nat->nat -> list nat -> iProp Σ) (α:loc) : + Lemma incr_counter_tape_spec_some2 N E c γ1 γ2 (Q:nat->iProp Σ) (α:loc) n ns: ↑N⊆E -> {{{ is_counter2 N c γ1 γ2 ∗ - ( |={E∖ ↑N, ∅}=> - ∃ n ns, (flip_tapes_frag (L:=L) γ1 α (expander (n::ns)) ∗ ⌜Forall (λ x, x<4) (n::ns)⌝) ∗ - (flip_tapes_frag (L:=L) γ1 α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝) ={∅, E∖↑N}=∗ - ∀ (z:nat), own γ2 (●F z) - ={E∖↑N}=∗ - own γ2 (●F (z+n)) ∗ Q z n ns) - + (flip_tapes_frag (L:=L) γ1 α (expander (n::ns)) ∗ ⌜Forall (λ x, x<4) (n::ns)⌝) ∗ + ( ∀ (z:nat), own γ2 (●F z) ={E∖↑N}=∗ + own γ2 (●F (z+n)) ∗ Q z) + }}} incr_counter_tape2 c #lbl:α @ E - {{{ (z n:nat), RET (#z, #n); ∃ ns, Q z n ns}}}. + {{{ (z:nat), RET (#z, #n); + (flip_tapes_frag (L:=L) γ1 α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝) ∗ Q z }}}. Proof. Admitted. (* iIntros (Hsubset Φ) "((#Hinv & #Hinv') & #Hvs & HP & [Hα %]) HΦ". *) @@ -250,19 +248,21 @@ Section impl2. (* by eapply Forall_inv_tail. *) (* Qed. *) - (* Lemma counter_presample_spec2 N E ns α *) - (* (ε2 : R -> nat -> R) *) - (* (P : iProp Σ) T γ1 γ2 γ3 c: *) - (* ↑N ⊆ E -> *) - (* (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> *) - (* (∀ (ε:R), 0<= ε ->SeriesC (λ n, if (bool_decide (n≤3%nat)) then 1 / (S 3%nat) * ε2 ε n else 0%R)%R <= ε)%R-> *) - (* (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ *) - (* inv (N.@ "counter") (counter_inv_pred2 c γ3)) -∗ *) - (* (□∀ (ε:R) n, (P ∗ flip_error_auth (L:=L) γ1 ε) ={E∖↑N}=∗ *) - (* (⌜(1<=ε2 ε n)%R⌝ ∨(flip_error_auth (L:=L) γ1 (ε2 ε n) ∗ T (n)))) *) - (* -∗ *) - (* P -∗ (flip_tapes_frag (L:=L) γ2 α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝) -∗ *) - (* wp_update E (∃ n, T (n) ∗ (flip_tapes_frag (L:=L) γ2 α (expander (ns++[n])) ∗ ⌜Forall (λ x, x<4) (ns++[n])⌝)). *) + Lemma counter_presample_spec2 NS E T γ1 γ2 c α ns: + ↑NS ⊆ E -> + is_counter2 NS c γ1 γ2 -∗ + (flip_tapes_frag (L:=L) γ1 α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝) -∗ + ( |={E∖↑NS,∅}=> + ∃ ε ε2 num, + ↯ ε ∗ + ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ + ⌜(SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R⌝ ∗ + (∀ ns', ↯ (ε2 ns') ={∅,E∖↑NS}=∗ + T ε ε2 num ns' + ))-∗ + wp_update E (∃ ε ε2 num ns', (flip_tapes_frag (L:=L) γ1 α (expander (ns++ns')) ∗ ⌜Forall (λ x, x<4) (ns++ns')⌝) ∗ T ε ε2 num ns'). + Proof. + Admitted. (* Proof. *) (* iIntros (Hsubset Hpos Hineq) "(#Hinv & #Hinv') #Hvs HP Hfrag". *) (* iApply fupd_wp_update. *) @@ -351,18 +351,18 @@ Section impl2. (* iFrame. *) (* Qed. *) - (* Lemma read_counter_spec2 N E c γ1 γ2 γ3 P Q: *) - (* ↑N ⊆ E -> *) - (* {{{ (is_flip (L:=L) (N.@"flip") γ1 γ2 ∗ *) - (* inv (N.@ "counter") (counter_inv_pred2 c γ3)) ∗ *) - (* □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ *) - (* own γ3 (●F z)∗ Q z) *) - (* ∗ P *) - (* }}} *) - (* read_counter2 c @ E *) - (* {{{ (n':nat), RET #n'; Q n' *) - (* }}}. *) - (* Proof. *) + Lemma read_counter_spec2 N E c γ1 γ2 Q: + ↑N ⊆ E -> + {{{ is_counter2 N c γ1 γ2 ∗ + (∀ (z:nat), own γ2 (●F z) ={E∖↑N}=∗ + own γ2 (●F z) ∗ Q z) + + }}} + read_counter2 c @ E + {{{ (n':nat), RET #n'; Q n' + }}}. + Proof. + Admitted. (* iIntros (Hsubset Φ) "((#Hinv & #Hinv') & #Hvs & HP) HΦ". *) (* rewrite /read_counter2. *) (* wp_pure. *) @@ -377,97 +377,85 @@ Section impl2. (* Qed. *) End impl2. - -(* Program Definition random_counter2 `{flip_spec Σ}: random_counter := *) -(* {| new_counter := new_counter2; *) -(* allocate_tape:= allocate_tape2; *) -(* incr_counter_tape := incr_counter_tape2; *) -(* read_counter:=read_counter2; *) -(* counterG := counterG2; *) -(* error_name := flip_error_name; *) -(* tape_name := flip_tape_name; *) -(* counter_name :=gname; *) -(* is_counter _ N c γ1 γ2 γ3 := (is_flip (N.@"flip") γ1 γ2 ∗ *) -(* inv (N.@ "counter") (counter_inv_pred2 c γ3))%I; *) -(* counter_error_auth _ γ x := flip_error_auth γ x; *) -(* counter_error_frag _ γ x := flip_error_frag γ x; *) -(* counter_tapes_auth _ γ m := flip_tapes_auth (L:=_) γ ((λ ns, expander ns)<$>m); *) -(* counter_tapes_frag _ γ α ns := (flip_tapes_frag (L:=_) γ α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝)%I; *) -(* counter_content_auth _ γ z := own γ (●F z); *) -(* counter_content_frag _ γ f z := own γ (◯F{f} z); *) -(* new_counter_spec _ := new_counter_spec2; *) -(* allocate_tape_spec _ :=allocate_tape_spec2; *) -(* incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some2; *) -(* counter_presample_spec _ :=counter_presample_spec2; *) -(* read_counter_spec _ :=read_counter_spec2 *) -(* |}. *) -(* Next Obligation. *) -(* intros ?? H ???. *) -(* apply counterG2_flipG. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (???????) "H1 H2". *) -(* iApply (flip_tapes_auth_exclusive with "[$][$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (???????) "H1 H2". *) -(* iCombine "H1 H2" gives "%K". by rewrite auth_auth_op_valid in K. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. iIntros (????? z z' ?) "H1 H2". *) -(* iCombine "H1 H2" gives "%K". *) -(* apply frac_auth_included_total in K. iPureIntro. *) -(* by apply nat_included. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (?????????). *) -(* rewrite frac_auth_frag_op. by rewrite own_op. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. iIntros (???????) "H1 H2". *) -(* iCombine "H1 H2" gives "%K". *) -(* iPureIntro. *) -(* by apply frac_auth_agree_L in K. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. iIntros (?????????) "H1 H2". *) -(* iMod (own_update_2 _ _ _ (_ ⋅ _) with "[$][$]") as "[$$]"; last done. *) -(* apply frac_auth_update. *) -(* apply nat_local_update. lia. *) -(* Qed. *) -(* Next Obligation. *) -(* intros ?? H ?. *) -(* apply counterG2_flipG. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (????????) "[? _] [? _]". *) -(* by iDestruct (flip_tapes_frag_exclusive with "[$][$]") as "%". *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros. *) -(* iApply (flip_error_update with "[$][$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* iIntros. *) -(* iApply (flip_error_agree with "[$][$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* iIntros. *) -(* iApply (flip_error_frag_valid with "[$]"). *) -(* Admitted. *) -(* Next Obligation. *) -(* Admitted. *) -(* Next Obligation. *) -(* Admitted. *) -(* Next Obligation. *) -(* Admitted. *) -(* Next Obligation. *) -(* simpl. *) -(* Admitted. *) -(* Next Obligation. *) -(* Admitted. *) + +Program Definition random_counter2 `{flip_spec Σ}: random_counter := + {| new_counter := new_counter2; + allocate_tape:= allocate_tape2; + incr_counter_tape := incr_counter_tape2; + read_counter:=read_counter2; + counterG := counterG2; + tape_name := flip_tape_name; + counter_name :=gname; + is_counter _ N c γ1 γ2 := is_counter2 N c γ1 γ2; + counter_tapes_auth _ γ m := flip_tapes_auth (L:=_) γ ((λ ns, expander ns)<$>m); + counter_tapes_frag _ γ α ns := (flip_tapes_frag (L:=_) γ α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝)%I; + counter_content_auth _ γ z := own γ (●F z); + counter_content_frag _ γ f z := own γ (◯F{f} z); + new_counter_spec _ := new_counter_spec2; + allocate_tape_spec _ :=allocate_tape_spec2; + incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some2; + counter_presample_spec _ :=counter_presample_spec2; + read_counter_spec _ :=read_counter_spec2 + |}. +Next Obligation. + intros ?? ? ???. + apply counterG2_flipG. +Qed. +Next Obligation. + simpl. + iIntros (???????) "H1 H2". + iApply (flip_tapes_auth_exclusive with "[$][$]"). +Qed. +Next Obligation. + simpl. + iIntros (???????) "H1 H2". + iCombine "H1 H2" gives "%K". by rewrite auth_auth_op_valid in K. +Qed. +Next Obligation. + simpl. iIntros (????? z z' ?) "H1 H2". + iCombine "H1 H2" gives "%K". + apply frac_auth_included_total in K. iPureIntro. + by apply nat_included. +Qed. +Next Obligation. + simpl. + iIntros (?????????). + rewrite frac_auth_frag_op. by rewrite own_op. +Qed. +Next Obligation. + simpl. iIntros (???????) "H1 H2". + iCombine "H1 H2" gives "%K". + iPureIntro. + by apply frac_auth_agree_L in K. +Qed. +Next Obligation. + simpl. iIntros (?????????) "H1 H2". + iMod (own_update_2 _ _ _ (_ ⋅ _) with "[$][$]") as "[$$]"; last done. + apply frac_auth_update. + apply nat_local_update. lia. +Qed. +Next Obligation. + intros ?? H ?. + apply counterG2_flipG. +Qed. +Next Obligation. + simpl. + iIntros (???????) "[? %]". + iPureIntro. + eapply Forall_impl; first done. + simpl. lia. +Qed. +Next Obligation. + simpl. + iIntros (????????) "[??] [??]". + iApply (flip_tapes_frag_exclusive with "[$][$]"). +Qed. +Next Obligation. + simpl. + iIntros (????????) "Hauth [Hfrag ?]". + iDestruct (flip_tapes_agree γ α ((λ ns0 : list nat, expander ns0) <$> m) (expander ns) with "[$][Hfrag]") as "%". + - (* HUH? *) +Admitted. +Next Obligation. + simpl. +Admitted. From 9367f48f3c89bf1e3f8f681ce74a705412fb9df6 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Mon, 30 Sep 2024 12:48:46 +0200 Subject: [PATCH 53/69] Progress in impl2 --- .../coneris/examples/random_counter/impl2.v | 96 ++++++++++++------- 1 file changed, 64 insertions(+), 32 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index a66fb394..84fa49a7 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -4,13 +4,34 @@ From clutch.coneris Require Import coneris hocap random_counter hocap_flip. Set Default Proof Using "Type*". - Local Definition expander (l:list nat):= l ≫= (λ x, [2<=?x; (Nat.odd x)]). Local Lemma expander_eta x: x<4->(Z.of_nat x= Z.of_nat 2*Z.b2z (2<=? x)%nat + Z.b2z (Nat.odd x))%Z. Proof. destruct x as [|[|[|[|]]]]; last lia; intros; simpl; lia. Qed. + +Local Lemma expander_inj l1 l2: Forall (λ x, x<4) l1 -> + Forall (λ y, y<4) l2 -> + expander l1 = expander l2 -> + l1 = l2. +Proof. + rewrite /expander. + revert l2. + induction l1 as [|x xs IH]. + - simpl. + by intros [] ???. + - simpl. + intros [|y ys]; simpl; first done. + rewrite !Forall_cons. + rewrite /Nat.odd. + intros [K1 H1] [K2 H2] H. + simplify_eq. + f_equal; last naive_solver. + repeat (destruct x as [|x]; try lia); + repeat (destruct y as [|y]; try lia); + simpl in *; simplify_eq. +Qed. (* Local Definition expander (l:list nat):= *) (* l ≫= (λ x, [Nat.div2 x; Nat.b2n (Nat.odd x)]). *) @@ -377,6 +398,11 @@ Section impl2. (* Qed. *) End impl2. + +Program Definition counterG2_to_flipG `{conerisGS Σ, !flip_spec, !counterG2 Σ} : flipG Σ. +Proof. + eapply counterG2_flipG. +Qed. Program Definition random_counter2 `{flip_spec Σ}: random_counter := {| new_counter := new_counter2; @@ -387,8 +413,8 @@ Program Definition random_counter2 `{flip_spec Σ}: random_counter := tape_name := flip_tape_name; counter_name :=gname; is_counter _ N c γ1 γ2 := is_counter2 N c γ1 γ2; - counter_tapes_auth _ γ m := flip_tapes_auth (L:=_) γ ((λ ns, expander ns)<$>m); - counter_tapes_frag _ γ α ns := (flip_tapes_frag (L:=_) γ α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝)%I; + counter_tapes_auth K γ m := (flip_tapes_auth (L:=counterG2_to_flipG) γ ((λ ns, expander ns)<$>m) ∗ ⌜map_Forall (λ _ ns, Forall (λ x, x<4) ns) m⌝)%I; + counter_tapes_frag K γ α ns := (flip_tapes_frag (L:=counterG2_to_flipG) γ α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝)%I; counter_content_auth _ γ z := own γ (●F z); counter_content_frag _ γ f z := own γ (◯F{f} z); new_counter_spec _ := new_counter_spec2; @@ -398,13 +424,44 @@ Program Definition random_counter2 `{flip_spec Σ}: random_counter := read_counter_spec _ :=read_counter_spec2 |}. Next Obligation. - intros ?? ? ???. - apply counterG2_flipG. + simpl. + iIntros (???????) "[H1 ?] [H2 ?]". + iApply (flip_tapes_auth_exclusive with "[$][$]"). Qed. Next Obligation. simpl. - iIntros (???????) "H1 H2". - iApply (flip_tapes_auth_exclusive with "[$][$]"). + iIntros (????????) "[??] [??]". + iApply (flip_tapes_frag_exclusive with "[$][$]"). +Qed. +Next Obligation. + simpl. + iIntros (????????) "[Hauth %H0] [Hfrag %]". + iDestruct (flip_tapes_agree γ α ((λ ns0 : list nat, expander ns0) <$> m) (expander ns) with "[$][$]") as "%K". + iPureIntro. + rewrite lookup_fmap_Some in K. destruct K as (?&K1&?). + replace ns with x; first done. + apply expander_inj; try done. + by eapply map_Forall_lookup_1 in H0. +Qed. +Next Obligation. + simpl. + iIntros (???????) "[? %]". + iPureIntro. + eapply Forall_impl; first done. + simpl. lia. +Qed. +Next Obligation. + simpl. + iIntros (??????????) "[H1 %] [H2 %]". + iMod (flip_tapes_update with "[$][$]") as "[??]". + iFrame. + iModIntro. + rewrite fmap_insert. iFrame. + iPureIntro. split. + - apply map_Forall_insert_2; last done. + eapply Forall_impl; first done. simpl; lia. + - eapply Forall_impl; first done. + simpl; lia. Qed. Next Obligation. simpl. @@ -434,28 +491,3 @@ Next Obligation. apply frac_auth_update. apply nat_local_update. lia. Qed. -Next Obligation. - intros ?? H ?. - apply counterG2_flipG. -Qed. -Next Obligation. - simpl. - iIntros (???????) "[? %]". - iPureIntro. - eapply Forall_impl; first done. - simpl. lia. -Qed. -Next Obligation. - simpl. - iIntros (????????) "[??] [??]". - iApply (flip_tapes_frag_exclusive with "[$][$]"). -Qed. -Next Obligation. - simpl. - iIntros (????????) "Hauth [Hfrag ?]". - iDestruct (flip_tapes_agree γ α ((λ ns0 : list nat, expander ns0) <$> m) (expander ns) with "[$][Hfrag]") as "%". - - (* HUH? *) -Admitted. -Next Obligation. - simpl. -Admitted. From 2a2a3b8966bd807de2d6aef3057d1924bb9a8486 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Mon, 30 Sep 2024 14:30:54 +0200 Subject: [PATCH 54/69] progress with impl2 --- .../coneris/examples/random_counter/impl2.v | 112 +++++++----------- theories/coneris/lib/hocap_flip.v | 4 +- 2 files changed, 48 insertions(+), 68 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index 84fa49a7..aa42b71d 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -215,59 +215,40 @@ Section impl2. {{{ (z:nat), RET (#z, #n); (flip_tapes_frag (L:=L) γ1 α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝) ∗ Q z }}}. Proof. - Admitted. - (* iIntros (Hsubset Φ) "((#Hinv & #Hinv') & #Hvs & HP & [Hα %]) HΦ". *) - (* rewrite /incr_counter_tape2. *) - (* wp_pures. *) - (* wp_apply (flip_tape_spec_some _ _ _ _ (flip_tapes_frag γ2 α (expander (n::ns)) ) (flip_tapes_frag γ2 α (drop 1%nat (expander (n::ns))))with "[$Hα]"); [|iSplit; first done|]. *) - (* { by apply nclose_subseteq'. } *) - (* { iModIntro. iIntros (m) "[Hfrag Hauth]". *) - (* iDestruct (flip_tapes_agree with "[$][$]") as "%". *) - (* simpl. *) - (* iMod (flip_tapes_pop with "[$][$]") as "[Hauth Htape]". *) - (* iModIntro. *) - (* simpl in *. *) - (* by iFrame. *) - (* } *) - (* iIntros "Hα". *) - (* wp_apply (conversion.wp_bool_to_int) as "_"; first done. *) - (* wp_pures. *) - (* wp_apply (flip_tape_spec_some _ _ _ _ (flip_tapes_frag γ2 α (drop 1%nat (expander (n::ns)))) (flip_tapes_frag γ2 α (expander ns))with "[$Hα]"); [|iSplit; first done|]. *) - (* { by apply nclose_subseteq'. } *) - (* { iModIntro. iIntros (m) "[Hfrag Hauth]". *) - (* iDestruct (flip_tapes_agree with "[$][$]") as "%". *) - (* simpl. *) - (* iMod (flip_tapes_pop with "[$][$]") as "[Hauth Htape]". *) - (* iModIntro. *) - (* simpl in *. *) - (* by iFrame. *) - (* } *) - (* iIntros "Hα". *) - (* wp_apply (conversion.wp_bool_to_int) as "_"; first done. *) - (* wp_pures. *) - (* wp_bind (FAA _ _). *) - (* iInv "Hinv'" as ">(%&%&->&?&?)" "Hclose". *) - (* wp_faa. simpl. *) - (* iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". *) - (* { apply difference_mono; [done|by apply nclose_subseteq']. } *) - (* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) - (* replace 2%Z with (Z.of_nat 2%nat) by done. *) - (* replace (_*_+_)%Z with (Z.of_nat n); last first. *) - (* { assert (n<4). *) - (* - by eapply (@Forall_inv _ (λ x, x<4)). *) - (* - by apply expander_eta. *) - (* } *) - (* replace (#(z+n)) with (#(z+n)%nat); last first. *) - (* { by rewrite Nat2Z.inj_add. } *) - (* iMod "Hclose'" as "_". *) - (* iMod ("Hclose" with "[-HΦ HQ Hα]") as "_"; first by iFrame. *) - (* iModIntro. *) - (* wp_pures. *) - (* iApply "HΦ". *) - (* iFrame. *) - (* iPureIntro. *) - (* by eapply Forall_inv_tail. *) - (* Qed. *) + iIntros (Hsubset Φ) "((#Hinv & #Hinv') & [Hα %] & Hvs) HΦ". + rewrite /incr_counter_tape2. + wp_pures. + wp_apply (flip_tape_spec_some with "[$]") as "Hα". + { by apply nclose_subseteq'. } + wp_apply (conversion.wp_bool_to_int) as "_"; first done. + wp_pures. + wp_apply (flip_tape_spec_some with "[$]") as "Hα". + { by apply nclose_subseteq'. } + wp_apply (conversion.wp_bool_to_int) as "_"; first done. + wp_pures. + wp_bind (FAA _ _). + iInv "Hinv'" as ">(%&%&->&?&?)" "Hclose". + wp_faa. simpl. + iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". + { apply difference_mono; [done|by apply nclose_subseteq']. } + iMod ("Hvs" with "[$]") as "[H6 HQ]". + replace 2%Z with (Z.of_nat 2%nat) by done. + replace (_*_+_)%Z with (Z.of_nat n); last first. + { assert (n<4). + - by eapply (@Forall_inv _ (λ x, x<4)). + - by apply expander_eta. + } + replace (#(z+n)) with (#(z+n)%nat); last first. + { by rewrite Nat2Z.inj_add. } + iMod "Hclose'" as "_". + iMod ("Hclose" with "[-HΦ HQ Hα]") as "_"; first by iFrame. + iModIntro. + wp_pures. + iApply "HΦ". + iFrame. + iPureIntro. + by eapply Forall_inv_tail. + Qed. Lemma counter_presample_spec2 NS E T γ1 γ2 c α ns: ↑NS ⊆ E -> @@ -383,19 +364,18 @@ Section impl2. {{{ (n':nat), RET #n'; Q n' }}}. Proof. - Admitted. - (* iIntros (Hsubset Φ) "((#Hinv & #Hinv') & #Hvs & HP) HΦ". *) - (* rewrite /read_counter2. *) - (* wp_pure. *) - (* iInv "Hinv'" as ">(%l&%z & -> & Hloc & Hcont)" "Hclose". *) - (* wp_load. *) - (* iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". *) - (* { apply difference_mono_l. by apply nclose_subseteq'. } *) - (* iMod ("Hvs" with "[$]") as "[Hcont HQ]". *) - (* iMod "Hclose'". *) - (* iMod ("Hclose" with "[ $Hloc $Hcont]"); first done. *) - (* iApply "HΦ". by iFrame. *) - (* Qed. *) + iIntros (Hsubset Φ) "((#Hinv & #Hinv') & Hvs) HΦ". + rewrite /read_counter2. + wp_pure. + iInv "Hinv'" as ">(%l&%z & -> & Hloc & Hcont)" "Hclose". + wp_load. + iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". + { apply difference_mono_l. by apply nclose_subseteq'. } + iMod ("Hvs" with "[$]") as "[Hcont HQ]". + iMod "Hclose'". + iMod ("Hclose" with "[ $Hloc $Hcont]"); first done. + iApply "HΦ". by iFrame. + Qed. End impl2. diff --git a/theories/coneris/lib/hocap_flip.v b/theories/coneris/lib/hocap_flip.v index 1a880185..a488b39b 100644 --- a/theories/coneris/lib/hocap_flip.v +++ b/theories/coneris/lib/hocap_flip.v @@ -49,7 +49,7 @@ Class flip_spec `{!conerisGS Σ} := FlipSpec ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ flip_tapes_frag (L:=L) γ1 α [] }}}; - flip_tape_spec_some {L: flipG Σ} N E γ1 (Q:bool -> list bool -> iProp Σ) (α:loc) n ns: + flip_tape_spec_some {L: flipG Σ} N E γ1 (α:loc) n ns: ↑N⊆E -> {{{ is_flip (L:=L) N γ1 ∗ flip_tapes_frag (L:=L) γ1 α (n::ns) }}} @@ -136,7 +136,7 @@ Section instantiate_flip. Qed. Next Obligation. simpl. - iIntros (???? Q ???? Φ) "(#Hinv & Hfrag) HΦ". + iIntros (???? ???? Φ) "(#Hinv & Hfrag) HΦ". wp_pures. wp_apply (rand_tape_spec_some with "[-HΦ]"); [done|..]. - by iFrame. From aa2d2f4865ef59fc8b2c5761e61a14c7582a35e6 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Mon, 30 Sep 2024 16:52:25 +0200 Subject: [PATCH 55/69] Progress with impl2 --- .../coneris/examples/random_counter/impl2.v | 232 +++++++++++------- 1 file changed, 145 insertions(+), 87 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index aa42b71d..abf383f4 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -32,6 +32,84 @@ Proof. repeat (destruct y as [|y]; try lia); simpl in *; simplify_eq. Qed. + +Local Fixpoint decoder (l:list bool) := + match l with + |[] => Some [] + | b::b'::ls => + res← decoder ls; + Some (((bool_to_nat b)*2+(bool_to_nat b'))::res) + | _ => None +end. + +(* Lemma decoder_unfold l: *) +(* decoder l = *) +(* match l with *) +(* |[] => Some [] *) +(* | b::b'::ls => *) +(* res← decoder ls; *) +(* Some (((bool_to_nat b)*2+(bool_to_nat b'))::res) *) +(* | _ => None end. *) +(* Proof. *) +(* induction l; by rewrite {1}/decoder. *) +(* Qed. *) + +Local Lemma decoder_correct bs ns: decoder bs = Some ns -> expander ns = bs. +Proof. + revert bs. + induction ns. + - intros bs H. simpl. destruct bs as [|?[|??]]; try done. + simpl in *. + rewrite bind_Some in H. + destruct H as (?&?&?). + simplify_eq. + - intros [|b[|b' ?]]; simpl; try done. + intros H. + rewrite bind_Some in H. + destruct H as (?&?&?). + simplify_eq. + repeat f_equal; last naive_solver; destruct b; destruct b'; done. +Qed. + + +Local Lemma decoder_inj x y z: decoder x = Some z -> decoder y = Some z -> x= y. +Proof. + intros H1 H2. + apply decoder_correct in H1, H2. + by subst. +Qed. + +Local Lemma decoder_ineq bs xs: decoder bs = Some xs -> Forall (λ x : nat, x < 4) xs. +Proof. + revert bs. + induction xs; first by (intros; apply Forall_nil). + intros [|b[|b'?]] H; try simplify_eq. + rewrite bind_Some in H. + destruct H as (?&?&?). + simplify_eq. + rewrite Forall_cons. + split; last naive_solver. + destruct b, b'; simpl; lia. +Qed. + +Local Lemma decoder_None p bs : length bs = p -> decoder bs = None -> ¬ ∃ num, length bs = 2 * num. +Proof. + revert bs. + induction (lt_wf p) as [x ? IH]. + destruct x. + - intros []??; simplify_eq. + - intros [|?[|??]]; intros Hlen H'; simplify_eq. + + simpl in *. simplify_eq. + intros (num&?). + destruct num; lia. + + simpl in *. + simplify_eq. + rewrite bind_None in H'. + destruct H' as [|(?&?&?)]; last done. + unshelve epose proof IH (length l) _ l _ _ as H1; try lia; try done. + intros (num & ?). apply H1. + exists (num-1). lia. +Qed. (* Local Definition expander (l:list nat):= *) (* l ≫= (λ x, [Nat.div2 x; Nat.b2n (Nat.odd x)]). *) @@ -264,94 +342,74 @@ Section impl2. ))-∗ wp_update E (∃ ε ε2 num ns', (flip_tapes_frag (L:=L) γ1 α (expander (ns++ns')) ∗ ⌜Forall (λ x, x<4) (ns++ns')⌝) ∗ T ε ε2 num ns'). Proof. + iIntros (Hsubset) "(#Hinv & #Hinv') [Hfrag %] Hvs". + iMod (flip_presample_spec _ _ _ _ _ + (λ ε num' ε2' ns', ∃ num ε2 ns, + ⌜(2 * num = num')%nat⌝∗ ⌜expander ns = ns'⌝ ∗ ⌜Forall (λ x, x<4) ns⌝ ∗ + (* ⌜ ∀ ls, Forall (λ x, x<4) ls -> ε2 ls = ε2' (expander ls)⌝ ∗ *) + T ε ε2 num ns + )%I with "[//][$][Hvs]") as "H"; first by apply nclose_subseteq'. + - iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose". + { apply difference_mono; [done|by apply nclose_subseteq']. } + iMod "Hvs" as "(%ε & %ε2 & %num & Herr & %Hpos & %Hineq & Hvs)". + iExists ε, (2*num). + iFrame. + iModIntro. + iExists (λ ls, match (decoder ls) with | Some ls' => ε2 ls' | _ => 1%R end). + repeat iSplit. + + iPureIntro. intros. case_match; [done|lra]. + + iPureIntro. + etrans; last exact. + rewrite !Rdiv_def -!SeriesC_scal_r. + etrans; last eapply (SeriesC_le_inj _ (λ bs, decoder bs)). + * apply Req_le. + apply SeriesC_ext. + intros bs. + destruct (decoder bs) eqn:K. + -- simpl. f_equal. + ++ case_match eqn:H0. + ** rewrite bool_decide_eq_true_2; first done. + erewrite elem_of_list_fmap. + rewrite Nat.eqb_eq in H0. + admit. + ** rewrite bool_decide_eq_false_2; first done. + rewrite elem_of_list_fmap. + intros (?&->&K'). + rewrite elem_of_enum_uniform_fin_list in K'. + rewrite Nat.eqb_neq in H0. + exfalso. apply H0. + admit. + ++ f_equal. + replace 4%R with (2^2)%R; last (simpl; lra). + by rewrite -pow_mult. + -- simpl. + eapply decoder_None in K; last done. + case_match eqn :H0; last lra. + rewrite Nat.eqb_eq in H0. + exfalso. + apply K. + exists num. lia. + * intros. rewrite -Rdiv_1_l. + replace 4%R with (INR 4); last (simpl; lra). + rewrite -pow_INR. apply Rmult_le_pos; last apply Rdiv_INR_ge_0. + case_bool_decide; [done|lra]. + * intros. by eapply decoder_inj. + * apply ex_seriesC_scal_r, ex_seriesC_list. + + iIntros (ns') "Herr". + case_match eqn :Hdecoder; last (by iDestruct (ec_contradict with "[$]") as "%"). + iMod ("Hvs" with "[$]") as "HT". + iMod "Hclose" as "_". + iFrame. + iPureIntro. repeat split. + * by apply decoder_correct. + * by eapply decoder_ineq. + - iDestruct "H" as "(%&%&%&%&?&%&%&%&<-&<-&%&?)". + iModIntro. + iFrame. + iSplit; last by rewrite Forall_app. + rewrite /expander. + by rewrite bind_app. Admitted. - (* Proof. *) - (* iIntros (Hsubset Hpos Hineq) "(#Hinv & #Hinv') #Hvs HP Hfrag". *) - (* iApply fupd_wp_update. *) - (* Admitted. *) - (* iApply flip_presample_spec. *) - (* iApply wp_update_state_step_coupl. *) - (* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) - (* iMod (inv_acc with "Hinv") as "[>(% & % & % & % & H1 & H2 & H3 & H4 & -> & H5 & H6) Hclose]"; [done|]. *) - (* iDestruct (hocap_tapes_agree' with "[$][$]") as "%". *) - (* erewrite <-(insert_delete m) at 1; last done. *) - (* rewrite big_sepM_insert; last apply lookup_delete. *) - (* simpl. *) - (* iDestruct "H3" as "[Htape H3]". *) - (* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%H1)". *) - (* iDestruct (ec_supply_bound with "[$][$]") as "%". *) - (* iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. *) - (* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) - (* iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang 2%nat); first done. *) - (* pose (f (l:list (fin 2)) := (match l with *) - (* | x::[x'] => 2*fin_to_nat x + fin_to_nat x' *) - (* | _ => 0 *) - (* end)). *) - (* unshelve iExists *) - (* (λ l, mknonnegreal (ε2 ε1' (f l) ) _). *) - (* { apply Hpos; simpl. auto. } *) - (* simpl. *) - (* iSplit. *) - (* { iPureIntro. *) - (* etrans; last apply Hineq; last auto. *) - (* pose (k:=(λ n, 1 / ((1 + 1) * ((1 + 1) * 1)) * ε2 ε1' (f n))%R). *) - (* erewrite (SeriesC_ext _ (λ x, if bool_decide (x ∈ enum_uniform_list 1%nat 2%nat) *) - (* then k x else 0%R))%R; last first. *) - (* - intros. case_bool_decide as K. *) - (* + rewrite elem_of_enum_uniform_list in K. by rewrite K /k. *) - (* + rewrite elem_of_enum_uniform_list in K. *) - (* case_match eqn:K'; [by rewrite Nat.eqb_eq in K'|done]. *) - (* - rewrite SeriesC_list; last apply NoDup_enum_uniform_list. *) - (* rewrite /k /=. rewrite SeriesC_nat_bounded'/=. lra. *) - (* } *) - (* iIntros (sample ?). *) - (* destruct (Rlt_decision (nonneg ε_rem + (ε2 ε1' (f sample)))%R 1%R) as [Hdec|Hdec]; last first. *) - (* { apply Rnot_lt_ge, Rge_le in Hdec. *) - (* iApply state_step_coupl_ret_err_ge_1. *) - (* simpl. simpl in *. lra. *) - (* } *) - (* iApply state_step_coupl_ret. *) - (* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 ε1' (f sample)) _) with "Hε_supply") as "[Hε_supply Hε]". *) - (* { apply Hpos. apply cond_nonneg. } *) - (* { simpl. done. } *) - (* assert (Nat.div2 (f sample)<2)%nat as K1. *) - (* { rewrite Nat.div2_div. apply Nat.Div0.div_lt_upper_bound. rewrite /f. *) - (* simpl. repeat case_match; try lia. pose proof fin_to_nat_lt t. pose proof fin_to_nat_lt t0. lia. *) - (* } *) - (* rewrite -H1. *) - (* iMod (tapeN_update_append _ _ _ _ (nat_to_fin K1) with "[$][$]") as "[Htapes Htape]". *) - (* assert (Nat.b2n (Nat.odd (f sample))<2)%nat as K2. *) - (* { edestruct (Nat.odd _); simpl; lia. } *) - (* rewrite -(list_fmap_singleton (fin_to_nat)) -fmap_app. *) - (* iMod (tapeN_update_append _ _ _ _ (nat_to_fin K2) with "[$][$]") as "[Htapes Htape]". *) - (* iMod (hocap_tapes_presample' _ _ _ _ _ (f sample) with "[$][$]") as "[H4 Hfrag]". *) - (* iMod "Hclose'" as "_". *) - (* iMod ("Hvs" with "[$]") as "[%|[H2 HT]]". *) - (* { iExFalso. iApply (ec_contradict with "[$]"). exact. } *) - (* rewrite insert_insert. *) - (* rewrite fmap_app -!app_assoc /=. *) - (* assert (([nat_to_fin K1;nat_to_fin K2]) = sample) as K. *) - (* { destruct sample as [|x xs]; first done. *) - (* destruct xs as [|x' xs]; first done. *) - (* destruct xs as [|]; last done. *) - (* pose proof fin_to_nat_lt x. pose proof fin_to_nat_lt x'. *) - (* repeat f_equal; apply fin_to_nat_inj; rewrite fin_to_nat_to_fin. *) - (* - rewrite /f. rewrite Nat.div2_div. *) - (* rewrite Nat.mul_comm Nat.div_add_l; last lia. rewrite Nat.div_small; lia. *) - (* - rewrite /f. rewrite Nat.add_comm Nat.odd_add_even. *) - (* + destruct (fin_to_nat x') as [|[|]]; simpl; lia. *) - (* + by econstructor. *) - (* } *) - (* iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4 $H5 $H6]") as "_". *) - (* { iNext. iSplit; last done. *) - (* rewrite big_sepM_insert_delete; iFrame. *) - (* simpl. rewrite !fin_to_nat_to_fin. *) - (* rewrite /expander bind_app -/(expander ns). simpl. by rewrite H1. *) - (* } *) - (* iApply fupd_mask_intro_subseteq; first set_solver. *) - (* rewrite K. *) - (* iFrame. *) - (* Qed. *) Lemma read_counter_spec2 N E c γ1 γ2 Q: ↑N ⊆ E -> From 9a1e51412f4e822e700abd24717a8fedf81189e5 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Tue, 1 Oct 2024 09:28:05 +0200 Subject: [PATCH 56/69] impl2 completed --- .../coneris/examples/random_counter/impl2.v | 30 +++++++++++++++++-- theories/prelude/fin.v | 16 ++++++++-- 2 files changed, 41 insertions(+), 5 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index abf383f4..e416fb06 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -110,6 +110,20 @@ Proof. intros (num & ?). apply H1. exists (num-1). lia. Qed. + +Local Lemma decoder_Some_length bs xs: decoder bs = Some xs -> length bs = 2 * length xs. +Proof. + revert bs. + induction xs as [|x xs IH]; intros [|?[|??]] H; simpl in *; simplify_eq; try done. + - rewrite bind_Some in H. + destruct H as (?&?&?). + simplify_eq. + - rewrite bind_Some in H. + destruct H as (?&?&?). + simplify_eq. + f_equal. rewrite IH; [lia|done]. +Qed. + (* Local Definition expander (l:list nat):= *) (* l ≫= (λ x, [Nat.div2 x; Nat.b2n (Nat.odd x)]). *) @@ -371,14 +385,23 @@ Section impl2. ** rewrite bool_decide_eq_true_2; first done. erewrite elem_of_list_fmap. rewrite Nat.eqb_eq in H0. - admit. + apply decoder_ineq in K as K'. + apply fin.nat_list_to_fin_list in K'. + destruct K' as [xs K']. + exists xs. split; first done. + rewrite elem_of_enum_uniform_fin_list. + apply decoder_Some_length in K. + apply (f_equal length) in K'. + rewrite fmap_length in K'. lia. ** rewrite bool_decide_eq_false_2; first done. rewrite elem_of_list_fmap. intros (?&->&K'). rewrite elem_of_enum_uniform_fin_list in K'. rewrite Nat.eqb_neq in H0. exfalso. apply H0. - admit. + apply decoder_Some_length in K. + rewrite fmap_length in K. + lia. ++ f_equal. replace 4%R with (2^2)%R; last (simpl; lra). by rewrite -pow_mult. @@ -409,7 +432,7 @@ Section impl2. iSplit; last by rewrite Forall_app. rewrite /expander. by rewrite bind_app. - Admitted. + Qed. Lemma read_counter_spec2 N E c γ1 γ2 Q: ↑N ⊆ E -> @@ -529,3 +552,4 @@ Next Obligation. apply frac_auth_update. apply nat_local_update. lia. Qed. + diff --git a/theories/prelude/fin.v b/theories/prelude/fin.v index 5bf91240..67d0c0a7 100644 --- a/theories/prelude/fin.v +++ b/theories/prelude/fin.v @@ -161,8 +161,20 @@ Section fin. rewrite f_inv_cancel_r //. Qed. - - + Lemma nat_list_to_fin_list xs N: + Forall (λ x, x < N) xs -> ∃ (xs' : list (fin N)), fin_to_nat <$> xs' = xs. + Proof. + induction xs. + - intros. by exists []. + - rewrite Forall_cons_iff. + intros [P H]. + apply IHxs in H. + destruct H as [xs' <-]. + eexists (nat_to_fin P :: _). + simpl. f_equal. + by rewrite fin_to_nat_to_fin. + Qed. + (* Adapted from Coq standard library FinFun *) Definition bFunNM n m (f:nat->nat) := forall x, x < n -> f x < m. From e196ae17e65c3f45133dfc87610951987c454282 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Tue, 1 Oct 2024 10:21:16 +0200 Subject: [PATCH 57/69] Progress with impl3 --- .../coneris/examples/random_counter/impl3.v | 772 +++++++++--------- 1 file changed, 393 insertions(+), 379 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl3.v b/theories/coneris/examples/random_counter/impl3.v index 94ca521b..59b1b94d 100644 --- a/theories/coneris/examples/random_counter/impl3.v +++ b/theories/coneris/examples/random_counter/impl3.v @@ -1,13 +1,23 @@ -(* From iris.algebra Require Import frac_auth. *) -(* From iris.base_logic.lib Require Import invariants. *) -(* From clutch.coneris Require Import coneris hocap random_counter. *) +From iris.algebra Require Import frac_auth. +From iris.base_logic.lib Require Import invariants. +From clutch.coneris Require Import coneris hocap hocap_rand random_counter. -(* Set Default Proof Using "Type*". *) +Set Default Proof Using "Type*". -(* Section filter. *) -(* Definition filter_f (n:nat):= bool_decide (n<4)%nat. *) -(* Definition filtered_list (l:list _) := filter filter_f l. *) -(* End filter. *) +Section filter. + Definition filter_f (n:nat):= (n<4)%nat. + Definition filtered_list (l:list _) := filter filter_f l. + + Lemma Forall_filter_f ns: + Forall (λ x : nat, x ≤ 3) ns -> filter filter_f ns = ns. + Proof. + induction ns; first done. + intros H. + rewrite Forall_cons in H. destruct H. + rewrite filter_cons_True; first f_equal; first naive_solver. + rewrite /filter_f. lia. + Qed. +End filter. (* Section lemmas. *) (* Context `{!conerisGS Σ}. *) @@ -19,386 +29,390 @@ (* iDestruct (big_sepM_lookup with "[$]") as "(%&%&?)"; first done. *) (* iExFalso. *) (* iApply (tapeN_tapeN_contradict with "[$][$]"). *) -(* Qed. *) +(* Qed. *) (* End lemmas. *) -(* Section impl3. *) +Section impl3. + Context `{H:conerisGS Σ, r1:@rand_spec Σ H, L:randG Σ, !inG Σ (frac_authR natR)}. -(* Definition new_counter3 : val:= λ: "_", ref #0. *) -(* Definition allocate_tape3 : val := λ: "_", AllocTape #4. *) -(* Definition incr_counter_tape3 :val := rec: "f" "l" "α":= *) -(* let: "n" := rand("α") #4 in *) -(* if: "n" < #4 *) -(* then (FAA "l" "n", "n") *) -(* else "f" "l" "α". *) -(* Definition read_counter3 : val := λ: "l", !"l". *) -(* Class counterG3 Σ := CounterG3 { counterG2_error::hocap_errorGS Σ; *) -(* counterG2_tapes:: hocap_tapesGS Σ; *) -(* counterG2_frac_authR:: inG Σ (frac_authR natR) }. *) + Definition new_counter3 : val:= λ: "_", ref #0. + Definition allocate_tape3 : val := λ: "_", rand_allocate_tape #4. + Definition incr_counter_tape3 :val := rec: "f" "l" "α":= + let: "n" := rand_tape "α" #4 in + if: "n" < #4 + then (FAA "l" "n", "n") + else "f" "l" "α". + Definition read_counter3 : val := λ: "l", !"l". + Class counterG3 Σ := CounterG3 { counterG3_randG:randG Σ; + counterG3_frac_authR:: inG Σ (frac_authR natR) }. -(* Context `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapesGS Σ, !inG Σ (frac_authR natR)}. *) -(* Definition counter_inv_pred3 (c:val) γ1 γ2 γ3:= *) -(* (∃ (ε:R) (m:gmap loc (nat * list nat)) (l:loc) (z:nat), *) -(* ↯ ε ∗ ●↯ ε @ γ1 ∗ *) -(* ([∗ map] α ↦ t ∈ m, ∃ (ls:list nat), ⌜ (filter filter_f ls) = t.2⌝ ∗ α ↪N ( 4%nat ; ls) ) *) -(* ∗ ●m@γ2 ∗ *) -(* ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ3 (●F z) *) -(* )%I. *) + Definition counter_inv_pred3 (c:val) γ2:= (∃ (l:loc) (z:nat), ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ2 (●F z) )%I. -(* Lemma new_counter_spec3 E ε N: *) -(* {{{ ↯ ε }}} *) -(* new_counter3 #() @ E *) -(* {{{ (c:val), RET c; *) -(* ∃ γ1 γ2 γ3, inv N (counter_inv_pred3 c γ1 γ2 γ3) ∗ *) -(* ◯↯ε @ γ1 ∗ own γ3 (◯F 0%nat) *) -(* }}}. *) -(* Proof. *) -(* rewrite /new_counter3. *) -(* iIntros (Φ) "Hε HΦ". *) -(* wp_pures. *) -(* wp_alloc l as "Hl". *) -(* iDestruct (ec_valid with "[$]") as "%". *) -(* unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "[%γ1 [H1 H2]]". *) -(* { lra. } *) -(* simpl. *) -(* iMod (hocap_tapes_alloc (∅:gmap _ _)) as "[%γ2 [H3 H4]]". *) -(* iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". *) -(* { by apply frac_auth_valid. } *) -(* replace (#0) with (#0%nat) by done. *) -(* iMod (inv_alloc N _ (counter_inv_pred3 (#l) γ1 γ2 γ3) with "[$Hε $Hl $H1 $H3 $H5]") as "#Hinv". *) -(* { iSplit; last done. by iApply big_sepM_empty. } *) -(* iApply "HΦ". *) -(* iExists _, _, _. by iFrame. *) -(* Qed. *) + Definition is_counter3 N (c:val) γ1 γ2:= + ((is_rand (L:=L) (N.@"rand") γ1) ∗ + inv (N.@"counter") (counter_inv_pred3 c γ2))%I. -(* Lemma allocate_tape_spec3 N E c γ1 γ2 γ3: *) -(* ↑N ⊆ E-> *) -(* {{{ inv N (counter_inv_pred3 c γ1 γ2 γ3) }}} *) -(* allocate_tape3 #() @ E *) -(* {{{ (v:val), RET v; *) -(* ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ α ◯↪N (3%nat; []) @ γ2 *) -(* }}}. *) -(* Proof. *) -(* iIntros (Hsubset Φ) "#Hinv HΦ". *) -(* rewrite /allocate_tape3. *) -(* wp_pures. *) -(* wp_alloctape α as "Hα". *) -(* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) -(* iDestruct (hocap_tapes_notin3 with "[$][$]") as "%". *) -(* iMod (hocap_tapes_new with "[$]") as "[H4 H7]"; first done. *) -(* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Hα]") as "_". *) -(* { iNext. iSplitL; last done. *) -(* rewrite big_sepM_insert; [simpl; iFrame|done]. *) -(* by rewrite filter_nil. *) -(* } *) -(* iApply "HΦ". *) -(* by iFrame. *) -(* Qed. *) + + (* (∃ (ε:R) (m:gmap loc (nat * list nat)) (l:loc) (z:nat), *) + (* ↯ ε ∗ ●↯ ε @ γ1 ∗ *) + (* ([∗ map] α ↦ t ∈ m, ∃ (ls:list nat), ⌜ (filter filter_f ls) = t.2⌝ ∗ α ↪N ( 4%nat ; ls) ) *) + (* ∗ ●m@γ2 ∗ *) + (* ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ3 (●F z) *) + (* )%I. *) -(* Lemma incr_counter_tape_spec_some3 N E c γ1 γ2 γ3 (P: iProp Σ) (Q:nat->iProp Σ) (α:loc) (n:nat) ns: *) -(* ↑N⊆E -> *) -(* {{{ inv N (counter_inv_pred3 c γ1 γ2 γ3) ∗ *) -(* □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ *) -(* own γ3 (●F(z+n)%nat)∗ Q z) ∗ *) -(* P ∗ α ◯↪N (3%nat; n::ns) @ γ2 *) -(* }}} *) -(* incr_counter_tape3 c #lbl:α @ E *) -(* {{{ (z:nat), RET (#z, #n); Q z ∗ α ◯↪N (3%nat; ns) @ γ2}}}. *) -(* Proof. *) -(* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP & Hα) HΦ". *) -(* rewrite /incr_counter_tape3. *) -(* iLöb as "IH". *) -(* wp_pures. *) -(* wp_bind (rand(_) _)%E. *) -(* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) -(* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) -(* erewrite <-(insert_delete m) at 1; last done. *) -(* rewrite big_sepM_insert; last apply lookup_delete. *) -(* simpl. *) -(* iDestruct "H3" as "[(%ls & %Hls & Htape) H3]". *) -(* destruct ls as [|x ls]. *) -(* { rewrite filter_nil in Hls. simplify_eq. } *) -(* wp_apply (wp_rand_tape with "[$]") as "[Htape %Hineq]". *) -(* rewrite Nat.le_lteq in Hineq. *) -(* destruct Hineq as [? | ->]. *) -(* - (* first value is valid *) *) -(* iMod (hocap_tapes_pop with "[$][$]") as "[H4 Hα]". *) -(* rewrite filter_cons /filter_f in Hls. *) -(* rewrite bool_decide_eq_true_2 in Hls; last done. simpl in *. *) -(* simplify_eq. *) -(* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". *) -(* { iSplitL; last done. *) -(* erewrite <-(insert_delete m) at 2; last done. *) -(* iNext. *) -(* rewrite insert_insert. *) -(* rewrite big_sepM_insert; last apply lookup_delete. by iFrame. *) -(* } *) -(* iModIntro. *) -(* wp_pures. *) -(* rewrite bool_decide_eq_true_2; last lia. *) -(* clear -Hsubset. *) -(* wp_pures. *) -(* wp_bind (FAA _ _). *) -(* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) -(* wp_faa. *) -(* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) -(* replace (#(z+n)) with (#(z+n)%nat); last first. *) -(* { by rewrite Nat2Z.inj_add. } *) -(* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) -(* iModIntro. wp_pures. *) -(* iApply "HΦ". *) -(* by iFrame. *) -(* - (* we get a 5, do iLöb induction *) *) -(* rewrite filter_cons /filter_f in Hls. *) -(* rewrite bool_decide_eq_false_2 in Hls; last lia. simpl in *. *) -(* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". *) -(* { iSplitL; last done. *) -(* erewrite <-(insert_delete m) at 2; last done. *) -(* iNext. *) -(* rewrite big_sepM_insert; last apply lookup_delete. by iFrame. *) -(* } *) -(* iModIntro. *) -(* do 4 wp_pure. *) -(* by iApply ("IH" with "[$][$]"). *) -(* Qed. *) + Lemma new_counter_spec3 E N: + {{{ True }}} + new_counter3 #() @ E + {{{ (c:val), RET c; + ∃ γ1 γ2, is_counter3 N c γ1 γ2 ∗ own γ2 (◯F 0%nat) + }}}. + Proof. + Admitted. + (* rewrite /new_counter3. *) + (* iIntros (Φ) "Hε HΦ". *) + (* wp_pures. *) + (* wp_alloc l as "Hl". *) + (* iDestruct (ec_valid with "[$]") as "%". *) + (* unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "[%γ1 [H1 H2]]". *) + (* { lra. } *) + (* simpl. *) + (* iMod (hocap_tapes_alloc (∅:gmap _ _)) as "[%γ2 [H3 H4]]". *) + (* iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". *) + (* { by apply frac_auth_valid. } *) + (* replace (#0) with (#0%nat) by done. *) + (* iMod (inv_alloc N _ (counter_inv_pred3 (#l) γ1 γ2 γ3) with "[$Hε $Hl $H1 $H3 $H5]") as "#Hinv". *) + (* { iSplit; last done. by iApply big_sepM_empty. } *) + (* iApply "HΦ". *) + (* iExists _, _, _. by iFrame. *) + (* Qed. *) -(* Lemma counter_presample_spec3 NS E ns α *) -(* (ε2 : R -> nat -> R) *) -(* (P : iProp Σ) T γ1 γ2 γ3 c: *) -(* ↑NS ⊆ E -> *) -(* (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> *) -(* (∀ (ε:R), 0<= ε ->SeriesC (λ n, if (bool_decide (n≤3%nat)) then 1 / (S 3%nat) * ε2 ε n else 0%R)%R <= ε)%R-> *) -(* inv NS (counter_inv_pred3 c γ1 γ2 γ3) -∗ *) -(* (□∀ (ε:R) n, (P ∗ ●↯ ε@ γ1) ={E∖↑NS}=∗ *) -(* (⌜(1<=ε2 ε n)%R⌝ ∨(●↯ (ε2 ε n) @ γ1 ∗ T (n)))) *) -(* -∗ *) -(* P -∗ α ◯↪N (3%nat; ns) @ γ2 -∗ *) -(* wp_update E (∃ n, T (n) ∗ α◯↪N (3%nat; ns++[n]) @ γ2). *) -(* Proof. *) -(* iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) -(* iMod wp_update_epsilon_err as "(%epsilon_err&%H&?)". *) -(* iApply wp_update_state_step_coupl. *) -(* iRevert "HP Hfrag". *) -(* iApply (ec_ind_amp _ 5%R with "[][$]"); [done|lra|]. *) -(* iModIntro. *) -(* clear epsilon_err H. *) -(* iIntros (epsilon_err ?) "#IH Hepsilon_err HP Hfrag". *) -(* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) -(* iMod (inv_acc with "Hinv") as "[>(%ε0 & % & % & % & H1 & H2 & H3 & H4 & -> & H5 & H6) Hclose]"; [done|]. *) -(* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) -(* erewrite <-(insert_delete m) at 1; last done. *) -(* rewrite big_sepM_insert; last apply lookup_delete. *) -(* simpl. *) -(* iDestruct "H3" as "[(%ls & %Hls & Htape) H3]". *) -(* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%)". *) -(* iDestruct (ec_valid with "H1") as "%". *) -(* iDestruct (ec_combine with "[$]") as "Htotal". *) -(* iDestruct (ec_supply_bound with "[$][$]") as "%". *) -(* iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. *) -(* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) -(* pose (ε2' := λ (x:nat), (if bool_decide (x<=3)%nat then ε2 ε0 x else *) -(* if bool_decide (x=4)%nat then ε0 + 5 * epsilon_err *) -(* else 0)%R *) -(* ). *) -(* iApply state_step_coupl_state_adv_comp_con_prob_lang; first done. *) -(* assert (forall x, 0<=ε2' x)%R. *) -(* { intros. rewrite /ε2'. repeat case_bool_decide; try done; first naive_solver. *) -(* apply Rplus_le_le_0_compat; simpl; [naive_solver|lra]. } *) -(* unshelve iExists (λ x, mknonnegreal (ε2' x) _); auto. *) -(* iSplit. *) -(* { iPureIntro. *) -(* unshelve epose proof (Hineq ε0 _) as H'; first (by naive_solver). *) -(* rewrite SeriesC_nat_bounded_fin in H'. *) -(* replace (INR 5%nat) with 5%R by (simpl; lra). *) -(* replace (INR 4%nat) with 4%R in H' by (simpl; lra). *) -(* rewrite /=/ε2' Hε1'. *) -(* rewrite SeriesC_finite_foldr/=. *) -(* rewrite SeriesC_finite_foldr/= in H'. *) -(* lra. *) -(* } *) -(* iIntros (sample). *) -(* simpl. *) -(* destruct (Rlt_decision (nonneg ε_rem + (ε2' sample))%R 1%R) as [Hdec|Hdec]; last first. *) -(* { apply Rnot_lt_ge, Rge_le in Hdec. *) -(* iApply state_step_coupl_ret_err_ge_1. *) -(* simpl. simpl in *. lra. *) -(* } *) -(* destruct (Nat.lt_total (fin_to_nat sample) 4%nat) as [|[|]]. *) -(* - (* we sample a value less than 4*) *) -(* iApply state_step_coupl_ret. *) -(* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2' sample) _) with "Hε_supply") as "[Hε_supply Hε]"; auto. *) -(* iMod (tapeN_update_append _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". *) -(* iMod (hocap_tapes_presample _ _ _ _ _ (fin_to_nat sample) with "[$][$]") as "[H4 Hfrag]". *) -(* iMod "Hclose'" as "_". *) -(* iMod ("Hvs" $! _ (fin_to_nat sample)%nat with "[$]") as "[%|[H2 HT]]". *) -(* { iExFalso. iApply (ec_contradict with "[$]"). *) -(* simpl. *) -(* rewrite /ε2'. *) -(* by rewrite bool_decide_eq_true_2; last lia. *) -(* } *) -(* simpl. *) -(* iMod ("Hclose" with "[$Hε H2 Htape H3 $H4 $H5 $H6]") as "_". *) -(* { iNext. *) -(* rewrite /ε2'. rewrite bool_decide_eq_true_2; last lia. *) -(* iFrame. *) -(* rewrite big_sepM_insert_delete; iFrame. simpl. iPureIntro. *) -(* split; last done. *) -(* rewrite filter_app. f_equal. *) -(* rewrite filter_cons. *) -(* rewrite decide_True; first done. *) -(* rewrite /filter_f. by apply bool_decide_pack. *) -(* } *) -(* iApply fupd_mask_intro_subseteq; first set_solver. *) -(* iFrame. *) -(* - (* we sampled a 4, and hence löb *) *) -(* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2' sample) _) with "Hε_supply") as "[Hε_supply Hε]"; auto. *) -(* iMod (tapeN_update_append _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". *) -(* iMod "Hclose'" as "_". *) -(* simpl. *) -(* rewrite {2}/ε2'. *) -(* rewrite bool_decide_eq_false_2; last lia. *) -(* rewrite bool_decide_eq_true_2; last lia. *) -(* iDestruct (ec_split with "[$]") as "[Hε Hampl]"; simpl; [naive_solver|lra|]. *) -(* simpl. *) -(* iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4 $H5 $H6]") as "_". *) -(* { iNext. *) -(* rewrite (big_sepM_delete _ m); [iFrame|done]. *) -(* simpl. iPureIntro; split; last done. *) -(* rewrite filter_app filter_cons decide_False; simpl. *) -(* - by rewrite filter_nil app_nil_r. *) -(* - rewrite /filter_f. case_bool_decide; [lia|auto]. *) -(* } *) -(* iDestruct ("IH" with "[$][$][$]") as "IH'". *) -(* by iMod ("IH'" $! (state_upd_tapes <[_:=_]> _) with "[$]") as "IH'". *) -(* - pose proof fin_to_nat_lt sample; lia. *) -(* Qed. *) + Lemma allocate_tape_spec3 N E c γ1 γ2: + ↑N ⊆ E-> + {{{ is_counter3 N c γ1 γ2 }}} + allocate_tape3 #() @ E + {{{ (v:val), RET v; + ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ (∃ ls, ⌜filter filter_f ls = []⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) + }}}. + Proof. + Admitted. + (* iIntros (Hsubset Φ) "#Hinv HΦ". *) + (* rewrite /allocate_tape3. *) + (* wp_pures. *) + (* wp_alloctape α as "Hα". *) + (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* iDestruct (hocap_tapes_notin3 with "[$][$]") as "%". *) + (* iMod (hocap_tapes_new with "[$]") as "[H4 H7]"; first done. *) + (* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Hα]") as "_". *) + (* { iNext. iSplitL; last done. *) + (* rewrite big_sepM_insert; [simpl; iFrame|done]. *) + (* by rewrite filter_nil. *) + (* } *) + (* iApply "HΦ". *) + (* by iFrame. *) + (* Qed. *) -(* Lemma read_counter_spec3 N E c γ1 γ2 γ3 P Q: *) -(* ↑N ⊆ E -> *) -(* {{{ inv N (counter_inv_pred3 c γ1 γ2 γ3) ∗ *) -(* □ (∀ (z:nat), P ∗ own γ3 (●F z) ={E∖↑N}=∗ *) -(* own γ3 (●F z)∗ Q z) *) -(* ∗ P *) -(* }}} *) -(* read_counter3 c @ E *) -(* {{{ (n':nat), RET #n'; Q n' *) -(* }}}. *) -(* Proof. *) -(* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP) HΦ". *) -(* rewrite /read_counter3. *) -(* wp_pure. *) -(* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) -(* wp_load. *) -(* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) -(* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]"); first done. *) -(* iApply ("HΦ" with "[$]"). *) -(* Qed. *) + Lemma incr_counter_tape_spec_some3 N E c γ1 γ2 (Q:nat->iProp Σ) (α:loc) n ns: + ↑N⊆E -> + {{{ is_counter3 N c γ1 γ2 ∗ + (∃ ls, ⌜filter filter_f ls = n::ns⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) ∗ + ( ∀ (z:nat), own γ2 (●F z) ={E∖↑N}=∗ + own γ2 (●F (z+n)) ∗ Q z) + + }}} + incr_counter_tape3 c #lbl:α @ E + {{{ (z:nat), RET (#z, #n); + (∃ ls, ⌜filter filter_f ls = ns⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) ∗ + Q z }}}. + Proof. + Admitted. + (* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP & Hα) HΦ". *) + (* rewrite /incr_counter_tape3. *) + (* iLöb as "IH". *) + (* wp_pures. *) + (* wp_bind (rand(_) _)%E. *) + (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) + (* erewrite <-(insert_delete m) at 1; last done. *) + (* rewrite big_sepM_insert; last apply lookup_delete. *) + (* simpl. *) + (* iDestruct "H3" as "[(%ls & %Hls & Htape) H3]". *) + (* destruct ls as [|x ls]. *) + (* { rewrite filter_nil in Hls. simplify_eq. } *) + (* wp_apply (wp_rand_tape with "[$]") as "[Htape %Hineq]". *) + (* rewrite Nat.le_lteq in Hineq. *) + (* destruct Hineq as [? | ->]. *) + (* - (* first value is valid *) *) + (* iMod (hocap_tapes_pop with "[$][$]") as "[H4 Hα]". *) + (* rewrite filter_cons /filter_f in Hls. *) + (* rewrite bool_decide_eq_true_2 in Hls; last done. simpl in *. *) + (* simplify_eq. *) + (* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". *) + (* { iSplitL; last done. *) + (* erewrite <-(insert_delete m) at 2; last done. *) + (* iNext. *) + (* rewrite insert_insert. *) + (* rewrite big_sepM_insert; last apply lookup_delete. by iFrame. *) + (* } *) + (* iModIntro. *) + (* wp_pures. *) + (* rewrite bool_decide_eq_true_2; last lia. *) + (* clear -Hsubset. *) + (* wp_pures. *) + (* wp_bind (FAA _ _). *) + (* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* wp_faa. *) + (* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) + (* replace (#(z+n)) with (#(z+n)%nat); last first. *) + (* { by rewrite Nat2Z.inj_add. } *) + (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) + (* iModIntro. wp_pures. *) + (* iApply "HΦ". *) + (* by iFrame. *) + (* - (* we get a 5, do iLöb induction *) *) + (* rewrite filter_cons /filter_f in Hls. *) + (* rewrite bool_decide_eq_false_2 in Hls; last lia. simpl in *. *) + (* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". *) + (* { iSplitL; last done. *) + (* erewrite <-(insert_delete m) at 2; last done. *) + (* iNext. *) + (* rewrite big_sepM_insert; last apply lookup_delete. by iFrame. *) + (* } *) + (* iModIntro. *) + (* do 4 wp_pure. *) + (* by iApply ("IH" with "[$][$]"). *) + (* Qed. *) + + Lemma counter_presample_spec3 NS E T γ1 γ2 c α ns: + ↑NS ⊆ E -> + is_counter3 NS c γ1 γ2 -∗ + (∃ ls, ⌜filter filter_f ls = ns⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) -∗ + ( |={E∖↑NS,∅}=> + ∃ ε ε2 num, + ↯ ε ∗ + ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ + ⌜(SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R⌝ ∗ + (∀ ns', ↯ (ε2 ns') ={∅,E∖↑NS}=∗ + T ε ε2 num ns' + ))-∗ + wp_update E (∃ ε ε2 num ns', (∃ ls, ⌜filter filter_f ls = ns++ns'⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) ∗ T ε ε2 num ns'). + Proof. + Admitted. + (* iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) + (* iMod wp_update_epsilon_err as "(%epsilon_err&%H&?)". *) + (* iApply wp_update_state_step_coupl. *) + (* iRevert "HP Hfrag". *) + (* iApply (ec_ind_amp _ 5%R with "[][$]"); [done|lra|]. *) + (* iModIntro. *) + (* clear epsilon_err H. *) + (* iIntros (epsilon_err ?) "#IH Hepsilon_err HP Hfrag". *) + (* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) + (* iMod (inv_acc with "Hinv") as "[>(%ε0 & % & % & % & H1 & H2 & H3 & H4 & -> & H5 & H6) Hclose]"; [done|]. *) + (* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) + (* erewrite <-(insert_delete m) at 1; last done. *) + (* rewrite big_sepM_insert; last apply lookup_delete. *) + (* simpl. *) + (* iDestruct "H3" as "[(%ls & %Hls & Htape) H3]". *) + (* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%)". *) + (* iDestruct (ec_valid with "H1") as "%". *) + (* iDestruct (ec_combine with "[$]") as "Htotal". *) + (* iDestruct (ec_supply_bound with "[$][$]") as "%". *) + (* iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. *) + (* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) + (* pose (ε2' := λ (x:nat), (if bool_decide (x<=3)%nat then ε2 ε0 x else *) + (* if bool_decide (x=4)%nat then ε0 + 5 * epsilon_err *) + (* else 0)%R *) + (* ). *) + (* iApply state_step_coupl_state_adv_comp_con_prob_lang; first done. *) + (* assert (forall x, 0<=ε2' x)%R. *) + (* { intros. rewrite /ε2'. repeat case_bool_decide; try done; first naive_solver. *) + (* apply Rplus_le_le_0_compat; simpl; [naive_solver|lra]. } *) + (* unshelve iExists (λ x, mknonnegreal (ε2' x) _); auto. *) + (* iSplit. *) + (* { iPureIntro. *) + (* unshelve epose proof (Hineq ε0 _) as H'; first (by naive_solver). *) + (* rewrite SeriesC_nat_bounded_fin in H'. *) + (* replace (INR 5%nat) with 5%R by (simpl; lra). *) + (* replace (INR 4%nat) with 4%R in H' by (simpl; lra). *) + (* rewrite /=/ε2' Hε1'. *) + (* rewrite SeriesC_finite_foldr/=. *) + (* rewrite SeriesC_finite_foldr/= in H'. *) + (* lra. *) + (* } *) + (* iIntros (sample). *) + (* simpl. *) + (* destruct (Rlt_decision (nonneg ε_rem + (ε2' sample))%R 1%R) as [Hdec|Hdec]; last first. *) + (* { apply Rnot_lt_ge, Rge_le in Hdec. *) + (* iApply state_step_coupl_ret_err_ge_1. *) + (* simpl. simpl in *. lra. *) + (* } *) + (* destruct (Nat.lt_total (fin_to_nat sample) 4%nat) as [|[|]]. *) + (* - (* we sample a value less than 4*) *) + (* iApply state_step_coupl_ret. *) + (* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2' sample) _) with "Hε_supply") as "[Hε_supply Hε]"; auto. *) + (* iMod (tapeN_update_append _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". *) + (* iMod (hocap_tapes_presample _ _ _ _ _ (fin_to_nat sample) with "[$][$]") as "[H4 Hfrag]". *) + (* iMod "Hclose'" as "_". *) + (* iMod ("Hvs" $! _ (fin_to_nat sample)%nat with "[$]") as "[%|[H2 HT]]". *) + (* { iExFalso. iApply (ec_contradict with "[$]"). *) + (* simpl. *) + (* rewrite /ε2'. *) + (* by rewrite bool_decide_eq_true_2; last lia. *) + (* } *) + (* simpl. *) + (* iMod ("Hclose" with "[$Hε H2 Htape H3 $H4 $H5 $H6]") as "_". *) + (* { iNext. *) + (* rewrite /ε2'. rewrite bool_decide_eq_true_2; last lia. *) + (* iFrame. *) + (* rewrite big_sepM_insert_delete; iFrame. simpl. iPureIntro. *) + (* split; last done. *) + (* rewrite filter_app. f_equal. *) + (* rewrite filter_cons. *) + (* rewrite decide_True; first done. *) + (* rewrite /filter_f. by apply bool_decide_pack. *) + (* } *) + (* iApply fupd_mask_intro_subseteq; first set_solver. *) + (* iFrame. *) + (* - (* we sampled a 4, and hence löb *) *) + (* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2' sample) _) with "Hε_supply") as "[Hε_supply Hε]"; auto. *) + (* iMod (tapeN_update_append _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". *) + (* iMod "Hclose'" as "_". *) + (* simpl. *) + (* rewrite {2}/ε2'. *) + (* rewrite bool_decide_eq_false_2; last lia. *) + (* rewrite bool_decide_eq_true_2; last lia. *) + (* iDestruct (ec_split with "[$]") as "[Hε Hampl]"; simpl; [naive_solver|lra|]. *) + (* simpl. *) + (* iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4 $H5 $H6]") as "_". *) + (* { iNext. *) + (* rewrite (big_sepM_delete _ m); [iFrame|done]. *) + (* simpl. iPureIntro; split; last done. *) + (* rewrite filter_app filter_cons decide_False; simpl. *) + (* - by rewrite filter_nil app_nil_r. *) + (* - rewrite /filter_f. case_bool_decide; [lia|auto]. *) + (* } *) + (* iDestruct ("IH" with "[$][$][$]") as "IH'". *) + (* by iMod ("IH'" $! (state_upd_tapes <[_:=_]> _) with "[$]") as "IH'". *) + (* - pose proof fin_to_nat_lt sample; lia. *) + (* Qed. *) + + Lemma read_counter_spec3 N E c γ1 γ2 Q: + ↑N ⊆ E -> + {{{ is_counter3 N c γ1 γ2 ∗ + (∀ (z:nat), own γ2 (●F z) ={E∖↑N}=∗ + own γ2 (●F z) ∗ Q z) + + }}} + read_counter3 c @ E + {{{ (n':nat), RET #n'; Q n' + }}}. + Proof. + Admitted. + (* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP) HΦ". *) + (* rewrite /read_counter3. *) + (* wp_pure. *) + (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) + (* wp_load. *) + (* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) + (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]"); first done. *) + (* iApply ("HΦ" with "[$]"). *) + (* Qed. *) -(* End impl3. *) +End impl3. +Program Definition random_counter3 `{F:rand_spec}: random_counter := + {| new_counter := new_counter3; + allocate_tape:= allocate_tape3; + incr_counter_tape := incr_counter_tape3; + read_counter:=read_counter3; + counterG := counterG3; + tape_name := rand_tape_name; + counter_name :=gname; + is_counter _ N c γ1 γ2 := is_counter3 N c γ1 γ2; + counter_tapes_auth _ γ m := (∃ (m':gmap loc (nat*list nat)), + ⌜(dom m = dom m')⌝ ∗ + ⌜∀ l xs' tb, m'!!l=Some (tb, xs') → tb = 4 ∧ m!!l=Some (filter filter_f xs')⌝ ∗ + rand_tapes_auth (L:=counterG3_randG) γ m')%I; + counter_tapes_frag _ γ α ns := + (∃ ls, ⌜filter filter_f ls = ns⌝ ∗ rand_tapes_frag (L:=counterG3_randG) γ α (4, ls))%I; + counter_content_auth _ γ z := own γ (●F z); + counter_content_frag _ γ f z := own γ (◯F{f} z); + new_counter_spec _ := new_counter_spec3; + allocate_tape_spec _ :=allocate_tape_spec3; + incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some3; + counter_presample_spec _ :=counter_presample_spec3; + read_counter_spec _ :=read_counter_spec3 + |}. -(* Program Definition random_counter3 `{!conerisGS Σ}: random_counter := *) -(* {| new_counter := new_counter3; *) -(* allocate_tape:= allocate_tape3; *) -(* incr_counter_tape := incr_counter_tape3; *) -(* read_counter:=read_counter3; *) -(* counterG := counterG3; *) -(* error_name := gname; *) -(* tape_name := gname; *) -(* counter_name :=gname; *) -(* is_counter _ N c γ1 γ2 γ3 := inv N (counter_inv_pred3 c γ1 γ2 γ3); *) -(* counter_error_auth _ γ x := ●↯ x @ γ; *) -(* counter_error_frag _ γ x := ◯↯ x @ γ; *) -(* counter_tapes_auth _ γ m := (●((λ ns, (3, ns))<$>m)@γ)%I; *) -(* counter_tapes_frag _ γ α ns := (α◯↪N (3%nat;ns) @ γ)%I; *) -(* counter_content_auth _ γ z := own γ (●F z); *) -(* counter_content_frag _ γ f z := own γ (◯F{f} z); *) -(* new_counter_spec _ := new_counter_spec3; *) -(* allocate_tape_spec _ :=allocate_tape_spec3; *) -(* incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some3; *) -(* counter_presample_spec _ :=counter_presample_spec3; *) -(* read_counter_spec _ :=read_counter_spec3 *) -(* |}. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (??????) "(%&<-&H1)(%&<-&H2)". *) -(* iCombine "H1 H2" gives "%H". by rewrite excl_auth.excl_auth_auth_op_valid in H. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (??????) "(%&<-&H1)(%&<-&H2)". *) -(* iCombine "H1 H2" gives "%H". by rewrite excl_auth.excl_auth_frag_op_valid in H. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (?????) "H". *) -(* iApply (hocap_error_auth_pos with "[$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (?????) "H". *) -(* iApply (hocap_error_frag_pos with "[$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (??????) "H1 H2". *) -(* iApply (hocap_error_agree with "[$][$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. iIntros (???????) "??". *) -(* iApply (hocap_error_update with "[$][$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (??????) "H1 H2". *) -(* by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[%H _]". *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (???????) "H1 H2". *) -(* iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. *) -(* rewrite dfrac_op_own dfrac_valid_own. by intro. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros. *) -(* iDestruct (hocap_tapes_agree with "[$][$]") as "%H". *) -(* iPureIntro. *) -(* rewrite lookup_fmap_Some in H. destruct H as (?&?&?). *) -(* by simplify_eq. *) -(* Qed. *) -(* Next Obligation. *) -(* iIntros. *) -(* iMod (hocap_tapes_presample with "[$][$]") as "[??]". *) -(* iModIntro. iFrame. *) -(* by rewrite fmap_insert. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (??????) "H1 H2". *) -(* iCombine "H1 H2" gives "%H". by rewrite auth_auth_op_valid in H. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. iIntros (???? z z' ?) "H1 H2". *) -(* iCombine "H1 H2" gives "%H". *) -(* apply frac_auth_included_total in H. iPureIntro. *) -(* by apply nat_included. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (????????). *) -(* rewrite frac_auth_frag_op. by rewrite own_op. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. iIntros (??????) "H1 H2". *) -(* iCombine "H1 H2" gives "%H". *) -(* iPureIntro. *) -(* by apply frac_auth_agree_L in H. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. iIntros (????????) "H1 H2". *) -(* iMod (own_update_2 _ _ _ (_ ⋅ _) with "[$][$]") as "[$$]"; last done. *) -(* apply frac_auth_update. *) -(* apply nat_local_update. lia. *) -(* Qed. *) - +Next Obligation. + simpl. + iIntros (???????) "(%&%&%&?) (%&%&%&?)". + iApply (rand_tapes_auth_exclusive with "[$][$]"). +Qed. +Next Obligation. + simpl. + iIntros (????????) "(%&%&?) (%&%&?)". + iApply (rand_tapes_frag_exclusive with "[$][$]"). +Qed. +Next Obligation. + simpl. + iIntros (????????) "(%&%&%K&?) (%&%&?)". + iDestruct (rand_tapes_agree γ α with "[$][$]") as "%K'". + iPureIntro. + apply K in K'. subst. naive_solver. +Qed. +Next Obligation. + simpl. + iIntros (???????) "(%&%&?)". + iPureIntro. + subst. + induction ls; first done. + rewrite filter_cons. + case_decide as H; last done. + rewrite Forall_cons_iff; split; last done. + rewrite /filter_f in H. lia. +Qed. +Next Obligation. + simpl. + iIntros (??????????) "(%&%&%&?) (%&%&?)". + iMod (rand_tapes_update _ _ _ _ (_,ns') with "[$][$]") as "[??]"; last iFrame. + - eapply Forall_impl; first done. simpl; lia. + - iPureIntro. split; [split; first (rewrite !dom_insert_L; set_solver)|]. + + intros ? xs' ? K. + rewrite lookup_insert_Some in K. + destruct K as [[??]|[??]]; simplify_eq. + * split; first done. + rewrite lookup_insert_Some; left. + split; first done. + by erewrite Forall_filter_f. + * rewrite lookup_insert_ne; last done. + naive_solver. + + by apply Forall_filter_f. +Qed. +Next Obligation. + simpl. + iIntros (???????) "H1 H2". + iCombine "H1 H2" gives "%K". by rewrite auth_auth_op_valid in K. +Qed. +Next Obligation. + simpl. iIntros (????? z z' ?) "H1 H2". + iCombine "H1 H2" gives "%K". + apply frac_auth_included_total in K. iPureIntro. + by apply nat_included. +Qed. +Next Obligation. + simpl. + iIntros (?????????). + rewrite frac_auth_frag_op. by rewrite own_op. +Qed. +Next Obligation. + simpl. iIntros (???????) "H1 H2". + iCombine "H1 H2" gives "%K". + iPureIntro. + by apply frac_auth_agree_L in K. +Qed. +Next Obligation. + simpl. iIntros (?????????) "H1 H2". + iMod (own_update_2 _ _ _ (_ ⋅ _) with "[$][$]") as "[$$]"; last done. + apply frac_auth_update. + apply nat_local_update. lia. +Qed. From 746032ee2611bcb119fcb85ceb8c2a6cd4588074 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Tue, 1 Oct 2024 10:30:23 +0200 Subject: [PATCH 58/69] impl3 create spec, allocate tape, and read from counter --- .../coneris/examples/random_counter/impl3.v | 82 ++++++++----------- 1 file changed, 36 insertions(+), 46 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl3.v b/theories/coneris/examples/random_counter/impl3.v index 59b1b94d..17018e82 100644 --- a/theories/coneris/examples/random_counter/impl3.v +++ b/theories/coneris/examples/random_counter/impl3.v @@ -36,9 +36,9 @@ Section impl3. Context `{H:conerisGS Σ, r1:@rand_spec Σ H, L:randG Σ, !inG Σ (frac_authR natR)}. Definition new_counter3 : val:= λ: "_", ref #0. - Definition allocate_tape3 : val := λ: "_", rand_allocate_tape #4. + Definition allocate_tape3 : val := λ: "_", rand_allocate_tape #4%nat. Definition incr_counter_tape3 :val := rec: "f" "l" "α":= - let: "n" := rand_tape "α" #4 in + let: "n" := rand_tape "α" #4%nat in if: "n" < #4 then (FAA "l" "n", "n") else "f" "l" "α". @@ -67,24 +67,20 @@ Section impl3. ∃ γ1 γ2, is_counter3 N c γ1 γ2 ∗ own γ2 (◯F 0%nat) }}}. Proof. - Admitted. - (* rewrite /new_counter3. *) - (* iIntros (Φ) "Hε HΦ". *) - (* wp_pures. *) - (* wp_alloc l as "Hl". *) - (* iDestruct (ec_valid with "[$]") as "%". *) - (* unshelve iMod (hocap_error_alloc (mknonnegreal ε _)) as "[%γ1 [H1 H2]]". *) - (* { lra. } *) - (* simpl. *) - (* iMod (hocap_tapes_alloc (∅:gmap _ _)) as "[%γ2 [H3 H4]]". *) - (* iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ3[H5 H6]]". *) - (* { by apply frac_auth_valid. } *) - (* replace (#0) with (#0%nat) by done. *) - (* iMod (inv_alloc N _ (counter_inv_pred3 (#l) γ1 γ2 γ3) with "[$Hε $Hl $H1 $H3 $H5]") as "#Hinv". *) - (* { iSplit; last done. by iApply big_sepM_empty. } *) - (* iApply "HΦ". *) - (* iExists _, _, _. by iFrame. *) - (* Qed. *) + rewrite /new_counter3. + iIntros (Φ) "_ HΦ". + wp_pures. + iMod (rand_inv_create_spec) as "(%γ1 & #Hinv)". + wp_alloc l as "Hl". + iMod (own_alloc (●F 0%nat ⋅ ◯F 0%nat)) as "[%γ2[H5 H6]]". + { by apply frac_auth_valid. } + replace (#0) with (#0%nat) by done. + iMod (inv_alloc _ _ (counter_inv_pred3 (#l) γ2) with "[$Hl $H5]") as "#Hinv'"; first done. + iApply "HΦ". + iFrame. + iModIntro. + iExists _. by iSplit. + Qed. Lemma allocate_tape_spec3 N E c γ1 γ2: ↑N ⊆ E-> @@ -94,22 +90,13 @@ Section impl3. ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ (∃ ls, ⌜filter filter_f ls = []⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) }}}. Proof. - Admitted. - (* iIntros (Hsubset Φ) "#Hinv HΦ". *) - (* rewrite /allocate_tape3. *) - (* wp_pures. *) - (* wp_alloctape α as "Hα". *) - (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) - (* iDestruct (hocap_tapes_notin3 with "[$][$]") as "%". *) - (* iMod (hocap_tapes_new with "[$]") as "[H4 H7]"; first done. *) - (* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Hα]") as "_". *) - (* { iNext. iSplitL; last done. *) - (* rewrite big_sepM_insert; [simpl; iFrame|done]. *) - (* by rewrite filter_nil. *) - (* } *) - (* iApply "HΦ". *) - (* by iFrame. *) - (* Qed. *) + iIntros (Hsubset Φ) "[#Hinv #Hinv'] HΦ". + rewrite /allocate_tape3. + wp_pures. + wp_apply rand_allocate_tape_spec as (v) "(%α & -> & Hfrag)"; [by eapply nclose_subseteq'|done|]. + iApply "HΦ". + by iFrame. + Qed. Lemma incr_counter_tape_spec_some3 N E c γ1 γ2 (Q:nat->iProp Σ) (α:loc) n ns: ↑N⊆E -> @@ -306,16 +293,19 @@ Section impl3. {{{ (n':nat), RET #n'; Q n' }}}. Proof. - Admitted. - (* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP) HΦ". *) - (* rewrite /read_counter3. *) - (* wp_pure. *) - (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) - (* wp_load. *) - (* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) - (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]"); first done. *) - (* iApply ("HΦ" with "[$]"). *) - (* Qed. *) + iIntros (Hsubset Φ) "([#Hinv #Hinv'] & Hvs) HΦ". + rewrite /read_counter3. + wp_pure. + iInv "Hinv'" as ">( %l & %z & -> & H5 & H6)" "Hclose". + wp_load. + iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". + { apply difference_mono_l. + by apply nclose_subseteq'. } + iMod ("Hvs" with "[$]") as "[? HQ]". + iMod "Hclose'" as "_". + iMod ("Hclose" with "[-HQ HΦ]"); first by iFrame. + iApply ("HΦ" with "[$]"). + Qed. End impl3. From 73f386850b989fb1e4efd1b3c6239ead183d5a55 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Tue, 1 Oct 2024 10:40:18 +0200 Subject: [PATCH 59/69] incr_counter_tape_spec_some3 --- .../coneris/examples/random_counter/impl3.v | 89 +++++++------------ 1 file changed, 31 insertions(+), 58 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl3.v b/theories/coneris/examples/random_counter/impl3.v index 17018e82..ba445c72 100644 --- a/theories/coneris/examples/random_counter/impl3.v +++ b/theories/coneris/examples/random_counter/impl3.v @@ -111,64 +111,37 @@ Section impl3. (∃ ls, ⌜filter filter_f ls = ns⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) ∗ Q z }}}. Proof. - Admitted. - (* iIntros (Hsubset Φ) "(#Hinv & #Hvs & HP & Hα) HΦ". *) - (* rewrite /incr_counter_tape3. *) - (* iLöb as "IH". *) - (* wp_pures. *) - (* wp_bind (rand(_) _)%E. *) - (* iInv N as ">(%ε & %m & %l & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) - (* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) - (* erewrite <-(insert_delete m) at 1; last done. *) - (* rewrite big_sepM_insert; last apply lookup_delete. *) - (* simpl. *) - (* iDestruct "H3" as "[(%ls & %Hls & Htape) H3]". *) - (* destruct ls as [|x ls]. *) - (* { rewrite filter_nil in Hls. simplify_eq. } *) - (* wp_apply (wp_rand_tape with "[$]") as "[Htape %Hineq]". *) - (* rewrite Nat.le_lteq in Hineq. *) - (* destruct Hineq as [? | ->]. *) - (* - (* first value is valid *) *) - (* iMod (hocap_tapes_pop with "[$][$]") as "[H4 Hα]". *) - (* rewrite filter_cons /filter_f in Hls. *) - (* rewrite bool_decide_eq_true_2 in Hls; last done. simpl in *. *) - (* simplify_eq. *) - (* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". *) - (* { iSplitL; last done. *) - (* erewrite <-(insert_delete m) at 2; last done. *) - (* iNext. *) - (* rewrite insert_insert. *) - (* rewrite big_sepM_insert; last apply lookup_delete. by iFrame. *) - (* } *) - (* iModIntro. *) - (* wp_pures. *) - (* rewrite bool_decide_eq_true_2; last lia. *) - (* clear -Hsubset. *) - (* wp_pures. *) - (* wp_bind (FAA _ _). *) - (* iInv N as ">(%ε & %m & % & %z & H1 & H2 & H3 & H4 & -> & H5 & H6)" "Hclose". *) - (* wp_faa. *) - (* iMod ("Hvs" with "[$]") as "[H6 HQ]". *) - (* replace (#(z+n)) with (#(z+n)%nat); last first. *) - (* { by rewrite Nat2Z.inj_add. } *) - (* iMod ("Hclose" with "[$H1 $H2 $H3 $H4 $H5 $H6]") as "_"; first done. *) - (* iModIntro. wp_pures. *) - (* iApply "HΦ". *) - (* by iFrame. *) - (* - (* we get a 5, do iLöb induction *) *) - (* rewrite filter_cons /filter_f in Hls. *) - (* rewrite bool_decide_eq_false_2 in Hls; last lia. simpl in *. *) - (* iMod ("Hclose" with "[$H1 $H2 H3 $H4 $H5 $H6 Htape]") as "_". *) - (* { iSplitL; last done. *) - (* erewrite <-(insert_delete m) at 2; last done. *) - (* iNext. *) - (* rewrite big_sepM_insert; last apply lookup_delete. by iFrame. *) - (* } *) - (* iModIntro. *) - (* do 4 wp_pure. *) - (* by iApply ("IH" with "[$][$]"). *) - (* Qed. *) - + iIntros (Hsubset Φ) "([#Hinv #Hinv'] & (%ls & %Hfilter & Hfrag) & Hvs) HΦ". + rewrite /incr_counter_tape3. + iLöb as "IH" forall (ls Hfilter Φ) "Hfrag". + wp_pures. + destruct ls as [|hd ls]; first simplify_eq. + wp_apply (rand_tape_spec_some with "[$]") as "Hfrag"; first by eapply nclose_subseteq'. + wp_pures. + case_bool_decide as K. + - wp_pures. + wp_bind (FAA _ _). + iInv "Hinv'" as ">(%&%&-> & ?&?)" "Hclose". + wp_faa. + iMod (fupd_mask_subseteq (E ∖ ↑N)) as "Hclose'". + + apply difference_mono_l. + by apply nclose_subseteq'. + + iMod ("Hvs" with "[$]") as "[? HQ]". + iMod "Hclose'" as "_". + rewrite -Nat2Z.inj_add. + rewrite filter_cons_True in Hfilter; last (rewrite /filter_f; lia). + simplify_eq. + iMod ("Hclose" with "[-HQ Hfrag HΦ]") as "_"; first by iFrame. + iModIntro. + wp_pures. + iApply "HΦ". by iFrame. + - wp_pure. + iApply ("IH" with "[][$][$][$]"). + iPureIntro. + rewrite filter_cons_False in Hfilter; first done. + rewrite /filter_f. lia. + Qed. + Lemma counter_presample_spec3 NS E T γ1 γ2 c α ns: ↑NS ⊆ E -> is_counter3 NS c γ1 γ2 -∗ From a0101d2cd61e36b334008204b2ff24d759ed2ea0 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Tue, 1 Oct 2024 11:17:34 +0200 Subject: [PATCH 60/69] NIT --- .../coneris/examples/random_counter/impl3.v | 21 +++++++++++++++++++ 1 file changed, 21 insertions(+) diff --git a/theories/coneris/examples/random_counter/impl3.v b/theories/coneris/examples/random_counter/impl3.v index ba445c72..3470a51c 100644 --- a/theories/coneris/examples/random_counter/impl3.v +++ b/theories/coneris/examples/random_counter/impl3.v @@ -156,6 +156,27 @@ Section impl3. ))-∗ wp_update E (∃ ε ε2 num ns', (∃ ls, ⌜filter filter_f ls = ns++ns'⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) ∗ T ε ε2 num ns'). Proof. + iIntros (Hsubset) "[#Hinv #Hinv'] (%ls & %Hfilter & Hfrag) Hvs". + iMod wp_update_epsilon_err as "(%eps&%Heps&Heps)". + (* iRevert (ls Hfilter) "Hfrag Hvs". *) + (* iApply (ec_ind_amp _ 5%R with "[][$]"); [done|lra|]. *) + (* clear eps Heps. *) + (* iModIntro. *) + (* iIntros (eps Heps) "#IH Heps %ls %Hfilter Hfrag Hvs". *) + iMod (rand_presample_spec _ _ _ _ _ _ (λ ε num ε2 ns', + _ + )%I with "[//][$][Heps Hvs]") as "H"; first by apply nclose_subseteq'. + - iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose". + { apply difference_mono_l. + by apply nclose_subseteq'. } + iMod "Hvs" as "(%ε & %ε2 & %num & Herr &%Hpos & %Hineq & Hrest)". + iFrame. + admit. + - iDestruct "H" as "(%ε & %num & %ε2 & %ns' & Hfrag & HT)". + iModIntro. + iFrame. + rewrite filter_app. + admit. Admitted. (* iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) (* iMod wp_update_epsilon_err as "(%epsilon_err&%H&?)". *) From 644c7bfe1d603f5c88af9fb9a3c79b8beeb55581 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Tue, 1 Oct 2024 12:32:51 +0200 Subject: [PATCH 61/69] add linear spec for hocap flip --- .../coneris/examples/random_counter/impl2.v | 71 -- .../coneris/examples/random_counter/impl3.v | 21 - theories/coneris/lib/hocap.v | 159 ---- theories/coneris/lib/hocap_flip.v | 95 +-- theories/coneris/lib/hocap_rand.v | 769 ------------------ 5 files changed, 38 insertions(+), 1077 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index e416fb06..98b1cee0 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -124,77 +124,6 @@ Proof. f_equal. rewrite IH; [lia|done]. Qed. -(* Local Definition expander (l:list nat):= *) -(* l ≫= (λ x, [Nat.div2 x; Nat.b2n (Nat.odd x)]). *) - -(* Class hocap_tapesGS' (Σ : gFunctors) := Hocap_tapesGS' { *) -(* hocap_tapesGS_inG' :: ghost_mapG Σ loc (bool* list nat) *) -(* }. *) -(* Definition hocap_tapesΣ' := ghost_mapΣ loc (bool *list nat). *) - -(* Notation "α ◯↪N ( b , ns ) @ γ":= (α ↪[ γ ] (b, ns))%I *) -(* (at level 20) : bi_scope. *) - -(* Notation "● m @ γ" := (ghost_map_auth γ 1 m) (at level 20) : bi_scope. *) - -(* Section tapes_lemmas. *) -(* Context `{!hocap_tapesGS' Σ}. *) - -(* Lemma hocap_tapes_alloc' m: *) -(* ⊢ |==>∃ γ, (● m @ γ) ∗ [∗ map] k↦v ∈ m, (k ◯↪N (v.1, v.2) @ γ). *) -(* Proof. *) -(* iMod ghost_map_alloc as (γ) "[??]". *) -(* iFrame. iModIntro. *) -(* iApply big_sepM_mono; last done. *) -(* by iIntros (?[??]). *) -(* Qed. *) - -(* Lemma hocap_tapes_agree' m b γ k ns: *) -(* (● m @ γ) -∗ (k ◯↪N (b, ns) @ γ) -∗ ⌜ m!!k = Some (b, ns) ⌝. *) -(* Proof. *) -(* iIntros "H1 H2". *) -(* by iCombine "H1 H2" gives "%". *) -(* Qed. *) - -(* Lemma hocap_tapes_new' γ m k ns b: *) -(* m!!k=None -> ⊢ (● m @ γ) ==∗ (● (<[k:=(b, ns)]>m) @ γ) ∗ (k ◯↪N (b, ns) @ γ). *) -(* Proof. *) -(* iIntros (Hlookup) "H". *) -(* by iApply ghost_map_insert. *) -(* Qed. *) - -(* Lemma hocap_tapes_presample' b γ m k ns n: *) -(* (● m @ γ) -∗ (k ◯↪N (b, ns) @ γ) ==∗ (● (<[k:=(b, ns++[n])]>m) @ γ) ∗ (k ◯↪N (b, ns++[n]) @ γ). *) -(* Proof. *) -(* iIntros "H1 H2". *) -(* iApply (ghost_map_update with "[$][$]"). *) -(* Qed. *) - -(* Lemma hocap_tapes_pop1' γ m k ns: *) -(* (● m @ γ) -∗ (k ◯↪N (true, ns) @ γ) ==∗ (● (<[k:=(false, ns)]>m) @ γ) ∗ (k ◯↪N (false, ns) @ γ). *) -(* Proof. *) -(* iIntros "H1 H2". *) -(* iApply (ghost_map_update with "[$][$]"). *) -(* Qed. *) - -(* Lemma hocap_tapes_pop2' γ m k ns n: *) -(* (● m @ γ) -∗ (k ◯↪N (false, n::ns) @ γ) ==∗ (● (<[k:=(true, ns)]>m) @ γ) ∗ (k ◯↪N (true, ns) @ γ). *) -(* Proof. *) -(* iIntros "H1 H2". *) -(* iApply (ghost_map_update with "[$][$]"). *) -(* Qed. *) - -(* Lemma hocap_tapes_notin' `{F: flip_spec Σ} {L:flipG Σ} α ns m γ (f:(bool*list nat)-> (list bool)) : *) -(* flip_tapes_frag (L:=L) γ α ns -∗ ([∗ map] α0↦t ∈ m, flip_tapes_frag (L:=L) γ α0 (f t)) -∗ ⌜m!!α=None ⌝. *) -(* Proof. *) -(* destruct (m!!α) eqn:Heqn; last by iIntros. *) -(* iIntros "Hα Hmap". *) -(* iDestruct (big_sepM_lookup with "[$]") as "?"; first done. *) -(* iExFalso. *) -(* by iDestruct (flip_tapes_frag_exclusive with "[$][$]") as "%". *) -(* Qed. *) - -(* End tapes_lemmas. *) Section impl2. Context `{F: flip_spec Σ}. diff --git a/theories/coneris/examples/random_counter/impl3.v b/theories/coneris/examples/random_counter/impl3.v index 3470a51c..02372bab 100644 --- a/theories/coneris/examples/random_counter/impl3.v +++ b/theories/coneris/examples/random_counter/impl3.v @@ -19,19 +19,6 @@ Section filter. Qed. End filter. -(* Section lemmas. *) -(* Context `{!conerisGS Σ}. *) -(* Lemma hocap_tapes_notin3 α N ns (m:gmap loc (nat*list nat)) : *) -(* α ↪N (N; ns) -∗ ([∗ map] α↦t ∈ m,∃ (ls:list nat), ⌜ (filter filter_f ls) = t.2⌝ ∗ α ↪N ( 4%nat ; ls)) -∗ ⌜m!!α=None ⌝. *) -(* Proof. *) -(* destruct (m!!α) eqn:Heqn; last by iIntros. *) -(* iIntros "Hα Hmap". *) -(* iDestruct (big_sepM_lookup with "[$]") as "(%&%&?)"; first done. *) -(* iExFalso. *) -(* iApply (tapeN_tapeN_contradict with "[$][$]"). *) -(* Qed. *) -(* End lemmas. *) - Section impl3. Context `{H:conerisGS Σ, r1:@rand_spec Σ H, L:randG Σ, !inG Σ (frac_authR natR)}. @@ -52,14 +39,6 @@ Section impl3. ((is_rand (L:=L) (N.@"rand") γ1) ∗ inv (N.@"counter") (counter_inv_pred3 c γ2))%I. - - (* (∃ (ε:R) (m:gmap loc (nat * list nat)) (l:loc) (z:nat), *) - (* ↯ ε ∗ ●↯ ε @ γ1 ∗ *) - (* ([∗ map] α ↦ t ∈ m, ∃ (ls:list nat), ⌜ (filter filter_f ls) = t.2⌝ ∗ α ↪N ( 4%nat ; ls) ) *) - (* ∗ ●m@γ2 ∗ *) - (* ⌜c=#l⌝ ∗ l ↦ #z ∗ own γ3 (●F z) *) - (* )%I. *) - Lemma new_counter_spec3 E N: {{{ True }}} new_counter3 #() @ E diff --git a/theories/coneris/lib/hocap.v b/theories/coneris/lib/hocap.v index 00cb1ec0..2613b48a 100644 --- a/theories/coneris/lib/hocap.v +++ b/theories/coneris/lib/hocap.v @@ -184,165 +184,6 @@ Section tapes_lemmas. iIntros "H1 H2". iApply (ghost_map_update with "[$][$]"). Qed. - - (* Lemma hocap_tapes_presample γ m k N ns n: *) - (* (● m @ γ) -∗ (k ◯↪N (N; ns) @ γ) ==∗ (● (<[k:=(N,ns++[n])]>m) @ γ) ∗ (k ◯↪N (N; ns++[n]) @ γ). *) - (* Proof. *) - (* iIntros "H1 H2". *) - (* iApply (ghost_map_update with "[$][$]"). *) - (* Qed. *) - - (* Lemma hocap_tapes_presample' γ m k N ns ns': *) - (* (● m @ γ) -∗ (k ◯↪N (N; ns) @ γ) ==∗ (● (<[k:=(N,ns++ns')]>m) @ γ) ∗ (k ◯↪N (N; ns++ns') @ γ). *) - (* Proof. *) - (* iIntros "H1 H2". *) - (* iApply (ghost_map_update with "[$][$]"). *) - (* Qed. *) - - (* Lemma hocap_tapes_pop γ m k N ns n: *) - (* (● m @ γ) -∗ (k ◯↪N (N; n::ns) @ γ) ==∗ (● (<[k:=(N,ns)]>m) @ γ) ∗ (k ◯↪N (N; ns) @ γ). *) - (* Proof. *) - (* iIntros "H1 H2". *) - (* iApply (ghost_map_update with "[$][$]"). *) - (* Qed. *) - End tapes_lemmas. - -(* Section HOCAP. *) - -(* Context `{!conerisGS Σ, !hocap_errorGS Σ}. *) - -(* Definition error_inv (γ :gname):= *) -(* inv hocap_error_nroot (∃ (ε:R), ↯ ε ∗ ●↯ ε @ γ). *) - -(* Lemma wp_hocap_rand_adv_comp (N : nat) z E *) -(* (ε2 : R -> fin (S N) -> R) *) -(* (P : iProp Σ) (Q : (fin (S N)) -> iProp Σ) γ: *) -(* TCEq N (Z.to_nat z) → *) -(* ↑hocap_error_nroot ⊆ E -> *) -(* (∀ ε n, (0<=ε -> 0<=ε2 ε n))%R -> *) -(* (∀ (ε:R), 0<=ε ->SeriesC (λ n, (1 / (S N)) * (ε2 ε n))%R <= (ε))%R → *) -(* {{{ error_inv γ∗ *) -(* □(∀ (ε:R) (n : fin (S N)), P ∗ ●↯ ε @ γ ={E∖↑hocap_error_nroot}=∗ *) -(* (⌜(1<=ε2 ε n)%R⌝ ∨(●↯ (ε2 ε n) @ γ ∗ Q (n))) ) ∗ *) -(* P }}} rand #z @ E {{{ n, RET #n; Q (n)}}}. *) -(* Proof. *) -(* iIntros (-> Hsubset ? Hineq) "%Φ [#Hinv [#Hchange HP]] HΦ". *) -(* iInv "Hinv" as ">(%ε & Hε & Hauth)" "Hclose". *) -(* iDestruct (ec_valid with "[$]") as "%". *) -(* wp_apply (wp_couple_rand_adv_comp1' _ _ _ _ (ε2 ε) with "[$]"); [naive_solver|naive_solver|]. *) -(* iIntros (n) "Hε". *) -(* iMod ("Hchange" $! _ n with "[$]") as "[%|[Hauth HQ]]"; last first. *) -(* { iMod ("Hclose" with "[Hauth Hε]") as "_". *) -(* - unshelve iExists (mknonnegreal _ _); last iFrame. *) -(* naive_solver. *) -(* - by iApply "HΦ". } *) -(* iExFalso. *) -(* by iApply (ec_contradict with "[$]"). *) -(* Qed. *) - -(* Lemma wp_hocap_flip_adv_comp E *) -(* (ε2 : R -> bool -> R) *) -(* (P : iProp Σ) (Q : bool -> iProp Σ) γ: *) -(* ↑hocap_error_nroot ⊆ E -> *) -(* (∀ ε b, 0<=ε -> 0<=ε2 ε b)%R -> *) -(* (∀ (ε:R), 0<=ε -> (((ε2 ε true) + (ε2 ε false))/2 <= (ε)))%R → *) -(* {{{ error_inv γ∗ *) -(* □(∀ (ε:R) (b : bool), P ∗ ●↯ ε @ γ ={E∖↑hocap_error_nroot}=∗ (⌜(1<=ε2 ε b)%R⌝ ∨ ●↯ (ε2 ε b) @ γ ∗ Q (b)) ) ∗ *) -(* P }}} flip @ E {{{ (b:bool), RET #b; Q (b)}}}. *) -(* Proof. *) -(* iIntros (Hsubset ? Hineq) "%Φ [#Hinv [#Hchange HP]] HΦ". *) -(* rewrite /flip/flipL. *) -(* wp_pures. *) -(* wp_apply (wp_hocap_rand_adv_comp _ _ _ (λ ε x, if fin_to_nat x =? 1%nat then ε2 ε true else ε2 ε false) P (λ x, Q (fin_to_nat x =? 1%nat)) with "[-HΦ]"); [done|..]. *) -(* - intros; case_match; naive_solver. *) -(* - intros ε. rewrite SeriesC_finite_foldr; simpl. specialize (Hineq ε). lra. *) -(* - iFrame. iSplitR; first iExact "Hinv". *) -(* iModIntro. *) -(* iIntros (ε n) "H". *) -(* iMod ("Hchange" $! _ (fin_to_nat n =? 1%nat)with "[$]") as "[%|[Hε HQ]]". *) -(* + iModIntro. iLeft. iPureIntro. by case_match. *) -(* + iModIntro. iRight. iFrame. by case_match. *) -(* - iIntros. *) -(* wp_apply (conversion.wp_int_to_bool); first done. *) -(* iIntros "_". *) -(* iApply "HΦ". *) -(* inv_fin n; simpl; first done. *) -(* intros n; inv_fin n; simpl; first done. *) -(* intros n; inv_fin n. *) -(* Qed. *) - -(* (* (** With tapes *) *) *) -(* (* Context `{!hocap_tapesGS Σ}. *) *) - -(* (* Definition tapes_inv (γ :gname):= *) *) -(* (* inv hocap_tapes_nroot (∃ m, ●m@γ ∗ [∗ map] α ↦ t ∈ m, α ↪N ( t.1 ; t.2 ) ). *) *) -(* (* Lemma wp_hocap_presample_adv_comp (N : nat) z E ns α *) *) -(* (* (ε2 : R -> fin (S N) -> R) *) *) -(* (* (P : iProp Σ) (Q : val-> iProp Σ) T γ γ': *) *) -(* (* TCEq N (Z.to_nat z) → *) *) -(* (* ↑hocap_error_nroot ⊆ E -> *) *) -(* (* ↑hocap_tapes_nroot ⊆ E -> *) *) -(* (* (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> *) *) -(* (* (∀ (ε:R), 0<= ε -> SeriesC (λ n, (1 / (S N)) * (ε2 ε n))%R <= ε)%R → *) *) -(* (* error_inv γ -∗ tapes_inv γ' -∗ *) *) -(* (* (□∀ (ε:R) m, (⌜m!!α = Some (N, ns)⌝ ∗ P ∗ ●↯ ε @ γ ∗ ●m@γ') *) *) -(* (* ={E∖↑hocap_error_nroot∖↑hocap_tapes_nroot}=∗ *) *) -(* (* ∃ n, (⌜(1<=ε2 ε n)%R⌝ ∨(●↯ (ε2 ε n) @ γ ∗ ●(<[α := (N, ns ++ [fin_to_nat n])]>m) @ γ' ∗ T (n)))) *) *) -(* (* -∗ P -∗ α ◯↪N (N; ns) @ γ' -∗ *) *) -(* (* wp_update E (∃ n, T (n) ∗ α◯↪N (N; ns++[fin_to_nat n]) @ γ'). *) *) -(* (* Proof. *) *) -(* (* iIntros (-> Hval Hsubset Hubset' Hineq) "#Hinv #Hinv' #Hshift HP Htape". *) *) -(* (* iApply fupd_wp_update. *) *) -(* (* iInv "Hinv" as ">(%ε' & Hε & Hauth)" "Hclose". *) *) -(* (* iInv "Hinv'" as ">(%m & Hm & Hmauth)" "Hclose'". *) *) -(* (* iDestruct (ec_valid with "[$]") as "%". *) *) -(* (* iApply wp_update_mono_fupd. iApply fupd_frame_l. *) *) -(* (* iSplitL "Hε". *) *) -(* (* - iApply (wp_update_presample_exp _ α _ ns ε' (ε2 ε) with "[$Hε]"). *) *) -(* (* + naive_solver. *) *) -(* (* + naive_solver. *) *) -(* (* + admit. *) *) -(* (* - iMod ("Hclose'" with "[$]"). *) *) -(* (* Abort. *) *) - -(* (* iApply (wp_presample_adv_comp); [done|exact|..]. *) *) -(* (* iApply fupd_frame_l. *) *) -(* (* - *) *) -(* (* Abort. *) *) - - - -(* End HOCAP. *) - -(* Section HOCAP_alt. *) -(* Context `{!conerisGS Σ}. *) - -(* Lemma wp_flip_exp_hocap (P : iProp Σ) (Q : bool → iProp Σ) E1 E2 : *) -(* □ (P ={E1, E2}=∗ *) -(* ∃ ε1 ε2, ⌜(( (ε2 true) + (ε2 false))/2 <= ε1)%R⌝ ∗ *) -(* ⌜(∀ b, 0<=ε2 b)%R⌝ ∗ *) -(* ↯ ε1 ∗ (∀ (b : bool), ↯ (ε2 b) ={E2, E1}=∗ Q b)) -∗ *) -(* {{{ P }}} flip @ E1 {{{ (b : bool), RET #b; Q b}}}. *) -(* Proof. *) -(* iIntros "#Hvs". iIntros (Ψ) "!# HP HΨ". *) -(* rewrite /flip /flipL. *) -(* wp_pures. *) -(* wp_bind (rand _)%E. *) -(* iMod ("Hvs" with "HP") as (ε1 ε2) "(% & % & Hε1 & HQ)". *) -(* set (ε2' := ε2 ∘ fin_to_bool). *) -(* iApply (wp_couple_rand_adv_comp1' _ _ _ _ ε2' with "Hε1"). *) -(* { intros; rewrite /ε2'. simpl. done. } *) -(* { rewrite SeriesC_finite_foldr /ε2' /=. lra. } *) -(* iIntros "!>" (n) "Hε2". *) -(* assert (ε2' n = ε2 (Z_to_bool n)) as ->. *) -(* { inv_fin n; [eauto|]. intros n. inv_fin n; [eauto|]. intros n. inv_fin n. } *) -(* iMod ("HQ" with "Hε2") as "HQ". iModIntro. *) -(* wp_apply (conversion.wp_int_to_bool); [done|]. *) -(* iIntros "_". *) -(* by iApply "HΨ". *) -(* Qed. *) - -(* End HOCAP_alt. *) diff --git a/theories/coneris/lib/hocap_flip.v b/theories/coneris/lib/hocap_flip.v index a488b39b..5b56f55a 100644 --- a/theories/coneris/lib/hocap_flip.v +++ b/theories/coneris/lib/hocap_flip.v @@ -74,12 +74,6 @@ Class flip_spec `{!conerisGS Σ} := FlipSpec (** Instantiate flip *) Section instantiate_flip. Context `{H: conerisGS Σ, r1:@rand_spec Σ H}. - (* Class flipG1 Σ := FlipG1 { flipG1_tapes:: hocap_tapesGS Σ}. *) - (* Local Definition flip_inv_pred1 `{!hocap_tapesGS Σ} γ1:= *) - (* (∃ (m:gmap loc (list bool)) , *) - (* ([∗ map] α ↦ t ∈ ((λ (ns:list bool), (1%nat, bool_to_nat<$>ns))<$>m), rand_tapes_auth γ1 α ( t.1 ; t.2 )) ∗ *) - (* ●(((λ (ns:list bool), (1, bool_to_nat<$>ns))<$>m):gmap _ _)@γ2 *) - (* )%I. *) #[local] Program Definition flip_spec1 : flip_spec := {| flip_allocate_tape:= (λ: <>, rand_allocate_tape #1%nat); @@ -210,54 +204,41 @@ Section instantiate_flip. Qed. End instantiate_flip. -(* Section test. *) -(* Context `{F:flip_spec}. *) -(* Lemma flip_presample_spec_simple {L: flipG Σ} NS E ns α *) -(* (ε2 : R -> bool -> R) *) -(* (P : iProp Σ) T γ1 γ2: *) -(* ↑NS ⊆ E -> *) -(* (∀ ε n, 0<= ε -> 0<=ε2 ε n)%R -> *) -(* (∀ (ε:R), 0<= ε -> (ε2 ε true + ε2 ε false)/2 <= ε)%R-> *) -(* is_flip (L:=L) NS γ1 γ2 -∗ *) -(* (□∀ (ε:R) n, (P ∗ flip_error_auth (L:=L) γ1 ε) ={E∖↑NS}=∗ *) -(* (⌜(1<=ε2 ε n)%R⌝ ∨ (flip_error_auth (L:=L) γ1 (ε2 ε n) ∗ T (n)))) *) -(* -∗ *) -(* P -∗ flip_tapes_frag (L:=L) γ2 α ns-∗ *) -(* wp_update E (∃ n, T (n) ∗ flip_tapes_frag (L:=L) γ2 α (ns++[n])). *) -(* Proof. *) -(* iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) -(* pose (ε2' ε l:= match l with *) -(* | [b]=> ε2 ε b *) -(* | _ => 1%R *) -(* end *) -(* ). *) -(* iMod (flip_presample_spec _ _ _ _ ε2' P 1%nat (λ l, match l with | [b] => T b | _ => False end )%I with "[//][][$][$]") as "(%l & [??])"; first done. *) -(* - rewrite /ε2'. *) -(* intros. repeat case_match; try done. naive_solver. *) -(* - intros. *) -(* etrans; last apply Hineq; try done. *) -(* erewrite (SeriesC_subset (λ x, x ∈ [[true]; [false]])); last first. *) -(* + intros ? H'. *) -(* case_match eqn:K; last done. *) -(* rewrite Nat.eqb_eq in K. *) -(* exfalso. *) -(* apply H'. *) -(* destruct a as [|[|] [|]]; try (simpl in *; done). *) -(* all: set_solver. *) -(* + rewrite SeriesC_list; last first. *) -(* { rewrite !NoDup_cons; repeat split; last apply NoDup_nil; set_solver. } *) -(* simpl. lra. *) -(* - iModIntro. iIntros (ε n) "?". *) -(* destruct n as [|b [|]]. *) -(* + iLeft. iPureIntro. *) -(* by rewrite /ε2'. *) -(* + iMod ("Hvs" $! ε b with "[$]") as "[% | ?]". *) -(* * iLeft. iPureIntro. by rewrite /ε2'. *) -(* * iRight. by rewrite /ε2'. *) -(* + iLeft. iPureIntro. *) -(* by rewrite /ε2'. *) -(* - repeat case_match; try done. *) -(* iModIntro. *) -(* iFrame. *) -(* Qed. *) -(* End test. *) +Section test. + Context `{F:flip_spec}. + Lemma flip_presample_spec_simple {L: flipG Σ} NS E γ1 α ns ε ε2: + ↑NS ⊆ E -> + (∀ n, 0<=ε2 n)%R -> + ((ε2 true + ε2 false)/2<=ε)%R -> + is_flip (L:=L) NS γ1 -∗ + flip_tapes_frag (L:=L) γ1 α ns -∗ + ↯ ε -∗ + wp_update E (∃ b, flip_tapes_frag (L:=L) γ1 α (ns ++ [b]) ∗ ↯ (ε2 b)). + Proof. + iIntros (Hsubset Hpos Hineq) "#Hinv Hfrag Herr". + pose (ε2' l:= match l with + | [b]=> ε2 b + | _ => 1%R + end + ). + iMod (flip_presample_spec _ _ _ _ _ + (λ ε' num ε2'' ns', ⌜ε=ε'⌝ ∗ ⌜num=1⌝ ∗ ⌜ε2' = ε2''⌝ ∗ + ∃ x, ⌜ns'=[x]⌝ ∗ ↯ (ε2 x) + )%I with "[//][$][Herr]") as "(%&%&%&%&?&%&%&%&%&%&?)"; first done. + - iFrame. + iApply fupd_mask_intro; first set_solver. + iIntros "Hclose". + iExists 1, ε2'. + repeat iSplit. + + iPureIntro. + intros [|?[|??]]; by simpl. + + iPureIntro. + setoid_rewrite <-elem_of_enum_uniform_list'. + rewrite SeriesC_list; last apply NoDup_enum_uniform_list. + simpl. lra. + + iIntros ([|?[|??]]) "Herr"; simpl; try (by iDestruct (ec_contradict with "[$]") as "%"). + iMod "Hclose" as "_". + by iFrame. + - subst. by iFrame. + Qed. +End test. diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index e29f8e24..4582418e 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -304,772 +304,3 @@ Section checks. Qed. End checks. - -(* Class rand_spec `{!conerisGS Σ} := RandSpec *) -(* { *) -(* (** * Operations *) *) -(* rand_allocate_tape : val; *) -(* rand_tape : val; *) -(* (** * Ghost state *) *) -(* (** The assumptions about [Σ] *) *) -(* randG : gFunctors → Type; *) - -(* rand_error_name : Type; *) -(* rand_tape_name: Type; *) -(* (** * Predicates *) *) -(* is_rand {L : randG Σ} (N:namespace) *) -(* (γ1: rand_error_name) (γ2: rand_tape_name): iProp Σ; *) -(* rand_error_auth {L : randG Σ} (γ: rand_error_name) (x:R): iProp Σ; *) -(* rand_error_frag {L : randG Σ} (γ: rand_error_name) (x:R): iProp Σ; *) -(* rand_tapes_auth {L : randG Σ} (γ: rand_tape_name) (m:gmap loc (nat * list nat)): iProp Σ; *) -(* rand_tapes_frag {L : randG Σ} (γ: rand_tape_name) (α:loc) (ns: (nat * list nat)): iProp Σ; *) -(* (** * General properties of the predicates *) *) -(* #[global] is_rand_persistent {L : randG Σ} N γ1 γ2 :: *) -(* Persistent (is_rand (L:=L) N γ1 γ2); *) -(* #[global] rand_error_auth_timeless {L : randG Σ} γ x :: *) -(* Timeless (rand_error_auth (L:=L) γ x); *) -(* #[global] rand_error_frag_timeless {L : randG Σ} γ x :: *) -(* Timeless (rand_error_frag (L:=L) γ x); *) -(* #[global] rand_tapes_auth_timeless {L : randG Σ} γ m :: *) -(* Timeless (rand_tapes_auth (L:=L) γ m); *) -(* #[global] rand_tapes_frag_timeless {L : randG Σ} γ α ns :: *) -(* Timeless (rand_tapes_frag (L:=L) γ α ns); *) -(* #[global] rand_error_name_inhabited:: *) -(* Inhabited rand_error_name; *) -(* #[global] rand_tape_name_inhabited:: *) -(* Inhabited rand_tape_name; *) - -(* rand_error_auth_exclusive {L : randG Σ} γ x1 x2: *) -(* rand_error_auth (L:=L) γ x1 -∗ rand_error_auth (L:=L) γ x2 -∗ False; *) -(* rand_error_frag_split {L : randG Σ} γ (x1 x2:nonnegreal): *) -(* rand_error_frag (L:=L) γ x1 ∗ rand_error_frag (L:=L) γ x2 ⊣⊢ *) -(* rand_error_frag (L:=L) γ (x1+x2)%R ; *) -(* rand_error_auth_valid {L : randG Σ} γ x: *) -(* rand_error_auth (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; *) -(* rand_error_frag_valid {L : randG Σ} γ x: *) -(* rand_error_frag (L:=L) γ x -∗ ⌜(0<=x<1)%R⌝; *) -(* rand_error_ineq {L : randG Σ} γ x1 x2: *) -(* rand_error_auth (L:=L) γ x1 -∗ rand_error_frag (L:=L) γ x2 -∗ ⌜(x2 <= x1)%R ⌝; *) -(* rand_error_decrease {L : randG Σ} γ (x1 x2:R): *) -(* rand_error_auth (L:=L) γ x1 -∗ rand_error_frag (L:=L) γ x2 ==∗ rand_error_auth (L:=L) γ (x1-x2)%R; *) -(* rand_error_increase {L : randG Σ} γ (x1:R) (x2:nonnegreal): *) -(* (x1+x2<1)%R -> ⊢ rand_error_auth (L:=L) γ x1 ==∗ *) -(* rand_error_auth (L:=L) γ (x1+x2) ∗ rand_error_frag (L:=L) γ x2; *) - -(* rand_tapes_auth_exclusive {L : randG Σ} γ m m': *) -(* rand_tapes_auth (L:=L) γ m -∗ rand_tapes_auth (L:=L) γ m' -∗ False; *) -(* rand_tapes_frag_exclusive {L : randG Σ} γ α ns ns': *) -(* rand_tapes_frag (L:=L) γ α ns -∗ rand_tapes_frag (L:=L) γ α ns' -∗ False; *) -(* rand_tapes_agree {L : randG Σ} γ α m ns: *) -(* rand_tapes_auth (L:=L) γ m -∗ rand_tapes_frag (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; *) -(* rand_tapes_valid {L : randG Σ} γ2 α tb ns: *) -(* rand_tapes_frag (L:=L) γ2 α (tb, ns) -∗ *) -(* ⌜Forall (λ n, n<=tb)%nat ns⌝; *) -(* rand_tapes_update {L : randG Σ} γ α m ns ns': *) -(* Forall (λ n, n<=ns'.1)%nat ns'.2 -> *) -(* rand_tapes_auth (L:=L) γ m -∗ rand_tapes_frag (L:=L) γ α ns ==∗ *) -(* rand_tapes_auth (L:=L) γ (<[α := ns']> m) ∗ rand_tapes_frag (L:=L) γ α ns'; *) - -(* (** * Program specs *) *) -(* rand_inv_create_spec {L : randG Σ} N E: *) -(* ⊢ wp_update E (∃ γ1 γ2, is_rand (L:=L) N γ1 γ2); *) -(* rand_err_convert_spec1 {L : randG Σ} N E ε γ1 γ2: *) -(* ↑N ⊆ E-> *) -(* is_rand (L:=L) N γ1 γ2 -∗ *) -(* ↯ ε -∗ *) -(* wp_update E (rand_error_frag (L:=L) γ1 ε); *) -(* rand_err_convert_spec2 {L : randG Σ} N E ε γ1 γ2: *) -(* ↑N ⊆ E-> *) -(* is_rand (L:=L) N γ1 γ2 -∗ *) -(* rand_error_frag (L:=L) γ1 ε -∗ *) -(* wp_update E (↯ ε); *) -(* rand_allocate_tape_spec {L: randG Σ} N E (tb:nat) γ1 γ2: *) -(* ↑N ⊆ E-> *) -(* {{{ is_rand (L:=L) N γ1 γ2 }}} *) -(* rand_allocate_tape #tb @ E *) -(* {{{ (v:val), RET v; *) -(* ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ rand_tapes_frag (L:=L) γ2 α (tb, []) *) -(* }}}; *) -(* rand_tape_spec_some {L: randG Σ} N E γ1 γ2 (P: iProp Σ) (Q:iProp Σ) (α:loc) (n:nat) (tb:nat) ns: *) -(* ↑N⊆E -> *) -(* {{{ is_rand (L:=L) N γ1 γ2 ∗ *) -(* □ (∀ (m:gmap loc (nat * list nat)), P ∗ *) -(* rand_tapes_auth (L:=L) γ2 m *) -(* ={E∖↑N}=∗ ⌜m!!α=Some (tb, n::ns)⌝ ∗ Q ∗ rand_tapes_auth (L:=L) γ2 (<[α:=(tb, ns)]> m)) ∗ *) -(* ▷P *) -(* }}} *) -(* rand_tape #lbl:α #tb @ E *) -(* {{{ RET #n; Q }}}; *) -(* rand_presample_spec {L: randG Σ} N E ns α (tb:nat) *) -(* (ε2 : list (fin (S tb)) -> R) *) -(* (P : iProp Σ) num T γ1 γ2 ε : *) -(* ↑N ⊆ E -> *) -(* (∀ l, length l = num -> 0<= ε2 l)%R -> *) -(* (SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R-> *) -(* is_rand (L:=L) N γ1 γ2 -∗ *) -(* (□∀ (ε':R) m, (P ∗ rand_error_auth (L:=L) γ1 (ε')%R ∗ rand_tapes_auth (L:=L) γ2 m ) ={E∖↑N}=∗ *) -(* ⌜(ε<=ε')%R⌝ ∗ ⌜(m!!α = Some (tb, ns))⌝ ∗ *) -(* (∀ ns', ⌜length ns' = num⌝ ={E∖↑N}=∗ *) -(* ⌜(1<=ε2 ns' + (ε'-ε))%R⌝ ∨ *) -(* (rand_error_auth (L:=L) γ1 (ε2 ns' + (ε'-ε))%R ∗ *) -(* rand_tapes_auth (L:=L) γ2 (<[α:= (tb, ns++(fin_to_nat <$> ns'))]> m) ∗ *) -(* T ns')))-∗ *) -(* P -∗ *) -(* wp_update E (∃ n, T n) *) -(* }. *) - -(* Section impl. *) -(* Local Opaque INR. *) -(* (* (** Instantiate rand *) *) *) -(* Class randG1 Σ := RandG1 { flipG1_error::hocap_errorGS Σ; *) -(* flipG1_tapes:: hocap_tapesGS Σ; *) -(* }. *) -(* Local Definition rand_inv_pred1 `{!conerisGS Σ, !hocap_errorGS Σ, !hocap_tapesGS Σ} γ1 γ2:= *) -(* (∃ (ε:R) (m:gmap loc (nat * list nat)) , *) -(* ↯ ε ∗ ●↯ ε @ γ1 ∗ *) -(* ([∗ map] α ↦ t ∈ m, α ↪N ( t.1 ; t.2 )) ∗ *) -(* ●m@γ2 *) -(* )%I. *) - -(* #[local] Program Definition rand_spec1 `{!conerisGS Σ}: rand_spec := *) -(* {| rand_allocate_tape:= (λ: "N", alloc "N"); *) -(* rand_tape:= (λ: "α" "N", rand("α") "N"); *) -(* randG := randG1; *) -(* rand_error_name := gname; *) -(* rand_tape_name := gname; *) -(* is_rand _ N γ1 γ2 := inv N (rand_inv_pred1 γ1 γ2); *) -(* rand_error_auth _ γ x := ●↯ x @ γ; *) -(* rand_error_frag _ γ x := ◯↯ x @ γ; *) -(* rand_tapes_auth _ γ m := (●m@γ)%I; *) -(* rand_tapes_frag _ γ α ns := ((α ◯↪N (ns.1; ns.2) @ γ) ∗ ⌜Forall (λ x, x<=ns.1) ns.2⌝)%I; *) -(* |}. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros. *) -(* iApply (hocap_error_auth_exclusive with "[$][$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros. *) -(* iApply (hocap_error_frag_split). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (?????) "H". *) -(* iApply (hocap_error_auth_valid with "[$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (?????) "H". *) -(* iApply (hocap_error_frag_valid with "[$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (??????) "H1 H2". *) -(* iApply (hocap_error_ineq with "[$][$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* iIntros. *) -(* iApply (hocap_error_decrease with "[$][$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* iIntros. *) -(* by iApply (hocap_error_increase with "[$]"). *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (??????) "H1 H2". *) -(* by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[%H _]". *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (???????) "[H1 %] [H2 %]". *) -(* iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. *) -(* rewrite dfrac_op_own dfrac_valid_own. by intro. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (??????[]) "? [??]". *) -(* by iDestruct (hocap_tapes_agree with "[$][$]") as "%H". *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* by iIntros (???????) "[? $]". *) -(* Qed. *) -(* Next Obligation. *) -(* iIntros (??????[??][??]?) "?[? %]". *) -(* iMod (hocap_tapes_update with "[$][$]") as "[??]". *) -(* by iFrame. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (?????). *) -(* iApply fupd_wp_update_ret. *) -(* unshelve iMod (hocap_error_alloc (nnreal_zero)) as "(%γ1 & ? & ?)"; simpl; [rewrite INR_0; lra|]. *) -(* iMod ec_zero as "?". *) -(* iMod (hocap_tapes_alloc ∅) as "(%γ2 & H4 & H5)". *) -(* iMod (inv_alloc _ _ (rand_inv_pred1 γ1 γ2) with "[-]"). *) -(* { iFrame. by iNext. } *) -(* by iFrame. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (????? ε ???) "#Hinv Herr". *) -(* iApply fupd_wp_update_ret. *) -(* iInv "Hinv" as ">(%ε'&%&H1&?&?&?)" "Hclose". *) -(* iDestruct (ec_valid with "[$Herr]") as "%". *) -(* iCombine "Herr H1" as "?". *) -(* iDestruct (ec_valid with "[$]") as "[% %]". *) -(* simpl in *. *) -(* unshelve iMod (hocap_error_increase _ _ (mknonnegreal ε _) with "[$]") as "[? Hfrag]"; [naive_solver|simpl;lra|]. *) -(* simpl. *) -(* iMod ("Hclose" with "[-Hfrag]") as "_"; last done. *) -(* iFrame. iApply (ec_eq with "[$]"). lra. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (????? ε ???) "#Hinv Herr". *) -(* iApply fupd_wp_update_ret. *) -(* iInv "Hinv" as ">(%ε'&%&H1&?&?&?)" "Hclose". *) -(* iDestruct (hocap_error_frag_valid with "[$]") as "%". *) -(* iDestruct (ec_valid with "[$]") as "%". *) -(* iDestruct (hocap_error_ineq with "[$][$]") as "%". *) -(* unshelve iMod (hocap_error_decrease with "[$][$]") as "?". *) -(* iDestruct (ec_eq with "[$]") as "?"; last first. *) -(* { iDestruct (ec_split with "[$]") as "[? H]"; last first. *) -(* - iMod ("Hclose" with "[-H]") as "_"; last done. *) -(* iFrame. *) -(* - naive_solver. *) -(* - lra. *) -(* } *) -(* lra. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (????????? Φ) "#Hinv HΦ". *) -(* wp_pures. *) -(* iInv "Hinv" as "(%&%&?&?&?&?)" "Hclose". *) -(* wp_apply (wp_alloc_tape); first done. *) -(* iIntros (α) "Hα". *) -(* iDestruct (hocap_tapes_notin with "[$][$]") as "%". *) -(* iMod (hocap_tapes_new _ _ α tb [] with "[$]") as "[?H]"; first done. *) -(* iMod ("Hclose" with "[-H HΦ]"). *) -(* { iModIntro. iFrame. *) -(* rewrite big_sepM_insert; [iFrame|done]. *) -(* } *) -(* iApply "HΦ". *) -(* by iFrame. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (?????????????? Φ) "(#Hinv & #Hvs & HP) HΦ". *) -(* wp_pures. *) -(* wp_bind (rand(_) _)%E. *) -(* iInv "Hinv" as ">(%&%&H1&H2&H3&H4)" "Hclose". *) -(* iMod ("Hvs" with "[$]") as "(%&HQ&Hauth)". *) -(* iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. *) -(* simpl. *) -(* wp_apply (wp_rand_tape with "[$]") as "[Htape %]". *) -(* iMod ("Hclose" with "[- HΦ HQ]") as "_". *) -(* { iExists _, (<[α:=_]> m). *) -(* iFrame. *) -(* iApply "H3". by iNext. *) -(* } *) -(* by iApply "HΦ". *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (??????????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP". *) -(* iApply wp_update_state_step_coupl. *) -(* iIntros (σ1 ε_supply) "((Hheap&Htapes)&Hε)". *) -(* iMod (inv_acc with "Hinv") as "[>(%ε' & % & H1 & H2 & H3 & H4 ) Hclose]"; [done|]. *) -(* iMod ("Hvs" with "[$]") as "(%&%&Hchange)". *) -(* (* iDestruct (hocap_tapes_agree with "[$][$]") as "%K". *) *) -(* iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. *) -(* simpl. *) -(* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". *) -(* iDestruct (ec_supply_bound with "[$][$]") as "%". *) -(* iMod (ec_supply_decrease with "[$][$]") as (ε'' ε_rem -> Hε') "Hε_supply". subst. *) -(* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) -(* iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. *) -(* unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 x + (ε''-ε) else 1)%R _). *) -(* { case_match; last (simpl;lra). *) -(* apply Rplus_le_le_0_compat; last (simpl in *; lra). *) -(* apply Hpos. by apply Nat.eqb_eq. *) -(* } *) -(* iSplit. *) -(* { iPureIntro. *) -(* simpl. *) -(* rewrite (SeriesC_subset (λ x, (x∈enum_uniform_fin_list tb num))); last first. *) -(* - intros a. intros K. *) -(* rewrite elem_of_enum_uniform_fin_list in K. *) -(* case_match eqn :H4; last done. *) -(* by rewrite Nat.eqb_eq in H4. *) -(* - erewrite (SeriesC_ext _ (λ a, 1/S tb ^ num * (if bool_decide (a ∈ enum_uniform_fin_list tb num) then ε2 a else 0) + (if bool_decide (a ∈ enum_uniform_fin_list tb num) then 1/S tb ^ num * (ε''-ε) else 0)))%R; last first. *) -(* + intros n. *) -(* case_bool_decide; last lra. *) -(* replace (_ =? _) with true; last first. *) -(* * symmetry. rewrite Nat.eqb_eq. by rewrite -elem_of_enum_uniform_fin_list. *) -(* * simpl. lra. *) -(* + rewrite SeriesC_plus; [|apply ex_seriesC_scal_l, ex_seriesC_list|apply ex_seriesC_list..]. *) -(* rewrite SeriesC_list_2; last apply NoDup_enum_uniform_fin_list. *) -(* rewrite SeriesC_scal_l enum_uniform_fin_list_length pow_INR. *) -(* replace (1/_*_*_)%R with (ε''-ε)%R; first (simpl; lra). *) -(* rewrite Rmult_comm -Rmult_assoc Rdiv_1_l Rmult_inv_r; first lra. *) -(* apply pow_nonzero. *) -(* pose proof pos_INR_S tb; lra. *) -(* } *) -(* iIntros (sample) "<-". *) -(* destruct (Rlt_decision (nonneg ε_rem + (ε2 sample + (ε''-ε))%R) 1%R) as [Hdec|Hdec]; last first. *) -(* { apply Rnot_lt_ge, Rge_le in Hdec. *) -(* iApply state_step_coupl_ret_err_ge_1. *) -(* simpl. simpl in *. rewrite Nat.eqb_refl. lra. *) -(* } *) -(* iApply state_step_coupl_ret. *) -(* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 sample + (ε''-ε))%R _) with "Hε_supply") as "[Hε_supply Hε]". *) -(* { apply Rplus_le_le_0_compat; simpl in *; [naive_solver|lra]. } *) -(* { simpl. done. } *) -(* simpl. *) -(* iMod (tapeN_update_append' _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". *) -(* (* iMod (hocap_tapes_update _ _ _ _ _ _ ((fin_to_nat <$> ns')++(fin_to_nat <$> sample)) with "[$][$]") as "[H4 Hfrag]". *) *) -(* iMod "Hclose'" as "_". *) -(* iMod ("Hchange" $! sample with "[//]") as "[%|(Hauth & H2 & HT)]". *) -(* { iExFalso. iApply (ec_contradict with "[$]"). exact. } *) -(* iMod ("Hclose" with "[$Hε $Hauth Htape H3 $H2]") as "_". *) -(* { iNext. *) -(* by iApply "H3". *) -(* } *) -(* iApply fupd_mask_intro_subseteq; first set_solver. *) -(* iFrame. *) -(* erewrite (nnreal_ext _ _); first done. *) -(* simpl. by rewrite Nat.eqb_refl. *) -(* Qed. *) - -(* End impl. *) - - -(* (** * EXPERIMENT *) *) - -(* Class rand_spec' `{!conerisGS Σ} := RandSpec' *) -(* { *) -(* (** * Operations *) *) -(* rand_allocate_tape' : val; *) -(* rand_tape' : val; *) -(* (** * Ghost state *) *) -(* (** The assumptions about [Σ] *) *) -(* randG' : gFunctors → Type; *) - -(* rand_tape_name': Type; *) -(* (** * Predicates *) *) -(* is_rand' {L : randG' Σ} (N:namespace) *) -(* (γ1: rand_tape_name') : iProp Σ; *) -(* rand_tapes_auth' {L : randG' Σ} (γ: rand_tape_name') (m:gmap loc (nat * list nat)): iProp Σ; *) -(* rand_tapes_frag' {L : randG' Σ} (γ: rand_tape_name') (α:loc) (ns: (nat * list nat)): iProp Σ; *) -(* (** * General properties of the predicates *) *) -(* #[global] is_rand_persistent' {L : randG' Σ} N γ1 :: *) -(* Persistent (is_rand' (L:=L) N γ1); *) -(* #[global] rand_tapes_auth_timeless' {L : randG' Σ} γ m :: *) -(* Timeless (rand_tapes_auth' (L:=L) γ m); *) -(* #[global] rand_tapes_frag_timeless' {L : randG' Σ} γ α ns :: *) -(* Timeless (rand_tapes_frag' (L:=L) γ α ns); *) -(* #[global] rand_tape_name_inhabited' :: *) -(* Inhabited rand_tape_name'; *) - -(* rand_tapes_auth_exclusive' {L : randG' Σ} γ m m': *) -(* rand_tapes_auth' (L:=L) γ m -∗ rand_tapes_auth' (L:=L) γ m' -∗ False; *) -(* rand_tapes_frag_exclusive' {L : randG' Σ} γ α ns ns': *) -(* rand_tapes_frag' (L:=L) γ α ns -∗ rand_tapes_frag' (L:=L) γ α ns' -∗ False; *) -(* rand_tapes_agree' {L : randG' Σ} γ α m ns: *) -(* rand_tapes_auth' (L:=L) γ m -∗ rand_tapes_frag' (L:=L) γ α ns -∗ ⌜ m!! α = Some (ns) ⌝; *) -(* rand_tapes_valid' {L : randG' Σ} γ2 α tb ns: *) -(* rand_tapes_frag' (L:=L) γ2 α (tb, ns) -∗ *) -(* ⌜Forall (λ n, n<=tb)%nat ns⌝ ; *) -(* rand_tapes_update' {L : randG' Σ} γ α m ns ns': *) -(* Forall (λ x, x<=ns'.1) ns'.2 -> *) -(* rand_tapes_auth' (L:=L) γ m -∗ rand_tapes_frag' (L:=L) γ α ns ==∗ *) -(* rand_tapes_auth' (L:=L) γ (<[α := ns']> m) ∗ rand_tapes_frag' (L:=L) γ α ns'; *) - -(* (** * Program specs *) *) -(* rand_inv_create_spec' {L : randG' Σ} N E: *) -(* ⊢ wp_update E (∃ γ1, is_rand' (L:=L) N γ1); *) -(* rand_allocate_tape_spec' {L: randG' Σ} N E (tb:nat) γ2: *) -(* ↑N ⊆ E-> *) -(* {{{ is_rand' (L:=L) N γ2 }}} *) -(* rand_allocate_tape' #tb @ E *) -(* {{{ (v:val), RET v; *) -(* ∃ (α:loc), ⌜v=#lbl:α⌝ ∗ rand_tapes_frag' (L:=L) γ2 α (tb, []) *) -(* }}}; *) -(* rand_tape_spec_some' {L: randG' Σ} N E γ2 (P: iProp Σ) (Q:iProp Σ) (α:loc) (n:nat) (tb:nat) ns: *) -(* ↑N⊆E -> *) -(* {{{ is_rand' (L:=L) N γ2 ∗ *) -(* □ (∀ (m:gmap loc (nat * list nat)), P ∗ *) -(* rand_tapes_auth' (L:=L) γ2 m *) -(* ={E∖↑N}=∗ ⌜m!!α=Some (tb, n::ns)⌝ ∗ Q ∗ rand_tapes_auth' (L:=L) γ2 (<[α:=(tb, ns)]> m)) ∗ *) -(* ▷ P *) -(* }}} *) -(* rand_tape' #lbl:α #tb @ E *) -(* {{{ RET #n; Q }}}; *) -(* rand_presample_spec' {L: randG' Σ} N E ns α (tb:nat) *) -(* (ε2 : list (fin (S tb)) -> R) *) -(* (P : iProp Σ) num T γ2 (ε:nonnegreal) : *) -(* ↑N ⊆ E -> *) -(* (∀ l, length l = num -> 0<= ε2 l)%R -> *) -(* (SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R-> *) -(* is_rand' (L:=L) N γ2 -∗ *) -(* (□∀ (ε': nonnegreal) m, (P ∗ err_interp ε' ∗ rand_tapes_auth' (L:=L) γ2 m *) -(* ) ={E∖↑N}=∗ ⌜(ε<=ε')%R⌝ ∗ ⌜(m!!α = Some (tb, ns))⌝ ∗ *) -(* (∀ ns', ⌜length ns' = num⌝ ={E∖↑N}=∗ ⌜(1<=ε2 ns' + (ε'-ε))%R⌝ ∨ ∃ (ε_final:nonnegreal), *) -(* ⌜(nonneg ε_final = ε2 ns' + (ε'-ε))%R ⌝ ∗ *) -(* (err_interp ε_final ∗ *) -(* rand_tapes_auth' (L:=L) γ2 (<[α:= (tb, ns++(fin_to_nat <$> ns'))]> m) ∗ T ns'))) *) -(* -∗ *) -(* P -∗ *) -(* wp_update E (∃ n, T n) *) -(* }. *) - -(* Section impl'. *) -(* Local Opaque INR. *) -(* (* (** Instantiate rand *) *) *) -(* Class randG1' Σ := RandG1' { *) -(* flipG1_tapes':: hocap_tapesGS Σ; *) -(* }. *) -(* Local Definition rand_inv_pred1' `{!conerisGS Σ, !hocap_tapesGS Σ} γ2:= *) -(* (∃ (m:gmap loc (nat * list nat)) , *) -(* ([∗ map] α ↦ t ∈ m, α ↪N ( t.1 ; t.2 )) ∗ *) -(* ●m@γ2 *) -(* )%I. *) - -(* #[local] Program Definition rand_spec1' `{!conerisGS Σ}: rand_spec' := *) -(* {| rand_allocate_tape':= (λ: "N", alloc "N"); *) -(* rand_tape':= (λ: "α" "N", rand("α") "N"); *) -(* randG' := randG1'; *) -(* rand_tape_name' := gname; *) -(* is_rand' _ N γ2 := inv N (rand_inv_pred1' γ2); *) -(* rand_tapes_auth' _ γ m := (●m@γ)%I; *) -(* rand_tapes_frag' _ γ α ns := (α ◯↪N (ns.1; ns.2) @ γ ∗ ⌜Forall (λ x, x<=ns.1) ns.2⌝)%I; *) -(* |}. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (??????) "H1 H2". *) -(* by iDestruct (ghost_map_auth_valid_2 with "[$][$]") as "[%H _]". *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (???????) "[H1 %] [H2 %]". *) -(* iDestruct (ghost_map_elem_frac_ne with "[$][$]") as "%"; last done. *) -(* rewrite dfrac_op_own dfrac_valid_own. by intro. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (??????[]) "?[? %]". *) -(* by iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* by iIntros (???????) "[? $]". *) -(* Qed. *) -(* Next Obligation. *) -(* iIntros (??????[][]?) "?[? %]". *) -(* iMod (hocap_tapes_update with "[$][$]") as "[??]". *) -(* by iFrame. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (?????). *) -(* iApply fupd_wp_update_ret. *) -(* iMod (hocap_tapes_alloc ∅) as "(%γ2 & H4 & H5)". *) -(* iMod (inv_alloc _ _ (rand_inv_pred1' γ2) with "[-]"). *) -(* { iFrame. by iNext. } *) -(* by iFrame. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (???????? Φ) "#Hinv HΦ". *) -(* wp_pures. *) -(* iInv "Hinv" as "(%&?&?)" "Hclose". *) -(* wp_apply (wp_alloc_tape); first done. *) -(* iIntros (α) "Hα". *) -(* iDestruct (hocap_tapes_notin with "[$][$]") as "%". *) -(* iMod (hocap_tapes_new _ _ α tb [] with "[$]") as "[?H]"; first done. *) -(* iMod ("Hclose" with "[-H HΦ]"). *) -(* { iModIntro. iFrame. *) -(* rewrite big_sepM_insert; [iFrame|done]. *) -(* } *) -(* iApply "HΦ". *) -(* by iFrame. *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (????????????? Φ) "(#Hinv & #Hvs & HP) HΦ". *) -(* wp_pures. *) -(* wp_bind (rand(_) _)%E. *) -(* iInv "Hinv" as ">(%&H3&H4)" "Hclose". *) -(* iMod ("Hvs" with "[$]") as "(%&HQ&Hauth)". *) -(* iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. *) -(* simpl. *) -(* wp_apply (wp_rand_tape with "[$]") as "[Htape %]". *) -(* iMod ("Hclose" with "[- HΦ HQ]") as "_". *) -(* { iExists (<[α:=_]> m). *) -(* iFrame. *) -(* iApply "H3". by iNext. *) -(* } *) -(* by iApply "HΦ". *) -(* Qed. *) -(* Next Obligation. *) -(* simpl. *) -(* iIntros (?????????????? Hsubset Hpos Hineq) "#Hinv #Hvs HP ". *) -(* iApply wp_update_state_step_coupl. *) -(* iIntros (σ1 ε_supply) "((Hheap&Htapes)&Hε)". *) -(* iMod (inv_acc with "Hinv") as "[>(% & H3 & H4 ) Hclose]"; [done|]. *) -(* iMod ("Hvs" with "[$]") as "(%&%&Hchange)". *) -(* iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. *) -(* simpl. *) -(* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". *) -(* (* iDestruct (ec_supply_bound with "[$][$]") as "%". *) *) -(* (* iMod (ec_supply_decrease with "[$][$]") as (ε'' ε_rem -> Hε') "Hε_supply". subst. *) *) -(* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) -(* unshelve epose (ε_rem := mknonnegreal (ε_supply-ε) _). *) -(* { simpl in *. lra. } *) -(* replace (ε_supply) with (ε_rem + ε)%NNR; last first. *) -(* { rewrite /ε_rem. apply nnreal_ext. simpl. lra. } *) -(* iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. *) -(* unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 x else 1)%R _). *) -(* { case_match; last (simpl;lra). apply Hpos. by rewrite -Nat.eqb_eq. } *) -(* iSplit. *) -(* { iPureIntro. *) -(* simpl. etrans; last exact. *) -(* rewrite (Rdiv_def (SeriesC _)(S tb ^num)%R) -SeriesC_scal_r. *) -(* apply SeriesC_le; last apply ex_seriesC_scal_r, ex_seriesC_list. *) -(* intros. rewrite elem_of_enum_uniform_fin_list'. *) -(* case_match; last lra. *) -(* split; last lra. *) -(* apply Rmult_le_pos; last (apply Hpos; by rewrite -Nat.eqb_eq). *) -(* rewrite -pow_INR. apply Rdiv_INR_ge_0. *) -(* } *) -(* iIntros (sample) "<-". *) -(* destruct (Rlt_decision (nonneg ε_rem + (ε2 sample)%R) 1%R) as [Hdec|Hdec]; last first. *) -(* { apply Rnot_lt_ge, Rge_le in Hdec. *) -(* iApply state_step_coupl_ret_err_ge_1. *) -(* simpl. simpl in *. rewrite Nat.eqb_refl. lra. *) -(* } *) -(* iApply state_step_coupl_ret. *) -(* simpl. *) -(* rewrite -Heq. *) -(* iMod (tapeN_update_append' _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". *) -(* iMod "Hclose'" as "_". *) -(* iMod ("Hchange" $! sample with "[//]") as "Hor". *) -(* iDestruct "Hor" as "[%|(%ε_final & %Heq' & Hε & Hauth & HT)]". *) -(* { exfalso. simpl in *. lra. } *) -(* iMod ("Hclose" with "[$Hauth Htape H3 ]") as "_". *) -(* { iNext. *) -(* by iApply "H3". *) -(* } *) -(* iApply fupd_mask_intro_subseteq; first set_solver. *) -(* iFrame. *) -(* erewrite (nnreal_ext (_+_)%NNR _); first done. *) -(* simpl. rewrite Nat.eqb_refl. simpl in *. lra. *) -(* Qed. *) - -(* End impl'. *) - -(* Section checks. *) -(* Context `{H: conerisGS Σ, r1:@rand_spec Σ H, r2:@rand_spec' Σ H, L:randG Σ, L': randG' Σ}. *) - -(* Lemma wp_rand_alloc_tape NS (N:nat) E γ1 γ2 : *) -(* ↑NS ⊆ E -> *) -(* {{{ is_rand (L:=L) NS γ1 γ2 }}} rand_allocate_tape #N @ E {{{ α, RET #lbl:α; rand_tapes_frag (L:=L) γ2 α (N,[]) }}}. *) -(* Proof. *) -(* iIntros (Hsubset Φ) "#Hinv HΦ". *) -(* wp_apply (rand_allocate_tape_spec with "[//]"); first exact. *) -(* iIntros (?) "(%&->&?)". *) -(* by iApply "HΦ". *) -(* Qed. *) - -(* Lemma wp_rand_alloc_tape' NS (N:nat) E γ2: *) -(* ↑NS ⊆ E -> *) -(* {{{ is_rand' (L:=L') NS γ2 }}} rand_allocate_tape' #N @ E {{{ α, RET #lbl:α; rand_tapes_frag' (L:=L') γ2 α (N,[]) }}}. *) -(* Proof. *) -(* iIntros (Hsubset Φ) "#Hinv HΦ". *) -(* wp_apply (rand_allocate_tape_spec' with "[//]"); first exact. *) -(* iIntros (?) "(%&->&?)". *) -(* by iApply "HΦ". *) -(* Qed. *) - -(* Lemma wp_rand_tape_1 NS (N:nat) E γ1 γ2 n ns α: *) -(* ↑NS ⊆ E -> *) -(* {{{ is_rand (L:=L) NS γ1 γ2 ∗ ▷ rand_tapes_frag (L:=L) γ2 α (N, n :: ns) }}} *) -(* rand_tape #lbl:α #N@ E *) -(* {{{ RET #(LitInt n); rand_tapes_frag (L:=L) γ2 α (N, ns) ∗ ⌜n <= N⌝ }}}. *) -(* Proof. *) -(* iIntros (Hsubset Φ) "(#Hinv & Hfrag) HΦ". *) -(* wp_apply (rand_tape_spec_some _ _ _ _ _ (⌜n<=N⌝ ∗ rand_tapes_frag _ _ _)%I with "[Hfrag]"); first exact. *) -(* - iSplit; first done. iSplit; last iApply "Hfrag". *) -(* iModIntro. *) -(* iIntros (?) "[Hfrag Hauth]". *) -(* iDestruct (rand_tapes_agree with "[$][$]") as "%". *) -(* iDestruct (rand_tapes_valid with "[$Hfrag]") as "%K". *) -(* inversion K; subst. *) -(* iMod (rand_tapes_update _ _ _ _ (N, ns) with "[$][$]") as "[Hauth Hfrag]"; first done. *) -(* iFrame "Hauth". iModIntro. *) -(* iFrame. *) -(* by iPureIntro. *) -(* - iIntros "[??]". iApply "HΦ". *) -(* iFrame. *) -(* Qed. *) - -(* Lemma wp_rand_tape_2 NS (N:nat) E γ2 n ns α: *) -(* ↑NS ⊆ E -> *) -(* {{{ is_rand' (L:=L') NS γ2 ∗ ▷ rand_tapes_frag' (L:=L') γ2 α (N, n :: ns) }}} *) -(* rand_tape' #lbl:α #N@ E *) -(* {{{ RET #(LitInt n); rand_tapes_frag' (L:=L') γ2 α (N, ns) ∗ ⌜n <= N⌝ }}}. *) -(* Proof. *) -(* iIntros (Hsubset Φ) "(#Hinv &>Hfrag) HΦ". *) -(* wp_apply (rand_tape_spec_some' _ _ _ _ (⌜n<=N⌝ ∗ rand_tapes_frag' _ _ _)%I with "[Hfrag]"); first exact. *) -(* - iSplit; first done. iSplit; last iApply "Hfrag". *) -(* iModIntro. *) -(* iIntros (?) "[Hfrag Hauth]". *) -(* iDestruct (rand_tapes_agree' with "[$][$]") as "%". *) -(* iDestruct (rand_tapes_valid' with "[$Hfrag]") as "%K". *) -(* inversion K; subst. *) -(* iMod (rand_tapes_update' _ _ _ _ (N, ns) with "[$][$]") as "[Hauth Hfrag]"; first done. *) -(* iFrame "Hauth". iModIntro. *) -(* iFrame. *) -(* by iPureIntro. *) -(* - iIntros "[??]". iApply "HΦ". *) -(* iFrame. *) -(* Qed. *) - -(* Lemma wp_presample_adv_comp_rand_tape (N : nat) NS E α ns (ε1 : R) (ε2 : fin (S N) -> R) γ1 γ2: *) -(* ↑NS ⊆ E -> *) -(* (∀ n, 0<=ε2 n)%R -> *) -(* (SeriesC (λ n, (1 / (S N)) * ε2 n)%R <= ε1)%R → *) -(* is_rand (L:=L) NS γ1 γ2 -∗ *) -(* ▷ rand_tapes_frag (L:=L) γ2 α (N, ns) -∗ *) -(* rand_error_frag (L:=L) γ1 ε1 -∗ *) -(* wp_update E (∃ n, rand_error_frag (L:=L) γ1 (ε2 n) ∗ rand_tapes_frag (L:=L) γ2 α (N, ns ++[fin_to_nat n]))%I. *) -(* Proof. *) -(* iIntros (Hsubset Hpos Hineq) "#Hinv >Htape Herr". *) -(* iMod (rand_presample_spec _ _ _ _ _ (λ l, match l with *) -(* | [x] => ε2 x *) -(* | _ => 1%R *) -(* end *) -(* ) *) -(* (rand_tapes_frag (L:=L) γ2 α (N, ns) ∗rand_error_frag (L:=L) γ1 ε1) 1%nat *) -(* (λ ls, ∃ n, ⌜ls=[n]⌝ ∗ rand_error_frag γ1 (ε2 n) ∗ rand_tapes_frag γ2 α (N, ns ++ [fin_to_nat n]))%I *) -(* with "[//][][$Htape $Herr]") as "(%l&%sample & -> & Herr &Htape)". *) -(* - done. *) -(* - by intros [|?[|]]. *) -(* - etrans; last exact. *) -(* etrans; last eapply (SeriesC_le_inj _ (λ l, match l with |[x] => Some x | _ => None end)). *) -(* + rewrite Rdiv_def -SeriesC_scal_r. apply Req_le_sym. *) -(* apply SeriesC_ext. *) -(* intros. case_match; subst. *) -(* * rewrite bool_decide_eq_false_2; first (simpl;lra). *) -(* by rewrite elem_of_enum_uniform_fin_list. *) -(* * case_match; last first. *) -(* { rewrite bool_decide_eq_false_2; first (simpl;lra). *) -(* by rewrite elem_of_enum_uniform_fin_list. } *) -(* rewrite bool_decide_eq_true_2; last by apply elem_of_enum_uniform_fin_list. *) -(* simpl. *) -(* rewrite Rdiv_def Rmult_1_r; lra. *) -(* + intros. apply Rmult_le_pos; last done. *) -(* apply Rdiv_INR_ge_0. *) -(* + intros. repeat case_match; by simplify_eq. *) -(* + apply ex_seriesC_finite. *) -(* - iModIntro. *) -(* iIntros (??) "([??]&?&?)". *) -(* iDestruct (rand_error_ineq with "[$][$]") as "%". *) -(* iDestruct (rand_tapes_agree with "[$][$]") as "%". *) -(* iModIntro. *) -(* iSplit; first (iPureIntro; lra). *) -(* iSplit; first done. *) -(* iIntros ([|sample[|]]?); try done. *) -(* destruct (Rlt_decision (ε2 sample + (ε' - ε1))%R 1%R). *) -(* + iRight. *) -(* iMod (rand_error_decrease with "[$][$]") as "?". *) -(* unshelve iMod (rand_error_increase _ _ (mknonnegreal (ε2 sample) _) with "[$]") as "[? ?]"; first done. *) -(* * simpl. lra. *) -(* * simpl. iFrame. *) -(* iDestruct (rand_tapes_valid with "[$]") as "%". *) -(* iMod (rand_tapes_update with "[$][$]") as "[$?]". *) -(* -- rewrite Forall_app; split; first done. *) -(* rewrite Forall_singleton. *) -(* pose proof fin_to_nat_lt sample; simpl. lia. *) -(* -- iFrame. *) -(* iModIntro. iSplit; last done. *) -(* replace (_-_+_)%R with (ε2 sample + (ε' - ε1))%R; first done. *) -(* lra. *) -(* + iLeft. *) -(* iPureIntro. lra. *) -(* - iModIntro. iFrame. *) -(* Qed. *) - -(* Lemma wp_presample_adv_comp_rand_tape' (N : nat) NS E α ns (ε1 : R) (ε2 : fin (S N) -> R) γ2: *) -(* ↑NS ⊆ E -> *) -(* (∀ n, 0<=ε2 n)%R -> *) -(* (SeriesC (λ n, (1 / (S N)) * ε2 n)%R <= ε1)%R → *) -(* is_rand' (L:=L') NS γ2 -∗ *) -(* ▷ rand_tapes_frag' (L:=L') γ2 α (N, ns) -∗ *) -(* ↯ ε1 -∗ *) -(* wp_update E (∃ n, ↯ (ε2 n) ∗ rand_tapes_frag' (L:=L') γ2 α (N, ns ++[fin_to_nat n]))%I. *) -(* Proof. *) -(* iIntros (Hsubset Hpos Hineq) "#Hinv >Htape Herr". *) -(* iDestruct (ec_valid with "[$]") as "%Hpos'". *) -(* destruct Hpos' as [Hpos' ?]. *) -(* iMod (rand_presample_spec' _ _ _ _ _ (λ l, match l with *) -(* | [x] => ε2 x *) -(* | _ => 1%R *) -(* end *) -(* ) *) -(* (rand_tapes_frag' (L:=L') γ2 α (N, ns) ∗ ↯ ε1) 1%nat *) -(* (λ ls, ∃ n, ⌜ls=[n]⌝ ∗ ↯ (ε2 n) ∗ rand_tapes_frag' γ2 α (N, ns ++ [fin_to_nat n]))%I *) -(* _ (mknonnegreal ε1 Hpos') *) -(* with "[//][][$Htape $Herr]") as "(%l&%sample & -> & Herr &Htape)". *) -(* - done. *) -(* - by intros [|?[|]]. *) -(* - etrans; last eapply Hineq. *) -(* etrans; last eapply (SeriesC_le_inj _ (λ l, match l with |[x] => Some x | _ => None end)). *) -(* + rewrite Rdiv_def -SeriesC_scal_r. apply Req_le_sym. *) -(* apply SeriesC_ext. *) -(* intros. case_match; subst. *) -(* * rewrite bool_decide_eq_false_2; first (simpl;lra). *) -(* by rewrite elem_of_enum_uniform_fin_list. *) -(* * case_match; last first. *) -(* { rewrite bool_decide_eq_false_2; first (simpl;lra). *) -(* by rewrite elem_of_enum_uniform_fin_list. } *) -(* rewrite bool_decide_eq_true_2; last by apply elem_of_enum_uniform_fin_list. *) -(* simpl. *) -(* rewrite Rdiv_def Rmult_1_r; lra. *) -(* + intros. apply Rmult_le_pos; last done. *) -(* apply Rdiv_INR_ge_0. *) -(* + intros. repeat case_match; by simplify_eq. *) -(* + apply ex_seriesC_finite. *) -(* - iModIntro. *) -(* iIntros (??) "([??]&?&?)". *) -(* iDestruct (ec_supply_bound with "[$][$]") as "%". *) -(* iDestruct (rand_tapes_agree' with "[$][$]") as "%". *) -(* iModIntro. *) -(* iSplit; first (iPureIntro; simpl; lra). *) -(* iSplit; first done. *) -(* iIntros ([|sample[|]]?); try done. *) -(* destruct (Rlt_decision (ε2 sample + (ε' - ε1))%R 1%R). *) -(* + iRight. *) -(* iMod (ec_supply_decrease with "[$][$]") as "(%&%&->&<-&?)". *) -(* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 sample) _) with "[$]") as "[? ?]"; first done. *) -(* * simpl. simpl in *. lra. *) -(* * simpl. iFrame. *) -(* iDestruct (rand_tapes_valid' with "[$]") as "%". *) -(* iMod (rand_tapes_update' with "[$][$]") as "[$?]". *) -(* -- rewrite Forall_app; split; first done. *) -(* rewrite Forall_singleton. *) -(* pose proof fin_to_nat_lt sample; simpl. lia. *) -(* -- iFrame. *) -(* iModIntro. iSplit; last done. *) -(* simpl. *) -(* iPureIntro. lra. *) -(* + iLeft. *) -(* iPureIntro. simpl; simpl in *. lra. *) -(* - iModIntro. iFrame. *) -(* Qed. *) - -(* End checks. *) From ab1e1008c8f6b1767dcd9437e84d1b2d05110839 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Wed, 2 Oct 2024 11:28:36 +0200 Subject: [PATCH 62/69] NIT --- .../coneris/examples/random_counter/impl3.v | 13 +-- theories/coneris/lib/hocap_rand.v | 97 +++++++++++-------- 2 files changed, 65 insertions(+), 45 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl3.v b/theories/coneris/examples/random_counter/impl3.v index 02372bab..383f3fe1 100644 --- a/theories/coneris/examples/random_counter/impl3.v +++ b/theories/coneris/examples/random_counter/impl3.v @@ -120,6 +120,7 @@ Section impl3. rewrite filter_cons_False in Hfilter; first done. rewrite /filter_f. lia. Qed. + Lemma counter_presample_spec3 NS E T γ1 γ2 c α ns: ↑NS ⊆ E -> @@ -135,13 +136,13 @@ Section impl3. ))-∗ wp_update E (∃ ε ε2 num ns', (∃ ls, ⌜filter filter_f ls = ns++ns'⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) ∗ T ε ε2 num ns'). Proof. - iIntros (Hsubset) "[#Hinv #Hinv'] (%ls & %Hfilter & Hfrag) Hvs". + iIntros (Hsubset) "[#Hinv #Hinv'] Hfrag Hvs". iMod wp_update_epsilon_err as "(%eps&%Heps&Heps)". - (* iRevert (ls Hfilter) "Hfrag Hvs". *) - (* iApply (ec_ind_amp _ 5%R with "[][$]"); [done|lra|]. *) - (* clear eps Heps. *) - (* iModIntro. *) - (* iIntros (eps Heps) "#IH Heps %ls %Hfilter Hfrag Hvs". *) + iRevert "Hfrag Hvs". + iApply (ec_ind_amp _ 5%R with "[][$]"); [done|lra|]. + clear eps Heps. + iModIntro. + iIntros (eps Heps) "#IH Heps (%ls & %Hfilter & Hfrag) Hvs". iMod (rand_presample_spec _ _ _ _ _ _ (λ ε num ε2 ns', _ )%I with "[//][$][Heps Hvs]") as "H"; first by apply nclose_subseteq'. diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index 4582418e..55829e2d 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -244,35 +244,39 @@ Section checks. Qed. Local Opaque enum_uniform_fin_list. - Lemma wp_presample_adv_comp_rand_tape (N : nat) NS E α ns (ε1 : R) (ε2 : fin (S N) -> R) γ1: - ↑NS ⊆ E -> - (∀ n, 0<=ε2 n)%R -> - (SeriesC (λ n, (1 / (S N)) * ε2 n)%R <= ε1)%R → - is_rand (L:=L) NS γ1 -∗ - ▷ rand_tapes_frag (L:=L) γ1 α (N, ns) -∗ - ↯ ε1 -∗ - wp_update E (∃ n, ↯ (ε2 n) ∗ rand_tapes_frag (L:=L) γ1 α (N, ns ++[fin_to_nat n]))%I. + + Lemma rand_presample_single_spec N E γ2 α (tb:nat) ns T: + ↑N ⊆ E -> + is_rand (L:=L) N γ2 -∗ + rand_tapes_frag (L:=L) γ2 α (tb, ns) -∗ + (|={E∖↑N, ∅}=> + ∃ ε (ε2 : fin (S tb) -> R), + ↯ ε ∗ + ⌜(∀ x, 0<= ε2 x)%R⌝ ∗ + ⌜(SeriesC ε2 /((S tb)) <= ε)%R⌝ ∗ + (∀ x, ↯ (ε2 x) ={∅, E∖↑N}=∗ T ε ε2 x)) -∗ + wp_update E (∃ ε ε2 x, rand_tapes_frag (L:=L) γ2 α (tb, ns ++ [fin_to_nat x]) ∗ + T ε ε2 x). Proof. - iIntros (Hsubset Hpos Hineq) "#Hinv >Htape Herr". - iDestruct (ec_valid with "[$]") as "%Hpos'". - destruct Hpos' as [Hpos' ?]. - iMod (rand_presample_spec _ _ _ _ _ _ - (λ ε' (num':nat) ε2' ns2, - ⌜ε1=ε'⌝ ∗ ⌜(1=num')%nat⌝ ∗ - ⌜∀ x y, fin_to_nat <$> x = fin_to_nat <$> y -> - ε2' x = (λ l, match l with |[x] => ε2 x | _ => 1%R end) y⌝ ∗ - ∃ (n:fin (S N)), ⌜fin_to_nat <$> ns2=[fin_to_nat n]⌝ ∗ ↯ (ε2 n))%I - with "[//][$ ][-]") as "H". - - done. - - iApply fupd_mask_intro; first set_solver. - iIntros "Hclose". - iFrame. - iExists 1, (λ l, match l with |[x] => ε2 x | _ => 1%R end). simpl. + iIntros (Hsubset) "#Hinv Hfrag Hvs". + iMod (rand_presample_spec _ _ _ _ _ _ (λ ε' num ε2' ns2, + ∃ x, ⌜ns2 = [x]⌝ ∗ T ε' (λ x, ε2' ([x])) x + ) with "[//][$][-]")%I as "H"; first done; last first. + - iDestruct "H" as "(%&%&%&%&Hfrag &%&%&HT)". + subst. + by iFrame. + - iMod "Hvs" as "(%&%ε2&$&%Hpos & %Hineq & Hrest)". + iModIntro. + iExists 1, (λ ls, match ls with |[x] => ε2 x | _ => 1 end)%R. repeat iSplit. + iPureIntro. by intros [|?[|]]. + iPureIntro. etrans; last eapply Hineq. + rewrite !Rdiv_def pow_1. + apply Rmult_le_compat_r. + { rewrite -Rdiv_1_l. + apply Rdiv_INR_ge_0. } etrans; last eapply (SeriesC_le_inj _ (λ l, match l with |[x] => Some x | _ => None end)). - * rewrite Rdiv_def -SeriesC_scal_r. apply Req_le_sym. + * apply Req_le_sym. apply SeriesC_ext. intros. case_match; subst. -- rewrite bool_decide_eq_false_2; first (simpl;lra). @@ -283,24 +287,39 @@ Section checks. by rewrite elem_of_enum_uniform_fin_list. } rewrite bool_decide_eq_true_2; last by apply elem_of_enum_uniform_fin_list. - simpl. - rewrite Rdiv_def Rmult_1_r; lra. - * intros. apply Rmult_le_pos; last done. - apply Rdiv_INR_ge_0. + done. + * done. * intros. repeat case_match; by simplify_eq. * apply ex_seriesC_finite. + iIntros ([|?[|]]) "?"; [by iDestruct (ec_contradict with "[$]") as "%"| |by iDestruct (ec_contradict with "[$]") as "%"]. iFrame. - iMod "Hclose". - iModIntro. - repeat iSplit; try done. - iPureIntro. - intros x y H'. - apply list_fmap_eq_inj in H'; first by simplify_eq. - apply fin_to_nat_inj. - - iModIntro. - iDestruct "H" as "(%&%&%&%&?&->&<-&%&%&->&?)". - iFrame. - Qed. + iMod ("Hrest" with "[$]"). + by iFrame. + Qed. + + Lemma wp_presample_adv_comp_rand_tape (N : nat) NS E α ns (ε1 : R) (ε2 : fin (S N) -> R) γ1: + ↑NS ⊆ E -> + (∀ n, 0<=ε2 n)%R -> + (SeriesC (λ n, (1 / (S N)) * ε2 n)%R <= ε1)%R → + is_rand (L:=L) NS γ1 -∗ + ▷ rand_tapes_frag (L:=L) γ1 α (N, ns) -∗ + ↯ ε1 -∗ + wp_update E (∃ n, ↯ (ε2 n) ∗ rand_tapes_frag (L:=L) γ1 α (N, ns ++[fin_to_nat n]))%I. + Proof. + iIntros (Hsubset Hpos Hineq) "#Hinv >Htape Herr". + iDestruct (ec_valid with "[$]") as "%Hpos'". + destruct Hpos' as [Hpos' ?]. + iMod (rand_presample_single_spec _ _ _ _ _ _ (λ ε' ε2' n, ↯ (ε2 n)) with "[//][$][-]") as "(%&%&%&?&?)"; first done; last by iFrame. + iFrame. + iApply fupd_mask_intro; first set_solver. + iIntros "Hclose". + iExists ε2. + repeat iSplit. + - done. + - iPureIntro. etrans; last apply Hineq. + rewrite SeriesC_scal_l. + lra. + - iIntros. iFrame. + Qed. End checks. From 531d4efd8181b2f899a51c1bab66ca4030fc08fa Mon Sep 17 00:00:00 2001 From: Hei Li Date: Wed, 2 Oct 2024 16:35:31 +0200 Subject: [PATCH 63/69] NIT --- .../coneris/examples/random_counter/impl3.v | 126 +++++++++++++++--- theories/coneris/lib/hocap_rand.v | 2 + 2 files changed, 107 insertions(+), 21 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl3.v b/theories/coneris/examples/random_counter/impl3.v index 383f3fe1..97d885cf 100644 --- a/theories/coneris/examples/random_counter/impl3.v +++ b/theories/coneris/examples/random_counter/impl3.v @@ -121,6 +121,68 @@ Section impl3. rewrite /filter_f. lia. Qed. + (** * test *) + Lemma counter_presample_spec3_test NS E T γ1 γ2 c α ns: + ↑NS ⊆ E -> + is_counter3 NS c γ1 γ2 -∗ + (∃ ls, ⌜filter filter_f ls = ns⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) -∗ + ( |={E∖↑NS,∅}=> + ∃ ε ε2, + ↯ ε ∗ + ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ + ⌜(SeriesC (λ x, if bool_decide (x<4)%nat then ε2 x else 0)%R / (4) <= ε)%R⌝ ∗ + (∀ x, ↯ (ε2 x) ={∅,E∖↑NS}=∗ + T ε ε2 x + ))-∗ + wp_update E (∃ ε ε2 x, (∃ ls, ⌜filter filter_f ls = ns++[x]⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) ∗ T ε ε2 x). + Proof. + iIntros (Hsubset) "[#Hinv #Hinv'] Hfrag Hvs". + iMod wp_update_epsilon_err as "(%eps&%Heps&Heps)". + (* iRevert "Hfrag Hvs". *) + (* iApply (ec_ind_amp _ 5%R with "[][$]"); [done|lra|]. *) + (* clear eps Heps. *) + (* iModIntro. *) + (* iIntros (eps Heps) "#IH Heps (%ls & <- & Hfrag) Hrest". *) + + iDestruct "Hfrag" as "(%ls & <- & Hfrag)". + + + iMod (rand_presample_single_spec _ _ _ _ _ _ + (λ ε' ε2' x', + ∃ (ε2:nat -> R), + ⌜ε2' = (λ x, if bool_decide (fin_to_nat x <4)%nat then ε2 x + else if bool_decide (fin_to_nat x = 4)%nat then ε'+5*eps + else 1)%R⌝ ∗ + if bool_decide (fin_to_nat x' = 4)%nat + then _ + else T (ε'-eps)%R ε2 (fin_to_nat x') + )%I with "[//][$][-]") as "Hvs". + - by apply nclose_subseteq'. + - iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose". + + apply difference_mono_l. + by apply nclose_subseteq'. + + iMod "Hvs" as "(%ε & %ε2 & Hε & %Hpos & %Hineq & Hvs)". + iModIntro. + iDestruct (ec_combine with "[$]") as "Herr". + iFrame. + iExists (λ x, if bool_decide (fin_to_nat x <4)%nat then ε2 x + else if bool_decide (fin_to_nat x = 4)%nat then ε+5*eps + else 1)%R. + repeat iSplit. + * admit. + * admit. + * iIntros (x) "Herr". + admit. + - admit. + Admitted. + + (* iRevert "Hfrag Hvs". *) + (* iApply (ec_ind_amp _ 5%R with "[][$]"); [done|lra|]. *) + (* clear eps Heps. *) + (* iModIntro. *) + (* iIntros (eps Heps) "#IH Heps (%ls & %Hfilter & Hfrag) Hvs". *) + (* iApply wp_update_state_step_coupl. *) + (* iIntros. *) Lemma counter_presample_spec3 NS E T γ1 γ2 c α ns: ↑NS ⊆ E -> @@ -138,26 +200,48 @@ Section impl3. Proof. iIntros (Hsubset) "[#Hinv #Hinv'] Hfrag Hvs". iMod wp_update_epsilon_err as "(%eps&%Heps&Heps)". - iRevert "Hfrag Hvs". - iApply (ec_ind_amp _ 5%R with "[][$]"); [done|lra|]. - clear eps Heps. - iModIntro. - iIntros (eps Heps) "#IH Heps (%ls & %Hfilter & Hfrag) Hvs". - iMod (rand_presample_spec _ _ _ _ _ _ (λ ε num ε2 ns', - _ - )%I with "[//][$][Heps Hvs]") as "H"; first by apply nclose_subseteq'. - - iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose". - { apply difference_mono_l. - by apply nclose_subseteq'. } - iMod "Hvs" as "(%ε & %ε2 & %num & Herr &%Hpos & %Hineq & Hrest)". - iFrame. - admit. - - iDestruct "H" as "(%ε & %num & %ε2 & %ns' & Hfrag & HT)". - iModIntro. - iFrame. - rewrite filter_app. - admit. - Admitted. + iApply fupd_wp_update. + iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose"; first set_solver. + iMod "Hvs" as "(%ε & %ε2 & %num & Herr & %Hpos & %Hineq & Hrest)". + iInduction num as [|num] "IH" forall (eps Heps ε ε2 Hpos Hineq T) "Heps Hfrag Herr Hrest Hclose". + { iDestruct (ec_weaken _ (ε2 []) with "[$Herr]") as "Hε". + - split; first done. + etrans; last exact. + rewrite -SeriesC_singleton_dependent. + rewrite Rdiv_1_r. + apply Req_le. + apply SeriesC_ext. + intros. + simpl. + repeat case_bool_decide; set_solver. + - iMod ("Hrest" with "[$]"). + iMod "Hclose" as "_". + iFrame. + rewrite app_nil_r. by iFrame. + } + + Admitted. + (* iMod "Hvs". *) + (* iRevert "Hfrag Hvs". *) + (* iApply (ec_ind_amp _ 5%R with "[][$]"); [done|lra|]. *) + (* clear eps Heps. *) + (* iModIntro. *) + (* iIntros (eps Heps) "#IH Heps (%ls & %Hfilter & Hfrag) Hvs". *) + (* iMod (rand_presample_spec _ _ _ _ _ _ (λ ε num ε2 ns', *) + (* _ *) + (* )%I with "[//][$][Heps Hvs]") as "H"; first by apply nclose_subseteq'. *) + (* - iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose". *) + (* { apply difference_mono_l. *) + (* by apply nclose_subseteq'. } *) + (* iMod "Hvs" as "(%ε & %ε2 & %num & Herr &%Hpos & %Hineq & Hrest)". *) + (* iFrame. *) + (* admit. *) + (* - iDestruct "H" as "(%ε & %num & %ε2 & %ns' & Hfrag & HT)". *) + (* iModIntro. *) + (* iFrame. *) + (* rewrite filter_app. *) + (* admit. *) + (* Admitted. *) (* iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) (* iMod wp_update_epsilon_err as "(%epsilon_err&%H&?)". *) (* iApply wp_update_state_step_coupl. *) @@ -348,7 +432,7 @@ Next Obligation. rewrite lookup_insert_Some; left. split; first done. by erewrite Forall_filter_f. - * rewrite lookup_insert_ne; last done. + * rewrite lookup_insert_ne; last done. naive_solver. + by apply Forall_filter_f. Qed. diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index 55829e2d..5f730560 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -322,4 +322,6 @@ Section checks. lra. - iIntros. iFrame. Qed. + + End checks. From 3710d4964d43574c6ff9d89ebcba9044148014b0 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Thu, 3 Oct 2024 12:31:10 +0200 Subject: [PATCH 64/69] hocap_error_coupl --- .../coneris/examples/random_counter/impl3.v | 7 - theories/coneris/lib/hocap.v | 322 ++++++++++++------ 2 files changed, 211 insertions(+), 118 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl3.v b/theories/coneris/examples/random_counter/impl3.v index 97d885cf..63898f43 100644 --- a/theories/coneris/examples/random_counter/impl3.v +++ b/theories/coneris/examples/random_counter/impl3.v @@ -176,13 +176,6 @@ Section impl3. - admit. Admitted. - (* iRevert "Hfrag Hvs". *) - (* iApply (ec_ind_amp _ 5%R with "[][$]"); [done|lra|]. *) - (* clear eps Heps. *) - (* iModIntro. *) - (* iIntros (eps Heps) "#IH Heps (%ls & %Hfilter & Hfrag) Hvs". *) - (* iApply wp_update_state_step_coupl. *) - (* iIntros. *) Lemma counter_presample_spec3 NS E T γ1 γ2 c α ns: ↑NS ⊆ E -> diff --git a/theories/coneris/lib/hocap.v b/theories/coneris/lib/hocap.v index 2613b48a..b85625f0 100644 --- a/theories/coneris/lib/hocap.v +++ b/theories/coneris/lib/hocap.v @@ -5,142 +5,241 @@ From clutch.coneris Require Import coneris. Set Default Proof Using "Type*". -Definition hocap_error_nroot:= nroot.@ "error". -Definition hocap_error_RA := authR nonnegrealR. +(* Definition hocap_error_nroot:= nroot.@ "error". *) +(* Definition hocap_error_RA := authR nonnegrealR. *) -Class hocap_errorGS (Σ : gFunctors) := Hocap_errorGS { - hocap_errorGS_inG :: inG Σ (hocap_error_RA); -}. +(* Class hocap_errorGS (Σ : gFunctors) := Hocap_errorGS { *) +(* hocap_errorGS_inG :: inG Σ (hocap_error_RA); *) +(* }. *) -Definition hocap_errorΣ := #[GFunctor (hocap_error_RA)]. +(* Definition hocap_errorΣ := #[GFunctor (hocap_error_RA)]. *) -Notation "'●↯' ε '@' γ" := (∃ (x : nonnegreal), ⌜x.(nonneg) = ε%R⌝ ∗ own γ (● x))%I - (at level 1). -Notation "'◯↯' ε '@' γ" := (∃ (x : nonnegreal), ⌜x.(nonneg) = ε%R⌝ ∗ own γ (◯ x))%I - (at level 1). +(* Notation "'●↯' ε '@' γ" := (∃ (x : nonnegreal), ⌜x.(nonneg) = ε%R⌝ ∗ own γ (● x))%I *) +(* (at level 1). *) +(* Notation "'◯↯' ε '@' γ" := (∃ (x : nonnegreal), ⌜x.(nonneg) = ε%R⌝ ∗ own γ (◯ x))%I *) +(* (at level 1). *) -Section error_lemmas. - Context `{!conerisGS Σ, !hocap_errorGS Σ}. +(* Section error_lemmas. *) +(* Context `{!conerisGS Σ, !hocap_errorGS Σ}. *) - Lemma hocap_error_auth_exclusive b b' γ: - ●↯ b @ γ -∗ ●↯ b' @ γ -∗ False. - Proof. - iIntros "[%[% H1]] [%[% H2]]". - iCombine "H1 H2" gives "%H1". - compute in H1. destruct H1. - exfalso. - apply H1. done. - Qed. +(* Lemma hocap_error_auth_exclusive b b' γ: *) +(* ●↯ b @ γ -∗ ●↯ b' @ γ -∗ False. *) +(* Proof. *) +(* iIntros "[%[% H1]] [%[% H2]]". *) +(* iCombine "H1 H2" gives "%H1". *) +(* compute in H1. destruct H1. *) +(* exfalso. *) +(* apply H1. done. *) +(* Qed. *) - Lemma hocap_error_frag_split (b b':nonnegreal) γ: - ◯↯ b @ γ ∗ ◯↯ b' @ γ ⊣⊢ ◯↯ (b+b') @ γ. - Proof. - iSplit. - - iIntros "[[%x1[% H1]] [%x2[% H2]]]". - iExists (x1 + x2)%NNR. - iSplit; [simpl; simpl in *; iPureIntro; lra|]. - rewrite auth_frag_op own_op. iFrame. - - iIntros "[%x [% H]]". - simpl in *. - replace x with (b+b')%NNR; last (apply nnreal_ext; simpl; lra). - rewrite auth_frag_op own_op. - by iDestruct "H" as "[$ $]". - Qed. +(* Lemma hocap_error_frag_split (b b':nonnegreal) γ: *) +(* ◯↯ b @ γ ∗ ◯↯ b' @ γ ⊣⊢ ◯↯ (b+b') @ γ. *) +(* Proof. *) +(* iSplit. *) +(* - iIntros "[[%x1[% H1]] [%x2[% H2]]]". *) +(* iExists (x1 + x2)%NNR. *) +(* iSplit; [simpl; simpl in *; iPureIntro; lra|]. *) +(* rewrite auth_frag_op own_op. iFrame. *) +(* - iIntros "[%x [% H]]". *) +(* simpl in *. *) +(* replace x with (b+b')%NNR; last (apply nnreal_ext; simpl; lra). *) +(* rewrite auth_frag_op own_op. *) +(* by iDestruct "H" as "[$ $]". *) +(* Qed. *) - (* Helpful lemmas *) - Lemma hocap_error_auth_valid (b:R) γ: - (●↯ b @ γ) -∗ ⌜(0<=b<1)%R⌝. - Proof. - iIntros "[%x[<- H]]". - iDestruct (own_valid with "[$]") as "%H". - iPureIntro. - rewrite auth_auth_valid in H. - destruct x. - compute in H. - split; simpl; lra. - Qed. - - Lemma hocap_error_frag_valid (b:R) γ: - (◯↯ b @ γ) -∗ ⌜(0<=b<1)%R⌝. - Proof. - iIntros "[%[<- H]]". - iDestruct (own_valid with "[$]") as "%H". - iPureIntro. - rewrite auth_frag_valid in H. - destruct x. - compute in H. - split; simpl; lra. - Qed. +(* (* Helpful lemmas *) *) +(* Lemma hocap_error_auth_valid (b:R) γ: *) +(* (●↯ b @ γ) -∗ ⌜(0<=b<1)%R⌝. *) +(* Proof. *) +(* iIntros "[%x[<- H]]". *) +(* iDestruct (own_valid with "[$]") as "%H". *) +(* iPureIntro. *) +(* rewrite auth_auth_valid in H. *) +(* destruct x. *) +(* compute in H. *) +(* split; simpl; lra. *) +(* Qed. *) + +(* Lemma hocap_error_frag_valid (b:R) γ: *) +(* (◯↯ b @ γ) -∗ ⌜(0<=b<1)%R⌝. *) +(* Proof. *) +(* iIntros "[%[<- H]]". *) +(* iDestruct (own_valid with "[$]") as "%H". *) +(* iPureIntro. *) +(* rewrite auth_frag_valid in H. *) +(* destruct x. *) +(* compute in H. *) +(* split; simpl; lra. *) +(* Qed. *) + +(* Lemma hocap_error_alloc (ε:nonnegreal): *) +(* (ε<1)%R -> ⊢ |==>∃ γ, (●↯ ε @ γ) ∗ (◯↯ ε @ γ). *) +(* Proof. *) +(* intros H. *) +(* iMod (own_alloc (● ε ⋅ ◯ ε)) as "[% [??]]". *) +(* - apply auth_both_valid_2. *) +(* + compute. destruct ε; simpl in H. lra. *) +(* + apply nonnegreal_included; lra. *) +(* - by eauto with iFrame. *) +(* Qed. *) + +(* Lemma hocap_error_ineq γ (b c:R) : *) +(* (●↯ b @ γ) -∗ (◯↯ c @ γ) -∗ ⌜ (c<=b)%R ⌝. *) +(* Proof. *) +(* iIntros "[% [<- Hγ●]] [% [<-Hγ◯]]". *) +(* iCombine "Hγ● Hγ◯" gives "%Hop". *) +(* by eapply auth_both_valid_discrete in Hop as [Hlt%nonnegreal_included ?]. *) +(* Qed. *) + +(* Lemma hocap_error_decrease γ (b' b:R) : *) +(* (●↯ (b) @ γ) -∗ (◯↯ b' @ γ) ==∗ (●↯ (b-b') @ γ). *) +(* Proof. *) +(* iIntros "H1 H2". *) +(* simpl. *) +(* iDestruct (hocap_error_ineq with "[$][$]") as "%". *) +(* iDestruct "H1" as "[%x [% H1]]". *) +(* iDestruct "H2" as "[%x' [% H2]]". *) +(* iMod (own_update_2 with "H1 H2") as "Hown". *) +(* { unshelve eapply (auth_update_dealloc _ _ ((x-x') _)%NNR), nonnegreal_local_update. *) +(* - lra. *) +(* - apply cond_nonneg. *) +(* - apply nnreal_ext =>/=. lra. } *) +(* iFrame. simpl. iPureIntro. *) +(* by subst. *) +(* Qed. *) + + +(* Lemma hocap_error_increase γ (b:R) (b':nonnegreal) : *) +(* (b+b'<1)%R -> ⊢ (●↯ b @ γ) ==∗ (●↯ (b+b')%R @ γ) ∗ (◯↯ b' @ γ). *) +(* Proof. *) +(* iIntros (Hineq) "[%x [% H]]". *) +(* iMod (own_update with "H") as "[$ $]"; last (iPureIntro; split; last done). *) +(* - apply auth_update_alloc. *) +(* apply (local_update_unital_discrete _ _ (x+b')%NNR) => z H1 H2. *) +(* split. *) +(* + destruct x, b'. compute. simpl in *. lra. *) +(* + apply nnreal_ext. simpl. *) +(* rewrite Rplus_comm. *) +(* apply Rplus_eq_compat_l. *) +(* simpl in *. rewrite H2. simpl. lra. *) +(* - simpl. lra. *) +(* Qed. *) + +(* Lemma hocap_error_auth_irrel γ (b c:R) : *) +(* (b=c)%R -> (●↯ b @ γ) -∗ (●↯ c @ γ). *) +(* Proof. *) +(* iIntros (->) "$". *) +(* Qed. *) + +(* Lemma hocap_error_frag_irrel γ (b c:R) : *) +(* (b=c)%R -> (◯↯ b @ γ) -∗ (◯↯ c @ γ). *) +(* Proof. *) +(* iIntros (->) "$". *) +(* Qed. *) + +(* End error_lemmas. *) + + +Section hocap_error_coupl. + Context `{conerisGS Σ}. + Context (tb:nat). + Local Canonical Structure finO := leibnizO (fin (S tb)). + Local Canonical Structure RO := leibnizO R. + Local Canonical Structure ε2O := leibnizO (list (fin(S tb))->R). + + + Definition hocap_error_coupl_pre Z Φ : (R * R * list ((list (fin (S tb)) -> R) * (list (fin (S tb)))) -> iProp Σ) := + (λ x, + let '(ε, ε_initial, ls) := x in + (Z ε ε_initial ls)∨ + ∃ num ε2, + ⌜(∀ l, 0<=ε2 l)%R⌝ ∗ + ⌜(SeriesC (λ l, if bool_decide (length l = num) then ε2 l else 0)/((S tb)^num) <= ε)%R⌝ ∗ + (∀ ns, Φ (ε2 ns, ε_initial, ls ++ [(ε2, ns)])) + )%I. - Lemma hocap_error_alloc (ε:nonnegreal): - (ε<1)%R -> ⊢ |==>∃ γ, (●↯ ε @ γ) ∗ (◯↯ ε @ γ). + Local Instance hocap_error_coupl_pre_ne Z Φ : + NonExpansive (hocap_error_coupl_pre Z Φ). Proof. - intros H. - iMod (own_alloc (● ε ⋅ ◯ ε)) as "[% [??]]". - - apply auth_both_valid_2. - + compute. destruct ε; simpl in H. lra. - + apply nonnegreal_included; lra. - - by eauto with iFrame. + solve_contractive. Qed. - Lemma hocap_error_ineq γ (b c:R) : - (●↯ b @ γ) -∗ (◯↯ c @ γ) -∗ ⌜ (c<=b)%R ⌝. + Local Instance hocap_error_coupl_pre_mono Z : BiMonoPred (hocap_error_coupl_pre Z). Proof. - iIntros "[% [<- Hγ●]] [% [<-Hγ◯]]". - iCombine "Hγ● Hγ◯" gives "%Hop". - by eapply auth_both_valid_discrete in Hop as [Hlt%nonnegreal_included ?]. + split; last apply _. + iIntros (Φ Ψ HNEΦ HNEΨ) "#Hwand". + iIntros ([[??]?]) "[H|(%&%&%&%&H)]". + - by iLeft. + - iRight. + iExists _, _. + repeat iSplit; try done. + iIntros. by iApply "Hwand". Qed. - Lemma hocap_error_decrease γ (b' b:R) : - (●↯ (b) @ γ) -∗ (◯↯ b' @ γ) ==∗ (●↯ (b-b') @ γ). + Definition hocap_error_coupl' Z := bi_least_fixpoint (hocap_error_coupl_pre Z). + Definition hocap_error_coupl ε_current ε_initial ls Z := hocap_error_coupl' Z (ε_current, ε_initial, ls). + + Lemma hocap_error_coupl_unfold ε ε_initial ls Z : + hocap_error_coupl ε ε_initial ls Z ≡ + ((Z ε ε_initial ls)∨ + ∃ num ε2, + ⌜(∀ l, 0<=ε2 l)%R⌝ ∗ + ⌜(SeriesC (λ l, if bool_decide (length l = num) then ε2 l else 0)/((S tb)^num) <= ε)%R⌝ ∗ + (∀ ns, hocap_error_coupl (ε2 ns) ε_initial (ls ++ [(ε2, ns)]) Z))%I. Proof. - iIntros "H1 H2". - simpl. - iDestruct (hocap_error_ineq with "[$][$]") as "%". - iDestruct "H1" as "[%x [% H1]]". - iDestruct "H2" as "[%x' [% H2]]". - iMod (own_update_2 with "H1 H2") as "Hown". - { unshelve eapply (auth_update_dealloc _ _ ((x-x') _)%NNR), nonnegreal_local_update. - - lra. - - apply cond_nonneg. - - apply nnreal_ext =>/=. lra. } - iFrame. simpl. iPureIntro. - by subst. - Qed. - + rewrite /hocap_error_coupl/hocap_error_coupl' least_fixpoint_unfold//. Qed. - Lemma hocap_error_increase γ (b:R) (b':nonnegreal) : - (b+b'<1)%R -> ⊢ (●↯ b @ γ) ==∗ (●↯ (b+b')%R @ γ) ∗ (◯↯ b' @ γ). + Lemma hocap_error_coupl_ret ε ε_initial ls Z: + Z ε ε_initial ls -∗ hocap_error_coupl ε ε_initial ls Z. Proof. - iIntros (Hineq) "[%x [% H]]". - iMod (own_update with "H") as "[$ $]"; last (iPureIntro; split; last done). - - apply auth_update_alloc. - apply (local_update_unital_discrete _ _ (x+b')%NNR) => z H1 H2. - split. - + destruct x, b'. compute. simpl in *. lra. - + apply nnreal_ext. simpl. - rewrite Rplus_comm. - apply Rplus_eq_compat_l. - simpl in *. rewrite H2. simpl. lra. - - simpl. lra. + iIntros. rewrite hocap_error_coupl_unfold. by iLeft. Qed. - - Lemma hocap_error_auth_irrel γ (b c:R) : - (b=c)%R -> (●↯ b @ γ) -∗ (●↯ c @ γ). + + Lemma hocap_error_coupl_rec ε ε_initial ls Z: + (∃ num ε2, + ⌜(∀ l, 0<=ε2 l)%R⌝ ∗ + ⌜(SeriesC (λ l, if bool_decide (length l = num) then ε2 l else 0)/((S tb)^num) <= ε)%R⌝ ∗ + (∀ ns, hocap_error_coupl (ε2 ns) ε_initial (ls ++ [(ε2, ns)]) Z)) -∗ hocap_error_coupl ε ε_initial ls Z. Proof. - iIntros (->) "$". + iIntros. rewrite hocap_error_coupl_unfold. by iRight. Qed. - - Lemma hocap_error_frag_irrel γ (b c:R) : - (b=c)%R -> (◯↯ b @ γ) -∗ (◯↯ c @ γ). + + Lemma hocap_error_coupl_ind (Ψ Z : R -> R -> list ((list (fin (S tb)) -> R) * (list (fin (S tb))))->iProp Σ): + ⊢ (□ (∀ ε ε_initial ls, + hocap_error_coupl_pre Z (λ '(ε', ε_initial', ls'), + Ψ ε' ε_initial' ls' ∧ hocap_error_coupl ε' ε_initial' ls' Z)%I (ε, ε_initial, ls) -∗ Ψ ε ε_initial ls) → + ∀ ε ε_initial ls, hocap_error_coupl ε ε_initial ls Z -∗ Ψ ε ε_initial ls)%I. Proof. - iIntros (->) "$". + iIntros "#IH" (ε ε_initial ls) "H". + set (Ψ' := (λ '(ε, ε_initial, ls), Ψ ε ε_initial ls) : (prodO _ _->iProp Σ)). + assert (NonExpansive Ψ') by solve_contractive. + iApply (least_fixpoint_ind _ Ψ' with "[] H"). + iIntros "!#" ([[??]?]) "H". by iApply "IH". Qed. -End error_lemmas. - + Lemma hocap_error_coupl_bind ε ε_initial ls Z1 Z2 : + (∀ ε' ε_initial' ls', Z1 ε' ε_initial' ls' -∗ hocap_error_coupl ε' ε_initial' ls' Z2) -∗ + hocap_error_coupl ε ε_initial ls Z1 -∗ + hocap_error_coupl ε ε_initial ls Z2. + Proof. + iIntros "HZ H". + iRevert "HZ". + iRevert (ε ε_initial ls) "H". + iApply hocap_error_coupl_ind. + iIntros "!#" (ε ε_initial ls) "[H|H] HZ". + - by iApply "HZ". + - iApply hocap_error_coupl_rec. + iDestruct "H" as "(%&%&%&%&Hrest)". + iExists _,_. repeat iSplit; try done. + iIntros (?). + iDestruct ("Hrest" $! _) as "[Hrest _]". + by iApply "Hrest". + Qed. + +End hocap_error_coupl. + Definition hocap_tapes_nroot:=nroot.@"tapes". Class hocap_tapesGS (Σ : gFunctors) := Hocap_tapesGS { hocap_tapesGS_inG :: ghost_mapG Σ loc (nat*list nat) @@ -187,3 +286,4 @@ Section tapes_lemmas. End tapes_lemmas. + From e68656cf2514be6f13096437a16a2c97355b381b Mon Sep 17 00:00:00 2001 From: Hei Li Date: Thu, 3 Oct 2024 13:03:00 +0200 Subject: [PATCH 65/69] NIT --- .../coneris/examples/random_counter/impl1.v | 69 +++--- theories/coneris/lib/hocap.v | 10 +- theories/coneris/lib/hocap_flip.v | 117 +++++----- theories/coneris/lib/hocap_rand.v | 214 ++++++++++-------- 4 files changed, 216 insertions(+), 194 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl1.v b/theories/coneris/examples/random_counter/impl1.v index e3945c42..9bd0af0e 100644 --- a/theories/coneris/examples/random_counter/impl1.v +++ b/theories/coneris/examples/random_counter/impl1.v @@ -144,40 +144,41 @@ Section impl1. wp_update E (∃ ε ε2 num ns', rand_tapes_frag (L:=L) γ1 α (3%nat, ns++ns') ∗ T ε ε2 num ns'). Proof. iIntros (Hsubset) "[#Hinv #Hinv'] Hfrag Hvs". - iMod (rand_presample_spec _ _ _ _ _ _ (λ ε num ε2 ns', - ∃ ε2', ⌜∀ xs ys, fin_to_nat <$> xs = ys -> ε2 xs = ε2' ys⌝ ∗ - T ε ε2' num (fin_to_nat <$> ns'))%I with "[//][$][Hvs]") as "H"; first by apply nclose_subseteq'. - - iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose". - { apply difference_mono_l. - by apply nclose_subseteq'. } - iMod "Hvs" as "(%ε & %ε2 & %num & Herr & %Hpos & %Hineq & Hrest)". - iFrame. - iModIntro. - iExists num, (λ ls, ε2 (fin_to_nat <$> ls)). - repeat iSplit. - + done. - + iPureIntro. - etrans; last exact. - apply Req_le. - replace (INR 4) with 4%R; last (simpl; lra). - f_equal. - rewrite !SeriesC_list. - * by rewrite foldr_fmap. - * apply NoDup_fmap. - -- apply list_fmap_eq_inj. - apply fin_to_nat_inj. - -- apply NoDup_enum_uniform_fin_list. - * apply NoDup_enum_uniform_fin_list. - + iIntros (ns') "Herr". - iMod ("Hrest" with "[$]") as "HT". - iMod "Hclose" as "_". - iModIntro. - iFrame. - iPureIntro. - by intros ??<-. - - iDestruct "H" as "(%ε & %num & %ε2 & %ns' & Hfrag & %ε2' & % & HT)". - iModIntro. iFrame. - Qed. + Admitted. + (* iMod (rand_presample_spec _ _ _ _ _ _ (λ ε num ε2 ns', *) + (* ∃ ε2', ⌜∀ xs ys, fin_to_nat <$> xs = ys -> ε2 xs = ε2' ys⌝ ∗ *) + (* T ε ε2' num (fin_to_nat <$> ns'))%I with "[//][$][Hvs]") as "H"; first by apply nclose_subseteq'. *) + (* - iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose". *) + (* { apply difference_mono_l. *) + (* by apply nclose_subseteq'. } *) + (* iMod "Hvs" as "(%ε & %ε2 & %num & Herr & %Hpos & %Hineq & Hrest)". *) + (* iFrame. *) + (* iModIntro. *) + (* iExists num, (λ ls, ε2 (fin_to_nat <$> ls)). *) + (* repeat iSplit. *) + (* + done. *) + (* + iPureIntro. *) + (* etrans; last exact. *) + (* apply Req_le. *) + (* replace (INR 4) with 4%R; last (simpl; lra). *) + (* f_equal. *) + (* rewrite !SeriesC_list. *) + (* * by rewrite foldr_fmap. *) + (* * apply NoDup_fmap. *) + (* -- apply list_fmap_eq_inj. *) + (* apply fin_to_nat_inj. *) + (* -- apply NoDup_enum_uniform_fin_list. *) + (* * apply NoDup_enum_uniform_fin_list. *) + (* + iIntros (ns') "Herr". *) + (* iMod ("Hrest" with "[$]") as "HT". *) + (* iMod "Hclose" as "_". *) + (* iModIntro. *) + (* iFrame. *) + (* iPureIntro. *) + (* by intros ??<-. *) + (* - iDestruct "H" as "(%ε & %num & %ε2 & %ns' & Hfrag & %ε2' & % & HT)". *) + (* iModIntro. iFrame. *) + (* Qed. *) Lemma read_counter_spec1 N E c γ1 γ2 Q: ↑N ⊆ E -> diff --git a/theories/coneris/lib/hocap.v b/theories/coneris/lib/hocap.v index b85625f0..b616d647 100644 --- a/theories/coneris/lib/hocap.v +++ b/theories/coneris/lib/hocap.v @@ -150,14 +150,14 @@ Section hocap_error_coupl. Local Canonical Structure ε2O := leibnizO (list (fin(S tb))->R). - Definition hocap_error_coupl_pre Z Φ : (R * R * list ((list (fin (S tb)) -> R) * (list (fin (S tb)))) -> iProp Σ) := + Definition hocap_error_coupl_pre Z Φ : (R * R * list (nat * (list (fin (S tb)) -> R) * (list (fin (S tb)))) -> iProp Σ) := (λ x, let '(ε, ε_initial, ls) := x in (Z ε ε_initial ls)∨ ∃ num ε2, ⌜(∀ l, 0<=ε2 l)%R⌝ ∗ ⌜(SeriesC (λ l, if bool_decide (length l = num) then ε2 l else 0)/((S tb)^num) <= ε)%R⌝ ∗ - (∀ ns, Φ (ε2 ns, ε_initial, ls ++ [(ε2, ns)])) + (∀ ns, Φ (ε2 ns, ε_initial, ls ++ [(num, ε2, ns)])) )%I. Local Instance hocap_error_coupl_pre_ne Z Φ : @@ -187,7 +187,7 @@ Section hocap_error_coupl. ∃ num ε2, ⌜(∀ l, 0<=ε2 l)%R⌝ ∗ ⌜(SeriesC (λ l, if bool_decide (length l = num) then ε2 l else 0)/((S tb)^num) <= ε)%R⌝ ∗ - (∀ ns, hocap_error_coupl (ε2 ns) ε_initial (ls ++ [(ε2, ns)]) Z))%I. + (∀ ns, hocap_error_coupl (ε2 ns) ε_initial (ls ++ [(num, ε2, ns)]) Z))%I. Proof. rewrite /hocap_error_coupl/hocap_error_coupl' least_fixpoint_unfold//. Qed. @@ -201,12 +201,12 @@ Section hocap_error_coupl. (∃ num ε2, ⌜(∀ l, 0<=ε2 l)%R⌝ ∗ ⌜(SeriesC (λ l, if bool_decide (length l = num) then ε2 l else 0)/((S tb)^num) <= ε)%R⌝ ∗ - (∀ ns, hocap_error_coupl (ε2 ns) ε_initial (ls ++ [(ε2, ns)]) Z)) -∗ hocap_error_coupl ε ε_initial ls Z. + (∀ ns, hocap_error_coupl (ε2 ns) ε_initial (ls ++ [(num, ε2, ns)]) Z)) -∗ hocap_error_coupl ε ε_initial ls Z. Proof. iIntros. rewrite hocap_error_coupl_unfold. by iRight. Qed. - Lemma hocap_error_coupl_ind (Ψ Z : R -> R -> list ((list (fin (S tb)) -> R) * (list (fin (S tb))))->iProp Σ): + Lemma hocap_error_coupl_ind (Ψ Z : R -> R -> list (nat * (list (fin (S tb)) -> R) * (list (fin (S tb))))->iProp Σ): ⊢ (□ (∀ ε ε_initial ls, hocap_error_coupl_pre Z (λ '(ε', ε_initial', ls'), Ψ ε' ε_initial' ls' ∧ hocap_error_coupl ε' ε_initial' ls' Z)%I (ε, ε_initial, ls) -∗ Ψ ε ε_initial ls) → diff --git a/theories/coneris/lib/hocap_flip.v b/theories/coneris/lib/hocap_flip.v index 5b56f55a..55b70bae 100644 --- a/theories/coneris/lib/hocap_flip.v +++ b/theories/coneris/lib/hocap_flip.v @@ -144,64 +144,65 @@ Section instantiate_flip. Next Obligation. simpl. iIntros (?????? T ?) "#Hinv Hfrag Hvs". - iMod (rand_presample_spec _ _ _ _ _ _ - (λ ε num ε2 ns', - ∃ bs' ε2', ⌜fmap (FMap:= list_fmap) bool_to_nat bs'=fin_to_nat <$> ns'⌝ ∗ - ⌜∀ xs ys, fmap (FMap:= list_fmap) bool_to_nat xs = fmap (FMap:= list_fmap)fin_to_nat ys -> ε2' xs = ε2 ys⌝ ∗ - T ε num ε2' bs' - )%I - with "[//][$][-]") as "(%&%&%&%&?&%&%&%K&%&?)"; [exact| |iModIntro; iFrame]; last first. - { by rewrite fmap_app K. } - iMod "Hvs" as "(%&%&%&Herr &%Hpos&%Hineq &Hrest)". - iFrame. - iExists num, (λ ls, ε2 (nat_to_bool<$>(fin_to_nat <$> ls))). - iModIntro. - repeat iSplit; try iPureIntro. - - intros. apply Hpos. - by rewrite !fmap_length. - - etrans; last exact. - replace 2%R with (INR 2); last (simpl; lra). - rewrite !Rdiv_def. - apply Rmult_le_compat_r. - { rewrite -Rdiv_1_l -pow_INR. apply Rdiv_INR_ge_0. } - etrans; last eapply (SeriesC_le_inj _ (λ l, Some ((λ x, nat_to_bool (fin_to_nat x)) <$> l))). - + apply Req_le_sym. - apply SeriesC_ext. - intros. simpl. rewrite fmap_length. - rewrite elem_of_enum_uniform_fin_list'. case_match; last lra. - rewrite -list_fmap_compose. - f_equal. - + intros. case_match; last lra. - apply Hpos. - by rewrite -Nat.eqb_eq. - + intros ??? H1 H2. - simplify_eq. - apply fmap_inj in H1; first done. - intros x y ?. - repeat (inv_fin x; simpl; try intros x); - by repeat (inv_fin y; simpl; try intros y). - + eapply ex_seriesC_list_length. intros. - case_match; last done. - by rewrite -Nat.eqb_eq. - - iIntros (ls) "H2". - iMod ("Hrest" with "[$H2 ]") as "?". - iFrame. - iModIntro. - iSplit; iPureIntro. - + rewrite -!list_fmap_compose. - erewrite (Forall_fmap_ext_1 (_∘_)); first done. - apply Forall_true. - intros x; by repeat (inv_fin x; simpl; try intros x). - + intros ?? <-. - f_equal. - rewrite -!list_fmap_compose. - rewrite -{1}(list_fmap_id xs). - erewrite (Forall_fmap_ext_1 (_∘_)); first done. - apply Forall_true. - intros []; simpl. - ** by rewrite nat_to_bool_neq_0. - ** by rewrite nat_to_bool_eq_0. - Qed. + Admitted. + (* iMod (rand_presample_spec _ _ _ _ _ _ *) + (* (λ ε num ε2 ns', *) + (* ∃ bs' ε2', ⌜fmap (FMap:= list_fmap) bool_to_nat bs'=fin_to_nat <$> ns'⌝ ∗ *) + (* ⌜∀ xs ys, fmap (FMap:= list_fmap) bool_to_nat xs = fmap (FMap:= list_fmap)fin_to_nat ys -> ε2' xs = ε2 ys⌝ ∗ *) + (* T ε num ε2' bs' *) + (* )%I *) + (* with "[//][$][-]") as "(%&%&%&%&?&%&%&%K&%&?)"; [exact| |iModIntro; iFrame]; last first. *) + (* { by rewrite fmap_app K. } *) + (* iMod "Hvs" as "(%&%&%&Herr &%Hpos&%Hineq &Hrest)". *) + (* iFrame. *) + (* iExists num, (λ ls, ε2 (nat_to_bool<$>(fin_to_nat <$> ls))). *) + (* iModIntro. *) + (* repeat iSplit; try iPureIntro. *) + (* - intros. apply Hpos. *) + (* by rewrite !fmap_length. *) + (* - etrans; last exact. *) + (* replace 2%R with (INR 2); last (simpl; lra). *) + (* rewrite !Rdiv_def. *) + (* apply Rmult_le_compat_r. *) + (* { rewrite -Rdiv_1_l -pow_INR. apply Rdiv_INR_ge_0. } *) + (* etrans; last eapply (SeriesC_le_inj _ (λ l, Some ((λ x, nat_to_bool (fin_to_nat x)) <$> l))). *) + (* + apply Req_le_sym. *) + (* apply SeriesC_ext. *) + (* intros. simpl. rewrite fmap_length. *) + (* rewrite elem_of_enum_uniform_fin_list'. case_match; last lra. *) + (* rewrite -list_fmap_compose. *) + (* f_equal. *) + (* + intros. case_match; last lra. *) + (* apply Hpos. *) + (* by rewrite -Nat.eqb_eq. *) + (* + intros ??? H1 H2. *) + (* simplify_eq. *) + (* apply fmap_inj in H1; first done. *) + (* intros x y ?. *) + (* repeat (inv_fin x; simpl; try intros x); *) + (* by repeat (inv_fin y; simpl; try intros y). *) + (* + eapply ex_seriesC_list_length. intros. *) + (* case_match; last done. *) + (* by rewrite -Nat.eqb_eq. *) + (* - iIntros (ls) "H2". *) + (* iMod ("Hrest" with "[$H2 ]") as "?". *) + (* iFrame. *) + (* iModIntro. *) + (* iSplit; iPureIntro. *) + (* + rewrite -!list_fmap_compose. *) + (* erewrite (Forall_fmap_ext_1 (_∘_)); first done. *) + (* apply Forall_true. *) + (* intros x; by repeat (inv_fin x; simpl; try intros x). *) + (* + intros ?? <-. *) + (* f_equal. *) + (* rewrite -!list_fmap_compose. *) + (* rewrite -{1}(list_fmap_id xs). *) + (* erewrite (Forall_fmap_ext_1 (_∘_)); first done. *) + (* apply Forall_true. *) + (* intros []; simpl. *) + (* ** by rewrite nat_to_bool_neq_0. *) + (* ** by rewrite nat_to_bool_eq_0. *) + (* Qed. *) End instantiate_flip. Section test. diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index 5f730560..48843129 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -63,13 +63,29 @@ Class rand_spec `{!conerisGS Σ} := RandSpec is_rand (L:=L) N γ2 -∗ rand_tapes_frag (L:=L) γ2 α (tb, ns) -∗ (|={E∖↑N, ∅}=> - ∃ ε num (ε2 : list (fin (S tb)) -> R), - ↯ ε ∗ - ⌜(∀ l, length l = num -> 0<= ε2 l)%R⌝ ∗ - ⌜(SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R⌝ ∗ - (∀ (ns':list (fin (S tb))), ↯ (ε2 ns') ={∅, E∖↑N}=∗ T ε num ε2 ns')) -∗ - wp_update E (∃ ε num ε2 ns', rand_tapes_frag (L:=L) γ2 α (tb, ns ++ (fin_to_nat <$> ns')) ∗ - T ε num ε2 ns') + ∃ ε, + ↯ ε ∗ + hocap_error_coupl tb ε ε [] + (λ ε_final ε_initial ls, + ↯ ε_final ={∅, E∖↑N}=∗ T ε_final ε_initial ls + ) + ) -∗ + wp_update E (∃ ε_final ε_initial (ls:list _), + rand_tapes_frag (L:=L) γ2 α (tb, ns ++ (fin_to_nat <$> (mbind (MBind := list_bind) snd ls))) ∗ + T ε_final ε_initial ls) + + (* rand_presample_spec {L: randG Σ} N E γ2 α (tb:nat) ns T: *) + (* ↑N ⊆ E -> *) + (* is_rand (L:=L) N γ2 -∗ *) + (* rand_tapes_frag (L:=L) γ2 α (tb, ns) -∗ *) + (* (|={E∖↑N, ∅}=> *) + (* ∃ ε num (ε2 : list (fin (S tb)) -> R), *) + (* ↯ ε ∗ *) + (* ⌜(∀ l, length l = num -> 0<= ε2 l)%R⌝ ∗ *) + (* ⌜(SeriesC (λ l, if bool_decide (l ∈ enum_uniform_fin_list tb num) then ε2 l else 0) /((S tb)^num) <= ε)%R⌝ ∗ *) + (* (∀ (ns':list (fin (S tb))), ↯ (ε2 ns') ={∅, E∖↑N}=∗ T ε num ε2 ns')) -∗ *) + (* wp_update E (∃ ε num ε2 ns', rand_tapes_frag (L:=L) γ2 α (tb, ns ++ (fin_to_nat <$> ns')) ∗ *) + (* T ε num ε2 ns') *) }. Section impl. @@ -167,59 +183,62 @@ Section impl. iApply wp_update_state_step_coupl. iIntros (σ1 ε_supply) "((Hheap&Htapes)&Hε)". iMod (inv_acc with "Hinv") as "[>(% & H3 & H4 ) Hclose]"; [done|]. - iMod "Hvs" as "(%err & %num & %ε2 & Herr & %Hpos & %Hineq & Hrest)". - iDestruct (hocap_tapes_agree with "[$][$]") as "%". - iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. - iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". - iDestruct (ec_supply_bound with "[$][$]") as "%". - iMod (ec_supply_decrease with "[$][$]") as (ε'' ε_rem -> Heq') "Hε_supply". - iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". - (* unshelve epose (ε_rem := mknonnegreal (ε_supply-ε) _). *) - (* { simpl in *. lra. } *) - (* replace (ε_supply) with (ε_rem + ε)%NNR; last first. *) - (* { rewrite /ε_rem. apply nnreal_ext. simpl. lra. } *) - iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. - unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 x else 1)%R _). - { case_match; last (simpl;lra). apply Hpos. by rewrite -Nat.eqb_eq. } - iSplit. - { iPureIntro. - simpl. rewrite Heq'. etrans; last exact. - rewrite (Rdiv_def (SeriesC _)(S tb ^num)%R) -SeriesC_scal_r. - apply SeriesC_le; last apply ex_seriesC_scal_r, ex_seriesC_list. - intros. rewrite elem_of_enum_uniform_fin_list'. - case_match; last lra. - split; last lra. - apply Rmult_le_pos; last (apply Hpos; by rewrite -Nat.eqb_eq). - rewrite -pow_INR. apply Rdiv_INR_ge_0. - } - iIntros (sample) "<-". - destruct (Rlt_decision (nonneg ε_rem + (ε2 sample)%R) 1%R) as [Hdec|Hdec]; last first. - { apply Rnot_lt_ge, Rge_le in Hdec. - iApply state_step_coupl_ret_err_ge_1. - simpl. simpl in *. rewrite Nat.eqb_refl. lra. - } - iApply state_step_coupl_ret. - simpl. simpl in *. - rewrite -Heq. - iMod (tapeN_update_append' _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". - iMod "Hclose'" as "_". - unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 sample) _) with "[$]") as "[Hsupply Herr]". - { naive_solver. } - { simpl; lra. } - iMod (hocap_tapes_update with "[$][$]") as "[Hauth Hfrag]". - iMod ("Hrest" $! sample with "[$Herr]") as "HT". - iFrame. - iMod ("Hclose" with "[-Hsupply]") as "_". - { iNext. iFrame. by iApply "H3". } - iApply fupd_mask_intro_subseteq; first set_solver. - iSplit. - - iApply ec_supply_eq; last done. - simpl. rewrite Nat.eqb_refl. lra. - - iPureIntro. - rewrite Forall_app; split; subst; first done. - eapply Forall_impl; first apply fin.fin_forall_leq. - simpl; intros; lia. - Qed. + iMod "Hvs" as "(%err & Herr & Hcoupl)". + Admitted. + + (* iMod "Hvs" as "(%err & %num & %ε2 & Herr & %Hpos & %Hineq & Hrest)". *) + (* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) + (* iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H3]"; first done. *) + (* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%Heq)". *) + (* iDestruct (ec_supply_bound with "[$][$]") as "%". *) + (* iMod (ec_supply_decrease with "[$][$]") as (ε'' ε_rem -> Heq') "Hε_supply". *) + (* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) + (* (* unshelve epose (ε_rem := mknonnegreal (ε_supply-ε) _). *) *) + (* (* { simpl in *. lra. } *) *) + (* (* replace (ε_supply) with (ε_rem + ε)%NNR; last first. *) *) + (* (* { rewrite /ε_rem. apply nnreal_ext. simpl. lra. } *) *) + (* iApply (state_step_coupl_iterM_state_adv_comp_con_prob_lang num α); first done. *) + (* unshelve iExists (λ x, mknonnegreal (if length x =? num then ε2 x else 1)%R _). *) + (* { case_match; last (simpl;lra). apply Hpos. by rewrite -Nat.eqb_eq. } *) + (* iSplit. *) + (* { iPureIntro. *) + (* simpl. rewrite Heq'. etrans; last exact. *) + (* rewrite (Rdiv_def (SeriesC _)(S tb ^num)%R) -SeriesC_scal_r. *) + (* apply SeriesC_le; last apply ex_seriesC_scal_r, ex_seriesC_list. *) + (* intros. rewrite elem_of_enum_uniform_fin_list'. *) + (* case_match; last lra. *) + (* split; last lra. *) + (* apply Rmult_le_pos; last (apply Hpos; by rewrite -Nat.eqb_eq). *) + (* rewrite -pow_INR. apply Rdiv_INR_ge_0. *) + (* } *) + (* iIntros (sample) "<-". *) + (* destruct (Rlt_decision (nonneg ε_rem + (ε2 sample)%R) 1%R) as [Hdec|Hdec]; last first. *) + (* { apply Rnot_lt_ge, Rge_le in Hdec. *) + (* iApply state_step_coupl_ret_err_ge_1. *) + (* simpl. simpl in *. rewrite Nat.eqb_refl. lra. *) + (* } *) + (* iApply state_step_coupl_ret. *) + (* simpl. simpl in *. *) + (* rewrite -Heq. *) + (* iMod (tapeN_update_append' _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". *) + (* iMod "Hclose'" as "_". *) + (* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 sample) _) with "[$]") as "[Hsupply Herr]". *) + (* { naive_solver. } *) + (* { simpl; lra. } *) + (* iMod (hocap_tapes_update with "[$][$]") as "[Hauth Hfrag]". *) + (* iMod ("Hrest" $! sample with "[$Herr]") as "HT". *) + (* iFrame. *) + (* iMod ("Hclose" with "[-Hsupply]") as "_". *) + (* { iNext. iFrame. by iApply "H3". } *) + (* iApply fupd_mask_intro_subseteq; first set_solver. *) + (* iSplit. *) + (* - iApply ec_supply_eq; last done. *) + (* simpl. rewrite Nat.eqb_refl. lra. *) + (* - iPureIntro. *) + (* rewrite Forall_app; split; subst; first done. *) + (* eapply Forall_impl; first apply fin.fin_forall_leq. *) + (* simpl; intros; lia. *) + (* Qed. *) End impl. @@ -259,43 +278,44 @@ Section checks. T ε ε2 x). Proof. iIntros (Hsubset) "#Hinv Hfrag Hvs". - iMod (rand_presample_spec _ _ _ _ _ _ (λ ε' num ε2' ns2, - ∃ x, ⌜ns2 = [x]⌝ ∗ T ε' (λ x, ε2' ([x])) x - ) with "[//][$][-]")%I as "H"; first done; last first. - - iDestruct "H" as "(%&%&%&%&Hfrag &%&%&HT)". - subst. - by iFrame. - - iMod "Hvs" as "(%&%ε2&$&%Hpos & %Hineq & Hrest)". - iModIntro. - iExists 1, (λ ls, match ls with |[x] => ε2 x | _ => 1 end)%R. - repeat iSplit. - + iPureIntro. by intros [|?[|]]. - + iPureIntro. etrans; last eapply Hineq. - rewrite !Rdiv_def pow_1. - apply Rmult_le_compat_r. - { rewrite -Rdiv_1_l. - apply Rdiv_INR_ge_0. } - etrans; last eapply (SeriesC_le_inj _ (λ l, match l with |[x] => Some x | _ => None end)). - * apply Req_le_sym. - apply SeriesC_ext. - intros. case_match; subst. - -- rewrite bool_decide_eq_false_2; first (simpl;lra). - by rewrite elem_of_enum_uniform_fin_list. - -- case_match; last first. - { rewrite bool_decide_eq_false_2; first (simpl;lra). - subst. - by rewrite elem_of_enum_uniform_fin_list. - } - rewrite bool_decide_eq_true_2; last by apply elem_of_enum_uniform_fin_list. - done. - * done. - * intros. repeat case_match; by simplify_eq. - * apply ex_seriesC_finite. - + iIntros ([|?[|]]) "?"; [by iDestruct (ec_contradict with "[$]") as "%"| |by iDestruct (ec_contradict with "[$]") as "%"]. - iFrame. - iMod ("Hrest" with "[$]"). - by iFrame. - Qed. + Admitted. + (* iMod (rand_presample_spec _ _ _ _ _ _ (λ ε' num ε2' ns2, *) + (* ∃ x, ⌜ns2 = [x]⌝ ∗ T ε' (λ x, ε2' ([x])) x *) + (* ) with "[//][$][-]")%I as "H"; first done; last first. *) + (* - iDestruct "H" as "(%&%&%&%&Hfrag &%&%&HT)". *) + (* subst. *) + (* by iFrame. *) + (* - iMod "Hvs" as "(%&%ε2&$&%Hpos & %Hineq & Hrest)". *) + (* iModIntro. *) + (* iExists 1, (λ ls, match ls with |[x] => ε2 x | _ => 1 end)%R. *) + (* repeat iSplit. *) + (* + iPureIntro. by intros [|?[|]]. *) + (* + iPureIntro. etrans; last eapply Hineq. *) + (* rewrite !Rdiv_def pow_1. *) + (* apply Rmult_le_compat_r. *) + (* { rewrite -Rdiv_1_l. *) + (* apply Rdiv_INR_ge_0. } *) + (* etrans; last eapply (SeriesC_le_inj _ (λ l, match l with |[x] => Some x | _ => None end)). *) + (* * apply Req_le_sym. *) + (* apply SeriesC_ext. *) + (* intros. case_match; subst. *) + (* -- rewrite bool_decide_eq_false_2; first (simpl;lra). *) + (* by rewrite elem_of_enum_uniform_fin_list. *) + (* -- case_match; last first. *) + (* { rewrite bool_decide_eq_false_2; first (simpl;lra). *) + (* subst. *) + (* by rewrite elem_of_enum_uniform_fin_list. *) + (* } *) + (* rewrite bool_decide_eq_true_2; last by apply elem_of_enum_uniform_fin_list. *) + (* done. *) + (* * done. *) + (* * intros. repeat case_match; by simplify_eq. *) + (* * apply ex_seriesC_finite. *) + (* + iIntros ([|?[|]]) "?"; [by iDestruct (ec_contradict with "[$]") as "%"| |by iDestruct (ec_contradict with "[$]") as "%"]. *) + (* iFrame. *) + (* iMod ("Hrest" with "[$]"). *) + (* by iFrame. *) + (* Qed. *) Lemma wp_presample_adv_comp_rand_tape (N : nat) NS E α ns (ε1 : R) (ε2 : fin (S N) -> R) γ1: From 9e2e5d11d8b039e859fe58d14b4a981e5e7fc268 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Fri, 4 Oct 2024 18:15:27 +0200 Subject: [PATCH 66/69] break build but useful state_update def --- theories/coneris/error_rules.v | 40 +++ .../examples/random_counter/random_counter.v | 150 +++++---- theories/coneris/lib/hocap.v | 175 ++++++----- theories/coneris/lib/hocap_flip.v | 78 +++-- theories/coneris/lib/hocap_rand.v | 114 ++++--- theories/coneris/wp_update.v | 293 +++++++++++++++++- 6 files changed, 616 insertions(+), 234 deletions(-) diff --git a/theories/coneris/error_rules.v b/theories/coneris/error_rules.v index 3ff7b576..571f9876 100644 --- a/theories/coneris/error_rules.v +++ b/theories/coneris/error_rules.v @@ -1151,6 +1151,46 @@ Qed. iFrame. Qed. + Lemma state_update_presample_exp E α N ns (ε1 : R) (ε2 : fin (S N) → R) : + (∀ n, 0<=ε2 n)%R -> + (SeriesC (λ n, 1 / (S N) * ε2 n)%R <= ε1)%R → + α ↪N (N; ns) -∗ ↯ ε1 -∗ state_update E (∃ n, α ↪N (N; ns ++ [fin_to_nat n]) ∗ ↯ (ε2 n)). + Proof. + rewrite state_update_unseal/state_update_def. + iIntros (Hpos Hsum) "Hα Hε". + iIntros (σ1 ε_now) "[(Hheap&Htapes) Hε_supply]". + iDestruct "Hα" as (ns') "(%Hmap & Hα)". + iDestruct (ghost_map_lookup with "Htapes Hα") as "%Hlookup". + iDestruct (ec_supply_bound with "Hε_supply Hε") as %Hε1_ub. + iMod (ec_supply_decrease with "Hε_supply Hε") as (ε1' ε_rem -> Hε1') "Hε_supply". + iApply fupd_mask_intro; [set_solver|]. + iIntros "Hclose". + subst. + iApply (state_step_coupl_state_adv_comp_con_prob_lang); first done. + iExists (λ x, mknonnegreal (ε2 x) _). + iSplit; first done. + iIntros (sample). + destruct (Rlt_decision (ε_rem + (ε2 sample))%R 1%R) as [Hdec|Hdec]; last first. + { apply Rnot_lt_ge, Rge_le in Hdec. + iApply state_step_coupl_ret_err_ge_1. + simpl. simpl in *. lra. + } + unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2 sample) _) with "Hε_supply") as "[Hε_supply Hε]"; first done. + { simplify_eq. simpl. done. } + iMod (ghost_map_update ((N; ns' ++ [sample]) : tape) with "Htapes Hα") as "[Htapes Hα]". + (* iSpecialize ("Hwp" $! sample). *) + (* rewrite pgl_wp_unfold /pgl_wp_pre. *) + (* simpl. *) + remember {| heap := heap (σ1); tapes := (<[α:=(N; ns' ++ [sample])]> (tapes σ1)) |} as σ2. + rewrite /=. + iModIntro. + iApply state_step_coupl_ret. + iMod "Hclose". + iFrame. + iPureIntro. rewrite fmap_app; by f_equal. + Qed. + + Lemma wp_couple_empty_tape_adv_comp E α N (ε1 : R) (ε2 : nat → R) : (∀ (n:nat), 0<= ε2 n)%R -> (SeriesC (λ n, if (bool_decide (n≤N)) then 1 / (S N) * ε2 n else 0%R)%R <= ε1)%R → diff --git a/theories/coneris/examples/random_counter/random_counter.v b/theories/coneris/examples/random_counter/random_counter.v index b6e22366..1bd7f647 100644 --- a/theories/coneris/examples/random_counter/random_counter.v +++ b/theories/coneris/examples/random_counter/random_counter.v @@ -47,6 +47,14 @@ Class random_counter `{!conerisGS Σ} := RandCounter Forall (λ x, x<=3%nat) ns'-> counter_tapes_auth (L:=L) γ m -∗ counter_tapes_frag (L:=L) γ α ns ==∗ counter_tapes_auth (L:=L) γ (<[α := ns']> m) ∗ counter_tapes_frag (L:=L) γ α (ns'); + counter_tapes_presample {L:counterG Σ} N E γ1 γ2 c α ns ε (ε2 : fin 4%nat -> R): + ↑N ⊆ E -> + (∀ x, 0<=ε2 x)%R -> + (SeriesC (λ n, 1 / 4 * ε2 n)%R <= ε)%R -> + is_counter(L:=L) N c γ1 γ2 -∗ + counter_tapes_frag (L:=L) γ1 α (ns) -∗ + ↯ ε -∗ + state_update E (∃ n, ↯ (ε2 n) ∗ counter_tapes_frag (L:=L) γ1 α (ns ++ [fin_to_nat n])); counter_content_auth_exclusive {L : counterG Σ} γ z1 z2: counter_content_auth (L:=L) γ z1 -∗ counter_content_auth (L:=L) γ z2 -∗ False; @@ -85,29 +93,29 @@ Class random_counter `{!conerisGS Σ} := RandCounter incr_counter_tape c #lbl:α @ E {{{ (z:nat), RET (#z, #n); counter_tapes_frag (L:=L) γ1 α ns ∗ Q z }}}; - counter_presample_spec {L: counterG Σ} NS E T γ1 γ2 c α ns: - ↑NS ⊆ E -> - is_counter (L:=L) NS c γ1 γ2 -∗ - counter_tapes_frag (L:=L) γ1 α ns -∗ - ( |={E∖↑NS,∅}=> - ∃ ε ε2 num, - ↯ ε ∗ - ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ - ⌜(SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R⌝ ∗ - (∀ ns', ↯ (ε2 ns') ={∅,E∖↑NS}=∗ - T ε ε2 num ns' - ))-∗ - wp_update E (∃ ε ε2 num ns', counter_tapes_frag (L:=L) γ1 α (ns++ns') ∗ T ε ε2 num ns'); - read_counter_spec {L: counterG Σ} N E c γ1 γ2 Q: - ↑N ⊆ E -> - {{{ is_counter (L:=L) N c γ1 γ2 ∗ - (∀ (z:nat), counter_content_auth (L:=L) γ2 z ={E∖↑N}=∗ - counter_content_auth (L:=L) γ2 z∗ Q z) + (* counter_presample_spec {L: counterG Σ} NS E T γ1 γ2 c α ns: *) + (* ↑NS ⊆ E -> *) + (* is_counter (L:=L) NS c γ1 γ2 -∗ *) + (* counter_tapes_frag (L:=L) γ1 α ns -∗ *) + (* ( |={E∖↑NS,∅}=> *) + (* ∃ ε ε2 num, *) + (* ↯ ε ∗ *) + (* ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ *) + (* ⌜(SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R⌝ ∗ *) + (* (∀ ns', ↯ (ε2 ns') ={∅,E∖↑NS}=∗ *) + (* T ε ε2 num ns' *) + (* ))-∗ *) + (* wp_update E (∃ ε ε2 num ns', counter_tapes_frag (L:=L) γ1 α (ns++ns') ∗ T ε ε2 num ns'); *) + (* read_counter_spec {L: counterG Σ} N E c γ1 γ2 Q: *) + (* ↑N ⊆ E -> *) + (* {{{ is_counter (L:=L) N c γ1 γ2 ∗ *) + (* (∀ (z:nat), counter_content_auth (L:=L) γ2 z ={E∖↑N}=∗ *) + (* counter_content_auth (L:=L) γ2 z∗ Q z) *) - }}} - read_counter c @ E - {{{ (n':nat), RET #n'; Q n' - }}} + (* }}} *) + (* read_counter c @ E *) + (* {{{ (n':nat), RET #n'; Q n' *) + (* }}} *) }. @@ -135,40 +143,70 @@ Section lemmas. ∃ ε ε2, Q z n ε ε2 }}}. Proof. iIntros (Hsubset Φ) "(#Hinv & Hfrag & Hvs) HΦ". - iMod (counter_presample_spec _ _ - (λ ε ε2 num ns', - ∃ n ε2', - ⌜num=1%nat⌝ ∗ ⌜ns'=[n]⌝ ∗ - ⌜(ε2' = λ x, ε2 [x])%R⌝ ∗ - ( ∀ (z:nat), counter_content_auth (L:=L) γ2 z - ={E∖↑N}=∗ - counter_content_auth (L:=L) γ2 (z+n) ∗ Q z n ε ε2') - ) %I with "[//][$][-HΦ]") as "(%ε & %ε2 & %num & %ns' & Hfrag & %n & %ε2' & -> & -> & -> & Hrest)"; first done. - - iMod "Hvs" as "(%ε & %ε2 & Herr & %Hpos & %Hineq & Hrest)". - iFrame. iModIntro. - iExists (λ l, match l with |[x] => ε2 x | _ => 1 end)%R, 1%nat. - repeat iSplit. - + iPureIntro. intros; repeat case_match; try lra. naive_solver. - + iPureIntro. etrans; last exact. - rewrite SeriesC_list. - * Local Transparent enum_uniform_fin_list. - simpl. lra. - * apply NoDup_fmap_2; last apply NoDup_enum_uniform_fin_list. - apply list_fmap_eq_inj, fin_to_nat_inj. - + iIntros (ns') "Herr". destruct (ns') as [|n[|??]]; [by iDestruct (ec_contradict with "[$]") as "%"| |by iDestruct (ec_contradict with "[$]") as "%"]. - iMod ("Hrest" $! n with "[$]") as "?". - iModIntro. - by iFrame. - - wp_apply (incr_counter_tape_spec_some _ _ _ _ _ (λ z, ∃ ε ε2, Q z n ε ε2)%I with "[$Hfrag Hrest]"); first exact. - + iSplit; first done. - iIntros (?) "?". - iMod ("Hrest" with "[$]") as "[$ ?]". - iModIntro. - iFrame. - + iIntros (?) "[?(%&%&?)]". - iApply "HΦ". - iFrame. - Qed. + iApply (state_update_wp _ (∃ ε ε2 n, + counter_tapes_frag γ1 α [n] ∗ + ∀ z : nat, + counter_content_auth γ2 z ={E ∖ ↑N}=∗ counter_content_auth γ2 (z + n) ∗ Q z n ε ε2) with "[Hvs Hfrag][HΦ]"); last first. + { iIntros "(%ε&%ε2&%n&?&?)". + wp_apply (incr_counter_tape_spec_some _ _ _ _ _ (λ z, Q z n ε ε2) with "[-HΦ]"); [exact|iFrame|]; first by iSplit. + iIntros (?) "[??]". + iApply "HΦ". + iFrame. + } + iMod (fupd_mask_frame _ (↑N) (E∖↑N) ∅ with "[Hvs]"); first set_solver. + - iMod "Hvs" as "H". + iModIntro. + rewrite left_id_L. + replace (_∖(_∖_)) with ((nclose N)); last first. + + set_unfold. clear -Hsubset. firstorder. set_solver. + + iModIntro. iExact "H". + - + (* iMod (fupd_mask_weaken (E1:=E) (E3:=E) (E∖↑N) with "[]"). *) + (* (* iModIntro. *) *) + (* (* iApply state_update_fupd. *) *) + (* iDestruct (fupd_mask_frame_r _ _ (↑N) with "[$Hvs]") as "Hvs"; first set_solver. *) + (* rewrite left_id_L. *) + (* replace (E∖_∪_) with E; last first. *) + (* { set_unfold. admit. } *) + (* iMod (counter_tapes_presample N E with "[$][$][Hvs]"). *) + Admitted. + + + (* iMod "Hvs". *) + (* iMod (counter_presample_spec _ _ *) + (* (λ ε ε2 num ns', *) + (* ∃ n ε2', *) + (* ⌜num=1%nat⌝ ∗ ⌜ns'=[n]⌝ ∗ *) + (* ⌜(ε2' = λ x, ε2 [x])%R⌝ ∗ *) + (* ( ∀ (z:nat), counter_content_auth (L:=L) γ2 z *) + (* ={E∖↑N}=∗ *) + (* counter_content_auth (L:=L) γ2 (z+n) ∗ Q z n ε ε2') *) + (* ) %I with "[//][$][-HΦ]") as "(%ε & %ε2 & %num & %ns' & Hfrag & %n & %ε2' & -> & -> & -> & Hrest)"; first done. *) + (* - iMod "Hvs" as "(%ε & %ε2 & Herr & %Hpos & %Hineq & Hrest)". *) + (* iFrame. iModIntro. *) + (* iExists (λ l, match l with |[x] => ε2 x | _ => 1 end)%R, 1%nat. *) + (* repeat iSplit. *) + (* + iPureIntro. intros; repeat case_match; try lra. naive_solver. *) + (* + iPureIntro. etrans; last exact. *) + (* rewrite SeriesC_list. *) + (* * Local Transparent enum_uniform_fin_list. *) + (* simpl. lra. *) + (* * apply NoDup_fmap_2; last apply NoDup_enum_uniform_fin_list. *) + (* apply list_fmap_eq_inj, fin_to_nat_inj. *) + (* + iIntros (ns') "Herr". destruct (ns') as [|n[|??]]; [by iDestruct (ec_contradict with "[$]") as "%"| |by iDestruct (ec_contradict with "[$]") as "%"]. *) + (* iMod ("Hrest" $! n with "[$]") as "?". *) + (* iModIntro. *) + (* by iFrame. *) + (* - wp_apply (incr_counter_tape_spec_some _ _ _ _ _ (λ z, ∃ ε ε2, Q z n ε ε2)%I with "[$Hfrag Hrest]"); first exact. *) + (* + iSplit; first done. *) + (* iIntros (?) "?". *) + (* iMod ("Hrest" with "[$]") as "[$ ?]". *) + (* iModIntro. *) + (* iFrame. *) + (* + iIntros (?) "[?(%&%&?)]". *) + (* iApply "HΦ". *) + (* iFrame. *) + (* Qed. *) Lemma incr_counter_tape_spec_none' N E c γ1 γ2 ε (ε2:nat -> R)(α:loc) (ns:list nat) (q:Qp) (z:nat): ↑N ⊆ E-> diff --git a/theories/coneris/lib/hocap.v b/theories/coneris/lib/hocap.v index b616d647..df41f837 100644 --- a/theories/coneris/lib/hocap.v +++ b/theories/coneris/lib/hocap.v @@ -142,103 +142,102 @@ Set Default Proof Using "Type*". (* End error_lemmas. *) -Section hocap_error_coupl. - Context `{conerisGS Σ}. - Context (tb:nat). - Local Canonical Structure finO := leibnizO (fin (S tb)). - Local Canonical Structure RO := leibnizO R. - Local Canonical Structure ε2O := leibnizO (list (fin(S tb))->R). - - - Definition hocap_error_coupl_pre Z Φ : (R * R * list (nat * (list (fin (S tb)) -> R) * (list (fin (S tb)))) -> iProp Σ) := - (λ x, - let '(ε, ε_initial, ls) := x in - (Z ε ε_initial ls)∨ - ∃ num ε2, - ⌜(∀ l, 0<=ε2 l)%R⌝ ∗ - ⌜(SeriesC (λ l, if bool_decide (length l = num) then ε2 l else 0)/((S tb)^num) <= ε)%R⌝ ∗ - (∀ ns, Φ (ε2 ns, ε_initial, ls ++ [(num, ε2, ns)])) - )%I. +(* Section hocap_error_coupl. *) +(* Context `{conerisGS Σ}. *) +(* Context (tb:nat). *) +(* Local Canonical Structure finO := leibnizO (fin (S tb)). *) +(* Local Canonical Structure RO := leibnizO R. *) +(* Local Canonical Structure ε2O := leibnizO (list (fin(S tb))->R). *) - Local Instance hocap_error_coupl_pre_ne Z Φ : - NonExpansive (hocap_error_coupl_pre Z Φ). - Proof. - solve_contractive. - Qed. +(* Definition hocap_error_coupl_pre Z Φ : (R * R * list (nat * (list (fin (S tb)) -> R) * (list (fin (S tb)))) -> iProp Σ) := *) +(* (λ x, *) +(* let '(ε, ε_initial, ls) := x in *) +(* (Z ε ε_initial ls)∨ *) +(* ∃ num ε2, *) +(* ⌜(∀ l, 0<=ε2 l)%R⌝ ∗ *) +(* ⌜(SeriesC (λ l, if bool_decide (length l = num) then ε2 l else 0)/((S tb)^num) <= ε)%R⌝ ∗ *) +(* (∀ ns, Φ (ε2 ns, ε_initial, ls ++ [(num, ε2, ns)])) *) +(* )%I. *) + +(* Local Instance hocap_error_coupl_pre_ne Z Φ : *) +(* NonExpansive (hocap_error_coupl_pre Z Φ). *) +(* Proof. *) +(* solve_contractive. *) +(* Qed. *) - Local Instance hocap_error_coupl_pre_mono Z : BiMonoPred (hocap_error_coupl_pre Z). - Proof. - split; last apply _. - iIntros (Φ Ψ HNEΦ HNEΨ) "#Hwand". - iIntros ([[??]?]) "[H|(%&%&%&%&H)]". - - by iLeft. - - iRight. - iExists _, _. - repeat iSplit; try done. - iIntros. by iApply "Hwand". - Qed. +(* Local Instance hocap_error_coupl_pre_mono Z : BiMonoPred (hocap_error_coupl_pre Z). *) +(* Proof. *) +(* split; last apply _. *) +(* iIntros (Φ Ψ HNEΦ HNEΨ) "#Hwand". *) +(* iIntros ([[??]?]) "[H|(%&%&%&%&H)]". *) +(* - by iLeft. *) +(* - iRight. *) +(* iExists _, _. *) +(* repeat iSplit; try done. *) +(* iIntros. by iApply "Hwand". *) +(* Qed. *) - Definition hocap_error_coupl' Z := bi_least_fixpoint (hocap_error_coupl_pre Z). - Definition hocap_error_coupl ε_current ε_initial ls Z := hocap_error_coupl' Z (ε_current, ε_initial, ls). +(* Definition hocap_error_coupl' Z := bi_least_fixpoint (hocap_error_coupl_pre Z). *) +(* Definition hocap_error_coupl ε_current ε_initial ls Z := hocap_error_coupl' Z (ε_current, ε_initial, ls). *) - Lemma hocap_error_coupl_unfold ε ε_initial ls Z : - hocap_error_coupl ε ε_initial ls Z ≡ - ((Z ε ε_initial ls)∨ - ∃ num ε2, - ⌜(∀ l, 0<=ε2 l)%R⌝ ∗ - ⌜(SeriesC (λ l, if bool_decide (length l = num) then ε2 l else 0)/((S tb)^num) <= ε)%R⌝ ∗ - (∀ ns, hocap_error_coupl (ε2 ns) ε_initial (ls ++ [(num, ε2, ns)]) Z))%I. - Proof. - rewrite /hocap_error_coupl/hocap_error_coupl' least_fixpoint_unfold//. Qed. +(* Lemma hocap_error_coupl_unfold ε ε_initial ls Z : *) +(* hocap_error_coupl ε ε_initial ls Z ≡ *) +(* ((Z ε ε_initial ls)∨ *) +(* ∃ num ε2, *) +(* ⌜(∀ l, 0<=ε2 l)%R⌝ ∗ *) +(* ⌜(SeriesC (λ l, if bool_decide (length l = num) then ε2 l else 0)/((S tb)^num) <= ε)%R⌝ ∗ *) +(* (∀ ns, hocap_error_coupl (ε2 ns) ε_initial (ls ++ [(num, ε2, ns)]) Z))%I. *) +(* Proof. *) +(* rewrite /hocap_error_coupl/hocap_error_coupl' least_fixpoint_unfold//. Qed. *) - Lemma hocap_error_coupl_ret ε ε_initial ls Z: - Z ε ε_initial ls -∗ hocap_error_coupl ε ε_initial ls Z. - Proof. - iIntros. rewrite hocap_error_coupl_unfold. by iLeft. - Qed. +(* Lemma hocap_error_coupl_ret ε ε_initial ls Z: *) +(* Z ε ε_initial ls -∗ hocap_error_coupl ε ε_initial ls Z. *) +(* Proof. *) +(* iIntros. rewrite hocap_error_coupl_unfold. by iLeft. *) +(* Qed. *) - Lemma hocap_error_coupl_rec ε ε_initial ls Z: - (∃ num ε2, - ⌜(∀ l, 0<=ε2 l)%R⌝ ∗ - ⌜(SeriesC (λ l, if bool_decide (length l = num) then ε2 l else 0)/((S tb)^num) <= ε)%R⌝ ∗ - (∀ ns, hocap_error_coupl (ε2 ns) ε_initial (ls ++ [(num, ε2, ns)]) Z)) -∗ hocap_error_coupl ε ε_initial ls Z. - Proof. - iIntros. rewrite hocap_error_coupl_unfold. by iRight. - Qed. +(* Lemma hocap_error_coupl_rec ε ε_initial ls Z: *) +(* (∃ num ε2, *) +(* ⌜(∀ l, 0<=ε2 l)%R⌝ ∗ *) +(* ⌜(SeriesC (λ l, if bool_decide (length l = num) then ε2 l else 0)/((S tb)^num) <= ε)%R⌝ ∗ *) +(* (∀ ns, hocap_error_coupl (ε2 ns) ε_initial (ls ++ [(num, ε2, ns)]) Z)) -∗ hocap_error_coupl ε ε_initial ls Z. *) +(* Proof. *) +(* iIntros. rewrite hocap_error_coupl_unfold. by iRight. *) +(* Qed. *) - Lemma hocap_error_coupl_ind (Ψ Z : R -> R -> list (nat * (list (fin (S tb)) -> R) * (list (fin (S tb))))->iProp Σ): - ⊢ (□ (∀ ε ε_initial ls, - hocap_error_coupl_pre Z (λ '(ε', ε_initial', ls'), - Ψ ε' ε_initial' ls' ∧ hocap_error_coupl ε' ε_initial' ls' Z)%I (ε, ε_initial, ls) -∗ Ψ ε ε_initial ls) → - ∀ ε ε_initial ls, hocap_error_coupl ε ε_initial ls Z -∗ Ψ ε ε_initial ls)%I. - Proof. - iIntros "#IH" (ε ε_initial ls) "H". - set (Ψ' := (λ '(ε, ε_initial, ls), Ψ ε ε_initial ls) : (prodO _ _->iProp Σ)). - assert (NonExpansive Ψ') by solve_contractive. - iApply (least_fixpoint_ind _ Ψ' with "[] H"). - iIntros "!#" ([[??]?]) "H". by iApply "IH". - Qed. +(* Lemma hocap_error_coupl_ind (Ψ Z : R -> R -> list (nat * (list (fin (S tb)) -> R) * (list (fin (S tb))))->iProp Σ): *) +(* ⊢ (□ (∀ ε ε_initial ls, *) +(* hocap_error_coupl_pre Z (λ '(ε', ε_initial', ls'), *) +(* Ψ ε' ε_initial' ls' ∧ hocap_error_coupl ε' ε_initial' ls' Z)%I (ε, ε_initial, ls) -∗ Ψ ε ε_initial ls) → *) +(* ∀ ε ε_initial ls, hocap_error_coupl ε ε_initial ls Z -∗ Ψ ε ε_initial ls)%I. *) +(* Proof. *) +(* iIntros "#IH" (ε ε_initial ls) "H". *) +(* set (Ψ' := (λ '(ε, ε_initial, ls), Ψ ε ε_initial ls) : (prodO _ _->iProp Σ)). *) +(* assert (NonExpansive Ψ') by solve_contractive. *) +(* iApply (least_fixpoint_ind _ Ψ' with "[] H"). *) +(* iIntros "!#" ([[??]?]) "H". by iApply "IH". *) +(* Qed. *) - Lemma hocap_error_coupl_bind ε ε_initial ls Z1 Z2 : - (∀ ε' ε_initial' ls', Z1 ε' ε_initial' ls' -∗ hocap_error_coupl ε' ε_initial' ls' Z2) -∗ - hocap_error_coupl ε ε_initial ls Z1 -∗ - hocap_error_coupl ε ε_initial ls Z2. - Proof. - iIntros "HZ H". - iRevert "HZ". - iRevert (ε ε_initial ls) "H". - iApply hocap_error_coupl_ind. - iIntros "!#" (ε ε_initial ls) "[H|H] HZ". - - by iApply "HZ". - - iApply hocap_error_coupl_rec. - iDestruct "H" as "(%&%&%&%&Hrest)". - iExists _,_. repeat iSplit; try done. - iIntros (?). - iDestruct ("Hrest" $! _) as "[Hrest _]". - by iApply "Hrest". - Qed. +(* Lemma hocap_error_coupl_bind ε ε_initial ls Z1 Z2 : *) +(* (∀ ε' ε_initial' ls', Z1 ε' ε_initial' ls' -∗ hocap_error_coupl ε' ε_initial' ls' Z2) -∗ *) +(* hocap_error_coupl ε ε_initial ls Z1 -∗ *) +(* hocap_error_coupl ε ε_initial ls Z2. *) +(* Proof. *) +(* iIntros "HZ H". *) +(* iRevert "HZ". *) +(* iRevert (ε ε_initial ls) "H". *) +(* iApply hocap_error_coupl_ind. *) +(* iIntros "!#" (ε ε_initial ls) "[H|H] HZ". *) +(* - by iApply "HZ". *) +(* - iApply hocap_error_coupl_rec. *) +(* iDestruct "H" as "(%&%&%&%&Hrest)". *) +(* iExists _,_. repeat iSplit; try done. *) +(* iIntros (?). *) +(* iDestruct ("Hrest" $! _) as "[Hrest _]". *) +(* by iApply "Hrest". *) +(* Qed. *) -End hocap_error_coupl. +(* End hocap_error_coupl. *) Definition hocap_tapes_nroot:=nroot.@"tapes". Class hocap_tapesGS (Σ : gFunctors) := Hocap_tapesGS { diff --git a/theories/coneris/lib/hocap_flip.v b/theories/coneris/lib/hocap_flip.v index 55b70bae..3df2d429 100644 --- a/theories/coneris/lib/hocap_flip.v +++ b/theories/coneris/lib/hocap_flip.v @@ -36,6 +36,14 @@ Class flip_spec `{!conerisGS Σ} := FlipSpec flip_tapes_update {L : flipG Σ} γ α m ns ns': flip_tapes_auth (L:=L) γ m -∗ flip_tapes_frag (L:=L) γ α ns ==∗ flip_tapes_auth (L:=L) γ (<[α := ns']> m) ∗ flip_tapes_frag (L:=L) γ α ns'; + flip_tapes_presample {L:flipG Σ} N E γ α ns ε (ε2 : bool -> R): + ↑N ⊆ E -> + (∀ x, 0<=ε2 x)%R -> + ((ε2 true + ε2 false)/2 <= ε)%R -> + is_flip (L:=L) N γ -∗ + flip_tapes_frag (L:=L) γ α (ns) -∗ + ↯ ε -∗ + state_update E (∃ n, ↯ (ε2 n) ∗ flip_tapes_frag (L:=L) γ α (ns ++ [n])); (** * Program specs *) flip_inv_create_spec {L : flipG Σ} N E: @@ -56,18 +64,18 @@ Class flip_spec `{!conerisGS Σ} := FlipSpec flip_tape #lbl:α @ E {{{ RET #n; flip_tapes_frag (L:=L) γ1 α ns }}}; - flip_presample_spec {L: flipG Σ} NS E γ1 α ns T: - ↑NS ⊆ E -> - is_flip (L:=L) NS γ1 -∗ - flip_tapes_frag (L:=L) γ1 α ns -∗ - (|={E∖↑NS, ∅}=> - ∃ ε num ε2, ↯ ε ∗ - ⌜(∀ l, length l = num -> 0<=ε2 l)%R⌝ ∗ - ⌜(SeriesC (λ l, if length l =? num then ε2 l else 0) /(2^num) <= ε)%R⌝∗ - (∀ ns', ↯ (ε2 ns') - ={∅, E∖↑NS}=∗ T ε num ε2 ns')) - -∗ - wp_update E (∃ ε num ε2 ns', flip_tapes_frag (L:=L) γ1 α (ns ++ ns') ∗ T ε num ε2 ns') + (* flip_presample_spec {L: flipG Σ} NS E γ1 α ns T: *) + (* ↑NS ⊆ E -> *) + (* is_flip (L:=L) NS γ1 -∗ *) + (* flip_tapes_frag (L:=L) γ1 α ns -∗ *) + (* (|={E∖↑NS, ∅}=> *) + (* ∃ ε num ε2, ↯ ε ∗ *) + (* ⌜(∀ l, length l = num -> 0<=ε2 l)%R⌝ ∗ *) + (* ⌜(SeriesC (λ l, if length l =? num then ε2 l else 0) /(2^num) <= ε)%R⌝∗ *) + (* (∀ ns', ↯ (ε2 ns') *) + (* ={∅, E∖↑NS}=∗ T ε num ε2 ns')) *) + (* -∗ *) + (* wp_update E (∃ ε num ε2 ns', flip_tapes_frag (L:=L) γ1 α (ns ++ ns') ∗ T ε num ε2 ns') *) }. @@ -117,6 +125,18 @@ Section instantiate_flip. - iModIntro. rewrite fmap_insert; iFrame. Qed. + Next Obligation. + simpl. + iIntros (???????????) "#Hinv Hfrag Hε". + iMod (rand_tapes_presample _ _ _ _ _ _ _ (λ x, ε2 (nat_to_bool (fin_to_nat x)))with "[$][$][$]") as "(%n&?&?)"; try done. + - rewrite SeriesC_finite_foldr/=. + rewrite nat_to_bool_eq_0 nat_to_bool_neq_0; last lia. + lra. + - iFrame. + iModIntro. + rewrite fmap_app. + by repeat (inv_fin n; try (intros n)); simpl. + Qed. Next Obligation. simpl. iIntros (???). @@ -141,10 +161,9 @@ Section instantiate_flip. + rewrite Z_to_bool_neq_0; lia. + rewrite Z_to_bool_eq_0; lia. Qed. - Next Obligation. - simpl. - iIntros (?????? T ?) "#Hinv Hfrag Hvs". - Admitted. + (* Next Obligation. *) + (* simpl. *) + (* iIntros (?????? T ?) "#Hinv Hfrag Hvs". *) (* iMod (rand_presample_spec _ _ _ _ _ _ *) (* (λ ε num ε2 ns', *) (* ∃ bs' ε2', ⌜fmap (FMap:= list_fmap) bool_to_nat bs'=fin_to_nat <$> ns'⌝ ∗ *) @@ -217,29 +236,8 @@ Section test. wp_update E (∃ b, flip_tapes_frag (L:=L) γ1 α (ns ++ [b]) ∗ ↯ (ε2 b)). Proof. iIntros (Hsubset Hpos Hineq) "#Hinv Hfrag Herr". - pose (ε2' l:= match l with - | [b]=> ε2 b - | _ => 1%R - end - ). - iMod (flip_presample_spec _ _ _ _ _ - (λ ε' num ε2'' ns', ⌜ε=ε'⌝ ∗ ⌜num=1⌝ ∗ ⌜ε2' = ε2''⌝ ∗ - ∃ x, ⌜ns'=[x]⌝ ∗ ↯ (ε2 x) - )%I with "[//][$][Herr]") as "(%&%&%&%&?&%&%&%&%&%&?)"; first done. - - iFrame. - iApply fupd_mask_intro; first set_solver. - iIntros "Hclose". - iExists 1, ε2'. - repeat iSplit. - + iPureIntro. - intros [|?[|??]]; by simpl. - + iPureIntro. - setoid_rewrite <-elem_of_enum_uniform_list'. - rewrite SeriesC_list; last apply NoDup_enum_uniform_list. - simpl. lra. - + iIntros ([|?[|??]]) "Herr"; simpl; try (by iDestruct (ec_contradict with "[$]") as "%"). - iMod "Hclose" as "_". - by iFrame. - - subst. by iFrame. + iApply wp_update_state_update. + iMod (flip_tapes_presample with "[$][$][$]") as "(%&?&?)"; try done. + by iFrame. Qed. End test. diff --git a/theories/coneris/lib/hocap_rand.v b/theories/coneris/lib/hocap_rand.v index 48843129..4e7f9073 100644 --- a/theories/coneris/lib/hocap_rand.v +++ b/theories/coneris/lib/hocap_rand.v @@ -40,6 +40,15 @@ Class rand_spec `{!conerisGS Σ} := RandSpec Forall (λ x, x<=ns'.1) ns'.2 -> rand_tapes_auth (L:=L) γ m -∗ rand_tapes_frag (L:=L) γ α ns ==∗ rand_tapes_auth (L:=L) γ (<[α := ns']> m) ∗ rand_tapes_frag (L:=L) γ α ns'; + rand_tapes_presample {L:randG Σ} N E γ α tb ns ε (ε2 : fin (S tb) -> R): + ↑N ⊆ E -> + (∀ x, 0<=ε2 x)%R -> + (SeriesC (λ n, 1 / (S tb) * ε2 n)%R <= ε)%R -> + is_rand(L:=L) N γ -∗ + rand_tapes_frag (L:=L) γ α (tb, ns) -∗ + ↯ ε -∗ + state_update E (∃ n, ↯ (ε2 n) ∗ rand_tapes_frag (L:=L) γ α (tb, ns ++ [fin_to_nat n])); + (** * Program specs *) rand_inv_create_spec {L : randG Σ} N E: @@ -58,21 +67,21 @@ Class rand_spec `{!conerisGS Σ} := RandSpec }}} rand_tape #lbl:α #tb @ E {{{ RET #n; rand_tapes_frag (L:=L) γ2 α (tb, ns) }}}; - rand_presample_spec {L: randG Σ} N E γ2 α (tb:nat) ns T: - ↑N ⊆ E -> - is_rand (L:=L) N γ2 -∗ - rand_tapes_frag (L:=L) γ2 α (tb, ns) -∗ - (|={E∖↑N, ∅}=> - ∃ ε, - ↯ ε ∗ - hocap_error_coupl tb ε ε [] - (λ ε_final ε_initial ls, - ↯ ε_final ={∅, E∖↑N}=∗ T ε_final ε_initial ls - ) - ) -∗ - wp_update E (∃ ε_final ε_initial (ls:list _), - rand_tapes_frag (L:=L) γ2 α (tb, ns ++ (fin_to_nat <$> (mbind (MBind := list_bind) snd ls))) ∗ - T ε_final ε_initial ls) + (* rand_presample_spec {L: randG Σ} N E γ2 α (tb:nat) ns T: *) + (* ↑N ⊆ E -> *) + (* is_rand (L:=L) N γ2 -∗ *) + (* rand_tapes_frag (L:=L) γ2 α (tb, ns) -∗ *) + (* (|={E∖↑N, ∅}=> *) + (* ∃ ε, *) + (* ↯ ε ∗ *) + (* hocap_error_coupl tb ε ε [] *) + (* (λ ε_final ε_initial ls, *) + (* ↯ ε_final ={∅, E∖↑N}=∗ T ε_final ε_initial ls *) + (* ) *) + (* ) -∗ *) + (* wp_update E (∃ ε_final ε_initial (ls:list _), *) + (* rand_tapes_frag (L:=L) γ2 α (tb, ns ++ (fin_to_nat <$> (mbind (MBind := list_bind) snd ls))) ∗ *) + (* T ε_final ε_initial ls) *) (* rand_presample_spec {L: randG Σ} N E γ2 α (tb:nat) ns T: *) (* ↑N ⊆ E -> *) @@ -134,6 +143,21 @@ Section impl. iMod (hocap_tapes_update with "[$][$]") as "[??]". by iFrame. Qed. + Next Obligation. + simpl. + iIntros (??????????????) "#H1 [H2 %] H3". + iInv "H1" as ">(%&?&?)" "Hclose". + iDestruct (hocap_tapes_agree with "[$][$]") as "%". + iDestruct (big_sepM_insert_acc with "[$]") as "[Htape H]"; first done. + iMod (state_update_presample_exp with "[$][$]") as "(%n & Htape & ?)"; try exact. + iMod (hocap_tapes_update with "[$][$]") as "[? ?]". + iModIntro. iFrame. + iMod ("Hclose" with "[-]"). + { iNext. iFrame. by iApply "H". } + iPureIntro. rewrite Forall_app. split; first done. + apply Forall_singleton. + pose proof fin_to_nat_lt n. simpl in *. lia. + Qed. Next Obligation. simpl. iIntros (?????). @@ -177,14 +201,14 @@ Section impl. iApply "HΦ". iFrame. iPureIntro. by eapply Forall_inv_tail. Qed. - Next Obligation. - simpl. - iIntros (???????????) "#Hinv [Hfrag %] Hvs". - iApply wp_update_state_step_coupl. - iIntros (σ1 ε_supply) "((Hheap&Htapes)&Hε)". - iMod (inv_acc with "Hinv") as "[>(% & H3 & H4 ) Hclose]"; [done|]. - iMod "Hvs" as "(%err & Herr & Hcoupl)". - Admitted. + (* Next Obligation. *) + (* simpl. *) + (* iIntros (???????????) "#Hinv [Hfrag %] Hvs". *) + (* iApply wp_update_state_step_coupl. *) + (* iIntros (σ1 ε_supply) "((Hheap&Htapes)&Hε)". *) + (* iMod (inv_acc with "Hinv") as "[>(% & H3 & H4 ) Hclose]"; [done|]. *) + (* iMod "Hvs" as "(%err & Herr & Hcoupl)". *) + (* Admitted. *) (* iMod "Hvs" as "(%err & %num & %ε2 & Herr & %Hpos & %Hineq & Hrest)". *) (* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) @@ -264,21 +288,22 @@ Section checks. Local Opaque enum_uniform_fin_list. - Lemma rand_presample_single_spec N E γ2 α (tb:nat) ns T: - ↑N ⊆ E -> - is_rand (L:=L) N γ2 -∗ - rand_tapes_frag (L:=L) γ2 α (tb, ns) -∗ - (|={E∖↑N, ∅}=> - ∃ ε (ε2 : fin (S tb) -> R), - ↯ ε ∗ - ⌜(∀ x, 0<= ε2 x)%R⌝ ∗ - ⌜(SeriesC ε2 /((S tb)) <= ε)%R⌝ ∗ - (∀ x, ↯ (ε2 x) ={∅, E∖↑N}=∗ T ε ε2 x)) -∗ - wp_update E (∃ ε ε2 x, rand_tapes_frag (L:=L) γ2 α (tb, ns ++ [fin_to_nat x]) ∗ - T ε ε2 x). - Proof. - iIntros (Hsubset) "#Hinv Hfrag Hvs". - Admitted. + (* Lemma rand_presample_single_spec N E γ2 α (tb:nat) ns T: *) + (* ↑N ⊆ E -> *) + (* is_rand (L:=L) N γ2 -∗ *) + (* rand_tapes_frag (L:=L) γ2 α (tb, ns) -∗ *) + (* (|={E∖↑N, ∅}=> *) + (* ∃ ε (ε2 : fin (S tb) -> R), *) + (* ↯ ε ∗ *) + (* ⌜(∀ x, 0<= ε2 x)%R⌝ ∗ *) + (* ⌜(SeriesC ε2 /((S tb)) <= ε)%R⌝ ∗ *) + (* (∀ x, ↯ (ε2 x) ={∅, E∖↑N}=∗ T ε ε2 x)) -∗ *) + (* wp_update E (∃ ε ε2 x, rand_tapes_frag (L:=L) γ2 α (tb, ns ++ [fin_to_nat x]) ∗ *) + (* T ε ε2 x). *) + (* Proof. *) + (* iIntros (Hsubset) "#Hinv Hfrag Hvs". *) + (* iApply wp_update_state_step_coupl. *) + (* iIntros (??) "[??]". *) (* iMod (rand_presample_spec _ _ _ _ _ _ (λ ε' num ε2' ns2, *) (* ∃ x, ⌜ns2 = [x]⌝ ∗ T ε' (λ x, ε2' ([x])) x *) (* ) with "[//][$][-]")%I as "H"; first done; last first. *) @@ -330,17 +355,8 @@ Section checks. iIntros (Hsubset Hpos Hineq) "#Hinv >Htape Herr". iDestruct (ec_valid with "[$]") as "%Hpos'". destruct Hpos' as [Hpos' ?]. - iMod (rand_presample_single_spec _ _ _ _ _ _ (λ ε' ε2' n, ↯ (ε2 n)) with "[//][$][-]") as "(%&%&%&?&?)"; first done; last by iFrame. - iFrame. - iApply fupd_mask_intro; first set_solver. - iIntros "Hclose". - iExists ε2. - repeat iSplit. - - done. - - iPureIntro. etrans; last apply Hineq. - rewrite SeriesC_scal_l. - lra. - - iIntros. iFrame. + iApply wp_update_state_update. + by iApply (rand_tapes_presample with "[$][$][$]"). Qed. diff --git a/theories/coneris/wp_update.v b/theories/coneris/wp_update.v index a6a80753..9d868f61 100644 --- a/theories/coneris/wp_update.v +++ b/theories/coneris/wp_update.v @@ -201,7 +201,7 @@ Section wp_update. iApply HP. iFrame. Qed. - Global Instance is_except_0_pgl_wp E Q : IsExcept0 (wp_update E Q). + Global Instance is_except_0_wp_update E Q : IsExcept0 (wp_update E Q). Proof. by rewrite /IsExcept0 -{2}fupd_wp_update -except_0_fupd -fupd_intro. Qed. @@ -238,3 +238,294 @@ Section wp_update. End wp_update. +Section state_update. + Context `{!conerisGS Σ}. + Definition state_update_def E P:= + (∀ σ1 ε1, + state_interp σ1 ∗ err_interp ε1 ={E, ∅}=∗ + state_step_coupl σ1 ε1 (λ σ2 ε2, + |={∅, E}=> state_interp σ2 ∗ err_interp ε2 ∗ P + ) + )%I. + + Local Definition state_update_aux : seal (@state_update_def). Proof. by eexists. Qed. + Definition state_update := state_update_aux.(unseal). + Lemma state_update_unseal : state_update = state_update_def. + Proof. rewrite -state_update_aux.(seal_eq) //. Qed. + + Lemma wp_update_state_update E P: + state_update E P -∗ wp_update E P. + Proof. + rewrite state_update_unseal/state_update_def. + iIntros. + by iApply wp_update_state_step_coupl. + Qed. + + Lemma state_update_ret E P: + P -∗ state_update E P. + Proof. + rewrite state_update_unseal/state_update_def. + iIntros. + iApply fupd_mask_intro; first set_solver. + iIntros "Hclose". + iApply state_step_coupl_ret. + iMod "Hclose". + iFrame. by iModIntro. + Qed. + + Global Instance from_modal_state_update_state_update P E : + FromModal True modality_id (state_update E P) (state_update E P) P. + Proof. iIntros (_) "HP /=". by iApply state_update_ret. Qed. + + Lemma state_update_mono_fupd E1 E2 P: + E1 ⊆ E2 -> (|={E2,E1}=> state_update E1 (|={E1, E2}=> P)) -∗ state_update E2 P. + Proof. + rewrite state_update_unseal/state_update_def. + iIntros (Hsubseteq) "Hvs". + iIntros (σ1 ε1) "[H1 H2]". + iMod ("Hvs" with "[$]") as ">?". + iModIntro. + iApply (state_step_coupl_mono with "[][$]"). + iIntros (??) ">(?&?&>?)". by iFrame. + Qed. + + + Lemma state_update_mono E P Q: + (P={E}=∗Q) -∗ state_update E P -∗ state_update E Q. + Proof. + rewrite state_update_unseal/state_update_def. + iIntros "Hvs H". + iIntros (σ1 ε1) "[H1 H2]". + iMod ("H" with "[$]") as "?". + iModIntro. + iApply (state_step_coupl_mono with "[Hvs][$]"). + iIntros (??) ">(?&?&?)". + iFrame. + by iApply "Hvs". + Qed. + + Lemma state_update_mono_fupd' E1 E2 P: + E1 ⊆ E2 -> (state_update E1 P) -∗ state_update E2 P. + Proof. + iIntros. + iApply state_update_mono_fupd; first done. + iApply fupd_mask_intro; first done. + iIntros "Hclose". + iApply (state_update_mono with "[Hclose][$]"). + iIntros. + iModIntro. + by iMod "Hclose". + Qed. + + Lemma state_update_fupd E P: + (|={E}=> state_update E P) -∗ state_update E P. + Proof. + iIntros "H". + iApply state_update_mono_fupd; first done. + iMod "H". + iApply state_update_mono; last done. + by iIntros. + Qed. + + Lemma state_update_fupd_change E1 E2 P Q: + (|={E1, E2}=> P) -∗ (P-∗ state_update E2 (|={E2, E1}=> Q)) -∗ state_update E1 Q. + Proof. + iIntros "H1 H2". + rewrite state_update_unseal/state_update_def. + iIntros (??) "[??]". + iMod "H1". + iMod ("H2" with "[$][$]") as "H2". + iModIntro. + iApply state_step_coupl_mono; last done. + simpl. + iIntros (??) ">(?&?&>?)". + by iFrame. + Qed. + + Global Instance elim_modal_bupd_state_update p E P Q : + ElimModal True p false (|==> P) P (state_update E Q) (state_update E Q). + Proof. + intros ?. + rewrite bi.intuitionistically_if_elim/=. + iIntros "[H1 H2]". + iApply state_update_fupd. + iMod "H1". + iModIntro. + by iApply "H2". + Qed. + + Global Instance elim_modal_fupd_state_update p E1 E2 P Q : + ElimModal (True) p false + (|={E1,E2}=> P) P + (state_update E1 Q) (state_update E2 (|={E2, E1}=> Q)). + Proof. + intros ?. + rewrite bi.intuitionistically_if_elim/=. + iIntros "[??]". + iApply (state_update_fupd_change with "[$][$]"). + Qed. + + Lemma state_update_bind E P Q: + state_update E P ∗ (P -∗ state_update E Q) ⊢ state_update E Q. + Proof. + rewrite state_update_unseal/state_update_def. + iIntros "[H1 H2]" (??) "[??]". + iMod ("H1" with "[$]") as "H1". + iModIntro. + iApply (state_step_coupl_bind with "[H2][$]"). + iIntros (??) "H1". + iApply fupd_state_step_coupl. + iMod "H1" as "(?&?&?)". + by iMod ("H2" with "[$][$]"). + Qed. + + Global Instance elim_modal_state_update_state_update p E1 E2 P Q : + ElimModal (E1 ⊆ E2) p false (state_update E1 Q) Q (state_update E2 P) (state_update E2 P). + Proof. + iIntros (?) "[H1 H2]". + rewrite bi.intuitionistically_if_elim. + iApply state_update_mono_fupd; first exact. + iApply fupd_mask_intro; first exact; iIntros "Hclose". + iApply state_update_bind; iFrame. + iIntros. + iApply (state_update_fupd_change with "[$]"). + iIntros. + iDestruct ("H2" with "[$]") as "H". + iApply (state_update_mono); last done. + iIntros. + iModIntro. + by iApply fupd_mask_intro_subseteq. + Qed. + + Global Instance elim_modal_state_update_wp_update p E1 E2 P Q : + ElimModal (E1 ⊆ E2) p false (state_update E1 Q) Q (wp_update E2 P) (wp_update E2 P). + Proof. + (* rewrite state_update_unseal/state_update_def. *) + iIntros (?) "[H1 H2]". + iApply (wp_update_bind). + iFrame. + iApply wp_update_state_update. + iIntros. + iApply state_update_mono_fupd; first exact. + iApply fupd_mask_intro; first done. + iIntros "Hclose". + destruct p; simpl. + - iDestruct (bi.intuitionistically_elim with "[$]") as "H1". + iMod "H1". + iMod "Hclose" as "_". + iModIntro. + by iApply fupd_mask_intro_subseteq. + - iMod "H1". + iMod "Hclose" as "_". + iModIntro. + by iApply fupd_mask_intro_subseteq. + Qed. + + Global Instance elim_modal_state_update_wp e Φ p E1 E2 P : + ElimModal (E1 ⊆ E2) p false (state_update E1 P) P (WP e @ E2 {{ Φ }}) (WP e @ E2 {{ Φ }}). + Proof. + destruct p. + all: iIntros (?); simpl; iIntros "[H1 H2]". + 1: iDestruct (bi.intuitionistically_elim with "[$]") as "H1". + all: iDestruct (state_update_mono_fupd with "[H1]") as "H1"; first exact; + last (iDestruct (wp_update_state_update with "[$]") as "H1"; iMod "H1"; by iApply "H2"). + all: iApply fupd_mask_intro; first exact; iIntros "Hclose"; iMod "H1"; iModIntro; by iMod "Hclose". + Qed. + + Lemma state_update_wp E P e Φ: + (state_update E P) -∗ (P -∗ WP e @ E {{Φ}}) -∗ WP e @ E {{Φ}}. + Proof. + iIntros ">? H". + by iApply "H". + Qed. + + Global Instance is_except_0_state_update E Q : IsExcept0 (state_update E Q). + Proof. + rewrite /IsExcept0. + iIntros. + iApply (state_update_fupd E Q). by rewrite -except_0_fupd -fupd_intro. + Qed. + + Lemma state_update_frame_l R E P : + R ∗ state_update E P ⊢ state_update E (P ∗ R). + Proof. + iIntros "[HR H]". + iMod "H". + iModIntro. + iFrame. + Qed. + + Global Instance frame_state_update p E R P Q: + Frame p R P Q → Frame p R (state_update E P) (state_update E Q). + Proof. + rewrite /Frame=> HR. + rewrite state_update_frame_l. + iIntros ">[??]". + iModIntro. + iApply HR; iFrame. + Qed. + + Global Instance from_pure_bupd_state_update b E P φ : + FromPure b P φ → FromPure b (state_update E P) φ. + Proof. + rewrite /FromPure=> HP. + iIntros "H !>". + by iApply HP. + Qed. + + Global Instance into_wand_state_update p q E R P Q : + IntoWand false false R P Q → IntoWand p q (state_update E R) (state_update E P) (state_update E Q). + Proof. + rewrite /IntoWand /= => HR. + rewrite !bi.intuitionistically_if_elim. + iIntros ">HR >HP !>". iApply (HR with "HR HP"). + Qed. + + Global Instance into_wand_bupd_persistent_state_update p q E R P Q : + IntoWand false q R P Q → IntoWand p q (state_update E R) P (state_update E Q). + Proof. + rewrite /IntoWand /= => HR. rewrite bi.intuitionistically_if_elim. + iIntros ">HR HP !>". + iApply (HR with "HR HP"). + Qed. + + Global Instance into_wand_bupd_args_state_update p q E R P Q : + IntoWand p false R P Q → IntoWand' p q R (state_update E P) (state_update E Q). + Proof. + rewrite /IntoWand' /IntoWand /= => ->. + rewrite bi.intuitionistically_if_elim. + iIntros "Hw HP". + iMod "HP". + iModIntro. + by iApply "Hw". + Qed. + + Global Instance from_sep_bupd_state_update E P Q1 Q2 : + FromSep P Q1 Q2 → FromSep (state_update E P) (state_update E Q1) (state_update E Q2). + Proof. + rewrite /FromSep=> HP. + iIntros "[>HQ1 >HQ2] !>". + iApply HP. iFrame. + Qed. + + Global Instance from_exist_state_update {B} P E (Φ : B → iProp Σ) : + FromExist P Φ → FromExist (state_update E P) (λ b, state_update E (Φ b))%I. + Proof. + rewrite /FromExist => HP. + iIntros "[%x >Hx] !>". + iApply HP. eauto. + Qed. + + Global Instance into_forall_state_update {B} P E (Φ : B → iProp Σ) : + IntoForall P Φ → IntoForall (state_update E P) (λ b, state_update E (Φ b))%I. + Proof. + rewrite /IntoForall=>HP. + iIntros "> H" (b) "!>". + iApply (HP with "H"). + Qed. + + Global Instance from_assumption_state_update p E P Q : + FromAssumption p P Q → KnownRFromAssumption p P (state_update E Q). + Proof. rewrite /KnownRFromAssumption /FromAssumption=>->. iApply state_update_ret. Qed. + +End state_update. From 9ae3eea64c802a875adacfbe3e6b635936030a65 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Mon, 7 Oct 2024 11:04:48 +0200 Subject: [PATCH 67/69] rand counter +impl1 +impl2 fix build --- .../coneris/examples/random_counter/impl1.v | 67 ++------- .../coneris/examples/random_counter/impl2.v | 132 ++++++----------- .../coneris/examples/random_counter/impl3.v | 137 ++++++------------ .../examples/random_counter/random_counter.v | 85 ++++------- theories/coneris/wp_update.v | 22 +++ 5 files changed, 151 insertions(+), 292 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl1.v b/theories/coneris/examples/random_counter/impl1.v index 9bd0af0e..4436b9ec 100644 --- a/theories/coneris/examples/random_counter/impl1.v +++ b/theories/coneris/examples/random_counter/impl1.v @@ -128,58 +128,23 @@ Section impl1. wp_pures. iApply "HΦ". by iFrame. Qed. - - Lemma counter_presample_spec1 NS E T γ1 γ2 c α ns: - ↑NS ⊆ E -> - is_counter1 NS c γ1 γ2 -∗ + + Lemma counter_tapes_presample1 N E γ1 γ2 c α ns ε (ε2 : fin 4%nat -> R): + ↑N ⊆ E -> + (∀ x, 0<=ε2 x)%R -> + (SeriesC (λ n, 1 / 4 * ε2 n)%R <= ε)%R -> + is_counter1 N c γ1 γ2 -∗ rand_tapes_frag (L:=L) γ1 α (3%nat, ns) -∗ - ( |={E∖↑NS,∅}=> - ∃ ε ε2 num, - ↯ ε ∗ - ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ - ⌜(SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R⌝ ∗ - (∀ ns', ↯ (ε2 ns') ={∅,E∖↑NS}=∗ - T ε ε2 num ns' - ))-∗ - wp_update E (∃ ε ε2 num ns', rand_tapes_frag (L:=L) γ1 α (3%nat, ns++ns') ∗ T ε ε2 num ns'). + ↯ ε -∗ + state_update E (∃ n, ↯ (ε2 n) ∗ rand_tapes_frag (L:=L) γ1 α (3%nat, ns++[fin_to_nat n])). Proof. - iIntros (Hsubset) "[#Hinv #Hinv'] Hfrag Hvs". - Admitted. - (* iMod (rand_presample_spec _ _ _ _ _ _ (λ ε num ε2 ns', *) - (* ∃ ε2', ⌜∀ xs ys, fin_to_nat <$> xs = ys -> ε2 xs = ε2' ys⌝ ∗ *) - (* T ε ε2' num (fin_to_nat <$> ns'))%I with "[//][$][Hvs]") as "H"; first by apply nclose_subseteq'. *) - (* - iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose". *) - (* { apply difference_mono_l. *) - (* by apply nclose_subseteq'. } *) - (* iMod "Hvs" as "(%ε & %ε2 & %num & Herr & %Hpos & %Hineq & Hrest)". *) - (* iFrame. *) - (* iModIntro. *) - (* iExists num, (λ ls, ε2 (fin_to_nat <$> ls)). *) - (* repeat iSplit. *) - (* + done. *) - (* + iPureIntro. *) - (* etrans; last exact. *) - (* apply Req_le. *) - (* replace (INR 4) with 4%R; last (simpl; lra). *) - (* f_equal. *) - (* rewrite !SeriesC_list. *) - (* * by rewrite foldr_fmap. *) - (* * apply NoDup_fmap. *) - (* -- apply list_fmap_eq_inj. *) - (* apply fin_to_nat_inj. *) - (* -- apply NoDup_enum_uniform_fin_list. *) - (* * apply NoDup_enum_uniform_fin_list. *) - (* + iIntros (ns') "Herr". *) - (* iMod ("Hrest" with "[$]") as "HT". *) - (* iMod "Hclose" as "_". *) - (* iModIntro. *) - (* iFrame. *) - (* iPureIntro. *) - (* by intros ??<-. *) - (* - iDestruct "H" as "(%ε & %num & %ε2 & %ns' & Hfrag & %ε2' & % & HT)". *) - (* iModIntro. iFrame. *) - (* Qed. *) - + iIntros (Hsubset Hpos Hineq) "[#Hinv #Hinv'] Hfrag Herr". + iMod (rand_tapes_presample with "[//][$][$]") as "(%&$&$)"; try done; first by apply nclose_subseteq'. + etrans; last exact. + apply Req_le. + apply SeriesC_ext; intros. simpl. lra. + Qed. + Lemma read_counter_spec1 N E c γ1 γ2 Q: ↑N ⊆ E -> {{{ is_counter1 N c γ1 γ2 ∗ @@ -220,10 +185,10 @@ Program Definition random_counter1 `{!conerisGS Σ, F:rand_spec}: random_counter counter_tapes_frag _ γ α ns := rand_tapes_frag (L:= counterG1_randG)γ α (3%nat, ns); counter_content_auth _ γ z := own γ (●F z); counter_content_frag _ γ f z := own γ (◯F{f} z); + counter_tapes_presample _ := counter_tapes_presample1; new_counter_spec _ := new_counter_spec1; allocate_tape_spec _ :=allocate_tape_spec1; incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some1; - counter_presample_spec _ :=counter_presample_spec1; read_counter_spec _ :=read_counter_spec1 |}. Next Obligation. diff --git a/theories/coneris/examples/random_counter/impl2.v b/theories/coneris/examples/random_counter/impl2.v index 98b1cee0..b4dfec13 100644 --- a/theories/coneris/examples/random_counter/impl2.v +++ b/theories/coneris/examples/random_counter/impl2.v @@ -271,98 +271,50 @@ Section impl2. by eapply Forall_inv_tail. Qed. - Lemma counter_presample_spec2 NS E T γ1 γ2 c α ns: - ↑NS ⊆ E -> - is_counter2 NS c γ1 γ2 -∗ + Lemma counter_tapes_presample2 N E γ1 γ2 c α ns ε (ε2 : fin 4%nat -> R): + ↑N ⊆ E -> + (∀ x, 0<=ε2 x)%R -> + (SeriesC (λ n, 1 / 4 * ε2 n)%R <= ε)%R -> + is_counter2 N c γ1 γ2 -∗ (flip_tapes_frag (L:=L) γ1 α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝) -∗ - ( |={E∖↑NS,∅}=> - ∃ ε ε2 num, - ↯ ε ∗ - ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ - ⌜(SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R⌝ ∗ - (∀ ns', ↯ (ε2 ns') ={∅,E∖↑NS}=∗ - T ε ε2 num ns' - ))-∗ - wp_update E (∃ ε ε2 num ns', (flip_tapes_frag (L:=L) γ1 α (expander (ns++ns')) ∗ ⌜Forall (λ x, x<4) (ns++ns')⌝) ∗ T ε ε2 num ns'). + ↯ ε -∗ + state_update E (∃ n, ↯ (ε2 n) ∗ + (flip_tapes_frag (L:=L) γ1 α (expander (ns++[fin_to_nat n])) ∗ ⌜Forall (λ x, x<4) (ns++[fin_to_nat n])⌝)). Proof. - iIntros (Hsubset) "(#Hinv & #Hinv') [Hfrag %] Hvs". - iMod (flip_presample_spec _ _ _ _ _ - (λ ε num' ε2' ns', ∃ num ε2 ns, - ⌜(2 * num = num')%nat⌝∗ ⌜expander ns = ns'⌝ ∗ ⌜Forall (λ x, x<4) ns⌝ ∗ - (* ⌜ ∀ ls, Forall (λ x, x<4) ls -> ε2 ls = ε2' (expander ls)⌝ ∗ *) - T ε ε2 num ns - )%I with "[//][$][Hvs]") as "H"; first by apply nclose_subseteq'. - - iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose". - { apply difference_mono; [done|by apply nclose_subseteq']. } - iMod "Hvs" as "(%ε & %ε2 & %num & Herr & %Hpos & %Hineq & Hvs)". - iExists ε, (2*num). - iFrame. - iModIntro. - iExists (λ ls, match (decoder ls) with | Some ls' => ε2 ls' | _ => 1%R end). - repeat iSplit. - + iPureIntro. intros. case_match; [done|lra]. - + iPureIntro. - etrans; last exact. - rewrite !Rdiv_def -!SeriesC_scal_r. - etrans; last eapply (SeriesC_le_inj _ (λ bs, decoder bs)). - * apply Req_le. - apply SeriesC_ext. - intros bs. - destruct (decoder bs) eqn:K. - -- simpl. f_equal. - ++ case_match eqn:H0. - ** rewrite bool_decide_eq_true_2; first done. - erewrite elem_of_list_fmap. - rewrite Nat.eqb_eq in H0. - apply decoder_ineq in K as K'. - apply fin.nat_list_to_fin_list in K'. - destruct K' as [xs K']. - exists xs. split; first done. - rewrite elem_of_enum_uniform_fin_list. - apply decoder_Some_length in K. - apply (f_equal length) in K'. - rewrite fmap_length in K'. lia. - ** rewrite bool_decide_eq_false_2; first done. - rewrite elem_of_list_fmap. - intros (?&->&K'). - rewrite elem_of_enum_uniform_fin_list in K'. - rewrite Nat.eqb_neq in H0. - exfalso. apply H0. - apply decoder_Some_length in K. - rewrite fmap_length in K. - lia. - ++ f_equal. - replace 4%R with (2^2)%R; last (simpl; lra). - by rewrite -pow_mult. - -- simpl. - eapply decoder_None in K; last done. - case_match eqn :H0; last lra. - rewrite Nat.eqb_eq in H0. - exfalso. - apply K. - exists num. lia. - * intros. rewrite -Rdiv_1_l. - replace 4%R with (INR 4); last (simpl; lra). - rewrite -pow_INR. apply Rmult_le_pos; last apply Rdiv_INR_ge_0. - case_bool_decide; [done|lra]. - * intros. by eapply decoder_inj. - * apply ex_seriesC_scal_r, ex_seriesC_list. - + iIntros (ns') "Herr". - case_match eqn :Hdecoder; last (by iDestruct (ec_contradict with "[$]") as "%"). - iMod ("Hvs" with "[$]") as "HT". - iMod "Hclose" as "_". - iFrame. - iPureIntro. repeat split. - * by apply decoder_correct. - * by eapply decoder_ineq. - - iDestruct "H" as "(%&%&%&%&?&%&%&%&<-&<-&%&?)". - iModIntro. - iFrame. - iSplit; last by rewrite Forall_app. - rewrite /expander. - by rewrite bind_app. + iIntros (Hsubset Hpos Hineq) "[#Hinv #Hinv'] [Hfrag %Hforall] Herr". + iMod (flip_tapes_presample _ _ _ _ _ _ (λ b, 1/2 *if b then (ε2 2%fin + ε2 3%fin) else (ε2 0%fin + ε2 1%fin))%R with "[//][$][$]") as "(%b & Herr & Hfrag)"; first by apply nclose_subseteq'. + { intros []; apply Rmult_le_pos; try lra. + all: by apply Rplus_le_le_0_compat. } + { etrans; last exact. + rewrite SeriesC_finite_foldr/=. lra. } + destruct b. + - iMod (flip_tapes_presample _ _ _ _ _ _ (λ b, if b then ε2 3%fin else ε2 2%fin)%R with "[//][$][$]") as "(%b & Herr & Hfrag)"; first by apply nclose_subseteq'. + { intros []; by try lra. } + { simpl. lra. } + destruct b. + + iFrame. rewrite /expander bind_app -!app_assoc/=. + rewrite Nat.OddT_odd; last (constructor 1 with (x:=1); lia). iFrame. + iPureIntro. + rewrite Forall_app; naive_solver. + + iFrame. rewrite /expander bind_app -!app_assoc/=. + rewrite Nat.odd_2. iFrame. + iPureIntro. + rewrite Forall_app; naive_solver. + - iMod (flip_tapes_presample _ _ _ _ _ _ (λ b, if b then ε2 1%fin else ε2 0%fin)%R with "[//][$][$]") as "(%b & Herr & Hfrag)"; first by apply nclose_subseteq'. + { intros []; by try lra. } + { simpl. lra. } + destruct b. + + iFrame. rewrite /expander bind_app -!app_assoc/=. + rewrite Nat.OddT_odd; last (constructor 1 with (x:=0); lia). iFrame. + iPureIntro. + rewrite Forall_app; naive_solver. + + iFrame. rewrite /expander bind_app -!app_assoc/=. + rewrite Nat.odd_0. iFrame. + iPureIntro. + rewrite Forall_app; split; first done. + rewrite Forall_singleton. lia. Qed. - + Lemma read_counter_spec2 N E c γ1 γ2 Q: ↑N ⊆ E -> {{{ is_counter2 N c γ1 γ2 ∗ @@ -407,10 +359,10 @@ Program Definition random_counter2 `{flip_spec Σ}: random_counter := counter_tapes_frag K γ α ns := (flip_tapes_frag (L:=counterG2_to_flipG) γ α (expander ns) ∗ ⌜Forall (λ x, x<4) ns⌝)%I; counter_content_auth _ γ z := own γ (●F z); counter_content_frag _ γ f z := own γ (◯F{f} z); + counter_tapes_presample _ := counter_tapes_presample2; new_counter_spec _ := new_counter_spec2; allocate_tape_spec _ :=allocate_tape_spec2; incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some2; - counter_presample_spec _ :=counter_presample_spec2; read_counter_spec _ :=read_counter_spec2 |}. Next Obligation. diff --git a/theories/coneris/examples/random_counter/impl3.v b/theories/coneris/examples/random_counter/impl3.v index 63898f43..c0292552 100644 --- a/theories/coneris/examples/random_counter/impl3.v +++ b/theories/coneris/examples/random_counter/impl3.v @@ -121,99 +121,54 @@ Section impl3. rewrite /filter_f. lia. Qed. - (** * test *) - Lemma counter_presample_spec3_test NS E T γ1 γ2 c α ns: - ↑NS ⊆ E -> - is_counter3 NS c γ1 γ2 -∗ - (∃ ls, ⌜filter filter_f ls = ns⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) -∗ - ( |={E∖↑NS,∅}=> - ∃ ε ε2, - ↯ ε ∗ - ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ - ⌜(SeriesC (λ x, if bool_decide (x<4)%nat then ε2 x else 0)%R / (4) <= ε)%R⌝ ∗ - (∀ x, ↯ (ε2 x) ={∅,E∖↑NS}=∗ - T ε ε2 x - ))-∗ - wp_update E (∃ ε ε2 x, (∃ ls, ⌜filter filter_f ls = ns++[x]⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) ∗ T ε ε2 x). - Proof. - iIntros (Hsubset) "[#Hinv #Hinv'] Hfrag Hvs". - iMod wp_update_epsilon_err as "(%eps&%Heps&Heps)". - (* iRevert "Hfrag Hvs". *) - (* iApply (ec_ind_amp _ 5%R with "[][$]"); [done|lra|]. *) - (* clear eps Heps. *) - (* iModIntro. *) - (* iIntros (eps Heps) "#IH Heps (%ls & <- & Hfrag) Hrest". *) - - iDestruct "Hfrag" as "(%ls & <- & Hfrag)". - - - iMod (rand_presample_single_spec _ _ _ _ _ _ - (λ ε' ε2' x', - ∃ (ε2:nat -> R), - ⌜ε2' = (λ x, if bool_decide (fin_to_nat x <4)%nat then ε2 x - else if bool_decide (fin_to_nat x = 4)%nat then ε'+5*eps - else 1)%R⌝ ∗ - if bool_decide (fin_to_nat x' = 4)%nat - then _ - else T (ε'-eps)%R ε2 (fin_to_nat x') - )%I with "[//][$][-]") as "Hvs". - - by apply nclose_subseteq'. - - iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose". - + apply difference_mono_l. - by apply nclose_subseteq'. - + iMod "Hvs" as "(%ε & %ε2 & Hε & %Hpos & %Hineq & Hvs)". - iModIntro. - iDestruct (ec_combine with "[$]") as "Herr". - iFrame. - iExists (λ x, if bool_decide (fin_to_nat x <4)%nat then ε2 x - else if bool_decide (fin_to_nat x = 4)%nat then ε+5*eps - else 1)%R. - repeat iSplit. - * admit. - * admit. - * iIntros (x) "Herr". - admit. - - admit. - Admitted. - - - Lemma counter_presample_spec3 NS E T γ1 γ2 c α ns: - ↑NS ⊆ E -> - is_counter3 NS c γ1 γ2 -∗ + Lemma counter_tapes_presample3 N E γ1 γ2 c α ns ε (ε2 : fin 4%nat -> R): + ↑N ⊆ E -> + (∀ x, 0<=ε2 x)%R -> + (SeriesC (λ n, 1 / 4 * ε2 n)%R <= ε)%R -> + is_counter3 N c γ1 γ2 -∗ (∃ ls, ⌜filter filter_f ls = ns⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) -∗ - ( |={E∖↑NS,∅}=> - ∃ ε ε2 num, - ↯ ε ∗ - ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ - ⌜(SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R⌝ ∗ - (∀ ns', ↯ (ε2 ns') ={∅,E∖↑NS}=∗ - T ε ε2 num ns' - ))-∗ - wp_update E (∃ ε ε2 num ns', (∃ ls, ⌜filter filter_f ls = ns++ns'⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) ∗ T ε ε2 num ns'). + ↯ ε -∗ + state_update E (∃ n, ↯ (ε2 n) ∗ + (∃ ls, ⌜filter filter_f ls = (ns++[fin_to_nat n])⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls))). Proof. - iIntros (Hsubset) "[#Hinv #Hinv'] Hfrag Hvs". - iMod wp_update_epsilon_err as "(%eps&%Heps&Heps)". - iApply fupd_wp_update. - iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose"; first set_solver. - iMod "Hvs" as "(%ε & %ε2 & %num & Herr & %Hpos & %Hineq & Hrest)". - iInduction num as [|num] "IH" forall (eps Heps ε ε2 Hpos Hineq T) "Heps Hfrag Herr Hrest Hclose". - { iDestruct (ec_weaken _ (ε2 []) with "[$Herr]") as "Hε". - - split; first done. - etrans; last exact. - rewrite -SeriesC_singleton_dependent. - rewrite Rdiv_1_r. - apply Req_le. - apply SeriesC_ext. - intros. - simpl. - repeat case_bool_decide; set_solver. - - iMod ("Hrest" with "[$]"). - iMod "Hclose" as "_". - iFrame. - rewrite app_nil_r. by iFrame. - } + Admitted. + (* Lemma counter_presample_spec3 NS E T γ1 γ2 c α ns: *) + (* ↑NS ⊆ E -> *) + (* is_counter3 NS c γ1 γ2 -∗ *) + (* (∃ ls, ⌜filter filter_f ls = ns⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) -∗ *) + (* ( |={E∖↑NS,∅}=> *) + (* ∃ ε ε2 num, *) + (* ↯ ε ∗ *) + (* ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ *) + (* ⌜(SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R⌝ ∗ *) + (* (∀ ns', ↯ (ε2 ns') ={∅,E∖↑NS}=∗ *) + (* T ε ε2 num ns' *) + (* ))-∗ *) + (* wp_update E (∃ ε ε2 num ns', (∃ ls, ⌜filter filter_f ls = ns++ns'⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) ∗ T ε ε2 num ns'). *) + (* Proof. *) + (* iIntros (Hsubset) "[#Hinv #Hinv'] Hfrag Hvs". *) + (* iMod wp_update_epsilon_err as "(%eps&%Heps&Heps)". *) + (* iApply fupd_wp_update. *) + (* iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose"; first set_solver. *) + (* iMod "Hvs" as "(%ε & %ε2 & %num & Herr & %Hpos & %Hineq & Hrest)". *) + (* iInduction num as [|num] "IH" forall (eps Heps ε ε2 Hpos Hineq T) "Heps Hfrag Herr Hrest Hclose". *) + (* { iDestruct (ec_weaken _ (ε2 []) with "[$Herr]") as "Hε". *) + (* - split; first done. *) + (* etrans; last exact. *) + (* rewrite -SeriesC_singleton_dependent. *) + (* rewrite Rdiv_1_r. *) + (* apply Req_le. *) + (* apply SeriesC_ext. *) + (* intros. *) + (* simpl. *) + (* repeat case_bool_decide; set_solver. *) + (* - iMod ("Hrest" with "[$]"). *) + (* iMod "Hclose" as "_". *) + (* iFrame. *) + (* rewrite app_nil_r. by iFrame. *) + (* } *) - Admitted. + (* Admitted. *) (* iMod "Hvs". *) (* iRevert "Hfrag Hvs". *) (* iApply (ec_ind_amp _ 5%R with "[][$]"); [done|lra|]. *) @@ -377,10 +332,10 @@ Program Definition random_counter3 `{F:rand_spec}: random_counter := (∃ ls, ⌜filter filter_f ls = ns⌝ ∗ rand_tapes_frag (L:=counterG3_randG) γ α (4, ls))%I; counter_content_auth _ γ z := own γ (●F z); counter_content_frag _ γ f z := own γ (◯F{f} z); + counter_tapes_presample _ := counter_tapes_presample3; new_counter_spec _ := new_counter_spec3; allocate_tape_spec _ :=allocate_tape_spec3; incr_counter_tape_spec_some _ :=incr_counter_tape_spec_some3; - counter_presample_spec _ :=counter_presample_spec3; read_counter_spec _ :=read_counter_spec3 |}. diff --git a/theories/coneris/examples/random_counter/random_counter.v b/theories/coneris/examples/random_counter/random_counter.v index 1bd7f647..0365a590 100644 --- a/theories/coneris/examples/random_counter/random_counter.v +++ b/theories/coneris/examples/random_counter/random_counter.v @@ -106,16 +106,16 @@ Class random_counter `{!conerisGS Σ} := RandCounter (* T ε ε2 num ns' *) (* ))-∗ *) (* wp_update E (∃ ε ε2 num ns', counter_tapes_frag (L:=L) γ1 α (ns++ns') ∗ T ε ε2 num ns'); *) - (* read_counter_spec {L: counterG Σ} N E c γ1 γ2 Q: *) - (* ↑N ⊆ E -> *) - (* {{{ is_counter (L:=L) N c γ1 γ2 ∗ *) - (* (∀ (z:nat), counter_content_auth (L:=L) γ2 z ={E∖↑N}=∗ *) - (* counter_content_auth (L:=L) γ2 z∗ Q z) *) + read_counter_spec {L: counterG Σ} N E c γ1 γ2 Q: + ↑N ⊆ E -> + {{{ is_counter (L:=L) N c γ1 γ2 ∗ + (∀ (z:nat), counter_content_auth (L:=L) γ2 z ={E∖↑N}=∗ + counter_content_auth (L:=L) γ2 z∗ Q z) - (* }}} *) - (* read_counter c @ E *) - (* {{{ (n':nat), RET #n'; Q n' *) - (* }}} *) + }}} + read_counter c @ E + {{{ (n':nat), RET #n'; Q n' + }}} }. @@ -153,60 +153,25 @@ Section lemmas. iApply "HΦ". iFrame. } - iMod (fupd_mask_frame _ (↑N) (E∖↑N) ∅ with "[Hvs]"); first set_solver. + iMod (fupd_mask_frame _ (↑N) (E∖↑N) ∅ with "[Hvs]") as "H"; first set_solver. - iMod "Hvs" as "H". iModIntro. + (* mask change *) rewrite left_id_L. - replace (_∖(_∖_)) with ((nclose N)); last first. - + set_unfold. clear -Hsubset. firstorder. set_solver. - + iModIntro. iExact "H". - - - (* iMod (fupd_mask_weaken (E1:=E) (E3:=E) (E∖↑N) with "[]"). *) - (* (* iModIntro. *) *) - (* (* iApply state_update_fupd. *) *) - (* iDestruct (fupd_mask_frame_r _ _ (↑N) with "[$Hvs]") as "Hvs"; first set_solver. *) - (* rewrite left_id_L. *) - (* replace (E∖_∪_) with E; last first. *) - (* { set_unfold. admit. } *) - (* iMod (counter_tapes_presample N E with "[$][$][Hvs]"). *) - Admitted. - - - (* iMod "Hvs". *) - (* iMod (counter_presample_spec _ _ *) - (* (λ ε ε2 num ns', *) - (* ∃ n ε2', *) - (* ⌜num=1%nat⌝ ∗ ⌜ns'=[n]⌝ ∗ *) - (* ⌜(ε2' = λ x, ε2 [x])%R⌝ ∗ *) - (* ( ∀ (z:nat), counter_content_auth (L:=L) γ2 z *) - (* ={E∖↑N}=∗ *) - (* counter_content_auth (L:=L) γ2 (z+n) ∗ Q z n ε ε2') *) - (* ) %I with "[//][$][-HΦ]") as "(%ε & %ε2 & %num & %ns' & Hfrag & %n & %ε2' & -> & -> & -> & Hrest)"; first done. *) - (* - iMod "Hvs" as "(%ε & %ε2 & Herr & %Hpos & %Hineq & Hrest)". *) - (* iFrame. iModIntro. *) - (* iExists (λ l, match l with |[x] => ε2 x | _ => 1 end)%R, 1%nat. *) - (* repeat iSplit. *) - (* + iPureIntro. intros; repeat case_match; try lra. naive_solver. *) - (* + iPureIntro. etrans; last exact. *) - (* rewrite SeriesC_list. *) - (* * Local Transparent enum_uniform_fin_list. *) - (* simpl. lra. *) - (* * apply NoDup_fmap_2; last apply NoDup_enum_uniform_fin_list. *) - (* apply list_fmap_eq_inj, fin_to_nat_inj. *) - (* + iIntros (ns') "Herr". destruct (ns') as [|n[|??]]; [by iDestruct (ec_contradict with "[$]") as "%"| |by iDestruct (ec_contradict with "[$]") as "%"]. *) - (* iMod ("Hrest" $! n with "[$]") as "?". *) - (* iModIntro. *) - (* by iFrame. *) - (* - wp_apply (incr_counter_tape_spec_some _ _ _ _ _ (λ z, ∃ ε ε2, Q z n ε ε2)%I with "[$Hfrag Hrest]"); first exact. *) - (* + iSplit; first done. *) - (* iIntros (?) "?". *) - (* iMod ("Hrest" with "[$]") as "[$ ?]". *) - (* iModIntro. *) - (* iFrame. *) - (* + iIntros (?) "[?(%&%&?)]". *) - (* iApply "HΦ". *) - (* iFrame. *) - (* Qed. *) + replace (_∖(_∖_)) with ((nclose N)); first (iModIntro; iExact "H"). + set_unfold. clear -Hsubset. firstorder. set_solver. + - iDestruct "H" as "(%ε & %ε2 & Herr & %Hpos & %Hineq & Hrest)". + iMod (counter_tapes_presample _ _ _ _ _ _ _ _ (λ x, ε2 (fin_to_nat x)) with "[//][$][$Herr]") as "(%n & Herr & ?)"; try done. + { rewrite SeriesC_finite_foldr/=. lra. } + iModIntro. + iMod (fupd_mask_frame _ (E) ∅ (E∖↑N) with "[Hrest Herr]") as "H"; first set_solver. + + iMod ("Hrest" with "[$]") as "H". iModIntro. + replace (_∖_∪_) with E; first (iModIntro; iExact "H"). + clear -Hsubset. + set_unfold. intros x. firstorder. + destruct (decide (x∈nclose N)); set_solver. + + iModIntro. iFrame. + Qed. Lemma incr_counter_tape_spec_none' N E c γ1 γ2 ε (ε2:nat -> R)(α:loc) (ns:list nat) (q:Qp) (z:nat): ↑N ⊆ E-> diff --git a/theories/coneris/wp_update.v b/theories/coneris/wp_update.v index 9d868f61..2b7197e7 100644 --- a/theories/coneris/wp_update.v +++ b/theories/coneris/wp_update.v @@ -327,6 +327,28 @@ Section state_update. by iIntros. Qed. + + Lemma state_update_epsilon_err E: + ⊢ state_update E (∃ ε, ⌜(0<ε)%R⌝ ∗ ↯ ε). + Proof. + rewrite state_update_unseal/state_update_def. + iIntros (? ε) "[Hstate Herr]". + iApply fupd_mask_intro; first set_solver. + iIntros "Hclose". + iApply state_step_coupl_ampl'. + iIntros (ε' ?). + iApply state_step_coupl_ret. + assert (ε<=ε')%R as H' by lra. + pose (diff :=((ε' - ε) H')%NNR). + replace (ε') with (ε + diff)%NNR; last (apply nnreal_ext; rewrite /diff; simpl; lra). + iMod (ec_supply_increase _ diff with "[$]") as "[??]". + { rewrite /diff. simpl. lra. } + iFrame. iMod "Hclose". iPureIntro. + rewrite /diff. + simpl. + lra. + Qed. + Lemma state_update_fupd_change E1 E2 P Q: (|={E1, E2}=> P) -∗ (P-∗ state_update E2 (|={E2, E1}=> Q)) -∗ state_update E1 Q. Proof. From c12d733aecbff028e435bd5d152ccbd621680718 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Mon, 7 Oct 2024 11:29:03 +0200 Subject: [PATCH 68/69] Finish impl3 --- .../coneris/examples/random_counter/impl3.v | 197 ++++-------------- 1 file changed, 40 insertions(+), 157 deletions(-) diff --git a/theories/coneris/examples/random_counter/impl3.v b/theories/coneris/examples/random_counter/impl3.v index c0292552..45799aa6 100644 --- a/theories/coneris/examples/random_counter/impl3.v +++ b/theories/coneris/examples/random_counter/impl3.v @@ -131,163 +131,46 @@ Section impl3. state_update E (∃ n, ↯ (ε2 n) ∗ (∃ ls, ⌜filter filter_f ls = (ns++[fin_to_nat n])⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls))). Proof. - Admitted. - (* Lemma counter_presample_spec3 NS E T γ1 γ2 c α ns: *) - (* ↑NS ⊆ E -> *) - (* is_counter3 NS c γ1 γ2 -∗ *) - (* (∃ ls, ⌜filter filter_f ls = ns⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) -∗ *) - (* ( |={E∖↑NS,∅}=> *) - (* ∃ ε ε2 num, *) - (* ↯ ε ∗ *) - (* ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ *) - (* ⌜(SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R⌝ ∗ *) - (* (∀ ns', ↯ (ε2 ns') ={∅,E∖↑NS}=∗ *) - (* T ε ε2 num ns' *) - (* ))-∗ *) - (* wp_update E (∃ ε ε2 num ns', (∃ ls, ⌜filter filter_f ls = ns++ns'⌝ ∗ rand_tapes_frag (L:=L) γ1 α (4, ls)) ∗ T ε ε2 num ns'). *) - (* Proof. *) - (* iIntros (Hsubset) "[#Hinv #Hinv'] Hfrag Hvs". *) - (* iMod wp_update_epsilon_err as "(%eps&%Heps&Heps)". *) - (* iApply fupd_wp_update. *) - (* iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose"; first set_solver. *) - (* iMod "Hvs" as "(%ε & %ε2 & %num & Herr & %Hpos & %Hineq & Hrest)". *) - (* iInduction num as [|num] "IH" forall (eps Heps ε ε2 Hpos Hineq T) "Heps Hfrag Herr Hrest Hclose". *) - (* { iDestruct (ec_weaken _ (ε2 []) with "[$Herr]") as "Hε". *) - (* - split; first done. *) - (* etrans; last exact. *) - (* rewrite -SeriesC_singleton_dependent. *) - (* rewrite Rdiv_1_r. *) - (* apply Req_le. *) - (* apply SeriesC_ext. *) - (* intros. *) - (* simpl. *) - (* repeat case_bool_decide; set_solver. *) - (* - iMod ("Hrest" with "[$]"). *) - (* iMod "Hclose" as "_". *) - (* iFrame. *) - (* rewrite app_nil_r. by iFrame. *) - (* } *) - - (* Admitted. *) - (* iMod "Hvs". *) - (* iRevert "Hfrag Hvs". *) - (* iApply (ec_ind_amp _ 5%R with "[][$]"); [done|lra|]. *) - (* clear eps Heps. *) - (* iModIntro. *) - (* iIntros (eps Heps) "#IH Heps (%ls & %Hfilter & Hfrag) Hvs". *) - (* iMod (rand_presample_spec _ _ _ _ _ _ (λ ε num ε2 ns', *) - (* _ *) - (* )%I with "[//][$][Heps Hvs]") as "H"; first by apply nclose_subseteq'. *) - (* - iMod (fupd_mask_subseteq (E ∖ ↑NS)) as "Hclose". *) - (* { apply difference_mono_l. *) - (* by apply nclose_subseteq'. } *) - (* iMod "Hvs" as "(%ε & %ε2 & %num & Herr &%Hpos & %Hineq & Hrest)". *) - (* iFrame. *) - (* admit. *) - (* - iDestruct "H" as "(%ε & %num & %ε2 & %ns' & Hfrag & HT)". *) - (* iModIntro. *) - (* iFrame. *) - (* rewrite filter_app. *) - (* admit. *) - (* Admitted. *) - (* iIntros (Hsubset Hpos Hineq) "#Hinv #Hvs HP Hfrag". *) - (* iMod wp_update_epsilon_err as "(%epsilon_err&%H&?)". *) - (* iApply wp_update_state_step_coupl. *) - (* iRevert "HP Hfrag". *) - (* iApply (ec_ind_amp _ 5%R with "[][$]"); [done|lra|]. *) - (* iModIntro. *) - (* clear epsilon_err H. *) - (* iIntros (epsilon_err ?) "#IH Hepsilon_err HP Hfrag". *) - (* iIntros (σ ε) "((Hheap&Htapes)&Hε)". *) - (* iMod (inv_acc with "Hinv") as "[>(%ε0 & % & % & % & H1 & H2 & H3 & H4 & -> & H5 & H6) Hclose]"; [done|]. *) - (* iDestruct (hocap_tapes_agree with "[$][$]") as "%". *) - (* erewrite <-(insert_delete m) at 1; last done. *) - (* rewrite big_sepM_insert; last apply lookup_delete. *) - (* simpl. *) - (* iDestruct "H3" as "[(%ls & %Hls & Htape) H3]". *) - (* iDestruct (tapeN_lookup with "[$][$]") as "(%&%&%)". *) - (* iDestruct (ec_valid with "H1") as "%". *) - (* iDestruct (ec_combine with "[$]") as "Htotal". *) - (* iDestruct (ec_supply_bound with "[$][$]") as "%". *) - (* iMod (ec_supply_decrease with "[$][$]") as (ε1' ε_rem -> Hε1') "Hε_supply". subst. *) - (* iApply fupd_mask_intro; [set_solver|]; iIntros "Hclose'". *) - (* pose (ε2' := λ (x:nat), (if bool_decide (x<=3)%nat then ε2 ε0 x else *) - (* if bool_decide (x=4)%nat then ε0 + 5 * epsilon_err *) - (* else 0)%R *) - (* ). *) - (* iApply state_step_coupl_state_adv_comp_con_prob_lang; first done. *) - (* assert (forall x, 0<=ε2' x)%R. *) - (* { intros. rewrite /ε2'. repeat case_bool_decide; try done; first naive_solver. *) - (* apply Rplus_le_le_0_compat; simpl; [naive_solver|lra]. } *) - (* unshelve iExists (λ x, mknonnegreal (ε2' x) _); auto. *) - (* iSplit. *) - (* { iPureIntro. *) - (* unshelve epose proof (Hineq ε0 _) as H'; first (by naive_solver). *) - (* rewrite SeriesC_nat_bounded_fin in H'. *) - (* replace (INR 5%nat) with 5%R by (simpl; lra). *) - (* replace (INR 4%nat) with 4%R in H' by (simpl; lra). *) - (* rewrite /=/ε2' Hε1'. *) - (* rewrite SeriesC_finite_foldr/=. *) - (* rewrite SeriesC_finite_foldr/= in H'. *) - (* lra. *) - (* } *) - (* iIntros (sample). *) - (* simpl. *) - (* destruct (Rlt_decision (nonneg ε_rem + (ε2' sample))%R 1%R) as [Hdec|Hdec]; last first. *) - (* { apply Rnot_lt_ge, Rge_le in Hdec. *) - (* iApply state_step_coupl_ret_err_ge_1. *) - (* simpl. simpl in *. lra. *) - (* } *) - (* destruct (Nat.lt_total (fin_to_nat sample) 4%nat) as [|[|]]. *) - (* - (* we sample a value less than 4*) *) - (* iApply state_step_coupl_ret. *) - (* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2' sample) _) with "Hε_supply") as "[Hε_supply Hε]"; auto. *) - (* iMod (tapeN_update_append _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". *) - (* iMod (hocap_tapes_presample _ _ _ _ _ (fin_to_nat sample) with "[$][$]") as "[H4 Hfrag]". *) - (* iMod "Hclose'" as "_". *) - (* iMod ("Hvs" $! _ (fin_to_nat sample)%nat with "[$]") as "[%|[H2 HT]]". *) - (* { iExFalso. iApply (ec_contradict with "[$]"). *) - (* simpl. *) - (* rewrite /ε2'. *) - (* by rewrite bool_decide_eq_true_2; last lia. *) - (* } *) - (* simpl. *) - (* iMod ("Hclose" with "[$Hε H2 Htape H3 $H4 $H5 $H6]") as "_". *) - (* { iNext. *) - (* rewrite /ε2'. rewrite bool_decide_eq_true_2; last lia. *) - (* iFrame. *) - (* rewrite big_sepM_insert_delete; iFrame. simpl. iPureIntro. *) - (* split; last done. *) - (* rewrite filter_app. f_equal. *) - (* rewrite filter_cons. *) - (* rewrite decide_True; first done. *) - (* rewrite /filter_f. by apply bool_decide_pack. *) - (* } *) - (* iApply fupd_mask_intro_subseteq; first set_solver. *) - (* iFrame. *) - (* - (* we sampled a 4, and hence löb *) *) - (* unshelve iMod (ec_supply_increase _ (mknonnegreal (ε2' sample) _) with "Hε_supply") as "[Hε_supply Hε]"; auto. *) - (* iMod (tapeN_update_append _ _ _ _ sample with "[$][$]") as "[Htapes Htape]". *) - (* iMod "Hclose'" as "_". *) - (* simpl. *) - (* rewrite {2}/ε2'. *) - (* rewrite bool_decide_eq_false_2; last lia. *) - (* rewrite bool_decide_eq_true_2; last lia. *) - (* iDestruct (ec_split with "[$]") as "[Hε Hampl]"; simpl; [naive_solver|lra|]. *) - (* simpl. *) - (* iMod ("Hclose" with "[$Hε $H2 Htape H3 $H4 $H5 $H6]") as "_". *) - (* { iNext. *) - (* rewrite (big_sepM_delete _ m); [iFrame|done]. *) - (* simpl. iPureIntro; split; last done. *) - (* rewrite filter_app filter_cons decide_False; simpl. *) - (* - by rewrite filter_nil app_nil_r. *) - (* - rewrite /filter_f. case_bool_decide; [lia|auto]. *) - (* } *) - (* iDestruct ("IH" with "[$][$][$]") as "IH'". *) - (* by iMod ("IH'" $! (state_upd_tapes <[_:=_]> _) with "[$]") as "IH'". *) - (* - pose proof fin_to_nat_lt sample; lia. *) - (* Qed. *) - + iIntros (Hsubset Hpos Hineq) "[#Hinv #Hinv'] Hfrag Herr". + iMod (state_update_epsilon_err) as "(%ep & %Heps & Heps)". + iRevert "Hfrag Herr". + iApply (ec_ind_amp _ (5/4)%R with "[][$]"); try lra. + clear ep Heps. + iModIntro. + iIntros (eps Heps) "#IH Heps (%ls & %Hfilter & Hfrag) Herr". + iDestruct (ec_valid with "[$]") as "%". + iCombine "Heps Herr" as "Herr'". + iMod (rand_tapes_presample _ _ _ _ _ _ _ + (λ x, match decide (fin_to_nat x < 4)%nat with + | left p => ε2 (nat_to_fin p) + | _ => ε + 5/4*eps + end + )%R with "[//][$][$]") as "(%n & Herr & Hfrag)"; first by apply nclose_subseteq'. + { intros. case_match; first done. + apply Rplus_le_le_0_compat; first naive_solver. + apply Rmult_le_pos; lra. + } + { rewrite SeriesC_finite_foldr/=. + rewrite SeriesC_finite_foldr/= in Hineq. + lra. + } + case_match eqn :K. + - (* accept *) + iFrame. + iPureIntro. + rewrite filter_app; f_equal; first done. + rewrite filter_cons /filter_f filter_nil K. + f_equal. by rewrite fin_to_nat_to_fin. + - iDestruct (ec_split with "[$]") as "[Hε Heps]"; first lra. + { apply Rmult_le_pos; lra. } + iApply ("IH" with "[$][Hfrag][$]"). + iFrame. + iPureIntro. + rewrite filter_app. + rewrite filter_cons_False; first by rewrite filter_nil app_nil_r. + rewrite /filter_f. lia. + Qed. + Lemma read_counter_spec3 N E c γ1 γ2 Q: ↑N ⊆ E -> {{{ is_counter3 N c γ1 γ2 ∗ From ea33c4c1c816633c987bbd72eb146de8e6ea21b5 Mon Sep 17 00:00:00 2001 From: Hei Li Date: Mon, 7 Oct 2024 12:24:35 +0200 Subject: [PATCH 69/69] Improve copset rewrites --- .../examples/random_counter/random_counter.v | 19 ++----------------- 1 file changed, 2 insertions(+), 17 deletions(-) diff --git a/theories/coneris/examples/random_counter/random_counter.v b/theories/coneris/examples/random_counter/random_counter.v index 0365a590..263b66a7 100644 --- a/theories/coneris/examples/random_counter/random_counter.v +++ b/theories/coneris/examples/random_counter/random_counter.v @@ -93,19 +93,6 @@ Class random_counter `{!conerisGS Σ} := RandCounter incr_counter_tape c #lbl:α @ E {{{ (z:nat), RET (#z, #n); counter_tapes_frag (L:=L) γ1 α ns ∗ Q z }}}; - (* counter_presample_spec {L: counterG Σ} NS E T γ1 γ2 c α ns: *) - (* ↑NS ⊆ E -> *) - (* is_counter (L:=L) NS c γ1 γ2 -∗ *) - (* counter_tapes_frag (L:=L) γ1 α ns -∗ *) - (* ( |={E∖↑NS,∅}=> *) - (* ∃ ε ε2 num, *) - (* ↯ ε ∗ *) - (* ⌜(∀ n, 0<=ε2 n)%R⌝ ∗ *) - (* ⌜(SeriesC (λ l, if bool_decide (l∈fmap (λ x, fmap (FMap:=list_fmap) fin_to_nat x) (enum_uniform_fin_list 3%nat num)) then ε2 l else 0%R) / (4^num) <= ε)%R⌝ ∗ *) - (* (∀ ns', ↯ (ε2 ns') ={∅,E∖↑NS}=∗ *) - (* T ε ε2 num ns' *) - (* ))-∗ *) - (* wp_update E (∃ ε ε2 num ns', counter_tapes_frag (L:=L) γ1 α (ns++ns') ∗ T ε ε2 num ns'); *) read_counter_spec {L: counterG Σ} N E c γ1 γ2 Q: ↑N ⊆ E -> {{{ is_counter (L:=L) N c γ1 γ2 ∗ @@ -159,7 +146,7 @@ Section lemmas. (* mask change *) rewrite left_id_L. replace (_∖(_∖_)) with ((nclose N)); first (iModIntro; iExact "H"). - set_unfold. clear -Hsubset. firstorder. set_solver. + rewrite difference_difference_r_L. set_solver. - iDestruct "H" as "(%ε & %ε2 & Herr & %Hpos & %Hineq & Hrest)". iMod (counter_tapes_presample _ _ _ _ _ _ _ _ (λ x, ε2 (fin_to_nat x)) with "[//][$][$Herr]") as "(%n & Herr & ?)"; try done. { rewrite SeriesC_finite_foldr/=. lra. } @@ -167,9 +154,7 @@ Section lemmas. iMod (fupd_mask_frame _ (E) ∅ (E∖↑N) with "[Hrest Herr]") as "H"; first set_solver. + iMod ("Hrest" with "[$]") as "H". iModIntro. replace (_∖_∪_) with E; first (iModIntro; iExact "H"). - clear -Hsubset. - set_unfold. intros x. firstorder. - destruct (decide (x∈nclose N)); set_solver. + by rewrite difference_empty_L union_comm_L -union_difference_L. + iModIntro. iFrame. Qed.