diff --git a/CHANGELOG.md b/CHANGELOG.md index 4fba95af8d295e6e5ab5da1d3fb693f96f89c32e..0dbe9daefc2eec3db16fd98aebd88952b54cf258 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -143,6 +143,7 @@ Changes in Coq: * Add the camera `ufrac` for unbounded fractions (i.e. without fractions that can be `> 1`) and the camera `ufrac_auth` for a variant of the authoritative fractional camera (`frac_auth`) with unbounded fractions. +* Changed `frac_auth` notation from `â—!`/`â—¯!` to `â—F`/`â—¯F`. ## Iris 3.1.0 (released 2017-12-19) diff --git a/theories/algebra/frac_auth.v b/theories/algebra/frac_auth.v index 8cae1cb777f887e22e5e049359c4567b770864fd..3f7d27c3cdcdd5066af33f9240ec3abafcfe715f 100644 --- a/theories/algebra/frac_auth.v +++ b/theories/algebra/frac_auth.v @@ -24,9 +24,9 @@ Typeclasses Opaque frac_auth_auth frac_auth_frag. Instance: Params (@frac_auth_auth) 1 := {}. Instance: Params (@frac_auth_frag) 2 := {}. -Notation "â—! a" := (frac_auth_auth a) (at level 10). -Notation "â—¯!{ q } a" := (frac_auth_frag q a) (at level 10, format "â—¯!{ q } a"). -Notation "â—¯! a" := (frac_auth_frag 1 a) (at level 10). +Notation "â—F a" := (frac_auth_auth a) (at level 10). +Notation "â—¯F{ q } a" := (frac_auth_frag q a) (at level 10, format "â—¯F{ q } a"). +Notation "â—¯F a" := (frac_auth_frag 1 a) (at level 10). Section frac_auth. Context {A : cmraT}. @@ -41,79 +41,79 @@ Section frac_auth. Global Instance frac_auth_frag_proper q : Proper ((≡) ==> (≡)) (@frac_auth_frag A q). Proof. solve_proper. Qed. - Global Instance frac_auth_auth_discrete a : Discrete a → Discrete (â—! a). + Global Instance frac_auth_auth_discrete a : Discrete a → Discrete (â—F a). Proof. intros; apply auth_auth_discrete; [apply Some_discrete|]; apply _. Qed. - Global Instance frac_auth_frag_discrete q a : Discrete a → Discrete (â—¯!{q} a). + Global Instance frac_auth_frag_discrete q a : Discrete a → Discrete (â—¯F{q} a). Proof. intros; apply auth_frag_discrete, Some_discrete; apply _. Qed. - Lemma frac_auth_validN n a : ✓{n} a → ✓{n} (â—! a â‹… â—¯! a). + Lemma frac_auth_validN n a : ✓{n} a → ✓{n} (â—F a â‹… â—¯F a). Proof. by rewrite auth_both_validN. Qed. - Lemma frac_auth_valid a : ✓ a → ✓ (â—! a â‹… â—¯! a). + Lemma frac_auth_valid a : ✓ a → ✓ (â—F a â‹… â—¯F a). Proof. intros. by apply auth_both_valid_2. Qed. - Lemma frac_auth_agreeN n a b : ✓{n} (â—! a â‹… â—¯! b) → a ≡{n}≡ b. + Lemma frac_auth_agreeN n a b : ✓{n} (â—F a â‹… â—¯F b) → a ≡{n}≡ b. Proof. rewrite auth_both_validN /= => -[Hincl Hvalid]. by move: Hincl=> /Some_includedN_exclusive /(_ Hvalid ) [??]. Qed. - Lemma frac_auth_agree a b : ✓ (â—! a â‹… â—¯! b) → a ≡ b. + Lemma frac_auth_agree a b : ✓ (â—F a â‹… â—¯F b) → a ≡ b. Proof. intros. apply equiv_dist=> n. by apply frac_auth_agreeN, cmra_valid_validN. Qed. - Lemma frac_auth_agreeL `{!LeibnizEquiv A} a b : ✓ (â—! a â‹… â—¯! b) → a = b. + Lemma frac_auth_agreeL `{!LeibnizEquiv A} a b : ✓ (â—F a â‹… â—¯F b) → a = b. Proof. intros. by apply leibniz_equiv, frac_auth_agree. Qed. - Lemma frac_auth_includedN n q a b : ✓{n} (â—! a â‹… â—¯!{q} b) → Some b ≼{n} Some a. + Lemma frac_auth_includedN n q a b : ✓{n} (â—F a â‹… â—¯F{q} b) → Some b ≼{n} Some a. Proof. by rewrite auth_both_validN /= => -[/Some_pair_includedN [_ ?] _]. Qed. Lemma frac_auth_included `{CmraDiscrete A} q a b : - ✓ (â—! a â‹… â—¯!{q} b) → Some b ≼ Some a. + ✓ (â—F a â‹… â—¯F{q} b) → Some b ≼ Some a. Proof. by rewrite auth_both_valid /= => -[/Some_pair_included [_ ?] _]. Qed. Lemma frac_auth_includedN_total `{CmraTotal A} n q a b : - ✓{n} (â—! a â‹… â—¯!{q} b) → b ≼{n} a. + ✓{n} (â—F a â‹… â—¯F{q} b) → b ≼{n} a. Proof. intros. by eapply Some_includedN_total, frac_auth_includedN. Qed. Lemma frac_auth_included_total `{CmraDiscrete A, CmraTotal A} q a b : - ✓ (â—! a â‹… â—¯!{q} b) → b ≼ a. + ✓ (â—F a â‹… â—¯F{q} b) → b ≼ a. Proof. intros. by eapply Some_included_total, frac_auth_included. Qed. - Lemma frac_auth_auth_validN n a : ✓{n} (â—! a) ↔ ✓{n} a. + Lemma frac_auth_auth_validN n a : ✓{n} (â—F a) ↔ ✓{n} a. Proof. rewrite auth_auth_frac_validN Some_validN. split. by intros [?[]]. intros. by split. Qed. - Lemma frac_auth_auth_valid a : ✓ (â—! a) ↔ ✓ a. + Lemma frac_auth_auth_valid a : ✓ (â—F a) ↔ ✓ a. Proof. rewrite !cmra_valid_validN. by setoid_rewrite frac_auth_auth_validN. Qed. - Lemma frac_auth_frag_validN n q a : ✓{n} (â—¯!{q} a) ↔ ✓{n} q ∧ ✓{n} a. + Lemma frac_auth_frag_validN n q a : ✓{n} (â—¯F{q} a) ↔ ✓{n} q ∧ ✓{n} a. Proof. done. Qed. - Lemma frac_auth_frag_valid q a : ✓ (â—¯!{q} a) ↔ ✓ q ∧ ✓ a. + Lemma frac_auth_frag_valid q a : ✓ (â—¯F{q} a) ↔ ✓ q ∧ ✓ a. Proof. done. Qed. - Lemma frac_auth_frag_op q1 q2 a1 a2 : â—¯!{q1+q2} (a1 â‹… a2) ≡ â—¯!{q1} a1 â‹… â—¯!{q2} a2. + Lemma frac_auth_frag_op q1 q2 a1 a2 : â—¯F{q1+q2} (a1 â‹… a2) ≡ â—¯F{q1} a1 â‹… â—¯F{q2} a2. Proof. done. Qed. - Lemma frac_auth_frag_validN_op_1_l n q a b : ✓{n} (â—¯!{1} a â‹… â—¯!{q} b) → False. + Lemma frac_auth_frag_validN_op_1_l n q a b : ✓{n} (â—¯F{1} a â‹… â—¯F{q} b) → False. Proof. rewrite -frac_auth_frag_op frac_auth_frag_validN=> -[/exclusiveN_l []]. Qed. - Lemma frac_auth_frag_valid_op_1_l q a b : ✓ (â—¯!{1} a â‹… â—¯!{q} b) → False. + Lemma frac_auth_frag_valid_op_1_l q a b : ✓ (â—¯F{1} a â‹… â—¯F{q} b) → False. Proof. rewrite -frac_auth_frag_op frac_auth_frag_valid=> -[/exclusive_l []]. Qed. Global Instance is_op_frac_auth (q q1 q2 : frac) (a a1 a2 : A) : - IsOp q q1 q2 → IsOp a a1 a2 → IsOp' (â—¯!{q} a) (â—¯!{q1} a1) (â—¯!{q2} a2). + IsOp q q1 q2 → IsOp a a1 a2 → IsOp' (â—¯F{q} a) (â—¯F{q1} a1) (â—¯F{q2} a2). Proof. by rewrite /IsOp' /IsOp=> /leibniz_equiv_iff -> ->. Qed. Global Instance is_op_frac_auth_core_id (q q1 q2 : frac) (a : A) : - CoreId a → IsOp q q1 q2 → IsOp' (â—¯!{q} a) (â—¯!{q1} a) (â—¯!{q2} a). + CoreId a → IsOp q q1 q2 → IsOp' (â—¯F{q} a) (â—¯F{q1} a) (â—¯F{q2} a). Proof. rewrite /IsOp' /IsOp=> ? /leibniz_equiv_iff ->. by rewrite -frac_auth_frag_op -core_id_dup. Qed. Lemma frac_auth_update q a b a' b' : - (a,b) ~l~> (a',b') → â—! a â‹… â—¯!{q} b ~~> â—! a' â‹… â—¯!{q} b'. + (a,b) ~l~> (a',b') → â—F a â‹… â—¯F{q} b ~~> â—F a' â‹… â—¯F{q} b'. Proof. intros. by apply auth_update, option_local_update, prod_local_update_2. Qed. - Lemma frac_auth_update_1 a b a' : ✓ a' → â—! a â‹… â—¯! b ~~> â—! a' â‹… â—¯! a'. + Lemma frac_auth_update_1 a b a' : ✓ a' → â—F a â‹… â—¯F b ~~> â—F a' â‹… â—¯F a'. Proof. intros. by apply auth_update, option_local_update, exclusive_local_update. Qed. diff --git a/theories/heap_lang/lib/counter.v b/theories/heap_lang/lib/counter.v index 3d46fbd6e5d1a2457353e6c340eb03feaebd9c3a..221a89f84769247d351f2b88be3096e86bdc80a7 100644 --- a/theories/heap_lang/lib/counter.v +++ b/theories/heap_lang/lib/counter.v @@ -97,13 +97,13 @@ Section contrib_spec. Context `{!heapG Σ, !ccounterG Σ} (N : namespace). Definition ccounter_inv (γ : gname) (l : loc) : iProp Σ := - (∃ n, own γ (â—! n) ∗ l ↦ #n)%I. + (∃ n, own γ (â—F n) ∗ l ↦ #n)%I. Definition ccounter_ctx (γ : gname) (l : loc) : iProp Σ := inv N (ccounter_inv γ l). Definition ccounter (γ : gname) (q : frac) (n : nat) : iProp Σ := - own γ (â—¯!{q} n). + own γ (â—¯F{q} n). (** The main proofs. *) Lemma ccounter_op γ q1 q2 n1 n2 : @@ -115,7 +115,7 @@ Section contrib_spec. {{{ γ l, RET #l; ccounter_ctx γ l ∗ ccounter γ 1 0 }}}. Proof. iIntros (Φ) "_ HΦ". rewrite -wp_fupd /newcounter /=. wp_lam. wp_alloc l as "Hl". - iMod (own_alloc (â—! O%nat â‹… â—¯! 0%nat)) as (γ) "[Hγ Hγ']"; + iMod (own_alloc (â—F O%nat â‹… â—¯F 0%nat)) as (γ) "[Hγ Hγ']"; first by apply auth_both_valid. iMod (inv_alloc N _ (ccounter_inv γ l) with "[Hl Hγ]"). { iNext. iExists 0%nat. by iFrame. }