package coq

  1. Overview
  2. Docs
Legend:
Library
Module
Module type
Parameter
Class
Class type

Untyped intermediate terms

glob_constr comes after constr_expr and before constr.

Resolution of names, insertion of implicit arguments placeholder, and notations are done, but coercions, inference of implicit arguments and pattern-matching compilation are not.

type existential_name = Names.Id.t

Sorts

type glob_sort_name =
  1. | GSProp
    (*

    representation of SProp literal

    *)
  2. | GProp
    (*

    representation of Prop level

    *)
  3. | GSet
    (*

    representation of Set level

    *)
  4. | GUniv of Univ.Level.t
  5. | GLocalUniv of Names.lident
    (*

    Locally bound universes (may also be nonstrict declaration)

    *)
  6. | GRawUniv of Univ.Level.t
    (*

    Hack for funind, DO NOT USE

    Note that producing the similar Constrexpr.CRawType for printing is OK, just don't try to reinterp it.

    *)
type 'a glob_sort_gen =
  1. | UAnonymous of {
    1. rigid : bool;
    }
    (*

    not rigid = unifiable by minimization

    *)
  2. | UNamed of 'a
type glob_level = glob_sort_name glob_sort_gen

levels, occurring in universe instances

type glob_sort = (glob_sort_name * int) list glob_sort_gen

sort expressions

type glob_recarg = int option
and glob_fix_kind =
  1. | GFix of glob_recarg array * int
  2. | GCoFix of int
type 'a cases_pattern_r =
  1. | PatVar of Names.Name.t
  2. | PatCstr of Names.constructor * 'a cases_pattern_g list * Names.Name.t
    (*

    PatCstr(p,C,l,x) = "|'C' 'l' as 'x'"

    *)

The kind of patterns that occurs in "match ... with ... end"

locs here refers to the ident's location, not whole pat

and 'a cases_pattern_g = ('a cases_pattern_r, 'a) DAst.t
type cases_pattern = [ `any ] cases_pattern_g
type binding_kind =
  1. | Explicit
  2. | MaxImplicit
  3. | NonMaxImplicit
type 'a glob_constr_r =
  1. | GRef of Names.GlobRef.t * glob_level list option
    (*

    An identifier that represents a reference to an object defined either in the (global) environment or in the (local) context.

    *)
  2. | GVar of Names.Id.t
    (*

    An identifier that cannot be regarded as "GRef". Bound variables are typically represented this way.

    *)
  3. | GEvar of existential_name CAst.t * (Names.lident * 'a glob_constr_g) list
  4. | GPatVar of Evar_kinds.matching_var_kind
    (*

    Used for patterns only

    *)
  5. | GApp of 'a glob_constr_g * 'a glob_constr_g list
  6. | GLambda of Names.Name.t * binding_kind * 'a glob_constr_g * 'a glob_constr_g
  7. | GProd of Names.Name.t * binding_kind * 'a glob_constr_g * 'a glob_constr_g
  8. | GLetIn of Names.Name.t * 'a glob_constr_g * 'a glob_constr_g option * 'a glob_constr_g
  9. | GCases of Constr.case_style * 'a glob_constr_g option * 'a tomatch_tuples_g * 'a cases_clauses_g
    (*

    GCases(style,r,tur,cc) = "match 'tur' return 'r' with 'cc'" (in MatchStyle)

    *)
  10. | GLetTuple of Names.Name.t list * Names.Name.t * 'a glob_constr_g option * 'a glob_constr_g * 'a glob_constr_g
  11. | GIf of 'a glob_constr_g * Names.Name.t * 'a glob_constr_g option * 'a glob_constr_g * 'a glob_constr_g
  12. | GRec of glob_fix_kind * Names.Id.t array * 'a glob_decl_g list array * 'a glob_constr_g array * 'a glob_constr_g array
  13. | GSort of glob_sort
  14. | GHole of Evar_kinds.t * Namegen.intro_pattern_naming_expr * Genarg.glob_generic_argument option
  15. | GCast of 'a glob_constr_g * Constr.cast_kind * 'a glob_constr_g
  16. | GProj of Names.Constant.t * glob_level list option * 'a glob_constr_g list * 'a glob_constr_g
  17. | GInt of Uint63.t
  18. | GFloat of Float64.t
  19. | GArray of glob_level list option * 'a glob_constr_g array * 'a glob_constr_g * 'a glob_constr_g

Representation of an internalized (or in other words globalized) term.

and 'a glob_constr_g = ('a glob_constr_r, 'a) DAst.t
and 'a glob_decl_g = Names.Name.t * binding_kind * 'a glob_constr_g option * 'a glob_constr_g
and 'a predicate_pattern_g = Names.Name.t * (Names.inductive * Names.Name.t list) CAst.t option

(na,id) = "as 'na' in 'id'" where if id is Some(l,I,k,args).

and 'a tomatch_tuple_g = 'a glob_constr_g * 'a predicate_pattern_g
and 'a tomatch_tuples_g = 'a tomatch_tuple_g list
and 'a cases_clause_g = (Names.Id.t list * 'a cases_pattern_g list * 'a glob_constr_g) CAst.t

(p,il,cl,t) = "|'cl' => 't'". Precondition: the free variables of t are members of il.

and 'a cases_clauses_g = 'a cases_clause_g list
type glob_constr = [ `any ] glob_constr_g
type tomatch_tuple = [ `any ] tomatch_tuple_g
type tomatch_tuples = [ `any ] tomatch_tuples_g
type cases_clause = [ `any ] cases_clause_g
type cases_clauses = [ `any ] cases_clauses_g
type glob_decl = [ `any ] glob_decl_g
type predicate_pattern = [ `any ] predicate_pattern_g
type any_glob_constr =
  1. | AnyGlobConstr : 'r glob_constr_g -> any_glob_constr
type 'a disjunctive_cases_clause_g = (Names.Id.t list * 'a cases_pattern_g list list * 'a glob_constr_g) CAst.t
type 'a disjunctive_cases_clauses_g = 'a disjunctive_cases_clause_g list
type 'a cases_pattern_disjunction_g = 'a cases_pattern_g list
type disjunctive_cases_clause = [ `any ] disjunctive_cases_clause_g
type disjunctive_cases_clauses = [ `any ] disjunctive_cases_clauses_g
type cases_pattern_disjunction = [ `any ] cases_pattern_disjunction_g
type 'a extended_glob_local_binder_r =
  1. | GLocalAssum of Names.Name.t * binding_kind * 'a glob_constr_g
  2. | GLocalDef of Names.Name.t * 'a glob_constr_g * 'a glob_constr_g option
  3. | GLocalPattern of 'a cases_pattern_disjunction_g * Names.Id.t list * Names.Id.t * binding_kind * 'a glob_constr_g
and 'a extended_glob_local_binder_g = ('a extended_glob_local_binder_r, 'a) DAst.t
type extended_glob_local_binder = [ `any ] extended_glob_local_binder_g