File: inductive.mli

package info (click to toggle)
coq 8.9.0-1
  • links: PTS, VCS
  • area: main
  • in suites: buster
  • size: 30,604 kB
  • sloc: ml: 192,230; sh: 2,585; python: 2,206; ansic: 1,878; makefile: 818; lisp: 202; xml: 24; sed: 2
file content (85 lines) | stat: -rw-r--r-- 3,104 bytes parent folder | download
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
(************************************************************************)
(*         *   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)         *)
(************************************************************************)

(*i*)
open Names
open Cic
open Environ
(*i*)

(*s Extracting an inductive type from a construction *)

val find_rectype : env -> constr -> pinductive * constr list

type mind_specif = mutual_inductive_body * one_inductive_body

(*s Fetching information in the environment about an inductive type.
    Raises [Not_found] if the inductive type is not found. *)
val lookup_mind_specif : env -> inductive -> mind_specif

val inductive_is_polymorphic : mutual_inductive_body -> bool

val inductive_is_cumulative : mutual_inductive_body -> bool

val type_of_inductive : env -> mind_specif puniverses -> constr

(* Return type as quoted by the user *)
val type_of_constructor : pconstructor -> mind_specif -> constr

val arities_of_specif : MutInd.t puniverses -> mind_specif -> constr array

(* [type_case_branches env (I,args) (p:A) c] computes useful types
   about the following Cases expression:
      <p>Cases (c :: (I args)) of b1..bn end
   It computes the type of every branch (pattern variables are
   introduced by products) and the type for the whole expression.
 *)
val type_case_branches :
  env -> pinductive * constr list -> constr * constr -> constr
    -> constr array * constr

(* Check a [case_info] actually correspond to a Case expression on the
   given inductive type. *)
val check_case_info : env -> inductive -> case_info -> unit

(*s Guard conditions for fix and cofix-points. *)
val check_fix : env -> fixpoint -> unit
val check_cofix : env -> cofixpoint -> unit

(*s Support for sort-polymorphic inductive types *)

val type_of_inductive_knowing_parameters :
  env -> mind_specif puniverses -> constr array -> constr

val max_inductive_sort : sorts array -> Univ.universe

val instantiate_universes : env -> rel_context ->
  template_arity -> constr array -> rel_context * sorts

(***************************************************************)
(* Debug *)

type size = Large | Strict
type subterm_spec =
    Subterm of (size * wf_paths)
  | Dead_code
  | Not_subterm
type guard_env =
  { env     : env;
    (* dB of last fixpoint *)
    rel_min : int;
    (* dB of variables denoting subterms *)
    genv    : subterm_spec Lazy.t list;
  }

type stack_element = |SClosure of guard_env*constr |SArg of subterm_spec Lazy.t
val subterm_specif : guard_env -> stack_element list -> constr -> subterm_spec
val branches_specif : guard_env -> subterm_spec Lazy.t -> case_info ->
  subterm_spec Lazy.t list array