File: tacred.mli

package info (click to toggle)
coq 8.20.1%2Bdfsg-1
  • links: PTS, VCS
  • area: main
  • in suites: sid, trixie
  • size: 44,116 kB
  • sloc: ml: 234,160; sh: 4,301; python: 3,270; ansic: 2,644; makefile: 882; lisp: 172; javascript: 63; xml: 24; sed: 2
file content (124 lines) | stat: -rw-r--r-- 4,851 bytes parent folder | download | duplicates (2)
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
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
(************************************************************************)
(*         *   The Coq Proof Assistant / The Coq Development Team       *)
(*  v      *         Copyright INRIA, CNRS and contributors             *)
(* <O___,, * (see version control and CREDITS file for authors & dates) *)
(*   \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)         *)
(************************************************************************)

open Names
open Environ
open Evd
open EConstr
open Reductionops
open Pattern
open Locus
open Ltac_pretype

(** Here the semantics is completely unclear.
   What does "Hint Unfold t" means when "t" is a parameter?
   Does the user mean "Unfold X.t" or does she mean "Unfold y"
   where X.t is later on instantiated with y? I choose the first
   interpretation (i.e. an evaluable reference is never expanded). *)
val subst_evaluable_reference :
  Mod_subst.substitution -> Evaluable.t -> Evaluable.t

type reduction_tactic_error =
    InvalidAbstraction of env * evar_map * constr * (env * Type_errors.type_error)

exception ReductionTacticError of reduction_tactic_error

(** {6 Reduction functions associated to tactics. } *)

(** Evaluable global reference *)

val is_evaluable : Environ.env -> Evaluable.t -> bool

exception NotEvaluableRef of GlobRef.t
val error_not_evaluable : ?loc:Loc.t -> GlobRef.t -> 'a

val evaluable_of_global_reference :
  Environ.env -> GlobRef.t -> Evaluable.t
(** Fails on opaque constants and variables
    (both those without bodies and those marked Opaque in the conversion oracle). *)

val soft_evaluable_of_global_reference :
  ?loc:Loc.t -> GlobRef.t -> Evaluable.t
(** Succeeds for any constant or variable even if marked opaque or otherwise not evaluable. *)

val global_of_evaluable_reference :
  Evaluable.t -> GlobRef.t

(** Red (returns None if nothing reducible) *)
val red_product : env -> evar_map -> constr -> constr option

(** Simpl *)
val simpl : reduction_function

(** Simpl only at the head *)
val whd_simpl : reduction_function

(** Hnf: like whd_simpl but force delta-reduction of constants that do
   not immediately hide a non reducible fix or cofix *)
val hnf_constr : reduction_function

(** Variant of the above that does not perform nf-βι *)
val hnf_constr0 : reduction_function

(** Unfold *)
val unfoldn :
  (occurrences * Evaluable.t) list ->  reduction_function

(** Fold *)
val fold_commands : constr list ->  reduction_function

(** Pattern *)
val pattern_occs : (occurrences * constr) list -> e_reduction_function

(** Rem: Lazy strategies are defined in Reduction *)

(** Call by value strategy (uses Closures) *)
val cbv_norm_flags : RedFlags.reds -> strong:bool -> reduction_function
  val cbv_beta : reduction_function
  val cbv_betaiota : reduction_function
  val cbv_betadeltaiota :  reduction_function
  val whd_compute :  reduction_function
  val compute :  reduction_function  (** = [cbv_betadeltaiota] *)

(** [reduce_to_atomic_ind env sigma t] puts [t] in the form [t'=(I args)]
   with [I] an inductive definition;
   returns [I] and [t'] or fails with a user error *)
val reduce_to_atomic_ind : env ->  evar_map -> types -> (inductive * EInstance.t) * types

(** [reduce_to_quantified_ind env sigma t] puts [t] in the form
   [t'=(x1:A1)..(xn:An)(I args)] with [I] an inductive definition;
   returns [I] and [t'] or fails with a user error *)
val reduce_to_quantified_ind : env ->  evar_map -> types -> (inductive * EInstance.t) * types

(** Same as {!reduce_to_quantified_ind} but more efficient because it does not
    return the normalized type. *)
val eval_to_quantified_ind : env -> evar_map -> types -> (inductive * EInstance.t)

(** [reduce_to_quantified_ref env sigma ref t] try to put [t] in the form
    [t'=(x1:A1)..(xn:An)(ref args)]. When this is not possible, if [allow_failure]
    is specified, [t] is unfolded until the point where this impossibility is plainly
    visible. Otherwise, it fails with user error. *)
val reduce_to_quantified_ref :
  ?allow_failure:bool -> env ->  evar_map -> GlobRef.t -> types -> types

val reduce_to_atomic_ref :
  ?allow_failure:bool -> env ->  evar_map -> GlobRef.t -> types -> types

val find_hnf_rectype :
  env ->  evar_map -> types -> (inductive * EInstance.t) * constr list

val contextually : bool -> occurrences * constr_pattern ->
  (patvar_map -> reduction_function) -> reduction_function

val e_contextually : bool -> occurrences * constr_pattern ->
  (patvar_map -> e_reduction_function) -> e_reduction_function

(** Errors if the inductive is not allowed for pattern-matching. **)
val check_privacy : env -> inductive -> unit