File: genredexpr.mli

package info (click to toggle)
coq 8.20.1%2Bdfsg-1
  • links: PTS, VCS
  • area: main
  • in suites: forky, 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 (82 lines) | stat: -rw-r--r-- 2,570 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
(************************************************************************)
(*         *   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)         *)
(************************************************************************)

(** Reduction expressions *)

(** The parsing produces initially a list of [red_atom] *)

type 'a red_atom =
  | FBeta
  | FMatch
  | FFix
  | FCofix
  | FZeta
  | FConst of 'a list
  | FDeltaBut of 'a list
  | FHead

(** This list of atoms is immediately converted to a [glob_red_flag] *)

type strength = Norm | Head
(* someday we may add NotUnderBinders *)

type 'a glob_red_flag = {
  rStrength : strength;
  rBeta : bool;
  rMatch : bool;
  rFix : bool;
  rCofix : bool;
  rZeta : bool;
  rDelta : bool; (** true = delta all but rConst; false = delta only on rConst*)
  rConst : 'a list
}

(** Generic kinds of reductions *)

type ('b,'c,'occvar) red_context = ('occvar Locus.occurrences_gen * ('b,'c) Util.union) option

type ('a, 'b, 'c, 'occvar, 'flags) red_expr_gen0 =
  | Red
  | Hnf
  | Simpl of 'flags * ('b, 'c, 'occvar) red_context
  | Cbv of 'flags
  | Cbn of 'flags
  | Lazy of 'flags
  | Unfold of ('occvar Locus.occurrences_gen * 'b) list
  | Fold of 'a list
  | Pattern of ('occvar Locus.occurrences_gen * 'a) list
  | ExtraRedExpr of string
  | CbvVm of ('b, 'c, 'occvar) red_context
  | CbvNative of ('b, 'c, 'occvar) red_context

type ('a, 'b, 'c, 'occvar) red_expr_gen =
  ('a, 'b, 'c, 'occvar, 'b glob_red_flag) red_expr_gen0

type ('a,'b,'c,'occvar) may_eval =
  | ConstrTerm of 'a
  | ConstrEval of ('a,'b,'c,'occvar) red_expr_gen * 'a
  | ConstrContext of Names.lident * 'a
  | ConstrTypeOf of 'a

open Constrexpr

type r_trm = constr_expr
type r_pat = constr_pattern_expr
type r_cst = Libnames.qualid or_by_notation

type raw_red_expr = (r_trm, r_cst, r_pat, int Locus.or_var) red_expr_gen

type 'a and_short_name = 'a * Names.lident option

type g_trm = Genintern.glob_constr_and_expr
type g_pat = Genintern.glob_constr_pattern_and_expr
type g_cst = Evaluable.t and_short_name Locus.or_var

type glob_red_expr = (g_trm, g_cst, g_trm, int Locus.or_var) red_expr_gen