diff --git a/algebra/upred.v b/algebra/upred.v
index f513350d5a264892712f98a909d05647067a4de5..1958f06d7f6a1cf0dd2b7085ff2d4ca837846049 100644
--- a/algebra/upred.v
+++ b/algebra/upred.v
@@ -123,15 +123,15 @@ Instance uPred_entails_rewrite_relation M : RewriteRelation (@uPred_entails M).
 Hint Resolve uPred_mono uPred_closed : uPred_def.
 
 (** logical connectives *)
-Program Definition uPred_const_def {M} (φ : Prop) : uPred M :=
+Program Definition uPred_pure_def {M} (φ : Prop) : uPred M :=
   {| uPred_holds n x := φ |}.
 Solve Obligations with done.
-Definition uPred_const_aux : { x | x = @uPred_const_def }. by eexists. Qed.
-Definition uPred_const {M} := proj1_sig uPred_const_aux M.
-Definition uPred_const_eq :
-  @uPred_const = @uPred_const_def := proj2_sig uPred_const_aux.
+Definition uPred_pure_aux : { x | x = @uPred_pure_def }. by eexists. Qed.
+Definition uPred_pure {M} := proj1_sig uPred_pure_aux M.
+Definition uPred_pure_eq :
+  @uPred_pure = @uPred_pure_def := proj2_sig uPred_pure_aux.
 
-Instance uPred_inhabited M : Inhabited (uPred M) := populate (uPred_const True).
+Instance uPred_inhabited M : Inhabited (uPred M) := populate (uPred_pure True).
 
 Program Definition uPred_and_def {M} (P Q : uPred M) : uPred M :=
   {| uPred_holds n x := P n x ∧ Q n x |}.
@@ -263,12 +263,12 @@ Notation "(⊢)" := uPred_entails (only parsing) : C_scope.
 Notation "P ⊣⊢ Q" := (equiv (A:=uPred _) P%I Q%I)
   (at level 95, no associativity) : C_scope.
 Notation "(⊣⊢)" := (equiv (A:=uPred _)) (only parsing) : C_scope.
-Notation "■ φ" := (uPred_const φ%C%type)
+Notation "■ φ" := (uPred_pure φ%C%type)
   (at level 20, right associativity) : uPred_scope.
-Notation "x = y" := (uPred_const (x%C%type = y%C%type)) : uPred_scope.
-Notation "x ⊥ y" := (uPred_const (x%C%type ⊥ y%C%type)) : uPred_scope.
-Notation "'False'" := (uPred_const False) : uPred_scope.
-Notation "'True'" := (uPred_const True) : uPred_scope.
+Notation "x = y" := (uPred_pure (x%C%type = y%C%type)) : uPred_scope.
+Notation "x ⊥ y" := (uPred_pure (x%C%type ⊥ y%C%type)) : uPred_scope.
+Notation "'False'" := (uPred_pure False) : uPred_scope.
+Notation "'True'" := (uPred_pure True) : uPred_scope.
 Infix "∧" := uPred_and : uPred_scope.
 Notation "(∧)" := uPred_and (only parsing) : uPred_scope.
 Infix "∨" := uPred_or : uPred_scope.
@@ -308,7 +308,7 @@ Arguments persistentP {_} _ {_}.
 
 Module uPred.
 Definition unseal :=
-  (uPred_const_eq, uPred_and_eq, uPred_or_eq, uPred_impl_eq, uPred_forall_eq,
+  (uPred_pure_eq, uPred_and_eq, uPred_or_eq, uPred_impl_eq, uPred_forall_eq,
   uPred_exist_eq, uPred_eq_eq, uPred_sep_eq, uPred_wand_eq, uPred_always_eq,
   uPred_later_eq, uPred_ownM_eq, uPred_valid_eq).
 Ltac unseal := rewrite !unseal /=.
@@ -353,7 +353,7 @@ Lemma entails_equiv_r (P Q R : uPred M) : (P ⊢ Q) → (Q ⊣⊢ R) → (P ⊢
 Proof. by intros ? <-. Qed.
 
 (** Non-expansiveness and setoid morphisms *)
-Global Instance const_proper : Proper (iff ==> (⊣⊢)) (@uPred_const M).
+Global Instance pure_proper : Proper (iff ==> (⊣⊢)) (@uPred_pure M).
 Proof. intros φ1 φ2 Hφ. by unseal; split=> -[|n] ?; try apply Hφ. Qed.
 Global Instance and_ne n : Proper (dist n ==> dist n ==> dist n) (@uPred_and M).
 Proof.
@@ -459,9 +459,9 @@ Global Instance iff_proper :
   Proper ((⊣⊢) ==> (⊣⊢) ==> (⊣⊢)) (@uPred_iff M) := ne_proper_2 _.
 
 (** Introduction and elimination rules *)
-Lemma const_intro φ P : φ → P ⊢ ■ φ.
+Lemma pure_intro φ P : φ → P ⊢ ■ φ.
 Proof. by intros ?; unseal; split. Qed.
-Lemma const_elim φ Q R : (Q ⊢ ■ φ) → (φ → Q ⊢ R) → Q ⊢ R.
+Lemma pure_elim φ Q R : (Q ⊢ ■ φ) → (φ → Q ⊢ R) → Q ⊢ R.
 Proof.
   unseal; intros HQP HQR; split=> n x ??; apply HQR; first eapply HQP; eauto.
 Qed.
@@ -517,9 +517,9 @@ Qed.
 
 (* Derived logical stuff *)
 Lemma False_elim P : False ⊢ P.
-Proof. by apply (const_elim False). Qed.
+Proof. by apply (pure_elim False). Qed.
 Lemma True_intro P : P ⊢ True.
-Proof. by apply const_intro. Qed.
+Proof. by apply pure_intro. Qed.
 Lemma and_elim_l' P Q R : (P ⊢ R) → P ∧ Q ⊢ R.
 Proof. by rewrite and_elim_l. Qed.
 Lemma and_elim_r' P Q R : (Q ⊢ R) → P ∧ Q ⊢ R.
@@ -562,8 +562,8 @@ Qed.
 Lemma equiv_iff P Q : (P ⊣⊢ Q) → True ⊢ P ↔ Q.
 Proof. intros ->; apply iff_refl. Qed.
 
-Lemma const_mono φ1 φ2 : (φ1 → φ2) → ■ φ1 ⊢ ■ φ2.
-Proof. intros; apply const_elim with φ1; eauto using const_intro. Qed.
+Lemma pure_mono φ1 φ2 : (φ1 → φ2) → ■ φ1 ⊢ ■ φ2.
+Proof. intros; apply pure_elim with φ1; eauto using pure_intro. Qed.
 Lemma and_mono P P' Q Q' : (P ⊢ Q) → (P' ⊢ Q') → P ∧ P' ⊢ Q ∧ Q'.
 Proof. auto. Qed.
 Lemma and_mono_l P P' Q : (P ⊢ Q) → P ∧ P' ⊢ Q ∧ P'.
@@ -589,8 +589,8 @@ Qed.
 Lemma exist_mono {A} (Φ Ψ : A → uPred M) :
   (∀ a, Φ a ⊢ Ψ a) → (∃ a, Φ a) ⊢ ∃ a, Ψ a.
 Proof. intros HΦ. apply exist_elim=> a; rewrite (HΦ a); apply exist_intro. Qed.
-Global Instance const_mono' : Proper (impl ==> (⊢)) (@uPred_const M).
-Proof. intros φ1 φ2; apply const_mono. Qed.
+Global Instance pure_mono' : Proper (impl ==> (⊢)) (@uPred_pure M).
+Proof. intros φ1 φ2; apply pure_mono. Qed.
 Global Instance and_mono' : Proper ((⊢) ==> (⊢) ==> (⊢)) (@uPred_and M).
 Proof. by intros P P' HP Q Q' HQ; apply and_mono. Qed.
 Global Instance and_flip_mono' :
@@ -673,18 +673,18 @@ Proof.
   rewrite -(comm _ P) and_exist_l. apply exist_proper=>a. by rewrite comm.
 Qed.
 
-Lemma const_intro_l φ Q R : φ → (■ φ ∧ Q ⊢ R) → Q ⊢ R.
-Proof. intros ? <-; auto using const_intro. Qed.
-Lemma const_intro_r φ Q R : φ → (Q ∧ ■ φ ⊢ R) → Q ⊢ R.
-Proof. intros ? <-; auto using const_intro. Qed.
-Lemma const_intro_impl φ Q R : φ → (Q ⊢ ■ φ → R) → Q ⊢ R.
-Proof. intros ? ->. eauto using const_intro_l, impl_elim_r. Qed.
-Lemma const_elim_l φ Q R : (φ → Q ⊢ R) → ■ φ ∧ Q ⊢ R.
-Proof. intros; apply const_elim with φ; eauto. Qed.
-Lemma const_elim_r φ Q R : (φ → Q ⊢ R) → Q ∧ ■ φ ⊢ R.
-Proof. intros; apply const_elim with φ; eauto. Qed.
-Lemma const_equiv (φ : Prop) : φ → ■ φ ⊣⊢ True.
-Proof. intros; apply (anti_symm _); auto using const_intro. Qed.
+Lemma pure_intro_l φ Q R : φ → (■ φ ∧ Q ⊢ R) → Q ⊢ R.
+Proof. intros ? <-; auto using pure_intro. Qed.
+Lemma pure_intro_r φ Q R : φ → (Q ∧ ■ φ ⊢ R) → Q ⊢ R.
+Proof. intros ? <-; auto using pure_intro. Qed.
+Lemma pure_intro_impl φ Q R : φ → (Q ⊢ ■ φ → R) → Q ⊢ R.
+Proof. intros ? ->. eauto using pure_intro_l, impl_elim_r. Qed.
+Lemma pure_elim_l φ Q R : (φ → Q ⊢ R) → ■ φ ∧ Q ⊢ R.
+Proof. intros; apply pure_elim with φ; eauto. Qed.
+Lemma pure_elim_r φ Q R : (φ → Q ⊢ R) → Q ∧ ■ φ ⊢ R.
+Proof. intros; apply pure_elim with φ; eauto. Qed.
+Lemma pure_equiv (φ : Prop) : φ → ■ φ ⊣⊢ True.
+Proof. intros; apply (anti_symm _); auto using pure_intro. Qed.
 
 Lemma eq_refl' {A : cofeT} (a : A) P : P ⊢ a ≡ a.
 Proof. rewrite (True_intro P). apply eq_refl. Qed.
@@ -822,10 +822,10 @@ Lemma sep_and P Q : (P ★ Q) ⊢ (P ∧ Q).
 Proof. auto. Qed.
 Lemma impl_wand P Q : (P → Q) ⊢ P -★ Q.
 Proof. apply wand_intro_r, impl_elim with P; auto. Qed.
-Lemma const_elim_sep_l φ Q R : (φ → Q ⊢ R) → ■ φ ★ Q ⊢ R.
-Proof. intros; apply const_elim with φ; eauto. Qed.
-Lemma const_elim_sep_r φ Q R : (φ → Q ⊢ R) → Q ★ ■ φ ⊢ R.
-Proof. intros; apply const_elim with φ; eauto. Qed.
+Lemma pure_elim_sep_l φ Q R : (φ → Q ⊢ R) → ■ φ ★ Q ⊢ R.
+Proof. intros; apply pure_elim with φ; eauto. Qed.
+Lemma pure_elim_sep_r φ Q R : (φ → Q ⊢ R) → Q ★ ■ φ ⊢ R.
+Proof. intros; apply pure_elim with φ; eauto. Qed.
 
 Global Instance sep_False : LeftAbsorb (⊣⊢) False%I (@uPred_sep M).
 Proof. intros P; apply (anti_symm _); auto. Qed.
@@ -858,7 +858,7 @@ Lemma sep_forall_r {A} (Φ : A → uPred M) Q : (∀ a, Φ a) ★ Q ⊢ ∀ a, 
 Proof. by apply forall_intro=> a; rewrite forall_elim. Qed.
 
 (* Always *)
-Lemma always_const φ : □ ■ φ ⊣⊢ ■ φ.
+Lemma always_pure φ : □ ■ φ ⊣⊢ ■ φ.
 Proof. by unseal. Qed.
 Lemma always_elim P : □ P ⊢ P.
 Proof.
@@ -910,7 +910,7 @@ Proof.
   apply (anti_symm (⊢)); auto using always_elim.
   apply (eq_rewrite a b (λ b, □ (a ≡ b))%I); auto.
   { intros n; solve_proper. }
-  rewrite -(eq_refl a) always_const; auto.
+  rewrite -(eq_refl a) always_pure; auto.
 Qed.
 Lemma always_and_sep P Q : □ (P ∧ Q) ⊣⊢ □ (P ★ Q).
 Proof. apply (anti_symm (⊢)); auto using always_and_sep_1. Qed.
@@ -1098,7 +1098,7 @@ Proof.
     apply HP, uPred_closed with n; eauto using cmra_validN_le.
 Qed.
 
-Global Instance const_timeless φ : TimelessP (■ φ : uPred M)%I.
+Global Instance pure_timeless φ : TimelessP (■ φ : uPred M)%I.
 Proof. by apply timelessP_spec; unseal => -[|n] x. Qed.
 Global Instance valid_timeless {A : cmraT} `{CMRADiscrete A} (a : A) :
    TimelessP (✓ a : uPred M)%I.
@@ -1141,7 +1141,7 @@ Qed.
 Global Instance always_timeless P : TimelessP P → TimelessP (□ P).
 Proof.
   intros ?; rewrite /TimelessP.
-  by rewrite -always_const -!always_later -always_or; apply always_mono.
+  by rewrite -always_pure -!always_later -always_or; apply always_mono.
 Qed.
 Global Instance always_if_timeless p P : TimelessP P → TimelessP (□?p P).
 Proof. destruct p; apply _. Qed.
@@ -1157,8 +1157,8 @@ Proof.
 Qed.
 
 (* Persistence *)
-Global Instance const_persistent φ : PersistentP (■ φ : uPred M)%I.
-Proof. by rewrite /PersistentP always_const. Qed.
+Global Instance pure_persistent φ : PersistentP (■ φ : uPred M)%I.
+Proof. by rewrite /PersistentP always_pure. Qed.
 Global Instance always_persistent P : PersistentP (â–¡ P).
 Proof. by intros; apply always_intro'. Qed.
 Global Instance and_persistent P Q :
@@ -1210,7 +1210,7 @@ Proof. by rewrite -(always_always Q); apply always_entails_r'. Qed.
 End uPred_logic.
 
 (* Hint DB for the logic *)
-Hint Resolve const_intro.
+Hint Resolve pure_intro.
 Hint Resolve or_elim or_intro_l' or_intro_r' : I.
 Hint Resolve and_intro and_elim_l' and_elim_r' : I.
 Hint Resolve always_mono : I.
diff --git a/algebra/upred_big_op.v b/algebra/upred_big_op.v
index 3427dd3c516f4a1c75cb41c79dca2ceefd1661cd..ed9b36a3c718076a71cb16c7fdfef5cb212d429c 100644
--- a/algebra/upred_big_op.v
+++ b/algebra/upred_big_op.v
@@ -223,7 +223,7 @@ Section gmap.
     (□ [★ map] k↦x ∈ m, Φ k x) ⊣⊢ ([★ map] k↦x ∈ m, □ Φ k x).
   Proof.
     rewrite /uPred_big_sepM.
-    induction (map_to_list m) as [|[i x] l IH]; csimpl; rewrite ?always_const //.
+    induction (map_to_list m) as [|[i x] l IH]; csimpl; rewrite ?always_pure //.
     by rewrite always_sep IH.
   Qed.
 
@@ -237,14 +237,14 @@ Section gmap.
   Proof.
     intros. apply (anti_symm _).
     { apply forall_intro=> k; apply forall_intro=> x.
-      apply impl_intro_l, const_elim_l=> ?; by apply big_sepM_lookup. }
+      apply impl_intro_l, pure_elim_l=> ?; by apply big_sepM_lookup. }
     rewrite /uPred_big_sepM. setoid_rewrite <-elem_of_map_to_list.
     induction (map_to_list m) as [|[i x] l IH]; csimpl; auto.
     rewrite -always_and_sep_l; apply and_intro.
-    - rewrite (forall_elim i) (forall_elim x) const_equiv; last by left.
+    - rewrite (forall_elim i) (forall_elim x) pure_equiv; last by left.
       by rewrite True_impl.
     - rewrite -IH. apply forall_mono=> k; apply forall_mono=> y.
-      apply impl_intro_l, const_elim_l=> ?. rewrite const_equiv; last by right.
+      apply impl_intro_l, pure_elim_l=> ?. rewrite pure_equiv; last by right.
       by rewrite True_impl.
   Qed.
 
@@ -253,7 +253,7 @@ Section gmap.
     ⊢ [★ map] k↦x ∈ m, Ψ k x.
   Proof.
     rewrite always_and_sep_l. do 2 setoid_rewrite always_forall.
-    setoid_rewrite always_impl; setoid_rewrite always_const.
+    setoid_rewrite always_impl; setoid_rewrite always_pure.
     rewrite -big_sepM_forall -big_sepM_sepM. apply big_sepM_mono; auto=> k x ?.
     by rewrite -always_wand_impl always_elim wand_elim_l.
   Qed.
@@ -345,7 +345,7 @@ Section gset.
   Lemma big_sepS_always Φ X : □ ([★ set] y ∈ X, Φ y) ⊣⊢ ([★ set] y ∈ X, □ Φ y).
   Proof.
     rewrite /uPred_big_sepS.
-    induction (elements X) as [|x l IH]; csimpl; first by rewrite ?always_const.
+    induction (elements X) as [|x l IH]; csimpl; first by rewrite ?always_pure.
     by rewrite always_sep IH.
   Qed.
 
@@ -358,13 +358,13 @@ Section gset.
   Proof.
     intros. apply (anti_symm _).
     { apply forall_intro=> x.
-      apply impl_intro_l, const_elim_l=> ?; by apply big_sepS_elem_of. }
+      apply impl_intro_l, pure_elim_l=> ?; by apply big_sepS_elem_of. }
     rewrite /uPred_big_sepS. setoid_rewrite <-elem_of_elements.
     induction (elements X) as [|x l IH]; csimpl; auto.
     rewrite -always_and_sep_l; apply and_intro.
-    - rewrite (forall_elim x) const_equiv; last by left. by rewrite True_impl.
+    - rewrite (forall_elim x) pure_equiv; last by left. by rewrite True_impl.
     - rewrite -IH. apply forall_mono=> y.
-      apply impl_intro_l, const_elim_l=> ?. rewrite const_equiv; last by right.
+      apply impl_intro_l, pure_elim_l=> ?. rewrite pure_equiv; last by right.
       by rewrite True_impl.
   Qed.
 
@@ -372,7 +372,7 @@ Section gset.
       □ (∀ x, ■ (x ∈ X) → Φ x → Ψ x) ∧ ([★ set] x ∈ X, Φ x) ⊢ [★ set] x ∈ X, Ψ x.
   Proof.
     rewrite always_and_sep_l always_forall.
-    setoid_rewrite always_impl; setoid_rewrite always_const.
+    setoid_rewrite always_impl; setoid_rewrite always_pure.
     rewrite -big_sepS_forall -big_sepS_sepS. apply big_sepS_mono; auto=> x ?.
     by rewrite -always_wand_impl always_elim wand_elim_l.
   Qed.
diff --git a/heap_lang/heap.v b/heap_lang/heap.v
index 66fd446848f132535ae2190245bd2bc16e8b5042..84509a58b6a24859e569ac2169d217f9324ffaa9 100644
--- a/heap_lang/heap.v
+++ b/heap_lang/heap.v
@@ -136,12 +136,12 @@ Section heap.
     l ↦{q1} v1 ★ l ↦{q2} v2 ⊣⊢ v1 = v2 ∧ l ↦{q1+q2} v1.
   Proof.
     destruct (decide (v1 = v2)) as [->|].
-    { by rewrite heap_mapsto_op_eq const_equiv // left_id. }
+    { by rewrite heap_mapsto_op_eq pure_equiv // left_id. }
     rewrite heap_mapsto_eq -auth_own_op op_singleton pair_op dec_agree_ne //.
-    apply (anti_symm (⊢)); last by apply const_elim_l.
+    apply (anti_symm (⊢)); last by apply pure_elim_l.
     rewrite auth_own_valid gmap_validI (forall_elim l) lookup_singleton.
     rewrite option_validI prod_validI frac_validI discrete_valid.
-    by apply const_elim_r.
+    by apply pure_elim_r.
   Qed.
 
   Lemma heap_mapsto_op_split l q v : l ↦{q} v ⊣⊢ (l ↦{q/2} v ★ l ↦{q/2} v).
diff --git a/heap_lang/lib/assert.v b/heap_lang/lib/assert.v
index 0e810633faa0612db856245dbde70045666da57b..236c7aacee4e636614530f0ca3362edd48c37949 100644
--- a/heap_lang/lib/assert.v
+++ b/heap_lang/lib/assert.v
@@ -19,5 +19,5 @@ Lemma wp_assert' {Σ} (Φ : val → iProp heap_lang Σ) e :
   WP e {{ v, v = #true ∧ ▷ Φ #() }} ⊢ WP Assert e {{ Φ }}.
 Proof.
   rewrite /Assert. wp_focus e; apply wp_mono=>v.
-  apply uPred.const_elim_l=>->. apply wp_assert.
+  apply uPred.pure_elim_l=>->. apply wp_assert.
 Qed.
diff --git a/program_logic/ghost_ownership.v b/program_logic/ghost_ownership.v
index 850a6a203bc8031db688d72394d1eb17e8838637..e227bd5ade1f622f7aedd927538dfe060f95ccb3 100644
--- a/program_logic/ghost_ownership.v
+++ b/program_logic/ghost_ownership.v
@@ -54,8 +54,8 @@ Proof.
     eapply pvs_ownG_updateP, (iprod_singleton_updateP_empty (inG_id i));
       first (eapply alloc_updateP_strong', cmra_transport_valid, Ha);
       naive_solver.
-  - apply exist_elim=>m; apply const_elim_l=>-[γ [Hfresh ->]].
-    by rewrite !own_eq /own_def -(exist_intro γ) const_equiv // left_id.
+  - apply exist_elim=>m; apply pure_elim_l=>-[γ [Hfresh ->]].
+    by rewrite !own_eq /own_def -(exist_intro γ) pure_equiv // left_id.
 Qed.
 Lemma own_alloc a E : ✓ a → True ={E}=> ∃ γ, own γ a.
 Proof.
@@ -70,14 +70,14 @@ Proof.
   - eapply pvs_ownG_updateP, iprod_singleton_updateP;
       first by (eapply singleton_updateP', cmra_transport_updateP', Ha).
     naive_solver.
-  - apply exist_elim=>m; apply const_elim_l=>-[a' [-> HP]].
-    rewrite -(exist_intro a'). by apply and_intro; [apply const_intro|].
+  - apply exist_elim=>m; apply pure_elim_l=>-[a' [-> HP]].
+    rewrite -(exist_intro a'). by apply and_intro; [apply pure_intro|].
 Qed.
 
 Lemma own_update γ a a' E : a ~~> a' → own γ a ={E}=> own γ a'.
 Proof.
   intros; rewrite (own_updateP (a' =)); last by apply cmra_update_updateP.
-  by apply pvs_mono, exist_elim=> a''; apply const_elim_l=> ->.
+  by apply pvs_mono, exist_elim=> a''; apply pure_elim_l=> ->.
 Qed.
 End global.
 
diff --git a/program_logic/lifting.v b/program_logic/lifting.v
index 5107d99ec1034f04683282e34c557a39e7f02b27..28a372fdfecd7f3f13e76a5bd0431d96ef45dcbc 100644
--- a/program_logic/lifting.v
+++ b/program_logic/lifting.v
@@ -76,9 +76,9 @@ Proof.
   apply forall_intro=>e2'; apply forall_intro=>σ2'.
   apply forall_intro=>ef; apply wand_intro_l.
   rewrite always_and_sep_l -assoc -always_and_sep_l.
-  apply const_elim_l=>-[[v2 Hv] ?] /=.
+  apply pure_elim_l=>-[[v2 Hv] ?] /=.
   rewrite -pvs_intro.
-  rewrite (forall_elim v2) (forall_elim σ2') (forall_elim ef) const_equiv //.
+  rewrite (forall_elim v2) (forall_elim σ2') (forall_elim ef) pure_equiv //.
   rewrite left_id wand_elim_r -(wp_value _ _ e2' v2) //.
   by erewrite of_to_val.
 Qed.
@@ -96,7 +96,7 @@ Proof.
   apply forall_intro=>e2'; apply forall_intro=>σ2'; apply forall_intro=>ef'.
   apply wand_intro_l.
   rewrite always_and_sep_l -assoc -always_and_sep_l to_of_val.
-  apply const_elim_l=>-[-> [[->] ->]] /=. by rewrite wand_elim_r.
+  apply pure_elim_l=>-[-> [[->] ->]] /=. by rewrite wand_elim_r.
 Qed.
 
 Lemma wp_lift_pure_det_step {E Φ} e1 e2 ef :
@@ -108,6 +108,6 @@ Proof.
   intros.
   rewrite -(wp_lift_pure_step E (λ e2' ef', e2 = e2' ∧ ef = ef') _ e1) //=.
   apply later_mono, forall_intro=>e'; apply forall_intro=>ef'.
-  by apply impl_intro_l, const_elim_l=>-[-> ->].
+  by apply impl_intro_l, pure_elim_l=>-[-> ->].
 Qed.
 End lifting.
diff --git a/program_logic/pviewshifts.v b/program_logic/pviewshifts.v
index 8e50db7ec874c8bcc5aeb5a8b1e9e68663ae0b2a..ea19b2a12b515d0968ebfb9ae271307d1e5b8ced 100644
--- a/program_logic/pviewshifts.v
+++ b/program_logic/pviewshifts.v
@@ -224,7 +224,7 @@ Proof. auto using pvs_mask_frame'. Qed.
 Lemma pvs_ownG_update E m m' : m ~~> m' → ownG m ={E}=> ownG m'.
 Proof.
   intros; rewrite (pvs_ownG_updateP E _ (m' =)); last by apply cmra_update_updateP.
-  by apply pvs_mono, uPred.exist_elim=> m''; apply uPred.const_elim_l=> ->.
+  by apply pvs_mono, uPred.exist_elim=> m''; apply uPred.pure_elim_l=> ->.
 Qed.
 End pvs.
 
diff --git a/proofmode/coq_tactics.v b/proofmode/coq_tactics.v
index 9349e964f72fbfcd7d2994f0d81127c56066a7e8..97a7004b32212b1189f548091699688fc8762f35 100644
--- a/proofmode/coq_tactics.v
+++ b/proofmode/coq_tactics.v
@@ -117,15 +117,15 @@ Qed.
 Lemma envs_lookup_sound Δ i p P :
   envs_lookup i Δ = Some (p,P) → Δ ⊢ □?p P ★ envs_delete i p Δ.
 Proof.
-  rewrite /envs_lookup /envs_delete /of_envs=>?; apply const_elim_sep_l=> Hwf.
+  rewrite /envs_lookup /envs_delete /of_envs=>?; apply pure_elim_sep_l=> Hwf.
   destruct Δ as [Γp Γs], (Γp !! i) eqn:?; simplify_eq/=.
   - rewrite (env_lookup_perm Γp) //= always_and_sep always_sep.
-    ecancel [□ [∧] _; □ P; [★] _]%I; apply const_intro.
+    ecancel [□ [∧] _; □ P; [★] _]%I; apply pure_intro.
     destruct Hwf; constructor;
       naive_solver eauto using env_delete_wf, env_delete_fresh.
   - destruct (Γs !! i) eqn:?; simplify_eq/=.
     rewrite (env_lookup_perm Γs) //=.
-    ecancel [□ [∧] _; P; [★] _]%I; apply const_intro.
+    ecancel [□ [∧] _; P; [★] _]%I; apply pure_intro.
     destruct Hwf; constructor;
       naive_solver eauto using env_delete_wf, env_delete_fresh.
 Qed.
@@ -141,13 +141,13 @@ Qed.
 Lemma envs_lookup_split Δ i p P :
   envs_lookup i Δ = Some (p,P) → Δ ⊢ □?p P ★ (□?p P -★ Δ).
 Proof.
-  rewrite /envs_lookup /of_envs=>?; apply const_elim_sep_l=> Hwf.
+  rewrite /envs_lookup /of_envs=>?; apply pure_elim_sep_l=> Hwf.
   destruct Δ as [Γp Γs], (Γp !! i) eqn:?; simplify_eq/=.
   - rewrite (env_lookup_perm Γp) //= always_and_sep always_sep.
-    rewrite const_equiv // left_id.
+    rewrite pure_equiv // left_id.
     cancel [â–¡ P]%I. apply wand_intro_l. solve_sep_entails.
   - destruct (Γs !! i) eqn:?; simplify_eq/=.
-    rewrite (env_lookup_perm Γs) //=. rewrite const_equiv // left_id.
+    rewrite (env_lookup_perm Γs) //=. rewrite pure_equiv // left_id.
     cancel [P]. apply wand_intro_l. solve_sep_entails.
 Qed.
 
@@ -160,11 +160,11 @@ Proof. intros [? ->]%envs_lookup_delete_Some. by apply envs_lookup_sound'. Qed.
 
 Lemma envs_app_sound Δ Δ' p Γ : envs_app p Γ Δ = Some Δ' → Δ ⊢ □?p [★] Γ -★ Δ'.
 Proof.
-  rewrite /of_envs /envs_app=> ?; apply const_elim_sep_l=> Hwf.
+  rewrite /of_envs /envs_app=> ?; apply pure_elim_sep_l=> Hwf.
   destruct Δ as [Γp Γs], p; simplify_eq/=.
   - destruct (env_app Γ Γs) eqn:Happ,
       (env_app Γ Γp) as [Γp'|] eqn:?; simplify_eq/=.
-    apply wand_intro_l, sep_intro_True_l; [apply const_intro|].
+    apply wand_intro_l, sep_intro_True_l; [apply pure_intro|].
     + destruct Hwf; constructor; simpl; eauto using env_app_wf.
       intros j. apply (env_app_disjoint _ _ _ j) in Happ.
       naive_solver eauto using env_app_fresh.
@@ -173,7 +173,7 @@ Proof.
       solve_sep_entails.
   - destruct (env_app Γ Γp) eqn:Happ,
       (env_app Γ Γs) as [Γs'|] eqn:?; simplify_eq/=.
-    apply wand_intro_l, sep_intro_True_l; [apply const_intro|].
+    apply wand_intro_l, sep_intro_True_l; [apply pure_intro|].
     + destruct Hwf; constructor; simpl; eauto using env_app_wf.
       intros j. apply (env_app_disjoint _ _ _ j) in Happ.
       naive_solver eauto using env_app_fresh.
@@ -185,10 +185,10 @@ Lemma envs_simple_replace_sound' Δ Δ' i p Γ :
   envs_delete i p Δ ⊢ □?p [★] Γ -★ Δ'.
 Proof.
   rewrite /envs_simple_replace /envs_delete /of_envs=> ?.
-  apply const_elim_sep_l=> Hwf. destruct Δ as [Γp Γs], p; simplify_eq/=.
+  apply pure_elim_sep_l=> Hwf. destruct Δ as [Γp Γs], p; simplify_eq/=.
   - destruct (env_app Γ Γs) eqn:Happ,
       (env_replace i Γ Γp) as [Γp'|] eqn:?; simplify_eq/=.
-    apply wand_intro_l, sep_intro_True_l; [apply const_intro|].
+    apply wand_intro_l, sep_intro_True_l; [apply pure_intro|].
     + destruct Hwf; constructor; simpl; eauto using env_replace_wf.
       intros j. apply (env_app_disjoint _ _ _ j) in Happ.
       destruct (decide (i = j)); try naive_solver eauto using env_replace_fresh.
@@ -197,7 +197,7 @@ Proof.
       solve_sep_entails.
   - destruct (env_app Γ Γp) eqn:Happ,
       (env_replace i Γ Γs) as [Γs'|] eqn:?; simplify_eq/=.
-    apply wand_intro_l, sep_intro_True_l; [apply const_intro|].
+    apply wand_intro_l, sep_intro_True_l; [apply pure_intro|].
     + destruct Hwf; constructor; simpl; eauto using env_replace_wf.
       intros j. apply (env_app_disjoint _ _ _ j) in Happ.
       destruct (decide (i = j)); try naive_solver eauto using env_replace_fresh.
@@ -225,19 +225,19 @@ Proof. intros. by rewrite envs_lookup_sound// envs_replace_sound'//. Qed.
 Lemma envs_split_sound Δ lr js Δ1 Δ2 :
   envs_split lr js Δ = Some (Δ1,Δ2) → Δ ⊢ Δ1 ★ Δ2.
 Proof.
-  rewrite /envs_split /of_envs=> ?; apply const_elim_sep_l=> Hwf.
+  rewrite /envs_split /of_envs=> ?; apply pure_elim_sep_l=> Hwf.
   destruct Δ as [Γp Γs], (env_split js _) as [[Γs1 Γs2]|] eqn:?; simplify_eq/=.
   rewrite (env_split_perm Γs) // big_sep_app {1}always_sep_dup'.
   destruct lr; simplify_eq/=; cancel [□ [∧] Γp; □ [∧] Γp; [★] Γs1; [★] Γs2]%I;
-    destruct Hwf; apply sep_intro_True_l; apply const_intro; constructor;
+    destruct Hwf; apply sep_intro_True_l; apply pure_intro; constructor;
       naive_solver eauto using env_split_wf_1, env_split_wf_2,
       env_split_fresh_1, env_split_fresh_2.
 Qed.
 
 Lemma envs_clear_spatial_sound Δ : Δ ⊢ envs_clear_spatial Δ ★ [★] env_spatial Δ.
 Proof.
-  rewrite /of_envs /envs_clear_spatial /=; apply const_elim_sep_l=> Hwf.
-  rewrite right_id -assoc; apply sep_intro_True_l; [apply const_intro|done].
+  rewrite /of_envs /envs_clear_spatial /=; apply pure_elim_sep_l=> Hwf.
+  rewrite right_id -assoc; apply sep_intro_True_l; [apply pure_intro|done].
   destruct Hwf; constructor; simpl; auto using Enil_wf.
 Qed.
 
@@ -270,8 +270,8 @@ Proof. intros [??] ?; constructor; eauto using env_Forall2_impl. Qed.
 Global Instance of_envs_mono : Proper (envs_Forall2 (⊢) ==> (⊢)) (@of_envs M).
 Proof.
   intros [Γp1 Γs1] [Γp2 Γs2] [Hp Hs]; unfold of_envs; simpl in *.
-  apply const_elim_sep_l=>Hwf. apply sep_intro_True_l.
-  - destruct Hwf; apply const_intro; constructor;
+  apply pure_elim_sep_l=>Hwf. apply sep_intro_True_l.
+  - destruct Hwf; apply pure_intro; constructor;
       naive_solver eauto using env_Forall2_wf, env_Forall2_fresh.
   - by repeat f_equiv.
 Qed.
@@ -287,8 +287,8 @@ Proof. by constructor. Qed.
 (** * Adequacy *)
 Lemma tac_adequate P : (Envs Enil Enil ⊢ P) → True ⊢ P.
 Proof.
-  intros <-. rewrite /of_envs /= always_const !right_id.
-  apply const_intro; repeat constructor.
+  intros <-. rewrite /of_envs /= always_pure !right_id.
+  apply pure_intro; repeat constructor.
 Qed.
 
 (** * Basic rules *)
@@ -329,7 +329,7 @@ Proof. by rewrite -(False_elim Q). Qed.
 (** * Pure *)
 Class ToPure (P : uPred M) (φ : Prop) := to_pure : P ⊣⊢ ■ φ.
 Arguments to_pure : clear implicits.
-Global Instance to_pure_const φ : ToPure (■ φ) φ.
+Global Instance to_pure_pure φ : ToPure (■ φ) φ.
 Proof. done. Qed.
 Global Instance to_pure_eq {A : cofeT} (a b : A) :
   Timeless a → ToPure (a ≡ b) (a ≡ b).
@@ -338,18 +338,18 @@ Global Instance to_pure_valid `{CMRADiscrete A} (a : A) : ToPure (✓ a) (✓ a)
 Proof. intros; red. by rewrite discrete_valid. Qed.
 
 Lemma tac_pure_intro Δ Q (φ : Prop) : ToPure Q φ → φ → Δ ⊢ Q.
-Proof. intros ->. apply const_intro. Qed.
+Proof. intros ->. apply pure_intro. Qed.
 
 Lemma tac_pure Δ Δ' i p P φ Q :
   envs_lookup_delete i Δ = Some (p, P, Δ') → ToPure P φ →
   (φ → Δ' ⊢ Q) → Δ ⊢ Q.
 Proof.
   intros ?? HQ. rewrite envs_lookup_delete_sound' //; simpl.
-  rewrite (to_pure P); by apply const_elim_sep_l.
+  rewrite (to_pure P); by apply pure_elim_sep_l.
 Qed.
 
 Lemma tac_pure_revert Δ φ Q : (Δ ⊢ ■ φ → Q) → (φ → Δ ⊢ Q).
-Proof. intros HΔ ?. by rewrite HΔ const_equiv // left_id. Qed.
+Proof. intros HΔ ?. by rewrite HΔ pure_equiv // left_id. Qed.
 
 (** * Later *)
 Class StripLaterEnv (Γ1 Γ2 : env (uPred M)) :=
@@ -373,7 +373,7 @@ Lemma strip_later_env_sound Δ1 Δ2 : StripLaterEnvs Δ1 Δ2 → Δ1 ⊢ ▷ Δ2
 Proof.
   intros [Hp Hs]; rewrite /of_envs /= !later_sep -always_later.
   repeat apply sep_mono; try apply always_mono.
-  - rewrite -later_intro; apply const_mono; destruct 1; constructor;
+  - rewrite -later_intro; apply pure_mono; destruct 1; constructor;
       naive_solver eauto using env_Forall2_wf, env_Forall2_fresh.
   - induction Hp; rewrite /= ?later_and; auto using and_mono, later_intro.
   - induction Hs; rewrite /= ?later_sep; auto using sep_mono, later_intro.
@@ -437,7 +437,7 @@ Proof.
 Qed.
 Lemma tac_impl_intro_pure Δ P φ Q : ToPure P φ → (φ → Δ ⊢ Q) → Δ ⊢ P → Q.
 Proof.
-  intros. by apply impl_intro_l; rewrite (to_pure P); apply const_elim_l.
+  intros. by apply impl_intro_l; rewrite (to_pure P); apply pure_elim_l.
 Qed.
 
 Lemma tac_wand_intro Δ Δ' i P Q :
@@ -455,7 +455,7 @@ Proof.
 Qed.
 Lemma tac_wand_intro_pure Δ P φ Q : ToPure P φ → (φ → Δ ⊢ Q) → Δ ⊢ P -★ Q.
 Proof.
-  intros. by apply wand_intro_l; rewrite (to_pure P); apply const_elim_sep_l.
+  intros. by apply wand_intro_l; rewrite (to_pure P); apply pure_elim_sep_l.
 Qed.
 
 Class ToWand (R P Q : uPred M) := to_wand : R ⊢ P -★ Q.
@@ -524,7 +524,7 @@ Lemma tac_specialize_pure Δ Δ' j q R P1 P2 φ Q :
   φ → (Δ' ⊢ Q) → Δ ⊢ Q.
 Proof.
   intros. rewrite envs_simple_replace_sound //; simpl.
-  by rewrite right_id (to_wand R) (to_pure P1) const_equiv // wand_True wand_elim_r.
+  by rewrite right_id (to_wand R) (to_pure P1) pure_equiv // wand_True wand_elim_r.
 Qed.
 
 Lemma tac_specialize_persistent Δ Δ' Δ'' j q P1 P2 R Q :
@@ -598,7 +598,7 @@ Lemma tac_pose_proof Δ Δ' j P1 P2 R Q :
   (Δ' ⊢ Q) → Δ ⊢ Q.
 Proof.
   intros HP ?? <-. rewrite envs_app_sound //; simpl.
-  by rewrite right_id -(to_pose_proof P1 P2 R) // always_const wand_True.
+  by rewrite right_id -(to_pose_proof P1 P2 R) // always_pure wand_True.
 Qed.
 
 Lemma tac_pose_proof_hyp Δ Δ' Δ'' i p j P Q :
diff --git a/proofmode/sts.v b/proofmode/sts.v
index a3b588195db07ae5709793a209167e5b40195e86..b4817fbac06b37a265d03596ccaee0953ed03f6f 100644
--- a/proofmode/sts.v
+++ b/proofmode/sts.v
@@ -22,7 +22,7 @@ Proof.
   rewrite // -always_and_sep_l. apply and_intro; first done.
   rewrite envs_lookup_sound //; simpl; apply sep_mono_r.
   apply forall_intro=>s; apply wand_intro_l.
-  rewrite -assoc; apply const_elim_sep_l=> Hs.
+  rewrite -assoc; apply pure_elim_sep_l=> Hs.
   destruct (HΔ' s) as (Δ'&?&?); clear HΔ'; auto.
   rewrite envs_simple_replace_sound' //; simpl.
   by rewrite right_id wand_elim_r.