diff --git a/algebra/frac.v b/algebra/frac.v
index 47fd97642304649210775ce3d4d22324fef5991f..2a10b1cd166c725eb0cc7a8498bed1632838ef46 100644
--- a/algebra/frac.v
+++ b/algebra/frac.v
@@ -1,6 +1,5 @@
 From Coq.QArith Require Import Qcanon.
 From iris.algebra Require Export cmra.
-From iris.algebra Require Import upred.
 
 Notation frac := Qp (only parsing).
 
@@ -20,12 +19,6 @@ Qed.
 Canonical Structure fracR := discreteR frac frac_ra_mixin.
 End frac.
 
-(** Internalized properties *)
-Lemma frac_equivI {M} (x y : frac) : x ≡ y ⊣⊢ (x = y : uPred M).
-Proof. by uPred.unseal. Qed.
-Lemma frac_validI {M} (x : frac) : ✓ x ⊣⊢ (■ (x ≤ 1)%Qc : uPred M).
-Proof. by uPred.unseal. Qed.
-
 (** Exclusive *)
 Global Instance frac_full_exclusive : Exclusive 1%Qp.
 Proof.
diff --git a/heap_lang/heap.v b/heap_lang/heap.v
index 2457a9577b084ff7878ee88dcd55b0e303e9bc60..597066ded0a206bd929ae9412eb29cc9c6ced61e 100644
--- a/heap_lang/heap.v
+++ b/heap_lang/heap.v
@@ -118,7 +118,7 @@ Section heap.
     rewrite heap_mapsto_eq -auth_own_op op_singleton pair_op dec_agree_ne //.
     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.
+    rewrite option_validI prod_validI !discrete_valid /=.
     by apply pure_elim_r.
   Qed.
 
diff --git a/program_logic/model.v b/program_logic/model.v
index 1beabfcd56a42a8029f4421783f416e44b0d31fb..212d2a740124f151f806cb9f5d8f660d3ec8d11c 100644
--- a/program_logic/model.v
+++ b/program_logic/model.v
@@ -71,9 +71,9 @@ Notation "#[ Σ1 ; .. ; Σn ]" :=
 (** * Subfunctors *)
 (** In order to make proofs in the Iris logic modular, they are not done with
 respect to some concrete list of functors [Σ], but are instead parametrized by
-an arbitrary list of functors [Σ] that contains atleast certain functors. For
-example, the lock library is parametrized by a functor [Σ] that should have:
-the functors corresponding to: the heap and the exclusive monoid to manage to
+an arbitrary list of functors [Σ] that contains at least certain functors. For
+example, the lock library is parameterized by a functor [Σ] that should have
+the functors corresponding to the heap and the exclusive monoid to manage to
 lock invariant.
 
 The contraints to can be expressed using the type class [subG Σ1 Σ2], which
@@ -109,7 +109,7 @@ Qed.
 
 
 (** * Solution of the recursive domain equation *)
-(** We first declare a module type and then an instance of it so as to seall of
+(** We first declare a module type and then an instance of it so as to seal all of
 the construction, this way we are sure we do not use any properties of the
 construction, and also avoid Coq from blindly unfolding it. *)
 Module Type iProp_solution_sig.