diff --git a/opam b/opam index e3f1c66f6607ab205aecc834a2b2ecb3cbb350f8..ff049a5080b450388b30ac3ccbc556ecdf9864e2 100644 --- a/opam +++ b/opam @@ -12,5 +12,5 @@ remove: ["rm" "-rf" "%{lib}%/coq/user-contrib/iris"] depends: [ "coq" { >= "8.6.1" & < "8.8~" } "coq-mathcomp-ssreflect" { (>= "1.6.1" & < "1.7~") | (= "dev") } - "coq-stdpp" { (= "dev.2017-10-28.7") | (= "dev") } + "coq-stdpp" { (= "dev.2017-11-11.0") | (= "dev") } ] diff --git a/theories/algebra/big_op.v b/theories/algebra/big_op.v index a18c18aeb4804e8fe5f97cda6cf11fb05bf65c7b..3729d5f48233a6347e7cb53a50703fd11e647fcb 100644 --- a/theories/algebra/big_op.v +++ b/theories/algebra/big_op.v @@ -30,10 +30,10 @@ Arguments big_opL {M} o {_ A} _ !_ /. Typeclasses Opaque big_opL. Notation "'[^' o 'list]' k ↦ x ∈ l , P" := (big_opL o (λ k x, P) l) (at level 200, o at level 1, l at level 10, k, x at level 1, right associativity, - format "[^ o list] k ↦ x ∈ l , P") : C_scope. + format "[^ o list] k ↦ x ∈ l , P") : stdpp_scope. Notation "'[^' o 'list]' x ∈ l , P" := (big_opL o (λ _ x, P) l) (at level 200, o at level 1, l at level 10, x at level 1, right associativity, - format "[^ o list] x ∈ l , P") : C_scope. + format "[^ o list] x ∈ l , P") : stdpp_scope. Definition big_opM `{Monoid M o} `{Countable K} {A} (f : K → A → M) (m : gmap K A) : M := big_opL o (λ _, curry f) (map_to_list m). @@ -42,10 +42,10 @@ Arguments big_opM {M} o {_ K _ _ A} _ _ : simpl never. Typeclasses Opaque big_opM. Notation "'[^' o 'map]' k ↦ x ∈ m , P" := (big_opM o (λ k x, P) m) (at level 200, o at level 1, m at level 10, k, x at level 1, right associativity, - format "[^ o map] k ↦ x ∈ m , P") : C_scope. + format "[^ o map] k ↦ x ∈ m , P") : stdpp_scope. Notation "'[^' o 'map]' x ∈ m , P" := (big_opM o (λ _ x, P) m) (at level 200, o at level 1, m at level 10, x at level 1, right associativity, - format "[^ o map] x ∈ m , P") : C_scope. + format "[^ o map] x ∈ m , P") : stdpp_scope. Definition big_opS `{Monoid M o} `{Countable A} (f : A → M) (X : gset A) : M := big_opL o (λ _, f) (elements X). @@ -54,7 +54,7 @@ Arguments big_opS {M} o {_ A _ _} _ _ : simpl never. Typeclasses Opaque big_opS. Notation "'[^' o 'set]' x ∈ X , P" := (big_opS o (λ x, P) X) (at level 200, o at level 1, X at level 10, x at level 1, right associativity, - format "[^ o set] x ∈ X , P") : C_scope. + format "[^ o set] x ∈ X , P") : stdpp_scope. Definition big_opMS `{Monoid M o} `{Countable A} (f : A → M) (X : gmultiset A) : M := big_opL o (λ _, f) (elements X). @@ -63,7 +63,7 @@ Arguments big_opMS {M} o {_ A _ _} _ _ : simpl never. Typeclasses Opaque big_opMS. Notation "'[^' o 'mset]' x ∈ X , P" := (big_opMS o (λ x, P) X) (at level 200, o at level 1, X at level 10, x at level 1, right associativity, - format "[^ o mset] x ∈ X , P") : C_scope. + format "[^ o mset] x ∈ X , P") : stdpp_scope. (** * Properties about big ops *) Section big_op. diff --git a/theories/algebra/cmra.v b/theories/algebra/cmra.v index b6c22f5ba318be9dae79d2b19d144a26013b32f1..3c3b1171ca0a40c770e295e5aaff79a7dcecd443 100644 --- a/theories/algebra/cmra.v +++ b/theories/algebra/cmra.v @@ -8,8 +8,8 @@ Instance: Params (@pcore) 2. Class Op (A : Type) := op : A → A → A. Hint Mode Op ! : typeclass_instances. Instance: Params (@op) 2. -Infix "â‹…" := op (at level 50, left associativity) : C_scope. -Notation "(â‹…)" := op (only parsing) : C_scope. +Infix "â‹…" := op (at level 50, left associativity) : stdpp_scope. +Notation "(â‹…)" := op (only parsing) : stdpp_scope. (* The inclusion quantifies over [A], not [option A]. This means we do not get reflexivity. However, if we used [option A], the following would no longer @@ -17,8 +17,8 @@ Notation "(â‹…)" := op (only parsing) : C_scope. x ≼ y ↔ x.1 ≼ y.1 ∧ x.2 ≼ y.2 *) Definition included `{Equiv A, Op A} (x y : A) := ∃ z, y ≡ x â‹… z. -Infix "≼" := included (at level 70) : C_scope. -Notation "(≼)" := included (only parsing) : C_scope. +Infix "≼" := included (at level 70) : stdpp_scope. +Notation "(≼)" := included (only parsing) : stdpp_scope. Hint Extern 0 (_ ≼ _) => reflexivity. Instance: Params (@included) 3. @@ -31,11 +31,11 @@ Notation "✓{ n } x" := (validN n x) Class Valid (A : Type) := valid : A → Prop. Hint Mode Valid ! : typeclass_instances. Instance: Params (@valid) 2. -Notation "✓ x" := (valid x) (at level 20) : C_scope. +Notation "✓ x" := (valid x) (at level 20) : stdpp_scope. Definition includedN `{Dist A, Op A} (n : nat) (x y : A) := ∃ z, y ≡{n}≡ x â‹… z. Notation "x ≼{ n } y" := (includedN n x y) - (at level 70, n at next level, format "x ≼{ n } y") : C_scope. + (at level 70, n at next level, format "x ≼{ n } y") : stdpp_scope. Instance: Params (@includedN) 4. Hint Extern 0 (_ ≼{_} _) => reflexivity. @@ -140,7 +140,7 @@ End cmra_mixin. Definition opM {A : cmraT} (x : A) (my : option A) := match my with Some y => x â‹… y | None => x end. -Infix "â‹…?" := opM (at level 50, left associativity) : C_scope. +Infix "â‹…?" := opM (at level 50, left associativity) : stdpp_scope. (** * CoreId elements *) Class CoreId {A : cmraT} (x : A) := core_id : pcore x ≡ Some x. diff --git a/theories/algebra/dra.v b/theories/algebra/dra.v index fcbccb2a249c522e4b8cdef15746551219d8b09b..dc459baf45b763368a8be51d8773630aef0541e4 100644 --- a/theories/algebra/dra.v +++ b/theories/algebra/dra.v @@ -199,7 +199,7 @@ Proof. split; naive_solver eauto using dra_op_valid. Qed. (* TODO: This has to be proven again. *) (* Lemma to_validity_included x y: - (✓ y ∧ to_validity x ≼ to_validity y)%C ↔ (✓ x ∧ x ≼ y). + (✓ y ∧ to_validity x ≼ to_validity y)%stdpp ↔ (✓ x ∧ x ≼ y). Proof. split. - move=>[Hvl [z [Hvxz EQ]]]. move:(Hvl)=>Hvl'. apply Hvxz in Hvl'. diff --git a/theories/base_logic/deprecated.v b/theories/base_logic/deprecated.v index 700c9e026e35b28b8c9c9ca957054f0ff61c14b6..e6cf27f42f870b11511531fa887fdc01e2fc1d7b 100644 --- a/theories/base_logic/deprecated.v +++ b/theories/base_logic/deprecated.v @@ -5,12 +5,12 @@ From iris.base_logic Require Import primitive. Set Default Proof Using "Type". (* Deprecated 2016-11-22. Use ⌜φ⌠instead. *) -Notation "■φ" := (uPred_pure φ%C%type) +Notation "■φ" := (uPred_pure φ%stdpp%type) (at level 20, right associativity, only parsing) : uPred_scope. (* Deprecated 2016-11-22. Use ⌜x = y⌠instead. *) -Notation "x = y" := (uPred_pure (x%C%type = y%C%type)) (only parsing) : uPred_scope. +Notation "x = y" := (uPred_pure (x%stdpp%type = y%stdpp%type)) (only parsing) : uPred_scope. (* Deprecated 2016-11-22. Use ⌜x ## y ⌠instead. *) -Notation "x ## y" := (uPred_pure (x%C%type ## y%C%type)) (only parsing) : uPred_scope. +Notation "x ## y" := (uPred_pure (x%stdpp%type ## y%stdpp%type)) (only parsing) : uPred_scope. *) diff --git a/theories/base_logic/double_negation.v b/theories/base_logic/double_negation.v index 6f2d19444e5bef72e7ea85ba7a412983f8183fa7..5f47c05ecbeff4185ac2249695703b7e93e54c3b 100644 --- a/theories/base_logic/double_negation.v +++ b/theories/base_logic/double_negation.v @@ -9,7 +9,7 @@ Definition uPred_nnupd {M} (P: uPred M) : uPred M := Notation "|=n=> Q" := (uPred_nnupd Q) (at level 99, Q at level 200, format "|=n=> Q") : bi_scope. Notation "P =n=> Q" := (P ⊢ |=n=> Q) - (at level 99, Q at level 200, only parsing) : C_scope. + (at level 99, Q at level 200, only parsing) : stdpp_scope. Notation "P =n=∗ Q" := (P -∗ |=n=> Q)%I (at level 99, Q at level 200, format "P =n=∗ Q") : bi_scope. diff --git a/theories/base_logic/lib/fancy_updates.v b/theories/base_logic/lib/fancy_updates.v index 0ce3e275ffa556b1bde0f10f200832fd051a744b..76ff311b4f0465224786d308461b513a9b5651da 100644 --- a/theories/base_logic/lib/fancy_updates.v +++ b/theories/base_logic/lib/fancy_updates.v @@ -23,7 +23,7 @@ Notation "P ={ E1 , E2 }=∗ Q" := (P -∗ |={E1,E2}=> Q)%I (at level 99, E1,E2 at level 50, Q at level 200, format "P ={ E1 , E2 }=∗ Q") : bi_scope. Notation "P ={ E1 , E2 }=∗ Q" := (P -∗ |={E1,E2}=> Q) - (at level 99, E1, E2 at level 50, Q at level 200, only parsing) : C_scope. + (at level 99, E1, E2 at level 50, Q at level 200, only parsing) : stdpp_scope. Notation "|={ E }=> Q" := (fupd E E Q) (at level 99, E at level 50, Q at level 200, @@ -32,7 +32,7 @@ Notation "P ={ E }=∗ Q" := (P -∗ |={E}=> Q)%I (at level 99, E at level 50, Q at level 200, format "P ={ E }=∗ Q") : bi_scope. Notation "P ={ E }=∗ Q" := (P -∗ |={E}=> Q) - (at level 99, E at level 50, Q at level 200, only parsing) : C_scope. + (at level 99, E at level 50, Q at level 200, only parsing) : stdpp_scope. Section fupd. Context `{invG Σ}. diff --git a/theories/base_logic/lib/namespaces.v b/theories/base_logic/lib/namespaces.v index 76857b335d1feb7ecf538fcb3e5e3976c475ddb7..dc226f276d5a29ec5eaf26085ca33c4bda87d040 100644 --- a/theories/base_logic/lib/namespaces.v +++ b/theories/base_logic/lib/namespaces.v @@ -21,8 +21,8 @@ Instance nclose : UpClose namespace coPset := unseal nclose_aux. Definition nclose_eq : @nclose = @nclose_def := seal_eq nclose_aux. Notation "N .@ x" := (ndot N x) - (at level 19, left associativity, format "N .@ x") : C_scope. -Notation "(.@)" := ndot (only parsing) : C_scope. + (at level 19, left associativity, format "N .@ x") : stdpp_scope. +Notation "(.@)" := ndot (only parsing) : stdpp_scope. Instance ndisjoint : Disjoint namespace := λ N1 N2, nclose N1 ## nclose N2. diff --git a/theories/base_logic/lib/viewshifts.v b/theories/base_logic/lib/viewshifts.v index 8636f0fba1bbbe5ca698452aabe301d5e32f958c..e45a87c0c9420662be8df21813459187b458b83d 100644 --- a/theories/base_logic/lib/viewshifts.v +++ b/theories/base_logic/lib/viewshifts.v @@ -16,10 +16,10 @@ Notation "P ={ E }=> Q" := (P ={E,E}=> Q)%I Notation "P ={ E1 , E2 }=> Q" := (P ={E1,E2}=> Q)%I (at level 99, E1,E2 at level 50, Q at level 200, - format "P ={ E1 , E2 }=> Q") : C_scope. + format "P ={ E1 , E2 }=> Q") : stdpp_scope. Notation "P ={ E }=> Q" := (P ={E}=> Q)%I (at level 99, E at level 50, Q at level 200, - format "P ={ E }=> Q") : C_scope. + format "P ={ E }=> Q") : stdpp_scope. Section vs. Context `{invG Σ}. diff --git a/theories/base_logic/upred.v b/theories/base_logic/upred.v index 46b102d5801faecbca24f4fb2c294b8b5d9a4d93..094577409a44b8e0762494449ba3a0ccb6b279fc 100644 --- a/theories/base_logic/upred.v +++ b/theories/base_logic/upred.v @@ -558,7 +558,7 @@ Notation "✓ x" := (uPred_cmra_valid x) (at level 20) : bi_scope. Notation "|==> Q" := (uPred_bupd Q) (at level 99, Q at level 200, format "|==> Q") : bi_scope. Notation "P ==∗ Q" := (P ⊢ |==> Q) - (at level 99, Q at level 200, only parsing) : C_scope. + (at level 99, Q at level 200, only parsing) : stdpp_scope. Notation "P ==∗ Q" := (P -∗ |==> Q)%I (at level 99, Q at level 200, format "P ==∗ Q") : bi_scope. diff --git a/theories/bi/interface.v b/theories/bi/interface.v index fc47bb496f15a977c2a8f513367b0961a68ceedf..ed0851fd9a2981fbfdbf792a3594492db7d495d4 100644 --- a/theories/bi/interface.v +++ b/theories/bi/interface.v @@ -28,7 +28,7 @@ Section bi_mixin. Local Notation "'emp'" := bi_emp. Local Notation "'True'" := (bi_pure True). Local Notation "'False'" := (bi_pure False). - Local Notation "'⌜' φ 'âŒ'" := (bi_pure φ%type%C). + Local Notation "'⌜' φ 'âŒ'" := (bi_pure φ%type%stdpp). Local Infix "∧" := bi_and. Local Infix "∨" := bi_or. Local Infix "→" := bi_impl. @@ -213,7 +213,7 @@ Arguments bi_dist : simpl never. Arguments bi_equiv : simpl never. Arguments bi_entails {PROP} _%I _%I : simpl never, rename. Arguments bi_emp {PROP} : simpl never, rename. -Arguments bi_pure {PROP} _%C : simpl never, rename. +Arguments bi_pure {PROP} _%stdpp : simpl never, rename. Arguments bi_and {PROP} _%I _%I : simpl never, rename. Arguments bi_or {PROP} _%I _%I : simpl never, rename. Arguments bi_impl {PROP} _%I _%I : simpl never, rename. @@ -255,7 +255,7 @@ Structure sbi := SBI { Arguments sbi_car : simpl never. Arguments sbi_entails {PROP} _%I _%I : simpl never, rename. Arguments bi_emp {PROP} : simpl never, rename. -Arguments bi_pure {PROP} _%C : simpl never, rename. +Arguments bi_pure {PROP} _%stdpp : simpl never, rename. Arguments bi_and {PROP} _%I _%I : simpl never, rename. Arguments bi_or {PROP} _%I _%I : simpl never, rename. Arguments bi_impl {PROP} _%I _%I : simpl never, rename. @@ -278,7 +278,7 @@ Arguments sbi_dist : simpl never. Arguments sbi_equiv : simpl never. Arguments sbi_entails {PROP} _%I _%I : simpl never, rename. Arguments sbi_emp {PROP} : simpl never, rename. -Arguments sbi_pure {PROP} _%C : simpl never, rename. +Arguments sbi_pure {PROP} _%stdpp : simpl never, rename. Arguments sbi_and {PROP} _%I _%I : simpl never, rename. Arguments sbi_or {PROP} _%I _%I : simpl never, rename. Arguments sbi_impl {PROP} _%I _%I : simpl never, rename. @@ -295,17 +295,17 @@ Hint Extern 0 (bi_entails _ _) => reflexivity. Instance bi_rewrite_relation (PROP : bi) : RewriteRelation (@bi_entails PROP). Instance bi_inhabited {PROP : bi} : Inhabited PROP := populate (bi_pure True). -Notation "P ⊢ Q" := (bi_entails P%I Q%I) : C_scope. -Notation "(⊢)" := bi_entails (only parsing) : C_scope. +Notation "P ⊢ Q" := (bi_entails P%I Q%I) : stdpp_scope. +Notation "(⊢)" := bi_entails (only parsing) : stdpp_scope. Notation "P ⊣⊢ Q" := (equiv (A:=bi_car _) P%I Q%I) - (at level 95, no associativity) : C_scope. -Notation "(⊣⊢)" := (equiv (A:=bi_car _)) (only parsing) : C_scope. + (at level 95, no associativity) : stdpp_scope. +Notation "(⊣⊢)" := (equiv (A:=bi_car _)) (only parsing) : stdpp_scope. -Notation "P -∗ Q" := (P ⊢ Q) : C_scope. +Notation "P -∗ Q" := (P ⊢ Q) : stdpp_scope. Notation "'emp'" := (bi_emp) : bi_scope. -Notation "'⌜' φ 'âŒ'" := (bi_pure φ%type%C) : bi_scope. +Notation "'⌜' φ 'âŒ'" := (bi_pure φ%type%stdpp) : bi_scope. Notation "'True'" := (bi_pure True) : bi_scope. Notation "'False'" := (bi_pure False) : bi_scope. Infix "∧" := bi_and : bi_scope. diff --git a/theories/program_logic/hoare.v b/theories/program_logic/hoare.v index 11d2e689d2d0f1414e30cc9a7b9e990c002f5586..b1656b7e6e4f8d82f1b99277b08f072fb9a55ec8 100644 --- a/theories/program_logic/hoare.v +++ b/theories/program_logic/hoare.v @@ -10,17 +10,17 @@ Instance: Params (@ht) 4. Notation "{{ P } } e @ E {{ Φ } }" := (ht E P%I e%E Φ%I) (at level 20, P, e, Φ at level 200, - format "{{ P } } e @ E {{ Φ } }") : C_scope. + format "{{ P } } e @ E {{ Φ } }") : stdpp_scope. Notation "{{ P } } e {{ Φ } }" := (ht ⊤ P%I e%E Φ%I) (at level 20, P, e, Φ at level 200, - format "{{ P } } e {{ Φ } }") : C_scope. + format "{{ P } } e {{ Φ } }") : stdpp_scope. Notation "{{ P } } e @ E {{ v , Q } }" := (ht E P%I e%E (λ v, Q)%I) (at level 20, P, e, Q at level 200, - format "{{ P } } e @ E {{ v , Q } }") : C_scope. + format "{{ P } } e @ E {{ v , Q } }") : stdpp_scope. Notation "{{ P } } e {{ v , Q } }" := (ht ⊤ P%I e%E (λ v, Q)%I) (at level 20, P, e, Q at level 200, - format "{{ P } } e {{ v , Q } }") : C_scope. + format "{{ P } } e {{ v , Q } }") : stdpp_scope. Section hoare. Context `{irisG Λ Σ}. diff --git a/theories/program_logic/weakestpre.v b/theories/program_logic/weakestpre.v index 08a7649d86d256e976d3ba1fe951b4a32d57f970..3f194d30dbda11e515cd48994fcb2e72c9b848ce 100644 --- a/theories/program_logic/weakestpre.v +++ b/theories/program_logic/weakestpre.v @@ -75,20 +75,20 @@ Notation "'{{{' P } } } e @ E {{{ x .. y , 'RET' pat ; Q } } }" := (∀ Φ : _ → uPred _, P -∗ â–· (∀ x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e @ E {{ Φ }}) (at level 20, x closed binder, y closed binder, - format "{{{ P } } } e @ E {{{ x .. y , RET pat ; Q } } }") : C_scope. + format "{{{ P } } } e @ E {{{ x .. y , RET pat ; Q } } }") : stdpp_scope. Notation "'{{{' P } } } e {{{ x .. y , 'RET' pat ; Q } } }" := (∀ Φ : _ → uPred _, P -∗ â–· (∀ x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e {{ Φ }}) (at level 20, x closed binder, y closed binder, - format "{{{ P } } } e {{{ x .. y , RET pat ; Q } } }") : C_scope. + format "{{{ P } } } e {{{ x .. y , RET pat ; Q } } }") : stdpp_scope. Notation "'{{{' P } } } e @ E {{{ 'RET' pat ; Q } } }" := (∀ Φ : _ → uPred _, P -∗ â–· (Q -∗ Φ pat%V) -∗ WP e @ E {{ Φ }}) (at level 20, - format "{{{ P } } } e @ E {{{ RET pat ; Q } } }") : C_scope. + format "{{{ P } } } e @ E {{{ RET pat ; Q } } }") : stdpp_scope. Notation "'{{{' P } } } e {{{ 'RET' pat ; Q } } }" := (∀ Φ : _ → uPred _, P -∗ â–· (Q -∗ Φ pat%V) -∗ WP e {{ Φ }}) (at level 20, - format "{{{ P } } } e {{{ RET pat ; Q } } }") : C_scope. + format "{{{ P } } } e {{{ RET pat ; Q } } }") : stdpp_scope. Section wp. Context `{irisG Λ Σ}. diff --git a/theories/proofmode/notation.v b/theories/proofmode/notation.v index 9ee7c2e7fefc17516a29c8ebc5f7af5a0547e1bd..01d2c37b7086f9ad898c1be8337d551dda096758 100644 --- a/theories/proofmode/notation.v +++ b/theories/proofmode/notation.v @@ -16,14 +16,14 @@ Notation "Γ '--------------------------------------' â–¡ Δ '------------------ (envs_entails (Envs Γ Δ) Q%I) (at level 1, Q at level 200, left associativity, format "Γ '--------------------------------------' â–¡ '//' Δ '--------------------------------------' ∗ '//' Q '//'", only printing) : - C_scope. + stdpp_scope. Notation "Δ '--------------------------------------' ∗ Q" := (envs_entails (Envs Enil Δ) Q%I) (at level 1, Q at level 200, left associativity, - format "Δ '--------------------------------------' ∗ '//' Q '//'", only printing) : C_scope. + format "Δ '--------------------------------------' ∗ '//' Q '//'", only printing) : stdpp_scope. Notation "Γ '--------------------------------------' â–¡ Q" := (envs_entails (Envs Γ Enil) Q%I) (at level 1, Q at level 200, left associativity, - format "Γ '--------------------------------------' â–¡ '//' Q '//'", only printing) : C_scope. + format "Γ '--------------------------------------' â–¡ '//' Q '//'", only printing) : stdpp_scope. Notation "'--------------------------------------' ∗ Q" := (envs_entails (Envs Enil Enil) Q%I) - (at level 1, Q at level 200, format "'--------------------------------------' ∗ '//' Q '//'", only printing) : C_scope. + (at level 1, Q at level 200, format "'--------------------------------------' ∗ '//' Q '//'", only printing) : stdpp_scope.