summaryrefslogtreecommitdiff
path: root/pretyping
diff options
context:
space:
mode:
authorGravatar Benjamin Barenblat <bbaren@debian.org>2018-12-29 14:31:27 -0500
committerGravatar Benjamin Barenblat <bbaren@debian.org>2018-12-29 14:31:27 -0500
commit9043add656177eeac1491a73d2f3ab92bec0013c (patch)
tree2b0092c84bfbf718eca10c81f60b2640dc8cab05 /pretyping
parenta4c7f8bd98be2a200489325ff7c5061cf80ab4f3 (diff)
Imported Upstream version 8.8.2upstream/8.8.2
Diffstat (limited to 'pretyping')
-rw-r--r--pretyping/arguments_renaming.ml25
-rw-r--r--pretyping/arguments_renaming.mli12
-rw-r--r--pretyping/cases.ml854
-rw-r--r--pretyping/cases.mli68
-rw-r--r--pretyping/cbv.ml149
-rw-r--r--pretyping/cbv.mli19
-rw-r--r--pretyping/classops.ml101
-rw-r--r--pretyping/classops.mli29
-rw-r--r--pretyping/coercion.ml176
-rw-r--r--pretyping/coercion.mli26
-rw-r--r--pretyping/constr_matching.ml343
-rw-r--r--pretyping/constr_matching.mli45
-rw-r--r--pretyping/detyping.ml642
-rw-r--r--pretyping/detyping.mli61
-rw-r--r--pretyping/doc.tex14
-rw-r--r--pretyping/evarconv.ml461
-rw-r--r--pretyping/evarconv.mli26
-rw-r--r--pretyping/evardefine.ml114
-rw-r--r--pretyping/evardefine.mli18
-rw-r--r--pretyping/evarsolve.ml747
-rw-r--r--pretyping/evarsolve.mli32
-rw-r--r--pretyping/find_subterm.ml51
-rw-r--r--pretyping/find_subterm.mli19
-rw-r--r--pretyping/geninterp.ml102
-rw-r--r--pretyping/geninterp.mli74
-rw-r--r--pretyping/glob_ops.ml727
-rw-r--r--pretyping/glob_ops.mli54
-rw-r--r--pretyping/indrec.ml138
-rw-r--r--pretyping/indrec.mli35
-rw-r--r--pretyping/inductiveops.ml223
-rw-r--r--pretyping/inductiveops.mli70
-rw-r--r--pretyping/inferCumulativity.ml212
-rw-r--r--pretyping/inferCumulativity.mli12
-rw-r--r--pretyping/locusops.ml16
-rw-r--r--pretyping/locusops.mli10
-rw-r--r--pretyping/ltac_pretype.ml68
-rw-r--r--pretyping/miscops.ml26
-rw-r--r--pretyping/miscops.mli15
-rw-r--r--pretyping/nativenorm.ml317
-rw-r--r--pretyping/nativenorm.mli20
-rw-r--r--pretyping/patternops.ml225
-rw-r--r--pretyping/patternops.mli19
-rw-r--r--pretyping/pretype_errors.ml119
-rw-r--r--pretyping/pretype_errors.mli99
-rw-r--r--pretyping/pretyping.ml836
-rw-r--r--pretyping/pretyping.mli90
-rw-r--r--pretyping/pretyping.mllib4
-rw-r--r--pretyping/program.ml55
-rw-r--r--pretyping/program.mli16
-rw-r--r--pretyping/recordops.ml128
-rw-r--r--pretyping/recordops.mli28
-rw-r--r--pretyping/redops.ml18
-rw-r--r--pretyping/redops.mli10
-rw-r--r--pretyping/reductionops.ml757
-rw-r--r--pretyping/reductionops.mli121
-rw-r--r--pretyping/retyping.ml165
-rw-r--r--pretyping/retyping.mli25
-rw-r--r--pretyping/tacred.ml466
-rw-r--r--pretyping/tacred.mli19
-rw-r--r--pretyping/typeclasses.ml215
-rw-r--r--pretyping/typeclasses.mli50
-rw-r--r--pretyping/typeclasses_errors.ml15
-rw-r--r--pretyping/typeclasses_errors.mli18
-rw-r--r--pretyping/typing.ml268
-rw-r--r--pretyping/typing.mli25
-rw-r--r--pretyping/unification.ml608
-rw-r--r--pretyping/unification.mli54
-rw-r--r--pretyping/univdecls.ml52
-rw-r--r--pretyping/univdecls.mli21
-rw-r--r--pretyping/vnorm.ml198
-rw-r--r--pretyping/vnorm.mli14
71 files changed, 6292 insertions, 4597 deletions
diff --git a/pretyping/arguments_renaming.ml b/pretyping/arguments_renaming.ml
index e18aece0..84295959 100644
--- a/pretyping/arguments_renaming.ml
+++ b/pretyping/arguments_renaming.ml
@@ -1,18 +1,23 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Globnames
open Term
+open Constr
open Environ
open Util
open Libobject
+
+module NamedDecl = Context.Named.Declaration
(*i*)
let name_table =
@@ -37,18 +42,12 @@ let subst_rename_args (subst, (_, (r, names as orig))) =
let r' = fst (subst_global subst r) in
if r==r' then orig else (r', names)
-let section_segment_of_reference = function
- | ConstRef con -> Lib.section_segment_of_constant con
- | IndRef (kn,_) | ConstructRef ((kn,_),_) ->
- Lib.section_segment_of_mutual_inductive kn
- | _ -> [], Univ.LMap.empty, Univ.UContext.empty
-
let discharge_rename_args = function
| _, (ReqGlobal (c, names), _ as req) ->
(try
- let vars,_,_ = section_segment_of_reference c in
+ let vars = Lib.variable_section_segment_of_reference c in
let c' = pop_global_reference c in
- let var_names = List.map (fun (id, _,_,_) -> Name id) vars in
+ let var_names = List.map (fst %> NamedDecl.get_id %> Name.mk_name) vars in
let names' = var_names @ names in
Some (ReqGlobal (c', names), (c', names'))
with Not_found -> Some req)
@@ -101,7 +100,7 @@ let rename_type_of_constructor env cstruct =
let rename_typing env c =
let j = Typeops.infer env c in
let j' =
- match kind_of_term c with
+ match kind c with
| Const (c,u) -> { j with uj_type = rename_type j.uj_type (ConstRef c) }
| Ind (i,u) -> { j with uj_type = rename_type j.uj_type (IndRef i) }
| Construct (k,u) -> { j with uj_type = rename_type j.uj_type (ConstructRef k) }
diff --git a/pretyping/arguments_renaming.mli b/pretyping/arguments_renaming.mli
index e123e778..65e3c3be 100644
--- a/pretyping/arguments_renaming.mli
+++ b/pretyping/arguments_renaming.mli
@@ -1,15 +1,17 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Globnames
open Environ
-open Term
+open Constr
val rename_arguments : bool -> global_reference -> Name.t list -> unit
diff --git a/pretyping/cases.ml b/pretyping/cases.ml
index 6e4d7270..fe0f20f8 100644
--- a/pretyping/cases.ml
+++ b/pretyping/cases.ml
@@ -1,23 +1,28 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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) *)
(************************************************************************)
+module CVars = Vars
+
open Pp
open CErrors
open Util
open Names
open Nameops
-open Term
-open Vars
+open Constr
open Termops
+open Environ
+open EConstr
+open Vars
open Namegen
open Declarations
open Inductiveops
-open Environ
open Reductionops
open Type_errors
open Glob_term
@@ -29,8 +34,11 @@ open Evardefine
open Evarsolve
open Evarconv
open Evd
-open Sigma.Notations
open Context.Rel.Declaration
+open Ltac_pretype
+
+module RelDecl = Context.Rel.Declaration
+module NamedDecl = Context.Named.Declaration
(* Pattern-matching errors *)
@@ -45,26 +53,26 @@ type pattern_matching_error =
exception PatternMatchingError of env * evar_map * pattern_matching_error
-let raise_pattern_matching_error (loc,env,sigma,te) =
- Loc.raise loc (PatternMatchingError(env,sigma,te))
+let raise_pattern_matching_error ?loc (env,sigma,te) =
+ Loc.raise ?loc (PatternMatchingError(env,sigma,te))
-let error_bad_pattern_loc loc env sigma cstr ind =
- raise_pattern_matching_error
- (loc, env, sigma, BadPattern (cstr,ind))
+let error_bad_pattern ?loc env sigma cstr ind =
+ raise_pattern_matching_error ?loc
+ (env, sigma, BadPattern (cstr,ind))
-let error_bad_constructor_loc loc env cstr ind =
- raise_pattern_matching_error
- (loc, env, Evd.empty, BadConstructor (cstr,ind))
+let error_bad_constructor ?loc env cstr ind =
+ raise_pattern_matching_error ?loc
+ (env, Evd.empty, BadConstructor (cstr,ind))
-let error_wrong_numarg_constructor_loc loc env c n =
- raise_pattern_matching_error (loc, env, Evd.empty, WrongNumargConstructor(c,n))
+let error_wrong_numarg_constructor ?loc env c n =
+ raise_pattern_matching_error ?loc (env, Evd.empty, WrongNumargConstructor(c,n))
-let error_wrong_numarg_inductive_loc loc env c n =
- raise_pattern_matching_error (loc, env, Evd.empty, WrongNumargInductive(c,n))
+let error_wrong_numarg_inductive ?loc env c n =
+ raise_pattern_matching_error ?loc (env, Evd.empty, WrongNumargInductive(c,n))
let list_try_compile f l =
let rec aux errors = function
- | [] -> if errors = [] then anomaly (str "try_find_f") else iraise (List.last errors)
+ | [] -> if errors = [] then anomaly (str "try_find_f.") else iraise (List.last errors)
| h::t ->
try f h
with UserError _ | TypeError _ | PretypeError _ | PatternMatchingError _ as e ->
@@ -89,33 +97,34 @@ let msg_may_need_inversion () =
(* Utils *)
let make_anonymous_patvars n =
- List.make n (PatVar (Loc.ghost,Anonymous))
+ List.make n (DAst.make @@ PatVar Anonymous)
(* We have x1:t1...xn:tn,xi':ti,y1..yk |- c and re-generalize
over xi:ti to get x1:t1...xn:tn,xi':ti,y1..yk |- c[xi:=xi'] *)
let relocate_rel n1 n2 k j = if Int.equal j (n1 + k) then n2+k else j
-let rec relocate_index n1 n2 k t = match kind_of_term t with
+let rec relocate_index sigma n1 n2 k t =
+ match EConstr.kind sigma t with
| Rel j when Int.equal j (n1 + k) -> mkRel (n2+k)
| Rel j when j < n1+k -> t
| Rel j when j > n1+k -> t
- | _ -> map_constr_with_binders succ (relocate_index n1 n2) k t
+ | _ -> EConstr.map_with_binders sigma succ (relocate_index sigma n1 n2) k t
(**********************************************************************)
(* Structures used in compiling pattern-matching *)
type 'a rhs =
{ rhs_env : env;
- rhs_vars : Id.t list;
- avoid_ids : Id.t list;
+ rhs_vars : Id.Set.t;
+ avoid_ids : Id.Set.t;
it : 'a option}
type 'a equation =
{ patterns : cases_pattern list;
rhs : 'a rhs;
alias_stack : Name.t list;
- eqn_loc : Loc.t;
+ eqn_loc : Loc.t option;
used : bool ref }
type 'a matrix = 'a equation list
@@ -135,7 +144,7 @@ type tomatch_status =
| Pushed of (bool*((constr * tomatch_type) * int list * Name.t))
| Alias of (bool*(Name.t * constr * (constr * types)))
| NonDepAlias
- | Abstract of int * Context.Rel.Declaration.t
+ | Abstract of int * rel_declaration
type tomatch_stack = tomatch_status list
@@ -155,9 +164,9 @@ let feed_history arg = function
| Continuation (n, l, h) when n>=1 ->
Continuation (n-1, arg :: l, h)
| Continuation (n, _, _) ->
- anomaly (str "Bad number of expected remaining patterns: " ++ int n)
+ anomaly (str "Bad number of expected remaining patterns: " ++ int n ++ str ".")
| Result _ ->
- anomaly (Pp.str "Exhausted pattern history")
+ anomaly (Pp.str "Exhausted pattern history.")
(* This is for non exhaustive error message *)
@@ -171,7 +180,7 @@ and build_glob_pattern args = function
| Top -> args
| MakeConstructor (pci, rh) ->
glob_pattern_of_partial_history
- [PatCstr (Loc.ghost, pci, args, Anonymous)] rh
+ [DAst.make @@ PatCstr (pci, args, Anonymous)] rh
let complete_history = glob_pattern_of_partial_history []
@@ -181,12 +190,12 @@ let pop_history_pattern = function
| Continuation (0, l, Top) ->
Result (List.rev l)
| Continuation (0, l, MakeConstructor (pci, rh)) ->
- feed_history (PatCstr (Loc.ghost,pci,List.rev l,Anonymous)) rh
+ feed_history (DAst.make @@ PatCstr (pci,List.rev l,Anonymous)) rh
| _ ->
- anomaly (Pp.str "Constructor not yet filled with its arguments")
+ anomaly (Pp.str "Constructor not yet filled with its arguments.")
let pop_history h =
- feed_history (PatVar (Loc.ghost, Anonymous)) h
+ feed_history (DAst.make @@ PatVar Anonymous) h
(* Builds a continuation expecting [n] arguments and building [ci] applied
to this [n] arguments *)
@@ -239,12 +248,13 @@ let push_history_pattern n pci cont =
type 'a pattern_matching_problem =
{ env : env;
+ lvar : Ltac_pretype.ltac_var_map;
evdref : evar_map ref;
pred : constr;
tomatch : tomatch_stack;
history : pattern_continuation;
mat : 'a matrix;
- caseloc : Loc.t;
+ caseloc : Loc.t option;
casestyle : case_style;
typing_function: type_constraint -> env -> evar_map ref -> 'a option -> unsafe_judgment }
@@ -266,24 +276,29 @@ type 'a pattern_matching_problem =
let rec find_row_ind = function
[] -> None
- | PatVar _ :: l -> find_row_ind l
- | PatCstr(loc,c,_,_) :: _ -> Some (loc,c)
+ | p :: l ->
+ match DAst.get p with
+ | PatVar _ -> find_row_ind l
+ | PatCstr(c,_,_) -> Some (p.CAst.loc,c)
let inductive_template evdref env tmloc ind =
let indu = evd_comb1 (Evd.fresh_inductive_instance env) evdref ind in
let arsign = inductive_alldecls_env env indu in
- let hole_source = match tmloc with
- | Some loc -> fun i -> (loc, Evar_kinds.TomatchTypeParameter (ind,i))
- | None -> fun _ -> (Loc.ghost, Evar_kinds.InternalHole) in
+ let indu = on_snd EInstance.make indu in
+ let hole_source i = match tmloc with
+ | Some loc -> Loc.tag ~loc @@ Evar_kinds.TomatchTypeParameter (ind,i)
+ | None -> Loc.tag @@ Evar_kinds.TomatchTypeParameter (ind,i) in
let (_,evarl,_) =
List.fold_right
(fun decl (subst,evarl,n) ->
match decl with
| LocalAssum (na,ty) ->
+ let ty = EConstr.of_constr ty in
let ty' = substl subst ty in
let e = e_new_evar env evdref ~src:(hole_source n) ty' in
(e::subst,e::evarl,n+1)
| LocalDef (na,b,ty) ->
+ let b = EConstr.of_constr b in
(substl subst b::subst,evarl,n+1))
arsign ([],[],1) in
applist (mkIndU indu,List.rev evarl)
@@ -307,9 +322,9 @@ let inh_coerce_to_ind evdref env loc ty tyi =
un inductif cela doit être égal *)
if not (e_cumul env evdref expected_typ ty) then evdref := sigma
-let binding_vars_of_inductive = function
+let binding_vars_of_inductive sigma = function
| NotInd _ -> []
- | IsInd (_,IndType(_,realargs),_) -> List.filter isRel realargs
+ | IsInd (_,IndType(_,realargs),_) -> List.filter (isRel sigma) realargs
let extract_inductive_data env sigma decl =
match decl with
@@ -317,7 +332,7 @@ let extract_inductive_data env sigma decl =
let tmtyp =
try try_find_ind env sigma t None
with Not_found -> NotInd (None,t) in
- let tmtypvars = binding_vars_of_inductive tmtyp in
+ let tmtypvars = binding_vars_of_inductive sigma tmtyp in
(tmtyp,tmtypvars)
| LocalDef (_,_,t) ->
(NotInd (None, t), [])
@@ -332,35 +347,55 @@ let unify_tomatch_with_patterns evdref env loc typ pats realnames =
let find_tomatch_tycon evdref env loc = function
(* Try if some 'in I ...' is present and can be used as a constraint *)
- | Some (_,ind,realnal) ->
+ | Some {CAst.v=(ind,realnal)} ->
mk_tycon (inductive_template evdref env loc ind),Some (List.rev realnal)
| None ->
empty_tycon,None
-let coerce_row typing_fun evdref env pats (tomatch,(_,indopt)) =
- let loc = Some (loc_of_glob_constr tomatch) in
+let make_return_predicate_ltac_lvar sigma na tm c lvar =
+ match na, DAst.get tm with
+ | Name id, (GVar id' | GRef (Globnames.VarRef id', _)) when Id.equal id id' ->
+ if Id.Map.mem id lvar.ltac_genargs then
+ let ltac_genargs = Id.Map.remove id lvar.ltac_genargs in
+ let ltac_idents = match kind sigma c with
+ | Var id' -> Id.Map.add id id' lvar.ltac_idents
+ | _ -> lvar.ltac_idents in
+ { lvar with ltac_genargs; ltac_idents }
+ else lvar
+ | _ -> lvar
+
+let ltac_interp_realnames lvar = function
+ | t, IsInd (ty,ind,realnal) -> t, IsInd (ty,ind,List.map (ltac_interp_name lvar) realnal)
+ | _ as x -> x
+
+let coerce_row typing_fun evdref env lvar pats (tomatch,(na,indopt)) =
+ let loc = loc_of_glob_constr tomatch in
let tycon,realnames = find_tomatch_tycon evdref env loc indopt in
- let j = typing_fun tycon env evdref tomatch in
- let evd, j = Coercion.inh_coerce_to_base (loc_of_glob_constr tomatch) env !evdref j in
+ let j = typing_fun tycon env evdref !lvar tomatch in
+ let evd, j = Coercion.inh_coerce_to_base ?loc:(loc_of_glob_constr tomatch) env !evdref j in
evdref := evd;
let typ = nf_evar !evdref j.uj_type in
+ lvar := make_return_predicate_ltac_lvar !evdref na tomatch j.uj_val !lvar;
let t =
try try_find_ind env !evdref typ realnames
with Not_found ->
unify_tomatch_with_patterns evdref env loc typ pats realnames in
(j.uj_val,t)
-let coerce_to_indtype typing_fun evdref env matx tomatchl =
+let coerce_to_indtype typing_fun evdref env lvar matx tomatchl =
let pats = List.map (fun r -> r.patterns) matx in
let matx' = match matrix_transpose pats with
| [] -> List.map (fun _ -> []) tomatchl (* no patterns at all *)
| m -> m in
- List.map2 (coerce_row typing_fun evdref env) matx' tomatchl
+ let lvar = ref lvar in
+ let tms = List.map2 (coerce_row typing_fun evdref env lvar) matx' tomatchl in
+ let tms = List.map (ltac_interp_realnames !lvar) tms in
+ !lvar,tms
(************************************************************************)
(* Utils *)
-let mkExistential env ?(src=(Loc.ghost,Evar_kinds.InternalHole)) evdref =
+let mkExistential env ?(src=(Loc.tag Evar_kinds.InternalHole)) evdref =
let e, u = e_new_type_evar env evdref univ_flexible_alg ~src:src in e
let evd_comb2 f evdref x y =
@@ -387,12 +422,12 @@ let adjust_tomatch_to_pattern pb ((current,typ),deps,dep) =
| Some (_,(ind,_)) ->
let indt = inductive_template pb.evdref pb.env None ind in
let current =
- if List.is_empty deps && isEvar typ then
+ if List.is_empty deps && isEvar !(pb.evdref) typ then
(* Don't insert coercions if dependent; only solve evars *)
let _ = e_cumul pb.env pb.evdref indt typ in
current
else
- (evd_comb2 (Coercion.inh_conv_coerce_to true Loc.ghost pb.env)
+ (evd_comb2 (Coercion.inh_conv_coerce_to true pb.env)
pb.evdref (make_judge current typ) indt).uj_val in
let sigma = !(pb.evdref) in
(current,try_find_ind pb.env sigma indt names))
@@ -406,7 +441,7 @@ let map_tomatch_type f = function
| IsInd (t,ind,names) -> IsInd (f t,map_inductive_type f ind,names)
| NotInd (c,t) -> NotInd (Option.map f c, f t)
-let liftn_tomatch_type n depth = map_tomatch_type (liftn n depth)
+let liftn_tomatch_type n depth = map_tomatch_type (Vars.liftn n depth)
let lift_tomatch_type n = liftn_tomatch_type n 1
(**********************************************************************)
@@ -415,11 +450,7 @@ let lift_tomatch_type n = liftn_tomatch_type n 1
let current_pattern eqn =
match eqn.patterns with
| pat::_ -> pat
- | [] -> anomaly (Pp.str "Empty list of patterns")
-
-let alias_of_pat = function
- | PatVar (_,name) -> name
- | PatCstr(_,_,_,name) -> name
+ | [] -> anomaly (Pp.str "Empty list of patterns.")
let remove_current_pattern eqn =
match eqn.patterns with
@@ -427,7 +458,7 @@ let remove_current_pattern eqn =
{ eqn with
patterns = pats;
alias_stack = alias_of_pat pat :: eqn.alias_stack }
- | [] -> anomaly (Pp.str "Empty list of patterns")
+ | [] -> anomaly (Pp.str "Empty list of patterns.")
let push_current_pattern (cur,ty) eqn =
match eqn.patterns with
@@ -436,7 +467,7 @@ let push_current_pattern (cur,ty) eqn =
{ eqn with
rhs = { eqn.rhs with rhs_env = rhs_env };
patterns = pats }
- | [] -> anomaly (Pp.str "Empty list of patterns")
+ | [] -> anomaly (Pp.str "Empty list of patterns.")
(* spiwack: like [push_current_pattern] but does not introduce an
alias in rhs_env. Aliasing binders are only useful for variables at
@@ -446,7 +477,7 @@ let push_noalias_current_pattern eqn =
match eqn.patterns with
| _::pats ->
{ eqn with patterns = pats }
- | [] -> anomaly (Pp.str "push_noalias_current_pattern: Empty list of patterns")
+ | [] -> anomaly (Pp.str "push_noalias_current_pattern: Empty list of patterns.")
@@ -458,17 +489,18 @@ let prepend_pattern tms eqn = {eqn with patterns = tms@eqn.patterns }
exception NotAdjustable
-let rec adjust_local_defs loc = function
+let rec adjust_local_defs ?loc = function
| (pat :: pats, LocalAssum _ :: decls) ->
- pat :: adjust_local_defs loc (pats,decls)
+ pat :: adjust_local_defs ?loc (pats,decls)
| (pats, LocalDef _ :: decls) ->
- PatVar (loc, Anonymous) :: adjust_local_defs loc (pats,decls)
+ (DAst.make ?loc @@ PatVar Anonymous) :: adjust_local_defs ?loc (pats,decls)
| [], [] -> []
| _ -> raise NotAdjustable
-let check_and_adjust_constructor env ind cstrs = function
- | PatVar _ as pat -> pat
- | PatCstr (loc,((_,i) as cstr),args,alias) as pat ->
+let check_and_adjust_constructor env ind cstrs pat = match DAst.get pat with
+ | PatVar _ -> pat
+ | PatCstr (((_,i) as cstr),args,alias) ->
+ let loc = pat.CAst.loc in
(* Check it is constructor of the right type *)
let ind' = inductive_of_constructor cstr in
if eq_ind ind' ind then
@@ -478,35 +510,37 @@ let check_and_adjust_constructor env ind cstrs = function
if Int.equal (List.length args) nb_args_constr then pat
else
try
- let args' = adjust_local_defs loc (args, List.rev ci.cs_args)
- in PatCstr (loc, cstr, args', alias)
+ let args' = adjust_local_defs ?loc (args, List.rev ci.cs_args)
+ in DAst.make ?loc @@ PatCstr (cstr, args', alias)
with NotAdjustable ->
- error_wrong_numarg_constructor_loc loc env cstr nb_args_constr
+ error_wrong_numarg_constructor ?loc env cstr nb_args_constr
else
(* Try to insert a coercion *)
try
- Coercion.inh_pattern_coerce_to loc env pat ind' ind
+ Coercion.inh_pattern_coerce_to ?loc env pat ind' ind
with Not_found ->
- error_bad_constructor_loc loc env cstr ind
+ error_bad_constructor ?loc env cstr ind
let check_all_variables env sigma typ mat =
List.iter
- (fun eqn -> match current_pattern eqn with
- | PatVar (_,id) -> ()
- | PatCstr (loc,cstr_sp,_,_) ->
- error_bad_pattern_loc loc env sigma cstr_sp typ)
+ (fun eqn ->
+ let pat = current_pattern eqn in
+ match DAst.get pat with
+ | PatVar id -> ()
+ | PatCstr (cstr_sp,_,_) ->
+ let loc = pat.CAst.loc in
+ error_bad_pattern ?loc env sigma cstr_sp typ)
mat
let check_unused_pattern env eqn =
if not !(eqn.used) then
- raise_pattern_matching_error
- (eqn.eqn_loc, env, Evd.empty, UnusedClause eqn.patterns)
+ raise_pattern_matching_error ?loc:eqn.eqn_loc (env, Evd.empty, UnusedClause eqn.patterns)
let set_used_pattern eqn = eqn.used := true
let extract_rhs pb =
match pb.mat with
- | [] -> errorlabstrm "build_leaf" (msg_may_need_inversion())
+ | [] -> user_err ~hdr:"build_leaf" (msg_may_need_inversion())
| eqn::_ ->
set_used_pattern eqn;
eqn.rhs
@@ -517,10 +551,10 @@ let extract_rhs pb =
let occur_in_rhs na rhs =
match na with
| Anonymous -> false
- | Name id -> Id.List.mem id rhs.rhs_vars
+ | Name id -> Id.Set.mem id rhs.rhs_vars
-let is_dep_patt_in eqn = function
- | PatVar (_,name) -> Flags.is_program_mode () || occur_in_rhs name eqn.rhs
+let is_dep_patt_in eqn pat = match DAst.get pat with
+ | PatVar name -> Flags.is_program_mode () || occur_in_rhs name eqn.rhs
| PatCstr _ -> true
let mk_dep_patt_row (pats,_,eqn) =
@@ -533,19 +567,19 @@ let dependencies_in_pure_rhs nargs eqns =
let deps_columns = matrix_transpose deps_rows in
List.map (List.exists (fun x -> x)) deps_columns
-let dependent_decl a =
+let dependent_decl sigma a =
function
- | LocalAssum (na,t) -> dependent a t
- | LocalDef (na,c,t) -> dependent a t || dependent a c
+ | LocalAssum (na,t) -> dependent sigma a t
+ | LocalDef (na,c,t) -> dependent sigma a t || dependent sigma a c
-let rec dep_in_tomatch n = function
- | (Pushed _ | Alias _ | NonDepAlias) :: l -> dep_in_tomatch n l
- | Abstract (_,d) :: l -> dependent_decl (mkRel n) d || dep_in_tomatch (n+1) l
+let rec dep_in_tomatch sigma n = function
+ | (Pushed _ | Alias _ | NonDepAlias) :: l -> dep_in_tomatch sigma n l
+ | Abstract (_,d) :: l -> dependent_decl sigma (mkRel n) d || dep_in_tomatch sigma (n+1) l
| [] -> false
-let dependencies_in_rhs nargs current tms eqns =
- match kind_of_term current with
- | Rel n when dep_in_tomatch n tms -> List.make nargs true
+let dependencies_in_rhs sigma nargs current tms eqns =
+ match EConstr.kind sigma current with
+ | Rel n when dep_in_tomatch sigma n tms -> List.make nargs true
| _ -> dependencies_in_pure_rhs nargs eqns
(* Computing the matrix of dependencies *)
@@ -554,31 +588,30 @@ let dependencies_in_rhs nargs current tms eqns =
declarations [d(i+1);...;dn] the term [tmi] is dependent in.
[find_dependencies_signature (used1,...,usedn) ((tm1,d1),...,(tmn,dn))]
- returns [(deps1,...,depsn)] where [depsi] is a subset of n,..,i+1
+ returns [(deps1,...,depsn)] where [depsi] is a subset of tm(i+1),..,tmn
denoting in which of the d(i+1)...dn, the term tmi is dependent.
- Dependencies are expressed by index, e.g. in dependency list
- [n-2;1], [1] points to [dn] and [n-2] to [d3]
*)
-let rec find_dependency_list tmblock = function
+let rec find_dependency_list sigma tmblock = function
| [] -> []
- | (used,tdeps,d)::rest ->
- let deps = find_dependency_list tmblock rest in
- if used && List.exists (fun x -> dependent_decl x d) tmblock
+ | (used,tdeps,tm,d)::rest ->
+ let deps = find_dependency_list sigma tmblock rest in
+ if used && List.exists (fun x -> dependent_decl sigma x d) tmblock
then
- List.add_set Int.equal
- (List.length rest + 1) (List.union Int.equal deps tdeps)
+ match EConstr.kind sigma tm with
+ | Rel n -> List.add_set Int.equal n (List.union Int.equal deps tdeps)
+ | _ -> List.union Int.equal deps tdeps
else deps
-let find_dependencies is_dep_or_cstr_in_rhs (tm,(_,tmtypleaves),d) nextlist =
- let deps = find_dependency_list (tm::tmtypleaves) nextlist in
+let find_dependencies sigma is_dep_or_cstr_in_rhs (tm,(_,tmtypleaves),d) nextlist =
+ let deps = find_dependency_list sigma (tm::tmtypleaves) nextlist in
if is_dep_or_cstr_in_rhs || not (List.is_empty deps)
- then ((true ,deps,d)::nextlist)
- else ((false,[] ,d)::nextlist)
+ then ((true ,deps,tm,d)::nextlist)
+ else ((false,[] ,tm,d)::nextlist)
-let find_dependencies_signature deps_in_rhs typs =
- let l = List.fold_right2 find_dependencies deps_in_rhs typs [] in
- List.map (fun (_,deps,_) -> deps) l
+let find_dependencies_signature sigma deps_in_rhs typs =
+ let l = List.fold_right2 (find_dependencies sigma) deps_in_rhs typs [] in
+ List.map (fun (_,deps,_,_) -> deps) l
(* Assume we had terms t1..tq to match in a context xp:Tp,...,x1:T1 |-
and xn:Tn has just been regeneralized into x:Tn so that the terms
@@ -591,31 +624,31 @@ let find_dependencies_signature deps_in_rhs typs =
[relocate_index_tomatch 1 n tomatch] will go the way back.
*)
-let relocate_index_tomatch n1 n2 =
+let relocate_index_tomatch sigma n1 n2 =
let rec genrec depth = function
| [] ->
[]
| Pushed (b,((c,tm),l,na)) :: rest ->
- let c = relocate_index n1 n2 depth c in
- let tm = map_tomatch_type (relocate_index n1 n2 depth) tm in
+ let c = relocate_index sigma n1 n2 depth c in
+ let tm = map_tomatch_type (relocate_index sigma n1 n2 depth) tm in
let l = List.map (relocate_rel n1 n2 depth) l in
Pushed (b,((c,tm),l,na)) :: genrec depth rest
| Alias (initial,(na,c,d)) :: rest ->
(* [c] is out of relocation scope *)
- Alias (initial,(na,c,map_pair (relocate_index n1 n2 depth) d)) :: genrec depth rest
+ Alias (initial,(na,c,map_pair (relocate_index sigma n1 n2 depth) d)) :: genrec depth rest
| NonDepAlias :: rest ->
NonDepAlias :: genrec depth rest
| Abstract (i,d) :: rest ->
let i = relocate_rel n1 n2 depth i in
- Abstract (i, map_constr (relocate_index n1 n2 depth) d)
+ Abstract (i, RelDecl.map_constr (fun c -> relocate_index sigma n1 n2 depth c) d)
:: genrec (depth+1) rest in
genrec 0
(* [replace_tomatch n c tomatch] replaces [Rel n] by [c] in [tomatch] *)
-let rec replace_term n c k t =
- if isRel t && Int.equal (destRel t) (n + k) then lift k c
- else map_constr_with_binders succ (replace_term n c) k t
+let rec replace_term sigma n c k t =
+ if isRel sigma t && Int.equal (destRel sigma t) (n + k) then Vars.lift k c
+ else EConstr.map_with_binders sigma succ (replace_term sigma n c) k t
let length_of_tomatch_type_sign na t =
let l = match na with
@@ -626,21 +659,21 @@ let length_of_tomatch_type_sign na t =
| NotInd _ -> l
| IsInd (_, _, names) -> List.length names + l
-let replace_tomatch n c =
+let replace_tomatch sigma n c =
let rec replrec depth = function
| [] -> []
| Pushed (initial,((b,tm),l,na)) :: rest ->
- let b = replace_term n c depth b in
- let tm = map_tomatch_type (replace_term n c depth) tm in
- List.iter (fun i -> if Int.equal i (n + depth) then anomaly (Pp.str "replace_tomatch")) l;
+ let b = replace_term sigma n c depth b in
+ let tm = map_tomatch_type (replace_term sigma n c depth) tm in
+ List.iter (fun i -> if Int.equal i (n + depth) then anomaly (Pp.str "replace_tomatch.")) l;
Pushed (initial,((b,tm),l,na)) :: replrec depth rest
| Alias (initial,(na,b,d)) :: rest ->
(* [b] is out of replacement scope *)
- Alias (initial,(na,b,map_pair (replace_term n c depth) d)) :: replrec depth rest
+ Alias (initial,(na,b,map_pair (replace_term sigma n c depth) d)) :: replrec depth rest
| NonDepAlias :: rest ->
NonDepAlias :: replrec depth rest
| Abstract (i,d) :: rest ->
- Abstract (i, map_constr (replace_term n c depth) d)
+ Abstract (i, RelDecl.map_constr (fun t -> replace_term sigma n c depth t) d)
:: replrec (depth+1) rest in
replrec 0
@@ -665,7 +698,7 @@ let rec liftn_tomatch_stack n depth = function
NonDepAlias :: liftn_tomatch_stack n depth rest
| Abstract (i,d)::rest ->
let i = if i<depth then i else i+n in
- Abstract (i, map_constr (liftn n depth) d)
+ Abstract (i, RelDecl.map_constr (liftn n depth) d)
::(liftn_tomatch_stack n (depth+1) rest)
let lift_tomatch_stack n = liftn_tomatch_stack n 1
@@ -700,7 +733,7 @@ let merge_name get_name obj = function
let merge_names get_name = List.map2 (merge_name get_name)
-let get_names env sign eqns =
+let get_names env sigma sign eqns =
let names1 = List.make (Context.Rel.length sign) Anonymous in
(* If any, we prefer names used in pats, from top to bottom *)
let names2,aliasname =
@@ -712,17 +745,17 @@ let get_names env sign eqns =
(* Otherwise, we take names from the parameters of the constructor but
avoiding conflicts with user ids *)
let allvars =
- List.fold_left (fun l (_,_,eqn) -> List.union Id.equal l eqn.rhs.avoid_ids)
- [] eqns in
+ List.fold_left (fun l (_,_,eqn) -> Id.Set.union l eqn.rhs.avoid_ids)
+ Id.Set.empty eqns in
let names3,_ =
List.fold_left2
(fun (l,avoid) d na ->
let na =
merge_name
- (fun (LocalAssum (na,t) | LocalDef (na,_,t)) -> Name (next_name_away (named_hd env t na) avoid))
+ (fun (LocalAssum (na,t) | LocalDef (na,_,t)) -> Name (next_name_away (named_hd env sigma t na) avoid))
d na
in
- (na::l,(out_name na)::avoid))
+ (na::l,Id.Set.add (Name.get_id na) avoid))
([],allvars) (List.rev sign) names2 in
names3,aliasname
@@ -733,7 +766,7 @@ let get_names env sign eqns =
(* We now replace the names y1 .. yn y by the actual names *)
(* xi1 .. xin xi to be found in the i-th clause of the matrix *)
-let recover_initial_subpattern_names = List.map2 set_name
+let recover_initial_subpattern_names = List.map2 RelDecl.set_name
let recover_and_adjust_alias_names names sign =
let rec aux = function
@@ -742,7 +775,7 @@ let recover_and_adjust_alias_names names sign =
| x::names, LocalAssum (_,t)::sign ->
(x, LocalAssum (alias_of_pat x,t)) :: aux (names,sign)
| names, (LocalDef (na,_,_) as decl)::sign ->
- (PatVar (Loc.ghost,na), decl) :: aux (names,sign)
+ (DAst.make @@ PatVar na, decl) :: aux (names,sign)
| _ -> assert false
in
List.split (aux (names,sign))
@@ -758,11 +791,11 @@ let push_rels_eqn_with_names sign eqn =
push_rels_eqn sign eqn
let push_generalized_decl_eqn env n decl eqn =
- match get_name decl with
+ match RelDecl.get_name decl with
| Anonymous ->
push_rels_eqn [decl] eqn
| Name _ ->
- push_rels_eqn [set_name (get_name (Environ.lookup_rel n eqn.rhs.rhs_env)) decl] eqn
+ push_rels_eqn [RelDecl.set_name (RelDecl.get_name (Environ.lookup_rel n eqn.rhs.rhs_env)) decl] eqn
let drop_alias_eqn eqn =
{ eqn with alias_stack = List.tl eqn.alias_stack }
@@ -770,7 +803,7 @@ let drop_alias_eqn eqn =
let push_alias_eqn alias eqn =
let aliasname = List.hd eqn.alias_stack in
let eqn = drop_alias_eqn eqn in
- let alias = set_name aliasname alias in
+ let alias = RelDecl.set_name aliasname alias in
push_rels_eqn [alias] eqn
(**********************************************************************)
@@ -830,13 +863,13 @@ let rec map_predicate f k ccl = function
| Abstract _ :: rest ->
map_predicate f (k+1) ccl rest
-let noccur_predicate_between n = map_predicate (noccur_between n)
+let noccur_predicate_between sigma n = map_predicate (noccur_between sigma n)
let liftn_predicate n = map_predicate (liftn n)
let lift_predicate n = liftn_predicate n 1
-let regeneralize_index_predicate n = map_predicate (relocate_index n 1) 0
+let regeneralize_index_predicate sigma n = map_predicate (relocate_index sigma n 1) 0
let substnl_predicate sigma = map_predicate (substnl sigma)
@@ -847,7 +880,7 @@ let subst_predicate (subst,copt) ccl tms =
| Some c -> c::subst in
substnl_predicate sigma 0 ccl tms
-let specialize_predicate_var (cur,typ,dep) tms ccl =
+let specialize_predicate_var (cur,typ,dep) env tms ccl =
let c = match dep with
| Anonymous -> None
| Name _ -> Some cur
@@ -855,7 +888,10 @@ let specialize_predicate_var (cur,typ,dep) tms ccl =
let l =
match typ with
| IsInd (_, IndType (_, _), []) -> []
- | IsInd (_, IndType (_, realargs), names) -> realargs
+ | IsInd (_, IndType (indf, realargs), names) ->
+ let arsign,_ = get_arity env indf in
+ let arsign = List.map EConstr.of_rel_decl arsign in
+ subst_of_rel_context_instance arsign realargs
| NotInd _ -> [] in
subst_predicate (l,c) ccl tms
@@ -868,13 +904,13 @@ let specialize_predicate_var (cur,typ,dep) tms ccl =
(* We first need to lift t(x) s.t. it is typed in Gamma, X:=rargs, x' *)
(* then we have to replace x by x' in t(x) and y by y' in P *)
(*****************************************************************************)
-let generalize_predicate (names,na) ny d tms ccl =
+let generalize_predicate sigma (names,na) ny d tms ccl =
let () = match na with
- | Anonymous -> anomaly (Pp.str "Undetected dependency")
+ | Anonymous -> anomaly (Pp.str "Undetected dependency.")
| _ -> () in
let p = List.length names + 1 in
let ccl = lift_predicate 1 ccl tms in
- regeneralize_index_predicate (ny+p+1) ccl tms
+ regeneralize_index_predicate sigma (ny+p+1) ccl tms
(*****************************************************************************)
(* We just matched over cur:ind(realargs) in the following matching problem *)
@@ -916,16 +952,16 @@ let rec extract_predicate ccl = function
ccl
let abstract_predicate env sigma indf cur realargs (names,na) tms ccl =
- let sign = make_arity_signature env true indf in
+ let sign = make_arity_signature env sigma true indf in
(* n is the number of real args + 1 (+ possible let-ins in sign) *)
let n = List.length sign in
(* Before abstracting we generalize over cur and on those realargs *)
(* that are rels, consistently with the specialization made in *)
(* build_branch *)
let tms = List.fold_right2 (fun par arg tomatch ->
- match kind_of_term par with
- | Rel i -> relocate_index_tomatch (i+n) (destRel arg) tomatch
- | _ -> tomatch) (realargs@[cur]) (Context.Rel.to_extended_list 0 sign)
+ match EConstr.kind sigma par with
+ | Rel i -> relocate_index_tomatch sigma (i+n) (destRel sigma arg) tomatch
+ | _ -> tomatch) (realargs@[cur]) (Context.Rel.to_extended_list EConstr.mkRel 0 sign)
(lift_tomatch_stack n tms) in
(* Pred is already dependent in the current term to match (if *)
(* (na<>Anonymous) and its realargs; we just need to adjust it to *)
@@ -937,7 +973,7 @@ let abstract_predicate env sigma indf cur realargs (names,na) tms ccl =
let pred = extract_predicate ccl tms in
(* Build the predicate properly speaking *)
let sign = List.map2 set_name (na::names) sign in
- it_mkLambda_or_LetIn_name env pred sign
+ it_mkLambda_or_LetIn_name env sigma pred sign
(* [expand_arg] is used by [specialize_predicate]
if Yk denotes [Xk;xk] or [Xk],
@@ -949,34 +985,37 @@ let expand_arg tms (p,ccl) ((_,t),_,na) =
let k = length_of_tomatch_type_sign na t in
(p+k,liftn_predicate (k-1) (p+1) ccl tms)
-
let use_unit_judge evd =
let j, ctx = coq_unit_judge () in
let evd' = Evd.merge_context_set Evd.univ_flexible_alg evd ctx in
evd', j
let add_assert_false_case pb tomatch =
- let pats = List.map (fun _ -> PatVar (Loc.ghost,Anonymous)) tomatch in
+ let pats = List.map (fun _ -> DAst.make @@ PatVar Anonymous) tomatch in
let aliasnames =
List.map_filter (function Alias _ | NonDepAlias -> Some Anonymous | _ -> None) tomatch
in
[ { patterns = pats;
rhs = { rhs_env = pb.env;
- rhs_vars = [];
- avoid_ids = [];
+ rhs_vars = Id.Set.empty;
+ avoid_ids = Id.Set.empty;
it = None };
alias_stack = Anonymous::aliasnames;
- eqn_loc = Loc.ghost;
+ eqn_loc = None;
used = ref false } ]
let adjust_impossible_cases pb pred tomatch submat =
match submat with
| [] ->
- begin match kind_of_term pred with
+ (** FIXME: This breaks if using evar-insensitive primitives. In particular,
+ this means that the Evd.define below may redefine an already defined
+ evar. See e.g. first definition of test for bug #3388. *)
+ let pred = EConstr.Unsafe.to_constr pred in
+ begin match Constr.kind pred with
| Evar (evk,_) when snd (evar_source evk !(pb.evdref)) == Evar_kinds.ImpossibleCase ->
if not (Evd.is_defined !(pb.evdref) evk) then begin
let evd, default = use_unit_judge !(pb.evdref) in
- pb.evdref := Evd.define evk default.uj_type evd
+ pb.evdref := Evd.define evk (EConstr.Unsafe.to_constr default.uj_type) evd
end;
add_assert_false_case pb tomatch
| _ ->
@@ -1022,12 +1061,13 @@ let specialize_predicate newtomatchs (names,depna) arsign cs tms ccl =
(* We prepare the substitution of X and x:I(X) *)
let realargsi =
if not (Int.equal nrealargs 0) then
- subst_of_rel_context_instance arsign (Array.to_list cs.cs_concl_realargs)
+ CVars.subst_of_rel_context_instance arsign (Array.to_list cs.cs_concl_realargs)
else
[] in
+ let realargsi = List.map EConstr.of_constr realargsi in
let copti = match depna with
| Anonymous -> None
- | Name _ -> Some (build_dependent_constructor cs)
+ | Name _ -> Some (EConstr.of_constr (build_dependent_constructor cs))
in
(* The substituends realargsi, copti are all defined in gamma, x1...xn *)
(* We need _parallel_ bindings to get gamma, x1...xn |- PI tms. ccl'' *)
@@ -1063,69 +1103,69 @@ let adjust_predicate_from_tomatch tomatch (current,typ as ct) pb =
(* Remove commutative cuts that turn out to be non-dependent after
some evars have been instantiated *)
-let rec ungeneralize n ng body =
- match kind_of_term body with
+let rec ungeneralize sigma n ng body =
+ match EConstr.kind sigma body with
| Lambda (_,_,c) when Int.equal ng 0 ->
subst1 (mkRel n) c
| Lambda (na,t,c) ->
(* We traverse an inner generalization *)
- mkLambda (na,t,ungeneralize (n+1) (ng-1) c)
+ mkLambda (na,t,ungeneralize sigma (n+1) (ng-1) c)
| LetIn (na,b,t,c) ->
(* We traverse an alias *)
- mkLetIn (na,b,t,ungeneralize (n+1) ng c)
+ mkLetIn (na,b,t,ungeneralize sigma (n+1) ng c)
| Case (ci,p,c,brs) ->
(* We traverse a split *)
let p =
- let sign,p = decompose_lam_assum p in
- let sign2,p = decompose_prod_n_assum ng p in
- let p = prod_applist p [mkRel (n+List.length sign+ng)] in
+ let sign,p = decompose_lam_assum sigma p in
+ let sign2,p = decompose_prod_n_assum sigma ng p in
+ let p = prod_applist sigma p [mkRel (n+List.length sign+ng)] in
it_mkLambda_or_LetIn (it_mkProd_or_LetIn p sign2) sign in
mkCase (ci,p,c,Array.map2 (fun q c ->
- let sign,b = decompose_lam_n_decls q c in
- it_mkLambda_or_LetIn (ungeneralize (n+q) ng b) sign)
+ let sign,b = decompose_lam_n_decls sigma q c in
+ it_mkLambda_or_LetIn (ungeneralize sigma (n+q) ng b) sign)
ci.ci_cstr_ndecls brs)
| App (f,args) ->
(* We traverse an inner generalization *)
- assert (isCase f);
- mkApp (ungeneralize n (ng+Array.length args) f,args)
+ assert (isCase sigma f);
+ mkApp (ungeneralize sigma n (ng+Array.length args) f,args)
| _ -> assert false
-let ungeneralize_branch n k (sign,body) cs =
- (sign,ungeneralize (n+cs.cs_nargs) k body)
+let ungeneralize_branch sigma n k (sign,body) cs =
+ (sign,ungeneralize sigma (n+cs.cs_nargs) k body)
-let rec is_dependent_generalization ng body =
- match kind_of_term body with
+let rec is_dependent_generalization sigma ng body =
+ match EConstr.kind sigma body with
| Lambda (_,_,c) when Int.equal ng 0 ->
- dependent (mkRel 1) c
+ not (noccurn sigma 1 c)
| Lambda (na,t,c) ->
(* We traverse an inner generalization *)
- is_dependent_generalization (ng-1) c
+ is_dependent_generalization sigma (ng-1) c
| LetIn (na,b,t,c) ->
(* We traverse an alias *)
- is_dependent_generalization ng c
+ is_dependent_generalization sigma ng c
| Case (ci,p,c,brs) ->
(* We traverse a split *)
Array.exists2 (fun q c ->
- let _,b = decompose_lam_n_decls q c in
- is_dependent_generalization ng b)
+ let _,b = decompose_lam_n_decls sigma q c in
+ is_dependent_generalization sigma ng b)
ci.ci_cstr_ndecls brs
| App (g,args) ->
(* We traverse an inner generalization *)
- assert (isCase g);
- is_dependent_generalization (ng+Array.length args) g
+ assert (isCase sigma g);
+ is_dependent_generalization sigma (ng+Array.length args) g
| _ -> assert false
-let is_dependent_branch k (_,br) =
- is_dependent_generalization k br
+let is_dependent_branch sigma k (_,br) =
+ is_dependent_generalization sigma k br
let postprocess_dependencies evd tocheck brs tomatch pred deps cs =
let rec aux k brs tomatch pred tocheck deps = match deps, tomatch with
| [], _ -> brs,tomatch,pred,[]
| n::deps, Abstract (i,d) :: tomatch ->
- let d = map_constr (nf_evar evd) d in
+ let d = map_constr (fun c -> nf_evar evd c) d in
let is_d = match d with LocalAssum _ -> false | LocalDef _ -> true in
- if is_d || List.exists (fun c -> dependent_decl (lift k c) d) tocheck
- && Array.exists (is_dependent_branch k) brs then
+ if is_d || List.exists (fun c -> dependent_decl evd (lift k c) d) tocheck
+ && Array.exists (is_dependent_branch evd k) brs then
(* Dependency in the current term to match and its dependencies is real *)
let brs,tomatch,pred,inst = aux (k+1) brs tomatch pred (mkRel n::tocheck) deps in
let inst = match d with
@@ -1138,9 +1178,9 @@ let postprocess_dependencies evd tocheck brs tomatch pred deps cs =
(* terms by its actual value in both the remaining terms to match and *)
(* the bodies of the Case *)
let pred = lift_predicate (-1) pred tomatch in
- let tomatch = relocate_index_tomatch 1 (n+1) tomatch in
+ let tomatch = relocate_index_tomatch evd 1 (n+1) tomatch in
let tomatch = lift_tomatch_stack (-1) tomatch in
- let brs = Array.map2 (ungeneralize_branch n k) brs cs in
+ let brs = Array.map2 (ungeneralize_branch evd n k) brs cs in
aux k brs tomatch pred tocheck deps
| _ -> assert false
in aux 0 brs tomatch pred tocheck deps
@@ -1148,9 +1188,9 @@ let postprocess_dependencies evd tocheck brs tomatch pred deps cs =
(************************************************************************)
(* Sorting equations by constructor *)
-let rec irrefutable env = function
- | PatVar (_,name) -> true
- | PatCstr (_,cstr,args,_) ->
+let rec irrefutable env pat = match DAst.get pat with
+ | PatVar name -> true
+ | PatCstr (cstr,args,_) ->
let ind = inductive_of_constructor cstr in
let (_,mip) = Inductive.lookup_mind_specif env ind in
let one_constr = Int.equal (Array.length mip.mind_user_lc) 1 in
@@ -1170,15 +1210,15 @@ let group_equations pb ind current cstrs mat =
(fun eqn () ->
let rest = remove_current_pattern eqn in
let pat = current_pattern eqn in
- match check_and_adjust_constructor pb.env ind cstrs pat with
- | PatVar (_,name) ->
+ match DAst.get (check_and_adjust_constructor pb.env ind cstrs pat) with
+ | PatVar name ->
(* This is a default clause that we expand *)
for i=1 to Array.length cstrs do
let args = make_anonymous_patvars cstrs.(i-1).cs_nargs in
brs.(i-1) <- (args, name, rest) :: brs.(i-1)
done;
if !only_default == None then only_default := Some true
- | PatCstr (loc,((_,i)),args,name) ->
+ | PatCstr (((_,i)),args,name) ->
(* This is a regular clause *)
only_default := Some false;
brs.(i-1) <- (args, name, rest) :: brs.(i-1)) mat () in
@@ -1192,17 +1232,17 @@ let rec generalize_problem names pb = function
| [] -> pb, []
| i::l ->
let pb',deps = generalize_problem names pb l in
- let d = map_constr (lift i) (Environ.lookup_rel i pb.env) in
+ let d = map_constr (lift i) (lookup_rel i pb.env) in
begin match d with
| LocalDef (Anonymous,_,_) -> pb', deps
| _ ->
(* for better rendering *)
- let d = map_type (whd_betaiota !(pb.evdref)) d in
+ let d = RelDecl.map_type (fun c -> whd_betaiota !(pb.evdref) c) d in
let tomatch = lift_tomatch_stack 1 pb'.tomatch in
- let tomatch = relocate_index_tomatch (i+1) 1 tomatch in
+ let tomatch = relocate_index_tomatch !(pb.evdref) (i+1) 1 tomatch in
{ pb' with
tomatch = Abstract (i,d) :: tomatch;
- pred = generalize_predicate names i d pb'.tomatch pb'.pred },
+ pred = generalize_predicate !(pb'.evdref) names i d pb'.tomatch pb'.pred },
i::deps
end
@@ -1224,10 +1264,17 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn
(* build the name x1..xn from the names present in the equations *)
(* that had matched constructor C *)
let cs_args = const_info.cs_args in
- let names,aliasname = get_names pb.env cs_args eqns in
- let typs = List.map2 set_name names cs_args
+ let cs_args = List.map (fun d -> map_rel_decl EConstr.of_constr d) cs_args in
+ let names,aliasname = get_names pb.env !(pb.evdref) cs_args eqns in
+ let typs = List.map2 RelDecl.set_name names cs_args
in
+ (* Beta-iota-normalize types to better compatibility of refine with 8.4 behavior *)
+ (* This is a bit too strong I think, in the sense that what we would *)
+ (* really like is to have beta-iota reduction only at the positions where *)
+ (* parameters are substituted *)
+ let typs = List.map (map_type (nf_betaiota pb.env !(pb.evdref))) typs in
+
(* We build the matrix obtained by expanding the matching on *)
(* "C x1..xn as x" followed by a residual matching on eqn into *)
(* a matching on "x1 .. xn eqn" *)
@@ -1247,35 +1294,35 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn
(* We compute over which of x(i+1)..xn and x matching on xi will need a *)
(* generalization *)
let dep_sign =
- find_dependencies_signature
- (dependencies_in_rhs const_info.cs_nargs current pb.tomatch eqns)
+ find_dependencies_signature !(pb.evdref)
+ (dependencies_in_rhs !(pb.evdref) const_info.cs_nargs current pb.tomatch eqns)
(List.rev typs') in
(* The dependent term to subst in the types of the remaining UnPushed
terms is relative to the current context enriched by topushs *)
- let ci = build_dependent_constructor const_info in
+ let ci = EConstr.of_constr (build_dependent_constructor const_info) in
(* Current context Gamma has the form Gamma1;cur:I(realargs);Gamma2 *)
(* We go from Gamma |- PI tms. pred to *)
(* Gamma;x1..xn;curalias:I(x1..xn) |- PI tms'. pred' *)
(* where, in tms and pred, those realargs that are vars are *)
(* replaced by the corresponding xi and cur replaced by curalias *)
- let cirealargs = Array.to_list const_info.cs_concl_realargs in
+ let cirealargs = Array.map_to_list EConstr.of_constr const_info.cs_concl_realargs in
(* Do the specialization for terms to match *)
let tomatch = List.fold_right2 (fun par arg tomatch ->
- match kind_of_term par with
- | Rel i -> replace_tomatch (i+const_info.cs_nargs) arg tomatch
+ match EConstr.kind !(pb.evdref) par with
+ | Rel i -> replace_tomatch !(pb.evdref) (i+const_info.cs_nargs) arg tomatch
| _ -> tomatch) (current::realargs) (ci::cirealargs)
(lift_tomatch_stack const_info.cs_nargs pb.tomatch) in
let pred_is_not_dep =
- noccur_predicate_between 1 (List.length realnames + 1) pb.pred tomatch in
+ noccur_predicate_between !(pb.evdref) 1 (List.length realnames + 1) pb.pred tomatch in
let typs' =
List.map2
(fun (tm, (tmtyp,_), decl) deps ->
- let na = get_name decl in
+ let na = RelDecl.get_name decl in
let na = match curname, na with
| Name _, Anonymous -> curname
| Name _, Name _ -> na
@@ -1296,10 +1343,10 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn
| Name _ ->
let cur_alias = lift const_info.cs_nargs current in
let ind =
- appvect (
- applist (mkIndU (inductive_of_constructor (fst const_info.cs_cstr), snd const_info.cs_cstr),
- List.map (lift const_info.cs_nargs) const_info.cs_params),
- const_info.cs_concl_realargs) in
+ mkApp (
+ applist (mkIndU (inductive_of_constructor (fst const_info.cs_cstr), EInstance.make (snd const_info.cs_cstr)),
+ List.map (EConstr.of_constr %> lift const_info.cs_nargs) const_info.cs_params),
+ Array.map EConstr.of_constr const_info.cs_concl_realargs) in
Alias (initial,(aliasname,cur_alias,(ci,ind))) in
let tomatch = List.rev_append (alias :: currents) tomatch in
@@ -1307,8 +1354,7 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn
let submat = adjust_impossible_cases pb pred tomatch submat in
let () = match submat with
| [] ->
- raise_pattern_matching_error
- (Loc.ghost, pb.env, Evd.empty, NonExhaustive (complete_history history))
+ raise_pattern_matching_error (pb.env, Evd.empty, NonExhaustive (complete_history history))
| _ -> ()
in
@@ -1366,23 +1412,24 @@ and match_current pb (initial,tomatch) =
(* We compile branches *)
let brvals = Array.map2 (compile_branch initial current realargs (names,dep) deps pb arsign) eqns cstrs in
(* We build the (elementary) case analysis *)
- let depstocheck = current::binding_vars_of_inductive typ in
+ let depstocheck = current::binding_vars_of_inductive !(pb.evdref) typ in
let brvals,tomatch,pred,inst =
postprocess_dependencies !(pb.evdref) depstocheck
brvals pb.tomatch pb.pred deps cstrs in
let brvals = Array.map (fun (sign,body) ->
+ let sign = List.map (map_name (ltac_interp_name pb.lvar)) sign in
it_mkLambda_or_LetIn body sign) brvals in
let (pred,typ) =
find_predicate pb.caseloc pb.env pb.evdref
pred current indt (names,dep) tomatch in
let ci = make_case_info pb.env (fst mind) pb.casestyle in
- let pred = nf_betaiota !(pb.evdref) pred in
+ let pred = nf_betaiota pb.env !(pb.evdref) pred in
let case =
- make_case_or_project pb.env indf ci pred current brvals
+ make_case_or_project pb.env !(pb.evdref) indf ci pred current brvals
in
Typing.check_allowed_sort pb.env !(pb.evdref) mind current pred;
{ uj_val = applist (case, inst);
- uj_type = prod_applist typ inst }
+ uj_type = prod_applist !(pb.evdref) typ inst }
(* Building the sub-problem when all patterns are variables. Case
@@ -1390,7 +1437,7 @@ and match_current pb (initial,tomatch) =
and shift_problem ((current,t),_,na) pb =
let ty = type_of_tomatch t in
let tomatch = lift_tomatch_stack 1 pb.tomatch in
- let pred = specialize_predicate_var (current,t,na) pb.tomatch pb.pred in
+ let pred = specialize_predicate_var (current,t,na) pb.env pb.tomatch pb.pred in
let pb =
{ pb with
env = push_rel (LocalDef (na,current,ty)) pb.env;
@@ -1407,7 +1454,7 @@ and shift_problem ((current,t),_,na) pb =
are already introduced in the context, we avoid creating aliases to
themselves by treating this case specially. *)
and pop_problem ((current,t),_,na) pb =
- let pred = specialize_predicate_var (current,t,na) pb.tomatch pb.pred in
+ let pred = specialize_predicate_var (current,t,na) pb.env pb.tomatch pb.pred in
let pb =
{ pb with
pred = pred;
@@ -1450,8 +1497,9 @@ and compile_alias initial pb (na,orig,(expanded,expanded_typ)) rest =
history = pop_history_pattern pb.history;
mat = List.map (push_alias_eqn alias) pb.mat } in
let j = compile pb in
+ let sigma = !(pb.evdref) in
{ uj_val =
- if isRel c || isVar c || count_occurrences (mkRel 1) j.uj_val <= 1 then
+ if isRel sigma c || isVar sigma c || count_occurrences sigma (mkRel 1) j.uj_val <= 1 then
subst1 c j.uj_val
else
mkLetIn (na,c,t,j.uj_val);
@@ -1476,10 +1524,10 @@ and compile_alias initial pb (na,orig,(expanded,expanded_typ)) rest =
evaluation; the drawback is that it might duplicate the instances
of the term to match when the corresponding variable is
substituted by a non-evaluated expression *)
- if not (Flags.is_program_mode ()) && (isRel orig || isVar orig) then
+ if not (Flags.is_program_mode ()) && (isRel sigma orig || isVar sigma orig) then
(* Try to compile first using non expanded alias *)
try
- if initial then f orig (Retyping.get_type_of pb.env !(pb.evdref) orig)
+ if initial then f orig (Retyping.get_type_of pb.env sigma orig)
else just_pop ()
with e when precatchable_exception e ->
(* Try then to compile using expanded alias *)
@@ -1515,15 +1563,15 @@ substituer après par les initiaux *)
(* builds the matrix of equations testing that each eqn has n patterns
* and linearizing the _ patterns.
- * Syntactic correctness has already been done in astterm *)
+ * Syntactic correctness has already been done in constrintern *)
let matx_of_eqns env eqns =
- let build_eqn (loc,ids,lpat,rhs) =
- let initial_lpat,initial_rhs = lpat,rhs in
- let initial_rhs = rhs in
+ let build_eqn {CAst.loc;v=(ids,initial_lpat,initial_rhs)} =
+ let avoid = ids_of_named_context_val (named_context_val env) in
+ let avoid = List.fold_left (fun accu id -> Id.Set.add id accu) avoid ids in
let rhs =
{ rhs_env = env;
rhs_vars = free_glob_vars initial_rhs;
- avoid_ids = ids@(ids_of_named_context (named_context env));
+ avoid_ids = avoid;
it = Some initial_rhs } in
{ patterns = initial_lpat;
alias_stack = [];
@@ -1560,7 +1608,7 @@ let matx_of_eqns env eqns =
returning True never happens and any inhabited type can be put instead).
*)
-let adjust_to_extended_env_and_remove_deps env extenv subst t =
+let adjust_to_extended_env_and_remove_deps env extenv sigma subst t =
let n = Context.Rel.length (rel_context env) in
let n' = Context.Rel.length (rel_context extenv) in
(* We first remove the bindings that are dependently typed (they are
@@ -1578,11 +1626,11 @@ let adjust_to_extended_env_and_remove_deps env extenv subst t =
let (p, _, _) = lookup_rel_id x (rel_context extenv) in
let rec traverse_local_defs p =
match lookup_rel p extenv with
- | LocalDef (_,c,_) -> assert (isRel c); traverse_local_defs (p + destRel c)
+ | LocalDef (_,c,_) -> assert (isRel sigma c); traverse_local_defs (p + destRel sigma c)
| LocalAssum _ -> p in
let p = traverse_local_defs p in
let u = lift (n' - n) u in
- try Some (p, u, expand_vars_in_term extenv u)
+ try Some (p, u, expand_vars_in_term extenv sigma u)
(* pedrot: does this really happen to raise [Failure _]? *)
with Failure _ -> None in
let subst0 = List.map_filter map subst in
@@ -1611,19 +1659,19 @@ let rec list_assoc_in_triple x = function
* similarly for each ti.
*)
-let abstract_tycon loc env evdref subst tycon extenv t =
- let t = nf_betaiota !evdref t in (* it helps in some cases to remove K-redex*)
- let src = match kind_of_term t with
- | Evar (evk,_) -> (loc,Evar_kinds.SubEvar evk)
- | _ -> (loc,Evar_kinds.CasesType true) in
- let subst0,t0 = adjust_to_extended_env_and_remove_deps env extenv subst t in
+let abstract_tycon ?loc env evdref subst tycon extenv t =
+ let t = nf_betaiota env !evdref t in (* it helps in some cases to remove K-redex*)
+ let src = match EConstr.kind !evdref t with
+ | Evar (evk,_) -> (Loc.tag ?loc @@ Evar_kinds.SubEvar evk)
+ | _ -> (Loc.tag ?loc @@ Evar_kinds.CasesType true) in
+ let subst0,t0 = adjust_to_extended_env_and_remove_deps env extenv !evdref subst t in
(* We traverse the type T of the original problem Xi looking for subterms
that match the non-constructor part of the constraints (this part
is in subst); these subterms are the "good" subterms and we replace them
by an evar that may depend (and only depend) on the corresponding
convertible subterms of the substitution *)
let rec aux (k,env,subst as x) t =
- let t = whd_evar !evdref t in match kind_of_term t with
+ match EConstr.kind !evdref t with
| Rel n when is_local_def (lookup_rel n env) -> t
| Evar ev ->
let ty = get_type_of env !evdref t in
@@ -1643,7 +1691,7 @@ let abstract_tycon loc env evdref subst tycon extenv t =
let good = List.filter (fun (_,u,_) -> is_conv_leq env !evdref t u) subst in
match good with
| [] ->
- map_constr_with_full_binders push_binder aux x t
+ map_constr_with_full_binders !evdref push_binder aux x t
| (_, _, u) :: _ -> (* u is in extenv *)
let vl = List.map pi1 good in
let ty =
@@ -1651,17 +1699,16 @@ let abstract_tycon loc env evdref subst tycon extenv t =
Evarutil.evd_comb1 (refresh_universes (Some false) env) evdref ty
in
let ty = lift (-k) (aux x ty) in
- let depvl = free_rels ty in
+ let depvl = free_rels !evdref ty in
let inst =
List.map_i
(fun i _ -> if Int.List.mem i vl then u else mkRel i) 1
(rel_context extenv) in
let rel_filter =
- List.map (fun a -> not (isRel a) || dependent a u
- || Int.Set.mem (destRel a) depvl) inst in
+ List.map (fun a -> not (isRel !evdref a) || dependent !evdref a u
+ || Int.Set.mem (destRel !evdref a) depvl) inst in
let named_filter =
- let open Context.Named.Declaration in
- List.map (fun d -> dependent (mkVar (get_id d)) u)
+ List.map (fun d -> local_occur_var !evdref (NamedDecl.get_id d) u)
(named_context extenv) in
let filter = Filter.make (rel_filter @ named_filter) in
let candidates = u :: List.map mkRel vl in
@@ -1670,7 +1717,7 @@ let abstract_tycon loc env evdref subst tycon extenv t =
in
aux (0,extenv,subst0) t0
-let build_tycon loc env tycon_env s subst tycon extenv evdref t =
+let build_tycon ?loc env tycon_env s subst tycon extenv evdref t =
let t,tt = match t with
| None ->
(* This is the situation we are building a return predicate and
@@ -1678,15 +1725,15 @@ let build_tycon loc env tycon_env s subst tycon extenv evdref t =
let n = Context.Rel.length (rel_context env) in
let n' = Context.Rel.length (rel_context tycon_env) in
let impossible_case_type, u =
- e_new_type_evar (reset_context env) evdref univ_flexible_alg ~src:(loc,Evar_kinds.ImpossibleCase) in
+ e_new_type_evar (reset_context env) evdref univ_flexible_alg ~src:(Loc.tag ?loc Evar_kinds.ImpossibleCase) in
(lift (n'-n) impossible_case_type, mkSort u)
| Some t ->
- let t = abstract_tycon loc tycon_env evdref subst tycon extenv t in
+ let t = abstract_tycon ?loc tycon_env evdref subst tycon extenv t in
let evd,tt = Typing.type_of extenv !evdref t in
evdref := evd;
(t,tt) in
let b = e_cumul env evdref tt (mkSort s) (* side effect *) in
- if not b then anomaly (Pp.str "Build_tycon: should be a type");
+ if not b then anomaly (Pp.str "Build_tycon: should be a type.");
{ uj_val = t; uj_type = tt }
(* For a multiple pattern-matching problem Xi on t1..tn with return
@@ -1701,26 +1748,26 @@ let build_tycon loc env tycon_env s subst tycon extenv evdref t =
let build_inversion_problem loc env sigma tms t =
let make_patvar t (subst,avoid) =
- let id = next_name_away (named_hd env t Anonymous) avoid in
- PatVar (Loc.ghost,Name id), ((id,t)::subst, id::avoid) in
+ let id = next_name_away (named_hd env sigma t Anonymous) avoid in
+ DAst.make @@ PatVar (Name id), ((id,t)::subst, Id.Set.add id avoid) in
let rec reveal_pattern t (subst,avoid as acc) =
- match kind_of_term (whd_all env sigma t) with
- | Construct (cstr,u) -> PatCstr (Loc.ghost,cstr,[],Anonymous), acc
- | App (f,v) when isConstruct f ->
- let cstr,u = destConstruct f in
+ match EConstr.kind sigma (whd_all env sigma t) with
+ | Construct (cstr,u) -> DAst.make (PatCstr (cstr,[],Anonymous)), acc
+ | App (f,v) when isConstruct sigma f ->
+ let cstr,u = destConstruct sigma f in
let n = constructor_nrealargs_env env cstr in
let l = List.lastn n (Array.to_list v) in
- let l,acc = List.fold_map' reveal_pattern l acc in
- PatCstr (Loc.ghost,cstr,l,Anonymous), acc
+ let l,acc = List.fold_right_map reveal_pattern l acc in
+ DAst.make (PatCstr (cstr,l,Anonymous)), acc
| _ -> make_patvar t acc in
let rec aux n env acc_sign tms acc =
match tms with
| [] -> [], acc_sign, acc
| (t, IsInd (_,IndType(indf,realargs),_)) :: tms ->
- let patl,acc = List.fold_map' reveal_pattern realargs acc in
+ let patl,acc = List.fold_right_map reveal_pattern realargs acc in
let pat,acc = make_patvar t acc in
let indf' = lift_inductive_family n indf in
- let sign = make_arity_signature env true indf' in
+ let sign = make_arity_signature env sigma true indf' in
let patl = pat :: List.rev patl in
let patl,sign = recover_and_adjust_alias_names patl sign in
let p = List.length patl in
@@ -1732,7 +1779,7 @@ let build_inversion_problem loc env sigma tms t =
let d = LocalAssum (alias_of_pat pat,typ) in
let patl,acc_sign,acc = aux (n+1) (push_rel d env) (d::acc_sign) tms acc in
pat::patl,acc_sign,acc in
- let avoid0 = ids_of_context env in
+ let avoid0 = vars_of_env env in
(* [patl] is a list of patterns revealing the substructure of
constructors present in the constraints on the type of the
multiple terms t1..tn that are matched in the original problem;
@@ -1753,11 +1800,11 @@ let build_inversion_problem loc env sigma tms t =
List.map (fun (c,d) -> (c,extract_inductive_data pb_env sigma d,d)) decls in
let decls = List.rev decls in
- let dep_sign = find_dependencies_signature (List.make n true) decls in
+ let dep_sign = find_dependencies_signature sigma (List.make n true) decls in
let sub_tms =
List.map2 (fun deps (tm, (tmtyp,_), decl) ->
- let na = if List.is_empty deps then Anonymous else force_name (get_name decl) in
+ let na = if List.is_empty deps then Anonymous else force_name (RelDecl.get_name decl) in
Pushed (true,((tm,tmtyp),deps,na)))
dep_sign decls in
let subst = List.map (fun (na,t) -> (na,lift n t)) subst in
@@ -1769,12 +1816,12 @@ let build_inversion_problem loc env sigma tms t =
let main_eqn =
{ patterns = patl;
alias_stack = [];
- eqn_loc = Loc.ghost;
+ eqn_loc = None;
used = ref false;
rhs = { rhs_env = pb_env;
(* we assume all vars are used; in practice we discard dependent
vars so that the field rhs_vars is normally not used *)
- rhs_vars = List.map fst subst;
+ rhs_vars = List.fold_left (fun accu (id, _) -> Id.Set.add id accu) Id.Set.empty subst;
avoid_ids = avoid;
it = Some (lift n t) } } in
(* [catch_all] is a catch-all default clause of the auxiliary
@@ -1787,22 +1834,23 @@ let build_inversion_problem loc env sigma tms t =
(* No need for a catch all clause *)
[]
else
- [ { patterns = List.map (fun _ -> PatVar (Loc.ghost,Anonymous)) patl;
+ [ { patterns = List.map (fun _ -> DAst.make @@ PatVar Anonymous) patl;
alias_stack = [];
- eqn_loc = Loc.ghost;
+ eqn_loc = None;
used = ref false;
rhs = { rhs_env = pb_env;
- rhs_vars = [];
+ rhs_vars = Id.Set.empty;
avoid_ids = avoid0;
it = None } } ] in
(* [pb] is the auxiliary pattern-matching serving as skeleton for the
return type of the original problem Xi *)
let s' = Retyping.get_sort_of env sigma t in
- let sigma, s = Evd.new_sort_variable univ_flexible_alg sigma in
+ let sigma, s = Evd.new_sort_variable univ_flexible sigma in
let sigma = Evd.set_leq_sort env sigma s' s in
let evdref = ref sigma in
let pb =
{ env = pb_env;
+ lvar = empty_lvar;
evdref = evdref;
pred = (*ty *) mkSort s;
tomatch = sub_tms;
@@ -1810,7 +1858,7 @@ let build_inversion_problem loc env sigma tms t =
mat = main_eqn :: catch_all_eqn;
caseloc = loc;
casestyle = RegularStyle;
- typing_function = build_tycon loc env pb_env s subst} in
+ typing_function = build_tycon ?loc env pb_env s subst} in
let pred = (compile pb).uj_val in
(!evdref,pred)
@@ -1820,52 +1868,62 @@ let build_initial_predicate arsign pred =
let rec buildrec n pred tmnames = function
| [] -> List.rev tmnames,pred
| (decl::realdecls)::lnames ->
- let na = get_name decl in
+ let na = RelDecl.get_name decl in
let n' = n + List.length realdecls in
buildrec (n'+1) pred (force_name na::tmnames) lnames
| _ -> assert false
in buildrec 0 pred [] (List.rev arsign)
-let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign =
+let extract_arity_signature ?(dolift=true) env0 lvar tomatchl tmsign =
let lift = if dolift then lift else fun n t -> t in
let get_one_sign n tm (na,t) =
match tm with
| NotInd (bo,typ) ->
(match t with
- | None -> (match bo with
+ | None -> let sign = match bo with
| None -> [LocalAssum (na, lift n typ)]
- | Some b -> [LocalDef (na, lift n b, lift n typ)])
- | Some (loc,_,_) ->
- user_err_loc (loc,"",
- str"Unexpected type annotation for a term of non inductive type."))
+ | Some b -> [LocalDef (na, lift n b, lift n typ)] in sign,sign
+ | Some {CAst.loc} ->
+ user_err ?loc
+ (str"Unexpected type annotation for a term of non inductive type."))
| IsInd (term,IndType(indf,realargs),_) ->
let indf' = if dolift then lift_inductive_family n indf else indf in
let ((ind,u),_) = dest_ind_family indf' in
let nrealargs_ctxt = inductive_nrealdecls_env env0 ind in
let arsign = fst (get_arity env0 indf') in
- let realnal =
+ let arsign = List.map (fun d -> map_rel_decl EConstr.of_constr d) arsign in
+ let realnal, realnal' =
match t with
- | Some (loc,ind',realnal) ->
+ | Some {CAst.loc;v=(ind',realnal)} ->
if not (eq_ind ind ind') then
- user_err_loc (loc,"",str "Wrong inductive type.");
+ user_err ?loc (str "Wrong inductive type.");
if not (Int.equal nrealargs_ctxt (List.length realnal)) then
- anomaly (Pp.str "Ill-formed 'in' clause in cases");
- List.rev realnal
- | None -> List.make nrealargs_ctxt Anonymous in
- LocalAssum (na, build_dependent_inductive env0 indf')
- ::(List.map2 set_name realnal arsign) in
+ anomaly (Pp.str "Ill-formed 'in' clause in cases.");
+ let realnal = List.rev realnal in
+ let realnal' = List.map (ltac_interp_name lvar) realnal in
+ realnal,realnal'
+ | None ->
+ let realnal = List.make nrealargs_ctxt Anonymous in
+ realnal, realnal in
+ let na' = ltac_interp_name lvar na in
+ let t = EConstr.of_constr (build_dependent_inductive env0 indf') in
+ (* Context with names for typing *)
+ let arsign1 = LocalAssum (na, t) :: List.map2 RelDecl.set_name realnal arsign in
+ (* Context with names for building the term *)
+ let arsign2 = LocalAssum (na', t) :: List.map2 RelDecl.set_name realnal' arsign in
+ arsign1,arsign2 in
let rec buildrec n = function
| [],[] -> []
| (_,tm)::ltm, (_,x)::tmsign ->
let l = get_one_sign n tm x in
- l :: buildrec (n + List.length l) (ltm,tmsign)
+ l :: buildrec (n + List.length (fst l)) (ltm,tmsign)
| _ -> assert false
in List.rev (buildrec 0 (tomatchl,tmsign))
-let inh_conv_coerce_to_tycon loc env evdref j tycon =
+let inh_conv_coerce_to_tycon ?loc env evdref j tycon =
match tycon with
| Some p ->
- let (evd',j) = Coercion.inh_conv_coerce_to true loc env !evdref j p in
+ let (evd',j) = Coercion.inh_conv_coerce_to ?loc true env !evdref j p in
evdref := evd';
j
| None -> j
@@ -1877,8 +1935,8 @@ let prepare_predicate_from_arsign_tycon env sigma loc tomatchs arsign c =
let subst, len =
List.fold_right2 (fun (tm, tmtype) sign (subst, len) ->
let signlen = List.length sign in
- match kind_of_term tm with
- | Rel n when dependent tm c
+ match EConstr.kind sigma tm with
+ | Rel n when dependent sigma tm c
&& Int.equal signlen 1 (* The term to match is not of a dependent type itself *) ->
((n, len) :: subst, len - signlen)
| Rel n when signlen > 1 (* The term is of a dependent type,
@@ -1889,21 +1947,21 @@ let prepare_predicate_from_arsign_tycon env sigma loc tomatchs arsign c =
let subst, len =
List.fold_left
(fun (subst, len) arg ->
- match kind_of_term arg with
- | Rel n when dependent arg c ->
+ match EConstr.kind sigma arg with
+ | Rel n when dependent sigma arg c ->
((n, len) :: subst, pred len)
| _ -> (subst, pred len))
(subst, len) realargs
in
let subst =
- if dependent tm c && List.for_all isRel realargs
+ if dependent sigma tm c && List.for_all (isRel sigma) realargs
then (n, len) :: subst else subst
in (subst, pred len))
| _ -> (subst, len - signlen))
(List.rev tomatchs) arsign ([], nar)
in
let rec predicate lift c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Rel n when n > lift ->
(try
(* Make the predicate dependent on the matched variable *)
@@ -1913,7 +1971,7 @@ let prepare_predicate_from_arsign_tycon env sigma loc tomatchs arsign c =
(* A variable that is not matched, lift over the arsign. *)
mkRel (n + nar))
| _ ->
- map_constr_with_binders succ predicate lift c
+ EConstr.map_with_binders sigma succ predicate lift c
in
assert (len == 0);
let p = predicate 0 c in
@@ -1933,7 +1991,22 @@ let prepare_predicate_from_arsign_tycon env sigma loc tomatchs arsign c =
* tycon to make the predicate if it is not closed.
*)
-let prepare_predicate loc typing_fun env sigma tomatchs arsign tycon pred =
+exception LocalOccur
+
+let noccur_with_meta sigma n m term =
+ let rec occur_rec n c = match EConstr.kind sigma c with
+ | Rel p -> if n<=p && p<n+m then raise LocalOccur
+ | App(f,cl) ->
+ (match EConstr.kind sigma f with
+ | Cast (c,_,_) when isMeta sigma c -> ()
+ | Meta _ -> ()
+ | _ -> EConstr.iter_with_binders sigma succ occur_rec n c)
+ | Evar (_, _) -> ()
+ | _ -> EConstr.iter_with_binders sigma succ occur_rec n c
+ in
+ try (occur_rec n term; true) with LocalOccur -> false
+
+let prepare_predicate ?loc typing_fun env sigma lvar tomatchs arsign tycon pred =
let refresh_tycon sigma t =
(** If we put the typing constraint in the term, it has to be
refreshed to preserve the invariant that no algebraic universe
@@ -1941,16 +2014,17 @@ let prepare_predicate loc typing_fun env sigma tomatchs arsign tycon pred =
refresh_universes ~status:Evd.univ_flexible ~onlyalg:true (Some true)
env sigma t
in
+ let typing_arsign,building_arsign = List.split arsign in
let preds =
match pred, tycon with
(* No return clause *)
- | None, Some t when not (noccur_with_meta 0 max_int t) ->
+ | None, Some t when not (noccur_with_meta sigma 0 max_int t) ->
(* If the tycon is not closed w.r.t real variables, we try *)
(* two different strategies *)
(* First strategy: we abstract the tycon wrt to the dependencies *)
let sigma, t = refresh_tycon sigma t in
let p1 =
- prepare_predicate_from_arsign_tycon env sigma loc tomatchs arsign t in
+ prepare_predicate_from_arsign_tycon env sigma loc tomatchs typing_arsign t in
(* Second strategy: we build an "inversion" predicate *)
let sigma2,pred2 = build_inversion_problem loc env sigma tomatchs t in
(match p1 with
@@ -1962,31 +2036,29 @@ let prepare_predicate loc typing_fun env sigma tomatchs arsign tycon pred =
let sigma,t = match tycon with
| Some t -> refresh_tycon sigma t
| None ->
- let sigma = Sigma.Unsafe.of_evar_map sigma in
- let Sigma ((t, _), sigma, _) =
- new_type_evar env sigma univ_flexible_alg ~src:(loc, Evar_kinds.CasesType false) in
- let sigma = Sigma.to_evar_map sigma in
+ let (sigma, (t, _)) =
+ new_type_evar env sigma univ_flexible_alg ~src:(Loc.tag ?loc @@ Evar_kinds.CasesType false) in
sigma, t
in
(* First strategy: we build an "inversion" predicate *)
let sigma1,pred1 = build_inversion_problem loc env sigma tomatchs t in
(* Second strategy: we directly use the evar as a non dependent pred *)
- let pred2 = lift (List.length (List.flatten arsign)) t in
+ let pred2 = lift (List.length (List.flatten typing_arsign)) t in
[sigma1, pred1; sigma, pred2]
(* Some type annotation *)
| Some rtntyp, _ ->
(* We extract the signature of the arity *)
- let envar = List.fold_right push_rel_context arsign env in
+ let envar = List.fold_right push_rel_context typing_arsign env in
let sigma, newt = new_sort_variable univ_flexible_alg sigma in
let evdref = ref sigma in
- let predcclj = typing_fun (mk_tycon (mkSort newt)) envar evdref rtntyp in
+ let predcclj = typing_fun (mk_tycon (mkSort newt)) envar evdref lvar rtntyp in
let sigma = !evdref in
- let predccl = (j_nf_evar sigma predcclj).uj_val in
+ let predccl = nf_evar sigma predcclj.uj_val in
[sigma, predccl]
in
List.map
(fun (sigma,pred) ->
- let (nal,pred) = build_initial_predicate arsign pred in
+ let (nal,pred) = build_initial_predicate building_arsign pred in
sigma,nal,pred)
preds
@@ -2011,7 +2083,7 @@ let prime avoid name =
let make_prime avoid prevname =
let previd, id = prime !avoid prevname in
- avoid := id :: !avoid;
+ avoid := Id.Set.add id !avoid;
previd, id
let eq_id avoid id =
@@ -2026,31 +2098,33 @@ let mk_JMeq evdref typ x typ' y =
let mk_JMeq_refl evdref typ x =
papp evdref coq_JMeq_refl [| typ; x |]
-let hole =
- GHole (Loc.ghost, Evar_kinds.QuestionMark (Evar_kinds.Define false),
+let hole na = DAst.make @@
+ GHole (Evar_kinds.QuestionMark (Evar_kinds.Define false,na),
Misctypes.IntroAnonymous, None)
let constr_of_pat env evdref arsign pat avoid =
let rec typ env (ty, realargs) pat avoid =
- match pat with
- | PatVar (l,name) ->
+ let loc = pat.CAst.loc in
+ match DAst.get pat with
+ | PatVar name ->
let name, avoid = match name with
Name n -> name, avoid
| Anonymous ->
let previd, id = prime avoid (Name (Id.of_string "wildcard")) in
- Name id, id :: avoid
+ Name id, Id.Set.add id avoid
in
- (PatVar (l, name), [LocalAssum (name, ty)] @ realargs, mkRel 1, ty,
+ ((DAst.make ?loc @@ PatVar name), [LocalAssum (name, ty)] @ realargs, mkRel 1, ty,
(List.map (fun x -> mkRel 1) realargs), 1, avoid)
- | PatCstr (l,((_, i) as cstr),args,alias) ->
+ | PatCstr (((_, i) as cstr),args,alias) ->
let cind = inductive_of_constructor cstr in
let IndType (indf, _) =
try find_rectype env ( !evdref) (lift (-(List.length realargs)) ty)
- with Not_found -> error_case_not_inductive env
+ with Not_found -> error_case_not_inductive env !evdref
{uj_val = ty; uj_type = Typing.unsafe_type_of env !evdref ty}
in
let (ind,u), params = dest_ind_family indf in
- if not (eq_ind ind cind) then error_bad_constructor_loc l env cstr ind;
+ let params = List.map EConstr.of_constr params in
+ if not (eq_ind ind cind) then error_bad_constructor ?loc env cstr ind;
let cstrs = get_constructors env indf in
let ci = cstrs.(i-1) in
let nb_args_constr = ci.cs_nargs in
@@ -2058,7 +2132,7 @@ let constr_of_pat env evdref arsign pat avoid =
let patargs, args, sign, env, n, m, avoid =
List.fold_right2
(fun decl ua (patargs, args, sign, env, n, m, avoid) ->
- let t = get_type decl in
+ let t = EConstr.of_constr (RelDecl.get_type decl) in
let pat', sign', arg', typ', argtypargs, n', avoid =
let liftt = liftn (List.length sign) (succ (List.length args)) t in
typ env (substl args liftt, []) ua avoid
@@ -2070,18 +2144,18 @@ let constr_of_pat env evdref arsign pat avoid =
in
let args = List.rev args in
let patargs = List.rev patargs in
- let pat' = PatCstr (l, cstr, patargs, alias) in
- let cstr = mkConstructU ci.cs_cstr in
- let app = applistc cstr (List.map (lift (List.length sign)) params) in
- let app = applistc app args in
+ let pat' = DAst.make ?loc @@ PatCstr (cstr, patargs, alias) in
+ let cstr = mkConstructU (on_snd EInstance.make ci.cs_cstr) in
+ let app = applist (cstr, List.map (lift (List.length sign)) params) in
+ let app = applist (app, args) in
let apptype = Retyping.get_type_of env ( !evdref) app in
- let IndType (indf, realargs) = find_rectype env ( !evdref) apptype in
+ let IndType (indf, realargs) = find_rectype env (!evdref) apptype in
match alias with
Anonymous ->
pat', sign, app, apptype, realargs, n, avoid
| Name id ->
let sign = LocalAssum (alias, lift m ty) :: sign in
- let avoid = id :: avoid in
+ let avoid = Id.Set.add id avoid in
let sign, i, avoid =
try
let env = push_rel_context sign env in
@@ -2092,58 +2166,61 @@ let constr_of_pat env evdref arsign pat avoid =
(lift 1 app) (* aliased term *)
in
let neq = eq_id avoid id in
- LocalDef (Name neq, mkRel 0, eq_t) :: sign, 2, neq :: avoid
+ LocalDef (Name neq, mkRel 0, eq_t) :: sign, 2, Id.Set.add neq avoid
with Reduction.NotConvertible -> sign, 1, avoid
in
(* Mark the equality as a hole *)
pat', sign, lift i app, lift i apptype, realargs, n + i, avoid
in
- let pat', sign, patc, patty, args, z, avoid = typ env (get_type (List.hd arsign), List.tl arsign) pat avoid in
- pat', (sign, patc, (get_type (List.hd arsign), args), pat'), avoid
+ let pat', sign, patc, patty, args, z, avoid = typ env (RelDecl.get_type (List.hd arsign), List.tl arsign) pat avoid in
+ pat', (sign, patc, (RelDecl.get_type (List.hd arsign), args), pat'), avoid
(* shadows functional version *)
let eq_id avoid id =
let hid = Id.of_string ("Heq_" ^ Id.to_string id) in
let hid' = next_ident_away hid !avoid in
- avoid := hid' :: !avoid;
+ avoid := Id.Set.add hid' !avoid;
hid'
-let is_topvar t =
-match kind_of_term t with
+let is_topvar sigma t =
+match EConstr.kind sigma t with
| Rel 0 -> true
| _ -> false
-let rels_of_patsign =
+let rels_of_patsign sigma =
List.map (fun decl ->
match decl with
- | LocalDef (na,t',t) when is_topvar t' -> LocalAssum (na,t)
+ | LocalDef (na,t',t) when is_topvar sigma t' -> LocalAssum (na,t)
| _ -> decl)
-let vars_of_ctx ctx =
+let vars_of_ctx sigma ctx =
let _, y =
List.fold_right (fun decl (prev, vars) ->
match decl with
- | LocalDef (na,t',t) when is_topvar t' ->
+ | LocalDef (na,t',t) when is_topvar sigma t' ->
prev,
- (GApp (Loc.ghost,
- (GRef (Loc.ghost, delayed_force coq_eq_refl_ref, None)),
- [hole; GVar (Loc.ghost, prev)])) :: vars
+ (DAst.make @@ GApp (
+ (DAst.make @@ GRef (delayed_force coq_eq_refl_ref, None)),
+ [hole na; DAst.make @@ GVar prev])) :: vars
| _ ->
- match get_name decl with
+ match RelDecl.get_name decl with
Anonymous -> invalid_arg "vars_of_ctx"
- | Name n -> n, GVar (Loc.ghost, n) :: vars)
+ | Name n -> n, (DAst.make @@ GVar n) :: vars)
ctx (Id.of_string "vars_of_ctx_error", [])
in List.rev y
let rec is_included x y =
- match x, y with
+ match DAst.get x, DAst.get y with
| PatVar _, _ -> true
| _, PatVar _ -> true
- | PatCstr (l, (_, i), args, alias), PatCstr (l', (_, i'), args', alias') ->
+ | PatCstr ((_, i), args, alias), PatCstr ((_, i'), args', alias') ->
if Int.equal i i' then List.for_all2 is_included args args'
else false
+let lift_rel_context n l =
+ map_rel_context_with_binders (liftn n) l
+
(* liftsign is the current pattern's complete signature length.
Hence pats is already typed in its
full signature. However prevpatterns are in the original one signature per pattern form.
@@ -2180,14 +2257,14 @@ let build_ineqs evdref prevpatterns pats liftsign =
(Some ([], 0, 0, [])) eqnpats pats
in match acc with
None -> c
- | Some (sign, len, _, c') ->
- let conj = it_mkProd_or_LetIn (mk_coq_not (mk_coq_and c'))
- (lift_rel_context liftsign sign)
- in
- conj :: c)
+ | Some (sign, len, _, c') ->
+ let sigma, conj = mk_coq_and !evdref c' in
+ let sigma, neg = mk_coq_not sigma conj in
+ let conj = it_mkProd_or_LetIn neg (lift_rel_context liftsign sign) in
+ evdref := sigma; conj :: c)
[] prevpatterns
in match diffs with [] -> None
- | _ -> Some (mk_coq_and diffs)
+ | _ -> Some (let sigma, conj = mk_coq_and !evdref diffs in evdref := sigma; conj)
let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity =
let i = ref 0 in
@@ -2199,7 +2276,7 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity =
(fun (idents, newpatterns, pats) pat arsign ->
let pat', cpat, idents = constr_of_pat env evdref arsign pat idents in
(idents, pat' :: newpatterns, cpat :: pats))
- ([], [], []) eqn.patterns sign
+ (Id.Set.empty, [], []) eqn.patterns sign
in
let newpatterns = List.rev newpatterns and opats = List.rev pats in
let rhs_rels, pats, signlen =
@@ -2218,12 +2295,12 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity =
(* lift to get outside of past patterns to get terms in the combined environment. *)
(fun (pats, n) (sign, c, (s, args), p) ->
let len = List.length sign in
- ((rels_of_patsign sign, lift n c,
+ ((rels_of_patsign !evdref sign, lift n c,
(s, List.map (lift n) args), p) :: pats, len + n))
([], 0) pats
in
let ineqs = build_ineqs evdref prevpatterns pats signlen in
- let rhs_rels' = rels_of_patsign rhs_rels in
+ let rhs_rels' = rels_of_patsign !evdref rhs_rels in
let _signenv = push_rel_context rhs_rels' env in
let arity =
let args, nargs =
@@ -2241,7 +2318,7 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity =
| Some ineqs ->
[LocalAssum (Anonymous, ineqs)], lift 1 arity
in
- let eqs_rels, arity = decompose_prod_n_assum neqs arity in
+ let eqs_rels, arity = decompose_prod_n_assum !evdref neqs arity in
eqs_rels @ neqs_rels @ rhs_rels', arity
in
let rhs_env = push_rel_context rhs_rels' env in
@@ -2252,13 +2329,13 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity =
let branch_name = Id.of_string ("program_branch_" ^ (string_of_int !i)) in
let branch_decl = LocalDef (Name branch_name, lift !i bbody, lift !i btype) in
let branch =
- let bref = GVar (Loc.ghost, branch_name) in
- match vars_of_ctx rhs_rels with
+ let bref = DAst.make @@ GVar branch_name in
+ match vars_of_ctx !evdref rhs_rels with
[] -> bref
- | l -> GApp (Loc.ghost, bref, l)
+ | l -> DAst.make @@ GApp (bref, l)
in
let branch = match ineqs with
- Some _ -> GApp (Loc.ghost, branch, [ hole ])
+ Some _ -> DAst.make @@ GApp (branch, [ hole Anonymous ])
| None -> branch
in
incr i;
@@ -2287,27 +2364,27 @@ let lift_ctx n ctx =
in ctx'
(* Turn matched terms into variables. *)
-let abstract_tomatch env tomatchs tycon =
+let abstract_tomatch env sigma tomatchs tycon =
let prev, ctx, names, tycon =
List.fold_left
(fun (prev, ctx, names, tycon) (c, t) ->
let lenctx = List.length ctx in
- match kind_of_term c with
+ match EConstr.kind sigma c with
Rel n -> (lift lenctx c, lift_tomatch_type lenctx t) :: prev, ctx, names, tycon
| _ ->
let tycon = Option.map
- (fun t -> subst_term (lift 1 c) (lift 1 t)) tycon in
+ (fun t -> subst_term sigma (lift 1 c) (lift 1 t)) tycon in
let name = next_ident_away (Id.of_string "filtered_var") names in
(mkRel 1, lift_tomatch_type (succ lenctx) t) :: lift_ctx 1 prev,
LocalDef (Name name, lift lenctx c, lift lenctx $ type_of_tomatch t) :: ctx,
- name :: names, tycon)
- ([], [], [], tycon) tomatchs
+ Id.Set.add name names, tycon)
+ ([], [], Id.Set.empty, tycon) tomatchs
in List.rev prev, ctx, tycon
let build_dependent_signature env evdref avoid tomatchs arsign =
let avoid = ref avoid in
let arsign = List.rev arsign in
- let allnames = List.rev_map (List.map get_name) arsign in
+ let allnames = List.rev_map (List.map RelDecl.get_name) arsign in
let nar = List.fold_left (fun n names -> List.length names + n) 0 allnames in
let eqs, neqs, refls, slift, arsign' =
List.fold_left2
@@ -2324,14 +2401,14 @@ let build_dependent_signature env evdref avoid tomatchs arsign =
as much as possible *)
let argsign = List.tl arsign in (* arguments in inverse application order *)
let app_decl = List.hd arsign in (* The matched argument *)
- let appn = get_name app_decl in
- let appt = get_type app_decl in
+ let appn = RelDecl.get_name app_decl in
+ let appt = RelDecl.get_type app_decl in
let argsign = List.rev argsign in (* arguments in application order *)
let env', nargeqs, argeqs, refl_args, slift, argsign' =
List.fold_left2
(fun (env, nargeqs, argeqs, refl_args, slift, argsign') arg decl ->
- let name = get_name decl in
- let t = get_type decl in
+ let name = RelDecl.get_name decl in
+ let t = RelDecl.get_type decl in
let argt = Retyping.get_type_of env !evdref arg in
let eq, refl_arg =
if Reductionops.is_conv env !evdref argt t then
@@ -2348,8 +2425,8 @@ let build_dependent_signature env evdref avoid tomatchs arsign =
in
let previd, id =
let name =
- match kind_of_term arg with
- Rel n -> get_name (lookup_rel n env)
+ match EConstr.kind !evdref arg with
+ Rel n -> RelDecl.get_name (lookup_rel n env)
| _ -> name
in
make_prime avoid name
@@ -2358,7 +2435,7 @@ let build_dependent_signature env evdref avoid tomatchs arsign =
(LocalAssum (Name (eq_id avoid previd), eq)) :: argeqs,
refl_arg :: refl_args,
pred slift,
- set_name (Name id) decl :: argsign'))
+ RelDecl.set_name (Name id) decl :: argsign'))
(env, neqs, [], [], slift, []) args argsign
in
let eq = mk_JMeq evdref
@@ -2373,13 +2450,13 @@ let build_dependent_signature env evdref avoid tomatchs arsign =
succ nargeqs,
refl_eq :: refl_args,
pred slift,
- ((set_name (Name id) app_decl :: argsign') :: arsigns))
+ ((RelDecl.set_name (Name id) app_decl :: argsign') :: arsigns))
| _ -> (* Non dependent inductive or not inductive, just use a regular equality *)
let decl = match arsign with [x] -> x | _ -> assert(false) in
- let name = get_name decl in
+ let name = RelDecl.get_name decl in
let previd, id = make_prime avoid name in
- let arsign' = set_name (Name id) decl in
+ let arsign' = RelDecl.set_name (Name id) decl in
let tomatch_ty = type_of_tomatch ty in
let eq =
mk_eq evdref (lift nar tomatch_ty)
@@ -2401,10 +2478,10 @@ let context_of_arsign l =
l ([], 0)
in x
-let compile_program_cases loc style (typing_function, evdref) tycon env
+let compile_program_cases ?loc style (typing_function, evdref) tycon env lvar
(predopt, tomatchl, eqns) =
let typing_fun tycon env = function
- | Some t -> typing_function tycon env evdref t
+ | Some t -> typing_function tycon env evdref lvar t
| None -> Evarutil.evd_comb0 use_unit_judge evdref in
(* We build the matrix of patterns and right-hand side *)
@@ -2412,29 +2489,30 @@ let compile_program_cases loc style (typing_function, evdref) tycon env
(* We build the vector of terms to match consistently with the *)
(* constructors found in patterns *)
- let tomatchs = coerce_to_indtype typing_function evdref env matx tomatchl in
+ let predlvar,tomatchs = coerce_to_indtype typing_function evdref env lvar matx tomatchl in
let tycon = valcon_of_tycon tycon in
- let tomatchs, tomatchs_lets, tycon' = abstract_tomatch env tomatchs tycon in
+ let tomatchs, tomatchs_lets, tycon' = abstract_tomatch env !evdref tomatchs tycon in
let env = push_rel_context tomatchs_lets env in
let len = List.length eqns in
let sign, allnames, signlen, eqs, neqs, args =
(* The arity signature *)
- let arsign = extract_arity_signature ~dolift:false env tomatchs tomatchl in
+ let arsign = extract_arity_signature ~dolift:false env predlvar tomatchs tomatchl in
+ let arsign = List.map fst arsign in (* Because no difference between the arity for typing and the arity for building *)
(* Build the dependent arity signature, the equalities which makes
the first part of the predicate and their instantiations. *)
- let avoid = [] in
+ let avoid = Id.Set.empty in
build_dependent_signature env evdref avoid tomatchs arsign
in
let tycon, arity =
+ let nar = List.fold_left (fun n sign -> List.length sign + n) 0 sign in
match tycon' with
- | None -> let ev = mkExistential env evdref in ev, ev
+ | None -> let ev = mkExistential env evdref in ev, lift nar ev
| Some t ->
let pred =
match prepare_predicate_from_arsign_tycon env !evdref loc tomatchs sign t with
| Some (evd, pred) -> evdref := evd; pred
| None ->
- let nar = List.fold_left (fun n sign -> List.length sign + n) 0 sign in
lift nar t
in Option.get tycon, pred
in
@@ -2468,25 +2546,26 @@ let compile_program_cases loc style (typing_function, evdref) tycon env
List.map (fun (c,d) -> (c,extract_inductive_data env !evdref d,d)) typs in
let dep_sign =
- find_dependencies_signature
+ find_dependencies_signature !evdref
(List.make (List.length typs) true)
typs in
let typs' =
List.map3
(fun (tm,tmt) deps na ->
- let deps = if not (isRel tm) then [] else deps in
+ let deps = if not (isRel !evdref tm) then [] else deps in
((tm,tmt),deps,na))
tomatchs dep_sign nal in
let initial_pushed = List.map (fun x -> Pushed (true,x)) typs' in
let typing_function tycon env evdref = function
- | Some t -> typing_function tycon env evdref t
+ | Some t -> typing_function tycon env evdref lvar t
| None -> evd_comb0 use_unit_judge evdref in
let pb =
{ env = env;
+ lvar = lvar;
evdref = evdref;
pred = pred;
tomatch = initial_pushed;
@@ -2499,19 +2578,19 @@ let compile_program_cases loc style (typing_function, evdref) tycon env
let j = compile pb in
(* We check for unused patterns *)
List.iter (check_unused_pattern env) matx;
- let body = it_mkLambda_or_LetIn (applistc j.uj_val args) lets in
+ let body = it_mkLambda_or_LetIn (applist (j.uj_val, args)) lets in
let j =
{ uj_val = it_mkLambda_or_LetIn body tomatchs_lets;
- uj_type = nf_evar !evdref tycon; }
+ uj_type = EConstr.of_constr (EConstr.to_constr !evdref tycon); }
in j
(**************************************************************************)
(* Main entry of the matching compilation *)
-let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, eqns) =
+let compile_cases ?loc style (typing_fun, evdref) tycon env lvar (predopt, tomatchl, eqns) =
if predopt == None && Flags.is_program_mode () && Program.is_program_cases () then
- compile_program_cases loc style (typing_fun, evdref)
- tycon env (predopt, tomatchl, eqns)
+ compile_program_cases ?loc style (typing_fun, evdref)
+ tycon env lvar (predopt, tomatchl, eqns)
else
(* We build the matrix of patterns and right-hand side *)
@@ -2519,15 +2598,15 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e
(* We build the vector of terms to match consistently with the *)
(* constructors found in patterns *)
- let tomatchs = coerce_to_indtype typing_fun evdref env matx tomatchl in
+ let predlvar,tomatchs = coerce_to_indtype typing_fun evdref env lvar matx tomatchl in
(* If an elimination predicate is provided, we check it is compatible
with the type of arguments to match; if none is provided, we
build alternative possible predicates *)
- let arsign = extract_arity_signature env tomatchs tomatchl in
- let preds = prepare_predicate loc typing_fun env !evdref tomatchs arsign tycon predopt in
+ let arsign = extract_arity_signature env predlvar tomatchs tomatchl in
+ let preds = prepare_predicate ?loc typing_fun env !evdref predlvar tomatchs arsign tycon predopt in
let compile_for_one_predicate (sigma,nal,pred) =
(* We push the initial terms to match and push their alias to rhs' envs *)
@@ -2543,14 +2622,14 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e
List.map (fun (c,d) -> (c,extract_inductive_data env sigma d,d)) typs in
let dep_sign =
- find_dependencies_signature
+ find_dependencies_signature !evdref
(List.make (List.length typs) true)
typs in
let typs' =
List.map3
(fun (tm,tmt) deps na ->
- let deps = if not (isRel tm) then [] else deps in
+ let deps = if not (isRel !evdref tm) then [] else deps in
((tm,tmt),deps,na))
tomatchs dep_sign nal in
@@ -2558,13 +2637,14 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e
(* A typing function that provides with a canonical term for absurd cases*)
let typing_fun tycon env evdref = function
- | Some t -> typing_fun tycon env evdref t
+ | Some t -> typing_fun tycon env evdref lvar t
| None -> evd_comb0 use_unit_judge evdref in
let myevdref = ref sigma in
let pb =
{ env = env;
+ lvar = lvar;
evdref = myevdref;
pred = pred;
tomatch = initial_pushed;
@@ -2577,7 +2657,7 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e
let j = compile pb in
(* We coerce to the tycon (if an elim predicate was provided) *)
- let j = inh_conv_coerce_to_tycon loc env myevdref j tycon in
+ let j = inh_conv_coerce_to_tycon ?loc env myevdref j tycon in
evdref := !myevdref;
j in
diff --git a/pretyping/cases.mli b/pretyping/cases.mli
index ee4148de..04a34646 100644
--- a/pretyping/cases.mli
+++ b/pretyping/cases.mli
@@ -1,18 +1,22 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
+open Constr
open Evd
open Environ
+open EConstr
open Inductiveops
open Glob_term
-open Evarutil
+open Ltac_pretype
+open Evardefine
(** {5 Compilation of pattern-matching } *)
@@ -28,43 +32,43 @@ type pattern_matching_error =
exception PatternMatchingError of env * evar_map * pattern_matching_error
-val error_wrong_numarg_constructor_loc : Loc.t -> env -> constructor -> int -> 'a
+val error_wrong_numarg_constructor : ?loc:Loc.t -> env -> constructor -> int -> 'a
-val error_wrong_numarg_inductive_loc : Loc.t -> env -> inductive -> int -> 'a
+val error_wrong_numarg_inductive : ?loc:Loc.t -> env -> inductive -> int -> 'a
val irrefutable : env -> cases_pattern -> bool
(** {6 Compilation primitive. } *)
val compile_cases :
- Loc.t -> case_style ->
- (type_constraint -> env -> evar_map ref -> glob_constr -> unsafe_judgment) * evar_map ref ->
+ ?loc:Loc.t -> case_style ->
+ (type_constraint -> env -> evar_map ref -> ltac_var_map -> glob_constr -> unsafe_judgment) * evar_map ref ->
type_constraint ->
- env -> glob_constr option * tomatch_tuples * cases_clauses ->
+ env -> ltac_var_map -> glob_constr option * tomatch_tuples * cases_clauses ->
unsafe_judgment
val constr_of_pat :
Environ.env ->
Evd.evar_map ref ->
- Context.Rel.Declaration.t list ->
+ rel_context ->
Glob_term.cases_pattern ->
- Names.Id.t list ->
+ Names.Id.Set.t ->
Glob_term.cases_pattern *
- (Context.Rel.Declaration.t list * Term.constr *
- (Term.types * Term.constr list) * Glob_term.cases_pattern) *
- Names.Id.t list
+ (rel_context * constr *
+ (types * constr list) * Glob_term.cases_pattern) *
+ Names.Id.Set.t
type 'a rhs =
{ rhs_env : env;
- rhs_vars : Id.t list;
- avoid_ids : Id.t list;
+ rhs_vars : Id.Set.t;
+ avoid_ids : Id.Set.t;
it : 'a option}
type 'a equation =
{ patterns : cases_pattern list;
rhs : 'a rhs;
alias_stack : Name.t list;
- eqn_loc : Loc.t;
+ eqn_loc : Loc.t option;
used : bool ref }
type 'a matrix = 'a equation list
@@ -84,7 +88,7 @@ type tomatch_status =
| Pushed of (bool*((constr * tomatch_type) * int list * Name.t))
| Alias of (bool * (Name.t * constr * (constr * types)))
| NonDepAlias
- | Abstract of int * Context.Rel.Declaration.t
+ | Abstract of int * rel_declaration
type tomatch_stack = tomatch_status list
@@ -100,25 +104,29 @@ and pattern_continuation =
type 'a pattern_matching_problem =
{ env : env;
+ lvar : Ltac_pretype.ltac_var_map;
evdref : evar_map ref;
pred : constr;
tomatch : tomatch_stack;
history : pattern_continuation;
mat : 'a matrix;
- caseloc : Loc.t;
+ caseloc : Loc.t option;
casestyle : case_style;
typing_function: type_constraint -> env -> evar_map ref -> 'a option -> unsafe_judgment }
-val compile : 'a pattern_matching_problem -> Environ.unsafe_judgment
+val compile : 'a pattern_matching_problem -> unsafe_judgment
-val prepare_predicate : Loc.t ->
- (Evarutil.type_constraint ->
- Environ.env -> Evd.evar_map ref -> glob_constr -> Environ.unsafe_judgment) ->
+val prepare_predicate : ?loc:Loc.t ->
+ (type_constraint ->
+ Environ.env -> Evd.evar_map ref -> ltac_var_map -> glob_constr -> unsafe_judgment) ->
Environ.env ->
Evd.evar_map ->
- (Term.types * tomatch_type) list ->
- Context.Rel.t list ->
- Constr.constr option ->
- glob_constr option ->
- (Evd.evar_map * Names.name list * Term.constr) list
+ Ltac_pretype.ltac_var_map ->
+ (types * tomatch_type) list ->
+ (rel_context * rel_context) list ->
+ constr option ->
+ glob_constr option -> (Evd.evar_map * Name.t list * constr) list
+
+val make_return_predicate_ltac_lvar : Evd.evar_map -> Name.t ->
+ Glob_term.glob_constr -> constr -> Ltac_pretype.ltac_var_map -> ltac_var_map
diff --git a/pretyping/cbv.ml b/pretyping/cbv.ml
index 84bf849e..cb0fc325 100644
--- a/pretyping/cbv.ml
+++ b/pretyping/cbv.ml
@@ -1,14 +1,16 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Util
open Names
-open Term
+open Constr
open Vars
open CClosure
open Esubst
@@ -45,7 +47,7 @@ type cbv_value =
| LAM of int * (Name.t * constr) list * constr * cbv_value subs
| FIXP of fixpoint * cbv_value subs * cbv_value array
| COFIXP of cofixpoint * cbv_value subs * cbv_value array
- | CONSTR of constructor puniverses * cbv_value array
+ | CONSTR of constructor Univ.puniverses * cbv_value array
(* type of terms with a hole. This hole can appear only under App or Case.
* TOP means the term is considered without context
@@ -69,7 +71,7 @@ and cbv_stack =
| TOP
| APP of cbv_value array * cbv_stack
| CASE of constr * constr array * case_info * cbv_value subs * cbv_stack
- | PROJ of projection * Declarations.projection_body * cbv_stack
+ | PROJ of Projection.t * Declarations.projection_body * cbv_stack
(* les vars pourraient etre des constr,
cela permet de retarder les lift: utile ?? *)
@@ -132,6 +134,7 @@ let mkSTACK = function
| STACK(0,v0,stk0), stk -> STACK(0,v0,stack_concat stk0 stk)
| v,stk -> STACK(0,v,stk)
+type cbv_infos = { tab : cbv_value infos_tab; infos : cbv_value infos; sigma : Evd.evar_map }
(* Change: zeta reduction cannot be avoided in CBV *)
@@ -170,11 +173,68 @@ let fixp_reducible flgs ((reci,i),_) stk =
let cofixp_reducible flgs _ stk =
if red_set flgs fCOFIX then
match stk with
- | (CASE _ | APP(_,CASE _)) -> true
+ | (CASE _ | PROJ _ | APP(_,CASE _) | APP(_,PROJ _)) -> true
| _ -> false
else
false
+let debug_cbv = ref false
+let _ = Goptions.declare_bool_option {
+ Goptions.optdepr = false;
+ Goptions.optname = "cbv visited constants display";
+ Goptions.optkey = ["Debug";"Cbv"];
+ Goptions.optread = (fun () -> !debug_cbv);
+ Goptions.optwrite = (fun a -> debug_cbv:=a);
+}
+
+let pr_key = function
+ | ConstKey (sp,_) -> Names.Constant.print sp
+ | VarKey id -> Names.Id.print id
+ | RelKey n -> Pp.(str "REL_" ++ int n)
+
+let rec reify_stack t = function
+ | TOP -> t
+ | APP (args,st) ->
+ reify_stack (mkApp(t,Array.map reify_value args)) st
+ | CASE (ty,br,ci,env,st) ->
+ reify_stack
+ (mkCase (ci, ty, t,br))
+ st
+ | PROJ (p, pinfo, st) ->
+ reify_stack (mkProj (p, t)) st
+
+and reify_value = function (* reduction under binders *)
+ | VAL (n,t) -> lift n t
+ | STACK (0,v,stk) ->
+ reify_stack (reify_value v) stk
+ | STACK (n,v,stk) ->
+ lift n (reify_stack (reify_value v) stk)
+ | CBN(t,env) ->
+ apply_env env t
+ | LAM (k,ctxt,b,env) ->
+ apply_env env @@
+ List.fold_left (fun c (n,t) ->
+ mkLambda (n, t, c)) b ctxt
+ | FIXP ((lij,(names,lty,bds)),env,args) ->
+ let fix = mkFix (lij, (names, lty, bds)) in
+ mkApp (apply_env env fix, Array.map reify_value args)
+ | COFIXP ((j,(names,lty,bds)),env,args) ->
+ let cofix = mkCoFix (j, (names,lty,bds)) in
+ mkApp (apply_env env cofix, Array.map reify_value args)
+ | CONSTR (c,args) ->
+ mkApp(mkConstructU c, Array.map reify_value args)
+
+and apply_env env t =
+ match kind t with
+ | Rel i ->
+ begin match expand_rel i env with
+ | Inl (k, v) ->
+ reify_value (shift_value k v)
+ | Inr (k,_) ->
+ mkRel k
+ end
+ | _ ->
+ map_with_binders subs_lift apply_env env t
(* The main recursive functions
*
@@ -189,7 +249,7 @@ let cofixp_reducible flgs _ stk =
let rec norm_head info env t stack =
(* no reduction under binders *)
- match kind_of_term t with
+ match kind t with
(* stack grows (remove casts) *)
| App (head,args) -> (* Applied terms are normalized immediately;
they could be computed when getting out of the stack *)
@@ -200,12 +260,12 @@ let rec norm_head info env t stack =
| Proj (p, c) ->
let p' =
- if red_set (info_flags info) (fCONST (Projection.constant p))
- && red_set (info_flags info) fBETA
+ if red_set (info_flags info.infos) (fCONST (Projection.constant p))
+ && red_set (info_flags info.infos) fBETA
then Projection.unfold p
else p
in
- let pinfo = Environ.lookup_projection p (info_env info) in
+ let pinfo = Environ.lookup_projection p (info_env info.infos) in
norm_head info env c (PROJ (p', pinfo, stack))
(* constants, axioms
@@ -220,14 +280,16 @@ let rec norm_head info env t stack =
| Var id -> norm_head_ref 0 info env stack (VarKey id)
- | Const sp -> norm_head_ref 0 info env stack (ConstKey sp)
+ | Const sp ->
+ Reductionops.reduction_effect_hook (env_of_infos info.infos) info.sigma t (lazy (reify_stack t stack));
+ norm_head_ref 0 info env stack (ConstKey sp)
| LetIn (_, b, _, c) ->
(* zeta means letin are contracted; delta without zeta means we *)
(* allow bindings but leave let's in place *)
- if red_set (info_flags info) fZETA then
+ if red_set (info_flags info.infos) fZETA then
(* New rule: for Cbv, Delta does not apply to locally bound variables
- or red_set (info_flags info) fDELTA
+ or red_set (info_flags info.infos) fDELTA
*)
let env' = subs_cons ([|cbv_stack_term info TOP env b|],env) in
norm_head info env' c stack
@@ -235,13 +297,16 @@ let rec norm_head info env t stack =
(CBN(t,env), stack) (* Should we consider a commutative cut ? *)
| Evar ev ->
- (match evar_value info.i_cache ev with
+ (match evar_value info.infos.i_cache ev with
Some c -> norm_head info env c stack
- | None -> (VAL(0, t), stack))
+ | None ->
+ let e, xs = ev in
+ let xs' = Array.map (apply_env env) xs in
+ (VAL(0, mkEvar (e,xs')), stack))
(* non-neutral cases *)
| Lambda _ ->
- let ctxt,b = decompose_lam t in
+ let ctxt,b = Term.decompose_lam t in
(LAM(List.length ctxt, List.rev ctxt,b,env), stack)
| Fix fix -> (FIXP(fix,env,[||]), stack)
| CoFix cofix -> (COFIXP(cofix,env,[||]), stack)
@@ -252,11 +317,19 @@ let rec norm_head info env t stack =
| Prod _ -> (CBN(t,env), stack)
and norm_head_ref k info env stack normt =
- if red_set_ref (info_flags info) normt then
- match ref_value_cache info normt with
- | Some body -> strip_appl (shift_value k body) stack
- | None -> (VAL(0,make_constr_ref k normt),stack)
- else (VAL(0,make_constr_ref k normt),stack)
+ if red_set_ref (info_flags info.infos) normt then
+ match ref_value_cache info.infos info.tab normt with
+ | Some body ->
+ if !debug_cbv then Feedback.msg_debug Pp.(str "Unfolding " ++ pr_key normt);
+ strip_appl (shift_value k body) stack
+ | None ->
+ if !debug_cbv then Feedback.msg_debug Pp.(str "Not unfolding " ++ pr_key normt);
+ (VAL(0,make_constr_ref k normt),stack)
+ else
+ begin
+ if !debug_cbv then Feedback.msg_debug Pp.(str "Not unfolding " ++ pr_key normt);
+ (VAL(0,make_constr_ref k normt),stack)
+ end
(* cbv_stack_term performs weak reduction on constr t under the subs
* env, with context stack, i.e. ([env]t stack). First computes weak
@@ -270,7 +343,7 @@ and cbv_stack_term info stack env t =
and cbv_stack_value info env = function
(* a lambda meets an application -> BETA *)
| (LAM (nlams,ctxt,b,env), APP (args, stk))
- when red_set (info_flags info) fBETA ->
+ when red_set (info_flags info.infos) fBETA ->
let nargs = Array.length args in
if nargs == nlams then
cbv_stack_term info stk (subs_cons(args,env)) b
@@ -284,31 +357,31 @@ and cbv_stack_value info env = function
(* a Fix applied enough -> IOTA *)
| (FIXP(fix,env,[||]), stk)
- when fixp_reducible (info_flags info) fix stk ->
+ when fixp_reducible (info_flags info.infos) fix stk ->
let (envf,redfix) = contract_fixp env fix in
cbv_stack_term info stk envf redfix
(* constructor guard satisfied or Cofix in a Case -> IOTA *)
| (COFIXP(cofix,env,[||]), stk)
- when cofixp_reducible (info_flags info) cofix stk->
+ when cofixp_reducible (info_flags info.infos) cofix stk->
let (envf,redfix) = contract_cofixp env cofix in
cbv_stack_term info stk envf redfix
(* constructor in a Case -> IOTA *)
| (CONSTR(((sp,n),u),[||]), APP(args,CASE(_,br,ci,env,stk)))
- when red_set (info_flags info) fMATCH ->
+ when red_set (info_flags info.infos) fMATCH ->
let cargs =
Array.sub args ci.ci_npar (Array.length args - ci.ci_npar) in
cbv_stack_term info (stack_app cargs stk) env br.(n-1)
(* constructor of arity 0 in a Case -> IOTA *)
| (CONSTR(((_,n),u),[||]), CASE(_,br,_,env,stk))
- when red_set (info_flags info) fMATCH ->
+ when red_set (info_flags info.infos) fMATCH ->
cbv_stack_term info stk env br.(n-1)
(* constructor in a Projection -> IOTA *)
| (CONSTR(((sp,n),u),[||]), APP(args,PROJ(p,pi,stk)))
- when red_set (info_flags info) fMATCH && Projection.unfolded p ->
+ when red_set (info_flags info.infos) fMATCH && Projection.unfolded p ->
let arg = args.(pi.Declarations.proj_npars + pi.Declarations.proj_arg) in
cbv_stack_value info env (strip_appl arg stk)
@@ -316,7 +389,7 @@ and cbv_stack_value info env = function
| (FIXP(fix,env,[||]), APP(appl,TOP)) -> FIXP(fix,env,appl)
| (COFIXP(cofix,env,[||]), APP(appl,TOP)) -> COFIXP(cofix,env,appl)
| (CONSTR(c,[||]), APP(appl,TOP)) -> CONSTR(c,appl)
-
+
(* definitely a value *)
| (head,stk) -> mkSTACK(head, stk)
@@ -350,12 +423,12 @@ and cbv_norm_value info = function (* reduction under binders *)
| STACK (n,v,stk) ->
lift n (apply_stack info (cbv_norm_value info v) stk)
| CBN(t,env) ->
- map_constr_with_binders subs_lift (cbv_norm_term info) env t
+ Constr.map_with_binders subs_lift (cbv_norm_term info) env t
| LAM (n,ctxt,b,env) ->
let nctxt =
List.map_i (fun i (x,ty) ->
(x,cbv_norm_term info (subs_liftn i env) ty)) 0 ctxt in
- compose_lam (List.rev nctxt) (cbv_norm_term info (subs_liftn n env) b)
+ Term.compose_lam (List.rev nctxt) (cbv_norm_term info (subs_liftn n env) b)
| FIXP ((lij,(names,lty,bds)),env,args) ->
mkApp
(mkFix (lij,
@@ -376,14 +449,14 @@ and cbv_norm_value info = function (* reduction under binders *)
(* with profiling *)
let cbv_norm infos constr =
- with_stats (lazy (cbv_norm_term infos (subs_id 0) constr))
-
-type cbv_infos = cbv_value infos
+ let constr = EConstr.Unsafe.to_constr constr in
+ EConstr.of_constr (with_stats (lazy (cbv_norm_term infos (subs_id 0) constr)))
(* constant bodies are normalized at the first expansion *)
let create_cbv_infos flgs env sigma =
- create
- (fun old_info c -> cbv_stack_term old_info TOP (subs_id 0) c)
+ let infos = create
+ (fun old_info tab c -> cbv_stack_term { tab; infos = old_info; sigma } TOP (subs_id 0) c)
flgs
env
- (Reductionops.safe_evar_value sigma)
+ (Reductionops.safe_evar_value sigma) in
+ { tab = CClosure.create_tab (); infos; sigma }
diff --git a/pretyping/cbv.mli b/pretyping/cbv.mli
index 87a03abb..cdaa39c5 100644
--- a/pretyping/cbv.mli
+++ b/pretyping/cbv.mli
@@ -1,13 +1,15 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
+open EConstr
open Environ
open CClosure
open Esubst
@@ -23,6 +25,9 @@ val cbv_norm : cbv_infos -> constr -> constr
(***********************************************************************
i This is for cbv debug *)
+
+open Constr
+
type cbv_value =
| VAL of int * constr
| STACK of int * cbv_value * cbv_stack
@@ -30,13 +35,13 @@ type cbv_value =
| LAM of int * (Name.t * constr) list * constr * cbv_value subs
| FIXP of fixpoint * cbv_value subs * cbv_value array
| COFIXP of cofixpoint * cbv_value subs * cbv_value array
- | CONSTR of constructor puniverses * cbv_value array
+ | CONSTR of constructor Univ.puniverses * cbv_value array
and cbv_stack =
| TOP
| APP of cbv_value array * cbv_stack
| CASE of constr * constr array * case_info * cbv_value subs * cbv_stack
- | PROJ of projection * Declarations.projection_body * cbv_stack
+ | PROJ of Projection.t * Declarations.projection_body * cbv_stack
val shift_value : int -> cbv_value -> cbv_value
diff --git a/pretyping/classops.ml b/pretyping/classops.ml
index 4f265e76..a0804b72 100644
--- a/pretyping/classops.ml
+++ b/pretyping/classops.ml
@@ -1,23 +1,23 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 CErrors
open Util
open Pp
-open Flags
open Names
+open Constr
open Libnames
open Globnames
open Nametab
open Environ
open Libobject
-open Term
-open Termops
open Mod_subst
(* usage qque peu general: utilise aussi dans record *)
@@ -29,9 +29,9 @@ type cl_typ =
| CL_SORT
| CL_FUN
| CL_SECVAR of variable
- | CL_CONST of constant
+ | CL_CONST of Constant.t
| CL_IND of inductive
- | CL_PROJ of constant
+ | CL_PROJ of Constant.t
type cl_info_typ = {
cl_param : int
@@ -45,12 +45,13 @@ type coe_info_typ = {
coe_value : constr;
coe_type : types;
coe_local : bool;
- coe_context : Univ.universe_context_set;
+ coe_context : Univ.ContextSet.t;
coe_is_identity : bool;
coe_is_projection : bool;
coe_param : int }
let coe_info_typ_equal c1 c2 =
+ let eq_constr c1 c2 = Termops.eq_constr Evd.empty (EConstr.of_constr c1) (EConstr.of_constr c2) in
eq_constr c1.coe_value c2.coe_value &&
eq_constr c1.coe_type c2.coe_type &&
c1.coe_local == c2.coe_local &&
@@ -60,8 +61,8 @@ let coe_info_typ_equal c1 c2 =
let cl_typ_ord t1 t2 = match t1, t2 with
| CL_SECVAR v1, CL_SECVAR v2 -> Id.compare v1 v2
- | CL_CONST c1, CL_CONST c2 -> con_ord c1 c2
- | CL_PROJ c1, CL_PROJ c2 -> con_ord c1 c2
+ | CL_CONST c1, CL_CONST c2 -> Constant.CanOrd.compare c1 c2
+ | CL_PROJ c1, CL_PROJ c2 -> Constant.CanOrd.compare c1 c2
| CL_IND i1, CL_IND i2 -> ind_ord i1 i2
| _ -> Pervasives.compare t1 t2 (** OK *)
@@ -89,7 +90,7 @@ sig
type t
val compare : t -> t -> int
val equal : t -> t -> bool
- val print : t -> std_ppcmds
+ val print : t -> Pp.t
end
type 'a t
val empty : 'a t
@@ -192,15 +193,16 @@ let coercion_exists coe = CoeTypMap.mem coe !coercion_tab
(* find_class_type : evar_map -> constr -> cl_typ * universe_list * constr list *)
let find_class_type sigma t =
+ let open EConstr in
let t', args = Reductionops.whd_betaiotazeta_stack sigma t in
- match kind_of_term t' with
- | Var id -> CL_SECVAR id, Univ.Instance.empty, args
+ match EConstr.kind sigma t' with
+ | Var id -> CL_SECVAR id, EInstance.empty, args
| Const (sp,u) -> CL_CONST sp, u, args
| Proj (p, c) when not (Projection.unfolded p) ->
- CL_PROJ (Projection.constant p), Univ.Instance.empty, c :: args
+ CL_PROJ (Projection.constant p), EInstance.empty, (c :: args)
| Ind (ind_sp,u) -> CL_IND ind_sp, u, args
- | Prod (_,_,_) -> CL_FUN, Univ.Instance.empty, []
- | Sort _ -> CL_SORT, Univ.Instance.empty, []
+ | Prod (_,_,_) -> CL_FUN, EInstance.empty, []
+ | Sort _ -> CL_SORT, EInstance.empty, []
| _ -> raise Not_found
@@ -214,7 +216,7 @@ let subst_cl_typ subst ct = match ct with
| CL_CONST c ->
let c',t = subst_con_kn subst c in
if c' == c then ct else
- pi1 (find_class_type Evd.empty t)
+ pi1 (find_class_type Evd.empty (EConstr.of_constr t))
| CL_IND i ->
let i' = subst_ind subst i in
if i' == i then ct else CL_IND i'
@@ -297,9 +299,9 @@ let lookup_path_to_sort_from env sigma s =
let get_coercion_constructor env coe =
let c, _ =
- Reductionops.whd_all_stack env Evd.empty coe.coe_value
+ Reductionops.whd_all_stack env Evd.empty (EConstr.of_constr coe.coe_value)
in
- match kind_of_term c with
+ match EConstr.kind Evd.empty (** FIXME *) c with
| Construct (cstr,u) ->
(cstr, Inductiveops.constructor_nrealargs cstr -1)
| _ ->
@@ -317,21 +319,21 @@ let coercion_value { coe_value = c; coe_type = t; coe_context = ctx;
let subst, ctx = Universes.fresh_universe_context_set_instance ctx in
let c' = Vars.subst_univs_level_constr subst c
and t' = Vars.subst_univs_level_constr subst t in
- (make_judge c' t', b, b'), ctx
+ (make_judge (EConstr.of_constr c') (EConstr.of_constr t'), b, b'), ctx
(* pretty-print functions are now in Pretty *)
(* rajouter une coercion dans le graphe *)
-let path_printer = ref (fun _ -> str "<a class path>"
- : (Bijint.Index.t * Bijint.Index.t) * inheritance_path -> std_ppcmds)
+let path_printer : (env -> Evd.evar_map -> (Bijint.Index.t * Bijint.Index.t) * inheritance_path -> Pp.t) ref =
+ ref (fun _ _ _ -> str "<a class path>")
let install_path_printer f = path_printer := f
-let print_path x = !path_printer x
+let print_path env sigma x = !path_printer env sigma x
-let message_ambig l =
- (str"Ambiguous paths:" ++ spc () ++
- prlist_with_sep fnl (fun ijp -> print_path ijp) l)
+let message_ambig env sigma l =
+ str"Ambiguous paths:" ++ spc () ++
+ prlist_with_sep fnl (fun ijp -> print_path env sigma ijp) l
(* add_coercion_in_graph : coe_index * cl_index * cl_index -> unit
coercion,source,target *)
@@ -344,8 +346,8 @@ let different_class_params i =
| CL_IND i -> Global.is_polymorphic (IndRef i)
| CL_CONST c -> Global.is_polymorphic (ConstRef c)
| _ -> false
-
-let add_coercion_in_graph (ic,source,target) =
+
+let add_coercion_in_graph env sigma (ic,source,target) =
let old_inheritance_graph = !inheritance_graph in
let ambig_paths =
(ref [] : ((cl_index * cl_index) * inheritance_path) list ref) in
@@ -386,8 +388,8 @@ let add_coercion_in_graph (ic,source,target) =
old_inheritance_graph
end;
let is_ambig = match !ambig_paths with [] -> false | _ -> true in
- if is_ambig && is_verbose () then
- Feedback.msg_info (message_ambig !ambig_paths)
+ if is_ambig && not !Flags.quiet then
+ Feedback.msg_info (message_ambig env sigma !ambig_paths)
type coercion = {
coercion_type : coe_typ;
@@ -402,8 +404,8 @@ type coercion = {
(* Computation of the class arity *)
let reference_arity_length ref =
- let t = Universes.unsafe_type_of_global ref in
- List.length (fst (Reductionops.splay_arity (Global.env()) Evd.empty t))
+ let t, _ = Global.type_of_global_in_context (Global.env ()) ref in
+ List.length (fst (Reductionops.splay_arity (Global.env()) Evd.empty (EConstr.of_constr t))) (** FIXME *)
let projection_arity_length p =
let len = reference_arity_length (ConstRef p) in
@@ -427,20 +429,20 @@ let automatically_import_coercions = ref false
open Goptions
let _ =
declare_bool_option
- { optsync = true;
- optdepr = false;
+ { optdepr = true; (* remove in 8.8 *)
optname = "automatic import of coercions";
optkey = ["Automatic";"Coercions";"Import"];
optread = (fun () -> !automatically_import_coercions);
optwrite = (:=) automatically_import_coercions }
-let cache_coercion (_, c) =
+let cache_coercion env sigma (_, c) =
let () = add_class c.coercion_source in
let () = add_class c.coercion_target in
let is, _ = class_info c.coercion_source in
let it, _ = class_info c.coercion_target in
- let value, ctx = Universes.fresh_global_instance (Global.env()) c.coercion_type in
- let typ = Retyping.get_type_of (Global.env ()) Evd.empty value in
+ let value, ctx = Universes.fresh_global_instance env c.coercion_type in
+ let typ = Retyping.get_type_of env sigma (EConstr.of_constr value) in
+ let typ = EConstr.Unsafe.to_constr typ in
let xf =
{ coe_value = value;
coe_type = typ;
@@ -450,19 +452,15 @@ let cache_coercion (_, c) =
coe_is_projection = c.coercion_is_proj;
coe_param = c.coercion_params } in
let () = add_new_coercion c.coercion_type xf in
- add_coercion_in_graph (xf,is,it)
+ add_coercion_in_graph env sigma (xf,is,it)
let load_coercion _ o =
- if
- !automatically_import_coercions || Flags.version_less_or_equal Flags.V8_2
- then
- cache_coercion o
+ if !automatically_import_coercions then
+ cache_coercion (Global.env ()) Evd.empty o
let open_coercion i o =
- if Int.equal i 1 && not
- (!automatically_import_coercions || Flags.version_less_or_equal Flags.V8_2)
- then
- cache_coercion o
+ if Int.equal i 1 && not !automatically_import_coercions then
+ cache_coercion (Global.env ()) Evd.empty o
let subst_coercion (subst, c) =
let coe = subst_coe_typ subst c.coercion_type in
@@ -501,7 +499,9 @@ let inCoercion : coercion -> obj =
declare_object {(default_object "COERCION") with
open_function = open_coercion;
load_function = load_coercion;
- cache_function = cache_coercion;
+ cache_function = (fun objn ->
+ let env = Global.env () in cache_coercion env Evd.empty objn
+ );
subst_function = subst_coercion;
classify_function = classify_coercion;
discharge_function = discharge_coercion }
@@ -538,7 +538,7 @@ let inheritance_graph () =
let coercion_of_reference r =
let ref = Nametab.global r in
if not (coercion_exists ref) then
- errorlabstrm "try_add_coercion"
+ user_err ~hdr:"try_add_coercion"
(Nametab.pr_global_env Id.Set.empty ref ++ str" is not a coercion.");
ref
@@ -554,7 +554,6 @@ module CoercionPrinting =
let member_message x b =
str "Explicit printing of coercion " ++ printer x ++
str (if b then " is set" else " is unset")
- let synchronous = true
end
module PrintingCoercion = Goptions.MakeRefTable(CoercionPrinting)
diff --git a/pretyping/classops.mli b/pretyping/classops.mli
index d509739c..f8600bbe 100644
--- a/pretyping/classops.mli
+++ b/pretyping/classops.mli
@@ -1,15 +1,17 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
-open Evd
open Environ
+open EConstr
+open Evd
open Mod_subst
(** {6 This is the type of class kinds } *)
@@ -17,9 +19,9 @@ type cl_typ =
| CL_SORT
| CL_FUN
| CL_SECVAR of variable
- | CL_CONST of constant
+ | CL_CONST of Constant.t
| CL_IND of inductive
- | CL_PROJ of constant
+ | CL_PROJ of Constant.t
(** Equality over [cl_typ] *)
val cl_typ_eq : cl_typ -> cl_typ -> bool
@@ -59,7 +61,7 @@ val class_info_from_index : cl_index -> cl_typ * cl_info_typ
(** [find_class_type env sigma c] returns the head reference of [c],
its universe instance and its arguments *)
-val find_class_type : evar_map -> types -> cl_typ * Univ.universe_instance * constr list
+val find_class_type : evar_map -> types -> cl_typ * EInstance.t * constr list
(** raises [Not_found] if not convertible to a class *)
val class_of : env -> evar_map -> types -> types * cl_index
@@ -95,16 +97,15 @@ val lookup_pattern_path_between :
(**/**)
(* Crade *)
-open Pp
val install_path_printer :
- ((cl_index * cl_index) * inheritance_path -> std_ppcmds) -> unit
+ (env -> evar_map -> (cl_index * cl_index) * inheritance_path -> Pp.t) -> unit
(**/**)
(** {6 This is for printing purpose } *)
val string_of_class : cl_typ -> string
-val pr_class : cl_typ -> std_ppcmds
-val pr_cl_index : cl_index -> std_ppcmds
-val get_coercion_value : coe_index -> constr
+val pr_class : cl_typ -> Pp.t
+val pr_cl_index : cl_index -> Pp.t
+val get_coercion_value : coe_index -> Constr.t
val inheritance_graph : unit -> ((cl_index * cl_index) * inheritance_path) list
val classes : unit -> cl_typ list
val coercions : unit -> coe_index list
diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml
index 913e80f3..04cb6a59 100644
--- a/pretyping/coercion.ml
+++ b/pretyping/coercion.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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) *)
(************************************************************************)
(* Created by Hugo Herbelin for Coq V7 by isolating the coercion
@@ -18,10 +20,10 @@ open CErrors
open Util
open Names
open Term
+open Environ
+open EConstr
open Vars
open Reductionops
-open Environ
-open Typeops
open Pretype_errors
open Classops
open Evarutil
@@ -33,14 +35,13 @@ open Globnames
let use_typeclasses_for_conversion = ref true
let _ =
- Goptions.declare_bool_option
- { Goptions.optsync = true;
- optdepr = false;
+ Goptions.(declare_bool_option
+ { optdepr = false;
optname = "use typeclass resolution during conversion";
optkey = ["Typeclass"; "Resolution"; "For"; "Conversion"];
optread = (fun () -> !use_typeclasses_for_conversion);
optwrite = (fun b -> use_typeclasses_for_conversion := b) }
-
+ )
(* Typing operations dealing with coercions *)
exception NoCoercion
@@ -52,7 +53,7 @@ let apply_coercion_args env evd check isproj argl funj =
let rec apply_rec acc typ = function
| [] ->
if isproj then
- let cst = fst (destConst (j_val funj)) in
+ let cst = fst (destConst !evdref (j_val funj)) in
let p = Projection.make cst false in
let pb = lookup_projection p env in
let args = List.skipn pb.Declarations.proj_npars argl in
@@ -63,35 +64,36 @@ let apply_coercion_args env evd check isproj argl funj =
{ uj_val = applist (j_val funj,argl);
uj_type = typ }
| h::restl -> (* On devrait pouvoir s'arranger pour qu'on n'ait pas a faire hnf_constr *)
- match kind_of_term (whd_all env evd typ) with
+ match EConstr.kind !evdref (whd_all env !evdref typ) with
| Prod (_,c1,c2) ->
- if check && not (e_cumul env evdref (Retyping.get_type_of env evd h) c1) then
+ if check && not (e_cumul env evdref (Retyping.get_type_of env !evdref h) c1) then
raise NoCoercion;
apply_rec (h::acc) (subst1 h c2) restl
- | _ -> anomaly (Pp.str "apply_coercion_args")
+ | _ -> anomaly (Pp.str "apply_coercion_args.")
in
let res = apply_rec [] funj.uj_type argl in
!evdref, res
(* appliquer le chemin de coercions de patterns p *)
-let apply_pattern_coercion loc pat p =
+let apply_pattern_coercion ?loc pat p =
List.fold_left
(fun pat (co,n) ->
- let f i = if i<n then Glob_term.PatVar (loc, Anonymous) else pat in
- Glob_term.PatCstr (loc, co, List.init (n+1) f, Anonymous))
+ let f i =
+ if i<n then (DAst.make ?loc @@ Glob_term.PatVar Anonymous) else pat in
+ DAst.make ?loc @@ Glob_term.PatCstr (co, List.init (n+1) f, Anonymous))
pat p
(* raise Not_found if no coercion found *)
-let inh_pattern_coerce_to loc env pat ind1 ind2 =
+let inh_pattern_coerce_to ?loc env pat ind1 ind2 =
let p = lookup_pattern_path_between env (ind1,ind2) in
- apply_pattern_coercion loc pat p
+ apply_pattern_coercion ?loc pat p
(* Program coercions *)
open Program
-let make_existential loc ?(opaque = not (get_proofs_transparency ())) env evdref c =
- let src = (loc, Evar_kinds.QuestionMark (Evar_kinds.Define opaque)) in
+let make_existential ?loc ?(opaque = not (get_proofs_transparency ())) na env evdref c =
+ let src = Loc.tag ?loc (Evar_kinds.QuestionMark (Evar_kinds.Define opaque,na)) in
Evarutil.e_new_evar env evdref ~src c
let app_opt env evdref f t =
@@ -99,10 +101,10 @@ let app_opt env evdref f t =
let pair_of_array a = (a.(0), a.(1))
-let disc_subset x =
- match kind_of_term x with
+let disc_subset sigma x =
+ match EConstr.kind sigma x with
| App (c, l) ->
- (match kind_of_term c with
+ (match EConstr.kind sigma c with
Ind (i,_) ->
let len = Array.length l in
let sigty = delayed_force sig_typ in
@@ -129,7 +131,7 @@ let lift_args n sign =
let mu env evdref t =
let rec aux v =
let v' = hnf env !evdref v in
- match disc_subset v' with
+ match disc_subset !evdref v' with
| Some (u, p) ->
let f, ct = aux u in
let p = hnf_nodelta env !evdref p in
@@ -140,8 +142,8 @@ let mu env evdref t =
| None -> (None, v)
in aux t
-and coerce loc env evdref (x : Term.constr) (y : Term.constr)
- : (Term.constr -> Term.constr) option
+and coerce ?loc env evdref (x : EConstr.constr) (y : EConstr.constr)
+ : (EConstr.constr -> EConstr.constr) option
=
let open Context.Rel.Declaration in
let rec coerce_unify env x y =
@@ -150,12 +152,11 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr)
evdref := the_conv_x_leq env x y !evdref;
None
with UnableToUnify _ -> coerce' env x y
- and coerce' env x y : (Term.constr -> Term.constr) option =
+ and coerce' env x y : (EConstr.constr -> EConstr.constr) option =
let subco () = subset_coerce env evdref x y in
let dest_prod c =
- let open Context.Rel.Declaration in
- match Reductionops.splay_prod_n env ( !evdref) 1 c with
- | [LocalAssum (na,t) | LocalDef (na,_,t)], c -> (na,t), c
+ match Reductionops.splay_prod_n env (!evdref) 1 c with
+ | [LocalAssum (na,t) | LocalDef (na,_,t)], c -> (na, t), c
| _ -> raise NoSubtacCoercion
in
let coerce_application typ typ' c c' l l' =
@@ -175,14 +176,14 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr)
with UnableToUnify _ -> raise NoSubtacCoercion
in
(* Disallow equalities on arities *)
- if Reduction.is_arity env eqT then raise NoSubtacCoercion;
+ if Reductionops.is_arity env !evdref eqT then raise NoSubtacCoercion;
let restargs = lift_args 1
(List.rev (Array.to_list (Array.sub l (succ i) (len - (succ i)))))
in
let args = List.rev (restargs @ mkRel 1 :: List.map (lift 1) tele) in
- let pred = mkLambda (n, eqT, applistc (lift 1 c) args) in
+ let pred = mkLambda (n, eqT, applist (lift 1 c, args)) in
let eq = papp evdref coq_eq_ind [| eqT; hdx; hdy |] in
- let evar = make_existential loc env evdref eq in
+ let evar = make_existential ?loc n env evdref eq in
let eq_app x = papp evdref coq_eq_rect
[| eqT; hdx; pred; x; hdy; evar|]
in
@@ -192,21 +193,21 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr)
let term = co x in
Typing.e_solve_evars env evdref term)
in
- if isEvar c || isEvar c' || not (Program.is_program_generalized_coercion ()) then
+ if isEvar !evdref c || isEvar !evdref c' || not (Program.is_program_generalized_coercion ()) then
(* Second-order unification needed. *)
raise NoSubtacCoercion;
aux [] typ typ' 0 (fun x -> x)
in
- match (kind_of_term x, kind_of_term y) with
+ match (EConstr.kind !evdref x, EConstr.kind !evdref y) with
| Sort s, Sort s' ->
- (match s, s' with
+ (match ESorts.kind !evdref s, ESorts.kind !evdref s' with
| Prop x, Prop y when x == y -> None
| Prop _, Type _ -> None
| Type x, Type y when Univ.Universe.equal x y -> None (* false *)
| _ -> subco ())
| Prod (name, a, b), Prod (name', a', b') ->
let name' =
- Name (Namegen.next_ident_away Namegen.default_dependent_ident (Termops.ids_of_context env))
+ Name (Namegen.next_ident_away Namegen.default_dependent_ident (Termops.vars_of_env env))
in
let env' = push_rel (LocalAssum (name', a')) env in
let c1 = coerce_unify env' (lift 1 a') (lift 1 a) in
@@ -225,7 +226,7 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr)
(mkApp (lift 1 f, [| coec1 |])))))
| App (c, l), App (c', l') ->
- (match kind_of_term c, kind_of_term c' with
+ (match EConstr.kind !evdref c, EConstr.kind !evdref c' with
Ind (i, u), Ind (i', u') -> (* Inductive types *)
let len = Array.length l in
let sigT = delayed_force sigT_typ in
@@ -242,16 +243,15 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr)
in
let c1 = coerce_unify env a a' in
let remove_head a c =
- match kind_of_term c with
+ match EConstr.kind !evdref c with
| Lambda (n, t, t') -> c, t'
| Evar (k, args) ->
let (evs, t) = Evardefine.define_evar_as_lambda env !evdref (k,args) in
evdref := evs;
- let (n, dom, rng) = destLambda t in
- let dom = whd_evar !evdref dom in
- if isEvar dom then
- let (domk, args) = destEvar dom in
- evdref := define domk a !evdref;
+ let (n, dom, rng) = destLambda !evdref t in
+ if isEvar !evdref dom then
+ let (domk, args) = destEvar !evdref dom in
+ evdref := define domk (EConstr.Unsafe.to_constr a) !evdref;
else ();
t, rng
| _ -> raise NoSubtacCoercion
@@ -302,7 +302,7 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr)
coerce_application typ typ' c c' l l')
else
subco ()
- | x, y when Constr.equal c c' ->
+ | x, y when EConstr.eq_constr !evdref c c' ->
if Int.equal (Array.length l) (Array.length l') then
let evm = !evdref in
let lam_type = Typing.unsafe_type_of env evm c in
@@ -313,20 +313,20 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr)
| _, _ -> subco ()
and subset_coerce env evdref x y =
- match disc_subset x with
+ match disc_subset !evdref x with
Some (u, p) ->
let c = coerce_unify env u y in
let f x =
app_opt env evdref c (papp evdref sig_proj1 [| u; p; x |])
in Some f
| None ->
- match disc_subset y with
+ match disc_subset !evdref y with
Some (u, p) ->
let c = coerce_unify env x u in
Some
(fun x ->
let cx = app_opt env evdref c x in
- let evar = make_existential loc env evdref (mkApp (p, [| cx |]))
+ let evar = make_existential ?loc Anonymous env evdref (mkApp (p, [| cx |]))
in
(papp evdref sig_intro [| u; p; cx; evar |]))
| None ->
@@ -340,9 +340,9 @@ let app_coercion env evdref coercion v =
let v' = Typing.e_solve_evars env evdref (f v) in
whd_betaiota !evdref v'
-let coerce_itf loc env evd v t c1 =
+let coerce_itf ?loc env evd v t c1 =
let evdref = ref evd in
- let coercion = coerce loc env evdref t c1 in
+ let coercion = coerce ?loc env evdref t c1 in
let t = Option.map (app_coercion env evdref coercion) v in
!evdref, t
@@ -370,12 +370,12 @@ let apply_coercion env sigma p hj typ_cl =
(hj,typ_cl,sigma) p
in evd, j
with NoCoercion as e -> raise e
- | e when CErrors.noncritical e -> anomaly (Pp.str "apply_coercion")
+ | e when CErrors.noncritical e -> anomaly (Pp.str "apply_coercion.")
(* Try to coerce to a funclass; raise NoCoercion if not possible *)
let inh_app_fun_core env evd j =
let t = whd_all env evd j.uj_type in
- match kind_of_term t with
+ match EConstr.kind evd t with
| Prod (_,_,_) -> (evd,j)
| Evar ev ->
let (evd',t) = Evardefine.define_evar_as_product evd ev in
@@ -405,36 +405,41 @@ let inh_app_fun resolve_tc env evd j =
try inh_app_fun_core env (saturate_evd env evd) j
with NoCoercion -> (evd, j)
-let inh_tosort_force loc env evd j =
+let type_judgment env sigma j =
+ match EConstr.kind sigma (whd_all env sigma j.uj_type) with
+ | Sort s -> {utj_val = j.uj_val; utj_type = ESorts.kind sigma s }
+ | _ -> error_not_a_type env sigma j
+
+let inh_tosort_force ?loc env evd j =
try
let t,p = lookup_path_to_sort_from env evd j.uj_type in
let evd,j1 = apply_coercion env evd p j t in
let j2 = on_judgment_type (whd_evar evd) j1 in
- (evd,type_judgment env j2)
+ (evd,type_judgment env evd j2)
with Not_found | NoCoercion ->
- error_not_a_type_loc loc env evd j
+ error_not_a_type ?loc env evd j
-let inh_coerce_to_sort loc env evd j =
+let inh_coerce_to_sort ?loc env evd j =
let typ = whd_all env evd j.uj_type in
- match kind_of_term typ with
- | Sort s -> (evd,{ utj_val = j.uj_val; utj_type = s })
- | Evar ev when not (is_defined evd (fst ev)) ->
+ match EConstr.kind evd typ with
+ | Sort s -> (evd,{ utj_val = j.uj_val; utj_type = ESorts.kind evd s })
+ | Evar ev ->
let (evd',s) = Evardefine.define_evar_as_sort env evd ev in
(evd',{ utj_val = j.uj_val; utj_type = s })
| _ ->
- inh_tosort_force loc env evd j
+ inh_tosort_force ?loc env evd j
-let inh_coerce_to_base loc env evd j =
+let inh_coerce_to_base ?loc env evd j =
if Flags.is_program_mode () then
let evdref = ref evd in
let ct, typ' = mu env evdref j.uj_type in
let res =
- { uj_val = app_coercion env evdref ct j.uj_val;
+ { uj_val = (app_coercion env evdref ct j.uj_val);
uj_type = typ' }
in !evdref, res
else (evd, j)
-let inh_coerce_to_prod loc env evd t =
+let inh_coerce_to_prod ?loc env evd t =
if Flags.is_program_mode () then
let evdref = ref evd in
let _, typ' = mu env evdref t in
@@ -442,7 +447,7 @@ let inh_coerce_to_prod loc env evd t =
else (evd, t)
let inh_coerce_to_fail env evd rigidonly v t c1 =
- if rigidonly && not (Heads.is_rigid env c1 && Heads.is_rigid env t)
+ if rigidonly && not (Heads.is_rigid env (EConstr.Unsafe.to_constr c1) && Heads.is_rigid env (EConstr.Unsafe.to_constr t))
then
raise NoCoercion
else
@@ -461,71 +466,72 @@ let inh_coerce_to_fail env evd rigidonly v t c1 =
try (the_conv_x_leq env t' c1 evd, v')
with UnableToUnify _ -> raise NoCoercion
-let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 =
+let rec inh_conv_coerce_to_fail ?loc env evd rigidonly v t c1 =
try (the_conv_x_leq env t c1 evd, v)
with UnableToUnify (best_failed_evd,e) ->
try inh_coerce_to_fail env evd rigidonly v t c1
with NoCoercion ->
match
- kind_of_term (whd_all env evd t),
- kind_of_term (whd_all env evd c1)
+ EConstr.kind evd (whd_all env evd t),
+ EConstr.kind evd (whd_all env evd c1)
with
| Prod (name,t1,t2), Prod (_,u1,u2) ->
(* Conversion did not work, we may succeed with a coercion. *)
(* We eta-expand (hence possibly modifying the original term!) *)
(* and look for a coercion c:u1->t1 s.t. fun x:u1 => v' (c x)) *)
(* has type forall (x:u1), u2 (with v' recursively obtained) *)
- (* Note: we retype the term because sort-polymorphism may have *)
- (* weaken its type *)
+ (* Note: we retype the term because template polymorphism may have *)
+ (* weakened its type *)
let name = match name with
| Anonymous -> Name Namegen.default_dependent_ident
| _ -> name in
let open Context.Rel.Declaration in
let env1 = push_rel (LocalAssum (name,u1)) env in
let (evd', v1) =
- inh_conv_coerce_to_fail loc env1 evd rigidonly
+ inh_conv_coerce_to_fail ?loc env1 evd rigidonly
(Some (mkRel 1)) (lift 1 u1) (lift 1 t1) in
let v1 = Option.get v1 in
- let v2 = Option.map (fun v -> beta_applist (lift 1 v,[v1])) v in
+ let v2 = Option.map (fun v -> beta_applist evd' (lift 1 v,[v1])) v in
let t2 = match v2 with
- | None -> subst_term v1 t2
+ | None -> subst_term evd' v1 t2
| Some v2 -> Retyping.get_type_of env1 evd' v2 in
- let (evd'',v2') = inh_conv_coerce_to_fail loc env1 evd' rigidonly v2 t2 u2 in
+ let (evd'',v2') = inh_conv_coerce_to_fail ?loc env1 evd' rigidonly v2 t2 u2 in
(evd'', Option.map (fun v2' -> mkLambda (name, u1, v2')) v2')
| _ -> raise (NoCoercionNoUnifier (best_failed_evd,e))
(* Look for cj' obtained from cj by inserting coercions, s.t. cj'.typ = t *)
-let inh_conv_coerce_to_gen resolve_tc rigidonly loc env evd cj t =
+let inh_conv_coerce_to_gen ?loc resolve_tc rigidonly env evd cj t =
let (evd', val') =
try
- inh_conv_coerce_to_fail loc env evd rigidonly (Some cj.uj_val) cj.uj_type t
+ inh_conv_coerce_to_fail ?loc env evd rigidonly (Some cj.uj_val) cj.uj_type t
with NoCoercionNoUnifier (best_failed_evd,e) ->
try
if Flags.is_program_mode () then
- coerce_itf loc env evd (Some cj.uj_val) cj.uj_type t
+ coerce_itf ?loc env evd (Some cj.uj_val) cj.uj_type t
else raise NoSubtacCoercion
with
| NoSubtacCoercion when not resolve_tc || not !use_typeclasses_for_conversion ->
- error_actual_type_loc loc env best_failed_evd cj t e
+ error_actual_type ?loc env best_failed_evd cj t e
| NoSubtacCoercion ->
let evd' = saturate_evd env evd in
try
if evd' == evd then
- error_actual_type_loc loc env best_failed_evd cj t e
+ error_actual_type ?loc env best_failed_evd cj t e
else
- inh_conv_coerce_to_fail loc env evd' rigidonly (Some cj.uj_val) cj.uj_type t
+ inh_conv_coerce_to_fail ?loc env evd' rigidonly (Some cj.uj_val) cj.uj_type t
with NoCoercionNoUnifier (_evd,_error) ->
- error_actual_type_loc loc env best_failed_evd cj t e
+ error_actual_type ?loc env best_failed_evd cj t e
in
let val' = match val' with Some v -> v | None -> assert(false) in
(evd',{ uj_val = val'; uj_type = t })
-let inh_conv_coerce_to resolve_tc = inh_conv_coerce_to_gen resolve_tc false
-let inh_conv_coerce_rigid_to resolve_tc = inh_conv_coerce_to_gen resolve_tc true
+let inh_conv_coerce_to ?loc resolve_tc = inh_conv_coerce_to_gen ?loc resolve_tc false
+
+let inh_conv_coerce_rigid_to ?loc resolve_tc = inh_conv_coerce_to_gen resolve_tc ?loc true
-let inh_conv_coerces_to loc env evd t t' =
+let inh_conv_coerces_to ?loc env evd t t' =
try
- fst (inh_conv_coerce_to_fail loc env evd true None t t')
+ fst (inh_conv_coerce_to_fail ?loc env evd true None t t')
with NoCoercion ->
evd (* Maybe not enough information to unify *)
diff --git a/pretyping/coercion.mli b/pretyping/coercion.mli
index 68f9a2e6..6cfd958b 100644
--- a/pretyping/coercion.mli
+++ b/pretyping/coercion.mli
@@ -1,15 +1,17 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Evd
open Names
-open Term
open Environ
+open EConstr
open Glob_term
(** {6 Coercions. } *)
@@ -25,17 +27,17 @@ val inh_app_fun : bool ->
(** [inh_coerce_to_sort env isevars j] coerces [j] to a type; i.e. it
inserts a coercion into [j], if needed, in such a way it gets as
type a sort; it fails if no coercion is applicable *)
-val inh_coerce_to_sort : Loc.t ->
+val inh_coerce_to_sort : ?loc:Loc.t ->
env -> evar_map -> unsafe_judgment -> evar_map * unsafe_type_judgment
(** [inh_coerce_to_base env isevars j] coerces [j] to its base type; i.e. it
inserts a coercion into [j], if needed, in such a way it gets as
type its base type (the notion depends on the coercion system) *)
-val inh_coerce_to_base : Loc.t ->
+val inh_coerce_to_base : ?loc:Loc.t ->
env -> evar_map -> unsafe_judgment -> evar_map * unsafe_judgment
(** [inh_coerce_to_prod env isevars t] coerces [t] to a product type *)
-val inh_coerce_to_prod : Loc.t ->
+val inh_coerce_to_prod : ?loc:Loc.t ->
env -> evar_map -> types -> evar_map * types
(** [inh_conv_coerce_to resolve_tc Loc.t env isevars j t] coerces [j] to an
@@ -43,20 +45,20 @@ val inh_coerce_to_prod : Loc.t ->
a way [t] and [j.uj_type] are convertible; it fails if no coercion is
applicable. resolve_tc=false disables resolving type classes (as the last
resort before failing) *)
-val inh_conv_coerce_to : bool -> Loc.t ->
+val inh_conv_coerce_to : ?loc:Loc.t -> bool ->
env -> evar_map -> unsafe_judgment -> types -> evar_map * unsafe_judgment
-val inh_conv_coerce_rigid_to : bool -> Loc.t ->
+val inh_conv_coerce_rigid_to : ?loc:Loc.t -> bool ->
env -> evar_map -> unsafe_judgment -> types -> evar_map * unsafe_judgment
(** [inh_conv_coerces_to loc env isevars t t'] checks if an object of type [t]
is coercible to an object of type [t'] adding evar constraints if needed;
it fails if no coercion exists *)
-val inh_conv_coerces_to : Loc.t ->
+val inh_conv_coerces_to : ?loc:Loc.t ->
env -> evar_map -> types -> types -> evar_map
(** [inh_pattern_coerce_to loc env isevars pat ind1 ind2] coerces the Cases
pattern [pat] typed in [ind1] into a pattern typed in [ind2];
raises [Not_found] if no coercion found *)
val inh_pattern_coerce_to :
- Loc.t -> env -> cases_pattern -> inductive -> inductive -> cases_pattern
+ ?loc:Loc.t -> env -> cases_pattern -> inductive -> inductive -> cases_pattern
diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml
index 886a9826..0413c6b6 100644
--- a/pretyping/constr_matching.ml
+++ b/pretyping/constr_matching.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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*)
@@ -12,15 +14,15 @@ open CErrors
open Util
open Names
open Globnames
-open Nameops
open Termops
-open Reductionops
open Term
+open EConstr
open Vars
open Pattern
open Patternops
open Misctypes
open Context.Rel.Declaration
+open Ltac_pretype
(*i*)
(* Given a term with second-order variables in it,
@@ -45,6 +47,7 @@ open Context.Rel.Declaration
*)
+type binding_bound_vars = Id.Set.t
type bound_ident_map = Id.t Id.Map.t
exception PatternMatchingFailure
@@ -52,67 +55,109 @@ exception PatternMatchingFailure
let warn_meta_collision =
CWarnings.create ~name:"meta-collision" ~category:"ltac"
(fun name ->
- strbrk "Collision between bound variable " ++ pr_id name ++
+ strbrk "Collision between bound variable " ++ Id.print name ++
strbrk " and a metavariable of same name.")
-let constrain n (ids, m as x) (names, terms as subst) =
+let constrain sigma n (ids, m) ((names,seen as names_seen), terms as subst) =
+ let open EConstr in
try
let (ids', m') = Id.Map.find n terms in
- if List.equal Id.equal ids ids' && eq_constr m m' then subst
+ if List.equal Id.equal ids ids' && eq_constr sigma m m' then subst
else raise PatternMatchingFailure
with Not_found ->
let () = if Id.Map.mem n names then warn_meta_collision n in
- (names, Id.Map.add n x terms)
+ (names_seen, Id.Map.add n (ids, m) terms)
-let add_binders na1 na2 binding_vars (names, terms as subst) =
+let add_binders na1 na2 binding_vars ((names,seen), terms as subst) =
match na1, na2 with
| Name id1, Name id2 when Id.Set.mem id1 binding_vars ->
if Id.Map.mem id1 names then
let () = Glob_ops.warn_variable_collision id1 in
- (names, terms)
+ subst
else
+ let id2 = Namegen.next_ident_away id2 seen in
let names = Id.Map.add id1 id2 names in
+ let seen = Id.Set.add id2 seen in
let () = if Id.Map.mem id1 terms then
warn_meta_collision id1 in
- (names, terms)
+ ((names,seen), terms)
| _ -> subst
-let rec build_lambda vars ctx m = match vars with
+let rec build_lambda sigma vars ctx m = match vars with
| [] ->
- let len = List.length ctx in
- lift (-1 * len) m
+ if Vars.closed0 sigma m then m else raise PatternMatchingFailure
| n :: vars ->
(* change [ x1 ... xn y z1 ... zm |- t ] into
[ x1 ... xn z1 ... zm |- lam y. t ] *)
- let len = List.length ctx in
- let init i =
- if i < pred n then mkRel (i + 2)
- else if Int.equal i (pred n) then mkRel 1
- else mkRel (i + 1)
- in
- let m = substl (List.init len init) m in
let pre, suf = List.chop (pred n) ctx in
- match suf with
+ let (na, t, suf) = match suf with
| [] -> assert false
- | (_, na, t) :: suf ->
- let map i = if i > n then pred i else i in
- let vars = List.map map vars in
- (** Check that the abstraction is legal *)
- let frels = free_rels t in
- let brels = List.fold_right Int.Set.add vars Int.Set.empty in
- let () = if not (Int.Set.subset frels brels) then raise PatternMatchingFailure in
- (** Create the abstraction *)
- let m = mkLambda (na, t, m) in
- build_lambda vars (pre @ suf) m
+ | (_, id, t) :: suf ->
+ (Name id, t, suf)
+ in
+ (** Check that the abstraction is legal by generating a transitive closure of
+ its dependencies. *)
+ let is_nondep t clear = match clear with
+ | [] -> true
+ | _ ->
+ let rels = free_rels sigma t in
+ let check i b = b || not (Int.Set.mem i rels) in
+ List.for_all_i check 1 clear
+ in
+ let fold (_, _, t) clear = is_nondep t clear :: clear in
+ (** Produce a list of booleans: true iff we keep the hypothesis *)
+ let clear = List.fold_right fold pre [false] in
+ let clear = List.drop_last clear in
+ (** If the conclusion depends on a variable we cleared, failure *)
+ let () = if not (is_nondep m clear) then raise PatternMatchingFailure in
+ (** Create the abstracted term *)
+ let fold (k, accu) keep =
+ if keep then
+ let k = succ k in
+ (k, Some k :: accu)
+ else (k, None :: accu)
+ in
+ let keep, shift = List.fold_left fold (0, []) clear in
+ let shift = List.rev shift in
+ let map = function
+ | None -> mkProp (** dummy term *)
+ | Some i -> mkRel (i + 1)
+ in
+ (** [x1 ... xn y z1 ... zm] -> [x1 ... xn f(z1) ... f(zm) y] *)
+ let subst =
+ List.map map shift @
+ mkRel 1 ::
+ List.mapi (fun i _ -> mkRel (i + keep + 2)) suf
+ in
+ let map i (na, id, c) =
+ let i = succ i in
+ let subst = List.skipn i subst in
+ let subst = List.map (fun c -> Vars.lift (- i) c) subst in
+ (na, id, substl subst c)
+ in
+ let pre = List.mapi map pre in
+ let pre = List.filter_with clear pre in
+ let m = substl subst m in
+ let map i =
+ if i > n then i - n + keep
+ else match List.nth shift (i - 1) with
+ | None ->
+ (** We cleared a variable that we wanted to abstract! *)
+ raise PatternMatchingFailure
+ | Some k -> k
+ in
+ let vars = List.map map vars in
+ (** Create the abstraction *)
+ let m = mkLambda (na, Vars.lift keep t, m) in
+ build_lambda sigma vars (pre @ suf) m
let rec extract_bound_aux k accu frels ctx = match ctx with
| [] -> accu
-| (na1, na2, _) :: ctx ->
+| (na, _, _) :: ctx ->
if Int.Set.mem k frels then
- begin match na1 with
+ begin match na with
| Name id ->
- let () = assert (match na2 with Anonymous -> false | Name _ -> true) in
let () = if Id.Set.mem id accu then raise PatternMatchingFailure in
extract_bound_aux (k + 1) (Id.Set.add id accu) frels ctx
| Anonymous -> raise PatternMatchingFailure
@@ -122,69 +167,78 @@ let rec extract_bound_aux k accu frels ctx = match ctx with
let extract_bound_vars frels ctx =
extract_bound_aux 1 Id.Set.empty frels ctx
-let dummy_constr = mkProp
+let dummy_constr = EConstr.mkProp
let make_renaming ids = function
-| (Name id, Name _, _) ->
+| (Name id, _, _) ->
begin
- try mkRel (List.index Id.equal id ids)
+ try EConstr.mkRel (List.index Id.equal id ids)
with Not_found -> dummy_constr
end
| _ -> dummy_constr
-let merge_binding allow_bound_rels ctx n cT subst =
+let push_binder na1 na2 t ctx =
+ let id2 = match na2 with
+ | Name id2 -> id2
+ | Anonymous ->
+ let avoid = Id.Set.of_list (List.map pi2 ctx) in
+ Namegen.next_ident_away Namegen.default_non_dependent_ident avoid in
+ (na1, id2, t) :: ctx
+
+let to_fix (idx, (nas, cs, ts)) =
+ let inj = EConstr.of_constr in
+ (idx, (nas, Array.map inj cs, Array.map inj ts))
+
+let merge_binding sigma allow_bound_rels ctx n cT subst =
let c = match ctx with
| [] -> (* Optimization *)
([], cT)
| _ ->
- let frels = free_rels cT in
+ let frels = free_rels sigma cT in
if allow_bound_rels then
let vars = extract_bound_vars frels ctx in
let ordered_vars = Id.Set.elements vars in
let rename binding = make_renaming ordered_vars binding in
let renaming = List.map rename ctx in
- (ordered_vars, substl renaming cT)
+ (ordered_vars, Vars.substl renaming cT)
else
let depth = List.length ctx in
let min_elt = try Int.Set.min_elt frels with Not_found -> succ depth in
if depth < min_elt then
- ([], lift (- depth) cT)
+ ([], Vars.lift (- depth) cT)
else raise PatternMatchingFailure
in
- constrain n c subst
+ constrain sigma n c subst
-let matches_core env sigma convert allow_partial_app allow_bound_rels
+let matches_core env sigma allow_bound_rels
(binding_vars,pat) c =
+ let open EConstr in
let convref ref c =
- match ref, kind_of_term c with
- | VarRef id, Var id' -> Names.id_eq id id'
- | ConstRef c, Const (c',_) -> Names.eq_constant c c'
+ match ref, EConstr.kind sigma c with
+ | VarRef id, Var id' -> Names.Id.equal id id'
+ | ConstRef c, Const (c',_) -> Constant.equal c c'
| IndRef i, Ind (i', _) -> Names.eq_ind i i'
| ConstructRef c, Construct (c',u) -> Names.eq_constructor c c'
- | _, _ ->
- (if convert then
- let sigma,c' = Evd.fresh_global env sigma ref in
- is_conv env sigma c' c
- else false)
+ | _, _ -> false
in
let rec sorec ctx env subst p t =
- let cT = strip_outer_cast t in
- match p,kind_of_term cT with
+ let cT = strip_outer_cast sigma t in
+ match p, EConstr.kind sigma cT with
| PSoApp (n,args),m ->
let fold (ans, seen) = function
| PRel n ->
- let () = if Int.Set.mem n seen then error "Non linear second-order pattern" in
+ let () = if Int.Set.mem n seen then user_err (str "Non linear second-order pattern") in
(n :: ans, Int.Set.add n seen)
- | _ -> error "Only bound indices allowed in second order pattern matching."
+ | _ -> user_err (str "Only bound indices allowed in second order pattern matching.")
in
let relargs, relset = List.fold_left fold ([], Int.Set.empty) args in
- let frels = free_rels cT in
+ let frels = free_rels sigma cT in
if Int.Set.subset frels relset then
- constrain n ([], build_lambda relargs ctx cT) subst
+ constrain sigma n ([], build_lambda sigma relargs ctx cT) subst
else
raise PatternMatchingFailure
- | PMeta (Some n), m -> merge_binding allow_bound_rels ctx n cT subst
+ | PMeta (Some n), m -> merge_binding sigma allow_bound_rels ctx n cT subst
| PMeta None, m -> subst
@@ -196,18 +250,21 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels
| PRel n1, Rel n2 when Int.equal n1 n2 -> subst
- | PSort GProp, Sort (Prop Null) -> subst
+ | PSort ps, Sort s ->
- | PSort GSet, Sort (Prop Pos) -> subst
-
- | PSort (GType _), Sort (Type _) -> subst
+ begin match ps, ESorts.kind sigma s with
+ | GProp, Prop Null -> subst
+ | GSet, Prop Pos -> subst
+ | GType _, Type _ -> subst
+ | _ -> raise PatternMatchingFailure
+ end
| PApp (p, [||]), _ -> sorec ctx env subst p t
| PApp (PApp (h, a1), a2), _ ->
sorec ctx env subst (PApp(h,Array.append a1 a2)) t
- | PApp (PMeta meta,args1), App (c2,args2) when allow_partial_app ->
+ | PApp (PMeta meta,args1), App (c2,args2) ->
(let diff = Array.length args2 - Array.length args1 in
if diff >= 0 then
let args21, args22 = Array.chop diff args2 in
@@ -215,10 +272,10 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels
let subst =
match meta with
| None -> subst
- | Some n -> merge_binding allow_bound_rels ctx n c subst in
+ | Some n -> merge_binding sigma allow_bound_rels ctx n c subst in
Array.fold_left2 (sorec ctx env) subst args1 args22
else (* Might be a projection on the right *)
- match kind_of_term c2 with
+ match EConstr.kind sigma c2 with
| Proj (pr, c) when not (Projection.unfolded pr) ->
(try let term = Retyping.expand_projection env sigma pr c (Array.to_list args2) in
sorec ctx env subst p term
@@ -226,8 +283,8 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels
| _ -> raise PatternMatchingFailure)
| PApp (c1,arg1), App (c2,arg2) ->
- (match c1, kind_of_term c2 with
- | PRef (ConstRef r), Proj (pr,c) when not (eq_constant r (Projection.constant pr))
+ (match c1, EConstr.kind sigma c2 with
+ | PRef (ConstRef r), Proj (pr,c) when not (Constant.equal r (Projection.constant pr))
|| Projection.unfolded pr ->
raise PatternMatchingFailure
| PProj (pr1,c1), Proj (pr,c) ->
@@ -244,7 +301,7 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels
with Invalid_argument _ -> raise PatternMatchingFailure)
| PApp (PRef (ConstRef c1), _), Proj (pr, c2)
- when Projection.unfolded pr || not (eq_constant c1 (Projection.constant pr)) ->
+ when Projection.unfolded pr || not (Constant.equal c1 (Projection.constant pr)) ->
raise PatternMatchingFailure
| PApp (c, args), Proj (pr, c2) ->
@@ -256,29 +313,33 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels
sorec ctx env subst c1 c2
| PProd (na1,c1,d1), Prod(na2,c2,d2) ->
- sorec ((na1,na2,c2)::ctx) (Environ.push_rel (LocalAssum (na2,c2)) env)
+ sorec (push_binder na1 na2 c2 ctx) (EConstr.push_rel (LocalAssum (na2,c2)) env)
(add_binders na1 na2 binding_vars (sorec ctx env subst c1 c2)) d1 d2
| PLambda (na1,c1,d1), Lambda(na2,c2,d2) ->
- sorec ((na1,na2,c2)::ctx) (Environ.push_rel (LocalAssum (na2,c2)) env)
+ sorec (push_binder na1 na2 c2 ctx) (EConstr.push_rel (LocalAssum (na2,c2)) env)
(add_binders na1 na2 binding_vars (sorec ctx env subst c1 c2)) d1 d2
- | PLetIn (na1,c1,d1), LetIn(na2,c2,t2,d2) ->
- sorec ((na1,na2,t2)::ctx) (Environ.push_rel (LocalDef (na2,c2,t2)) env)
+ | PLetIn (na1,c1,Some t1,d1), LetIn(na2,c2,t2,d2) ->
+ sorec (push_binder na1 na2 t2 ctx) (EConstr.push_rel (LocalDef (na2,c2,t2)) env)
+ (add_binders na1 na2 binding_vars (sorec ctx env (sorec ctx env subst c1 c2) t1 t2)) d1 d2
+
+ | PLetIn (na1,c1,None,d1), LetIn(na2,c2,t2,d2) ->
+ sorec (push_binder na1 na2 t2 ctx) (EConstr.push_rel (LocalDef (na2,c2,t2)) env)
(add_binders na1 na2 binding_vars (sorec ctx env subst c1 c2)) d1 d2
| PIf (a1,b1,b1'), Case (ci,_,a2,[|b2;b2'|]) ->
- let ctx_b2,b2 = decompose_lam_n_decls ci.ci_cstr_ndecls.(0) b2 in
- let ctx_b2',b2' = decompose_lam_n_decls ci.ci_cstr_ndecls.(1) b2' in
+ let ctx_b2,b2 = decompose_lam_n_decls sigma ci.ci_cstr_ndecls.(0) b2 in
+ let ctx_b2',b2' = decompose_lam_n_decls sigma ci.ci_cstr_ndecls.(1) b2' in
let n = Context.Rel.length ctx_b2 in
let n' = Context.Rel.length ctx_b2' in
- if noccur_between 1 n b2 && noccur_between 1 n' b2' then
- let f l (LocalAssum (na,t) | LocalDef (na,_,t)) = (Anonymous,na,t)::l in
+ if Vars.noccur_between sigma 1 n b2 && Vars.noccur_between sigma 1 n' b2' then
+ let f l (LocalAssum (na,t) | LocalDef (na,_,t)) = push_binder Anonymous na t l in
let ctx_br = List.fold_left f ctx ctx_b2 in
let ctx_br' = List.fold_left f ctx ctx_b2' in
let b1 = lift_pattern n b1 and b1' = lift_pattern n' b1' in
- sorec ctx_br' (Environ.push_rel_context ctx_b2' env)
- (sorec ctx_br (Environ.push_rel_context ctx_b2 env)
+ sorec ctx_br' (push_rel_context ctx_b2' env)
+ (sorec ctx_br (push_rel_context ctx_b2 env)
(sorec ctx env subst a1 a2) b1 b2) b1' b2'
else
raise PatternMatchingFailure
@@ -305,21 +366,27 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels
let chk_head = sorec ctx env (sorec ctx env subst a1 a2) p1 p2 in
List.fold_left chk_branch chk_head br1
- | PFix c1, Fix _ when eq_constr (mkFix c1) cT -> subst
- | PCoFix c1, CoFix _ when eq_constr (mkCoFix c1) cT -> subst
- | _ -> raise PatternMatchingFailure
+ | PFix c1, Fix _ when eq_constr sigma (mkFix (to_fix c1)) cT -> subst
+ | PCoFix c1, CoFix _ when eq_constr sigma (mkCoFix (to_fix c1)) cT -> subst
+ | PEvar (c1,args1), Evar (c2,args2) when Evar.equal c1 c2 ->
+ Array.fold_left2 (sorec ctx env) subst args1 args2
+ | (PRef _ | PVar _ | PRel _ | PApp _ | PProj _ | PLambda _
+ | PProd _ | PLetIn _ | PSort _ | PIf _ | PCase _
+ | PFix _ | PCoFix _| PEvar _), _ -> raise PatternMatchingFailure
in
- sorec [] env (Id.Map.empty, Id.Map.empty) pat c
+ sorec [] env ((Id.Map.empty,Id.Set.empty), Id.Map.empty) pat c
-let matches_core_closed env sigma convert allow_partial_app pat c =
- let names, subst = matches_core env sigma convert allow_partial_app false pat c in
- (names, Id.Map.map snd subst)
+let matches_core_closed env sigma pat c =
+ let names, subst = matches_core env sigma false pat c in
+ (fst names, Id.Map.map snd subst)
-let extended_matches env sigma = matches_core env sigma false true true
+let extended_matches env sigma pat c =
+ let (names,_), subst = matches_core env sigma true pat c in
+ names, subst
let matches env sigma pat c =
- snd (matches_core_closed env sigma false true (Id.Set.empty,pat) c)
+ snd (matches_core_closed env sigma (Id.Set.empty,pat) c)
let special_meta = (-1)
@@ -332,8 +399,9 @@ let mkresult s c n = IStream.Cons ( { m_sub=s; m_ctx=c; } , (IStream.thunk n) )
let isPMeta = function PMeta _ -> true | _ -> false
let matches_head env sigma pat c =
+ let open EConstr in
let head =
- match pat, kind_of_term c with
+ match pat, EConstr.kind sigma c with
| PApp (c1,arg1), App (c2,arg2) ->
if isPMeta c1 then c else
let n1 = Array.length arg1 in
@@ -343,10 +411,10 @@ let matches_head env sigma pat c =
matches env sigma pat head
(* Tells if it is an authorized occurrence and if the instance is closed *)
-let authorized_occ env sigma partial_app closed pat c mk_ctx =
+let authorized_occ env sigma closed pat c mk_ctx =
try
- let subst = matches_core_closed env sigma false partial_app pat c in
- if closed && Id.Map.exists (fun _ c -> not (closed0 c)) (snd subst)
+ let subst = matches_core_closed env sigma pat c in
+ if closed && Id.Map.exists (fun _ c -> not (closed0 sigma c)) (snd subst)
then (fun next -> next ())
else (fun next -> mkresult subst (mk_ctx (mkMeta special_meta)) next)
with PatternMatchingFailure -> (fun next -> next ())
@@ -354,10 +422,11 @@ let authorized_occ env sigma partial_app closed pat c mk_ctx =
let subargs env v = Array.map_to_list (fun c -> (env, c)) v
(* Tries to match a subterm of [c] with [pat] *)
-let sub_match ?(partial_app=false) ?(closed=true) env sigma pat c =
+let sub_match ?(closed=true) env sigma pat c =
+ let open EConstr in
let rec aux env c mk_ctx next =
- let here = authorized_occ env sigma partial_app closed pat c mk_ctx in
- let next () = match kind_of_term c with
+ let here = authorized_occ env sigma closed pat c mk_ctx in
+ let next () = match EConstr.kind sigma c with
| Cast (c1,k,c2) ->
let next_mk_ctx = function
| [c1] -> mk_ctx (mkCast (c1, k, c2))
@@ -369,51 +438,29 @@ let sub_match ?(partial_app=false) ?(closed=true) env sigma pat c =
| [c1; c2] -> mk_ctx (mkLambda (x, c1, c2))
| _ -> assert false
in
- let env' = Environ.push_rel (LocalAssum (x,c1)) env in
+ let env' = EConstr.push_rel (LocalAssum (x,c1)) env in
try_aux [(env, c1); (env', c2)] next_mk_ctx next
| Prod (x,c1,c2) ->
let next_mk_ctx = function
| [c1; c2] -> mk_ctx (mkProd (x, c1, c2))
| _ -> assert false
in
- let env' = Environ.push_rel (LocalAssum (x,c1)) env in
+ let env' = EConstr.push_rel (LocalAssum (x,c1)) env in
try_aux [(env, c1); (env', c2)] next_mk_ctx next
| LetIn (x,c1,t,c2) ->
let next_mk_ctx = function
| [c1; c2] -> mk_ctx (mkLetIn (x, c1, t, c2))
| _ -> assert false
in
- let env' = Environ.push_rel (LocalDef (x,c1,t)) env in
+ let env' = EConstr.push_rel (LocalDef (x,c1,t)) env in
try_aux [(env, c1); (env', c2)] next_mk_ctx next
| App (c1,lc) ->
- let topdown = true in
- if partial_app then
- if topdown then
- let lc1 = Array.sub lc 0 (Array.length lc - 1) in
- let app = mkApp (c1,lc1) in
- let mk_ctx = function
- | [app';c] -> mk_ctx (mkApp (app',[|c|]))
- | _ -> assert false in
- try_aux [(env, app); (env, Array.last lc)] mk_ctx next
- else
- let rec aux2 app args next =
- match args with
- | [] ->
- let mk_ctx le =
- mk_ctx (mkApp (List.hd le, Array.of_list (List.tl le))) in
- let sub = (env, c1) :: subargs env lc in
- try_aux sub mk_ctx next
- | arg :: args ->
- let app = mkApp (app,[|arg|]) in
- let next () = aux2 app args next in
- let mk_ctx ce = mk_ctx (mkApp (ce, Array.of_list args)) in
- aux env app mk_ctx next in
- aux2 c1 (Array.to_list lc) next
- else
- let mk_ctx le =
- mk_ctx (mkApp (List.hd le, Array.of_list (List.tl le))) in
- let sub = (env, c1) :: subargs env lc in
- try_aux sub mk_ctx next
+ let lc1 = Array.sub lc 0 (Array.length lc - 1) in
+ let app = mkApp (c1,lc1) in
+ let mk_ctx = function
+ | [app';c] -> mk_ctx (mkApp (app',[|c|]))
+ | _ -> assert false in
+ try_aux [(env, app); (env, Array.last lc)] mk_ctx next
| Case (ci,hd,c1,lc) ->
let next_mk_ctx = function
| c1 :: hd :: lc -> mk_ctx (mkCase (ci,hd,c1,Array.of_list lc))
@@ -421,29 +468,28 @@ let sub_match ?(partial_app=false) ?(closed=true) env sigma pat c =
in
let sub = (env, c1) :: (env, hd) :: subargs env lc in
try_aux sub next_mk_ctx next
- | Fix (indx,(names,types,bodies)) ->
+ | Fix (indx,(names,types,bodies as recdefs)) ->
let nb_fix = Array.length types in
let next_mk_ctx le =
let (ntypes,nbodies) = CList.chop nb_fix le in
mk_ctx (mkFix (indx,(names, Array.of_list ntypes, Array.of_list nbodies))) in
- let sub = subargs env types @ subargs env bodies in
+ let env' = push_rec_types recdefs env in
+ let sub = subargs env types @ subargs env' bodies in
try_aux sub next_mk_ctx next
- | CoFix (i,(names,types,bodies)) ->
+ | CoFix (i,(names,types,bodies as recdefs)) ->
let nb_fix = Array.length types in
let next_mk_ctx le =
let (ntypes,nbodies) = CList.chop nb_fix le in
mk_ctx (mkCoFix (i,(names, Array.of_list ntypes, Array.of_list nbodies))) in
- let sub = subargs env types @ subargs env bodies in
+ let env' = push_rec_types recdefs env in
+ let sub = subargs env types @ subargs env' bodies in
try_aux sub next_mk_ctx next
| Proj (p,c') ->
- let next_mk_ctx le = mk_ctx (mkProj (p,List.hd le)) in
- if partial_app then
- try
- let term = Retyping.expand_projection env sigma p c' [] in
- aux env term mk_ctx next
- with Retyping.RetypeError _ -> next ()
- else
- try_aux [env, c'] next_mk_ctx next
+ begin try
+ let term = Retyping.expand_projection env sigma p c' [] in
+ aux env term mk_ctx next
+ with Retyping.RetypeError _ -> next ()
+ end
| Construct _| Ind _|Evar _|Const _ | Rel _|Meta _|Var _|Sort _ ->
next ()
in
@@ -464,13 +510,7 @@ let sub_match ?(partial_app=false) ?(closed=true) env sigma pat c =
let result () = aux env c (fun x -> x) lempty in
IStream.thunk result
-let match_subterm env sigma pat c = sub_match env sigma (Id.Set.empty,pat) c
-
-let match_appsubterm env sigma pat c =
- sub_match ~partial_app:true env sigma (Id.Set.empty,pat) c
-
-let match_subterm_gen env sigma app pat c =
- sub_match ~partial_app:app env sigma pat c
+let match_subterm env sigma pat c = sub_match env sigma pat c
let is_matching env sigma pat c =
try let _ = matches env sigma pat c in true
@@ -482,12 +522,5 @@ let is_matching_head env sigma pat c =
let is_matching_appsubterm ?(closed=true) env sigma pat c =
let pat = (Id.Set.empty,pat) in
- let results = sub_match ~partial_app:true ~closed env sigma pat c in
+ let results = sub_match ~closed env sigma pat c in
not (IStream.is_empty results)
-
-let matches_conv env sigma p c =
- snd (matches_core_closed env sigma true false (Id.Set.empty,p) c)
-
-let is_matching_conv env sigma pat n =
- try let _ = matches_conv env sigma pat n in true
- with PatternMatchingFailure -> false
diff --git a/pretyping/constr_matching.mli b/pretyping/constr_matching.mli
index 8d8166f2..3c2c7391 100644
--- a/pretyping/constr_matching.mli
+++ b/pretyping/constr_matching.mli
@@ -1,17 +1,23 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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) *)
(************************************************************************)
(** This module implements pattern-matching on terms *)
open Names
-open Term
+open Constr
+open EConstr
open Environ
open Pattern
+open Ltac_pretype
+
+type binding_bound_vars = Id.Set.t
(** [PatternMatchingFailure] is the exception raised when pattern
matching fails *)
@@ -41,7 +47,7 @@ val matches_head : env -> Evd.evar_map -> constr_pattern -> constr -> patvar_map
variables or metavariables have the same name, the metavariable,
or else the rightmost bound variable, takes precedence *)
val extended_matches :
- env -> Evd.evar_map -> Tacexpr.binding_bound_vars * constr_pattern ->
+ env -> Evd.evar_map -> binding_bound_vars * constr_pattern ->
constr -> bound_ident_map * extended_patvar_map
(** [is_matching pat c] just tells if [c] matches against [pat] *)
@@ -51,38 +57,19 @@ val is_matching : env -> Evd.evar_map -> constr_pattern -> constr -> bool
prefix of it matches against [pat] *)
val is_matching_head : env -> Evd.evar_map -> constr_pattern -> constr -> bool
-(** [matches_conv env sigma] matches up to conversion in environment
- [(env,sigma)] when constants in pattern are concerned; it raises
- [PatternMatchingFailure] if not matchable; bindings are given in
- increasing order based on the numbers given in the pattern *)
-val matches_conv : env -> Evd.evar_map -> constr_pattern -> constr -> patvar_map
-
(** The type of subterm matching results: a substitution + a context
(whose hole is denoted here with [special_meta]) *)
type matching_result =
{ m_sub : bound_ident_map * patvar_map;
- m_ctx : constr }
-
-(** [match_subterm n pat c] returns the substitution and the context
- corresponding to each **closed** subterm of [c] matching [pat]. *)
-val match_subterm : env -> Evd.evar_map -> constr_pattern -> constr -> matching_result IStream.t
+ m_ctx : EConstr.t }
-(** [match_appsubterm pat c] returns the substitution and the context
+(** [match_subterm pat c] returns the substitution and the context
corresponding to each **closed** subterm of [c] matching [pat],
considering application contexts as well. *)
-val match_appsubterm : env -> Evd.evar_map -> constr_pattern -> constr -> matching_result IStream.t
-
-(** [match_subterm_gen] calls either [match_subterm] or [match_appsubterm] *)
-val match_subterm_gen : env -> Evd.evar_map ->
- bool (** true = with app context *) ->
- Tacexpr.binding_bound_vars * constr_pattern -> constr ->
+val match_subterm : env -> Evd.evar_map ->
+ binding_bound_vars * constr_pattern -> constr ->
matching_result IStream.t
(** [is_matching_appsubterm pat c] tells if a subterm of [c] matches
against [pat] taking partial subterms into consideration *)
val is_matching_appsubterm : ?closed:bool -> env -> Evd.evar_map -> constr_pattern -> constr -> bool
-
-(** [is_matching_conv env sigma pat c] tells if [c] matches against [pat]
- up to conversion for constants in patterns *)
-val is_matching_conv :
- env -> Evd.evar_map -> constr_pattern -> constr -> bool
diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml
index 85125a50..9ba5949a 100644
--- a/pretyping/detyping.ml
+++ b/pretyping/detyping.ml
@@ -1,19 +1,23 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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) *)
(************************************************************************)
+module CVars = Vars
+
open Pp
open CErrors
open Util
open Names
open Term
+open EConstr
open Vars
open Inductiveops
-open Environ
open Glob_term
open Glob_ops
open Termops
@@ -25,8 +29,11 @@ open Mod_subst
open Misctypes
open Decl_kinds
open Context.Named.Declaration
+open Ltac_pretype
-let dl = Loc.ghost
+type _ delay =
+| Now : 'a delay
+| Later : [ `thunk ] delay
(** Should we keep details of universes during detyping ? *)
let print_universes = Flags.univ_print
@@ -64,24 +71,24 @@ let has_two_constructors lc =
let isomorphic_to_tuple lc = Int.equal (Array.length lc) 1
-let encode_bool r =
+let encode_bool ({CAst.loc} as r) =
let (x,lc) = encode_inductive r in
if not (has_two_constructors lc) then
- user_err_loc (loc_of_reference r,"encode_if",
- str "This type has not exactly two constructors.");
+ user_err ?loc ~hdr:"encode_if"
+ (str "This type has not exactly two constructors.");
x
-let encode_tuple r =
+let encode_tuple ({CAst.loc} as r) =
let (x,lc) = encode_inductive r in
if not (isomorphic_to_tuple lc) then
- user_err_loc (loc_of_reference r,"encode_tuple",
- str "This type cannot be seen as a tuple type.");
+ user_err ?loc ~hdr:"encode_tuple"
+ (str "This type cannot be seen as a tuple type.");
x
module PrintingInductiveMake =
functor (Test : sig
val encode : reference -> inductive
- val member_message : std_ppcmds -> bool -> std_ppcmds
+ val member_message : Pp.t -> bool -> Pp.t
val field : string
val title : string
end) ->
@@ -133,8 +140,7 @@ let wildcard_value = ref true
let force_wildcard () = !wildcard_value
let _ = declare_bool_option
- { optsync = true;
- optdepr = false;
+ { optdepr = false;
optname = "forced wildcard";
optkey = ["Printing";"Wildcard"];
optread = force_wildcard;
@@ -144,8 +150,7 @@ let synth_type_value = ref true
let synthetize_type () = !synth_type_value
let _ = declare_bool_option
- { optsync = true;
- optdepr = false;
+ { optdepr = false;
optname = "pattern matching return type synthesizability";
optkey = ["Printing";"Synth"];
optread = synthetize_type;
@@ -155,30 +160,27 @@ let reverse_matching_value = ref true
let reverse_matching () = !reverse_matching_value
let _ = declare_bool_option
- { optsync = true;
- optdepr = false;
+ { optdepr = false;
optname = "pattern-matching reversibility";
optkey = ["Printing";"Matching"];
optread = reverse_matching;
optwrite = (:=) reverse_matching_value }
-let print_primproj_params_value = ref true
+let print_primproj_params_value = ref false
let print_primproj_params () = !print_primproj_params_value
let _ = declare_bool_option
- { optsync = true;
- optdepr = false;
+ { optdepr = false;
optname = "printing of primitive projection parameters";
optkey = ["Printing";"Primitive";"Projection";"Parameters"];
optread = print_primproj_params;
optwrite = (:=) print_primproj_params_value }
-let print_primproj_compatibility_value = ref true
+let print_primproj_compatibility_value = ref false
let print_primproj_compatibility () = !print_primproj_compatibility_value
let _ = declare_bool_option
- { optsync = true;
- optdepr = false;
+ { optdepr = false;
optname = "backwards-compatible printing of primitive projections";
optkey = ["Printing";"Primitive";"Projection";"Compatibility"];
optread = print_primproj_compatibility;
@@ -188,7 +190,7 @@ let _ = declare_bool_option
(* Auxiliary function for MutCase printing *)
(* [computable] tries to tell if the predicate typing the result is inferable*)
-let computable p k =
+let computable sigma p k =
(* We first remove as many lambda as the arity, then we look
if it remains a lambda for a dependent elimination. This function
works for normal eta-expanded term. For non eta-expanded or
@@ -205,29 +207,29 @@ let computable p k =
sinon on perd la réciprocité de la synthèse (qui, lui,
engendrera un prédicat non dépendant) *)
- let sign,ccl = decompose_lam_assum p in
+ let sign,ccl = decompose_lam_assum sigma p in
Int.equal (Context.Rel.length sign) (k + 1)
&&
- noccur_between 1 (k+1) ccl
+ noccur_between sigma 1 (k+1) ccl
-let lookup_name_as_displayed env t s =
- let rec lookup avoid n c = match kind_of_term c with
+let lookup_name_as_displayed env sigma t s =
+ let rec lookup avoid n c = match EConstr.kind sigma c with
| Prod (name,_,c') ->
- (match compute_displayed_name_in RenamingForGoal avoid name c' with
+ (match compute_displayed_name_in sigma RenamingForGoal avoid name c' with
| (Name id,avoid') -> if Id.equal id s then Some n else lookup avoid' (n+1) c'
| (Anonymous,avoid') -> lookup avoid' (n+1) (pop c'))
| LetIn (name,_,_,c') ->
- (match compute_displayed_name_in RenamingForGoal avoid name c' with
+ (match compute_displayed_name_in sigma RenamingForGoal avoid name c' with
| (Name id,avoid') -> if Id.equal id s then Some n else lookup avoid' (n+1) c'
| (Anonymous,avoid') -> lookup avoid' (n+1) (pop c'))
| Cast (c,_,_) -> lookup avoid n c
| _ -> None
- in lookup (ids_of_named_context (named_context env)) 1 t
+ in lookup (Environ.ids_of_named_context_val (Environ.named_context_val env)) 1 t
-let lookup_index_as_renamed env t n =
- let rec lookup n d c = match kind_of_term c with
+let lookup_index_as_renamed env sigma t n =
+ let rec lookup n d c = match EConstr.kind sigma c with
| Prod (name,_,c') ->
- (match compute_displayed_name_in RenamingForGoal [] name c' with
+ (match compute_displayed_name_in sigma RenamingForGoal Id.Set.empty name c' with
(Name _,_) -> lookup n (d+1) c'
| (Anonymous,_) ->
if Int.equal n 0 then
@@ -237,7 +239,7 @@ let lookup_index_as_renamed env t n =
else
lookup (n-1) (d+1) c')
| LetIn (name,_,_,c') ->
- (match compute_displayed_name_in RenamingForGoal [] name c' with
+ (match compute_displayed_name_in sigma RenamingForGoal Id.Set.empty name c' with
| (Name _,_) -> lookup n (d+1) c'
| (Anonymous,_) ->
if Int.equal n 0 then
@@ -252,22 +254,109 @@ let lookup_index_as_renamed env t n =
in lookup n 1 t
(**********************************************************************)
+(* Factorization of match patterns *)
+
+let print_factorize_match_patterns = ref true
+
+let _ =
+ let open Goptions in
+ declare_bool_option
+ { optdepr = false;
+ optname = "factorization of \"match\" patterns in printing";
+ optkey = ["Printing";"Factorizable";"Match";"Patterns"];
+ optread = (fun () -> !print_factorize_match_patterns);
+ optwrite = (fun b -> print_factorize_match_patterns := b) }
+
+let print_allow_match_default_clause = ref true
+
+let _ =
+ let open Goptions in
+ declare_bool_option
+ { optdepr = false;
+ optname = "possible use of \"match\" default pattern in printing";
+ optkey = ["Printing";"Allow";"Match";"Default";"Clause"];
+ optread = (fun () -> !print_allow_match_default_clause);
+ optwrite = (fun b -> print_allow_match_default_clause := b) }
+
+let rec join_eqns (ids,rhs as x) patll = function
+ | ({CAst.loc; v=(ids',patl',rhs')} as eqn')::rest ->
+ if not !Flags.raw_print && !print_factorize_match_patterns &&
+ List.eq_set Id.equal ids ids' && glob_constr_eq rhs rhs'
+ then
+ join_eqns x (patl'::patll) rest
+ else
+ let eqn,rest = join_eqns x patll rest in
+ eqn, eqn'::rest
+ | [] ->
+ patll, []
+
+let number_of_patterns {CAst.v=(_ids,patll,_rhs)} = List.length patll
+
+let is_default_candidate {CAst.v=(ids,_patll,_rhs)} = ids = []
+
+let rec move_more_factorized_default_candidate_to_end eqn n = function
+ | eqn' :: eqns ->
+ let set,get = set_temporary_memory () in
+ if is_default_candidate eqn' && set (number_of_patterns eqn') >= n then
+ let isbest, dft, eqns = move_more_factorized_default_candidate_to_end eqn' (get ()) eqns in
+ if isbest then false, dft, eqns else false, dft, eqn' :: eqns
+ else
+ let isbest, dft, eqns = move_more_factorized_default_candidate_to_end eqn n eqns in
+ isbest, dft, eqn' :: eqns
+ | [] -> true, Some eqn, []
+
+let rec select_default_clause = function
+ | eqn :: eqns ->
+ let set,get = set_temporary_memory () in
+ if is_default_candidate eqn && set (number_of_patterns eqn) > 1 then
+ let isbest, dft, eqns = move_more_factorized_default_candidate_to_end eqn (get ()) eqns in
+ if isbest then dft, eqns else dft, eqn :: eqns
+ else
+ let dft, eqns = select_default_clause eqns in dft, eqn :: eqns
+ | [] -> None, []
+
+let factorize_eqns eqns =
+ let open CAst in
+ let rec aux found = function
+ | {loc;v=(ids,patl,rhs)}::rest ->
+ let patll,rest = join_eqns (ids,rhs) [patl] rest in
+ aux (CAst.make ?loc (ids,patll,rhs)::found) rest
+ | [] ->
+ found in
+ let eqns = aux [] (List.rev eqns) in
+ let mk_anon patl = List.map (fun _ -> DAst.make @@ PatVar Anonymous) patl in
+ let open CAst in
+ if not !Flags.raw_print && !print_allow_match_default_clause && eqns <> [] then
+ match select_default_clause eqns with
+ (* At least two clauses and the last one is disjunctive with no variables *)
+ | Some {loc=gloc;v=([],patl::_::_,rhs)}, (_::_ as eqns) ->
+ eqns@[CAst.make ?loc:gloc ([],[mk_anon patl],rhs)]
+ (* Only one clause which is disjunctive with no variables: we keep at least one constructor *)
+ (* so that it is not interpreted as a dummy "match" *)
+ | Some {loc=gloc;v=([],patl::patl'::_,rhs)}, [] ->
+ [CAst.make ?loc:gloc ([],[patl;mk_anon patl'],rhs)]
+ | Some {v=((_::_,_,_ | _,([]|[_]),_))}, _ -> assert false
+ | None, eqns -> eqns
+ else
+ eqns
+
+(**********************************************************************)
(* Fragile algorithm to reverse pattern-matching compilation *)
-let update_name na ((_,(e,_)),c) =
+let update_name sigma na ((_,(e,_)),c) =
match na with
- | Name _ when force_wildcard () && noccurn (List.index Name.equal na e) c ->
+ | Name _ when force_wildcard () && noccurn sigma (List.index Name.equal na e) c ->
Anonymous
| _ ->
na
-let rec decomp_branch tags nal b (avoid,env as e) c =
+let rec decomp_branch tags nal b (avoid,env as e) sigma c =
let flag = if b then RenamingForGoal else RenamingForCasesPattern (fst env,c) in
match tags with
| [] -> (List.rev nal,(e,c))
| b::tags ->
let na,c,f,body,t =
- match kind_of_term (strip_outer_cast c), b with
+ match EConstr.kind sigma (strip_outer_cast sigma c), b with
| Lambda (na,t,c),false -> na,c,compute_displayed_let_name_in,None,Some t
| LetIn (na,b,t,c),true ->
na,c,compute_displayed_name_in,Some b,Some t
@@ -277,60 +366,60 @@ let rec decomp_branch tags nal b (avoid,env as e) c =
| _, true ->
Anonymous,lift 1 c,compute_displayed_name_in,None,None
in
- let na',avoid' = f flag avoid na c in
+ let na',avoid' = f sigma flag avoid na c in
decomp_branch tags (na'::nal) b
- (avoid', add_name_opt na' body t env) c
+ (avoid', add_name_opt na' body t env) sigma c
-let rec build_tree na isgoal e ci cl =
- let mkpat n rhs pl = PatCstr(dl,(ci.ci_ind,n+1),pl,update_name na rhs) in
+let rec build_tree na isgoal e sigma ci cl =
+ let mkpat n rhs pl = DAst.make @@ PatCstr((ci.ci_ind,n+1),pl,update_name sigma na rhs) in
let cnl = ci.ci_pp_info.cstr_tags in
- let cna = ci.ci_cstr_nargs in
List.flatten
(List.init (Array.length cl)
- (fun i -> contract_branch isgoal e (cnl.(i),cna.(i),mkpat i,cl.(i))))
+ (fun i -> contract_branch isgoal e sigma (cnl.(i),mkpat i,cl.(i))))
-and align_tree nal isgoal (e,c as rhs) = match nal with
- | [] -> [[],rhs]
+and align_tree nal isgoal (e,c as rhs) sigma = match nal with
+ | [] -> [Id.Set.empty,[],rhs]
| na::nal ->
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Case (ci,p,c,cl) when
- eq_constr c (mkRel (List.index Name.equal na (fst (snd e))))
+ eq_constr sigma c (mkRel (List.index Name.equal na (fst (snd e))))
&& not (Int.equal (Array.length cl) 0)
&& (* don't contract if p dependent *)
- computable p (List.length ci.ci_pp_info.ind_tags) (* FIXME: can do better *) ->
- let clauses = build_tree na isgoal e ci cl in
+ computable sigma p (List.length ci.ci_pp_info.ind_tags) (* FIXME: can do better *) ->
+ let clauses = build_tree na isgoal e sigma ci cl in
List.flatten
- (List.map (fun (pat,rhs) ->
- let lines = align_tree nal isgoal rhs in
- List.map (fun (hd,rest) -> pat::hd,rest) lines)
+ (List.map (fun (ids,pat,rhs) ->
+ let lines = align_tree nal isgoal rhs sigma in
+ List.map (fun (ids',hd,rest) -> Id.Set.fold Id.Set.add ids ids',pat::hd,rest) lines)
clauses)
| _ ->
- let pat = PatVar(dl,update_name na rhs) in
- let mat = align_tree nal isgoal rhs in
- List.map (fun (hd,rest) -> pat::hd,rest) mat
+ let na = update_name sigma na rhs in
+ let pat = DAst.make @@ PatVar na in
+ let mat = align_tree nal isgoal rhs sigma in
+ List.map (fun (ids,hd,rest) -> Nameops.Name.fold_right Id.Set.add na ids,pat::hd,rest) mat
-and contract_branch isgoal e (cdn,can,mkpat,b) =
- let nal,rhs = decomp_branch cdn [] isgoal e b in
- let mat = align_tree nal isgoal rhs in
- List.map (fun (hd,rhs) -> (mkpat rhs hd,rhs)) mat
+and contract_branch isgoal e sigma (cdn,mkpat,rhs) =
+ let nal,rhs = decomp_branch cdn [] isgoal e sigma rhs in
+ let mat = align_tree nal isgoal rhs sigma in
+ List.map (fun (ids,hd,rhs) -> ids,mkpat rhs hd,rhs) mat
(**********************************************************************)
(* Transform internal representation of pattern-matching into list of *)
(* clauses *)
-let is_nondep_branch c l =
+let is_nondep_branch sigma c l =
try
(* FIXME: do better using tags from l *)
- let sign,ccl = decompose_lam_n_decls (List.length l) c in
- noccur_between 1 (Context.Rel.length sign) ccl
+ let sign,ccl = decompose_lam_n_decls sigma (List.length l) c in
+ noccur_between sigma 1 (Context.Rel.length sign) ccl
with e when CErrors.noncritical e -> (* Not eta-expanded or not reduced *)
false
let extract_nondep_branches test c b l =
let rec strip l r =
- match r,l with
- | r, [] -> r
- | GLambda (_,_,_,_,t), false::l -> strip l t
+ match DAst.get r, l with
+ | r', [] -> r
+ | GLambda (_,_,_,t), false::l -> strip l t
| GLetIn (_,_,_,t), true::l -> strip l t
(* FIXME: do we need adjustment? *)
| _,_ -> assert false in
@@ -338,10 +427,10 @@ let extract_nondep_branches test c b l =
let it_destRLambda_or_LetIn_names l c =
let rec aux l nal c =
- match c, l with
+ match DAst.get c, l with
| _, [] -> (List.rev nal,c)
- | GLambda (_,na,_,_,c), false::l -> aux l (na::nal) c
- | GLetIn (_,na,_,c), true::l -> aux l (na::nal) c
+ | GLambda (na,_,_,c), false::l -> aux l (na::nal) c
+ | GLetIn (na,_,_,c), true::l -> aux l (na::nal) c
| _, true::l -> (* let-expansion *) aux l (Anonymous :: nal) c
| _, false::l ->
(* eta-expansion *)
@@ -352,11 +441,11 @@ let it_destRLambda_or_LetIn_names l c =
x
in
let x = next (free_glob_vars c) in
- let a = GVar (dl,x) in
+ let a = DAst.make @@ GVar x in
aux l (Name x :: nal)
- (match c with
- | GApp (loc,p,l) -> GApp (loc,p,l@[a])
- | _ -> (GApp (dl,c,[a])))
+ (match DAst.get c with
+ | GApp (p,l) -> DAst.make ?loc:c.CAst.loc @@ GApp (p,l@[a])
+ | _ -> DAst.make @@ GApp (c,[a]))
in aux l [] c
let detype_case computable detype detype_eqns testdep avoid data p c bl =
@@ -368,17 +457,15 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl =
then
Anonymous, None, None
else
- match Option.map detype p with
- | None -> Anonymous, None, None
- | Some p ->
- let nl,typ = it_destRLambda_or_LetIn_names k p in
- let n,typ = match typ with
- | GLambda (_,x,_,t,c) -> x, c
- | _ -> Anonymous, typ in
- let aliastyp =
- if List.for_all (Name.equal Anonymous) nl then None
- else Some (dl,indsp,nl) in
- n, aliastyp, Some typ
+ let p = detype p in
+ let nl,typ = it_destRLambda_or_LetIn_names k p in
+ let n,typ = match DAst.get typ with
+ | GLambda (x,_,t,c) -> x, c
+ | _ -> Anonymous, typ in
+ let aliastyp =
+ if List.for_all (Name.equal Anonymous) nl then None
+ else Some (CAst.make (indsp,nl)) in
+ n, aliastyp, Some typ
in
let constructs = Array.init (Array.length bl) (fun i -> (indsp,i+1)) in
let tag =
@@ -399,20 +486,24 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl =
| LetStyle, None ->
let bl' = Array.map detype bl in
let (nal,d) = it_destRLambda_or_LetIn_names constagsl.(0) bl'.(0) in
- GLetTuple (dl,nal,(alias,pred),tomatch,d)
+ GLetTuple (nal,(alias,pred),tomatch,d)
| IfStyle, None ->
let bl' = Array.map detype bl in
let nondepbrs =
Array.map3 (extract_nondep_branches testdep) bl bl' constagsl in
if Array.for_all ((!=) None) nondepbrs then
- GIf (dl,tomatch,(alias,pred),
+ GIf (tomatch,(alias,pred),
Option.get nondepbrs.(0),Option.get nondepbrs.(1))
else
let eqnl = detype_eqns constructs constagsl bl in
- GCases (dl,tag,pred,[tomatch,(alias,aliastyp)],eqnl)
+ GCases (tag,pred,[tomatch,(alias,aliastyp)],eqnl)
| _ ->
let eqnl = detype_eqns constructs constagsl bl in
- GCases (dl,tag,pred,[tomatch,(alias,aliastyp)],eqnl)
+ GCases (tag,pred,[tomatch,(alias,aliastyp)],eqnl)
+
+let detype_universe sigma u =
+ let fn (l, n) = Some (Termops.reference_of_level sigma l, n) in
+ Univ.Universe.map fn u
let detype_sort sigma = function
| Prop Null -> GProp
@@ -420,7 +511,7 @@ let detype_sort sigma = function
| Type u ->
GType
(if !print_universes
- then [dl, Pp.string_of_ppcmds (Univ.Universe.pr_with (Evd.pr_evd_level sigma) u)]
+ then detype_universe sigma u
else [])
type binder_kind = BProd | BLambda | BLetIn
@@ -428,72 +519,85 @@ type binder_kind = BProd | BLambda | BLetIn
(**********************************************************************)
(* Main detyping function *)
-let detype_anonymous = ref (fun loc n -> anomaly ~label:"detype" (Pp.str "index to an anonymous variable"))
+let detype_anonymous = ref (fun ?loc n -> anomaly ~label:"detype" (Pp.str "index to an anonymous variable."))
let set_detype_anonymous f = detype_anonymous := f
let detype_level sigma l =
- GType (Some (dl, Pp.string_of_ppcmds (Evd.pr_evd_level sigma l)))
+ let l = Termops.reference_of_level sigma l in
+ GType (UNamed l)
let detype_instance sigma l =
+ let l = EInstance.kind sigma l in
if Univ.Instance.is_empty l then None
else Some (List.map (detype_level sigma) (Array.to_list (Univ.Instance.to_array l)))
-let rec detype flags avoid env sigma t =
- match kind_of_term (collapse_appl t) with
+let delay (type a) (d : a delay) (f : a delay -> _ -> _ -> _ -> _ -> _ -> a glob_constr_r) flags env avoid sigma t : a glob_constr_g =
+ match d with
+ | Now -> DAst.make (f d flags env avoid sigma t)
+ | Later -> DAst.delay (fun () -> f d flags env avoid sigma t)
+
+let rec detype d flags avoid env sigma t =
+ delay d detype_r flags avoid env sigma t
+
+and detype_r d flags avoid env sigma t =
+ match EConstr.kind sigma (collapse_appl sigma t) with
| Rel n ->
(try match lookup_name_of_rel n (fst env) with
- | Name id -> GVar (dl, id)
- | Anonymous -> !detype_anonymous dl n
+ | Name id -> GVar id
+ | Anonymous -> GVar (!detype_anonymous n)
with Not_found ->
let s = "_UNBOUND_REL_"^(string_of_int n)
- in GVar (dl, Id.of_string s))
+ in GVar (Id.of_string s))
| Meta n ->
(* Meta in constr are not user-parsable and are mapped to Evar *)
- (* using numbers to be unparsable *)
- GEvar (dl, Id.of_string ("M" ^ string_of_int n), [])
+ if n = Constr_matching.special_meta then
+ (* Using a dash to be unparsable *)
+ GEvar (Id.of_string_soft "CONTEXT-HOLE", [])
+ else
+ GEvar (Id.of_string_soft ("M" ^ string_of_int n), [])
| Var id ->
- (try let _ = Global.lookup_named id in GRef (dl, VarRef id, None)
- with Not_found -> GVar (dl, id))
- | Sort s -> GSort (dl,detype_sort sigma s)
+ (try let _ = Global.lookup_named id in GRef (VarRef id, None)
+ with Not_found -> GVar id)
+ | Sort s -> GSort (detype_sort sigma (ESorts.kind sigma s))
| Cast (c1,REVERTcast,c2) when not !Flags.raw_print ->
- detype flags avoid env sigma c1
+ DAst.get (detype d flags avoid env sigma c1)
| Cast (c1,k,c2) ->
- let d1 = detype flags avoid env sigma c1 in
- let d2 = detype flags avoid env sigma c2 in
+ let d1 = detype d flags avoid env sigma c1 in
+ let d2 = detype d flags avoid env sigma c2 in
let cast = match k with
| VMcast -> CastVM d2
| NATIVEcast -> CastNative d2
| _ -> CastConv d2
in
- GCast(dl,d1,cast)
- | Prod (na,ty,c) -> detype_binder flags BProd avoid env sigma na None ty c
- | Lambda (na,ty,c) -> detype_binder flags BLambda avoid env sigma na None ty c
- | LetIn (na,b,ty,c) -> detype_binder flags BLetIn avoid env sigma na (Some b) ty c
+ GCast(d1,cast)
+ | Prod (na,ty,c) -> detype_binder d flags BProd avoid env sigma na None ty c
+ | Lambda (na,ty,c) -> detype_binder d flags BLambda avoid env sigma na None ty c
+ | LetIn (na,b,ty,c) -> detype_binder d flags BLetIn avoid env sigma na (Some b) ty c
| App (f,args) ->
let mkapp f' args' =
- match f' with
- | GApp (dl',f',args'') ->
- GApp (dl,f',args''@args')
- | _ -> GApp (dl,f',args')
+ match DAst.get f' with
+ | GApp (f',args'') ->
+ GApp (f',args''@args')
+ | _ -> GApp (f',args')
in
- mkapp (detype flags avoid env sigma f)
- (Array.map_to_list (detype flags avoid env sigma) args)
- | Const (sp,u) -> GRef (dl, ConstRef sp, detype_instance sigma u)
+ mkapp (detype d flags avoid env sigma f)
+ (Array.map_to_list (detype d flags avoid env sigma) args)
+ | Const (sp,u) -> GRef (ConstRef sp, detype_instance sigma u)
| Proj (p,c) ->
let noparams () =
let pb = Environ.lookup_projection p (snd env) in
let pars = pb.Declarations.proj_npars in
- let hole = GHole(Loc.ghost,Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None) in
+ let hole = DAst.make @@ GHole(Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None) in
let args = List.make pars hole in
- GApp (dl, GRef (dl, ConstRef (Projection.constant p), None),
- (args @ [detype flags avoid env sigma c]))
+ GApp (DAst.make @@ GRef (ConstRef (Projection.constant p), None),
+ (args @ [detype d flags avoid env sigma c]))
in
if fst flags || !Flags.in_debugger || !Flags.in_toplevel then
try noparams ()
with _ ->
(* lax mode, used by debug printers only *)
- GApp (dl, GRef (dl, ConstRef (Projection.constant p), None),
- [detype flags avoid env sigma c])
+ GApp (DAst.make @@ GRef (ConstRef (Projection.constant p), None),
+ [detype d flags avoid env sigma c])
else
if print_primproj_compatibility () && Projection.unfolded p then
(** Print the compatibility match version *)
@@ -504,16 +608,18 @@ let rec detype flags avoid env sigma t =
let ty = Retyping.get_type_of (snd env) sigma c in
let ((ind,u), args) = Inductiveops.find_mrectype (snd env) sigma ty in
let body' = strip_lam_assum body in
- let body' = subst_instance_constr u body' in
+ let u = EInstance.kind sigma u in
+ let body' = CVars.subst_instance_constr u body' in
+ let body' = EConstr.of_constr body' in
substl (c :: List.rev args) body'
with Retyping.RetypeError _ | Not_found ->
anomaly (str"Cannot detype an unfolded primitive projection.")
- in detype flags avoid env sigma c'
+ in DAst.get (detype d flags avoid env sigma c')
else
if print_primproj_params () then
try
let c = Retyping.expand_projection (snd env) sigma p c [] in
- detype flags avoid env sigma c
+ DAst.get (detype d flags avoid env sigma c)
with Retyping.RetypeError _ -> noparams ()
else noparams ()
@@ -523,133 +629,134 @@ let rec detype flags avoid env sigma t =
| LocalDef _ -> true
| LocalAssum (id,_) ->
try let n = List.index Name.equal (Name id) (fst env) in
- isRelN n c
- with Not_found -> isVarId id c
+ isRelN sigma n c
+ with Not_found -> isVarId sigma id c
in
let id,l =
try
let id = match Evd.evar_ident evk sigma with
- | None -> Evd.pr_evar_suggested_name evk sigma
+ | None -> Termops.pr_evar_suggested_name evk sigma
| Some id -> id
in
let l = Evd.evar_instance_array bound_to_itself_or_letin (Evd.find sigma evk) cl in
- let fvs,rels = List.fold_left (fun (fvs,rels) (_,c) -> match kind_of_term c with Rel n -> (fvs,Int.Set.add n rels) | Var id -> (Id.Set.add id fvs,rels) | _ -> (fvs,rels)) (Id.Set.empty,Int.Set.empty) l in
- let l = Evd.evar_instance_array (fun d c -> not !print_evar_arguments && (bound_to_itself_or_letin d c && not (isRel c && Int.Set.mem (destRel c) rels || isVar c && (Id.Set.mem (destVar c) fvs)))) (Evd.find sigma evk) cl in
+ let fvs,rels = List.fold_left (fun (fvs,rels) (_,c) -> match EConstr.kind sigma c with Rel n -> (fvs,Int.Set.add n rels) | Var id -> (Id.Set.add id fvs,rels) | _ -> (fvs,rels)) (Id.Set.empty,Int.Set.empty) l in
+ let l = Evd.evar_instance_array (fun d c -> not !print_evar_arguments && (bound_to_itself_or_letin d c && not (isRel sigma c && Int.Set.mem (destRel sigma c) rels || isVar sigma c && (Id.Set.mem (destVar sigma c) fvs)))) (Evd.find sigma evk) cl in
id,l
with Not_found ->
- Id.of_string ("X" ^ string_of_int (Evar.repr evk)),
+ Id.of_string ("X" ^ string_of_int (Evar.repr evk)),
(Array.map_to_list (fun c -> (Id.of_string "__",c)) cl)
in
- GEvar (dl,id,
- List.map (on_snd (detype flags avoid env sigma)) l)
+ GEvar (id,
+ List.map (on_snd (detype d flags avoid env sigma)) l)
| Ind (ind_sp,u) ->
- GRef (dl, IndRef ind_sp, detype_instance sigma u)
+ GRef (IndRef ind_sp, detype_instance sigma u)
| Construct (cstr_sp,u) ->
- GRef (dl, ConstructRef cstr_sp, detype_instance sigma u)
+ GRef (ConstructRef cstr_sp, detype_instance sigma u)
| Case (ci,p,c,bl) ->
- let comp = computable p (List.length (ci.ci_pp_info.ind_tags)) in
- detype_case comp (detype flags avoid env sigma)
- (detype_eqns flags avoid env sigma ci comp)
- is_nondep_branch avoid
+ let comp = computable sigma p (List.length (ci.ci_pp_info.ind_tags)) in
+ detype_case comp (detype d flags avoid env sigma)
+ (detype_eqns d flags avoid env sigma ci comp)
+ (is_nondep_branch sigma) avoid
(ci.ci_ind,ci.ci_pp_info.style,
ci.ci_pp_info.cstr_tags,ci.ci_pp_info.ind_tags)
- (Some p) c bl
- | Fix (nvn,recdef) -> detype_fix flags avoid env sigma nvn recdef
- | CoFix (n,recdef) -> detype_cofix flags avoid env sigma n recdef
+ p c bl
+ | Fix (nvn,recdef) -> detype_fix d flags avoid env sigma nvn recdef
+ | CoFix (n,recdef) -> detype_cofix d flags avoid env sigma n recdef
-and detype_fix flags avoid env sigma (vn,_ as nvn) (names,tys,bodies) =
+and detype_fix d flags avoid env sigma (vn,_ as nvn) (names,tys,bodies) =
let def_avoid, def_env, lfi =
Array.fold_left2
(fun (avoid, env, l) na ty ->
let id = next_name_away na avoid in
- (id::avoid, add_name (Name id) None ty env, id::l))
+ (Id.Set.add id avoid, add_name (Name id) None ty env, id::l))
(avoid, env, []) names tys in
let n = Array.length tys in
let v = Array.map3
- (fun c t i -> share_names flags (i+1) [] def_avoid def_env sigma c (lift n t))
+ (fun c t i -> share_names d flags (i+1) [] def_avoid def_env sigma c (lift n t))
bodies tys vn in
- GRec(dl,GFix (Array.map (fun i -> Some i, GStructRec) (fst nvn), snd nvn),Array.of_list (List.rev lfi),
+ GRec(GFix (Array.map (fun i -> Some i, GStructRec) (fst nvn), snd nvn),Array.of_list (List.rev lfi),
Array.map (fun (bl,_,_) -> bl) v,
Array.map (fun (_,_,ty) -> ty) v,
Array.map (fun (_,bd,_) -> bd) v)
-and detype_cofix flags avoid env sigma n (names,tys,bodies) =
+and detype_cofix d flags avoid env sigma n (names,tys,bodies) =
let def_avoid, def_env, lfi =
Array.fold_left2
(fun (avoid, env, l) na ty ->
let id = next_name_away na avoid in
- (id::avoid, add_name (Name id) None ty env, id::l))
+ (Id.Set.add id avoid, add_name (Name id) None ty env, id::l))
(avoid, env, []) names tys in
let ntys = Array.length tys in
let v = Array.map2
- (fun c t -> share_names flags 0 [] def_avoid def_env sigma c (lift ntys t))
+ (fun c t -> share_names d flags 0 [] def_avoid def_env sigma c (lift ntys t))
bodies tys in
- GRec(dl,GCoFix n,Array.of_list (List.rev lfi),
+ GRec(GCoFix n,Array.of_list (List.rev lfi),
Array.map (fun (bl,_,_) -> bl) v,
Array.map (fun (_,_,ty) -> ty) v,
Array.map (fun (_,bd,_) -> bd) v)
-and share_names flags n l avoid env sigma c t =
- match kind_of_term c, kind_of_term t with
+and share_names d flags n l avoid env sigma c t =
+ match EConstr.kind sigma c, EConstr.kind sigma t with
(* factorize even when not necessary to have better presentation *)
| Lambda (na,t,c), Prod (na',t',c') ->
let na = match (na,na') with
Name _, _ -> na
| _, Name _ -> na'
| _ -> na in
- let t' = detype flags avoid env sigma t in
+ let t' = detype d flags avoid env sigma t in
let id = next_name_away na avoid in
- let avoid = id::avoid and env = add_name (Name id) None t env in
- share_names flags (n-1) ((Name id,Explicit,None,t')::l) avoid env sigma c c'
+ let avoid = Id.Set.add id avoid and env = add_name (Name id) None t env in
+ share_names d flags (n-1) ((Name id,Explicit,None,t')::l) avoid env sigma c c'
(* May occur for fix built interactively *)
| LetIn (na,b,t',c), _ when n > 0 ->
- let t'' = detype flags avoid env sigma t' in
- let b' = detype flags avoid env sigma b in
+ let t'' = detype d flags avoid env sigma t' in
+ let b' = detype d flags avoid env sigma b in
let id = next_name_away na avoid in
- let avoid = id::avoid and env = add_name (Name id) (Some b) t' env in
- share_names flags n ((Name id,Explicit,Some b',t'')::l) avoid env sigma c (lift 1 t)
+ let avoid = Id.Set. add id avoid and env = add_name (Name id) (Some b) t' env in
+ share_names d flags n ((Name id,Explicit,Some b',t'')::l) avoid env sigma c (lift 1 t)
(* Only if built with the f/n notation or w/o let-expansion in types *)
| _, LetIn (_,b,_,t) when n > 0 ->
- share_names flags n l avoid env sigma c (subst1 b t)
+ share_names d flags n l avoid env sigma c (subst1 b t)
(* If it is an open proof: we cheat and eta-expand *)
| _, Prod (na',t',c') when n > 0 ->
- let t'' = detype flags avoid env sigma t' in
+ let t'' = detype d flags avoid env sigma t' in
let id = next_name_away na' avoid in
- let avoid = id::avoid and env = add_name (Name id) None t' env in
+ let avoid = Id.Set.add id avoid and env = add_name (Name id) None t' env in
let appc = mkApp (lift 1 c,[|mkRel 1|]) in
- share_names flags (n-1) ((Name id,Explicit,None,t'')::l) avoid env sigma appc c'
+ share_names d flags (n-1) ((Name id,Explicit,None,t'')::l) avoid env sigma appc c'
(* If built with the f/n notation: we renounce to share names *)
| _ ->
if n>0 then Feedback.msg_debug (strbrk "Detyping.detype: cannot factorize fix enough");
- let c = detype flags avoid env sigma c in
- let t = detype flags avoid env sigma t in
+ let c = detype d flags avoid env sigma c in
+ let t = detype d flags avoid env sigma t in
(List.rev l,c,t)
-and detype_eqns flags avoid env sigma ci computable constructs consnargsl bl =
+and detype_eqns d flags avoid env sigma ci computable constructs consnargsl bl =
try
if !Flags.raw_print || not (reverse_matching ()) then raise Exit;
- let mat = build_tree Anonymous (snd flags) (avoid,env) ci bl in
- List.map (fun (pat,((avoid,env),c)) -> (dl,[],[pat],detype flags avoid env sigma c))
+ let mat = build_tree Anonymous (snd flags) (avoid,env) sigma ci bl in
+ List.map (fun (ids,pat,((avoid,env),c)) ->
+ CAst.make (Id.Set.elements ids,[pat],detype d flags avoid env sigma c))
mat
with e when CErrors.noncritical e ->
Array.to_list
- (Array.map3 (detype_eqn flags avoid env sigma) constructs consnargsl bl)
+ (Array.map3 (detype_eqn d flags avoid env sigma) constructs consnargsl bl)
-and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs branch =
+and detype_eqn d (lax,isgoal as flags) avoid env sigma constr construct_nargs branch =
let make_pat x avoid env b body ty ids =
- if force_wildcard () && noccurn 1 b then
- PatVar (dl,Anonymous),avoid,(add_name Anonymous body ty env),ids
+ if force_wildcard () && noccurn sigma 1 b then
+ DAst.make @@ PatVar Anonymous,avoid,(add_name Anonymous body ty env),ids
else
let flag = if isgoal then RenamingForGoal else RenamingForCasesPattern (fst env,b) in
- let na,avoid' = compute_displayed_name_in flag avoid x b in
- PatVar (dl,na),avoid',(add_name na body ty env),add_vname ids na
+ let na,avoid' = compute_displayed_name_in sigma flag avoid x b in
+ DAst.make (PatVar na),avoid',(add_name na body ty env),add_vname ids na
in
let rec buildrec ids patlist avoid env l b =
- match kind_of_term b, l with
- | _, [] ->
- (dl, Id.Set.elements ids,
- [PatCstr(dl, constr, List.rev patlist,Anonymous)],
- detype flags avoid env sigma b)
+ match EConstr.kind sigma b, l with
+ | _, [] -> CAst.make @@
+ (Id.Set.elements ids,
+ [DAst.make @@ PatCstr(constr, List.rev patlist,Anonymous)],
+ detype d flags avoid env sigma b)
| Lambda (x,t,b), false::l ->
let pat,new_avoid,new_env,new_ids = make_pat x avoid env b None t ids in
buildrec new_ids (pat::patlist) new_avoid new_env l b
@@ -662,7 +769,7 @@ and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs bran
buildrec ids patlist avoid env l c
| _, true::l ->
- let pat = PatVar (dl,Anonymous) in
+ let pat = DAst.make @@ PatVar Anonymous in
buildrec ids (pat::patlist) avoid env l b
| _, false::l ->
@@ -677,25 +784,24 @@ and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs bran
in
buildrec Id.Set.empty [] avoid env construct_nargs branch
-and detype_binder (lax,isgoal as flags) bk avoid env sigma na body ty c =
+and detype_binder d (lax,isgoal as flags) bk avoid env sigma na body ty c =
let flag = if isgoal then RenamingForGoal else RenamingElsewhereFor (fst env,c) in
let na',avoid' = match bk with
- | BLetIn -> compute_displayed_let_name_in flag avoid na c
- | _ -> compute_displayed_name_in flag avoid na c in
- let r = detype flags avoid' (add_name na' body ty env) sigma c in
+ | BLetIn -> compute_displayed_let_name_in sigma flag avoid na c
+ | _ -> compute_displayed_name_in sigma flag avoid na c in
+ let r = detype d flags avoid' (add_name na' body ty env) sigma c in
match bk with
- | BProd -> GProd (dl, na',Explicit,detype (lax,false) avoid env sigma ty, r)
- | BLambda -> GLambda (dl, na',Explicit,detype (lax,false) avoid env sigma ty, r)
+ | BProd -> GProd (na',Explicit,detype d (lax,false) avoid env sigma ty, r)
+ | BLambda -> GLambda (na',Explicit,detype d (lax,false) avoid env sigma ty, r)
| BLetIn ->
- let c = detype (lax,false) avoid env sigma (Option.get body) in
+ let c = detype d (lax,false) avoid env sigma (Option.get body) in
(* Heuristic: we display the type if in Prop *)
let s = try Retyping.get_sort_family_of (snd env) sigma ty with _ when !Flags.in_debugger || !Flags.in_toplevel -> InType (* Can fail because of sigma missing in debugger *) in
- let c = if s != InProp then c else
- GCast (dl, c, CastConv (detype (lax,false) avoid env sigma ty)) in
- GLetIn (dl, na', c, r)
+ let t = if s != InProp && not !Flags.raw_print then None else Some (detype d (lax,false) avoid env sigma ty) in
+ GLetIn (na', c, t, r)
-let detype_rel_context ?(lax=false) where avoid env sigma sign =
- let where = Option.map (fun c -> it_mkLambda_or_LetIn c sign) where in
+let detype_rel_context d ?(lax=false) where avoid env sigma sign =
+ let where = Option.map (fun c -> EConstr.it_mkLambda_or_LetIn c sign) where in
let rec aux avoid env = function
| [] -> []
| decl::rest ->
@@ -707,26 +813,30 @@ let detype_rel_context ?(lax=false) where avoid env sigma sign =
| None -> na,avoid
| Some c ->
if is_local_def decl then
- compute_displayed_let_name_in
+ compute_displayed_let_name_in sigma
(RenamingElsewhereFor (fst env,c)) avoid na c
else
- compute_displayed_name_in
+ compute_displayed_name_in sigma
(RenamingElsewhereFor (fst env,c)) avoid na c in
let b = match decl with
| LocalAssum _ -> None
| LocalDef (_,b,_) -> Some b
in
- let b' = Option.map (detype (lax,false) avoid env sigma) b in
- let t' = detype (lax,false) avoid env sigma t in
+ let b' = Option.map (detype d (lax,false) avoid env sigma) b in
+ let t' = detype d (lax,false) avoid env sigma t in
(na',Explicit,b',t') :: aux avoid' (add_name na' b t env) rest
in aux avoid env (List.rev sign)
let detype_names isgoal avoid nenv env sigma t =
- detype (false,isgoal) avoid (nenv,env) sigma t
-let detype ?(lax=false) isgoal avoid env sigma t =
- detype (lax,isgoal) avoid (names_of_rel_context env, env) sigma t
+ detype Now (false,isgoal) avoid (nenv,env) sigma t
+let detype d ?(lax=false) isgoal avoid env sigma t =
+ detype d (lax,isgoal) avoid (names_of_rel_context env, env) sigma t
+
+let detype_rel_context d ?lax where avoid env sigma sign =
+ detype_rel_context d ?lax where avoid env sigma sign
let detype_closed_glob ?lax isgoal avoid env sigma t =
+ let open Context.Rel.Declaration in
let convert_id cl id =
try Id.Map.find id cl.idents
with Not_found -> id
@@ -735,11 +845,11 @@ let detype_closed_glob ?lax isgoal avoid env sigma t =
| Name id -> Name (convert_id cl id)
| Anonymous -> Anonymous
in
- let rec detype_closed_glob cl = function
- | GVar (loc,id) ->
+ let rec detype_closed_glob cl cg : Glob_term.glob_constr = DAst.map (function
+ | GVar id ->
(* if [id] is bound to a name. *)
begin try
- GVar(loc,Id.Map.find id cl.idents)
+ GVar(Id.Map.find id cl.idents)
(* if [id] is bound to a typed term *)
with Not_found -> try
(* assumes [detype] does not raise [Not_found] exceptions *)
@@ -747,127 +857,131 @@ let detype_closed_glob ?lax isgoal avoid env sigma t =
(* spiwack: I'm not sure it is the right thing to do,
but I'm computing the detyping environment like
[Printer.pr_constr_under_binders_env] does. *)
- let assums = List.map (fun id -> (Name id,(* dummy *) mkProp)) b in
- let env = Termops.push_rels_assum assums env in
- detype ?lax isgoal avoid env sigma c
+ let assums = List.map (fun id -> LocalAssum (Name id,(* dummy *) mkProp)) b in
+ let env = push_rel_context assums env in
+ DAst.get (detype Now ?lax isgoal avoid env sigma c)
(* if [id] is bound to a [closed_glob_constr]. *)
with Not_found -> try
let {closure;term} = Id.Map.find id cl.untyped in
- detype_closed_glob closure term
+ DAst.get (detype_closed_glob closure term)
(* Otherwise [id] stands for itself *)
with Not_found ->
- GVar(loc,id)
+ GVar id
end
- | GLambda (loc,id,k,t,c) ->
+ | GLambda (id,k,t,c) ->
let id = convert_name cl id in
- GLambda(loc,id,k,detype_closed_glob cl t, detype_closed_glob cl c)
- | GProd (loc,id,k,t,c) ->
+ GLambda(id,k,detype_closed_glob cl t, detype_closed_glob cl c)
+ | GProd (id,k,t,c) ->
let id = convert_name cl id in
- GProd(loc,id,k,detype_closed_glob cl t, detype_closed_glob cl c)
- | GLetIn (loc,id,b,e) ->
+ GProd(id,k,detype_closed_glob cl t, detype_closed_glob cl c)
+ | GLetIn (id,b,t,e) ->
let id = convert_name cl id in
- GLetIn(loc,id,detype_closed_glob cl b, detype_closed_glob cl e)
- | GLetTuple (loc,ids,(n,r),b,e) ->
+ GLetIn(id,detype_closed_glob cl b, Option.map (detype_closed_glob cl) t, detype_closed_glob cl e)
+ | GLetTuple (ids,(n,r),b,e) ->
let ids = List.map (convert_name cl) ids in
let n = convert_name cl n in
- GLetTuple (loc,ids,(n,r),detype_closed_glob cl b, detype_closed_glob cl e)
- | GCases (loc,sty,po,tml,eqns) ->
+ GLetTuple (ids,(n,r),detype_closed_glob cl b, detype_closed_glob cl e)
+ | GCases (sty,po,tml,eqns) ->
let (tml,eqns) =
Glob_ops.map_pattern_binders (fun na -> convert_name cl na) tml eqns
in
let (tml,eqns) =
Glob_ops.map_pattern (fun c -> detype_closed_glob cl c) tml eqns
in
- GCases(loc,sty,po,tml,eqns)
+ GCases(sty,po,tml,eqns)
| c ->
- Glob_ops.map_glob_constr (detype_closed_glob cl) c
+ DAst.get (Glob_ops.map_glob_constr (detype_closed_glob cl) cg)
+ ) cg
in
detype_closed_glob t.closure t.term
(**********************************************************************)
(* Module substitution: relies on detyping *)
-let rec subst_cases_pattern subst pat =
- match pat with
- | PatVar _ -> pat
- | PatCstr (loc,((kn,i),j),cpl,n) ->
+let rec subst_cases_pattern subst = DAst.map (function
+ | PatVar _ as pat -> pat
+ | PatCstr (((kn,i),j),cpl,n) as pat ->
let kn' = subst_mind subst kn
and cpl' = List.smartmap (subst_cases_pattern subst) cpl in
if kn' == kn && cpl' == cpl then pat else
- PatCstr (loc,((kn',i),j),cpl',n)
+ PatCstr (((kn',i),j),cpl',n)
+ )
let (f_subst_genarg, subst_genarg_hook) = Hook.make ()
-let rec subst_glob_constr subst raw =
- match raw with
- | GRef (loc,ref,u) ->
+let rec subst_glob_constr subst = DAst.map (function
+ | GRef (ref,u) as raw ->
let ref',t = subst_global subst ref in
if ref' == ref then raw else
- detype false [] (Global.env()) Evd.empty t
+ DAst.get (detype Now false Id.Set.empty (Global.env()) Evd.empty (EConstr.of_constr t))
- | GVar _ -> raw
- | GEvar _ -> raw
- | GPatVar _ -> raw
+ | GSort _
+ | GVar _
+ | GEvar _
+ | GPatVar _ as raw -> raw
- | GApp (loc,r,rl) ->
+ | GApp (r,rl) as raw ->
let r' = subst_glob_constr subst r
and rl' = List.smartmap (subst_glob_constr subst) rl in
if r' == r && rl' == rl then raw else
- GApp(loc,r',rl')
+ GApp(r',rl')
- | GLambda (loc,n,bk,r1,r2) ->
+ | GLambda (n,bk,r1,r2) as raw ->
let r1' = subst_glob_constr subst r1 and r2' = subst_glob_constr subst r2 in
if r1' == r1 && r2' == r2 then raw else
- GLambda (loc,n,bk,r1',r2')
+ GLambda (n,bk,r1',r2')
- | GProd (loc,n,bk,r1,r2) ->
+ | GProd (n,bk,r1,r2) as raw ->
let r1' = subst_glob_constr subst r1 and r2' = subst_glob_constr subst r2 in
if r1' == r1 && r2' == r2 then raw else
- GProd (loc,n,bk,r1',r2')
+ GProd (n,bk,r1',r2')
- | GLetIn (loc,n,r1,r2) ->
- let r1' = subst_glob_constr subst r1 and r2' = subst_glob_constr subst r2 in
- if r1' == r1 && r2' == r2 then raw else
- GLetIn (loc,n,r1',r2')
+ | GLetIn (n,r1,t,r2) as raw ->
+ let r1' = subst_glob_constr subst r1 in
+ let r2' = subst_glob_constr subst r2 in
+ let t' = Option.smartmap (subst_glob_constr subst) t in
+ if r1' == r1 && t == t' && r2' == r2 then raw else
+ GLetIn (n,r1',t',r2')
- | GCases (loc,sty,rtno,rl,branches) ->
+ | GCases (sty,rtno,rl,branches) as raw ->
+ let open CAst in
let rtno' = Option.smartmap (subst_glob_constr subst) rtno
and rl' = List.smartmap (fun (a,x as y) ->
let a' = subst_glob_constr subst a in
let (n,topt) = x in
let topt' = Option.smartmap
- (fun (loc,(sp,i),y as t) ->
+ (fun ({loc;v=((sp,i),y)} as t) ->
let sp' = subst_mind subst sp in
- if sp == sp' then t else (loc,(sp',i),y)) topt in
+ if sp == sp' then t else CAst.(make ?loc ((sp',i),y))) topt in
if a == a' && topt == topt' then y else (a',(n,topt'))) rl
and branches' = List.smartmap
- (fun (loc,idl,cpl,r as branch) ->
+ (fun ({loc;v=(idl,cpl,r)} as branch) ->
let cpl' =
List.smartmap (subst_cases_pattern subst) cpl
and r' = subst_glob_constr subst r in
if cpl' == cpl && r' == r then branch else
- (loc,idl,cpl',r'))
+ CAst.(make ?loc (idl,cpl',r')))
branches
in
if rtno' == rtno && rl' == rl && branches' == branches then raw else
- GCases (loc,sty,rtno',rl',branches')
+ GCases (sty,rtno',rl',branches')
- | GLetTuple (loc,nal,(na,po),b,c) ->
+ | GLetTuple (nal,(na,po),b,c) as raw ->
let po' = Option.smartmap (subst_glob_constr subst) po
and b' = subst_glob_constr subst b
and c' = subst_glob_constr subst c in
if po' == po && b' == b && c' == c then raw else
- GLetTuple (loc,nal,(na,po'),b',c')
+ GLetTuple (nal,(na,po'),b',c')
- | GIf (loc,c,(na,po),b1,b2) ->
+ | GIf (c,(na,po),b1,b2) as raw ->
let po' = Option.smartmap (subst_glob_constr subst) po
and b1' = subst_glob_constr subst b1
and b2' = subst_glob_constr subst b2
and c' = subst_glob_constr subst c in
if c' == c && po' == po && b1' == b1 && b2' == b2 then raw else
- GIf (loc,c',(na,po'),b1',b2')
+ GIf (c',(na,po'),b1',b2')
- | GRec (loc,fix,ida,bl,ra1,ra2) ->
+ | GRec (fix,ida,bl,ra1,ra2) as raw ->
let ra1' = Array.smartmap (subst_glob_constr subst) ra1
and ra2' = Array.smartmap (subst_glob_constr subst) ra2 in
let bl' = Array.smartmap
@@ -877,11 +991,9 @@ let rec subst_glob_constr subst raw =
if ty'==ty && obd'==obd then dcl else (na,k,obd',ty')))
bl in
if ra1' == ra1 && ra2' == ra2 && bl'==bl then raw else
- GRec (loc,fix,ida,bl',ra1',ra2')
-
- | GSort _ -> raw
+ GRec (fix,ida,bl',ra1',ra2')
- | GHole (loc, knd, naming, solve) ->
+ | GHole (knd, naming, solve) as raw ->
let nknd = match knd with
| Evar_kinds.ImplicitArg (ref, i, b) ->
let nref, _ = subst_global subst ref in
@@ -890,25 +1002,25 @@ let rec subst_glob_constr subst raw =
in
let nsolve = Option.smartmap (Hook.get f_subst_genarg subst) solve in
if nsolve == solve && nknd == knd then raw
- else GHole (loc, nknd, naming, nsolve)
+ else GHole (nknd, naming, nsolve)
- | GCast (loc,r1,k) ->
+ | GCast (r1,k) as raw ->
let r1' = subst_glob_constr subst r1 in
let k' = Miscops.smartmap_cast_type (subst_glob_constr subst) k in
- if r1' == r1 && k' == k then raw else GCast (loc,r1',k')
+ if r1' == r1 && k' == k then raw else GCast (r1',k')
+ )
(* Utilities to transform kernel cases to simple pattern-matching problem *)
let simple_cases_matrix_of_branches ind brs =
List.map (fun (i,n,b) ->
let nal,c = it_destRLambda_or_LetIn_names n b in
- let mkPatVar na = PatVar (Loc.ghost,na) in
- let p = PatCstr (Loc.ghost,(ind,i+1),List.map mkPatVar nal,Anonymous) in
- let map name = try Some (Nameops.out_name name) with Failure _ -> None in
- let ids = List.map_filter map nal in
- (Loc.ghost,ids,[p],c))
+ let mkPatVar na = DAst.make @@ PatVar na in
+ let p = DAst.make @@ PatCstr ((ind,i+1),List.map mkPatVar nal,Anonymous) in
+ let ids = List.map_filter Nameops.Name.to_option nal in
+ CAst.make @@ (ids,[p],c))
brs
let return_type_of_predicate ind nrealargs_tags pred =
let nal,p = it_destRLambda_or_LetIn_names (nrealargs_tags@[false]) pred in
- (List.hd nal, Some (Loc.ghost, ind, List.tl nal)), Some p
+ (List.hd nal, Some (CAst.make (ind, List.tl nal))), Some p
diff --git a/pretyping/detyping.mli b/pretyping/detyping.mli
index c51cb0f4..32b94e1b 100644
--- a/pretyping/detyping.mli
+++ b/pretyping/detyping.mli
@@ -1,19 +1,26 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
open Environ
+open EConstr
open Glob_term
open Termops
open Mod_subst
open Misctypes
open Evd
+open Ltac_pretype
+
+type _ delay =
+| Now : 'a delay
+| Later : [ `thunk ] delay
(** Should we keep details of universes during detyping ? *)
val print_universes : bool ref
@@ -21,39 +28,43 @@ val print_universes : bool ref
(** If true, prints full local context of evars *)
val print_evar_arguments : bool ref
+(** If true, contract branches with same r.h.s. and same matching
+ variables in a disjunctive pattern *)
+val print_factorize_match_patterns : bool ref
+
+(** If true and the last non unique clause of a "match" is a
+ variable-free disjunctive pattern, turn it into a catch-call case *)
+val print_allow_match_default_clause : bool ref
+
val subst_cases_pattern : substitution -> cases_pattern -> cases_pattern
val subst_glob_constr : substitution -> glob_constr -> glob_constr
+val factorize_eqns : 'a cases_clauses_g -> 'a disjunctive_cases_clauses_g
+
(** [detype isgoal avoid ctx c] turns a closed [c], into a glob_constr
de Bruijn indexes are turned to bound names, avoiding names in [avoid]
[isgoal] tells if naming must avoid global-level synonyms as intro does
[ctx] gives the names of the free variables *)
-val detype_names : bool -> Id.t list -> names_context -> env -> evar_map -> constr -> glob_constr
-
-val detype : ?lax:bool -> bool -> Id.t list -> env -> evar_map -> constr -> glob_constr
+val detype_names : bool -> Id.Set.t -> names_context -> env -> evar_map -> constr -> glob_constr
-val detype_case :
- bool -> (constr -> glob_constr) ->
- (constructor array -> bool list array -> constr array ->
- (Loc.t * Id.t list * cases_pattern list * glob_constr) list) ->
- (constr -> bool list -> bool) ->
- Id.t list -> inductive * case_style * bool list array * bool list ->
- constr option -> constr -> constr array -> glob_constr
+val detype : 'a delay -> ?lax:bool -> bool -> Id.Set.t -> env -> evar_map -> constr -> 'a glob_constr_g
-val detype_sort : evar_map -> sorts -> glob_sort
+val detype_sort : evar_map -> Sorts.t -> glob_sort
-val detype_rel_context : ?lax:bool -> constr option -> Id.t list -> (names_context * env) ->
- evar_map -> Context.Rel.t -> glob_decl list
+val detype_rel_context : 'a delay -> ?lax:bool -> constr option -> Id.Set.t -> (names_context * env) ->
+ evar_map -> rel_context -> 'a glob_decl_g list
-val detype_closed_glob : ?lax:bool -> bool -> Id.t list -> env -> evar_map -> closed_glob_constr -> glob_constr
+val detype_closed_glob : ?lax:bool -> bool -> Id.Set.t -> env -> evar_map -> closed_glob_constr -> glob_constr
(** look for the index of a named var or a nondep var as it is renamed *)
-val lookup_name_as_displayed : env -> constr -> Id.t -> int option
-val lookup_index_as_renamed : env -> constr -> int -> int option
+val lookup_name_as_displayed : env -> evar_map -> constr -> Id.t -> int option
+val lookup_index_as_renamed : env -> evar_map -> constr -> int -> int option
+
+(* XXX: This is a hack and should go away *)
+val set_detype_anonymous : (?loc:Loc.t -> int -> Id.t) -> unit
-val set_detype_anonymous : (Loc.t -> int -> glob_constr) -> unit
val force_wildcard : unit -> bool
val synthetize_type : unit -> bool
@@ -71,7 +82,7 @@ val subst_genarg_hook :
module PrintingInductiveMake :
functor (Test : sig
val encode : Libnames.reference -> Names.inductive
- val member_message : Pp.std_ppcmds -> bool -> Pp.std_ppcmds
+ val member_message : Pp.t -> bool -> Pp.t
val field : string
val title : string
end) ->
@@ -80,9 +91,9 @@ module PrintingInductiveMake :
val compare : t -> t -> int
val encode : Libnames.reference -> Names.inductive
val subst : substitution -> t -> t
- val printer : t -> Pp.std_ppcmds
+ val printer : t -> Pp.t
val key : Goptions.option_name
val title : string
- val member_message : t -> bool -> Pp.std_ppcmds
+ val member_message : t -> bool -> Pp.t
val synchronous : bool
end
diff --git a/pretyping/doc.tex b/pretyping/doc.tex
deleted file mode 100644
index d92a027e..00000000
--- a/pretyping/doc.tex
+++ /dev/null
@@ -1,14 +0,0 @@
-
-\newpage
-\section*{Pre-typing}
-
-\ocwsection \label{pretyping}
-
-\bigskip
-\begin{center}\epsfig{file=pretyping.dep.ps,width=\linewidth}\end{center}
-
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: t
-%%% End:
diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml
index 9fd55a48..d37090a6 100644
--- a/pretyping/evarconv.ml
+++ b/pretyping/evarconv.ml
@@ -1,37 +1,41 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 CErrors
open Util
open Names
-open Term
+open Constr
+open Termops
+open Environ
+open EConstr
open Vars
open CClosure
open Reduction
open Reductionops
-open Termops
-open Environ
open Recordops
open Evarutil
open Evardefine
open Evarsolve
-open Globnames
open Evd
open Pretype_errors
-open Sigma.Notations
-open Context.Rel.Declaration
+open Context.Named.Declaration
+
+module RelDecl = Context.Rel.Declaration
+module NamedDecl = Context.Named.Declaration
type unify_fun = transparent_state ->
- env -> evar_map -> conv_pb -> constr -> constr -> Evarsolve.unification_result
+ env -> evar_map -> conv_pb -> EConstr.constr -> EConstr.constr -> Evarsolve.unification_result
let debug_unification = ref (false)
let _ = Goptions.declare_bool_option {
- Goptions.optsync = true; Goptions.optdepr = false;
+ Goptions.optdepr = false;
Goptions.optname =
"Print states sent to Evarconv unification";
Goptions.optkey = ["Debug";"Unification"];
@@ -39,6 +43,31 @@ let _ = Goptions.declare_bool_option {
Goptions.optwrite = (fun a -> debug_unification:=a);
}
+(*******************************************)
+(* Functions to deal with impossible cases *)
+(*******************************************)
+(* XXX: we would like to search for this with late binding
+ "data.id.type" etc... *)
+let impossible_default_case () =
+ let c, ctx = Universes.fresh_global_instance (Global.env()) (Globnames.ConstRef Coqlib.id) in
+ let (_, u) = Constr.destConst c in
+ Some (c, Constr.mkConstU (Coqlib.type_of_id, u), ctx)
+
+let coq_unit_judge =
+ let open Environ in
+ let make_judge c t = make_judge (EConstr.of_constr c) (EConstr.of_constr t) in
+ let na1 = Name (Id.of_string "A") in
+ let na2 = Name (Id.of_string "H") in
+ fun () ->
+ match impossible_default_case () with
+ | Some (id, type_of_id, ctx) ->
+ make_judge id type_of_id, ctx
+ | None ->
+ (* In case the constants id/ID are not defined *)
+ Environ.make_judge (mkLambda (na1,mkProp,mkLambda(na2,mkRel 1,mkRel 1)))
+ (mkProd (na1,mkProp,mkArrow (mkRel 1) (mkRel 2))),
+ Univ.ContextSet.empty
+
let unfold_projection env evd ts p c =
let cst = Projection.constant p in
if is_transparent_constant ts cst then
@@ -46,21 +75,20 @@ let unfold_projection env evd ts p c =
else None
let eval_flexible_term ts env evd c =
- match kind_of_term c with
- | Const (c,u as cu) ->
+ match EConstr.kind evd c with
+ | Const (c, u) ->
if is_transparent_constant ts c
- then constant_opt_value_in env cu
+ then Option.map EConstr.of_constr (constant_opt_value_in env (c, EInstance.kind evd u))
else None
| Rel n ->
(try match lookup_rel n env with
- | LocalAssum _ -> None
- | LocalDef (_,v,_) -> Some (lift n v)
+ | RelDecl.LocalAssum _ -> None
+ | RelDecl.LocalDef (_,v,_) -> Some (lift n v)
with Not_found -> None)
| Var id ->
(try
if is_transparent_variable ts id then
- let open Context.Named.Declaration in
- lookup_named id env |> get_value
+ env |> lookup_named id |> NamedDecl.get_value
else None
with Not_found -> None)
| LetIn (_,b,_,c) -> Some (subst1 b c)
@@ -72,11 +100,11 @@ let eval_flexible_term ts env evd c =
type flex_kind_of_term =
| Rigid
- | MaybeFlexible of Constr.t (* reducible but not necessarily reduced *)
- | Flexible of existential
+ | MaybeFlexible of EConstr.t (* reducible but not necessarily reduced *)
+ | Flexible of EConstr.existential
let flex_kind_of_term ts env evd c sk =
- match kind_of_term c with
+ match EConstr.kind evd c with
| LetIn _ | Rel _ | Const _ | Var _ | Proj _ ->
Option.cata (fun x -> MaybeFlexible x) Rigid (eval_flexible_term ts env evd c)
| Lambda _ when not (Option.is_empty (Stack.decomp sk)) -> MaybeFlexible c
@@ -86,10 +114,13 @@ let flex_kind_of_term ts env evd c sk =
| Fix _ -> Rigid (* happens when the fixpoint is partially applied *)
| Cast _ | App _ | Case _ -> assert false
+let add_conv_pb (pb, env, x, y) sigma =
+ Evd.add_conv_pb (pb, env, EConstr.Unsafe.to_constr x, EConstr.Unsafe.to_constr y) sigma
+
let apprec_nohdbeta ts env evd c =
let (t,sk as appr) = Reductionops.whd_nored_state evd (c, []) in
if Stack.not_purely_applicative sk
- then Stack.zip (fst (whd_betaiota_deltazeta_for_iota_state
+ then Stack.zip evd (fst (whd_betaiota_deltazeta_for_iota_state
ts env evd Cst_stack.empty appr))
else c
@@ -99,7 +130,7 @@ let position_problem l2r = function
let occur_rigidly (evk,_ as ev) evd t =
let rec aux t =
- match kind_of_term (whd_evar evd t) with
+ match EConstr.kind evd t with
| App (f, c) -> if aux f then Array.exists aux c else false
| Construct _ | Ind _ | Sort _ | Meta _ | Fix _ | CoFix _ -> true
| Proj (p, c) -> not (aux c)
@@ -109,7 +140,7 @@ let occur_rigidly (evk,_ as ev) evd t =
| Const _ -> false
| Prod (_, b, t) -> ignore(aux b || aux t); true
| Rel _ | Var _ -> false
- | Case (_,_,c,_) -> if eq_constr (mkEvar ev) c then raise Occur else false
+ | Case (_,_,c,_) -> if eq_constr evd (mkEvar ev) c then raise Occur else false
in try ignore(aux t); false with Occur -> true
(* [check_conv_record env sigma (t1,stack1) (t2,stack2)] tries to decompose
@@ -133,20 +164,28 @@ let occur_rigidly (evk,_ as ev) evd t =
projection would have been reduced) *)
let check_conv_record env sigma (t1,sk1) (t2,sk2) =
- let (proji, u), arg = Universes.global_app_of_constr t1 in
+ let (proji, u), arg = Termops.global_app_of_constr sigma t1 in
let canon_s,sk2_effective =
try
- match kind_of_term t2 with
+ match EConstr.kind sigma t2 with
Prod (_,a,b) -> (* assert (l2=[]); *)
- let _, a, b = destProd (Evarutil.nf_evar sigma t2) in
- if dependent (mkRel 1) b then raise Not_found
- else lookup_canonical_conversion (proji, Prod_cs),
+ let _, a, b = destProd sigma t2 in
+ if noccurn sigma 1 b then
+ lookup_canonical_conversion (proji, Prod_cs),
(Stack.append_app [|a;pop b|] Stack.empty)
+ else raise Not_found
| Sort s ->
+ let s = ESorts.kind sigma s in
lookup_canonical_conversion
- (proji, Sort_cs (family_of_sort s)),[]
+ (proji, Sort_cs (Sorts.family s)),[]
+ | Proj (p, c) ->
+ let c2 = Globnames.ConstRef (Projection.constant p) in
+ let c = Retyping.expand_projection env sigma p c [] in
+ let _, args = destApp sigma c in
+ let sk2 = Stack.append_app args sk2 in
+ lookup_canonical_conversion (proji, Const_cs c2), sk2
| _ ->
- let c2 = global_of_constr t2 in
+ let (c2, _) = Termops.global_of_constr sigma t2 in
lookup_canonical_conversion (proji, Const_cs c2),sk2
with Not_found ->
let (c, cs) = lookup_canonical_conversion (proji,Default_cs) in
@@ -154,12 +193,14 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) =
in
let t', { o_DEF = c; o_CTX = ctx; o_INJ=n; o_TABS = bs;
o_TPARAMS = params; o_NPARAMS = nparams; o_TCOMPS = us } = canon_s in
+ let us = List.map EConstr.of_constr us in
+ let params = List.map EConstr.of_constr params in
let params1, c1, extra_args1 =
match arg with
| Some c -> (* A primitive projection applied to c *)
let ty = Retyping.get_type_of ~lax:true env sigma c in
let (i,u), ind_args =
- try Inductiveops.find_mrectype env sigma ty
+ try Inductiveops.find_mrectype env sigma ty
with _ -> raise Not_found
in Stack.append_app_list ind_args Stack.empty, c, sk1
| None ->
@@ -172,14 +213,19 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) =
else match (Stack.strip_n_app (l_us-1) sk2_effective) with
| None -> raise Not_found
| Some (l',el,s') -> (l'@Stack.append_app [|el|] Stack.empty,s') in
- let subst, ctx' = Universes.fresh_universe_context_set_instance ctx in
+ let u, ctx' = Universes.fresh_instance_from ctx None in
+ let subst = Univ.make_inverse_instance_subst u in
+ let c = EConstr.of_constr c in
let c' = subst_univs_level_constr subst c in
+ let t' = EConstr.of_constr t' in
let t' = subst_univs_level_constr subst t' in
- let bs' = List.map (subst_univs_level_constr subst) bs in
- let h, _ = decompose_app_vect t' in
+ let bs' = List.map (EConstr.of_constr %> subst_univs_level_constr subst) bs in
+ let params = List.map (fun c -> subst_univs_level_constr subst c) params in
+ let us = List.map (fun c -> subst_univs_level_constr subst c) us in
+ let h, _ = decompose_app_vect sigma t' in
ctx',(h, t2),c',bs',(Stack.append_app_list params Stack.empty,params1),
(Stack.append_app_list us Stack.empty,us2),(extra_args1,extra_args2),c1,
- (n,Stack.zip(t2,sk2))
+ (n, Stack.zip sigma (t2,sk2))
(* Precondition: one of the terms of the pb is an uninstantiated evar,
* possibly applied to arguments. *)
@@ -250,7 +296,7 @@ let ise_stack2 no_app env evd f sk1 sk2 =
| UnifFailure _ as x -> fail x)
| UnifFailure _ as x -> fail x)
| Stack.Proj (n1,a1,p1,_)::q1, Stack.Proj (n2,a2,p2,_)::q2 ->
- if eq_constant (Projection.constant p1) (Projection.constant p2)
+ if Constant.equal (Projection.constant p1) (Projection.constant p2)
then ise_stack2 true i q1 q2
else fail (UnifFailure (i, NotSameHead))
| Stack.Fix (((li1, i1),(_,tys1,bds1 as recdef1)),a1,_)::q1,
@@ -263,8 +309,6 @@ let ise_stack2 no_app env evd f sk1 sk2 =
| Success i' -> ise_stack2 true i' q1 q2
| UnifFailure _ as x -> fail x
else fail (UnifFailure (i,NotSameHead))
- | Stack.Update _ :: _, _ | Stack.Shift _ :: _, _
- | _, Stack.Update _ :: _ | _, Stack.Shift _ :: _ -> assert false
| Stack.App _ :: _, Stack.App _ :: _ ->
if no_app && deep then fail ((*dummy*)UnifFailure(i,NotSameHead)) else
begin match ise_app_stack2 env f i sk1 sk2 with
@@ -294,11 +338,9 @@ let exact_ise_stack2 env evd f sk1 sk2 =
(fun i -> ise_stack2 i a1 a2)]
else UnifFailure (i,NotSameHead)
| Stack.Proj (n1,a1,p1,_)::q1, Stack.Proj (n2,a2,p2,_)::q2 ->
- if eq_constant (Projection.constant p1) (Projection.constant p2)
+ if Constant.equal (Projection.constant p1) (Projection.constant p2)
then ise_stack2 i q1 q2
else (UnifFailure (i, NotSameHead))
- | Stack.Update _ :: _, _ | Stack.Shift _ :: _, _
- | _, Stack.Update _ :: _ | _, Stack.Shift _ :: _ -> assert false
| Stack.App _ :: _, Stack.App _ :: _ ->
begin match ise_app_stack2 env f i sk1 sk2 with
|_,(UnifFailure _ as x) -> x
@@ -310,6 +352,14 @@ let exact_ise_stack2 env evd f sk1 sk2 =
ise_stack2 evd (List.rev sk1) (List.rev sk2)
else UnifFailure (evd, (* Dummy *) NotSameHead)
+(* Add equality constraints for covariant/invariant positions. For
+ irrelevant positions, unify universes when flexible. *)
+let compare_cumulative_instances evd variances u u' =
+ match Evarutil.compare_cumulative_instances CONV variances u u' evd with
+ | Inl evd ->
+ Success evd
+ | Inr p -> UnifFailure (evd, UnifUnivInconsistency p)
+
let rec evar_conv_x ts env evd pbty term1 term2 =
let term1 = whd_head_evar evd term1 in
let term2 = whd_head_evar evd term2 in
@@ -321,7 +371,7 @@ let rec evar_conv_x ts env evd pbty term1 term2 =
let e =
try
let evd, b = infer_conv ~catch_incon:false ~pb:pbty ~ts:(fst ts)
- env evd term1 term2
+ env evd term1 term2
in
if b then Success evd
else UnifFailure (evd, ConversionFailed (env,term1,term2))
@@ -335,7 +385,7 @@ let rec evar_conv_x ts env evd pbty term1 term2 =
match ground_test with
| Some result -> result
| None ->
- (* Until pattern-unification is used consistently, use nohdbeta to not
+ (* Until pattern-unification is used consistently, use nohdbeta to not
destroy beta-redexes that can be used for 1st-order unification *)
let term1 = apprec_nohdbeta (fst ts) env evd term1 in
let term2 = apprec_nohdbeta (fst ts) env evd term2 in
@@ -344,16 +394,16 @@ let rec evar_conv_x ts env evd pbty term1 term2 =
(whd_nored_state evd (term1,Stack.empty), Cst_stack.empty)
(whd_nored_state evd (term2,Stack.empty), Cst_stack.empty)
in
- begin match kind_of_term term1, kind_of_term term2 with
+ begin match EConstr.kind evd term1, EConstr.kind evd term2 with
| Evar ev, _ when Evd.is_undefined evd (fst ev) ->
(match solve_simple_eqn (evar_conv_x ts) env evd
- (position_problem true pbty,ev,term2) with
+ (position_problem true pbty,ev, term2) with
| UnifFailure (_,OccurCheck _) ->
(* Eta-expansion might apply *) default ()
| x -> x)
| _, Evar ev when Evd.is_undefined evd (fst ev) ->
(match solve_simple_eqn (evar_conv_x ts) env evd
- (position_problem false pbty,ev,term1) with
+ (position_problem false pbty,ev, term1) with
| UnifFailure (_, OccurCheck _) ->
(* Eta-expansion might apply *) default ()
| x -> x)
@@ -369,8 +419,8 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
match is_unification_pattern_evar env evd ev lF tM with
| None -> fallback ()
| Some l1' -> (* Miller-Pfenning's patterns unification *)
- let t2 = nf_evar evd tM in
- let t2 = solve_pattern_eqn env l1' t2 in
+ let t2 = tM in
+ let t2 = solve_pattern_eqn env evd l1' t2 in
solve_simple_eqn (evar_conv_x ts) env evd
(position_problem on_left pbty,ev,t2)
in
@@ -379,34 +429,82 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
let not_only_app = Stack.not_purely_applicative skO in
match switch (ise_stack2 not_only_app env evd (evar_conv_x ts)) skF skO with
|Some (l,r), Success i' when on_left && (not_only_app || List.is_empty l) ->
- switch (evar_conv_x ts env i' pbty) (Stack.zip(termF,l)) (Stack.zip(termO,r))
+ switch (evar_conv_x ts env i' pbty) (Stack.zip evd (termF,l)) (Stack.zip evd (termO,r))
|Some (r,l), Success i' when not on_left && (not_only_app || List.is_empty l) ->
- switch (evar_conv_x ts env i' pbty) (Stack.zip(termF,l)) (Stack.zip(termO,r))
+ switch (evar_conv_x ts env i' pbty) (Stack.zip evd (termF,l)) (Stack.zip evd (termO,r))
|None, Success i' -> switch (evar_conv_x ts env i' pbty) termF termO
|_, (UnifFailure _ as x) -> x
|Some _, _ -> UnifFailure (evd,NotSameArgSize) in
let eta env evd onleft sk term sk' term' =
assert (match sk with [] -> true | _ -> false);
- let (na,c1,c'1) = destLambda term in
+ let (na,c1,c'1) = destLambda evd term in
let c = nf_evar evd c1 in
- let env' = push_rel (LocalAssum (na,c)) env in
+ let env' = push_rel (RelDecl.LocalAssum (na,c)) env in
let out1 = whd_betaiota_deltazeta_for_iota_state
(fst ts) env' evd Cst_stack.empty (c'1, Stack.empty) in
let out2 = whd_nored_state evd
- (Stack.zip (term', sk' @ [Stack.Shift 1]), Stack.append_app [|mkRel 1|] Stack.empty),
+ (lift 1 (Stack.zip evd (term', sk')), Stack.append_app [|EConstr.mkRel 1|] Stack.empty),
Cst_stack.empty in
if onleft then evar_eqappr_x ts env' evd CONV out1 out2
else evar_eqappr_x ts env' evd CONV out2 out1
in
let rigids env evd sk term sk' term' =
- let b,univs = Universes.eq_constr_universes term term' in
- if b then
- ise_and evd [(fun i ->
- let cstrs = Universes.to_constraints (Evd.universes i) univs in
- try Success (Evd.add_constraints i cstrs)
- with Univ.UniverseInconsistency p -> UnifFailure (i, UnifUnivInconsistency p));
- (fun i -> exact_ise_stack2 env i (evar_conv_x ts) sk sk')]
- else UnifFailure (evd,NotSameHead)
+ let check_strict evd u u' =
+ let cstrs = Univ.enforce_eq_instances u u' Univ.Constraint.empty in
+ try Success (Evd.add_constraints evd cstrs)
+ with Univ.UniverseInconsistency p -> UnifFailure (evd, UnifUnivInconsistency p)
+ in
+ let compare_heads evd =
+ match EConstr.kind evd term, EConstr.kind evd term' with
+ | Const (c, u), Const (c', u') when Constant.equal c c' ->
+ let u = EInstance.kind evd u and u' = EInstance.kind evd u' in
+ check_strict evd u u'
+ | Const _, Const _ -> UnifFailure (evd, NotSameHead)
+ | Ind ((mi,i) as ind , u), Ind (ind', u') when Names.eq_ind ind ind' ->
+ if EInstance.is_empty u && EInstance.is_empty u' then Success evd
+ else
+ let u = EInstance.kind evd u and u' = EInstance.kind evd u' in
+ let mind = Environ.lookup_mind mi env in
+ let open Declarations in
+ begin match mind.mind_universes with
+ | Monomorphic_ind _ -> assert false
+ | Polymorphic_ind _ -> check_strict evd u u'
+ | Cumulative_ind cumi ->
+ let nparamsaplied = Stack.args_size sk in
+ let nparamsaplied' = Stack.args_size sk' in
+ let needed = Reduction.inductive_cumulativity_arguments (mind,i) in
+ if not (Int.equal nparamsaplied needed && Int.equal nparamsaplied' needed)
+ then check_strict evd u u'
+ else
+ compare_cumulative_instances evd (Univ.ACumulativityInfo.variance cumi) u u'
+ end
+ | Ind _, Ind _ -> UnifFailure (evd, NotSameHead)
+ | Construct (((mi,ind),ctor as cons), u), Construct (cons', u')
+ when Names.eq_constructor cons cons' ->
+ if EInstance.is_empty u && EInstance.is_empty u' then Success evd
+ else
+ let u = EInstance.kind evd u and u' = EInstance.kind evd u' in
+ let mind = Environ.lookup_mind mi env in
+ let open Declarations in
+ begin match mind.mind_universes with
+ | Monomorphic_ind _ -> assert false
+ | Polymorphic_ind _ -> check_strict evd u u'
+ | Cumulative_ind cumi ->
+ let nparamsaplied = Stack.args_size sk in
+ let nparamsaplied' = Stack.args_size sk' in
+ let needed = Reduction.constructor_cumulativity_arguments (mind,ind,ctor) in
+ if not (Int.equal nparamsaplied needed && Int.equal nparamsaplied' needed)
+ then check_strict evd u u'
+ else
+ Success (compare_constructor_instances evd u u')
+ end
+ | Construct _, Construct _ -> UnifFailure (evd, NotSameHead)
+ | _, _ -> anomaly (Pp.str "")
+ in
+ ise_and evd [(fun i ->
+ try compare_heads i
+ with Univ.UniverseInconsistency p -> UnifFailure (i, UnifUnivInconsistency p));
+ (fun i -> exact_ise_stack2 env i (evar_conv_x ts) sk sk')]
in
let flex_maybeflex on_left ev ((termF,skF as apprF),cstsF) ((termM, skM as apprM),cstsM) vM =
let switch f a b = if on_left then f a b else f b a in
@@ -415,10 +513,10 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
match Stack.list_of_app_stack skF with
| None -> quick_fail evd
| Some lF ->
- let tM = Stack.zip apprM in
+ let tM = Stack.zip evd apprM in
miller_pfenning on_left
(fun () -> if not_only_app then (* Postpone the use of an heuristic *)
- switch (fun x y -> Success (add_conv_pb (pbty,env,x,y) i)) (Stack.zip apprF) tM
+ switch (fun x y -> Success (add_conv_pb (pbty,env,x,y) i)) (Stack.zip evd apprF) tM
else quick_fail i)
ev lF tM i
and consume (termF,skF as apprF) (termM,skM as apprM) i =
@@ -432,7 +530,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
in
let default i = ise_try i [f1; consume apprF apprM; delta]
in
- match kind_of_term termM with
+ match EConstr.kind evd termM with
| Proj (p, c) when not (Stack.is_empty skF) ->
(* Might be ?X args = p.c args', and we have to eta-expand the
primitive projection if |args| >= |args'|+1. *)
@@ -462,7 +560,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
let flex_rigid on_left ev (termF, skF as apprF) (termR, skR as apprR) =
let switch f a b = if on_left then f a b else f b a in
let eta evd =
- match kind_of_term termR with
+ match EConstr.kind evd termR with
| Lambda _ when (* if ever problem is ill-typed: *) List.is_empty skR ->
eta env evd false skR termR skF termF
| Construct u -> eta_constructor ts env evd skR u skF termF
@@ -472,7 +570,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
| None ->
ise_try evd [consume_stack on_left apprF apprR; eta]
| Some lF ->
- let tR = Stack.zip apprR in
+ let tR = Stack.zip evd apprR in
miller_pfenning on_left
(fun () ->
ise_try evd
@@ -480,12 +578,12 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
(fun i ->
if not (occur_rigidly ev i tR) then
let i,tF =
- if isRel tR || isVar tR then
+ if isRel i tR || isVar i tR then
(* Optimization so as to generate candidates *)
let i,ev = evar_absorb_arguments env i ev lF in
i,mkEvar ev
else
- i,Stack.zip apprF in
+ i,Stack.zip evd apprF in
switch (fun x y -> Success (add_conv_pb (pbty,env,x,y) i))
tF tR
else
@@ -496,8 +594,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
(* Evar must be undefined since we have flushed evars *)
let () = if !debug_unification then
let open Pp in
- Feedback.msg_notice (v 0 (pr_state appr1 ++ cut () ++ pr_state appr2 ++ cut ())
- ++ fnl ()) in
+ Feedback.msg_notice (v 0 (pr_state appr1 ++ cut () ++ pr_state appr2 ++ cut ())) in
match (flex_kind_of_term (fst ts) env evd term1 sk1,
flex_kind_of_term (fst ts) env evd term2 sk2) with
| Flexible (sp1,al1 as ev1), Flexible (sp2,al2 as ev2) ->
@@ -509,9 +606,9 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
(* We do have sk1[] = sk2[]: we now unify ?ev1 and ?ev2 *)
(* Note that ?ev1 and ?ev2, may have been instantiated in the meantime *)
let ev1' = whd_evar i' (mkEvar ev1) in
- if isEvar ev1' then
+ if isEvar i' ev1' then
solve_simple_eqn (evar_conv_x ts) env i'
- (position_problem true pbty,destEvar ev1',term2)
+ (position_problem true pbty,destEvar i' ev1', term2)
else
evar_eqappr_x ts env evd pbty
((ev1', sk1), csts1) ((term2, sk2), csts2)
@@ -519,9 +616,9 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
(* We have sk1'[] = sk2[] for some sk1' s.t. sk1[]=sk1'[r[]] *)
(* we now unify r[?ev1] and ?ev2 *)
let ev2' = whd_evar i' (mkEvar ev2) in
- if isEvar ev2' then
+ if isEvar i' ev2' then
solve_simple_eqn (evar_conv_x ts) env i'
- (position_problem false pbty,destEvar ev2',Stack.zip(term1,r))
+ (position_problem false pbty,destEvar i' ev2',Stack.zip evd (term1,r))
else
evar_eqappr_x ts env evd pbty
((ev2', sk1), csts1) ((term2, sk2), csts2)
@@ -530,9 +627,9 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
(* We have sk1[] = sk2'[] for some sk2' s.t. sk2[]=sk2'[r[]] *)
(* we now unify ?ev1 and r[?ev2] *)
let ev1' = whd_evar i' (mkEvar ev1) in
- if isEvar ev1' then
+ if isEvar i' ev1' then
solve_simple_eqn (evar_conv_x ts) env i'
- (position_problem true pbty,destEvar ev1',Stack.zip(term2,r))
+ (position_problem true pbty,destEvar i' ev1',Stack.zip evd (term2,r))
else evar_eqappr_x ts env evd pbty
((ev1', sk1), csts1) ((term2, sk2), csts2)
| None, (UnifFailure _ as x) ->
@@ -585,7 +682,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
flex_maybeflex false ev2 (appr2,csts2) (appr1,csts1) v1
| MaybeFlexible v1, MaybeFlexible v2 -> begin
- match kind_of_term term1, kind_of_term term2 with
+ match EConstr.kind evd term1, EConstr.kind evd term2 with
| LetIn (na1,b1,t1,c'1), LetIn (na2,b2,t2,c'2) ->
let f1 i = (* FO *)
ise_and i
@@ -596,8 +693,8 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
(fun i ->
let b = nf_evar i b1 in
let t = nf_evar i t1 in
- let na = Nameops.name_max na1 na2 in
- evar_conv_x ts (push_rel (LocalDef (na,b,t)) env) i pbty c'1 c'2);
+ let na = Nameops.Name.pick na1 na2 in
+ evar_conv_x ts (push_rel (RelDecl.LocalDef (na,b,t)) env) i pbty c'1 c'2);
(fun i -> exact_ise_stack2 env i (evar_conv_x ts) sk1 sk2)]
and f2 i =
let out1 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts1 (v1,sk1)
@@ -620,9 +717,9 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
ise_try evd [f1; f2]
(* Catch the p.c ~= p c' cases *)
- | Proj (p,c), Const (p',u) when eq_constant (Projection.constant p) p' ->
+ | Proj (p,c), Const (p',u) when Constant.equal (Projection.constant p) p' ->
let res =
- try Some (destApp (Retyping.expand_projection env evd p c []))
+ try Some (destApp evd (Retyping.expand_projection env evd p c []))
with Retyping.RetypeError _ -> None
in
(match res with
@@ -631,9 +728,9 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
(appr2,csts2)
| None -> UnifFailure (evd,NotSameHead))
- | Const (p,u), Proj (p',c') when eq_constant p (Projection.constant p') ->
+ | Const (p,u), Proj (p',c') when Constant.equal p (Projection.constant p') ->
let res =
- try Some (destApp (Retyping.expand_projection env evd p' c' []))
+ try Some (destApp evd (Retyping.expand_projection env evd p' c' []))
with Retyping.RetypeError _ -> None
in
(match res with
@@ -648,14 +745,16 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
allow this identification (first-order unification of universes). Otherwise
fallback to unfolding.
*)
- let b,univs = Universes.eq_constr_universes term1 term2 in
- if b then
+ let univs = EConstr.eq_constr_universes env evd term1 term2 in
+ match univs with
+ | Some univs ->
ise_and i [(fun i ->
try Success (Evd.add_universe_constraints i univs)
with UniversesDiffer -> UnifFailure (i,NotSameHead)
| Univ.UniverseInconsistency p -> UnifFailure (i, UnifUnivInconsistency p));
(fun i -> exact_ise_stack2 env i (evar_conv_x ts) sk1 sk2)]
- else UnifFailure (i,NotSameHead)
+ | None ->
+ UnifFailure (i,NotSameHead)
and f2 i =
(try
if not (snd ts) then raise Not_found
@@ -668,7 +767,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
if the first argument is a beta-redex (expand a constant
only if necessary) or the second argument is potentially
usable as a canonical projection or canonical value *)
- let rec is_unnamed (hd, args) = match kind_of_term hd with
+ let rec is_unnamed (hd, args) = match EConstr.kind i hd with
| (Var _|Construct _|Ind _|Const _|Prod _|Sort _) ->
Stack.not_purely_applicative args
| (CoFix _|Meta _|Rel _)-> true
@@ -689,7 +788,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
let rhs_is_already_stuck =
rhs_is_already_stuck || rhs_is_stuck_and_unnamed () in
- if (isLambda term1 || rhs_is_already_stuck)
+ if (EConstr.isLambda i term1 || rhs_is_already_stuck)
&& (not (Stack.not_purely_applicative sk1)) then
evar_eqappr_x ~rhs_is_already_stuck ts env i pbty
(whd_betaiota_deltazeta_for_iota_state
@@ -703,16 +802,16 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
ise_try evd [f1; f2; f3]
end
- | Rigid, Rigid when isLambda term1 && isLambda term2 ->
- let (na1,c1,c'1) = destLambda term1 in
- let (na2,c2,c'2) = destLambda term2 in
+ | Rigid, Rigid when EConstr.isLambda evd term1 && EConstr.isLambda evd term2 ->
+ let (na1,c1,c'1) = EConstr.destLambda evd term1 in
+ let (na2,c2,c'2) = EConstr.destLambda evd term2 in
assert app_empty;
ise_and evd
[(fun i -> evar_conv_x ts env i CONV c1 c2);
(fun i ->
let c = nf_evar i c1 in
- let na = Nameops.name_max na1 na2 in
- evar_conv_x ts (push_rel (LocalAssum (na,c)) env) i CONV c'1 c'2)]
+ let na = Nameops.Name.pick na1 na2 in
+ evar_conv_x ts (push_rel (RelDecl.LocalAssum (na,c)) env) i CONV c'1 c'2)]
| Flexible ev1, Rigid -> flex_rigid true ev1 appr1 appr2
| Rigid, Flexible ev2 -> flex_rigid false ev2 appr2 appr1
@@ -745,17 +844,19 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
ise_try evd [f3; f4]
(* Eta-expansion *)
- | Rigid, _ when isLambda term1 && (* if ever ill-typed: *) List.is_empty sk1 ->
+ | Rigid, _ when isLambda evd term1 && (* if ever ill-typed: *) List.is_empty sk1 ->
eta env evd true sk1 term1 sk2 term2
- | _, Rigid when isLambda term2 && (* if ever ill-typed: *) List.is_empty sk2 ->
+ | _, Rigid when isLambda evd term2 && (* if ever ill-typed: *) List.is_empty sk2 ->
eta env evd false sk2 term2 sk1 term1
| Rigid, Rigid -> begin
- match kind_of_term term1, kind_of_term term2 with
+ match EConstr.kind evd term1, EConstr.kind evd term2 with
| Sort s1, Sort s2 when app_empty ->
(try
+ let s1 = ESorts.kind evd s1 in
+ let s2 = ESorts.kind evd s2 in
let evd' =
if pbty == CONV
then Evd.set_eq_sort env evd s1 s2
@@ -770,8 +871,8 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
[(fun i -> evar_conv_x ts env i CONV c1 c2);
(fun i ->
let c = nf_evar i c1 in
- let na = Nameops.name_max n1 n2 in
- evar_conv_x ts (push_rel (LocalAssum (na,c)) env) i pbty c'1 c'2)]
+ let na = Nameops.Name.pick n1 n2 in
+ evar_conv_x ts (push_rel (RelDecl.LocalAssum (na,c)) env) i pbty c'1 c'2)]
| Rel x1, Rel x2 ->
if Int.equal x1 x2 then
@@ -818,7 +919,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
begin match ise_stack2 true env evd (evar_conv_x ts) sk1 sk2 with
|_, (UnifFailure _ as x) -> x
|None, Success i' -> evar_conv_x ts env i' CONV term1 term2
- |Some (sk1',sk2'), Success i' -> evar_conv_x ts env i' CONV (Stack.zip (term1,sk1')) (Stack.zip (term2,sk2'))
+ |Some (sk1',sk2'), Success i' -> evar_conv_x ts env i' CONV (Stack.zip i' (term1,sk1')) (Stack.zip i' (term2,sk2'))
end
| (Ind _ | Sort _ | Prod _ | CoFix _ | Fix _ | Rel _ | Var _ | Const _), _ ->
@@ -866,10 +967,8 @@ and conv_record trs env evd (ctx,(h,h2),c,bs,(params,params1),(us,us2),(sk1,sk2)
let test i = evar_conv_x trs env i CUMUL ty (substl ks b) in
(i,t2::ks, m-1, test)
else
- let dloc = (Loc.ghost,Evar_kinds.InternalHole) in
- let i = Sigma.Unsafe.of_evar_map i in
- let Sigma (ev, i', _) = Evarutil.new_evar env i ~src:dloc (substl ks b) in
- let i' = Sigma.to_evar_map i' in
+ let dloc = Loc.tag Evar_kinds.InternalHole in
+ let (i', ev) = Evarutil.new_evar env i ~src:dloc (substl ks b) in
(i', ev :: ks, m - 1,test))
(evd,[],List.length bs,fun i -> Success i) bs
in
@@ -887,19 +986,19 @@ and conv_record trs env evd (ctx,(h,h2),c,bs,(params,params1),(us,us2),(sk1,sk2)
(fun i -> exact_ise_stack2 env i (evar_conv_x trs) sk1 sk2);
test;
(fun i -> evar_conv_x trs env i CONV h2
- (fst (decompose_app_vect (substl ks h))))]
+ (fst (decompose_app_vect i (substl ks h))))]
else UnifFailure(evd,(*dummy*)NotSameHead)
and eta_constructor ts env evd sk1 ((ind, i), u) sk2 term2 =
let mib = lookup_mind (fst ind) env in
match mib.Declarations.mind_record with
- | Some (Some (id, projs, pbs)) when mib.Declarations.mind_finite == Decl_kinds.BiFinite ->
+ | Some (Some (id, projs, pbs)) when mib.Declarations.mind_finite == Declarations.BiFinite ->
let pars = mib.Declarations.mind_nparams in
(try
let l1' = Stack.tail pars sk1 in
let l2' =
- let term = Stack.zip (term2,sk2) in
- List.map (fun p -> mkProj (Projection.make p false, term)) (Array.to_list projs)
+ let term = Stack.zip evd (term2,sk2) in
+ List.map (fun p -> EConstr.mkProj (Projection.make p false, term)) (Array.to_list projs)
in
exact_ise_stack2 env evd (evar_conv_x (fst ts, false)) l1'
(Stack.append_app_list l2' Stack.empty)
@@ -914,8 +1013,8 @@ let evar_conv_x ts = evar_conv_x (ts, true)
(* Profiling *)
let evar_conv_x =
if Flags.profile then
- let evar_conv_xkey = Profile.declare_profile "evar_conv_x" in
- Profile.profile6 evar_conv_xkey evar_conv_x
+ let evar_conv_xkey = CProfile.declare_profile "evar_conv_x" in
+ CProfile.profile6 evar_conv_xkey evar_conv_x
else evar_conv_x
let evar_conv_hook_get, evar_conv_hook_set = Hook.make ~default:evar_conv_x ()
@@ -943,50 +1042,57 @@ let first_order_unification ts env evd (ev1,l1) (term2,l2) =
let choose_less_dependent_instance evk evd term args =
let evi = Evd.find_undefined evd evk in
let subst = make_pure_subst evi args in
- let subst' = List.filter (fun (id,c) -> Term.eq_constr c term) subst in
+ let subst' = List.filter (fun (id,c) -> EConstr.eq_constr evd c term) subst in
match subst' with
| [] -> None
- | (id, _) :: _ -> Some (Evd.define evk (mkVar id) evd)
+ | (id, _) :: _ -> Some (Evd.define evk (Constr.mkVar id) evd)
-open Context.Named.Declaration
let apply_on_subterm env evdref f c t =
let rec applyrec (env,(k,c) as acc) t =
(* By using eq_constr, we make an approximation, for instance, we *)
(* could also be interested in finding a term u convertible to t *)
(* such that c occurs in u *)
- if e_eq_constr_univs evdref c t then f k
+ let eq_constr c1 c2 = match EConstr.eq_constr_universes env !evdref c1 c2 with
+ | None -> false
+ | Some cstr ->
+ try ignore (Evd.add_universe_constraints !evdref cstr); true
+ with UniversesDiffer -> false
+ in
+ if eq_constr c t then f k
else
- match kind_of_term t with
- | Evar (evk,args) when Evd.is_undefined !evdref evk ->
+ match EConstr.kind !evdref t with
+ | Evar (evk,args) ->
let ctx = evar_filtered_context (Evd.find_undefined !evdref evk) in
let g decl a = if is_local_assum decl then applyrec acc a else a in
mkEvar (evk, Array.of_list (List.map2 g ctx (Array.to_list args)))
| _ ->
- map_constr_with_binders_left_to_right
+ map_constr_with_binders_left_to_right !evdref
(fun d (env,(k,c)) -> (push_rel d env, (k+1,lift 1 c)))
applyrec acc t
in
applyrec (env,(0,c)) t
-let filter_possible_projections c ty ctxt args =
+let filter_possible_projections evd c ty ctxt args =
(* Since args in the types will be replaced by holes, we count the
fv of args to have a well-typed filter; don't know how necessary
- it is however to have a well-typed filter here *)
- let fv1 = free_rels (mkApp (c,args)) (* Hack: locally untyped *) in
- let fv2 = collect_vars (mkApp (c,args)) in
+ it is however to have a well-typed filter here *)
+ let fv1 = free_rels evd (mkApp (c,args)) (* Hack: locally untyped *) in
+ let fv2 = collect_vars evd (mkApp (c,args)) in
let len = Array.length args in
- let tyvars = collect_vars ty in
+ let tyvars = collect_vars evd ty in
List.map_i (fun i decl ->
let () = assert (i < len) in
let a = Array.unsafe_get args i in
- (match decl with LocalAssum _ -> false | LocalDef (_,c,_) -> not (isRel c || isVar c)) ||
+ (match decl with
+ | NamedDecl.LocalAssum _ -> false
+ | NamedDecl.LocalDef (_,c,_) -> not (isRel evd (EConstr.of_constr c) || isVar evd (EConstr.of_constr c))) ||
a == c ||
(* Here we make an approximation, for instance, we could also be *)
(* interested in finding a term u convertible to c such that a occurs *)
(* in u *)
- isRel a && Int.Set.mem (destRel a) fv1 ||
- isVar a && Id.Set.mem (destVar a) fv2 ||
- Id.Set.mem (get_id decl) tyvars)
+ isRel evd a && Int.Set.mem (destRel evd a) fv1 ||
+ isVar evd a && Id.Set.mem (destVar evd a) fv2 ||
+ Id.Set.mem (NamedDecl.get_id decl) tyvars)
0 ctxt
let solve_evars = ref (fun _ -> failwith "solve_evars not installed")
@@ -1017,39 +1123,39 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs =
let env_evar = evar_filtered_env evi in
let sign = named_context_val env_evar in
let ctxt = evar_filtered_context evi in
- let instance = List.map mkVar (List.map get_id ctxt) in
+ let instance = List.map mkVar (List.map NamedDecl.get_id ctxt) in
let rec make_subst = function
- | decl'::ctxt', c::l, occs::occsl when isVarId (get_id decl') c ->
+ | decl'::ctxt', c::l, occs::occsl when isVarId evd (NamedDecl.get_id decl') c ->
begin match occs with
| Some _ ->
- error "Cannot force abstraction on identity instance."
+ user_err Pp.(str "Cannot force abstraction on identity instance.")
| None ->
make_subst (ctxt',l,occsl)
end
| decl'::ctxt', c::l, occs::occsl ->
- let (id,_,t) = to_tuple decl' in
+ let id = NamedDecl.get_id decl' in
+ let t = EConstr.of_constr (NamedDecl.get_type decl') in
let evs = ref [] in
let ty = Retyping.get_type_of env_rhs evd c in
- let filter' = filter_possible_projections c ty ctxt args in
+ let filter' = filter_possible_projections evd c ty ctxt args in
(id,t,c,ty,evs,Filter.make filter',occs) :: make_subst (ctxt',l,occsl)
| _, _, [] -> []
- | _ -> anomaly (Pp.str "Signature or instance are shorter than the occurrences list") in
+ | _ -> anomaly (Pp.str "Signature or instance are shorter than the occurrences list.") in
let rec set_holes evdref rhs = function
| (id,_,c,cty,evsref,filter,occs)::subst ->
let set_var k =
match occs with
| Some Locus.AllOccurrences -> mkVar id
- | Some _ -> error "Selection of specific occurrences not supported"
+ | Some _ -> user_err Pp.(str "Selection of specific occurrences not supported")
| None ->
let evty = set_holes evdref cty subst in
let instance = Filter.filter_list filter instance in
- let evd = Sigma.Unsafe.of_evar_map !evdref in
- let Sigma (ev, evd, _) = new_evar_instance sign evd evty ~filter instance in
- let evd = Sigma.to_evar_map evd in
+ let evd = !evdref in
+ let (evd, ev) = new_evar_instance sign evd evty ~filter instance in
evdref := evd;
- evsref := (fst (destEvar ev),evty)::!evsref;
+ evsref := (fst (destEvar !evdref ev),evty)::!evsref;
ev in
set_holes evdref (apply_on_subterm env_rhs evdref set_var c rhs) subst
| [] -> rhs in
@@ -1077,12 +1183,12 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs =
(* We force abstraction over this unconstrained occurrence *)
(* and we use typing to propagate this instantiation *)
(* This is an arbitrary choice *)
- let evd = Evd.define evk (mkVar id) evd in
+ let evd = Evd.define evk (Constr.mkVar id) evd in
match evar_conv_x ts env_evar evd CUMUL idty evty with
- | UnifFailure _ -> error "Cannot find an instance"
+ | UnifFailure _ -> user_err Pp.(str "Cannot find an instance")
| Success evd ->
match reconsider_unif_constraints (evar_conv_x ts) evd with
- | UnifFailure _ -> error "Cannot find an instance"
+ | UnifFailure _ -> user_err Pp.(str "Cannot find an instance")
| Success evd ->
evd
else
@@ -1095,15 +1201,18 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs =
force_instantiation evd !evsref
| [] ->
let evd =
- try Evarsolve.check_evar_instance evd evk rhs
+ try Evarsolve.check_evar_instance evd evk rhs
(evar_conv_x full_transparent_state)
with IllTypedInstance _ -> raise (TypingFailed evd)
in
- Evd.define evk rhs evd
+ Evd.define evk (EConstr.Unsafe.to_constr rhs) evd
in
abstract_free_holes evd subst, true
with TypingFailed evd -> evd, false
+let to_pb (pb, env, t1, t2) =
+ (pb, env, EConstr.Unsafe.to_constr t1, EConstr.Unsafe.to_constr t2)
+
let second_order_matching_with_args ts env evd pbty ev l t =
(*
let evd,ev = evar_absorb_arguments env evd ev l in
@@ -1113,18 +1222,22 @@ let second_order_matching_with_args ts env evd pbty ev l t =
else UnifFailure (evd, ConversionFailed (env,mkApp(mkEvar ev,l),t))
if b then Success evd else
*)
- let pb = (pbty,env,mkApp(mkEvar ev,l),t) in
+ let pb = to_pb (pbty,env,mkApp(mkEvar ev,l),t) in
UnifFailure (evd, CannotSolveConstraint (pb,ProblemBeyondCapabilities))
let apply_conversion_problem_heuristic ts env evd pbty t1 t2 =
let t1 = apprec_nohdbeta ts env evd (whd_head_evar evd t1) in
let t2 = apprec_nohdbeta ts env evd (whd_head_evar evd t2) in
- let (term1,l1 as appr1) = try destApp t1 with DestKO -> (t1, [||]) in
- let (term2,l2 as appr2) = try destApp t2 with DestKO -> (t2, [||]) in
+ let (term1,l1 as appr1) = try destApp evd t1 with DestKO -> (t1, [||]) in
+ let (term2,l2 as appr2) = try destApp evd t2 with DestKO -> (t2, [||]) in
+ let () = if !debug_unification then
+ let open Pp in
+ Feedback.msg_notice (v 0 (str "Heuristic:" ++ spc () ++ print_constr t1
+ ++ cut () ++ print_constr t2 ++ cut ())) in
let app_empty = Array.is_empty l1 && Array.is_empty l2 in
- match kind_of_term term1, kind_of_term term2 with
+ match EConstr.kind evd term1, EConstr.kind evd term2 with
| Evar (evk1,args1), (Rel _|Var _) when app_empty
- && List.for_all (fun a -> Term.eq_constr a term2 || isEvar a)
+ && List.for_all (fun a -> EConstr.eq_constr evd a term2 || isEvar evd a)
(remove_instance_local_defs evd evk1 args1) ->
(* The typical kind of constraint coming from pattern-matching return
type inference *)
@@ -1132,9 +1245,9 @@ let apply_conversion_problem_heuristic ts env evd pbty t1 t2 =
| Some evd -> Success evd
| None ->
let reason = ProblemBeyondCapabilities in
- UnifFailure (evd, CannotSolveConstraint ((pbty,env,t1,t2),reason)))
+ UnifFailure (evd, CannotSolveConstraint (to_pb (pbty,env,t1,t2),reason)))
| (Rel _|Var _), Evar (evk2,args2) when app_empty
- && List.for_all (fun a -> Term.eq_constr a term1 || isEvar a)
+ && List.for_all (fun a -> EConstr.eq_constr evd a term1 || isEvar evd a)
(remove_instance_local_defs evd evk2 args2) ->
(* The typical kind of constraint coming from pattern-matching return
type inference *)
@@ -1142,7 +1255,7 @@ let apply_conversion_problem_heuristic ts env evd pbty t1 t2 =
| Some evd -> Success evd
| None ->
let reason = ProblemBeyondCapabilities in
- UnifFailure (evd, CannotSolveConstraint ((pbty,env,t1,t2),reason)))
+ UnifFailure (evd, CannotSolveConstraint (to_pb (pbty,env,t1,t2),reason)))
| Evar (evk1,args1), Evar (evk2,args2) when Evar.equal evk1 evk2 ->
let f env evd pbty x y = is_fconv ~reds:ts pbty env evd x y in
Success (solve_refl ~can_drop:true f env evd
@@ -1176,29 +1289,31 @@ let apply_conversion_problem_heuristic ts env evd pbty t1 t2 =
evar_conv_x ts env evd pbty t1 t2
let error_cannot_unify env evd pb ?reason t1 t2 =
- Pretype_errors.error_cannot_unify_loc
- (loc_of_conv_pb evd pb) env
+ Pretype_errors.error_cannot_unify
+ ?loc:(loc_of_conv_pb evd pb) env
evd ?reason (t1, t2)
let check_problems_are_solved env evd =
match snd (extract_all_conv_pbs evd) with
- | (pbty,env,t1,t2) as pb::_ -> error_cannot_unify env evd pb t1 t2
+ | (pbty,env,t1,t2) as pb::_ -> error_cannot_unify env evd pb (EConstr.of_constr t1) (EConstr.of_constr t2)
| _ -> ()
+exception MaxUndefined of (Evar.t * evar_info * Constr.t list)
+
let max_undefined_with_candidates evd =
- (* If evar were ordered with highest index first, fold_undefined
- would be going decreasingly and we could use fold_undefined to
- find the undefined evar of maximum index (alternatively,
- max_bindings from ocaml 3.12 could be used); instead we traverse
- the whole map *)
- let l = Evd.fold_undefined
- (fun evk ev_info evars ->
- match ev_info.evar_candidates with
- | None -> evars
- | Some l -> (evk,ev_info,l)::evars) evd [] in
- match l with
- | [] -> None
- | a::l -> Some (List.last (a::l))
+ let fold evk evi () = match evi.evar_candidates with
+ | None -> ()
+ | Some l -> raise (MaxUndefined (evk, evi, l))
+ in
+ (** [fold_right] traverses the undefined map in decreasing order of indices.
+ The evar with candidates of maximum index is thus the first evar with
+ candidates found by a [fold_right] traversal. This has a significant impact on
+ performance. *)
+ try
+ let () = Evar.Map.fold_right fold (Evd.undefined_map evd) () in
+ None
+ with MaxUndefined ans ->
+ Some ans
let rec solve_unconstrained_evars_with_candidates ts evd =
(* max_undefined is supposed to return the most recent, hence
@@ -1207,11 +1322,11 @@ let rec solve_unconstrained_evars_with_candidates ts evd =
| None -> evd
| Some (evk,ev_info,l) ->
let rec aux = function
- | [] -> error "Unsolvable existential variables."
+ | [] -> user_err Pp.(str "Unsolvable existential variables.")
| a::l ->
try
let conv_algo = evar_conv_x ts in
- let evd = check_evar_instance evd evk a conv_algo in
+ let evd = check_evar_instance evd evk (EConstr.of_constr a) conv_algo in
let evd = Evd.define evk a evd in
match reconsider_unif_constraints conv_algo evd with
| Success evd -> solve_unconstrained_evars_with_candidates ts evd
@@ -1229,11 +1344,11 @@ let solve_unconstrained_impossible_cases env evd =
match ev_info.evar_source with
| loc,Evar_kinds.ImpossibleCase ->
let j, ctx = coq_unit_judge () in
- let evd' = Evd.merge_context_set Evd.univ_flexible_alg ~loc evd' ctx in
+ let evd' = Evd.merge_context_set Evd.univ_flexible_alg ?loc evd' ctx in
let ty = j_type j in
let conv_algo = evar_conv_x full_transparent_state in
let evd' = check_evar_instance evd' evk ty conv_algo in
- Evd.define evk ty evd'
+ Evd.define evk (EConstr.Unsafe.to_constr ty) evd'
| _ -> evd') evd evd
let solve_unif_constraints_with_heuristics env
@@ -1242,6 +1357,8 @@ let solve_unif_constraints_with_heuristics env
let rec aux evd pbs progress stuck =
match pbs with
| (pbty,env,t1,t2 as pb) :: pbs ->
+ let t1 = EConstr.of_constr t1 in
+ let t2 = EConstr.of_constr t2 in
(match apply_conversion_problem_heuristic ts env evd pbty t1 t2 with
| Success evd' ->
let (evd', rest) = extract_all_conv_pbs evd' in
@@ -1258,6 +1375,8 @@ let solve_unif_constraints_with_heuristics env
match stuck with
| [] -> (* We're finished *) evd
| (pbty,env,t1,t2 as pb) :: _ ->
+ let t1 = EConstr.of_constr t1 in
+ let t2 = EConstr.of_constr t2 in
(* There remains stuck problems *)
error_cannot_unify env evd pb t1 t2
in
diff --git a/pretyping/evarconv.mli b/pretyping/evarconv.mli
index 2231e5bc..9270d6e3 100644
--- a/pretyping/evarconv.mli
+++ b/pretyping/evarconv.mli
@@ -1,13 +1,15 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
+open EConstr
open Environ
open Reductionops
open Evd
@@ -36,7 +38,7 @@ val e_cumul : env -> ?ts:transparent_state -> evar_map ref -> constr -> constr -
val solve_unif_constraints_with_heuristics : env -> ?ts:transparent_state -> evar_map -> evar_map
val consider_remaining_unif_problems : env -> ?ts:transparent_state -> evar_map -> evar_map
-(** @deprecated Alias for [solve_unif_constraints_with_heuristics] *)
+[@@ocaml.deprecated "Alias for [solve_unif_constraints_with_heuristics]"]
(** Check all pending unification problems are solved and raise an
error otherwise *)
@@ -46,18 +48,18 @@ val check_problems_are_solved : env -> evar_map -> unit
(** Check if a canonical structure is applicable *)
val check_conv_record : env -> evar_map ->
- constr * types Stack.t -> constr * types Stack.t ->
- Univ.universe_context_set * (constr * constr)
+ state -> state ->
+ Univ.ContextSet.t * (constr * constr)
* constr * constr list * (constr Stack.t * constr Stack.t) *
- (constr Stack.t * types Stack.t) *
- (constr Stack.t * types Stack.t) * constr *
+ (constr Stack.t * constr Stack.t) *
+ (constr Stack.t * constr Stack.t) * constr *
(int option * constr)
(** Try to solve problems of the form ?x[args] = c by second-order
matching, using typing to select occurrences *)
val second_order_matching : transparent_state -> env -> evar_map ->
- existential -> occurrences option list -> constr -> evar_map * bool
+ EConstr.existential -> occurrences option list -> constr -> evar_map * bool
(** Declare function to enforce evars resolution by using typing constraints *)
@@ -80,3 +82,5 @@ val evar_eqappr_x : ?rhs_is_already_stuck:bool -> transparent_state * bool ->
Evarsolve.unification_result
(**/**)
+(** {6 Functions to deal with impossible cases } *)
+val coq_unit_judge : unit -> EConstr.unsafe_judgment Univ.in_universe_context_set
diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml
index f9ab75ce..03f40ad9 100644
--- a/pretyping/evardefine.ml
+++ b/pretyping/evardefine.ml
@@ -1,47 +1,45 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Sorts
open Util
open Pp
open Names
-open Term
-open Vars
+open Constr
open Termops
+open EConstr
+open Vars
open Namegen
-open Environ
open Evd
open Evarutil
open Pretype_errors
-open Sigma.Notations
-let new_evar_unsafe env evd ?src ?filter ?candidates ?store ?naming ?principal typ =
- let evd = Sigma.Unsafe.of_evar_map evd in
- let Sigma (evk, evd, _) = new_evar env evd ?src ?filter ?candidates ?store ?naming ?principal typ in
- (Sigma.to_evar_map evd, evk)
+module RelDecl = Context.Rel.Declaration
let env_nf_evar sigma env =
- let open Context.Rel.Declaration in
+ let nf_evar c = nf_evar sigma c in
process_rel_context
- (fun d e -> push_rel (map_constr (nf_evar sigma) d) e) env
+ (fun d e -> push_rel (RelDecl.map_constr nf_evar d) e) env
let env_nf_betaiotaevar sigma env =
- let open Context.Rel.Declaration in
process_rel_context
- (fun d e ->
- push_rel (map_constr (Reductionops.nf_betaiota sigma) d) e) env
+ (fun d env ->
+ push_rel (RelDecl.map_constr (fun c -> Reductionops.nf_betaiota env sigma c) d) env) env
(****************************************)
(* Operations on value/type constraints *)
(****************************************)
-type type_constraint = types option
+type type_constraint = EConstr.types option
-type val_constraint = constr option
+type val_constraint = EConstr.constr option
(* Old comment...
* Basically, we have the following kind of constraints (in increasing
@@ -77,34 +75,30 @@ let define_pure_evar_as_product evd evk =
let open Context.Named.Declaration in
let evi = Evd.find_undefined evd evk in
let evenv = evar_env evi in
- let id = next_ident_away idx (ids_of_named_context (evar_context evi)) in
- let concl = Reductionops.whd_all evenv evd evi.evar_concl in
- let s = destSort concl in
+ let id = next_ident_away idx (Environ.ids_of_named_context_val evi.evar_hyps) in
+ let concl = Reductionops.whd_all evenv evd (EConstr.of_constr evi.evar_concl) in
+ let s = destSort evd concl in
let evd1,(dom,u1) =
- let evd = Sigma.Unsafe.of_evar_map evd in
- let Sigma (e, evd1, _) = new_type_evar evenv evd univ_flexible_alg ~filter:(evar_filter evi) in
- (Sigma.to_evar_map evd1, e)
+ new_type_evar evenv evd univ_flexible_alg ~filter:(evar_filter evi)
in
let evd2,rng =
let newenv = push_named (LocalAssum (id, dom)) evenv in
let src = evar_source evk evd1 in
let filter = Filter.extend 1 (evar_filter evi) in
- if is_prop_sort s then
+ if Sorts.is_prop (ESorts.kind evd1 s) then
(* Impredicative product, conclusion must fall in [Prop]. *)
- new_evar_unsafe newenv evd1 concl ~src ~filter
+ new_evar newenv evd1 concl ~src ~filter
else
let status = univ_flexible_alg in
let evd3, (rng, srng) =
- let evd1 = Sigma.Unsafe.of_evar_map evd1 in
- let Sigma (e, evd3, _) = new_type_evar newenv evd1 status ~src ~filter in
- (Sigma.to_evar_map evd3, e)
+ new_type_evar newenv evd1 status ~src ~filter
in
let prods = Univ.sup (univ_of_sort u1) (univ_of_sort srng) in
- let evd3 = Evd.set_leq_sort evenv evd3 (Type prods) s in
+ let evd3 = Evd.set_leq_sort evenv evd3 (Type prods) (ESorts.kind evd1 s) in
evd3, rng
in
let prod = mkProd (Name id, dom, subst_var id rng) in
- let evd3 = Evd.define evk prod evd2 in
+ let evd3 = Evd.define evk (EConstr.Unsafe.to_constr prod) evd2 in
evd3,prod
(* Refine an applied evar to a product and returns its instantiation *)
@@ -112,11 +106,11 @@ let define_pure_evar_as_product evd evk =
let define_evar_as_product evd (evk,args) =
let evd,prod = define_pure_evar_as_product evd evk in
(* Quick way to compute the instantiation of evk with args *)
- let na,dom,rng = destProd prod in
- let evdom = mkEvar (fst (destEvar dom), args) in
+ let na,dom,rng = destProd evd prod in
+ let evdom = mkEvar (fst (destEvar evd dom), args) in
let evrngargs = Array.cons (mkRel 1) (Array.map (lift 1) args) in
- let evrng = mkEvar (fst (destEvar rng), evrngargs) in
- evd,mkProd (na, evdom, evrng)
+ let evrng = mkEvar (fst (destEvar evd rng), evrngargs) in
+ evd, mkProd (na, evdom, evrng)
(* Refine an evar with an abstraction
@@ -131,36 +125,36 @@ let define_pure_evar_as_lambda env evd evk =
let open Context.Named.Declaration in
let evi = Evd.find_undefined evd evk in
let evenv = evar_env evi in
- let typ = Reductionops.whd_all evenv evd (evar_concl evi) in
- let evd1,(na,dom,rng) = match kind_of_term typ with
+ let typ = Reductionops.whd_all evenv evd (EConstr.of_constr (evar_concl evi)) in
+ let evd1,(na,dom,rng) = match EConstr.kind evd typ with
| Prod (na,dom,rng) -> (evd,(na,dom,rng))
- | Evar ev' -> let evd,typ = define_evar_as_product evd ev' in evd,destProd typ
- | _ -> error_not_product_loc Loc.ghost env evd typ in
- let avoid = ids_of_named_context (evar_context evi) in
+ | Evar ev' -> let evd,typ = define_evar_as_product evd ev' in evd,destProd evd typ
+ | _ -> error_not_product env evd typ in
+ let avoid = Environ.ids_of_named_context_val evi.evar_hyps in
let id =
next_name_away_with_default_using_types "x" na avoid (Reductionops.whd_evar evd dom) in
let newenv = push_named (LocalAssum (id, dom)) evenv in
let filter = Filter.extend 1 (evar_filter evi) in
let src = evar_source evk evd1 in
- let evd2,body = new_evar_unsafe newenv evd1 ~src (subst1 (mkVar id) rng) ~filter in
+ let evd2,body = new_evar newenv evd1 ~src (subst1 (mkVar id) rng) ~filter in
let lam = mkLambda (Name id, dom, subst_var id body) in
- Evd.define evk lam evd2, lam
+ Evd.define evk (EConstr.Unsafe.to_constr lam) evd2, lam
let define_evar_as_lambda env evd (evk,args) =
let evd,lam = define_pure_evar_as_lambda env evd evk in
(* Quick way to compute the instantiation of evk with args *)
- let na,dom,body = destLambda lam in
+ let na,dom,body = destLambda evd lam in
let evbodyargs = Array.cons (mkRel 1) (Array.map (lift 1) args) in
- let evbody = mkEvar (fst (destEvar body), evbodyargs) in
- evd,mkLambda (na, dom, evbody)
+ let evbody = mkEvar (fst (destEvar evd body), evbodyargs) in
+ evd, mkLambda (na, dom, evbody)
let rec evar_absorb_arguments env evd (evk,args as ev) = function
| [] -> evd,ev
| a::l ->
(* TODO: optimize and avoid introducing intermediate evars *)
let evd,lam = define_pure_evar_as_lambda env evd evk in
- let _,_,body = destLambda lam in
- let evk = fst (destEvar body) in
+ let _,_,body = destLambda evd lam in
+ let evk = fst (destEvar evd body) in
evar_absorb_arguments env evd (evk, Array.cons a args) l
(* Refining an evar to a sort *)
@@ -169,29 +163,29 @@ let define_evar_as_sort env evd (ev,args) =
let evd, u = new_univ_variable univ_rigid evd in
let evi = Evd.find_undefined evd ev in
let s = Type u in
- let concl = Reductionops.whd_all (evar_env evi) evd evi.evar_concl in
- let sort = destSort concl in
- let evd' = Evd.define ev (mkSort s) evd in
- Evd.set_leq_sort env evd' (Type (Univ.super u)) sort, s
+ let concl = Reductionops.whd_all (evar_env evi) evd (EConstr.of_constr evi.evar_concl) in
+ let sort = destSort evd concl in
+ let evd' = Evd.define ev (Constr.mkSort s) evd in
+ Evd.set_leq_sort env evd' (Type (Univ.super u)) (ESorts.kind evd' sort), s
(* Propagation of constraints through application and abstraction:
Given a type constraint on a functional term, returns the type
constraint on its domain and codomain. If the input constraint is
an evar instantiate it with the product of 2 new evars. *)
-let split_tycon loc env evd tycon =
+let split_tycon ?loc env evd tycon =
let rec real_split evd c =
let t = Reductionops.whd_all env evd c in
- match kind_of_term t with
+ match EConstr.kind evd t with
| Prod (na,dom,rng) -> evd, (na, dom, rng)
| Evar ev (* ev is undefined because of whd_all *) ->
let (evd',prod) = define_evar_as_product evd ev in
- let (_,dom,rng) = destProd prod in
+ let (_,dom,rng) = destProd evd prod in
evd',(Anonymous, dom, rng)
- | App (c,args) when isEvar c ->
- let (evd',lam) = define_evar_as_lambda env evd (destEvar c) in
+ | App (c,args) when isEvar evd c ->
+ let (evd',lam) = define_evar_as_lambda env evd (destEvar evd c) in
real_split evd' (mkApp (lam,args))
- | _ -> error_not_product_loc loc env evd c
+ | _ -> error_not_product ?loc env evd c
in
match tycon with
| None -> evd,(Anonymous,None,None)
@@ -202,6 +196,6 @@ let split_tycon loc env evd tycon =
let valcon_of_tycon x = x
let lift_tycon n = Option.map (lift n)
-let pr_tycon env = function
+let pr_tycon env sigma = function
None -> str "None"
- | Some t -> Termops.print_constr_env env t
+ | Some t -> Termops.print_constr_env env sigma t
diff --git a/pretyping/evardefine.mli b/pretyping/evardefine.mli
index 07b0e69d..cd23f9c6 100644
--- a/pretyping/evardefine.mli
+++ b/pretyping/evardefine.mli
@@ -1,13 +1,15 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
+open EConstr
open Evd
open Environ
@@ -30,7 +32,7 @@ val evar_absorb_arguments : env -> evar_map -> existential -> constr list ->
evar_map * existential
val split_tycon :
- Loc.t -> env -> evar_map -> type_constraint ->
+ ?loc:Loc.t -> env -> evar_map -> type_constraint ->
evar_map * (Name.t * type_constraint * type_constraint)
val valcon_of_tycon : type_constraint -> val_constraint
@@ -38,9 +40,9 @@ val lift_tycon : int -> type_constraint -> type_constraint
val define_evar_as_product : evar_map -> existential -> evar_map * types
val define_evar_as_lambda : env -> evar_map -> existential -> evar_map * types
-val define_evar_as_sort : env -> evar_map -> existential -> evar_map * sorts
+val define_evar_as_sort : env -> evar_map -> existential -> evar_map * Sorts.t
(** {6 debug pretty-printer:} *)
-val pr_tycon : env -> type_constraint -> Pp.std_ppcmds
+val pr_tycon : env -> evar_map -> type_constraint -> Pp.t
diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml
index 4fd03084..d0721439 100644
--- a/pretyping/evarsolve.ml
+++ b/pretyping/evarsolve.ml
@@ -1,45 +1,42 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Sorts
open Util
open CErrors
open Names
-open Term
-open Vars
+open Constr
open Environ
open Termops
open Evd
+open EConstr
+open Vars
open Namegen
open Retyping
open Reductionops
open Evarutil
open Pretype_errors
-open Sigma.Notations
let normalize_evar evd ev =
- match kind_of_term (whd_evar evd (mkEvar ev)) with
+ match EConstr.kind evd (mkEvar ev) with
| Evar (evk,args) -> (evk,args)
| _ -> assert false
-let get_polymorphic_positions f =
+let get_polymorphic_positions sigma f =
let open Declarations in
- match kind_of_term f with
+ match EConstr.kind sigma f with
| Ind (ind, u) | Construct ((ind, _), u) ->
let mib,oib = Global.lookup_inductive ind in
(match oib.mind_arity with
| RegularArity _ -> assert false
| TemplateArity templ -> templ.template_param_levels)
- | Const (cst, u) ->
- let cb = Global.lookup_constant cst in
- (match cb.const_type with
- | RegularArity _ -> assert false
- | TemplateArity (_, templ) ->
- templ.template_param_levels)
| _ -> assert false
let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false)
@@ -48,6 +45,7 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false)
let modified = ref false in
(* direction: true for fresh universes lower than the existing ones *)
let refresh_sort status ~direction s =
+ let s = ESorts.kind !evdref s in
let s' = evd_comb0 (new_sort_variable status) evdref in
let evd =
if direction then set_leq_sort env !evdref s' s
@@ -56,8 +54,10 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false)
modified := true; evdref := evd; mkSort s'
in
let rec refresh ~onlyalg status ~direction t =
- match kind_of_term t with
- | Sort (Type u as s) ->
+ match EConstr.kind !evdref t with
+ | Sort s ->
+ begin match ESorts.kind !evdref s with
+ | Type u ->
(match Univ.universe_level u with
| None -> refresh_sort status ~direction s
| Some l ->
@@ -67,31 +67,33 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false)
else t
| UnivFlexible alg ->
if onlyalg && alg then
- (evdref := Evd.make_flexible_variable !evdref false l; t)
+ (evdref := Evd.make_flexible_variable !evdref ~algebraic:false l; t)
else t))
- | Sort (Prop Pos as s) when refreshset && not direction ->
+ | Prop Pos when refreshset && not direction ->
(* Cannot make a universe "lower" than "Set",
only refreshing when we want higher universes. *)
refresh_sort status ~direction s
+ | _ -> t
+ end
| Prod (na,u,v) ->
mkProd (na, u, refresh ~onlyalg status ~direction v)
| _ -> t
(** Refresh the types of evars under template polymorphic references *)
and refresh_term_evars onevars top t =
- match kind_of_term (whd_evar !evdref t) with
- | App (f, args) when is_template_polymorphic env f ->
- let pos = get_polymorphic_positions f in
+ match EConstr.kind !evdref t with
+ | App (f, args) when is_template_polymorphic env !evdref f ->
+ let pos = get_polymorphic_positions !evdref f in
refresh_polymorphic_positions args pos
- | App (f, args) when top && isEvar f ->
+ | App (f, args) when top && isEvar !evdref f ->
refresh_term_evars true false f;
Array.iter (refresh_term_evars onevars false) args
| Evar (ev, a) when onevars ->
let evi = Evd.find !evdref ev in
- let ty' = refresh ~onlyalg univ_flexible ~direction:true evi.evar_concl in
+ let ty' = refresh ~onlyalg univ_flexible ~direction:true (EConstr.of_constr evi.evar_concl) in
if !modified then
- evdref := Evd.add !evdref ev {evi with evar_concl = ty'}
+ evdref := Evd.add !evdref ev {evi with evar_concl = EConstr.Unsafe.to_constr ty'}
else ()
- | _ -> Constr.iter (refresh_term_evars onevars false) t
+ | _ -> EConstr.iter !evdref (refresh_term_evars onevars false) t
and refresh_polymorphic_positions args pos =
let rec aux i = function
| Some l :: ls ->
@@ -106,7 +108,7 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false)
in aux 0 pos
in
let t' =
- if isArity t then
+ if isArity !evdref t then
match pbty with
| None ->
(* No cumulativity needed, but we still need to refresh the algebraics *)
@@ -135,6 +137,8 @@ let test_success conv_algo env evd c c' rhs =
is_success (conv_algo env evd c c' rhs)
let add_conv_oriented_pb ?(tail=true) (pbty,env,t1,t2) evd =
+ let t1 = EConstr.Unsafe.to_constr t1 in
+ let t2 = EConstr.Unsafe.to_constr t2 in
match pbty with
| Some true -> add_conv_pb ~tail (Reduction.CUMUL,env,t1,t2) evd
| Some false -> add_conv_pb ~tail (Reduction.CUMUL,env,t2,t1) evd
@@ -142,18 +146,18 @@ let add_conv_oriented_pb ?(tail=true) (pbty,env,t1,t2) evd =
(* We retype applications to ensure the universe constraints are collected *)
-exception IllTypedInstance of env * types * types
+exception IllTypedInstance of env * EConstr.types * EConstr.types
let recheck_applications conv_algo env evdref t =
let rec aux env t =
- match kind_of_term t with
+ match EConstr.kind !evdref t with
| App (f, args) ->
let () = aux env f in
let fty = Retyping.get_type_of env !evdref f in
let argsty = Array.map (fun x -> aux env x; Retyping.get_type_of env !evdref x) args in
let rec aux i ty =
if i < Array.length argsty then
- match kind_of_term (whd_all env !evdref ty) with
+ match EConstr.kind !evdref (whd_all env !evdref ty) with
| Prod (na, dom, codom) ->
(match conv_algo env !evdref Reduction.CUMUL argsty.(i) dom with
| Success evd -> evdref := evd;
@@ -164,7 +168,7 @@ let recheck_applications conv_algo env evdref t =
else ()
in aux 0 fty
| _ ->
- iter_constr_with_full_binders (fun d env -> push_rel d env) aux env t
+ iter_with_full_binders !evdref (fun d env -> push_rel d env) aux env t
in aux env t
@@ -177,7 +181,7 @@ type 'a update =
| NoUpdate
open Context.Named.Declaration
-let inst_of_vars sign = Array.map_of_list (mkVar % get_id) sign
+let inst_of_vars sign = Array.map_of_list (get_id %> mkVar) sign
let restrict_evar_key evd evk filter candidates =
match filter, candidates with
@@ -193,11 +197,9 @@ let restrict_evar_key evd evk filter candidates =
| None -> evar_filter evi
| Some filter -> filter in
let candidates = match candidates with
- | NoUpdate -> evi.evar_candidates
+ | NoUpdate -> Option.map (fun l -> List.map EConstr.of_constr l) evi.evar_candidates
| UpdateWith c -> Some c in
- let sigma = Sigma.Unsafe.of_evar_map evd in
- let Sigma (evk, sigma, _) = restrict_evar sigma evk filter candidates in
- (Sigma.to_evar_map sigma, evk)
+ restrict_evar evd evk filter candidates
end
(* Restrict an applied evar and returns its restriction in the same context *)
@@ -226,25 +228,22 @@ open Context.Rel.Declaration
let noccur_evar env evd evk c =
let cache = ref Int.Set.empty (* cache for let-ins *) in
let rec occur_rec check_types (k, env as acc) c =
- match kind_of_term c with
+ match EConstr.kind evd c with
| Evar (evk',args' as ev') ->
- (match safe_evar_value evd ev' with
- | Some c -> occur_rec check_types acc c
- | None ->
- if Evar.equal evk evk' then raise Occur
- else (if check_types then
- occur_rec false acc (existential_type evd ev');
- Array.iter (occur_rec check_types acc) args'))
+ if Evar.equal evk evk' then raise Occur
+ else (if check_types then
+ occur_rec false acc (existential_type evd ev');
+ Array.iter (occur_rec check_types acc) args')
| Rel i when i > k ->
if not (Int.Set.mem (i-k) !cache) then
let decl = Environ.lookup_rel i env in
if check_types then
- (cache := Int.Set.add (i-k) !cache; occur_rec false acc (lift i (get_type decl)));
+ (cache := Int.Set.add (i-k) !cache; occur_rec false acc (lift i (EConstr.of_constr (get_type decl))));
(match decl with
| LocalAssum _ -> ()
- | LocalDef (_,b,_) -> cache := Int.Set.add (i-k) !cache; occur_rec false acc (lift i b))
+ | LocalDef (_,b,_) -> cache := Int.Set.add (i-k) !cache; occur_rec false acc (lift i (EConstr.of_constr b)))
| Proj (p,c) -> occur_rec true acc c
- | _ -> iter_constr_with_full_binders (fun rd (k,env) -> (succ k, push_rel rd env))
+ | _ -> iter_with_full_binders evd (fun rd (k,env) -> (succ k, push_rel rd env))
(occur_rec check_types) acc c
in
try occur_rec false (0,env) c; true with Occur -> false
@@ -253,158 +252,209 @@ let noccur_evar env evd evk c =
(* Managing chains of local definitons *)
(***************************************)
+type alias =
+| RelAlias of int
+| VarAlias of Id.t
+
+let of_alias = function
+| RelAlias n -> mkRel n
+| VarAlias id -> mkVar id
+
+let to_alias sigma c = match EConstr.kind sigma c with
+| Rel n -> Some (RelAlias n)
+| Var id -> Some (VarAlias id)
+| _ -> None
+
+let is_alias sigma c alias = match EConstr.kind sigma c, alias with
+| Var id, VarAlias id' -> Id.equal id id'
+| Rel n, RelAlias n' -> Int.equal n n'
+| _ -> false
+
+let eq_alias a b = match a, b with
+| RelAlias n, RelAlias m -> Int.equal m n
+| VarAlias id1, VarAlias id2 -> Id.equal id1 id2
+| _ -> false
+
+type aliasing = EConstr.t option * alias list
+
+let empty_aliasing = None, []
+let make_aliasing c = Some c, []
+let push_alias (alias, l) a = (alias, a :: l)
+let lift_aliasing n (alias, l) =
+ let map a = match a with
+ | VarAlias _ -> a
+ | RelAlias m -> RelAlias (m + n)
+ in
+ (Option.map (fun c -> lift n c) alias, List.map map l)
+
+type aliases = {
+ rel_aliases : aliasing Int.Map.t;
+ var_aliases : aliasing Id.Map.t;
+ (** Only contains [VarAlias] *)
+}
+
(* Expand rels and vars that are bound to other rels or vars so that
dependencies in variables are canonically associated to the most ancient
variable in its family of aliased variables *)
-let compute_var_aliases sign =
+let compute_var_aliases sign sigma =
let open Context.Named.Declaration in
List.fold_right (fun decl aliases ->
let id = get_id decl in
match decl with
| LocalDef (_,t,_) ->
- (match kind_of_term t with
+ (match EConstr.kind sigma t with
| Var id' ->
let aliases_of_id =
- try Id.Map.find id' aliases with Not_found -> [] in
- Id.Map.add id (aliases_of_id@[t]) aliases
+ try Id.Map.find id' aliases with Not_found -> empty_aliasing in
+ Id.Map.add id (push_alias aliases_of_id (VarAlias id')) aliases
| _ ->
- Id.Map.add id [t] aliases)
+ Id.Map.add id (make_aliasing t) aliases)
| LocalAssum _ -> aliases)
sign Id.Map.empty
-let compute_rel_aliases var_aliases rels =
+let compute_rel_aliases var_aliases rels sigma =
snd (List.fold_right
(fun decl (n,aliases) ->
(n-1,
match decl with
| LocalDef (_,t,u) ->
- (match kind_of_term t with
+ (match EConstr.kind sigma t with
| Var id' ->
let aliases_of_n =
- try Id.Map.find id' var_aliases with Not_found -> [] in
- Int.Map.add n (aliases_of_n@[t]) aliases
+ try Id.Map.find id' var_aliases with Not_found -> empty_aliasing in
+ Int.Map.add n (push_alias aliases_of_n (VarAlias id')) aliases
| Rel p ->
let aliases_of_n =
- try Int.Map.find (p+n) aliases with Not_found -> [] in
- Int.Map.add n (aliases_of_n@[mkRel (p+n)]) aliases
+ try Int.Map.find (p+n) aliases with Not_found -> empty_aliasing in
+ Int.Map.add n (push_alias aliases_of_n (RelAlias (p+n))) aliases
| _ ->
- Int.Map.add n [lift n (mkCast(t,DEFAULTcast,u))] aliases)
+ Int.Map.add n (make_aliasing (lift n (mkCast(t,DEFAULTcast,u)))) aliases)
| LocalAssum _ -> aliases)
)
rels
(List.length rels,Int.Map.empty))
-let make_alias_map env =
+let make_alias_map env sigma =
(* We compute the chain of aliases for each var and rel *)
- let var_aliases = compute_var_aliases (named_context env) in
- let rel_aliases = compute_rel_aliases var_aliases (rel_context env) in
- (var_aliases,rel_aliases)
+ let var_aliases = compute_var_aliases (named_context env) sigma in
+ let rel_aliases = compute_rel_aliases var_aliases (rel_context env) sigma in
+ { var_aliases; rel_aliases }
-let lift_aliases n (var_aliases,rel_aliases as aliases) =
+let lift_aliases n aliases =
if Int.equal n 0 then aliases else
- (var_aliases,
- Int.Map.fold (fun p l -> Int.Map.add (p+n) (List.map (lift n) l))
- rel_aliases Int.Map.empty)
-
-let get_alias_chain_of aliases x = match kind_of_term x with
- | Rel n -> (try Int.Map.find n (snd aliases) with Not_found -> [])
- | Var id -> (try Id.Map.find id (fst aliases) with Not_found -> [])
- | _ -> []
-
-let normalize_alias_opt aliases x =
- match get_alias_chain_of aliases x with
- | [] -> None
- | a::_ when isRel a || isVar a -> Some a
- | [_] -> None
- | _::a::_ -> Some a
-
-let normalize_alias aliases x =
- match normalize_alias_opt aliases x with
+ let rel_aliases =
+ Int.Map.fold (fun p l -> Int.Map.add (p+n) (lift_aliasing n l))
+ aliases.rel_aliases Int.Map.empty
+ in
+ { aliases with rel_aliases }
+
+let get_alias_chain_of sigma aliases x = match x with
+ | RelAlias n -> (try Int.Map.find n aliases.rel_aliases with Not_found -> empty_aliasing)
+ | VarAlias id -> (try Id.Map.find id aliases.var_aliases with Not_found -> empty_aliasing)
+
+let normalize_alias_opt_alias sigma aliases x =
+ match get_alias_chain_of sigma aliases x with
+ | _, [] -> None
+ | _, a :: _ -> Some a
+
+let normalize_alias_opt sigma aliases x = match to_alias sigma x with
+| None -> None
+| Some a -> normalize_alias_opt_alias sigma aliases a
+
+let normalize_alias sigma aliases x =
+ match normalize_alias_opt_alias sigma aliases x with
| Some a -> a
| None -> x
-let normalize_alias_var var_aliases id =
- destVar (normalize_alias (var_aliases,Int.Map.empty) (mkVar id))
+let normalize_alias_var sigma var_aliases id =
+ let aliases = { var_aliases; rel_aliases = Int.Map.empty } in
+ match normalize_alias sigma aliases (VarAlias id) with
+ | VarAlias id -> id
+ | RelAlias _ -> assert false (** var only aliases to variables *)
-let extend_alias decl (var_aliases,rel_aliases) =
+let extend_alias sigma decl { var_aliases; rel_aliases } =
let rel_aliases =
- Int.Map.fold (fun n l -> Int.Map.add (n+1) (List.map (lift 1) l))
+ Int.Map.fold (fun n l -> Int.Map.add (n+1) (lift_aliasing 1 l))
rel_aliases Int.Map.empty in
let rel_aliases =
match decl with
| LocalDef(_,t,_) ->
- (match kind_of_term t with
+ (match EConstr.kind sigma t with
| Var id' ->
let aliases_of_binder =
- try Id.Map.find id' var_aliases with Not_found -> [] in
- Int.Map.add 1 (aliases_of_binder@[t]) rel_aliases
+ try Id.Map.find id' var_aliases with Not_found -> empty_aliasing in
+ Int.Map.add 1 (push_alias aliases_of_binder (VarAlias id')) rel_aliases
| Rel p ->
let aliases_of_binder =
- try Int.Map.find (p+1) rel_aliases with Not_found -> [] in
- Int.Map.add 1 (aliases_of_binder@[mkRel (p+1)]) rel_aliases
+ try Int.Map.find (p+1) rel_aliases with Not_found -> empty_aliasing in
+ Int.Map.add 1 (push_alias aliases_of_binder (RelAlias (p+1))) rel_aliases
| _ ->
- Int.Map.add 1 [lift 1 t] rel_aliases)
+ Int.Map.add 1 (make_aliasing (lift 1 t)) rel_aliases)
| LocalAssum _ -> rel_aliases in
- (var_aliases, rel_aliases)
-
-let expand_alias_once aliases x =
- match get_alias_chain_of aliases x with
- | [] -> None
- | l -> Some (List.last l)
-
-let expansions_of_var aliases x =
- match get_alias_chain_of aliases x with
- | [] -> [x]
- | a::_ as l when isRel a || isVar a -> x :: List.rev l
- | _::l -> x :: List.rev l
-
-let expansion_of_var aliases x =
- match get_alias_chain_of aliases x with
- | [] -> x
- | a::_ -> a
-
-let rec expand_vars_in_term_using aliases t = match kind_of_term t with
- | Rel _ | Var _ ->
- normalize_alias aliases t
+ { var_aliases; rel_aliases }
+
+let expand_alias_once sigma aliases x =
+ match get_alias_chain_of sigma aliases x with
+ | None, [] -> None
+ | Some a, [] -> Some a
+ | _, l -> Some (of_alias (List.last l))
+
+let expansions_of_var sigma aliases x =
+ let (_, l) = get_alias_chain_of sigma aliases x in
+ x :: List.rev l
+
+let expansion_of_var sigma aliases x =
+ match get_alias_chain_of sigma aliases x with
+ | None, [] -> (false, of_alias x)
+ | Some a, _ -> (true, a)
+ | None, a :: _ -> (true, of_alias a)
+
+let rec expand_vars_in_term_using sigma aliases t = match EConstr.kind sigma t with
+ | Rel n -> of_alias (normalize_alias sigma aliases (RelAlias n))
+ | Var id -> of_alias (normalize_alias sigma aliases (VarAlias id))
| _ ->
- map_constr_with_full_binders
- extend_alias expand_vars_in_term_using aliases t
+ let self aliases c = expand_vars_in_term_using sigma aliases c in
+ map_constr_with_full_binders sigma (extend_alias sigma) self aliases t
-let expand_vars_in_term env = expand_vars_in_term_using (make_alias_map env)
+let expand_vars_in_term env sigma = expand_vars_in_term_using sigma (make_alias_map env sigma)
-let free_vars_and_rels_up_alias_expansion aliases c =
+let free_vars_and_rels_up_alias_expansion sigma aliases c =
let acc1 = ref Int.Set.empty and acc2 = ref Id.Set.empty in
let acc3 = ref Int.Set.empty and acc4 = ref Id.Set.empty in
let cache_rel = ref Int.Set.empty and cache_var = ref Id.Set.empty in
let is_in_cache depth = function
- | Rel n -> Int.Set.mem (n-depth) !cache_rel
- | Var s -> Id.Set.mem s !cache_var
- | _ -> false in
+ | RelAlias n -> Int.Set.mem (n-depth) !cache_rel
+ | VarAlias s -> Id.Set.mem s !cache_var
+ in
let put_in_cache depth = function
- | Rel n -> cache_rel := Int.Set.add (n-depth) !cache_rel
- | Var s -> cache_var := Id.Set.add s !cache_var
- | _ -> () in
+ | RelAlias n -> cache_rel := Int.Set.add (n-depth) !cache_rel
+ | VarAlias s -> cache_var := Id.Set.add s !cache_var
+ in
let rec frec (aliases,depth) c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Rel _ | Var _ as ck ->
+ let ck = match ck with
+ | Rel n -> RelAlias n
+ | Var id -> VarAlias id
+ | _ -> assert false
+ in
if is_in_cache depth ck then () else begin
put_in_cache depth ck;
- let c' = expansion_of_var aliases c in
- (if c != c' then (* expansion, hence a let-in *)
- match kind_of_term c with
- | Var id -> acc4 := Id.Set.add id !acc4
- | Rel n -> if n >= depth+1 then acc3 := Int.Set.add (n-depth) !acc3
- | _ -> ());
- match kind_of_term c' with
+ let expanded, c' = expansion_of_var sigma aliases ck in
+ (if expanded then (* expansion, hence a let-in *)
+ match ck with
+ | VarAlias id -> acc4 := Id.Set.add id !acc4
+ | RelAlias n -> if n >= depth+1 then acc3 := Int.Set.add (n-depth) !acc3);
+ match EConstr.kind sigma c' with
| Var id -> acc2 := Id.Set.add id !acc2
| Rel n -> if n >= depth+1 then acc1 := Int.Set.add (n-depth) !acc1
| _ -> frec (aliases,depth) c end
| Const _ | Ind _ | Construct _ ->
- acc2 := Id.Set.union (vars_of_global (Global.env()) c) !acc2
+ acc2 := Id.Set.union (vars_of_global (Global.env()) (EConstr.to_constr sigma c)) !acc2
| _ ->
- iter_constr_with_full_binders
- (fun d (aliases,depth) -> (extend_alias d aliases,depth+1))
+ iter_with_full_binders sigma
+ (fun d (aliases,depth) -> (extend_alias sigma d aliases,depth+1))
frec (aliases,depth) c
in
frec (aliases,0) c;
@@ -414,42 +464,35 @@ let free_vars_and_rels_up_alias_expansion aliases c =
(* Managing pattern-unification *)
(********************************)
-let rec expand_and_check_vars aliases = function
- | [] -> []
- | a::l when isRel a || isVar a ->
- let a = expansion_of_var aliases a in
- if isRel a || isVar a then a :: expand_and_check_vars aliases l
- else raise Exit
- | _ ->
- raise Exit
-
-module Constrhash = Hashtbl.Make
- (struct type t = constr
- let equal = Term.eq_constr
- let hash = hash_constr
- end)
-
-let constr_list_distinct l =
- let visited = Constrhash.create 23 in
- let rec loop = function
- | h::t ->
- if Constrhash.mem visited h then false
- else (Constrhash.add visited h h; loop t)
- | [] -> true
- in loop l
-
-let get_actual_deps aliases l t =
- if occur_meta_or_existential t then
+let expand_and_check_vars sigma aliases l =
+ let map a = match get_alias_chain_of sigma aliases a with
+ | None, [] -> Some a
+ | None, a :: _ -> Some a
+ | Some _, _ -> None
+ in
+ Option.List.map map l
+
+let alias_distinct l =
+ let rec check (rels, vars) = function
+ | [] -> true
+ | RelAlias n :: l ->
+ not (Int.Set.mem n rels) && check (Int.Set.add n rels, vars) l
+ | VarAlias id :: l ->
+ not (Id.Set.mem id vars) && check (rels, Id.Set.add id vars) l
+ in
+ check (Int.Set.empty, Id.Set.empty) l
+
+let get_actual_deps evd aliases l t =
+ if occur_meta_or_existential evd t then
(* Probably no restrictions on allowed vars in presence of evars *)
l
else
(* Probably strong restrictions coming from t being evar-closed *)
- let (fv_rels,fv_ids,_,_) = free_vars_and_rels_up_alias_expansion aliases t in
- List.filter (fun c ->
- match kind_of_term c with
- | Var id -> Id.Set.mem id fv_ids
- | Rel n -> Int.Set.mem n fv_rels
- | _ -> assert false) l
+ let (fv_rels,fv_ids,_,_) = free_vars_and_rels_up_alias_expansion evd aliases t in
+ List.filter (function
+ | VarAlias id -> Id.Set.mem id fv_ids
+ | RelAlias n -> Int.Set.mem n fv_rels
+ ) l
open Context.Named.Declaration
let remove_instance_local_defs evd evk args =
@@ -468,35 +511,42 @@ let remove_instance_local_defs evd evk args =
(* Check if an applied evar "?X[args] l" is a Miller's pattern *)
-let find_unification_pattern_args env l t =
- if List.for_all (fun x -> isRel x || isVar x) l (* common failure case *) then
- let aliases = make_alias_map env in
- match (try Some (expand_and_check_vars aliases l) with Exit -> None) with
- | Some l as x when constr_list_distinct (get_actual_deps aliases l t) -> x
- | _ -> None
- else
- None
+let find_unification_pattern_args env evd l t =
+ let aliases = make_alias_map env evd in
+ match expand_and_check_vars evd aliases l with
+ | Some l as x when alias_distinct (get_actual_deps evd aliases l t) -> x
+ | _ -> None
-let is_unification_pattern_meta env nb m l t =
+let is_unification_pattern_meta env evd nb m l t =
(* Variables from context and rels > nb are implicitly all there *)
(* so we need to be a rel <= nb *)
- if List.for_all (fun x -> isRel x && destRel x <= nb) l then
- match find_unification_pattern_args env l t with
- | Some _ as x when not (dependent (mkMeta m) t) -> x
+ let map a = match EConstr.kind evd a with
+ | Rel n -> if n <= nb then Some (RelAlias n) else None
+ | _ -> None
+ in
+ match Option.List.map map l with
+ | Some l ->
+ begin match find_unification_pattern_args env evd l t with
+ | Some _ as x when not (dependent evd (mkMeta m) t) -> x
| _ -> None
- else
+ end
+ | None ->
None
let is_unification_pattern_evar env evd (evk,args) l t =
- if List.for_all (fun x -> isRel x || isVar x) l
- && noccur_evar env evd evk t
- then
+ match Option.List.map (fun c -> to_alias evd c) l with
+ | Some l when noccur_evar env evd evk t ->
let args = remove_instance_local_defs evd evk args in
+ let args = Option.List.map (fun c -> to_alias evd c) args in
+ begin match args with
+ | None -> None
+ | Some args ->
let n = List.length args in
- match find_unification_pattern_args env (args @ l) t with
+ match find_unification_pattern_args env evd (args @ l) t with
| Some l -> Some (List.skipn n l)
| _ -> None
- else None
+ end
+ | _ -> None
let is_unification_pattern_pure_evar env evd (evk,args) t =
let is_ev = is_unification_pattern_evar env evd (evk,args) [] t in
@@ -505,8 +555,8 @@ let is_unification_pattern_pure_evar env evd (evk,args) t =
| Some _ -> true
let is_unification_pattern (env,nb) evd f l t =
- match kind_of_term f with
- | Meta m -> is_unification_pattern_meta env nb m l t
+ match EConstr.kind evd f with
+ | Meta m -> is_unification_pattern_meta env evd nb m l t
| Evar ev -> is_unification_pattern_evar env evd ev l t
| _ -> None
@@ -517,18 +567,18 @@ let is_unification_pattern (env,nb) evd f l t =
*implicitly* depend on Vars but lambda abstraction will not reflect this
dependency: ?X x = ?1 (?1 is a meta) will return \_.?1 while it should
return \y. ?1{x\y} (non constant function if ?1 depends on x) (BB) *)
-let solve_pattern_eqn env l c =
+let solve_pattern_eqn env sigma l c =
let c' = List.fold_right (fun a c ->
- let c' = subst_term (lift 1 a) (lift 1 c) in
- match kind_of_term a with
+ let c' = subst_term sigma (lift 1 (of_alias a)) (lift 1 c) in
+ match a with
(* Rem: if [a] links to a let-in, do as if it were an assumption *)
- | Rel n ->
+ | RelAlias n ->
let open Context.Rel.Declaration in
let d = map_constr (lift n) (lookup_rel n env) in
mkLambda_or_LetIn d c'
- | Var id ->
+ | VarAlias id ->
let d = lookup_named id env in mkNamedLambda_or_LetIn d c'
- | _ -> assert false)
+ )
l c in
(* Warning: we may miss some opportunity to eta-reduce more since c'
is not in normal form *)
@@ -550,36 +600,35 @@ let solve_pattern_eqn env l c =
let make_projectable_subst aliases sigma evi args =
let sign = evar_filtered_context evi in
- let evar_aliases = compute_var_aliases sign in
+ let sign = List.map (fun d -> map_named_decl EConstr.of_constr d) sign in
+ let evar_aliases = compute_var_aliases sign sigma in
let (_,full_subst,cstr_subst) =
List.fold_right
(fun decl (args,all,cstrs) ->
match decl,args with
| LocalAssum (id,c), a::rest ->
- let a = whd_evar sigma a in
let cstrs =
- let a',args = decompose_app_vect a in
- match kind_of_term a' with
+ let a',args = decompose_app_vect sigma a in
+ match EConstr.kind sigma a' with
| Construct cstr ->
let l = try Constrmap.find (fst cstr) cstrs with Not_found -> [] in
Constrmap.add (fst cstr) ((args,id)::l) cstrs
| _ -> cstrs in
- (rest,Id.Map.add id [a,normalize_alias_opt aliases a,id] all,cstrs)
+ (rest,Id.Map.add id [a,normalize_alias_opt sigma aliases a,id] all,cstrs)
| LocalDef (id,c,_), a::rest ->
- let a = whd_evar sigma a in
- (match kind_of_term c with
+ (match EConstr.kind sigma c with
| Var id' ->
- let idc = normalize_alias_var evar_aliases id' in
+ let idc = normalize_alias_var sigma evar_aliases id' in
let sub = try Id.Map.find idc all with Not_found -> [] in
- if List.exists (fun (c,_,_) -> Term.eq_constr a c) sub then
+ if List.exists (fun (c,_,_) -> EConstr.eq_constr sigma a c) sub then
(rest,all,cstrs)
else
(rest,
- Id.Map.add idc ((a,normalize_alias_opt aliases a,id)::sub) all,
+ Id.Map.add idc ((a,normalize_alias_opt sigma aliases a,id)::sub) all,
cstrs)
| _ ->
- (rest,Id.Map.add id [a,normalize_alias_opt aliases a,id] all,cstrs))
- | _ -> anomaly (Pp.str "Instance does not match its signature"))
+ (rest,Id.Map.add id [a,normalize_alias_opt sigma aliases a,id] all,cstrs))
+ | _ -> anomaly (Pp.str "Instance does not match its signature."))
sign (Array.rev_to_list args,Id.Map.empty,Constrmap.empty) in
(full_subst,cstr_subst)
@@ -594,14 +643,13 @@ let make_projectable_subst aliases sigma evi args =
*)
let define_evar_from_virtual_equation define_fun env evd src t_in_env ty_t_in_sign sign filter inst_in_env =
- let evd = Sigma.Unsafe.of_evar_map evd in
- let Sigma (evar_in_env, evd, _) = new_evar_instance sign evd ty_t_in_sign ~filter ~src inst_in_env in
- let evd = Sigma.to_evar_map evd in
+ let (evd, evar_in_env) = new_evar_instance sign evd ty_t_in_sign ~filter ~src inst_in_env in
let t_in_env = whd_evar evd t_in_env in
- let evd = define_fun env evd None (destEvar evar_in_env) t_in_env in
+ let (evk, _) = destEvar evd evar_in_env in
+ let evd = define_fun env evd None (destEvar evd evar_in_env) t_in_env in
let ctxt = named_context_of_val sign in
let inst_in_sign = inst_of_vars (Filter.filter_list filter ctxt) in
- let evar_in_sign = mkEvar (fst (destEvar evar_in_env), inst_in_sign) in
+ let evar_in_sign = mkEvar (evk, inst_in_sign) in
(evd,whd_evar evd evar_in_sign)
(* We have x1..xq |- ?e1 : τ and had to solve something like
@@ -627,13 +675,14 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env =
if Evd.is_defined evd evk1 then
(* Some circularity somewhere (see e.g. #3209) *)
raise MorePreciseOccurCheckNeeeded;
- let (evk1,args1) = destEvar (whd_evar evd (mkEvar (evk1,args1))) in
+ let (evk1,args1) = destEvar evd (mkEvar (evk1,args1)) in
let evi1 = Evd.find_undefined evd evk1 in
let env1,rel_sign = env_rel_context_chop k env in
let sign1 = evar_hyps evi1 in
let filter1 = evar_filter evi1 in
let src = subterm_source evk1 evi1.evar_source in
let ids1 = List.map get_id (named_context_of_val sign1) in
+ let avoid = Environ.ids_of_named_context_val sign1 in
let inst_in_sign = List.map mkVar (Filter.filter_list filter1 ids1) in
let open Context.Rel.Declaration in
let (sign2,filter2,inst2_in_env,inst2_in_sign,_,evd,_) =
@@ -646,18 +695,18 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env =
~status:univ_flexible (Some false) env evd (mkSort s) in
define_evar_from_virtual_equation define_fun env evd src t_in_env
ty_t_in_sign sign filter inst_in_env in
- let evd,b_in_sign = match d with
- | LocalAssum _ -> evd,None
+ let evd,d' = match d with
+ | LocalAssum _ -> evd, Context.Named.Declaration.LocalAssum (id,t_in_sign)
| LocalDef (_,b,_) ->
let evd,b = define_evar_from_virtual_equation define_fun env evd src b
t_in_sign sign filter inst_in_env in
- evd,Some b in
- (push_named_context_val (Context.Named.Declaration.of_tuple (id,b_in_sign,t_in_sign)) sign, Filter.extend 1 filter,
+ evd, Context.Named.Declaration.LocalDef (id,b,t_in_sign) in
+ (push_named_context_val d' sign, Filter.extend 1 filter,
(mkRel 1)::(List.map (lift 1) inst_in_env),
(mkRel 1)::(List.map (lift 1) inst_in_sign),
- push_rel d env,evd,id::avoid))
+ push_rel d env,evd,Id.Set.add id avoid))
rel_sign
- (sign1,filter1,Array.to_list args1,inst_in_sign,env1,evd,ids1)
+ (sign1,filter1,Array.to_list args1,inst_in_sign,env1,evd,avoid)
in
let evd,ev2ty_in_sign =
let s = Retyping.get_sort_of env evd ty_in_env in
@@ -665,11 +714,9 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env =
~status:univ_flexible (Some false) env evd (mkSort s) in
define_evar_from_virtual_equation define_fun env evd src ty_in_env
ty_t_in_sign sign2 filter2 inst2_in_env in
- let evd = Sigma.Unsafe.of_evar_map evd in
- let Sigma (ev2_in_sign, evd, _) =
+ let (evd, ev2_in_sign) =
new_evar_instance sign2 evd ev2ty_in_sign ~filter:filter2 ~src inst2_in_sign in
- let evd = Sigma.to_evar_map evd in
- let ev2_in_env = (fst (destEvar ev2_in_sign), Array.of_list inst2_in_env) in
+ let ev2_in_env = (fst (destEvar evd ev2_in_sign), Array.of_list inst2_in_env) in
(evd, ev2_in_sign, ev2_in_env)
let restrict_upon_filter evd evk p args =
@@ -690,7 +737,7 @@ let find_projectable_constructor env evd cstr k args cstr_subst =
List.filter (fun (args',id) ->
(* is_conv is maybe too strong (and source of useless computation) *)
(* (at least expansion of aliases is needed) *)
- Array.for_all2 (is_conv env evd) args args') l in
+ Array.for_all2 (fun c1 c2 -> is_conv env evd c1 c2) args args') l in
List.map snd l
with Not_found ->
[]
@@ -729,7 +776,7 @@ let find_projectable_constructor env evd cstr k args cstr_subst =
type evar_projection =
| ProjectVar
-| ProjectEvar of existential * evar_info * Id.t * evar_projection
+| ProjectEvar of EConstr.existential * evar_info * Id.t * evar_projection
exception NotUnique
exception NotUniqueInType of (Id.t * evar_projection) list
@@ -737,19 +784,18 @@ exception NotUniqueInType of (Id.t * evar_projection) list
let rec assoc_up_to_alias sigma aliases y yc = function
| [] -> raise Not_found
| (c,cc,id)::l ->
- let c' = whd_evar sigma c in
- if Term.eq_constr y c' then id
+ if is_alias sigma c y then id
else
match l with
| _ :: _ -> assoc_up_to_alias sigma aliases y yc l
| [] ->
(* Last chance, we reason up to alias conversion *)
- match (if c == c' then cc else normalize_alias_opt aliases c') with
- | Some cc when Term.eq_constr yc cc -> id
- | _ -> if Term.eq_constr yc c then id else raise Not_found
+ match (normalize_alias_opt sigma aliases c) with
+ | Some cc when eq_alias yc cc -> id
+ | _ -> if is_alias sigma c yc then id else raise Not_found
let rec find_projectable_vars with_evars aliases sigma y subst =
- let yc = normalize_alias aliases y in
+ let yc = normalize_alias sigma aliases y in
let is_projectable idc idcl subst' =
(* First test if some [id] aliased to [idc] is bound to [y] in [subst] *)
try
@@ -759,12 +805,12 @@ let rec find_projectable_vars with_evars aliases sigma y subst =
(* Then test if [idc] is (indirectly) bound in [subst] to some evar *)
(* projectable on [y] *)
if with_evars then
- let f (c,_,id) = isEvar c && is_undefined sigma (fst (destEvar c)) in
+ let f (c,_,id) = isEvar sigma c in
let idcl' = List.filter f idcl in
match idcl' with
| [c,_,id] ->
begin
- let (evk,argsv as t) = destEvar c in
+ let (evk,argsv as t) = destEvar sigma c in
let evi = Evd.find sigma evk in
let subst,_ = make_projectable_subst aliases sigma evi argsv in
let l = find_projectable_vars with_evars aliases sigma y subst in
@@ -773,7 +819,7 @@ let rec find_projectable_vars with_evars aliases sigma y subst =
| _ -> subst'
end
| [] -> subst'
- | _ -> anomaly (Pp.str "More than one non var in aliases class of evar instance")
+ | _ -> anomaly (Pp.str "More than one non var in aliases class of evar instance.")
else
subst' in
Id.Map.fold is_projectable subst []
@@ -799,6 +845,25 @@ let rec find_solution_type evarenv = function
| (id,ProjectEvar _)::l -> find_solution_type evarenv l
| [] -> assert false
+let is_preferred_projection_over sign (id,p) (id',p') =
+ (* We give priority to projection of variables over instantiation of
+ an evar considering that the latter is a stronger decision which
+ may even procude an incorrect (ill-typed) solution *)
+ match p, p' with
+ | ProjectEvar _, ProjectVar -> false
+ | ProjectVar, ProjectEvar _ -> true
+ | _, _ ->
+ List.index Id.equal id sign < List.index Id.equal id' sign
+
+let choose_projection evi sols =
+ let sign = List.map get_id (evar_filtered_context evi) in
+ match sols with
+ | y::l ->
+ List.fold_right (fun (id,p as x) (id',_ as y) ->
+ if is_preferred_projection_over sign x y then x else y)
+ l y
+ | _ -> assert false
+
(* In case the solution to a projection problem requires the instantiation of
* subsidiary evars, [do_projection_effects] performs them; it
* also try to instantiate the type of those subsidiary evars if their
@@ -812,19 +877,18 @@ let rec find_solution_type evarenv = function
let rec do_projection_effects define_fun env ty evd = function
| ProjectVar -> evd
| ProjectEvar ((evk,argsv),evi,id,p) ->
- let evd = Evd.define evk (mkVar id) evd in
+ let evd = Evd.define evk (Constr.mkVar id) evd in
(* TODO: simplify constraints involving evk *)
let evd = do_projection_effects define_fun env ty evd p in
let ty = whd_all env evd (Lazy.force ty) in
- if not (isSort ty) then
+ if not (isSort evd ty) then
(* Don't try to instantiate if a sort because if evar_concl is an
evar it may commit to a univ level which is not the right
one (however, regarding coercions, because t is obtained by
unif, we know that no coercion can be inserted) *)
let subst = make_pure_subst evi argsv in
- let ty' = replace_vars subst evi.evar_concl in
- let ty' = whd_evar evd ty' in
- if isEvar ty' then define_fun env evd (Some false) (destEvar ty') ty else evd
+ let ty' = replace_vars subst (EConstr.of_constr evi.evar_concl) in
+ if isEvar evd ty' then define_fun env evd (Some false) (destEvar evd ty') ty else evd
else
evd
@@ -850,7 +914,7 @@ let rec do_projection_effects define_fun env ty evd = function
type projectibility_kind =
| NoUniqueProjection
- | UniqueProjection of constr * evar_projection list
+ | UniqueProjection of EConstr.constr * evar_projection list
type projectibility_status =
| CannotInvert
@@ -859,17 +923,16 @@ type projectibility_status =
let invert_arg_from_subst evd aliases k0 subst_in_env_extended_with_k_binders c_in_env_extended_with_k_binders =
let effects = ref [] in
let rec aux k t =
- let t = whd_evar evd t in
- match kind_of_term t with
- | Rel i when i>k0+k -> aux' k (mkRel (i-k))
- | Var id -> aux' k t
- | _ -> map_constr_with_binders succ aux k t
+ match EConstr.kind evd t with
+ | Rel i when i>k0+k -> aux' k (RelAlias (i-k))
+ | Var id -> aux' k (VarAlias id)
+ | _ -> map_with_binders evd succ aux k t
and aux' k t =
try project_with_effects aliases evd effects t subst_in_env_extended_with_k_binders
with Not_found ->
- match expand_alias_once aliases t with
+ match expand_alias_once evd aliases t with
| None -> raise Not_found
- | Some c -> aux k c in
+ | Some c -> aux k (lift k c) in
try
let c = aux 0 c_in_env_extended_with_k_binders in
Invertible (UniqueProjection (c,!effects))
@@ -931,23 +994,23 @@ let invert_invertible_arg fullenv evd aliases k (evk,argsv) args' =
let set_of_evctx l =
List.fold_left (fun s decl -> Id.Set.add (get_id decl) s) Id.Set.empty l
-let filter_effective_candidates evi filter candidates =
+let filter_effective_candidates evd evi filter candidates =
match filter with
| None -> candidates
| Some filter ->
let ids = set_of_evctx (Filter.filter_list filter (evar_context evi)) in
- List.filter (fun a -> Id.Set.subset (collect_vars a) ids) candidates
+ List.filter (fun a -> Id.Set.subset (collect_vars evd a) ids) candidates
let filter_candidates evd evk filter candidates_update =
let evi = Evd.find_undefined evd evk in
let candidates = match candidates_update with
- | NoUpdate -> evi.evar_candidates
+ | NoUpdate -> Option.map (fun l -> List.map EConstr.of_constr l) evi.evar_candidates
| UpdateWith c -> Some c
in
match candidates with
| None -> NoUpdate
| Some l ->
- let l' = filter_effective_candidates evi filter l in
+ let l' = filter_effective_candidates evd evi filter l in
if List.length l = List.length l' && candidates_update = NoUpdate then
NoUpdate
else
@@ -960,13 +1023,14 @@ let closure_of_filter evd evk = function
| None -> None
| Some filter ->
let evi = Evd.find_undefined evd evk in
- let vars = collect_vars (Evarutil.nf_evar evd (evar_concl evi)) in
- let test b decl = b || Idset.mem (get_id decl) vars ||
+ let vars = collect_vars evd (EConstr.of_constr (evar_concl evi)) in
+ let test b decl = b || Id.Set.mem (get_id decl) vars ||
match decl with
| LocalAssum _ ->
false
| LocalDef (_,c,_) ->
- not (isRel c || isVar c)
+ let c = EConstr.of_constr c in
+ not (isRel evd c || isVar evd c)
in
let newfilter = Filter.map_along test filter (evar_context evi) in
(* Now ensure that restriction is at least what is was originally *)
@@ -989,17 +1053,17 @@ let restrict_hyps evd evk filter candidates =
let typablefilter = closure_of_filter evd evk (Some filter) in
(typablefilter,candidates)
-exception EvarSolvedWhileRestricting of evar_map * constr
+exception EvarSolvedWhileRestricting of evar_map * EConstr.constr
let do_restrict_hyps evd (evk,args as ev) filter candidates =
let filter,candidates = match filter with
| None -> None,candidates
| Some filter -> restrict_hyps evd evk filter candidates in
match candidates,filter with
- | UpdateWith [], _ -> error "Not solvable."
+ | UpdateWith [], _ -> user_err Pp.(str "Not solvable.")
| UpdateWith [nc],_ ->
- let evd = Evd.define evk nc evd in
- raise (EvarSolvedWhileRestricting (evd,whd_evar evd (mkEvar ev)))
+ let evd = Evd.define evk (EConstr.Unsafe.to_constr nc) evd in
+ raise (EvarSolvedWhileRestricting (evd,mkEvar ev))
| NoUpdate, None -> evd,ev
| _ -> restrict_applied_evar evd ev filter candidates
@@ -1007,7 +1071,7 @@ let do_restrict_hyps evd (evk,args as ev) filter candidates =
(* ?e is assumed to have no candidates *)
let postpone_non_unique_projection env evd pbty (evk,argsv as ev) sols rhs =
- let rhs = expand_vars_in_term env rhs in
+ let rhs = expand_vars_in_term env evd rhs in
let filter =
restrict_upon_filter evd evk
(* Keep only variables that occur in rhs *)
@@ -1017,8 +1081,8 @@ let postpone_non_unique_projection env evd pbty (evk,argsv as ev) sols rhs =
(* that says that the body is hidden. Note that expand_vars_in_term *)
(* expands only rels and vars aliases, not rels or vars bound to an *)
(* arbitrary complex term *)
- (fun a -> not (isRel a || isVar a)
- || dependent a rhs || List.exists (fun (id,_) -> isVarId id a) sols)
+ (fun a -> not (isRel evd a || isVar evd a)
+ || dependent evd a rhs || List.exists (fun (id,_) -> isVarId evd id a) sols)
argsv in
let filter = closure_of_filter evd evk filter in
let candidates = extract_candidates sols in
@@ -1049,6 +1113,9 @@ let postpone_non_unique_projection env evd pbty (evk,argsv as ev) sols rhs =
* Note: argument f is the function used to instantiate evars.
*)
+let instantiate_evar_array evi c args =
+ EConstr.of_constr (instantiate_evar_array evi (EConstr.Unsafe.to_constr c) (Array.map EConstr.Unsafe.to_constr args))
+
let filter_compatible_candidates conv_algo env evd evi args rhs c =
let c' = instantiate_evar_array evi c args in
match conv_algo env evd Reduction.CONV rhs c' with
@@ -1068,7 +1135,9 @@ let restrict_candidates conv_algo env evd filter1 (evk1,argsv1) (evk2,argsv2) =
| _, None -> filter_candidates evd evk1 filter1 NoUpdate
| None, Some _ -> raise DoesNotPreserveCandidateRestriction
| Some l1, Some l2 ->
- let l1 = filter_effective_candidates evi1 filter1 l1 in
+ let l1 = List.map EConstr.of_constr l1 in
+ let l2 = List.map EConstr.of_constr l2 in
+ let l1 = filter_effective_candidates evd evi1 filter1 l1 in
let l1' = List.filter (fun c1 ->
let c1' = instantiate_evar_array evi1 c1 argsv1 in
let filter c2 =
@@ -1082,7 +1151,7 @@ let restrict_candidates conv_algo env evd filter1 (evk1,argsv1) (evk2,argsv2) =
if Int.equal (List.length l1) (List.length l1') then NoUpdate
else UpdateWith l1'
-exception CannotProject of evar_map * existential
+exception CannotProject of evar_map * EConstr.existential
(* Assume that FV(?n[x1:=t1..xn:=tn]) belongs to some set U.
Can ?n be instantiated by a term u depending essentially on xi such that the
@@ -1099,10 +1168,8 @@ exception CannotProject of evar_map * existential
*)
let rec is_constrainable_in top evd k (ev,(fv_rels,fv_ids) as g) t =
- let f,args2 = decompose_app_vect t in
- let f,args1 = decompose_app_vect (whd_evar evd f) in
- let args = Array.append args1 args2 in
- match kind_of_term f with
+ let f,args = decompose_app_vect evd t in
+ match EConstr.kind evd f with
| Construct ((ind,_),u) ->
let n = Inductiveops.inductive_nparams ind in
if n > Array.length args then true (* We don't try to be more clever *)
@@ -1118,30 +1185,32 @@ let rec is_constrainable_in top evd k (ev,(fv_rels,fv_ids) as g) t =
| _ -> (* We don't try to be more clever *) true
let has_constrainable_free_vars env evd aliases force k ev (fv_rels,fv_ids,let_rels,let_ids) t =
- let t' = expansion_of_var aliases t in
- if t' != t then
+ match to_alias evd t with
+ | Some t ->
+ let expanded, t' = expansion_of_var evd aliases t in
+ if expanded then
(* t is a local definition, we keep it only if appears in the list *)
(* of let-in variables effectively occurring on the right-hand side, *)
(* which is the only reason to keep it when inverting arguments *)
- match kind_of_term t with
- | Var id -> Id.Set.mem id let_ids
- | Rel n -> Int.Set.mem n let_rels
- | _ -> assert false
- else
+ match t with
+ | VarAlias id -> Id.Set.mem id let_ids
+ | RelAlias n -> Int.Set.mem n let_rels
+ else begin match t with
+ | VarAlias id -> Id.Set.mem id fv_ids
+ | RelAlias n -> n <= k || Int.Set.mem n fv_rels
+ end
+ | None ->
(* t is an instance for a proper variable; we filter it along *)
(* the free variables allowed to occur *)
- match kind_of_term t with
- | Var id -> Id.Set.mem id fv_ids
- | Rel n -> n <= k || Int.Set.mem n fv_rels
- | _ -> (not force || noccur_evar env evd ev t) && is_constrainable_in true evd k (ev,(fv_rels,fv_ids)) t
+ (not force || noccur_evar env evd ev t) && is_constrainable_in true evd k (ev,(fv_rels,fv_ids)) t
-exception EvarSolvedOnTheFly of evar_map * constr
+exception EvarSolvedOnTheFly of evar_map * EConstr.constr
(* Try to project evk1[argsv1] on evk2[argsv2], if [ev1] is a pattern on
the common domain of definition *)
let project_evar_on_evar force g env evd aliases k2 pbty (evk1,argsv1 as ev1) (evk2,argsv2 as ev2) =
(* Apply filtering on ev1 so that fvs(ev1) are in fvs(ev2). *)
- let fvs2 = free_vars_and_rels_up_alias_expansion aliases (mkEvar ev2) in
+ let fvs2 = free_vars_and_rels_up_alias_expansion evd aliases (mkEvar ev2) in
let filter1 = restrict_upon_filter evd evk1
(has_constrainable_free_vars env evd aliases force k2 evk2 fvs2)
argsv1 in
@@ -1171,11 +1240,11 @@ let check_evar_instance evd evk1 body conv_algo =
(* This happens in practice, cf MathClasses build failure on 2013-3-15 *)
let ty =
try Retyping.get_type_of ~lax:true evenv evd body
- with Retyping.RetypeError _ -> error "Ill-typed evar instance"
+ with Retyping.RetypeError _ -> user_err Pp.(str "Ill-typed evar instance")
in
- match conv_algo evenv evd Reduction.CUMUL ty evi.evar_concl with
+ match conv_algo evenv evd Reduction.CUMUL ty (EConstr.of_constr evi.evar_concl) with
| Success evd -> evd
- | UnifFailure _ -> raise (IllTypedInstance (evenv,ty,evi.evar_concl))
+ | UnifFailure _ -> raise (IllTypedInstance (evenv,ty,EConstr.of_constr evi.evar_concl))
let update_evar_source ev1 ev2 evd =
let loc, evs2 = evar_source ev2 evd in
@@ -1188,8 +1257,8 @@ let update_evar_source ev1 ev2 evd =
let solve_evar_evar_l2r force f g env evd aliases pbty ev1 (evk2,_ as ev2) =
try
let evd,body = project_evar_on_evar force g env evd aliases 0 pbty ev1 ev2 in
- let evd' = Evd.define evk2 body evd in
- let evd' = update_evar_source (fst (destEvar body)) evk2 evd' in
+ let evd' = Evd.define evk2 (EConstr.Unsafe.to_constr body) evd in
+ let evd' = update_evar_source (fst (destEvar evd body)) evk2 evd' in
check_evar_instance evd' evk2 body g
with EvarSolvedOnTheFly (evd,c) ->
f env evd pbty ev2 c
@@ -1206,7 +1275,7 @@ let preferred_orientation evd evk1 evk2 =
| _ -> true
let solve_evar_evar_aux force f g env evd pbty (evk1,args1 as ev1) (evk2,args2 as ev2) =
- let aliases = make_alias_map env in
+ let aliases = make_alias_map env evd in
if preferred_orientation evd evk1 evk2 then
try solve_evar_evar_l2r force f g env evd aliases (opp_problem pbty) ev2 ev1
with CannotProject (evd,ev2) ->
@@ -1223,15 +1292,18 @@ let solve_evar_evar_aux force f g env evd pbty (evk1,args1 as ev1) (evk2,args2 a
let solve_evar_evar ?(force=false) f g env evd pbty (evk1,args1 as ev1) (evk2,args2 as ev2) =
let pbty = if force then None else pbty in
let evi = Evd.find evd evk1 in
+ let downcast evk t evd = downcast evk (EConstr.Unsafe.to_constr t) evd in
let evd =
try
(* ?X : Π Δ. Type i = ?Y : Π Δ'. Type j.
The body of ?X and ?Y just has to be of type Π Δ. Type k for some k <= i, j. *)
let evienv = Evd.evar_env evi in
let ctx1, i = Reduction.dest_arity evienv evi.evar_concl in
+ let ctx1 = List.map (fun c -> map_rel_decl EConstr.of_constr c) ctx1 in
let evi2 = Evd.find evd evk2 in
let evi2env = Evd.evar_env evi2 in
let ctx2, j = Reduction.dest_arity evi2env evi2.evar_concl in
+ let ctx2 = List.map (fun c -> map_rel_decl EConstr.of_constr c) ctx2 in
let ui, uj = univ_of_sort i, univ_of_sort j in
if i == j || Evd.check_eq evd ui uj
then (* Shortcut, i = j *)
@@ -1253,10 +1325,10 @@ let solve_evar_evar ?(force=false) f g env evd pbty (evk1,args1 as ev1) (evk2,ar
solve_evar_evar_aux force f g env evd pbty ev1 ev2
type conv_fun =
- env -> evar_map -> conv_pb -> constr -> constr -> unification_result
+ env -> evar_map -> conv_pb -> EConstr.constr -> EConstr.constr -> unification_result
type conv_fun_bool =
- env -> evar_map -> conv_pb -> constr -> constr -> bool
+ env -> evar_map -> conv_pb -> EConstr.constr -> EConstr.constr -> bool
(* Solve pbs ?e[t1..tn] = ?e[u1..un] which arise often in fixpoint
* definitions. We try to unify the ti with the ui pairwise. The pairs
@@ -1265,7 +1337,13 @@ type conv_fun_bool =
let solve_refl ?(can_drop=false) conv_algo env evd pbty evk argsv1 argsv2 =
let evdref = ref evd in
- if Array.equal (e_eq_constr_univs evdref) argsv1 argsv2 then !evdref else
+ let eq_constr c1 c2 = match EConstr.eq_constr_universes env !evdref c1 c2 with
+ | None -> false
+ | Some cstr ->
+ try ignore (Evd.add_universe_constraints !evdref cstr); true
+ with UniversesDiffer -> false
+ in
+ if Array.equal eq_constr argsv1 argsv2 then !evdref else
(* Filter and restrict if needed *)
let args = Array.map2 (fun a1 a2 -> (a1, a2)) argsv1 argsv2 in
let untypedfilter =
@@ -1297,14 +1375,14 @@ let solve_candidates conv_algo env evd (evk,argsv) rhs =
| Some l ->
let l' =
List.map_filter
- (filter_compatible_candidates conv_algo env evd evi argsv rhs) l in
+ (fun c -> filter_compatible_candidates conv_algo env evd evi argsv rhs (EConstr.of_constr c)) l in
match l' with
| [] -> raise IncompatibleCandidates
| [c,evd] ->
(* solve_candidates might have been called recursively in the mean *)
(* time and the evar been solved by the filtering process *)
if Evd.is_undefined evd evk then
- let evd' = Evd.define evk c evd in
+ let evd' = Evd.define evk (EConstr.Unsafe.to_constr c) evd in
check_evar_instance evd' evk c conv_algo
else evd
| l when List.length l < List.length l' ->
@@ -1313,8 +1391,10 @@ let solve_candidates conv_algo env evd (evk,argsv) rhs =
| l -> evd
let occur_evar_upto_types sigma n c =
+ let c = EConstr.Unsafe.to_constr c in
let seen = ref Evar.Set.empty in
- let rec occur_rec c = match kind_of_term c with
+ (** FIXME: Is that supposed to be evar-insensitive? *)
+ let rec occur_rec c = match Constr.kind c with
| Evar (sp,_) when Evar.equal sp n -> raise Occur
| Evar (sp,args as e) ->
if Evar.Set.mem sp !seen then
@@ -1322,7 +1402,7 @@ let occur_evar_upto_types sigma n c =
else (
seen := Evar.Set.add sp !seen;
Option.iter occur_rec (existential_opt_value sigma e);
- occur_rec (existential_type sigma e))
+ occur_rec (Evd.existential_type sigma e))
| _ -> Constr.iter occur_rec c
in
try occur_rec c; false with Occur -> true
@@ -1350,14 +1430,14 @@ let occur_evar_upto_types sigma n c =
* Σ; Γ, y1..yk |- φ(u1..un) = c /\ x1..xn |- φ(x1..xn)
*)
-exception NotInvertibleUsingOurAlgorithm of constr
+exception NotInvertibleUsingOurAlgorithm of EConstr.constr
exception NotEnoughInformationToProgress of (Id.t * evar_projection) list
-exception NotEnoughInformationEvarEvar of constr
-exception OccurCheckIn of evar_map * constr
+exception NotEnoughInformationEvarEvar of EConstr.constr
+exception OccurCheckIn of evar_map * EConstr.constr
exception MetaOccurInBodyInternal
let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
- let aliases = make_alias_map env in
+ let aliases = make_alias_map env evd in
let evdref = ref evd in
let progress = ref false in
let evi = Evd.find evd evk in
@@ -1371,15 +1451,19 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
let c, p = match sols with
| [] -> raise Not_found
| [id,p] -> (mkVar id, p)
- | (id,p)::_::_ ->
- if choose then (mkVar id, p) else raise (NotUniqueInType sols)
+ | _ ->
+ if choose then
+ let (id,p) = choose_projection evi sols in
+ (mkVar id, p)
+ else
+ raise (NotUniqueInType sols)
in
- let ty = lazy (Retyping.get_type_of env !evdref t) in
+ let ty = lazy (Retyping.get_type_of env !evdref (of_alias t)) in
let evd = do_projection_effects (evar_define conv_algo ~choose) env ty !evdref p in
evdref := evd;
c
with
- | Not_found -> raise (NotInvertibleUsingOurAlgorithm t)
+ | Not_found -> raise (NotInvertibleUsingOurAlgorithm (of_alias t))
| NotUniqueInType sols ->
if not !progress then
raise (NotEnoughInformationToProgress sols);
@@ -1389,15 +1473,15 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
let ty' = instantiate_evar_array evi ty argsv in
let (evd,evar,(evk',argsv' as ev')) =
materialize_evar (evar_define conv_algo ~choose) env !evdref 0 ev ty' in
- let ts = expansions_of_var aliases t in
- let test c = isEvar c || List.mem_f Constr.equal c ts in
+ let ts = expansions_of_var evd aliases t in
+ let test c = isEvar evd c || List.exists (is_alias evd c) ts in
let filter = restrict_upon_filter evd evk test argsv' in
let filter = closure_of_filter evd evk' filter in
let candidates = extract_candidates sols in
let evd = match candidates with
| NoUpdate ->
let evd, ev'' = restrict_applied_evar evd ev' filter NoUpdate in
- Evd.add_conv_pb (Reduction.CONV,env,mkEvar ev'',t) evd
+ add_conv_oriented_pb ~tail:false (None,env,mkEvar ev'',of_alias t) evd
| UpdateWith _ ->
restrict_evar evd evk' filter candidates
in
@@ -1405,21 +1489,20 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
evar in
let rec imitate (env',k as envk) t =
- let t = whd_evar !evdref t in
- match kind_of_term t with
+ match EConstr.kind !evdref t with
| Rel i when i>k ->
let open Context.Rel.Declaration in
(match Environ.lookup_rel (i-k) env' with
- | LocalAssum _ -> project_variable (mkRel (i-k))
+ | LocalAssum _ -> project_variable (RelAlias (i-k))
| LocalDef (_,b,_) ->
- try project_variable (mkRel (i-k))
- with NotInvertibleUsingOurAlgorithm _ -> imitate envk (lift i b))
+ try project_variable (RelAlias (i-k))
+ with NotInvertibleUsingOurAlgorithm _ -> imitate envk (lift i (EConstr.of_constr b)))
| Var id ->
(match Environ.lookup_named id env' with
- | LocalAssum _ -> project_variable t
+ | LocalAssum _ -> project_variable (VarAlias id)
| LocalDef (_,b,_) ->
- try project_variable t
- with NotInvertibleUsingOurAlgorithm _ -> imitate envk b)
+ try project_variable (VarAlias id)
+ with NotInvertibleUsingOurAlgorithm _ -> imitate envk (EConstr.of_constr b))
| LetIn (na,b,u,c) ->
imitate envk (subst1 b c)
| Evar (evk',args' as ev') ->
@@ -1446,7 +1529,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
(* Try to project (a restriction of) the left evar ... *)
try
let evd,body = project_evar_on_evar false conv_algo env' evd aliases 0 None ev'' ev' in
- let evd = Evd.define evk' body evd in
+ let evd = Evd.define evk' (EConstr.Unsafe.to_constr body) evd in
check_evar_instance evd evk' body conv_algo
with
| EvarSolvedOnTheFly _ -> assert false (* ev has no candidates *)
@@ -1458,9 +1541,9 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
| _ ->
progress := true;
match
- let c,args = decompose_app_vect t in
- match kind_of_term c with
- | Construct (cstr,u) when noccur_between 1 k t ->
+ let c,args = decompose_app_vect !evdref t in
+ match EConstr.kind !evdref c with
+ | Construct (cstr,u) when noccur_between !evdref 1 k t ->
(* This is common case when inferring the return clause of match *)
(* (currently rudimentary: we do not treat the case of multiple *)
(* possible inversions; we do not treat overlap with a possible *)
@@ -1475,7 +1558,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
let candidates =
try
let t =
- map_constr_with_full_binders (fun d (env,k) -> push_rel d env, k+1)
+ map_constr_with_full_binders !evdref (fun d (env,k) -> push_rel d env, k+1)
imitate envk t in
t::l
with e when CErrors.noncritical e -> l in
@@ -1488,28 +1571,28 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
evar'')
| None ->
(* Evar/Rigid problem (or assimilated if not normal): we "imitate" *)
- map_constr_with_full_binders (fun d (env,k) -> push_rel d env, k+1)
+ map_constr_with_full_binders !evdref (fun d (env,k) -> push_rel d env, k+1)
imitate envk t
in
let rhs = whd_beta evd rhs (* heuristic *) in
let fast rhs =
let filter_ctxt = evar_filtered_context evi in
- let names = ref Idset.empty in
+ let names = ref Id.Set.empty in
let rec is_id_subst ctxt s =
match ctxt, s with
| (decl :: ctxt'), (c :: s') ->
let id = get_id decl in
- names := Idset.add id !names;
- isVarId id c && is_id_subst ctxt' s'
+ names := Id.Set.add id !names;
+ isVarId evd id c && is_id_subst ctxt' s'
| [], [] -> true
| _ -> false
in
is_id_subst filter_ctxt (Array.to_list argsv) &&
- closed0 rhs &&
- Idset.subset (collect_vars rhs) !names
+ closed0 evd rhs &&
+ Id.Set.subset (collect_vars evd rhs) !names
in
let body =
- if fast rhs then nf_evar evd rhs
+ if fast rhs then EConstr.of_constr (EConstr.to_constr evd rhs) (** FIXME? *)
else
let t' = imitate (env,0) rhs in
if !progress then
@@ -1525,7 +1608,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
*)
and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs =
- match kind_of_term rhs with
+ match EConstr.kind evd rhs with
| Evar (evk2,argsv2 as ev2) ->
if Evar.equal evk evk2 then
solve_refl ~can_drop:choose
@@ -1538,7 +1621,7 @@ and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs =
with NoCandidates ->
try
let (evd',body) = invert_definition conv_algo choose env evd pbty ev rhs in
- if occur_meta body then raise MetaOccurInBodyInternal;
+ if occur_meta evd' body then raise MetaOccurInBodyInternal;
(* invert_definition may have instantiate some evars of rhs with evk *)
(* so we recheck acyclicity *)
if occur_evar_upto_types evd' evk body then raise (OccurCheckIn (evd',body));
@@ -1561,7 +1644,7 @@ and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs =
print_constr body);
raise e in*)
let evd' = check_evar_instance evd' evk body conv_algo in
- Evd.define evk body evd'
+ Evd.define evk (EConstr.Unsafe.to_constr body) evd'
with
| NotEnoughInformationToProgress sols ->
postpone_non_unique_projection env evd pbty ev sols rhs
@@ -1574,7 +1657,7 @@ and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs =
| OccurCheckIn (evd,rhs) ->
(* last chance: rhs actually reduces to ev *)
let c = whd_all env evd rhs in
- match kind_of_term c with
+ match EConstr.kind evd c with
| Evar (evk',argsv2) when Evar.equal evk evk' ->
solve_refl (fun env sigma pb c c' -> is_fconv pb env sigma c c')
env evd pbty evk argsv argsv2
@@ -1607,17 +1690,19 @@ and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs =
* ass.
*)
-let status_changed lev (pbty,_,t1,t2) =
- (try Evar.Set.mem (head_evar t1) lev with NoHeadEvar -> false) ||
- (try Evar.Set.mem (head_evar t2) lev with NoHeadEvar -> false)
+let status_changed evd lev (pbty,_,t1,t2) =
+ let t1 = EConstr.of_constr t1 in
+ let t2 = EConstr.of_constr t2 in
+ (try Evar.Set.mem (head_evar evd t1) lev with NoHeadEvar -> false) ||
+ (try Evar.Set.mem (head_evar evd t2) lev with NoHeadEvar -> false)
let reconsider_unif_constraints conv_algo evd =
- let (evd,pbs) = extract_changed_conv_pbs evd status_changed in
+ let (evd,pbs) = extract_changed_conv_pbs evd (status_changed evd) in
List.fold_left
(fun p (pbty,env,t1,t2 as x) ->
match p with
| Success evd ->
- (match conv_algo env evd pbty t1 t2 with
+ (match conv_algo env evd pbty (EConstr.of_constr t1) (EConstr.of_constr t2) with
| Success _ as x -> x
| UnifFailure (i,e) -> UnifFailure (i,CannotSolveConstraint (x,e)))
| UnifFailure _ as x -> x)
diff --git a/pretyping/evarsolve.mli b/pretyping/evarsolve.mli
index b6bdc078..3f05c58c 100644
--- a/pretyping/evarsolve.mli
+++ b/pretyping/evarsolve.mli
@@ -1,15 +1,21 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
+open EConstr
open Evd
open Environ
+type alias
+
+val of_alias : alias -> EConstr.t
+
type unification_result =
| Success of evar_map
| UnifFailure of evar_map * Pretype_errors.unification_error
@@ -18,7 +24,7 @@ val is_success : unification_result -> bool
(** Replace the vars and rels that are aliases to other vars and rels by
their representative that is most ancient in the context *)
-val expand_vars_in_term : env -> constr -> constr
+val expand_vars_in_term : env -> evar_map -> constr -> constr
(** [evar_define choose env ev c] try to instantiate [ev] with [c] (typed in [env]),
possibly solving related unification problems, possibly leaving open
@@ -44,7 +50,7 @@ val refresh_universes :
env -> evar_map -> types -> evar_map * types
val solve_refl : ?can_drop:bool -> conv_fun_bool -> env -> evar_map ->
- bool option -> existential_key -> constr array -> constr array -> evar_map
+ bool option -> Evar.t -> constr array -> constr array -> evar_map
val solve_evar_evar : ?force:bool ->
(env -> evar_map -> bool option -> existential -> constr -> evar_map) ->
@@ -57,15 +63,15 @@ val solve_simple_eqn : conv_fun -> ?choose:bool -> env -> evar_map ->
val reconsider_unif_constraints : conv_fun -> evar_map -> unification_result
val reconsider_conv_pbs : conv_fun -> evar_map -> unification_result
-(** @deprecated Alias for [reconsider_unif_constraints] *)
+[@@ocaml.deprecated "Alias for [reconsider_unif_constraints]"]
val is_unification_pattern_evar : env -> evar_map -> existential -> constr list ->
- constr -> constr list option
+ constr -> alias list option
val is_unification_pattern : env * int -> evar_map -> constr -> constr list ->
- constr -> constr list option
+ constr -> alias list option
-val solve_pattern_eqn : env -> constr list -> constr -> constr
+val solve_pattern_eqn : env -> evar_map -> alias list -> constr -> constr
val noccur_evar : env -> evar_map -> Evar.t -> constr -> bool
@@ -73,10 +79,10 @@ exception IllTypedInstance of env * types * types
(* May raise IllTypedInstance if types are not convertible *)
val check_evar_instance :
- evar_map -> existential_key -> constr -> conv_fun -> evar_map
+ evar_map -> Evar.t -> constr -> conv_fun -> evar_map
val remove_instance_local_defs :
- evar_map -> existential_key -> constr array -> constr list
+ evar_map -> Evar.t -> 'a array -> 'a list
val get_type_of_refresh :
?polyprop:bool -> ?lax:bool -> env -> evar_map -> constr -> evar_map * types
diff --git a/pretyping/find_subterm.ml b/pretyping/find_subterm.ml
index 4caa1e99..b1608703 100644
--- a/pretyping/find_subterm.ml
+++ b/pretyping/find_subterm.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Pp
@@ -11,11 +13,12 @@ open Util
open CErrors
open Names
open Locus
-open Term
-open Nameops
+open EConstr
open Termops
open Pretype_errors
+module NamedDecl = Context.Named.Declaration
+
(** Processing occurrences *)
type occurrence_error =
@@ -28,14 +31,14 @@ let explain_invalid_occurrence l =
++ prlist_with_sep spc int l ++ str "."
let explain_incorrect_in_value_occurrence id =
- pr_id id ++ str " has no value."
+ Id.print id ++ str " has no value."
let explain_occurrence_error = function
| InvalidOccurrence l -> explain_invalid_occurrence l
| IncorrectInValueOccurrence id -> explain_incorrect_in_value_occurrence id
let error_occurrences_error e =
- errorlabstrm "" (explain_occurrence_error e)
+ user_err (explain_occurrence_error e)
let error_invalid_occurrence occ =
error_occurrences_error (InvalidOccurrence occ)
@@ -61,7 +64,10 @@ let proceed_with_occurrences f occs x =
let map_named_declaration_with_hyploc f hyploc acc decl =
let open Context.Named.Declaration in
- let f = f (Some (get_id decl, hyploc)) in
+ let f acc typ =
+ let acc, typ = f (Some (NamedDecl.get_id decl, hyploc)) acc typ in
+ acc, typ
+ in
match decl,hyploc with
| LocalAssum (id,_), InHypValueOnly ->
error_occurrences_error (IncorrectInValueOccurrence id)
@@ -80,10 +86,10 @@ let map_named_declaration_with_hyploc f hyploc acc decl =
exception SubtermUnificationError of subterm_unification_error
-exception NotUnifiable of (constr * constr * unification_error) option
+exception NotUnifiable of (EConstr.t * EConstr.t * unification_error) option
type 'a testing_function = {
- match_fun : 'a -> constr -> 'a;
+ match_fun : 'a -> EConstr.constr -> 'a;
merge_fun : 'a -> 'a -> 'a;
mutable testing_state : 'a;
mutable last_found : position_reporting option
@@ -94,7 +100,7 @@ type 'a testing_function = {
(b,l), b=true means no occurrence except the ones in l and b=false,
means all occurrences except the ones in l *)
-let replace_term_occ_gen_modulo occs like_first test bywhat cl occ t =
+let replace_term_occ_gen_modulo sigma occs like_first test bywhat cl occ t =
let (nowhere_except_in,locs) = Locusops.convert_occs occs in
let maxocc = List.fold_right max locs 0 in
let pos = ref occ in
@@ -128,23 +134,23 @@ let replace_term_occ_gen_modulo occs like_first test bywhat cl occ t =
with NotUnifiable _ ->
subst_below k t
and subst_below k t =
- map_constr_with_binders_left_to_right (fun d k -> k+1) substrec k t
+ map_constr_with_binders_left_to_right sigma (fun d k -> k+1) substrec k t
in
let t' = substrec 0 t in
(!pos, t')
-let replace_term_occ_modulo occs test bywhat t =
+let replace_term_occ_modulo evd occs test bywhat t =
let occs',like_first =
match occs with AtOccs occs -> occs,false | LikeFirst -> AllOccurrences,true in
proceed_with_occurrences
- (replace_term_occ_gen_modulo occs' like_first test bywhat None) occs' t
+ (replace_term_occ_gen_modulo evd occs' like_first test bywhat None) occs' t
-let replace_term_occ_decl_modulo occs test bywhat d =
+let replace_term_occ_decl_modulo evd occs test bywhat d =
let (plocs,hyploc),like_first =
match occs with AtOccs occs -> occs,false | LikeFirst -> (AllOccurrences,InHyp),true in
proceed_with_occurrences
(map_named_declaration_with_hyploc
- (replace_term_occ_gen_modulo plocs like_first test bywhat)
+ (replace_term_occ_gen_modulo evd plocs like_first test bywhat)
hyploc)
plocs d
@@ -152,11 +158,12 @@ let replace_term_occ_decl_modulo occs test bywhat d =
let make_eq_univs_test env evd c =
{ match_fun = (fun evd c' ->
- let b, cst = Universes.eq_constr_universes_proj env c c' in
- if b then
+ match EConstr.eq_constr_universes_proj env evd c c' with
+ | None -> raise (NotUnifiable None)
+ | Some cst ->
try Evd.add_universe_constraints evd cst
with Evd.UniversesDiffer -> raise (NotUnifiable None)
- else raise (NotUnifiable None));
+ );
merge_fun = (fun evd _ -> evd);
testing_state = evd;
last_found = None
@@ -165,7 +172,7 @@ let make_eq_univs_test env evd c =
let subst_closed_term_occ env evd occs c t =
let test = make_eq_univs_test env evd c in
let bywhat () = mkRel 1 in
- let t' = replace_term_occ_modulo occs test bywhat t in
+ let t' = replace_term_occ_modulo evd occs test bywhat t in
t', test.testing_state
let subst_closed_term_occ_decl env evd occs c d =
@@ -175,6 +182,6 @@ let subst_closed_term_occ_decl env evd occs c d =
let bywhat () = mkRel 1 in
proceed_with_occurrences
(map_named_declaration_with_hyploc
- (fun _ -> replace_term_occ_gen_modulo plocs like_first test bywhat None)
+ (fun _ -> replace_term_occ_gen_modulo evd plocs like_first test bywhat None)
hyploc) plocs d,
test.testing_state
diff --git a/pretyping/find_subterm.mli b/pretyping/find_subterm.mli
index c741ab04..9ba63b4f 100644
--- a/pretyping/find_subterm.mli
+++ b/pretyping/find_subterm.mli
@@ -1,16 +1,18 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Locus
-open Term
open Evd
open Pretype_errors
open Environ
+open EConstr
(** Finding subterms, possibly up to some unification function,
possibly at some given occurrences *)
@@ -41,15 +43,16 @@ val make_eq_univs_test : env -> evar_map -> constr -> evar_map testing_function
matching subterms at the indicated occurrences [occl] with [mk
()]; it turns a NotUnifiable exception raised by the testing
function into a SubtermUnificationError. *)
-val replace_term_occ_modulo : occurrences or_like_first ->
+val replace_term_occ_modulo : evar_map -> occurrences or_like_first ->
'a testing_function -> (unit -> constr) -> constr -> constr
(** [replace_term_occ_decl_modulo] is similar to
[replace_term_occ_modulo] but for a named_declaration. *)
val replace_term_occ_decl_modulo :
+ evar_map ->
(occurrences * hyp_location_flag) or_like_first ->
'a testing_function -> (unit -> constr) ->
- Context.Named.Declaration.t -> Context.Named.Declaration.t
+ named_declaration -> named_declaration
(** [subst_closed_term_occ occl c d] replaces occurrences of
closed [c] at positions [occl] by [Rel 1] in [d] (see also Note OCC),
@@ -61,7 +64,7 @@ val subst_closed_term_occ : env -> evar_map -> occurrences or_like_first ->
closed [c] at positions [occl] by [Rel 1] in [decl]. *)
val subst_closed_term_occ_decl : env -> evar_map ->
(occurrences * hyp_location_flag) or_like_first ->
- constr -> Context.Named.Declaration.t -> Context.Named.Declaration.t * evar_map
+ constr -> named_declaration -> named_declaration * evar_map
(** Miscellaneous *)
val error_invalid_occurrence : int list -> 'a
diff --git a/pretyping/geninterp.ml b/pretyping/geninterp.ml
new file mode 100644
index 00000000..1f8b9263
--- /dev/null
+++ b/pretyping/geninterp.ml
@@ -0,0 +1,102 @@
+(************************************************************************)
+(* * 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) *)
+(************************************************************************)
+
+open Names
+open Genarg
+
+module TacStore = Store.Make ()
+
+(** Dynamic toplevel values *)
+
+module ValT = Dyn.Make ()
+
+module Val =
+struct
+
+ type 'a typ = 'a ValT.tag
+
+ type _ tag =
+ | Base : 'a typ -> 'a tag
+ | List : 'a tag -> 'a list tag
+ | Opt : 'a tag -> 'a option tag
+ | Pair : 'a tag * 'b tag -> ('a * 'b) tag
+
+ type t = Dyn : 'a typ * 'a -> t
+
+ let eq = ValT.eq
+ let repr = ValT.repr
+ let create = ValT.create
+
+ let pr : type a. a typ -> Pp.t = fun t -> Pp.str (repr t)
+
+ let typ_list = ValT.create "list"
+ let typ_opt = ValT.create "option"
+ let typ_pair = ValT.create "pair"
+
+ let rec inject : type a. a tag -> a -> t = fun tag x -> match tag with
+ | Base t -> Dyn (t, x)
+ | List tag -> Dyn (typ_list, List.map (fun x -> inject tag x) x)
+ | Opt tag -> Dyn (typ_opt, Option.map (fun x -> inject tag x) x)
+ | Pair (tag1, tag2) ->
+ Dyn (typ_pair, (inject tag1 (fst x), inject tag2 (snd x)))
+
+end
+
+module ValTMap = ValT.Map
+
+module ValReprObj =
+struct
+ type ('raw, 'glb, 'top) obj = 'top Val.tag
+ let name = "valrepr"
+ let default _ = None
+end
+
+module ValRepr = Register(ValReprObj)
+
+let rec val_tag : type a b c. (a, b, c) genarg_type -> c Val.tag = function
+| ListArg t -> Val.List (val_tag t)
+| OptArg t -> Val.Opt (val_tag t)
+| PairArg (t1, t2) -> Val.Pair (val_tag t1, val_tag t2)
+| ExtraArg s -> ValRepr.obj (ExtraArg s)
+
+let val_tag = function Topwit t -> val_tag t
+
+let register_val0 wit tag =
+ let tag = match tag with
+ | None ->
+ let name = match wit with
+ | ExtraArg s -> ArgT.repr s
+ | _ -> assert false
+ in
+ Val.Base (Val.create name)
+ | Some tag -> tag
+ in
+ ValRepr.register0 wit tag
+
+(** Interpretation functions *)
+
+type interp_sign = {
+ lfun : Val.t Id.Map.t;
+ extra : TacStore.t }
+
+type ('glb, 'top) interp_fun = interp_sign -> 'glb -> 'top Ftactic.t
+
+module InterpObj =
+struct
+ type ('raw, 'glb, 'top) obj = ('glb, Val.t) interp_fun
+ let name = "interp"
+ let default _ = None
+end
+
+module Interp = Register(InterpObj)
+
+let interp = Interp.obj
+
+let register_interp0 = Interp.register0
diff --git a/pretyping/geninterp.mli b/pretyping/geninterp.mli
new file mode 100644
index 00000000..fa522e9c
--- /dev/null
+++ b/pretyping/geninterp.mli
@@ -0,0 +1,74 @@
+(************************************************************************)
+(* * 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) *)
+(************************************************************************)
+
+(** Interpretation functions for generic arguments and interpreted Ltac
+ values. *)
+
+open Names
+open Genarg
+
+(** {6 Dynamic toplevel values} *)
+
+module Val :
+sig
+ type 'a typ
+
+ val create : string -> 'a typ
+
+ type _ tag =
+ | Base : 'a typ -> 'a tag
+ | List : 'a tag -> 'a list tag
+ | Opt : 'a tag -> 'a option tag
+ | Pair : 'a tag * 'b tag -> ('a * 'b) tag
+
+ type t = Dyn : 'a typ * 'a -> t
+
+ val eq : 'a typ -> 'b typ -> ('a, 'b) CSig.eq option
+ val repr : 'a typ -> string
+ val pr : 'a typ -> Pp.t
+
+ val typ_list : t list typ
+ val typ_opt : t option typ
+ val typ_pair : (t * t) typ
+
+ val inject : 'a tag -> 'a -> t
+
+end
+
+module ValTMap (M : Dyn.TParam) :
+ Dyn.MapS with type 'a obj = 'a M.t with type 'a key = 'a Val.typ
+
+(** Dynamic types for toplevel values. While the generic types permit to relate
+ objects at various levels of interpretation, toplevel values are wearing
+ their own type regardless of where they came from. This allows to use the
+ same runtime representation for several generic types. *)
+
+val val_tag : 'a typed_abstract_argument_type -> 'a Val.tag
+(** Retrieve the dynamic type associated to a toplevel genarg. *)
+
+val register_val0 : ('raw, 'glb, 'top) genarg_type -> 'top Val.tag option -> unit
+(** Register the representation of a generic argument. If no tag is given as
+ argument, a new fresh tag with the same name as the argument is associated
+ to the generic type. *)
+
+(** {6 Interpretation functions} *)
+
+module TacStore : Store.S
+
+type interp_sign = {
+ lfun : Val.t Id.Map.t;
+ extra : TacStore.t }
+
+type ('glb, 'top) interp_fun = interp_sign -> 'glb -> 'top Ftactic.t
+
+val interp : ('raw, 'glb, 'top) genarg_type -> ('glb, Val.t) interp_fun
+
+val register_interp0 :
+ ('raw, 'glb, 'top) genarg_type -> ('glb, Val.t) interp_fun -> unit
diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml
index 51660818..eda45f3d 100644
--- a/pretyping/glob_ops.ml
+++ b/pretyping/glob_ops.ml
@@ -1,33 +1,46 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Util
+open CAst
open Names
open Nameops
open Globnames
open Misctypes
open Glob_term
+open Evar_kinds
+open Ltac_pretype
(* Untyped intermediate terms, after ASTs and before constr. *)
-let cases_pattern_loc = function
- PatVar(loc,_) -> loc
- | PatCstr(loc,_,_,_) -> loc
+let cases_pattern_loc c = c.CAst.loc
+
+let alias_of_pat pat = DAst.with_val (function
+ | PatVar name -> name
+ | PatCstr(_,_,name) -> name
+ ) pat
+
+let set_pat_alias id = DAst.map (function
+ | PatVar Anonymous -> PatVar (Name id)
+ | PatCstr (cstr,patl,Anonymous) -> PatCstr (cstr,patl,Name id)
+ | pat -> assert false)
let cases_predicate_names tml =
List.flatten (List.map (function
| (tm,(na,None)) -> [na]
- | (tm,(na,Some (_,_,nal))) -> na::nal) tml)
+ | (tm,(na,Some {v=(_,nal)})) -> na::nal) tml)
-let mkGApp loc p t =
- match p with
- | GApp (loc,f,l) -> GApp (loc,f,l@[t])
- | _ -> GApp (loc,p,[t])
+let mkGApp ?loc p t = DAst.make ?loc @@
+ match DAst.get p with
+ | GApp (f,l) -> GApp (f,l@[t])
+ | _ -> GApp (p,[t])
let map_glob_decl_left_to_right f (na,k,obd,ty) =
let comp1 = Option.map f obd in
@@ -35,367 +48,280 @@ let map_glob_decl_left_to_right f (na,k,obd,ty) =
(na,k,comp1,comp2)
let binding_kind_eq bk1 bk2 = match bk1, bk2 with
-| Decl_kinds.Explicit, Decl_kinds.Explicit -> true
-| Decl_kinds.Implicit, Decl_kinds.Implicit -> true
-| _ -> false
+ | Decl_kinds.Explicit, Decl_kinds.Explicit -> true
+ | Decl_kinds.Implicit, Decl_kinds.Implicit -> true
+ | (Decl_kinds.Explicit | Decl_kinds.Implicit), _ -> false
let case_style_eq s1 s2 = match s1, s2 with
-| LetStyle, LetStyle -> true
-| IfStyle, IfStyle -> true
-| LetPatternStyle, LetPatternStyle -> true
-| MatchStyle, MatchStyle -> true
-| RegularStyle, RegularStyle -> true
-| _ -> false
-
-let rec cases_pattern_eq p1 p2 = match p1, p2 with
-| PatVar (_, na1), PatVar (_, na2) -> Name.equal na1 na2
-| PatCstr (_, c1, pl1, na1), PatCstr (_, c2, pl2, na2) ->
- eq_constructor c1 c2 && List.equal cases_pattern_eq pl1 pl2 &&
- Name.equal na1 na2
-| _ -> false
+ | LetStyle, LetStyle -> true
+ | IfStyle, IfStyle -> true
+ | LetPatternStyle, LetPatternStyle -> true
+ | MatchStyle, MatchStyle -> true
+ | RegularStyle, RegularStyle -> true
+ | (LetStyle | IfStyle | LetPatternStyle | MatchStyle | RegularStyle), _ -> false
+
+let rec cases_pattern_eq p1 p2 = match DAst.get p1, DAst.get p2 with
+ | PatVar na1, PatVar na2 -> Name.equal na1 na2
+ | PatCstr (c1, pl1, na1), PatCstr (c2, pl2, na2) ->
+ eq_constructor c1 c2 && List.equal cases_pattern_eq pl1 pl2 &&
+ Name.equal na1 na2
+ | (PatVar _ | PatCstr _), _ -> false
let cast_type_eq eq t1 t2 = match t1, t2 with
-| CastConv t1, CastConv t2 -> eq t1 t2
-| CastVM t1, CastVM t2 -> eq t1 t2
-| CastCoerce, CastCoerce -> true
-| CastNative t1, CastNative t2 -> eq t1 t2
-| _ -> false
-
-let rec glob_constr_eq c1 c2 = match c1, c2 with
-| GRef (_, gr1, _), GRef (_, gr2, _) -> eq_gr gr1 gr2
-| GVar (_, id1), GVar (_, id2) -> Id.equal id1 id2
-| GEvar (_, id1, arg1), GEvar (_, id2, arg2) ->
- Id.equal id1 id2 &&
- List.equal instance_eq arg1 arg2
-| GPatVar (_, (b1, pat1)), GPatVar (_, (b2, pat2)) ->
- (b1 : bool) == b2 && Id.equal pat1 pat2
-| GApp (_, f1, arg1), GApp (_, f2, arg2) ->
- glob_constr_eq f1 f2 && List.equal glob_constr_eq arg1 arg2
-| GLambda (_, na1, bk1, t1, c1), GLambda (_, na2, bk2, t2, c2) ->
- Name.equal na1 na2 && binding_kind_eq bk1 bk2 &&
- glob_constr_eq t1 t2 && glob_constr_eq c1 c2
-| GProd (_, na1, bk1, t1, c1), GProd (_, na2, bk2, t2, c2) ->
- Name.equal na1 na2 && binding_kind_eq bk1 bk2 &&
- glob_constr_eq t1 t2 && glob_constr_eq c1 c2
-| GLetIn (_, na1, t1, c1), GLetIn (_, na2, t2, c2) ->
- Name.equal na1 na2 && glob_constr_eq t1 t2 && glob_constr_eq c1 c2
-| GCases (_, st1, c1, tp1, cl1), GCases (_, st2, c2, tp2, cl2) ->
- case_style_eq st1 st2 && Option.equal glob_constr_eq c1 c2 &&
- List.equal tomatch_tuple_eq tp1 tp2 &&
- List.equal cases_clause_eq cl1 cl2
-| GLetTuple (_, na1, (n1, p1), c1, t1), GLetTuple (_, na2, (n2, p2), c2, t2) ->
- List.equal Name.equal na1 na2 && Name.equal n1 n2 &&
- Option.equal glob_constr_eq p1 p2 && glob_constr_eq c1 c2 &&
- glob_constr_eq t1 t2
-| GIf (_, m1, (pat1, p1), c1, t1), GIf (_, m2, (pat2, p2), c2, t2) ->
- glob_constr_eq m1 m2 && Name.equal pat1 pat2 &&
- Option.equal glob_constr_eq p1 p2 && glob_constr_eq c1 c2 &&
- glob_constr_eq t1 t2
-| GRec (_, kn1, id1, decl1, c1, t1), GRec (_, kn2, id2, decl2, c2, t2) ->
- fix_kind_eq kn1 kn2 && Array.equal Id.equal id1 id2 &&
- Array.equal (fun l1 l2 -> List.equal glob_decl_eq l1 l2) decl1 decl2 &&
- Array.equal glob_constr_eq c1 c2 &&
- Array.equal glob_constr_eq t1 t2
-| GSort (_, s1), GSort (_, s2) -> Miscops.glob_sort_eq s1 s2
-| GHole (_, kn1, nam1, gn1), GHole (_, kn2, nam2, gn2) ->
- Option.equal (==) gn1 gn2 (** Only thing sensible *) &&
- Miscops.intro_pattern_naming_eq nam1 nam2
-| GCast (_, c1, t1), GCast (_, c2, t2) ->
- glob_constr_eq c1 c2 && cast_type_eq glob_constr_eq t1 t2
-| _ -> false
-
-and tomatch_tuple_eq (c1, p1) (c2, p2) =
- let eqp (_, i1, na1) (_, i2, na2) =
+ | CastConv t1, CastConv t2 -> eq t1 t2
+ | CastVM t1, CastVM t2 -> eq t1 t2
+ | CastCoerce, CastCoerce -> true
+ | CastNative t1, CastNative t2 -> eq t1 t2
+ | (CastConv _ | CastVM _ | CastCoerce | CastNative _), _ -> false
+
+let matching_var_kind_eq k1 k2 = match k1, k2 with
+| FirstOrderPatVar ido1, FirstOrderPatVar ido2 -> Id.equal ido1 ido2
+| SecondOrderPatVar id1, SecondOrderPatVar id2 -> Id.equal id1 id2
+| (FirstOrderPatVar _ | SecondOrderPatVar _), _ -> false
+
+let tomatch_tuple_eq f (c1, p1) (c2, p2) =
+ let eqp {CAst.v=(i1, na1)} {CAst.v=(i2, na2)} =
eq_ind i1 i2 && List.equal Name.equal na1 na2
in
let eq_pred (n1, o1) (n2, o2) = Name.equal n1 n2 && Option.equal eqp o1 o2 in
- glob_constr_eq c1 c2 && eq_pred p1 p2
+ f c1 c2 && eq_pred p1 p2
-and cases_clause_eq (_, id1, p1, c1) (_, id2, p2, c2) =
- List.equal Id.equal id1 id2 && List.equal cases_pattern_eq p1 p2 &&
- glob_constr_eq c1 c2
+and cases_clause_eq f {CAst.v=(id1, p1, c1)} {CAst.v=(id2, p2, c2)} =
+ List.equal Id.equal id1 id2 && List.equal cases_pattern_eq p1 p2 && f c1 c2
-and glob_decl_eq (na1, bk1, c1, t1) (na2, bk2, c2, t2) =
+let glob_decl_eq f (na1, bk1, c1, t1) (na2, bk2, c2, t2) =
Name.equal na1 na2 && binding_kind_eq bk1 bk2 &&
- Option.equal glob_constr_eq c1 c2 &&
- glob_constr_eq t1 t2
-
-and fix_kind_eq k1 k2 = match k1, k2 with
-| GFix (a1, i1), GFix (a2, i2) ->
- let eq (i1, o1) (i2, o2) =
- Option.equal Int.equal i1 i2 && fix_recursion_order_eq o1 o2
- in
- Int.equal i1 i2 && Array.equal eq a1 a1
-| GCoFix i1, GCoFix i2 -> Int.equal i1 i2
-| _ -> false
-
-and fix_recursion_order_eq o1 o2 = match o1, o2 with
-| GStructRec, GStructRec -> true
-| GWfRec c1, GWfRec c2 -> glob_constr_eq c1 c2
-| GMeasureRec (c1, o1), GMeasureRec (c2, o2) ->
- glob_constr_eq c1 c2 && Option.equal glob_constr_eq o1 o2
-| _ -> false
-
-and instance_eq (x1,c1) (x2,c2) =
- Id.equal x1 x2 && glob_constr_eq c1 c2
-
-let map_glob_constr_left_to_right f = function
- | GApp (loc,g,args) ->
+ Option.equal f c1 c2 && f t1 t2
+
+let fix_recursion_order_eq f o1 o2 = match o1, o2 with
+ | GStructRec, GStructRec -> true
+ | GWfRec c1, GWfRec c2 -> f c1 c2
+ | GMeasureRec (c1, o1), GMeasureRec (c2, o2) ->
+ f c1 c2 && Option.equal f o1 o2
+ | (GStructRec | GWfRec _ | GMeasureRec _), _ -> false
+
+let fix_kind_eq f k1 k2 = match k1, k2 with
+ | GFix (a1, i1), GFix (a2, i2) ->
+ let eq (i1, o1) (i2, o2) =
+ Option.equal Int.equal i1 i2 && fix_recursion_order_eq f o1 o2
+ in
+ Int.equal i1 i2 && Array.equal eq a1 a1
+ | GCoFix i1, GCoFix i2 -> Int.equal i1 i2
+ | (GFix _ | GCoFix _), _ -> false
+
+let instance_eq f (x1,c1) (x2,c2) =
+ Id.equal x1 x2 && f c1 c2
+
+let mk_glob_constr_eq f c1 c2 = match DAst.get c1, DAst.get c2 with
+ | GRef (gr1, _), GRef (gr2, _) -> eq_gr gr1 gr2
+ | GVar id1, GVar id2 -> Id.equal id1 id2
+ | GEvar (id1, arg1), GEvar (id2, arg2) ->
+ Id.equal id1 id2 && List.equal (instance_eq f) arg1 arg2
+ | GPatVar k1, GPatVar k2 -> matching_var_kind_eq k1 k2
+ | GApp (f1, arg1), GApp (f2, arg2) ->
+ f f1 f2 && List.equal f arg1 arg2
+ | GLambda (na1, bk1, t1, c1), GLambda (na2, bk2, t2, c2) ->
+ Name.equal na1 na2 && binding_kind_eq bk1 bk2 && f t1 t2 && f c1 c2
+ | GProd (na1, bk1, t1, c1), GProd (na2, bk2, t2, c2) ->
+ Name.equal na1 na2 && binding_kind_eq bk1 bk2 && f t1 t2 && f c1 c2
+ | GLetIn (na1, b1, t1, c1), GLetIn (na2, b2, t2, c2) ->
+ Name.equal na1 na2 && f b1 b2 && Option.equal f t1 t2 && f c1 c2
+ | GCases (st1, c1, tp1, cl1), GCases (st2, c2, tp2, cl2) ->
+ case_style_eq st1 st2 && Option.equal f c1 c2 &&
+ List.equal (tomatch_tuple_eq f) tp1 tp2 &&
+ List.equal (cases_clause_eq f) cl1 cl2
+ | GLetTuple (na1, (n1, p1), c1, t1), GLetTuple (na2, (n2, p2), c2, t2) ->
+ List.equal Name.equal na1 na2 && Name.equal n1 n2 &&
+ Option.equal f p1 p2 && f c1 c2 && f t1 t2
+ | GIf (m1, (pat1, p1), c1, t1), GIf (m2, (pat2, p2), c2, t2) ->
+ f m1 m2 && Name.equal pat1 pat2 &&
+ Option.equal f p1 p2 && f c1 c2 && f t1 t2
+ | GRec (kn1, id1, decl1, c1, t1), GRec (kn2, id2, decl2, c2, t2) ->
+ fix_kind_eq f kn1 kn2 && Array.equal Id.equal id1 id2 &&
+ Array.equal (fun l1 l2 -> List.equal (glob_decl_eq f) l1 l2) decl1 decl2 &&
+ Array.equal f c1 c2 && Array.equal f t1 t2
+ | GSort s1, GSort s2 -> Miscops.glob_sort_eq s1 s2
+ | GHole (kn1, nam1, gn1), GHole (kn2, nam2, gn2) ->
+ Option.equal (==) gn1 gn2 (** Only thing sensible *) &&
+ Miscops.intro_pattern_naming_eq nam1 nam2
+ | GCast (c1, t1), GCast (c2, t2) ->
+ f c1 c2 && cast_type_eq f t1 t2
+ | (GRef _ | GVar _ | GEvar _ | GPatVar _ | GApp _ | GLambda _ | GProd _ | GLetIn _ |
+ GCases _ | GLetTuple _ | GIf _ | GRec _ | GSort _ | GHole _ | GCast _), _ -> false
+
+let rec glob_constr_eq c = mk_glob_constr_eq glob_constr_eq c
+
+let map_glob_constr_left_to_right f = DAst.map (function
+ | GApp (g,args) ->
let comp1 = f g in
let comp2 = Util.List.map_left f args in
- GApp (loc,comp1,comp2)
- | GLambda (loc,na,bk,ty,c) ->
+ GApp (comp1,comp2)
+ | GLambda (na,bk,ty,c) ->
let comp1 = f ty in
let comp2 = f c in
- GLambda (loc,na,bk,comp1,comp2)
- | GProd (loc,na,bk,ty,c) ->
+ GLambda (na,bk,comp1,comp2)
+ | GProd (na,bk,ty,c) ->
let comp1 = f ty in
let comp2 = f c in
- GProd (loc,na,bk,comp1,comp2)
- | GLetIn (loc,na,b,c) ->
+ GProd (na,bk,comp1,comp2)
+ | GLetIn (na,b,t,c) ->
let comp1 = f b in
+ let compt = Option.map f t in
let comp2 = f c in
- GLetIn (loc,na,comp1,comp2)
- | GCases (loc,sty,rtntypopt,tml,pl) ->
+ GLetIn (na,comp1,compt,comp2)
+ | GCases (sty,rtntypopt,tml,pl) ->
let comp1 = Option.map f rtntypopt in
let comp2 = Util.List.map_left (fun (tm,x) -> (f tm,x)) tml in
- let comp3 = Util.List.map_left (fun (loc,idl,p,c) -> (loc,idl,p,f c)) pl in
- GCases (loc,sty,comp1,comp2,comp3)
- | GLetTuple (loc,nal,(na,po),b,c) ->
+ let comp3 = Util.List.map_left (CAst.map (fun (idl,p,c) -> (idl,p,f c))) pl in
+ GCases (sty,comp1,comp2,comp3)
+ | GLetTuple (nal,(na,po),b,c) ->
let comp1 = Option.map f po in
let comp2 = f b in
let comp3 = f c in
- GLetTuple (loc,nal,(na,comp1),comp2,comp3)
- | GIf (loc,c,(na,po),b1,b2) ->
+ GLetTuple (nal,(na,comp1),comp2,comp3)
+ | GIf (c,(na,po),b1,b2) ->
let comp1 = Option.map f po in
let comp2 = f b1 in
let comp3 = f b2 in
- GIf (loc,f c,(na,comp1),comp2,comp3)
- | GRec (loc,fk,idl,bl,tyl,bv) ->
+ GIf (f c,(na,comp1),comp2,comp3)
+ | GRec (fk,idl,bl,tyl,bv) ->
let comp1 = Array.map (Util.List.map_left (map_glob_decl_left_to_right f)) bl in
let comp2 = Array.map f tyl in
let comp3 = Array.map f bv in
- GRec (loc,fk,idl,comp1,comp2,comp3)
- | GCast (loc,c,k) ->
+ GRec (fk,idl,comp1,comp2,comp3)
+ | GCast (c,k) ->
let comp1 = f c in
let comp2 = Miscops.map_cast_type f k in
- GCast (loc,comp1,comp2)
+ GCast (comp1,comp2)
| (GVar _ | GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) as x -> x
+ )
let map_glob_constr = map_glob_constr_left_to_right
let fold_return_type f acc (na,tyopt) = Option.fold_left f acc tyopt
-let fold_glob_constr f acc = function
+let fold_glob_constr f acc = DAst.with_val (function
| GVar _ -> acc
- | GApp (_,c,args) -> List.fold_left f (f acc c) args
- | GLambda (_,_,_,b,c) | GProd (_,_,_,b,c) | GLetIn (_,_,b,c) ->
+ | GApp (c,args) -> List.fold_left f (f acc c) args
+ | GLambda (_,_,b,c) | GProd (_,_,b,c) ->
f (f acc b) c
- | GCases (_,_,rtntypopt,tml,pl) ->
- let fold_pattern acc (_,idl,p,c) = f acc c in
+ | GLetIn (_,b,t,c) ->
+ f (Option.fold_left f (f acc b) t) c
+ | GCases (_,rtntypopt,tml,pl) ->
+ let fold_pattern acc {CAst.v=(idl,p,c)} = f acc c in
List.fold_left fold_pattern
(List.fold_left f (Option.fold_left f acc rtntypopt) (List.map fst tml))
pl
- | GLetTuple (_,_,rtntyp,b,c) ->
+ | GLetTuple (_,rtntyp,b,c) ->
f (f (fold_return_type f acc rtntyp) b) c
- | GIf (_,c,rtntyp,b1,b2) ->
+ | GIf (c,rtntyp,b1,b2) ->
f (f (f (fold_return_type f acc rtntyp) c) b1) b2
- | GRec (_,_,_,bl,tyl,bv) ->
+ | GRec (_,_,bl,tyl,bv) ->
let acc = Array.fold_left
(List.fold_left (fun acc (na,k,bbd,bty) ->
f (Option.fold_left f acc bbd) bty)) acc bl in
Array.fold_left f (Array.fold_left f acc tyl) bv
- | GCast (_,c,k) ->
+ | GCast (c,k) ->
let acc = match k with
| CastConv t | CastVM t | CastNative t -> f acc t | CastCoerce -> acc in
f acc c
| (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> acc
+ )
-let iter_glob_constr f = fold_glob_constr (fun () -> f) ()
+let fold_return_type_with_binders f g v acc (na,tyopt) =
+ Option.fold_left (f (Name.fold_right g na v)) acc tyopt
-let same_id na id = match na with
-| Anonymous -> false
-| Name id' -> Id.equal id id'
+let fold_glob_constr_with_binders g f v acc = DAst.(with_val (function
+ | GVar _ -> acc
+ | GApp (c,args) -> List.fold_left (f v) (f v acc c) args
+ | GLambda (na,_,b,c) | GProd (na,_,b,c) ->
+ f (Name.fold_right g na v) (f v acc b) c
+ | GLetIn (na,b,t,c) ->
+ f (Name.fold_right g na v) (Option.fold_left (f v) (f v acc b) t) c
+ | GCases (_,rtntypopt,tml,pl) ->
+ let fold_pattern acc {v=(idl,p,c)} = f (List.fold_right g idl v) acc c in
+ let fold_tomatch (v',acc) (tm,(na,onal)) =
+ ((if rtntypopt = None then v' else
+ Option.fold_left (fun v'' {v=(_,nal)} -> List.fold_right (Name.fold_right g) nal v'')
+ (Name.fold_right g na v') onal),
+ f v acc tm) in
+ let (v',acc) = List.fold_left fold_tomatch (v,acc) tml in
+ let acc = Option.fold_left (f v') acc rtntypopt in
+ List.fold_left fold_pattern acc pl
+ | GLetTuple (nal,rtntyp,b,c) ->
+ f (List.fold_right (Name.fold_right g) nal v)
+ (f v (fold_return_type_with_binders f g v acc rtntyp) b) c
+ | GIf (c,rtntyp,b1,b2) ->
+ f v (f v (f v (fold_return_type_with_binders f g v acc rtntyp) c) b1) b2
+ | GRec (_,idl,bll,tyl,bv) ->
+ let v' = Array.fold_right g idl v in
+ let f' i acc fid =
+ let v,acc =
+ List.fold_left
+ (fun (v,acc) (na,k,bbd,bty) ->
+ (Name.fold_right g na v, f v (Option.fold_left (f v) acc bbd) bty))
+ (v,acc)
+ bll.(i) in
+ f v' (f v acc tyl.(i)) (bv.(i)) in
+ Array.fold_left_i f' acc idl
+ | GCast (c,k) ->
+ let acc = match k with
+ | CastConv t | CastVM t | CastNative t -> f v acc t | CastCoerce -> acc in
+ f v acc c
+ | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> acc))
+
+let iter_glob_constr f = fold_glob_constr (fun () -> f) ()
let occur_glob_constr id =
- let rec occur = function
- | GVar (loc,id') -> Id.equal id id'
- | GApp (loc,f,args) -> (occur f) || (List.exists occur args)
- | GLambda (loc,na,bk,ty,c) ->
- (occur ty) || (not (same_id na id) && (occur c))
- | GProd (loc,na,bk,ty,c) ->
- (occur ty) || (not (same_id na id) && (occur c))
- | GLetIn (loc,na,b,c) ->
- (occur b) || (not (same_id na id) && (occur c))
- | GCases (loc,sty,rtntypopt,tml,pl) ->
- (occur_option rtntypopt)
- || (List.exists (fun (tm,_) -> occur tm) tml)
- || (List.exists occur_pattern pl)
- | GLetTuple (loc,nal,rtntyp,b,c) ->
- occur_return_type rtntyp id
- || (occur b) || (not (List.mem_f Name.equal (Name id) nal) && (occur c))
- | GIf (loc,c,rtntyp,b1,b2) ->
- occur_return_type rtntyp id || (occur c) || (occur b1) || (occur b2)
- | GRec (loc,fk,idl,bl,tyl,bv) ->
- not (Array.for_all4 (fun fid bl ty bd ->
- let rec occur_fix = function
- [] -> not (occur ty) && (Id.equal fid id || not(occur bd))
- | (na,k,bbd,bty)::bl ->
- not (occur bty) &&
- (match bbd with
- Some bd -> not (occur bd)
- | _ -> true) &&
- (match na with Name id' -> Id.equal id id' | _ -> not (occur_fix bl)) in
- occur_fix bl)
- idl bl tyl bv)
- | GCast (loc,c,k) -> (occur c) || (match k with CastConv t
- | CastVM t | CastNative t -> occur t | CastCoerce -> false)
- | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> false
-
- and occur_pattern (loc,idl,p,c) = not (Id.List.mem id idl) && (occur c)
-
- and occur_option = function None -> false | Some p -> occur p
-
- and occur_return_type (na,tyopt) id = not (same_id na id) && occur_option tyopt
-
- in occur
-
-
-let add_name_to_ids set na =
- match na with
- | Anonymous -> set
- | Name id -> Id.Set.add id set
+ let rec occur barred acc c = match DAst.get c with
+ | GVar id' -> Id.equal id id'
+ | _ ->
+ (* [g] looks if [id] appears in a binding position, in which
+ case, we don't have to look in the corresponding subterm *)
+ let g id' barred = barred || Id.equal id id' in
+ let f barred acc c = acc || not barred && occur false acc c in
+ fold_glob_constr_with_binders g f barred acc c in
+ occur false false
let free_glob_vars =
- let rec vars bounded vs = function
- | GVar (loc,id') -> if Id.Set.mem id' bounded then vs else Id.Set.add id' vs
- | GApp (loc,f,args) -> List.fold_left (vars bounded) vs (f::args)
- | GLambda (loc,na,_,ty,c) | GProd (loc,na,_,ty,c) | GLetIn (loc,na,ty,c) ->
- let vs' = vars bounded vs ty in
- let bounded' = add_name_to_ids bounded na in
- vars bounded' vs' c
- | GCases (loc,sty,rtntypopt,tml,pl) ->
- let vs1 = vars_option bounded vs rtntypopt in
- let vs2 = List.fold_left (fun vs (tm,_) -> vars bounded vs tm) vs1 tml in
- List.fold_left (vars_pattern bounded) vs2 pl
- | GLetTuple (loc,nal,rtntyp,b,c) ->
- let vs1 = vars_return_type bounded vs rtntyp in
- let vs2 = vars bounded vs1 b in
- let bounded' = List.fold_left add_name_to_ids bounded nal in
- vars bounded' vs2 c
- | GIf (loc,c,rtntyp,b1,b2) ->
- let vs1 = vars_return_type bounded vs rtntyp in
- let vs2 = vars bounded vs1 c in
- let vs3 = vars bounded vs2 b1 in
- vars bounded vs3 b2
- | GRec (loc,fk,idl,bl,tyl,bv) ->
- let bounded' = Array.fold_right Id.Set.add idl bounded in
- let vars_fix i vs fid =
- let vs1,bounded1 =
- List.fold_left
- (fun (vs,bounded) (na,k,bbd,bty) ->
- let vs' = vars_option bounded vs bbd in
- let vs'' = vars bounded vs' bty in
- let bounded' = add_name_to_ids bounded na in
- (vs'',bounded')
- )
- (vs,bounded')
- bl.(i)
- in
- let vs2 = vars bounded1 vs1 tyl.(i) in
- vars bounded1 vs2 bv.(i)
- in
- Array.fold_left_i vars_fix vs idl
- | GCast (loc,c,k) -> let v = vars bounded vs c in
- (match k with CastConv t | CastVM t | CastNative t -> vars bounded v t | _ -> v)
- | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> vs
-
- and vars_pattern bounded vs (loc,idl,p,c) =
- let bounded' = List.fold_right Id.Set.add idl bounded in
- vars bounded' vs c
-
- and vars_option bounded vs = function None -> vs | Some p -> vars bounded vs p
-
- and vars_return_type bounded vs (na,tyopt) =
- let bounded' = add_name_to_ids bounded na in
- vars_option bounded' vs tyopt
- in
+ let rec vars bound vs c = match DAst.get c with
+ | GVar id' -> if Id.Set.mem id' bound then vs else Id.Set.add id' vs
+ | _ -> fold_glob_constr_with_binders Id.Set.add vars bound vs c in
fun rt ->
let vs = vars Id.Set.empty Id.Set.empty rt in
- Id.Set.elements vs
+ vs
let glob_visible_short_qualid c =
- let rec aux acc = function
- | GRef (_,c,_) ->
+ let rec aux acc c = match DAst.get c with
+ | GRef (c,_) ->
let qualid = Nametab.shortest_qualid_of_global Id.Set.empty c in
let dir,id = Libnames.repr_qualid qualid in
- if DirPath.is_empty dir then id :: acc else acc
- | c ->
+ if DirPath.is_empty dir then Id.Set.add id acc else acc
+ | _ ->
fold_glob_constr aux acc c
- in aux [] c
+ in aux Id.Set.empty c
let warn_variable_collision =
let open Pp in
CWarnings.create ~name:"variable-collision" ~category:"ltac"
(fun name ->
- strbrk "Collision between bound variables of name " ++ pr_id name)
+ strbrk "Collision between bound variables of name " ++ Id.print name)
let add_and_check_ident id set =
if Id.Set.mem id set then warn_variable_collision id;
Id.Set.add id set
let bound_glob_vars =
- let rec vars bound = function
- | GLambda (_,na,_,_,_) | GProd (_,na,_,_,_) | GLetIn (_,na,_,_) as c ->
- let bound = name_fold add_and_check_ident na bound in
- fold_glob_constr vars bound c
- | GCases (loc,sty,rtntypopt,tml,pl) ->
- let bound = vars_option bound rtntypopt in
- let bound =
- List.fold_left (fun bound (tm,_) -> vars bound tm) bound tml in
- List.fold_left vars_pattern bound pl
- | GLetTuple (loc,nal,rtntyp,b,c) ->
- let bound = vars_return_type bound rtntyp in
- let bound = vars bound b in
- let bound = List.fold_right (name_fold add_and_check_ident) nal bound in
- vars bound c
- | GIf (loc,c,rtntyp,b1,b2) ->
- let bound = vars_return_type bound rtntyp in
- let bound = vars bound c in
- let bound = vars bound b1 in
- vars bound b2
- | GRec (loc,fk,idl,bl,tyl,bv) ->
- let bound = Array.fold_right Id.Set.add idl bound in
- let vars_fix i bound fid =
- let bound =
- List.fold_left
- (fun bound (na,k,bbd,bty) ->
- let bound = vars_option bound bbd in
- let bound = vars bound bty in
- name_fold add_and_check_ident na bound
- )
- bound
- bl.(i)
- in
- let bound = vars bound tyl.(i) in
- vars bound bv.(i)
- in
- Array.fold_left_i vars_fix bound idl
- | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GVar _) -> bound
- | GApp _ | GCast _ as c -> fold_glob_constr vars bound c
-
- and vars_pattern bound (loc,idl,p,c) =
- let bound = List.fold_right add_and_check_ident idl bound in
- vars bound c
-
- and vars_option bound = function None -> bound | Some p -> vars bound p
-
- and vars_return_type bound (na,tyopt) =
- let bound = name_fold add_and_check_ident na bound in
- vars_option bound tyopt
+ let rec vars bound =
+ fold_glob_constr_with_binders
+ (fun id () -> bound := add_and_check_ident id !bound)
+ (fun () () -> vars bound)
+ () ()
in
fun rt ->
- vars Id.Set.empty rt
+ let bound = ref Id.Set.empty in
+ vars bound rt;
+ !bound
(** Mapping of names in binders *)
@@ -405,37 +331,38 @@ let bound_glob_vars =
probably be no significant penalty in doing reallocation as
pattern-matching expressions are usually rather small. *)
-let map_inpattern_binders f ((loc,id,nal) as x) =
+let map_inpattern_binders f ({loc;v=(id,nal)} as x) =
let r = CList.smartmap f nal in
if r == nal then x
- else loc,id,r
+ else CAst.make ?loc (id,r)
let map_tomatch_binders f ((c,(na,inp)) as x) : tomatch_tuple =
let r = Option.smartmap (fun p -> map_inpattern_binders f p) inp in
if r == inp then x
else c,(f na, r)
-let rec map_case_pattern_binders f = function
- | PatVar (loc,na) as x ->
+let rec map_case_pattern_binders f = DAst.map (function
+ | PatVar na as x ->
let r = f na in
if r == na then x
- else PatVar (loc,r)
- | PatCstr (loc,c,ps,na) as x ->
+ else PatVar r
+ | PatCstr (c,ps,na) as x ->
let rna = f na in
let rps =
CList.smartmap (fun p -> map_case_pattern_binders f p) ps
in
if rna == na && rps == ps then x
- else PatCstr(loc,c,rps,rna)
+ else PatCstr(c,rps,rna)
+ )
-let map_cases_branch_binders f ((loc,il,cll,rhs) as x) : cases_clause =
+let map_cases_branch_binders f ({CAst.loc;v=(il,cll,rhs)} as x) : cases_clause =
(* spiwack: not sure if I must do something with the list of idents.
It is intended to be a superset of the free variable of the
right-hand side, if I understand correctly. But I'm not sure when
or how they are used. *)
let r = List.smartmap (fun cl -> map_case_pattern_binders f cl) cll in
if r == cll then x
- else loc,il,r,rhs
+ else CAst.make ?loc (il,r,rhs)
let map_pattern_binders f tomatch branches =
CList.smartmap (fun tm -> map_tomatch_binders f tm) tomatch,
@@ -445,41 +372,28 @@ let map_pattern_binders f tomatch branches =
let map_tomatch f (c,pp) : tomatch_tuple = f c , pp
-let map_cases_branch f (loc,il,cll,rhs) : cases_clause =
- loc , il , cll , f rhs
+let map_cases_branch f =
+ CAst.map (fun (il,cll,rhs) -> (il , cll , f rhs))
let map_pattern f tomatch branches =
List.map (fun tm -> map_tomatch f tm) tomatch,
List.map (fun br -> map_cases_branch f br) branches
-let loc_of_glob_constr = function
- | GRef (loc,_,_) -> loc
- | GVar (loc,_) -> loc
- | GEvar (loc,_,_) -> loc
- | GPatVar (loc,_) -> loc
- | GApp (loc,_,_) -> loc
- | GLambda (loc,_,_,_,_) -> loc
- | GProd (loc,_,_,_,_) -> loc
- | GLetIn (loc,_,_,_) -> loc
- | GCases (loc,_,_,_,_) -> loc
- | GLetTuple (loc,_,_,_,_) -> loc
- | GIf (loc,_,_,_,_) -> loc
- | GRec (loc,_,_,_,_,_) -> loc
- | GSort (loc,_) -> loc
- | GHole (loc,_,_,_) -> loc
- | GCast (loc,_,_) -> loc
+let loc_of_glob_constr c = c.CAst.loc
(**********************************************************************)
(* Alpha-renaming *)
+exception UnsoundRenaming
+
let collide_id l id = List.exists (fun (id',id'') -> Id.equal id id' || Id.equal id id'') l
-let test_id l id = if collide_id l id then raise Not_found
-let test_na l na = name_iter (test_id l) na
+let test_id l id = if collide_id l id then raise UnsoundRenaming
+let test_na l na = Name.iter (test_id l) na
let update_subst na l =
let in_range id l = List.exists (fun (_,id') -> Id.equal id id') l in
- let l' = name_fold Id.List.remove_assoc na l in
- name_fold
+ let l' = Name.fold_right Id.List.remove_assoc na l in
+ Name.fold_right
(fun id _ ->
if in_range id l' then
let id' = Namegen.next_ident_away_from id (fun id' -> in_range id' l') in
@@ -487,8 +401,6 @@ let update_subst na l =
else na,l)
na (na,l)
-exception UnsoundRenaming
-
let rename_var l id =
try
let id' = Id.List.assoc id l in
@@ -499,77 +411,160 @@ let rename_var l id =
if List.exists (fun (_,id') -> Id.equal id id') l then raise UnsoundRenaming
else id
-let rec rename_glob_vars l = function
- | GVar (loc,id) as r ->
+let force c = DAst.make ?loc:c.CAst.loc (DAst.get c)
+
+let rec rename_glob_vars l c = force @@ DAst.map_with_loc (fun ?loc -> function
+ | GVar id as r ->
let id' = rename_var l id in
- if id == id' then r else GVar (loc,id')
- | GRef (_,VarRef id,_) as r ->
+ if id == id' then r else GVar id'
+ | GRef (VarRef id,_) as r ->
if List.exists (fun (_,id') -> Id.equal id id') l then raise UnsoundRenaming
else r
- | GProd (loc,na,bk,t,c) ->
+ | GProd (na,bk,t,c) ->
let na',l' = update_subst na l in
- GProd (loc,na,bk,rename_glob_vars l t,rename_glob_vars l' c)
- | GLambda (loc,na,bk,t,c) ->
+ GProd (na,bk,rename_glob_vars l t,rename_glob_vars l' c)
+ | GLambda (na,bk,t,c) ->
let na',l' = update_subst na l in
- GLambda (loc,na',bk,rename_glob_vars l t,rename_glob_vars l' c)
- | GLetIn (loc,na,b,c) ->
+ GLambda (na',bk,rename_glob_vars l t,rename_glob_vars l' c)
+ | GLetIn (na,b,t,c) ->
let na',l' = update_subst na l in
- GLetIn (loc,na',rename_glob_vars l b,rename_glob_vars l' c)
+ GLetIn (na',rename_glob_vars l b,Option.map (rename_glob_vars l) t,rename_glob_vars l' c)
(* Lazy strategy: we fail if a collision with renaming occurs, rather than renaming further *)
- | GCases (loc,ci,po,tomatchl,cls) ->
+ | GCases (ci,po,tomatchl,cls) ->
let test_pred_pat (na,ino) =
- test_na l na; Option.iter (fun (_,_,nal) -> List.iter (test_na l) nal) ino in
+ test_na l na; Option.iter (fun {v=(_,nal)} -> List.iter (test_na l) nal) ino in
let test_clause idl = List.iter (test_id l) idl in
let po = Option.map (rename_glob_vars l) po in
let tomatchl = Util.List.map_left (fun (tm,x) -> test_pred_pat x; (rename_glob_vars l tm,x)) tomatchl in
- let cls = Util.List.map_left (fun (loc,idl,p,c) -> test_clause idl; (loc,idl,p,rename_glob_vars l c)) cls in
- GCases (loc,ci,po,tomatchl,cls)
- | GLetTuple (loc,nal,(na,po),c,b) ->
+ let cls = Util.List.map_left (CAst.map (fun (idl,p,c) -> test_clause idl; (idl,p,rename_glob_vars l c))) cls in
+ GCases (ci,po,tomatchl,cls)
+ | GLetTuple (nal,(na,po),c,b) ->
List.iter (test_na l) (na::nal);
- GLetTuple (loc,nal,(na,Option.map (rename_glob_vars l) po),
+ GLetTuple (nal,(na,Option.map (rename_glob_vars l) po),
rename_glob_vars l c,rename_glob_vars l b)
- | GIf (loc,c,(na,po),b1,b2) ->
+ | GIf (c,(na,po),b1,b2) ->
test_na l na;
- GIf (loc,rename_glob_vars l c,(na,Option.map (rename_glob_vars l) po),
+ GIf (rename_glob_vars l c,(na,Option.map (rename_glob_vars l) po),
rename_glob_vars l b1,rename_glob_vars l b2)
- | GRec (loc,k,idl,decls,bs,ts) ->
+ | GRec (k,idl,decls,bs,ts) ->
Array.iter (test_id l) idl;
- GRec (loc,k,idl,
+ GRec (k,idl,
Array.map (List.map (fun (na,k,bbd,bty) ->
test_na l na; (na,k,Option.map (rename_glob_vars l) bbd,rename_glob_vars l bty))) decls,
Array.map (rename_glob_vars l) bs,
Array.map (rename_glob_vars l) ts)
- | r -> map_glob_constr (rename_glob_vars l) r
+ | _ -> DAst.get (map_glob_constr (rename_glob_vars l) c)
+ ) c
(**********************************************************************)
(* Conversion from glob_constr to cases pattern, if possible *)
-let rec cases_pattern_of_glob_constr na = function
- | GVar (loc,id) ->
+let is_gvar id c = match DAst.get c with
+| GVar id' -> Id.equal id id'
+| _ -> false
+
+let rec cases_pattern_of_glob_constr na = DAst.map (function
+ | GVar id ->
begin match na with
| Name _ ->
(* Unable to manage the presence of both an alias and a variable *)
raise Not_found
- | Anonymous -> PatVar (loc,Name id)
+ | Anonymous -> PatVar (Name id)
+ end
+ | GHole (_,_,_) -> PatVar na
+ | GRef (ConstructRef cstr,_) -> PatCstr (cstr,[],na)
+ | GApp (c, l) ->
+ begin match DAst.get c with
+ | GRef (ConstructRef cstr,_) ->
+ PatCstr (cstr,List.map (cases_pattern_of_glob_constr Anonymous) l,na)
+ | _ -> raise Not_found
end
- | GHole (loc,_,_,_) -> PatVar (loc,na)
- | GRef (loc,ConstructRef cstr,_) ->
- PatCstr (loc,cstr,[],na)
- | GApp (loc,GRef (_,ConstructRef cstr,_),l) ->
- PatCstr (loc,cstr,List.map (cases_pattern_of_glob_constr Anonymous) l,na)
+ | GLetIn (Name id as na',b,None,e) when is_gvar id e && na = Anonymous ->
+ (* A canonical encoding of aliases *)
+ DAst.get (cases_pattern_of_glob_constr na' b)
| _ -> raise Not_found
+ )
+
+open Declarations
+open Term
+open Context
+
+(* Keep only patterns which are not bound to a local definitions *)
+let drop_local_defs typi args =
+ let (decls,_) = decompose_prod_assum typi in
+ let rec aux decls args =
+ match decls, args with
+ | [], [] -> []
+ | Rel.Declaration.LocalDef _ :: decls, pat :: args ->
+ begin
+ match DAst.get pat with
+ | PatVar Anonymous -> aux decls args
+ | _ -> raise Not_found (* The pattern is used, one cannot drop it *)
+ end
+ | Rel.Declaration.LocalAssum _ :: decls, a :: args -> a :: aux decls args
+ | _ -> assert false in
+ aux (List.rev decls) args
+
+let add_patterns_for_params_remove_local_defs (ind,j) l =
+ let (mib,mip) = Global.lookup_inductive ind in
+ let nparams = mib.Declarations.mind_nparams in
+ let l =
+ if mip.mind_consnrealdecls.(j-1) = mip.mind_consnrealargs.(j-1) then
+ (* Optimisation *) l
+ else
+ let typi = mip.mind_nf_lc.(j-1) in
+ let (_,typi) = decompose_prod_n_assum (Rel.length mib.mind_params_ctxt) typi in
+ drop_local_defs typi l in
+ Util.List.addn nparams (DAst.make @@ PatVar Anonymous) l
+
+let add_alias ?loc na c =
+ match na with
+ | Anonymous -> c
+ | Name id -> GLetIn (na,DAst.make ?loc c,None,DAst.make ?loc (GVar id))
(* Turn a closed cases pattern into a glob_constr *)
-let rec glob_constr_of_closed_cases_pattern_aux = function
- | PatCstr (loc,cstr,[],Anonymous) ->
- GRef (loc,ConstructRef cstr,None)
- | PatCstr (loc,cstr,l,Anonymous) ->
- let ref = GRef (loc,ConstructRef cstr,None) in
- GApp (loc,ref, List.map glob_constr_of_closed_cases_pattern_aux l)
+let rec glob_constr_of_cases_pattern_aux isclosed x = DAst.map_with_loc (fun ?loc -> function
+ | PatCstr (cstr,[],na) -> add_alias ?loc na (GRef (ConstructRef cstr,None))
+ | PatCstr (cstr,l,na) ->
+ let ref = DAst.make ?loc @@ GRef (ConstructRef cstr,None) in
+ let l = add_patterns_for_params_remove_local_defs cstr l in
+ add_alias ?loc na (GApp (ref, List.map (glob_constr_of_cases_pattern_aux isclosed) l))
+ | PatVar (Name id) when not isclosed ->
+ GVar id
+ | PatVar Anonymous when not isclosed ->
+ GHole (Evar_kinds.QuestionMark (Define false,Anonymous),Misctypes.IntroAnonymous,None)
| _ -> raise Not_found
+ ) x
-let glob_constr_of_closed_cases_pattern = function
- | PatCstr (loc,cstr,l,na) ->
- na,glob_constr_of_closed_cases_pattern_aux (PatCstr (loc,cstr,l,Anonymous))
+let glob_constr_of_closed_cases_pattern p = match DAst.get p with
+ | PatCstr (cstr,l,na) ->
+ let loc = p.CAst.loc in
+ na,glob_constr_of_cases_pattern_aux true (DAst.make ?loc @@ PatCstr (cstr,l,Anonymous))
| _ ->
raise Not_found
+
+let glob_constr_of_cases_pattern p = glob_constr_of_cases_pattern_aux false p
+
+(**********************************************************************)
+(* Interpreting ltac variables *)
+
+open Pp
+open CErrors
+
+let ltac_interp_name { ltac_idents ; ltac_genargs } = function
+ | Anonymous -> Anonymous
+ | Name id as n ->
+ try Name (Id.Map.find id ltac_idents)
+ with Not_found ->
+ if Id.Map.mem id ltac_genargs then
+ user_err (str"Ltac variable"++spc()++ Id.print id ++
+ spc()++str"is not bound to an identifier."++spc()++
+ str"It cannot be used in a binder.")
+ else n
+
+let empty_lvar : ltac_var_map = {
+ ltac_constrs = Id.Map.empty;
+ ltac_uconstrs = Id.Map.empty;
+ ltac_idents = Id.Map.empty;
+ ltac_genargs = Id.Map.empty;
+}
diff --git a/pretyping/glob_ops.mli b/pretyping/glob_ops.mli
index 55e6b653..124440f5 100644
--- a/pretyping/glob_ops.mli
+++ b/pretyping/glob_ops.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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
@@ -11,21 +13,25 @@ open Glob_term
(** Equalities *)
-val cases_pattern_eq : cases_pattern -> cases_pattern -> bool
+val cases_pattern_eq : 'a cases_pattern_g -> 'a cases_pattern_g -> bool
+
+val alias_of_pat : 'a cases_pattern_g -> Name.t
+
+val set_pat_alias : Id.t -> 'a cases_pattern_g -> 'a cases_pattern_g
val cast_type_eq : ('a -> 'a -> bool) ->
'a Misctypes.cast_type -> 'a Misctypes.cast_type -> bool
-val glob_constr_eq : glob_constr -> glob_constr -> bool
+val glob_constr_eq : 'a glob_constr_g -> 'a glob_constr_g -> bool
(** Operations on [glob_constr] *)
-val cases_pattern_loc : cases_pattern -> Loc.t
+val cases_pattern_loc : 'a cases_pattern_g -> Loc.t option
-val cases_predicate_names : tomatch_tuples -> Name.t list
+val cases_predicate_names : 'a tomatch_tuples_g -> Name.t list
(** Apply one argument to a glob_constr *)
-val mkGApp : Loc.t -> glob_constr -> glob_constr -> glob_constr
+val mkGApp : ?loc:Loc.t -> 'a glob_constr_g -> 'a glob_constr_g -> 'a glob_constr_g
val map_glob_constr :
(glob_constr -> glob_constr) -> glob_constr -> glob_constr
@@ -36,13 +42,18 @@ val map_glob_constr_left_to_right :
val warn_variable_collision : ?loc:Loc.t -> Id.t -> unit
+val mk_glob_constr_eq : (glob_constr -> glob_constr -> bool) ->
+ glob_constr -> glob_constr -> bool
+
val fold_glob_constr : ('a -> glob_constr -> 'a) -> 'a -> glob_constr -> 'a
+val fold_glob_constr_with_binders : (Id.t -> 'a -> 'a) -> ('a -> 'b -> glob_constr -> 'b) -> 'a -> 'b -> glob_constr -> 'b
val iter_glob_constr : (glob_constr -> unit) -> glob_constr -> unit
-val occur_glob_constr : Id.t -> glob_constr -> bool
-val free_glob_vars : glob_constr -> Id.t list
+val occur_glob_constr : Id.t -> 'a glob_constr_g -> bool
+val free_glob_vars : 'a glob_constr_g -> Id.Set.t
val bound_glob_vars : glob_constr -> Id.Set.t
-val loc_of_glob_constr : glob_constr -> Loc.t
-val glob_visible_short_qualid : glob_constr -> Id.t list
+(* Obsolete *)
+val loc_of_glob_constr : 'a glob_constr_g -> Loc.t option
+val glob_visible_short_qualid : 'a glob_constr_g -> Id.Set.t
(* Renaming free variables using a renaming map; fails with
[UnsoundRenaming] if applying the renaming would introduce
@@ -52,14 +63,14 @@ val glob_visible_short_qualid : glob_constr -> Id.t list
exception UnsoundRenaming
val rename_var : (Id.t * Id.t) list -> Id.t -> Id.t
-val rename_glob_vars : (Id.t * Id.t) list -> glob_constr -> glob_constr
+val rename_glob_vars : (Id.t * Id.t) list -> 'a glob_constr_g -> 'a glob_constr_g
(** [map_pattern_binders f m c] applies [f] to all the binding names
in a pattern-matching expression ({!Glob_term.GCases}) represented
here by its relevant components [m] and [c]. It is used to
interpret Ltac-bound names both in pretyping and printing of
terms. *)
-val map_pattern_binders : (name -> name) ->
+val map_pattern_binders : (Name.t -> Name.t) ->
tomatch_tuples -> cases_clauses -> (tomatch_tuples*cases_clauses)
(** [map_pattern f m c] applies [f] to the return predicate and the
@@ -73,6 +84,15 @@ val map_pattern : (glob_constr -> glob_constr) ->
Take the current alias as parameter,
@raise Not_found if translation is impossible *)
-val cases_pattern_of_glob_constr : Name.t -> glob_constr -> cases_pattern
+val cases_pattern_of_glob_constr : Name.t -> 'a glob_constr_g -> 'a cases_pattern_g
+
+val glob_constr_of_closed_cases_pattern : 'a cases_pattern_g -> Name.t * 'a glob_constr_g
+
+(** A canonical encoding of cases pattern into constr such that
+ composed with [cases_pattern_of_glob_constr Anonymous] gives identity *)
+val glob_constr_of_cases_pattern : 'a cases_pattern_g -> 'a glob_constr_g
+
+val add_patterns_for_params_remove_local_defs : constructor -> 'a cases_pattern_g list -> 'a cases_pattern_g list
-val glob_constr_of_closed_cases_pattern : cases_pattern -> Name.t * glob_constr
+val ltac_interp_name : Ltac_pretype.ltac_var_map -> Name.t -> Name.t
+val empty_lvar : Ltac_pretype.ltac_var_map
diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml
index 39aeb41f..3143f8a5 100644
--- a/pretyping/indrec.ml
+++ b/pretyping/indrec.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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) *)
(************************************************************************)
(* File initially created by Christine Paulin, 1996 *)
@@ -18,6 +20,7 @@ open Libnames
open Globnames
open Nameops
open Term
+open Constr
open Vars
open Namegen
open Declarations
@@ -27,19 +30,30 @@ open Inductiveops
open Environ
open Reductionops
open Nametab
-open Sigma.Notations
open Context.Rel.Declaration
type dep_flag = bool
(* Errors related to recursors building *)
type recursion_scheme_error =
- | NotAllowedCaseAnalysis of (*isrec:*) bool * sorts * pinductive
+ | NotAllowedCaseAnalysis of (*isrec:*) bool * Sorts.t * pinductive
| NotMutualInScheme of inductive * inductive
| NotAllowedDependentAnalysis of (*isrec:*) bool * inductive
exception RecursionSchemeError of recursion_scheme_error
+let named_hd env t na = named_hd env Evd.empty (EConstr.of_constr t) na
+let name_assumption env = function
+| LocalAssum (na,t) -> LocalAssum (named_hd env t na, t)
+| LocalDef (na,c,t) -> LocalDef (named_hd env c na, c, t)
+
+let mkLambda_or_LetIn_name env d b = mkLambda_or_LetIn (name_assumption env d) b
+let mkProd_or_LetIn_name env d b = mkProd_or_LetIn (name_assumption env d) b
+let mkLambda_name env (n,a,b) = mkLambda_or_LetIn_name env (LocalAssum (n,a)) b
+let mkProd_name env (n,a,b) = mkProd_or_LetIn_name env (LocalAssum (n,a)) b
+let it_mkProd_or_LetIn_name env b l = List.fold_left (fun c d -> mkProd_or_LetIn_name env d c) b l
+let it_mkLambda_or_LetIn_name env b l = List.fold_left (fun c d -> mkLambda_or_LetIn_name env d c) b l
+
let make_prod_dep dep env = if dep then mkProd_name env else mkProd
let mkLambda_string s t c = mkLambda (Name (Id.of_string s), t, c)
@@ -55,7 +69,7 @@ let is_private mib =
let check_privacy_block mib =
if is_private mib then
- errorlabstrm ""(str"case analysis on a private inductive type")
+ user_err (str"case analysis on a private inductive type")
(**********************************************************************)
(* Building case analysis schemes *)
@@ -63,7 +77,7 @@ let check_privacy_block mib =
let mis_make_case_com dep env sigma (ind, u as pind) (mib,mip as specif) kind =
let lnamespar = Vars.subst_instance_context u mib.mind_params_ctxt in
- let indf = make_ind_family(pind, Context.Rel.to_extended_list 0 lnamespar) in
+ let indf = make_ind_family(pind, Context.Rel.to_extended_list mkRel 0 lnamespar) in
let constrs = get_constructors env indf in
let projs = get_projections env indf in
@@ -93,8 +107,8 @@ let mis_make_case_com dep env sigma (ind, u as pind) (mib,mip as specif) kind =
let pbody =
appvect
(mkRel (ndepar + nbprod),
- if dep then Context.Rel.to_extended_vect 0 deparsign
- else Context.Rel.to_extended_vect 1 arsign) in
+ if dep then Context.Rel.to_extended_vect mkRel 0 deparsign
+ else Context.Rel.to_extended_vect mkRel 1 arsign) in
let p =
it_mkLambda_or_LetIn_name env'
((if dep then mkLambda_name env' else mkLambda)
@@ -118,18 +132,19 @@ let mis_make_case_com dep env sigma (ind, u as pind) (mib,mip as specif) kind =
it_mkLambda_or_LetIn_name env' obj deparsign
else
let cs = lift_constructor (k+1) constrs.(k) in
- let t = build_branch_type env dep (mkRel (k+1)) cs in
+ let t = build_branch_type env sigma dep (mkRel (k+1)) cs in
mkLambda_string "f" t
(add_branch (push_rel (LocalAssum (Anonymous, t)) env) (k+1))
in
- let Sigma (s, sigma, p) = Sigma.fresh_sort_in_family ~rigid:Evd.univ_flexible_alg env sigma kind in
- let typP = make_arity env' dep indf s in
+ let (sigma, s) = Evd.fresh_sort_in_family ~rigid:Evd.univ_flexible_alg env sigma kind in
+ let typP = make_arity env' sigma dep indf s in
+ let typP = EConstr.Unsafe.to_constr typP in
let c =
it_mkLambda_or_LetIn_name env
(mkLambda_string "P" typP
(add_branch (push_rel (LocalAssum (Anonymous,typP)) env') 0)) lnamespar
in
- Sigma (c, sigma, p)
+ (sigma, c)
(* check if the type depends recursively on one of the inductive scheme *)
@@ -153,8 +168,10 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
let nparams = List.length vargs in
let process_pos env depK pk =
let rec prec env i sign p =
- let p',largs = whd_allnolet_stack env sigma p in
- match kind_of_term p' with
+ let p',largs = whd_allnolet_stack env sigma (EConstr.of_constr p) in
+ let p' = EConstr.Unsafe.to_constr p' in
+ let largs = List.map EConstr.Unsafe.to_constr largs in
+ match kind p' with
| Prod (n,t,c) ->
let d = LocalAssum (n,t) in
make_prod env (n,t,prec (push_rel d env) (i+1) (d::sign) c)
@@ -166,18 +183,19 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
let base = applist (lift i pk,realargs) in
if depK then
Reduction.beta_appvect
- base [|applist (mkRel (i+1), Context.Rel.to_extended_list 0 sign)|]
+ base [|applist (mkRel (i+1), Context.Rel.to_extended_list mkRel 0 sign)|]
else
base
| _ ->
- let t' = whd_all env sigma p in
- if Term.eq_constr p' t' then assert false
+ let t' = whd_all env sigma (EConstr.of_constr p) in
+ let t' = EConstr.Unsafe.to_constr t' in
+ if Constr.equal p' t' then assert false
else prec env i sign t'
in
prec env 0 []
in
let rec process_constr env i c recargs nhyps li =
- if nhyps > 0 then match kind_of_term c with
+ if nhyps > 0 then match kind c with
| Prod (n,t,c_0) ->
let (optionpos,rest) =
match recargs with
@@ -229,8 +247,10 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
let process_pos env fk =
let rec prec env i hyps p =
- let p',largs = whd_allnolet_stack env sigma p in
- match kind_of_term p' with
+ let p',largs = whd_allnolet_stack env sigma (EConstr.of_constr p) in
+ let p' = EConstr.Unsafe.to_constr p' in
+ let largs = List.map EConstr.Unsafe.to_constr largs in
+ match kind p' with
| Prod (n,t,c) ->
let d = LocalAssum (n,t) in
mkLambda_name env (n,t,prec (push_rel d env) (i+1) (d::hyps) c)
@@ -239,11 +259,12 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
mkLetIn (n,b,t,prec (push_rel d env) (i+1) (d::hyps) c)
| Ind _ ->
let realargs = List.skipn nparrec largs
- and arg = appvect (mkRel (i+1), Context.Rel.to_extended_vect 0 hyps) in
+ and arg = appvect (mkRel (i+1), Context.Rel.to_extended_vect mkRel 0 hyps) in
applist(lift i fk,realargs@[arg])
| _ ->
- let t' = whd_all env sigma p in
- if Term.eq_constr t' p' then assert false
+ let t' = whd_all env sigma (EConstr.of_constr p) in
+ let t' = EConstr.Unsafe.to_constr t' in
+ if Constr.equal t' p' then assert false
else prec env i hyps t'
in
prec env 0 []
@@ -261,7 +282,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
| None ->
mkLambda_name env
(n,t,process_constr (push_rel d env) (i+1)
- (whd_beta Evd.empty (applist (lift 1 f, [(mkRel 1)])))
+ (EConstr.Unsafe.to_constr (whd_beta Evd.empty (EConstr.of_constr (applist (lift 1 f, [(mkRel 1)])))))
(cprest,rest))
| Some(_,f_0) ->
let nF = lift (i+1+decF) f_0 in
@@ -269,7 +290,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
let arg = process_pos env' nF (lift 1 t) in
mkLambda_name env
(n,t,process_constr env' (i+1)
- (whd_beta Evd.empty (applist (lift 1 f, [(mkRel 1); arg])))
+ (EConstr.Unsafe.to_constr (whd_beta Evd.empty (EConstr.of_constr (applist (lift 1 f, [(mkRel 1); arg])))))
(cprest,rest)))
| (LocalDef (n,c,t) as d)::cprest, rest ->
mkLetIn
@@ -277,13 +298,13 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
process_constr (push_rel d env) (i+1) (lift 1 f)
(cprest,rest))
| [],[] -> f
- | _,[] | [],_ -> anomaly (Pp.str "process_constr")
+ | _,[] | [],_ -> anomaly (Pp.str "process_constr.")
in
process_constr env 0 f (List.rev cstr.cs_args, recargs)
(* Main function *)
-let mis_make_indrec env sigma listdepkind mib u =
+let mis_make_indrec env sigma ?(force_mutual=false) listdepkind mib u =
let nparams = mib.mind_nparams in
let nparrec = mib.mind_nparams_rec in
let evdref = ref sigma in
@@ -317,7 +338,7 @@ let mis_make_indrec env sigma listdepkind mib u =
(* arity in the context of the fixpoint, i.e.
P1..P_nrec f1..f_nbconstruct *)
- let args = Context.Rel.to_extended_list (nrec+nbconstruct) lnamesparrec in
+ let args = Context.Rel.to_extended_list mkRel (nrec+nbconstruct) lnamesparrec in
let indf = make_ind_family((indi,u),args) in
let arsign,_ = get_arity env indf in
@@ -331,15 +352,15 @@ let mis_make_indrec env sigma listdepkind mib u =
(* constructors in context of the Cases expr, i.e.
P1..P_nrec f1..f_nbconstruct F_1..F_nrec a_1..a_nar x:I *)
- let args' = Context.Rel.to_extended_list (dect+nrec) lnamesparrec in
- let args'' = Context.Rel.to_extended_list ndepar lnonparrec in
+ let args' = Context.Rel.to_extended_list mkRel (dect+nrec) lnamesparrec in
+ let args'' = Context.Rel.to_extended_list mkRel ndepar lnonparrec in
let indf' = make_ind_family((indi,u),args'@args'') in
let branches =
let constrs = get_constructors env indf' in
let fi = Termops.rel_vect (dect-i-nctyi) nctyi in
let vecfi = Array.map
- (fun f -> appvect (f, Context.Rel.to_extended_vect ndepar lnonparrec))
+ (fun f -> appvect (f, Context.Rel.to_extended_vect mkRel ndepar lnonparrec))
fi
in
Array.map3
@@ -360,9 +381,9 @@ let mis_make_indrec env sigma listdepkind mib u =
let deparsign' = LocalAssum (Anonymous,depind')::arsign' in
let pargs =
- let nrpar = Context.Rel.to_extended_list (2*ndepar) lnonparrec
- and nrar = if dep then Context.Rel.to_extended_list 0 deparsign'
- else Context.Rel.to_extended_list 1 arsign'
+ let nrpar = Context.Rel.to_extended_list mkRel (2*ndepar) lnonparrec
+ and nrar = if dep then Context.Rel.to_extended_list mkRel 0 deparsign'
+ else Context.Rel.to_extended_list mkRel 1 arsign'
in nrpar@nrar
in
@@ -378,9 +399,10 @@ let mis_make_indrec env sigma listdepkind mib u =
arsign'
in
let obj =
- Inductiveops.make_case_or_project env indf ci pred
- (mkRel 1) branches
+ Inductiveops.make_case_or_project env !evdref indf ci (EConstr.of_constr pred)
+ (EConstr.mkRel 1) (Array.map EConstr.of_constr branches)
in
+ let obj = EConstr.to_constr !evdref obj in
it_mkLambda_or_LetIn_name env obj
(Termops.lift_rel_context nrec deparsign)
in
@@ -389,8 +411,8 @@ let mis_make_indrec env sigma listdepkind mib u =
let typtyi =
let concl =
- let pargs = if dep then Context.Rel.to_extended_vect 0 deparsign
- else Context.Rel.to_extended_vect 1 arsign
+ let pargs = if dep then Context.Rel.to_extended_vect mkRel 0 deparsign
+ else Context.Rel.to_extended_vect mkRel 1 arsign
in appvect (mkRel (nbconstruct+ndepar+nonrecpar+j),pargs)
in it_mkProd_or_LetIn_name env
concl
@@ -417,7 +439,7 @@ let mis_make_indrec env sigma listdepkind mib u =
else
let recarg = (dest_subterms recargsvec.(tyi)).(j) in
let recarg = recargpar@recarg in
- let vargs = Context.Rel.to_extended_list (nrec+i+j) lnamesparrec in
+ let vargs = Context.Rel.to_extended_list mkRel (nrec+i+j) lnamesparrec in
let cs = get_constructor ((indi,u),mibi,mipi,vargs) (j+1) in
let p_0 =
type_rec_branch
@@ -431,12 +453,13 @@ let mis_make_indrec env sigma listdepkind mib u =
in
let rec put_arity env i = function
| ((indi,u),_,_,dep,kinds)::rest ->
- let indf = make_ind_family ((indi,u), Context.Rel.to_extended_list i lnamesparrec) in
+ let indf = make_ind_family ((indi,u), Context.Rel.to_extended_list mkRel i lnamesparrec) in
let s =
Evarutil.evd_comb1 (Evd.fresh_sort_in_family ~rigid:Evd.univ_flexible_alg env)
evdref kinds
in
- let typP = make_arity env dep indf s in
+ let typP = make_arity env !evdref dep indf s in
+ let typP = EConstr.Unsafe.to_constr typP in
mkLambda_string "P" typP
(put_arity (push_rel (LocalAssum (Anonymous,typP)) env) (i+1) rest)
| [] ->
@@ -446,7 +469,7 @@ let mis_make_indrec env sigma listdepkind mib u =
(* Body on make_one_rec *)
let ((indi,u),mibi,mipi,dep,kind) = List.nth listdepkind p in
- if (mis_is_recursive_subset
+ if force_mutual || (mis_is_recursive_subset
(List.map (fun ((indi,u),_,_,_,_) -> snd indi) listdepkind)
mipi.mind_recargs)
then
@@ -454,10 +477,9 @@ let mis_make_indrec env sigma listdepkind mib u =
it_mkLambda_or_LetIn_name env (put_arity env' 0 listdepkind)
lnamesparrec
else
- let sigma = Sigma.Unsafe.of_evar_map !evdref in
- let Sigma (c, sigma, _) = mis_make_case_com dep env sigma (indi,u) (mibi,mipi) kind in
- let evd' = Sigma.to_evar_map sigma in
- evdref := evd'; c
+ let evd = !evdref in
+ let (evd, c) = mis_make_case_com dep env evd (indi,u) (mibi,mipi) kind in
+ evdref := evd; c
in
(* Body of mis_make_indrec *)
!evdref, List.init nrec make_one_rec
@@ -486,7 +508,7 @@ let build_case_analysis_scheme_default env sigma pity kind =
[rec] by [s] *)
let change_sort_arity sort =
- let rec drec a = match kind_of_term a with
+ let rec drec a = match kind a with
| Cast (c,_,_) -> drec c
| Prod (n,t,c) -> let s, c' = drec c in s, mkProd (n, t, c')
| LetIn (n,b,t,c) -> let s, c' = drec c in s, mkLetIn (n,b,t,c')
@@ -500,7 +522,7 @@ let change_sort_arity sort =
let weaken_sort_scheme env evd set sort npars term ty =
let evdref = ref evd in
let rec drec np elim =
- match kind_of_term elim with
+ match kind elim with
| Prod (n,t,c) ->
if Int.equal np 0 then
let osort, t' = change_sort_arity sort t in
@@ -512,7 +534,7 @@ let weaken_sort_scheme env evd set sort npars term ty =
mkProd (n, t, c'), mkLambda (n, t, term')
| LetIn (n,b,t,c) -> let c',term' = drec np c in
mkLetIn (n,b,t,c'), mkLetIn (n,b,t,term')
- | _ -> anomaly ~label:"weaken_sort_scheme" (Pp.str "wrong elimination type")
+ | _ -> anomaly ~label:"weaken_sort_scheme" (Pp.str "wrong elimination type.")
in
let ty, term = drec npars ty in
!evdref, ty, term
@@ -536,7 +558,7 @@ let check_arities env listdepkind =
[] listdepkind
in true
-let build_mutual_induction_scheme env sigma = function
+let build_mutual_induction_scheme env sigma ?(force_mutual=false) = function
| ((mind,u),dep,s)::lrecspec ->
let (mib,mip) = lookup_mind_specif env mind in
if dep && not (Inductiveops.has_dependent_elim mib) then
@@ -547,7 +569,7 @@ let build_mutual_induction_scheme env sigma = function
(List.map
(function ((mind',u'),dep',s') ->
let (sp',_) = mind' in
- if eq_mind sp sp' then
+ if MutInd.equal sp sp' then
let (mibi',mipi') = lookup_mind_specif env mind' in
((mind',u'),mibi',mipi',dep',s')
else
@@ -555,8 +577,8 @@ let build_mutual_induction_scheme env sigma = function
lrecspec)
in
let _ = check_arities env listdepkind in
- mis_make_indrec env sigma listdepkind mib u
- | _ -> anomaly (Pp.str "build_induction_scheme expects a non empty list of inductive types")
+ mis_make_indrec env sigma ~force_mutual listdepkind mib u
+ | _ -> anomaly (Pp.str "build_induction_scheme expects a non empty list of inductive types.")
let build_induction_scheme env sigma pind dep kind =
let (mib,mip) = lookup_mind_specif env (fst pind) in
@@ -586,7 +608,7 @@ let lookup_eliminator ind_sp s =
(* Try first to get an eliminator defined in the same section as the *)
(* inductive type *)
try
- let cst =Global.constant_of_delta_kn (make_kn mp dp (Label.of_id id)) in
+ let cst =Global.constant_of_delta_kn (KerName.make mp dp (Label.of_id id)) in
let _ = Global.lookup_constant cst in
ConstRef cst
with Not_found ->
@@ -594,9 +616,9 @@ let lookup_eliminator ind_sp s =
(* using short name (e.g. for "eq_rec") *)
try Nametab.locate (qualid_of_ident id)
with Not_found ->
- errorlabstrm "default_elim"
+ user_err ~hdr:"default_elim"
(strbrk "Cannot find the elimination combinator " ++
- pr_id id ++ strbrk ", the elimination of the inductive definition " ++
+ Id.print id ++ strbrk ", the elimination of the inductive definition " ++
pr_global_env Id.Set.empty (IndRef ind_sp) ++
strbrk " on sort " ++ Termops.pr_sort_family s ++
strbrk " is probably not allowed.")
diff --git a/pretyping/indrec.mli b/pretyping/indrec.mli
index 192b64a5..7a68443b 100644
--- a/pretyping/indrec.mli
+++ b/pretyping/indrec.mli
@@ -1,20 +1,22 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
+open Constr
open Environ
open Evd
(** Errors related to recursors building *)
type recursion_scheme_error =
- | NotAllowedCaseAnalysis of (*isrec:*) bool * sorts * pinductive
+ | NotAllowedCaseAnalysis of (*isrec:*) bool * Sorts.t * pinductive
| NotMutualInScheme of inductive * inductive
| NotAllowedDependentAnalysis of (*isrec:*) bool * inductive
@@ -26,26 +28,27 @@ type dep_flag = bool
(** Build a case analysis elimination scheme in some sort family *)
-val build_case_analysis_scheme : env -> 'r Sigma.t -> pinductive ->
- dep_flag -> sorts_family -> (constr, 'r) Sigma.sigma
+val build_case_analysis_scheme : env -> Evd.evar_map -> pinductive ->
+ dep_flag -> Sorts.family -> evar_map * Constr.t
(** Build a dependent case elimination predicate unless type is in Prop
or is a recursive record with primitive projections. *)
-val build_case_analysis_scheme_default : env -> 'r Sigma.t -> pinductive ->
- sorts_family -> (constr, 'r) Sigma.sigma
+val build_case_analysis_scheme_default : env -> evar_map -> pinductive ->
+ Sorts.family -> evar_map * Constr.t
(** Builds a recursive induction scheme (Peano-induction style) in the same
sort family as the inductive family; it is dependent if not in Prop
or a recursive record with primitive projections. *)
val build_induction_scheme : env -> evar_map -> pinductive ->
- dep_flag -> sorts_family -> evar_map * constr
+ dep_flag -> Sorts.family -> evar_map * constr
(** Builds mutual (recursive) induction schemes *)
val build_mutual_induction_scheme :
- env -> evar_map -> (pinductive * dep_flag * sorts_family) list -> evar_map * constr list
+ env -> evar_map -> ?force_mutual:bool ->
+ (pinductive * dep_flag * Sorts.family) list -> evar_map * constr list
(** Scheme combinators *)
@@ -54,13 +57,13 @@ val build_mutual_induction_scheme :
scheme quantified on sort [s]. [set] asks for [s] be declared equal to [i],
otherwise just less or equal to [i]. *)
-val weaken_sort_scheme : env -> evar_map -> bool -> sorts -> int -> constr -> types ->
+val weaken_sort_scheme : env -> evar_map -> bool -> Sorts.t -> int -> constr -> types ->
evar_map * types * constr
(** Recursor names utilities *)
-val lookup_eliminator : inductive -> sorts_family -> Globnames.global_reference
-val elimination_suffix : sorts_family -> string
-val make_elimination_ident : Id.t -> sorts_family -> Id.t
+val lookup_eliminator : inductive -> Sorts.family -> Globnames.global_reference
+val elimination_suffix : Sorts.family -> string
+val make_elimination_ident : Id.t -> Sorts.family -> Id.t
val case_suffix : string
diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml
index 214e19fe..97aa82e4 100644
--- a/pretyping/inductiveops.ml
+++ b/pretyping/inductiveops.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 CErrors
@@ -11,6 +13,7 @@ open Util
open Names
open Univ
open Term
+open Constr
open Vars
open Termops
open Declarations
@@ -24,14 +27,14 @@ open Context.Rel.Declaration
let type_of_inductive env (ind,u) =
let (mib,_ as specif) = Inductive.lookup_mind_specif env ind in
- Typeops.check_hyps_inclusion env (mkInd ind) mib.mind_hyps;
+ Typeops.check_hyps_inclusion env mkInd ind mib.mind_hyps;
Inductive.type_of_inductive env (specif,u)
(* Return type as quoted by the user *)
let type_of_constructor env (cstr,u) =
let (mib,_ as specif) =
Inductive.lookup_mind_specif env (inductive_of_constructor cstr) in
- Typeops.check_hyps_inclusion env (mkConstruct cstr) mib.mind_hyps;
+ Typeops.check_hyps_inclusion env mkConstruct cstr mib.mind_hyps;
Inductive.type_of_constructor (cstr,u) specif
(* Return constructor types in user form *)
@@ -58,21 +61,24 @@ let lift_inductive_family n = liftn_inductive_family n 1
let substnl_ind_family l n = map_ind_family (substnl l n)
-type inductive_type = IndType of inductive_family * constr list
+type inductive_type = IndType of inductive_family * EConstr.constr list
let make_ind_type (indf, realargs) = IndType (indf,realargs)
let dest_ind_type (IndType (indf,realargs)) = (indf,realargs)
let map_inductive_type f (IndType (indf, realargs)) =
- IndType (map_ind_family f indf, List.map f realargs)
+ let f' c = EConstr.Unsafe.to_constr (f (EConstr.of_constr c)) in
+ IndType (map_ind_family f' indf, List.map f realargs)
-let liftn_inductive_type n d = map_inductive_type (liftn n d)
+let liftn_inductive_type n d = map_inductive_type (EConstr.Vars.liftn n d)
let lift_inductive_type n = liftn_inductive_type n 1
-let substnl_ind_type l n = map_inductive_type (substnl l n)
+let substnl_ind_type l n = map_inductive_type (EConstr.Vars.substnl l n)
let mkAppliedInd (IndType ((ind,params), realargs)) =
- applist (mkIndU ind,params@realargs)
+ let open EConstr in
+ let ind = on_snd EInstance.make ind in
+ applist (mkIndU ind, (List.map EConstr.of_constr params)@realargs)
(* Does not consider imbricated or mutually recursive types *)
let mis_is_recursive_subset listind rarg =
@@ -80,7 +86,7 @@ let mis_is_recursive_subset listind rarg =
List.exists
(fun ra ->
match dest_recarg ra with
- | Mrec (_,i) -> Int.List.mem i listind
+ | Mrec (_,i) -> Int.List.mem i listind
| _ -> false) rvec
in
Array.exists one_is_rec (dest_subterms rarg)
@@ -94,7 +100,7 @@ let mis_nf_constructor_type ((ind,u),mib,mip) j =
and ntypes = mib.mind_ntypes
and nconstr = Array.length mip.mind_consnames in
let make_Ik k = mkIndU (((fst ind),ntypes-k-1),u) in
- if j > nconstr then error "Not enough constructors in the type.";
+ if j > nconstr then user_err Pp.(str "Not enough constructors in the type.");
substl (List.init ntypes make_Ik) (subst_instance_constr u specif.(j-1))
(* Number of constructors *)
@@ -271,7 +277,7 @@ let projection_nparams p = projection_nparams_env (Global.env ()) p
let has_dependent_elim mib =
match mib.mind_record with
- | Some (Some _) -> mib.mind_finite == Decl_kinds.BiFinite
+ | Some (Some _) -> mib.mind_finite == BiFinite
| _ -> true
(* Annotation for cases *)
@@ -343,34 +349,35 @@ let get_projections env (ind,params) =
| Some (Some (id, projs, pbs)) -> Some projs
| _ -> None
-let make_case_or_project env indf ci pred c branches =
+let make_case_or_project env sigma indf ci pred c branches =
+ let open EConstr in
let projs = get_projections env indf in
match projs with
| None -> (mkCase (ci, pred, c, branches))
| Some ps ->
assert(Array.length branches == 1);
let () =
- let _, _, t = destLambda pred in
+ let _, _, t = destLambda sigma pred in
let (ind, _), _ = dest_ind_family indf in
let mib, _ = Inductive.lookup_mind_specif env ind in
- if (* dependent *) not (noccurn 1 t) &&
+ if (* dependent *) not (Vars.noccurn sigma 1 t) &&
not (has_dependent_elim mib) then
- errorlabstrm "make_case_or_project"
- Pp.(str"Dependent case analysis not allowed" ++
- str" on inductive type " ++ Names.MutInd.print (fst ind))
+ user_err ~hdr:"make_case_or_project"
+ Pp.(str"Dependent case analysis not allowed" ++
+ str" on inductive type " ++ Names.MutInd.print (fst ind))
in
let branch = branches.(0) in
- let ctx, br = decompose_lam_n_assum (Array.length ps) branch in
+ let ctx, br = decompose_lam_n_assum sigma (Array.length ps) branch in
let n, subst =
List.fold_right
(fun decl (i, subst) ->
- match decl with
- | LocalAssum (na, t) ->
- let t = mkProj (Projection.make ps.(i) true, c) in
- (i + 1, t :: subst)
- | LocalDef (na, b, t) -> (i, substl subst b :: subst))
- ctx (0, [])
- in substl subst br
+ match decl with
+ | LocalAssum (na, t) ->
+ let t = mkProj (Projection.make ps.(i) true, c) in
+ (i + 1, t :: subst)
+ | LocalDef (na, b, t) -> (i, Vars.substl subst b :: subst))
+ ctx (0, [])
+ in Vars.substl subst br
(* substitution in a signature *)
@@ -393,8 +400,8 @@ let get_arity env ((ind,u),params) =
mib.mind_params_ctxt
else begin
assert (Int.equal nparams mib.mind_nparams_rec);
- let nnonrecparamdecls = List.length mib.mind_params_ctxt - mib.mind_nparams_rec in
- snd (List.chop nnonrecparamdecls mib.mind_params_ctxt)
+ let nnonrecparamdecls = mib.mind_nparams - mib.mind_nparams_rec in
+ snd (Termops.context_chop nnonrecparamdecls mib.mind_params_ctxt)
end in
let parsign = Vars.subst_instance_context u parsign in
let arproperlength = List.length mip.mind_arity_ctxt - List.length parsign in
@@ -408,51 +415,55 @@ let build_dependent_constructor cs =
applist
(mkConstructU cs.cs_cstr,
(List.map (lift cs.cs_nargs) cs.cs_params)
- @(Context.Rel.to_extended_list 0 cs.cs_args))
+ @(Context.Rel.to_extended_list mkRel 0 cs.cs_args))
let build_dependent_inductive env ((ind, params) as indf) =
let arsign,_ = get_arity env indf in
let nrealargs = List.length arsign in
applist
(mkIndU ind,
- (List.map (lift nrealargs) params)@(Context.Rel.to_extended_list 0 arsign))
+ (List.map (lift nrealargs) params)@(Context.Rel.to_extended_list mkRel 0 arsign))
(* builds the arity of an elimination predicate in sort [s] *)
-let make_arity_signature env dep indf =
+let make_arity_signature env sigma dep indf =
let (arsign,_) = get_arity env indf in
+ let arsign = List.map (fun d -> Termops.map_rel_decl EConstr.of_constr d) arsign in
if dep then
(* We need names everywhere *)
- Namegen.name_context env
- ((LocalAssum (Anonymous,build_dependent_inductive env indf))::arsign)
+ Namegen.name_context env sigma
+ ((LocalAssum (Anonymous,EConstr.of_constr (build_dependent_inductive env indf)))::arsign)
(* Costly: would be better to name once for all at definition time *)
else
(* No need to enforce names *)
arsign
-let make_arity env dep indf s = mkArity (make_arity_signature env dep indf, s)
+let make_arity env sigma dep indf s =
+ let open EConstr in
+ it_mkProd_or_LetIn (mkSort s) (make_arity_signature env sigma dep indf)
(* [p] is the predicate and [cs] a constructor summary *)
-let build_branch_type env dep p cs =
+let build_branch_type env sigma dep p cs =
let base = appvect (lift cs.cs_nargs p, cs.cs_concl_realargs) in
if dep then
- Namegen.it_mkProd_or_LetIn_name env
- (applist (base,[build_dependent_constructor cs]))
- cs.cs_args
+ EConstr.Unsafe.to_constr (Namegen.it_mkProd_or_LetIn_name env sigma
+ (EConstr.of_constr (applist (base,[build_dependent_constructor cs])))
+ (List.map (fun d -> Termops.map_rel_decl EConstr.of_constr d) cs.cs_args))
else
- it_mkProd_or_LetIn base cs.cs_args
+ Term.it_mkProd_or_LetIn base cs.cs_args
(**************************************************)
-let extract_mrectype t =
- let (t, l) = decompose_app t in
- match kind_of_term t with
+let extract_mrectype sigma t =
+ let open EConstr in
+ let (t, l) = decompose_app sigma t in
+ match EConstr.kind sigma t with
| Ind ind -> (ind, l)
| _ -> raise Not_found
let find_mrectype_vect env sigma c =
- let (t, l) = decompose_appvect (whd_all env sigma c) in
- match kind_of_term t with
+ let (t, l) = Termops.decompose_app_vect sigma (whd_all env sigma c) in
+ match EConstr.kind sigma t with
| Ind ind -> (ind, l)
| _ -> raise Not_found
@@ -460,28 +471,35 @@ let find_mrectype env sigma c =
let (ind, v) = find_mrectype_vect env sigma c in (ind, Array.to_list v)
let find_rectype env sigma c =
- let (t, l) = decompose_app (whd_all env sigma c) in
- match kind_of_term t with
- | Ind (ind,u as indu) ->
+ let open EConstr in
+ let (t, l) = decompose_app sigma (whd_all env sigma c) in
+ match EConstr.kind sigma t with
+ | Ind (ind,u) ->
let (mib,mip) = Inductive.lookup_mind_specif env ind in
if mib.mind_nparams > List.length l then raise Not_found;
+ let l = List.map EConstr.Unsafe.to_constr l in
let (par,rargs) = List.chop mib.mind_nparams l in
- IndType((indu, par),rargs)
+ let indu = (ind, EInstance.kind sigma u) in
+ IndType((indu, par),List.map EConstr.of_constr rargs)
| _ -> raise Not_found
let find_inductive env sigma c =
- let (t, l) = decompose_app (whd_all env sigma c) in
- match kind_of_term t with
+ let open EConstr in
+ let (t, l) = decompose_app sigma (whd_all env sigma c) in
+ match EConstr.kind sigma t with
| Ind ind
- when (fst (Inductive.lookup_mind_specif env (fst ind))).mind_finite <> Decl_kinds.CoFinite ->
+ when (fst (Inductive.lookup_mind_specif env (fst ind))).mind_finite <> CoFinite ->
+ let l = List.map EConstr.Unsafe.to_constr l in
(ind, l)
| _ -> raise Not_found
let find_coinductive env sigma c =
- let (t, l) = decompose_app (whd_all env sigma c) in
- match kind_of_term t with
+ let open EConstr in
+ let (t, l) = decompose_app sigma (whd_all env sigma c) in
+ match EConstr.kind sigma t with
| Ind ind
- when (fst (Inductive.lookup_mind_specif env (fst ind))).mind_finite == Decl_kinds.CoFinite ->
+ when (fst (Inductive.lookup_mind_specif env (fst ind))).mind_finite == CoFinite ->
+ let l = List.map EConstr.Unsafe.to_constr l in
(ind, l)
| _ -> raise Not_found
@@ -490,32 +508,32 @@ let find_coinductive env sigma c =
(* find appropriate names for pattern variables. Useful in the Case
and Inversion (case_then_using et case_nodep_then_using) tactics. *)
-let is_predicate_explicitly_dep env pred arsign =
+let is_predicate_explicitly_dep env sigma pred arsign =
let rec srec env pval arsign =
- let pv' = whd_all env Evd.empty pval in
- match kind_of_term pv', arsign with
+ let pv' = whd_all env sigma pval in
+ match EConstr.kind sigma pv', arsign with
| Lambda (na,t,b), (LocalAssum _)::arsign ->
- srec (push_rel_assum (na,t) env) b arsign
+ srec (push_rel_assum (na, t) env) b arsign
| Lambda (na,_,t), _ ->
(* The following code has an impact on the introduction names
- given by the tactics "case" and "inversion": when the
- elimination is not dependent, "case" uses Anonymous for
- inductive types in Prop and names created by mkProd_name for
- inductive types in Set/Type while "inversion" uses anonymous
- for inductive types both in Prop and Set/Type !!
-
- Previously, whether names were created or not relied on
- whether the predicate created in Indrec.make_case_com had a
- dependent arity or not. To avoid different predicates
- printed the same in v8, all predicates built in indrec.ml
- got a dependent arity (Aug 2004). The new way to decide
- whether names have to be created or not is to use an
- Anonymous or Named variable to enforce the expected
- dependency status (of course, Anonymous implies non
- dependent, but not conversely).
-
- From Coq > 8.2, using or not the the effective dependency of
+ given by the tactics "case" and "inversion": when the
+ elimination is not dependent, "case" uses Anonymous for
+ inductive types in Prop and names created by mkProd_name for
+ inductive types in Set/Type while "inversion" uses anonymous
+ for inductive types both in Prop and Set/Type !!
+
+ Previously, whether names were created or not relied on
+ whether the predicate created in Indrec.make_case_com had a
+ dependent arity or not. To avoid different predicates
+ printed the same in v8, all predicates built in indrec.ml
+ got a dependent arity (Aug 2004). The new way to decide
+ whether names have to be created or not is to use an
+ Anonymous or Named variable to enforce the expected
+ dependency status (of course, Anonymous implies non
+ dependent, but not conversely).
+
+ From Coq > 8.2, using or not the effective dependency of
the predicate is parametrable! *)
begin match na with
@@ -523,19 +541,20 @@ let is_predicate_explicitly_dep env pred arsign =
| Name _ -> true
end
- | _ -> anomaly (Pp.str "Non eta-expanded dep-expanded \"match\" predicate")
+ | _ -> anomaly (Pp.str "Non eta-expanded dep-expanded \"match\" predicate.")
in
- srec env pred arsign
+ srec env (EConstr.of_constr pred) arsign
-let is_elim_predicate_explicitly_dependent env pred indf =
+let is_elim_predicate_explicitly_dependent env sigma pred indf =
let arsign,_ = get_arity env indf in
- is_predicate_explicitly_dep env pred arsign
+ is_predicate_explicitly_dep env sigma pred arsign
-let set_names env n brty =
- let (ctxt,cl) = decompose_prod_n_assum n brty in
- Namegen.it_mkProd_or_LetIn_name env cl ctxt
+let set_names env sigma n brty =
+ let open EConstr in
+ let (ctxt,cl) = decompose_prod_n_assum sigma n brty in
+ EConstr.Unsafe.to_constr (Namegen.it_mkProd_or_LetIn_name env sigma cl ctxt)
-let set_pattern_names env ind brv =
+let set_pattern_names env sigma ind brv =
let (mib,mip) = Inductive.lookup_mind_specif env ind in
let arities =
Array.map
@@ -543,10 +562,11 @@ let set_pattern_names env ind brv =
Context.Rel.length ((prod_assum c)) -
mib.mind_nparams)
mip.mind_nf_lc in
- Array.map2 (set_names env) arities brv
+ Array.map2 (set_names env sigma) arities brv
-let type_case_branches_with_names env indspec p c =
+let type_case_branches_with_names env sigma indspec p c =
let (ind,args) = indspec in
+ let args = List.map EConstr.Unsafe.to_constr args in
let (mib,mip as specif) = Inductive.lookup_mind_specif env (fst ind) in
let nparams = mib.mind_nparams in
let (params,realargs) = List.chop nparams args in
@@ -554,8 +574,8 @@ let type_case_branches_with_names env indspec p c =
(* Build case type *)
let conclty = lambda_appvect_assum (mip.mind_nrealdecls+1) p (Array.of_list (realargs@[c])) in
(* Adjust names *)
- if is_elim_predicate_explicitly_dependent env p (ind,params) then
- (set_pattern_names env (fst ind) lbrty, conclty)
+ if is_elim_predicate_explicitly_dependent env sigma p (ind,params) then
+ (set_pattern_names env sigma (fst ind) (Array.map EConstr.of_constr lbrty), conclty)
else (lbrty, conclty)
(* Type of Case predicates *)
@@ -563,7 +583,7 @@ let arity_of_case_predicate env (ind,params) dep k =
let arsign,_ = get_arity env (ind,params) in
let mind = build_dependent_inductive env (ind,params) in
let concl = if dep then mkArrow mind (mkSort k) else mkSort k in
- it_mkProd_or_LetIn concl arsign
+ Term.it_mkProd_or_LetIn concl arsign
(***********************************************)
(* Inferring the sort of parameters of a polymorphic inductive type
@@ -582,15 +602,15 @@ let rec instantiate_universes env evdref scl is = function
let ctx,_ = Reduction.dest_arity env ty in
let u = Univ.Universe.make l in
let s =
- (* Does the sort of parameter [u] appear in (or equal)
+ (* Does the sort of parameter [u] appear in (or equal)
the sort of inductive [is] ? *)
if univ_level_mem l is then
scl (* constrained sort: replace by scl *)
else
(* unconstrained sort: replace by fresh universe *)
let evm, s = Evd.new_sort_variable Evd.univ_flexible !evdref in
- let evm = Evd.set_leq_sort env evm s (Sorts.sort_of_univ u) in
- evdref := evm; s
+ let evm = Evd.set_leq_sort env evm s (Sorts.sort_of_univ u) in
+ evdref := evm; s
in
(LocalAssum (na,mkArity(ctx,s))) :: instantiate_universes env evdref scl is (sign, exp)
| sign, [] -> sign (* Uniform parameters are exhausted *)
@@ -598,24 +618,26 @@ let rec instantiate_universes env evdref scl is = function
let type_of_inductive_knowing_conclusion env sigma ((mib,mip),u) conclty =
match mip.mind_arity with
- | RegularArity s -> sigma, subst_instance_constr u s.mind_user_arity
+ | RegularArity s -> sigma, EConstr.of_constr (subst_instance_constr u s.mind_user_arity)
| TemplateArity ar ->
- let _,scl = Reduction.dest_arity env conclty in
+ let _,scl = splay_arity env sigma conclty in
+ let scl = EConstr.ESorts.kind sigma scl in
let ctx = List.rev mip.mind_arity_ctxt in
let evdref = ref sigma in
let ctx =
instantiate_universes
env evdref scl ar.template_level (ctx,ar.template_param_levels) in
- !evdref, mkArity (List.rev ctx,scl)
+ !evdref, EConstr.of_constr (mkArity (List.rev ctx,scl))
let type_of_projection_knowing_arg env sigma p c ty =
+ let c = EConstr.Unsafe.to_constr c in
let IndType(pars,realargs) =
try find_rectype env sigma ty
with Not_found ->
raise (Invalid_argument "type_of_projection_knowing_arg_type: not an inductive type")
in
let (_,u), pars = dest_ind_family pars in
- substl (c :: List.rev pars) (Typeops.type_of_projection env (p,u))
+ substl (c :: List.rev pars) (Typeops.type_of_projection_constant env (p,u))
(***********************************************)
(* Guard condition *)
@@ -623,8 +645,9 @@ let type_of_projection_knowing_arg env sigma p c ty =
(* A function which checks that a term well typed verifies both
syntactic conditions *)
-let control_only_guard env c =
- let check_fix_cofix e c = match kind_of_term c with
+let control_only_guard env sigma c =
+ let check_fix_cofix e c =
+ match kind (EConstr.to_constr sigma c) with
| CoFix (_,(_,_,_) as cofix) ->
Inductive.check_cofix e cofix
| Fix (_,(_,_,_) as fix) ->
@@ -633,6 +656,6 @@ let control_only_guard env c =
in
let rec iter env c =
check_fix_cofix env c;
- iter_constr_with_full_binders push_rel iter env c
+ iter_constr_with_full_binders sigma EConstr.push_rel iter env c
in
iter env c
diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli
index 7bd61659..b0d714b0 100644
--- a/pretyping/inductiveops.mli
+++ b/pretyping/inductiveops.mli
@@ -1,13 +1,15 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
+open Constr
open Declarations
open Environ
open Evd
@@ -28,8 +30,8 @@ val arities_of_constructors : env -> pinductive -> types array
reasoning either with only recursively uniform parameters or with all
parameters including the recursively non-uniform ones *)
type inductive_family
-val make_ind_family : inductive puniverses * constr list -> inductive_family
-val dest_ind_family : inductive_family -> inductive puniverses * constr list
+val make_ind_family : inductive Univ.puniverses * constr list -> inductive_family
+val dest_ind_family : inductive_family -> inductive Univ.puniverses * constr list
val map_ind_family : (constr -> constr) -> inductive_family -> inductive_family
val liftn_inductive_family : int -> int -> inductive_family -> inductive_family
val lift_inductive_family : int -> inductive_family -> inductive_family
@@ -37,15 +39,15 @@ val substnl_ind_family :
constr list -> int -> inductive_family -> inductive_family
(** An inductive type with its parameters and real arguments *)
-type inductive_type = IndType of inductive_family * constr list
-val make_ind_type : inductive_family * constr list -> inductive_type
-val dest_ind_type : inductive_type -> inductive_family * constr list
-val map_inductive_type : (constr -> constr) -> inductive_type -> inductive_type
+type inductive_type = IndType of inductive_family * EConstr.constr list
+val make_ind_type : inductive_family * EConstr.constr list -> inductive_type
+val dest_ind_type : inductive_type -> inductive_family * EConstr.constr list
+val map_inductive_type : (EConstr.constr -> EConstr.constr) -> inductive_type -> inductive_type
val liftn_inductive_type : int -> int -> inductive_type -> inductive_type
val lift_inductive_type : int -> inductive_type -> inductive_type
-val substnl_ind_type : constr list -> int -> inductive_type -> inductive_type
+val substnl_ind_type : EConstr.constr list -> int -> inductive_type -> inductive_type
-val mkAppliedInd : inductive_type -> constr
+val mkAppliedInd : inductive_type -> EConstr.constr
val mis_is_recursive_subset : int list -> wf_paths -> bool
val mis_is_recursive :
inductive * mutual_inductive_body * one_inductive_body -> bool
@@ -120,17 +122,17 @@ val constructor_nrealdecls_env : env -> constructor -> int
val constructor_has_local_defs : constructor -> bool
val inductive_has_local_defs : inductive -> bool
-val allowed_sorts : env -> inductive -> sorts_family list
+val allowed_sorts : env -> inductive -> Sorts.family list
(** (Co)Inductive records with primitive projections do not have eta-conversion,
hence no dependent elimination. *)
val has_dependent_elim : mutual_inductive_body -> bool
(** Primitive projections *)
-val projection_nparams : projection -> int
-val projection_nparams_env : env -> projection -> int
+val projection_nparams : Projection.t -> int
+val projection_nparams_env : env -> Projection.t -> int
val type_of_projection_knowing_arg : env -> evar_map -> Projection.t ->
- constr -> types -> types
+ EConstr.t -> EConstr.types -> types
(** Extract information from an inductive family *)
@@ -147,35 +149,35 @@ val get_constructor :
pinductive * mutual_inductive_body * one_inductive_body * constr list ->
int -> constructor_summary
val get_constructors : env -> inductive_family -> constructor_summary array
-val get_projections : env -> inductive_family -> constant array option
+val get_projections : env -> inductive_family -> Constant.t array option
(** [get_arity] returns the arity of the inductive family instantiated
with the parameters; if recursively non-uniform parameters are not
part of the inductive family, they appears in the arity *)
-val get_arity : env -> inductive_family -> Context.Rel.t * sorts_family
+val get_arity : env -> inductive_family -> Context.Rel.t * Sorts.family
val build_dependent_constructor : constructor_summary -> constr
val build_dependent_inductive : env -> inductive_family -> constr
-val make_arity_signature : env -> bool -> inductive_family -> Context.Rel.t
-val make_arity : env -> bool -> inductive_family -> sorts -> types
-val build_branch_type : env -> bool -> constr -> constructor_summary -> types
+val make_arity_signature : env -> evar_map -> bool -> inductive_family -> EConstr.rel_context
+val make_arity : env -> evar_map -> bool -> inductive_family -> Sorts.t -> EConstr.types
+val build_branch_type : env -> evar_map -> bool -> constr -> constructor_summary -> types
(** Raise [Not_found] if not given a valid inductive type *)
-val extract_mrectype : constr -> pinductive * constr list
-val find_mrectype : env -> evar_map -> types -> pinductive * constr list
-val find_mrectype_vect : env -> evar_map -> types -> pinductive * constr array
-val find_rectype : env -> evar_map -> types -> inductive_type
-val find_inductive : env -> evar_map -> types -> pinductive * constr list
-val find_coinductive : env -> evar_map -> types -> pinductive * constr list
+val extract_mrectype : evar_map -> EConstr.t -> (inductive * EConstr.EInstance.t) * EConstr.constr list
+val find_mrectype : env -> evar_map -> EConstr.types -> (inductive * EConstr.EInstance.t) * EConstr.constr list
+val find_mrectype_vect : env -> evar_map -> EConstr.types -> (inductive * EConstr.EInstance.t) * EConstr.constr array
+val find_rectype : env -> evar_map -> EConstr.types -> inductive_type
+val find_inductive : env -> evar_map -> EConstr.types -> (inductive * EConstr.EInstance.t) * constr list
+val find_coinductive : env -> evar_map -> EConstr.types -> (inductive * EConstr.EInstance.t) * constr list
(********************)
(** Builds the case predicate arity (dependent or not) *)
val arity_of_case_predicate :
- env -> inductive_family -> bool -> sorts -> types
+ env -> inductive_family -> bool -> Sorts.t -> types
val type_case_branches_with_names :
- env -> pinductive * constr list -> constr -> constr -> types array * types
+ env -> evar_map -> pinductive * EConstr.constr list -> constr -> constr -> types array * types
(** Annotation for cases *)
val make_case_info : env -> inductive -> case_style -> case_info
@@ -185,8 +187,8 @@ val make_case_info : env -> inductive -> case_style -> case_info
Fail with an error if the elimination is dependent while the
inductive type does not allow dependent elimination. *)
val make_case_or_project :
- env -> inductive_family -> case_info ->
- (* pred *) constr -> (* term *) constr -> (* branches *) constr array -> constr
+ env -> evar_map -> inductive_family -> case_info ->
+ (* pred *) EConstr.constr -> (* term *) EConstr.constr -> (* branches *) EConstr.constr array -> EConstr.constr
(*i Compatibility
val make_default_case_info : env -> case_style -> inductive -> case_info
@@ -195,7 +197,7 @@ i*)
(********************)
val type_of_inductive_knowing_conclusion :
- env -> evar_map -> Inductive.mind_specif puniverses -> types -> evar_map * types
+ env -> evar_map -> Inductive.mind_specif Univ.puniverses -> EConstr.types -> evar_map * EConstr.types
(********************)
-val control_only_guard : env -> types -> unit
+val control_only_guard : env -> Evd.evar_map -> EConstr.types -> unit
diff --git a/pretyping/inferCumulativity.ml b/pretyping/inferCumulativity.ml
new file mode 100644
index 00000000..be79b8b0
--- /dev/null
+++ b/pretyping/inferCumulativity.ml
@@ -0,0 +1,212 @@
+(************************************************************************)
+(* * 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) *)
+(************************************************************************)
+
+open Reduction
+open Declarations
+open Constr
+open Univ
+open Util
+
+(** Throughout this module we modify a map [variances] from local
+ universes to [Variance.t]. It starts as a trivial mapping to
+ [Irrelevant] and every time we encounter a local universe we
+ restrict it accordingly. *)
+
+let infer_level_eq u variances =
+ if LMap.mem u variances
+ then LMap.set u Variance.Invariant variances
+ else variances
+
+let infer_level_leq u variances =
+ match LMap.find u variances with
+ | exception Not_found -> variances
+ | varu -> LMap.set u (Variance.sup varu Variance.Covariant) variances
+
+let infer_generic_instance_eq variances u =
+ Array.fold_left (fun variances u -> infer_level_eq u variances)
+ variances (Instance.to_array u)
+
+let variance_pb cv_pb var =
+ let open Variance in
+ match cv_pb, var with
+ | _, Irrelevant -> Irrelevant
+ | _, Invariant -> Invariant
+ | CONV, Covariant -> Invariant
+ | CUMUL, Covariant -> Covariant
+
+let infer_cumulative_ind_instance cv_pb cumi variances u =
+ Array.fold_left2 (fun variances varu u ->
+ match LMap.find u variances with
+ | exception Not_found -> variances
+ | varu' ->
+ LMap.set u (Variance.sup varu' (variance_pb cv_pb varu)) variances)
+ variances (ACumulativityInfo.variance cumi) (Instance.to_array u)
+
+let infer_inductive_instance cv_pb env variances ind nargs u =
+ let mind = Environ.lookup_mind (fst ind) env in
+ match mind.mind_universes with
+ | Monomorphic_ind _ -> assert (Instance.is_empty u); variances
+ | Polymorphic_ind _ -> infer_generic_instance_eq variances u
+ | Cumulative_ind cumi ->
+ if not (Int.equal (inductive_cumulativity_arguments (mind,snd ind)) nargs)
+ then infer_generic_instance_eq variances u
+ else infer_cumulative_ind_instance cv_pb cumi variances u
+
+let infer_constructor_instance_eq env variances ((mi,ind),ctor) nargs u =
+ let mind = Environ.lookup_mind mi env in
+ match mind.mind_universes with
+ | Monomorphic_ind _ -> assert (Instance.is_empty u); variances
+ | Polymorphic_ind _ -> infer_generic_instance_eq variances u
+ | Cumulative_ind cumi ->
+ if not (Int.equal (constructor_cumulativity_arguments (mind,ind,ctor)) nargs)
+ then infer_generic_instance_eq variances u
+ else infer_cumulative_ind_instance CONV cumi variances u
+
+let infer_sort cv_pb variances s =
+ match cv_pb with
+ | CONV ->
+ LSet.fold infer_level_eq (Universe.levels (Sorts.univ_of_sort s)) variances
+ | CUMUL ->
+ LSet.fold infer_level_leq (Universe.levels (Sorts.univ_of_sort s)) variances
+
+let infer_table_key infos variances c =
+ let open Names in
+ match c with
+ | ConstKey (_, u) ->
+ infer_generic_instance_eq variances u
+ | VarKey _ | RelKey _ -> variances
+
+let whd_stack (infos, tab) hd stk = CClosure.whd_stack infos tab hd stk
+
+let rec infer_fterm cv_pb infos variances hd stk =
+ Control.check_for_interrupt ();
+ let hd,stk = whd_stack infos hd stk in
+ let open CClosure in
+ match fterm_of hd with
+ | FAtom a ->
+ begin match kind a with
+ | Sort s -> infer_sort cv_pb variances s
+ | Meta _ -> infer_stack infos variances stk
+ | _ -> assert false
+ end
+ | FEvar ((_,args),e) ->
+ let variances = infer_stack infos variances stk in
+ infer_vect infos variances (Array.map (mk_clos e) args)
+ | FRel _ -> infer_stack infos variances stk
+ | FFlex fl ->
+ let variances = infer_table_key infos variances fl in
+ infer_stack infos variances stk
+ | FProj (_,c) ->
+ let variances = infer_fterm CONV infos variances c [] in
+ infer_stack infos variances stk
+ | FLambda _ ->
+ let (_,ty,bd) = destFLambda mk_clos hd in
+ let variances = infer_fterm CONV infos variances ty [] in
+ infer_fterm CONV infos variances bd []
+ | FProd (_,dom,codom) ->
+ let variances = infer_fterm CONV infos variances dom [] in
+ infer_fterm cv_pb infos variances codom []
+ | FInd (ind, u) ->
+ let variances =
+ if Instance.is_empty u then variances
+ else
+ let nargs = stack_args_size stk in
+ infer_inductive_instance cv_pb (info_env (fst infos)) variances ind nargs u
+ in
+ infer_stack infos variances stk
+ | FConstruct (ctor,u) ->
+ let variances =
+ if Instance.is_empty u then variances
+ else
+ let nargs = stack_args_size stk in
+ infer_constructor_instance_eq (info_env (fst infos)) variances ctor nargs u
+ in
+ infer_stack infos variances stk
+ | FFix ((_,(_,tys,cl)),e) | FCoFix ((_,(_,tys,cl)),e) ->
+ let n = Array.length cl in
+ let variances = infer_vect infos variances (Array.map (mk_clos e) tys) in
+ let le = Esubst.subs_liftn n e in
+ let variances = infer_vect infos variances (Array.map (mk_clos le) cl) in
+ infer_stack infos variances stk
+
+ (* Removed by whnf *)
+ | FLOCKED | FCaseT _ | FCast _ | FLetIn _ | FApp _ | FLIFT _ | FCLOS _ -> assert false
+
+and infer_stack infos variances (stk:CClosure.stack) =
+ match stk with
+ | [] -> variances
+ | z :: stk ->
+ let open CClosure in
+ let variances = match z with
+ | Zapp v -> infer_vect infos variances v
+ | Zproj _ -> variances
+ | Zfix (fx,a) ->
+ let variances = infer_fterm CONV infos variances fx [] in
+ infer_stack infos variances a
+ | ZcaseT (ci,p,br,e) ->
+ let variances = infer_fterm CONV infos variances (mk_clos e p) [] in
+ infer_vect infos variances (Array.map (mk_clos e) br)
+ | Zshift _ -> variances
+ | Zupdate _ -> variances
+ in
+ infer_stack infos variances stk
+
+and infer_vect infos variances v =
+ Array.fold_left (fun variances c -> infer_fterm CONV infos variances c []) variances v
+
+let infer_term cv_pb env variances c =
+ let open CClosure in
+ let infos = (create_clos_infos all env, create_tab ()) in
+ infer_fterm cv_pb infos variances (CClosure.inject c) []
+
+let infer_arity_constructor is_arity env variances arcn =
+ let infer_typ typ (env,variances) =
+ match typ with
+ | Context.Rel.Declaration.LocalAssum (_, typ') ->
+ (Environ.push_rel typ env, infer_term CUMUL env variances typ')
+ | Context.Rel.Declaration.LocalDef _ -> assert false
+ in
+ let typs, codom = Reduction.dest_prod env arcn in
+ let env, variances = Context.Rel.fold_outside infer_typ typs ~init:(env, variances) in
+ (* If we have Inductive foo@{i j} : ... -> Type@{i} := C : ... -> foo Type@{j}
+ i is irrelevant, j is invariant. *)
+ if not is_arity then infer_term CUMUL env variances codom else variances
+
+let infer_inductive env mie =
+ let open Entries in
+ let { mind_entry_params = params;
+ mind_entry_inds = entries; } = mie
+ in
+ let univs =
+ match mie.mind_entry_universes with
+ | Monomorphic_ind_entry _
+ | Polymorphic_ind_entry _ as univs -> univs
+ | Cumulative_ind_entry cumi ->
+ let uctx = CumulativityInfo.univ_context cumi in
+ let uarray = Instance.to_array @@ UContext.instance uctx in
+ let env = Environ.push_context uctx env in
+ let variances =
+ Array.fold_left (fun variances u -> LMap.add u Variance.Irrelevant variances)
+ LMap.empty uarray
+ in
+ let env, _ = Typeops.infer_local_decls env params in
+ let variances = List.fold_left (fun variances entry ->
+ let variances = infer_arity_constructor true
+ env variances entry.mind_entry_arity
+ in
+ List.fold_left (infer_arity_constructor false env)
+ variances entry.mind_entry_lc)
+ variances
+ entries
+ in
+ let variances = Array.map (fun u -> LMap.find u variances) uarray in
+ Cumulative_ind_entry (CumulativityInfo.make (uctx, variances))
+ in
+ { mie with mind_entry_universes = univs }
diff --git a/pretyping/inferCumulativity.mli b/pretyping/inferCumulativity.mli
new file mode 100644
index 00000000..a0c8d339
--- /dev/null
+++ b/pretyping/inferCumulativity.mli
@@ -0,0 +1,12 @@
+(************************************************************************)
+(* * 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) *)
+(************************************************************************)
+
+val infer_inductive : Environ.env -> Entries.mutual_inductive_entry ->
+ Entries.mutual_inductive_entry
diff --git a/pretyping/locusops.ml b/pretyping/locusops.ml
index e4fbf8d5..1664e68f 100644
--- a/pretyping/locusops.ml
+++ b/pretyping/locusops.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Locus
@@ -50,9 +52,9 @@ let is_nowhere = function
let simple_clause_of enum_hyps cl =
let error_occurrences () =
- CErrors.error "This tactic does not support occurrences selection" in
+ CErrors.user_err Pp.(str "This tactic does not support occurrences selection") in
let error_body_selection () =
- CErrors.error "This tactic does not support body selection" in
+ CErrors.user_err Pp.(str "This tactic does not support body selection") in
let hyps =
match cl.onhyps with
| None ->
@@ -84,7 +86,7 @@ let concrete_clause_of enum_hyps cl =
(** Miscellaneous functions *)
let out_arg = function
- | Misctypes.ArgVar _ -> CErrors.anomaly (Pp.str "Unevaluated or_var variable")
+ | Misctypes.ArgVar _ -> CErrors.anomaly (Pp.str "Unevaluated or_var variable.")
| Misctypes.ArgArg x -> x
let occurrences_of_hyp id cls =
diff --git a/pretyping/locusops.mli b/pretyping/locusops.mli
index c7661239..a07c018c 100644
--- a/pretyping/locusops.mli
+++ b/pretyping/locusops.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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
diff --git a/pretyping/ltac_pretype.ml b/pretyping/ltac_pretype.ml
new file mode 100644
index 00000000..be8579c2
--- /dev/null
+++ b/pretyping/ltac_pretype.ml
@@ -0,0 +1,68 @@
+open Names
+open Glob_term
+
+(** {5 Maps of pattern variables} *)
+
+(** Type [constr_under_binders] is for representing the term resulting
+ of a matching. Matching can return terms defined in a some context
+ of named binders; in the context, variable names are ordered by
+ (<) and referred to by index in the term Thanks to the canonical
+ ordering, a matching problem like
+
+ [match ... with [(fun x y => ?p,fun y x => ?p)] => [forall x y => p]]
+
+ will be accepted. Thanks to the reference by index, a matching
+ problem like
+
+ [match ... with [(fun x => ?p)] => [forall x => p]]
+
+ will work even if [x] is also the name of an existing goal
+ variable.
+
+ Note: we do not keep types in the signature. Besides simplicity,
+ the main reason is that it would force to close the signature over
+ binders that occur only in the types of effective binders but not
+ in the term itself (e.g. for a term [f x] with [f:A -> True] and
+ [x:A]).
+
+ On the opposite side, by not keeping the types, we loose
+ opportunity to propagate type informations which otherwise would
+ not be inferable, as e.g. when matching [forall x, x = 0] with
+ pattern [forall x, ?h = 0] and using the solution "x|-h:=x" in
+ expression [forall x, h = x] where nothing tells how the type of x
+ could be inferred. We also loose the ability of typing ltac
+ variables before calling the right-hand-side of ltac matching clauses. *)
+
+type constr_under_binders = Id.t list * EConstr.constr
+
+(** Types of substitutions with or w/o bound variables *)
+
+type patvar_map = EConstr.constr Id.Map.t
+type extended_patvar_map = constr_under_binders Id.Map.t
+
+(** A globalised term together with a closure representing the value
+ of its free variables. Intended for use when these variables are taken
+ from the Ltac environment. *)
+type closure = {
+ idents:Id.t Id.Map.t;
+ typed: constr_under_binders Id.Map.t ;
+ untyped:closed_glob_constr Id.Map.t }
+and closed_glob_constr = {
+ closure: closure;
+ term: glob_constr }
+
+(** Ltac variable maps *)
+type var_map = constr_under_binders Id.Map.t
+type uconstr_var_map = closed_glob_constr Id.Map.t
+type unbound_ltac_var_map = Geninterp.Val.t Id.Map.t
+
+type ltac_var_map = {
+ ltac_constrs : var_map;
+ (** Ltac variables bound to constrs *)
+ ltac_uconstrs : uconstr_var_map;
+ (** Ltac variables bound to untyped constrs *)
+ ltac_idents: Id.t Id.Map.t;
+ (** Ltac variables bound to identifiers *)
+ ltac_genargs : unbound_ltac_var_map;
+ (** Ltac variables bound to other kinds of arguments *)
+}
diff --git a/pretyping/miscops.ml b/pretyping/miscops.ml
index 142e430f..0f0af540 100644
--- a/pretyping/miscops.ml
+++ b/pretyping/miscops.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Util
@@ -30,7 +32,8 @@ let smartmap_cast_type f c =
let glob_sort_eq g1 g2 = match g1, g2 with
| GProp, GProp -> true
| GSet, GSet -> true
-| GType l1, GType l2 -> List.equal (fun x y -> CString.equal (snd x) (snd y)) l1 l2
+| GType l1, GType l2 ->
+ List.equal (Option.equal (fun (x,m) (y,n) -> Libnames.eq_reference x y && Int.equal m n)) l1 l2
| _ -> false
let intro_pattern_naming_eq nam1 nam2 = match nam1, nam2 with
@@ -58,3 +61,16 @@ let map_red_expr_gen f g h = function
| CbvNative occs_o -> CbvNative (Option.map (map_occs (map_union g h)) occs_o)
| Cbn flags -> Cbn (map_flags g flags)
| ExtraRedExpr _ | Red _ | Hnf as x -> x
+
+(** Mapping bindings *)
+
+let map_explicit_bindings f l =
+ let map = CAst.map (fun (hyp, x) -> (hyp, f x)) in
+ List.map map l
+
+let map_bindings f = function
+| ImplicitBindings l -> ImplicitBindings (List.map f l)
+| ExplicitBindings expl -> ExplicitBindings (map_explicit_bindings f expl)
+| NoBindings -> NoBindings
+
+let map_with_bindings f (x, bl) = (f x, map_bindings f bl)
diff --git a/pretyping/miscops.mli b/pretyping/miscops.mli
index 337473a6..abe817fe 100644
--- a/pretyping/miscops.mli
+++ b/pretyping/miscops.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Misctypes
@@ -27,3 +29,8 @@ val intro_pattern_naming_eq :
val map_red_expr_gen : ('a -> 'd) -> ('b -> 'e) -> ('c -> 'f) ->
('a,'b,'c) red_expr_gen -> ('d,'e,'f) red_expr_gen
+
+(** Mapping bindings *)
+
+val map_bindings : ('a -> 'b) -> 'a bindings -> 'b bindings
+val map_with_bindings : ('a -> 'b) -> 'a with_bindings -> 'b with_bindings
diff --git a/pretyping/nativenorm.ml b/pretyping/nativenorm.ml
index 0dd64697..1ed4d21b 100644
--- a/pretyping/nativenorm.ml
+++ b/pretyping/nativenorm.ml
@@ -1,17 +1,18 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Pp
open CErrors
open Term
+open Constr
open Vars
open Environ
open Reduction
-open Univ
open Declarations
open Names
open Inductive
@@ -20,10 +21,65 @@ open Nativecode
open Nativevalues
open Context.Rel.Declaration
+module NamedDecl = Context.Named.Declaration
+
(** This module implements normalization by evaluation to OCaml code *)
exception Find_at of int
+(* profiling *)
+
+let profiling_enabled = ref false
+
+(* for supported platforms, filename for profiler results *)
+
+let profile_filename = ref "native_compute_profile.data"
+
+let profiler_platform () =
+ match [@warning "-8"] Sys.os_type with
+ | "Unix" ->
+ let in_ch = Unix.open_process_in "uname" in
+ let uname = input_line in_ch in
+ let _ = close_in in_ch in
+ Format.sprintf "Unix (%s)" uname
+ | "Win32" -> "Windows (Win32)"
+ | "Cygwin" -> "Windows (Cygwin)"
+
+let get_profile_filename () = !profile_filename
+
+let set_profile_filename fn =
+ profile_filename := fn
+
+(* find unused profile filename *)
+let get_available_profile_filename () =
+ let profile_filename = get_profile_filename () in
+ let dir = Filename.dirname profile_filename in
+ let base = Filename.basename profile_filename in
+ (* starting with OCaml 4.04, could use Filename.remove_extension and Filename.extension, which
+ gets rid of need for exception-handling here
+ *)
+ let (name,ext) =
+ try
+ let nm = Filename.chop_extension base in
+ let nm_len = String.length nm in
+ let ex = String.sub base nm_len (String.length base - nm_len) in
+ (nm,ex)
+ with Invalid_argument _ -> (base,"")
+ in
+ try
+ (* unlikely race: fn deleted, another process uses fn *)
+ Filename.temp_file ~temp_dir:dir (name ^ "_") ext
+ with Sys_error s ->
+ let msg = "When trying to find native_compute profile output file: " ^ s in
+ let _ = Feedback.msg_info (Pp.str msg) in
+ assert false
+
+let get_profiling_enabled () =
+ !profiling_enabled
+
+let set_profiling_enabled b =
+ profiling_enabled := b
+
let invert_tag cst tag reloc_tbl =
try
for j = 0 to Array.length reloc_tbl - 1 do
@@ -37,7 +93,7 @@ let invert_tag cst tag reloc_tbl =
let decompose_prod env t =
let (name,dom,codom as res) = destProd (whd_all env t) in
match name with
- | Anonymous -> (Name (id_of_string "x"),dom,codom)
+ | Anonymous -> (Name (Id.of_string "x"),dom,codom)
| _ -> res
let app_type env c =
@@ -47,7 +103,7 @@ let app_type env c =
let find_rectype_a env c =
let (t, l) = app_type env c in
- match kind_of_term t with
+ match kind t with
| Ind ind -> (ind, l)
| _ -> raise Not_found
@@ -80,7 +136,7 @@ let construct_of_constr_notnative const env tag (mind, _ as ind) u allargs =
let construct_of_constr const env tag typ =
let t, l = app_type env typ in
- match kind_of_term t with
+ match kind t with
| Ind (ind,u) ->
construct_of_constr_notnative const env tag ind u l
| _ -> assert false
@@ -90,13 +146,15 @@ let construct_of_constr_const env tag typ =
let construct_of_constr_block = construct_of_constr false
-let build_branches_type env (mind,_ as _ind) mib mip u params dep p =
+let build_branches_type env sigma (mind,_ as _ind) mib mip u params dep p =
let rtbl = mip.mind_reloc_tbl in
(* [build_one_branch i cty] construit le type de la ieme branche (commence
a 0) et les lambda correspondant aux realargs *)
let build_one_branch i cty =
let typi = type_constructor mind mib u cty params in
- let decl,indapp = Reductionops.splay_prod env Evd.empty typi in
+ let decl,indapp = Reductionops.splay_prod env sigma (EConstr.of_constr typi) in
+ let decl = List.map (on_snd EConstr.Unsafe.to_constr) decl in
+ let indapp = EConstr.Unsafe.to_constr indapp in
let decl_with_letin,_ = decompose_prod_assum typi in
let ind,cargs = find_rectype_a env indapp in
let nparams = Array.length params in
@@ -120,44 +178,6 @@ let build_case_type dep p realargs c =
if dep then mkApp(mkApp(p, realargs), [|c|])
else mkApp(p, realargs)
-(* TODO move this function *)
-let type_of_rel env n =
- lookup_rel n env |> get_type |> lift n
-
-let type_of_prop = mkSort type1_sort
-
-let type_of_sort s =
- match s with
- | Prop _ -> type_of_prop
- | Type u -> mkType (Univ.super u)
-
-let type_of_var env id =
- let open Context.Named.Declaration in
- try lookup_named id env |> get_type
- with Not_found ->
- anomaly ~label:"type_of_var" (str "variable " ++ Id.print id ++ str " unbound")
-
-let sort_of_product env domsort rangsort =
- match (domsort, rangsort) with
- (* Product rule (s,Prop,Prop) *)
- | (_, Prop Null) -> rangsort
- (* Product rule (Prop/Set,Set,Set) *)
- | (Prop _, Prop Pos) -> rangsort
- (* Product rule (Type,Set,?) *)
- | (Type u1, Prop Pos) ->
- if is_impredicative_set env then
- (* Rule is (Type,Set,Set) in the Set-impredicative calculus *)
- rangsort
- else
- (* Rule is (Type_i,Set,Type_i) in the Set-predicative calculus *)
- Type (sup u1 type0_univ)
- (* Product rule (Prop,Type_i,Type_i) *)
- | (Prop Pos, Type u2) -> Type (sup type0_univ u2)
- (* Product rule (Prop,Type_i,Type_i) *)
- | (Prop Null, Type _) -> rangsort
- (* Product rule (Type_i,Type_i,Type_i) *)
- | (Type u1, Type u2) -> Type (sup u1 u2)
-
(* normalisation of values *)
let branch_of_switch lvl ans bs =
@@ -170,9 +190,9 @@ let branch_of_switch lvl ans bs =
bs ci in
Array.init (Array.length tbl) branch
-let rec nf_val env v typ =
+let rec nf_val env sigma v typ =
match kind_of_value v with
- | Vaccu accu -> nf_accu env accu
+ | Vaccu accu -> nf_accu env sigma accu
| Vfun f ->
let lvl = nb_rel env in
let name,dom,codom =
@@ -182,44 +202,44 @@ let rec nf_val env v typ =
(Pp.strbrk "Returned a functional value in a type not recognized as a product type.")
in
let env = push_rel (LocalAssum (name,dom)) env in
- let body = nf_val env (f (mk_rel_accu lvl)) codom in
+ let body = nf_val env sigma (f (mk_rel_accu lvl)) codom in
mkLambda(name,dom,body)
| Vconst n -> construct_of_constr_const env n typ
| Vblock b ->
let capp,ctyp = construct_of_constr_block env (block_tag b) typ in
- let args = nf_bargs env b ctyp in
+ let args = nf_bargs env sigma b ctyp in
mkApp(capp,args)
-and nf_type env v =
+and nf_type env sigma v =
match kind_of_value v with
- | Vaccu accu -> nf_accu env accu
+ | Vaccu accu -> nf_accu env sigma accu
| _ -> assert false
-and nf_type_sort env v =
+and nf_type_sort env sigma v =
match kind_of_value v with
| Vaccu accu ->
- let t,s = nf_accu_type env accu in
+ let t,s = nf_accu_type env sigma accu in
let s = try destSort s with DestKO -> assert false in
t, s
| _ -> assert false
-and nf_accu env accu =
+and nf_accu env sigma accu =
let atom = atom_of_accu accu in
- if Int.equal (accu_nargs accu) 0 then nf_atom env atom
+ if Int.equal (accu_nargs accu) 0 then nf_atom env sigma atom
else
- let a,typ = nf_atom_type env atom in
- let _, args = nf_args env accu typ in
+ let a,typ = nf_atom_type env sigma atom in
+ let _, args = nf_args env sigma (args_of_accu accu) typ in
mkApp(a,Array.of_list args)
-and nf_accu_type env accu =
+and nf_accu_type env sigma accu =
let atom = atom_of_accu accu in
- if Int.equal (accu_nargs accu) 0 then nf_atom_type env atom
+ if Int.equal (accu_nargs accu) 0 then nf_atom_type env sigma atom
else
- let a,typ = nf_atom_type env atom in
- let t, args = nf_args env accu typ in
+ let a,typ = nf_atom_type env sigma atom in
+ let t, args = nf_args env sigma (args_of_accu accu) typ in
mkApp(a,Array.of_list args), t
-and nf_args env accu t =
+and nf_args env sigma args t =
let aux arg (t,l) =
let _,dom,codom =
try decompose_prod env t with
@@ -227,13 +247,13 @@ and nf_args env accu t =
CErrors.anomaly
(Pp.strbrk "Returned a functional value in a type not recognized as a product type.")
in
- let c = nf_val env arg dom in
+ let c = nf_val env sigma arg dom in
(subst1 c codom, c::l)
in
- let t,l = List.fold_right aux (args_of_accu accu) (t,[]) in
+ let t,l = Array.fold_right aux args (t,[]) in
t, List.rev l
-and nf_bargs env b t =
+and nf_bargs env sigma b t =
let t = ref t in
let len = block_size b in
Array.init len
@@ -244,10 +264,10 @@ and nf_bargs env b t =
CErrors.anomaly
(Pp.strbrk "Returned a functional value in a type not recognized as a product type.")
in
- let c = nf_val env (block_field b i) dom in
+ let c = nf_val env sigma (block_field b i) dom in
t := subst1 c codom; c)
-and nf_atom env atom =
+and nf_atom env sigma atom =
match atom with
| Arel i -> mkRel (nb_rel env - i)
| Aconstant cst -> mkConstU cst
@@ -255,33 +275,32 @@ and nf_atom env atom =
| Asort s -> mkSort s
| Avar id -> mkVar id
| Aprod(n,dom,codom) ->
- let dom = nf_type env dom in
+ let dom = nf_type env sigma dom in
let vn = mk_rel_accu (nb_rel env) in
let env = push_rel (LocalAssum (n,dom)) env in
- let codom = nf_type env (codom vn) in
+ let codom = nf_type env sigma (codom vn) in
mkProd(n,dom,codom)
| Ameta (mv,_) -> mkMeta mv
- | Aevar (ev,_) -> mkEvar ev
| Aproj(p,c) ->
- let c = nf_accu env c in
+ let c = nf_accu env sigma c in
mkProj(Projection.make p true,c)
- | _ -> fst (nf_atom_type env atom)
+ | _ -> fst (nf_atom_type env sigma atom)
-and nf_atom_type env atom =
+and nf_atom_type env sigma atom =
match atom with
| Arel i ->
let n = (nb_rel env - i) in
- mkRel n, type_of_rel env n
+ mkRel n, Typeops.type_of_relative env n
| Aconstant cst ->
mkConstU cst, Typeops.type_of_constant_in env cst
| Aind ind ->
mkIndU ind, Inductiveops.type_of_inductive env ind
| Asort s ->
- mkSort s, type_of_sort s
+ mkSort s, Typeops.type_of_sort s
| Avar id ->
- mkVar id, type_of_var env id
+ mkVar id, Typeops.type_of_variable env id
| Acase(ans,accu,p,bs) ->
- let a,ta = nf_accu_type env accu in
+ let a,ta = nf_accu_type env sigma accu in
let ((mind,_),u as ind),allargs = find_rectype_a env ta in
let (mib,mip) = Inductive.lookup_mind_specif env (fst ind) in
let nparams = mib.mind_nparams in
@@ -290,14 +309,14 @@ and nf_atom_type env atom =
hnf_prod_applist env
(Inductiveops.type_of_inductive env ind) (Array.to_list params) in
let pT = whd_all env pT in
- let dep, p = nf_predicate env ind mip params p pT in
+ let dep, p = nf_predicate env sigma ind mip params p pT in
(* Calcul du type des branches *)
- let btypes = build_branches_type env (fst ind) mib mip u params dep p in
+ let btypes = build_branches_type env sigma (fst ind) mib mip u params dep p in
(* calcul des branches *)
let bsw = branch_of_switch (nb_rel env) ans bs in
let mkbranch i v =
let decl,decl_with_letin,codom = btypes.(i) in
- let b = nf_val (Termops.push_rels_assum decl env) v codom in
+ let b = nf_val (Termops.push_rels_assum decl env) sigma v codom in
Termops.it_mkLambda_or_LetIn_from_no_LetIn b decl_with_letin
in
let branchs = Array.mapi mkbranch bsw in
@@ -305,8 +324,8 @@ and nf_atom_type env atom =
let ci = ans.asw_ci in
mkCase(ci, p, a, branchs), tcase
| Afix(tt,ft,rp,s) ->
- let tt = Array.map (fun t -> nf_type env t) tt in
- let name = Array.map (fun _ -> (Name (id_of_string "Ffix"))) tt in
+ let tt = Array.map (fun t -> nf_type env sigma t) tt in
+ let name = Array.map (fun _ -> (Name (Id.of_string "Ffix"))) tt in
let lvl = nb_rel env in
let nbfix = Array.length ft in
let fargs = mk_rels_accu lvl (Array.length ft) in
@@ -314,38 +333,37 @@ and nf_atom_type env atom =
let env = push_rec_types (name,tt,[||]) env in
(* We lift here because the types of arguments (in tt) will be evaluated
in an environment where the fixpoints have been pushed *)
- let norm_body i v = nf_val env (napply v fargs) (lift nbfix tt.(i)) in
+ let norm_body i v = nf_val env sigma (napply v fargs) (lift nbfix tt.(i)) in
let ft = Array.mapi norm_body ft in
mkFix((rp,s),(name,tt,ft)), tt.(s)
| Acofix(tt,ft,s,_) | Acofixe(tt,ft,s,_) ->
- let tt = Array.map (nf_type env) tt in
- let name = Array.map (fun _ -> (Name (id_of_string "Fcofix"))) tt in
+ let tt = Array.map (nf_type env sigma) tt in
+ let name = Array.map (fun _ -> (Name (Id.of_string "Fcofix"))) tt in
let lvl = nb_rel env in
let fargs = mk_rels_accu lvl (Array.length ft) in
let env = push_rec_types (name,tt,[||]) env in
- let ft = Array.mapi (fun i v -> nf_val env (napply v fargs) tt.(i)) ft in
+ let ft = Array.mapi (fun i v -> nf_val env sigma (napply v fargs) tt.(i)) ft in
mkCoFix(s,(name,tt,ft)), tt.(s)
| Aprod(n,dom,codom) ->
- let dom,s1 = nf_type_sort env dom in
+ let dom,s1 = nf_type_sort env sigma dom in
let vn = mk_rel_accu (nb_rel env) in
let env = push_rel (LocalAssum (n,dom)) env in
- let codom,s2 = nf_type_sort env (codom vn) in
- mkProd(n,dom,codom), mkSort (sort_of_product env s1 s2)
- | Aevar(ev,ty) ->
- let ty = nf_type env ty in
- mkEvar ev, ty
+ let codom,s2 = nf_type_sort env sigma (codom vn) in
+ mkProd(n,dom,codom), Typeops.type_of_product env n s1 s2
+ | Aevar(evk, _, args) ->
+ nf_evar env sigma evk args
| Ameta(mv,ty) ->
- let ty = nf_type env ty in
+ let ty = nf_type env sigma ty in
mkMeta mv, ty
| Aproj(p,c) ->
- let c,tc = nf_accu_type env c in
+ let c,tc = nf_accu_type env sigma c in
let cj = make_judge c tc in
let uj = Typeops.judge_of_projection env (Projection.make p true) cj in
uj.uj_val, uj.uj_type
-and nf_predicate env ind mip params v pT =
- match kind_of_value v, kind_of_term pT with
+and nf_predicate env sigma ind mip params v pT =
+ match kind_of_value v, kind pT with
| Vfun f, Prod _ ->
let k = nb_rel env in
let vb = f (mk_rel_accu k) in
@@ -356,51 +374,124 @@ and nf_predicate env ind mip params v pT =
(Pp.strbrk "Returned a functional value in a type not recognized as a product type.")
in
let dep,body =
- nf_predicate (push_rel (LocalAssum (name,dom)) env) ind mip params vb codom in
+ nf_predicate (push_rel (LocalAssum (name,dom)) env) sigma ind mip params vb codom in
dep, mkLambda(name,dom,body)
| Vfun f, _ ->
let k = nb_rel env in
let vb = f (mk_rel_accu k) in
- let name = Name (id_of_string "c") in
+ let name = Name (Id.of_string "c") in
let n = mip.mind_nrealargs in
let rargs = Array.init n (fun i -> mkRel (n-i)) in
let params = if Int.equal n 0 then params else Array.map (lift n) params in
let dom = mkApp(mkIndU ind,Array.append params rargs) in
- let body = nf_type (push_rel (LocalAssum (name,dom)) env) vb in
+ let body = nf_type (push_rel (LocalAssum (name,dom)) env) sigma vb in
true, mkLambda(name,dom,body)
- | _, _ -> false, nf_type env v
+ | _, _ -> false, nf_type env sigma v
+
+and nf_evar env sigma evk args =
+ let evi = try Evd.find sigma evk with Not_found -> assert false in
+ let hyps = Environ.named_context_of_val (Evd.evar_filtered_hyps evi) in
+ let ty = Evd.evar_concl evi in
+ if List.is_empty hyps then begin
+ assert (Int.equal (Array.length args) 0);
+ mkEvar (evk, [||]), ty
+ end
+ else
+ (** Let-bound arguments are present in the evar arguments but not in the
+ type, so we turn the let into a product. *)
+ let drop_body = function
+ | NamedDecl.LocalAssum _ as d -> d
+ | NamedDecl.LocalDef (na, _, t) -> NamedDecl.LocalAssum (na, t)
+ in
+ let hyps = List.map drop_body hyps in
+ let fold accu d = Term.mkNamedProd_or_LetIn d accu in
+ let t = List.fold_left fold ty hyps in
+ let ty, args = nf_args env sigma args t in
+ (** nf_args takes arguments in the reverse order but produces them in the
+ correct one, so we have to reverse them again for the evar node *)
+ mkEvar (evk, Array.rev_of_list args), ty
let evars_of_evar_map sigma =
{ Nativelambda.evars_val = Evd.existential_opt_value sigma;
Nativelambda.evars_typ = Evd.existential_type sigma;
Nativelambda.evars_metas = Evd.meta_type sigma }
-let native_norm env sigma c ty =
- if Coq_config.no_native_compiler then
- error "Native_compute reduction has been disabled at configure time."
+(* fork perf process, return profiler's process id *)
+let start_profiler_linux profile_fn =
+ let coq_pid = Unix.getpid () in (* pass pid of running coqtop *)
+ (* we don't want to see perf's console output *)
+ let dev_null = Unix.descr_of_out_channel (open_out_bin "/dev/null") in
+ let _ = Feedback.msg_info (Pp.str ("Profiling to file " ^ profile_fn)) in
+ let perf = "perf" in
+ let profiler_pid =
+ Unix.create_process
+ perf
+ [|perf; "record"; "-g"; "-o"; profile_fn; "-p"; string_of_int coq_pid |]
+ Unix.stdin dev_null dev_null
+ in
+ (* doesn't seem to be a way to test whether process creation succeeded *)
+ if !Flags.debug then
+ Feedback.msg_debug (Pp.str (Format.sprintf "Native compute profiler started, pid = %d, output to: %s" profiler_pid profile_fn));
+ Some profiler_pid
+
+(* kill profiler via SIGINT *)
+let stop_profiler_linux m_pid =
+ match m_pid with
+ | Some pid -> (
+ let _ = if !Flags.debug then Feedback.msg_debug (Pp.str "Stopping native code profiler") in
+ try
+ Unix.kill pid Sys.sigint;
+ let _ = Unix.waitpid [] pid in ()
+ with Unix.Unix_error (Unix.ESRCH,"kill","") ->
+ Feedback.msg_info (Pp.str "Could not stop native code profiler, no such process")
+ )
+ | None -> ()
+
+let start_profiler () =
+ let profile_fn = get_available_profile_filename () in
+ match profiler_platform () with
+ "Unix (Linux)" -> start_profiler_linux profile_fn
+ | _ ->
+ let _ = Feedback.msg_info
+ (Pp.str (Format.sprintf "Native_compute profiling not supported on the platform: %s"
+ (profiler_platform ()))) in
+ None
+
+let stop_profiler m_pid =
+ match profiler_platform() with
+ "Unix (Linux)" -> stop_profiler_linux m_pid
+ | _ -> ()
+
+let native_norm env sigma c ty =
+ let c = EConstr.Unsafe.to_constr c in
+ let ty = EConstr.Unsafe.to_constr ty in
+ if not Coq_config.native_compiler then
+ user_err Pp.(str "Native_compute reduction has been disabled at configure time.")
else
let penv = Environ.pre_env env in
- let sigma = evars_of_evar_map sigma in
(*
Format.eprintf "Numbers of free variables (named): %i\n" (List.length vl1);
Format.eprintf "Numbers of free variables (rel): %i\n" (List.length vl2);
*)
let ml_filename, prefix = Nativelib.get_ml_filename () in
- let code, upd = mk_norm_code penv sigma prefix c in
- match Nativelib.compile ml_filename code with
+ let code, upd = mk_norm_code penv (evars_of_evar_map sigma) prefix c in
+ let profile = get_profiling_enabled () in
+ match Nativelib.compile ml_filename code ~profile:profile with
| true, fn ->
if !Flags.debug then Feedback.msg_debug (Pp.str "Running norm ...");
+ let profiler_pid = if profile then start_profiler () else None in
let t0 = Sys.time () in
Nativelib.call_linker ~fatal:true prefix fn (Some upd);
let t1 = Sys.time () in
+ if profile then stop_profiler profiler_pid;
let time_info = Format.sprintf "Evaluation done in %.5f@." (t1 -. t0) in
if !Flags.debug then Feedback.msg_debug (Pp.str time_info);
- let res = nf_val env !Nativelib.rt1 ty in
+ let res = nf_val env sigma !Nativelib.rt1 ty in
let t2 = Sys.time () in
let time_info = Format.sprintf "Reification done in %.5f@." (t2 -. t1) in
if !Flags.debug then Feedback.msg_debug (Pp.str time_info);
- res
- | _ -> anomaly (Pp.str "Compilation failure")
+ EConstr.of_constr res
+ | _ -> anomaly (Pp.str "Compilation failure.")
let native_conv_generic pb sigma t =
Nativeconv.native_conv_gen pb (evars_of_evar_map sigma) t
diff --git a/pretyping/nativenorm.mli b/pretyping/nativenorm.mli
index 0b1ce8e5..67b7a2a4 100644
--- a/pretyping/nativenorm.mli
+++ b/pretyping/nativenorm.mli
@@ -1,16 +1,26 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
+
+open EConstr
open Environ
open Evd
(** This module implements normalization by evaluation to OCaml code *)
+val get_profile_filename : unit -> string
+val set_profile_filename : string -> unit
+
+val get_profiling_enabled : unit -> bool
+val set_profiling_enabled : bool -> unit
+
+
val native_norm : env -> evar_map -> constr -> types -> constr
(** Conversion with inference of universe constraints *)
diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml
index fe73b610..c9925087 100644
--- a/pretyping/patternops.ml
+++ b/pretyping/patternops.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 CErrors
@@ -12,15 +14,14 @@ open Names
open Globnames
open Nameops
open Term
+open Constr
open Vars
open Glob_term
-open Glob_ops
open Pp
open Mod_subst
open Misctypes
open Decl_kinds
open Pattern
-open Evd
open Environ
let case_info_pattern_eq i1 i2 =
@@ -44,8 +45,9 @@ let rec constr_pattern_eq p1 p2 = match p1, p2 with
Name.equal v1 v2 && constr_pattern_eq t1 t2 && constr_pattern_eq b1 b2
| PProd (v1, t1, b1), PProd (v2, t2, b2) ->
Name.equal v1 v2 && constr_pattern_eq t1 t2 && constr_pattern_eq b1 b2
-| PLetIn (v1, t1, b1), PLetIn (v2, t2, b2) ->
- Name.equal v1 v2 && constr_pattern_eq t1 t2 && constr_pattern_eq b1 b2
+| PLetIn (v1, b1, t1, c1), PLetIn (v2, b2, t2, c2) ->
+ Name.equal v1 v2 && constr_pattern_eq b1 b2 &&
+ Option.equal constr_pattern_eq t1 t2 && constr_pattern_eq c1 c2
| PSort s1, PSort s2 -> Miscops.glob_sort_eq s1 s2
| PMeta m1, PMeta m2 -> Option.equal Id.equal m1 m2
| PIf (t1, l1, r1), PIf (t2, l2, r2) ->
@@ -59,7 +61,11 @@ let rec constr_pattern_eq p1 p2 = match p1, p2 with
fixpoint_eq f1 f2
| PCoFix f1, PCoFix f2 ->
cofixpoint_eq f1 f2
-| _ -> false
+| PProj (p1, t1), PProj (p2, t2) ->
+ Projection.equal p1 p2 && constr_pattern_eq t1 t2
+| (PRef _ | PVar _ | PEvar _ | PRel _ | PApp _ | PSoApp _
+ | PLambda _ | PProd _ | PLetIn _ | PSort _ | PMeta _
+ | PIf _ | PCase _ | PFix _ | PCoFix _ | PProj _), _ -> false
(** FIXME: fixpoint and cofixpoint should be relativized to pattern *)
and pattern_eq (i1, j1, p1) (i2, j2, p2) =
@@ -76,8 +82,8 @@ and cofixpoint_eq (i1, r1) (i2, r2) =
and rec_declaration_eq (n1, c1, r1) (n2, c2, r2) =
Array.equal Name.equal n1 n2 &&
- Array.equal Term.eq_constr c1 c2 &&
- Array.equal Term.eq_constr r1 r2
+ Array.equal Constr.equal c1 c2 &&
+ Array.equal Constr.equal r1 r2
let rec occur_meta_pattern = function
| PApp (f,args) ->
@@ -85,7 +91,8 @@ let rec occur_meta_pattern = function
| PProj (_,arg) -> occur_meta_pattern arg
| PLambda (na,t,c) -> (occur_meta_pattern t) || (occur_meta_pattern c)
| PProd (na,t,c) -> (occur_meta_pattern t) || (occur_meta_pattern c)
- | PLetIn (na,t,c) -> (occur_meta_pattern t) || (occur_meta_pattern c)
+ | PLetIn (na,b,t,c) ->
+ Option.fold_left (fun b t -> b || occur_meta_pattern t) (occur_meta_pattern b) t || (occur_meta_pattern c)
| PIf (c,c1,c2) ->
(occur_meta_pattern c) ||
(occur_meta_pattern c1) || (occur_meta_pattern c2)
@@ -96,12 +103,37 @@ let rec occur_meta_pattern = function
| PMeta _ | PSoApp _ -> true
| PEvar _ | PVar _ | PRef _ | PRel _ | PSort _ | PFix _ | PCoFix _ -> false
+let rec occurn_pattern n = function
+ | PRel p -> Int.equal n p
+ | PApp (f,args) ->
+ (occurn_pattern n f) || (Array.exists (occurn_pattern n) args)
+ | PProj (_,arg) -> occurn_pattern n arg
+ | PLambda (na,t,c) -> (occurn_pattern n t) || (occurn_pattern (n+1) c)
+ | PProd (na,t,c) -> (occurn_pattern n t) || (occurn_pattern (n+1) c)
+ | PLetIn (na,b,t,c) ->
+ Option.fold_left (fun b t -> b || occurn_pattern n t) (occurn_pattern n b) t ||
+ (occurn_pattern (n+1) c)
+ | PIf (c,c1,c2) ->
+ (occurn_pattern n c) ||
+ (occurn_pattern n c1) || (occurn_pattern n c2)
+ | PCase(_,p,c,br) ->
+ (occurn_pattern n p) ||
+ (occurn_pattern n c) ||
+ (List.exists (fun (_,_,p) -> occurn_pattern n p) br)
+ | PMeta _ | PSoApp _ -> true
+ | PEvar (_,args) -> Array.exists (occurn_pattern n) args
+ | PVar _ | PRef _ | PSort _ -> false
+ | PFix fix -> not (noccurn n (mkFix fix))
+ | PCoFix cofix -> not (noccurn n (mkCoFix cofix))
+
+let noccurn_pattern n c = not (occurn_pattern n c)
+
exception BoundPattern;;
let rec head_pattern_bound t =
match t with
| PProd (_,_,b) -> head_pattern_bound b
- | PLetIn (_,_,b) -> head_pattern_bound b
+ | PLetIn (_,_,_,b) -> head_pattern_bound b
| PApp (c,args) -> head_pattern_bound c
| PIf (c,_,_) -> head_pattern_bound c
| PCase (_,p,c,br) -> head_pattern_bound c
@@ -112,19 +144,19 @@ let rec head_pattern_bound t =
-> raise BoundPattern
(* Perhaps they were arguments, but we don't beta-reduce *)
| PLambda _ -> raise BoundPattern
- | PCoFix _ -> anomaly ~label:"head_pattern_bound" (Pp.str "not a type")
+ | PCoFix _ -> anomaly ~label:"head_pattern_bound" (Pp.str "not a type.")
-let head_of_constr_reference c = match kind_of_term c with
+let head_of_constr_reference sigma c = match EConstr.kind sigma c with
| Const (sp,_) -> ConstRef sp
| Construct (sp,_) -> ConstructRef sp
| Ind (sp,_) -> IndRef sp
| Var id -> VarRef id
- | _ -> anomaly (Pp.str "Not a rigid reference")
+ | _ -> anomaly (Pp.str "Not a rigid reference.")
let pattern_of_constr env sigma t =
let rec pattern_of_constr env t =
let open Context.Rel.Declaration in
- match kind_of_term t with
+ match kind t with
| Rel n -> PRel n
| Meta n -> PMeta (Some (Id.of_string ("META" ^ string_of_int n)))
| Var id -> PVar id
@@ -132,7 +164,7 @@ let pattern_of_constr env sigma t =
| Sort (Prop Pos) -> PSort GSet
| Sort (Type _) -> PSort (GType [])
| Cast (c,_,_) -> pattern_of_constr env c
- | LetIn (na,c,t,b) -> PLetIn (na,pattern_of_constr env c,
+ | LetIn (na,c,t,b) -> PLetIn (na,pattern_of_constr env c,Some (pattern_of_constr env t),
pattern_of_constr (push_rel (LocalDef (na,c,t)) env) b)
| Prod (na,c,b) -> PProd (na,pattern_of_constr env c,
pattern_of_constr (push_rel (LocalAssum (na, c)) env) b)
@@ -140,31 +172,31 @@ let pattern_of_constr env sigma t =
pattern_of_constr (push_rel (LocalAssum (na, c)) env) b)
| App (f,a) ->
(match
- match kind_of_term f with
+ match kind f with
| Evar (evk,args) ->
(match snd (Evd.evar_source evk sigma) with
- Evar_kinds.MatchingVar (true,id) -> Some id
+ Evar_kinds.MatchingVar (Evar_kinds.SecondOrderPatVar id) -> Some id
| _ -> None)
| _ -> None
with
| Some n -> PSoApp (n,Array.to_list (Array.map (pattern_of_constr env) a))
| None -> PApp (pattern_of_constr env f,Array.map (pattern_of_constr env) a))
- | Const (sp,u) -> PRef (ConstRef (constant_of_kn(canonical_con sp)))
+ | Const (sp,u) -> PRef (ConstRef (Constant.make1 (Constant.canonical sp)))
| Ind (sp,u) -> PRef (canonical_gr (IndRef sp))
| Construct (sp,u) -> PRef (canonical_gr (ConstructRef sp))
| Proj (p, c) ->
- pattern_of_constr env (Retyping.expand_projection env sigma p c [])
+ pattern_of_constr env (EConstr.Unsafe.to_constr (Retyping.expand_projection env sigma p (EConstr.of_constr c) []))
| Evar (evk,ctxt as ev) ->
(match snd (Evd.evar_source evk sigma) with
- | Evar_kinds.MatchingVar (b,id) ->
- let ty = Evarutil.nf_evar sigma (existential_type sigma ev) in
- let () = ignore (pattern_of_constr env ty) in
- assert (not b); PMeta (Some id)
- | Evar_kinds.GoalEvar ->
- PEvar (evk,Array.map (pattern_of_constr env) ctxt)
+ | Evar_kinds.MatchingVar (Evar_kinds.FirstOrderPatVar id) ->
+ PMeta (Some id)
+ | Evar_kinds.GoalEvar | Evar_kinds.VarInstance _ ->
+ (* These are the two evar kinds used for existing goals *)
+ (* see Proofview.mark_in_evm *)
+ if Evd.is_defined sigma evk then pattern_of_constr env (Evd.existential_value sigma ev)
+ else PEvar (evk,Array.map (pattern_of_constr env) ctxt)
+ | Evar_kinds.MatchingVar (Evar_kinds.SecondOrderPatVar ido) -> assert false
| _ ->
- let ty = Evarutil.nf_evar sigma (existential_type sigma ev) in
- let () = ignore (pattern_of_constr env ty) in
PMeta None)
| Case (ci,p,a,br) ->
let cip =
@@ -189,7 +221,7 @@ let map_pattern_with_binders g f l = function
| PSoApp (n,pl) -> PSoApp (n, List.map (f l) pl)
| PLambda (n,a,b) -> PLambda (n,f l a,f (g n l) b)
| PProd (n,a,b) -> PProd (n,f l a,f (g n l) b)
- | PLetIn (n,a,b) -> PLetIn (n,f l a,f (g n l) b)
+ | PLetIn (n,a,t,b) -> PLetIn (n,f l a,Option.map (f l) t,f (g n l) b)
| PIf (c,b1,b2) -> PIf (f l c,f l b1,f l b2)
| PCase (ci,po,p,pl) ->
PCase (ci,f l po,f l p, List.map (fun (i,n,c) -> (i,n,f l c)) pl)
@@ -204,11 +236,13 @@ let error_instantiate_pattern id l =
| [_] -> "is"
| _ -> "are"
in
- errorlabstrm "" (str "Cannot substitute the term bound to " ++ pr_id id
- ++ strbrk " in pattern because the term refers to " ++ pr_enum pr_id l
+ user_err (str "Cannot substitute the term bound to " ++ Id.print id
+ ++ strbrk " in pattern because the term refers to " ++ pr_enum Id.print l
++ strbrk " which " ++ str is ++ strbrk " not bound in the pattern.")
let instantiate_pattern env sigma lvar c =
+ let open EConstr in
+ let open Vars in
let rec aux vars = function
| PVar id as x ->
(try
@@ -220,14 +254,16 @@ let instantiate_pattern env sigma lvar c =
ctx
in
let c = substl inst c in
- pattern_of_constr env sigma c
+ (** FIXME: Stupid workaround to pattern_of_constr being evar sensitive *)
+ let c = Evarutil.nf_evar sigma c in
+ pattern_of_constr env sigma (EConstr.Unsafe.to_constr c)
with Not_found (* List.index failed *) ->
let vars =
List.map_filter (function Name id -> Some id | _ -> None) vars in
error_instantiate_pattern id (List.subtract Id.equal ctx vars)
with Not_found (* Map.find failed *) ->
x)
- | (PFix _ | PCoFix _) -> error ("Non instantiable pattern.")
+ | (PFix _ | PCoFix _) -> user_err Pp.(str "Non instantiable pattern.")
| c ->
map_pattern_with_binders (fun id vars -> id::vars) aux vars c in
aux [] c
@@ -274,11 +310,12 @@ let rec subst_pattern subst pat =
let c2' = subst_pattern subst c2 in
if c1' == c1 && c2' == c2 then pat else
PProd (name,c1',c2')
- | PLetIn (name,c1,c2) ->
+ | PLetIn (name,c1,t,c2) ->
let c1' = subst_pattern subst c1 in
+ let t' = Option.smartmap (subst_pattern subst) t in
let c2' = subst_pattern subst c2 in
- if c1' == c1 && c2' == c2 then pat else
- PLetIn (name,c1',c2')
+ if c1' == c1 && t' == t && c2' == c2 then pat else
+ PLetIn (name,c1',t',c2')
| PSort _
| PMeta _ -> pat
| PIf (c,c1,c2) ->
@@ -315,52 +352,56 @@ let rec subst_pattern subst pat =
let mkPLambda na b = PLambda(na,PMeta None,b)
let rev_it_mkPLambda = List.fold_right mkPLambda
-let err loc pp = user_err_loc (loc,"pattern_of_glob_constr", pp)
+let err ?loc pp = user_err ?loc ~hdr:"pattern_of_glob_constr" pp
let warn_cast_in_pattern =
CWarnings.create ~name:"cast-in-pattern" ~category:"automation"
(fun () -> Pp.strbrk "Casts are ignored in patterns")
-let rec pat_of_raw metas vars = function
- | GVar (_,id) ->
+let rec pat_of_raw metas vars = DAst.with_loc_val (fun ?loc -> function
+ | GVar id ->
(try PRel (List.index Name.equal (Name id) vars)
with Not_found -> PVar id)
- | GPatVar (_,(false,n)) ->
+ | GPatVar (Evar_kinds.FirstOrderPatVar n) ->
metas := n::!metas; PMeta (Some n)
- | GRef (_,gr,_) ->
+ | GRef (gr,_) ->
PRef (canonical_gr gr)
(* Hack to avoid rewriting a complete interpretation of patterns *)
- | GApp (_, GPatVar (_,(true,n)), cl) ->
+ | GApp (c, cl) ->
+ begin match DAst.get c with
+ | GPatVar (Evar_kinds.SecondOrderPatVar n) ->
metas := n::!metas; PSoApp (n, List.map (pat_of_raw metas vars) cl)
- | GApp (_,c,cl) ->
+ | _ ->
PApp (pat_of_raw metas vars c,
Array.of_list (List.map (pat_of_raw metas vars) cl))
- | GLambda (_,na,bk,c1,c2) ->
- name_iter (fun n -> metas := n::!metas) na;
+ end
+ | GLambda (na,bk,c1,c2) ->
+ Name.iter (fun n -> metas := n::!metas) na;
PLambda (na, pat_of_raw metas vars c1,
pat_of_raw metas (na::vars) c2)
- | GProd (_,na,bk,c1,c2) ->
- name_iter (fun n -> metas := n::!metas) na;
+ | GProd (na,bk,c1,c2) ->
+ Name.iter (fun n -> metas := n::!metas) na;
PProd (na, pat_of_raw metas vars c1,
pat_of_raw metas (na::vars) c2)
- | GLetIn (_,na,c1,c2) ->
- name_iter (fun n -> metas := n::!metas) na;
+ | GLetIn (na,c1,t,c2) ->
+ Name.iter (fun n -> metas := n::!metas) na;
PLetIn (na, pat_of_raw metas vars c1,
+ Option.map (pat_of_raw metas vars) t,
pat_of_raw metas (na::vars) c2)
- | GSort (_,s) ->
+ | GSort s ->
PSort s
| GHole _ ->
PMeta None
- | GCast (_,c,_) ->
+ | GCast (c,_) ->
warn_cast_in_pattern ();
pat_of_raw metas vars c
- | GIf (_,c,(_,None),b1,b2) ->
+ | GIf (c,(_,None),b1,b2) ->
PIf (pat_of_raw metas vars c,
pat_of_raw metas vars b1,pat_of_raw metas vars b2)
- | GLetTuple (loc,nal,(_,None),b,c) ->
- let mkGLambda c na =
- GLambda (loc,na,Explicit,GHole (loc,Evar_kinds.InternalHole, IntroAnonymous, None),c) in
- let c = List.fold_left mkGLambda c nal in
+ | GLetTuple (nal,(_,None),b,c) ->
+ let mkGLambda na c = DAst.make ?loc @@
+ GLambda (na,Explicit, DAst.make @@ GHole (Evar_kinds.InternalHole, IntroAnonymous, None),c) in
+ let c = List.fold_right mkGLambda nal c in
let cip =
{ cip_style = LetStyle;
cip_ind = None;
@@ -370,24 +411,31 @@ let rec pat_of_raw metas vars = function
let tags = List.map (fun _ -> false) nal (* Approximation which can be without let-ins... *) in
PCase (cip, PMeta None, pat_of_raw metas vars b,
[0,tags,pat_of_raw metas vars c])
- | GCases (loc,sty,p,[c,(na,indnames)],brs) ->
+ | GCases (sty,p,[c,(na,indnames)],brs) ->
+ let get_ind p = match DAst.get p with
+ | PatCstr((ind,_),_,_) -> Some ind
+ | _ -> None
+ in
let get_ind = function
- | (_,_,[PatCstr(_,(ind,_),_,_)],_)::_ -> Some ind
+ | {CAst.v=(_,[p],_)}::_ -> get_ind p
| _ -> None
in
let ind_tags,ind = match indnames with
- | Some (_,ind,nal) -> Some (List.length nal), Some ind
+ | Some {CAst.v=(ind,nal)} -> Some (List.length nal), Some ind
| None -> None, get_ind brs
in
let ext,brs = pats_of_glob_branches loc metas vars ind brs
in
let pred = match p,indnames with
- | Some p, Some (_,_,nal) ->
+ | Some p, Some {CAst.v=(_,nal)} ->
let nvars = na :: List.rev nal @ vars in
rev_it_mkPLambda nal (mkPLambda na (pat_of_raw metas nvars p))
- | (None | Some (GHole _)), _ -> PMeta None
+ | None, _ -> PMeta None
| Some p, None ->
- user_err_loc (loc,"",strbrk "Clause \"in\" expected in patterns over \"match\" expressions with an explicit \"return\" clause.")
+ match DAst.get p with
+ | GHole _ -> PMeta None
+ | _ ->
+ user_err ?loc (strbrk "Clause \"in\" expected in patterns over \"match\" expressions with an explicit \"return\" clause.")
in
let info =
{ cip_style = sty;
@@ -400,32 +448,41 @@ let rec pat_of_raw metas vars = function
one non-trivial branch. These facts are used in [Constrextern]. *)
PCase (info, pred, pat_of_raw metas vars c, brs)
- | r -> err (loc_of_glob_constr r) (Pp.str "Non supported pattern.")
+ | GPatVar _ | GIf _ | GLetTuple _ | GCases _ | GEvar _ | GRec _ ->
+ err ?loc (Pp.str "Non supported pattern."))
and pats_of_glob_branches loc metas vars ind brs =
- let get_arg = function
- | PatVar(_,na) -> na
- | PatCstr(loc,_,_,_) -> err loc (Pp.str "Non supported pattern.")
+ let get_arg p = match DAst.get p with
+ | PatVar na ->
+ Name.iter (fun n -> metas := n::!metas) na;
+ na
+ | PatCstr(_,_,_) -> err ?loc:p.CAst.loc (Pp.str "Non supported pattern.")
in
let rec get_pat indexes = function
| [] -> false, []
- | [(_,_,[PatVar(_,Anonymous)],GHole _)] -> true, [] (* ends with _ => _ *)
- | (_,_,[PatCstr(_,(indsp,j),lv,_)],br) :: brs ->
- let () = match ind with
- | Some sp when eq_ind sp indsp -> ()
+ | {CAst.loc=loc';v=(_,[p], br)} :: brs ->
+ begin match DAst.get p, DAst.get br, brs with
+ | PatVar Anonymous, GHole _, [] ->
+ true, [] (* ends with _ => _ *)
+ | PatCstr((indsp,j),lv,_), _, _ ->
+ let () = match ind with
+ | Some sp when eq_ind sp indsp -> ()
+ | _ ->
+ err ?loc (Pp.str "All constructors must be in the same inductive type.")
+ in
+ if Int.Set.mem (j-1) indexes then
+ err ?loc
+ (str "No unique branch for " ++ int j ++ str"-th constructor.");
+ let lna = List.map get_arg lv in
+ let vars' = List.rev lna @ vars in
+ let pat = rev_it_mkPLambda lna (pat_of_raw metas vars' br) in
+ let ext,pats = get_pat (Int.Set.add (j-1) indexes) brs in
+ let tags = List.map (fun _ -> false) lv (* approximation, w/o let-in *) in
+ ext, ((j-1, tags, pat) :: pats)
| _ ->
- err loc (Pp.str "All constructors must be in the same inductive type.")
- in
- if Int.Set.mem (j-1) indexes then
- err loc
- (str "No unique branch for " ++ int j ++ str"-th constructor.");
- let lna = List.map get_arg lv in
- let vars' = List.rev lna @ vars in
- let pat = rev_it_mkPLambda lna (pat_of_raw metas vars' br) in
- let ext,pats = get_pat (Int.Set.add (j-1) indexes) brs in
- let tags = List.map (fun _ -> false) lv (* approximation, w/o let-in *) in
- ext, ((j-1, tags, pat) :: pats)
- | (loc,_,_,_) :: _ -> err loc (Pp.str "Non supported pattern.")
+ err ?loc:loc' (Pp.str "Non supported pattern.")
+ end
+ | {CAst.loc;v=(_,_,_)} :: _ -> err ?loc (Pp.str "Non supported pattern.")
in
get_pat Int.Set.empty brs
diff --git a/pretyping/patternops.mli b/pretyping/patternops.mli
index 1f63565d..9f087857 100644
--- a/pretyping/patternops.mli
+++ b/pretyping/patternops.mli
@@ -1,17 +1,20 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
+open EConstr
open Globnames
open Glob_term
open Mod_subst
open Misctypes
open Pattern
+open Ltac_pretype
(** {5 Functions on patterns} *)
@@ -21,6 +24,8 @@ val occur_meta_pattern : constr_pattern -> bool
val subst_pattern : substitution -> constr_pattern -> constr_pattern
+val noccurn_pattern : int -> constr_pattern -> bool
+
exception BoundPattern
(** [head_pattern_bound t] extracts the head variable/constant of the
@@ -32,13 +37,13 @@ val head_pattern_bound : constr_pattern -> global_reference
(** [head_of_constr_reference c] assumes [r] denotes a reference and
returns its label; raises an anomaly otherwise *)
-val head_of_constr_reference : Term.constr -> global_reference
+val head_of_constr_reference : Evd.evar_map -> constr -> global_reference
(** [pattern_of_constr c] translates a term [c] with metavariables into
a pattern; currently, no destructor (Cases, Fix, Cofix) and no
existential variable are allowed in [c] *)
-val pattern_of_constr : Environ.env -> Evd.evar_map -> constr -> constr_pattern
+val pattern_of_constr : Environ.env -> Evd.evar_map -> Constr.constr -> constr_pattern
(** [pattern_of_glob_constr l c] translates a term [c] with metavariables into
a pattern; variables bound in [l] are replaced by the pattern to which they
diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml
index 00b6100c..278a4761 100644
--- a/pretyping/pretype_errors.ml
+++ b/pretyping/pretype_errors.ml
@@ -1,26 +1,27 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Util
open Names
-open Term
open Environ
+open EConstr
open Type_errors
type unification_error =
- | OccurCheck of existential_key * constr
+ | OccurCheck of Evar.t * constr
| NotClean of existential * env * constr (* Constr is a variable not in scope *)
| NotSameArgSize
| NotSameHead
| NoCanonicalStructure
| ConversionFailed of env * constr * constr (* Non convertible closed terms *)
- | MetaOccurInBody of existential_key
- | InstanceNotSameType of existential_key * env * types * types
+ | MetaOccurInBody of Evar.t
+ | InstanceNotSameType of Evar.t * env * types * types
| UnifUnivInconsistency of Univ.univ_inconsistency
| CannotSolveConstraint of Evd.evar_constraint * unification_error
| ProblemBeyondCapabilities
@@ -31,14 +32,16 @@ type position_reporting = (position * int) * constr
type subterm_unification_error = bool * position_reporting * position_reporting * (constr * constr * unification_error) option
+type type_error = (constr, types) ptype_error
+
type pretype_error =
(* Old Case *)
| CantFindCaseType of constr
(* Type inference unification *)
| ActualTypeNotCoercible of unsafe_judgment * types * unification_error
(* Tactic unification *)
- | UnifOccurCheck of existential_key * constr
- | UnsolvableImplicit of existential_key * Evd.unsolvability_explanation option
+ | UnifOccurCheck of Evar.t * constr
+ | UnsolvableImplicit of Evar.t * Evd.unsolvability_explanation option
| CannotUnify of constr * constr * unification_error option
| CannotUnifyLocal of constr * constr * constr
| CannotUnifyBindingType of constr * constr
@@ -55,7 +58,7 @@ type pretype_error =
| TypingError of type_error
| CannotUnifyOccurrences of subterm_unification_error
| UnsatisfiableConstraints of
- (existential_key * Evar_kinds.t) option * Evar.Set.t option
+ (Evar.t * Evar_kinds.t) option * Evar.Set.t option
exception PretypeError of env * Evd.evar_map * pretype_error
@@ -64,43 +67,54 @@ let precatchable_exception = function
| Nametab.GlobalizationError _ -> true
| _ -> false
-let raise_pretype_error (loc,env,sigma,te) =
- Loc.raise loc (PretypeError(env,sigma,te))
+let raise_pretype_error ?loc (env,sigma,te) =
+ Loc.raise ?loc (PretypeError(env,sigma,te))
-let raise_located_type_error (loc,env,sigma,te) =
- Loc.raise loc (PretypeError(env,sigma,TypingError te))
+let raise_type_error ?loc (env,sigma,te) =
+ Loc.raise ?loc (PretypeError(env,sigma,TypingError te))
+let error_actual_type ?loc env sigma {uj_val=c;uj_type=actty} expty reason =
+ let j = {uj_val=c;uj_type=actty} in
+ raise_pretype_error ?loc
+ (env, sigma, ActualTypeNotCoercible (j, expty, reason))
-let error_actual_type_loc loc env sigma {uj_val=c;uj_type=actty} expty reason =
+let error_actual_type_core ?loc env sigma {uj_val=c;uj_type=actty} expty =
let j = {uj_val=c;uj_type=actty} in
- raise_pretype_error
- (loc, env, sigma, ActualTypeNotCoercible (j, expty, reason))
+ raise_type_error ?loc
+ (env, sigma, ActualType (j, expty))
+
+let error_cant_apply_not_functional ?loc env sigma rator randl =
+ raise_type_error ?loc
+ (env, sigma, CantApplyNonFunctional (rator, randl))
-let error_cant_apply_not_functional_loc loc env sigma rator randl =
- raise_located_type_error
- (loc, env, sigma, CantApplyNonFunctional (rator, Array.of_list randl))
+let error_cant_apply_bad_type ?loc env sigma (n,c,t) rator randl =
+ raise_type_error ?loc
+ (env, sigma,
+ CantApplyBadType ((n,c,t), rator, randl))
-let error_cant_apply_bad_type_loc loc env sigma (n,c,t) rator randl =
- raise_located_type_error
- (loc, env, sigma,
- CantApplyBadType ((n,c,t), rator, Array.of_list randl))
+let error_ill_formed_branch ?loc env sigma c i actty expty =
+ raise_type_error
+ ?loc (env, sigma, IllFormedBranch (c, i, actty, expty))
-let error_ill_formed_branch_loc loc env sigma c i actty expty =
- raise_located_type_error
- (loc, env, sigma, IllFormedBranch (c, i, actty, expty))
+let error_number_branches ?loc env sigma cj expn =
+ raise_type_error ?loc (env, sigma, NumberBranches (cj, expn))
-let error_number_branches_loc loc env sigma cj expn =
- raise_located_type_error (loc, env, sigma, NumberBranches (cj, expn))
+let error_case_not_inductive ?loc env sigma cj =
+ raise_type_error ?loc (env, sigma, CaseNotInductive cj)
-let error_case_not_inductive_loc loc env sigma cj =
- raise_located_type_error (loc, env, sigma, CaseNotInductive cj)
+let error_ill_typed_rec_body ?loc env sigma i na jl tys =
+ raise_type_error ?loc
+ (env, sigma, IllTypedRecBody (i, na, jl, tys))
-let error_ill_typed_rec_body_loc loc env sigma i na jl tys =
- raise_located_type_error
- (loc, env, sigma, IllTypedRecBody (i, na, jl, tys))
+let error_elim_arity ?loc env sigma pi s c j a =
+ raise_type_error ?loc
+ (env, sigma, ElimArity (pi, s, c, j, a))
-let error_not_a_type_loc loc env sigma j =
- raise_located_type_error (loc, env, sigma, NotAType j)
+let error_not_a_type ?loc env sigma j =
+ raise_type_error ?loc (env, sigma, NotAType j)
+
+let error_assumption ?loc env sigma j =
+ raise_type_error ?loc (env, sigma, BadAssumption j)
(*s Implicit arguments synthesis errors. It is hard to find
a precise location. *)
@@ -108,15 +122,12 @@ let error_not_a_type_loc loc env sigma j =
let error_occur_check env sigma ev c =
raise (PretypeError (env, sigma, UnifOccurCheck (ev,c)))
-let error_unsolvable_implicit loc env sigma evk explain =
- Loc.raise loc
+let error_unsolvable_implicit ?loc env sigma evk explain =
+ Loc.raise ?loc
(PretypeError (env, sigma, UnsolvableImplicit (evk, explain)))
-let error_cannot_unify_loc loc env sigma ?reason (m,n) =
- Loc.raise loc (PretypeError (env, sigma,CannotUnify (m,n,reason)))
-
-let error_cannot_unify env sigma ?reason (m,n) =
- raise (PretypeError (env, sigma,CannotUnify (m,n,reason)))
+let error_cannot_unify ?loc env sigma ?reason (m,n) =
+ Loc.raise ?loc (PretypeError (env, sigma,CannotUnify (m,n,reason)))
let error_cannot_unify_local env sigma (m,n,sn) =
raise (PretypeError (env, sigma,CannotUnifyLocal (m,n,sn)))
@@ -140,21 +151,21 @@ let error_non_linear_unification env sigma hdmeta t =
(*s Ml Case errors *)
-let error_cant_find_case_type_loc loc env sigma expr =
- raise_pretype_error (loc, env, sigma, CantFindCaseType expr)
+let error_cant_find_case_type ?loc env sigma expr =
+ raise_pretype_error ?loc (env, sigma, CantFindCaseType expr)
(*s Pretyping errors *)
-let error_unexpected_type_loc loc env sigma actty expty =
- raise_pretype_error (loc, env, sigma, UnexpectedType (actty, expty))
+let error_unexpected_type ?loc env sigma actty expty =
+ raise_pretype_error ?loc (env, sigma, UnexpectedType (actty, expty))
-let error_not_product_loc loc env sigma c =
- raise_pretype_error (loc, env, sigma, NotProduct c)
+let error_not_product ?loc env sigma c =
+ raise_pretype_error ?loc (env, sigma, NotProduct c)
(*s Error in conversion from AST to glob_constr *)
-let error_var_not_found_loc loc s =
- raise_pretype_error (loc, empty_env, Evd.empty, VarNotFound s)
+let error_var_not_found ?loc s =
+ raise_pretype_error ?loc (empty_env, Evd.empty, VarNotFound s)
(*s Typeclass errors *)
@@ -166,7 +177,7 @@ let unsatisfiable_constraints env evd ev comp =
| Some ev ->
let loc, kind = Evd.evar_source ev evd in
let err = UnsatisfiableConstraints (Some (ev, kind), comp) in
- Loc.raise loc (PretypeError (env,evd,err))
+ Loc.raise ?loc (PretypeError (env,evd,err))
let unsatisfiable_exception exn =
match exn with
diff --git a/pretyping/pretype_errors.mli b/pretyping/pretype_errors.mli
index 880f48e5..6f14d025 100644
--- a/pretyping/pretype_errors.mli
+++ b/pretyping/pretype_errors.mli
@@ -1,27 +1,30 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
+open Constr
open Environ
+open EConstr
open Type_errors
(** {6 The type of errors raised by the pretyper } *)
type unification_error =
- | OccurCheck of existential_key * constr
+ | OccurCheck of Evar.t * constr
| NotClean of existential * env * constr
| NotSameArgSize
| NotSameHead
| NoCanonicalStructure
| ConversionFailed of env * constr * constr
- | MetaOccurInBody of existential_key
- | InstanceNotSameType of existential_key * env * types * types
+ | MetaOccurInBody of Evar.t
+ | InstanceNotSameType of Evar.t * env * types * types
| UnifUnivInconsistency of Univ.univ_inconsistency
| CannotSolveConstraint of Evd.evar_constraint * unification_error
| ProblemBeyondCapabilities
@@ -32,14 +35,16 @@ type position_reporting = (position * int) * constr
type subterm_unification_error = bool * position_reporting * position_reporting * (constr * constr * unification_error) option
+type type_error = (constr, types) ptype_error
+
type pretype_error =
(** Old Case *)
| CantFindCaseType of constr
(** Type inference unification *)
| ActualTypeNotCoercible of unsafe_judgment * types * unification_error
(** Tactic Unification *)
- | UnifOccurCheck of existential_key * constr
- | UnsolvableImplicit of existential_key * Evd.unsolvability_explanation option
+ | UnifOccurCheck of Evar.t * constr
+ | UnsolvableImplicit of Evar.t * Evd.unsolvability_explanation option
| CannotUnify of constr * constr * unification_error option
| CannotUnifyLocal of constr * constr * constr
| CannotUnifyBindingType of constr * constr
@@ -56,7 +61,7 @@ type pretype_error =
| TypingError of type_error
| CannotUnifyOccurrences of subterm_unification_error
| UnsatisfiableConstraints of
- (existential_key * Evar_kinds.t) option * Evar.Set.t option
+ (Evar.t * Evar_kinds.t) option * Evar.Set.t option
(** unresolvable evar, connex component *)
exception PretypeError of env * Evd.evar_map * pretype_error
@@ -64,52 +69,60 @@ exception PretypeError of env * Evd.evar_map * pretype_error
val precatchable_exception : exn -> bool
(** Raising errors *)
-val error_actual_type_loc :
- Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> constr ->
+val error_actual_type :
+ ?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> constr ->
unification_error -> 'b
-val error_cant_apply_not_functional_loc :
- Loc.t -> env -> Evd.evar_map ->
- unsafe_judgment -> unsafe_judgment list -> 'b
+val error_actual_type_core :
+ ?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> constr -> 'b
+
+val error_cant_apply_not_functional :
+ ?loc:Loc.t -> env -> Evd.evar_map ->
+ unsafe_judgment -> unsafe_judgment array -> 'b
-val error_cant_apply_bad_type_loc :
- Loc.t -> env -> Evd.evar_map -> int * constr * constr ->
- unsafe_judgment -> unsafe_judgment list -> 'b
+val error_cant_apply_bad_type :
+ ?loc:Loc.t -> env -> Evd.evar_map -> int * constr * constr ->
+ unsafe_judgment -> unsafe_judgment array -> 'b
-val error_case_not_inductive_loc :
- Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> 'b
+val error_case_not_inductive :
+ ?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> 'b
-val error_ill_formed_branch_loc :
- Loc.t -> env -> Evd.evar_map ->
+val error_ill_formed_branch :
+ ?loc:Loc.t -> env -> Evd.evar_map ->
constr -> pconstructor -> constr -> constr -> 'b
-val error_number_branches_loc :
- Loc.t -> env -> Evd.evar_map ->
+val error_number_branches :
+ ?loc:Loc.t -> env -> Evd.evar_map ->
unsafe_judgment -> int -> 'b
-val error_ill_typed_rec_body_loc :
- Loc.t -> env -> Evd.evar_map ->
+val error_ill_typed_rec_body :
+ ?loc:Loc.t -> env -> Evd.evar_map ->
int -> Name.t array -> unsafe_judgment array -> types array -> 'b
-val error_not_a_type_loc :
- Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> 'b
+val error_elim_arity :
+ ?loc:Loc.t -> env -> Evd.evar_map ->
+ pinductive -> Sorts.family list -> constr ->
+ unsafe_judgment -> (Sorts.family * Sorts.family * arity_error) option -> 'b
+
+val error_not_a_type :
+ ?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> 'b
+
+val error_assumption :
+ ?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> 'b
val error_cannot_coerce : env -> Evd.evar_map -> constr * constr -> 'b
(** {6 Implicit arguments synthesis errors } *)
-val error_occur_check : env -> Evd.evar_map -> existential_key -> constr -> 'b
+val error_occur_check : env -> Evd.evar_map -> Evar.t -> constr -> 'b
val error_unsolvable_implicit :
- Loc.t -> env -> Evd.evar_map -> existential_key ->
+ ?loc:Loc.t -> env -> Evd.evar_map -> Evar.t ->
Evd.unsolvability_explanation option -> 'b
-val error_cannot_unify_loc : Loc.t -> env -> Evd.evar_map ->
+val error_cannot_unify : ?loc:Loc.t -> env -> Evd.evar_map ->
?reason:unification_error -> constr * constr -> 'b
-val error_cannot_unify : env -> Evd.evar_map -> ?reason:unification_error ->
- constr * constr -> 'b
-
val error_cannot_unify_local : env -> Evd.evar_map -> constr * constr * constr -> 'b
val error_cannot_find_well_typed_abstraction : env -> Evd.evar_map ->
@@ -126,24 +139,24 @@ val error_non_linear_unification : env -> Evd.evar_map ->
(** {6 Ml Case errors } *)
-val error_cant_find_case_type_loc :
- Loc.t -> env -> Evd.evar_map -> constr -> 'b
+val error_cant_find_case_type :
+ ?loc:Loc.t -> env -> Evd.evar_map -> constr -> 'b
(** {6 Pretyping errors } *)
-val error_unexpected_type_loc :
- Loc.t -> env -> Evd.evar_map -> constr -> constr -> 'b
+val error_unexpected_type :
+ ?loc:Loc.t -> env -> Evd.evar_map -> constr -> constr -> 'b
-val error_not_product_loc :
- Loc.t -> env -> Evd.evar_map -> constr -> 'b
+val error_not_product :
+ ?loc:Loc.t -> env -> Evd.evar_map -> constr -> 'b
(** {6 Error in conversion from AST to glob_constr } *)
-val error_var_not_found_loc : Loc.t -> Id.t -> 'b
+val error_var_not_found : ?loc:Loc.t -> Id.t -> 'b
(** {6 Typeclass errors } *)
-val unsatisfiable_constraints : env -> Evd.evar_map -> Evd.evar option ->
+val unsatisfiable_constraints : env -> Evd.evar_map -> Evar.t option ->
Evar.Set.t option -> 'a
val unsatisfiable_exception : exn -> bool
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index 4b6d10c6..c7db802f 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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) *)
(************************************************************************)
(* This file contains the syntax-directed part of the type inference
@@ -27,12 +29,13 @@ open Util
open Names
open Evd
open Term
-open Vars
open Termops
-open Reductionops
open Environ
+open EConstr
+open Vars
+open Reductionops
open Type_errors
-open Typeops
+open Typing
open Globnames
open Nameops
open Evarutil
@@ -41,25 +44,12 @@ open Pretype_errors
open Glob_term
open Glob_ops
open Evarconv
-open Pattern
open Misctypes
-open Sigma.Notations
-open Context.Named.Declaration
+open Ltac_pretype
+
+module NamedDecl = Context.Named.Declaration
type typing_constraint = OfType of types | IsType | WithoutTypeConstraint
-type var_map = constr_under_binders Id.Map.t
-type uconstr_var_map = Glob_term.closed_glob_constr Id.Map.t
-type unbound_ltac_var_map = Geninterp.Val.t Id.Map.t
-type ltac_var_map = {
- ltac_constrs : var_map;
- ltac_uconstrs : uconstr_var_map;
- ltac_idents: Id.t Id.Map.t;
- ltac_genargs : unbound_ltac_var_map;
-}
-type glob_constr_ltac_closure = ltac_var_map * glob_constr
-type pure_open_constr = evar_map * constr
-type 'a delayed_open =
- { delayed : 'r. Environ.env -> 'r Sigma.t -> ('a, 'r) Sigma.sigma }
(************************************************************************)
(* This concerns Cases *)
@@ -77,47 +67,46 @@ type t = {
(** Delay the computation of the evar extended environment *)
}
-let get_extra env =
+let get_extra env sigma =
let open Context.Named.Declaration in
let ids = List.map get_id (named_context env) in
let avoid = List.fold_right Id.Set.add ids Id.Set.empty in
- Context.Rel.fold_outside push_rel_decl_to_named_context
- (Environ.rel_context env) ~init:(empty_csubst, [], avoid, named_context env)
+ Context.Rel.fold_outside (fun d acc -> push_rel_decl_to_named_context sigma d acc)
+ (rel_context env) ~init:(empty_csubst, avoid, named_context env)
-let make_env env = { env = env; extra = lazy (get_extra env) }
+let make_env env sigma = { env = env; extra = lazy (get_extra env sigma) }
let rel_context env = rel_context env.env
-let push_rel d env = {
+let push_rel sigma d env = {
env = push_rel d env.env;
- extra = lazy (push_rel_decl_to_named_context d (Lazy.force env.extra));
+ extra = lazy (push_rel_decl_to_named_context sigma d (Lazy.force env.extra));
}
-let pop_rel_context n env = make_env (pop_rel_context n env.env)
+let pop_rel_context n env sigma = make_env (pop_rel_context n env.env) sigma
-let push_rel_context ctx env = {
+let push_rel_context sigma ctx env = {
env = push_rel_context ctx env.env;
- extra = lazy (List.fold_right push_rel_decl_to_named_context ctx (Lazy.force env.extra));
+ extra = lazy (List.fold_right (fun d acc -> push_rel_decl_to_named_context sigma d acc) ctx (Lazy.force env.extra));
}
let lookup_named id env = lookup_named id env.env
let e_new_evar env evdref ?src ?naming typ =
- let subst2 subst vsubst c = csubst_subst subst (replace_vars vsubst c) in
let open Context.Named.Declaration in
- let inst_vars = List.map (fun d -> mkVar (get_id d)) (named_context env.env) in
+ let inst_vars = List.map (get_id %> mkVar) (named_context env.env) in
let inst_rels = List.rev (rel_list 0 (nb_rel env.env)) in
- let (subst, vsubst, _, nc) = Lazy.force env.extra in
- let typ' = subst2 subst vsubst typ in
+ let (subst, _, nc) = Lazy.force env.extra in
+ let typ' = csubst_subst subst typ in
let instance = inst_rels @ inst_vars in
let sign = val_of_named_context nc in
- let sigma = Sigma.Unsafe.of_evar_map !evdref in
- let Sigma (e, sigma, _) = new_evar_instance sign sigma typ' ?src ?naming instance in
- evdref := Sigma.to_evar_map sigma;
+ let sigma = !evdref in
+ let (sigma, e) = new_evar_instance sign sigma typ' ?src ?naming instance in
+ evdref := sigma;
e
-let push_rec_types (lna,typarray,_) env =
+let push_rec_types sigma (lna,typarray,_) env =
let ctxt = Array.map2_i (fun i na t -> Context.Rel.Declaration.LocalAssum (na, lift i t)) lna typarray in
- Array.fold_left (fun e assum -> push_rel assum e) env ctxt
+ Array.fold_left (fun e assum -> push_rel sigma assum e) env ctxt
end
@@ -127,7 +116,11 @@ open ExtraEnv
exception Found of int array
-let search_guard loc env possible_indexes fixdefs =
+let nf_fix sigma (nas, cs, ts) =
+ let inj c = EConstr.to_constr sigma c in
+ (nas, Array.map inj cs, Array.map inj ts)
+
+let search_guard ?loc env possible_indexes fixdefs =
(* Standard situation with only one possibility for each fix. *)
(* We treat it separately in order to get proper error msg. *)
let is_singleton = function [_] -> true | _ -> false in
@@ -137,7 +130,7 @@ let search_guard loc env possible_indexes fixdefs =
(try check_fix env fix
with reraise ->
let (e, info) = CErrors.push reraise in
- let info = Loc.add_loc info loc in
+ let info = Option.cata (fun loc -> Loc.add_loc info loc) info loc in
iraise (e, info));
indexes
else
@@ -160,7 +153,7 @@ let search_guard loc env possible_indexes fixdefs =
with TypeError _ -> ())
(List.combinations possible_indexes);
let errmsg = "Cannot guess decreasing argument of fix." in
- user_err_loc (loc,"search_guard", Pp.str errmsg)
+ user_err ?loc ~hdr:"search_guard" (Pp.str errmsg)
with Found indexes -> indexes)
(* To force universe name declaration before use *)
@@ -170,8 +163,7 @@ let is_strict_universe_declarations () = !strict_universe_declarations
let _ =
Goptions.(declare_bool_option
- { optsync = true;
- optdepr = false;
+ { optdepr = false;
optname = "strict universe declaration";
optkey = ["Strict";"Universe";"Declaration"];
optread = is_strict_universe_declarations;
@@ -179,67 +171,86 @@ let _ =
let _ =
Goptions.(declare_bool_option
- { optsync = true;
- optdepr = false;
+ { optdepr = false;
optname = "minimization to Set";
optkey = ["Universe";"Minimization";"ToSet"];
optread = Universes.is_set_minimization;
optwrite = (:=) Universes.set_minimization })
-
+
(** Miscellaneous interpretation functions *)
-let interp_universe_level_name evd (loc,s) =
- let names, _ = Universes.global_universe_names () in
- if CString.string_contains s "." then
- match List.rev (CString.split '.' s) with
- | [] -> anomaly (str"Invalid universe name " ++ str s)
- | n :: dp ->
- let num = int_of_string n in
- let dp = DirPath.make (List.map Id.of_string dp) in
- let level = Univ.Level.make dp num in
- let evd =
- try Evd.add_global_univ evd level
- with UGraph.AlreadyDeclared -> evd
- in evd, level
- else
- try
- let level = Evd.universe_of_name evd s in
- evd, level
- with Not_found ->
- try
- let id = try Id.of_string s with _ -> raise Not_found in
- evd, snd (Idmap.find id names)
- with Not_found ->
- if not (is_strict_universe_declarations ()) then
- new_univ_level_variable ~loc ~name:s univ_rigid evd
- else user_err_loc (loc, "interp_universe_level_name",
- Pp.(str "Undeclared universe: " ++ str s))
+
+let interp_known_universe_level evd r =
+ let qid = Libnames.qualid_of_reference r in
+ try
+ match r.CAst.v with
+ | Libnames.Ident id -> Evd.universe_of_name evd id
+ | Libnames.Qualid _ -> raise Not_found
+ with Not_found ->
+ let univ, k = Nametab.locate_universe qid.CAst.v in
+ Univ.Level.make univ k
+
+let interp_universe_level_name ~anon_rigidity evd r =
+ try evd, interp_known_universe_level evd r
+ with Not_found ->
+ match r with (* Qualified generated name *)
+ | {CAst.loc; v=Libnames.Qualid qid} ->
+ let dp, i = Libnames.repr_qualid qid in
+ let num =
+ try int_of_string (Id.to_string i)
+ with Failure _ ->
+ user_err ?loc ~hdr:"interp_universe_level_name"
+ (Pp.(str "Undeclared global universe: " ++ Libnames.pr_reference r))
+ in
+ let level = Univ.Level.make dp num in
+ let evd =
+ try Evd.add_global_univ evd level
+ with UGraph.AlreadyDeclared -> evd
+ in evd, level
+ | {CAst.loc; v=Libnames.Ident id} -> (* Undeclared *)
+ if not (is_strict_universe_declarations ()) then
+ new_univ_level_variable ?loc ~name:id univ_rigid evd
+ else user_err ?loc ~hdr:"interp_universe_level_name"
+ (Pp.(str "Undeclared universe: " ++ Id.print id))
let interp_universe ?loc evd = function
| [] -> let evd, l = new_univ_level_variable ?loc univ_rigid evd in
evd, Univ.Universe.make l
| l ->
- List.fold_left (fun (evd, u) l ->
- let evd', l = interp_universe_level_name evd l in
- (evd', Univ.sup u (Univ.Universe.make l)))
+ List.fold_left (fun (evd, u) l ->
+ let evd', u' =
+ match l with
+ | Some (l,n) ->
+ (* [univ_flexible_alg] can produce algebraic universes in terms *)
+ let anon_rigidity = univ_flexible in
+ let evd', l = interp_universe_level_name ~anon_rigidity evd l in
+ let u' = Univ.Universe.make l in
+ (match n with
+ | 0 -> evd', u'
+ | 1 -> evd', Univ.Universe.super u'
+ | _ ->
+ user_err ?loc ~hdr:"interp_universe"
+ (Pp.(str "Cannot interpret universe increment +" ++ int n)))
+ | None ->
+ let evd, l = new_univ_level_variable ?loc univ_flexible evd in
+ evd, Univ.Universe.make l
+ in (evd', Univ.sup u u'))
(evd, Univ.Universe.type0m) l
-let interp_universe_level loc evd = function
- | None -> new_univ_level_variable ~loc univ_rigid evd
- | Some (loc,s) -> interp_universe_level_name evd (loc,s)
-
-let interp_sort ?loc evd = function
- | GProp -> evd, Prop Null
- | GSet -> evd, Prop Pos
- | GType n ->
- let evd, u = interp_universe ?loc evd n in
- evd, Type u
+let interp_known_level_info ?loc evd = function
+ | UUnknown | UAnonymous ->
+ user_err ?loc ~hdr:"interp_known_level_info"
+ (str "Anonymous universes not allowed here.")
+ | UNamed ref ->
+ try interp_known_universe_level evd ref
+ with Not_found ->
+ user_err ?loc ~hdr:"interp_known_level_info" (str "Undeclared universe " ++ Libnames.pr_reference ref)
-let interp_elimination_sort = function
- | GProp -> InProp
- | GSet -> InSet
- | GType _ -> InType
+let interp_level_info ?loc evd : Misctypes.level_info -> _ = function
+ | UUnknown -> new_univ_level_variable ?loc univ_rigid evd
+ | UAnonymous -> new_univ_level_variable ?loc univ_flexible evd
+ | UNamed s -> interp_universe_level_name ~anon_rigidity:univ_flexible evd s
-type inference_hook = env -> evar_map -> evar -> evar_map * constr
+type inference_hook = env -> evar_map -> Evar.t -> evar_map * constr
type inference_flags = {
use_typeclasses : bool;
@@ -257,16 +268,36 @@ type inference_flags = {
[sigma'] into those already in [sigma] or deriving from an evar in
[sigma] by restriction, and the evars properly created in [sigma'] *)
+type frozen =
+| FrozenId of evar_info Evar.Map.t
+ (** No pending evars. We do not put a set here not to reallocate like crazy,
+ but the actual data of the map is not used, only keys matter. All
+ functions operating on this type must have the same behaviour on
+ [FrozenId map] and [FrozenProgress (Evar.Map.domain map, Evar.Set.empty)] *)
+| FrozenProgress of (Evar.Set.t * Evar.Set.t) Lazy.t
+ (** Proper partition of the evar map as described above. *)
+
let frozen_and_pending_holes (sigma, sigma') =
- let add_derivative_of evk evi acc =
- match advance sigma' evk with None -> acc | Some evk' -> Evar.Set.add evk' acc in
- let frozen = Evd.fold_undefined add_derivative_of sigma Evar.Set.empty in
- let fold evk _ accu = if not (Evar.Set.mem evk frozen) then Evar.Set.add evk accu else accu in
- let pending = Evd.fold_undefined fold sigma' Evar.Set.empty in
- (frozen,pending)
+ let undefined0 = Evd.undefined_map sigma in
+ (** Fast path when the undefined evars where not modified *)
+ if undefined0 == Evd.undefined_map sigma' then
+ FrozenId undefined0
+ else
+ let data = lazy begin
+ let add_derivative_of evk evi acc =
+ match advance sigma' evk with None -> acc | Some evk' -> Evar.Set.add evk' acc in
+ let frozen = Evar.Map.fold add_derivative_of undefined0 Evar.Set.empty in
+ let fold evk _ accu = if not (Evar.Set.mem evk frozen) then Evar.Set.add evk accu else accu in
+ let pending = Evd.fold_undefined fold sigma' Evar.Set.empty in
+ (frozen, pending)
+ end in
+ FrozenProgress data
let apply_typeclasses env evdref frozen fail_evar =
- let filter_frozen evk = Evar.Set.mem evk frozen in
+ let filter_frozen = match frozen with
+ | FrozenId map -> fun evk -> Evar.Map.mem evk map
+ | FrozenProgress (lazy (frozen, _)) -> fun evk -> Evar.Set.mem evk frozen
+ in
evdref := Typeclasses.resolve_typeclasses
~filter:(if Flags.is_program_mode ()
then (fun evk evi -> Typeclasses.no_goals_or_obligations evk evi && not (filter_frozen evk))
@@ -276,13 +307,15 @@ let apply_typeclasses env evdref frozen fail_evar =
evdref := Typeclasses.resolve_typeclasses
~filter:(fun evk evi -> Typeclasses.all_evars evk evi && not (filter_frozen evk)) ~split:true ~fail:false env !evdref
-let apply_inference_hook hook evdref pending =
+let apply_inference_hook hook evdref frozen = match frozen with
+| FrozenId _ -> ()
+| FrozenProgress (lazy (_, pending)) ->
evdref := Evar.Set.fold (fun evk sigma ->
if Evd.is_undefined sigma evk (* in particular not defined by side-effect *)
then
try
let sigma, c = hook sigma evk in
- Evd.define evk c sigma
+ Evd.define evk (EConstr.Unsafe.to_constr c) sigma
with Exit ->
sigma
else
@@ -299,7 +332,9 @@ let check_typeclasses_instances_are_solved env current_sigma frozen =
(* Naive way, call resolution again with failure flag *)
apply_typeclasses env (ref current_sigma) frozen true
-let check_extra_evars_are_solved env current_sigma pending =
+let check_extra_evars_are_solved env current_sigma frozen = match frozen with
+| FrozenId _ -> ()
+| FrozenProgress (lazy (_, pending)) ->
Evar.Set.iter
(fun evk ->
if not (Evd.is_defined current_sigma evk) then
@@ -307,67 +342,76 @@ let check_extra_evars_are_solved env current_sigma pending =
match k with
| Evar_kinds.ImplicitArg (gr, (i, id), false) -> ()
| _ ->
- error_unsolvable_implicit loc env current_sigma evk None) pending
+ error_unsolvable_implicit ?loc env current_sigma evk None) pending
(* [check_evars] fails if some unresolved evar remains *)
let check_evars env initial_sigma sigma c =
let rec proc_rec c =
- match kind_of_term c with
- | Evar (evk,_ as ev) ->
- (match existential_opt_value sigma ev with
- | Some c -> proc_rec c
- | None ->
- if not (Evd.mem initial_sigma evk) then
- let (loc,k) = evar_source evk sigma in
- match k with
- | Evar_kinds.ImplicitArg (gr, (i, id), false) -> ()
- | _ -> Pretype_errors.error_unsolvable_implicit loc env sigma evk None)
- | _ -> Constr.iter proc_rec c
+ match EConstr.kind sigma c with
+ | Evar (evk, _) ->
+ if not (Evd.mem initial_sigma evk) then
+ let (loc,k) = evar_source evk sigma in
+ begin match k with
+ | Evar_kinds.ImplicitArg (gr, (i, id), false) -> ()
+ | _ -> Pretype_errors.error_unsolvable_implicit ?loc env sigma evk None
+ end
+ | _ -> EConstr.iter sigma proc_rec c
in proc_rec c
-let check_evars_are_solved env current_sigma frozen pending =
+let check_evars_are_solved env current_sigma frozen =
check_typeclasses_instances_are_solved env current_sigma frozen;
check_problems_are_solved env current_sigma;
- check_extra_evars_are_solved env current_sigma pending
+ check_extra_evars_are_solved env current_sigma frozen
(* Try typeclasses, hooks, unification heuristics ... *)
-let solve_remaining_evars flags env current_sigma pending =
- let frozen,pending = frozen_and_pending_holes pending in
+let solve_remaining_evars flags env current_sigma init_sigma =
+ let frozen = frozen_and_pending_holes (init_sigma, current_sigma) in
let evdref = ref current_sigma in
if flags.use_typeclasses then apply_typeclasses env evdref frozen false;
if Option.has_some flags.use_hook then
- apply_inference_hook (Option.get flags.use_hook env) evdref pending;
+ apply_inference_hook (Option.get flags.use_hook env) evdref frozen;
if flags.solve_unification_constraints then apply_heuristics env evdref false;
- if flags.fail_evar then check_evars_are_solved env !evdref frozen pending;
+ if flags.fail_evar then check_evars_are_solved env !evdref frozen;
!evdref
-let check_evars_are_solved env current_sigma pending =
- let frozen,pending = frozen_and_pending_holes pending in
- check_evars_are_solved env current_sigma frozen pending
+let check_evars_are_solved env current_sigma init_sigma =
+ let frozen = frozen_and_pending_holes (init_sigma, current_sigma) in
+ check_evars_are_solved env current_sigma frozen
-let process_inference_flags flags env initial_sigma (sigma,c) =
- let sigma = solve_remaining_evars flags env sigma (initial_sigma, sigma) in
+let process_inference_flags flags env initial_sigma (sigma,c,cty) =
+ let sigma = solve_remaining_evars flags env sigma initial_sigma in
let c = if flags.expand_evars then nf_evar sigma c else c in
- sigma,c
-
-(* Allow references to syntactically nonexistent variables (i.e., if applied on an inductive) *)
-let allow_anonymous_refs = ref false
+ sigma,c,cty
+
+let adjust_evar_source evdref na c =
+ match na, kind !evdref c with
+ | Name id, Evar (evk,args) ->
+ let evi = Evd.find !evdref evk in
+ begin match evi.evar_source with
+ | loc, Evar_kinds.QuestionMark (b,Anonymous) ->
+ let src = (loc,Evar_kinds.QuestionMark (b,na)) in
+ let (evd, evk') = restrict_evar !evdref evk (evar_filter evi) ~src None in
+ evdref := evd;
+ mkEvar (evk',args)
+ | _ -> c
+ end
+ | _, _ -> c
(* coerce to tycon if any *)
-let inh_conv_coerce_to_tycon resolve_tc loc env evdref j = function
+let inh_conv_coerce_to_tycon ?loc resolve_tc env evdref j = function
| None -> j
| Some t ->
- evd_comb2 (Coercion.inh_conv_coerce_to resolve_tc loc env.ExtraEnv.env) evdref j t
+ evd_comb2 (Coercion.inh_conv_coerce_to ?loc resolve_tc env.ExtraEnv.env) evdref j t
let check_instance loc subst = function
| [] -> ()
| (id,_) :: _ ->
if List.mem_assoc id subst then
- user_err_loc (loc,"",pr_id id ++ str "appears more than once.")
+ user_err ?loc (Id.print id ++ str "appears more than once.")
else
- user_err_loc (loc,"",str "No such variable in the signature of the existential variable: " ++ pr_id id ++ str ".")
+ user_err ?loc (str "No such variable in the signature of the existential variable: " ++ Id.print id ++ str ".")
(* used to enforce a name in Lambda when the type constraints itself
is named, hence possibly dependent *)
@@ -376,18 +420,7 @@ let orelse_name name name' = match name with
| Anonymous -> name'
| _ -> name
-let ltac_interp_name { ltac_idents ; ltac_genargs } = function
- | Anonymous -> Anonymous
- | Name id as n ->
- try Name (Id.Map.find id ltac_idents)
- with Not_found ->
- if Id.Map.mem id ltac_genargs then
- errorlabstrm "" (str"Ltac variable"++spc()++ pr_id id ++
- spc()++str"is not bound to an identifier."++spc()++
- str"It cannot be used in a binder.")
- else n
-
-let ltac_interp_name_env k0 lvar env =
+let ltac_interp_name_env k0 lvar env sigma =
(* envhd is the initial part of the env when pretype was called first *)
(* (in practice is is probably 0, but we have to grant the
specification of pretype which accepts to start with a non empty
@@ -398,20 +431,20 @@ let ltac_interp_name_env k0 lvar env =
let open Context.Rel.Declaration in
let ctxt' = List.smartmap (map_name (ltac_interp_name lvar)) ctxt in
if List.equal (fun d1 d2 -> Name.equal (get_name d1) (get_name d2)) ctxt ctxt' then env
- else push_rel_context ctxt' (pop_rel_context n env)
+ else push_rel_context sigma ctxt' (pop_rel_context n env sigma)
let invert_ltac_bound_name lvar env id0 id =
let id' = Id.Map.find id lvar.ltac_idents in
try mkRel (pi1 (lookup_rel_id id' (rel_context env)))
with Not_found ->
- errorlabstrm "" (str "Ltac variable " ++ pr_id id0 ++
- str " depends on pattern variable name " ++ pr_id id ++
+ user_err (str "Ltac variable " ++ Id.print id0 ++
+ str " depends on pattern variable name " ++ Id.print id ++
str " which is not bound in current context.")
let protected_get_type_of env sigma c =
try Retyping.get_type_of ~lax:true env.ExtraEnv.env sigma c
with Retyping.RetypeError _ ->
- errorlabstrm ""
+ user_err
(str "Cannot reinterpret " ++ quote (print_constr c) ++
str " in the current environment.")
@@ -422,7 +455,7 @@ let pretype_id pretype k0 loc env evdref lvar id =
let (n,_,typ) = lookup_rel_id id (rel_context env) in
{ uj_val = mkRel n; uj_type = lift n typ }
with Not_found ->
- let env = ltac_interp_name_env k0 lvar env in
+ let env = ltac_interp_name_env k0 lvar env !evdref in
(* Check if [id] is an ltac variable *)
try
let (ids,c) = Id.Map.find id lvar.ltac_constrs in
@@ -447,136 +480,154 @@ let pretype_id pretype k0 loc env evdref lvar id =
(* and build a nice error message *)
if Id.Map.mem id lvar.ltac_genargs then begin
let Geninterp.Val.Dyn (typ, _) = Id.Map.find id lvar.ltac_genargs in
- user_err_loc (loc,"",
- str "Variable " ++ pr_id id ++ str " should be bound to a term but is \
+ user_err ?loc
+ (str "Variable " ++ Id.print id ++ str " should be bound to a term but is \
bound to a " ++ Geninterp.Val.pr typ ++ str ".")
end;
(* Check if [id] is a section or goal variable *)
try
- { uj_val = mkVar id; uj_type = (get_type (lookup_named id env)) }
+ { uj_val = mkVar id; uj_type = NamedDecl.get_type (lookup_named id env) }
with Not_found ->
(* [id] not found, standard error message *)
- error_var_not_found_loc loc id
-
-let evar_kind_of_term sigma c =
- kind_of_term (whd_evar sigma c)
+ error_var_not_found ?loc id
(*************************************************************************)
(* Main pretyping function *)
-let interp_universe_level_name loc evd l =
- match l with
+let interp_known_glob_level ?loc evd = function
+ | GProp -> Univ.Level.prop
+ | GSet -> Univ.Level.set
+ | GType s -> interp_known_level_info ?loc evd s
+
+let interp_glob_level ?loc evd : Misctypes.glob_level -> _ = function
| GProp -> evd, Univ.Level.prop
| GSet -> evd, Univ.Level.set
- | GType s -> interp_universe_level loc evd s
+ | GType s -> interp_level_info ?loc evd s
+
+let interp_instance ?loc evd ~len l =
+ if len != List.length l then
+ user_err ?loc ~hdr:"pretype"
+ (str "Universe instance should have length " ++ int len)
+ else
+ let evd, l' =
+ List.fold_left
+ (fun (evd, univs) l ->
+ let evd, l = interp_glob_level ?loc evd l in
+ (evd, l :: univs)) (evd, [])
+ l
+ in
+ if List.exists (fun l -> Univ.Level.is_prop l) l' then
+ user_err ?loc ~hdr:"pretype"
+ (str "Universe instances cannot contain Prop, polymorphic" ++
+ str " universe instances must be greater or equal to Set.");
+ evd, Some (Univ.Instance.of_array (Array.of_list (List.rev l')))
-let pretype_global loc rigid env evd gr us =
+let pretype_global ?loc rigid env evd gr us =
let evd, instance =
match us with
| None -> evd, None
| Some l ->
- let _, ctx = Universes.unsafe_constr_of_global gr in
- let arr = Univ.Instance.to_array (Univ.UContext.instance ctx) in
- let len = Array.length arr in
- if len != List.length l then
- user_err_loc (loc, "pretype",
- str "Universe instance should have length " ++ int len)
- else
- let evd, l' = List.fold_left (fun (evd, univs) l ->
- let evd, l = interp_universe_level_name loc evd l in
- (evd, l :: univs)) (evd, []) l
- in
- if List.exists (fun l -> Univ.Level.is_prop l) l' then
- user_err_loc (loc, "pretype",
- str "Universe instances cannot contain Prop, polymorphic" ++
- str " universe instances must be greater or equal to Set.");
- evd, Some (Univ.Instance.of_array (Array.of_list (List.rev l')))
+ let _, ctx = Global.constr_of_global_in_context env.ExtraEnv.env gr in
+ let len = Univ.AUContext.size ctx in
+ interp_instance ?loc evd ~len l
in
- Evd.fresh_global ~loc ~rigid ?names:instance env.ExtraEnv.env evd gr
+ let (sigma, c) = Evd.fresh_global ?loc ~rigid ?names:instance env.ExtraEnv.env evd gr in
+ (sigma, EConstr.of_constr c)
-let pretype_ref loc evdref env ref us =
+let pretype_ref ?loc evdref env ref us =
match ref with
| VarRef id ->
(* Section variable *)
- (try make_judge (mkVar id) (get_type (lookup_named id env))
+ (try make_judge (mkVar id) (NamedDecl.get_type (lookup_named id env))
with Not_found ->
(* This may happen if env is a goal env and section variables have
been cleared - section variables should be different from goal
variables *)
- Pretype_errors.error_var_not_found_loc loc id)
+ Pretype_errors.error_var_not_found ?loc id)
| ref ->
- let evd, c = pretype_global loc univ_flexible env !evdref ref us in
+ let evd, c = pretype_global ?loc univ_flexible env !evdref ref us in
let () = evdref := evd in
- let ty = Typing.unsafe_type_of env.ExtraEnv.env evd c in
+ let ty = unsafe_type_of env.ExtraEnv.env evd c in
make_judge c ty
-let judge_of_Type loc evd s =
- let evd, s = interp_universe ~loc evd s in
+let judge_of_Type ?loc evd s =
+ let evd, s = interp_universe ?loc evd s in
let judge =
{ uj_val = mkSort (Type s); uj_type = mkSort (Type (Univ.super s)) }
in
evd, judge
-let pretype_sort loc evdref = function
+let pretype_sort ?loc evdref = function
| GProp -> judge_of_prop
| GSet -> judge_of_set
- | GType s -> evd_comb1 (judge_of_Type loc) evdref s
+ | GType s -> evd_comb1 (judge_of_Type ?loc) evdref s
let new_type_evar env evdref loc =
- let sigma = Sigma.Unsafe.of_evar_map !evdref in
- let Sigma ((e, _), sigma, _) =
+ let sigma = !evdref in
+ let (sigma, (e, _)) =
Evarutil.new_type_evar env.ExtraEnv.env sigma
univ_flexible_alg ~src:(loc,Evar_kinds.InternalHole)
in
- evdref := Sigma.to_evar_map sigma;
+ evdref := sigma;
e
-let (f_genarg_interp, genarg_interp_hook) = Hook.make ()
+module ConstrInterpObj =
+struct
+ type ('r, 'g, 't) obj =
+ unbound_ltac_var_map -> env -> evar_map -> types -> 'g -> constr * evar_map
+ let name = "constr_interp"
+ let default _ = None
+end
+
+module ConstrInterp = Genarg.Register(ConstrInterpObj)
+
+let register_constr_interp0 = ConstrInterp.register0
(* [pretype tycon env evdref lvar lmeta cstr] attempts to type [cstr] *)
(* in environment [env], with existential variables [evdref] and *)
(* the type constraint tycon *)
let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdref (lvar : ltac_var_map) t =
- let inh_conv_coerce_to_tycon = inh_conv_coerce_to_tycon resolve_tc in
+ let inh_conv_coerce_to_tycon ?loc = inh_conv_coerce_to_tycon ?loc resolve_tc in
let pretype_type = pretype_type k0 resolve_tc in
let pretype = pretype k0 resolve_tc in
let open Context.Rel.Declaration in
- match t with
- | GRef (loc,ref,u) ->
- inh_conv_coerce_to_tycon loc env evdref
- (pretype_ref loc evdref env ref u)
+ let loc = t.CAst.loc in
+ match DAst.get t with
+ | GRef (ref,u) ->
+ inh_conv_coerce_to_tycon ?loc env evdref
+ (pretype_ref ?loc evdref env ref u)
tycon
- | GVar (loc, id) ->
- inh_conv_coerce_to_tycon loc env evdref
+ | GVar id ->
+ inh_conv_coerce_to_tycon ?loc env evdref
(pretype_id (fun e r l t -> pretype tycon e r l t) k0 loc env evdref lvar id)
tycon
- | GEvar (loc, id, inst) ->
+ | GEvar (id, inst) ->
(* Ne faudrait-il pas s'assurer que hyps est bien un
sous-contexte du contexte courant, et qu'il n'y a pas de Rel "caché" *)
let evk =
try Evd.evar_key id !evdref
with Not_found ->
- user_err_loc (loc,"",str "Unknown existential variable.") in
+ user_err ?loc (str "Unknown existential variable.") in
let hyps = evar_filtered_context (Evd.find !evdref evk) in
let args = pretype_instance k0 resolve_tc env evdref lvar loc hyps evk inst in
let c = mkEvar (evk, args) in
let j = (Retyping.get_judgment_of env.ExtraEnv.env !evdref c) in
- inh_conv_coerce_to_tycon loc env evdref j tycon
+ inh_conv_coerce_to_tycon ?loc env evdref j tycon
- | GPatVar (loc,(someta,n)) ->
- let env = ltac_interp_name_env k0 lvar env in
+ | GPatVar kind ->
+ let env = ltac_interp_name_env k0 lvar env !evdref in
let ty =
match tycon with
| Some ty -> ty
| None -> new_type_evar env evdref loc in
- let k = Evar_kinds.MatchingVar (someta,n) in
+ let k = Evar_kinds.MatchingVar kind in
{ uj_val = e_new_evar env evdref ~src:(loc,k) ty; uj_type = ty }
- | GHole (loc, k, naming, None) ->
- let env = ltac_interp_name_env k0 lvar env in
+ | GHole (k, naming, None) ->
+ let env = ltac_interp_name_env k0 lvar env !evdref in
let ty =
match tycon with
| Some ty -> ty
@@ -584,37 +635,40 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
new_type_evar env evdref loc in
{ uj_val = e_new_evar env evdref ~src:(loc,k) ~naming ty; uj_type = ty }
- | GHole (loc, k, _naming, Some arg) ->
- let env = ltac_interp_name_env k0 lvar env in
+ | GHole (k, _naming, Some arg) ->
+ let env = ltac_interp_name_env k0 lvar env !evdref in
let ty =
match tycon with
| Some ty -> ty
| None ->
new_type_evar env evdref loc in
+ let open Genarg in
let ist = lvar.ltac_genargs in
- let (c, sigma) = Hook.get f_genarg_interp ty env.ExtraEnv.env !evdref ist arg in
+ let GenArg (Glbwit tag, arg) = arg in
+ let interp = ConstrInterp.obj tag in
+ let (c, sigma) = interp ist env.ExtraEnv.env !evdref ty arg in
let () = evdref := sigma in
{ uj_val = c; uj_type = ty }
- | GRec (loc,fixkind,names,bl,lar,vdef) ->
+ | GRec (fixkind,names,bl,lar,vdef) ->
let rec type_bl env ctxt = function
[] -> ctxt
| (na,bk,None,ty)::bl ->
let ty' = pretype_type empty_valcon env evdref lvar ty in
let dcl = LocalAssum (na, ty'.utj_val) in
let dcl' = LocalAssum (ltac_interp_name lvar na,ty'.utj_val) in
- type_bl (push_rel dcl env) (Context.Rel.add dcl' ctxt) bl
+ type_bl (push_rel !evdref dcl env) (Context.Rel.add dcl' ctxt) bl
| (na,bk,Some bd,ty)::bl ->
let ty' = pretype_type empty_valcon env evdref lvar ty in
let bd' = pretype (mk_tycon ty'.utj_val) env evdref lvar bd in
let dcl = LocalDef (na, bd'.uj_val, ty'.utj_val) in
let dcl' = LocalDef (ltac_interp_name lvar na, bd'.uj_val, ty'.utj_val) in
- type_bl (push_rel dcl env) (Context.Rel.add dcl' ctxt) bl in
+ type_bl (push_rel !evdref dcl env) (Context.Rel.add dcl' ctxt) bl in
let ctxtv = Array.map (type_bl env Context.Rel.empty) bl in
let larj =
Array.map2
(fun e ar ->
- pretype_type empty_valcon (push_rel_context e env) evdref lvar ar)
+ pretype_type empty_valcon (push_rel_context !evdref e env) evdref lvar ar)
ctxtv lar in
let lara = Array.map (fun a -> a.utj_val) larj in
let ftys = Array.map2 (fun e a -> it_mkProd_or_LetIn a e) ctxtv lara in
@@ -630,23 +684,24 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
| None -> true
in
(* Note: bodies are not used by push_rec_types, so [||] is safe *)
- let newenv = push_rec_types (names,ftys,[||]) env in
+ let newenv = push_rec_types !evdref (names,ftys,[||]) env in
let vdefj =
Array.map2_i
(fun i ctxt def ->
(* we lift nbfix times the type in tycon, because of
* the nbfix variables pushed to newenv *)
let (ctxt,ty) =
- decompose_prod_n_assum (Context.Rel.length ctxt)
+ decompose_prod_n_assum !evdref (Context.Rel.length ctxt)
(lift nbfix ftys.(i)) in
- let nenv = push_rel_context ctxt newenv in
+ let nenv = push_rel_context !evdref ctxt newenv in
let j = pretype (mk_tycon ty) nenv evdref lvar def in
{ uj_val = it_mkLambda_or_LetIn j.uj_val ctxt;
uj_type = it_mkProd_or_LetIn j.uj_type ctxt })
ctxtv vdef in
- Typing.check_type_fixpoint loc env.ExtraEnv.env evdref names ftys vdefj;
- let ftys = Array.map (nf_evar !evdref) ftys in
- let fdefs = Array.map (fun x -> nf_evar !evdref (j_val x)) vdefj in
+ Typing.check_type_fixpoint ?loc env.ExtraEnv.env evdref names ftys vdefj;
+ let nf c = nf_evar !evdref c in
+ let ftys = Array.map nf ftys in (** FIXME *)
+ let fdefs = Array.map (fun x -> nf (j_val x)) vdefj in
let fixj = match fixkind with
| GFix (vn,i) ->
(* First, let's find the guard indexes. *)
@@ -665,25 +720,26 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
let fixdecls = (names,ftys,fdefs) in
let indexes =
search_guard
- loc env.ExtraEnv.env possible_indexes fixdecls
+ ?loc env.ExtraEnv.env possible_indexes (nf_fix !evdref fixdecls)
in
make_judge (mkFix ((indexes,i),fixdecls)) ftys.(i)
| GCoFix i ->
- let cofix = (i,(names,ftys,fdefs)) in
- (try check_cofix env.ExtraEnv.env cofix
+ let fixdecls = (names,ftys,fdefs) in
+ let cofix = (i, fixdecls) in
+ (try check_cofix env.ExtraEnv.env (i, nf_fix !evdref fixdecls)
with reraise ->
let (e, info) = CErrors.push reraise in
- let info = Loc.add_loc info loc in
+ let info = Option.cata (Loc.add_loc info) info loc in
iraise (e, info));
make_judge (mkCoFix cofix) ftys.(i)
in
- inh_conv_coerce_to_tycon loc env evdref fixj tycon
+ inh_conv_coerce_to_tycon ?loc env evdref fixj tycon
- | GSort (loc,s) ->
- let j = pretype_sort loc evdref s in
- inh_conv_coerce_to_tycon loc env evdref j tycon
+ | GSort s ->
+ let j = pretype_sort ?loc evdref s in
+ inh_conv_coerce_to_tycon ?loc env evdref j tycon
- | GApp (loc,f,args) ->
+ | GApp (f,args) ->
let fj = pretype empty_tycon env evdref lvar f in
let floc = loc_of_glob_constr f in
let length = List.length args in
@@ -691,24 +747,24 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
(* Bidirectional typechecking hint:
parameters of a constructor are completely determined
by a typing constraint *)
- if Flags.is_program_mode () && length > 0 && isConstruct fj.uj_val then
+ if Flags.is_program_mode () && length > 0 && isConstruct !evdref fj.uj_val then
match tycon with
| None -> []
| Some ty ->
- let ((ind, i), u) = destConstruct fj.uj_val in
+ let ((ind, i), u) = destConstruct !evdref fj.uj_val in
let npars = inductive_nparams ind in
if Int.equal npars 0 then []
else
try
let IndType (indf, args) = find_rectype env.ExtraEnv.env !evdref ty in
let ((ind',u'),pars) = dest_ind_family indf in
- if eq_ind ind ind' then pars
+ if eq_ind ind ind' then List.map EConstr.of_constr pars
else (* Let the usual code throw an error *) []
with Not_found -> []
else []
in
let app_f =
- match kind_of_term fj.uj_val with
+ match EConstr.kind !evdref fj.uj_val with
| Const (p, u) when Environ.is_projection p env.ExtraEnv.env ->
let p = Projection.make p false in
let pb = Environ.lookup_projection p env.ExtraEnv.env in
@@ -724,7 +780,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
let argloc = loc_of_glob_constr c in
let resj = evd_comb1 (Coercion.inh_app_fun resolve_tc env.ExtraEnv.env) evdref resj in
let resty = whd_all env.ExtraEnv.env !evdref resj.uj_type in
- match kind_of_term resty with
+ match EConstr.kind !evdref resty with
| Prod (na,c1,c2) ->
let tycon = Some c1 in
let hj = pretype tycon env evdref lvar c in
@@ -736,57 +792,56 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
args, nf_evar !evdref (j_val hj)
else [], j_val hj
in
+ let ujval = adjust_evar_source evdref na ujval in
let value, typ = app_f n (j_val resj) ujval, subst1 ujval c2 in
let j = { uj_val = value; uj_type = typ } in
apply_rec env (n+1) j candargs rest
| _ ->
let hj = pretype empty_tycon env evdref lvar c in
- error_cant_apply_not_functional_loc
- (Loc.merge floc argloc) env.ExtraEnv.env !evdref
- resj [hj]
+ error_cant_apply_not_functional
+ ?loc:(Loc.merge_opt floc argloc) env.ExtraEnv.env !evdref
+ resj [|hj|]
in
let resj = apply_rec env 1 fj candargs args in
let resj =
- match evar_kind_of_term !evdref resj.uj_val with
+ match EConstr.kind !evdref resj.uj_val with
| App (f,args) ->
- let f = whd_evar !evdref f in
- if is_template_polymorphic env.ExtraEnv.env f then
+ if is_template_polymorphic env.ExtraEnv.env !evdref f then
(* Special case for inductive type applications that must be
refreshed right away. *)
- let sigma = !evdref in
- let c = mkApp (f,Array.map (whd_evar sigma) args) in
+ let c = mkApp (f, args) in
let c = evd_comb1 (Evarsolve.refresh_universes (Some true) env.ExtraEnv.env) evdref c in
let t = Retyping.get_type_of env.ExtraEnv.env !evdref c in
make_judge c (* use this for keeping evars: resj.uj_val *) t
else resj
| _ -> resj
in
- inh_conv_coerce_to_tycon loc env evdref resj tycon
+ inh_conv_coerce_to_tycon ?loc env evdref resj tycon
- | GLambda(loc,name,bk,c1,c2) ->
+ | GLambda(name,bk,c1,c2) ->
let tycon' = evd_comb1
(fun evd tycon ->
match tycon with
| None -> evd, tycon
| Some ty ->
- let evd, ty' = Coercion.inh_coerce_to_prod loc env.ExtraEnv.env evd ty in
+ let evd, ty' = Coercion.inh_coerce_to_prod ?loc env.ExtraEnv.env evd ty in
evd, Some ty')
evdref tycon
in
- let (name',dom,rng) = evd_comb1 (split_tycon loc env.ExtraEnv.env) evdref tycon' in
+ let (name',dom,rng) = evd_comb1 (split_tycon ?loc env.ExtraEnv.env) evdref tycon' in
let dom_valcon = valcon_of_tycon dom in
let j = pretype_type dom_valcon env evdref lvar c1 in
(* The name specified by ltac is used also to create bindings. So
the substitution must also be applied on variables before they are
looked up in the rel context. *)
let var = LocalAssum (name, j.utj_val) in
- let j' = pretype rng (push_rel var env) evdref lvar c2 in
+ let j' = pretype rng (push_rel !evdref var env) evdref lvar c2 in
let name = ltac_interp_name lvar name in
let resj = judge_of_abstraction env.ExtraEnv.env (orelse_name name name') j j' in
- inh_conv_coerce_to_tycon loc env evdref resj tycon
+ inh_conv_coerce_to_tycon ?loc env evdref resj tycon
- | GProd(loc,name,bk,c1,c2) ->
+ | GProd(name,bk,c1,c2) ->
let j = pretype_type empty_valcon env evdref lvar c1 in
(* The name specified by ltac is used also to create bindings. So
the substitution must also be applied on variables before they are
@@ -797,7 +852,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
{ j with utj_val = lift 1 j.utj_val }
| Name _ ->
let var = LocalAssum (name, j.utj_val) in
- let env' = push_rel var env in
+ let env' = push_rel !evdref var env in
pretype_type empty_valcon env' evdref lvar c2
in
let name = ltac_interp_name lvar name in
@@ -806,18 +861,18 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
judge_of_product env.ExtraEnv.env name j j'
with TypeError _ as e ->
let (e, info) = CErrors.push e in
- let info = Loc.add_loc info loc in
+ let info = Option.cata (Loc.add_loc info) info loc in
iraise (e, info) in
- inh_conv_coerce_to_tycon loc env evdref resj tycon
+ inh_conv_coerce_to_tycon ?loc env evdref resj tycon
- | GLetIn(loc,name,c1,c2) ->
- let j =
- match c1 with
- | GCast (loc, c, CastConv t) ->
- let tj = pretype_type empty_valcon env evdref lvar t in
- pretype (mk_tycon tj.utj_val) env evdref lvar c
- | _ -> pretype empty_tycon env evdref lvar c1
- in
+ | GLetIn(name,c1,t,c2) ->
+ let tycon1 =
+ match t with
+ | Some t ->
+ mk_tycon (pretype_type empty_valcon env evdref lvar t).utj_val
+ | None ->
+ empty_tycon in
+ let j = pretype tycon1 env evdref lvar c1 in
let t = evd_comb1 (Evarsolve.refresh_universes
~onlyalg:true ~status:Evd.univ_flexible (Some false) env.ExtraEnv.env)
evdref j.uj_type in
@@ -826,28 +881,29 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
looked up in the rel context. *)
let var = LocalDef (name, j.uj_val, t) in
let tycon = lift_tycon 1 tycon in
- let j' = pretype tycon (push_rel var env) evdref lvar c2 in
+ let j' = pretype tycon (push_rel !evdref var env) evdref lvar c2 in
let name = ltac_interp_name lvar name in
{ uj_val = mkLetIn (name, j.uj_val, t, j'.uj_val) ;
uj_type = subst1 j.uj_val j'.uj_type }
- | GLetTuple (loc,nal,(na,po),c,d) ->
+ | GLetTuple (nal,(na,po),c,d) ->
let cj = pretype empty_tycon env evdref lvar c in
let (IndType (indf,realargs)) =
try find_rectype env.ExtraEnv.env !evdref cj.uj_type
with Not_found ->
let cloc = loc_of_glob_constr c in
- error_case_not_inductive_loc cloc env.ExtraEnv.env !evdref cj
+ error_case_not_inductive ?loc:cloc env.ExtraEnv.env !evdref cj
in
let cstrs = get_constructors env.ExtraEnv.env indf in
if not (Int.equal (Array.length cstrs) 1) then
- user_err_loc (loc,"",str "Destructing let is only for inductive types" ++
+ user_err ?loc (str "Destructing let is only for inductive types" ++
str " with one constructor.");
let cs = cstrs.(0) in
if not (Int.equal (List.length nal) cs.cs_nargs) then
- user_err_loc (loc,"", str "Destructing let on this type expects " ++
+ user_err ?loc:loc (str "Destructing let on this type expects " ++
int cs.cs_nargs ++ str " variables.");
let fsign, record =
+ let set_name na d = set_name na (map_rel_decl EConstr.of_constr d) in
match get_projections env.ExtraEnv.env indf with
| None ->
List.map2 set_name (List.rev nal) cs.cs_args, false
@@ -855,6 +911,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
let rec aux n k names l =
match names, l with
| na :: names, (LocalAssum (_,t) :: l) ->
+ let t = EConstr.of_constr t in
let proj = Projection.make ps.(cs.cs_nargs - k) true in
LocalDef (na, lift (cs.cs_nargs - n) (mkProj (proj, cj.uj_val)), t)
:: aux (n+1) (k + 1) names l
@@ -863,6 +920,9 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
| [], [] -> []
| _ -> assert false
in aux 1 1 (List.rev nal) cs.cs_args, true in
+ let fsign = if Flags.version_strictly_greater Flags.V8_6
+ then Context.Rel.map (whd_betaiota !evdref) fsign
+ else fsign (* beta-iota-normalization regression in 8.5 and 8.6 *) in
let obj ind p v f =
if not record then
let nal = List.map (fun na -> ltac_interp_name lvar na) nal in
@@ -873,26 +933,29 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
mkCase (ci, p, cj.uj_val,[|f|])
else it_mkLambda_or_LetIn f fsign
in
- let env_f = push_rel_context fsign env in
+ let env_f = push_rel_context !evdref fsign env in
(* Make dependencies from arity signature impossible *)
let arsgn =
let arsgn,_ = get_arity env.ExtraEnv.env indf in
- if not !allow_anonymous_refs then
- List.map (set_name Anonymous) arsgn
- else arsgn
+ List.map (set_name Anonymous) arsgn
in
- let psign = LocalAssum (na, build_dependent_inductive env.ExtraEnv.env indf) :: arsgn in
+ let indt = build_dependent_inductive env.ExtraEnv.env indf in
+ let psign = LocalAssum (na, indt) :: arsgn in (* For locating names in [po] *)
+ let predlvar = Cases.make_return_predicate_ltac_lvar !evdref na c cj.uj_val lvar in
+ let psign' = LocalAssum (ltac_interp_name predlvar na, indt) :: arsgn in
+ let psign' = List.map (fun d -> map_rel_decl EConstr.of_constr d) psign' in
+ let psign' = Namegen.name_context env.ExtraEnv.env !evdref psign' in (* For naming abstractions in [po] *)
+ let psign = List.map (fun d -> map_rel_decl EConstr.of_constr d) psign in
let nar = List.length arsgn in
(match po with
| Some p ->
- let env_p = push_rel_context psign env in
- let pj = pretype_type empty_valcon env_p evdref lvar p in
+ let env_p = push_rel_context !evdref psign env in
+ let pj = pretype_type empty_valcon env_p evdref predlvar p in
let ccl = nf_evar !evdref pj.utj_val in
- let psign = make_arity_signature env.ExtraEnv.env true indf in (* with names *)
- let p = it_mkLambda_or_LetIn ccl psign in
+ let p = it_mkLambda_or_LetIn ccl psign' in
let inst =
- (Array.to_list cs.cs_concl_realargs)
- @[build_dependent_constructor cs] in
+ (Array.map_to_list EConstr.of_constr cs.cs_concl_realargs)
+ @[EConstr.of_constr (build_dependent_constructor cs)] in
let lp = lift cs.cs_nargs p in
let fty = hnf_lam_applist env.ExtraEnv.env !evdref lp inst in
let fj = pretype (mk_tycon fty) env_f evdref lvar d in
@@ -901,80 +964,84 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
Typing.check_allowed_sort env.ExtraEnv.env !evdref ind cj.uj_val p;
obj ind p cj.uj_val fj.uj_val
in
- { uj_val = v; uj_type = substl (realargs@[cj.uj_val]) ccl }
+ { uj_val = v; uj_type = (substl (realargs@[cj.uj_val]) ccl) }
| None ->
let tycon = lift_tycon cs.cs_nargs tycon in
- let fj = pretype tycon env_f evdref lvar d in
+ let fj = pretype tycon env_f evdref predlvar d in
let ccl = nf_evar !evdref fj.uj_type in
let ccl =
- if noccur_between 1 cs.cs_nargs ccl then
+ if noccur_between !evdref 1 cs.cs_nargs ccl then
lift (- cs.cs_nargs) ccl
else
- error_cant_find_case_type_loc loc env.ExtraEnv.env !evdref
+ error_cant_find_case_type ?loc env.ExtraEnv.env !evdref
cj.uj_val in
(* let ccl = refresh_universes ccl in *)
- let p = it_mkLambda_or_LetIn (lift (nar+1) ccl) psign in
+ let p = it_mkLambda_or_LetIn (lift (nar+1) ccl) psign' in
let v =
let ind,_ = dest_ind_family indf in
Typing.check_allowed_sort env.ExtraEnv.env !evdref ind cj.uj_val p;
obj ind p cj.uj_val fj.uj_val
in { uj_val = v; uj_type = ccl })
- | GIf (loc,c,(na,po),b1,b2) ->
+ | GIf (c,(na,po),b1,b2) ->
let cj = pretype empty_tycon env evdref lvar c in
let (IndType (indf,realargs)) =
try find_rectype env.ExtraEnv.env !evdref cj.uj_type
with Not_found ->
let cloc = loc_of_glob_constr c in
- error_case_not_inductive_loc cloc env.ExtraEnv.env !evdref cj in
+ error_case_not_inductive ?loc:cloc env.ExtraEnv.env !evdref cj in
let cstrs = get_constructors env.ExtraEnv.env indf in
if not (Int.equal (Array.length cstrs) 2) then
- user_err_loc (loc,"",
- str "If is only for inductive types with two constructors.");
+ user_err ?loc
+ (str "If is only for inductive types with two constructors.");
let arsgn =
let arsgn,_ = get_arity env.ExtraEnv.env indf in
- if not !allow_anonymous_refs then
- (* Make dependencies from arity signature impossible *)
- List.map (set_name Anonymous) arsgn
- else arsgn
+ (* Make dependencies from arity signature impossible *)
+ List.map (set_name Anonymous) arsgn
in
let nar = List.length arsgn in
- let psign = LocalAssum (na, build_dependent_inductive env.ExtraEnv.env indf) :: arsgn in
+ let indt = build_dependent_inductive env.ExtraEnv.env indf in
+ let psign = LocalAssum (na, indt) :: arsgn in (* For locating names in [po] *)
+ let predlvar = Cases.make_return_predicate_ltac_lvar !evdref na c cj.uj_val lvar in
+ let psign' = LocalAssum (ltac_interp_name predlvar na, indt) :: arsgn in
+ let psign' = List.map (fun d -> map_rel_decl EConstr.of_constr d) psign' in
+ let psign' = Namegen.name_context env.ExtraEnv.env !evdref psign' in (* For naming abstractions in [po] *)
+ let psign = List.map (fun d -> map_rel_decl EConstr.of_constr d) psign in
let pred,p = match po with
| Some p ->
- let env_p = push_rel_context psign env in
- let pj = pretype_type empty_valcon env_p evdref lvar p in
+ let env_p = push_rel_context !evdref psign env in
+ let pj = pretype_type empty_valcon env_p evdref predlvar p in
let ccl = nf_evar !evdref pj.utj_val in
- let pred = it_mkLambda_or_LetIn ccl psign in
- let typ = lift (- nar) (beta_applist (pred,[cj.uj_val])) in
+ let pred = it_mkLambda_or_LetIn ccl psign' in
+ let typ = lift (- nar) (beta_applist !evdref (pred,[cj.uj_val])) in
pred, typ
| None ->
let p = match tycon with
| Some ty -> ty
| None ->
- let env = ltac_interp_name_env k0 lvar env in
+ let env = ltac_interp_name_env k0 lvar env !evdref in
new_type_evar env evdref loc
in
- it_mkLambda_or_LetIn (lift (nar+1) p) psign, p in
+ it_mkLambda_or_LetIn (lift (nar+1) p) psign', p in
let pred = nf_evar !evdref pred in
let p = nf_evar !evdref p in
let f cs b =
let n = Context.Rel.length cs.cs_args in
let pi = lift n pred in (* liftn n 2 pred ? *)
- let pi = beta_applist (pi, [build_dependent_constructor cs]) in
+ let pi = beta_applist !evdref (pi, [EConstr.of_constr (build_dependent_constructor cs)]) in
+ let cs_args = List.map (fun d -> map_rel_decl EConstr.of_constr d) cs.cs_args in
+ let cs_args =
+ if Flags.version_strictly_greater Flags.V8_6
+ then Context.Rel.map (whd_betaiota !evdref) cs_args
+ else cs_args (* beta-iota-normalization regression in 8.5 and 8.6 *) in
let csgn =
- if not !allow_anonymous_refs then
- List.map (set_name Anonymous) cs.cs_args
- else
- List.map (map_name (function Name _ as n -> n
- | Anonymous -> Name Namegen.default_non_dependent_ident))
- cs.cs_args
- in
- let env_c = push_rel_context csgn env in
+ List.map (set_name Anonymous) cs_args
+ in
+ let env_c = push_rel_context !evdref csgn env in
let bj = pretype (mk_tycon pi) env_c evdref lvar b in
- it_mkLambda_or_LetIn bj.uj_val cs.cs_args in
+ it_mkLambda_or_LetIn bj.uj_val cs_args in
let b1 = f cstrs.(0) b1 in
let b2 = f cstrs.(1) b2 in
let v =
@@ -985,19 +1052,19 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
mkCase (ci, pred, cj.uj_val, [|b1;b2|])
in
let cj = { uj_val = v; uj_type = p } in
- inh_conv_coerce_to_tycon loc env evdref cj tycon
+ inh_conv_coerce_to_tycon ?loc env evdref cj tycon
- | GCases (loc,sty,po,tml,eqns) ->
- Cases.compile_cases loc sty
- ((fun vtyc env evdref -> pretype vtyc (make_env env) evdref lvar),evdref)
- tycon env.ExtraEnv.env (* loc *) (po,tml,eqns)
+ | GCases (sty,po,tml,eqns) ->
+ Cases.compile_cases ?loc sty
+ ((fun vtyc env evdref -> pretype vtyc (make_env env !evdref) evdref),evdref)
+ tycon env.ExtraEnv.env (* loc *) lvar (po,tml,eqns)
- | GCast (loc,c,k) ->
+ | GCast (c,k) ->
let cj =
match k with
| CastCoerce ->
let cj = pretype empty_tycon env evdref lvar c in
- evd_comb1 (Coercion.inh_coerce_to_base loc env.ExtraEnv.env) evdref cj
+ evd_comb1 (Coercion.inh_coerce_to_base ?loc env.ExtraEnv.env) evdref cj
| CastConv t | CastVM t | CastNative t ->
let k = (match k with CastVM _ -> VMcast | CastNative _ -> NATIVEcast | _ -> DEFAULTcast) in
let tj = pretype_type empty_valcon env evdref lvar t in
@@ -1009,13 +1076,13 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
| VMcast ->
let cj = pretype empty_tycon env evdref lvar c in
let cty = nf_evar !evdref cj.uj_type and tval = nf_evar !evdref tval in
- if not (occur_existential cty || occur_existential tval) then
+ if not (occur_existential !evdref cty || occur_existential !evdref tval) then
let (evd,b) = Reductionops.vm_infer_conv env.ExtraEnv.env !evdref cty tval in
if b then (evdref := evd; cj, tval)
else
- error_actual_type_loc loc env.ExtraEnv.env !evdref cj tval
+ error_actual_type ?loc env.ExtraEnv.env !evdref cj tval
(ConversionFailed (env.ExtraEnv.env,cty,tval))
- else user_err_loc (loc,"",str "Cannot check cast with vm: " ++
+ else user_err ?loc (str "Cannot check cast with vm: " ++
str "unresolved arguments remain.")
| NATIVEcast ->
let cj = pretype empty_tycon env evdref lvar c in
@@ -1024,7 +1091,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
let (evd,b) = Nativenorm.native_infer_conv env.ExtraEnv.env !evdref cty tval in
if b then (evdref := evd; cj, tval)
else
- error_actual_type_loc loc env.ExtraEnv.env !evdref cj tval
+ error_actual_type ?loc env.ExtraEnv.env !evdref cj tval
(ConversionFailed (env.ExtraEnv.env,cty,tval))
end
| _ ->
@@ -1032,12 +1099,12 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
in
let v = mkCast (cj.uj_val, k, tval) in
{ uj_val = v; uj_type = tval }
- in inh_conv_coerce_to_tycon loc env evdref cj tycon
+ in inh_conv_coerce_to_tycon ?loc env evdref cj tycon
and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update =
let f decl (subst,update) =
- let id = get_id decl in
- let t = replace_vars subst (get_type decl) in
+ let id = NamedDecl.get_id decl in
+ let t = replace_vars subst (EConstr.of_constr (NamedDecl.get_type decl)) in
let c, update =
try
let c = List.assoc id update in
@@ -1046,66 +1113,75 @@ and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update =
with Not_found ->
try
let (n,_,t') = lookup_rel_id id (rel_context env) in
- if is_conv env.ExtraEnv.env !evdref t t' then mkRel n, update else raise Not_found
+ if is_conv env.ExtraEnv.env !evdref t (lift n t') then mkRel n, update else raise Not_found
with Not_found ->
try
- let t' = lookup_named id env |> get_type in
+ let t' = env |> lookup_named id |> NamedDecl.get_type in
if is_conv env.ExtraEnv.env !evdref t t' then mkVar id, update else raise Not_found
with Not_found ->
- user_err_loc (loc,"",str "Cannot interpret " ++
+ user_err ?loc (str "Cannot interpret " ++
pr_existential_key !evdref evk ++
- str " in current context: no binding for " ++ pr_id id ++ str ".") in
+ str " in current context: no binding for " ++ Id.print id ++ str ".") in
((id,c)::subst, update) in
let subst,inst = List.fold_right f hyps ([],update) in
check_instance loc subst inst;
Array.map_of_list snd subst
(* [pretype_type valcon env evdref lvar c] coerces [c] into a type *)
-and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function
- | GHole (loc, knd, naming, None) ->
+and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar c = match DAst.get c with
+ | GHole (knd, naming, None) ->
+ let loc = loc_of_glob_constr c in
(match valcon with
| Some v ->
let s =
let sigma = !evdref in
let t = Retyping.get_type_of env.ExtraEnv.env sigma v in
- match kind_of_term (whd_all env.ExtraEnv.env sigma t) with
- | Sort s -> s
- | Evar ev when is_Type (existential_type sigma ev) ->
+ match EConstr.kind sigma (whd_all env.ExtraEnv.env sigma t) with
+ | Sort s -> ESorts.kind sigma s
+ | Evar ev when is_Type sigma (existential_type sigma ev) ->
evd_comb1 (define_evar_as_sort env.ExtraEnv.env) evdref ev
- | _ -> anomaly (Pp.str "Found a type constraint which is not a type")
+ | _ -> anomaly (Pp.str "Found a type constraint which is not a type.")
in
- { utj_val = v;
- utj_type = s }
+ (* Correction of bug #5315 : we need to define an evar for *all* holes *)
+ let evkt = e_new_evar env evdref ~src:(loc, knd) ~naming (mkSort s) in
+ let ev,_ = destEvar !evdref evkt in
+ evdref := Evd.define ev (to_constr !evdref v) !evdref;
+ (* End of correction of bug #5315 *)
+ { utj_val = v;
+ utj_type = s }
| None ->
- let env = ltac_interp_name_env k0 lvar env in
+ let env = ltac_interp_name_env k0 lvar env !evdref in
let s = evd_comb0 (new_sort_variable univ_flexible_alg) evdref in
{ utj_val = e_new_evar env evdref ~src:(loc, knd) ~naming (mkSort s);
utj_type = s})
- | c ->
+ | _ ->
let j = pretype k0 resolve_tc empty_tycon env evdref lvar c in
let loc = loc_of_glob_constr c in
- let tj = evd_comb1 (Coercion.inh_coerce_to_sort loc env.ExtraEnv.env) evdref j in
+ let tj = evd_comb1 (Coercion.inh_coerce_to_sort ?loc env.ExtraEnv.env) evdref j in
match valcon with
| None -> tj
| Some v ->
if e_cumul env.ExtraEnv.env evdref v tj.utj_val then tj
else
- error_unexpected_type_loc
- (loc_of_glob_constr c) env.ExtraEnv.env !evdref tj.utj_val v
+ error_unexpected_type
+ ?loc:(loc_of_glob_constr c) env.ExtraEnv.env !evdref tj.utj_val v
let ise_pretype_gen flags env sigma lvar kind c =
- let env = make_env env in
+ let env = make_env env sigma in
let evdref = ref sigma in
let k0 = Context.Rel.length (rel_context env) in
- let c' = match kind with
+ let c', c'_ty = match kind with
| WithoutTypeConstraint ->
- (pretype k0 flags.use_typeclasses empty_tycon env evdref lvar c).uj_val
+ let j = pretype k0 flags.use_typeclasses empty_tycon env evdref lvar c in
+ j.uj_val, j.uj_type
| OfType exptyp ->
- (pretype k0 flags.use_typeclasses (mk_tycon exptyp) env evdref lvar c).uj_val
+ let j = pretype k0 flags.use_typeclasses (mk_tycon exptyp) env evdref lvar c in
+ j.uj_val, j.uj_type
| IsType ->
- (pretype_type k0 flags.use_typeclasses empty_valcon env evdref lvar c).utj_val
+ let tj = pretype_type k0 flags.use_typeclasses empty_valcon env evdref lvar c in
+ tj.utj_val, mkSort tj.utj_type
in
- process_inference_flags flags env.ExtraEnv.env sigma (!evdref,c')
+ process_inference_flags flags env.ExtraEnv.env sigma (!evdref,c',c'_ty)
let default_inference_flags fail = {
use_typeclasses = true;
@@ -1124,40 +1200,9 @@ let no_classes_no_fail_inference_flags = {
let all_and_fail_flags = default_inference_flags true
let all_no_fail_flags = default_inference_flags false
-let empty_lvar : ltac_var_map = {
- ltac_constrs = Id.Map.empty;
- ltac_uconstrs = Id.Map.empty;
- ltac_idents = Id.Map.empty;
- ltac_genargs = Id.Map.empty;
-}
-
-let on_judgment f j =
- let c = mkCast(j.uj_val,DEFAULTcast, j.uj_type) in
- let (c,_,t) = destCast (f c) in
- {uj_val = c; uj_type = t}
-
-let understand_judgment env sigma c =
- let env = make_env env in
- let evdref = ref sigma in
- let k0 = Context.Rel.length (rel_context env) in
- let j = pretype k0 true empty_tycon env evdref empty_lvar c in
- let j = on_judgment (fun c ->
- let evd, c = process_inference_flags all_and_fail_flags env.ExtraEnv.env sigma (!evdref,c) in
- evdref := evd; c) j
- in j, Evd.evar_universe_context !evdref
-
-let understand_judgment_tcc env evdref c =
- let env = make_env env in
- let k0 = Context.Rel.length (rel_context env) in
- let j = pretype k0 true empty_tycon env evdref empty_lvar c in
- on_judgment (fun c ->
- let (evd,c) = process_inference_flags all_no_fail_flags env.ExtraEnv.env Evd.empty (!evdref,c) in
- evdref := evd; c) j
-
let ise_pretype_gen_ctx flags env sigma lvar kind c =
- let evd, c = ise_pretype_gen flags env sigma lvar kind c in
- let evd, f = Evarutil.nf_evars_and_universes evd in
- f c, Evd.evar_universe_context evd
+ let evd, c, _ = ise_pretype_gen flags env sigma lvar kind c in
+ c, Evd.evar_universe_context evd
(** Entry points of the high-level type synthesis algorithm *)
@@ -1167,42 +1212,19 @@ let understand
env sigma c =
ise_pretype_gen_ctx flags env sigma empty_lvar expected_type c
-let understand_tcc ?(flags=all_no_fail_flags) env sigma ?(expected_type=WithoutTypeConstraint) c =
+let understand_tcc_ty ?(flags=all_no_fail_flags) env sigma ?(expected_type=WithoutTypeConstraint) c =
ise_pretype_gen flags env sigma empty_lvar expected_type c
-let understand_tcc_evars ?(flags=all_no_fail_flags) env evdref ?(expected_type=WithoutTypeConstraint) c =
- let sigma, c = ise_pretype_gen flags env !evdref empty_lvar expected_type c in
- evdref := sigma;
- c
+let understand_tcc ?flags env sigma ?expected_type c =
+ let sigma, c, _ = understand_tcc_ty ?flags env sigma ?expected_type c in
+ sigma, c
let understand_ltac flags env sigma lvar kind c =
- ise_pretype_gen flags env sigma lvar kind c
-
-let constr_flags = {
- use_typeclasses = true;
- solve_unification_constraints = true;
- use_hook = None;
- fail_evar = true;
- expand_evars = true }
-
-(* Fully evaluate an untyped constr *)
-let type_uconstr ?(flags = constr_flags)
- ?(expected_type = WithoutTypeConstraint) ist c =
- { delayed = begin fun env sigma ->
- let { closure; term } = c in
- let vars = {
- ltac_constrs = closure.typed;
- ltac_uconstrs = closure.untyped;
- ltac_idents = closure.idents;
- ltac_genargs = Id.Map.empty;
- } in
- let sigma = Sigma.to_evar_map sigma in
- let (sigma, c) = understand_ltac flags env sigma vars expected_type term in
- Sigma.Unsafe.of_pair (c, sigma)
- end }
+ let (sigma, c, _) = ise_pretype_gen flags env sigma lvar kind c in
+ (sigma, c)
let pretype k0 resolve_tc typcon env evdref lvar t =
- pretype k0 resolve_tc typcon (make_env env) evdref lvar t
+ pretype k0 resolve_tc typcon (make_env env !evdref) evdref lvar t
let pretype_type k0 resolve_tc valcon env evdref lvar t =
- pretype_type k0 resolve_tc valcon (make_env env) evdref lvar t
+ pretype_type k0 resolve_tc valcon (make_env env !evdref) evdref lvar t
diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli
index 0f3f7c3c..415c4e17 100644
--- a/pretyping/pretyping.mli
+++ b/pretyping/pretyping.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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) *)
(************************************************************************)
(** This file implements type inference. It maps [glob_constr]
@@ -12,42 +14,24 @@
into elementary ones, insertion of coercions and resolution of
implicit arguments. *)
-open Names
-open Term
open Environ
open Evd
+open EConstr
open Glob_term
-open Evarutil
-open Misctypes
+open Ltac_pretype
+open Evardefine
+
+val interp_known_glob_level : ?loc:Loc.t -> Evd.evar_map ->
+ Misctypes.glob_level -> Univ.Level.t
(** An auxiliary function for searching for fixpoint guard indexes *)
val search_guard :
- Loc.t -> env -> int list list -> rec_declaration -> int array
+ ?loc:Loc.t -> env -> int list list -> Constr.rec_declaration -> int array
type typing_constraint = OfType of types | IsType | WithoutTypeConstraint
-type var_map = Pattern.constr_under_binders Id.Map.t
-type uconstr_var_map = Glob_term.closed_glob_constr Id.Map.t
-type unbound_ltac_var_map = Geninterp.Val.t Id.Map.t
-
-type ltac_var_map = {
- ltac_constrs : var_map;
- (** Ltac variables bound to constrs *)
- ltac_uconstrs : uconstr_var_map;
- (** Ltac variables bound to untyped constrs *)
- ltac_idents: Id.t Id.Map.t;
- (** Ltac variables bound to identifiers *)
- ltac_genargs : unbound_ltac_var_map;
- (** Ltac variables bound to other kinds of arguments *)
-}
-
-val empty_lvar : ltac_var_map
-
-type glob_constr_ltac_closure = ltac_var_map * glob_constr
-type pure_open_constr = evar_map * constr
-
-type inference_hook = env -> evar_map -> evar -> evar_map * constr
+type inference_hook = env -> evar_map -> Evar.t -> evar_map * constr
type inference_flags = {
use_typeclasses : bool;
@@ -57,9 +41,6 @@ type inference_flags = {
expand_evars : bool
}
-type 'a delayed_open =
- { delayed : 'r. Environ.env -> 'r Sigma.t -> ('a, 'r) Sigma.sigma }
-
val default_inference_flags : bool -> inference_flags
val no_classes_no_fail_inference_flags : inference_flags
@@ -68,9 +49,6 @@ val all_no_fail_flags : inference_flags
val all_and_fail_flags : inference_flags
-(** Allow references to syntactically nonexistent variables (i.e., if applied on an inductive) *)
-val allow_anonymous_refs : bool ref
-
(** Generic calls to the interpreter from glob_constr to open_constr;
by default, inference_flags tell to use type classes and
heuristics (but no external tactic solver hooks), as well as to
@@ -79,10 +57,12 @@ val allow_anonymous_refs : bool ref
evar_map is modified explicitly or by side-effect. *)
val understand_tcc : ?flags:inference_flags -> env -> evar_map ->
- ?expected_type:typing_constraint -> glob_constr -> open_constr
+ ?expected_type:typing_constraint -> glob_constr -> evar_map * constr
-val understand_tcc_evars : ?flags:inference_flags -> env -> evar_map ref ->
- ?expected_type:typing_constraint -> glob_constr -> constr
+(** As [understand_tcc] but also returns the type of the elaborated term.
+ The [expand_evars] flag is not applied to the type (only to the term). *)
+val understand_tcc_ty : ?flags:inference_flags -> env -> evar_map ->
+ ?expected_type:typing_constraint -> glob_constr -> evar_map * constr * types
(** More general entry point with evars from ltac *)
@@ -98,7 +78,7 @@ val understand_tcc_evars : ?flags:inference_flags -> env -> evar_map ref ->
val understand_ltac : inference_flags ->
env -> evar_map -> ltac_var_map ->
- typing_constraint -> glob_constr -> pure_open_constr
+ typing_constraint -> glob_constr -> evar_map * EConstr.t
(** Standard call to get a constr from a glob_constr, resolving
implicit arguments and coercions, and compiling pattern-matching;
@@ -106,24 +86,9 @@ val understand_ltac : inference_flags ->
heuristics (but no external tactic solver hook), as well as to
ensure that conversion problems are all solved and that no
unresolved evar remains, expanding evars. *)
-
val understand : ?flags:inference_flags -> ?expected_type:typing_constraint ->
env -> evar_map -> glob_constr -> constr Evd.in_evar_universe_context
-(** Idem but returns the judgment of the understood term *)
-
-val understand_judgment : env -> evar_map ->
- glob_constr -> unsafe_judgment Evd.in_evar_universe_context
-
-(** Idem but do not fail on unresolved evars (type cl*)
-val understand_judgment_tcc : env -> evar_map ref ->
- glob_constr -> unsafe_judgment
-
-val type_uconstr :
- ?flags:inference_flags ->
- ?expected_type:typing_constraint ->
- Geninterp.interp_sign -> Glob_term.closed_glob_constr -> constr delayed_open
-
(** Trying to solve remaining evars and remaining conversion problems
possibly using type classes, heuristics, external tactic solver
hook depending on given flags. *)
@@ -132,13 +97,13 @@ val type_uconstr :
[pending], however, it can contain more evars than the pending ones. *)
val solve_remaining_evars : inference_flags ->
- env -> (* initial map *) evar_map -> (* map to solve *) pending -> evar_map
+ env -> (* current map *) evar_map -> (* initial map *) evar_map -> evar_map
(** Checking evars and pending conversion problems are all solved,
reporting an appropriate error message *)
val check_evars_are_solved :
- env -> (* current map: *) evar_map -> (* map to check: *) pending -> unit
+ env -> (* current map: *) evar_map -> (* initial map: *) evar_map -> unit
(** [check_evars env initial_sigma extended_sigma c] fails if some
new unresolved evar remains in [c] *)
@@ -156,15 +121,12 @@ val pretype_type :
val ise_pretype_gen :
inference_flags -> env -> evar_map ->
- ltac_var_map -> typing_constraint -> glob_constr -> evar_map * constr
+ ltac_var_map -> typing_constraint -> glob_constr -> evar_map * constr * types
(**/**)
(** To embed constr in glob_constr *)
-val interp_sort : ?loc:Loc.t -> evar_map -> glob_sort -> evar_map * sorts
-val interp_elimination_sort : glob_sort -> sorts_family
-
-val genarg_interp_hook :
- (types -> env -> evar_map -> unbound_ltac_var_map ->
- Genarg.glob_generic_argument -> constr * evar_map) Hook.t
+val register_constr_interp0 :
+ ('r, 'g, 't) Genarg.genarg_type ->
+ (unbound_ltac_var_map -> env -> evar_map -> types -> 'g -> constr * evar_map) -> unit
diff --git a/pretyping/pretyping.mllib b/pretyping/pretyping.mllib
index c8b3307d..ae4ad0be 100644
--- a/pretyping/pretyping.mllib
+++ b/pretyping/pretyping.mllib
@@ -1,7 +1,10 @@
+Geninterp
+Ltac_pretype
Locusops
Pretype_errors
Reductionops
Inductiveops
+InferCumulativity
Vnorm
Arguments_renaming
Nativenorm
@@ -29,3 +32,4 @@ Indrec
Cases
Pretyping
Unification
+Univdecls
diff --git a/pretyping/program.ml b/pretyping/program.ml
index 62aedcfb..52d940d8 100644
--- a/pretyping/program.ml
+++ b/pretyping/program.ml
@@ -1,34 +1,20 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Pp
open CErrors
open Util
-open Names
-open Term
-let make_dir l = DirPath.make (List.rev_map Id.of_string l)
-
-let find_reference locstr dir s =
- let dp = make_dir dir in
- let sp = Libnames.make_path dp (Id.of_string s) in
- try Nametab.global_of_path sp
- with Not_found ->
- user_err_loc (Loc.ghost, "", str "Library " ++ Libnames.pr_dirpath dp ++
- str " has to be required first.")
-
-let coq_reference locstr dir s = find_reference locstr ("Coq"::dir) s
-let coq_constant locstr dir s = Universes.constr_of_global (coq_reference locstr dir s)
-
-let init_constant dir s () = coq_constant "Program" dir s
-let init_reference dir s () = coq_reference "Program" dir s
+let init_reference dir s () = Coqlib.coq_reference "Program" dir s
let papp evdref r args =
+ let open EConstr in
let gr = delayed_force r in
mkApp (Evarutil.e_new_global evdref gr, args)
@@ -54,20 +40,22 @@ let coq_eq_rect = init_reference ["Init"; "Logic"] "eq_rect"
let coq_JMeq_ind = init_reference ["Logic";"JMeq"] "JMeq"
let coq_JMeq_refl = init_reference ["Logic";"JMeq"] "JMeq_refl"
-let coq_not = init_constant ["Init";"Logic"] "not"
-let coq_and = init_constant ["Init";"Logic"] "and"
+let coq_not = init_reference ["Init";"Logic"] "not"
+let coq_and = init_reference ["Init";"Logic"] "and"
-let mk_coq_not x = mkApp (delayed_force coq_not, [| x |])
+let mk_coq_not sigma x =
+ let sigma, notc = Evarutil.new_global sigma (coq_not ()) in
+ sigma, EConstr.mkApp (notc, [| x |])
let unsafe_fold_right f = function
hd :: tl -> List.fold_right f tl hd
| [] -> invalid_arg "unsafe_fold_right"
-let mk_coq_and l =
- let and_typ = delayed_force coq_and in
- unsafe_fold_right
+let mk_coq_and sigma l =
+ let sigma, and_typ = Evarutil.new_global sigma (coq_and ()) in
+ sigma, unsafe_fold_right
(fun c conj ->
- mkApp (and_typ, [| c ; conj |]))
+ EConstr.mkApp (and_typ, [| c ; conj |]))
l
(* true = transparent by default, false = opaque if possible *)
@@ -85,8 +73,7 @@ open Goptions
let _ =
declare_bool_option
- { optsync = true;
- optdepr = false;
+ { optdepr = false;
optname = "preferred transparency of Program obligations";
optkey = ["Transparent";"Obligations"];
optread = get_proofs_transparency;
@@ -94,8 +81,7 @@ let _ =
let _ =
declare_bool_option
- { optsync = true;
- optdepr = false;
+ { optdepr = false;
optname = "program cases";
optkey = ["Program";"Cases"];
optread = (fun () -> !program_cases);
@@ -103,8 +89,7 @@ let _ =
let _ =
declare_bool_option
- { optsync = true;
- optdepr = false;
+ { optdepr = false;
optname = "program generalized coercion";
optkey = ["Program";"Generalized";"Coercion"];
optread = (fun () -> !program_generalized_coercion);
diff --git a/pretyping/program.mli b/pretyping/program.mli
index 023ff8ca..df0848ba 100644
--- a/pretyping/program.mli
+++ b/pretyping/program.mli
@@ -1,12 +1,14 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
+open EConstr
open Globnames
(** A bunch of Coq constants used by Progam *)
@@ -32,8 +34,8 @@ val coq_eq_rect : unit -> global_reference
val coq_JMeq_ind : unit -> global_reference
val coq_JMeq_refl : unit -> global_reference
-val mk_coq_and : constr list -> constr
-val mk_coq_not : constr -> constr
+val mk_coq_and : Evd.evar_map -> constr list -> Evd.evar_map * constr
+val mk_coq_not : Evd.evar_map -> constr -> Evd.evar_map * constr
(** Polymorphic application of delayed references *)
val papp : Evd.evar_map ref -> (unit -> global_reference) -> constr array -> constr
diff --git a/pretyping/recordops.ml b/pretyping/recordops.ml
index 284af0cb..d070edea 100644
--- a/pretyping/recordops.ml
+++ b/pretyping/recordops.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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) *)
(************************************************************************)
(* Created by Amokrane Saïbi, Dec 1998 *)
@@ -19,7 +21,7 @@ open Pp
open Names
open Globnames
open Nametab
-open Term
+open Constr
open Libobject
open Mod_subst
open Reductionops
@@ -37,7 +39,7 @@ type struc_typ = {
s_CONST : constructor;
s_EXPECTEDPARAM : int;
s_PROJKIND : (Name.t * bool) list;
- s_PROJ : constant option list }
+ s_PROJ : Constant.t option list }
let structure_table =
Summary.ref (Indmap.empty : struc_typ Indmap.t) ~name:"record-structs"
@@ -48,7 +50,7 @@ let projection_table =
is the inductive always (fst constructor) ? It seems so... *)
type struc_tuple =
- inductive * constructor * (Name.t * bool) list * constant option list
+ inductive * constructor * (Name.t * bool) list * Constant.t option list
let load_structure i (_,(ind,id,kl,projs)) =
let n = (fst (Global.lookup_inductive ind)).Declarations.mind_nparams in
@@ -134,7 +136,7 @@ let find_projection = function
type obj_typ = {
o_DEF : constr;
- o_CTX : Univ.ContextSet.t;
+ o_CTX : Univ.AUContext.t;
o_INJ : int option; (* position of trivial argument if any *)
o_TABS : constr list; (* ordered *)
o_TPARAMS : constr list; (* ordered *)
@@ -144,7 +146,7 @@ type obj_typ = {
type cs_pattern =
Const_cs of global_reference
| Prod_cs
- | Sort_cs of sorts_family
+ | Sort_cs of Sorts.family
| Default_cs
let eq_cs_pattern p1 p2 = match p1, p2 with
@@ -171,16 +173,20 @@ let keep_true_projections projs kinds =
let filter (p, (_, b)) = if b then Some p else None in
List.map_filter filter (List.combine projs kinds)
-let cs_pattern_of_constr t =
- match kind_of_term t with
+let cs_pattern_of_constr env t =
+ match kind t with
App (f,vargs) ->
begin
try Const_cs (global_of_constr f) , None, Array.to_list vargs
with e when CErrors.noncritical e -> raise Not_found
end
| Rel n -> Default_cs, Some n, []
- | Prod (_,a,b) when not (Termops.dependent (mkRel 1) b) -> Prod_cs, None, [a; Termops.pop b]
- | Sort s -> Sort_cs (family_of_sort s), None, []
+ | Prod (_,a,b) when Vars.noccurn 1 b -> Prod_cs, None, [a; Vars.lift (-1) b]
+ | Proj (p, c) ->
+ let { Environ.uj_type = ty } = Typeops.infer env c in
+ let _, params = Inductive.find_rectype env ty in
+ Const_cs (ConstRef (Projection.constant p)), None, params @ [c]
+ | Sort s -> Sort_cs (Sorts.family s), None, []
| _ ->
begin
try Const_cs (global_of_constr t) , None, []
@@ -189,27 +195,33 @@ let cs_pattern_of_constr t =
let warn_projection_no_head_constant =
CWarnings.create ~name:"projection-no-head-constant" ~category:"typechecker"
- (fun (t,con_pp,proji_sp_pp) ->
+ (fun (sign,env,t,con,proji_sp) ->
+ let env = Termops.push_rels_assum sign env in
+ let con_pp = Nametab.pr_global_env Id.Set.empty (ConstRef con) in
+ let proji_sp_pp = Nametab.pr_global_env Id.Set.empty (ConstRef proji_sp) in
+ let term_pp = Termops.print_constr_env env Evd.empty (EConstr.of_constr t) in
strbrk "Projection value has no head constant: "
- ++ Termops.print_constr t ++ strbrk " in canonical instance "
+ ++ term_pp ++ strbrk " in canonical instance "
++ con_pp ++ str " of " ++ proji_sp_pp ++ strbrk ", ignoring it.")
(* Intended to always succeed *)
let compute_canonical_projections warn (con,ind) =
let env = Global.env () in
- let ctx = Univ.instantiate_univ_context (Environ.constant_context env con) in
- let u = Univ.UContext.instance ctx in
+ let ctx = Environ.constant_context env con in
+ let u = Univ.make_abstract_instance ctx in
let v = (mkConstU (con,u)) in
- let ctx = Univ.ContextSet.of_context ctx in
let c = Environ.constant_value_in env (con,u) in
- let lt,t = Reductionops.splay_lam env Evd.empty c in
- let lt = List.rev_map snd lt in
+ let sign,t = Reductionops.splay_lam env Evd.empty (EConstr.of_constr c) in
+ let sign = List.map (on_snd EConstr.Unsafe.to_constr) sign in
+ let t = EConstr.Unsafe.to_constr t in
+ let lt = List.rev_map snd sign in
let args = snd (decompose_app t) in
let { s_EXPECTEDPARAM = p; s_PROJ = lpj; s_PROJKIND = kl } =
lookup_structure ind in
let params, projs = List.chop p args in
let lpj = keep_true_projections lpj kl in
let lps = List.combine lpj projs in
+ let nenv = Termops.push_rels_assum sign env in
let comp =
List.fold_left
(fun l (spopt,t) -> (* comp=components *)
@@ -217,12 +229,10 @@ let compute_canonical_projections warn (con,ind) =
| Some proji_sp ->
begin
try
- let patt, n , args = cs_pattern_of_constr t in
+ let patt, n , args = cs_pattern_of_constr nenv t in
((ConstRef proji_sp, patt, t, n, args) :: l)
with Not_found ->
- let con_pp = Nametab.pr_global_env Id.Set.empty (ConstRef con)
- and proji_sp_pp = Nametab.pr_global_env Id.Set.empty (ConstRef proji_sp) in
- if warn then warn_projection_no_head_constant (t,con_pp,proji_sp_pp);
+ if warn then warn_projection_no_head_constant (sign,env,t,con,proji_sp);
l
end
| _ -> l)
@@ -255,8 +265,8 @@ let add_canonical_structure warn o =
in match ocs with
| None -> object_table := Refmap.add proj ((pat,s)::l) !object_table;
| Some (c, cs) ->
- let old_can_s = (Termops.print_constr cs.o_DEF)
- and new_can_s = (Termops.print_constr s.o_DEF) in
+ let old_can_s = (Termops.print_constr (EConstr.of_constr cs.o_DEF))
+ and new_can_s = (Termops.print_constr (EConstr.of_constr s.o_DEF)) in
let prj = (Nametab.pr_global_env Id.Set.empty proj)
and hd_val = (pr_cs_pattern cs_pat) in
if warn then warn_redundant_canonical_projection (hd_val,prj,new_can_s,old_can_s))
@@ -278,7 +288,7 @@ let subst_canonical_structure (subst,(cst,ind as obj)) =
let discharge_canonical_structure (_,(cst,ind)) =
Some (Lib.discharge_con cst,Lib.discharge_inductive ind)
-let inCanonStruc : constant * inductive -> obj =
+let inCanonStruc : Constant.t * inductive -> obj =
declare_object {(default_object "CANONICAL-STRUCTURE") with
open_function = open_canonical_structure;
cache_function = cache_canonical_structure;
@@ -290,29 +300,40 @@ let add_canonical_structure x = Lib.add_anonymous_leaf (inCanonStruc x)
(*s High-level declaration of a canonical structure *)
-let error_not_structure ref =
- errorlabstrm "object_declare"
- (Nameops.pr_id (basename_of_global ref) ++ str" is not a structure object.")
+let error_not_structure ref description =
+ user_err ~hdr:"object_declare"
+ (str"Could not declare a canonical structure " ++
+ (Id.print (basename_of_global ref) ++ str"." ++ spc() ++
+ str(description)))
let check_and_decompose_canonical_structure ref =
- let sp = match ref with ConstRef sp -> sp | _ -> error_not_structure ref in
+ let sp =
+ match ref with
+ ConstRef sp -> sp
+ | _ -> error_not_structure ref "Expected an instance of a record or structure."
+ in
let env = Global.env () in
- let ctx = Environ.constant_context env sp in
- let u = Univ.UContext.instance ctx in
+ let u = Univ.make_abstract_instance (Environ.constant_context env sp) in
let vc = match Environ.constant_opt_value_in env (sp, u) with
| Some vc -> vc
- | None -> error_not_structure ref in
- let body = snd (splay_lam (Global.env()) Evd.empty vc) in
- let f,args = match kind_of_term body with
+ | None -> error_not_structure ref "Could not find its value in the global environment." in
+ let body = snd (splay_lam (Global.env()) Evd.empty (EConstr.of_constr vc)) (** FIXME *) in
+ let body = EConstr.Unsafe.to_constr body in
+ let f,args = match kind body with
| App (f,args) -> f,args
- | _ -> error_not_structure ref in
- let indsp = match kind_of_term f with
+ | _ ->
+ error_not_structure ref "Expected a record or structure constructor applied to arguments." in
+ let indsp = match kind f with
| Construct ((indsp,1),u) -> indsp
- | _ -> error_not_structure ref in
- let s = try lookup_structure indsp with Not_found -> error_not_structure ref in
+ | _ -> error_not_structure ref "Expected an instance of a record or structure." in
+ let s =
+ try lookup_structure indsp
+ with Not_found ->
+ error_not_structure ref
+ ("Could not find the record or structure " ^ (MutInd.to_string (fst indsp))) in
let ntrue_projs = List.count snd s.s_PROJKIND in
if s.s_EXPECTEDPARAM + ntrue_projs > Array.length args then
- error_not_structure ref;
+ error_not_structure ref "Got too few arguments to the record or structure constructor.";
(sp,indsp)
let declare_canonical_structure ref =
@@ -321,15 +342,26 @@ let declare_canonical_structure ref =
let lookup_canonical_conversion (proj,pat) =
assoc_pat pat (Refmap.find proj !object_table)
+let decompose_projection sigma c args =
+ match EConstr.kind sigma c with
+ | Const (c, u) ->
+ let n = find_projection_nparams (ConstRef c) in
+ (** Check if there is some canonical projection attached to this structure *)
+ let _ = Refmap.find (ConstRef c) !object_table in
+ let arg = Stack.nth args n in
+ arg
+ | Proj (p, c) ->
+ let _ = Refmap.find (ConstRef (Projection.constant p)) !object_table in
+ c
+ | _ -> raise Not_found
+
let is_open_canonical_projection env sigma (c,args) =
+ let open EConstr in
try
- let ref = global_of_constr c in
- let n = find_projection_nparams ref in
- (** Check if there is some canonical projection attached to this structure *)
- let _ = Refmap.find ref !object_table in
+ let arg = decompose_projection sigma c args in
try
- let arg = whd_all env sigma (Stack.nth args n) in
- let hd = match kind_of_term arg with App (hd, _) -> hd | _ -> arg in
- not (isConstruct hd)
+ let arg = whd_all env sigma arg in
+ let hd = match EConstr.kind sigma arg with App (hd, _) -> hd | _ -> arg in
+ not (isConstruct sigma hd)
with Failure _ -> false
with Not_found -> false
diff --git a/pretyping/recordops.mli b/pretyping/recordops.mli
index a6a90c75..1f7b23c0 100644
--- a/pretyping/recordops.mli
+++ b/pretyping/recordops.mli
@@ -1,13 +1,15 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
+open Constr
open Globnames
(** Operations concerning records and canonical structures *)
@@ -20,10 +22,10 @@ type struc_typ = {
s_CONST : constructor;
s_EXPECTEDPARAM : int;
s_PROJKIND : (Name.t * bool) list;
- s_PROJ : constant option list }
+ s_PROJ : Constant.t option list }
type struc_tuple =
- inductive * constructor * (Name.t * bool) list * constant option list
+ inductive * constructor * (Name.t * bool) list * Constant.t option list
val declare_structure : struc_tuple -> unit
@@ -35,7 +37,7 @@ val lookup_structure : inductive -> struc_typ
(** [lookup_projections isp] returns the projections associated to the
inductive path [isp] if it corresponds to a structure, otherwise
it fails with [Not_found] *)
-val lookup_projections : inductive -> constant option list
+val lookup_projections : inductive -> Constant.t option list
(** raise [Not_found] if not a projection *)
val find_projection_nparams : global_reference -> int
@@ -52,12 +54,12 @@ val find_projection : global_reference -> struc_typ
type cs_pattern =
Const_cs of global_reference
| Prod_cs
- | Sort_cs of sorts_family
+ | Sort_cs of Sorts.family
| Default_cs
type obj_typ = {
o_DEF : constr;
- o_CTX : Univ.ContextSet.t;
+ o_CTX : Univ.AUContext.t;
o_INJ : int option; (** position of trivial argument *)
o_TABS : constr list; (** ordered *)
o_TPARAMS : constr list; (** ordered *)
@@ -65,13 +67,13 @@ type obj_typ = {
o_TCOMPS : constr list } (** ordered *)
(** Return the form of the component of a canonical structure *)
-val cs_pattern_of_constr : constr -> cs_pattern * int option * constr list
+val cs_pattern_of_constr : Environ.env -> constr -> cs_pattern * int option * constr list
-val pr_cs_pattern : cs_pattern -> Pp.std_ppcmds
+val pr_cs_pattern : cs_pattern -> Pp.t
val lookup_canonical_conversion : (global_reference * cs_pattern) -> constr * obj_typ
val declare_canonical_structure : global_reference -> unit
val is_open_canonical_projection :
- Environ.env -> Evd.evar_map -> (constr * constr Reductionops.Stack.t) -> bool
+ Environ.env -> Evd.evar_map -> Reductionops.state -> bool
val canonical_projections : unit ->
((global_reference * cs_pattern) * obj_typ) list
diff --git a/pretyping/redops.ml b/pretyping/redops.ml
index 7d65925e..90c3bdfa 100644
--- a/pretyping/redops.ml
+++ b/pretyping/redops.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Genredexpr
@@ -20,13 +22,13 @@ let make_red_flag l =
| FZeta :: lf -> add_flag { red with rZeta = true } lf
| FConst l :: lf ->
if red.rDelta then
- CErrors.error
- "Cannot set both constants to unfold and constants not to unfold";
+ CErrors.user_err Pp.(str
+ "Cannot set both constants to unfold and constants not to unfold");
add_flag { red with rConst = union_consts red.rConst l } lf
| FDeltaBut l :: lf ->
if red.rConst <> [] && not red.rDelta then
- CErrors.error
- "Cannot set both constants to unfold and constants not to unfold";
+ CErrors.user_err Pp.(str
+ "Cannot set both constants to unfold and constants not to unfold");
add_flag
{ red with rConst = union_consts red.rConst l; rDelta = true }
lf
diff --git a/pretyping/redops.mli b/pretyping/redops.mli
index f6d4d808..285931ec 100644
--- a/pretyping/redops.mli
+++ b/pretyping/redops.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Genredexpr
diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml
index 297f0a1a..360c6e86 100644
--- a/pretyping/reductionops.ml
+++ b/pretyping/reductionops.ml
@@ -1,20 +1,23 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 CErrors
open Util
open Names
-open Term
-open Vars
+open Constr
open Termops
open Univ
open Evd
open Environ
+open EConstr
+open Vars
open Context.Rel.Declaration
exception Elimconst
@@ -26,18 +29,63 @@ exception Elimconst
their parameters in its stack.
*)
-let refolding_in_reduction = ref false
let _ = Goptions.declare_bool_option {
- Goptions.optsync = true; Goptions.optdepr = false;
+ Goptions.optdepr = false;
Goptions.optname =
- "Perform refolding of fixpoints/constants like cbn during reductions";
- Goptions.optkey = ["Refolding";"Reduction"];
- Goptions.optread = (fun () -> !refolding_in_reduction);
- Goptions.optwrite = (fun a -> refolding_in_reduction:=a);
+ "Generate weak constraints between Irrelevant universes";
+ Goptions.optkey = ["Cumulativity";"Weak";"Constraints"];
+ Goptions.optread = (fun () -> not !UState.drop_weak_constraints);
+ Goptions.optwrite = (fun a -> UState.drop_weak_constraints:=not a);
}
-let get_refolding_in_reduction () = !refolding_in_reduction
-let set_refolding_in_reduction = (:=) refolding_in_reduction
+
+(** Support for reduction effects *)
+
+open Mod_subst
+open Libobject
+
+type effect_name = string
+
+(** create a persistent set to store effect functions *)
+module ConstrMap = Map.Make (Constr)
+
+(* Table bindings a constant to an effect *)
+let constant_effect_table = Summary.ref ~name:"reduction-side-effect" ConstrMap.empty
+
+(* Table bindings function key to effective functions *)
+let effect_table = Summary.ref ~name:"reduction-function-effect" String.Map.empty
+
+(** a test to know whether a constant is actually the effect function *)
+let reduction_effect_hook env sigma termkey c =
+ try
+ let funkey = ConstrMap.find termkey !constant_effect_table in
+ let effect = String.Map.find funkey !effect_table in
+ effect env sigma (Lazy.force c)
+ with Not_found -> ()
+
+let cache_reduction_effect (_,(termkey,funkey)) =
+ constant_effect_table := ConstrMap.add termkey funkey !constant_effect_table
+
+let subst_reduction_effect (subst,(termkey,funkey)) =
+ (subst_mps subst termkey,funkey)
+
+let inReductionEffect : Constr.constr * string -> obj =
+ declare_object {(default_object "REDUCTION-EFFECT") with
+ cache_function = cache_reduction_effect;
+ open_function = (fun i o -> if Int.equal i 1 then cache_reduction_effect o);
+ subst_function = subst_reduction_effect;
+ classify_function = (fun o -> Substitute o) }
+
+let declare_reduction_effect funkey f =
+ if String.Map.mem funkey !effect_table then
+ CErrors.anomaly Pp.(str "Cannot redeclare effect function " ++ qstring funkey ++ str ".");
+ effect_table := String.Map.add funkey f !effect_table
+
+(** A function to set the value of the print function *)
+let set_reduction_effect x funkey =
+ let termkey = Universes.constr_of_global x in
+ Lib.add_anonymous_leaf (inReductionEffect (termkey,funkey))
+
(** Machinery to custom the behavior of the reduction *)
module ReductionBehaviour = struct
@@ -72,10 +120,10 @@ module ReductionBehaviour = struct
let r' = fst (subst_global subst r) in if r==r' then orig else (r',o)
let discharge = function
- | _,(ReqGlobal (ConstRef c, req), (_, b)) ->
+ | _,(ReqGlobal (ConstRef c as gr, req), (_, b)) ->
let b =
- if Lib.is_in_section (ConstRef c) then
- let vars, _, _ = Lib.section_segment_of_constant c in
+ if Lib.is_in_section gr then
+ let vars = Lib.variable_section_segment_of_reference gr in
let extra = List.length vars in
let nargs' =
if b.b_nargs = max_int then max_int
@@ -156,6 +204,7 @@ end
(** Machinery about stack of unfolded constants *)
module Cst_stack = struct
+ open EConstr
(** constant * params * args
- constant applied to params = term in head applied to args
@@ -191,24 +240,24 @@ module Cst_stack = struct
| (cst,params,[])::_ -> Some(cst,params)
| _ -> None
- let reference t = match best_cst t with
- | Some (c, _) when Term.isConst c -> Some (fst (Term.destConst c))
+ let reference sigma t = match best_cst t with
+ | Some (c, _) when isConst sigma c -> Some (fst (destConst sigma c))
| _ -> None
(** [best_replace d cst_l c] makes the best replacement for [d]
by [cst_l] in [c] *)
- let best_replace d cst_l c =
+ let best_replace sigma d cst_l c =
let reconstruct_head = List.fold_left
(fun t (i,args) -> mkApp (t,Array.sub args i (Array.length args - i))) in
List.fold_right
- (fun (cst,params,args) t -> Termops.replace_term
+ (fun (cst,params,args) t -> Termops.replace_term sigma
(reconstruct_head d args)
(applist (cst, List.rev params))
t) cst_l c
let pr l =
let open Pp in
- let p_c = Termops.print_constr in
+ let p_c c = Termops.print_constr c in
prlist_with_sep pr_semicolon
(fun (c,params,args) ->
hov 1 (str"(" ++ p_c c ++ str ")" ++ spc () ++ pr_sequence p_c params ++ spc () ++ str "(args:" ++
@@ -220,34 +269,36 @@ end
(** The type of (machine) stacks (= lambda-bar-calculus' contexts) *)
module Stack :
sig
+ open EConstr
type 'a app_node
- val pr_app_node : ('a -> Pp.std_ppcmds) -> 'a app_node -> Pp.std_ppcmds
+ val pr_app_node : ('a -> Pp.t) -> 'a app_node -> Pp.t
type cst_member =
| Cst_const of pconstant
- | Cst_proj of projection
+ | Cst_proj of Projection.t
type 'a member =
| App of 'a app_node
| Case of case_info * 'a * 'a array * Cst_stack.t
- | Proj of int * int * projection * Cst_stack.t
- | Fix of fixpoint * 'a t * Cst_stack.t
+ | Proj of int * int * Projection.t * Cst_stack.t
+ | Fix of ('a, 'a) pfixpoint * 'a t * Cst_stack.t
| Cst of cst_member * int * int list * 'a t * Cst_stack.t
- | Shift of int
- | Update of 'a
and 'a t = 'a member list
- val pr : ('a -> Pp.std_ppcmds) -> 'a t -> Pp.std_ppcmds
+
+ exception IncompatibleFold2
+
+ val pr : ('a -> Pp.t) -> 'a t -> Pp.t
val empty : 'a t
val is_empty : 'a t -> bool
val append_app : 'a array -> 'a t -> 'a t
val decomp : 'a t -> ('a * 'a t) option
val decomp_node_last : 'a app_node -> 'a t -> ('a * 'a t)
- val equal : ('a * int -> 'a * int -> bool) -> (fixpoint * int -> fixpoint * int -> bool)
- -> 'a t -> 'a t -> (int * int) option
+ val equal : ('a -> 'a -> bool) -> (('a, 'a) pfixpoint -> ('a, 'a) pfixpoint -> bool)
+ -> 'a t -> 'a t -> bool
val compare_shape : 'a t -> 'a t -> bool
- val map : (constr -> constr) -> constr t -> constr t
+ val map : ('a -> 'a) -> 'a t -> 'a t
val fold2 : ('a -> constr -> constr -> 'a) -> 'a ->
- constr t -> constr t -> 'a * int * int
+ constr t -> constr t -> 'a
val append_app_list : 'a list -> 'a t -> 'a t
val strip_app : 'a t -> 'a t * 'a t
val strip_n_app : int -> 'a t -> ('a t * 'a * 'a t) option
@@ -258,10 +309,11 @@ sig
val args_size : 'a t -> int
val tail : int -> 'a t -> 'a t
val nth : 'a t -> int -> 'a
- val best_state : constr * constr t -> Cst_stack.t -> constr * constr t
- val zip : ?refold:bool -> constr * constr t -> constr
+ val best_state : evar_map -> constr * constr t -> Cst_stack.t -> constr * constr t
+ val zip : ?refold:bool -> evar_map -> constr * constr t -> constr
end =
struct
+ open EConstr
type 'a app_node = int * 'a array * int
(* first releavnt position, arguments, last relevant position *)
@@ -280,16 +332,14 @@ struct
type cst_member =
| Cst_const of pconstant
- | Cst_proj of projection
+ | Cst_proj of Projection.t
type 'a member =
| App of 'a app_node
- | Case of Term.case_info * 'a * 'a array * Cst_stack.t
- | Proj of int * int * projection * Cst_stack.t
- | Fix of fixpoint * 'a t * Cst_stack.t
+ | Case of case_info * 'a * 'a array * Cst_stack.t
+ | Proj of int * int * Projection.t * Cst_stack.t
+ | Fix of ('a, 'a) pfixpoint * 'a t * Cst_stack.t
| Cst of cst_member * int * int list * 'a t * Cst_stack.t
- | Shift of int
- | Update of 'a
and 'a t = 'a member list
let rec pr_member pr_c member =
@@ -303,17 +353,15 @@ struct
++ str ")"
| Proj (n,m,p,cst) ->
str "ZProj(" ++ int n ++ pr_comma () ++ int m ++
- pr_comma () ++ pr_con (Projection.constant p) ++ str ")"
+ pr_comma () ++ Constant.print (Projection.constant p) ++ str ")"
| Fix (f,args,cst) ->
- str "ZFix(" ++ Termops.pr_fix Termops.print_constr f
+ str "ZFix(" ++ Termops.pr_fix pr_c f
++ pr_comma () ++ pr pr_c args ++ str ")"
| Cst (mem,curr,remains,params,cst_l) ->
str "ZCst(" ++ pr_cst_member pr_c mem ++ pr_comma () ++ int curr
++ pr_comma () ++
prlist_with_sep pr_semicolon int remains ++
pr_comma () ++ pr pr_c params ++ str ")"
- | Shift i -> str "ZShift(" ++ int i ++ str ")"
- | Update t -> str "ZUpdate(" ++ pr_c t ++ str ")"
and pr pr_c l =
let open Pp in
prlist_with_sep pr_semicolon (fun x -> hov 1 (pr_member pr_c x)) l
@@ -348,54 +396,42 @@ struct
else (l.(j), sk)
let equal f f_fix sk1 sk2 =
- let equal_cst_member x lft1 y lft2 =
+ let equal_cst_member x y =
match x, y with
| Cst_const (c1,u1), Cst_const (c2, u2) ->
- Constant.equal c1 c2 && Univ.Instance.equal u1 u2
+ Constant.equal c1 c2 && Univ.Instance.equal u1 u2
| Cst_proj p1, Cst_proj p2 ->
- Constant.equal (Projection.constant p1) (Projection.constant p2)
+ Constant.equal (Projection.constant p1) (Projection.constant p2)
| _, _ -> false
in
- let rec equal_rec sk1 lft1 sk2 lft2 =
+ let rec equal_rec sk1 sk2 =
match sk1,sk2 with
- | [],[] -> Some (lft1,lft2)
- | (Update _ :: _, _ | _, Update _ :: _) -> assert false
- | Shift k :: s1, _ -> equal_rec s1 (lft1+k) sk2 lft2
- | _, Shift k :: s2 -> equal_rec sk1 lft1 s2 (lft2+k)
+ | [],[] -> true
| App a1 :: s1, App a2 :: s2 ->
- let t1,s1' = decomp_node_last a1 s1 in
- let t2,s2' = decomp_node_last a2 s2 in
- if f (t1,lft1) (t2,lft2) then equal_rec s1' lft1 s2' lft2 else None
+ let t1,s1' = decomp_node_last a1 s1 in
+ let t2,s2' = decomp_node_last a2 s2 in
+ (f t1 t2) && (equal_rec s1' s2')
| Case (_,t1,a1,_) :: s1, Case (_,t2,a2,_) :: s2 ->
- if f (t1,lft1) (t2,lft2) && CArray.equal (fun x y -> f (x,lft1) (y,lft2)) a1 a2
- then equal_rec s1 lft1 s2 lft2
- else None
+ f t1 t2 && CArray.equal (fun x y -> f x y) a1 a2 && equal_rec s1 s2
| (Proj (n1,m1,p,_)::s1, Proj(n2,m2,p2,_)::s2) ->
- if Int.equal n1 n2 && Int.equal m1 m2
- && Constant.equal (Projection.constant p) (Projection.constant p2)
- then equal_rec s1 lft1 s2 lft2
- else None
+ Int.equal n1 n2 && Int.equal m1 m2
+ && Constant.equal (Projection.constant p) (Projection.constant p2)
+ && equal_rec s1 s2
| Fix (f1,s1,_) :: s1', Fix (f2,s2,_) :: s2' ->
- if f_fix (f1,lft1) (f2,lft2) then
- match equal_rec (List.rev s1) lft1 (List.rev s2) lft2 with
- | None -> None
- | Some (lft1',lft2') -> equal_rec s1' lft1' s2' lft2'
- else None
+ f_fix f1 f2
+ && equal_rec (List.rev s1) (List.rev s2)
+ && equal_rec s1' s2'
| Cst (c1,curr1,remains1,params1,_)::s1', Cst (c2,curr2,remains2,params2,_)::s2' ->
- if equal_cst_member c1 lft1 c2 lft2 then
- match equal_rec (List.rev params1) lft1 (List.rev params2) lft2 with
- | Some (lft1',lft2') -> equal_rec s1' lft1' s2' lft2'
- | None -> None
- else None
- | ((App _|Case _|Proj _|Fix _|Cst _)::_|[]), _ -> None
- in equal_rec (List.rev sk1) 0 (List.rev sk2) 0
+ equal_cst_member c1 c2
+ && equal_rec (List.rev params1) (List.rev params2)
+ && equal_rec s1' s2'
+ | ((App _|Case _|Proj _|Fix _|Cst _)::_|[]), _ -> false
+ in equal_rec (List.rev sk1) (List.rev sk2)
let compare_shape stk1 stk2 =
let rec compare_rec bal stk1 stk2 =
match (stk1,stk2) with
([],[]) -> Int.equal bal 0
- | ((Update _|Shift _)::s1, _) -> compare_rec bal s1 stk2
- | (_, (Update _|Shift _)::s2) -> compare_rec bal stk1 s2
| (App (i,_,j)::s1, _) -> compare_rec (bal + j + 1 - i) s1 stk2
| (_, App (i,_,j)::s2) -> compare_rec (bal - j - 1 + i) stk1 s2
| (Case(c1,_,_,_)::s1, Case(c2,_,_,_)::s2) ->
@@ -409,41 +445,31 @@ struct
| (_,_) -> false in
compare_rec 0 stk1 stk2
+ exception IncompatibleFold2
let fold2 f o sk1 sk2 =
- let rec aux o lft1 sk1 lft2 sk2 =
- let fold_array =
- Array.fold_left2 (fun a x y -> f a (Vars.lift lft1 x) (Vars.lift lft2 y))
- in
+ let rec aux o sk1 sk2 =
match sk1,sk2 with
- | [], [] -> o,lft1,lft2
- | Shift n :: q1, _ -> aux o (lft1+n) q1 lft2 sk2
- | _, Shift n :: q2 -> aux o lft1 sk1 (lft2+n) q2
+ | [], [] -> o
| App n1 :: q1, App n2 :: q2 ->
- let t1,l1 = decomp_node_last n1 q1 in
- let t2,l2 = decomp_node_last n2 q2 in
- aux (f o (Vars.lift lft1 t1) (Vars.lift lft2 t2))
- lft1 l1 lft2 l2
+ let t1,l1 = decomp_node_last n1 q1 in
+ let t2,l2 = decomp_node_last n2 q2 in
+ aux (f o t1 t2) l1 l2
| Case (_,t1,a1,_) :: q1, Case (_,t2,a2,_) :: q2 ->
- aux (fold_array
- (f o (Vars.lift lft1 t1) (Vars.lift lft2 t2))
- a1 a2) lft1 q1 lft2 q2
+ aux (Array.fold_left2 f (f o t1 t2) a1 a2) q1 q2
| Proj (n1,m1,p1,_) :: q1, Proj (n2,m2,p2,_) :: q2 ->
- aux o lft1 q1 lft2 q2
+ aux o q1 q2
| Fix ((_,(_,a1,b1)),s1,_) :: q1, Fix ((_,(_,a2,b2)),s2,_) :: q2 ->
- let (o',lft1',lft2') = aux (fold_array (fold_array o b1 b2) a1 a2)
- lft1 (List.rev s1) lft2 (List.rev s2) in
- aux o' lft1' q1 lft2' q2
+ let o' = aux (Array.fold_left2 f (Array.fold_left2 f o b1 b2) a1 a2) (List.rev s1) (List.rev s2) in
+ aux o' q1 q2
| Cst (cst1,_,_,params1,_) :: q1, Cst (cst2,_,_,params2,_) :: q2 ->
- let (o',lft1',lft2') =
- aux o lft1 (List.rev params1) lft2 (List.rev params2)
- in aux o' lft1' q1 lft2' q2
- | (((Update _|App _|Case _|Proj _|Fix _| Cst _) :: _|[]), _) ->
- raise (Invalid_argument "Reductionops.Stack.fold2")
- in aux o 0 (List.rev sk1) 0 (List.rev sk2)
+ let o' = aux o (List.rev params1) (List.rev params2) in
+ aux o' q1 q2
+ | (((App _|Case _|Proj _|Fix _| Cst _) :: _|[]), _) ->
+ raise IncompatibleFold2
+ in aux o (List.rev sk1) (List.rev sk2)
let rec map f x = List.map (function
- | Update _ -> assert false
- | (Proj (_,_,_,_) | Shift _) as e -> e
+ | (Proj (_,_,_,_)) as e -> e
| App (i,a,j) ->
let le = j - i + 1 in
App (0,Array.map f (Array.sub a i le), le-1)
@@ -460,18 +486,15 @@ struct
let rec args_size = function
| App (i,_,j)::s -> j + 1 - i + args_size s
- | Shift(_)::s -> args_size s
- | Update(_)::s -> args_size s
| (Case _|Fix _|Proj _|Cst _)::_ | [] -> 0
let strip_app s =
let rec aux out = function
- | ( App _ | Shift _ as e) :: s -> aux (e :: out) s
+ | ( App _ as e) :: s -> aux (e :: out) s
| s -> List.rev out,s
in aux [] s
let strip_n_app n s =
let rec aux n out = function
- | Shift k as e :: s -> aux n (e :: out) s
| App (i,a,j) as e :: s ->
let nb = j - i + 1 in
if n >= nb then
@@ -496,14 +519,12 @@ struct
let list_of_app_stack s =
let rec aux = function
| App (i,a,j) :: s ->
- let (k,(args',s')) = aux s in
- let a' = Array.map (Vars.lift k) (Array.sub a i (j - i + 1)) in
- k,(Array.fold_right (fun x y -> x::y) a' args', s')
- | Shift n :: s ->
- let (k,(args',s')) = aux s in (k+n,(args', s'))
- | s -> (0,([],s)) in
- let (lft,(out,s')) = aux s in
- let init = match s' with [] when Int.equal lft 0 -> true | _ -> false in
+ let (args',s') = aux s in
+ let a' = Array.sub a i (j - i + 1) in
+ (Array.fold_right (fun x y -> x::y) a' args', s')
+ | s -> ([],s) in
+ let (out,s') = aux s in
+ let init = match s' with [] -> true | _ -> false in
Option.init init out
let assign s p c =
@@ -512,20 +533,18 @@ struct
| None -> assert false
let tail n0 s0 =
- let rec aux lft n s =
- let out s = if Int.equal lft 0 then s else Shift lft :: s in
- if Int.equal n 0 then out s else
+ let rec aux n s =
+ if Int.equal n 0 then s else
match s with
| App (i,a,j) :: s ->
let nb = j - i + 1 in
if n >= nb then
- aux lft (n - nb) s
+ aux (n - nb) s
else
let p = i+n in
if j >= p then App(p,a,j)::s else s
- | Shift k :: s' -> aux (lft+k) n s'
| _ -> raise (Invalid_argument "Reductionops.Stack.tail")
- in aux 0 n0 s0
+ in aux n0 s0
let nth s p =
match strip_n_app p s with
@@ -533,11 +552,11 @@ struct
| None -> raise Not_found
(** This function breaks the abstraction of Cst_stack ! *)
- let best_state (_,sk as s) l =
+ let best_state sigma (_,sk as s) l =
let rec aux sk def = function
|(cst, params, []) -> (cst, append_app_list (List.rev params) sk)
|(cst, params, (i,t)::q) -> match decomp sk with
- | Some (el,sk') when Constr.equal el t.(i) ->
+ | Some (el,sk') when EConstr.eq_constr sigma el t.(i) ->
if i = pred (Array.length t)
then aux sk' def (cst, params, q)
else aux sk' def (cst, params, (succ i,t)::q)
@@ -546,59 +565,62 @@ struct
let constr_of_cst_member f sk =
match f with
- | Cst_const (c, u) -> mkConstU (c,u), sk
+ | Cst_const (c, u) -> mkConstU (c, EInstance.make u), sk
| Cst_proj p ->
match decomp sk with
| Some (hd, sk) -> mkProj (p, hd), sk
| None -> assert false
- let rec zip ?(refold=false) = function
+ let zip ?(refold=false) sigma s =
+ let rec zip = function
| f, [] -> f
| f, (App (i,a,j) :: s) ->
let a' = if Int.equal i 0 && Int.equal j (Array.length a - 1)
then a
else Array.sub a i (j - i + 1) in
- zip ~refold (mkApp (f, a'), s)
+ zip (mkApp (f, a'), s)
| f, (Case (ci,rt,br,cst_l)::s) when refold ->
- zip ~refold (best_state (mkCase (ci,rt,f,br), s) cst_l)
- | f, (Case (ci,rt,br,_)::s) -> zip ~refold (mkCase (ci,rt,f,br), s)
+ zip (best_state sigma (mkCase (ci,rt,f,br), s) cst_l)
+ | f, (Case (ci,rt,br,_)::s) -> zip (mkCase (ci,rt,f,br), s)
| f, (Fix (fix,st,cst_l)::s) when refold ->
- zip ~refold (best_state (mkFix fix, st @ (append_app [|f|] s)) cst_l)
- | f, (Fix (fix,st,_)::s) -> zip ~refold
+ zip (best_state sigma (mkFix fix, st @ (append_app [|f|] s)) cst_l)
+ | f, (Fix (fix,st,_)::s) -> zip
(mkFix fix, st @ (append_app [|f|] s))
| f, (Cst (cst,_,_,params,cst_l)::s) when refold ->
- zip ~refold (best_state (constr_of_cst_member cst (params @ (append_app [|f|] s))) cst_l)
+ zip (best_state sigma (constr_of_cst_member cst (params @ (append_app [|f|] s))) cst_l)
| f, (Cst (cst,_,_,params,_)::s) ->
- zip ~refold (constr_of_cst_member cst (params @ (append_app [|f|] s)))
- | f, (Shift n::s) -> zip ~refold (lift n f, s)
+ zip (constr_of_cst_member cst (params @ (append_app [|f|] s)))
| f, (Proj (n,m,p,cst_l)::s) when refold ->
- zip ~refold (best_state (mkProj (p,f),s) cst_l)
- | f, (Proj (n,m,p,_)::s) -> zip ~refold (mkProj (p,f),s)
- | _, (Update _::_) -> assert false
+ zip (best_state sigma (mkProj (p,f),s) cst_l)
+ | f, (Proj (n,m,p,_)::s) -> zip (mkProj (p,f),s)
+ in
+ zip s
+
end
(** The type of (machine) states (= lambda-bar-calculus' cuts) *)
type state = constr * constr Stack.t
-type contextual_reduction_function = env -> evar_map -> constr -> constr
-type reduction_function = contextual_reduction_function
+type contextual_reduction_function = env -> evar_map -> constr -> constr
+type reduction_function = contextual_reduction_function
type local_reduction_function = evar_map -> constr -> constr
-type e_reduction_function = { e_redfun : 'r. env -> 'r Sigma.t -> constr -> (constr, 'r) Sigma.sigma }
+type e_reduction_function = env -> evar_map -> constr -> evar_map * constr
-type contextual_stack_reduction_function =
- env -> evar_map -> constr -> constr * constr list
-type stack_reduction_function = contextual_stack_reduction_function
+type contextual_stack_reduction_function =
+ env -> evar_map -> constr -> constr * constr list
+type stack_reduction_function = contextual_stack_reduction_function
type local_stack_reduction_function =
evar_map -> constr -> constr * constr list
-type contextual_state_reduction_function =
- env -> evar_map -> state -> state
-type state_reduction_function = contextual_state_reduction_function
+type contextual_state_reduction_function =
+ env -> evar_map -> state -> state
+type state_reduction_function = contextual_state_reduction_function
type local_state_reduction_function = evar_map -> state -> state
let pr_state (tm,sk) =
let open Pp in
- h 0 (Termops.print_constr tm ++ str "|" ++ cut () ++ Stack.pr Termops.print_constr sk)
+ let pr c = Termops.print_constr c in
+ h 0 (pr tm ++ str "|" ++ cut () ++ Stack.pr pr sk)
(*************************************)
(*** Reduction Functions Operators ***)
@@ -612,16 +634,16 @@ let safe_meta_value sigma ev =
let strong whdfun env sigma t =
let rec strongrec env t =
- map_constr_with_full_binders push_rel strongrec env (whdfun env sigma t) in
+ map_constr_with_full_binders sigma push_rel strongrec env (whdfun env sigma t) in
strongrec env t
let local_strong whdfun sigma =
- let rec strongrec t = Constr.map strongrec (whdfun sigma t) in
+ let rec strongrec t = EConstr.map sigma strongrec (whdfun sigma t) in
strongrec
let rec strong_prodspine redfun sigma c =
let x = redfun sigma c in
- match kind_of_term x with
+ match EConstr.kind sigma x with
| Prod (na,a,b) -> mkProd (na,a,strong_prodspine redfun sigma b)
| _ -> x
@@ -633,20 +655,25 @@ let eta = CClosure.RedFlags.mkflags [CClosure.RedFlags.fETA]
(* Beta Reduction tools *)
-let apply_subst recfun env refold cst_l t stack =
+let apply_subst recfun env sigma refold cst_l t stack =
let rec aux env cst_l t stack =
- match (Stack.decomp stack,kind_of_term t) with
+ match (Stack.decomp stack, EConstr.kind sigma t) with
| Some (h,stacktl), Lambda (_,_,c) ->
let cst_l' = if refold then Cst_stack.add_param h cst_l else cst_l in
aux (h::env) cst_l' c stacktl
- | _ -> recfun cst_l (substl env t, stack)
+ | _ -> recfun sigma cst_l (substl env t, stack)
in aux env cst_l t stack
-let stacklam recfun env t stack =
- apply_subst (fun _ -> recfun) env false Cst_stack.empty t stack
+let stacklam recfun env sigma t stack =
+ apply_subst (fun _ _ s -> recfun s) env sigma false Cst_stack.empty t stack
+
+let beta_app sigma (c,l) =
+ let zip s = Stack.zip sigma s in
+ stacklam zip [] sigma c (Stack.append_app l Stack.empty)
-let beta_applist (c,l) =
- stacklam Stack.zip [] c (Stack.append_app_list l Stack.empty)
+let beta_applist sigma (c,l) =
+ let zip s = Stack.zip sigma s in
+ stacklam zip [] sigma c (Stack.append_app_list l Stack.empty)
(* Iota reduction tools *)
@@ -657,7 +684,7 @@ type 'a miota_args = {
mcargs : 'a list; (* the constructor's arguments *)
mlf : 'a array } (* the branch code vector *)
-let reducible_mind_case c = match kind_of_term c with
+let reducible_mind_case sigma c = match EConstr.kind sigma c with
| Construct _ | CoFix _ -> true
| _ -> false
@@ -671,27 +698,38 @@ let reducible_mind_case c = match kind_of_term c with
f x := t. End M. Definition f := u. and say goodbye to any hope
of refolding M.f this way ...
*)
-let magicaly_constant_of_fixbody env reference bd = function
+let magicaly_constant_of_fixbody env sigma reference bd = function
| Name.Anonymous -> bd
| Name.Name id ->
+ let open Universes in
try
let (cst_mod,cst_sect,_) = Constant.repr3 reference in
let cst = Constant.make3 cst_mod cst_sect (Label.of_id id) in
- let (cst, u), ctx = Universes.fresh_constant_instance env cst in
+ let (cst, u), ctx = fresh_constant_instance env cst in
match constant_opt_value_in env (cst,u) with
| None -> bd
| Some t ->
- let b, csts = Universes.eq_constr_universes t bd in
- let subst = Universes.Constraints.fold (fun (l,d,r) acc ->
- Univ.LMap.add (Option.get (Universe.level l)) (Option.get (Universe.level r)) acc)
- csts Univ.LMap.empty
- in
- let inst = Instance.subst_fn (fun u -> Univ.LMap.find u subst) u in
- if b then mkConstU (cst,inst) else bd
+ let csts = EConstr.eq_constr_universes env sigma (EConstr.of_constr t) bd in
+ begin match csts with
+ | Some csts ->
+ let subst = Constraints.fold (fun cst acc ->
+ let l, r = match cst with
+ | ULub (u, v) | UWeak (u, v) -> u, v
+ | UEq (u, v) | ULe (u, v) ->
+ let get u = Option.get (Universe.level u) in
+ get u, get v
+ in
+ Univ.LMap.add l r acc)
+ csts Univ.LMap.empty
+ in
+ let inst = Instance.subst_fn (fun u -> Univ.LMap.find u subst) u in
+ mkConstU (cst, EInstance.make inst)
+ | None -> bd
+ end
with
| Not_found -> bd
-let contract_cofix ?env ?reference (bodynum,(names,types,bodies as typedbodies)) =
+let contract_cofix ?env sigma ?reference (bodynum,(names,types,bodies as typedbodies)) =
let nbodies = Array.length bodies in
let make_Fi j =
let ind = nbodies-j-1 in
@@ -703,37 +741,37 @@ let contract_cofix ?env ?reference (bodynum,(names,types,bodies as typedbodies))
| Some e ->
match reference with
| None -> bd
- | Some r -> magicaly_constant_of_fixbody e r bd names.(ind) in
+ | Some r -> magicaly_constant_of_fixbody e sigma r bd names.(ind) in
let closure = List.init nbodies make_Fi in
substl closure bodies.(bodynum)
(** Similar to the "fix" case below *)
-let reduce_and_refold_cofix recfun env refold cst_l cofix sk =
+let reduce_and_refold_cofix recfun env sigma refold cst_l cofix sk =
let raw_answer =
let env = if refold then Some env else None in
- contract_cofix ?env ?reference:(Cst_stack.reference cst_l) cofix in
+ contract_cofix ?env sigma ?reference:(Cst_stack.reference sigma cst_l) cofix in
apply_subst
- (fun x (t,sk') ->
+ (fun sigma x (t,sk') ->
let t' =
- if refold then Cst_stack.best_replace (mkCoFix cofix) cst_l t else t in
+ if refold then Cst_stack.best_replace sigma (mkCoFix cofix) cst_l t else t in
recfun x (t',sk'))
- [] refold Cst_stack.empty raw_answer sk
+ [] sigma refold Cst_stack.empty raw_answer sk
-let reduce_mind_case mia =
- match kind_of_term mia.mconstr with
+let reduce_mind_case sigma mia =
+ match EConstr.kind sigma mia.mconstr with
| Construct ((ind_sp,i),u) ->
(* let ncargs = (fst mia.mci).(i-1) in*)
let real_cargs = List.skipn mia.mci.ci_npar mia.mcargs in
applist (mia.mlf.(i-1),real_cargs)
| CoFix cofix ->
- let cofix_def = contract_cofix cofix in
+ let cofix_def = contract_cofix sigma cofix in
mkCase (mia.mci, mia.mP, applist(cofix_def,mia.mcargs), mia.mlf)
| _ -> assert false
(* contracts fix==FIX[nl;i](A1...Ak;[F1...Fk]{B1....Bk}) to produce
Bi[Fj --> FIX[nl;j](A1...Ak;[F1...Fk]{B1...Bk})] *)
-let contract_fix ?env ?reference ((recindices,bodynum),(names,types,bodies as typedbodies)) =
+let contract_fix ?env sigma ?reference ((recindices,bodynum),(names,types,bodies as typedbodies)) =
let nbodies = Array.length recindices in
let make_Fi j =
let ind = nbodies-j-1 in
@@ -745,7 +783,7 @@ let contract_fix ?env ?reference ((recindices,bodynum),(names,types,bodies as ty
| Some e ->
match reference with
| None -> bd
- | Some r -> magicaly_constant_of_fixbody e r bd names.(ind) in
+ | Some r -> magicaly_constant_of_fixbody e sigma r bd names.(ind) in
let closure = List.init nbodies make_Fi in
substl closure bodies.(bodynum)
@@ -753,18 +791,18 @@ let contract_fix ?env ?reference ((recindices,bodynum),(names,types,bodies as ty
replace the fixpoint by the best constant from [cst_l]
Other rels are directly substituted by constants "magically found from the
context" in contract_fix *)
-let reduce_and_refold_fix recfun env refold cst_l fix sk =
+let reduce_and_refold_fix recfun env sigma refold cst_l fix sk =
let raw_answer =
- let env = if refold then None else Some env in
- contract_fix ?env ?reference:(Cst_stack.reference cst_l) fix in
+ let env = if refold then Some env else None in
+ contract_fix ?env sigma ?reference:(Cst_stack.reference sigma cst_l) fix in
apply_subst
- (fun x (t,sk') ->
+ (fun sigma x (t,sk') ->
let t' =
if refold then
- Cst_stack.best_replace (mkFix fix) cst_l t
+ Cst_stack.best_replace sigma (mkFix fix) cst_l t
else t
in recfun x (t',sk'))
- [] refold Cst_stack.empty raw_answer sk
+ [] sigma refold Cst_stack.empty raw_answer sk
let fix_recarg ((recindices,bodynum),_) stack =
assert (0 <= bodynum && bodynum < Array.length recindices);
@@ -789,7 +827,7 @@ let fix_recarg ((recindices,bodynum),_) stack =
let debug_RAKAM = ref (false)
let _ = Goptions.declare_bool_option {
- Goptions.optsync = true; Goptions.optdepr = false;
+ Goptions.optdepr = false;
Goptions.optname =
"Print states of the Reductionops abstract machine";
Goptions.optkey = ["Debug";"RAKAM"];
@@ -797,30 +835,30 @@ let _ = Goptions.declare_bool_option {
Goptions.optwrite = (fun a -> debug_RAKAM:=a);
}
-let equal_stacks (x, l) (y, l') =
- let f_equal (x,lft1) (y,lft2) = Constr.equal (Vars.lift lft1 x) (Vars.lift lft2 y) in
- let eq_fix (a,b) (c,d) = f_equal (Constr.mkFix a, b) (Constr.mkFix c, d) in
- match Stack.equal f_equal eq_fix l l' with
- | None -> false
- | Some (lft1,lft2) -> f_equal (x, lft1) (y, lft2)
+let equal_stacks sigma (x, l) (y, l') =
+ let f_equal x y = eq_constr sigma x y in
+ let eq_fix a b = f_equal (mkFix a) (mkFix b) in
+ Stack.equal f_equal eq_fix l l' && f_equal x y
let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
let open Context.Named.Declaration in
- let rec whrec cst_l (x, stack as s) =
+ let rec whrec cst_l (x, stack) =
let () = if !debug_RAKAM then
let open Pp in
+ let pr c = Termops.print_constr c in
Feedback.msg_notice
- (h 0 (str "<<" ++ Termops.print_constr x ++
+ (h 0 (str "<<" ++ pr x ++
str "|" ++ cut () ++ Cst_stack.pr cst_l ++
- str "|" ++ cut () ++ Stack.pr Termops.print_constr stack ++
+ str "|" ++ cut () ++ Stack.pr pr stack ++
str ">>"))
in
+ let c0 = EConstr.kind sigma x in
let fold () =
let () = if !debug_RAKAM then
let open Pp in Feedback.msg_notice (str "<><><><><>") in
- (s,cst_l)
+ ((EConstr.of_kind c0, stack),cst_l)
in
- match kind_of_term x with
+ match c0 with
| Rel n when CClosure.RedFlags.red_set flags CClosure.RedFlags.fDELTA ->
(match lookup_rel n env with
| LocalDef (_,body,_) -> whrec Cst_stack.empty (lift n body, stack)
@@ -830,18 +868,20 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
| LocalDef (_,body,_) ->
whrec (if refold then Cst_stack.add_cst (mkVar id) cst_l else cst_l) (body, stack)
| _ -> fold ())
- | Evar ev ->
- (match safe_evar_value sigma ev with
- | Some body -> whrec cst_l (body, stack)
- | None -> fold ())
+ | Evar ev -> fold ()
| Meta ev ->
(match safe_meta_value sigma ev with
- | Some body -> whrec cst_l (body, stack)
+ | Some body -> whrec cst_l (EConstr.of_constr body, stack)
| None -> fold ())
- | Const (c,u as const) when CClosure.RedFlags.red_set flags (CClosure.RedFlags.fCONST c) ->
- (match constant_opt_value_in env const with
+ | Const (c,u as const) ->
+ reduction_effect_hook env sigma (EConstr.to_constr sigma x)
+ (lazy (EConstr.to_constr sigma (Stack.zip sigma (x,stack))));
+ if CClosure.RedFlags.red_set flags (CClosure.RedFlags.fCONST c) then
+ let u' = EInstance.kind sigma u in
+ (match constant_opt_value_in env (c, u') with
| None -> fold ()
| Some body ->
+ let body = EConstr.of_constr body in
if not tactic_mode
then whrec (if refold then Cst_stack.add_cst (mkConstU const) cst_l else cst_l)
(body, stack)
@@ -858,12 +898,12 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
let (tm',sk'),cst_l' =
whrec (Cst_stack.add_cst (mkConstU const) cst_l) (body, app_sk)
in
- let rec is_case x = match kind_of_term x with
+ let rec is_case x = match EConstr.kind sigma x with
| Lambda (_,_, x) | LetIn (_,_,_, x) | Cast (x, _,_) -> is_case x
| App (hd, _) -> is_case hd
| Case _ -> true
| _ -> false in
- if equal_stacks (x, app_sk) (tm', sk')
+ if equal_stacks sigma (x, app_sk) (tm', sk')
|| Stack.will_expose_iota sk'
|| is_case tm'
then fold ()
@@ -877,8 +917,8 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
| None -> fold ()
| Some (bef,arg,s') ->
whrec Cst_stack.empty
- (arg,Stack.Cst(Stack.Cst_const const,curr,remains,bef,cst_l)::s')
- )
+ (arg,Stack.Cst(Stack.Cst_const (fst const, u'),curr,remains,bef,cst_l)::s')
+ ) else fold ()
| Proj (p, c) when CClosure.RedFlags.red_projection flags p ->
(let pb = lookup_projection p env in
let kn = Projection.constant p in
@@ -891,7 +931,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
| None ->
let stack' = (c, Stack.Proj (npars, arg, p, cst_l) :: stack) in
let stack'', csts = whrec Cst_stack.empty stack' in
- if equal_stacks stack' stack'' then fold ()
+ if equal_stacks sigma stack' stack'' then fold ()
else stack'', csts
| Some (recargs, nargs, flags) ->
if (List.mem `ReductionNeverUnfold flags
@@ -921,7 +961,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
Stack.append_app [|c|] bef,cst_l)::s'))
| LetIn (_,b,_,c) when CClosure.RedFlags.red_set flags CClosure.RedFlags.fZETA ->
- apply_subst whrec [b] refold cst_l c stack
+ apply_subst (fun _ -> whrec) [b] sigma refold cst_l c stack
| Cast (c,_,_) -> whrec cst_l (c, stack)
| App (f,cl) ->
whrec
@@ -930,20 +970,20 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
| Lambda (na,t,c) ->
(match Stack.decomp stack with
| Some _ when CClosure.RedFlags.red_set flags CClosure.RedFlags.fBETA ->
- apply_subst whrec [] refold cst_l x stack
+ apply_subst (fun _ -> whrec) [] sigma refold cst_l x stack
| None when CClosure.RedFlags.red_set flags CClosure.RedFlags.fETA ->
- let env' = push_rel (LocalAssum (na,t)) env in
+ let env' = push_rel (LocalAssum (na, t)) env in
let whrec' = whd_state_gen ~refold ~tactic_mode flags env' sigma in
- (match kind_of_term (Stack.zip ~refold (fst (whrec' (c, Stack.empty)))) with
+ (match EConstr.kind sigma (Stack.zip ~refold sigma (fst (whrec' (c, Stack.empty)))) with
| App (f,cl) ->
let napp = Array.length cl in
if napp > 0 then
let (x', l'),_ = whrec' (Array.last cl, Stack.empty) in
- match kind_of_term x', l' with
+ match EConstr.kind sigma x', l' with
| Rel 1, [] ->
let lc = Array.sub cl 0 (napp-1) in
- let u = if Int.equal napp 1 then f else appvect (f,lc) in
- if noccurn 1 u then (pop u,Stack.empty),Cst_stack.empty else fold ()
+ let u = if Int.equal napp 1 then f else mkApp (f,lc) in
+ if noccurn sigma 1 u then (pop u,Stack.empty),Cst_stack.empty else fold ()
| _ -> fold ()
else fold ()
| _ -> fold ())
@@ -968,11 +1008,11 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
|args, (Stack.Proj (n,m,p,_)::s') when use_match ->
whrec Cst_stack.empty (Stack.nth args (n+m), s')
|args, (Stack.Fix (f,s',cst_l)::s'') when use_fix ->
- let x' = Stack.zip(x,args) in
+ let x' = Stack.zip sigma (x, args) in
let out_sk = s' @ (Stack.append_app [|x'|] s'') in
- reduce_and_refold_fix whrec env refold cst_l f out_sk
+ reduce_and_refold_fix whrec env sigma refold cst_l f out_sk
|args, (Stack.Cst (const,curr,remains,s',cst_l) :: s'') ->
- let x' = Stack.zip(x,args) in
+ let x' = Stack.zip sigma (x, args) in
begin match remains with
| [] ->
(match const with
@@ -980,6 +1020,8 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
(match constant_opt_value_in env const with
| None -> fold ()
| Some body ->
+ let const = (fst const, EInstance.make (snd const)) in
+ let body = EConstr.of_constr body in
whrec (if refold then Cst_stack.add_cst (mkConstU const) cst_l else cst_l)
(body, s' @ (Stack.append_app [|x'|] s'')))
| Stack.Cst_proj p ->
@@ -998,7 +1040,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
(arg,
Stack.Cst (const,next,remains',s' @ (Stack.append_app [|x'|] bef),cst_l) :: s''')
end
- |_, (Stack.App _|Stack.Update _|Stack.Shift _)::_ -> assert false
+ |_, (Stack.App _)::_ -> assert false
|_, _ -> fold ()
else fold ()
@@ -1006,40 +1048,42 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
if CClosure.RedFlags.red_set flags CClosure.RedFlags.fCOFIX then
match Stack.strip_app stack with
|args, ((Stack.Case _ |Stack.Proj _)::s') ->
- reduce_and_refold_cofix whrec env refold cst_l cofix stack
+ reduce_and_refold_cofix whrec env sigma refold cst_l cofix stack
|_ -> fold ()
else fold ()
- | Rel _ | Var _ | Const _ | LetIn _ | Proj _ -> fold ()
+ | Rel _ | Var _ | LetIn _ | Proj _ -> fold ()
| Sort _ | Ind _ | Prod _ -> fold ()
in
fun xs ->
let (s,cst_l as res) = whrec (Option.default Cst_stack.empty csts) xs in
- if tactic_mode then (Stack.best_state s cst_l,Cst_stack.empty) else res
+ if tactic_mode then (Stack.best_state sigma s cst_l,Cst_stack.empty) else res
(** reduction machine without global env and refold machinery *)
let local_whd_state_gen flags sigma =
- let rec whrec (x, stack as s) =
- match kind_of_term x with
+ let rec whrec (x, stack) =
+ let c0 = EConstr.kind sigma x in
+ let s = (EConstr.of_kind c0, stack) in
+ match c0 with
| LetIn (_,b,_,c) when CClosure.RedFlags.red_set flags CClosure.RedFlags.fZETA ->
- stacklam whrec [b] c stack
+ stacklam whrec [b] sigma c stack
| Cast (c,_,_) -> whrec (c, stack)
| App (f,cl) -> whrec (f, Stack.append_app cl stack)
| Lambda (_,_,c) ->
(match Stack.decomp stack with
| Some (a,m) when CClosure.RedFlags.red_set flags CClosure.RedFlags.fBETA ->
- stacklam whrec [a] c m
+ stacklam whrec [a] sigma c m
| None when CClosure.RedFlags.red_set flags CClosure.RedFlags.fETA ->
- (match kind_of_term (Stack.zip (whrec (c, Stack.empty))) with
+ (match EConstr.kind sigma (Stack.zip sigma (whrec (c, Stack.empty))) with
| App (f,cl) ->
let napp = Array.length cl in
if napp > 0 then
let x', l' = whrec (Array.last cl, Stack.empty) in
- match kind_of_term x', l' with
+ match EConstr.kind sigma x', l' with
| Rel 1, [] ->
let lc = Array.sub cl 0 (napp-1) in
- let u = if Int.equal napp 1 then f else appvect (f,lc) in
- if noccurn 1 u then (pop u,Stack.empty) else s
+ let u = if Int.equal napp 1 then f else mkApp (f,lc) in
+ if noccurn sigma 1 u then (pop u,Stack.empty) else s
| _ -> s
else s
| _ -> s)
@@ -1059,14 +1103,10 @@ let local_whd_state_gen flags sigma =
|None -> s
|Some (bef,arg,s') -> whrec (arg, Stack.Fix(f,bef,Cst_stack.empty)::s'))
- | Evar ev ->
- (match safe_evar_value sigma ev with
- Some c -> whrec (c,stack)
- | None -> s)
-
+ | Evar ev -> s
| Meta ev ->
(match safe_meta_value sigma ev with
- Some c -> whrec (c,stack)
+ Some c -> whrec (EConstr.of_constr c,stack)
| None -> s)
| Construct ((ind,c),u) ->
@@ -1079,9 +1119,9 @@ let local_whd_state_gen flags sigma =
|args, (Stack.Proj (n,m,p,_) :: s') when use_match ->
whrec (Stack.nth args (n+m), s')
|args, (Stack.Fix (f,s',cst)::s'') when use_fix ->
- let x' = Stack.zip(x,args) in
- whrec (contract_fix f, s' @ (Stack.append_app [|x'|] s''))
- |_, (Stack.App _|Stack.Update _|Stack.Shift _|Stack.Cst _)::_ -> assert false
+ let x' = Stack.zip sigma (x,args) in
+ whrec (contract_fix sigma f, s' @ (Stack.append_app [|x'|] s''))
+ |_, (Stack.App _|Stack.Cst _)::_ -> assert false
|_, _ -> s
else s
@@ -1089,32 +1129,35 @@ let local_whd_state_gen flags sigma =
if CClosure.RedFlags.red_set flags CClosure.RedFlags.fCOFIX then
match Stack.strip_app stack with
|args, ((Stack.Case _ | Stack.Proj _)::s') ->
- whrec (contract_cofix cofix, stack)
+ whrec (contract_cofix sigma cofix, stack)
|_ -> s
else s
- | x -> s
+ | Rel _ | Var _ | Sort _ | Prod _ | LetIn _ | Const _ | Ind _ | Proj _ -> s
+
in
whrec
let raw_whd_state_gen flags env =
- let f sigma s = fst (whd_state_gen (get_refolding_in_reduction ()) false flags env sigma s) in
+ let f sigma s = fst (whd_state_gen ~refold:false
+ ~tactic_mode:false
+ flags env sigma s) in
f
let stack_red_of_state_red f =
- let f sigma x = decompose_app (Stack.zip (f sigma (x, Stack.empty))) in
+ let f sigma x = EConstr.decompose_app sigma (Stack.zip sigma (f sigma (x, Stack.empty))) in
f
(* Drops the Cst_stack *)
let iterate_whd_gen refold flags env sigma s =
let rec aux t =
- let (hd,sk),_ = whd_state_gen refold false flags env sigma (t,Stack.empty) in
+ let (hd,sk),_ = whd_state_gen ~refold ~tactic_mode:false flags env sigma (t,Stack.empty) in
let whd_sk = Stack.map aux sk in
- Stack.zip ~refold (hd,whd_sk)
+ Stack.zip sigma ~refold (hd,whd_sk)
in aux s
let red_of_state_red f sigma x =
- Stack.zip (f sigma (x,Stack.empty))
+ Stack.zip sigma (f sigma (x,Stack.empty))
(* 0. No Reduction Functions *)
@@ -1169,7 +1212,7 @@ let whd_allnolet env =
(* 4. Ad-hoc eta reduction, does not subsitute evars *)
-let shrink_eta c = Stack.zip (local_whd_state_gen eta Evd.empty (c,Stack.empty))
+let shrink_eta c = Stack.zip Evd.empty (local_whd_state_gen eta Evd.empty (c,Stack.empty))
(* 5. Zeta Reduction Functions *)
@@ -1191,14 +1234,24 @@ let nf_evar = Evarutil.nf_evar
let clos_norm_flags flgs env sigma t =
try
let evars ev = safe_evar_value sigma ev in
- CClosure.norm_val
+ EConstr.of_constr (CClosure.norm_val
+ (CClosure.create_clos_infos ~evars flgs env)
+ (CClosure.create_tab ())
+ (CClosure.inject (EConstr.Unsafe.to_constr t)))
+ with e when is_anomaly e -> user_err Pp.(str "Tried to normalize ill-typed term")
+
+let clos_whd_flags flgs env sigma t =
+ try
+ let evars ev = safe_evar_value sigma ev in
+ EConstr.of_constr (CClosure.whd_val
(CClosure.create_clos_infos ~evars flgs env)
- (CClosure.inject t)
- with e when is_anomaly e -> error "Tried to normalize ill-typed term"
+ (CClosure.create_tab ())
+ (CClosure.inject (EConstr.Unsafe.to_constr t)))
+ with e when is_anomaly e -> user_err Pp.(str "Tried to normalize ill-typed term")
-let nf_beta = clos_norm_flags CClosure.beta (Global.env ())
-let nf_betaiota = clos_norm_flags CClosure.betaiota (Global.env ())
-let nf_betaiotazeta = clos_norm_flags CClosure.betaiotazeta (Global.env ())
+let nf_beta = clos_norm_flags CClosure.beta
+let nf_betaiota = clos_norm_flags CClosure.betaiota
+let nf_betaiotazeta = clos_norm_flags CClosure.betaiotazeta
let nf_all env sigma =
clos_norm_flags CClosure.all env sigma
@@ -1207,11 +1260,11 @@ let nf_all env sigma =
(* Conversion *)
(********************************************************************)
(*
-let fkey = Profile.declare_profile "fhnf";;
-let fhnf info v = Profile.profile2 fkey fhnf info v;;
+let fkey = CProfile.declare_profile "fhnf";;
+let fhnf info v = CProfile.profile2 fkey fhnf info v;;
-let fakey = Profile.declare_profile "fhnf_apply";;
-let fhnf_apply info k h a = Profile.profile4 fakey fhnf_apply info k h a;;
+let fakey = CProfile.declare_profile "fhnf_apply";;
+let fhnf_apply info k h a = CProfile.profile4 fakey fhnf_apply info k h a;;
*)
let is_transparent e k =
@@ -1221,7 +1274,7 @@ let is_transparent e k =
(* Conversion utility functions *)
-type conversion_test = constraints -> constraints
+type conversion_test = Constraint.t -> Constraint.t
let pb_is_equal pb = pb == Reduction.CONV
@@ -1229,11 +1282,21 @@ let pb_equal = function
| Reduction.CUMUL -> Reduction.CONV
| Reduction.CONV -> Reduction.CONV
-let report_anomaly _ =
- let e = UserError ("", Pp.str "Conversion test raised an anomaly") in
+let report_anomaly e =
+ let msg = Pp.(str "Conversion test raised an anomaly:" ++
+ spc () ++ CErrors.print e) in
+ let e = UserError (None,msg) in
let e = CErrors.push e in
iraise e
+let f_conv ?l2r ?reds env ?evars x y =
+ let inj = EConstr.Unsafe.to_constr in
+ Reduction.conv ?l2r ?reds env ?evars (inj x) (inj y)
+
+let f_conv_leq ?l2r ?reds env ?evars x y =
+ let inj = EConstr.Unsafe.to_constr in
+ Reduction.conv_leq ?l2r ?reds env ?evars (inj x) (inj y)
+
let test_trans_conversion (f: constr Reduction.extended_conversion_function) reds env sigma x y =
try
let evars ev = safe_evar_value sigma ev in
@@ -1242,16 +1305,16 @@ let test_trans_conversion (f: constr Reduction.extended_conversion_function) red
with Reduction.NotConvertible -> false
| e when is_anomaly e -> report_anomaly e
-let is_conv ?(reds=full_transparent_state) env sigma = test_trans_conversion Reduction.conv reds env sigma
-let is_conv_leq ?(reds=full_transparent_state) env sigma = test_trans_conversion Reduction.conv_leq reds env sigma
+let is_conv ?(reds=full_transparent_state) env sigma = test_trans_conversion f_conv reds env sigma
+let is_conv_leq ?(reds=full_transparent_state) env sigma = test_trans_conversion f_conv_leq reds env sigma
let is_fconv ?(reds=full_transparent_state) = function
| Reduction.CONV -> is_conv ~reds
| Reduction.CUMUL -> is_conv_leq ~reds
let check_conv ?(pb=Reduction.CUMUL) ?(ts=full_transparent_state) env sigma x y =
let f = match pb with
- | Reduction.CONV -> Reduction.conv
- | Reduction.CUMUL -> Reduction.conv_leq
+ | Reduction.CONV -> f_conv
+ | Reduction.CUMUL -> f_conv_leq
in
try f ~reds:ts env ~evars:(safe_evar_value sigma, Evd.universes sigma) x y; true
with Reduction.NotConvertible -> false
@@ -1269,23 +1332,34 @@ let sigma_compare_instances ~flex i0 i1 sigma =
| Univ.UniverseInconsistency _ ->
raise Reduction.NotConvertible
+let sigma_check_inductive_instances cv_pb variance u1 u2 sigma =
+ match Evarutil.compare_cumulative_instances cv_pb variance u1 u2 sigma with
+ | Inl sigma -> sigma
+ | Inr _ ->
+ raise Reduction.NotConvertible
+
let sigma_univ_state =
- { Reduction.compare = sigma_compare_sorts;
- Reduction.compare_instances = sigma_compare_instances }
+ let open Reduction in
+ { compare_sorts = sigma_compare_sorts;
+ compare_instances = sigma_compare_instances;
+ compare_cumul_instances = sigma_check_inductive_instances; }
let infer_conv_gen conv_fun ?(catch_incon=true) ?(pb=Reduction.CUMUL)
?(ts=full_transparent_state) env sigma x y =
+ (** FIXME *)
try
- let fold cstr sigma =
- try Some (Evd.add_universe_constraints sigma cstr)
- with Univ.UniverseInconsistency _ | Evd.UniversesDiffer -> None
- in
let b, sigma =
let ans =
if pb == Reduction.CUMUL then
- Universes.leq_constr_univs_infer (Evd.universes sigma) fold x y sigma
+ EConstr.leq_constr_universes env sigma x y
else
- Universes.eq_constr_univs_infer (Evd.universes sigma) fold x y sigma
+ EConstr.eq_constr_universes env sigma x y
+ in
+ let ans = match ans with
+ | None -> None
+ | Some cstr ->
+ try Some (Evd.add_universe_constraints sigma cstr)
+ with Univ.UniverseInconsistency _ | Evd.UniversesDiffer -> None
in
match ans with
| None -> false, sigma
@@ -1293,6 +1367,8 @@ let infer_conv_gen conv_fun ?(catch_incon=true) ?(pb=Reduction.CUMUL)
in
if b then sigma, true
else
+ let x = EConstr.Unsafe.to_constr x in
+ let y = EConstr.Unsafe.to_constr y in
let sigma' =
conv_fun pb ~l2r:false sigma ts
env (sigma, sigma_univ_state) x y in
@@ -1315,37 +1391,37 @@ let vm_infer_conv ?(pb=Reduction.CUMUL) env t1 t2 =
(* Special-Purpose Reduction *)
(********************************************************************)
-let whd_meta sigma c = match kind_of_term c with
- | Meta p -> (try meta_value sigma p with Not_found -> c)
+let whd_meta sigma c = match EConstr.kind sigma c with
+ | Meta p -> (try EConstr.of_constr (meta_value sigma p) with Not_found -> c)
| _ -> c
let default_plain_instance_ident = Id.of_string "H"
(* Try to replace all metas. Does not replace metas in the metas' values
* Differs from (strong whd_meta). *)
-let plain_instance s c =
- let rec irec n u = match kind_of_term u with
+let plain_instance sigma s c =
+ let rec irec n u = match EConstr.kind sigma u with
| Meta p -> (try lift n (Metamap.find p s) with Not_found -> u)
- | App (f,l) when isCast f ->
- let (f,_,t) = destCast f in
+ | App (f,l) when isCast sigma f ->
+ let (f,_,t) = destCast sigma f in
let l' = CArray.Fun1.smartmap irec n l in
- (match kind_of_term f with
+ (match EConstr.kind sigma f with
| Meta p ->
(* Don't flatten application nodes: this is used to extract a
proof-term from a proof-tree and we want to keep the structure
of the proof-tree *)
(try let g = Metamap.find p s in
- match kind_of_term g with
+ match EConstr.kind sigma g with
| App _ ->
let l' = CArray.Fun1.smartmap lift 1 l' in
mkLetIn (Name default_plain_instance_ident,g,t,mkApp(mkRel 1, l'))
| _ -> mkApp (g,l')
with Not_found -> mkApp (f,l'))
| _ -> mkApp (irec n f,l'))
- | Cast (m,_,_) when isMeta m ->
- (try lift n (Metamap.find (destMeta m) s) with Not_found -> u)
+ | Cast (m,_,_) when isMeta sigma m ->
+ (try lift n (Metamap.find (destMeta sigma m) s) with Not_found -> u)
| _ ->
- map_constr_with_binders succ irec n u
+ map_with_binders sigma succ irec n u
in
if Metamap.is_empty s then c
else irec 0 c
@@ -1386,7 +1462,7 @@ let plain_instance s c =
let instance sigma s c =
(* if s = [] then c else *)
- local_strong whd_betaiota sigma (plain_instance s c)
+ local_strong whd_betaiota sigma (plain_instance sigma s c)
(* pseudo-reduction rule:
* [hnf_prod_app env s (Prod(_,B)) N --> B[N]
@@ -1395,34 +1471,33 @@ let instance sigma s c =
* error message. *)
let hnf_prod_app env sigma t n =
- match kind_of_term (whd_all env sigma t) with
+ match EConstr.kind sigma (whd_all env sigma t) with
| Prod (_,_,b) -> subst1 n b
- | _ -> anomaly ~label:"hnf_prod_app" (Pp.str "Need a product")
+ | _ -> anomaly ~label:"hnf_prod_app" (Pp.str "Need a product.")
let hnf_prod_appvect env sigma t nl =
- Array.fold_left (hnf_prod_app env sigma) t nl
+ Array.fold_left (fun acc t -> hnf_prod_app env sigma acc t) t nl
let hnf_prod_applist env sigma t nl =
- List.fold_left (hnf_prod_app env sigma) t nl
+ List.fold_left (fun acc t -> hnf_prod_app env sigma acc t) t nl
let hnf_lam_app env sigma t n =
- match kind_of_term (whd_all env sigma t) with
+ match EConstr.kind sigma (whd_all env sigma t) with
| Lambda (_,_,b) -> subst1 n b
- | _ -> anomaly ~label:"hnf_lam_app" (Pp.str "Need an abstraction")
+ | _ -> anomaly ~label:"hnf_lam_app" (Pp.str "Need an abstraction.")
let hnf_lam_appvect env sigma t nl =
- Array.fold_left (hnf_lam_app env sigma) t nl
+ Array.fold_left (fun acc t -> hnf_lam_app env sigma acc t) t nl
let hnf_lam_applist env sigma t nl =
- List.fold_left (hnf_lam_app env sigma) t nl
+ List.fold_left (fun acc t -> hnf_lam_app env sigma acc t) t nl
let splay_prod env sigma =
let rec decrec env m c =
let t = whd_all env sigma c in
- match kind_of_term t with
+ match EConstr.kind sigma t with
| Prod (n,a,c0) ->
- decrec (push_rel (LocalAssum (n,a)) env)
- ((n,a)::m) c0
+ decrec (push_rel (LocalAssum (n,a)) env) ((n,a)::m) c0
| _ -> m,t
in
decrec env []
@@ -1430,10 +1505,9 @@ let splay_prod env sigma =
let splay_lam env sigma =
let rec decrec env m c =
let t = whd_all env sigma c in
- match kind_of_term t with
+ match EConstr.kind sigma t with
| Lambda (n,a,c0) ->
- decrec (push_rel (LocalAssum (n,a)) env)
- ((n,a)::m) c0
+ decrec (push_rel (LocalAssum (n,a)) env) ((n,a)::m) c0
| _ -> m,t
in
decrec env []
@@ -1441,7 +1515,7 @@ let splay_lam env sigma =
let splay_prod_assum env sigma =
let rec prodec_rec env l c =
let t = whd_allnolet env sigma c in
- match kind_of_term t with
+ match EConstr.kind sigma t with
| Prod (x,t,c) ->
prodec_rec (push_rel (LocalAssum (x,t)) env)
(Context.Rel.add (LocalAssum (x,t)) l) c
@@ -1451,14 +1525,14 @@ let splay_prod_assum env sigma =
| Cast (c,_,_) -> prodec_rec env l c
| _ ->
let t' = whd_all env sigma t in
- if Term.eq_constr t t' then l,t
+ if EConstr.eq_constr sigma t t' then l,t
else prodec_rec env l t'
in
prodec_rec env Context.Rel.empty
let splay_arity env sigma c =
let l, c = splay_prod env sigma c in
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Sort s -> l,s
| _ -> invalid_arg "splay_arity"
@@ -1466,7 +1540,7 @@ let sort_of_arity env sigma c = snd (splay_arity env sigma c)
let splay_prod_n env sigma n =
let rec decrec env m ln c = if Int.equal m 0 then (ln,c) else
- match kind_of_term (whd_all env sigma c) with
+ match EConstr.kind sigma (whd_all env sigma c) with
| Prod (n,a,c0) ->
decrec (push_rel (LocalAssum (n,a)) env)
(m-1) (Context.Rel.add (LocalAssum (n,a)) ln) c0
@@ -1476,7 +1550,7 @@ let splay_prod_n env sigma n =
let splay_lam_n env sigma n =
let rec decrec env m ln c = if Int.equal m 0 then (ln,c) else
- match kind_of_term (whd_all env sigma c) with
+ match EConstr.kind sigma (whd_all env sigma c) with
| Lambda (n,a,c0) ->
decrec (push_rel (LocalAssum (n,a)) env)
(m-1) (Context.Rel.add (LocalAssum (n,a)) ln) c0
@@ -1485,7 +1559,7 @@ let splay_lam_n env sigma n =
decrec env n Context.Rel.empty
let is_sort env sigma t =
- match kind_of_term (whd_all env sigma t) with
+ match EConstr.kind sigma (whd_all env sigma t) with
| Sort s -> true
| _ -> false
@@ -1493,7 +1567,7 @@ let is_sort env sigma t =
of case/fix (heuristic used by evar_conv) *)
let whd_betaiota_deltazeta_for_iota_state ts env sigma csts s =
- let refold = get_refolding_in_reduction () in
+ let refold = false in
let tactic_mode = false in
let rec whrec csts s =
let (t, stack as s),csts' = whd_state_gen ~csts ~refold ~tactic_mode CClosure.betaiota env sigma s in
@@ -1501,24 +1575,24 @@ let whd_betaiota_deltazeta_for_iota_state ts env sigma csts s =
|args, (Stack.Case _ :: _ as stack') ->
let (t_o,stack_o),csts_o = whd_state_gen ~csts:csts' ~refold ~tactic_mode
(CClosure.RedFlags.red_add_transparent CClosure.all ts) env sigma (t,args) in
- if reducible_mind_case t_o then whrec csts_o (t_o, stack_o@stack') else s,csts'
+ if reducible_mind_case sigma t_o then whrec csts_o (t_o, stack_o@stack') else s,csts'
|args, (Stack.Fix _ :: _ as stack') ->
let (t_o,stack_o),csts_o = whd_state_gen ~csts:csts' ~refold ~tactic_mode
(CClosure.RedFlags.red_add_transparent CClosure.all ts) env sigma (t,args) in
- if isConstruct t_o then whrec csts_o (t_o, stack_o@stack') else s,csts'
+ if isConstruct sigma t_o then whrec csts_o (t_o, stack_o@stack') else s,csts'
|args, (Stack.Proj (n,m,p,_) :: stack'') ->
let (t_o,stack_o),csts_o = whd_state_gen ~csts:csts' ~refold ~tactic_mode
(CClosure.RedFlags.red_add_transparent CClosure.all ts) env sigma (t,args) in
- if isConstruct t_o then
+ if isConstruct sigma t_o then
whrec Cst_stack.empty (Stack.nth stack_o (n+m), stack'')
else s,csts'
- |_, ((Stack.App _| Stack.Shift _|Stack.Update _|Stack.Cst _) :: _|[]) -> s,csts'
+ |_, ((Stack.App _|Stack.Cst _) :: _|[]) -> s,csts'
in whrec csts s
let find_conclusion env sigma =
let rec decrec env c =
let t = whd_all env sigma c in
- match kind_of_term t with
+ match EConstr.kind sigma t with
| Prod (x,t,c0) -> decrec (push_rel (LocalAssum (x,t)) env) c0
| Lambda (x,t,c0) -> decrec (push_rel (LocalAssum (x,t)) env) c0
| t -> t
@@ -1538,7 +1612,7 @@ let meta_value evd mv =
match meta_opt_fvalue evd mv with
| Some (b,_) ->
let metas = Metamap.bind valrec b.freemetas in
- instance evd metas b.rebus
+ instance evd metas (EConstr.of_constr b.rebus)
| None -> mkMeta mv
in
valrec mv
@@ -1550,7 +1624,10 @@ let meta_instance sigma b =
let c_sigma = Metamap.bind (fun mv -> meta_value sigma mv) fm in
instance sigma c_sigma b.rebus
-let nf_meta sigma c = meta_instance sigma (mk_freelisted c)
+let nf_meta sigma c =
+ let c = EConstr.Unsafe.to_constr c in
+ let cl = mk_freelisted c in
+ meta_instance sigma { cl with rebus = EConstr.of_constr cl.rebus }
(* Instantiate metas that create beta/iota redexes *)
@@ -1564,78 +1641,80 @@ let meta_reducible_instance evd b =
in
let metas = Metaset.fold fold fm Metamap.empty in
let rec irec u =
- let u = whd_betaiota Evd.empty u in
- match kind_of_term u with
- | Case (ci,p,c,bl) when isMeta (strip_outer_cast c) ->
- let m = destMeta (strip_outer_cast c) in
+ let u = whd_betaiota Evd.empty u (** FIXME *) in
+ match EConstr.kind evd u with
+ | Case (ci,p,c,bl) when EConstr.isMeta evd (strip_outer_cast evd c) ->
+ let m = destMeta evd (strip_outer_cast evd c) in
(match
try
let g, s = Metamap.find m metas in
+ let g = EConstr.of_constr g in
let is_coerce = match s with CoerceToType -> true | _ -> false in
- if isConstruct g || not is_coerce then Some g else None
+ if isConstruct evd g || not is_coerce then Some g else None
with Not_found -> None
with
| Some g -> irec (mkCase (ci,p,g,bl))
| None -> mkCase (ci,irec p,c,Array.map irec bl))
- | App (f,l) when isMeta (strip_outer_cast f) ->
- let m = destMeta (strip_outer_cast f) in
+ | App (f,l) when EConstr.isMeta evd (strip_outer_cast evd f) ->
+ let m = destMeta evd (strip_outer_cast evd f) in
(match
try
let g, s = Metamap.find m metas in
+ let g = EConstr.of_constr g in
let is_coerce = match s with CoerceToType -> true | _ -> false in
- if isLambda g || not is_coerce then Some g else None
+ if isLambda evd g || not is_coerce then Some g else None
with Not_found -> None
with
| Some g -> irec (mkApp (g,l))
| None -> mkApp (f,Array.map irec l))
| Meta m ->
(try let g, s = Metamap.find m metas in
+ let g = EConstr.of_constr g in
let is_coerce = match s with CoerceToType -> true | _ -> false in
if not is_coerce then irec g else u
with Not_found -> u)
- | Proj (p,c) when isMeta c || isCast c && isMeta (pi1 (destCast c)) ->
- let m = try destMeta c with _ -> destMeta (pi1 (destCast c)) in
+ | Proj (p,c) when isMeta evd c || isCast evd c && isMeta evd (pi1 (destCast evd c)) (* What if two nested casts? *) ->
+ let m = try destMeta evd c with _ -> destMeta evd (pi1 (destCast evd c)) (* idem *) in
(match
try
let g, s = Metamap.find m metas in
+ let g = EConstr.of_constr g in
let is_coerce = match s with CoerceToType -> true | _ -> false in
- if isConstruct g || not is_coerce then Some g else None
+ if isConstruct evd g || not is_coerce then Some g else None
with Not_found -> None
with
| Some g -> irec (mkProj (p,g))
| None -> mkProj (p,c))
- | _ -> Constr.map irec u
+ | _ -> EConstr.map evd irec u
in
if Metaset.is_empty fm then (* nf_betaiota? *) b.rebus
else irec b.rebus
-let head_unfold_under_prod ts env _ c =
- let unfold (cst,u as cstu) =
+let head_unfold_under_prod ts env sigma c =
+ let unfold (cst,u) =
+ let cstu = (cst, EInstance.kind sigma u) in
if Cpred.mem cst (snd ts) then
match constant_opt_value_in env cstu with
- | Some c -> c
- | None -> mkConstU cstu
- else mkConstU cstu in
+ | Some c -> EConstr.of_constr c
+ | None -> mkConstU (cst, u)
+ else mkConstU (cst, u) in
let rec aux c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Prod (n,t,c) -> mkProd (n,aux t, aux c)
| _ ->
- let (h,l) = decompose_app c in
- match kind_of_term h with
- | Const cst -> beta_applist (unfold cst,l)
+ let (h,l) = decompose_app_vect sigma c in
+ match EConstr.kind sigma h with
+ | Const cst -> beta_app sigma (unfold cst, l)
| _ -> c in
aux c
let betazetaevar_applist sigma n c l =
let rec stacklam n env t stack =
if Int.equal n 0 then applist (substl env t, stack) else
- match kind_of_term t, stack with
+ match EConstr.kind sigma t, stack with
| Lambda(_,_,c), arg::stacktl -> stacklam (n-1) (arg::env) c stacktl
| LetIn(_,b,_,c), _ -> stacklam (n-1) (substl env b::env) c stack
- | Evar ev, _ ->
- (match safe_evar_value sigma ev with
- | Some body -> stacklam n env body stack
- | None -> applist (substl env t, stack))
- | _ -> anomaly (Pp.str "Not enough lambda/let's") in
+ | Evar _, _ -> applist (substl env t, stack)
+ | _ -> anomaly (Pp.str "Not enough lambda/let's.") in
stacklam n [] c l
diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli
index 4cd7a2a8..b8ac085a 100644
--- a/pretyping/reductionops.mli
+++ b/pretyping/reductionops.mli
@@ -1,13 +1,16 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
+open Constr
+open EConstr
open Univ
open Evd
open Environ
@@ -25,13 +28,24 @@ module ReductionBehaviour : sig
bool -> Globnames.global_reference -> (int list * int * flag list) -> unit
val get :
Globnames.global_reference -> (int list * int * flag list) option
- val print : Globnames.global_reference -> Pp.std_ppcmds
+ val print : Globnames.global_reference -> Pp.t
end
-(** Option telling if reduction should use the refolding machinery of cbn
- (off by default) *)
-val get_refolding_in_reduction : unit -> bool
-val set_refolding_in_reduction : bool -> unit
+(** {6 Support for reduction effects } *)
+
+type effect_name = string
+
+(* [declare_reduction_effect name f] declares [f] under key [name];
+ [name] must be a unique in "world". *)
+val declare_reduction_effect : effect_name ->
+ (Environ.env -> Evd.evar_map -> Constr.constr -> unit) -> unit
+
+(* [set_reduction_effect cst name] declares effect [name] to be called when [cst] is found *)
+val set_reduction_effect : Globnames.global_reference -> effect_name -> unit
+
+(* [effect_hook env sigma key term] apply effect associated to [key] on [term] *)
+val reduction_effect_hook : Environ.env -> Evd.evar_map -> Constr.constr ->
+ Constr.constr Lazy.t -> unit
(** {6 Machinery about a stack of unfolded constant }
@@ -44,33 +58,30 @@ module Cst_stack : sig
val add_args : constr array -> t -> t
val add_cst : constr -> t -> t
val best_cst : t -> (constr * constr list) option
- val best_replace : constr -> t -> constr -> constr
- val reference : t -> Constant.t option
- val pr : t -> Pp.std_ppcmds
+ val best_replace : Evd.evar_map -> constr -> t -> constr -> constr
+ val reference : Evd.evar_map -> t -> Constant.t option
+ val pr : t -> Pp.t
end
-
module Stack : sig
type 'a app_node
- val pr_app_node : ('a -> Pp.std_ppcmds) -> 'a app_node -> Pp.std_ppcmds
+ val pr_app_node : ('a -> Pp.t) -> 'a app_node -> Pp.t
type cst_member =
| Cst_const of pconstant
- | Cst_proj of projection
+ | Cst_proj of Projection.t
type 'a member =
| App of 'a app_node
| Case of case_info * 'a * 'a array * Cst_stack.t
- | Proj of int * int * projection * Cst_stack.t
- | Fix of fixpoint * 'a t * Cst_stack.t
+ | Proj of int * int * Projection.t * Cst_stack.t
+ | Fix of ('a, 'a) pfixpoint * 'a t * Cst_stack.t
| Cst of cst_member * int (** current foccussed arg *) * int list (** remaining args *)
* 'a t * Cst_stack.t
- | Shift of int
- | Update of 'a
and 'a t = 'a member list
- val pr : ('a -> Pp.std_ppcmds) -> 'a t -> Pp.std_ppcmds
+ val pr : ('a -> Pp.t) -> 'a t -> Pp.t
val empty : 'a t
val is_empty : 'a t -> bool
@@ -80,15 +91,18 @@ module Stack : sig
val decomp_node_last : 'a app_node -> 'a t -> ('a * 'a t)
val compare_shape : 'a t -> 'a t -> bool
+
+ exception IncompatibleFold2
(** [fold2 f x sk1 sk2] folds [f] on any pair of term in [(sk1,sk2)].
- @return the result and the lifts to apply on the terms *)
- val fold2 : ('a -> Term.constr -> Term.constr -> 'a) -> 'a ->
- Term.constr t -> Term.constr t -> 'a * int * int
- val map : (Term.constr -> Term.constr) -> Term.constr t -> Term.constr t
+ @return the result and the lifts to apply on the terms
+ @raise IncompatibleFold2 when [sk1] and [sk2] have incompatible shapes *)
+ val fold2 : ('a -> constr -> constr -> 'a) -> 'a ->
+ constr t -> constr t -> 'a
+ val map : ('a -> 'a) -> 'a t -> 'a t
val append_app_list : 'a list -> 'a t -> 'a t
(** if [strip_app s] = [(a,b)], then [s = a @ b] and [b] does not
- start by App or Shift *)
+ start by App *)
val strip_app : 'a t -> 'a t * 'a t
(** @return (the nth first elements, the (n+1)th element, the remaining stack) *)
val strip_n_app : int -> 'a t -> ('a t * 'a * 'a t) option
@@ -101,8 +115,8 @@ module Stack : sig
val tail : int -> 'a t -> 'a t
val nth : 'a t -> int -> 'a
- val best_state : constr * constr t -> Cst_stack.t -> constr * constr t
- val zip : ?refold:bool -> constr * constr t -> constr
+ val best_state : evar_map -> constr * constr t -> Cst_stack.t -> constr * constr t
+ val zip : ?refold:bool -> evar_map -> constr * constr t -> constr
end
(************************************************************************)
@@ -113,7 +127,7 @@ type contextual_reduction_function = env -> evar_map -> constr -> constr
type reduction_function = contextual_reduction_function
type local_reduction_function = evar_map -> constr -> constr
-type e_reduction_function = { e_redfun : 'r. env -> 'r Sigma.t -> constr -> (constr, 'r) Sigma.sigma }
+type e_reduction_function = env -> evar_map -> constr -> evar_map * constr
type contextual_stack_reduction_function =
env -> evar_map -> constr -> constr * constr list
@@ -126,7 +140,7 @@ type contextual_state_reduction_function =
type state_reduction_function = contextual_state_reduction_function
type local_state_reduction_function = evar_map -> state -> state
-val pr_state : state -> Pp.std_ppcmds
+val pr_state : state -> Pp.t
(** {6 Reduction Function Operators } *)
@@ -137,22 +151,23 @@ val strong_prodspine : local_reduction_function -> local_reduction_function
val stack_reduction_of_reduction :
'a reduction_function -> 'a state_reduction_function
i*)
-val stacklam : (state -> 'a) -> constr list -> constr -> constr Stack.t -> 'a
+val stacklam : (state -> 'a) -> constr list -> evar_map -> constr -> constr Stack.t -> 'a
val whd_state_gen : ?csts:Cst_stack.t -> refold:bool -> tactic_mode:bool ->
CClosure.RedFlags.reds -> Environ.env -> Evd.evar_map -> state -> state * Cst_stack.t
val iterate_whd_gen : bool -> CClosure.RedFlags.reds ->
- Environ.env -> Evd.evar_map -> Term.constr -> Term.constr
+ Environ.env -> Evd.evar_map -> constr -> constr
(** {6 Generic Optimized Reduction Function using Closures } *)
val clos_norm_flags : CClosure.RedFlags.reds -> reduction_function
+val clos_whd_flags : CClosure.RedFlags.reds -> reduction_function
(** Same as [(strong whd_beta[delta][iota])], but much faster on big terms *)
-val nf_beta : local_reduction_function
-val nf_betaiota : local_reduction_function
-val nf_betaiotazeta : local_reduction_function
+val nf_beta : reduction_function
+val nf_betaiota : reduction_function
+val nf_betaiotazeta : reduction_function
val nf_all : reduction_function
val nf_evar : evar_map -> constr -> constr
@@ -200,9 +215,9 @@ val shrink_eta : constr -> constr
(** Various reduction functions *)
-val safe_evar_value : evar_map -> existential -> constr option
+val safe_evar_value : evar_map -> Constr.existential -> Constr.constr option
-val beta_applist : constr * constr list -> constr
+val beta_applist : evar_map -> constr * constr list -> constr
val hnf_prod_app : env -> evar_map -> constr -> constr -> constr
val hnf_prod_appvect : env -> evar_map -> constr -> constr array -> constr
@@ -213,12 +228,12 @@ val hnf_lam_applist : env -> evar_map -> constr -> constr list -> constr
val splay_prod : env -> evar_map -> constr -> (Name.t * constr) list * constr
val splay_lam : env -> evar_map -> constr -> (Name.t * constr) list * constr
-val splay_arity : env -> evar_map -> constr -> (Name.t * constr) list * sorts
-val sort_of_arity : env -> evar_map -> constr -> sorts
-val splay_prod_n : env -> evar_map -> int -> constr -> Context.Rel.t * constr
-val splay_lam_n : env -> evar_map -> int -> constr -> Context.Rel.t * constr
+val splay_arity : env -> evar_map -> constr -> (Name.t * constr) list * ESorts.t
+val sort_of_arity : env -> evar_map -> constr -> ESorts.t
+val splay_prod_n : env -> evar_map -> int -> constr -> rel_context * constr
+val splay_lam_n : env -> evar_map -> int -> constr -> rel_context * constr
val splay_prod_assum :
- env -> evar_map -> constr -> Context.Rel.t * constr
+ env -> evar_map -> constr -> rel_context * constr
type 'a miota_args = {
mP : constr; (** the result type *)
@@ -227,22 +242,22 @@ type 'a miota_args = {
mcargs : 'a list; (** the constructor's arguments *)
mlf : 'a array } (** the branch code vector *)
-val reducible_mind_case : constr -> bool
-val reduce_mind_case : constr miota_args -> constr
+val reducible_mind_case : evar_map -> constr -> bool
+val reduce_mind_case : evar_map -> constr miota_args -> constr
-val find_conclusion : env -> evar_map -> constr -> (constr,constr) kind_of_term
+val find_conclusion : env -> evar_map -> constr -> (constr, constr, ESorts.t, EInstance.t) kind_of_term
val is_arity : env -> evar_map -> constr -> bool
val is_sort : env -> evar_map -> types -> bool
-val contract_fix : ?env:Environ.env -> ?reference:Constant.t -> fixpoint -> constr
-val fix_recarg : fixpoint -> constr Stack.t -> (int * constr) option
+val contract_fix : ?env:Environ.env -> evar_map -> ?reference:Constant.t -> fixpoint -> constr
+val fix_recarg : ('a, 'a) pfixpoint -> 'b Stack.t -> (int * 'b) option
(** {6 Querying the kernel conversion oracle: opaque/transparent constants } *)
-val is_transparent : Environ.env -> constant tableKey -> bool
+val is_transparent : Environ.env -> Constant.t tableKey -> bool
(** {6 Conversion Functions (uses closures, lazy strategy) } *)
-type conversion_test = constraints -> constraints
+type conversion_test = Constraint.t -> Constraint.t
val pb_is_equal : conv_pb -> bool
val pb_equal : conv_pb -> conv_pb
@@ -258,7 +273,7 @@ val check_conv : ?pb:conv_pb -> ?ts:transparent_state -> env -> evar_map -> con
(** [infer_conv] Adds necessary universe constraints to the evar map.
pb defaults to CUMUL and ts to a full transparent state.
- @raises UniverseInconsistency iff catch_incon is set to false,
+ @raise UniverseInconsistency iff catch_incon is set to false,
otherwise returns false in that case.
*)
val infer_conv : ?catch_incon:bool -> ?pb:conv_pb -> ?ts:transparent_state ->
@@ -274,14 +289,14 @@ val vm_infer_conv : ?pb:conv_pb -> env -> evar_map -> constr -> constr ->
(** [infer_conv_gen] behaves like [infer_conv] but is parametrized by a
conversion function. Used to pretype vm and native casts. *)
val infer_conv_gen : (conv_pb -> l2r:bool -> evar_map -> transparent_state ->
- (constr, evar_map) Reduction.generic_conversion_function) ->
+ (Constr.constr, evar_map) Reduction.generic_conversion_function) ->
?catch_incon:bool -> ?pb:conv_pb -> ?ts:transparent_state -> env ->
evar_map -> constr -> constr -> evar_map * bool
(** {6 Special-Purpose Reduction Functions } *)
-val whd_meta : evar_map -> constr -> constr
-val plain_instance : constr Metamap.t -> constr -> constr
+val whd_meta : local_reduction_function
+val plain_instance : evar_map -> constr Metamap.t -> constr -> constr
val instance : evar_map -> constr Metamap.t -> constr -> constr
val head_unfold_under_prod : transparent_state -> reduction_function
val betazetaevar_applist : evar_map -> int -> constr -> constr list -> constr
diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml
index 98b36fb9..3582b644 100644
--- a/pretyping/retyping.ml
+++ b/pretyping/retyping.ml
@@ -1,25 +1,32 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Pp
open CErrors
open Util
open Term
-open Vars
+open Constr
open Inductive
open Inductiveops
open Names
open Reductionops
open Environ
open Termops
+open EConstr
+open Vars
open Arguments_renaming
open Context.Rel.Declaration
+module RelDecl = Context.Rel.Declaration
+module NamedDecl = Context.Named.Declaration
+
type retype_error =
| NotASort
| NotAnArity
@@ -44,14 +51,14 @@ let retype_error re = raise (RetypeError re)
let anomaly_on_error f x =
try f x
- with RetypeError e -> anomaly ~label:"retyping" (print_retype_error e)
+ with RetypeError e -> anomaly ~label:"retyping" (print_retype_error e ++ str ".")
let get_type_from_constraints env sigma t =
- if isEvar (fst (decompose_app t)) then
+ if isEvar sigma (fst (decompose_app_vect sigma t)) then
match
List.map_filter (fun (pbty,env,t1,t2) ->
- if is_fconv Reduction.CONV env sigma t t1 then Some t2
- else if is_fconv Reduction.CONV env sigma t t2 then Some t1
+ if is_fconv Reduction.CONV env sigma t (EConstr.of_constr t1) then Some t2
+ else if is_fconv Reduction.CONV env sigma t (EConstr.of_constr t2) then Some t1
else None)
(snd (Evd.extract_all_conv_pbs sigma))
with
@@ -62,7 +69,7 @@ let get_type_from_constraints env sigma t =
let rec subst_type env sigma typ = function
| [] -> typ
| h::rest ->
- match kind_of_term (whd_all env sigma typ) with
+ match EConstr.kind sigma (whd_all env sigma typ) with
| Prod (na,c1,c2) -> subst_type env sigma (subst1 h c2) rest
| _ -> retype_error NonFunctionalConstruction
@@ -71,49 +78,50 @@ let rec subst_type env sigma typ = function
let sort_of_atomic_type env sigma ft args =
let rec concl_of_arity env n ar args =
- match kind_of_term (whd_all env sigma ar), args with
+ match EConstr.kind sigma (whd_all env sigma ar), args with
| Prod (na, t, b), h::l -> concl_of_arity (push_rel (LocalDef (na, lift n h, t)) env) (n + 1) b l
- | Sort s, [] -> s
+ | Sort s, [] -> ESorts.kind sigma s
| _ -> retype_error NotASort
in concl_of_arity env 0 ft (Array.to_list args)
let type_of_var env id =
- let open Context.Named.Declaration in
- try get_type (lookup_named id env)
+ try NamedDecl.get_type (lookup_named id env)
with Not_found -> retype_error (BadVariable id)
let decomp_sort env sigma t =
- match kind_of_term (whd_all env sigma t) with
- | Sort s -> s
+ match EConstr.kind sigma (whd_all env sigma t) with
+ | Sort s -> ESorts.kind sigma s
| _ -> retype_error NotASort
+let destSort sigma s = ESorts.kind sigma (destSort sigma s)
+
let retype ?(polyprop=true) sigma =
let rec type_of env cstr =
- match kind_of_term cstr with
+ match EConstr.kind sigma cstr with
| Meta n ->
- (try strip_outer_cast (Evd.meta_ftype sigma n).Evd.rebus
+ (try strip_outer_cast sigma (EConstr.of_constr (Evd.meta_ftype sigma n).Evd.rebus)
with Not_found -> retype_error (BadMeta n))
| Rel n ->
- let ty = get_type (lookup_rel n env) in
+ let ty = RelDecl.get_type (lookup_rel n env) in
lift n ty
| Var id -> type_of_var env id
- | Const cst -> rename_type_of_constant env cst
- | Evar ev -> Evd.existential_type sigma ev
- | Ind ind -> rename_type_of_inductive env ind
- | Construct cstr -> rename_type_of_constructor env cstr
+ | Const (cst, u) -> EConstr.of_constr (rename_type_of_constant env (cst, EInstance.kind sigma u))
+ | Evar ev -> existential_type sigma ev
+ | Ind (ind, u) -> EConstr.of_constr (rename_type_of_inductive env (ind, EInstance.kind sigma u))
+ | Construct (cstr, u) -> EConstr.of_constr (rename_type_of_constructor env (cstr, EInstance.kind sigma u))
| Case (_,p,c,lf) ->
let Inductiveops.IndType(indf,realargs) =
let t = type_of env c in
try Inductiveops.find_rectype env sigma t
with Not_found ->
try
- let t = get_type_from_constraints env sigma t in
+ let t = EConstr.of_constr (get_type_from_constraints env sigma t) in
Inductiveops.find_rectype env sigma t
with Not_found -> retype_error BadRecursiveType
in
let n = inductive_nrealdecls_env env (fst (fst (dest_ind_family indf))) in
let t = betazetaevar_applist sigma n p realargs in
- (match kind_of_term (whd_all env sigma (type_of env t)) with
+ (match EConstr.kind sigma (whd_all env sigma (type_of env t)) with
| Prod _ -> whd_beta sigma (applist (t, [c]))
| _ -> t)
| Lambda (name,c1,c2) ->
@@ -122,25 +130,28 @@ let retype ?(polyprop=true) sigma =
subst1 b (type_of (push_rel (LocalDef (name,b,c1)) env) c2)
| Fix ((_,i),(_,tys,_)) -> tys.(i)
| CoFix (i,(_,tys,_)) -> tys.(i)
- | App(f,args) when is_template_polymorphic env f ->
+ | App(f,args) when is_template_polymorphic env sigma f ->
let t = type_of_global_reference_knowing_parameters env f args in
- strip_outer_cast (subst_type env sigma t (Array.to_list args))
+ strip_outer_cast sigma (subst_type env sigma t (Array.to_list args))
| App(f,args) ->
- strip_outer_cast
+ strip_outer_cast sigma
(subst_type env sigma (type_of env f) (Array.to_list args))
| Proj (p,c) ->
let ty = type_of env c in
- (try
+ EConstr.of_constr (try
Inductiveops.type_of_projection_knowing_arg env sigma p c ty
with Invalid_argument _ -> retype_error BadRecursiveType)
| Cast (c,_, t) -> t
| Sort _ | Prod _ -> mkSort (sort_of env cstr)
and sort_of env t =
- match kind_of_term t with
- | Cast (c,_, s) when isSort s -> destSort s
- | Sort (Prop c) -> type1_sort
- | Sort (Type u) -> Type (Univ.super u)
+ match EConstr.kind sigma t with
+ | Cast (c,_, s) when isSort sigma s -> destSort sigma s
+ | Sort s ->
+ begin match ESorts.kind sigma s with
+ | Prop _ -> Sorts.type1
+ | Type u -> Type (Univ.super u)
+ end
| Prod (name,t,c2) ->
(match (sort_of env t, sort_of (push_rel (LocalAssum (name,t)) env) c2) with
| _, (Prop Null as s) -> s
@@ -150,69 +161,67 @@ let retype ?(polyprop=true) sigma =
| Prop Pos, (Type u2) -> Type (Univ.sup Univ.type0_univ u2)
| Prop Null, (Type _ as s) -> s
| Type u1, Type u2 -> Type (Univ.sup u1 u2))
- | App(f,args) when is_template_polymorphic env f ->
+ | App(f,args) when is_template_polymorphic env sigma f ->
let t = type_of_global_reference_knowing_parameters env f args in
sort_of_atomic_type env sigma t args
| App(f,args) -> sort_of_atomic_type env sigma (type_of env f) args
| Lambda _ | Fix _ | Construct _ -> retype_error NotAType
| _ -> decomp_sort env sigma (type_of env t)
- and sort_family_of env t =
- match kind_of_term t with
- | Cast (c,_, s) when isSort s -> family_of_sort (destSort s)
- | Sort (Prop c) -> InType
- | Sort (Type u) -> InType
+ and type_of_global_reference_knowing_parameters env c args =
+ let argtyps =
+ Array.map (fun c -> lazy (EConstr.to_constr sigma (type_of env c))) args in
+ match EConstr.kind sigma c with
+ | Ind (ind, u) ->
+ let u = EInstance.kind sigma u in
+ let mip = lookup_mind_specif env ind in
+ EConstr.of_constr (try Inductive.type_of_inductive_knowing_parameters
+ ~polyprop env (mip, u) argtyps
+ with Reduction.NotArity -> retype_error NotAnArity)
+ | Construct (cstr, u) ->
+ let u = EInstance.kind sigma u in
+ EConstr.of_constr (type_of_constructor env (cstr, u))
+ | _ -> assert false
+
+ in type_of, sort_of, type_of_global_reference_knowing_parameters
+
+let get_sort_family_of ?(truncation_style=false) ?(polyprop=true) env sigma t =
+ let type_of,_,type_of_global_reference_knowing_parameters = retype ~polyprop sigma in
+ let rec sort_family_of env t =
+ match EConstr.kind sigma t with
+ | Cast (c,_, s) when isSort sigma s -> Sorts.family (destSort sigma s)
+ | Sort _ -> InType
| Prod (name,t,c2) ->
let s2 = sort_family_of (push_rel (LocalAssum (name,t)) env) c2 in
if not (is_impredicative_set env) &&
s2 == InSet && sort_family_of env t == InType then InType else s2
- | App(f,args) when is_template_polymorphic env f ->
+ | App(f,args) when is_template_polymorphic env sigma f ->
+ if truncation_style then InType else
let t = type_of_global_reference_knowing_parameters env f args in
- family_of_sort (sort_of_atomic_type env sigma t args)
+ Sorts.family (sort_of_atomic_type env sigma t args)
| App(f,args) ->
- family_of_sort (sort_of_atomic_type env sigma (type_of env f) args)
+ Sorts.family (sort_of_atomic_type env sigma (type_of env f) args)
| Lambda _ | Fix _ | Construct _ -> retype_error NotAType
+ | Ind _ when truncation_style && is_template_polymorphic env sigma t -> InType
| _ ->
- family_of_sort (decomp_sort env sigma (type_of env t))
-
- and type_of_global_reference_knowing_parameters env c args =
- let argtyps =
- Array.map (fun c -> lazy (nf_evar sigma (type_of env c))) args in
- match kind_of_term c with
- | Ind ind ->
- let mip = lookup_mind_specif env (fst ind) in
- (try Inductive.type_of_inductive_knowing_parameters
- ~polyprop env (mip,snd ind) argtyps
- with Reduction.NotArity -> retype_error NotAnArity)
- | Const cst ->
- (try Typeops.type_of_constant_knowing_parameters_in env cst argtyps
- with Reduction.NotArity -> retype_error NotAnArity)
- | Var id -> type_of_var env id
- | Construct cstr -> type_of_constructor env cstr
- | _ -> assert false
-
- in type_of, sort_of, sort_family_of,
- type_of_global_reference_knowing_parameters
+ Sorts.family (decomp_sort env sigma (type_of env t))
+ in sort_family_of env t
let get_sort_of ?(polyprop=true) env sigma t =
- let _,f,_,_ = retype ~polyprop sigma in anomaly_on_error (f env) t
-let get_sort_family_of ?(polyprop=true) env sigma c =
- let _,_,f,_ = retype ~polyprop sigma in anomaly_on_error (f env) c
+ let _,f,_ = retype ~polyprop sigma in anomaly_on_error (f env) t
let type_of_global_reference_knowing_parameters env sigma c args =
- let _,_,_,f = retype sigma in anomaly_on_error (f env c) args
+ let _,_,f = retype sigma in anomaly_on_error (f env c) args
let type_of_global_reference_knowing_conclusion env sigma c conclty =
- let conclty = nf_evar sigma conclty in
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Ind (ind,u) ->
let spec = Inductive.lookup_mind_specif env ind in
- type_of_inductive_knowing_conclusion env sigma (spec,u) conclty
- | Const cst ->
- let t = constant_type_in env cst in
- (* TODO *)
- sigma, Typeops.type_of_constant_type_knowing_parameters env t [||]
+ type_of_inductive_knowing_conclusion env sigma (spec, EInstance.kind sigma u) conclty
+ | Const (cst, u) ->
+ let t = constant_type_in env (cst, EInstance.kind sigma u) in
+ sigma, EConstr.of_constr t
| Var id -> sigma, type_of_var env id
- | Construct cstr -> sigma, type_of_constructor env cstr
+ | Construct (cstr, u) -> sigma, EConstr.of_constr (type_of_constructor env (cstr, EInstance.kind sigma u))
| _ -> assert false
(* Profiling *)
@@ -220,14 +229,14 @@ let type_of_global_reference_knowing_conclusion env sigma c conclty =
(* let f,_,_,_ = retype ~polyprop sigma in *)
(* if lax then f env c else anomaly_on_error (f env) c *)
-(* let get_type_of_key = Profile.declare_profile "get_type_of" *)
-(* let get_type_of = Profile.profile5 get_type_of_key get_type_of *)
+(* let get_type_of_key = CProfile.declare_profile "get_type_of" *)
+(* let get_type_of = CProfile.profile5 get_type_of_key get_type_of *)
(* let get_type_of ?(polyprop=true) ?(lax=false) env sigma c = *)
(* get_type_of polyprop lax env sigma c *)
let get_type_of ?(polyprop=true) ?(lax=false) env sigma c =
- let f,_,_,_ = retype ~polyprop sigma in
+ let f,_,_ = retype ~polyprop sigma in
if lax then f env c else anomaly_on_error (f env) c
(* Makes an unsafe judgment from a constr *)
@@ -239,7 +248,7 @@ let sorts_of_context env evc ctxt =
| [] -> env,[]
| d :: ctxt ->
let env,sorts = aux ctxt in
- let s = get_sort_of env evc (get_type d) in
+ let s = get_sort_of env evc (RelDecl.get_type d) in
(push_rel d env,s::sorts) in
snd (aux ctxt)
diff --git a/pretyping/retyping.mli b/pretyping/retyping.mli
index 8ca40f82..2aff0c77 100644
--- a/pretyping/retyping.mli
+++ b/pretyping/retyping.mli
@@ -1,14 +1,16 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
open Evd
open Environ
+open EConstr
(** This family of functions assumes its constr argument is known to be
well-typable. It does not type-check, just recompute the type
@@ -29,10 +31,13 @@ val get_type_of :
?polyprop:bool -> ?lax:bool -> env -> evar_map -> constr -> types
val get_sort_of :
- ?polyprop:bool -> env -> evar_map -> types -> sorts
+ ?polyprop:bool -> env -> evar_map -> types -> Sorts.t
+(* When [truncation_style] is [true], tells if the type has been explicitly
+ truncated to Prop or (impredicative) Set; in particular, singleton type and
+ small inductive types, which have all eliminations to Type, are in Type *)
val get_sort_family_of :
- ?polyprop:bool -> env -> evar_map -> types -> sorts_family
+ ?truncation_style:bool -> ?polyprop:bool -> env -> evar_map -> types -> Sorts.family
(** Makes an unsafe judgment from a constr *)
val get_judgment_of : env -> evar_map -> constr -> unsafe_judgment
@@ -43,8 +48,8 @@ val type_of_global_reference_knowing_parameters : env -> evar_map -> constr ->
val type_of_global_reference_knowing_conclusion :
env -> evar_map -> constr -> types -> evar_map * types
-val sorts_of_context : env -> evar_map -> Context.Rel.t -> sorts list
+val sorts_of_context : env -> evar_map -> rel_context -> Sorts.t list
-val expand_projection : env -> evar_map -> Names.projection -> constr -> constr list -> constr
+val expand_projection : env -> evar_map -> Names.Projection.t -> constr -> constr list -> constr
-val print_retype_error : retype_error -> Pp.std_ppcmds
+val print_retype_error : retype_error -> Pp.t
diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml
index 820a81b5..518d2f60 100644
--- a/pretyping/tacred.ml
+++ b/pretyping/tacred.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Pp
@@ -11,24 +13,27 @@ open CErrors
open Util
open Names
open Term
-open Vars
open Libnames
open Globnames
open Termops
+open Environ
+open EConstr
+open Vars
open Find_subterm
open Namegen
-open Environ
open CClosure
open Reductionops
open Cbv
open Patternops
open Locus
-open Sigma.Notations
+
+module RelDecl = Context.Rel.Declaration
+module NamedDecl = Context.Named.Declaration
(* Errors *)
type reduction_tactic_error =
- InvalidAbstraction of env * Evd.evar_map * constr * (env * Type_errors.type_error)
+ InvalidAbstraction of env * Evd.evar_map * EConstr.constr * (env * Type_errors.type_error)
exception ReductionTacticError of reduction_tactic_error
@@ -38,7 +43,7 @@ exception Elimconst
exception Redelimination
let error_not_evaluable r =
- errorlabstrm "error_not_evaluable"
+ user_err ~hdr:"error_not_evaluable"
(str "Cannot coerce" ++ spc () ++ Nametab.pr_global_env Id.Set.empty r ++
spc () ++ str "to an evaluable reference.")
@@ -54,13 +59,11 @@ let is_evaluable env = function
| EvalVarRef id -> is_evaluable_var env id
let value_of_evaluable_ref env evref u =
- let open Context.Named.Declaration in
match evref with
| EvalConstRef con ->
- (try constant_value_in env (con,u)
- with NotEvaluableConst IsProj ->
- raise (Invalid_argument "value_of_evaluable_ref"))
- | EvalVarRef id -> lookup_named id env |> get_value |> Option.get
+ let u = Unsafe.to_instance u in
+ EConstr.of_constr (constant_value_in env (con, u))
+ | EvalVarRef id -> env |> lookup_named id |> NamedDecl.get_value |> Option.get
let evaluable_of_global_reference env = function
| ConstRef cst when is_evaluable_const env cst -> EvalConstRef cst
@@ -72,17 +75,17 @@ let global_of_evaluable_reference = function
| EvalVarRef id -> VarRef id
type evaluable_reference =
- | EvalConst of constant
+ | EvalConst of Constant.t
| EvalVar of Id.t
| EvalRel of int
- | EvalEvar of existential
+ | EvalEvar of EConstr.existential
-let evaluable_reference_eq r1 r2 = match r1, r2 with
-| EvalConst c1, EvalConst c2 -> eq_constant c1 c2
+let evaluable_reference_eq sigma r1 r2 = match r1, r2 with
+| EvalConst c1, EvalConst c2 -> Constant.equal c1 c2
| EvalVar id1, EvalVar id2 -> Id.equal id1 id2
| EvalRel i1, EvalRel i2 -> Int.equal i1 i2
| EvalEvar (e1, ctx1), EvalEvar (e2, ctx2) ->
- Evar.equal e1 e2 && Array.equal eq_constr ctx1 ctx2
+ Evar.equal e1 e2 && Array.equal (EConstr.eq_constr sigma) ctx1 ctx2
| _ -> false
let mkEvalRef ref u =
@@ -90,45 +93,49 @@ let mkEvalRef ref u =
| EvalConst cst -> mkConstU (cst,u)
| EvalVar id -> mkVar id
| EvalRel n -> mkRel n
- | EvalEvar ev -> mkEvar ev
+ | EvalEvar ev -> EConstr.mkEvar ev
-let isEvalRef env c = match kind_of_term c with
+let isEvalRef env sigma c = match EConstr.kind sigma c with
| Const (sp,_) -> is_evaluable env (EvalConstRef sp)
| Var id -> is_evaluable env (EvalVarRef id)
| Rel _ | Evar _ -> true
| _ -> false
-let destEvalRefU c = match kind_of_term c with
+let destEvalRefU sigma c = match EConstr.kind sigma c with
| Const (cst,u) -> EvalConst cst, u
- | Var id -> (EvalVar id, Univ.Instance.empty)
- | Rel n -> (EvalRel n, Univ.Instance.empty)
- | Evar ev -> (EvalEvar ev, Univ.Instance.empty)
- | _ -> anomaly (Pp.str "Not an unfoldable reference")
+ | Var id -> (EvalVar id, EInstance.empty)
+ | Rel n -> (EvalRel n, EInstance.empty)
+ | Evar ev -> (EvalEvar ev, EInstance.empty)
+ | _ -> anomaly (Pp.str "Not an unfoldable reference.")
let unsafe_reference_opt_value env sigma eval =
match eval with
| EvalConst cst ->
(match (lookup_constant cst env).Declarations.const_body with
- | Declarations.Def c -> Some (Mod_subst.force_constr c)
+ | Declarations.Def c -> Some (EConstr.of_constr (Mod_subst.force_constr c))
| _ -> None)
| EvalVar id ->
- let open Context.Named.Declaration in
- lookup_named id env |> get_value
+ env |> lookup_named id |> NamedDecl.get_value
| EvalRel n ->
- let open Context.Rel.Declaration in
- lookup_rel n env |> map_value (lift n) |> get_value
- | EvalEvar ev -> Evd.existential_opt_value sigma ev
+ env |> lookup_rel n |> RelDecl.get_value |> Option.map (lift n)
+ | EvalEvar ev ->
+ match EConstr.kind sigma (mkEvar ev) with
+ | Evar _ -> None
+ | c -> Some (EConstr.of_kind c)
let reference_opt_value env sigma eval u =
match eval with
- | EvalConst cst -> constant_opt_value_in env (cst,u)
+ | EvalConst cst ->
+ let u = EInstance.kind sigma u in
+ Option.map EConstr.of_constr (constant_opt_value_in env (cst,u))
| EvalVar id ->
- let open Context.Named.Declaration in
- lookup_named id env |> get_value
+ env |> lookup_named id |> NamedDecl.get_value
| EvalRel n ->
- let open Context.Rel.Declaration in
- lookup_rel n env |> map_value (lift n) |> get_value
- | EvalEvar ev -> Evd.existential_opt_value sigma ev
+ env |> lookup_rel n |> RelDecl.get_value |> Option.map (lift n)
+ | EvalEvar ev ->
+ match EConstr.kind sigma (mkEvar ev) with
+ | Evar _ -> None
+ | c -> Some (EConstr.of_kind c)
exception NotEvaluable
let reference_value env sigma c u =
@@ -179,18 +186,18 @@ let eval_table = Summary.ref (Cmap.empty : frozen) ~name:"evaluation"
the xp..x1.
*)
-let check_fix_reversibility labs args ((lv,i),(_,tys,bds)) =
+let check_fix_reversibility sigma labs args ((lv,i),(_,tys,bds)) =
let n = List.length labs in
let nargs = List.length args in
if nargs > n then raise Elimconst;
let nbfix = Array.length bds in
let li =
List.map
- (function d -> match kind_of_term d with
+ (function d -> match EConstr.kind sigma d with
| Rel k ->
if
- Array.for_all (noccurn k) tys
- && Array.for_all (noccurn (k+nbfix)) bds
+ Array.for_all (Vars.noccurn sigma k) tys
+ && Array.for_all (Vars.noccurn sigma (k+nbfix)) bds
&& k <= n
then
(k, List.nth labs (k-1))
@@ -204,7 +211,7 @@ let check_fix_reversibility labs args ((lv,i),(_,tys,bds)) =
raise Elimconst;
List.iteri (fun i t_i ->
if not (Int.List.mem_assoc (i+1) li) then
- let fvs = List.map ((+) (i+1)) (Int.Set.elements (free_rels t_i)) in
+ let fvs = List.map ((+) (i+1)) (Int.Set.elements (free_rels sigma t_i)) in
match List.intersect Int.equal fvs reversible_rels with
| [] -> ()
| _ -> raise Elimconst)
@@ -233,17 +240,18 @@ let invert_name labs l na0 env sigma ref = function
| EvalRel _ | EvalEvar _ -> None
| EvalVar id' -> Some (EvalVar id)
| EvalConst kn ->
- Some (EvalConst (con_with_label kn (Label.of_id id))) in
+ Some (EvalConst (Constant.change_label kn (Label.of_id id))) in
match refi with
| None -> None
| Some ref ->
try match unsafe_reference_opt_value env sigma ref with
| None -> None
| Some c ->
- let labs',ccl = decompose_lam c in
+ let labs',ccl = decompose_lam sigma c in
let _, l' = whd_betalet_stack sigma ccl in
let labs' = List.map snd labs' in
(** ppedrot: there used to be generic equality on terms here *)
+ let eq_constr c1 c2 = EConstr.eq_constr sigma c1 c2 in
if List.equal eq_constr labs' labs &&
List.equal eq_constr l l' then Some (minfxargs,ref)
else None
@@ -258,16 +266,16 @@ let invert_name labs l na0 env sigma ref = function
let compute_consteval_direct env sigma ref =
let rec srec env n labs onlyproj c =
let c',l = whd_betadeltazeta_stack env sigma c in
- match kind_of_term c' with
+ match EConstr.kind sigma c' with
| Lambda (id,t,g) when List.is_empty l && not onlyproj ->
let open Context.Rel.Declaration in
srec (push_rel (LocalAssum (id,t)) env) (n+1) (t::labs) onlyproj g
| Fix fix when not onlyproj ->
- (try check_fix_reversibility labs l fix
+ (try check_fix_reversibility sigma labs l fix
with Elimconst -> NotAnElimination)
- | Case (_,_,d,_) when isRel d && not onlyproj -> EliminationCases n
+ | Case (_,_,d,_) when isRel sigma d && not onlyproj -> EliminationCases n
| Case (_,_,d,_) -> srec env n labs true d
- | Proj (p, d) when isRel d -> EliminationProj n
+ | Proj (p, d) when isRel sigma d -> EliminationProj n
| _ -> NotAnElimination
in
match unsafe_reference_opt_value env sigma ref with
@@ -278,7 +286,7 @@ let compute_consteval_mutual_fix env sigma ref =
let rec srec env minarg labs ref c =
let c',l = whd_betalet_stack sigma c in
let nargs = List.length l in
- match kind_of_term c' with
+ match EConstr.kind sigma c' with
| Lambda (na,t,g) when List.is_empty l ->
let open Context.Rel.Declaration in
srec (push_rel (LocalAssum (na,t)) env) (minarg+1) (t::labs) ref g
@@ -294,11 +302,11 @@ let compute_consteval_mutual_fix env sigma ref =
let new_minarg = max (minarg'+minarg-nargs) minarg' in
EliminationMutualFix (new_minarg,ref,(refs,infos))
| _ -> assert false)
- | _ when isEvalRef env c' ->
+ | _ when isEvalRef env sigma c' ->
(* Forget all \'s and args and do as if we had started with c' *)
- let ref,_ = destEvalRefU c' in
+ let ref,_ = destEvalRefU sigma c' in
(match unsafe_reference_opt_value env sigma ref with
- | None -> anomaly (Pp.str "Should have been trapped by compute_direct")
+ | None -> anomaly (Pp.str "Should have been trapped by compute_direct.")
| Some c -> srec env (minarg-nargs) [] ref c)
| _ -> (* Should not occur *) NotAnElimination
in
@@ -355,17 +363,17 @@ let make_elim_fun (names,(nbfix,lv,n)) u largs =
(* k from the comment is q+1 *)
try mkRel (p+1-(List.index Int.equal (n-q) lyi))
with Not_found -> aq)
- 0 (List.map (lift p) lu)
+ 0 (List.map (Vars.lift p) lu)
in
fun i ->
match names.(i) with
| None -> None
| Some (minargs,ref) ->
- let body = applistc (mkEvalRef ref u) la in
+ let body = applist (mkEvalRef ref u, la) in
let g =
List.fold_left_i (fun q (* j = n+1-q *) c (ij,tij) ->
- let subst = List.map (lift (-q)) (List.firstn (n-ij) la) in
- let tij' = substl (List.rev subst) tij in
+ let subst = List.map (Vars.lift (-q)) (List.firstn (n-ij) la) in
+ let tij' = Vars.substl (List.rev subst) tij in
mkLambda (x,tij',c)) 1 body (List.rev lv)
in Some (minargs,g)
@@ -385,21 +393,20 @@ let substl_with_function subst sigma constr =
let evd = ref sigma in
let minargs = ref Evar.Map.empty in
let v = Array.of_list subst in
- let rec subst_total k c = match kind_of_term c with
+ let rec subst_total k c = match EConstr.kind sigma c with
| Rel i when k < i ->
if i <= k + Array.length v then
match v.(i-k-1) with
| (fx, Some (min, ref)) ->
- let sigma = Sigma.Unsafe.of_evar_map !evd in
- let Sigma (evk, sigma, _) = Evarutil.new_pure_evar venv sigma dummy in
- let sigma = Sigma.to_evar_map sigma in
+ let sigma = !evd in
+ let (sigma, evk) = Evarutil.new_pure_evar venv sigma dummy in
evd := sigma;
minargs := Evar.Map.add evk min !minargs;
- lift k (mkEvar (evk, [|fx;ref|]))
- | (fx, None) -> lift k fx
+ Vars.lift k (mkEvar (evk, [|fx;ref|]))
+ | (fx, None) -> Vars.lift k fx
else mkRel (i - Array.length v)
| _ ->
- map_constr_with_binders succ subst_total k c in
+ map_with_binders sigma succ subst_total k c in
let c = subst_total 0 constr in
(c, !evd, !minargs)
@@ -409,28 +416,28 @@ exception Partial
reduction is solved by the expanded fix term. *)
let solve_arity_problem env sigma fxminargs c =
let evm = ref sigma in
- let set_fix i = evm := Evd.define i (mkVar vfx) !evm in
+ let set_fix i = evm := Evd.define i (Constr.mkVar vfx) !evm in
let rec check strict c =
let c' = whd_betaiotazeta sigma c in
- let (h,rcargs) = decompose_app c' in
- match kind_of_term h with
+ let (h,rcargs) = decompose_app_vect sigma c' in
+ match EConstr.kind sigma h with
Evar(i,_) when Evar.Map.mem i fxminargs && not (Evd.is_defined !evm i) ->
let minargs = Evar.Map.find i fxminargs in
- if List.length rcargs < minargs then
+ if Array.length rcargs < minargs then
if strict then set_fix i
else raise Partial;
- List.iter (check strict) rcargs
- | (Var _|Const _) when isEvalRef env h ->
- (let ev, u = destEvalRefU h in
+ Array.iter (check strict) rcargs
+ | (Var _|Const _) when isEvalRef env sigma h ->
+ (let ev, u = destEvalRefU sigma h in
match reference_opt_value env sigma ev u with
| Some h' ->
let bak = !evm in
- (try List.iter (check false) rcargs
+ (try Array.iter (check false) rcargs
with Partial ->
evm := bak;
- check strict (applist(h',rcargs)))
- | None -> List.iter (check strict) rcargs)
- | _ -> iter_constr (check strict) c' in
+ check strict (mkApp(h',rcargs)))
+ | None -> Array.iter (check strict) rcargs)
+ | _ -> EConstr.iter sigma (check strict) c' in
check true c;
!evm
@@ -441,16 +448,18 @@ let substl_checking_arity env subst sigma c =
let sigma' = solve_arity_problem env sigma minargs body in
(* we propagate the constraints: solved problems are substituted;
the other ones are replaced by the function symbol *)
- let rec nf_fix c =
- match kind_of_term c with
- Evar(i,[|fx;f|] as ev) when Evar.Map.mem i minargs ->
- (match Evd.existential_opt_value sigma' ev with
- Some c' -> c'
- | None -> f)
- | _ -> map_constr nf_fix c in
+ let rec nf_fix c = match EConstr.kind sigma c with
+ | Evar (i,[|fx;f|]) when Evar.Map.mem i minargs ->
+ (** FIXME: find a less hackish way of doing this *)
+ begin match EConstr.kind sigma' c with
+ | Evar _ -> f
+ | c -> EConstr.of_kind c
+ end
+ | _ -> EConstr.map sigma nf_fix c
+ in
nf_fix body
-type fix_reduction_result = NotReducible | Reduced of (constr*constr list)
+type fix_reduction_result = NotReducible | Reduced of (constr * constr list)
let reduce_fix whdfun sigma fix stack =
match fix_recarg fix (Stack.append_app_list stack Stack.empty) with
@@ -458,8 +467,8 @@ let reduce_fix whdfun sigma fix stack =
| Some (recargnum,recarg) ->
let (recarg'hd,_ as recarg') = whdfun sigma recarg in
let stack' = List.assign stack recargnum (applist recarg') in
- (match kind_of_term recarg'hd with
- | Construct _ -> Reduced (contract_fix fix, stack')
+ (match EConstr.kind sigma recarg'hd with
+ | Construct _ -> Reduced (contract_fix sigma fix, stack')
| _ -> NotReducible)
let contract_fix_use_function env sigma f
@@ -467,20 +476,20 @@ let contract_fix_use_function env sigma f
let nbodies = Array.length recindices in
let make_Fi j = (mkFix((recindices,j),typedbodies), f j) in
let lbodies = List.init nbodies make_Fi in
- substl_checking_arity env (List.rev lbodies) sigma (nf_beta sigma bodies.(bodynum))
+ substl_checking_arity env (List.rev lbodies) sigma (nf_beta env sigma bodies.(bodynum))
let reduce_fix_use_function env sigma f whfun fix stack =
match fix_recarg fix (Stack.append_app_list stack Stack.empty) with
| None -> NotReducible
| Some (recargnum,recarg) ->
let (recarg'hd,_ as recarg') =
- if isRel recarg then
+ if EConstr.isRel sigma recarg then
(* The recarg cannot be a local def, no worry about the right env *)
(recarg, [])
else
whfun recarg in
let stack' = List.assign stack recargnum (applist recarg') in
- (match kind_of_term recarg'hd with
+ (match EConstr.kind sigma recarg'hd with
| Construct _ ->
Reduced (contract_fix_use_function env sigma f fix,stack')
| _ -> NotReducible)
@@ -491,16 +500,16 @@ let contract_cofix_use_function env sigma f
let make_Fi j = (mkCoFix(j,typedbodies), f j) in
let subbodies = List.init nbodies make_Fi in
substl_checking_arity env (List.rev subbodies)
- sigma (nf_beta sigma bodies.(bodynum))
+ sigma (nf_beta env sigma bodies.(bodynum))
let reduce_mind_case_use_function func env sigma mia =
- match kind_of_term mia.mconstr with
+ match EConstr.kind sigma mia.mconstr with
| Construct ((ind_sp,i),u) ->
let real_cargs = List.skipn mia.mci.ci_npar mia.mcargs in
applist (mia.mlf.(i-1), real_cargs)
| CoFix (bodynum,(names,_,_) as cofix) ->
let build_cofix_name =
- if isConst func then
+ if isConst sigma func then
let minargs = List.length mia.mcargs in
fun i ->
if Int.equal i bodynum then Some (minargs,func)
@@ -511,13 +520,13 @@ let reduce_mind_case_use_function func env sigma mia =
mutual inductive, try to reuse the global name if
the block was indeed initially built as a global
definition *)
- let kn = map_puniverses (fun x -> con_with_label x (Label.of_id id))
- (destConst func)
- in
- try match constant_opt_value_in env kn with
+ let (kn, u) = destConst sigma func in
+ let kn = Constant.change_label kn (Label.of_id id) in
+ let cst = (kn, EInstance.kind sigma u) in
+ try match constant_opt_value_in env cst with
| None -> None
(* TODO: check kn is correct *)
- | Some _ -> Some (minargs,mkConstU kn)
+ | Some _ -> Some (minargs,mkConstU (kn, u))
with Not_found -> None
else
fun _ -> None in
@@ -527,45 +536,56 @@ let reduce_mind_case_use_function func env sigma mia =
| _ -> assert false
-let match_eval_ref env constr =
- match kind_of_term constr with
- | Const (sp, u) when is_evaluable env (EvalConstRef sp) ->
- Some (EvalConst sp, u)
- | Var id when is_evaluable env (EvalVarRef id) -> Some (EvalVar id, Univ.Instance.empty)
- | Rel i -> Some (EvalRel i, Univ.Instance.empty)
- | Evar ev -> Some (EvalEvar ev, Univ.Instance.empty)
+let match_eval_ref env sigma constr stack =
+ match EConstr.kind sigma constr with
+ | Const (sp, u) ->
+ reduction_effect_hook env sigma (EConstr.to_constr sigma constr)
+ (lazy (EConstr.to_constr sigma (applist (constr,stack))));
+ if is_evaluable env (EvalConstRef sp) then Some (EvalConst sp, u) else None
+ | Var id when is_evaluable env (EvalVarRef id) -> Some (EvalVar id, EInstance.empty)
+ | Rel i -> Some (EvalRel i, EInstance.empty)
+ | Evar ev -> Some (EvalEvar ev, EInstance.empty)
| _ -> None
-let match_eval_ref_value env sigma constr =
- match kind_of_term constr with
- | Const (sp, u) when is_evaluable env (EvalConstRef sp) ->
- Some (constant_value_in env (sp, u))
+let match_eval_ref_value env sigma constr stack =
+ match EConstr.kind sigma constr with
+ | Const (sp, u) ->
+ reduction_effect_hook env sigma (EConstr.to_constr sigma constr)
+ (lazy (EConstr.to_constr sigma (applist (constr,stack))));
+ if is_evaluable env (EvalConstRef sp) then
+ let u = EInstance.kind sigma u in
+ Some (EConstr.of_constr (constant_value_in env (sp, u)))
+ else
+ None
+ | Proj (p, c) when not (Projection.unfolded p) ->
+ reduction_effect_hook env sigma (EConstr.to_constr sigma constr)
+ (lazy (EConstr.to_constr sigma (applist (constr,stack))));
+ if is_evaluable env (EvalConstRef (Projection.constant p)) then
+ Some (mkProj (Projection.unfold p, c))
+ else None
| Var id when is_evaluable env (EvalVarRef id) ->
- let open Context.Named.Declaration in
- lookup_named id env |> get_value
+ env |> lookup_named id |> NamedDecl.get_value
| Rel n ->
- let open Context.Rel.Declaration in
- lookup_rel n env |> map_value (lift n) |> get_value
- | Evar ev -> Evd.existential_opt_value sigma ev
+ env |> lookup_rel n |> RelDecl.get_value |> Option.map (lift n)
| _ -> None
let special_red_case env sigma whfun (ci, p, c, lf) =
let rec redrec s =
let (constr, cargs) = whfun s in
- match match_eval_ref env constr with
+ match match_eval_ref env sigma constr cargs with
| Some (ref, u) ->
(match reference_opt_value env sigma ref u with
| None -> raise Redelimination
| Some gvalue ->
- if reducible_mind_case gvalue then
+ if reducible_mind_case sigma gvalue then
reduce_mind_case_use_function constr env sigma
{mP=p; mconstr=gvalue; mcargs=cargs;
mci=ci; mlf=lf}
else
redrec (applist(gvalue, cargs)))
| None ->
- if reducible_mind_case constr then
- reduce_mind_case
+ if reducible_mind_case sigma constr then
+ reduce_mind_case sigma
{mP=p; mconstr=constr; mcargs=cargs;
mci=ci; mlf=lf}
else
@@ -578,7 +598,7 @@ let recargs = function
| EvalConst c -> ReductionBehaviour.get (ConstRef c)
let reduce_projection env sigma pb (recarg'hd,stack') stack =
- (match kind_of_term recarg'hd with
+ (match EConstr.kind sigma recarg'hd with
| Construct _ ->
let proj_narg =
pb.Declarations.proj_npars + pb.Declarations.proj_arg
@@ -587,11 +607,11 @@ let reduce_projection env sigma pb (recarg'hd,stack') stack =
let reduce_proj env sigma whfun whfun' c =
let rec redrec s =
- match kind_of_term s with
+ match EConstr.kind sigma s with
| Proj (proj, c) ->
let c' = try redrec c with Redelimination -> c in
let constr, cargs = whfun c' in
- (match kind_of_term constr with
+ (match EConstr.kind sigma constr with
| Construct _ ->
let proj_narg =
let pb = lookup_projection proj env in
@@ -608,7 +628,7 @@ let reduce_proj env sigma whfun whfun' c =
let whd_nothing_for_iota env sigma s =
let rec whrec (x, stack as s) =
- match kind_of_term x with
+ match EConstr.kind sigma x with
| Rel n ->
let open Context.Rel.Declaration in
(match lookup_rel n env with
@@ -619,27 +639,26 @@ let whd_nothing_for_iota env sigma s =
(match lookup_named id env with
| LocalDef (_,body,_) -> whrec (body, stack)
| _ -> s)
- | Evar ev ->
- (try whrec (Evd.existential_value sigma ev, stack)
- with Evd.NotInstantiatedEvar | Not_found -> s)
+ | Evar ev -> s
| Meta ev ->
- (try whrec (Evd.meta_value sigma ev, stack)
+ (try whrec (EConstr.of_constr (Evd.meta_value sigma ev), stack)
with Not_found -> s)
- | Const const when is_transparent_constant full_transparent_state (fst const) ->
- (match constant_opt_value_in env const with
- | Some body -> whrec (body, stack)
+ | Const (const, u) when is_transparent_constant full_transparent_state const ->
+ let u = EInstance.kind sigma u in
+ (match constant_opt_value_in env (const, u) with
+ | Some body -> whrec (EConstr.of_constr body, stack)
| None -> s)
- | LetIn (_,b,_,c) -> stacklam whrec [b] c stack
+ | LetIn (_,b,_,c) -> stacklam whrec [b] sigma c stack
| Cast (c,_,_) -> whrec (c, stack)
| App (f,cl) -> whrec (f, Stack.append_app cl stack)
| Lambda (na,t,c) ->
(match Stack.decomp stack with
- | Some (a,m) -> stacklam whrec [a] c m
+ | Some (a,m) -> stacklam whrec [a] sigma c m
| _ -> s)
| x -> s
in
- decompose_app (Stack.zip (whrec (s,Stack.empty)))
+ EConstr.decompose_app sigma (Stack.zip sigma (whrec (s,Stack.empty)))
(* [red_elim_const] contracts iota/fix/cofix redexes hidden behind
constants by keeping the name of the constants in the recursive calls;
@@ -664,7 +683,7 @@ let rec red_elim_const env sigma ref u largs =
let c = reference_value env sigma ref u in
let c', lrest = whd_nothing_for_iota env sigma (applist(c,largs)) in
let whfun = whd_simpl_stack env sigma in
- (special_red_case env sigma whfun (destCase c'), lrest), nocase
+ (special_red_case env sigma whfun (EConstr.destCase sigma c'), lrest), nocase
| EliminationProj n when nargs >= n ->
let c = reference_value env sigma ref u in
let c', lrest = whd_nothing_for_iota env sigma (applist(c,largs)) in
@@ -676,24 +695,24 @@ let rec red_elim_const env sigma ref u largs =
let d, lrest = whd_nothing_for_iota env sigma (applist(c,largs)) in
let f = make_elim_fun ([|Some (minfxargs,ref)|],infos) u largs in
let whfun = whd_construct_stack env sigma in
- (match reduce_fix_use_function env sigma f whfun (destFix d) lrest with
+ (match reduce_fix_use_function env sigma f whfun (destFix sigma d) lrest with
| NotReducible -> raise Redelimination
- | Reduced (c,rest) -> (nf_beta sigma c, rest), nocase)
+ | Reduced (c,rest) -> (nf_beta env sigma c, rest), nocase)
| EliminationMutualFix (min,refgoal,refinfos) when nargs >= min ->
let rec descend (ref,u) args =
let c = reference_value env sigma ref u in
- if evaluable_reference_eq ref refgoal then
+ if evaluable_reference_eq sigma ref refgoal then
(c,args)
else
let c', lrest = whd_betalet_stack sigma (applist(c,args)) in
- descend (destEvalRefU c') lrest in
+ descend (destEvalRefU sigma c') lrest in
let (_, midargs as s) = descend (ref,u) largs in
let d, lrest = whd_nothing_for_iota env sigma (applist s) in
let f = make_elim_fun refinfos u midargs in
let whfun = whd_construct_stack env sigma in
- (match reduce_fix_use_function env sigma f whfun (destFix d) lrest with
+ (match reduce_fix_use_function env sigma f whfun (destFix sigma d) lrest with
| NotReducible -> raise Redelimination
- | Reduced (c,rest) -> (nf_beta sigma c, rest), nocase)
+ | Reduced (c,rest) -> (nf_beta env sigma c, rest), nocase)
| NotAnElimination when unfold_nonelim ->
let c = reference_value env sigma ref u in
(whd_betaiotazeta sigma (applist (c, largs)), []), nocase
@@ -709,7 +728,7 @@ and reduce_params env sigma stack l =
else
let arg = List.nth stack i in
let rarg = whd_construct_stack env sigma arg in
- match kind_of_term (fst rarg) with
+ match EConstr.kind sigma (fst rarg) with
| Construct _ -> List.assign stack i (applist rarg)
| _ -> raise Redelimination)
stack l
@@ -720,13 +739,15 @@ and reduce_params env sigma stack l =
and whd_simpl_stack env sigma =
let rec redrec s =
- let (x, stack as s') = decompose_app s in
- match kind_of_term x with
+ let (x, stack) = decompose_app_vect sigma s in
+ let stack = Array.to_list stack in
+ let s' = (x, stack) in
+ match EConstr.kind sigma x with
| Lambda (na,t,c) ->
(match stack with
| [] -> s'
- | a :: rest -> redrec (beta_applist (x,stack)))
- | LetIn (n,b,t,c) -> redrec (applist (substl [b] c, stack))
+ | a :: rest -> redrec (beta_applist sigma (x, stack)))
+ | LetIn (n,b,t,c) -> redrec (applist (Vars.substl [b] c, stack))
| App (f,cl) -> redrec (applist(f, (Array.to_list cl)@stack))
| Cast (c,_,_) -> redrec (applist(c, stack))
| Case (ci,p,c,lf) ->
@@ -766,12 +787,12 @@ and whd_simpl_stack env sigma =
with Redelimination -> s')
| _ ->
- match match_eval_ref env x with
+ match match_eval_ref env sigma x stack with
| Some (ref, u) ->
(try
let sapp, nocase = red_elim_const env sigma ref u stack in
let hd, _ as s'' = redrec (applist(sapp)) in
- let rec is_case x = match kind_of_term x with
+ let rec is_case x = match EConstr.kind sigma x with
| Lambda (_,_, x) | LetIn (_,_,_, x) | Cast (x, _,_) -> is_case x
| App (hd, _) -> is_case hd
| Case _ -> true
@@ -787,8 +808,8 @@ and whd_simpl_stack env sigma =
and whd_construct_stack env sigma s =
let (constr, cargs as s') = whd_simpl_stack env sigma s in
- if reducible_mind_case constr then s'
- else match match_eval_ref env constr with
+ if reducible_mind_case sigma constr then s'
+ else match match_eval_ref env sigma constr cargs with
| Some (ref, u) ->
(match reference_opt_value env sigma ref u with
| None -> raise Redelimination
@@ -804,12 +825,12 @@ and whd_construct_stack env sigma s =
*)
let try_red_product env sigma c =
- let simpfun = clos_norm_flags betaiotazeta env sigma in
+ let simpfun c = clos_norm_flags betaiotazeta env sigma c in
let rec redrec env x =
let x = whd_betaiota sigma x in
- match kind_of_term x with
+ match EConstr.kind sigma x with
| App (f,l) ->
- (match kind_of_term f with
+ (match EConstr.kind sigma f with
| Fix fix ->
let stack = Stack.append_app l Stack.empty in
(match fix_recarg fix stack with
@@ -817,17 +838,17 @@ let try_red_product env sigma c =
| Some (recargnum,recarg) ->
let recarg' = redrec env recarg in
let stack' = Stack.assign stack recargnum recarg' in
- simpfun (Stack.zip (f,stack')))
- | _ -> simpfun (appvect (redrec env f, l)))
+ simpfun (Stack.zip sigma (f,stack')))
+ | _ -> simpfun (mkApp (redrec env f, l)))
| Cast (c,_,_) -> redrec env c
| Prod (x,a,b) ->
let open Context.Rel.Declaration in
- mkProd (x, a, redrec (push_rel (LocalAssum (x,a)) env) b)
- | LetIn (x,a,b,t) -> redrec env (subst1 a t)
+ mkProd (x, a, redrec (push_rel (LocalAssum (x, a)) env) b)
+ | LetIn (x,a,b,t) -> redrec env (Vars.subst1 a t)
| Case (ci,p,d,lf) -> simpfun (mkCase (ci,p,redrec env d,lf))
| Proj (p, c) ->
let c' =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Construct _ -> c
| _ -> redrec env c
in
@@ -836,7 +857,7 @@ let try_red_product env sigma c =
| Reduced s -> simpfun (applist s)
| NotReducible -> raise Redelimination)
| _ ->
- (match match_eval_ref env x with
+ (match match_eval_ref env sigma x [] with
| Some (ref, u) ->
(* TO DO: re-fold fixpoints after expansion *)
(* to get true one-step reductions *)
@@ -848,7 +869,7 @@ let try_red_product env sigma c =
let red_product env sigma c =
try try_red_product env sigma c
- with Redelimination -> error "No head constant to reduce."
+ with Redelimination -> user_err (str "No head constant to reduce.")
(*
(* This old version of hnf uses betadeltaiota instead of itself (resp
@@ -906,8 +927,8 @@ let whd_simpl_orelse_delta_but_fix_old env sigma c =
let whd_simpl_stack =
if Flags.profile then
- let key = Profile.declare_profile "whd_simpl_stack" in
- Profile.profile3 key whd_simpl_stack
+ let key = CProfile.declare_profile "whd_simpl_stack" in
+ CProfile.profile3 key whd_simpl_stack
else whd_simpl_stack
(* Same as [whd_simpl] but also reduces constants that do not hide a
@@ -917,13 +938,13 @@ let whd_simpl_stack =
let whd_simpl_orelse_delta_but_fix env sigma c =
let rec redrec s =
let (constr, stack as s') = whd_simpl_stack env sigma s in
- match match_eval_ref_value env sigma constr with
+ match match_eval_ref_value env sigma constr stack with
| Some c ->
- (match kind_of_term (strip_lam c) with
+ (match EConstr.kind sigma (snd (decompose_lam sigma c)) with
| CoFix _ | Fix _ -> s'
| Proj (p,t) when
- (match kind_of_term constr with
- | Const (c', _) -> eq_constant (Projection.constant p) c'
+ (match EConstr.kind sigma constr with
+ | Const (c', _) -> Constant.equal (Projection.constant p) c'
| _ -> false) ->
let pb = Environ.lookup_projection p env in
if List.length stack <= pb.Declarations.proj_npars then
@@ -948,9 +969,9 @@ let simpl env sigma c = strong whd_simpl env sigma c
(* Reduction at specific subterms *)
let matches_head env sigma c t =
- match kind_of_term t with
+ match EConstr.kind sigma t with
| App (f,_) -> Constr_matching.matches env sigma c f
- | Proj (p, _) -> Constr_matching.matches env sigma c (mkConst (Projection.constant p))
+ | Proj (p, _) -> Constr_matching.matches env sigma c (mkConstU (Projection.constant p, EInstance.empty))
| _ -> raise Constr_matching.PatternMatchingFailure
(** FIXME: Specific function to handle projections: it ignores what happens on the
@@ -958,26 +979,25 @@ let matches_head env sigma c t =
of the projection and its eta expanded form.
*)
let change_map_constr_with_binders_left_to_right g f (env, l as acc) sigma c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Proj (p, r) -> (* Treat specially for partial applications *)
let t = Retyping.expand_projection env sigma p r [] in
- let hdf, al = destApp t in
+ let hdf, al = destApp sigma t in
let a = al.(Array.length al - 1) in
let app = (mkApp (hdf, Array.sub al 0 (Array.length al - 1))) in
let app' = f acc app in
let a' = f acc a in
- (match kind_of_term app' with
+ (match EConstr.kind sigma app' with
| App (hdf', al') when hdf' == hdf ->
(* Still the same projection, we ignore the change in parameters *)
mkProj (p, a')
| _ -> mkApp (app', [| a' |]))
- | _ -> map_constr_with_binders_left_to_right g f acc c
+ | _ -> map_constr_with_binders_left_to_right sigma g f acc c
-let e_contextually byhead (occs,c) f = { e_redfun = begin fun env sigma t ->
+let e_contextually byhead (occs,c) f = begin fun env sigma t ->
let (nowhere_except_in,locs) = Locusops.convert_occs occs in
let maxocc = List.fold_right max locs 0 in
let pos = ref 1 in
- let sigma = Sigma.to_evar_map sigma in
(** FIXME: we do suspicious things with this evarmap *)
let evd = ref sigma in
let rec traverse nested (env,c as envc) t =
@@ -993,12 +1013,12 @@ let e_contextually byhead (occs,c) f = { e_redfun = begin fun env sigma t ->
incr pos;
if ok then begin
if Option.has_some nested then
- errorlabstrm "" (str "The subterm at occurrence " ++ int (Option.get nested) ++ str " overlaps with the subterm at occurrence " ++ int (!pos-1) ++ str ".");
+ user_err (str "The subterm at occurrence " ++ int (Option.get nested) ++ str " overlaps with the subterm at occurrence " ++ int (!pos-1) ++ str ".");
(* Skip inner occurrences for stable counting of occurrences *)
if locs != [] then
ignore (traverse_below (Some (!pos-1)) envc t);
- let Sigma (t, evm, _) = (f subst).e_redfun env (Sigma.Unsafe.of_evar_map !evd) t in
- (evd := Sigma.to_evar_map evm; t)
+ let (evm, t) = (f subst) env !evd t in
+ (evd := evm; t)
end
else
traverse_below nested envc t
@@ -1007,7 +1027,7 @@ let e_contextually byhead (occs,c) f = { e_redfun = begin fun env sigma t ->
and traverse_below nested envc t =
(* when byhead, find other occurrences without matching again partial
application with same head *)
- match kind_of_term t with
+ match EConstr.kind !evd t with
| App (f,l) when byhead -> mkApp (f, Array.map_left (traverse nested envc) l)
| Proj (p,c) when byhead -> mkProj (p,traverse nested envc c)
| _ ->
@@ -1017,24 +1037,21 @@ let e_contextually byhead (occs,c) f = { e_redfun = begin fun env sigma t ->
in
let t' = traverse None (env,c) t in
if List.exists (fun o -> o >= !pos) locs then error_invalid_occurrence locs;
- Sigma.Unsafe.of_pair (t', !evd)
- end }
+ (!evd, t')
+ end
let contextually byhead occs f env sigma t =
- let f' subst = { e_redfun = begin fun env sigma t ->
- Sigma.here (f subst env (Sigma.to_evar_map sigma) t) sigma
- end } in
- let Sigma (c, _, _) = (e_contextually byhead occs f').e_redfun env (Sigma.Unsafe.of_evar_map sigma) t in
- c
+ let f' subst env sigma t = sigma, f subst env sigma t in
+ snd (e_contextually byhead occs f' env sigma t)
(* linear bindings (following pretty-printer) of the value of name in c.
* n is the number of the next occurrence of name.
* ol is the occurrence list to find. *)
let match_constr_evaluable_ref sigma c evref =
- match kind_of_term c, evref with
- | Const (c,u), EvalConstRef c' when eq_constant c c' -> Some u
- | Var id, EvalVarRef id' when id_eq id id' -> Some Univ.Instance.empty
+ match EConstr.kind sigma c, evref with
+ | Const (c,u), EvalConstRef c' when Constant.equal c c' -> Some u
+ | Var id, EvalVarRef id' when Id.equal id id' -> Some EInstance.empty
| _, _ -> None
let substlin env sigma evalref n (nowhere_except_in,locs) c =
@@ -1053,7 +1070,7 @@ let substlin env sigma evalref n (nowhere_except_in,locs) c =
incr pos;
if ok then value u else c
| None ->
- map_constr_with_binders_left_to_right
+ map_constr_with_binders_left_to_right sigma
(fun _ () -> ())
substrec () c
in
@@ -1066,11 +1083,11 @@ let string_of_evaluable_ref env = function
string_of_qualid
(Nametab.shortest_qualid_of_global (vars_of_env env) (ConstRef kn))
-let unfold env sigma name =
+let unfold env sigma name c =
if is_evaluable env name then
- clos_norm_flags (unfold_red name) env sigma
+ clos_norm_flags (unfold_red name) env sigma c
else
- error (string_of_evaluable_ref env name^" is opaque.")
+ user_err Pp.(str (string_of_evaluable_ref env name^" is opaque."))
(* [unfoldoccs : (readable_constraints -> (int list * full_path) -> constr -> constr)]
* Unfolds the constant name in a term c following a list of occurrences occl.
@@ -1080,13 +1097,13 @@ let unfoldoccs env sigma (occs,name) c =
let unfo nowhere_except_in locs =
let (nbocc,uc) = substlin env sigma name 1 (nowhere_except_in,locs) c in
if Int.equal nbocc 1 then
- error ((string_of_evaluable_ref env name)^" does not occur.");
+ user_err Pp.(str ((string_of_evaluable_ref env name)^" does not occur."));
let rest = List.filter (fun o -> o >= nbocc) locs in
let () = match rest with
| [] -> ()
| _ -> error_invalid_occurrence rest
in
- nf_betaiotazeta sigma uc
+ nf_betaiotazeta env sigma uc
in
match occs with
| NoOccurrences -> c
@@ -1102,21 +1119,21 @@ let unfoldn loccname env sigma c =
let fold_one_com com env sigma c =
let rcom =
try red_product env sigma com
- with Redelimination -> error "Not reducible." in
+ with Redelimination -> user_err Pp.(str "Not reducible.") in
(* Reason first on the beta-iota-zeta normal form of the constant as
unfold produces it, so that the "unfold f; fold f" configuration works
to refold fix expressions *)
- let a = subst_term (clos_norm_flags unfold_side_red env sigma rcom) c in
- if not (eq_constr a c) then
- subst1 com a
+ let a = subst_term sigma (clos_norm_flags unfold_side_red env sigma rcom) c in
+ if not (EConstr.eq_constr sigma a c) then
+ Vars.subst1 com a
else
(* Then reason on the non beta-iota-zeta form for compatibility -
even if it is probably a useless configuration *)
- let a = subst_term rcom c in
- subst1 com a
+ let a = subst_term sigma rcom c in
+ Vars.subst1 com a
let fold_commands cl env sigma c =
- List.fold_right (fun com -> fold_one_com com env sigma) (List.rev cl) c
+ List.fold_right (fun com c -> fold_one_com com env sigma c) (List.rev cl) c
(* call by value reduction functions *)
@@ -1134,38 +1151,37 @@ let compute = cbv_betadeltaiota
(* gives [na:ta]c' such that c converts to ([na:ta]c' a), abstracting only
* the specified occurrences. *)
-let abstract_scheme env (locc,a) (c, sigma) =
+let abstract_scheme env sigma (locc,a) (c, sigma) =
let ta = Retyping.get_type_of env sigma a in
- let na = named_hd env ta Anonymous in
- if occur_meta ta then error "Cannot find a type for the generalisation.";
- if occur_meta a then
+ let na = named_hd env sigma ta Anonymous in
+ if occur_meta sigma ta then user_err Pp.(str "Cannot find a type for the generalisation.");
+ if occur_meta sigma a then
mkLambda (na,ta,c), sigma
else
let c', sigma' = subst_closed_term_occ env sigma (AtOccs locc) a c in
mkLambda (na,ta,c'), sigma'
-let pattern_occs loccs_trm = { e_redfun = begin fun env sigma c ->
- let sigma = Sigma.to_evar_map sigma in
- let abstr_trm, sigma = List.fold_right (abstract_scheme env) loccs_trm (c,sigma) in
+let pattern_occs loccs_trm = begin fun env sigma c ->
+ let abstr_trm, sigma = List.fold_right (abstract_scheme env sigma) loccs_trm (c,sigma) in
try
let _ = Typing.unsafe_type_of env sigma abstr_trm in
- Sigma.Unsafe.of_pair (applist(abstr_trm, List.map snd loccs_trm), sigma)
+ (sigma, applist(abstr_trm, List.map snd loccs_trm))
with Type_errors.TypeError (env',t) ->
raise (ReductionTacticError (InvalidAbstraction (env,sigma,abstr_trm,(env',t))))
- end }
+ end
(* Used in several tactics. *)
let check_privacy env ind =
let spec = Inductive.lookup_mind_specif env (fst ind) in
if Inductive.is_private spec then
- errorlabstrm "" (str "case analysis on a private type.")
+ user_err (str "case analysis on a private type.")
else ind
let check_not_primitive_record env ind =
let spec = Inductive.lookup_mind_specif env (fst ind) in
if Inductive.is_primitive_record spec then
- errorlabstrm "" (str "case analysis on a primitive record type: " ++
+ user_err (str "case analysis on a primitive record type: " ++
str "use projections or let instead.")
else ind
@@ -1175,30 +1191,30 @@ let check_not_primitive_record env ind =
let reduce_to_ind_gen allow_product env sigma t =
let rec elimrec env t l =
let t = hnf_constr env sigma t in
- match kind_of_term (fst (decompose_app t)) with
+ match EConstr.kind sigma (fst (decompose_app_vect sigma t)) with
| Ind ind-> (check_privacy env ind, it_mkProd_or_LetIn t l)
| Prod (n,ty,t') ->
let open Context.Rel.Declaration in
if allow_product then
elimrec (push_rel (LocalAssum (n,ty)) env) t' ((LocalAssum (n,ty))::l)
else
- errorlabstrm "" (str"Not an inductive definition.")
+ user_err (str"Not an inductive definition.")
| _ ->
(* Last chance: we allow to bypass the Opaque flag (as it
was partially the case between V5.10 and V8.1 *)
let t' = whd_all env sigma t in
- match kind_of_term (fst (decompose_app t')) with
+ match EConstr.kind sigma (fst (decompose_app_vect sigma t')) with
| Ind ind-> (check_privacy env ind, it_mkProd_or_LetIn t' l)
- | _ -> errorlabstrm "" (str"Not an inductive product.")
+ | _ -> user_err (str"Not an inductive product.")
in
elimrec env t []
-let reduce_to_quantified_ind x = reduce_to_ind_gen true x
-let reduce_to_atomic_ind x = reduce_to_ind_gen false x
+let reduce_to_quantified_ind env sigma c = reduce_to_ind_gen true env sigma c
+let reduce_to_atomic_ind env sigma c = reduce_to_ind_gen false env sigma c
let find_hnf_rectype env sigma t =
let ind,t = reduce_to_atomic_ind env sigma t in
- ind, snd (decompose_app t)
+ ind, snd (decompose_app sigma t)
(* Reduce the weak-head redex [beta,iota/fix/cofix[all],cast,zeta,simpl/delta]
or raise [NotStepReducible] if not a weak-head redex *)
@@ -1207,13 +1223,13 @@ exception NotStepReducible
let one_step_reduce env sigma c =
let rec redrec (x, stack) =
- match kind_of_term x with
+ match EConstr.kind sigma x with
| Lambda (n,t,c) ->
(match stack with
| [] -> raise NotStepReducible
- | a :: rest -> (subst1 a c, rest))
+ | a :: rest -> (Vars.subst1 a c, rest))
| App (f,cl) -> redrec (f, (Array.to_list cl)@stack)
- | LetIn (_,f,_,cl) -> (subst1 f cl,stack)
+ | LetIn (_,f,_,cl) -> (Vars.subst1 f cl,stack)
| Cast (c,_,_) -> redrec (c,stack)
| Case (ci,p,c,lf) ->
(try
@@ -1225,8 +1241,8 @@ let one_step_reduce env sigma c =
| Reduced s' -> s'
| NotReducible -> raise NotStepReducible
with Redelimination -> raise NotStepReducible)
- | _ when isEvalRef env x ->
- let ref,u = destEvalRefU x in
+ | _ when isEvalRef env sigma x ->
+ let ref,u = destEvalRefU sigma x in
(try
fst (red_elim_const env sigma ref u stack)
with Redelimination ->
@@ -1239,7 +1255,7 @@ let one_step_reduce env sigma c =
applist (redrec (c,[]))
let error_cannot_recognize ref =
- errorlabstrm ""
+ user_err
(str "Cannot recognize a statement based on " ++
Nametab.pr_global_env Id.Set.empty ref ++ str".")
@@ -1253,8 +1269,8 @@ let reduce_to_ref_gen allow_product env sigma ref t =
else
(* lazily reduces to match the head of [t] with the expected [ref] *)
let rec elimrec env t l =
- let c, _ = decompose_appvect (Reductionops.whd_nored sigma t) in
- match kind_of_term c with
+ let c, _ = decompose_app_vect sigma t in
+ match EConstr.kind sigma c with
| Prod (n,ty,t') ->
if allow_product then
let open Context.Rel.Declaration in
@@ -1263,12 +1279,12 @@ let reduce_to_ref_gen allow_product env sigma ref t =
error_cannot_recognize ref
| _ ->
try
- if eq_gr (global_of_constr c) ref
+ if eq_gr (fst (global_of_constr sigma c)) ref
then it_mkProd_or_LetIn t l
else raise Not_found
with Not_found ->
try
- let t' = nf_betaiota sigma (one_step_reduce env sigma t) in
+ let t' = nf_betaiota env sigma (one_step_reduce env sigma t) in
elimrec env t' l
with NotStepReducible -> error_cannot_recognize ref
in
diff --git a/pretyping/tacred.mli b/pretyping/tacred.mli
index f8dfe1ad..aa7604f5 100644
--- a/pretyping/tacred.mli
+++ b/pretyping/tacred.mli
@@ -1,20 +1,23 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
open Environ
open Evd
+open EConstr
open Reductionops
open Pattern
open Globnames
open Locus
open Univ
+open Ltac_pretype
type reduction_tactic_error =
InvalidAbstraction of env * evar_map * constr * (env * Type_errors.type_error)
@@ -75,12 +78,12 @@ val cbv_norm_flags : CClosure.RedFlags.reds -> reduction_function
(** [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 -> pinductive * types
+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 -> pinductive * types
+val reduce_to_quantified_ind : env -> evar_map -> types -> (inductive * EInstance.t) * types
(** [reduce_to_quantified_ref env sigma ref t] try to put [t] in the form
[t'=(x1:A1)..(xn:An)(ref args)] and fails with user error if not possible *)
@@ -91,7 +94,7 @@ val reduce_to_atomic_ref :
env -> evar_map -> global_reference -> types -> types
val find_hnf_rectype :
- env -> evar_map -> types -> pinductive * constr list
+ env -> evar_map -> types -> (inductive * EInstance.t) * constr list
val contextually : bool -> occurrences * constr_pattern ->
(patvar_map -> reduction_function) -> reduction_function
diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml
index b8da6b68..08051fd3 100644
--- a/pretyping/typeclasses.ml
+++ b/pretyping/typeclasses.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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*)
@@ -11,12 +13,16 @@ open Names
open Globnames
open Decl_kinds
open Term
+open Constr
open Vars
open Evd
open Util
open Typeclasses_errors
open Libobject
open Context.Rel.Declaration
+
+module RelDecl = Context.Rel.Declaration
+module NamedDecl = Context.Named.Declaration
(*i*)
let typeclasses_unique_solutions = ref false
@@ -27,8 +33,7 @@ open Goptions
let _ =
declare_bool_option
- { optsync = true;
- optdepr = false;
+ { optdepr = false;
optname = "check that typeclasses proof search returns unique solutions";
optkey = ["Typeclasses";"Unique";"Solutions"];
optread = get_typeclasses_unique_solutions;
@@ -55,18 +60,21 @@ type direction = Forward | Backward
(* This module defines type-classes *)
type typeclass = {
+ (* Universe quantification *)
+ cl_univs : Univ.AUContext.t;
+
(* The class implementation *)
cl_impl : global_reference;
(* Context in which the definitions are typed. Includes both typeclass parameters and superclasses. *)
- cl_context : (global_reference * bool) option list * Context.Rel.t;
+ cl_context : global_reference option list * Context.Rel.t;
(* Context of definitions and properties on defs, will not be shared *)
cl_props : Context.Rel.t;
(* The method implementaions as projections. *)
cl_projs : (Name.t * (direction * Vernacexpr.hint_info_expr) option
- * constant option) list;
+ * Constant.t option) list;
cl_strict : bool;
@@ -79,9 +87,8 @@ type instance = {
is_class: global_reference;
is_info: Vernacexpr.hint_info_expr;
(* Sections where the instance should be redeclared,
- -1 for discard, 0 for none. *)
- is_global: int;
- is_poly: bool;
+ None for discard, Some 0 for none. *)
+ is_global: int option;
is_impl: global_reference;
}
@@ -91,15 +98,16 @@ let instance_impl is = is.is_impl
let hint_priority is = is.is_info.Vernacexpr.hint_priority
-let new_instance cl info glob poly impl =
+let new_instance cl info glob impl =
let global =
- if glob then Lib.sections_depth ()
- else -1
+ if glob then Some (Lib.sections_depth ())
+ else None
in
+ if match global with Some n -> n>0 && isVarRef impl | _ -> false then
+ CErrors.user_err (Pp.str "Cannot set Global an instance referring to a section variable.");
{ is_class = cl.cl_impl;
is_info = info ;
is_global = global ;
- is_poly = poly;
is_impl = impl }
(*
@@ -109,64 +117,48 @@ let new_instance cl info glob poly impl =
let classes : typeclasses ref = Summary.ref Refmap.empty ~name:"classes"
let instances : instances ref = Summary.ref Refmap.empty ~name:"instances"
-open Declarations
-
-let typeclass_univ_instance (cl,u') =
- let subst =
- let u =
- match cl.cl_impl with
- | ConstRef c ->
- let cb = Global.lookup_constant c in
- if cb.const_polymorphic then Univ.UContext.instance cb.const_universes
- else Univ.Instance.empty
- | IndRef c ->
- let mib,oib = Global.lookup_inductive c in
- if mib.mind_polymorphic then Univ.UContext.instance mib.mind_universes
- else Univ.Instance.empty
- | _ -> Univ.Instance.empty
- in Array.fold_left2 (fun subst u u' -> Univ.LMap.add u u' subst)
- Univ.LMap.empty (Univ.Instance.to_array u) (Univ.Instance.to_array u')
- in
- let subst_ctx = Context.Rel.map (subst_univs_level_constr subst) in
- { cl with cl_context = fst cl.cl_context, subst_ctx (snd cl.cl_context);
- cl_props = subst_ctx cl.cl_props}, u'
+let typeclass_univ_instance (cl, u) =
+ assert (Univ.AUContext.size cl.cl_univs == Univ.Instance.length u);
+ let subst_ctx c = Context.Rel.map (subst_instance_constr u) c in
+ { cl with cl_context = on_snd subst_ctx cl.cl_context;
+ cl_props = subst_ctx cl.cl_props}
let class_info c =
try Refmap.find c !classes
- with Not_found -> not_a_class (Global.env()) (printable_constr_of_global c)
+ with Not_found -> not_a_class (Global.env()) (EConstr.of_constr (printable_constr_of_global c))
-let global_class_of_constr env c =
- try let gr, u = Universes.global_of_constr c in
+let global_class_of_constr env sigma c =
+ try let gr, u = Termops.global_of_constr sigma c in
class_info gr, u
with Not_found -> not_a_class env c
-let dest_class_app env c =
- let cl, args = decompose_app c in
- global_class_of_constr env cl, args
+let dest_class_app env sigma c =
+ let cl, args = EConstr.decompose_app sigma c in
+ global_class_of_constr env sigma cl, (List.map EConstr.Unsafe.to_constr args)
-let dest_class_arity env c =
- let rels, c = decompose_prod_assum c in
- rels, dest_class_app env c
+let dest_class_arity env sigma c =
+ let open EConstr in
+ let rels, c = decompose_prod_assum sigma c in
+ rels, dest_class_app env sigma c
-let class_of_constr c =
- try Some (dest_class_arity (Global.env ()) c)
+let class_of_constr sigma c =
+ try Some (dest_class_arity (Global.env ()) sigma c)
with e when CErrors.noncritical e -> None
-let is_class_constr c =
- try let gr, u = Universes.global_of_constr c in
+let is_class_constr sigma c =
+ try let gr, u = Termops.global_of_constr sigma c in
Refmap.mem gr !classes
with Not_found -> false
let rec is_class_type evd c =
- let c, args = decompose_app c in
- match kind_of_term c with
+ let c, _ = Termops.decompose_app_vect evd c in
+ match EConstr.kind evd c with
| Prod (_, _, t) -> is_class_type evd t
- | Evar (e, _) when Evd.is_defined evd e ->
- is_class_type evd (Evarutil.whd_head_evar evd c)
- | _ -> is_class_constr c
+ | Cast (t, _, _) -> is_class_type evd t
+ | _ -> is_class_constr evd c
let is_class_evar evd evi =
- is_class_type evd evi.Evd.evar_concl
+ is_class_type evd (EConstr.of_constr evi.Evd.evar_concl)
(*
* classes persistent object
@@ -181,13 +173,14 @@ let subst_class (subst,cl) =
let do_subst_con c = Mod_subst.subst_constant subst c
and do_subst c = Mod_subst.subst_mps subst c
and do_subst_gr gr = fst (subst_global subst gr) in
- let do_subst_ctx = List.smartmap (map_constr do_subst) in
+ let do_subst_ctx = List.smartmap (RelDecl.map_constr do_subst) in
let do_subst_context (grs,ctx) =
- List.smartmap (Option.smartmap (fun (gr,b) -> do_subst_gr gr, b)) grs,
+ List.smartmap (Option.smartmap do_subst_gr) grs,
do_subst_ctx ctx in
let do_subst_projs projs = List.smartmap (fun (x, y, z) ->
(x, y, Option.smartmap do_subst_con z)) projs in
- { cl_impl = do_subst_gr cl.cl_impl;
+ { cl_univs = cl.cl_univs;
+ cl_impl = do_subst_gr cl.cl_impl;
cl_context = do_subst_context cl.cl_context;
cl_props = do_subst_ctx cl.cl_props;
cl_projs = do_subst_projs cl.cl_projs;
@@ -197,22 +190,18 @@ let subst_class (subst,cl) =
let discharge_class (_,cl) =
let repl = Lib.replacement_context () in
let rel_of_variable_context ctx = List.fold_right
- ( fun (n,_,b,t) (ctx', subst) ->
- let decl = match b with
- | None -> LocalAssum (Name n, substn_vars 1 subst t)
- | Some b -> LocalDef (Name n, substn_vars 1 subst b, substn_vars 1 subst t)
- in
- (decl :: ctx', n :: subst)
+ ( fun (decl,_) (ctx', subst) ->
+ let decl' = decl |> NamedDecl.map_constr (substn_vars 1 subst) |> NamedDecl.to_rel_decl in
+ (decl' :: ctx', NamedDecl.get_id decl :: subst)
) ctx ([], []) in
- let discharge_rel_context subst n rel =
+ let discharge_rel_context (subst, usubst) n rel =
let rel = Context.Rel.map (Cooking.expmod_constr repl) rel in
- let ctx, _ =
- List.fold_right
- (fun decl (ctx, k) ->
- map_constr (substn_vars k subst) decl :: ctx, succ k
- )
- rel ([], n)
- in ctx
+ let fold decl (ctx, k) =
+ let map c = subst_univs_level_constr usubst (substn_vars k subst c) in
+ RelDecl.map_constr map decl :: ctx, succ k
+ in
+ let ctx, _ = List.fold_right fold rel ([], n) in
+ ctx
in
let abs_context cl =
match cl.cl_impl with
@@ -222,22 +211,25 @@ let discharge_class (_,cl) =
let discharge_context ctx' subst (grs, ctx) =
let grs' =
let newgrs = List.map (fun decl ->
- match decl |> get_type |> class_of_constr with
+ match decl |> RelDecl.get_type |> EConstr.of_constr |> class_of_constr Evd.empty with
| None -> None
- | Some (_, ((tc,_), _)) -> Some (tc.cl_impl, true))
+ | Some (_, ((tc,_), _)) -> Some tc.cl_impl)
ctx'
in
- List.smartmap (Option.smartmap (fun (gr, b) -> Lib.discharge_global gr, b)) grs
+ List.smartmap (Option.smartmap Lib.discharge_global) grs
@ newgrs
in grs', discharge_rel_context subst 1 ctx @ ctx' in
let cl_impl' = Lib.discharge_global cl.cl_impl in
if cl_impl' == cl.cl_impl then cl else
- let ctx, usubst, uctx = abs_context cl in
+ let info = abs_context cl in
+ let ctx = info.Lib.abstr_ctx in
let ctx, subst = rel_of_variable_context ctx in
- let context = discharge_context ctx subst cl.cl_context in
- let props = discharge_rel_context subst (succ (List.length (fst cl.cl_context))) cl.cl_props in
+ let usubst, cl_univs' = Lib.discharge_abstract_universe_context info cl.cl_univs in
+ let context = discharge_context ctx (subst, usubst) cl.cl_context in
+ let props = discharge_rel_context (subst, usubst) (succ (List.length (fst cl.cl_context))) cl.cl_props in
let discharge_proj (x, y, z) = x, y, Option.smartmap Lib.discharge_con z in
- { cl_impl = cl_impl';
+ { cl_univs = cl_univs';
+ cl_impl = cl_impl';
cl_context = context;
cl_props = props;
cl_projs = List.smartmap discharge_proj cl.cl_projs;
@@ -284,15 +276,18 @@ let build_subclasses ~check env sigma glob { hint_priority = pri } =
Nameops.add_suffix _id ("_subinstance_" ^ string_of_int !i))
in
let ty, ctx = Global.type_of_global_in_context env glob in
- let sigma = Evd.merge_context_set Evd.univ_rigid sigma (Univ.ContextSet.of_context ctx) in
+ let inst, ctx = Universes.fresh_instance_from ctx None in
+ let ty = Vars.subst_instance_constr inst ty in
+ let ty = EConstr.of_constr ty in
+ let sigma = Evd.merge_context_set Evd.univ_rigid sigma ctx in
let rec aux pri c ty path =
- let ty = Evarutil.nf_evar sigma ty in
- match class_of_constr ty with
+ match class_of_constr sigma ty with
| None -> []
| Some (rels, ((tc,u), args)) ->
let instapp =
- Reductionops.whd_beta sigma (appvectc c (Context.Rel.to_extended_vect 0 rels))
+ Reductionops.whd_beta sigma (EConstr.of_constr (appvectc c (Context.Rel.to_extended_vect mkRel 0 rels)))
in
+ let instapp = EConstr.Unsafe.to_constr instapp in
let projargs = Array.of_list (args @ [instapp]) in
let projs = List.map_filter
(fun (n, b, proj) ->
@@ -301,8 +296,10 @@ let build_subclasses ~check env sigma glob { hint_priority = pri } =
| Some (Backward, _) -> None
| Some (Forward, info) ->
let proj = Option.get proj in
+ let rels = List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) rels in
+ let u = EConstr.EInstance.kind sigma u in
let body = it_mkLambda_or_LetIn (mkApp (mkConstU (proj,u), projargs)) rels in
- if check && check_instance env sigma body then None
+ if check && check_instance env sigma (EConstr.of_constr body) then None
else
let newpri =
match pri, info.hint_priority with
@@ -314,12 +311,12 @@ let build_subclasses ~check env sigma glob { hint_priority = pri } =
in
let declare_proj hints (cref, info, body) =
let path' = cref :: path in
- let ty = Retyping.get_type_of env sigma body in
+ let ty = Retyping.get_type_of env sigma (EConstr.of_constr body) in
let rest = aux pri body ty path' in
hints @ (path', info, body) :: rest
in List.fold_left declare_proj [] projs
in
- let term = Universes.constr_of_global_univ (glob,Univ.UContext.instance ctx) in
+ let term = Universes.constr_of_global_univ (glob, inst) in
(*FIXME subclasses should now get substituted for each particular instance of
the polymorphic superclass *)
aux pri term ty [glob]
@@ -357,22 +354,34 @@ let subst_instance (subst, (action, inst)) = action,
is_impl = fst (subst_global subst inst.is_impl) }
let discharge_instance (_, (action, inst)) =
- if inst.is_global <= 0 then None
- else Some (action,
+ match inst.is_global with
+ | None -> None
+ | Some n ->
+ assert (not (isVarRef inst.is_impl));
+ Some (action,
{ inst with
- is_global = pred inst.is_global;
+ is_global = Some (pred n);
is_class = Lib.discharge_global inst.is_class;
is_impl = Lib.discharge_global inst.is_impl })
-let is_local i = Int.equal i.is_global (-1)
+let is_local i = (i.is_global == None)
+
+let is_local_for_hint i =
+ match i.is_global with
+ | None -> true (* i.e. either no Global keyword not in section, or in section *)
+ | Some n -> n <> 0 (* i.e. in a section, declare the hint as local
+ since discharge is managed by rebuild_instance which calls again
+ add_instance_hint; don't ask hints to take discharge into account
+ itself *)
let add_instance check inst =
let poly = Global.is_polymorphic inst.is_impl in
- add_instance_hint (IsGlobal inst.is_impl) [inst.is_impl] (is_local inst)
+ let local = is_local_for_hint inst in
+ add_instance_hint (IsGlobal inst.is_impl) [inst.is_impl] local
inst.is_info poly;
List.iter (fun (path, pri, c) -> add_instance_hint (IsConstr c) path
- (is_local inst) pri poly)
+ local pri poly)
(build_subclasses ~check:(check && not (isVarRef inst.is_impl))
(Global.env ()) (Evd.from_env (Global.env ())) inst.is_impl inst.is_info)
@@ -407,11 +416,12 @@ let remove_instance i =
remove_instance_hint i.is_impl
let declare_instance info local glob =
- let ty = Global.type_of_global_unsafe glob in
+ let ty, _ = Global.type_of_global_in_context (Global.env ()) glob in
let info = Option.default {hint_priority = None; hint_pattern = None} info in
- match class_of_constr ty with
+ match class_of_constr Evd.empty (EConstr.of_constr ty) with
| Some (rels, ((tc,_), args) as _cl) ->
- add_instance (new_instance tc info (not local) (Flags.use_polymorphic_flag ()) glob)
+ assert (not (isVarRef glob) || local);
+ add_instance (new_instance tc info (not local) glob)
| None -> ()
let add_class cl =
@@ -420,7 +430,7 @@ let add_class cl =
match inst with
| Some (Backward, info) ->
(match body with
- | None -> CErrors.error "Non-definable projection can not be declared as a subinstance"
+ | None -> CErrors.user_err Pp.(str "Non-definable projection can not be declared as a subinstance")
| Some b -> declare_instance (Some info) false (ConstRef b))
| _ -> ())
cl.cl_projs
@@ -433,19 +443,20 @@ let add_class cl =
let instance_constructor (cl,u) args =
let lenpars = List.count is_local_assum (snd cl.cl_context) in
+ let open EConstr in
let pars = fst (List.chop lenpars args) in
match cl.cl_impl with
| IndRef ind ->
let ind = ind, u in
- (Some (applistc (mkConstructUi (ind, 1)) args),
- applistc (mkIndU ind) pars)
+ (Some (applist (mkConstructUi (ind, 1), args)),
+ applist (mkIndU ind, pars))
| ConstRef cst ->
let cst = cst, u in
let term = match args with
| [] -> None
| _ -> Some (List.last args)
in
- (term, applistc (mkConstU cst) pars)
+ (term, applist (mkConstU cst, pars))
| _ -> assert false
let typeclasses () = Refmap.fold (fun _ l c -> l :: c) !classes []
@@ -512,7 +523,7 @@ let mark_unresolvable evi = mark_resolvability false evi
let mark_resolvable evi = mark_resolvability true evi
open Evar_kinds
-type evar_filter = existential_key -> Evar_kinds.t -> bool
+type evar_filter = Evar.t -> Evar_kinds.t -> bool
let all_evars _ _ = true
let all_goals _ = function VarInstance _ | GoalEvar -> true | _ -> false
@@ -543,8 +554,8 @@ let solve_all_instances env evd filter unique split fail =
Hook.get get_solve_all_instances env evd filter unique split fail
(** Profiling resolution of typeclasses *)
-(* let solve_classeskey = Profile.declare_profile "solve_typeclasses" *)
-(* let solve_problem = Profile.profile5 solve_classeskey solve_problem *)
+(* let solve_classeskey = CProfile.declare_profile "solve_typeclasses" *)
+(* let solve_problem = CProfile.profile5 solve_classeskey solve_problem *)
let resolve_typeclasses ?(fast_path = true) ?(filter=no_goals) ?(unique=get_typeclasses_unique_solutions ())
?(split=true) ?(fail=true) env evd =
diff --git a/pretyping/typeclasses.mli b/pretyping/typeclasses.mli
index 620bc367..b80c2871 100644
--- a/pretyping/typeclasses.mli
+++ b/pretyping/typeclasses.mli
@@ -1,14 +1,16 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Globnames
-open Term
+open Constr
open Evd
open Environ
@@ -16,14 +18,17 @@ type direction = Forward | Backward
(** This module defines type-classes *)
type typeclass = {
+ (** The toplevel universe quantification in which the typeclass lives. In
+ particular, [cl_props] and [cl_context] are quantified over it. *)
+ cl_univs : Univ.AUContext.t;
+
(** The class implementation: a record parameterized by the context with defs in it or a definition if
the class is a singleton. This acts as the class' global identifier. *)
cl_impl : global_reference;
(** Context in which the definitions are typed. Includes both typeclass parameters and superclasses.
- The boolean indicates if the typeclass argument is a direct superclass and the global reference
- gives a direct link to the class itself. *)
- cl_context : (global_reference * bool) option list * Context.Rel.t;
+ The global reference gives a direct link to the class itself. *)
+ cl_context : global_reference option list * Context.Rel.t;
(** Context of definitions and properties on defs, will not be shared *)
cl_props : Context.Rel.t;
@@ -32,7 +37,7 @@ type typeclass = {
Some may be undefinable due to sorting restrictions or simply undefined if
no name is provided. The [int option option] indicates subclasses whose hint has
the given priority. *)
- cl_projs : (Name.t * (direction * Vernacexpr.hint_info_expr) option * constant option) list;
+ cl_projs : (Name.t * (direction * Vernacexpr.hint_info_expr) option * Constant.t option) list;
(** Whether we use matching or full unification during resolution *)
cl_strict : bool;
@@ -50,7 +55,7 @@ val all_instances : unit -> instance list
val add_class : typeclass -> unit
-val new_instance : typeclass -> Vernacexpr.hint_info_expr -> bool -> Decl_kinds.polymorphic ->
+val new_instance : typeclass -> Vernacexpr.hint_info_expr -> bool ->
global_reference -> instance
val add_instance : instance -> unit
val remove_instance : instance -> unit
@@ -61,13 +66,13 @@ val class_info : global_reference -> typeclass (** raises a UserError if not a c
(** These raise a UserError if not a class.
Caution: the typeclass structures is not instantiated w.r.t. the universe instance.
This is done separately by typeclass_univ_instance. *)
-val dest_class_app : env -> constr -> typeclass puniverses * constr list
+val dest_class_app : env -> evar_map -> EConstr.constr -> (typeclass * EConstr.EInstance.t) * constr list
(** Get the instantiated typeclass structure for a given universe instance. *)
-val typeclass_univ_instance : typeclass puniverses -> typeclass puniverses
+val typeclass_univ_instance : typeclass Univ.puniverses -> typeclass
(** Just return None if not a class *)
-val class_of_constr : constr -> (Context.Rel.t * (typeclass puniverses * constr list)) option
+val class_of_constr : evar_map -> EConstr.constr -> (EConstr.rel_context * ((typeclass * EConstr.EInstance.t) * constr list)) option
val instance_impl : instance -> global_reference
@@ -79,18 +84,23 @@ val is_instance : global_reference -> bool
(** Returns the term and type for the given instance of the parameters and fields
of the type class. *)
-val instance_constructor : typeclass puniverses -> constr list ->
- constr option * types
+val instance_constructor : typeclass EConstr.puniverses -> EConstr.t list ->
+ EConstr.t option * EConstr.t
(** Filter which evars to consider for resolution. *)
-type evar_filter = existential_key -> Evar_kinds.t -> bool
+type evar_filter = Evar.t -> Evar_kinds.t -> bool
val all_evars : evar_filter
val all_goals : evar_filter
val no_goals : evar_filter
val no_goals_or_obligations : evar_filter
(** Resolvability.
- Only undefined evars can be marked or checked for resolvability. *)
+ Only undefined evars can be marked or checked for resolvability.
+ They represent type-class search roots.
+
+ A resolvable evar is an evar the type-class engine may try to solve
+ An unresolvable evar is an evar the type-class engine will NOT try to solve
+*)
val set_resolvable : Evd.Store.t -> bool -> Evd.Store.t
val is_resolvable : evar_info -> bool
@@ -99,11 +109,11 @@ val mark_unresolvables : ?filter:evar_filter -> evar_map -> evar_map
val mark_resolvables : ?filter:evar_filter -> evar_map -> evar_map
val mark_resolvable : evar_info -> evar_info
val is_class_evar : evar_map -> evar_info -> bool
-val is_class_type : evar_map -> types -> bool
+val is_class_type : evar_map -> EConstr.types -> bool
val resolve_typeclasses : ?fast_path:bool -> ?filter:evar_filter -> ?unique:bool ->
?split:bool -> ?fail:bool -> env -> evar_map -> evar_map
-val resolve_one_typeclass : ?unique:bool -> env -> evar_map -> types -> open_constr
+val resolve_one_typeclass : ?unique:bool -> env -> evar_map -> EConstr.types -> evar_map * EConstr.constr
val set_typeclass_transparency_hook : (evaluable_global_reference -> bool (*local?*) -> bool -> unit) Hook.t
val set_typeclass_transparency : evaluable_global_reference -> bool -> bool -> unit
@@ -120,7 +130,7 @@ val add_instance_hint : global_reference_or_constr -> global_reference list ->
val remove_instance_hint : global_reference -> unit
val solve_all_instances_hook : (env -> evar_map -> evar_filter -> bool -> bool -> bool -> evar_map) Hook.t
-val solve_one_instance_hook : (env -> evar_map -> types -> bool -> open_constr) Hook.t
+val solve_one_instance_hook : (env -> evar_map -> EConstr.types -> bool -> evar_map * EConstr.constr) Hook.t
val declare_instance : Vernacexpr.hint_info_expr option -> bool -> global_reference -> unit
diff --git a/pretyping/typeclasses_errors.ml b/pretyping/typeclasses_errors.ml
index b1dfb19a..e10c81c2 100644
--- a/pretyping/typeclasses_errors.ml
+++ b/pretyping/typeclasses_errors.ml
@@ -1,14 +1,15 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
+open EConstr
open Environ
open Constrexpr
open Globnames
@@ -18,7 +19,7 @@ type contexts = Parameters | Properties
type typeclass_error =
| NotAClass of constr
- | UnboundMethod of global_reference * Id.t Loc.located (* Class name, method *)
+ | UnboundMethod of global_reference * Misctypes.lident (* Class name, method *)
| MismatchedContextInstance of contexts * constr_expr list * Context.Rel.t (* found, expected *)
exception TypeClassError of env * typeclass_error
diff --git a/pretyping/typeclasses_errors.mli b/pretyping/typeclasses_errors.mli
index ee76f638..d9829565 100644
--- a/pretyping/typeclasses_errors.mli
+++ b/pretyping/typeclasses_errors.mli
@@ -1,14 +1,14 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Loc
-open Names
-open Term
+open EConstr
open Environ
open Constrexpr
open Globnames
@@ -17,14 +17,14 @@ type contexts = Parameters | Properties
type typeclass_error =
| NotAClass of constr
- | UnboundMethod of global_reference * Id.t located (** Class name, method *)
+ | UnboundMethod of global_reference * Misctypes.lident (** Class name, method *)
| MismatchedContextInstance of contexts * constr_expr list * Context.Rel.t (** found, expected *)
exception TypeClassError of env * typeclass_error
val not_a_class : env -> constr -> 'a
-val unbound_method : env -> global_reference -> Id.t located -> 'a
+val unbound_method : env -> global_reference -> Misctypes.lident -> 'a
val mismatched_ctx_inst : env -> contexts -> constr_expr list -> Context.Rel.t -> 'a
diff --git a/pretyping/typing.ml b/pretyping/typing.ml
index 9e9997f7..4c834f2f 100644
--- a/pretyping/typing.ml
+++ b/pretyping/typing.ml
@@ -1,52 +1,81 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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) *)
(************************************************************************)
+module CVars = Vars
+
open Pp
open CErrors
open Util
open Term
-open Vars
open Environ
+open EConstr
+open Vars
open Reductionops
-open Type_errors
open Inductive
open Inductiveops
open Typeops
open Arguments_renaming
+open Pretype_errors
open Context.Rel.Declaration
let meta_type evd mv =
let ty =
try Evd.meta_ftype evd mv
- with Not_found -> anomaly (str "unknown meta ?" ++ str (Nameops.string_of_meta mv)) in
+ with Not_found -> anomaly (str "unknown meta ?" ++ str (Nameops.string_of_meta mv) ++ str ".") in
+ let ty = Evd.map_fl EConstr.of_constr ty in
meta_instance evd ty
-let constant_type_knowing_parameters env cst jl =
- let paramstyp = Array.map (fun j -> lazy j.uj_type) jl in
- type_of_constant_knowing_parameters_in env cst paramstyp
-
-let inductive_type_knowing_parameters env (ind,u) jl =
+let inductive_type_knowing_parameters env sigma (ind,u) jl =
+ let u = Unsafe.to_instance u in
let mspec = lookup_mind_specif env ind in
- let paramstyp = Array.map (fun j -> lazy j.uj_type) jl in
+ let paramstyp = Array.map (fun j -> lazy (EConstr.to_constr sigma j.uj_type)) jl in
Inductive.type_of_inductive_knowing_parameters env (mspec,u) paramstyp
let e_type_judgment env evdref j =
- match kind_of_term (whd_all env !evdref j.uj_type) with
- | Sort s -> {utj_val = j.uj_val; utj_type = s }
+ match EConstr.kind !evdref (whd_all env !evdref j.uj_type) with
+ | Sort s -> {utj_val = j.uj_val; utj_type = ESorts.kind !evdref s }
| Evar ev ->
let (evd,s) = Evardefine.define_evar_as_sort env !evdref ev in
evdref := evd; { utj_val = j.uj_val; utj_type = s }
- | _ -> error_not_type env j
+ | _ -> error_not_a_type env !evdref j
let e_assumption_of_judgment env evdref j =
try (e_type_judgment env evdref j).utj_val
- with TypeError _ ->
- error_assumption env j
+ with Type_errors.TypeError _ | PretypeError _ ->
+ error_assumption env !evdref j
+
+let e_judge_of_applied_inductive_knowing_parameters env evdref funj ind argjv =
+ let rec apply_rec n typ = function
+ | [] ->
+ { uj_val = mkApp (j_val funj, Array.map j_val argjv);
+ uj_type =
+ let ar = inductive_type_knowing_parameters env !evdref ind argjv in
+ hnf_prod_appvect env !evdref (EConstr.of_constr ar) (Array.map j_val argjv) }
+ | hj::restjl ->
+ let (c1,c2) =
+ match EConstr.kind !evdref (whd_all env !evdref typ) with
+ | Prod (_,c1,c2) -> (c1,c2)
+ | Evar ev ->
+ let (evd',t) = Evardefine.define_evar_as_product !evdref ev in
+ evdref := evd';
+ let (_,c1,c2) = destProd evd' t in
+ (c1,c2)
+ | _ ->
+ error_cant_apply_not_functional env !evdref funj argjv
+ in
+ if Evarconv.e_cumul env evdref hj.uj_type c1 then
+ apply_rec (n+1) (subst1 hj.uj_val c2) restjl
+ else
+ error_cant_apply_bad_type env !evdref (n, c1, hj.uj_type) funj argjv
+ in
+ apply_rec 1 funj.uj_type (Array.to_list argjv)
let e_judge_of_apply env evdref funj argjv =
let rec apply_rec n typ = function
@@ -54,28 +83,30 @@ let e_judge_of_apply env evdref funj argjv =
{ uj_val = mkApp (j_val funj, Array.map j_val argjv);
uj_type = typ }
| hj::restjl ->
- match kind_of_term (whd_all env !evdref typ) with
- | Prod (_,c1,c2) ->
- if Evarconv.e_cumul env evdref hj.uj_type c1 then
- apply_rec (n+1) (subst1 hj.uj_val c2) restjl
- else
- error_cant_apply_bad_type env (n,c1, hj.uj_type) funj argjv
- | Evar ev ->
- let (evd',t) = Evardefine.define_evar_as_product !evdref ev in
- evdref := evd';
- let (_,_,c2) = destProd t in
- apply_rec (n+1) (subst1 hj.uj_val c2) restjl
- | _ ->
- error_cant_apply_not_functional env funj argjv
+ let (c1,c2) =
+ match EConstr.kind !evdref (whd_all env !evdref typ) with
+ | Prod (_,c1,c2) -> (c1,c2)
+ | Evar ev ->
+ let (evd',t) = Evardefine.define_evar_as_product !evdref ev in
+ evdref := evd';
+ let (_,c1,c2) = destProd evd' t in
+ (c1,c2)
+ | _ ->
+ error_cant_apply_not_functional env !evdref funj argjv
+ in
+ if Evarconv.e_cumul env evdref hj.uj_type c1 then
+ apply_rec (n+1) (subst1 hj.uj_val c2) restjl
+ else
+ error_cant_apply_bad_type env !evdref (n, c1, hj.uj_type) funj argjv
in
apply_rec 1 funj.uj_type (Array.to_list argjv)
let e_check_branch_types env evdref (ind,u) cj (lfj,explft) =
if not (Int.equal (Array.length lfj) (Array.length explft)) then
- error_number_branches env cj (Array.length explft);
+ error_number_branches env !evdref cj (Array.length explft);
for i = 0 to Array.length explft - 1 do
if not (Evarconv.e_cumul env evdref lfj.(i).uj_type explft.(i)) then
- error_ill_formed_branch env cj.uj_val ((ind,i+1),u) lfj.(i).uj_type explft.(i)
+ error_ill_formed_branch env !evdref cj.uj_val ((ind,i+1),u) lfj.(i).uj_type explft.(i)
done
let max_sort l =
@@ -83,21 +114,22 @@ let max_sort l =
if Sorts.List.mem InSet l then InSet else InProp
let e_is_correct_arity env evdref c pj ind specif params =
- let arsign = make_arity_signature env true (make_ind_family (ind,params)) in
+ let arsign = make_arity_signature env !evdref true (make_ind_family (ind,params)) in
let allowed_sorts = elim_sorts specif in
- let error () = error_elim_arity env ind allowed_sorts c pj None in
+ let error () = Pretype_errors.error_elim_arity env !evdref ind allowed_sorts c pj None in
let rec srec env pt ar =
let pt' = whd_all env !evdref pt in
- match kind_of_term pt', ar with
+ match EConstr.kind !evdref pt', ar with
| Prod (na1,a1,t), (LocalAssum (_,a1'))::ar' ->
if not (Evarconv.e_cumul env evdref a1 a1') then error ();
srec (push_rel (LocalAssum (na1,a1)) env) t ar'
| Sort s, [] ->
+ let s = ESorts.kind !evdref s in
if not (Sorts.List.mem (Sorts.family s) allowed_sorts)
then error ()
| Evar (ev,_), [] ->
let evd, s = Evd.fresh_sort_in_family env !evdref (max_sort allowed_sorts) in
- evdref := Evd.define ev (mkSort s) evd
+ evdref := Evd.define ev (Constr.mkSort s) evd
| _, (LocalDef _ as d)::ar' ->
srec (push_rel d env) (lift 1 pt') ar'
| _ ->
@@ -105,52 +137,66 @@ let e_is_correct_arity env evdref c pj ind specif params =
in
srec env pj.uj_type (List.rev arsign)
+let lambda_applist_assum sigma n c l =
+ let rec app n subst t l =
+ if Int.equal n 0 then
+ if l == [] then substl subst t
+ else anomaly (Pp.str "Not enough arguments.")
+ else match EConstr.kind sigma t, l with
+ | Lambda(_,_,c), arg::l -> app (n-1) (arg::subst) c l
+ | LetIn(_,b,_,c), _ -> app (n-1) (substl subst b::subst) c l
+ | _ -> anomaly (Pp.str "Not enough lambda/let's.") in
+ app n [] c l
+
let e_type_case_branches env evdref (ind,largs) pj c =
let specif = lookup_mind_specif env (fst ind) in
let nparams = inductive_params specif in
let (params,realargs) = List.chop nparams largs in
let p = pj.uj_val in
+ let params = List.map EConstr.Unsafe.to_constr params in
let () = e_is_correct_arity env evdref c pj ind specif params in
- let lc = build_branches_type ind specif params p in
+ let lc = build_branches_type ind specif params (EConstr.to_constr !evdref p) in
+ let lc = Array.map EConstr.of_constr lc in
let n = (snd specif).Declarations.mind_nrealdecls in
- let ty = whd_betaiota !evdref (lambda_applist_assum (n+1) p (realargs@[c])) in
+ let ty = whd_betaiota !evdref (lambda_applist_assum !evdref (n+1) p (realargs@[c])) in
(lc, ty)
let e_judge_of_case env evdref ci pj cj lfj =
- let indspec =
+ let ((ind, u), spec) =
try find_mrectype env !evdref cj.uj_type
- with Not_found -> error_case_not_inductive env cj in
+ with Not_found -> error_case_not_inductive env !evdref cj in
+ let indspec = ((ind, EInstance.kind !evdref u), spec) in
let _ = check_case_info env (fst indspec) ci in
let (bty,rslty) = e_type_case_branches env evdref indspec pj cj.uj_val in
e_check_branch_types env evdref (fst indspec) cj (lfj,bty);
{ uj_val = mkCase (ci, pj.uj_val, cj.uj_val, Array.map j_val lfj);
uj_type = rslty }
-let check_type_fixpoint loc env evdref lna lar vdefj =
+let check_type_fixpoint ?loc env evdref lna lar vdefj =
let lt = Array.length vdefj in
if Int.equal (Array.length lar) lt then
for i = 0 to lt-1 do
if not (Evarconv.e_cumul env evdref (vdefj.(i)).uj_type
(lift lt lar.(i))) then
- Pretype_errors.error_ill_typed_rec_body_loc loc env !evdref
+ error_ill_typed_rec_body ?loc env !evdref
i lna vdefj lar
done
(* FIXME: might depend on the level of actual parameters!*)
let check_allowed_sort env sigma ind c p =
let pj = Retyping.get_judgment_of env sigma p in
- let ksort = family_of_sort (sort_of_arity env sigma pj.uj_type) in
+ let ksort = Sorts.family (ESorts.kind sigma (sort_of_arity env sigma pj.uj_type)) in
let specif = Global.lookup_inductive (fst ind) in
let sorts = elim_sorts specif in
if not (List.exists ((==) ksort) sorts) then
let s = inductive_sort_family (snd specif) in
- error_elim_arity env ind sorts c pj
- (Some(ksort,s,error_elim_explain ksort s))
+ error_elim_arity env sigma ind sorts c pj
+ (Some(ksort,s,Type_errors.error_elim_explain ksort s))
let e_judge_of_cast env evdref cj k tj =
let expected_type = tj.utj_val in
if not (Evarconv.e_cumul env evdref cj.uj_type expected_type) then
- error_actual_type env cj expected_type;
+ error_actual_type_core env !evdref cj expected_type;
{ uj_val = mkCast (cj.uj_val, k, expected_type);
uj_type = expected_type }
@@ -159,19 +205,78 @@ let enrich_env env evdref =
let penv' = Pre_env.({ penv with env_stratification =
{ penv.env_stratification with env_universes = Evd.universes !evdref } }) in
Environ.env_of_pre_env penv'
-
+
+let check_fix env sigma pfix =
+ let inj c = EConstr.to_constr sigma c in
+ let (idx, (ids, cs, ts)) = pfix in
+ check_fix env (idx, (ids, Array.map inj cs, Array.map inj ts))
+
+let check_cofix env sigma pcofix =
+ let inj c = EConstr.to_constr sigma c in
+ let (idx, (ids, cs, ts)) = pcofix in
+ check_cofix env (idx, (ids, Array.map inj cs, Array.map inj ts))
+
(* The typing machine with universes and existential variables. *)
+let judge_of_prop =
+ { uj_val = EConstr.mkProp;
+ uj_type = EConstr.mkSort Sorts.type1 }
+
+let judge_of_set =
+ { uj_val = EConstr.mkSet;
+ uj_type = EConstr.mkSort Sorts.type1 }
+
+let judge_of_prop_contents = function
+ | Null -> judge_of_prop
+ | Pos -> judge_of_set
+
+let judge_of_type u =
+ let uu = Univ.Universe.super u in
+ { uj_val = EConstr.mkType u;
+ uj_type = EConstr.mkType uu }
+
+let judge_of_relative env v =
+ Termops.on_judgment EConstr.of_constr (judge_of_relative env v)
+
+let judge_of_variable env id =
+ Termops.on_judgment EConstr.of_constr (judge_of_variable env id)
+
+let judge_of_projection env sigma p cj =
+ let pb = lookup_projection p env in
+ let (ind,u), args =
+ try find_mrectype env sigma cj.uj_type
+ with Not_found -> error_case_not_inductive env sigma cj
+ in
+ let u = EInstance.kind sigma u in
+ let ty = EConstr.of_constr (CVars.subst_instance_constr u pb.Declarations.proj_type) in
+ let ty = substl (cj.uj_val :: List.rev args) ty in
+ {uj_val = EConstr.mkProj (p,cj.uj_val);
+ uj_type = ty}
+
+let judge_of_abstraction env name var j =
+ { uj_val = mkLambda (name, var.utj_val, j.uj_val);
+ uj_type = mkProd (name, var.utj_val, j.uj_type) }
+
+let judge_of_product env name t1 t2 =
+ let s = sort_of_product env t1.utj_type t2.utj_type in
+ { uj_val = mkProd (name, t1.utj_val, t2.utj_val);
+ uj_type = mkSort s }
+
+let judge_of_letin env name defj typj j =
+ { uj_val = mkLetIn (name, defj.uj_val, typj.utj_val, j.uj_val) ;
+ uj_type = subst1 defj.uj_val j.uj_type }
+
(* cstr must be in n.f. w.r.t. evars and execute returns a judgement
where both the term and type are in n.f. *)
let rec execute env evdref cstr =
- match kind_of_term cstr with
+ let cstr = whd_evar !evdref cstr in
+ match EConstr.kind !evdref cstr with
| Meta n ->
{ uj_val = cstr; uj_type = meta_type !evdref n }
| Evar ev ->
- let ty = Evd.existential_type !evdref ev in
- let jty = execute env evdref (whd_evar !evdref ty) in
+ let ty = EConstr.existential_type !evdref ev in
+ let jty = execute env evdref ty in
let jty = e_assumption_of_judgment env evdref jty in
{ uj_val = cstr; uj_type = jty }
@@ -181,14 +286,17 @@ let rec execute env evdref cstr =
| Var id ->
judge_of_variable env id
- | Const c ->
- make_judge cstr (rename_type_of_constant env c)
+ | Const (c, u) ->
+ let u = EInstance.kind !evdref u in
+ make_judge cstr (EConstr.of_constr (rename_type_of_constant env (c, u)))
- | Ind ind ->
- make_judge cstr (rename_type_of_inductive env ind)
+ | Ind (ind, u) ->
+ let u = EInstance.kind !evdref u in
+ make_judge cstr (EConstr.of_constr (rename_type_of_inductive env (ind, u)))
- | Construct cstruct ->
- make_judge cstr (rename_type_of_constructor env cstruct)
+ | Construct (cstruct, u) ->
+ let u = EInstance.kind !evdref u in
+ make_judge cstr (EConstr.of_constr (rename_type_of_constructor env (cstruct, u)))
| Case (ci,p,c,lf) ->
let cj = execute env evdref c in
@@ -199,43 +307,37 @@ let rec execute env evdref cstr =
| Fix ((vn,i as vni),recdef) ->
let (_,tys,_ as recdef') = execute_recdef env evdref recdef in
let fix = (vni,recdef') in
- check_fix env fix;
+ check_fix env !evdref fix;
make_judge (mkFix fix) tys.(i)
| CoFix (i,recdef) ->
let (_,tys,_ as recdef') = execute_recdef env evdref recdef in
let cofix = (i,recdef') in
- check_cofix env cofix;
+ check_cofix env !evdref cofix;
make_judge (mkCoFix cofix) tys.(i)
- | Sort (Prop c) ->
- judge_of_prop_contents c
-
- | Sort (Type u) ->
+ | Sort s ->
+ begin match ESorts.kind !evdref s with
+ | Prop c ->
+ judge_of_prop_contents c
+ | Type u ->
judge_of_type u
+ end
| Proj (p, c) ->
let cj = execute env evdref c in
- judge_of_projection env p (Evarutil.j_nf_evar !evdref cj)
+ judge_of_projection env !evdref p cj
| App (f,args) ->
let jl = execute_array env evdref args in
- let j =
- match kind_of_term f with
- | Ind ind when Environ.template_polymorphic_pind ind env ->
- (* Sort-polymorphism of inductive types *)
- make_judge f
- (inductive_type_knowing_parameters env ind
- (Evarutil.jv_nf_evar !evdref jl))
- | Const cst when Environ.template_polymorphic_pconstant cst env ->
- (* Sort-polymorphism of inductive types *)
- make_judge f
- (constant_type_knowing_parameters env cst
- (Evarutil.jv_nf_evar !evdref jl))
- | _ ->
- execute env evdref f
- in
- e_judge_of_apply env evdref j jl
+ (match EConstr.kind !evdref f with
+ | Ind (ind, u) when EInstance.is_empty u && Environ.template_polymorphic_ind ind env ->
+ let fj = execute env evdref f in
+ e_judge_of_applied_inductive_knowing_parameters env evdref fj (ind, u) jl
+ | _ ->
+ (* No template polymorphism *)
+ let fj = execute env evdref f in
+ e_judge_of_apply env evdref fj jl)
| Lambda (name,c1,c2) ->
let j = execute env evdref c1 in
@@ -273,7 +375,7 @@ and execute_recdef env evdref (names,lar,vdef) =
let env1 = push_rec_types (names,lara,vdef) env in
let vdefj = execute_array env1 evdref vdef in
let vdefv = Array.map j_val vdefj in
- let _ = check_type_fixpoint Loc.ghost env1 evdref names lara vdefj in
+ let _ = check_type_fixpoint env1 evdref names lara vdefj in
(names,lara,vdefv)
and execute_array env evdref = Array.map (execute env evdref)
@@ -282,7 +384,7 @@ let e_check env evdref c t =
let env = enrich_env env evdref in
let j = execute env evdref c in
if not (Evarconv.e_cumul env evdref j.uj_type t) then
- error_actual_type env j (nf_evar !evdref t)
+ error_actual_type_core env !evdref j t
(* Type of a constr *)
@@ -327,4 +429,4 @@ let e_solve_evars env evdref c =
(* side-effect on evdref *)
nf_evar !evdref c
-let _ = Evarconv.set_solve_evars e_solve_evars
+let _ = Evarconv.set_solve_evars (fun env evdref c -> e_solve_evars env evdref c)
diff --git a/pretyping/typing.mli b/pretyping/typing.mli
index 04e5e40b..0ff724a1 100644
--- a/pretyping/typing.mli
+++ b/pretyping/typing.mli
@@ -1,13 +1,17 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
+open Names
+open Constr
open Environ
+open EConstr
open Evd
(** This module provides the typing machine with existential variables
@@ -24,7 +28,7 @@ val type_of : ?refresh:bool -> env -> evar_map -> constr -> evar_map * types
val e_type_of : ?refresh:bool -> env -> evar_map ref -> constr -> types
(** Typecheck a type and return its sort *)
-val e_sort_of : env -> evar_map ref -> types -> sorts
+val e_sort_of : env -> evar_map ref -> types -> Sorts.t
(** Typecheck a term has a given type (assuming the type is OK) *)
val e_check : env -> evar_map ref -> constr -> types -> unit
@@ -42,5 +46,12 @@ val check_allowed_sort : env -> evar_map -> pinductive -> constr -> constr ->
(** Raise an error message if bodies have types not unifiable with the
expected ones *)
-val check_type_fixpoint : Loc.t -> env -> evar_map ref ->
+val check_type_fixpoint : ?loc:Loc.t -> env -> evar_map ref ->
Names.Name.t array -> types array -> unsafe_judgment array -> unit
+
+val judge_of_prop : unsafe_judgment
+val judge_of_set : unsafe_judgment
+val judge_of_abstraction : Environ.env -> Name.t ->
+ unsafe_type_judgment -> unsafe_judgment -> unsafe_judgment
+val judge_of_product : Environ.env -> Name.t ->
+ unsafe_type_judgment -> unsafe_type_judgment -> unsafe_judgment
diff --git a/pretyping/unification.ml b/pretyping/unification.ml
index 25931869..6f594186 100644
--- a/pretyping/unification.ml
+++ b/pretyping/unification.ml
@@ -1,20 +1,23 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 CErrors
open Pp
open Util
open Names
-open Term
-open Vars
+open Constr
open Termops
-open Namegen
open Environ
+open EConstr
+open Vars
+open Namegen
open Evd
open Reduction
open Reductionops
@@ -28,12 +31,20 @@ open Recordops
open Locus
open Locusops
open Find_subterm
-open Sigma.Notations
-open Context.Named.Declaration
+
+type metabinding = (metavariable * EConstr.constr * (instance_constraint * instance_typing_status))
+
+type subst0 =
+ (evar_map *
+ metabinding list *
+ (Environ.env * EConstr.existential * EConstr.t) list)
+
+module RelDecl = Context.Rel.Declaration
+module NamedDecl = Context.Named.Declaration
let keyed_unification = ref (false)
let _ = Goptions.declare_bool_option {
- Goptions.optsync = true; Goptions.optdepr = false;
+ Goptions.optdepr = false;
Goptions.optname = "Unification is keyed";
Goptions.optkey = ["Keyed";"Unification"];
Goptions.optread = (fun () -> !keyed_unification);
@@ -44,7 +55,7 @@ let is_keyed_unification () = !keyed_unification
let debug_unification = ref (false)
let _ = Goptions.declare_bool_option {
- Goptions.optsync = true; Goptions.optdepr = false;
+ Goptions.optdepr = false;
Goptions.optname =
"Print states sent to tactic unification";
Goptions.optkey = ["Debug";"Tactic";"Unification"];
@@ -52,8 +63,23 @@ let _ = Goptions.declare_bool_option {
Goptions.optwrite = (fun a -> debug_unification:=a);
}
+(** Making this unification algorithm correct w.r.t. the evar-map abstraction
+ breaks too much stuff. So we redefine incorrect functions here. *)
+
+let unsafe_occur_meta_or_existential c =
+ let c = EConstr.Unsafe.to_constr c in
+ let rec occrec c = match Constr.kind c with
+ | Evar _ -> raise Occur
+ | Meta _ -> raise Occur
+ | _ -> Constr.iter occrec c
+ in try occrec c; false with Occur -> true
+
+
let occur_meta_or_undefined_evar evd c =
- let rec occrec c = match kind_of_term c with
+ (** This is performance-critical. Using the evar-insensitive API changes the
+ resulting heuristic. *)
+ let c = EConstr.Unsafe.to_constr c in
+ let rec occrec c = match Constr.kind c with
| Meta _ -> raise Occur
| Evar (ev,args) ->
(match evar_body (Evd.find evd ev) with
@@ -66,27 +92,29 @@ let occur_meta_or_undefined_evar evd c =
let occur_meta_evd sigma mv c =
let rec occrec c =
(* Note: evars are not instantiated by terms with metas *)
- let c = whd_evar sigma (whd_meta sigma c) in
- match kind_of_term c with
+ let c = whd_meta sigma c in
+ match EConstr.kind sigma c with
| Meta mv' when Int.equal mv mv' -> raise Occur
- | _ -> Constr.iter occrec c
+ | _ -> EConstr.iter sigma occrec c
in try occrec c; false with Occur -> true
(* if lname_typ is [xn,An;..;x1,A1] and l is a list of terms,
gives [x1:A1]..[xn:An]c' such that c converts to ([x1:A1]..[xn:An]c' l) *)
let abstract_scheme env evd c l lname_typ =
+ let mkLambda_name env (n,a,b) =
+ mkLambda (named_hd env evd a n, a, b)
+ in
List.fold_left2
(fun (t,evd) (locc,a) decl ->
- let open Context.Rel.Declaration in
- let na = get_name decl in
- let ta = get_type decl in
- let na = match kind_of_term a with Var id -> Name id | _ -> na in
+ let na = RelDecl.get_name decl in
+ let ta = RelDecl.get_type decl in
+ let na = match EConstr.kind evd a with Var id -> Name id | _ -> na in
(* [occur_meta ta] test removed for support of eelim/ecase but consequences
are unclear...
if occur_meta ta then error "cannot find a type for the generalisation"
else *)
- if occur_meta a then mkLambda_name env (na,ta,t), evd
+ if occur_meta evd a then mkLambda_name env (na,ta,t), evd
else
let t', evd' = Find_subterm.subst_closed_term_occ env evd locc a t in
mkLambda_name env (na,ta,t'), evd')
@@ -106,6 +134,9 @@ let abstract_list_all env evd typ c l =
| UserError _ ->
error_cannot_find_well_typed_abstraction env evd p l None
| Type_errors.TypeError (env',x) ->
+ (** FIXME: plug back the typing information *)
+ error_cannot_find_well_typed_abstraction env evd p l None
+ | Pretype_errors.PretypeError (env',evd,TypingError x) ->
error_cannot_find_well_typed_abstraction env evd p l (Some (env',x)) in
evd,(p,typp)
@@ -113,20 +144,18 @@ let set_occurrences_of_last_arg args =
Some AllOccurrences :: List.tl (Array.map_to_list (fun _ -> None) args)
let abstract_list_all_with_dependencies env evd typ c l =
- let evd = Sigma.Unsafe.of_evar_map evd in
- let Sigma (ev, evd, _) = new_evar env evd typ in
- let evd = Sigma.to_evar_map evd in
- let evd,ev' = evar_absorb_arguments env evd (destEvar ev) l in
+ let (evd, ev) = new_evar env evd typ in
+ let evd,ev' = evar_absorb_arguments env evd (destEvar evd ev) l in
let n = List.length l in
let argoccs = set_occurrences_of_last_arg (Array.sub (snd ev') 0 n) in
let evd,b =
Evarconv.second_order_matching empty_transparent_state
env evd ev' argoccs c in
if b then
- let p = nf_evar evd (existential_value evd (destEvar ev)) in
+ let p = nf_evar evd ev in
evd, p
else error_cannot_find_well_typed_abstraction env evd
- (nf_evar evd c) l None
+ c l None
(**)
@@ -146,50 +175,57 @@ let extract_instance_status = function
| CUMUL -> add_type_status (IsSubType, IsSuperType)
| CONV -> add_type_status (Conv, Conv)
-let rec subst_meta_instances bl c =
- match kind_of_term c with
+let rec subst_meta_instances sigma bl c =
+ match EConstr.kind sigma c with
| Meta i ->
let select (j,_,_) = Int.equal i j in
(try pi2 (List.find select bl) with Not_found -> c)
- | _ -> Constr.map (subst_meta_instances bl) c
+ | _ -> EConstr.map sigma (subst_meta_instances sigma bl) c
(** [env] should be the context in which the metas live *)
let pose_all_metas_as_evars env evd t =
let evdref = ref evd in
- let rec aux t = match kind_of_term t with
+ let rec aux t = match EConstr.kind !evdref t with
| Meta mv ->
(match Evd.meta_opt_fvalue !evdref mv with
- | Some ({rebus=c},_) -> c
+ | Some ({rebus=c},_) -> EConstr.of_constr c
| None ->
let {rebus=ty;freemetas=mvs} = Evd.meta_ftype evd mv in
+ let ty = EConstr.of_constr ty in
let ty = if Evd.Metaset.is_empty mvs then ty else aux ty in
+ let ty =
+ if Flags.version_strictly_greater Flags.V8_6
+ then nf_betaiota env evd ty (* How it was in Coq <= 8.4 (but done in logic.ml at this time) *)
+ else ty (* some beta-iota-normalization "regression" in 8.5 and 8.6 *) in
let src = Evd.evar_source_of_meta mv !evdref in
let ev = Evarutil.e_new_evar env evdref ~src ty in
- evdref := meta_assign mv (ev,(Conv,TypeNotProcessed)) !evdref;
+ evdref := meta_assign mv (EConstr.Unsafe.to_constr ev,(Conv,TypeNotProcessed)) !evdref;
ev)
| _ ->
- Constr.map aux t in
+ EConstr.map !evdref aux t in
let c = aux t in
(* side-effect *)
(!evdref, c)
-let solve_pattern_eqn_array (env,nb) f l c (sigma,metasubst,evarsubst) =
- match kind_of_term f with
+let solve_pattern_eqn_array (env,nb) f l c (sigma,metasubst,evarsubst : subst0) =
+ match EConstr.kind sigma f with
| Meta k ->
(* We enforce that the Meta does not depend on the [nb]
extra assumptions added by unification to the context *)
let env' = pop_rel_context nb env in
let sigma,c = pose_all_metas_as_evars env' sigma c in
- let c = solve_pattern_eqn env l c in
+ let c = solve_pattern_eqn env sigma l c in
let pb = (Conv,TypeNotProcessed) in
- if noccur_between 1 nb c then
+ if noccur_between sigma 1 nb c then
sigma,(k,lift (-nb) c,pb)::metasubst,evarsubst
- else error_cannot_unify_local env sigma (applist (f, l),c,c)
+ else
+ let l = List.map of_alias l in
+ error_cannot_unify_local env sigma (applist (f, l),c,c)
| Evar ev ->
let env' = pop_rel_context nb env in
let sigma,c = pose_all_metas_as_evars env' sigma c in
- sigma,metasubst,(env,ev,solve_pattern_eqn env l c)::evarsubst
+ sigma,metasubst,(env,ev,solve_pattern_eqn env sigma l c)::evarsubst
| _ -> assert false
let push d (env,n) = (push_rel_assum d env,n+1)
@@ -214,33 +250,6 @@ let unify_r2l x = x
let sort_eqns = unify_r2l
*)
-let global_pattern_unification_flag = ref true
-
-(* Compatibility option introduced and activated in Coq 8.3 whose
- syntax is now deprecated. *)
-
-open Goptions
-let _ =
- declare_bool_option
- { optsync = true;
- optdepr = true;
- optname = "pattern-unification for existential variables in tactics";
- optkey = ["Tactic";"Evars";"Pattern";"Unification"];
- optread = (fun () -> !global_pattern_unification_flag);
- optwrite = (:=) global_pattern_unification_flag }
-
-(* Compatibility option superseding the previous one, introduced and
- activated in Coq 8.4 *)
-
-let _ =
- declare_bool_option
- { optsync = true;
- optdepr = false;
- optname = "pattern-unification for existential variables in tactics";
- optkey = ["Tactic";"Pattern";"Unification"];
- optread = (fun () -> !global_pattern_unification_flag);
- optwrite = (:=) global_pattern_unification_flag }
-
type core_unify_flags = {
modulo_conv_on_closed_terms : Names.transparent_state option;
(* What this flag controls was activated with all constants transparent, *)
@@ -264,12 +273,10 @@ type core_unify_flags = {
use_pattern_unification : bool;
(* This solves pattern "?n x1 ... xn = t" when the xi are distinct rels *)
- (* This says if pattern unification is tried; can be overwritten with *)
- (* option "Set Tactic Pattern Unification" *)
+ (* This says if pattern unification is tried *)
use_meta_bound_pattern_unification : bool;
- (* This is implied by use_pattern_unification (though deactivated *)
- (* by unsetting Tactic Pattern Unification); has no particular *)
+ (* This is implied by use_pattern_unification; has no particular *)
(* reasons to be set differently than use_pattern_unification *)
(* except for compatibility of "auto". *)
(* This was on for all tactics, including auto, since Sep 2006 for 8.1 *)
@@ -450,18 +457,16 @@ let set_flags_for_type flags = { flags with
}
let use_evars_pattern_unification flags =
- !global_pattern_unification_flag && flags.use_pattern_unification
- && Flags.version_strictly_greater Flags.V8_2
+ flags.use_pattern_unification
-let use_metas_pattern_unification flags nb l =
- !global_pattern_unification_flag && flags.use_pattern_unification
- || (Flags.version_less_or_equal Flags.V8_3 ||
- flags.use_meta_bound_pattern_unification) &&
- Array.for_all (fun c -> isRel c && destRel c <= nb) l
+let use_metas_pattern_unification sigma flags nb l =
+ flags.use_pattern_unification
+ || flags.use_meta_bound_pattern_unification &&
+ Array.for_all (fun c -> isRel sigma c && destRel sigma c <= nb) l
type key =
| IsKey of CClosure.table_key
- | IsProj of projection * constr
+ | IsProj of Projection.t * EConstr.constr
let expand_table_key env = function
| ConstKey cst -> constant_opt_value_in env cst
@@ -477,13 +482,17 @@ let unfold_projection env p stk =
| None -> assert false)
let expand_key ts env sigma = function
- | Some (IsKey k) -> expand_table_key env k
+ | Some (IsKey k) -> Option.map EConstr.of_constr (expand_table_key env k)
| Some (IsProj (p, c)) ->
- let red = Stack.zip (fst (whd_betaiota_deltazeta_for_iota_state ts env sigma
+ let red = Stack.zip sigma (fst (whd_betaiota_deltazeta_for_iota_state ts env sigma
Cst_stack.empty (c, unfold_projection env p [])))
- in if Term.eq_constr (mkProj (p, c)) red then None else Some red
+ in if EConstr.eq_constr sigma (EConstr.mkProj (p, c)) red then None else Some red
| None -> None
+let isApp_or_Proj sigma c =
+ match kind sigma c with
+ | App _ | Proj _ -> true
+ | _ -> false
type unirec_flags = {
at_top: bool;
@@ -494,12 +503,13 @@ type unirec_flags = {
let subterm_restriction opt flags =
not opt.at_top && flags.restrict_conv_on_strict_subterms
-let key_of env b flags f =
+let key_of env sigma b flags f =
if subterm_restriction b flags then None else
- match kind_of_term f with
+ match EConstr.kind sigma f with
| Const (cst, u) when is_transparent env (ConstKey cst) &&
(Cpred.mem cst (snd flags.modulo_delta)
|| Environ.is_projection cst env) ->
+ let u = EInstance.kind sigma u in
Some (IsKey (ConstKey (cst, u)))
| Var id when is_transparent env (VarKey id) &&
Id.Pred.mem id (fst flags.modulo_delta) ->
@@ -532,74 +542,85 @@ let oracle_order env cf1 cf2 =
| Some k2 ->
match k1, k2 with
| IsProj (p, _), IsKey (ConstKey (p',_))
- when eq_constant (Projection.constant p) p' ->
+ when Constant.equal (Projection.constant p) p' ->
Some (not (Projection.unfolded p))
| IsKey (ConstKey (p,_)), IsProj (p', _)
- when eq_constant p (Projection.constant p') ->
+ when Constant.equal p (Projection.constant p') ->
Some (Projection.unfolded p')
| _ ->
Some (Conv_oracle.oracle_order (fun x -> x)
(Environ.oracle env) false (translate_key k1) (translate_key k2))
-let is_rigid_head flags t =
- match kind_of_term t with
+let is_rigid_head sigma flags t =
+ match EConstr.kind sigma t with
| Const (cst,u) -> not (Cpred.mem cst (snd flags.modulo_delta))
| Ind (i,u) -> true
| Construct _ -> true
| Fix _ | CoFix _ -> true
- | _ -> false
-
-let force_eqs c =
- Universes.Constraints.fold
- (fun ((l,d,r) as c) acc ->
- let c' = if d == Universes.ULub then (l,Universes.UEq,r) else c in
- Universes.Constraints.add c' acc)
- c Universes.Constraints.empty
-
-let constr_cmp pb sigma flags t u =
- let b, cstrs =
- if pb == Reduction.CONV then Universes.eq_constr_universes t u
- else Universes.leq_constr_universes t u
+ | Rel _ | Var _ | Meta _ | Evar _ | Sort _ | Cast (_, _, _) | Prod (_, _, _)
+ | Lambda (_, _, _) | LetIn (_, _, _, _) | App (_, _) | Case (_, _, _, _)
+ | Proj (_, _) -> false (* Why aren't Prod, Sort rigid heads ? *)
+
+let force_eqs c =
+ let open Universes in
+ Constraints.fold
+ (fun c acc ->
+ let c' = match c with
+ (* Should we be forcing weak constraints? *)
+ | ULub (l, r) | UWeak (l, r) -> UEq (Univ.Universe.make l,Univ.Universe.make r)
+ | ULe _ | UEq _ -> c
+ in
+ Constraints.add c' acc)
+ c Constraints.empty
+
+let constr_cmp pb env sigma flags t u =
+ let cstrs =
+ if pb == Reduction.CONV then EConstr.eq_constr_universes env sigma t u
+ else EConstr.leq_constr_universes env sigma t u
in
- if b then
- try Evd.add_universe_constraints sigma cstrs, b
+ match cstrs with
+ | Some cstrs ->
+ begin try Evd.add_universe_constraints sigma cstrs, true
with Univ.UniverseInconsistency _ -> sigma, false
| Evd.UniversesDiffer ->
- if is_rigid_head flags t then
- try Evd.add_universe_constraints sigma (force_eqs cstrs), b
+ if is_rigid_head sigma flags t then
+ try Evd.add_universe_constraints sigma (force_eqs cstrs), true
with Univ.UniverseInconsistency _ -> sigma, false
else sigma, false
- else sigma, b
+ end
+ | None ->
+ sigma, false
let do_reduce ts (env, nb) sigma c =
- Stack.zip (fst (whd_betaiota_deltazeta_for_iota_state
+ Stack.zip sigma (fst (whd_betaiota_deltazeta_for_iota_state
ts env sigma Cst_stack.empty (c, Stack.empty)))
-let use_full_betaiota flags =
- flags.modulo_betaiota && Flags.version_strictly_greater Flags.V8_3
-
-let isAllowedEvar flags c = match kind_of_term c with
+let isAllowedEvar sigma flags c = match EConstr.kind sigma c with
| Evar (evk,_) -> not (Evar.Set.mem evk flags.frozen_evars)
| _ -> false
-let subst_defined_metas_evars (bl,el) c =
- let rec substrec c = match kind_of_term c with
+let subst_defined_metas_evars sigma (bl,el) c =
+ (** This seems to be performance-critical, and using the evar-insensitive
+ primitives blow up the time passed in this function. *)
+ let c = EConstr.Unsafe.to_constr c in
+ let rec substrec c = match Constr.kind c with
| Meta i ->
let select (j,_,_) = Int.equal i j in
- substrec (pi2 (List.find select bl))
+ substrec (EConstr.Unsafe.to_constr (pi2 (List.find select bl)))
| Evar (evk,args) ->
- let select (_,(evk',args'),_) = Evar.equal evk evk' && Array.equal Constr.equal args args' in
- (try substrec (pi3 (List.find select el))
+ let eq c1 c2 = Constr.equal c1 (EConstr.Unsafe.to_constr c2) in
+ let select (_,(evk',args'),_) = Evar.equal evk evk' && Array.for_all2 eq args args' in
+ (try substrec (EConstr.Unsafe.to_constr (pi3 (List.find select el)))
with Not_found -> Constr.map substrec c)
| _ -> Constr.map substrec c
- in try Some (substrec c) with Not_found -> None
+ in try Some (EConstr.of_constr (substrec c)) with Not_found -> None
-let check_compatibility env pbty flags (sigma,metasubst,evarsubst) tyM tyN =
- match subst_defined_metas_evars (metasubst,[]) tyM with
+let check_compatibility env pbty flags (sigma,metasubst,evarsubst : subst0) tyM tyN =
+ match subst_defined_metas_evars sigma (metasubst,[]) tyM with
| None -> sigma
| Some m ->
- match subst_defined_metas_evars (metasubst,[]) tyN with
+ match subst_defined_metas_evars sigma (metasubst,[]) tyN with
| None -> sigma
| Some n ->
if is_ground_term sigma m && is_ground_term sigma n then
@@ -609,9 +630,9 @@ let check_compatibility env pbty flags (sigma,metasubst,evarsubst) tyM tyN =
else sigma
-let rec is_neutral env ts t =
- let (f, l) = decompose_appvect t in
- match kind_of_term f with
+let rec is_neutral env sigma ts t =
+ let (f, l) = decompose_app_vect sigma t in
+ match EConstr.kind sigma f with
| Const (c, u) ->
not (Environ.evaluable_constant c env) ||
not (is_transparent env (ConstKey c)) ||
@@ -622,24 +643,27 @@ let rec is_neutral env ts t =
not (Id.Pred.mem id (fst ts))
| Rel n -> true
| Evar _ | Meta _ -> true
- | Case (_, p, c, cl) -> is_neutral env ts c
- | Proj (p, c) -> is_neutral env ts c
- | _ -> false
-
-let is_eta_constructor_app env ts f l1 term =
- match kind_of_term f with
+ | Case (_, p, c, cl) -> is_neutral env sigma ts c
+ | Proj (p, c) -> is_neutral env sigma ts c
+ | Lambda _ | LetIn _ | Construct _ | CoFix _ -> false
+ | Sort _ | Cast (_, _, _) | Prod (_, _, _) | Ind _ -> false (* Really? *)
+ | Fix _ -> false (* This is an approximation *)
+ | App _ -> assert false
+
+let is_eta_constructor_app env sigma ts f l1 term =
+ match EConstr.kind sigma f with
| Construct (((_, i as ind), j), u) when i == 0 && j == 1 ->
let mib = lookup_mind (fst ind) env in
(match mib.Declarations.mind_record with
- | Some (Some (_,exp,projs)) when mib.Declarations.mind_finite == Decl_kinds.BiFinite &&
+ | Some (Some (_,exp,projs)) when mib.Declarations.mind_finite == Declarations.BiFinite &&
Array.length projs == Array.length l1 - mib.Declarations.mind_nparams ->
(** Check that the other term is neutral *)
- is_neutral env ts term
+ is_neutral env sigma ts term
| _ -> false)
| _ -> false
-let eta_constructor_app env f l1 term =
- match kind_of_term f with
+let eta_constructor_app env sigma f l1 term =
+ match EConstr.kind sigma f with
| Construct (((_, i as ind), j), u) ->
let mib = lookup_mind (fst ind) env in
(match mib.Declarations.mind_record with
@@ -652,15 +676,15 @@ let eta_constructor_app env f l1 term =
| _ -> assert false)
| _ -> assert false
-let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb flags m n =
- let rec unirec_rec (curenv,nb as curenvnb) pb opt ((sigma,metasubst,evarsubst) as substn) curm curn =
+let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top env cv_pb flags m n =
+ let rec unirec_rec (curenv,nb as curenvnb) pb opt ((sigma,metasubst,evarsubst) as substn : subst0) curm curn =
let cM = Evarutil.whd_head_evar sigma curm
and cN = Evarutil.whd_head_evar sigma curn in
let () =
if !debug_unification then
- Feedback.msg_debug (Termops.print_constr_env curenv cM ++ str" ~= " ++ Termops.print_constr_env curenv cN)
+ Feedback.msg_debug (print_constr_env curenv sigma cM ++ str" ~= " ++ print_constr_env curenv sigma cN)
in
- match (kind_of_term cM,kind_of_term cN) with
+ match (EConstr.kind sigma cM, EConstr.kind sigma cN) with
| Meta k1, Meta k2 ->
if Int.equal k1 k2 then substn else
let stM,stN = extract_instance_status pb in
@@ -675,7 +699,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
if k2 < k1 then sigma,(k1,cN,stN)::metasubst,evarsubst
else sigma,(k2,cM,stM)::metasubst,evarsubst
| Meta k, _
- when not (dependent cM cN) (* helps early trying alternatives *) ->
+ when not (dependent sigma cM cN) (* helps early trying alternatives *) ->
let sigma =
if opt.with_types && flags.check_applied_meta_types then
(try
@@ -689,13 +713,13 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
(* Here we check that [cN] does not contain any local variables *)
if Int.equal nb 0 then
sigma,(k,cN,snd (extract_instance_status pb))::metasubst,evarsubst
- else if noccur_between 1 nb cN then
+ else if noccur_between sigma 1 nb cN then
(sigma,
(k,lift (-nb) cN,snd (extract_instance_status pb))::metasubst,
evarsubst)
else error_cannot_unify_local curenv sigma (m,n,cN)
| _, Meta k
- when not (dependent cN cM) (* helps early trying alternatives *) ->
+ when not (dependent sigma cN cM) (* helps early trying alternatives *) ->
let sigma =
if opt.with_types && flags.check_applied_meta_types then
(try
@@ -709,7 +733,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
(* Here we check that [cM] does not contain any local variables *)
if Int.equal nb 0 then
(sigma,(k,cM,fst (extract_instance_status pb))::metasubst,evarsubst)
- else if noccur_between 1 nb cM
+ else if noccur_between sigma 1 nb cM
then
(sigma,(k,lift (-nb) cM,fst (extract_instance_status pb))::metasubst,
evarsubst)
@@ -717,27 +741,29 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
| Evar (evk,_ as ev), Evar (evk',_)
when not (Evar.Set.mem evk flags.frozen_evars)
&& Evar.equal evk evk' ->
- let sigma',b = constr_cmp cv_pb sigma flags cM cN in
+ let sigma',b = constr_cmp cv_pb env sigma flags cM cN in
if b then
sigma',metasubst,evarsubst
else
sigma,metasubst,((curenv,ev,cN)::evarsubst)
| Evar (evk,_ as ev), _
when not (Evar.Set.mem evk flags.frozen_evars)
- && not (occur_evar evk cN) ->
- let cmvars = free_rels cM and cnvars = free_rels cN in
+ && not (occur_evar sigma evk cN) ->
+ let cmvars = free_rels sigma cM and cnvars = free_rels sigma cN in
if Int.Set.subset cnvars cmvars then
sigma,metasubst,((curenv,ev,cN)::evarsubst)
else error_cannot_unify_local curenv sigma (m,n,cN)
| _, Evar (evk,_ as ev)
when not (Evar.Set.mem evk flags.frozen_evars)
- && not (occur_evar evk cM) ->
- let cmvars = free_rels cM and cnvars = free_rels cN in
+ && not (occur_evar sigma evk cM) ->
+ let cmvars = free_rels sigma cM and cnvars = free_rels sigma cN in
if Int.Set.subset cmvars cnvars then
sigma,metasubst,((curenv,ev,cM)::evarsubst)
else error_cannot_unify_local curenv sigma (m,n,cN)
| Sort s1, Sort s2 ->
(try
+ let s1 = ESorts.kind sigma s1 in
+ let s2 = ESorts.kind sigma s2 in
let sigma' =
if pb == CUMUL
then Evd.set_leq_sort curenv sigma s1 s2
@@ -756,7 +782,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
| _, LetIn (_,a,_,c) -> unirec_rec curenvnb pb opt substn cM (subst1 a c)
(** Fast path for projections. *)
- | Proj (p1,c1), Proj (p2,c2) when eq_constant
+ | Proj (p1,c1), Proj (p2,c2) when Constant.equal
(Projection.constant p1) (Projection.constant p2) ->
(try unify_same_proj curenvnb cv_pb {opt with at_top = true}
substn c1 c2
@@ -775,30 +801,30 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
| App (f1, l1), _ when flags.modulo_eta &&
(* This ensures cN is an evar, meta or irreducible constant/variable
and not a constructor. *)
- is_eta_constructor_app curenv flags.modulo_delta f1 l1 cN ->
+ is_eta_constructor_app curenv sigma flags.modulo_delta f1 l1 cN ->
(try
- let l1', l2' = eta_constructor_app curenv f1 l1 cN in
+ let l1', l2' = eta_constructor_app curenv sigma f1 l1 cN in
let opt' = {opt with at_top = true; with_cs = false} in
Array.fold_left2 (unirec_rec curenvnb CONV opt') substn l1' l2'
with ex when precatchable_exception ex ->
- match kind_of_term cN with
+ match EConstr.kind sigma cN with
| App(f2,l2) when
- (isMeta f2 && use_metas_pattern_unification flags nb l2
- || use_evars_pattern_unification flags && isAllowedEvar flags f2) ->
+ (isMeta sigma f2 && use_metas_pattern_unification sigma flags nb l2
+ || use_evars_pattern_unification flags && isAllowedEvar sigma flags f2) ->
unify_app_pattern false curenvnb pb opt substn cM f1 l1 cN f2 l2
| _ -> raise ex)
| _, App (f2, l2) when flags.modulo_eta &&
- is_eta_constructor_app curenv flags.modulo_delta f2 l2 cM ->
+ is_eta_constructor_app curenv sigma flags.modulo_delta f2 l2 cM ->
(try
- let l2', l1' = eta_constructor_app curenv f2 l2 cM in
+ let l2', l1' = eta_constructor_app curenv sigma f2 l2 cM in
let opt' = {opt with at_top = true; with_cs = false} in
Array.fold_left2 (unirec_rec curenvnb CONV opt') substn l1' l2'
with ex when precatchable_exception ex ->
- match kind_of_term cM with
+ match EConstr.kind sigma cM with
| App(f1,l1) when
- (isMeta f1 && use_metas_pattern_unification flags nb l1
- || use_evars_pattern_unification flags && isAllowedEvar flags f1) ->
+ (isMeta sigma f1 && use_metas_pattern_unification sigma flags nb l1
+ || use_evars_pattern_unification flags && isAllowedEvar sigma flags f1) ->
unify_app_pattern true curenvnb pb opt substn cM f1 l1 cN f2 l2
| _ -> raise ex)
@@ -813,13 +839,13 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
reduce curenvnb pb opt substn cM cN)
| App (f1,l1), _ when
- (isMeta f1 && use_metas_pattern_unification flags nb l1
- || use_evars_pattern_unification flags && isAllowedEvar flags f1) ->
+ (isMeta sigma f1 && use_metas_pattern_unification sigma flags nb l1
+ || use_evars_pattern_unification flags && isAllowedEvar sigma flags f1) ->
unify_app_pattern true curenvnb pb opt substn cM f1 l1 cN cN [||]
| _, App (f2,l2) when
- (isMeta f2 && use_metas_pattern_unification flags nb l2
- || use_evars_pattern_unification flags && isAllowedEvar flags f2) ->
+ (isMeta sigma f2 && use_metas_pattern_unification sigma flags nb l2
+ || use_evars_pattern_unification flags && isAllowedEvar sigma flags f2) ->
unify_app_pattern false curenvnb pb opt substn cM cM [||] cN f2 l2
| App (f1,l1), App (f2,l2) ->
@@ -834,11 +860,11 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
| _ ->
unify_not_same_head curenvnb pb opt substn cM cN
- and unify_app_pattern dir curenvnb pb opt substn cM f1 l1 cN f2 l2 =
+ and unify_app_pattern dir curenvnb pb opt (sigma, _, _ as substn) cM f1 l1 cN f2 l2 =
let f, l, t = if dir then f1, l1, cN else f2, l2, cM in
match is_unification_pattern curenvnb sigma f (Array.to_list l) t with
| None ->
- (match kind_of_term t with
+ (match EConstr.kind sigma t with
| App (f',l') ->
if dir then unify_app curenvnb pb opt substn cM f1 l1 t f' l'
else unify_app curenvnb pb opt substn t f' l' cN f2 l2
@@ -847,19 +873,19 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
| Some l ->
solve_pattern_eqn_array curenvnb f l t substn
- and unify_app (curenv, nb as curenvnb) pb opt (sigma, metas, evars as substn) cM f1 l1 cN f2 l2 =
+ and unify_app (curenv, nb as curenvnb) pb opt (sigma, metas, evars as substn : subst0) cM f1 l1 cN f2 l2 =
try
let needs_expansion p c' =
- match kind_of_term c' with
+ match EConstr.kind sigma c' with
| Meta _ -> true
| Evar _ -> true
| Const (c, u) -> Constant.equal c (Projection.constant p)
| _ -> false
in
let expand_proj c c' l =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Proj (p, t) when not (Projection.unfolded p) && needs_expansion p c' ->
- (try destApp (Retyping.expand_projection curenv sigma p t (Array.to_list l))
+ (try destApp sigma (Retyping.expand_projection curenv sigma p t (Array.to_list l))
with RetypeError _ -> (** Unification can be called on ill-typed terms, due
to FO and eta in particular, fail gracefully in that case *)
(c, l))
@@ -894,33 +920,33 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
ty1 ty2
with RetypeError _ -> substn
- and unify_not_same_head curenvnb pb opt (sigma, metas, evars as substn) cM cN =
+ and unify_not_same_head curenvnb pb opt (sigma, metas, evars as substn : subst0) cM cN =
try canonical_projections curenvnb pb opt cM cN substn
with ex when precatchable_exception ex ->
- let sigma', b = constr_cmp cv_pb sigma flags cM cN in
+ let sigma', b = constr_cmp cv_pb env sigma flags cM cN in
if b then (sigma', metas, evars)
else
try reduce curenvnb pb opt substn cM cN
with ex when precatchable_exception ex ->
let (f1,l1) =
- match kind_of_term cM with App (f,l) -> (f,l) | _ -> (cM,[||]) in
+ match EConstr.kind sigma cM with App (f,l) -> (f,l) | _ -> (cM,[||]) in
let (f2,l2) =
- match kind_of_term cN with App (f,l) -> (f,l) | _ -> (cN,[||]) in
+ match EConstr.kind sigma cN with App (f,l) -> (f,l) | _ -> (cN,[||]) in
expand curenvnb pb opt substn cM f1 l1 cN f2 l2
and reduce curenvnb pb opt (sigma, metas, evars as substn) cM cN =
- if use_full_betaiota flags && not (subterm_restriction opt flags) then
+ if flags.modulo_betaiota && not (subterm_restriction opt flags) then
let cM' = do_reduce flags.modulo_delta curenvnb sigma cM in
- if not (Term.eq_constr cM cM') then
+ if not (EConstr.eq_constr sigma cM cM') then
unirec_rec curenvnb pb opt substn cM' cN
else
let cN' = do_reduce flags.modulo_delta curenvnb sigma cN in
- if not (Term.eq_constr cN cN') then
+ if not (EConstr.eq_constr sigma cN cN') then
unirec_rec curenvnb pb opt substn cM cN'
else error_cannot_unify (fst curenvnb) sigma (cM,cN)
else error_cannot_unify (fst curenvnb) sigma (cM,cN)
- and expand (curenv,_ as curenvnb) pb opt (sigma,metasubst,evarsubst as substn) cM f1 l1 cN f2 l2 =
+ and expand (curenv,_ as curenvnb) pb opt (sigma,metasubst,evarsubst as substn : subst0) cM f1 l1 cN f2 l2 =
let res =
(* Try full conversion on meta-free terms. *)
(* Back to 1995 (later on called trivial_unify in 2002), the
@@ -939,10 +965,10 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
| None -> None
| Some convflags ->
let subst = ((if flags.use_metas_eagerly_in_conv_on_closed_terms then metasubst else ms), (if flags.use_evars_eagerly_in_conv_on_closed_terms then evarsubst else es)) in
- match subst_defined_metas_evars subst cM with
+ match subst_defined_metas_evars sigma subst cM with
| None -> (* some undefined Metas in cM *) None
| Some m1 ->
- match subst_defined_metas_evars subst cN with
+ match subst_defined_metas_evars sigma subst cN with
| None -> (* some undefined Metas in cN *) None
| Some n1 ->
(* No subterm restriction there, too much incompatibilities *)
@@ -966,7 +992,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
match res with
| Some substn -> substn
| None ->
- let cf1 = key_of curenv opt flags f1 and cf2 = key_of curenv opt flags f2 in
+ let cf1 = key_of curenv sigma opt flags f1 and cf2 = key_of curenv sigma opt flags f2 in
match oracle_order curenv cf1 cf2 with
| None -> error_cannot_unify curenv sigma (cM,cN)
| Some true ->
@@ -996,7 +1022,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
and canonical_projections (curenv, _ as curenvnb) pb opt cM cN (sigma,_,_ as substn) =
let f1 () =
- if isApp cM then
+ if isApp_or_Proj sigma cM then
let f1l1 = whd_nored_state sigma (cM,Stack.empty) in
if is_open_canonical_projection curenv sigma f1l1 then
let f2l2 = whd_nored_state sigma (cN,Stack.empty) in
@@ -1012,7 +1038,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
error_cannot_unify (fst curenvnb) sigma (cM,cN)
else
try f1 () with e when precatchable_exception e ->
- if isApp cN then
+ if isApp_or_Proj sigma cN then
let f2l2 = whd_nored_state sigma (cN, Stack.empty) in
if is_open_canonical_projection curenv sigma f2l2 then
let f1l1 = whd_nored_state sigma (cM, Stack.empty) in
@@ -1034,23 +1060,23 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
(evd,t2::ks, m-1)
else
let mv = new_meta () in
- let evd' = meta_declare mv (substl ks b) evd in
+ let evd' = meta_declare mv (EConstr.Unsafe.to_constr (substl ks b)) evd in
(evd', mkMeta mv :: ks, m - 1))
(sigma,[],List.length bs) bs
in
try
let opt' = {opt with with_types = false} in
- let (substn,_,_) = Reductionops.Stack.fold2
+ let substn = Reductionops.Stack.fold2
(fun s u1 u -> unirec_rec curenvnb pb opt' s u1 (substl ks u))
(evd,ms,es) us2 us in
- let (substn,_,_) = Reductionops.Stack.fold2
+ let substn = Reductionops.Stack.fold2
(fun s u1 u -> unirec_rec curenvnb pb opt' s u1 (substl ks u))
substn params1 params in
- let (substn,_,_) = Reductionops.Stack.fold2 (unirec_rec curenvnb pb opt') substn ts ts1 in
+ let substn = Reductionops.Stack.fold2 (fun s u1 u2 -> unirec_rec curenvnb pb opt' s u1 u2) substn ts ts1 in
let app = mkApp (c, Array.rev_of_list ks) in
(* let substn = unirec_rec curenvnb pb b false substn t cN in *)
unirec_rec curenvnb pb opt' substn c1 app
- with Invalid_argument "Reductionops.Stack.fold2" ->
+ with Reductionops.Stack.IncompatibleFold2 ->
error_cannot_unify (fst curenvnb) sigma (cM,cN)
else error_cannot_unify (fst curenvnb) sigma (cM,cN)
in
@@ -1066,7 +1092,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
else
let sigma, b = match flags.modulo_conv_on_closed_terms with
| Some convflags -> infer_conv ~pb:cv_pb ~ts:convflags env sigma m n
- | _ -> constr_cmp cv_pb sigma flags m n in
+ | _ -> constr_cmp cv_pb env sigma flags m n in
if b then Some sigma
else if (match flags.modulo_conv_on_closed_terms, flags.modulo_delta with
| Some (cv_id, cv_k), (dl_id, dl_k) ->
@@ -1093,7 +1119,7 @@ let right = false
let rec unify_with_eta keptside flags env sigma c1 c2 =
(* Question: try whd_all on ci if not two lambdas? *)
- match kind_of_term c1, kind_of_term c2 with
+ match EConstr.kind sigma c1, EConstr.kind sigma c2 with
| (Lambda (na,t1,c1'), Lambda (_,t2,c2')) ->
let env' = push_rel_assum (na,t1) env in
let sigma,metas,evars = unify_0 env sigma CONV flags t1 t2 in
@@ -1196,22 +1222,22 @@ let merge_instances env sigma flags st1 st2 c1 c2 =
* close it off. But this might not always work,
* since other metavars might also need to be resolved. *)
-let applyHead env (type r) (evd : r Sigma.t) n c =
- let rec apprec : type s. _ -> _ -> _ -> (r, s) Sigma.le -> s Sigma.t -> (constr, r) Sigma.sigma =
- fun n c cty p evd ->
+let applyHead env evd n c =
+ let rec apprec n c cty evd =
if Int.equal n 0 then
- Sigma (c, evd, p)
+ (evd, c)
else
- match kind_of_term (whd_all env (Sigma.to_evar_map evd) cty) with
+ match EConstr.kind evd (whd_all env evd cty) with
| Prod (_,c1,c2) ->
- let Sigma (evar, evd', q) = Evarutil.new_evar env evd ~src:(Loc.ghost,Evar_kinds.GoalEvar) c1 in
- apprec (n-1) (mkApp(c,[|evar|])) (subst1 evar c2) (p +> q) evd'
- | _ -> error "Apply_Head_Then"
+ let (evd',evar) =
+ Evarutil.new_evar env evd ~src:(Loc.tag Evar_kinds.GoalEvar) c1 in
+ apprec (n-1) (mkApp(c,[|evar|])) (subst1 evar c2) evd'
+ | _ -> user_err Pp.(str "Apply_Head_Then")
in
- apprec n c (Typing.unsafe_type_of env (Sigma.to_evar_map evd) c) Sigma.refl evd
+ apprec n c (Typing.unsafe_type_of env evd c) evd
-let is_mimick_head ts f =
- match kind_of_term f with
+let is_mimick_head sigma ts f =
+ match EConstr.kind sigma f with
| Const (c,u) -> not (CClosure.is_transparent_constant ts c)
| Var id -> not (CClosure.is_transparent_variable ts id)
| (Rel _|Construct _|Ind _) -> true
@@ -1219,9 +1245,9 @@ let is_mimick_head ts f =
let try_to_coerce env evd c cty tycon =
let j = make_judge c cty in
- let (evd',j') = inh_conv_coerce_rigid_to true Loc.ghost env evd j tycon in
+ let (evd',j') = inh_conv_coerce_rigid_to true env evd j tycon in
let evd' = Evarconv.solve_unif_constraints_with_heuristics env evd' in
- let evd' = Evd.map_metas_fvalue (nf_evar evd') evd' in
+ let evd' = Evd.map_metas_fvalue (fun c -> EConstr.Unsafe.to_constr (nf_evar evd' (EConstr.of_constr c))) evd' in
(evd',j'.uj_val)
let w_coerce_to_type env evd c cty mvty =
@@ -1242,7 +1268,7 @@ let w_coerce env evd mv c =
let unify_to_type env sigma flags c status u =
let sigma, c = refresh_universes (Some false) env sigma c in
let t = get_type_of env sigma (nf_meta sigma c) in
- let t = nf_betaiota sigma (nf_meta sigma t) in
+ let t = nf_betaiota env sigma (nf_meta sigma t) in
unify_0 env sigma CUMUL flags t u
let unify_type env sigma flags mv status c =
@@ -1269,36 +1295,31 @@ let solve_simple_evar_eqn ts env evd ev rhs =
match solve_simple_eqn (Evarconv.evar_conv_x ts) env evd (None,ev,rhs) with
| UnifFailure (evd,reason) ->
error_cannot_unify env evd ~reason (mkEvar ev,rhs);
- | Success evd ->
- if Flags.version_less_or_equal Flags.V8_5 then
- (* We used to force solving unrelated problems at arbitrary times *)
- Evarconv.solve_unif_constraints_with_heuristics env evd
- else (* solve_simple_eqn calls reconsider_unif_constraints itself *)
- evd
+ | Success evd -> evd
(* [w_merge env sigma b metas evars] merges common instances in metas
or in evars, possibly generating new unification problems; if [b]
is true, unification of types of metas is required *)
-let w_merge env with_types flags (evd,metas,evars) =
+let w_merge env with_types flags (evd,metas,evars : subst0) =
let rec w_merge_rec evd metas evars eqns =
(* Process evars *)
match evars with
| (curenv,(evk,_ as ev),rhs)::evars' ->
if Evd.is_defined evd evk then
- let v = Evd.existential_value evd ev in
+ let v = mkEvar ev in
let (evd,metas',evars'') =
unify_0 curenv evd CONV flags rhs v in
w_merge_rec evd (metas'@metas) (evars''@evars') eqns
else begin
(* This can make rhs' ill-typed if metas are *)
- let rhs' = subst_meta_instances metas rhs in
- match kind_of_term rhs with
- | App (f,cl) when occur_meta rhs' ->
- if occur_evar evk rhs' then
+ let rhs' = subst_meta_instances evd metas rhs in
+ match EConstr.kind evd rhs with
+ | App (f,cl) when occur_meta evd rhs' ->
+ if occur_evar evd evk rhs' then
error_occur_check curenv evd evk rhs';
- if is_mimick_head flags.modulo_delta f then
+ if is_mimick_head evd flags.modulo_delta f then
let evd' =
mimick_undefined_evar evd flags f (Array.length cl) evk in
w_merge_rec evd' metas evars eqns
@@ -1338,20 +1359,20 @@ let w_merge env with_types flags (evd,metas,evars) =
if meta_defined evd mv then
let {rebus=c'},(status',_) = meta_fvalue evd mv in
let (take_left,st,(evd,metas',evars')) =
- merge_instances env evd flags status' status c' c
+ merge_instances env evd flags status' status (EConstr.of_constr c') c
in
let evd' =
if take_left then evd
- else meta_reassign mv (c,(st,TypeProcessed)) evd
+ else meta_reassign mv (EConstr.Unsafe.to_constr c,(st,TypeProcessed)) evd
in
w_merge_rec evd' (metas'@metas@metas'') (evars'@evars'') eqns
else
let evd' =
if occur_meta_evd evd mv c then
- if isMetaOf mv (whd_all env evd c) then evd
+ if isMetaOf evd mv (whd_all env evd c) then evd
else error_cannot_unify env evd (mkMeta mv,c)
else
- meta_assign mv (c,(status,TypeProcessed)) evd in
+ meta_assign mv (EConstr.Unsafe.to_constr c,(status,TypeProcessed)) evd in
w_merge_rec evd' (metas''@metas) evars'' eqns
| [] ->
(* Process type eqns *)
@@ -1371,23 +1392,21 @@ let w_merge env with_types flags (evd,metas,evars) =
and mimick_undefined_evar evd flags hdc nargs sp =
let ev = Evd.find_undefined evd sp in
- let sp_env = Global.env_of_context ev.evar_hyps in
- let evd = Sigma.Unsafe.of_evar_map evd in
- let Sigma (c, evd', _) = applyHead sp_env evd nargs hdc in
- let evd' = Sigma.to_evar_map evd' in
+ let sp_env = Global.env_of_context (evar_filtered_hyps ev) in
+ let (evd', c) = applyHead sp_env evd nargs hdc in
let (evd'',mc,ec) =
unify_0 sp_env evd' CUMUL flags
- (get_type_of sp_env evd' c) ev.evar_concl in
+ (get_type_of sp_env evd' c) (EConstr.of_constr ev.evar_concl) in
let evd''' = w_merge_rec evd'' mc ec [] in
if evd' == evd'''
- then Evd.define sp c evd'''
- else Evd.define sp (Evarutil.nf_evar evd''' c) evd''' in
+ then Evd.define sp (EConstr.Unsafe.to_constr c) evd'''
+ else Evd.define sp (EConstr.Unsafe.to_constr (Evarutil.nf_evar evd''' c)) evd''' in
let check_types evd =
let metas = Evd.meta_list evd in
let eqns = List.fold_left (fun acc (mv, b) ->
match b with
- | Clval (n, (t, (c, TypeNotProcessed)), v) -> (mv, c, t.rebus) :: acc
+ | Clval (n, (t, (c, TypeNotProcessed)), v) -> (mv, c, EConstr.of_constr t.rebus) :: acc
| _ -> acc) [] metas
in w_merge_rec evd [] [] eqns
in
@@ -1398,6 +1417,11 @@ let w_merge env with_types flags (evd,metas,evars) =
in
if with_types then check_types res else res
+let retract_coercible_metas evd =
+ let (metas, evd) = retract_coercible_metas evd in
+ let map (mv, c, st) = (mv, EConstr.of_constr c, st) in
+ (List.map map metas, evd)
+
let w_unify_meta_types env ?(flags=default_unify_flags ()) evd =
let metas,evd = retract_coercible_metas evd in
w_merge env true flags.merge_unify_flags (evd,metas,[])
@@ -1415,13 +1439,17 @@ let w_unify_meta_types env ?(flags=default_unify_flags ()) evd =
let head_app sigma m =
fst (whd_nored_state sigma (m, Stack.empty))
+let isEvar_or_Meta sigma c = match EConstr.kind sigma c with
+| Evar _ | Meta _ -> true
+| _ -> false
+
let check_types env flags (sigma,_,_ as subst) m n =
- if isEvar_or_Meta (head_app sigma m) then
+ if isEvar_or_Meta sigma (head_app sigma m) then
unify_0_with_initial_metas subst true env CUMUL
flags
(get_type_of env sigma n)
(get_type_of env sigma m)
- else if isEvar_or_Meta (head_app sigma n) then
+ else if isEvar_or_Meta sigma (head_app sigma n) then
unify_0_with_initial_metas subst true env CUMUL
flags
(get_type_of env sigma m)
@@ -1463,7 +1491,7 @@ let w_typed_unify_array env evd flags f1 l1 f2 l2 =
let iter_fail f a =
let n = Array.length a in
let rec ffail i =
- if Int.equal i n then error "iter_fail"
+ if Int.equal i n then user_err Pp.(str "iter_fail")
else
try f a.(i)
with ex when precatchable_exception ex -> ffail (i+1)
@@ -1472,22 +1500,18 @@ let iter_fail f a =
(* make_abstraction: a variant of w_unify_to_subterm which works on
contexts, with evars, and possibly with occurrences *)
-let indirectly_dependent c d decls =
- not (isVar c) &&
+let indirectly_dependent sigma c d decls =
+ not (isVar sigma c) &&
(* This test is not needed if the original term is a variable, but
it is needed otherwise, as e.g. when abstracting over "2" in
"forall H:0=2, H=H:>(0=1+1) -> 0=2." where there is now obvious
way to see that the second hypothesis depends indirectly over 2 *)
- List.exists (fun d' -> dependent_in_decl (mkVar (get_id d')) d) decls
-
-let indirect_dependency d decls =
- decls |> List.filter (fun d' -> dependent_in_decl (mkVar (get_id d')) d) |> List.hd |> get_id
+ List.exists (fun d' -> dependent_in_decl sigma (EConstr.mkVar (NamedDecl.get_id d')) d) decls
let finish_evar_resolution ?(flags=Pretyping.all_and_fail_flags) env current_sigma (pending,c) =
- let current_sigma = Sigma.to_evar_map current_sigma in
let sigma = Pretyping.solve_remaining_evars flags env current_sigma pending in
let sigma, subst = nf_univ_variables sigma in
- Sigma.Unsafe.of_pair (subst_univs_constr subst (nf_evar sigma c), sigma)
+ (sigma, EConstr.of_constr (Universes.subst_univs_constr subst (EConstr.Unsafe.to_constr (nf_evar sigma c))))
let default_matching_core_flags sigma =
let ts = Names.full_transparent_state in {
@@ -1516,7 +1540,7 @@ let default_matching_merge_flags sigma =
use_pattern_unification = true;
}
-let default_matching_flags (sigma,_) =
+let default_matching_flags sigma =
let flags = default_matching_core_flags sigma in {
core_unify_flags = flags;
merge_unify_flags = default_matching_merge_flags sigma;
@@ -1539,17 +1563,17 @@ let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) =
modulo_conv_on_closed_terms = Some Names.full_transparent_state;
restrict_conv_on_strict_subterms = true } }
else default_matching_flags pending in
- let n = List.length (snd (decompose_app c)) in
+ let n = Array.length (snd (decompose_app_vect sigma c)) in
let matching_fun _ t =
try
let t',l2 =
if from_prefix_of_ind then
(* We check for fully applied subterms of the form "u u1 .. un" *)
(* of inductive type knowing only a prefix "u u1 .. ui" *)
- let t,l = decompose_app t in
+ let t,l = decompose_app sigma t in
let l1,l2 =
try List.chop n l with Failure _ -> raise (NotUnifiable None) in
- if not (List.for_all closed0 l2) then raise (NotUnifiable None)
+ if not (List.for_all (fun c -> Vars.closed0 sigma c) l2) then raise (NotUnifiable None)
else
applist (t,l1), l2
else t, [] in
@@ -1575,9 +1599,9 @@ let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) =
(fun test -> match test.testing_state with
| None -> None
| Some (sigma,_,l) ->
- let c = applist (nf_evar sigma (local_strong whd_meta sigma c),l) in
+ let c = applist (nf_evar sigma (local_strong whd_meta sigma c), l) in
let univs, subst = nf_univ_variables sigma in
- Some (sigma,subst_univs_constr subst c))
+ Some (sigma,EConstr.of_constr (Universes.subst_univs_constr subst (EConstr.Unsafe.to_constr c))))
let make_eq_test env evd c =
let out cstr =
@@ -1588,27 +1612,28 @@ let make_eq_test env evd c =
let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl =
let id =
let t = match ty with Some t -> t | None -> get_type_of env sigma c in
- let x = id_of_name_using_hdchar (Global.env()) t name in
- let ids = ids_of_named_context (named_context env) in
+ let x = id_of_name_using_hdchar (Global.env()) sigma t name in
+ let ids = Environ.ids_of_named_context_val (named_context_val env) in
if name == Anonymous then next_ident_away_in_goal x ids else
if mem_named_context_val x (named_context_val env) then
- errorlabstrm "Unification.make_abstraction_core"
- (str "The variable " ++ Nameops.pr_id x ++ str " is already declared.")
+ user_err ~hdr:"Unification.make_abstraction_core"
+ (str "The variable " ++ Id.print x ++ str " is already declared.")
else
x
in
let likefirst = clause_with_generic_occurrences occs in
- let mkvarid () = mkVar id in
+ let mkvarid () = EConstr.mkVar id in
let compute_dependency _ d (sign,depdecls) =
- let hyp = get_id d in
+ let d = map_named_decl EConstr.of_constr d in
+ let hyp = NamedDecl.get_id d in
match occurrences_of_hyp hyp occs with
| NoOccurrences, InHyp ->
(push_named_context_val d sign,depdecls)
| AllOccurrences, InHyp as occ ->
let occ = if likefirst then LikeFirst else AtOccs occ in
- let newdecl = replace_term_occ_decl_modulo occ test mkvarid d in
- if Context.Named.Declaration.equal d newdecl
- && not (indirectly_dependent c d depdecls)
+ let newdecl = replace_term_occ_decl_modulo sigma occ test mkvarid d in
+ if Context.Named.Declaration.equal (EConstr.eq_constr sigma) d newdecl
+ && not (indirectly_dependent sigma c d depdecls)
then
if check_occs && not (in_every_hyp occs)
then raise (PretypeError (env,sigma,NoOccurrenceFound (c,Some hyp)))
@@ -1617,7 +1642,7 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl =
(push_named_context_val newdecl sign, newdecl :: depdecls)
| occ ->
(* There are specific occurrences, hence not like first *)
- let newdecl = replace_term_occ_decl_modulo (AtOccs occ) test mkvarid d in
+ let newdecl = replace_term_occ_decl_modulo sigma (AtOccs occ) test mkvarid d in
(push_named_context_val newdecl sign, newdecl :: depdecls) in
try
let sign,depdecls =
@@ -1627,13 +1652,13 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl =
| NoOccurrences -> concl
| occ ->
let occ = if likefirst then LikeFirst else AtOccs occ in
- replace_term_occ_modulo occ test mkvarid concl
+ replace_term_occ_modulo sigma occ test mkvarid concl
in
let lastlhyp =
- if List.is_empty depdecls then None else Some (get_id (List.last depdecls)) in
+ if List.is_empty depdecls then None else Some (NamedDecl.get_id (List.last depdecls)) in
let res = match out test with
| None -> None
- | Some (sigma, c) -> Some (Sigma.Unsafe.of_pair (c, sigma))
+ | Some (sigma, c) -> Some (sigma,c)
in
(id,sign,depdecls,lastlhyp,ccl,res)
with
@@ -1654,16 +1679,15 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl =
type prefix_of_inductive_support_flag = bool
type abstraction_request =
-| AbstractPattern of prefix_of_inductive_support_flag * (types -> bool) * Name.t * pending_constr * clause * bool
+| AbstractPattern of prefix_of_inductive_support_flag * (types -> bool) * Name.t * (evar_map * constr) * clause * bool
| AbstractExact of Name.t * constr * types option * clause * bool
type 'r abstraction_result =
Names.Id.t * named_context_val *
- Context.Named.Declaration.t list * Names.Id.t option *
- types * (constr, 'r) Sigma.sigma option
+ named_declaration list * Names.Id.t option *
+ types * (evar_map * constr) option
let make_abstraction env evd ccl abs =
- let evd = Sigma.to_evar_map evd in
match abs with
| AbstractPattern (from_prefix,check,name,c,occs,check_occs) ->
make_abstraction_core name
@@ -1681,7 +1705,7 @@ let keyed_unify env evd kop =
| None -> fun _ -> true
| Some kop ->
fun cl ->
- let kc = Keys.constr_key cl in
+ let kc = Keys.constr_key (fun c -> EConstr.kind evd c) cl in
match kc with
| None -> false
| Some kc -> Keys.equiv_keys kop kc
@@ -1691,22 +1715,22 @@ let keyed_unify env evd kop =
Fails if no match is found *)
let w_unify_to_subterm env evd ?(flags=default_unify_flags ()) (op,cl) =
let bestexn = ref None in
- let kop = Keys.constr_key op in
+ let kop = Keys.constr_key (fun c -> EConstr.kind evd c) op in
let rec matchrec cl =
- let cl = strip_outer_cast cl in
+ let cl = strip_outer_cast evd cl in
(try
- if closed0 cl && not (isEvar cl) && keyed_unify env evd kop cl then
+ if closed0 evd cl && not (isEvar evd cl) && keyed_unify env evd kop cl then
(try
if !keyed_unification then
- let f1, l1 = decompose_app_vect op in
- let f2, l2 = decompose_app_vect cl in
+ let f1, l1 = decompose_app_vect evd op in
+ let f2, l2 = decompose_app_vect evd cl in
w_typed_unify_array env evd flags f1 l1 f2 l2,cl
else w_typed_unify env evd CONV flags op cl,cl
with ex when Pretype_errors.unsatisfiable_exception ex ->
- bestexn := Some ex; error "Unsat")
- else error "Bound 1"
+ bestexn := Some ex; user_err Pp.(str "Unsat"))
+ else user_err Pp.(str "Bound 1")
with ex when precatchable_exception ex ->
- (match kind_of_term cl with
+ (match EConstr.kind evd cl with
| App (f,args) ->
let n = Array.length args in
assert (n>0);
@@ -1753,7 +1777,9 @@ let w_unify_to_subterm env evd ?(flags=default_unify_flags ()) (op,cl) =
with ex when precatchable_exception ex ->
matchrec c)
- | _ -> error "Match_subterm"))
+ | Cast (_, _, _) (* Is this expected? *)
+ | Rel _ | Var _ | Meta _ | Evar _ | Sort _ | Const _ | Ind _
+ | Construct _ -> user_err Pp.(str "Match_subterm")))
in
try matchrec cl
with ex when precatchable_exception ex ->
@@ -1767,9 +1793,9 @@ let w_unify_to_subterm env evd ?(flags=default_unify_flags ()) (op,cl) =
let w_unify_to_subterm_all env evd ?(flags=default_unify_flags ()) (op,cl) =
let return a b =
let (evd,c as a) = a () in
- if List.exists (fun (evd',c') -> Term.eq_constr c c') b then b else a :: b
+ if List.exists (fun (evd',c') -> EConstr.eq_constr evd' c c') b then b else a :: b
in
- let fail str _ = error str in
+ let fail str _ = user_err (Pp.str str) in
let bind f g a =
let a1 = try f a
with ex
@@ -1786,12 +1812,12 @@ let w_unify_to_subterm_all env evd ?(flags=default_unify_flags ()) (op,cl) =
in ffail 0
in
let rec matchrec cl =
- let cl = strip_outer_cast cl in
+ let cl = strip_outer_cast evd cl in
(bind
- (if closed0 cl
+ (if closed0 evd cl
then return (fun () -> w_typed_unify env evd CONV flags op cl,cl)
else fail "Bound 1")
- (match kind_of_term cl with
+ (match EConstr.kind evd cl with
| App (f,args) ->
let n = Array.length args in
assert (n>0);
@@ -1819,7 +1845,11 @@ let w_unify_to_subterm_all env evd ?(flags=default_unify_flags ()) (op,cl) =
| Lambda (_,t,c) ->
bind (matchrec t) (matchrec c)
- | _ -> fail "Match_subterm"))
+ | Cast (_, _, _) -> fail "Match_subterm" (* Is this expected? *)
+
+ | Rel _ | Var _ | Meta _ | Evar _ | Sort _ | Const _ | Ind _
+ | Construct _ -> fail "Match_subterm"))
+
in
let res = matchrec cl [] in
match res with
@@ -1831,13 +1861,13 @@ let w_unify_to_subterm_list env evd flags hdmeta oplist t =
List.fold_right
(fun op (evd,l) ->
let op = whd_meta evd op in
- if isMeta op then
+ if isMeta evd op then
if flags.allow_K_in_toplevel_higher_order_unification then (evd,op::l)
- else error_abstraction_over_meta env evd hdmeta (destMeta op)
+ else error_abstraction_over_meta env evd hdmeta (destMeta evd op)
else
let allow_K = flags.allow_K_in_toplevel_higher_order_unification in
let flags =
- if occur_meta_or_existential op || !keyed_unification then
+ if unsafe_occur_meta_or_existential op || !keyed_unification then
(* This is up to delta for subterms w/o metas ... *)
flags
else
@@ -1846,7 +1876,7 @@ let w_unify_to_subterm_list env evd flags hdmeta oplist t =
unify pre-existing non frozen evars of the goal or of the
pattern *)
set_no_delta_flags flags in
- let t' = (strip_outer_cast op,t) in
+ let t' = (strip_outer_cast evd op,t) in
let (evd',cl) =
try
if is_keyed_unification () then
@@ -1862,11 +1892,11 @@ let w_unify_to_subterm_list env evd flags hdmeta oplist t =
(* w_unify_to_subterm does not go through evars, so
the next step, which was already in <= 8.4, is
needed at least for compatibility of rewrite *)
- dependent op t -> (evd,op)
+ dependent evd op t -> (evd,op)
in
if not allow_K &&
(* ensure we found a different instance *)
- List.exists (fun op -> Term.eq_constr op cl) l
+ List.exists (fun op -> EConstr.eq_constr evd' op cl) l
then error_non_linear_unification env evd hdmeta cl
else (evd',cl::l))
oplist
@@ -1898,14 +1928,14 @@ let secondOrderAbstractionAlgo dep =
let w_unify2 env evd flags dep cv_pb ty1 ty2 =
let c1, oplist1 = whd_nored_stack evd ty1 in
let c2, oplist2 = whd_nored_stack evd ty2 in
- match kind_of_term c1, kind_of_term c2 with
+ match EConstr.kind evd c1, EConstr.kind evd c2 with
| Meta p1, _ ->
(* Find the predicate *)
- secondOrderAbstractionAlgo dep env evd flags ty2 (p1,oplist1)
+ secondOrderAbstractionAlgo dep env evd flags ty2 (p1, oplist1)
| _, Meta p2 ->
(* Find the predicate *)
secondOrderAbstractionAlgo dep env evd flags ty1 (p2, oplist2)
- | _ -> error "w_unify2"
+ | _ -> user_err Pp.(str "w_unify2")
(* The unique unification algorithm works like this: If the pattern is
flexible, and the goal has a lambda-abstraction at the head, then
@@ -1928,11 +1958,11 @@ let w_unify2 env evd flags dep cv_pb ty1 ty2 =
convertible and first-order otherwise. But if failed if e.g. the type of
Meta(1) had meta-variables in it. *)
let w_unify env evd cv_pb ?(flags=default_unify_flags ()) ty1 ty2 =
- let hd1,l1 = decompose_appvect (whd_nored evd ty1) in
- let hd2,l2 = decompose_appvect (whd_nored evd ty2) in
+ let hd1,l1 = decompose_app_vect evd (whd_nored evd ty1) in
+ let hd2,l2 = decompose_app_vect evd (whd_nored evd ty2) in
let is_empty1 = Array.is_empty l1 in
let is_empty2 = Array.is_empty l2 in
- match kind_of_term hd1, not is_empty1, kind_of_term hd2, not is_empty2 with
+ match EConstr.kind evd hd1, not is_empty1, EConstr.kind evd hd2, not is_empty2 with
(* Pattern case *)
| (Meta _, true, Lambda _, _ | Lambda _, _, Meta _, true)
when Int.equal (Array.length l1) (Array.length l2) ->
@@ -1969,8 +1999,8 @@ let w_unify env evd cv_pb flags ty1 ty2 =
let w_unify =
if Flags.profile then
- let wunifkey = Profile.declare_profile "w_unify" in
- Profile.profile6 wunifkey w_unify
+ let wunifkey = CProfile.declare_profile "w_unify" in
+ CProfile.profile6 wunifkey w_unify
else w_unify
let w_unify env evd cv_pb ?(flags=default_unify_flags ()) ty1 ty2 =
diff --git a/pretyping/unification.mli b/pretyping/unification.mli
index 0ad882a9..16ce5c93 100644
--- a/pretyping/unification.mli
+++ b/pretyping/unification.mli
@@ -1,12 +1,15 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
+open Constr
+open EConstr
open Environ
open Evd
@@ -44,7 +47,7 @@ val elim_no_delta_flags : unit -> unify_flags
val is_keyed_unification : unit -> bool
-(** The "unique" unification fonction *)
+(** The "unique" unification function *)
val w_unify :
env -> evar_map -> conv_pb -> ?flags:unify_flags -> constr -> constr -> evar_map
@@ -71,18 +74,18 @@ exception PatternNotFound
type prefix_of_inductive_support_flag = bool
type abstraction_request =
-| AbstractPattern of prefix_of_inductive_support_flag * (types -> bool) * Names.Name.t * pending_constr * Locus.clause * bool
+| AbstractPattern of prefix_of_inductive_support_flag * (types -> bool) * Names.Name.t * (evar_map * constr) * Locus.clause * bool
| AbstractExact of Names.Name.t * constr * types option * Locus.clause * bool
val finish_evar_resolution : ?flags:Pretyping.inference_flags ->
- env -> 'r Sigma.t -> pending_constr -> (constr, 'r) Sigma.sigma
+ env -> evar_map -> (evar_map * constr) -> evar_map * constr
type 'r abstraction_result =
Names.Id.t * named_context_val *
- Context.Named.Declaration.t list * Names.Id.t option *
- types * (constr, 'r) Sigma.sigma option
+ named_declaration list * Names.Id.t option *
+ types * (evar_map * constr) option
-val make_abstraction : env -> 'r Sigma.t -> constr ->
+val make_abstraction : env -> evar_map -> constr ->
abstraction_request -> 'r abstraction_result
val pose_all_metas_as_evars : env -> evar_map -> constr -> evar_map * constr
@@ -97,28 +100,29 @@ val abstract_list_all :
(* For tracing *)
-val w_merge : env -> bool -> core_unify_flags -> evar_map *
- (metavariable * constr * (instance_constraint * instance_typing_status)) list *
- (env * types pexistential * types) list -> evar_map
+type metabinding = (metavariable * constr * (instance_constraint * instance_typing_status))
+
+type subst0 =
+ (evar_map *
+ metabinding list *
+ (Environ.env * existential * t) list)
+
+val w_merge : env -> bool -> core_unify_flags -> subst0 -> evar_map
val unify_0 : Environ.env ->
Evd.evar_map ->
Evd.conv_pb ->
core_unify_flags ->
- Term.types ->
- Term.types ->
- Evd.evar_map * Evd.metabinding list *
- (Environ.env * Term.types Term.pexistential * Term.constr) list
+ types ->
+ types ->
+ subst0
val unify_0_with_initial_metas :
- Evd.evar_map * Evd.metabinding list *
- (Environ.env * Term.types Term.pexistential * Term.constr) list ->
+ subst0 ->
bool ->
Environ.env ->
Evd.conv_pb ->
core_unify_flags ->
- Term.types ->
- Term.types ->
- Evd.evar_map * Evd.metabinding list *
- (Environ.env * Term.types Term.pexistential * Term.constr) list
-
+ types ->
+ types ->
+ subst0
diff --git a/pretyping/univdecls.ml b/pretyping/univdecls.ml
new file mode 100644
index 00000000..8864be57
--- /dev/null
+++ b/pretyping/univdecls.ml
@@ -0,0 +1,52 @@
+(************************************************************************)
+(* * 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) *)
+(************************************************************************)
+
+open CErrors
+
+(** Local universes and constraints declarations *)
+type universe_decl =
+ (Misctypes.lident list, Univ.Constraint.t) Misctypes.gen_universe_decl
+
+let default_univ_decl =
+ let open Misctypes in
+ { univdecl_instance = [];
+ univdecl_extensible_instance = true;
+ univdecl_constraints = Univ.Constraint.empty;
+ univdecl_extensible_constraints = true }
+
+let interp_univ_constraints env evd cstrs =
+ let interp (evd,cstrs) (u, d, u') =
+ let ul = Pretyping.interp_known_glob_level evd u in
+ let u'l = Pretyping.interp_known_glob_level evd u' in
+ let cstr = (ul,d,u'l) in
+ let cstrs' = Univ.Constraint.add cstr cstrs in
+ try let evd = Evd.add_constraints evd (Univ.Constraint.singleton cstr) in
+ evd, cstrs'
+ with Univ.UniverseInconsistency e ->
+ user_err ~hdr:"interp_constraint"
+ (Univ.explain_universe_inconsistency (Termops.pr_evd_level evd) e)
+ in
+ List.fold_left interp (evd,Univ.Constraint.empty) cstrs
+
+let interp_univ_decl env decl =
+ let open Misctypes in
+ let pl : lident list = decl.univdecl_instance in
+ let evd = Evd.from_ctx (UState.make_with_initial_binders (Environ.universes env) pl) in
+ let evd, cstrs = interp_univ_constraints env evd decl.univdecl_constraints in
+ let decl = { univdecl_instance = pl;
+ univdecl_extensible_instance = decl.univdecl_extensible_instance;
+ univdecl_constraints = cstrs;
+ univdecl_extensible_constraints = decl.univdecl_extensible_constraints }
+ in evd, decl
+
+let interp_univ_decl_opt env l =
+ match l with
+ | None -> Evd.from_env env, default_univ_decl
+ | Some decl -> interp_univ_decl env decl
diff --git a/pretyping/univdecls.mli b/pretyping/univdecls.mli
new file mode 100644
index 00000000..305d045b
--- /dev/null
+++ b/pretyping/univdecls.mli
@@ -0,0 +1,21 @@
+(************************************************************************)
+(* * 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) *)
+(************************************************************************)
+
+(** Local universe and constraint declarations. *)
+type universe_decl =
+ (Misctypes.lident list, Univ.Constraint.t) Misctypes.gen_universe_decl
+
+val default_univ_decl : universe_decl
+
+val interp_univ_decl : Environ.env -> Constrexpr.universe_decl_expr ->
+ Evd.evar_map * universe_decl
+
+val interp_univ_decl_opt : Environ.env -> Constrexpr.universe_decl_expr option ->
+ Evd.evar_map * universe_decl
diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml
index e281f22d..c53305e2 100644
--- a/pretyping/vnorm.ml
+++ b/pretyping/vnorm.ml
@@ -1,22 +1,29 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Util
open Names
open Declarations
open Term
+open Constr
open Vars
open Environ
open Inductive
open Reduction
+open Vmvalues
open Vm
open Context.Rel.Declaration
+module RelDecl = Context.Rel.Declaration
+module NamedDecl = Context.Named.Declaration
+
(*******************************************)
(* Calcul de la forme normal d'un terme *)
(*******************************************)
@@ -48,7 +55,7 @@ let invert_tag cst tag reloc_tbl =
let find_rectype_a env c =
let (t, l) = decompose_appvect (whd_all env c) in
- match kind_of_term t with
+ match kind t with
| Ind ind -> (ind, l)
| _ -> assert false
@@ -96,13 +103,15 @@ let construct_of_constr_block = construct_of_constr false
let type_of_ind env (ind, u) =
type_of_inductive env (Inductive.lookup_mind_specif env ind, u)
-let build_branches_type env (mind,_ as _ind) mib mip u params dep p =
+let build_branches_type env sigma (mind,_ as _ind) mib mip u params dep p =
let rtbl = mip.mind_reloc_tbl in
(* [build_one_branch i cty] construit le type de la ieme branche (commence
a 0) et les lambda correspondant aux realargs *)
let build_one_branch i cty =
let typi = type_constructor mind mib u cty params in
- let decl,indapp = Reductionops.splay_prod env Evd.empty typi in
+ let decl,indapp = Reductionops.splay_prod env sigma (EConstr.of_constr typi) in
+ let decl = List.map (on_snd EConstr.Unsafe.to_constr) decl in
+ let indapp = EConstr.Unsafe.to_constr indapp in
let decl_with_letin,_ = decompose_prod_assum typi in
let ((ind,u),cargs) = find_rectype_a env indapp in
let nparams = Array.length params in
@@ -128,28 +137,27 @@ let build_case_type dep p realargs c =
(* La fonction de normalisation *)
-let rec nf_val env v t = nf_whd env (whd_val v) t
+let rec nf_val env sigma v t = nf_whd env sigma (Vmvalues.whd_val v) t
-and nf_vtype env v = nf_val env v crazy_type
+and nf_vtype env sigma v = nf_val env sigma v crazy_type
-and nf_whd env whd typ =
+and nf_whd env sigma whd typ =
match whd with
- | Vsort s -> mkSort s
| Vprod p ->
- let dom = nf_vtype env (dom p) in
+ let dom = nf_vtype env sigma (dom p) in
let name = Name (Id.of_string "x") in
- let vc = body_of_vfun (nb_rel env) (codom p) in
- let codom = nf_vtype (push_rel (LocalAssum (name,dom)) env) vc in
+ let vc = reduce_fun (nb_rel env) (codom p) in
+ let codom = nf_vtype (push_rel (LocalAssum (name,dom)) env) sigma vc in
mkProd(name,dom,codom)
- | Vfun f -> nf_fun env f typ
- | Vfix(f,None) -> nf_fix env f
- | Vfix(f,Some vargs) -> fst (nf_fix_app env f vargs)
- | Vcofix(cf,_,None) -> nf_cofix env cf
+ | Vfun f -> nf_fun env sigma f typ
+ | Vfix(f,None) -> nf_fix env sigma f
+ | Vfix(f,Some vargs) -> fst (nf_fix_app env sigma f vargs)
+ | Vcofix(cf,_,None) -> nf_cofix env sigma cf
| Vcofix(cf,_,Some vargs) ->
- let cfd = nf_cofix env cf in
+ let cfd = nf_cofix env sigma cf in
let i,(_,ta,_) = destCoFix cfd in
let t = ta.(i) in
- let _, args = nf_args env vargs t in
+ let _, args = nf_args env sigma vargs t in
mkApp(cfd,args)
| Vconstr_const n ->
construct_of_constr_const env n typ
@@ -162,149 +170,176 @@ and nf_whd env whd typ =
| _ -> assert false
else (tag, 0) in
let capp,ctyp = construct_of_constr_block env tag typ in
- let args = nf_bargs env b ofs ctyp in
+ let args = nf_bargs env sigma b ofs ctyp in
mkApp(capp,args)
| Vatom_stk(Aid idkey, stk) ->
- constr_type_of_idkey env idkey stk
+ constr_type_of_idkey env sigma idkey stk
| Vatom_stk(Aind ((mi,i) as ind), stk) ->
let mib = Environ.lookup_mind mi env in
let nb_univs =
- if mib.mind_polymorphic then Univ.UContext.size mib.mind_universes
- else 0
+ Univ.AUContext.size (Declareops.inductive_polymorphic_context mib)
in
let mk u =
let pind = (ind, u) in (mkIndU pind, type_of_ind env pind)
in
- nf_univ_args ~nb_univs mk env stk
- | Vatom_stk(Atype u, stk) -> assert false
+ nf_univ_args ~nb_univs mk env sigma stk
+ | Vatom_stk(Asort s, stk) ->
+ assert (List.is_empty stk); mkSort s
| Vuniv_level lvl ->
assert false
-and nf_univ_args ~nb_univs mk env stk =
+and nf_univ_args ~nb_univs mk env sigma stk =
let u =
if Int.equal nb_univs 0 then Univ.Instance.empty
else match stk with
| Zapp args :: _ ->
let inst =
- Array.init nb_univs (fun i -> Vm.uni_lvl_val (arg args i))
+ Array.init nb_univs (fun i -> uni_lvl_val (arg args i))
in
Univ.Instance.of_array inst
| _ -> assert false
in
let (t,ty) = mk u in
- nf_stk ~from:nb_univs env t ty stk
-
-and constr_type_of_idkey env (idkey : Vars.id_key) stk =
+ nf_stk ~from:nb_univs env sigma t ty stk
+
+and nf_evar env sigma evk stk =
+ let evi = try Evd.find sigma evk with Not_found -> assert false in
+ let hyps = Environ.named_context_of_val (Evd.evar_filtered_hyps evi) in
+ let concl = Evd.evar_concl evi in
+ if List.is_empty hyps then
+ nf_stk env sigma (mkEvar (evk, [||])) concl stk
+ else match stk with
+ | Zapp args :: stk ->
+ (** We assume that there is no consecutive Zapp nodes in a VM stack. Is that
+ really an invariant? *)
+ (** Let-bound arguments are present in the evar arguments but not in the
+ type, so we turn the let into a product. *)
+ let drop_body = function
+ | NamedDecl.LocalAssum _ as d -> d
+ | NamedDecl.LocalDef (na, _, t) -> NamedDecl.LocalAssum (na, t)
+ in
+ let hyps = List.map drop_body hyps in
+ let fold accu d = Term.mkNamedProd_or_LetIn d accu in
+ let t = List.fold_left fold concl hyps in
+ let t, args = nf_args env sigma args t in
+ let inst, args = Array.chop (List.length hyps) args in
+ let c = mkApp (mkEvar (evk, inst), args) in
+ nf_stk env sigma c t stk
+ | _ ->
+ CErrors.anomaly (Pp.str "Argument size mismatch when decompiling an evar")
+
+and constr_type_of_idkey env sigma (idkey : Vmvalues.id_key) stk =
match idkey with
| ConstKey cst ->
let cbody = Environ.lookup_constant cst env in
let nb_univs =
- if cbody.const_polymorphic then Univ.UContext.size cbody.const_universes
- else 0
+ Univ.AUContext.size (Declareops.constant_polymorphic_context cbody)
in
let mk u =
let pcst = (cst, u) in (mkConstU pcst, Typeops.type_of_constant_in env pcst)
in
- nf_univ_args ~nb_univs mk env stk
+ nf_univ_args ~nb_univs mk env sigma stk
| VarKey id ->
- let open Context.Named.Declaration in
- let ty = get_type (lookup_named id env) in
- nf_stk env (mkVar id) ty stk
+ let ty = NamedDecl.get_type (lookup_named id env) in
+ nf_stk env sigma (mkVar id) ty stk
| RelKey i ->
let n = (nb_rel env - i) in
- let ty = get_type (lookup_rel n env) in
- nf_stk env (mkRel n) (lift n ty) stk
+ let ty = RelDecl.get_type (lookup_rel n env) in
+ nf_stk env sigma (mkRel n) (lift n ty) stk
+ | EvarKey evk ->
+ nf_evar env sigma evk stk
-and nf_stk ?from:(from=0) env c t stk =
+and nf_stk ?from:(from=0) env sigma c t stk =
match stk with
| [] -> c
| Zapp vargs :: stk ->
if nargs vargs >= from then
- let t, args = nf_args ~from:from env vargs t in
- nf_stk env (mkApp(c,args)) t stk
+ let t, args = nf_args ~from:from env sigma vargs t in
+ nf_stk env sigma (mkApp(c,args)) t stk
else
let rest = from - nargs vargs in
- nf_stk ~from:rest env c t stk
+ nf_stk ~from:rest env sigma c t stk
| Zfix (f,vargs) :: stk ->
assert (from = 0) ;
- let fa, typ = nf_fix_app env f vargs in
+ let fa, typ = nf_fix_app env sigma f vargs in
let _,_,codom = decompose_prod env typ in
- nf_stk env (mkApp(fa,[|c|])) (subst1 c codom) stk
+ nf_stk env sigma (mkApp(fa,[|c|])) (subst1 c codom) stk
| Zswitch sw :: stk ->
assert (from = 0) ;
let ((mind,_ as ind), u), allargs = find_rectype_a env t in
let (mib,mip) = Inductive.lookup_mind_specif env ind in
let nparams = mib.mind_nparams in
let params,realargs = Util.Array.chop nparams allargs in
+ let nparamdecls = Context.Rel.length (Inductive.inductive_paramdecls (mib,u)) in
let pT =
- hnf_prod_applist env (type_of_ind env (ind,u)) (Array.to_list params) in
+ hnf_prod_applist_assum env nparamdecls (type_of_ind env (ind,u)) (Array.to_list params) in
let pT = whd_all env pT in
- let dep, p = nf_predicate env (ind,u) mip params (type_of_switch sw) pT in
+ let dep, p = nf_predicate env sigma (ind,u) mip params (type_of_switch sw) pT in
(* Calcul du type des branches *)
- let btypes = build_branches_type env ind mib mip u params dep p in
+ let btypes = build_branches_type env sigma ind mib mip u params dep p in
(* calcul des branches *)
let bsw = branch_of_switch (nb_rel env) sw in
let mkbranch i (n,v) =
let decl,decl_with_letin,codom = btypes.(i) in
- let b = nf_val (Termops.push_rels_assum decl env) v codom in
+ let b = nf_val (Termops.push_rels_assum decl env) sigma v codom in
Termops.it_mkLambda_or_LetIn_from_no_LetIn b decl_with_letin
in
let branchs = Array.mapi mkbranch bsw in
let tcase = build_case_type dep p realargs c in
- let ci = case_info sw in
- nf_stk env (mkCase(ci, p, c, branchs)) tcase stk
+ let ci = sw.sw_annot.Cbytecodes.ci in
+ nf_stk env sigma (mkCase(ci, p, c, branchs)) tcase stk
| Zproj p :: stk ->
assert (from = 0) ;
let p' = Projection.make p true in
- let ty = Inductiveops.type_of_projection_knowing_arg env Evd.empty p' c t in
- nf_stk env (mkProj(p',c)) ty stk
+ let ty = Inductiveops.type_of_projection_knowing_arg env sigma p' (EConstr.of_constr c) (EConstr.of_constr t) in
+ nf_stk env sigma (mkProj(p',c)) ty stk
-and nf_predicate env ind mip params v pT =
- match whd_val v, kind_of_term pT with
+and nf_predicate env sigma ind mip params v pT =
+ match whd_val v, kind pT with
| Vfun f, Prod _ ->
let k = nb_rel env in
- let vb = body_of_vfun k f in
+ let vb = reduce_fun k f in
let name,dom,codom = decompose_prod env pT in
let dep,body =
- nf_predicate (push_rel (LocalAssum (name,dom)) env) ind mip params vb codom in
+ nf_predicate (push_rel (LocalAssum (name,dom)) env) sigma ind mip params vb codom in
dep, mkLambda(name,dom,body)
| Vfun f, _ ->
let k = nb_rel env in
- let vb = body_of_vfun k f in
+ let vb = reduce_fun k f in
let name = Name (Id.of_string "c") in
let n = mip.mind_nrealargs in
let rargs = Array.init n (fun i -> mkRel (n-i)) in
let params = if Int.equal n 0 then params else Array.map (lift n) params in
let dom = mkApp(mkIndU ind,Array.append params rargs) in
- let body = nf_vtype (push_rel (LocalAssum (name,dom)) env) vb in
+ let body = nf_vtype (push_rel (LocalAssum (name,dom)) env) sigma vb in
true, mkLambda(name,dom,body)
- | _, _ -> false, nf_val env v crazy_type
+ | _, _ -> false, nf_val env sigma v crazy_type
-and nf_args env vargs ?from:(f=0) t =
+and nf_args env sigma vargs ?from:(f=0) t =
let t = ref t in
let len = nargs vargs - f in
let args =
Array.init len
(fun i ->
let _,dom,codom = decompose_prod env !t in
- let c = nf_val env (arg vargs (f+i)) dom in
+ let c = nf_val env sigma (arg vargs (f+i)) dom in
t := subst1 c codom; c) in
!t,args
-and nf_bargs env b ofs t =
+and nf_bargs env sigma b ofs t =
let t = ref t in
let len = bsize b - ofs in
let args =
Array.init len
(fun i ->
let _,dom,codom = decompose_prod env !t in
- let c = nf_val env (bfield b (i+ofs)) dom in
+ let c = nf_val env sigma (bfield b (i+ofs)) dom in
t := subst1 c codom; c) in
args
-and nf_fun env f typ =
+and nf_fun env sigma f typ =
let k = nb_rel env in
- let vb = body_of_vfun k f in
+ let vb = reduce_fun k f in
let name,dom,codom =
try decompose_prod env typ
with DestKO ->
@@ -312,49 +347,54 @@ and nf_fun env f typ =
CErrors.anomaly
(Pp.strbrk "Returned a functional value in a type not recognized as a product type.")
in
- let body = nf_val (push_rel (LocalAssum (name,dom)) env) vb codom in
+ let body = nf_val (push_rel (LocalAssum (name,dom)) env) sigma vb codom in
mkLambda(name,dom,body)
-and nf_fix env f =
+and nf_fix env sigma f =
let init = current_fix f in
let rec_args = rec_args f in
let k = nb_rel env in
let vb, vt = reduce_fix k f in
let ndef = Array.length vt in
- let ft = Array.map (fun v -> nf_val env v crazy_type) vt in
+ let ft = Array.map (fun v -> nf_val env sigma v crazy_type) vt in
let name = Array.init ndef (fun _ -> (Name (Id.of_string "Ffix"))) in
(* Third argument of the tuple is ignored by push_rec_types *)
let env = push_rec_types (name,ft,ft) env in
(* We lift here because the types of arguments (in tt) will be evaluated
in an environment where the fixpoints have been pushed *)
- let norm_vb v t = nf_fun env v (lift ndef t) in
+ let norm_vb v t = nf_fun env sigma v (lift ndef t) in
let fb = Util.Array.map2 norm_vb vb ft in
mkFix ((rec_args,init),(name,ft,fb))
-and nf_fix_app env f vargs =
- let fd = nf_fix env f in
+and nf_fix_app env sigma f vargs =
+ let fd = nf_fix env sigma f in
let (_,i),(_,ta,_) = destFix fd in
let t = ta.(i) in
- let t, args = nf_args env vargs t in
+ let t, args = nf_args env sigma vargs t in
mkApp(fd,args),t
-and nf_cofix env cf =
+and nf_cofix env sigma cf =
let init = current_cofix cf in
let k = nb_rel env in
let vb,vt = reduce_cofix k cf in
let ndef = Array.length vt in
- let cft = Array.map (fun v -> nf_val env v crazy_type) vt in
+ let cft = Array.map (fun v -> nf_val env sigma v crazy_type) vt in
let name = Array.init ndef (fun _ -> (Name (Id.of_string "Fcofix"))) in
let env = push_rec_types (name,cft,cft) env in
- let cfb = Util.Array.map2 (fun v t -> nf_val env v t) vb cft in
+ let cfb = Util.Array.map2 (fun v t -> nf_val env sigma v t) vb cft in
mkCoFix (init,(name,cft,cfb))
-let cbv_vm env c t =
+let cbv_vm env sigma c t =
+ if Termops.occur_meta sigma c then
+ CErrors.user_err Pp.(str "vm_compute does not support metas.");
+ (** This evar-normalizes terms beforehand *)
+ let c = EConstr.to_constr sigma c in
+ let t = EConstr.to_constr sigma t in
let v = Vconv.val_of_constr env c in
- nf_val env v t
+ EConstr.of_constr (nf_val env sigma v t)
let vm_infer_conv ?(pb=Reduction.CUMUL) env sigma t1 t2 =
Reductionops.infer_conv_gen (fun pb ~l2r sigma ts -> Vconv.vm_conv_gen pb)
~catch_incon:true ~pb env sigma t1 t2
-let _ = Reductionops.set_vm_infer_conv vm_infer_conv
+let _ = if Coq_config.bytecode_compiler then Reductionops.set_vm_infer_conv vm_infer_conv
diff --git a/pretyping/vnorm.mli b/pretyping/vnorm.mli
index 58f5b14e..3e0eabb0 100644
--- a/pretyping/vnorm.mli
+++ b/pretyping/vnorm.mli
@@ -1,13 +1,15 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * 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 *)
+(* // * 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 Term
+open EConstr
open Environ
(** {6 Reduction functions } *)
-val cbv_vm : env -> constr -> types -> constr
+val cbv_vm : env -> Evd.evar_map -> constr -> types -> constr