diff --git a/theories/base.v b/theories/base.v
index 859c6f009c545ee34b790e35bda610872105c3d9..92606a94b39483093fc2900afbeff5687503ae0f 100644
--- a/theories/base.v
+++ b/theories/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/bset.v b/theories/bset.v
index 91501e33703fdda2297814a3d9f5ea63124cf7d9..8f9e369bad93dc6f0cb0271ab5402a156206a652 100644
--- a/theories/bset.v
+++ b/theories/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 stdpp 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/coPset.v b/theories/coPset.v
index 49e323349baf4cb6a68ec7123834a6a5fe16e337..a36339f4c9697aa4690177ab8b94eb1ffe7900c8 100644
--- a/theories/coPset.v
+++ b/theories/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 stdpp Require Export collections.
 From stdpp 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/countable.v b/theories/countable.v
index 3ac06848c3a7c575623691e6c71ef6a0a9655697..c4e153736e929b7f74ddd84e7387ec2065d0a1d3 100644
--- a/theories/countable.v
+++ b/theories/countable.v
@@ -1,7 +1,7 @@
 (* Copyright (c) 2012-2015, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
 From stdpp 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/finite.v b/theories/finite.v
index 9a4d20a4b55196c6d7e022524eb8721e286a3fe2..793435cb1a6c05fd2b3cb783249152f7b7f3e7fb 100644
--- a/theories/finite.v
+++ b/theories/finite.v
@@ -1,7 +1,7 @@
 (* Copyright (c) 2012-2015, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
 From stdpp 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/functions.v b/theories/functions.v
index 4c37187d2a8146e1bb0dacd986d8ffab3a0b50f8..868a430b336e13894adca08da2b4b6c598c34159 100644
--- a/theories/functions.v
+++ b/theories/functions.v
@@ -1,5 +1,5 @@
 From stdpp 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/gmap.v b/theories/gmap.v
index a1bf157f471bca4c9c295a82c720bbbaa68d94a2..7787305d90d44b8e04a9fa3bfb3a9ea93ed65f41 100644
--- a/theories/gmap.v
+++ b/theories/gmap.v
@@ -4,7 +4,7 @@
 type. The implementation is based on [Pmap]s, radix-2 search trees. *)
 From stdpp Require Export countable fin_maps fin_map_dom.
 From stdpp 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/gmultiset.v b/theories/gmultiset.v
index 7f1f83eb762e96661a72fd215a1bc8f52f142abe..3918fd9b3c3dd42523e07b8aa70cb293ca3e4bcd 100644
--- a/theories/gmultiset.v
+++ b/theories/gmultiset.v
@@ -1,7 +1,7 @@
 (* Copyright (c) 2012-2016, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
 From stdpp 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/hashset.v b/theories/hashset.v
index ffed2bd15a4a66b36bbd27d1ddce94f93d57be80..215479cc3ba30ffab6346c95c690e341db1940ce 100644
--- a/theories/hashset.v
+++ b/theories/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 stdpp Require Export fin_maps listset.
 From stdpp 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/hlist.v b/theories/hlist.v
index 39ad98053bf35cf7dd8ca1633b567eb5b0c53134..7c0dff7fca533278e18f4c4ee889b18bb817a800 100644
--- a/theories/hlist.v
+++ b/theories/hlist.v
@@ -1,5 +1,5 @@
 From stdpp 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/lexico.v b/theories/lexico.v
index f5402ebd63331014aecad124d372c68d0ffe73d7..7ab91c2b5c330782064a379b105825614df9b69e 100644
--- a/theories/lexico.v
+++ b/theories/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 stdpp Require Import numbers.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Notation cast_trichotomy T :=
   match T with
diff --git a/theories/listset.v b/theories/listset.v
index 8966cc2e86df5251be0505163e81bdb8e1b3dd77..f86a0d83378f1394ec5cdb5ad1bab44144eb4bb9 100644
--- a/theories/listset.v
+++ b/theories/listset.v
@@ -3,7 +3,7 @@
 (** This file implements finite set as unordered lists without duplicates
 removed. This implementation forms a monad. *)
 From stdpp 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/listset_nodup.v b/theories/listset_nodup.v
index e3b76d511b0bc4854471608d97d07dca6f08e2b2..5f62b6ff4a8bde2143e6fe7cbd89b0a86fd7b731 100644
--- a/theories/listset_nodup.v
+++ b/theories/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 stdpp 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/natmap.v b/theories/natmap.v
index d109590bbc65a0980f430860139a51b2db69f72a..31e7c088e2d9c8a0e055aead238bbb8e22e12137 100644
--- a/theories/natmap.v
+++ b/theories/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 stdpp 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/nmap.v b/theories/nmap.v
index 12645efbaa269de49be4cc9c195d25001a6a5c5f..a37a081727bec43d673abd265bc781255e3bae9c 100644
--- a/theories/nmap.v
+++ b/theories/nmap.v
@@ -4,7 +4,7 @@
 maps whose keys range over Coq's data type of binary naturals [N]. *)
 From stdpp Require Import pmap mapset.
 From stdpp Require Export prelude fin_maps.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Local Open Scope N_scope.
 
diff --git a/theories/numbers.v b/theories/numbers.v
index 034a4ca14ce014e8e918fadf55eb19fdf28752d1..ad31094f43a81b893226f981ad687698759641a8 100644
--- a/theories/numbers.v
+++ b/theories/numbers.v
@@ -6,7 +6,7 @@ notations. *)
 From Coq Require Export EqdepFacts PArith NArith ZArith NPeano.
 From Coq Require Import QArith Qcanon.
 From stdpp 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/option.v b/theories/option.v
index 53fe1bf5c98b5ee81a4054e091a3bfc60a06b64a..610fe2c6df85de21c861873ec7ac79428eec7aea 100644
--- a/theories/option.v
+++ b/theories/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 stdpp 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/orders.v b/theories/orders.v
index bc301900b0f68d31fe893caaaba650a2e11194e6..f8897d0bb3b67c01aa617090da923b5a079531e7 100644
--- a/theories/orders.v
+++ b/theories/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 stdpp Require Export tactics.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 Section orders.
   Context {A} {R : relation A}.
diff --git a/theories/pmap.v b/theories/pmap.v
index a6966f39ef92ad83321ac192ca5704a3e310fafc..ccb29e9af339c2deb2e56dc08ec0466b25e9a188 100644
--- a/theories/pmap.v
+++ b/theories/pmap.v
@@ -10,7 +10,7 @@ Leibniz equality to become extensional. *)
 From Coq Require Import PArith.
 From stdpp Require Import mapset countable.
 From stdpp 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/pretty.v b/theories/pretty.v
index afd2b6899b6b6f2925c00aa6a49d34f92450315d..febf1048644f248a4b83337523a7f570e8608a42 100644
--- a/theories/pretty.v
+++ b/theories/pretty.v
@@ -3,7 +3,7 @@
 From stdpp Require Export strings.
 From stdpp 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/proof_irrel.v b/theories/proof_irrel.v
index 4d9c6d7d83bf459912342375e766e2b18dde9d97..f4fc9738df5c9eb7d1008f8a32fb6b92e8edc207 100644
--- a/theories/proof_irrel.v
+++ b/theories/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 stdpp 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/relations.v b/theories/relations.v
index 3152cc7888f61f35ecc4a89339d569fdb575976c..84f2ff9cee24c81ce42f8284ac001d136413d863 100644
--- a/theories/relations.v
+++ b/theories/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 stdpp Require Export tactics base.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 
 (** * Definitions *)
 Section definitions.
diff --git a/theories/set.v b/theories/set.v
index 27a979751c390c47c7cf780e27ce6913858e6f38..0e5b5a96a8c15e6e16b0e5f3d9c9f1705225a146 100644
--- a/theories/set.v
+++ b/theories/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 stdpp 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/sorting.v b/theories/sorting.v
index 46d7eea554cea9d54489b5f0f803ae3875be5f07..da10080278df33755fe82617cd9a3ec308bb27b7 100644
--- a/theories/sorting.v
+++ b/theories/sorting.v
@@ -4,7 +4,7 @@
 standard library, but without using the module system. *)
 From Coq Require Export Sorted.
 From stdpp 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/streams.v b/theories/streams.v
index 725a59158e64a09efab7bd04928ecd819535e789..4b6fbe54e4c9337e956a2990ba10940db1b9f4a9 100644
--- a/theories/streams.v
+++ b/theories/streams.v
@@ -1,7 +1,7 @@
 (* Copyright (c) 2012-2015, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
 From stdpp 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/stringmap.v b/theories/stringmap.v
index 9de316cae7b019ad4195f5458ef3cf5d8b25a63a..b5f13b21f92658101dee3f3c63a6abdd9d99fc24 100644
--- a/theories/stringmap.v
+++ b/theories/stringmap.v
@@ -6,7 +6,7 @@ search trees (uncompressed Patricia trees) as implemented in the file [pmap]
 and guarantees logarithmic-time operations. *)
 From stdpp Require Export fin_maps pretty.
 From stdpp 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/strings.v b/theories/strings.v
index 93718899c741b6b8987ec498d3406cd415cbf25f..5e00a2ad7baf984551b38c45f60ff46b73d264dd 100644
--- a/theories/strings.v
+++ b/theories/strings.v
@@ -4,7 +4,7 @@ From Coq Require Import Ascii.
 From Coq Require Export String.
 From stdpp Require Export list.
 From stdpp 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/tactics.v b/theories/tactics.v
index f104f381e9c200d323e5d9a0eb578bb50ee9f8c0..c6ca7aa5f71f93191a601f84df73e828dd0765a6 100644
--- a/theories/tactics.v
+++ b/theories/tactics.v
@@ -5,7 +5,7 @@ the development. *)
 From Coq Require Import Omega.
 From Coq Require Export Lia.
 From stdpp 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/vector.v b/theories/vector.v
index b0d6d326916c6085de6d98a3b7e08407950e60c0..ef6d2d4bb266cffb6c0571bac87986badb7b8bcb 100644
--- a/theories/vector.v
+++ b/theories/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 stdpp Require Export list.
-Set Default Proof Using "Type*".
+Set Default Proof Using "Type".
 Open Scope vector_scope.
 
 (** * The fin type *)
diff --git a/theories/zmap.v b/theories/zmap.v
index 32689e5859a04a323fb65af130b9a75e51f293e9..227a420d8fa6740c052532e9a7a7d1bc9bf8b96f 100644
--- a/theories/zmap.v
+++ b/theories/zmap.v
@@ -4,7 +4,7 @@
 maps whose keys range over Coq's data type of binary naturals [Z]. *)
 From stdpp Require Import pmap mapset.
 From stdpp 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 :=