diff --git a/theories/assoc.v b/theories/assoc.v index c9e609ba16b5cafec7ab87a6167f03a30638e971..3b8730a5f8c955a2c2d152113a91993019ee0606 100644 --- a/theories/assoc.v +++ b/theories/assoc.v @@ -6,8 +6,8 @@ main advantage of these association lists compared to search trees, is that it has canonical representatives and thus extensional Leibniz equality. Compared to a naive unordered association list, the merge operation (used for unions, intersections, and difference) is also linear-time. *) -Require Import mapset. -Require Export fin_maps. +Require Import prelude.mapset. +Require Export prelude.fin_maps. (** Because the association list is sorted using [strict lexico] instead of [lexico], it automatically guarantees that no duplicates exist. *) diff --git a/theories/collections.v b/theories/collections.v index eaff59e01b0132b99b2a8bf59f6a423aecabfece..7caafe69dc78efb92f897d064d37501fb4b92054 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 base tactics orders. +Require Export prelude.base prelude.tactics prelude.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 c62eb23510f7b757cf97080083bb9909c6a4d995..628aa5523fc1716c95672d69ae0b4d54e4ee61eb 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 list. +Require Export prelude.list. Local Obligation Tactic := idtac. Local Open Scope positive. diff --git a/theories/decidable.v b/theories/decidable.v index 96720b24efcc5c7a8fe929f6c2b2df02e9ec8cdc..184372c8bf37317d0d2e82ca40da4ae3a3ed1f6c 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 proof_irrel. +Require Export prelude.proof_irrel. Hint Extern 200 (Decision _) => progress (lazy beta) : typeclass_instances. diff --git a/theories/error.v b/theories/error.v index 46fe0c57845e6c5aaddd5b4cb35a2315bdb3e488..0eb6c33f51b4baaf128cacc1fb15a47c210c6fd3 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 list. +Require Export prelude.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 f4e63f636e3e26fa67d75ea4b60305d15703dc2e..28591e60ab45be8aaccc7203a1d2a7f7193f87e2 100644 --- a/theories/fin_collections.v +++ b/theories/fin_collections.v @@ -3,8 +3,8 @@ (** 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 relations listset. -Require Export numbers collections. +Require Import Permutation prelude.relations prelude.listset. +Require Export prelude.numbers prelude.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 bbdbb8ea6077d104071e28c8d379cf3d8cc66fab..f69fabecb9bf92c308dd7b4982dd1530d5b22c34 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 collections fin_maps. +Require Export prelude.collections prelude.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 b77edcd478d8e9eb4b1c6ec55cd71dfbbe524ec1..7e12d0e4e70edff8153ef1c7d6afedf57953d163 100644 --- a/theories/fin_maps.v +++ b/theories/fin_maps.v @@ -5,7 +5,7 @@ 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 relations vector orders. +Require Export prelude.relations prelude.vector prelude.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 9aa8a459ebe3a1503bd2652bff77136442b7a222..7824001163d1cb787b7641ee2cf5dcb50ebd298b 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 countable list. +Require Export prelude.countable prelude.list. Obligation Tactic := idtac. Class Finite A `{∀ x y : A, Decision (x = y)} := { diff --git a/theories/hashset.v b/theories/hashset.v index 345ba837ae1c8a26a6024d87679384cb05260abf..f517f0e76b62a4d9a6e1e60948c852c6ff396a42 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 fin_maps listset. -Require Import zmap. +Require Export prelude.fin_maps prelude.listset. +Require Import prelude.zmap. Record hashset {A} (hash : A → Z) := Hashset { hashset_car : Zmap (list A); diff --git a/theories/lexico.v b/theories/lexico.v index e36b552ad473a263aea6638af86e2d73a1d5d6bd..31ddc33ba9194a8e9100381dd3af53672cfbaed4 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 numbers. +Require Import prelude.numbers. Notation cast_trichotomy T := match T with diff --git a/theories/list.v b/theories/list.v index 25fe629b709e003cf0389e400427f548d33bd597..2b39ffe56c550dfcebd5ce99aafbfb995a63d303 100644 --- a/theories/list.v +++ b/theories/list.v @@ -3,7 +3,7 @@ (** This file collects general purpose definitions and theorems on lists that are not in the Coq standard library. *) Require Export Permutation. -Require Export numbers base decidable option. +Require Export prelude.numbers prelude.base prelude.decidable prelude.option. Arguments length {_} _. Arguments cons {_} _ _. diff --git a/theories/listset.v b/theories/listset.v index 5aa61a4cd498b79ba180fc597a76f49d2fd1bcf1..147715aa2814890804fbbccc81cffc867f46ca6f 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 base decidable collections list. +Require Export prelude.base prelude.decidable prelude.collections prelude.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 7328d21ef36a128c008adbe11775367e318934b0..ab4341a837bcff16421b1dc8a2958a5ca10fa5ed 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 base decidable collections list. +Require Export prelude.base prelude.decidable prelude.collections prelude.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 d7e7c445e0c842b116282ebc20c22bd829cc8a9e..586a099aea70debe4fcfd024030a994d457073b3 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 fin_map_dom. +Require Export prelude.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 320b2cb87c2c3dcad9de9912591f8c0f03164279..309520af7969a9d98a271d83bfc578680560edb6 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 fin_maps mapset. +Require Import prelude.fin_maps prelude.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 d833c6b909991c9ad2d0fa07736fb949c55c4e8e..1adc1a636600c5cf8f45062e781b11b323879092 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 pmap mapset. -Require Export prelude fin_maps. +Require Import prelude.pmap prelude.mapset. +Require Export prelude.prelude prelude.fin_maps. Local Open Scope N_scope. diff --git a/theories/numbers.v b/theories/numbers.v index 886d9f768c3918bbc5f9bd862e70182b2bebd278..4694aaf73f08d9ec875a94bb1d5cef6d6ba4282c 100644 --- a/theories/numbers.v +++ b/theories/numbers.v @@ -5,7 +5,7 @@ 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 base decidable. +Require Export prelude.base prelude.decidable. Open Scope nat_scope. Coercion Z.of_nat : nat >-> Z. diff --git a/theories/option.v b/theories/option.v index a67260f46c918ef0a0d92701e03dd42435e6a6e3..ed2392de07f899703615aa5167b9505f04ea6839 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 base tactics decidable. +Require Export prelude.base prelude.tactics prelude.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/optionmap.v b/theories/optionmap.v index 614f96aa2dd4bece37dad3cc77e7e4ab03ea9929..4873518b86ae30691e5a1434c95442e313e3b726 100644 --- a/theories/optionmap.v +++ b/theories/optionmap.v @@ -1,7 +1,7 @@ (* Copyright (c) 2012-2015, Robbert Krebbers. *) (* This file is distributed under the terms of the BSD license. *) -Require Import mapset. -Require Export prelude fin_maps. +Require Import prelude.mapset. +Require Export prelude.prelude prelude.fin_maps. Record optionmap (M : Type → Type) (A : Type) : Type := OptionMap { optionmap_None : option A; optionmap_Some : M A }. diff --git a/theories/orders.v b/theories/orders.v index 7156831810861dc7029177ff31aa9990de134f8f..09ae83d099cc2ecacef0c52329196401c4817807 100644 --- a/theories/orders.v +++ b/theories/orders.v @@ -3,7 +3,7 @@ (** 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 base decidable tactics list. +Require Export prelude.base prelude.decidable prelude.tactics prelude.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 3cf4f3aa959657c7c89943f1c904856616a16fca..f11ab53d006e6e8a7d80fdd3942fa6be62333ed9 100644 --- a/theories/pmap.v +++ b/theories/pmap.v @@ -7,8 +7,8 @@ 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 mapset. -Require Export fin_maps. +Require Import PArith prelude.mapset. +Require Export prelude.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 0c00f512508add9f675a56174ec66706eedae65f..020b6535b59d0ef3cb41240e9b465d2c2c36fa3d 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 - base - tactics - decidable - orders - option - vector - numbers - relations - collections - fin_collections - listset - list - lexico. + 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. diff --git a/theories/pretty.v b/theories/pretty.v index 2bd7f21d685547ab973338a6bba9ee974c5a3b7c..f76b4e3ee4a8792ae663f7afbdddf808b4aa9e50 100644 --- a/theories/pretty.v +++ b/theories/pretty.v @@ -1,7 +1,7 @@ (* Copyright (c) 2012-2015, Robbert Krebbers. *) (* This file is distributed under the terms of the BSD license. *) -Require Export numbers option. -Require Import Ascii String relations. +Require Export prelude.numbers prelude.option. +Require Import Ascii String prelude.relations. Infix "+:+" := String.append (at level 60, right associativity) : C_scope. Arguments String.append _ _ : simpl never. diff --git a/theories/proof_irrel.v b/theories/proof_irrel.v index 6f065020aaad26eb2ab2e44ed4e0fda43122558a..d1eab74108d6ae67712723a6bbbc4dbcd1eea6a2 100644 --- a/theories/proof_irrel.v +++ b/theories/proof_irrel.v @@ -1,7 +1,7 @@ (* 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 tactics. +Require Export Eqdep_dec prelude.tactics. Hint Extern 200 (ProofIrrel _) => progress (lazy beta) : typeclass_instances. diff --git a/theories/relations.v b/theories/relations.v index da9523984bda004747e021bacb66f56b56925e0e..cd81c198f12413881742bf04fc758313e5ca68bc 100644 --- a/theories/relations.v +++ b/theories/relations.v @@ -5,7 +5,7 @@ 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 tactics base. +Require Export prelude.tactics prelude.base. (** * Definitions *) Section definitions. diff --git a/theories/sets.v b/theories/sets.v index 9e3d2d210d210d6e262e7ed360b490d931e7943b..bd80b8eeb58b3374962750361f4591785cf0210d 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. +Require Export prelude.prelude. Record set (A : Type) : Type := mkSet { set_car : A → Prop }. Arguments mkSet {_} _. @@ -28,4 +28,4 @@ Instance set_join : MJoin set := λ A (XX : set (set A)), Instance set_collection_monad : CollectionMonad set. Proof. by split; try apply _. Qed. -Global Opaque set_union set_intersection. \ No newline at end of file +Global Opaque set_union set_intersection. diff --git a/theories/streams.v b/theories/streams.v index af231e02f2a693092099986973b685ca499c43c6..c660562cf3debb5545aaa6f1b901877e30cd7e5d 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 tactics. +Require Export prelude.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 c519eff821418f160913a6ae6d342cbc43a75b99..b7fafeeca4cce7d0830c7e31029d4060fc50d1b7 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 fin_maps pretty. -Require Import Ascii String list pmap mapset. +Require Export prelude.fin_maps prelude.pretty. +Require Import Ascii String prelude.list prelude.pmap prelude.mapset. (** * Encoding and decoding *) (** In order to reuse or existing implementation of radix-2 search trees over diff --git a/theories/tactics.v b/theories/tactics.v index 70f4f94120ef13483dbbe32604226c0422c0f910..ece653ff4ad57a3b12c9aa2e85e7e70ebd3c5cd5 100644 --- a/theories/tactics.v +++ b/theories/tactics.v @@ -3,7 +3,7 @@ (** This file collects general purpose tactics that are used throughout the development. *) Require Export Psatz. -Require Export base. +Require Export prelude.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 6cacf063e7d7215aa41658b8fceed1876f63ae0a..da77f8fd936010db1d0f13fb1ec49117126231e1 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 list finite. +Require Import prelude.list prelude.finite. Open Scope vector_scope. (** * The fin type *) diff --git a/theories/zmap.v b/theories/zmap.v index f91c7fb4ef76d308cdc97ad9c518fad5cdff79b4..04f016e1d6331c5fa019175a3dd1387a726b3fbf 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 pmap mapset. -Require Export prelude fin_maps. +Require Import prelude.pmap prelude.mapset. +Require Export prelude.prelude prelude.fin_maps. Local Open Scope Z_scope. Record Zmap (A : Type) : Type :=