diff --git a/theories/base.v b/theories/base.v
index f711649e7014bff9f132e3546a9b6df8753e904a..63ddc65098ed010a4007c896e2059016780385f5 100644
--- a/theories/base.v
+++ b/theories/base.v
@@ -7,7 +7,7 @@ structures. *)
 Global Generalizable All Variables.
 Global Set Automatic Coercions Import.
 Global Set Asymmetric Patterns.
-Require Export Morphisms RelationClasses List Bool Utf8 Program Setoid.
+From Coq Require Export Morphisms RelationClasses List Bool Utf8 Program Setoid.
 Obligation Tactic := idtac.
 
 (** * General *)
diff --git a/theories/bsets.v b/theories/bsets.v
index 87e1c50eaffba71dd7e8aa4ed35d953b7ad3d461..8ae132f1f432046baa268c02ab2a64609d4b0a72 100644
--- a/theories/bsets.v
+++ b/theories/bsets.v
@@ -1,7 +1,7 @@
 (* Copyright (c) 2012-2015, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
 (** This file implements bsets as functions into Prop. *)
-Require Export prelude.prelude.
+From stdpp Require Export prelude.
 
 Record bset (A : Type) : Type := mkBSet { bset_car : A → bool }.
 Arguments mkBSet {_} _.
diff --git a/theories/co_pset.v b/theories/co_pset.v
index 3b49cb373059d3d8b1c1dda5aee2b73e14af5587..85e80a4665b84b99ee4bfde4ab3bd6a1ac2af3d3 100644
--- a/theories/co_pset.v
+++ b/theories/co_pset.v
@@ -2,8 +2,8 @@
 (* This file is distributed under the terms of the BSD license. *)
 (** This files implements an efficient implementation of finite/cofinite sets
 of positive binary naturals [positive]. *)
-Require Export prelude.collections.
-Require Import prelude.pmap prelude.gmap prelude.mapset.
+From stdpp Require Export collections.
+From stdpp Require Import pmap gmap mapset.
 Local Open Scope positive_scope.
 
 (** * The tree data structure *)
diff --git a/theories/collections.v b/theories/collections.v
index b34b5cdc32165f44049f1b0eb4b7fb0c510a129c..f3ed2dd95d42d6e1a6c0db253dc7e4d2d804b1d9 100644
--- a/theories/collections.v
+++ b/theories/collections.v
@@ -3,7 +3,7 @@
 (** This file collects definitions and theorems on collections. Most
 importantly, it implements some tactics to automatically solve goals involving
 collections. *)
-Require Export prelude.base prelude.tactics prelude.orders.
+From stdpp Require Export base tactics orders.
 
 Instance collection_subseteq `{ElemOf A C} : SubsetEq C := λ X Y,
   ∀ x, x ∈ X → x ∈ Y.
diff --git a/theories/countable.v b/theories/countable.v
index 1635bac2995c67b00ab1cde29730f0d63e4f2d62..041da1952c4c682fcf4db1c5ccdbb65497b88012 100644
--- a/theories/countable.v
+++ b/theories/countable.v
@@ -1,6 +1,6 @@
 (* Copyright (c) 2012-2015, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
-Require Export prelude.list.
+From stdpp Require Export list.
 Local Open Scope positive.
 
 Class Countable A `{∀ x y : A, Decision (x = y)} := {
diff --git a/theories/decidable.v b/theories/decidable.v
index 817357de205b1063d3b2454c72a9a5b4d2fcb5c6..aecd65d0e3296783e38e81284479d21630a83c95 100644
--- a/theories/decidable.v
+++ b/theories/decidable.v
@@ -3,7 +3,7 @@
 (** This file collects theorems, definitions, tactics, related to propositions
 with a decidable equality. Such propositions are collected by the [Decision]
 type class. *)
-Require Export prelude.proof_irrel.
+From stdpp Require Export proof_irrel.
 
 Hint Extern 200 (Decision _) => progress (lazy beta) : typeclass_instances.
 
diff --git a/theories/error.v b/theories/error.v
index b2eb18522711f044d7ffeed32bc2d86f4860b10b..6a6ddbdb23f72d29d9d1bfc412004a969163396a 100644
--- a/theories/error.v
+++ b/theories/error.v
@@ -1,6 +1,6 @@
 (* Copyright (c) 2012-2015, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
-Require Export prelude.list.
+From stdpp Require Export list.
 
 Definition error (S E A : Type) : Type := S → E + (A * S).
 
diff --git a/theories/fin_collections.v b/theories/fin_collections.v
index 7fe2529002d66e54dfd9ef5eb0d1376fc485a107..7dbbb4b48f15e4b85511bb699cbaa4ce3f47739b 100644
--- a/theories/fin_collections.v
+++ b/theories/fin_collections.v
@@ -3,8 +3,9 @@
 (** This file collects definitions and theorems on finite collections. Most
 importantly, it implements a fold and size function and some useful induction
 principles on finite collections . *)
-Require Import Permutation prelude.relations prelude.listset.
-Require Export prelude.numbers prelude.collections.
+From Coq Require Import Permutation.
+From stdpp Require Import relations listset.
+From stdpp Require Export numbers collections.
 
 Instance collection_size `{Elements A C} : Size C := length ∘ elements.
 Definition collection_fold `{Elements A C} {B}
diff --git a/theories/fin_map_dom.v b/theories/fin_map_dom.v
index 7c4e76974aeb81104a670882203b2ecaef846b60..343ea76649af6eaa6c407260d6954d29ef6039f1 100644
--- a/theories/fin_map_dom.v
+++ b/theories/fin_map_dom.v
@@ -3,7 +3,7 @@
 (** This file provides an axiomatization of the domain function of finite
 maps. We provide such an axiomatization, instead of implementing the domain
 function in a generic way, to allow more efficient implementations. *)
-Require Export prelude.collections prelude.fin_maps.
+From stdpp Require Export collections fin_maps.
 
 Class FinMapDom K M D `{FMap M,
     ∀ A, Lookup K A (M A), ∀ A, Empty (M A), ∀ A, PartialAlter K A (M A),
diff --git a/theories/fin_maps.v b/theories/fin_maps.v
index 9ed4f6986413411e56105cccc07f146e53852ce0..1699bd0e29b57211423ca2d201eff452de00b725 100644
--- a/theories/fin_maps.v
+++ b/theories/fin_maps.v
@@ -4,8 +4,8 @@
 finite maps and collects some theory on it. Most importantly, it proves useful
 induction principles for finite maps and implements the tactic
 [simplify_map_equality] to simplify goals involving finite maps. *)
-Require Import Permutation.
-Require Export prelude.relations prelude.vector prelude.orders.
+From Coq Require Import Permutation.
+From stdpp Require Export relations vector orders.
 
 (** * Axiomatization of finite maps *)
 (** We require Leibniz equality to be extensional on finite maps. This of
diff --git a/theories/finite.v b/theories/finite.v
index a529a1a040d7aeedf44ac7736e61a18ef337a88d..11586869d1c7e3aa04d6a607fdaf403c8988858d 100644
--- a/theories/finite.v
+++ b/theories/finite.v
@@ -1,6 +1,6 @@
 (* Copyright (c) 2012-2015, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
-Require Export prelude.countable prelude.list.
+From stdpp Require Export countable list.
 
 Class Finite A `{∀ x y : A, Decision (x = y)} := {
   enum : list A;
diff --git a/theories/gmap.v b/theories/gmap.v
index a3466c307cd76edcf60291c4f6e04c7eee7bf1a3..dd86e07c5450100107f180a466acd576c10a689a 100644
--- a/theories/gmap.v
+++ b/theories/gmap.v
@@ -2,8 +2,8 @@
 (* This file is distributed under the terms of the BSD license. *)
 (** This file implements finite maps and finite sets with keys of any countable
 type. The implementation is based on [Pmap]s, radix-2 search trees. *)
-Require Export prelude.countable prelude.fin_maps prelude.fin_map_dom.
-Require Import prelude.pmap prelude.mapset.
+From stdpp Require Export countable fin_maps fin_map_dom.
+From stdpp Require Import pmap mapset.
 
 (** * The data structure *)
 (** We pack a [Pmap] together with a proof that ensures that all keys correspond
diff --git a/theories/hashset.v b/theories/hashset.v
index c6327dad01118aa1f8b86fd9ea91aa17e1fdd4ce..aa375074a449c30d82958662eea587aaf0721b3a 100644
--- a/theories/hashset.v
+++ b/theories/hashset.v
@@ -3,8 +3,8 @@
 (** This file implements finite set using hash maps. Hash sets are represented
 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. *)
-Require Export prelude.fin_maps prelude.listset.
-Require Import prelude.zmap.
+From stdpp Require Export fin_maps listset.
+From stdpp Require Import zmap.
 
 Record hashset {A} (hash : A → Z) := Hashset {
   hashset_car : Zmap (list A);
diff --git a/theories/lexico.v b/theories/lexico.v
index 31ddc33ba9194a8e9100381dd3af53672cfbaed4..05c5313842650a5da292ea9fe45bc910842524f7 100644
--- a/theories/lexico.v
+++ b/theories/lexico.v
@@ -2,7 +2,7 @@
 (* This file is distributed under the terms of the BSD license. *)
 (** 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. *)
-Require Import prelude.numbers.
+From stdpp Require Import numbers.
 
 Notation cast_trichotomy T :=
   match T with
diff --git a/theories/list.v b/theories/list.v
index 34645eeb4909690a848d83e13bfa71310b0250c0..0614f0abd28b24ae418bd08c61bfc8e2df7c41df 100644
--- a/theories/list.v
+++ b/theories/list.v
@@ -2,8 +2,8 @@
 (* This file is distributed under the terms of the BSD license. *)
 (** This file collects general purpose definitions and theorems on lists that
 are not in the Coq standard library. *)
-Require Export Permutation.
-Require Export prelude.numbers prelude.base prelude.decidable prelude.option.
+From Coq Require Export Permutation.
+From stdpp Require Export numbers base decidable option.
 
 Arguments length {_} _.
 Arguments cons {_} _ _.
diff --git a/theories/listset.v b/theories/listset.v
index 147715aa2814890804fbbccc81cffc867f46ca6f..bcf8ed894286c1cda135db82537952687c7719ad 100644
--- a/theories/listset.v
+++ b/theories/listset.v
@@ -2,7 +2,7 @@
 (* This file is distributed under the terms of the BSD license. *)
 (** This file implements finite set as unordered lists without duplicates
 removed. This implementation forms a monad. *)
-Require Export prelude.base prelude.decidable prelude.collections prelude.list.
+From stdpp Require Export base decidable collections list.
 
 Record listset A := Listset { listset_car: list A }.
 Arguments listset_car {_} _.
diff --git a/theories/listset_nodup.v b/theories/listset_nodup.v
index ab4341a837bcff16421b1dc8a2958a5ca10fa5ed..d685086268ca78ca308e4206f313043a229924f1 100644
--- a/theories/listset_nodup.v
+++ b/theories/listset_nodup.v
@@ -3,7 +3,7 @@
 (** This file implements finite as unordered lists without duplicates.
 Although this implementation is slow, it is very useful as decidable equality
 is the only constraint on the carrier set. *)
-Require Export prelude.base prelude.decidable prelude.collections prelude.list.
+From stdpp Require Export base decidable collections list.
 
 Record listset_nodup A := ListsetNoDup {
   listset_nodup_car : list A; listset_nodup_prf : NoDup listset_nodup_car
diff --git a/theories/mapset.v b/theories/mapset.v
index 317b6fc7b7295ee1a802e1ffb8fe5803de9ee3a0..bdf6c6bfc1d94fe893ee1c9c9ac252a3dde3b20e 100644
--- a/theories/mapset.v
+++ b/theories/mapset.v
@@ -3,7 +3,7 @@
 (** This files gives an implementation of finite sets using finite maps with
 elements of the unit type. Since maps enjoy extensional equality, the
 constructed finite sets do so as well. *)
-Require Export prelude.fin_map_dom.
+From stdpp Require Export fin_map_dom.
 
 Record mapset (M : Type → Type) : Type :=
   Mapset { mapset_car: M (unit : Type) }.
diff --git a/theories/natmap.v b/theories/natmap.v
index 309520af7969a9d98a271d83bfc578680560edb6..a8ebf6e2e445dc74f1621722adb81d08acc70759 100644
--- a/theories/natmap.v
+++ b/theories/natmap.v
@@ -3,7 +3,7 @@
 (** This files implements a type [natmap A] of finite maps whose keys range
 over Coq's data type of unary natural numbers [nat]. The implementation equips
 a list with a proof of canonicity. *)
-Require Import prelude.fin_maps prelude.mapset.
+From stdpp Require Import fin_maps mapset.
 
 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 1adc1a636600c5cf8f45062e781b11b323879092..f9f143413abc416fba9c4e0636ccf552603128d0 100644
--- a/theories/nmap.v
+++ b/theories/nmap.v
@@ -2,8 +2,8 @@
 (* This file is distributed under the terms of the BSD license. *)
 (** This files extends the implementation of finite over [positive] to finite
 maps whose keys range over Coq's data type of binary naturals [N]. *)
-Require Import prelude.pmap prelude.mapset.
-Require Export prelude.prelude prelude.fin_maps.
+From stdpp Require Import pmap mapset.
+From stdpp Require Export prelude fin_maps.
 
 Local Open Scope N_scope.
 
diff --git a/theories/numbers.v b/theories/numbers.v
index e2e728b3ffc7c56126783c42c4033561dcdb1115..d688ae71e4b72a9c15ab90f9cf832d4afa7fb7c7 100644
--- a/theories/numbers.v
+++ b/theories/numbers.v
@@ -3,9 +3,9 @@
 (** This file collects some trivial facts on the Coq types [nat] and [N] for
 natural numbers, and the type [Z] for integers. It also declares some useful
 notations. *)
-Require Export Eqdep PArith NArith ZArith NPeano.
-Require Import QArith Qcanon.
-Require Export prelude.base prelude.decidable prelude.option.
+From Coq Require Export Eqdep PArith NArith ZArith NPeano.
+From Coq Require Import QArith Qcanon.
+From stdpp Require Export base decidable option.
 Open Scope nat_scope.
 
 Coercion Z.of_nat : nat >-> Z.
@@ -50,7 +50,7 @@ Proof.
     * clear nat_le_pi. intros; exfalso; auto with lia.
     * injection 1. intros Hy. by case (nat_le_pi x y p y' q Hy). }
   intros x y p q.
-  by apply (eq_dep_eq_dec (λ x y, decide (x = y))), aux.
+  by apply (Eqdep_dec.eq_dep_eq_dec (λ x y, decide (x = y))), aux.
 Qed.
 Instance nat_lt_pi: ∀ x y : nat, ProofIrrel (x < y).
 Proof. apply _. Qed.
diff --git a/theories/option.v b/theories/option.v
index f6e217602c1ec45725c6192d87cfd5c5d4f18f1e..e7f14fd1cf7ca1ad798a824c28918e3fcf491263 100644
--- a/theories/option.v
+++ b/theories/option.v
@@ -2,7 +2,7 @@
 (* This file is distributed under the terms of the BSD license. *)
 (** This file collects general purpose definitions and theorems on the option
 data type that are not in the Coq standard library. *)
-Require Export prelude.base prelude.tactics prelude.decidable.
+From stdpp Require Export base tactics decidable.
 
 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 2203764820e562bfe45d07183b24bd03b5003f2f..cf17a85deb55f231fa537a7442a7250c09f35e46 100644
--- a/theories/orders.v
+++ b/theories/orders.v
@@ -2,8 +2,8 @@
 (* This file is distributed under the terms of the BSD license. *)
 (** This file collects common properties of pre-orders and semi lattices. This
 theory will mainly be used for the theory on collections and finite maps. *)
-Require Export Sorted.
-Require Export prelude.base prelude.decidable prelude.tactics prelude.list.
+From Coq Require Export Sorted.
+From stdpp Require Export base decidable tactics list.
 
 (** * Arbitrary pre-, parial and total orders *)
 (** Properties about arbitrary pre-, partial, and total orders. We do not use
diff --git a/theories/pmap.v b/theories/pmap.v
index 00cd7e874b9352e3c70555374a3070de2a2c2b14..5875a799125930d54fadd166570ab0ddb1265c69 100644
--- a/theories/pmap.v
+++ b/theories/pmap.v
@@ -7,8 +7,9 @@ trees (uncompressed Patricia trees) and guarantees logarithmic-time operations.
 However, we extend Leroy's implementation by packing the trees into a Sigma
 type such that canonicity of representation is ensured. This is necesarry for
 Leibniz equality to become extensional. *)
-Require Import PArith prelude.mapset.
-Require Export prelude.fin_maps.
+From Coq Require Import PArith.
+From stdpp Require Import mapset.
+From stdpp Require Export fin_maps.
 
 Local Open Scope positive_scope.
 Local Hint Extern 0 (@eq positive _ _) => congruence.
diff --git a/theories/prelude.v b/theories/prelude.v
index 020b6535b59d0ef3cb41240e9b465d2c2c36fa3d..34c65d5a15868eae9d6d3f6c8d10be0a4a086f31 100644
--- a/theories/prelude.v
+++ b/theories/prelude.v
@@ -1,16 +1,16 @@
 (* Copyright (c) 2012-2015, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
-Require Export
-  prelude.base
-  prelude.tactics
-  prelude.decidable
-  prelude.orders
-  prelude.option
-  prelude.vector
-  prelude.numbers
-  prelude.relations
-  prelude.collections
-  prelude.fin_collections
-  prelude.listset
-  prelude.list
-  prelude.lexico.
+From stdpp Require Export
+  base
+  tactics
+  decidable
+  orders
+  option
+  vector
+  numbers
+  relations
+  collections
+  fin_collections
+  listset
+  list
+  lexico.
diff --git a/theories/pretty.v b/theories/pretty.v
index b246ec428a2cfb16d918d51ed796d59efc90e10b..c7b8d50e16c7e3edc2456a636b0998dfe7c4552f 100644
--- a/theories/pretty.v
+++ b/theories/pretty.v
@@ -1,8 +1,8 @@
 (* Copyright (c) 2012-2015, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
-Require Export prelude.strings.
-Require Import prelude.relations.
-Require Import Ascii.
+From stdpp Require Export strings.
+From stdpp Require Import relations.
+From Coq Require Import Ascii.
 
 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 d1eab74108d6ae67712723a6bbbc4dbcd1eea6a2..afb9297a926119002f2c9b94da543f1628be5604 100644
--- a/theories/proof_irrel.v
+++ b/theories/proof_irrel.v
@@ -1,7 +1,8 @@
 (* Copyright (c) 2012-2015, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
 (** This file collects facts on proof irrelevant types/propositions. *)
-Require Export Eqdep_dec prelude.tactics.
+From Coq Require Import Eqdep_dec.
+From stdpp Require Export tactics.
 
 Hint Extern 200 (ProofIrrel _) => progress (lazy beta) : typeclass_instances.
 
diff --git a/theories/relations.v b/theories/relations.v
index cd81c198f12413881742bf04fc758313e5ca68bc..aa94487f5c0388513d55939ce841bc588bfd9841 100644
--- a/theories/relations.v
+++ b/theories/relations.v
@@ -4,8 +4,8 @@
 These are particularly useful as we define the operational semantics as a
 small step semantics. This file defines a hint database [ars] containing
 some theorems on abstract rewriting systems. *)
-Require Import Wf_nat.
-Require Export prelude.tactics prelude.base.
+From Coq Require Import Wf_nat.
+From stdpp Require Export tactics base.
 
 (** * Definitions *)
 Section definitions.
diff --git a/theories/sets.v b/theories/sets.v
index bd80b8eeb58b3374962750361f4591785cf0210d..f2543ebeb5bda3ac1a01ac7d98056d59f09b1eae 100644
--- a/theories/sets.v
+++ b/theories/sets.v
@@ -1,7 +1,7 @@
 (* Copyright (c) 2012-2015, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
 (** This file implements sets as functions into Prop. *)
-Require Export prelude.prelude.
+From stdpp Require Export prelude.
 
 Record set (A : Type) : Type := mkSet { set_car : A → Prop }.
 Arguments mkSet {_} _.
diff --git a/theories/streams.v b/theories/streams.v
index c660562cf3debb5545aaa6f1b901877e30cd7e5d..8fe09cc5ad1cfa1ed4c50ae3195e87e724a988c4 100644
--- a/theories/streams.v
+++ b/theories/streams.v
@@ -1,6 +1,6 @@
 (* Copyright (c) 2012-2015, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
-Require Export prelude.tactics.
+From stdpp Require Export tactics.
 
 CoInductive stream (A : Type) : Type := scons : A → stream A → stream A.
 Arguments scons {_} _ _.
diff --git a/theories/stringmap.v b/theories/stringmap.v
index 58c8862db6ee2fa19e0fa0622f26ffc496b49bd6..e320a190ff354545b3445518daf39439a1bf4350 100644
--- a/theories/stringmap.v
+++ b/theories/stringmap.v
@@ -4,8 +4,8 @@
 range over Coq's data type of strings [string]. The implementation uses radix-2
 search trees (uncompressed Patricia trees) as implemented in the file [pmap]
 and guarantees logarithmic-time operations. *)
-Require Export prelude.fin_maps prelude.pretty.
-Require Import prelude.gmap.
+From stdpp Require Export fin_maps pretty.
+From stdpp Require Import gmap.
 
 Notation stringmap := (gmap string).
 Notation stringset := (gset string).
@@ -58,4 +58,4 @@ Fixpoint fresh_strings_of_set
   | S n =>
      let x := fresh_string_of_set s X in
      x :: fresh_strings_of_set s n ({[ x ]} ∪ X)
-  end%nat.
\ No newline at end of file
+  end%nat.
diff --git a/theories/strings.v b/theories/strings.v
index f2074f2488ae2cd40f474b2d1beca5c45efd0236..9fa1549957f2b8671bc40f4fe8904545d91f33fc 100644
--- a/theories/strings.v
+++ b/theories/strings.v
@@ -1,7 +1,8 @@
 (* Copyright (c) 2012-2015, Robbert Krebbers. *)
 (* This file is distributed under the terms of the BSD license. *)
-Require Import Ascii.
-Require Export String prelude.countable.
+From Coq Require Import Ascii.
+From Coq Require Export String.
+From stdpp Require Export countable.
 
 (** * Fix scopes *)
 Open Scope string_scope.
diff --git a/theories/tactics.v b/theories/tactics.v
index 357f4fbde4633748810a9bcf6c4982add7b316d9..aee81d73cc70a4a7b25f6952a2f72c33578a852d 100644
--- a/theories/tactics.v
+++ b/theories/tactics.v
@@ -2,9 +2,9 @@
 (* This file is distributed under the terms of the BSD license. *)
 (** This file collects general purpose tactics that are used throughout
 the development. *)
-Require Import Omega.
-Require Export Psatz.
-Require Export prelude.base.
+From Coq Require Import Omega.
+From Coq Require Export Psatz.
+From stdpp Require Export base.
 
 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 4b11efe4a387abef0b512d38a160589da779b483..18547be4bb82f705049ef2cd386bffac7d41ff49 100644
--- a/theories/vector.v
+++ b/theories/vector.v
@@ -5,7 +5,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. *)
-Require Import prelude.list prelude.finite.
+From stdpp Require Import list finite.
 Open Scope vector_scope.
 
 (** * The fin type *)
diff --git a/theories/zmap.v b/theories/zmap.v
index 04f016e1d6331c5fa019175a3dd1387a726b3fbf..d2621ba987941f84ec7cc9d11975d1aaa6610f97 100644
--- a/theories/zmap.v
+++ b/theories/zmap.v
@@ -2,8 +2,8 @@
 (* This file is distributed under the terms of the BSD license. *)
 (** This files extends the implementation of finite over [positive] to finite
 maps whose keys range over Coq's data type of binary naturals [Z]. *)
-Require Import prelude.pmap prelude.mapset.
-Require Export prelude.prelude prelude.fin_maps.
+From stdpp Require Import pmap mapset.
+From stdpp Require Export prelude fin_maps.
 Local Open Scope Z_scope.
 
 Record Zmap (A : Type) : Type :=