diff --git a/theories/prelude/base.v b/theories/prelude/base.v
index 859c6f009c545ee34b790e35bda610872105c3d9..92606a94b39483093fc2900afbeff5687503ae0f 100644
--- a/theories/prelude/base.v
+++ b/theories/prelude/base.v
@@ -9,7 +9,7 @@ Global Set Automatic Coercions Import.
 Global Set Asymmetric Patterns.
 Global Unset Transparent Obligations.
 From Coq Require Export Morphisms RelationClasses List Bool Utf8 Setoid.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 Export ListNotations.
 From Coq.Program Require Export Basics Syntax.
 Obligation Tactic := idtac.
diff --git a/theories/prelude/bset.v b/theories/prelude/bset.v
index a431ca3cfc713502d05f1e0038eab093e9731d22..e384b05c5cafeab3ec0fea31e33bb03ba5badcd2 100644
--- a/theories/prelude/bset.v
+++ b/theories/prelude/bset.v
@@ -2,7 +2,7 @@
 (* This file is distributed under the terms of the BSD license. *)
 (** This file implements bsets as functions into Prop. *)
 From iris.prelude Require Export prelude.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Record bset (A : Type) : Type := mkBSet { bset_car : A → bool }.
 Arguments mkBSet {_} _.
diff --git a/theories/prelude/coPset.v b/theories/prelude/coPset.v
index 0ad02732ddc81f2ab981ca3e29573792401c7d65..63be7ed49f02ccce094f5f964ae078f0df229e5c 100644
--- a/theories/prelude/coPset.v
+++ b/theories/prelude/coPset.v
@@ -13,7 +13,7 @@ Since [positive]s are bitstrings, we encode [coPset]s as trees that correspond
 to the decision function that map bitstrings to bools. *)
 From iris.prelude Require Export collections.
 From iris.prelude Require Import pmap gmap mapset.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 Local Open Scope positive_scope.
 
 (** * The tree data structure *)
diff --git a/theories/prelude/countable.v b/theories/prelude/countable.v
index 37f0d922da96e6670680327884c1050bf830c177..ac733dac2ae52dcc65b47ad139b70d0350bc51b1 100644
--- a/theories/prelude/countable.v
+++ b/theories/prelude/countable.v
@@ -1,7 +1,7 @@
 (* Copyright (c) 2012-2015, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
 From iris.prelude Require Export list.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 Local Open Scope positive.
 
 Class Countable A `{EqDecision A} := {
diff --git a/theories/prelude/finite.v b/theories/prelude/finite.v
index 9c51d7cbf7442d4a67ce05b818f57dc27841544f..041e3b4b0830803d9efca1253efa4255a187a90a 100644
--- a/theories/prelude/finite.v
+++ b/theories/prelude/finite.v
@@ -1,7 +1,7 @@
 (* Copyright (c) 2012-2015, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
 From iris.prelude Require Export countable vector.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Class Finite A `{EqDecision A} := {
   enum : list A;
@@ -181,12 +181,12 @@ Section forall_exists.
   Context `{∀ x, Decision (P x)}.
 
   Global Instance forall_dec: Decision (∀ x, P x).
-  Proof.
+  Proof using Type*.
    refine (cast_if (decide (Forall P (enum A))));
     abstract by rewrite <-Forall_finite.
   Defined.
   Global Instance exists_dec: Decision (∃ x, P x).
-  Proof.
+  Proof using Type*.
    refine (cast_if (decide (Exists P (enum A))));
     abstract by rewrite <-Exists_finite.
   Defined.
diff --git a/theories/prelude/functions.v b/theories/prelude/functions.v
index da46a43a468911e654cd1b3ad5e63eda4876f353..901a3dbca9f1586ffd8900c2531f275b83c4ed9f 100644
--- a/theories/prelude/functions.v
+++ b/theories/prelude/functions.v
@@ -1,5 +1,5 @@
 From iris.prelude Require Export base tactics.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Section definitions.
   Context {A T : Type} `{EqDecision A}.
diff --git a/theories/prelude/gmap.v b/theories/prelude/gmap.v
index 4e5d2f7746cdea3be3aef26871052ca5a49b2d1f..c47984417a43c7a6c32d4ca6b4972461ee765cf5 100644
--- a/theories/prelude/gmap.v
+++ b/theories/prelude/gmap.v
@@ -4,7 +4,7 @@
 type. The implementation is based on [Pmap]s, radix-2 search trees. *)
 From iris.prelude Require Export countable fin_maps fin_map_dom.
 From iris.prelude Require Import pmap mapset set.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 (** * The data structure *)
 (** We pack a [Pmap] together with a proof that ensures that all keys correspond
diff --git a/theories/prelude/gmultiset.v b/theories/prelude/gmultiset.v
index 7782da2455494db701aa985a4b0840b5f6060b4b..25eb03ae01bf806e5634bebca24405cb5c5e7b61 100644
--- a/theories/prelude/gmultiset.v
+++ b/theories/prelude/gmultiset.v
@@ -1,7 +1,7 @@
 (* Copyright (c) 2012-2016, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
 From iris.prelude Require Import gmap.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Record gmultiset A `{Countable A} := GMultiSet { gmultiset_car : gmap A nat }.
 Arguments GMultiSet {_ _ _} _.
diff --git a/theories/prelude/hashset.v b/theories/prelude/hashset.v
index dcfd1dfe38084e94a4db0106f0f7a5b27ee52da9..bfa30fe7cc9cc4911be846761ea8e2a24056b21e 100644
--- a/theories/prelude/hashset.v
+++ b/theories/prelude/hashset.v
@@ -5,7 +5,7 @@ using radix-2 search trees. Each hash bucket is thus indexed using an binary
 integer of type [Z], and contains an unordered list without duplicates. *)
 From iris.prelude Require Export fin_maps listset.
 From iris.prelude Require Import zmap.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Record hashset {A} (hash : A → Z) := Hashset {
   hashset_car : Zmap (list A);
diff --git a/theories/prelude/hlist.v b/theories/prelude/hlist.v
index 907864fef88284a254f9ef075ff4979a622102f0..26077a1d539ee32dc03cef1a03a669a1591b3827 100644
--- a/theories/prelude/hlist.v
+++ b/theories/prelude/hlist.v
@@ -1,5 +1,5 @@
 From iris.prelude Require Import tactics.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 Local Set Universe Polymorphism.
 
 (* Not using [list Type] in order to avoid universe inconsistencies *)
diff --git a/theories/prelude/lexico.v b/theories/prelude/lexico.v
index 6a7b868a34c36033fa9abe9e3d12f5b3e605090e..2d836a85ff0d51e1bf354b225d42d0068b2ac264 100644
--- a/theories/prelude/lexico.v
+++ b/theories/prelude/lexico.v
@@ -3,7 +3,7 @@
 (** This files defines a lexicographic order on various common data structures
 and proves that it is a partial order having a strong variant of trichotomy. *)
 From iris.prelude Require Import numbers.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Notation cast_trichotomy T :=
   match T with
diff --git a/theories/prelude/listset.v b/theories/prelude/listset.v
index a3fa10badd7fc78430b96f36106a20ae7133f3b8..88bda4f61caa1e4b6ecc4d83c2897e11b9d0ed70 100644
--- a/theories/prelude/listset.v
+++ b/theories/prelude/listset.v
@@ -3,7 +3,7 @@
 (** This file implements finite set as unordered lists without duplicates
 removed. This implementation forms a monad. *)
 From iris.prelude Require Export collections list.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Record listset A := Listset { listset_car: list A }.
 Arguments listset_car {_} _.
diff --git a/theories/prelude/listset_nodup.v b/theories/prelude/listset_nodup.v
index b9aedede369f901865088713eea5127f3b042f4f..960f39b633d206c6a80e439ad558812b7cde730c 100644
--- a/theories/prelude/listset_nodup.v
+++ b/theories/prelude/listset_nodup.v
@@ -4,7 +4,7 @@
 Although this implementation is slow, it is very useful as decidable equality
 is the only constraint on the carrier set. *)
 From iris.prelude Require Export collections list.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Record listset_nodup A := ListsetNoDup {
   listset_nodup_car : list A; listset_nodup_prf : NoDup listset_nodup_car
diff --git a/theories/prelude/natmap.v b/theories/prelude/natmap.v
index b6e78446d5fc8626faec5077db1a9c887846e04a..2bf78d1ea612792636be074815d00ad720dc16cb 100644
--- a/theories/prelude/natmap.v
+++ b/theories/prelude/natmap.v
@@ -4,7 +4,7 @@
 over Coq's data type of unary natural numbers [nat]. The implementation equips
 a list with a proof of canonicity. *)
 From iris.prelude Require Import fin_maps mapset.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Notation natmap_raw A := (list (option A)).
 Definition natmap_wf {A} (l : natmap_raw A) :=
diff --git a/theories/prelude/nmap.v b/theories/prelude/nmap.v
index 0c965df3cd7b7a3007f26a9ce279599d5b0b0748..7957b666c861a945b0696540c7f77a354b940486 100644
--- a/theories/prelude/nmap.v
+++ b/theories/prelude/nmap.v
@@ -4,7 +4,7 @@
 maps whose keys range over Coq's data type of binary naturals [N]. *)
 From iris.prelude Require Import pmap mapset.
 From iris.prelude Require Export prelude fin_maps.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Local Open Scope N_scope.
 
diff --git a/theories/prelude/numbers.v b/theories/prelude/numbers.v
index c8eacf886eabd5b113530ad6e00a5a85c7b05726..e26e0e73ad6c596d4ba2b8288b63dfd7ffd91ad0 100644
--- a/theories/prelude/numbers.v
+++ b/theories/prelude/numbers.v
@@ -6,7 +6,7 @@ notations. *)
 From Coq Require Export EqdepFacts PArith NArith ZArith NPeano.
 From Coq Require Import QArith Qcanon.
 From iris.prelude Require Export base decidable option.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 Open Scope nat_scope.
 
 Coercion Z.of_nat : nat >-> Z.
diff --git a/theories/prelude/option.v b/theories/prelude/option.v
index f6cc09cc598c032aa003eb7aa2dd9040d72e90d0..76f72c91299e63b858ae93230ba918a69267e783 100644
--- a/theories/prelude/option.v
+++ b/theories/prelude/option.v
@@ -3,7 +3,7 @@
 (** This file collects general purpose definitions and theorems on the option
 data type that are not in the Coq standard library. *)
 From iris.prelude Require Export tactics.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Inductive option_reflect {A} (P : A → Prop) (Q : Prop) : option A → Type :=
   | ReflectSome x : P x → option_reflect P Q (Some x)
diff --git a/theories/prelude/orders.v b/theories/prelude/orders.v
index 9e6e2ca20ecf7245a4634d2b5adfc23aeef5addc..a28f3ed551144154f3751a02a3fb07a92e01b294 100644
--- a/theories/prelude/orders.v
+++ b/theories/prelude/orders.v
@@ -3,7 +3,7 @@
 (** Properties about arbitrary pre-, partial, and total orders. We do not use
 the relation [⊆] because we often have multiple orders on the same structure *)
 From iris.prelude Require Export tactics.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Section orders.
   Context {A} {R : relation A}.
diff --git a/theories/prelude/pmap.v b/theories/prelude/pmap.v
index 60ded55a1d0e21c0e5f59ae0a4a6f199b56ea801..4d56bc2597318ba73270d7270341a62e213ea121 100644
--- a/theories/prelude/pmap.v
+++ b/theories/prelude/pmap.v
@@ -10,7 +10,7 @@ Leibniz equality to become extensional. *)
 From Coq Require Import PArith.
 From iris.prelude Require Import mapset countable.
 From iris.prelude Require Export fin_maps.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Local Open Scope positive_scope.
 Local Hint Extern 0 (@eq positive _ _) => congruence.
diff --git a/theories/prelude/pretty.v b/theories/prelude/pretty.v
index 06924729b0dc363cd6ee1b30d55d809c96660ff5..61dddb24067afa367f3d7734f4346cf66a7b1047 100644
--- a/theories/prelude/pretty.v
+++ b/theories/prelude/pretty.v
@@ -3,7 +3,7 @@
 From iris.prelude Require Export strings.
 From iris.prelude Require Import relations.
 From Coq Require Import Ascii.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Class Pretty A := pretty : A → string.
 Definition pretty_N_char (x : N) : ascii :=
diff --git a/theories/prelude/proof_irrel.v b/theories/prelude/proof_irrel.v
index b409854626c30804db5f1ae82938064c3e47cddc..6b785c8de6b6b049eca8ab4492ea1f420a53ec21 100644
--- a/theories/prelude/proof_irrel.v
+++ b/theories/prelude/proof_irrel.v
@@ -2,7 +2,7 @@
 (* This file is distributed under the terms of the BSD license. *)
 (** This file collects facts on proof irrelevant types/propositions. *)
 From iris.prelude Require Export base.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Hint Extern 200 (ProofIrrel _) => progress (lazy beta) : typeclass_instances.
 
diff --git a/theories/prelude/relations.v b/theories/prelude/relations.v
index 798245df8fb37adc2a8138c0fd0de11752cfb9fa..9984fe671d52d66ce63ecfceaca36f4d0bc72ee4 100644
--- a/theories/prelude/relations.v
+++ b/theories/prelude/relations.v
@@ -6,7 +6,7 @@ small step semantics. This file defines a hint database [ars] containing
 some theorems on abstract rewriting systems. *)
 From Coq Require Import Wf_nat.
 From iris.prelude Require Export tactics base.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 (** * Definitions *)
 Section definitions.
diff --git a/theories/prelude/set.v b/theories/prelude/set.v
index 7e9385e8096e1a5969e779135e5104d396979cb3..22e8278f896c31aa0bc4499434bbdda668bfa6ea 100644
--- a/theories/prelude/set.v
+++ b/theories/prelude/set.v
@@ -2,7 +2,7 @@
 (* This file is distributed under the terms of the BSD license. *)
 (** This file implements sets as functions into Prop. *)
 From iris.prelude Require Export collections.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Record set (A : Type) : Type := mkSet { set_car : A → Prop }.
 Add Printing Constructor set.
diff --git a/theories/prelude/sorting.v b/theories/prelude/sorting.v
index 2048c3017f678c7025d88c44b03013692642a765..4ae478e012c9ba605449ad86054f07d38d491fcf 100644
--- a/theories/prelude/sorting.v
+++ b/theories/prelude/sorting.v
@@ -4,7 +4,7 @@
 standard library, but without using the module system. *)
 From Coq Require Export Sorted.
 From iris.prelude Require Export orders list.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Section merge_sort.
   Context {A} (R : relation A) `{∀ x y, Decision (R x y)}.
diff --git a/theories/prelude/streams.v b/theories/prelude/streams.v
index 1c9ef9aa95f544b04ca7b4dbda7d5b9bfc4bdf93..2b56722d8a1dd92e93b9cb233bfec0836718a91a 100644
--- a/theories/prelude/streams.v
+++ b/theories/prelude/streams.v
@@ -1,7 +1,7 @@
 (* Copyright (c) 2012-2015, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
 From iris.prelude Require Export tactics.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 CoInductive stream (A : Type) : Type := scons : A → stream A → stream A.
 Arguments scons {_} _ _.
diff --git a/theories/prelude/stringmap.v b/theories/prelude/stringmap.v
index 41fc9b8cbcc8353a0ccb57014432c1a07a71ee22..909b65b5b42e130c277ff521d5767582db30c06d 100644
--- a/theories/prelude/stringmap.v
+++ b/theories/prelude/stringmap.v
@@ -6,7 +6,7 @@ search trees (uncompressed Patricia trees) as implemented in the file [pmap]
 and guarantees logarithmic-time operations. *)
 From iris.prelude Require Export fin_maps pretty.
 From iris.prelude Require Import gmap.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Notation stringmap := (gmap string).
 Notation stringset := (gset string).
diff --git a/theories/prelude/strings.v b/theories/prelude/strings.v
index dca3404c10e93169afc142ad2cb03d6ecc202697..3c72afe6a42b98178888ed23970cd1fecec81b27 100644
--- a/theories/prelude/strings.v
+++ b/theories/prelude/strings.v
@@ -4,7 +4,7 @@ From Coq Require Import Ascii.
 From Coq Require Export String.
 From iris.prelude Require Export list.
 From iris.prelude Require Import countable.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 (* To avoid randomly ending up with String.length because this module is
 imported hereditarily somewhere. *)
diff --git a/theories/prelude/tactics.v b/theories/prelude/tactics.v
index 21d0122a1e4331438f3d2692d338f0bbfefab258..7133ba584bd8ed403aabff782260220da49d00ef 100644
--- a/theories/prelude/tactics.v
+++ b/theories/prelude/tactics.v
@@ -5,7 +5,7 @@ the development. *)
 From Coq Require Import Omega.
 From Coq Require Export Lia.
 From iris.prelude Require Export decidable.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Lemma f_equal_dep {A B} (f g : ∀ x : A, B x) x : f = g → f x = g x.
 Proof. intros ->; reflexivity. Qed.
diff --git a/theories/prelude/vector.v b/theories/prelude/vector.v
index 4b34264582b455e6894fa705c7aaf29c8d7dcf93..6eb2cadab3a9d6e903789435bc63b54c40e10684 100644
--- a/theories/prelude/vector.v
+++ b/theories/prelude/vector.v
@@ -6,7 +6,7 @@ definitions from the standard library, but renames or changes their notations,
 so that it becomes more consistent with the naming conventions in this
 development. *)
 From iris.prelude Require Export list.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 Open Scope vector_scope.
 
 (** * The fin type *)
diff --git a/theories/prelude/zmap.v b/theories/prelude/zmap.v
index c9c1759196e8d53c45d850971e6b7ca81ae52cdb..cffb198abb1260301add5de05baf941d9228b703 100644
--- a/theories/prelude/zmap.v
+++ b/theories/prelude/zmap.v
@@ -4,7 +4,7 @@
 maps whose keys range over Coq's data type of binary naturals [Z]. *)
 From iris.prelude Require Import pmap mapset.
 From iris.prelude Require Export prelude fin_maps.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 Local Open Scope Z_scope.
 
 Record Zmap (A : Type) : Type :=