module-system.lagda.rst 12 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288
  1. ..
  2. ::
  3. {-# OPTIONS --allow-unsolved-metas --rewriting --sized-types #-}
  4. module language.module-system where
  5. open import language.built-ins
  6. .. _module-system:
  7. *************
  8. Module System
  9. *************
  10. .. _module-application:
  11. Module application
  12. ------------------
  13. .. _anonymous-modules:
  14. Anonymous modules
  15. -----------------
  16. .. _module-basics:
  17. Basics
  18. ------
  19. First let us introduce some terminology. A **definition** is a syntactic construction defining an entity such as a function or a datatype. A name is a string used to identify definitions. The same definition can have many names and at different points in the program it will have different names. It may also be the case that two definitions have the same name. In this case there will be an error if the name is used.
  20. The main purpose of the module system is to structure the way names are used in a program. This is done by organising the program in an hierarchical structure of modules where each module contains a number of definitions and submodules. For instance,
  21. ::
  22. module Main where
  23. module B where
  24. f : Nat → Nat
  25. f n = suc n
  26. g : Nat → Nat → Nat
  27. g n m = m
  28. Note that we use indentation to indicate which definitions are part of a module. In the example ``f`` is in the module ``Main.B`` and ``g`` is in ``Main``. How to refer to a particular definition is determined by where it is located in the module hierarchy. Definitions from an enclosing module are referred to by their given names as seen in the type of f above. To access a definition from outside its defining module a qualified name has to be used.
  29. ::
  30. module Main₂ where
  31. module B where
  32. f : Nat → Nat
  33. f n = suc n
  34. ff : Nat → Nat
  35. ff x = B.f (B.f x)
  36. To be able to use the short names for definitions in a module the module has to be opened.
  37. ::
  38. module Main₃ where
  39. module B where
  40. f : Nat → Nat
  41. f n = suc n
  42. open B
  43. ff : Nat → Nat
  44. ff x = f (f x)
  45. If ``A.qname`` refers to a definition ``d``, then after ``open A``, ``qname`` will also refer to ``d``. Note that ``qname`` can itself be a qualified name. Opening a module only introduces new names for a definition, it never removes the old names. The policy is to allow the introduction of ambiguous names, but give an error if an ambiguous name is used.
  46. Modules can also be opened within a local scope by putting the ``open B`` within a ``where`` clause:
  47. ::
  48. ff₁ : Nat → Nat
  49. ff₁ x = f (f x) where open B
  50. Private definitions
  51. -------------------
  52. To make a definition inaccessible outside its defining module it can be declared ``private``. A private definition is treated as a normal definition inside the module that defines it, but outside the module the definition has no name. In a dependently type setting there are some problems with private definitions---since the type checker performs computations, private names might show up in goals and error messages. Consider the following (contrived) example
  53. ::
  54. module Main₄ where
  55. module A where
  56. private
  57. IsZero’ : Nat → Set
  58. IsZero’ zero = ⊤
  59. IsZero’ (suc n) = ⊥
  60. IsZero : Nat → Set
  61. IsZero n = IsZero’ n
  62. open A
  63. prf : (n : Nat) → IsZero n
  64. prf n = ?
  65. The type of the goal ``?0`` is ``IsZero n`` which normalises to ``IsZero’ n``. The question is how to display this normal form to the user. At the point of ``?0`` there is no name for ``IsZero’``. One option could be try to fold the term and print ``IsZero n``. This is a very hard problem in general, so rather than trying to do this we make it clear to the user that ``IsZero’`` is something that is not in scope and print the goal as ``;Main₄.A.IsZero’ n``. The leading semicolon indicates that the entity is not in scope. The same technique is used for definitions that only have ambiguous names.
  66. In effect using private definitions means that, from the user’s perspective, we do not have subject reduction. This is just an illusion, however---the type checker has full access to all definitions.
  67. Name modifiers
  68. --------------
  69. An alternative to making definitions private is to exert finer control over what names are introduced when opening a module. This is done by qualifying an ``open`` statement with one or more of the modifiers ``using``, ``hiding``, or ``renaming``. You can combine both ``using`` and ``hiding`` with ``renaming``, but not with each other. The effect of
  70. .. code-block:: agda
  71. open A using (xs) renaming (ys to zs)
  72. is to introduce the names ``xs`` and ``zs`` where ``xs`` refers to the same definition as ``A.xs`` and ``zs`` refers to ``A.ys``. We do not permit ``xs``, ``ys`` and ``zs`` to overlap. The other forms of opening are defined in terms of this one.
  73. An omitted ``renaming`` modifier is equivalent to an empty renaming.
  74. To refer to a module ``M`` inside ``A`` you write ``module M``. For instance,
  75. .. code-block:: agda
  76. open A using (module M)
  77. Since 2.6.1: The fixity of an operator can be set or changed in a ``renaming`` directive::
  78. module ExampleRenamingFixity where
  79. module ArithFoo where
  80. postulate
  81. A : Set
  82. _&_ _^_ : A → A → A
  83. infixr 10 _&_
  84. open ArithFoo renaming (_&_ to infixl 8 _+_; _^_ to infixl 10 _^_)
  85. Here, we change the fixity of ``_&_`` while renaming it to ``_+_``, and assign a new fixity to ``_^_`` which has the default fixity in module ``ArithFoo``.
  86. Re-exporting names
  87. ------------------
  88. A useful feature is the ability to re-export names from another module. For instance, one may want to create a module to collect the definitions from several other modules. This is achieved by qualifying the open statement with the public keyword:
  89. ::
  90. module Example where
  91. module Nat₁ where
  92. data Nat₁ : Set where
  93. zero : Nat₁
  94. suc : Nat₁ → Nat₁
  95. module Bool₁ where
  96. data Bool₁ : Set where
  97. true false : Bool₁
  98. module Prelude where
  99. open Nat₁ public
  100. open Bool₁ public
  101. isZero : Nat₁ → Bool₁
  102. isZero zero = true
  103. isZero (suc _) = false
  104. The module ``Prelude`` above exports the names ``Nat``, ``zero``, ``Bool``, etc., in addition to ``isZero``.
  105. Parameterised modules
  106. ---------------------
  107. So far, the module system features discussed have dealt solely with scope manipulation. We now turn our attention to some more advanced features.
  108. It is sometimes useful to be able to work temporarily in a given signature. For instance, when defining functions for sorting lists it is convenient to assume a set of list elements ``A`` and an ordering over ``A``. In Coq this can be done in two ways: using a functor, which is essentially a function between modules, or using a section. A section allows you to abstract some arguments from several definitions at once. We introduce parameterised modules analogous to sections in Coq. When declaring a module you can give a telescope of module parameters which are abstracted from all the definitions in the module. For instance, a simple implementation of a sorting function looks like this:
  109. ::
  110. module Sort (A : Set)(_≤_ : A → A → Bool) where
  111. insert : A → List A → List A
  112. insert x [] = x ∷ []
  113. insert x (y ∷ ys) with x ≤ y
  114. insert x (y ∷ ys) | true = x ∷ y ∷ ys
  115. insert x (y ∷ ys) | false = y ∷ insert x ys
  116. sort : List A → List A
  117. sort [] = []
  118. sort (x ∷ xs) = insert x (sort xs)
  119. As mentioned parametrising a module has the effect of abstracting the parameters over the definitions in the module, so outside the Sort module we have
  120. .. code-block:: agda
  121. Sort.insert : (A : Set)(_≤_ : A → A → Bool) →
  122. A → List A → List A
  123. Sort.sort : (A : Set)(_≤_ : A → A → Bool) →
  124. List A → List A
  125. For function definitions, explicit module parameter become explicit arguments to the abstracted function, and implicit parameters become implicit arguments. For constructors, however, the parameters are always implicit arguments. This is a consequence of the fact that module parameters are turned into datatype parameters, and the datatype parameters are implicit arguments to the constructors. It also happens to be the reasonable thing to do.
  126. Something which you cannot do in Coq is to apply a section to its arguments. We allow this through the module application statement. In our example:
  127. .. code-block:: agda
  128. module SortNat = Sort Nat leqNat
  129. This will define a new module SortNat as follows
  130. .. code-block:: agda
  131. module SortNat where
  132. insert : Nat → List Nat → List Nat
  133. insert = Sort.insert Nat leqNat
  134. sort : List Nat → List Nat
  135. sort = Sort.sort Nat leqNat
  136. The new module can also be parameterised, and you can use name modifiers to control what definitions from the original module are applied and what names they have in the new module. The general form of a module application is
  137. .. code-block:: agda
  138. module M1 Δ = M2 terms modifiers
  139. A common pattern is to apply a module to its arguments and then open the resulting module. To simplify this we introduce the short-hand
  140. .. code-block:: agda
  141. open module M1 Δ = M2 terms [public] mods
  142. for
  143. .. code-block:: agda
  144. module M1 Δ = M2 terms mods
  145. open M1 [public]
  146. Splitting a program over multiple files
  147. ---------------------------------------
  148. When building large programs it is crucial to be able to split the program over multiple files and to not have to type check and compile all the files for every change. The module system offers a structured way to do this. We define a program to be a collection of modules, each module being defined in a separate file. To gain access to a module defined in a different file you can import the module:
  149. .. code-block:: agda
  150. import M
  151. In order to implement this we must be able to find the file in which a module is defined. To do this we require that the top-level module ``A.B.C`` is defined in the file ``C.agda`` in the directory ``A/B/``. One could imagine instead to give a file name to the import statement, but this would mean cluttering the program with details about the file system which is not very nice.
  152. When importing a module ``M``, the module and its contents are brought into scope as if the module had been defined in the current file. In order to get access to the unqualified names of the module contents it has to be opened. Similarly to module application we introduce the short-hand
  153. .. code-block:: agda
  154. open import M
  155. for
  156. .. code-block:: agda
  157. import M
  158. open M
  159. Sometimes the name of an imported module clashes with a local module. In this case it is possible to import the module under a different name.
  160. .. code-block:: agda
  161. import M as M’
  162. It is also possible to attach modifiers to import statements, limiting or changing what names are visible from inside the module.
  163. Note that modifiers attached to ``open import`` statements apply to the ``open`` statement and not the ``import`` statement.
  164. Datatype modules and record modules
  165. -----------------------------------
  166. When you define a datatype it also defines a module so constructors can now be referred to qualified by their data type.
  167. For instance, given::
  168. module DatatypeModules where
  169. data Nat₂ : Set where
  170. zero : Nat₂
  171. suc : Nat₂ → Nat₂
  172. data Fin : Nat₂ → Set where
  173. zero : ∀ {n} → Fin (suc n)
  174. suc : ∀ {n} → Fin n → Fin (suc n)
  175. you can refer to the constructors unambiguously as ``Nat₂.zero``, ``Nat₂.suc``, ``Fin.zero``, and ``Fin.suc`` (``Nat₂`` and ``Fin`` are modules containing the respective constructors). Example:
  176. ::
  177. inj : (n m : Nat₂) → Nat₂.suc n ≡ suc m → n ≡ m
  178. inj .m m refl = refl
  179. Previously you had to write something like
  180. ::
  181. inj₁ : (n m : Nat₂) → _≡_ {A = Nat₂} (suc n) (suc m) → n ≡ m
  182. inj₁ .m m refl = refl
  183. to make the type checker able to figure out that you wanted the natural number suc in this case.
  184. Also record declarations define a corresponding module, see
  185. :ref:`record-modules`.