aboutsummaryrefslogtreecommitdiffhomepage
path: root/library/decl_kinds.ml
diff options
context:
space:
mode:
authorGravatar Emilio Jesus Gallego Arias <e+git@x80.org>2017-12-28 00:16:46 +0100
committerGravatar Emilio Jesus Gallego Arias <e+git@x80.org>2018-04-23 16:23:49 +0200
commit7dfac786626f8f6775dadc0df85360759584c976 (patch)
tree9c0355fb0dba4a48e14d0e5b316c66dfd416d685 /library/decl_kinds.ml
parent5c34cfa54ec1959758baa3dd283e2e30853380db (diff)
[api] Relocate `intf` modules according to dependency-order.
In a component-based source code organization of Coq `intf` doesn't fit very well, as it sits in bit of "limbo" between different components, and indeed, encourages developers to place types in sometimes random places wrt the hierarchy. For example, lower parts of the system reference `Vernacexpr`, which morally lives in a pretty higher part of the system. We move all the files in `intf` to the lowest place their dependencies can accommodate: - `Misctypes`: is used by Declaremod, thus it has to go in `library` or below. Ideally, this file would disappear. - `Evar_kinds`: it is used by files in `engine`, so that seems its proper placement. - `Decl_kinds`: used in `library`, seems like the right place. [could also be merged. - `Glob_term`: belongs to pretyping, where it is placed. - `Locus`: ditto. - `Pattern`: ditto. - `Genredexpr`: depended by a few modules in `pretyping`, seems like the righ place. - `Constrexpr`: used in `pretyping`, the use is a bit unfortunate and could be fixed, as this module should be declared in `interp` which is the one eliminating it. - `Vernacexpr`: same problem than `Constrexpr`; this one can be fixed as it contains stuff it shouldn't. The right place should be `parsing`. - `Extend`: Is placed in `pretyping` due to being used by `Vernacexpr`. - `Notation_term`: First place used is `interp`, seems like the right place. Additionally, for some files it could be worth to merge files of the form `Foo` with `Foo_ops` in the medium term, as to create proper ADT modules as done in the kernel with `Name`, etc...
Diffstat (limited to 'library/decl_kinds.ml')
-rw-r--r--library/decl_kinds.ml87
1 files changed, 87 insertions, 0 deletions
diff --git a/library/decl_kinds.ml b/library/decl_kinds.ml
new file mode 100644
index 000000000..0d3285311
--- /dev/null
+++ b/library/decl_kinds.ml
@@ -0,0 +1,87 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+(** Informal mathematical status of declarations *)
+
+type discharge = DoDischarge | NoDischarge
+
+type locality = Discharge | Local | Global
+
+type binding_kind = Explicit | Implicit
+
+type polymorphic = bool
+
+type private_flag = bool
+
+type cumulative_inductive_flag = bool
+
+type theorem_kind =
+ | Theorem
+ | Lemma
+ | Fact
+ | Remark
+ | Property
+ | Proposition
+ | Corollary
+
+type definition_object_kind =
+ | Definition
+ | Coercion
+ | SubClass
+ | CanonicalStructure
+ | Example
+ | Fixpoint
+ | CoFixpoint
+ | Scheme
+ | StructureComponent
+ | IdentityCoercion
+ | Instance
+ | Method
+ | Let
+
+type assumption_object_kind = Definitional | Logical | Conjectural
+
+(* [assumption_kind]
+
+ | Local | Global
+ ------------------------------------
+ Definitional | Variable | Parameter
+ Logical | Hypothesis | Axiom
+
+*)
+type assumption_kind = locality * polymorphic * assumption_object_kind
+
+type definition_kind = locality * polymorphic * definition_object_kind
+
+(** Kinds used in proofs *)
+
+type goal_object_kind =
+ | DefinitionBody of definition_object_kind
+ | Proof of theorem_kind
+
+type goal_kind = locality * polymorphic * goal_object_kind
+
+(** Kinds used in library *)
+
+type logical_kind =
+ | IsAssumption of assumption_object_kind
+ | IsDefinition of definition_object_kind
+ | IsProof of theorem_kind
+
+(** Recursive power of type declarations *)
+
+type recursivity_kind = Declarations.recursivity_kind =
+ | Finite (** = inductive *)
+ [@ocaml.deprecated "Please use [Declarations.Finite"]
+ | CoFinite (** = coinductive *)
+ [@ocaml.deprecated "Please use [Declarations.CoFinite"]
+ | BiFinite (** = non-recursive, like in "Record" definitions *)
+ [@ocaml.deprecated "Please use [Declarations.BiFinite"]
+[@@ocaml.deprecated "Please use [Declarations.recursivity_kind"]