From 8a8c14056242d9990901fdc372fec0e9783859f8 Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Wed, 20 Feb 2019 18:06:47 +0100 Subject: [PATCH] Bump stdpp. --- opam | 2 +- tests/proofmode.v | 2 +- theories/algebra/big_op.v | 8 ++--- theories/algebra/cmra_big_op.v | 2 +- theories/algebra/coPset.v | 2 +- theories/algebra/gmultiset.v | 2 +- theories/algebra/gset.v | 2 +- theories/algebra/sts.v | 40 ++++++++++++------------- theories/base_logic/lib/invariants.v | 6 ++-- theories/base_logic/lib/na_invariants.v | 6 ++-- theories/bi/big_op.v | 8 ++--- theories/heap_lang/lib/ticket_lock.v | 6 ++-- theories/heap_lang/proph_map.v | 2 +- 13 files changed, 44 insertions(+), 44 deletions(-) diff --git a/opam b/opam index b75665f01..1e123ae37 100644 --- a/opam +++ b/opam @@ -11,5 +11,5 @@ install: [make "install"] remove: ["rm" "-rf" "%{lib}%/coq/user-contrib/iris"] depends: [ "coq" { (>= "8.7.1" & < "8.10~") | (= "dev") } - "coq-stdpp" { (= "dev.2019-02-07.1.824e9723") | (= "dev") } + "coq-stdpp" { (= "dev.2019-02-20.2.c8c298d4") | (= "dev") } ] diff --git a/tests/proofmode.v b/tests/proofmode.v index 7fd712ccb..b2e87333f 100644 --- a/tests/proofmode.v +++ b/tests/proofmode.v @@ -152,7 +152,7 @@ Qed. Lemma test_iExist_coercion (P : Z → PROP) : (∀ x, P x) -∗ ∃ x, P x. Proof. iIntros "HP". iExists (0:nat). iApply ("HP" $! (0:nat)). Qed. -Lemma test_iExist_tc `{Collection A C} P : (∃ x1 x2 : gset positive, P -∗ P)%I. +Lemma test_iExist_tc `{Set_ A C} P : (∃ x1 x2 : gset positive, P -∗ P)%I. Proof. iExists {[ 1%positive ]}, ∅. auto. Qed. Lemma test_iSpecialize_tc P : (∀ x y z : gset positive, P) -∗ P. diff --git a/theories/algebra/big_op.v b/theories/algebra/big_op.v index be8030e1d..e559caedf 100644 --- a/theories/algebra/big_op.v +++ b/theories/algebra/big_op.v @@ -314,7 +314,7 @@ Section gset. X ## Y → ([^o set] y ∈ X ∪ Y, f y) ≡ ([^o set] y ∈ X, f y) `o` ([^o set] y ∈ Y, f y). Proof. - intros. induction X as [|x X ? IH] using collection_ind_L. + intros. induction X as [|x X ? IH] using set_ind_L. { by rewrite left_id_L big_opS_empty left_id. } rewrite -assoc_L !big_opS_insert; [|set_solver..]. by rewrite -assoc IH; last set_solver. @@ -332,7 +332,7 @@ Section gset. Lemma big_opS_unit X : ([^o set] y ∈ X, monoid_unit) ≡ (monoid_unit : M). Proof. - induction X using collection_ind_L; rewrite /= ?big_opS_insert ?left_id //. + induction X using set_ind_L; rewrite /= ?big_opS_insert ?left_id //. Qed. Lemma big_opS_opS f g X : @@ -458,7 +458,7 @@ Section homomorphisms. `{!MonoidHomomorphism o1 o2 R h} (f : A → M1) X : R (h ([^o1 set] x ∈ X, f x)) ([^o2 set] x ∈ X, h (f x)). Proof. - intros. induction X as [|x X ? IH] using collection_ind_L. + intros. induction X as [|x X ? IH] using set_ind_L. - by rewrite !big_opS_empty monoid_homomorphism_unit. - by rewrite !big_opS_insert // monoid_homomorphism -IH. Qed. @@ -466,7 +466,7 @@ Section homomorphisms. `{!WeakMonoidHomomorphism o1 o2 R h} (f : A → M1) X : X ≠∅ → R (h ([^o1 set] x ∈ X, f x)) ([^o2 set] x ∈ X, h (f x)). Proof. - intros. induction X as [|x X ? IH] using collection_ind_L; [done|]. + intros. induction X as [|x X ? IH] using set_ind_L; [done|]. destruct (decide (X = ∅)) as [->|]. - by rewrite !big_opS_insert // !big_opS_empty !right_id. - by rewrite !big_opS_insert // monoid_homomorphism -IH //. diff --git a/theories/algebra/cmra_big_op.v b/theories/algebra/cmra_big_op.v index 54122de3e..64f5df9f8 100644 --- a/theories/algebra/cmra_big_op.v +++ b/theories/algebra/cmra_big_op.v @@ -23,7 +23,7 @@ Qed. Lemma big_opS_None {M : cmraT} `{Countable A} (f : A → option M) X : ([^op set] x ∈ X, f x) = None ↔ ∀ x, x ∈ X → f x = None. Proof. - induction X as [|x X ? IH] using collection_ind_L; [done|]. + induction X as [|x X ? IH] using set_ind_L; [done|]. rewrite -equiv_None big_opS_insert // equiv_None op_None IH. set_solver. Qed. Lemma big_opMS_None {M : cmraT} `{Countable A} (f : A → option M) X : diff --git a/theories/algebra/coPset.v b/theories/algebra/coPset.v index 763ca1e29..445ea071a 100644 --- a/theories/algebra/coPset.v +++ b/theories/algebra/coPset.v @@ -1,6 +1,6 @@ From iris.algebra Require Export cmra. From iris.algebra Require Import updates local_updates. -From stdpp Require Export collections coPset. +From stdpp Require Export sets coPset. Set Default Proof Using "Type". (** This is pretty much the same as algebra/gset, but I was not able to generalize the construction without breaking canonical structures. *) diff --git a/theories/algebra/gmultiset.v b/theories/algebra/gmultiset.v index 1916b8f40..5fc7d0e4e 100644 --- a/theories/algebra/gmultiset.v +++ b/theories/algebra/gmultiset.v @@ -1,6 +1,6 @@ From iris.algebra Require Export cmra. From iris.algebra Require Import updates local_updates. -From stdpp Require Export collections gmultiset countable. +From stdpp Require Export sets gmultiset countable. Set Default Proof Using "Type". (* The multiset union CMRA *) diff --git a/theories/algebra/gset.v b/theories/algebra/gset.v index 83d923854..ee2c85116 100644 --- a/theories/algebra/gset.v +++ b/theories/algebra/gset.v @@ -1,6 +1,6 @@ From iris.algebra Require Export cmra. From iris.algebra Require Import updates local_updates. -From stdpp Require Export collections gmap mapset. +From stdpp Require Export sets gmap mapset. Set Default Proof Using "Type". (* The union CMRA *) diff --git a/theories/algebra/sts.v b/theories/algebra/sts.v index 0e6442a92..fc51c1906 100644 --- a/theories/algebra/sts.v +++ b/theories/algebra/sts.v @@ -1,4 +1,4 @@ -From stdpp Require Export set. +From stdpp Require Export propset. From iris.algebra Require Export cmra. From iris.algebra Require Import dra. Set Default Proof Using "Type". @@ -12,13 +12,13 @@ Structure stsT := Sts { state : Type; token : Type; prim_step : relation state; - tok : state → set token; + tok : state → propset token; }. Arguments Sts {_ _} _ _. Arguments prim_step {_} _ _. Arguments tok {_} _. -Notation states sts := (set (state sts)). -Notation tokens sts := (set (token sts)). +Notation states sts := (propset (state sts)). +Notation tokens sts := (propset (token sts)). (** * Theory and definitions *) Section sts. @@ -49,8 +49,8 @@ Definition up_set (S : states sts) (T : tokens sts) : states sts := (** Tactic setup *) Hint Resolve Step : core. -Hint Extern 50 (equiv (A:=set _) _ _) => set_solver : sts. -Hint Extern 50 (¬equiv (A:=set _) _ _) => set_solver : sts. +Hint Extern 50 (equiv (A:=propset _) _ _) => set_solver : sts. +Hint Extern 50 (¬equiv (A:=propset _) _ _) => set_solver : sts. Hint Extern 50 (_ ∈ _) => set_solver : sts. Hint Extern 50 (_ ⊆ _) => set_solver : sts. Hint Extern 50 (_ ## _) => set_solver : sts. @@ -62,7 +62,7 @@ Proof. eauto with sts; set_solver. Qed. Global Instance frame_step_proper : Proper ((≡) ==> (=) ==> (=) ==> iff) frame_step. -Proof. move=> ?? /collection_equiv_spec [??]; split; by apply frame_step_mono. Qed. +Proof. move=> ?? /set_equiv_spec [??]; split; by apply frame_step_mono. Qed. Instance closed_proper' : Proper ((≡) ==> (≡) ==> impl) closed. Proof. destruct 3; constructor; intros; setoid_subst; eauto. Qed. Global Instance closed_proper : Proper ((≡) ==> (≡) ==> iff) closed. @@ -71,11 +71,11 @@ Global Instance up_preserving : Proper ((=) ==> flip (⊆) ==> (⊆)) up. Proof. intros s ? <- T T' HT ; apply elem_of_subseteq. induction 1 as [|s1 s2 s3 [T1 T2]]; [constructor|]. - eapply elem_of_mkSet, rtc_l; [eapply Frame_step with T1 T2|]; eauto with sts. + eapply elem_of_PropSet, rtc_l; [eapply Frame_step with T1 T2|]; eauto with sts. Qed. Global Instance up_proper : Proper ((=) ==> (≡) ==> (≡)) up. Proof. - by move=> ??? ?? /collection_equiv_spec [??]; split; apply up_preserving. + by move=> ??? ?? /set_equiv_spec [??]; split; apply up_preserving. Qed. Global Instance up_set_preserving : Proper ((⊆) ==> flip (⊆) ==> (⊆)) up_set. Proof. @@ -84,7 +84,7 @@ Proof. Qed. Global Instance up_set_proper : Proper ((≡) ==> (≡) ==> (≡)) up_set. Proof. - move=> S1 S2 /collection_equiv_spec [??] T1 T2 /collection_equiv_spec [??]; + move=> S1 S2 /set_equiv_spec [??] T1 T2 /set_equiv_spec [??]; split; by apply up_set_preserving. Qed. @@ -127,7 +127,7 @@ Proof. intros s ?; apply elem_of_bind; eauto using elem_of_up. Qed. Lemma elem_of_up_set S T s : s ∈ S → s ∈ up_set S T. Proof. apply subseteq_up_set. Qed. Lemma up_up_set s T : up s T ≡ up_set {[ s ]} T. -Proof. by rewrite /up_set collection_bind_singleton. Qed. +Proof. by rewrite /up_set set_bind_singleton. Qed. Lemma closed_up_set S T : (∀ s, s ∈ S → tok s ## T) → closed (up_set S T) T. Proof. intros HS; unfold up_set; split. @@ -140,7 +140,7 @@ Proof. Qed. Lemma closed_up s T : tok s ## T → closed (up s T) T. Proof. - intros; rewrite -(collection_bind_singleton (λ s, up s T) s). + intros; rewrite -(set_bind_singleton (λ s, up s T) s). apply closed_up_set; set_solver. Qed. Lemma closed_up_set_empty S : closed (up_set S ∅) ∅. @@ -149,7 +149,7 @@ Lemma closed_up_empty s : closed (up s ∅) ∅. Proof. eauto using closed_up with sts. Qed. Lemma up_closed S T : closed S T → up_set S T ≡ S. Proof. - intros ?; apply collection_equiv_spec; split; auto using subseteq_up_set. + intros ?; apply set_equiv_spec; split; auto using subseteq_up_set. intros s; unfold up_set; rewrite elem_of_bind; intros (s'&Hstep&?). induction Hstep; eauto using closed_step. Qed. @@ -159,7 +159,7 @@ Lemma up_set_subseteq S1 T S2 : closed S2 T → S1 ⊆ S2 → sts.up_set S1 T Proof. move=> ?? s [s' [? ?]]. eauto using closed_steps. Qed. Lemma up_op s T1 T2 : up s (T1 ∪ T2) ⊆ up s T1 ∩ up s T2. Proof. (* Notice that the other direction does not hold. *) - intros x Hx. split; eapply elem_of_mkSet, rtc_subrel; try exact Hx. + intros x Hx. split; eapply elem_of_PropSet, rtc_subrel; try exact Hx. - intros; eapply frame_step_mono; last first; try done. set_solver+. - intros; eapply frame_step_mono; last first; try done. set_solver+. Qed. @@ -171,8 +171,8 @@ Notation frame_steps T := (rtc (frame_step T)). (* The type of bounds we can give to the state of an STS. This is the type that we equip with an RA structure. *) Inductive car (sts : stsT) := - | auth : state sts → set (token sts) → car sts - | frag : set (state sts) → set (token sts ) → car sts. + | auth : state sts → propset (token sts) → car sts + | frag : propset (state sts) → propset (token sts) → car sts. Arguments auth {_} _ _. Arguments frag {_} _ _. End sts. @@ -215,7 +215,7 @@ Instance sts_op : Op (car sts) := λ x1 x2, | auth s T1, auth _ T2 => auth s (T1 ∪ T2)(* never happens *) end. -Hint Extern 50 (equiv (A:=set _) _ _) => set_solver : sts. +Hint Extern 50 (equiv (A:=propset _) _ _) => set_solver : sts. Hint Extern 50 (∃ s : state sts, _) => set_solver : sts. Hint Extern 50 (_ ∈ _) => set_solver : sts. Hint Extern 50 (_ ⊆ _) => set_solver : sts. @@ -388,7 +388,7 @@ Lemma sts_up_set_intersection S1 Sf Tf : Proof. intros Hclf. apply (anti_symm (⊆)). - move=>s [HS1 HSf]. split. by apply HS1. by apply subseteq_up_set. - - move=>s [HS1 [s' [/elem_of_mkSet Hsup Hs']]]. split; first done. + - move=>s [HS1 [s' [/elem_of_PropSet Hsup Hs']]]. split; first done. eapply closed_steps, Hsup; first done. set_solver +Hs'. Qed. @@ -449,10 +449,10 @@ Structure stsT := Sts { }. Arguments Sts {_} _. Arguments prim_step {_} _ _. -Notation states sts := (set (state sts)). +Notation states sts := (propset (state sts)). Definition stsT_token := Empty_set. -Definition stsT_tok {sts : stsT} (_ : state sts) : set stsT_token := ∅. +Definition stsT_tok {sts : stsT} (_ : state sts) : propset stsT_token := ∅. Canonical Structure sts_notok (sts : stsT) : sts.stsT := sts.Sts (@prim_step sts) stsT_tok. diff --git a/theories/base_logic/lib/invariants.v b/theories/base_logic/lib/invariants.v index ae1ffaf5f..4e0c9016e 100644 --- a/theories/base_logic/lib/invariants.v +++ b/theories/base_logic/lib/invariants.v @@ -43,11 +43,11 @@ Qed. Lemma fresh_inv_name (E : gset positive) N : ∃ i, i ∉ E ∧ i ∈ (↑N:coPset). Proof. - exists (coPpick (↑ N ∖ coPset.of_gset E)). - rewrite -coPset.elem_of_of_gset (comm and) -elem_of_difference. + exists (coPpick (↑ N ∖ gset_to_coPset E)). + rewrite -elem_of_gset_to_coPset (comm and) -elem_of_difference. apply coPpick_elem_of=> Hfin. eapply nclose_infinite, (difference_finite_inv _ _), Hfin. - apply of_gset_finite. + apply gset_to_coPset_finite. Qed. Lemma inv_alloc N E P : ▷ P ={E}=∗ inv N P. diff --git a/theories/base_logic/lib/na_invariants.v b/theories/base_logic/lib/na_invariants.v index b39356a19..41b81ed33 100644 --- a/theories/base_logic/lib/na_invariants.v +++ b/theories/base_logic/lib/na_invariants.v @@ -80,11 +80,11 @@ Section proofs. iMod (own_updateP with "Hempty") as ([m1 m2]) "[Hm Hown]". { apply prod_updateP'. apply cmra_updateP_id, (reflexivity (R:=eq)). apply (gset_disj_alloc_empty_updateP_strong' (λ i, i ∈ (↑N:coPset))). - intros Ef. exists (coPpick (↑ N ∖ coPset.of_gset Ef)). - rewrite -coPset.elem_of_of_gset comm -elem_of_difference. + intros Ef. exists (coPpick (↑ N ∖ gset_to_coPset Ef)). + rewrite -elem_of_gset_to_coPset comm -elem_of_difference. apply coPpick_elem_of=> Hfin. eapply nclose_infinite, (difference_finite_inv _ _), Hfin. - apply of_gset_finite. } + apply gset_to_coPset_finite. } simpl. iDestruct "Hm" as %(<- & i & -> & ?). rewrite /na_inv. iMod (inv_alloc N with "[-]"); last (iModIntro; iExists i; eauto). diff --git a/theories/bi/big_op.v b/theories/bi/big_op.v index 9ff0864e0..200b7e802 100644 --- a/theories/bi/big_op.v +++ b/theories/bi/big_op.v @@ -1,6 +1,6 @@ From iris.algebra Require Export big_op. From iris.bi Require Import derived_laws_sbi plainly. -From stdpp Require Import countable fin_collections functions. +From stdpp Require Import countable fin_sets functions. Set Default Proof Using "Type". Import interface.bi derived_laws_bi.bi derived_laws_sbi.bi. @@ -788,7 +788,7 @@ Section gset. ([∗ set] y ∈ filter P X, Φ y) ⊣⊢ ([∗ set] y ∈ X, if decide (P y) then Φ y else emp). Proof. - induction X as [|x X ? IH] using collection_ind_L. + induction X as [|x X ? IH] using set_ind_L. { by rewrite filter_empty_L !big_sepS_empty. } destruct (decide (P x)). - rewrite filter_union_L filter_singleton_L //. @@ -841,7 +841,7 @@ Section gset. intros. apply (anti_symm _). { apply forall_intro=> x. apply impl_intro_l, pure_elim_l=> ?; by apply: big_sepS_elem_of. } - induction X as [|x X ? IH] using collection_ind_L; auto using big_sepS_empty'. + induction X as [|x X ? IH] using set_ind_L; auto using big_sepS_empty'. rewrite big_sepS_insert // -persistent_and_sep. apply and_intro. - by rewrite (forall_elim x) pure_True ?True_impl; last set_solver. - rewrite -IH. apply forall_mono=> y. apply impl_intro_l, pure_elim_l=> ?. @@ -853,7 +853,7 @@ Section gset. □ (∀ x, ⌜x ∈ X⌠→ Φ x -∗ Ψ x) -∗ [∗ set] x ∈ X, Ψ x. Proof. - apply wand_intro_l. induction X as [|x X ? IH] using collection_ind_L. + apply wand_intro_l. induction X as [|x X ? IH] using set_ind_L. { by rewrite sep_elim_r. } rewrite !big_sepS_insert // intuitionistically_sep_dup. rewrite -assoc [(□ _ ∗ _)%I]comm -!assoc assoc. apply sep_mono. diff --git a/theories/heap_lang/lib/ticket_lock.v b/theories/heap_lang/lib/ticket_lock.v index 0770df87c..017d256c7 100644 --- a/theories/heap_lang/lib/ticket_lock.v +++ b/theories/heap_lang/lib/ticket_lock.v @@ -42,7 +42,7 @@ Section proof. Definition lock_inv (γ : gname) (lo ln : loc) (R : iProp Σ) : iProp Σ := (∃ o n : nat, lo ↦ #o ∗ ln ↦ #n ∗ - own γ (◠(Excl' o, GSet (seq_set 0 n))) ∗ + own γ (◠(Excl' o, GSet (set_seq 0 n))) ∗ ((own γ (◯ (Excl' o, GSet ∅)) ∗ R) ∨ own γ (◯ (ε, GSet {[ o ]}))))%I. Definition is_lock (γ : gname) (lk : val) (R : iProp Σ) : iProp Σ := @@ -117,8 +117,8 @@ Section proof. - iMod (own_update with "Hauth") as "[Hauth Hofull]". { eapply auth_update_alloc, prod_local_update_2. eapply (gset_disj_alloc_empty_local_update _ {[ n ]}). - apply (seq_set_S_disjoint 0). } - rewrite -(seq_set_S_union_L 0). + apply (set_seq_S_end_disjoint 0). } + rewrite -(set_seq_S_end_union_L 0). wp_cas_suc. iModIntro. iSplitL "Hlo' Hln' Haown Hauth". { iNext. iExists o', (S n). rewrite Nat2Z.inj_succ -Z.add_1_r. by iFrame. } diff --git a/theories/heap_lang/proph_map.v b/theories/heap_lang/proph_map.v index cb03d7717..39308f1a9 100644 --- a/theories/heap_lang/proph_map.v +++ b/theories/heap_lang/proph_map.v @@ -35,7 +35,7 @@ Section definitions. (** The first resolve for [p] in [pvs] *) Definition first_resolve (pvs : proph_val_list P V) (p : P) : option V := - (map_of_list pvs : gmap P V) !! p. + (list_to_map pvs : gmap P V) !! p. Definition first_resolve_in_list (R : proph_map P V) (pvs : proph_val_list P V) := ∀ p v, p ∈ dom (gset _) R → -- GitLab