From 4d94262b23f4c02b532b0474a3aebb316cf6c68d Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Tue, 16 Feb 2016 18:40:12 +0100 Subject: [PATCH] Use keywords in the notation of big ops on uPred. MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We now have: Π★{map Q } ... Π★{set Q } ... to differentiate between sets and maps. --- algebra/upred_big_op.v | 63 +++++++++++++++++++++--------------------- heap_lang/heap.v | 2 +- 2 files changed, 32 insertions(+), 33 deletions(-) diff --git a/algebra/upred_big_op.v b/algebra/upred_big_op.v index 9b5e59594..d7f75b7ec 100644 --- a/algebra/upred_big_op.v +++ b/algebra/upred_big_op.v @@ -3,28 +3,27 @@ From prelude Require Import fin_maps fin_collections. (** * Big ops over lists *) (* These are the basic building blocks for other big ops *) -Fixpoint uPred_list_and {M} (Ps : list (uPred M)) : uPred M:= - match Ps with [] => True | P :: Ps => P ∧ uPred_list_and Ps end%I. -Instance: Params (@uPred_list_and) 1. -Notation "'Π∧' Ps" := (uPred_list_and Ps) (at level 20) : uPred_scope. -Fixpoint uPred_list_sep {M} (Ps : list (uPred M)) : uPred M := - match Ps with [] => True | P :: Ps => P ★ uPred_list_sep Ps end%I. -Instance: Params (@uPred_list_sep) 1. -Notation "'Π★' Ps" := (uPred_list_sep Ps) (at level 20) : uPred_scope. +Fixpoint uPred_big_and {M} (Ps : list (uPred M)) : uPred M:= + match Ps with [] => True | P :: Ps => P ∧ uPred_big_and Ps end%I. +Instance: Params (@uPred_big_and) 1. +Notation "'Π∧' Ps" := (uPred_big_and Ps) (at level 20) : uPred_scope. +Fixpoint uPred_big_sep {M} (Ps : list (uPred M)) : uPred M := + match Ps with [] => True | P :: Ps => P ★ uPred_big_sep Ps end%I. +Instance: Params (@uPred_big_sep) 1. +Notation "'Π★' Ps" := (uPred_big_sep Ps) (at level 20) : uPred_scope. (** * Other big ops *) (** We use a type class to obtain overloaded notations *) -Class UPredBigSep (M : cmraT) (A B : Type) := - uPred_big_sep : A → B → uPred M. -Instance: Params (@uPred_big_sep) 4. -Notation "'Π★{' x } P" := (uPred_big_sep x P) - (at level 20, x at level 10, format "Π★{ x } P") : uPred_scope. +Definition uPred_big_sepM {M} `{FinMapToList K A MA} + (m : MA) (P : K → A → uPred M) : uPred M := + uPred_big_sep (curry P <$> map_to_list m). +Notation "'Π★{map' m } P" := (uPred_big_sepM m P) + (at level 20, m at level 10, format "Π★{map m } P") : uPred_scope. -Instance uPred_big_sepM {M} `{FinMapToList K A MA} : - UPredBigSep M MA (K → A → uPred M) := λ m P, - uPred_list_sep (curry P <$> map_to_list m). -Instance uPred_big_sepC {M} `{Elements A C} : - UPredBigSep M C (A → uPred M) := λ X P, uPred_list_sep (P <$> elements X). +Definition uPred_big_sepS {M} `{Elements A C} + (X : C) (P : A → uPred M) : uPred M := uPred_big_sep (P <$> elements X). +Notation "'Π★{set' X } P" := (uPred_big_sepS X P) + (at level 20, X at level 10, format "Π★{set X } P") : uPred_scope. (** * Always stability for lists *) Class AlwaysStableL {M} (Ps : list (uPred M)) := @@ -38,47 +37,47 @@ Implicit Types Ps Qs : list (uPred M). Implicit Types A : Type. (* Big ops *) -Global Instance list_and_proper : Proper ((≡) ==> (≡)) (@uPred_list_and M). +Global Instance big_and_proper : Proper ((≡) ==> (≡)) (@uPred_big_and M). Proof. by induction 1 as [|P Q Ps Qs HPQ ? IH]; rewrite /= ?HPQ ?IH. Qed. -Global Instance list_sep_proper : Proper ((≡) ==> (≡)) (@uPred_list_sep M). +Global Instance big_sep_proper : Proper ((≡) ==> (≡)) (@uPred_big_sep M). Proof. by induction 1 as [|P Q Ps Qs HPQ ? IH]; rewrite /= ?HPQ ?IH. Qed. -Global Instance list_and_perm : Proper ((≡ₚ) ==> (≡)) (@uPred_list_and M). +Global Instance big_and_perm : Proper ((≡ₚ) ==> (≡)) (@uPred_big_and M). Proof. induction 1 as [|P Ps Qs ? IH|P Q Ps|]; simpl; auto. * by rewrite IH. * by rewrite !assoc (comm _ P). * etransitivity; eauto. Qed. -Global Instance list_sep_perm : Proper ((≡ₚ) ==> (≡)) (@uPred_list_sep M). +Global Instance big_sep_perm : Proper ((≡ₚ) ==> (≡)) (@uPred_big_sep M). Proof. induction 1 as [|P Ps Qs ? IH|P Q Ps|]; simpl; auto. * by rewrite IH. * by rewrite !assoc (comm _ P). * etransitivity; eauto. Qed. -Lemma list_and_app Ps Qs : (Π∧ (Ps ++ Qs))%I ≡ (Π∧ Ps ∧ Π∧ Qs)%I. +Lemma big_and_app Ps Qs : (Π∧ (Ps ++ Qs))%I ≡ (Π∧ Ps ∧ Π∧ Qs)%I. Proof. by induction Ps as [|?? IH]; rewrite /= ?left_id -?assoc ?IH. Qed. -Lemma list_sep_app Ps Qs : (Π★ (Ps ++ Qs))%I ≡ (Π★ Ps ★ Π★ Qs)%I. +Lemma big_sep_app Ps Qs : (Π★ (Ps ++ Qs))%I ≡ (Π★ Ps ★ Π★ Qs)%I. Proof. by induction Ps as [|?? IH]; rewrite /= ?left_id -?assoc ?IH. Qed. -Lemma list_sep_and Ps : (Π★ Ps) ⊑ (Π∧ Ps). +Lemma big_sep_and Ps : (Π★ Ps) ⊑ (Π∧ Ps). Proof. by induction Ps as [|P Ps IH]; simpl; auto with I. Qed. -Lemma list_and_elem_of Ps P : P ∈ Ps → (Π∧ Ps) ⊑ P. +Lemma big_and_elem_of Ps P : P ∈ Ps → (Π∧ Ps) ⊑ P. Proof. induction 1; simpl; auto with I. Qed. -Lemma list_sep_elem_of Ps P : P ∈ Ps → (Π★ Ps) ⊑ P. +Lemma big_sep_elem_of Ps P : P ∈ Ps → (Π★ Ps) ⊑ P. Proof. induction 1; simpl; auto with I. Qed. (* Big ops over finite maps *) Section fin_map. Context `{FinMap K Ma} {A} (P : K → A → uPred M). - Lemma big_sepM_empty : (Π★{∅} P)%I ≡ True%I. + Lemma big_sepM_empty : (Π★{map ∅} P)%I ≡ True%I. Proof. by rewrite /uPred_big_sep /uPred_big_sepM map_to_list_empty. Qed. Lemma big_sepM_insert (m : Ma A) i x : - m !! i = None → (Π★{<[i:=x]> m} P)%I ≡ (P i x ★ Π★{m} P)%I. + m !! i = None → (Π★{map <[i:=x]> m} P)%I ≡ (P i x ★ Π★{map m} P)%I. Proof. intros ?; by rewrite /uPred_big_sep /uPred_big_sepM map_to_list_insert. Qed. - Lemma big_sepM_singleton i x : (Π★{{[i ↦ x]}} P)%I ≡ (P i x)%I. + Lemma big_sepM_singleton i x : (Π★{map {[i ↦ x]}} P)%I ≡ (P i x)%I. Proof. rewrite -insert_empty big_sepM_insert/=; last auto using lookup_empty. by rewrite big_sepM_empty right_id. @@ -88,9 +87,9 @@ End fin_map. (* Always stable *) Local Notation AS := AlwaysStable. Local Notation ASL := AlwaysStableL. -Global Instance list_and_always_stable Ps : ASL Ps → AS (Π∧ Ps). +Global Instance big_and_always_stable Ps : ASL Ps → AS (Π∧ Ps). Proof. induction 1; apply _. Qed. -Global Instance list_sep_always_stable Ps : ASL Ps → AS (Π★ Ps). +Global Instance big_sep_always_stable Ps : ASL Ps → AS (Π★ Ps). Proof. induction 1; apply _. Qed. Global Instance nil_always_stable : ASL (@nil (uPred M)). diff --git a/heap_lang/heap.v b/heap_lang/heap.v index 7bf9a2d05..a3a890a45 100644 --- a/heap_lang/heap.v +++ b/heap_lang/heap.v @@ -72,7 +72,7 @@ Section heap. Qed. Lemma heap_alloc N σ : - ownP σ ⊑ pvs N N (∃ γ, heap_ctx HeapI γ N ∧ Π★{σ} heap_mapsto HeapI γ). + ownP σ ⊑ pvs N N (∃ γ, heap_ctx HeapI γ N ∧ Π★{map σ} heap_mapsto HeapI γ). Proof. rewrite -{1}(from_to_heap σ); etransitivity; first apply (auth_alloc (ownP ∘ of_heap) N (to_heap σ)), to_heap_valid. -- GitLab