From 10edecb60edf7dc62c43bdde1cd39c403335c46c Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Thu, 5 Apr 2018 11:23:06 +0200 Subject: [PATCH] Notation `##@{A}`. --- theories/base.v | 10 ++++++++-- theories/coPset.v | 2 +- theories/collections.v | 4 ++-- theories/mapset.v | 2 +- 4 files changed, 12 insertions(+), 6 deletions(-) diff --git a/theories/base.v b/theories/base.v index 801af5bf..9bbc2f3b 100644 --- a/theories/base.v +++ b/theories/base.v @@ -866,6 +866,10 @@ Infix "##" := disjoint (at level 70) : stdpp_scope. Notation "(##)" := disjoint (only parsing) : stdpp_scope. Notation "( X ##.)" := (disjoint X) (only parsing) : stdpp_scope. Notation "(.## X )" := (λ Y, Y ## X) (only parsing) : stdpp_scope. + +Infix "##@{ A }" := (@disjoint A _) (at level 70, only parsing) : stdpp_scope. +Notation "(##@{ A } )" := (@disjoint A _) (only parsing) : stdpp_scope. + Infix "##*" := (Forall2 (##)) (at level 70) : stdpp_scope. Notation "(##*)" := (Forall2 (##)) (only parsing) : stdpp_scope. Infix "##**" := (Forall2 (##*)) (at level 70) : stdpp_scope. @@ -897,17 +901,19 @@ Class DisjointList A := disjoint_list : list A → Prop. Hint Mode DisjointList ! : typeclass_instances. Instance: Params (@disjoint_list) 2. Notation "## Xs" := (disjoint_list Xs) (at level 20, format "## Xs") : stdpp_scope. +Notation "##@{ A } Xs" := + (@disjoint_list A _ Xs) (at level 20, only parsing) : stdpp_scope. Section disjoint_list. Context `{Disjoint A, Union A, Empty A}. Implicit Types X : A. Inductive disjoint_list_default : DisjointList A := - | disjoint_nil_2 : ## (@nil A) + | disjoint_nil_2 : ##@{A} [] | disjoint_cons_2 (X : A) (Xs : list A) : X ## ⋃ Xs → ## Xs → ## (X :: Xs). Global Existing Instance disjoint_list_default. - Lemma disjoint_list_nil : ## @nil A ↔ True. + Lemma disjoint_list_nil : ##@{A} [] ↔ True. Proof. split; constructor. Qed. Lemma disjoint_list_cons X Xs : ## (X :: Xs) ↔ X ## ⋃ Xs ∧ ## Xs. Proof. split. inversion_clear 1; auto. intros [??]. constructor; auto. Qed. diff --git a/theories/coPset.v b/theories/coPset.v index 77b0d2b2..171197ca 100644 --- a/theories/coPset.v +++ b/theories/coPset.v @@ -194,7 +194,7 @@ Instance coPset_elem_of_dec : RelDecision (∈@{coPset}). Proof. solve_decision. Defined. Instance coPset_equiv_dec : RelDecision (≡@{coPset}). Proof. refine (λ X Y, cast_if (decide (X = Y))); abstract (by fold_leibniz). Defined. -Instance mapset_disjoint_dec : RelDecision (@disjoint coPset _). +Instance mapset_disjoint_dec : RelDecision (##@{coPset}). Proof. refine (λ X Y, cast_if (decide (X ∩ Y = ∅))); abstract (by rewrite disjoint_intersection_L). diff --git a/theories/collections.v b/theories/collections.v index cc4b9cc7..050615a4 100644 --- a/theories/collections.v +++ b/theories/collections.v @@ -30,7 +30,7 @@ Section setoids_simple. Proof. apply _. Qed. Global Instance elem_of_proper : Proper ((=) ==> (≡) ==> iff) (∈@{C}) | 5. Proof. by intros x ? <- X Y. Qed. - Global Instance disjoint_proper: Proper ((≡) ==> (≡) ==> iff) (@disjoint C _). + Global Instance disjoint_proper: Proper ((≡) ==> (≡) ==> iff) (##@{C}). Proof. intros X1 X2 HX Y1 Y2 HY; apply forall_proper; intros x. by rewrite HX, HY. Qed. @@ -407,7 +407,7 @@ Section simple_collection. Lemma elem_of_disjoint X Y : X ## Y ↔ ∀ x, x ∈ X → x ∈ Y → False. Proof. done. Qed. - Global Instance disjoint_sym : Symmetric (@disjoint C _). + Global Instance disjoint_sym : Symmetric (##@{C}). Proof. intros X Y. set_solver. Qed. Lemma disjoint_empty_l Y : ∅ ## Y. Proof. set_solver. Qed. diff --git a/theories/mapset.v b/theories/mapset.v index fa5e02e5..14da0ad6 100644 --- a/theories/mapset.v +++ b/theories/mapset.v @@ -80,7 +80,7 @@ Section deciders. Proof. refine (λ X1 X2, cast_if (decide (X1 = X2))); abstract (by fold_leibniz). Defined. Global Instance mapset_elem_of_dec : RelDecision (∈@{mapset M}) | 1. Proof. refine (λ x X, cast_if (decide (mapset_car X !! x = Some ()))); done. Defined. - Global Instance mapset_disjoint_dec : RelDecision (@disjoint (mapset M) _). + Global Instance mapset_disjoint_dec : RelDecision (##@{mapset M}). Proof. refine (λ X1 X2, cast_if (decide (X1 ∩ X2 = ∅))); abstract (by rewrite disjoint_intersection_L). -- GitLab