Skip to content
Snippets Groups Projects
Commit 5a623192 authored by Ralf Jung's avatar Ralf Jung
Browse files

Use Coq #[deprecated] attribute

parent fea0c2de
No related branches found
No related tags found
No related merge requests found
From iris.algebra Require Import ofe cmra. From iris.algebra Require Import ofe cmra.
Set Default Proof Using "Type". Set Default Proof Using "Type".
Local Set Warnings "-deprecated".
(* Old notation for backwards compatibility. *) (* Old notation for backwards compatibility. *)
(* Deprecated 2016-11-22. Use ofeT instead. *) (* Deprecated 2016-11-22. Use ofeT instead. *)
#[deprecated(note = "Use ofeT instead.")]
Notation cofeT := ofeT (only parsing). Notation cofeT := ofeT (only parsing).
(* Deprecated 2016-12-09. Use agree instead. *) (* Deprecated 2016-12-09. Use agree instead. *)
...@@ -14,13 +17,15 @@ Local Arguments op _ _ _ !_ /. ...@@ -14,13 +17,15 @@ Local Arguments op _ _ _ !_ /.
Local Arguments pcore _ _ !_ /. Local Arguments pcore _ _ !_ /.
(* This is isomorphic to option, but has a very different RA structure. *) (* This is isomorphic to option, but has a very different RA structure. *)
Inductive dec_agree (A : Type) : Type := Inductive dec_agree_ (A : Type) : Type :=
| DecAgree : A dec_agree A | DecAgree : A dec_agree_ A
| DecAgreeBot : dec_agree A. | DecAgreeBot : dec_agree_ A.
Arguments DecAgree {_} _. Arguments DecAgree {_} _.
Arguments DecAgreeBot {_}. Arguments DecAgreeBot {_}.
Instance maybe_DecAgree {A} : Maybe (@DecAgree A) := λ x, Instance maybe_DecAgree {A} : Maybe (@DecAgree A) := λ x,
match x with DecAgree a => Some a | _ => None end. match x with DecAgree a => Some a | _ => None end.
#[deprecated(note = "Use agree instead.")]
Notation dec_agree := dec_agree_.
Section dec_agree. Section dec_agree.
Context `{EqDecision A}. Context `{EqDecision A}.
...@@ -38,6 +43,7 @@ Instance dec_agree_op : Op (dec_agree A) := λ x y, ...@@ -38,6 +43,7 @@ Instance dec_agree_op : Op (dec_agree A) := λ x y,
end. end.
Instance dec_agree_pcore : PCore (dec_agree A) := Some. Instance dec_agree_pcore : PCore (dec_agree A) := Some.
#[deprecated(note = "Use agree instead.")]
Definition dec_agree_ra_mixin : RAMixin (dec_agree A). Definition dec_agree_ra_mixin : RAMixin (dec_agree A).
Proof. Proof.
apply ra_total_mixin; apply _ || eauto. apply ra_total_mixin; apply _ || eauto.
...@@ -47,6 +53,7 @@ Proof. ...@@ -47,6 +53,7 @@ Proof.
- by intros [?|] [?|] ?. - by intros [?|] [?|] ?.
Qed. Qed.
#[deprecated(note = "Use agree instead.")]
Canonical Structure dec_agreeR : cmraT := Canonical Structure dec_agreeR : cmraT :=
discreteR (dec_agree A) dec_agree_ra_mixin. discreteR (dec_agree A) dec_agree_ra_mixin.
...@@ -59,15 +66,19 @@ Proof. intros x. by exists x. Qed. ...@@ -59,15 +66,19 @@ Proof. intros x. by exists x. Qed.
Global Instance dec_agree_core_id (x : dec_agreeR) : CoreId x. Global Instance dec_agree_core_id (x : dec_agreeR) : CoreId x.
Proof. by constructor. Qed. Proof. by constructor. Qed.
#[deprecated(note = "Use agree instead.")]
Lemma dec_agree_ne a b : a b DecAgree a DecAgree b = DecAgreeBot. Lemma dec_agree_ne a b : a b DecAgree a DecAgree b = DecAgreeBot.
Proof. intros. by rewrite /= decide_False. Qed. Proof. intros. by rewrite /= decide_False. Qed.
#[deprecated(note = "Use agree instead.")]
Lemma dec_agree_idemp (x : dec_agree A) : x x = x. Lemma dec_agree_idemp (x : dec_agree A) : x x = x.
Proof. destruct x; by rewrite /= ?decide_True. Qed. Proof. destruct x; by rewrite /= ?decide_True. Qed.
#[deprecated(note = "Use agree instead.")]
Lemma dec_agree_op_inv (x1 x2 : dec_agree A) : (x1 x2) x1 = x2. Lemma dec_agree_op_inv (x1 x2 : dec_agree A) : (x1 x2) x1 = x2.
Proof. destruct x1, x2; by repeat (simplify_eq/= || case_match). Qed. Proof. destruct x1, x2; by repeat (simplify_eq/= || case_match). Qed.
#[deprecated(note = "Use agree instead.")]
Lemma DecAgree_included a b : DecAgree a DecAgree b a = b. Lemma DecAgree_included a b : DecAgree a DecAgree b a = b.
Proof. Proof.
split. intros [[c|] [=]%leibniz_equiv]. by simplify_option_eq. by intros ->. split. intros [[c|] [=]%leibniz_equiv]. by simplify_option_eq. by intros ->.
......
...@@ -1368,6 +1368,7 @@ Tactic Notation "iModIntro" uconstr(sel) := ...@@ -1368,6 +1368,7 @@ Tactic Notation "iModIntro" uconstr(sel) :=
Tactic Notation "iModIntro" := iModIntro _. Tactic Notation "iModIntro" := iModIntro _.
(** DEPRECATED *) (** DEPRECATED *)
#[deprecated(note = "Use iModIntro instead")]
Tactic Notation "iAlways" := iModIntro. Tactic Notation "iAlways" := iModIntro.
(** * Later *) (** * Later *)
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment