summaryrefslogtreecommitdiff
path: root/pretyping
diff options
context:
space:
mode:
authorGravatar Enrico Tassi <gareuselesinge@debian.org>2015-01-25 14:42:51 +0100
committerGravatar Enrico Tassi <gareuselesinge@debian.org>2015-01-25 14:42:51 +0100
commit7cfc4e5146be5666419451bdd516f1f3f264d24a (patch)
treee4197645da03dc3c7cc84e434cc31d0a0cca7056 /pretyping
parent420f78b2caeaaddc6fe484565b2d0e49c66888e5 (diff)
Imported Upstream version 8.5~beta1+dfsg
Diffstat (limited to 'pretyping')
-rw-r--r--pretyping/arguments_renaming.ml50
-rw-r--r--pretyping/arguments_renaming.mli14
-rw-r--r--pretyping/cases.ml1274
-rw-r--r--pretyping/cases.mli108
-rw-r--r--pretyping/cbv.ml80
-rw-r--r--pretyping/cbv.mli7
-rw-r--r--pretyping/classops.ml370
-rw-r--r--pretyping/classops.mli30
-rw-r--r--pretyping/coercion.ml724
-rw-r--r--pretyping/coercion.mli105
-rw-r--r--pretyping/constr_matching.ml494
-rw-r--r--pretyping/constr_matching.mli (renamed from pretyping/matching.mli)49
-rw-r--r--pretyping/detyping.ml661
-rw-r--r--pretyping/detyping.mli51
-rw-r--r--pretyping/evarconv.ml1386
-rw-r--r--pretyping/evarconv.mli82
-rw-r--r--pretyping/evarsolve.ml1548
-rw-r--r--pretyping/evarsolve.mli74
-rw-r--r--pretyping/evarutil.ml2027
-rw-r--r--pretyping/evarutil.mli187
-rw-r--r--pretyping/evd.ml1932
-rw-r--r--pretyping/evd.mli613
-rw-r--r--pretyping/find_subterm.ml179
-rw-r--r--pretyping/find_subterm.mli69
-rw-r--r--pretyping/glob_ops.ml434
-rw-r--r--pretyping/glob_ops.mli61
-rw-r--r--pretyping/glob_term.ml418
-rw-r--r--pretyping/glob_term.mli167
-rw-r--r--pretyping/indrec.ml281
-rw-r--r--pretyping/indrec.mli39
-rw-r--r--pretyping/inductiveops.ml357
-rw-r--r--pretyping/inductiveops.mli118
-rw-r--r--pretyping/locusops.ml125
-rw-r--r--pretyping/locusops.mli46
-rw-r--r--pretyping/matching.ml357
-rw-r--r--pretyping/miscops.ml60
-rw-r--r--pretyping/miscops.mli29
-rw-r--r--pretyping/namegen.ml204
-rw-r--r--pretyping/namegen.mli68
-rw-r--r--pretyping/nativenorm.ml404
-rw-r--r--pretyping/nativenorm.mli17
-rw-r--r--pretyping/pattern.mli126
-rw-r--r--pretyping/patternops.ml (renamed from pretyping/pattern.ml)262
-rw-r--r--pretyping/patternops.mli55
-rw-r--r--pretyping/pretype_errors.ml153
-rw-r--r--pretyping/pretype_errors.mli122
-rw-r--r--pretyping/pretyping.ml1644
-rw-r--r--pretyping/pretyping.mli152
-rw-r--r--pretyping/pretyping.mllib21
-rw-r--r--pretyping/program.ml69
-rw-r--r--pretyping/program.mli39
-rw-r--r--pretyping/recordops.ml221
-rw-r--r--pretyping/recordops.mli32
-rw-r--r--pretyping/redops.ml38
-rw-r--r--pretyping/redops.mli13
-rw-r--r--pretyping/reductionops.ml1676
-rw-r--r--pretyping/reductionops.mli181
-rw-r--r--pretyping/retyping.ml182
-rw-r--r--pretyping/retyping.mli21
-rw-r--r--pretyping/tacred.ml1009
-rw-r--r--pretyping/tacred.mli45
-rw-r--r--pretyping/term_dnet.ml402
-rw-r--r--pretyping/term_dnet.mli108
-rw-r--r--pretyping/termops.ml554
-rw-r--r--pretyping/termops.mli158
-rw-r--r--pretyping/typeclasses.ml382
-rw-r--r--pretyping/typeclasses.mli89
-rw-r--r--pretyping/typeclasses_errors.ml34
-rw-r--r--pretyping/typeclasses_errors.mli24
-rw-r--r--pretyping/typing.ml93
-rw-r--r--pretyping/typing.mli17
-rw-r--r--pretyping/unification.ml1382
-rw-r--r--pretyping/unification.mli77
-rw-r--r--pretyping/vnorm.ml119
-rw-r--r--pretyping/vnorm.mli4
75 files changed, 15359 insertions, 9444 deletions
diff --git a/pretyping/arguments_renaming.ml b/pretyping/arguments_renaming.ml
index 78290e03..3cfc0dc8 100644
--- a/pretyping/arguments_renaming.ml
+++ b/pretyping/arguments_renaming.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,32 +8,20 @@
(*i*)
open Names
-open Libnames
-open Decl_kinds
+open Globnames
open Term
-open Sign
-open Evd
open Environ
-open Nametab
-open Mod_subst
open Util
-open Pp
open Libobject
-open Nameops
(*i*)
-let empty_name_table = (Refmap.empty : name list list Refmap.t)
-let name_table = ref empty_name_table
-
-let _ =
- Summary.declare_summary "rename-arguments"
- { Summary.freeze_function = (fun () -> !name_table);
- Summary.unfreeze_function = (fun r -> name_table := r);
- Summary.init_function = (fun () -> name_table := empty_name_table) }
+let name_table =
+ Summary.ref (Refmap.empty : Name.t list list Refmap.t)
+ ~name:"rename-arguments"
type req =
| ReqLocal
- | ReqGlobal of global_reference * name list list
+ | ReqGlobal of global_reference * Name.t list list
let load_rename_args _ (_, (_, (r, names))) =
name_table := Refmap.add r names !name_table
@@ -53,12 +41,12 @@ 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,_,_ = 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 names' = List.map (fun l -> var_names @ l) names in
@@ -99,22 +87,24 @@ let rename_type ty ref =
with Not_found -> ty
let rename_type_of_constant env c =
- let ty = Typeops.type_of_constant env c in
- rename_type ty (ConstRef c)
+ let ty = Typeops.type_of_constant_in env c in
+ rename_type ty (ConstRef (fst c))
let rename_type_of_inductive env ind =
let ty = Inductiveops.type_of_inductive env ind in
- rename_type ty (IndRef ind)
+ rename_type ty (IndRef (fst ind))
let rename_type_of_constructor env cstruct =
let ty = Inductiveops.type_of_constructor env cstruct in
- rename_type ty (ConstructRef cstruct)
+ rename_type ty (ConstructRef (fst cstruct))
let rename_typing env c =
- let j = Typeops.typing env c in
- match kind_of_term c with
- | Const c -> { j with uj_type = rename_type j.uj_type (ConstRef c) }
- | Ind i -> { j with uj_type = rename_type j.uj_type (IndRef i) }
- | Construct k -> { j with uj_type = rename_type j.uj_type (ConstructRef k) }
- | _ -> j
+ let j = Typeops.infer env c in
+ let j' =
+ match kind_of_term 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) }
+ | _ -> j
+ in j'
diff --git a/pretyping/arguments_renaming.mli b/pretyping/arguments_renaming.mli
index 74c4cdc3..290bfc59 100644
--- a/pretyping/arguments_renaming.mli
+++ b/pretyping/arguments_renaming.mli
@@ -1,22 +1,22 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Names
-open Libnames
+open Globnames
open Environ
open Term
-val rename_arguments : bool -> global_reference -> name list list -> unit
+val rename_arguments : bool -> global_reference -> Name.t list list -> unit
(** [Not_found] is raised is no names are defined for [r] *)
-val arguments_names : global_reference -> name list list
+val arguments_names : global_reference -> Name.t list list
-val rename_type_of_constant : env -> constant -> types
-val rename_type_of_inductive : env -> inductive -> types
-val rename_type_of_constructor : env -> constructor -> types
+val rename_type_of_constant : env -> pconstant -> types
+val rename_type_of_inductive : env -> pinductive -> types
+val rename_type_of_constructor : env -> pconstructor -> types
val rename_typing : env -> constr -> unsafe_judgment
diff --git a/pretyping/cases.ml b/pretyping/cases.ml
index 57857351..fdb19d37 100644
--- a/pretyping/cases.ml
+++ b/pretyping/cases.ml
@@ -1,29 +1,32 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Compat
+open Pp
+open Errors
open Util
open Names
open Nameops
open Term
+open Vars
+open Context
open Termops
open Namegen
open Declarations
open Inductiveops
open Environ
-open Sign
open Reductionops
-open Typeops
open Type_errors
open Glob_term
+open Glob_ops
open Retyping
open Pretype_errors
open Evarutil
+open Evarsolve
open Evarconv
open Evd
@@ -34,56 +37,39 @@ type pattern_matching_error =
| BadConstructor of constructor * inductive
| WrongNumargConstructor of constructor * int
| WrongNumargInductive of inductive * int
- | WrongPredicateArity of constr * constr * constr
- | NeedsInversion of constr * constr
| UnusedClause of cases_pattern list
| NonExhaustive of cases_pattern list
| CannotInferPredicate of (constr * types) array
-exception PatternMatchingError of env * pattern_matching_error
+exception PatternMatchingError of env * evar_map * pattern_matching_error
-let raise_pattern_matching_error (loc,ctx,te) =
- Loc.raise loc (PatternMatchingError(ctx,te))
+let raise_pattern_matching_error (loc,env,sigma,te) =
+ Loc.raise loc (PatternMatchingError(env,sigma,te))
-let error_bad_pattern_loc loc cstr ind =
- raise_pattern_matching_error (loc, Global.env(), BadPattern (cstr,ind))
+let error_bad_pattern_loc loc env sigma cstr ind =
+ raise_pattern_matching_error
+ (loc, env, sigma, BadPattern (cstr,ind))
-let error_bad_constructor_loc loc cstr ind =
- raise_pattern_matching_error (loc, Global.env(), BadConstructor (cstr,ind))
+let error_bad_constructor_loc 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, WrongNumargConstructor(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, WrongNumargInductive(c,n))
-
-let error_wrong_predicate_arity_loc loc env c n1 n2 =
- raise_pattern_matching_error (loc, env, WrongPredicateArity (c,n1,n2))
-
-let error_needs_inversion env x t =
- raise (PatternMatchingError (env, NeedsInversion (x,t)))
-
-module type S = sig
- val compile_cases :
- loc -> case_style ->
- (type_constraint -> env -> evar_map ref -> glob_constr -> unsafe_judgment) * evar_map ref ->
- type_constraint ->
- env -> glob_constr option * tomatch_tuples * cases_clauses ->
- unsafe_judgment
-end
+ raise_pattern_matching_error (loc, env, Evd.empty, WrongNumargInductive(c,n))
let rec list_try_compile f = function
| [a] -> f a
- | [] -> anomaly "try_find_f"
+ | [] -> anomaly (str "try_find_f")
| h::t ->
try f h
- with UserError _ | TypeError _ | PretypeError _ | PatternMatchingError _
- | Loc.Exc_located
- (_, (UserError _ | TypeError _ | PretypeError _ | PatternMatchingError _)) ->
+ with UserError _ | TypeError _ | PretypeError _ | PatternMatchingError _ ->
list_try_compile f t
let force_name =
- let nx = Name (id_of_string "x") in function Anonymous -> nx | na -> na
+ let nx = Name default_dependent_ident in function Anonymous -> nx | na -> na
(************************************************************************)
(* Pattern-matching compilation (Cases) *)
@@ -99,18 +85,15 @@ let msg_may_need_inversion () =
(* Utils *)
let make_anonymous_patvars n =
- list_make n (PatVar (dummy_loc,Anonymous))
-
-(* Environment management *)
-let push_rels vars env = List.fold_right push_rel vars env
+ List.make n (PatVar (Loc.ghost,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 j = n1+k then n2+k else j
+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
- | Rel j when j = n1+k -> mkRel (n2+k)
+ | 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
@@ -120,27 +103,33 @@ let rec relocate_index n1 n2 k t = match kind_of_term t with
type 'a rhs =
{ rhs_env : env;
- rhs_vars : identifier list;
- avoid_ids : identifier list;
+ rhs_vars : Id.t list;
+ avoid_ids : Id.t list;
it : 'a option}
type 'a equation =
{ patterns : cases_pattern list;
rhs : 'a rhs;
- alias_stack : name list;
- eqn_loc : loc;
+ alias_stack : Name.t list;
+ eqn_loc : Loc.t;
used : bool ref }
type 'a matrix = 'a equation list
(* 1st argument of IsInd is the original ind before extracting the summary *)
type tomatch_type =
- | IsInd of types * inductive_type * name list
+ | IsInd of types * inductive_type * Name.t list
| NotInd of constr option * types
+(* spiwack: The first argument of [Pushed] is [true] for initial
+ Pushed and [false] otherwise. Used to decide whether the term being
+ matched on must be aliased in the variable case (only initial
+ Pushed need to be aliased). The first argument of [Alias] is [true]
+ if the alias was introduced by an initial pushed and [false]
+ otherwise.*)
type tomatch_status =
- | Pushed of ((constr * tomatch_type) * int list * name)
- | Alias of (name * constr * (constr * types))
+ | Pushed of (bool*((constr * tomatch_type) * int list * Name.t))
+ | Alias of (bool*(Name.t * constr * (constr * types)))
| NonDepAlias
| Abstract of int * rel_declaration
@@ -162,9 +151,9 @@ let feed_history arg = function
| Continuation (n, l, h) when n>=1 ->
Continuation (n-1, arg :: l, h)
| Continuation (n, _, _) ->
- anomaly ("Bad number of expected remaining patterns: "^(string_of_int n))
+ anomaly (str "Bad number of expected remaining patterns: " ++ int n)
| Result _ ->
- anomaly "Exhausted pattern history"
+ anomaly (Pp.str "Exhausted pattern history")
(* This is for non exhaustive error message *)
@@ -178,22 +167,22 @@ and build_glob_pattern args = function
| Top -> args
| MakeConstructor (pci, rh) ->
glob_pattern_of_partial_history
- [PatCstr (dummy_loc, pci, args, Anonymous)] rh
+ [PatCstr (Loc.ghost, pci, args, Anonymous)] rh
let complete_history = glob_pattern_of_partial_history []
(* This is to build glued pattern-matching history and alias bodies *)
-let rec pop_history_pattern = function
+let pop_history_pattern = function
| Continuation (0, l, Top) ->
Result (List.rev l)
| Continuation (0, l, MakeConstructor (pci, rh)) ->
- feed_history (PatCstr (dummy_loc,pci,List.rev l,Anonymous)) rh
+ feed_history (PatCstr (Loc.ghost,pci,List.rev l,Anonymous)) rh
| _ ->
- anomaly "Constructor not yet filled with its arguments"
+ anomaly (Pp.str "Constructor not yet filled with its arguments")
let pop_history h =
- feed_history (PatVar (dummy_loc, Anonymous)) h
+ feed_history (PatVar (Loc.ghost, Anonymous)) h
(* Builds a continuation expecting [n] arguments and building [ci] applied
to this [n] arguments *)
@@ -251,7 +240,7 @@ type 'a pattern_matching_problem =
tomatch : tomatch_stack;
history : pattern_continuation;
mat : 'a matrix;
- caseloc : loc;
+ caseloc : Loc.t;
casestyle : case_style;
typing_function: type_constraint -> env -> evar_map ref -> 'a option -> unsafe_judgment }
@@ -277,60 +266,67 @@ let rec find_row_ind = function
| PatCstr(loc,c,_,_) :: _ -> Some (loc,c)
let inductive_template evdref env tmloc ind =
- let arsign = get_full_arity_sign env ind in
+ 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, TomatchTypeParameter (ind,i))
- | None -> fun _ -> (dummy_loc, InternalHole) in
+ | Some loc -> fun i -> (loc, Evar_kinds.TomatchTypeParameter (ind,i))
+ | None -> fun _ -> (Loc.ghost, Evar_kinds.InternalHole) in
let (_,evarl,_) =
List.fold_right
(fun (na,b,ty) (subst,evarl,n) ->
match b with
| None ->
let ty' = substl subst ty in
- let e = e_new_evar evdref env ~src:(hole_source n) ty' in
+ let e = e_new_evar env evdref ~src:(hole_source n) ty' in
(e::subst,e::evarl,n+1)
| Some b ->
(substl subst b::subst,evarl,n+1))
arsign ([],[],1) in
- applist (mkInd ind,List.rev evarl)
+ applist (mkIndU indu,List.rev evarl)
let try_find_ind env sigma typ realnames =
let (IndType(_,realargs) as ind) = find_rectype env sigma typ in
let names =
match realnames with
| Some names -> names
- | None -> list_make (List.length realargs) Anonymous in
+ | None -> List.make (List.length realargs) Anonymous in
IsInd (typ,ind,names)
-let inh_coerce_to_ind evdref env ty tyi =
- let expected_typ = inductive_template evdref env None tyi in
- (* devrait être indifférent d'exiger leq ou pas puisque pour
- un inductif cela doit être égal *)
- let _ = e_cumul env evdref expected_typ ty in ()
+let inh_coerce_to_ind evdref env loc ty tyi =
+ let sigma = !evdref in
+ let expected_typ = inductive_template evdref env loc tyi in
+ (* Try to refine the type with inductive information coming from the
+ constructor and renounce if not able to give more information *)
+ (* devrait être indifférent d'exiger leq ou pas puisque pour
+ un inductif cela doit être égal *)
+ if not (e_cumul env evdref expected_typ ty) then evdref := sigma
let binding_vars_of_inductive = function
| NotInd _ -> []
| IsInd (_,IndType(_,realargs),_) -> List.filter isRel realargs
let extract_inductive_data env sigma (_,b,t) =
- if b<>None then (NotInd (None,t),[]) else
- 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
- (tmtyp,tmtypvars)
+ match b with
+ | None ->
+ 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
+ (tmtyp,tmtypvars)
+ | Some _ ->
+ (NotInd (None, t), [])
let unify_tomatch_with_patterns evdref env loc typ pats realnames =
match find_row_ind pats with
| None -> NotInd (None,typ)
| Some (_,(ind,_)) ->
- inh_coerce_to_ind evdref env typ ind;
+ inh_coerce_to_ind evdref env loc typ ind;
try try_find_ind env !evdref typ realnames
with Not_found -> NotInd (None,typ)
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 (_,ind,realnal) ->
mk_tycon (inductive_template evdref env loc ind),Some (List.rev realnal)
| None ->
empty_tycon,None
@@ -339,6 +335,8 @@ let coerce_row typing_fun evdref env pats (tomatch,(_,indopt)) =
let loc = Some (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
+ evdref := evd;
let typ = nf_evar !evdref j.uj_type in
let t =
try try_find_ind env !evdref typ realnames
@@ -356,17 +354,14 @@ let coerce_to_indtype typing_fun evdref env matx tomatchl =
(************************************************************************)
(* Utils *)
-let mkExistential env ?(src=(dummy_loc,InternalHole)) evdref =
- e_new_evar evdref env ~src:src (new_Type ())
+let mkExistential env ?(src=(Loc.ghost,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 =
let (evd',y) = f !evdref x y in
evdref := evd';
y
-
-module Cases_F(Coercion : Coercion.S) : S = struct
-
let adjust_tomatch_to_pattern pb ((current,typ),deps,dep) =
(* Ideally, we could find a common inductive type to which both the
term to match and the patterns coerce *)
@@ -386,13 +381,13 @@ 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 deps = [] & isEvar typ then
+ if List.is_empty deps && isEvar 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 dummy_loc pb.env)
- pb.evdref (make_judge current typ) (mk_tycon_type indt)).uj_val in
+ (evd_comb2 (Coercion.inh_conv_coerce_to true Loc.ghost 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))
| _ -> (current,tmtyp)
@@ -401,10 +396,6 @@ let type_of_tomatch = function
| IsInd (t,_,_) -> t
| NotInd (_,t) -> t
-let mkDeclTomatch na = function
- | IsInd (t,_,_) -> (na,None,t)
- | NotInd (c,t) -> (na,c,t)
-
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)
@@ -418,7 +409,7 @@ let lift_tomatch_type n = liftn_tomatch_type n 1
let current_pattern eqn =
match eqn.patterns with
| pat::_ -> pat
- | [] -> anomaly "Empty list of patterns"
+ | [] -> anomaly (Pp.str "Empty list of patterns")
let alias_of_pat = function
| PatVar (_,name) -> name
@@ -430,7 +421,7 @@ let remove_current_pattern eqn =
{ eqn with
patterns = pats;
alias_stack = alias_of_pat pat :: eqn.alias_stack }
- | [] -> anomaly "Empty list of patterns"
+ | [] -> anomaly (Pp.str "Empty list of patterns")
let push_current_pattern (cur,ty) eqn =
match eqn.patterns with
@@ -439,7 +430,19 @@ let push_current_pattern (cur,ty) eqn =
{ eqn with
rhs = { eqn.rhs with rhs_env = rhs_env };
patterns = pats }
- | [] -> anomaly "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
+ the root of a pattern matching problem (initial push), so we
+ distinguish the cases. *)
+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")
+
+
let prepend_pattern tms eqn = {eqn with patterns = tms@eqn.patterns }
@@ -466,33 +469,32 @@ let check_and_adjust_constructor env ind cstrs = function
(* Check the constructor has the right number of args *)
let ci = cstrs.(i-1) in
let nb_args_constr = ci.cs_nargs in
- if List.length args = nb_args_constr then pat
+ 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)
with NotAdjustable ->
- error_wrong_numarg_constructor_loc loc (Global.env())
- cstr nb_args_constr
+ error_wrong_numarg_constructor_loc loc env cstr nb_args_constr
else
(* Try to insert a coercion *)
try
- Coercion.inh_pattern_coerce_to loc pat ind' ind
+ Coercion.inh_pattern_coerce_to loc env pat ind' ind
with Not_found ->
- error_bad_constructor_loc loc cstr ind
+ error_bad_constructor_loc loc env cstr ind
-let check_all_variables typ mat =
+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 cstr_sp typ)
+ error_bad_pattern_loc 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, UnusedClause eqn.patterns)
+ (eqn.eqn_loc, env, Evd.empty, UnusedClause eqn.patterns)
let set_used_pattern eqn = eqn.used := true
@@ -509,20 +511,21 @@ let extract_rhs pb =
let occur_in_rhs na rhs =
match na with
| Anonymous -> false
- | Name id -> List.mem id rhs.rhs_vars
+ | Name id -> Id.List.mem id rhs.rhs_vars
let is_dep_patt_in eqn = function
- | PatVar (_,name) -> occur_in_rhs name eqn.rhs
+ | PatVar (_,name) -> Flags.is_program_mode () || occur_in_rhs name eqn.rhs
| PatCstr _ -> true
let mk_dep_patt_row (pats,_,eqn) =
List.map (is_dep_patt_in eqn) pats
let dependencies_in_pure_rhs nargs eqns =
- if eqns = [] then list_make nargs false (* Only "_" patts *) else
+ if List.is_empty eqns then
+ List.make nargs (not (Flags.is_program_mode ())) (* Only "_" patts *) else
let deps_rows = List.map mk_dep_patt_row eqns in
let deps_columns = matrix_transpose deps_rows in
- List.map (List.exists ((=) true)) deps_columns
+ List.map (List.exists (fun x -> x)) deps_columns
let dependent_decl a = function
| (na,None,t) -> dependent a t
@@ -530,12 +533,12 @@ let dependent_decl a = function
let rec dep_in_tomatch n = function
| (Pushed _ | Alias _ | NonDepAlias) :: l -> dep_in_tomatch n l
- | Abstract (_,d) :: l -> dependent_decl (mkRel n) d or dep_in_tomatch (n+1) l
+ | Abstract (_,d) :: l -> dependent_decl (mkRel n) d || dep_in_tomatch (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
+ | Rel n when dep_in_tomatch n tms -> List.make nargs true
| _ -> dependencies_in_pure_rhs nargs eqns
(* Computing the matrix of dependencies *)
@@ -555,12 +558,14 @@ let rec find_dependency_list 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
- then list_add_set (List.length rest + 1) (list_union deps tdeps)
+ then
+ List.add_set Int.equal
+ (List.length rest + 1) (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
- if is_dep_or_cstr_in_rhs || deps <> []
+ if is_dep_or_cstr_in_rhs || not (List.is_empty deps)
then ((true ,deps,d)::nextlist)
else ((false,[] ,d)::nextlist)
@@ -583,14 +588,14 @@ let relocate_index_tomatch n1 n2 =
let rec genrec depth = function
| [] ->
[]
- | Pushed ((c,tm),l,na) :: rest ->
+ | 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 l = List.map (relocate_rel n1 n2 depth) l in
- Pushed ((c,tm),l,na) :: genrec depth rest
- | Alias (na,c,d) :: rest ->
+ Pushed (b,((c,tm),l,na)) :: genrec depth rest
+ | Alias (initial,(na,c,d)) :: rest ->
(* [c] is out of relocation scope *)
- Alias (na,c,map_pair (relocate_index n1 n2 depth) d) :: genrec depth rest
+ Alias (initial,(na,c,map_pair (relocate_index n1 n2 depth) d)) :: genrec depth rest
| NonDepAlias :: rest ->
NonDepAlias :: genrec depth rest
| Abstract (i,d) :: rest ->
@@ -602,24 +607,29 @@ let relocate_index_tomatch n1 n2 =
(* [replace_tomatch n c tomatch] replaces [Rel n] by [c] in [tomatch] *)
let rec replace_term n c k t =
- if isRel t && destRel t = n+k then lift k c
+ 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 length_of_tomatch_type_sign na = function
- | NotInd _ -> if na<>Anonymous then 1 else 0
- | IsInd (_,_,names) -> List.length names + if na<>Anonymous then 1 else 0
+let length_of_tomatch_type_sign na t =
+ let l = match na with
+ | Anonymous -> 0
+ | Name _ -> 1
+ in
+ match t with
+ | NotInd _ -> l
+ | IsInd (_, _, names) -> List.length names + l
let replace_tomatch n c =
let rec replrec depth = function
| [] -> []
- | Pushed ((b,tm),l,na) :: rest ->
+ | 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 i=n+depth then anomaly "replace_tomatch") l;
- Pushed ((b,tm),l,na) :: replrec depth rest
- | Alias (na,b,d) :: rest ->
+ 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 (na,b,map_pair (replace_term n c depth) d) :: replrec depth rest
+ Alias (initial,(na,b,map_pair (replace_term n c depth) d)) :: replrec depth rest
| NonDepAlias :: rest ->
NonDepAlias :: replrec depth rest
| Abstract (i,d) :: rest ->
@@ -636,13 +646,13 @@ let replace_tomatch n c =
let rec liftn_tomatch_stack n depth = function
| [] -> []
- | Pushed ((c,tm),l,na)::rest ->
+ | Pushed (initial,((c,tm),l,na))::rest ->
let c = liftn n depth c in
let tm = liftn_tomatch_type n depth tm in
let l = List.map (fun i -> if i<depth then i else i+n) l in
- Pushed ((c,tm),l,na)::(liftn_tomatch_stack n depth rest)
- | Alias (na,c,d)::rest ->
- Alias (na,liftn n depth c,map_pair (liftn n depth) d)
+ Pushed (initial,((c,tm),l,na))::(liftn_tomatch_stack n depth rest)
+ | Alias (initial,(na,c,d))::rest ->
+ Alias (initial,(na,liftn n depth c,map_pair (liftn n depth) d))
::(liftn_tomatch_stack n depth rest)
| NonDepAlias :: rest ->
NonDepAlias :: liftn_tomatch_stack n depth rest
@@ -684,7 +694,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 names1 = list_make (List.length sign) Anonymous in
+ let names1 = List.make (List.length sign) Anonymous in
(* If any, we prefer names used in pats, from top to bottom *)
let names2,aliasname =
List.fold_right
@@ -695,7 +705,7 @@ 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 l eqn.rhs.avoid_ids)
+ List.fold_left (fun l (_,_,eqn) -> List.union Id.equal l eqn.rhs.avoid_ids)
[] eqns in
let names3,_ =
List.fold_left2
@@ -723,10 +733,11 @@ let recover_initial_subpattern_names = List.map2 set_declaration_name
let recover_alias_names get_name = List.map2 (fun x (_,c,t) ->(get_name x,c,t))
let push_rels_eqn sign eqn =
- {eqn with rhs = {eqn.rhs with rhs_env = push_rels sign eqn.rhs.rhs_env} }
+ {eqn with
+ rhs = {eqn.rhs with rhs_env = push_rel_context sign eqn.rhs.rhs_env} }
let push_rels_eqn_with_names sign eqn =
- let subpats = List.rev (list_firstn (List.length sign) eqn.patterns) in
+ let subpats = List.rev (List.firstn (List.length sign) eqn.patterns) in
let subpatnames = List.map alias_of_pat subpats in
let sign = recover_initial_subpattern_names subpatnames sign in
push_rels_eqn sign eqn
@@ -795,7 +806,7 @@ Some hints:
let rec map_predicate f k ccl = function
| [] -> f k ccl
- | Pushed ((_,tm),_,na) :: rest ->
+ | Pushed (_,((_,tm),_,na)) :: rest ->
let k' = length_of_tomatch_type_sign na tm in
map_predicate f (k+k') ccl rest
| (Alias _ | NonDepAlias) :: rest ->
@@ -821,10 +832,14 @@ let subst_predicate (args,copt) ccl tms =
substnl_predicate sigma 0 ccl tms
let specialize_predicate_var (cur,typ,dep) tms ccl =
- let c = if dep<>Anonymous then Some cur else None in
+ let c = match dep with
+ | Anonymous -> None
+ | Name _ -> Some cur
+ in
let l =
match typ with
- | IsInd (_,IndType(_,realargs),names) -> if names<>[] then realargs else []
+ | IsInd (_, IndType (_, _), []) -> []
+ | IsInd (_, IndType (_, realargs), names) -> realargs
| NotInd _ -> [] in
subst_predicate (l,c) ccl tms
@@ -838,7 +853,9 @@ let specialize_predicate_var (cur,typ,dep) tms ccl =
(* 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 =
- if na=Anonymous then anomaly "Undetected dependency";
+ let () = match na with
+ | 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
@@ -862,16 +879,23 @@ let rec extract_predicate ccl = function
extract_predicate ccl tms
| Abstract (i,d)::tms ->
mkProd_wo_LetIn d (extract_predicate ccl tms)
- | Pushed ((cur,NotInd _),_,na)::tms ->
- let tms = if na<>Anonymous then lift_tomatch_stack 1 tms else tms in
- let pred = extract_predicate ccl tms in
- if na<>Anonymous then subst1 cur pred else pred
- | Pushed ((cur,IsInd (_,IndType(_,realargs),_)),_,na)::tms ->
+ | Pushed (_,((cur,NotInd _),_,na))::tms ->
+ begin match na with
+ | Anonymous -> extract_predicate ccl tms
+ | Name _ ->
+ let tms = lift_tomatch_stack 1 tms in
+ let pred = extract_predicate ccl tms in
+ subst1 cur pred
+ end
+ | Pushed (_,((cur,IsInd (_,IndType(_,realargs),_)),_,na))::tms ->
let realargs = List.rev realargs in
- let k = if na<>Anonymous then 1 else 0 in
+ let k, nrealargs = match na with
+ | Anonymous -> 0, realargs
+ | Name _ -> 1, (cur :: realargs)
+ in
let tms = lift_tomatch_stack (List.length realargs + k) tms in
let pred = extract_predicate ccl tms in
- substl (if na<>Anonymous then cur::realargs else realargs) pred
+ substl nrealargs pred
| [] ->
ccl
@@ -890,7 +914,10 @@ let abstract_predicate env sigma indf cur realargs (names,na) tms ccl =
(* Pred is already dependent in the current term to match (if *)
(* (na<>Anonymous) and its realargs; we just need to adjust it to *)
(* full sign if dep in cur is not taken into account *)
- let ccl = if na <> Anonymous then ccl else lift_predicate 1 ccl tms in
+ let ccl = match na with
+ | Anonymous -> lift_predicate 1 ccl tms
+ | Name _ -> ccl
+ in
let pred = extract_predicate ccl tms in
(* Build the predicate properly speaking *)
let sign = List.map2 set_declaration_name (na::names) sign in
@@ -906,28 +933,40 @@ 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 aliasnames =
+ List.map_filter (function Alias _ | NonDepAlias -> Some Anonymous | _ -> None) tomatch
+ in
+ [ { patterns = pats;
+ rhs = { rhs_env = pb.env;
+ rhs_vars = [];
+ avoid_ids = [];
+ it = None };
+ alias_stack = Anonymous::aliasnames;
+ eqn_loc = Loc.ghost;
+ used = ref false } ]
+
let adjust_impossible_cases pb pred tomatch submat =
- if submat = [] then
- match kind_of_term (whd_evar !(pb.evdref) pred) with
- | Evar (evk,_) when snd (evar_source evk !(pb.evdref)) = ImpossibleCase ->
- let default = (coq_unit_judge ()).uj_type in
- pb.evdref := Evd.define evk default !(pb.evdref);
- (* we add an "assert false" case *)
- let pats = List.map (fun _ -> PatVar (dummy_loc,Anonymous)) tomatch in
- let aliasnames =
- map_succeed (function Alias _ | NonDepAlias -> Anonymous | _ -> failwith"") tomatch
- in
- [ { patterns = pats;
- rhs = { rhs_env = pb.env;
- rhs_vars = [];
- avoid_ids = [];
- it = None };
- alias_stack = Anonymous::aliasnames;
- eqn_loc = dummy_loc;
- used = ref false } ]
+ match submat with
+ | [] ->
+ begin match kind_of_term 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
+ end;
+ add_assert_false_case pb tomatch
| _ ->
submat
- else
+ end
+ | _ ->
submat
(*****************************************************************************)
@@ -957,7 +996,8 @@ let adjust_impossible_cases pb pred tomatch submat =
let specialize_predicate newtomatchs (names,depna) arsign cs tms ccl =
(* Assume some gamma st: gamma |- PI [X,x:I(X)]. PI tms. ccl *)
let nrealargs = List.length names in
- let k = nrealargs + (if depna<>Anonymous then 1 else 0) in
+ let l = match depna with Anonymous -> 0 | Name _ -> 1 in
+ let k = nrealargs + l in
(* We adjust pred st: gamma, x1..xn |- PI [X,x:I(X)]. PI tms. ccl' *)
(* so that x can later be instantiated by Ci(x1..xn) *)
(* and X by the realargs for Ci *)
@@ -965,12 +1005,14 @@ let specialize_predicate newtomatchs (names,depna) arsign cs tms ccl =
let ccl' = liftn_predicate n (k+1) ccl tms in
(* We prepare the substitution of X and x:I(X) *)
let realargsi =
- if nrealargs <> 0 then
+ if not (Int.equal nrealargs 0) then
adjust_subst_to_rel_context arsign (Array.to_list cs.cs_concl_realargs)
else
[] in
- let copti =
- if depna<>Anonymous then Some (build_dependent_constructor cs) else None in
+ let copti = match depna with
+ | Anonymous -> None
+ | Name _ -> Some (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'' *)
(* Note: applying the substitution in tms is not important (is it sure?) *)
@@ -992,7 +1034,10 @@ let adjust_predicate_from_tomatch tomatch (current,typ as ct) pb =
let ((_,oldtyp),deps,na) = tomatch in
match typ, oldtyp with
| IsInd (_,_,names), NotInd _ ->
- let k = if na <> Anonymous then 2 else 1 in
+ let k = match na with
+ | Anonymous -> 1
+ | Name _ -> 2
+ in
let n = List.length names in
{ pb with pred = liftn_predicate n k pb.pred pb.tomatch },
(ct,List.map (fun i -> if i >= k then i+n else i) deps,na)
@@ -1004,7 +1049,7 @@ let adjust_predicate_from_tomatch tomatch (current,typ as ct) pb =
let rec ungeneralize n ng body =
match kind_of_term body with
- | Lambda (_,_,c) when ng = 0 ->
+ | Lambda (_,_,c) when Int.equal ng 0 ->
subst1 (mkRel n) c
| Lambda (na,t,c) ->
(* We traverse an inner generalization *)
@@ -1019,7 +1064,7 @@ let rec ungeneralize n ng body =
let sign2,p = decompose_prod_n_assum ng p in
let p = prod_applist 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 ->
+ mkCase (ci,p,c,Array.map2 (fun q c ->
let sign,b = decompose_lam_n_assum q c in
it_mkLambda_or_LetIn (ungeneralize (n+q) ng b) sign)
ci.ci_cstr_ndecls brs)
@@ -1032,15 +1077,44 @@ let rec ungeneralize n ng body =
let ungeneralize_branch n k (sign,body) cs =
(sign,ungeneralize (n+cs.cs_nargs) k body)
+let rec is_dependent_generalization ng body =
+ match kind_of_term body with
+ | Lambda (_,_,c) when Int.equal ng 0 ->
+ dependent (mkRel 1) c
+ | Lambda (na,t,c) ->
+ (* We traverse an inner generalization *)
+ is_dependent_generalization (ng-1) c
+ | LetIn (na,b,t,c) ->
+ (* We traverse an alias *)
+ is_dependent_generalization ng c
+ | Case (ci,p,c,brs) ->
+ (* We traverse a split *)
+ Array.exists2 (fun q c ->
+ let _,b = decompose_lam_n_assum q c in is_dependent_generalization 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 false
+
+let is_dependent_branch k (_,br) =
+ is_dependent_generalization 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_rel_declaration (nf_evar evd) d in
- if List.exists (fun c -> dependent_decl (lift k c) d) tocheck || pi2 d <> None then
+ let is_d = match d with (_, None, _) -> false | _ -> true in
+ if is_d || List.exists (fun c -> dependent_decl (lift k c) d) tocheck
+ && Array.exists (is_dependent_branch 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 = if pi2 d = None then mkRel n::inst else inst in
+ let inst = match d with
+ | (_, None, _) -> mkRel n :: inst
+ | _ -> inst
+ in
brs, Abstract (i,d) :: tomatch, pred, inst
else
(* Finally, no dependency remains, so, we can replace the generalized *)
@@ -1049,7 +1123,7 @@ let postprocess_dependencies evd tocheck brs tomatch pred deps cs =
let pred = lift_predicate (-1) pred tomatch in
let tomatch = relocate_index_tomatch 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 n k) brs cs in
aux k brs tomatch pred tocheck deps
| _ -> assert false
in aux 0 brs tomatch pred tocheck deps
@@ -1062,8 +1136,8 @@ let rec irrefutable env = function
| PatCstr (_,cstr,args,_) ->
let ind = inductive_of_constructor cstr in
let (_,mip) = Inductive.lookup_mind_specif env ind in
- let one_constr = Array.length mip.mind_user_lc = 1 in
- one_constr & List.for_all (irrefutable env) args
+ let one_constr = Int.equal (Array.length mip.mind_user_lc) 1 in
+ one_constr && List.for_all (irrefutable env) args
let first_clause_irrefutable env = function
| eqn::mat -> List.for_all (irrefutable env) eqn.patterns
@@ -1072,8 +1146,8 @@ let first_clause_irrefutable env = function
let group_equations pb ind current cstrs mat =
let mat =
if first_clause_irrefutable pb.env mat then [List.hd mat] else mat in
- let brs = Array.create (Array.length cstrs) [] in
- let only_default = ref true in
+ let brs = Array.make (Array.length cstrs) [] in
+ let only_default = ref None in
let _ =
List.fold_right (* To be sure it's from bottom to top *)
(fun eqn () ->
@@ -1085,12 +1159,13 @@ let group_equations pb ind current cstrs mat =
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
+ done;
+ if !only_default == None then only_default := Some true
| PatCstr (loc,((_,i)),args,name) ->
(* This is a regular clause *)
- only_default := false;
+ only_default := Some false;
brs.(i-1) <- (args, name, rest) :: brs.(i-1)) mat () in
- (brs,!only_default)
+ (brs,Option.default false !only_default)
(************************************************************************)
(* Here starts the pattern-matching compilation algorithm *)
@@ -1101,14 +1176,17 @@ let rec generalize_problem names pb = function
| i::l ->
let (na,b,t as d) = map_rel_declaration (lift i) (Environ.lookup_rel i pb.env) in
let pb',deps = generalize_problem names pb l in
- if na = Anonymous & b <> None then pb',deps else
- let d = on_pi3 (whd_betaiota !(pb.evdref)) d in (* for better rendering *)
- let tomatch = lift_tomatch_stack 1 pb'.tomatch in
- let tomatch = relocate_index_tomatch (i+1) 1 tomatch in
- { pb' with
- tomatch = Abstract (i,d) :: tomatch;
- pred = generalize_predicate names i d pb'.tomatch pb'.pred },
- i::deps
+ begin match (na, b) with
+ | Anonymous, Some _ -> pb', deps
+ | _ ->
+ let d = on_pi3 (whd_betaiota !(pb.evdref)) d in (* for better rendering *)
+ let tomatch = lift_tomatch_stack 1 pb'.tomatch in
+ let tomatch = relocate_index_tomatch (i+1) 1 tomatch in
+ { pb' with
+ tomatch = Abstract (i,d) :: tomatch;
+ pred = generalize_predicate names i d pb'.tomatch pb'.pred },
+ i::deps
+ end
(* No more patterns: typing the right-hand side of equations *)
let build_leaf pb =
@@ -1117,10 +1195,12 @@ let build_leaf pb =
j_nf_evar !(pb.evdref) j
(* Build the sub-pattern-matching problem for a given branch "C x1..xn as x" *)
-let build_branch current realargs deps (realnames,curname) pb arsign eqns const_info =
+(* spiwack: the [initial] argument keeps track whether the branch is a
+ toplevel branch ([true]) or a deep one ([false]). *)
+let build_branch initial current realargs deps (realnames,curname) pb arsign eqns const_info =
(* We remember that we descend through constructor C *)
let history =
- push_history_pattern const_info.cs_nargs const_info.cs_cstr pb.history in
+ push_history_pattern const_info.cs_nargs (fst const_info.cs_cstr) pb.history in
(* We prepare the matching on x1:T1 .. xn:Tn using some heuristic to *)
(* build the name x1..xn from the names present in the equations *)
@@ -1137,9 +1217,9 @@ let build_branch current realargs deps (realnames,curname) pb arsign eqns const_
(* We adjust the terms to match in the context they will be once the *)
(* context [x1:T1,..,xn:Tn] will have been pushed on the current env *)
let typs' =
- list_map_i (fun i d -> (mkRel i,map_rel_declaration (lift i) d)) 1 typs in
+ List.map_i (fun i d -> (mkRel i,map_rel_declaration (lift i) d)) 1 typs in
- let extenv = push_rels typs pb.env in
+ let extenv = push_rel_context typs pb.env in
let typs' =
List.map (fun (c,d) ->
@@ -1176,10 +1256,11 @@ let build_branch current realargs deps (realnames,curname) pb arsign eqns const_
let typs' =
List.map2
(fun (tm,(tmtyp,_),(na,_,_)) deps ->
- let na = match curname with
- | Name _ -> (if na <> Anonymous then na else curname)
- | Anonymous ->
- if deps = [] && pred_is_not_dep then Anonymous else force_name na in
+ let na = match curname, na with
+ | Name _, Anonymous -> curname
+ | Name _, Name _ -> na
+ | Anonymous, _ ->
+ if List.is_empty deps && pred_is_not_dep then Anonymous else force_name na in
((tm,tmtyp),deps,na))
typs' (List.rev dep_sign) in
@@ -1187,26 +1268,29 @@ let build_branch current realargs deps (realnames,curname) pb arsign eqns const_
let pred =
specialize_predicate typs' (realnames,curname) arsign const_info tomatch pb.pred in
- let currents = List.map (fun x -> Pushed x) typs' in
+ let currents = List.map (fun x -> Pushed (false,x)) typs' in
- let alias =
- if aliasname = Anonymous then
+ let alias = match aliasname with
+ | Anonymous ->
NonDepAlias
- else
+ | Name _ ->
let cur_alias = lift const_info.cs_nargs current in
let ind =
appvect (
- applist (mkInd (inductive_of_constructor const_info.cs_cstr),
+ 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
- Alias (aliasname,cur_alias,(ci,ind)) in
+ Alias (initial,(aliasname,cur_alias,(ci,ind))) in
let tomatch = List.rev_append (alias :: currents) tomatch in
let submat = adjust_impossible_cases pb pred tomatch submat in
- if submat = [] then
+ let () = match submat with
+ | [] ->
raise_pattern_matching_error
- (dummy_loc, pb.env, NonExhaustive (complete_history history));
+ (Loc.ghost, pb.env, Evd.empty, NonExhaustive (complete_history history))
+ | _ -> ()
+ in
typs,
{ pb with
@@ -1227,38 +1311,48 @@ let build_branch current realargs deps (realnames,curname) pb arsign eqns const_
*)
+let mk_case pb (ci,pred,c,brs) =
+ let mib = lookup_mind (fst ci.ci_ind) pb.env in
+ match mib.mind_record with
+ | Some (Some (_, cs, pbs)) ->
+ Reduction.beta_appvect brs.(0)
+ (Array.map (fun p -> mkProj (Projection.make p true, c)) cs)
+ | _ -> mkCase (ci,pred,c,brs)
+
(**********************************************************************)
(* Main compiling descent *)
let rec compile pb =
match pb.tomatch with
| Pushed cur :: rest -> match_current { pb with tomatch = rest } cur
- | Alias x :: rest -> compile_alias pb x rest
+ | Alias (initial,x) :: rest -> compile_alias initial pb x rest
| NonDepAlias :: rest -> compile_non_dep_alias pb rest
| Abstract (i,d) :: rest -> compile_generalization pb i d rest
| [] -> build_leaf pb
(* Case splitting *)
-and match_current pb tomatch =
+and match_current pb (initial,tomatch) =
let tm = adjust_tomatch_to_pattern pb tomatch in
let pb,tomatch = adjust_predicate_from_tomatch tomatch tm pb in
let ((current,typ),deps,dep) = tomatch in
match typ with
| NotInd (_,typ) ->
- check_all_variables typ pb.mat;
- shift_problem tomatch pb
+ check_all_variables pb.env !(pb.evdref) typ pb.mat;
+ compile_all_variables initial tomatch pb
| IsInd (_,(IndType(indf,realargs) as indt),names) ->
let mind,_ = dest_ind_family indf in
+ let mind = Tacred.check_privacy pb.env mind in
let cstrs = get_constructors pb.env indf in
let arsign, _ = get_arity pb.env indf in
- let eqns,onlydflt = group_equations pb mind current cstrs pb.mat in
- if (Array.length cstrs <> 0 or pb.mat <> []) & onlydflt then
- shift_problem tomatch pb
+ let eqns,onlydflt = group_equations pb (fst mind) current cstrs pb.mat in
+ let no_cstr = Int.equal (Array.length cstrs) 0 in
+ if (not no_cstr || not (List.is_empty pb.mat)) && onlydflt then
+ compile_all_variables initial tomatch pb
else
(* We generalize over terms depending on current term to match *)
let pb,deps = generalize_problem (names,dep) pb deps in
(* We compile branches *)
- let brvals = array_map2 (compile_branch current realargs (names,dep) deps pb arsign) eqns cstrs in
+ 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 brvals,tomatch,pred,inst =
@@ -1269,14 +1363,16 @@ and match_current pb tomatch =
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 mind pb.casestyle in
+ let ci = make_case_info pb.env (fst mind) pb.casestyle in
let pred = nf_betaiota !(pb.evdref) pred in
- let case = mkCase (ci,pred,current,brvals) in
+ let case = mk_case pb (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 }
-(* Building the sub-problem when all patterns are variables *)
+
+(* Building the sub-problem when all patterns are variables. Case
+ where [current] is an intially pushed term. *)
and shift_problem ((current,t),_,na) pb =
let ty = type_of_tomatch t in
let tomatch = lift_tomatch_stack 1 pb.tomatch in
@@ -1292,9 +1388,27 @@ and shift_problem ((current,t),_,na) pb =
{ uj_val = subst1 current j.uj_val;
uj_type = subst1 current j.uj_type }
+(* Building the sub-problem when all patterns are variables,
+ non-initial case. Variables which appear as subterms of constructor
+ 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 pb =
+ { pb with
+ pred = pred;
+ history = pop_history pb.history;
+ mat = List.map push_noalias_current_pattern pb.mat } in
+ compile pb
+
+(* Building the sub-problem when all patterns are variables. *)
+and compile_all_variables initial cur pb =
+ if initial then shift_problem cur pb
+ else pop_problem cur pb
+
(* Building the sub-problem when all patterns are variables *)
-and compile_branch current realargs names deps pb arsign eqns cstr =
- let sign, pb = build_branch current realargs deps names pb arsign eqns cstr in
+and compile_branch initial current realargs names deps pb arsign eqns cstr =
+ let sign, pb = build_branch initial current realargs deps names pb arsign eqns cstr in
sign, (compile pb).uj_val
(* Abstract over a declaration before continuing splitting *)
@@ -1308,7 +1422,10 @@ and compile_generalization pb i d rest =
{ uj_val = mkLambda_or_LetIn d j.uj_val;
uj_type = mkProd_wo_LetIn d j.uj_type }
-and compile_alias pb (na,orig,(expanded,expanded_typ)) rest =
+(* spiwack: the [initial] argument keeps track whether the alias has
+ been introduced by a toplevel branch ([true]) or a deep one
+ ([false]). *)
+and compile_alias initial pb (na,orig,(expanded,expanded_typ)) rest =
let f c t =
let alias = (na,Some c,t) in
let pb =
@@ -1325,18 +1442,35 @@ and compile_alias pb (na,orig,(expanded,expanded_typ)) rest =
else
mkLetIn (na,c,t,j.uj_val);
uj_type = subst1 c j.uj_type } in
- if isRel orig or isVar orig then
+ (* spiwack: when an alias appears on a deep branch, its non-expanded
+ form is automatically a variable of the same name. We avoid
+ introducing such superfluous aliases so that refines are elegant. *)
+ let just_pop () =
+ let pb =
+ { pb with
+ tomatch = rest;
+ history = pop_history_pattern pb.history;
+ mat = List.map drop_alias_eqn pb.mat } in
+ compile pb
+ in
+ let sigma = !(pb.evdref) in
+ if not (Flags.is_program_mode ()) && (isRel orig || isVar orig) then
(* Try to compile first using non expanded alias *)
- try f orig (Retyping.get_type_of pb.env !(pb.evdref) orig)
+ try
+ if initial then f orig (Retyping.get_type_of pb.env !(pb.evdref) orig)
+ else just_pop ()
with e when precatchable_exception e ->
(* Try then to compile using expanded alias *)
+ pb.evdref := sigma;
f expanded expanded_typ
else
(* Try to compile first using expanded alias *)
try f expanded expanded_typ
with e when precatchable_exception e ->
(* Try then to compile using non expanded alias *)
- f orig (Retyping.get_type_of pb.env !(pb.evdref) orig)
+ pb.evdref := sigma;
+ if initial then f orig (Retyping.get_type_of pb.env !(pb.evdref) orig)
+ else just_pop ()
(* Remember that a non-trivial pattern has been consumed *)
@@ -1357,7 +1491,7 @@ 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 *)
-let matx_of_eqns env tomatchl eqns =
+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
@@ -1412,27 +1546,30 @@ let adjust_to_extended_env_and_remove_deps env extenv subst t =
- [subst0] is made of items [(p,u,(u,ty))] where [ty] is the type of [u]
and both are adjusted to [extenv] while [p] is the index of [id] in
[extenv] (after expansion of the aliases) *)
- let subst0 = map_succeed (fun (x,u) ->
+ let map (x, u) =
(* d1 ... dn dn+1 ... dn'-p+1 ... dn' *)
(* \--env-/ (= x:ty) *)
(* \--------------extenv------------/ *)
- let (p,_,_) = lookup_rel_id x (rel_context extenv) in
+ let (p, _, _) = lookup_rel_id x (rel_context extenv) in
let rec traverse_local_defs p =
match pi2 (lookup_rel p extenv) with
| Some c -> assert (isRel c); traverse_local_defs (p + destRel c)
| None -> p in
let p = traverse_local_defs p in
- let u = lift (n'-n) u in
- (p,u,expand_vars_in_term extenv u)) subst in
- let t0 = lift (n'-n) t in
- (subst0,t0)
+ let u = lift (n' - n) u in
+ try Some (p, u, expand_vars_in_term extenv u)
+ (* pedrot: does this really happen to raise [Failure _]? *)
+ with Failure _ -> None in
+ let subst0 = List.map_filter map subst in
+ let t0 = lift (n' - n) t in
+ (subst0, t0)
let push_binder d (k,env,subst) =
(k+1,push_rel d env,List.map (fun (na,u,d) -> (na,lift 1 u,d)) subst)
let rec list_assoc_in_triple x = function
[] -> raise Not_found
- | (a,b,_)::l -> if compare a x = 0 then b else list_assoc_in_triple x l
+ | (a, b, _)::l -> if Int.equal a x then b else list_assoc_in_triple x l
(* Let vijk and ti be a set of dependent terms and T a type, all
* defined in some environment env. The vijk and ti are supposed to be
@@ -1449,9 +1586,11 @@ let rec list_assoc_in_triple x = function
* similarly for each ti.
*)
-let abstract_tycon loc env evdref subst _tycon extenv t =
- let sigma = !evdref in
- let t = nf_betaiota sigma t in (* it helps in some cases to remove K-redex *)
+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
(* We traverse the type T of the original problem Xi looking for subterms
that match the non-constructor part of the constraints (this part
@@ -1460,42 +1599,49 @@ let abstract_tycon loc env evdref subst _tycon extenv t =
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
- | Rel n when pi2 (lookup_rel n env) <> None ->
- map_constr_with_full_binders push_binder aux x t
+ | Rel n when pi2 (lookup_rel n env) != None -> t
| Evar ev ->
- let ty = get_type_of env sigma t in
+ let ty = get_type_of env !evdref t in
+ let ty = Evarutil.evd_comb1 (refresh_universes (Some false) env) evdref ty in
let inst =
- list_map_i
+ List.map_i
(fun i _ ->
try list_assoc_in_triple i subst0 with Not_found -> mkRel i)
1 (rel_context env) in
- let ev = e_new_evar evdref env ~src:(loc, CasesType) ty in
- evdref := add_conv_pb (Reduction.CONV,env,substl inst ev,t) !evdref;
- ev
+ let ev' = e_new_evar env evdref ~src ty in
+ begin match solve_simple_eqn (evar_conv_x full_transparent_state) env !evdref (None,ev,substl inst ev') with
+ | Success evd -> evdref := evd
+ | UnifFailure _ -> assert false
+ end;
+ ev'
| _ ->
- let good = List.filter (fun (_,u,_) -> is_conv_leq env sigma t u) subst in
- if good <> [] then
- let u = pi3 (List.hd good) in (* u is in extenv *)
+ 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
+ | (_, _, u) :: _ -> (* u is in extenv *)
let vl = List.map pi1 good in
- let ty = lift (-k) (aux x (get_type_of env !evdref t)) in
+ let ty =
+ let ty = get_type_of env !evdref t in
+ 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 inst =
- list_map_i
- (fun i _ -> if List.mem i vl then u else mkRel i) 1
+ 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
- || Intset.mem (destRel a) depvl) inst in
+ || Int.Set.mem (destRel a) depvl) inst in
let named_filter =
List.map (fun (id,_,_) -> dependent (mkVar id) u)
(named_context extenv) in
- let filter = rel_filter@named_filter in
+ let filter = Filter.make (rel_filter @ named_filter) in
let candidates = u :: List.map mkRel vl in
- let ev =
- e_new_evar evdref extenv ~src:(loc, CasesType) ~filter ~candidates ty in
+ let ev = e_new_evar extenv evdref ~src ~filter ~candidates ty in
lift k ev
- else
- map_constr_with_full_binders push_binder aux x t in
+ in
aux (0,extenv,subst0) t0
let build_tycon loc env tycon_env subst tycon extenv evdref t =
@@ -1505,10 +1651,9 @@ let build_tycon loc env tycon_env subst tycon extenv evdref t =
we are in an impossible branch *)
let n = rel_context_length (rel_context env) in
let n' = rel_context_length (rel_context tycon_env) in
- let tt = new_Type () in
- let impossible_case_type =
- e_new_evar evdref env ~src:(loc,ImpossibleCase) tt in
- (lift (n'-n) impossible_case_type, tt)
+ let impossible_case_type, u =
+ e_new_type_evar (reset_context env) evdref univ_flexible_alg ~src:(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 evd,tt = Typing.e_type_of extenv !evdref t in
@@ -1529,33 +1674,33 @@ let build_tycon loc env tycon_env 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 (dummy_loc,Name id), ((id,t)::subst, id::avoid) in
+ PatVar (Loc.ghost,Name id), ((id,t)::subst, id::avoid) in
let rec reveal_pattern t (subst,avoid as acc) =
match kind_of_term (whd_betadeltaiota env sigma t) with
- | Construct cstr -> PatCstr (dummy_loc,cstr,[],Anonymous), acc
+ | Construct (cstr,u) -> PatCstr (Loc.ghost,cstr,[],Anonymous), acc
| App (f,v) when isConstruct f ->
- let cstr = destConstruct f in
- let n = constructor_nrealargs 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 (dummy_loc,cstr,l,Anonymous), acc
+ let cstr,u = destConstruct 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
| _ -> 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_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 = recover_alias_names alias_of_pat (pat :: List.rev patl) sign in
let p = List.length realargs in
- let env' = push_rels sign env in
+ let env' = push_rel_context sign env in
let patl',acc_sign,acc = aux (n+p+1) env' (sign@acc_sign) tms acc in
patl@pat::patl',acc_sign,acc
| (t, NotInd (bo,typ)) :: tms ->
let pat,acc = make_patvar t acc in
- let d = (alias_of_pat pat,None,t) in
+ let d = (alias_of_pat pat,None,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
@@ -1572,19 +1717,19 @@ let build_inversion_problem loc env sigma tms t =
let n = List.length sign in
let decls =
- list_map_i (fun i d -> (mkRel i,map_rel_declaration (lift i) d)) 1 sign in
+ List.map_i (fun i d -> (mkRel i,map_rel_declaration (lift i) d)) 1 sign in
- let pb_env = push_rels sign env in
+ let pb_env = push_rel_context sign env in
let decls =
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 (List.make n true) decls in
let sub_tms =
List.map2 (fun deps (tm,(tmtyp,_),(na,b,t)) ->
- let na = if deps = [] then Anonymous else force_name na in
- Pushed ((tm,tmtyp),deps,na))
+ let na = if List.is_empty deps then Anonymous else force_name na in
+ Pushed (true,((tm,tmtyp),deps,na)))
dep_sign decls in
let subst = List.map (fun (na,t) -> (na,lift n t)) subst in
(* [eqn1] is the first clause of the auxiliary pattern-matching that
@@ -1595,7 +1740,7 @@ let build_inversion_problem loc env sigma tms t =
let eqn1 =
{ patterns = patl;
alias_stack = [];
- eqn_loc = dummy_loc;
+ eqn_loc = Loc.ghost;
used = ref false;
rhs = { rhs_env = pb_env;
(* we assume all vars are used; in practice we discard dependent
@@ -1608,9 +1753,9 @@ let build_inversion_problem loc env sigma tms t =
type constraints are incompatible with the constraints on the
inductive types of the multiple terms matched in Xi *)
let eqn2 =
- { patterns = List.map (fun _ -> PatVar (dummy_loc,Anonymous)) patl;
+ { patterns = List.map (fun _ -> PatVar (Loc.ghost,Anonymous)) patl;
alias_stack = [];
- eqn_loc = dummy_loc;
+ eqn_loc = Loc.ghost;
used = ref false;
rhs = { rhs_env = pb_env;
rhs_vars = [];
@@ -1618,11 +1763,18 @@ let build_inversion_problem loc env sigma tms t =
it = None } } in
(* [pb] is the auxiliary pattern-matching serving as skeleton for the
return type of the original problem Xi *)
+ (* let sigma, s = Evd.new_sort_variable sigma in *)
+(*FIXME TRY *)
+ (* let sigma, s = Evd.new_sort_variable univ_flexible sigma in *)
+ let s' = Retyping.get_sort_of env sigma t in
+ let sigma, s = Evd.new_sort_variable univ_flexible_alg sigma in
+ let sigma = Evd.set_leq_sort env sigma s' s in
let evdref = ref sigma in
+ (* let ty = evd_comb1 (refresh_universes false) evdref ty in *)
let pb =
{ env = pb_env;
evdref = evdref;
- pred = new_Type();
+ pred = (*ty *) mkSort s;
tomatch = sub_tms;
history = start_history n;
mat = [eqn1;eqn2];
@@ -1643,30 +1795,30 @@ let build_initial_predicate arsign pred =
| _ -> assert false
in buildrec 0 pred [] (List.rev arsign)
-let extract_arity_signature env0 tomatchl tmsign =
+let extract_arity_signature ?(dolift=true) env0 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 -> [na,Option.map (lift n) bo,lift n typ]
- | Some (loc,_,_,_) ->
+ | Some (loc,_,_) ->
user_err_loc (loc,"",
str"Unexpected type annotation for a term of non inductive type."))
| IsInd (term,IndType(indf,realargs),_) ->
- let indf' = lift_inductive_family n indf in
- let (ind,_) = dest_ind_family indf' in
- let nparams_ctxt,nrealargs_ctxt = inductive_nargs env0 ind in
+ 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 =
match t with
- | Some (loc,ind',nparams,realnal) ->
- if ind <> ind' then
+ | Some (loc,ind',realnal) ->
+ if not (eq_ind ind ind') then
user_err_loc (loc,"",str "Wrong inductive type.");
- if nparams_ctxt <> nparams
- or nrealargs_ctxt <> List.length realnal then
- anomaly "Ill-formed 'in' clause in cases";
+ 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
+ | None -> List.make nrealargs_ctxt Anonymous in
(na,None,build_dependent_inductive env0 indf')
::(List.map2 (fun x (_,c,t) ->(x,c,t)) realnal arsign) in
let rec buildrec n = function
@@ -1694,7 +1846,7 @@ let prepare_predicate_from_arsign_tycon loc tomatchs arsign c =
let signlen = List.length sign in
match kind_of_term tm with
| Rel n when dependent tm c
- && signlen = 1 (* The term to match is not of a dependent type itself *) ->
+ && 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,
maybe some variable in its type appears in the tycon. *) ->
@@ -1720,7 +1872,7 @@ let prepare_predicate_from_arsign_tycon loc tomatchs arsign c =
| Rel n when n > lift ->
(try
(* Make the predicate dependent on the matched variable *)
- let idx = List.assoc (n - lift) subst in
+ let idx = Int.List.assoc (n - lift) subst in
mkRel (idx + lift)
with Not_found ->
(* A variable that is not matched, lift over the arsign. *)
@@ -1741,11 +1893,11 @@ let prepare_predicate_from_arsign_tycon loc tomatchs arsign c =
* tycon to make the predicate if it is not closed.
*)
-let prepare_predicate loc typing_fun sigma env tomatchs arsign tycon pred =
+let prepare_predicate loc typing_fun env sigma tomatchs arsign tycon pred =
let preds =
match pred, tycon with
(* No type annotation *)
- | None, Some (None, t) when not (noccur_with_meta 0 max_int t) ->
+ | None, Some t when not (noccur_with_meta 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 *)
@@ -1758,8 +1910,12 @@ let prepare_predicate loc typing_fun sigma env tomatchs arsign tycon pred =
(* No dependent type constraint, or no constraints at all: *)
(* we use two strategies *)
let sigma,t = match tycon with
- | Some (None, t) -> sigma,t
- | _ -> new_type_evar sigma env ~src:(loc, CasesType) in
+ | Some t -> sigma,t
+ | None ->
+ let sigma, (t, _) =
+ new_type_evar env sigma univ_flexible_alg ~src:(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 *)
@@ -1768,16 +1924,17 @@ let prepare_predicate loc typing_fun sigma env tomatchs arsign tycon pred =
(* Some type annotation *)
| Some rtntyp, _ ->
(* We extract the signature of the arity *)
- let envar = List.fold_right push_rels arsign env in
- let sigma, newt = new_sort_variable sigma in
+ let envar = List.fold_right push_rel_context 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 sigma = Option.cata (fun tycon ->
- let na = Name (id_of_string "x") in
- let tms = List.map (fun tm -> Pushed(tm,[],na)) tomatchs in
- let predinst = extract_predicate predcclj.uj_val tms in
- Coercion.inh_conv_coerces_to loc env !evdref predinst tycon)
- !evdref tycon in
+ let sigma = !evdref in
+ (* let sigma = Option.cata (fun tycon -> *)
+ (* let na = Name (Id.of_string "x") in *)
+ (* let tms = List.map (fun tm -> Pushed(tm,[],na)) tomatchs in *)
+ (* let predinst = extract_predicate predcclj.uj_val tms in *)
+ (* Coercion.inh_conv_coerce_to loc env !evdref predinst tycon) *)
+ (* !evdref tycon in *)
let predccl = (j_nf_evar sigma predcclj).uj_val in
[sigma, predccl]
in
@@ -1787,23 +1944,537 @@ let prepare_predicate loc typing_fun sigma env tomatchs arsign tycon pred =
sigma,nal,pred)
preds
+(** Program cases *)
+
+open Program
+
+let ($) f x = f x
+
+let string_of_name name =
+ match name with
+ | Anonymous -> "anonymous"
+ | Name n -> Id.to_string n
+
+let make_prime_id name =
+ let str = string_of_name name in
+ Id.of_string str, Id.of_string (str ^ "'")
+
+let prime avoid name =
+ let previd, id = make_prime_id name in
+ previd, next_ident_away id avoid
+
+let make_prime avoid prevname =
+ let previd, id = prime !avoid prevname in
+ avoid := id :: !avoid;
+ previd, id
+
+let eq_id avoid id =
+ let hid = Id.of_string ("Heq_" ^ Id.to_string id) in
+ let hid' = next_ident_away hid avoid in
+ hid'
+
+let mk_eq evdref typ x y = papp evdref coq_eq_ind [| typ; x ; y |]
+let mk_eq_refl evdref typ x = papp evdref coq_eq_refl [| typ; x |]
+let mk_JMeq evdref typ x typ' y =
+ papp evdref coq_JMeq_ind [| 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 true), 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 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
+ in
+ (PatVar (l, name), [name, None, ty] @ realargs, mkRel 1, ty,
+ (List.map (fun x -> mkRel 1) realargs), 1, avoid)
+ | PatCstr (l,((_, 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
+ {uj_val = ty; uj_type = Typing.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 cstrs = get_constructors env indf in
+ let ci = cstrs.(i-1) in
+ let nb_args_constr = ci.cs_nargs in
+ assert (Int.equal nb_args_constr (List.length args));
+ let patargs, args, sign, env, n, m, avoid =
+ List.fold_right2
+ (fun (na, c, t) ua (patargs, args, sign, env, n, m, avoid) ->
+ 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
+ in
+ let args' = arg' :: List.map (lift n') args in
+ let env' = push_rel_context sign' env in
+ (pat' :: patargs, args', sign' @ sign, env', n' + n, succ m, avoid))
+ ci.cs_args (List.rev args) ([], [], [], env, 0, 0, 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 apptype = Retyping.get_type_of env ( !evdref) app 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 = (alias, None, lift m ty) :: sign in
+ let avoid = id :: avoid in
+ let sign, i, avoid =
+ try
+ let env = push_rel_context sign env in
+ evdref := the_conv_x_leq (push_rel_context sign env)
+ (lift (succ m) ty) (lift 1 apptype) !evdref;
+ let eq_t = mk_eq evdref (lift (succ m) ty)
+ (mkRel 1) (* alias *)
+ (lift 1 app) (* aliased term *)
+ in
+ let neq = eq_id avoid id in
+ (Name neq, Some (mkRel 0), eq_t) :: sign, 2, 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 (pi3 (List.hd arsign), List.tl arsign) pat avoid in
+ pat', (sign, patc, (pi3 (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;
+ hid'
+
+let is_topvar t =
+match kind_of_term t with
+| Rel 0 -> true
+| _ -> false
+
+let rels_of_patsign l =
+ List.map (fun ((na, b, t) as x) ->
+ match b with
+ | Some t' when is_topvar t' -> (na, None, t)
+ | _ -> x) l
+
+let vars_of_ctx ctx =
+ let _, y =
+ List.fold_right (fun (na, b, t) (prev, vars) ->
+ match b with
+ | Some t' when is_topvar t' ->
+ prev,
+ (GApp (Loc.ghost,
+ (GRef (Loc.ghost, delayed_force coq_eq_refl_ref, None)),
+ [hole; GVar (Loc.ghost, prev)])) :: vars
+ | _ ->
+ match na with
+ Anonymous -> invalid_arg "vars_of_ctx"
+ | Name n -> n, GVar (Loc.ghost, n) :: vars)
+ ctx (Id.of_string "vars_of_ctx_error", [])
+ in List.rev y
+
+let rec is_included x y =
+ match x, y with
+ | PatVar _, _ -> true
+ | _, PatVar _ -> true
+ | PatCstr (l, (_, i), args, alias), PatCstr (l', (_, i'), args', alias') ->
+ if Int.equal i i' then List.for_all2 is_included args args'
+ else false
+
+(* 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.
+ *)
+let build_ineqs evdref prevpatterns pats liftsign =
+ let _tomatchs = List.length pats in
+ let diffs =
+ List.fold_left
+ (fun c eqnpats ->
+ let acc = List.fold_left2
+ (* ppat is the pattern we are discriminating against, curpat is the current one. *)
+ (fun acc (ppat_sign, ppat_c, (ppat_ty, ppat_tyargs), ppat)
+ (curpat_sign, curpat_c, (curpat_ty, curpat_tyargs), curpat) ->
+ match acc with
+ None -> None
+ | Some (sign, len, n, c) -> (* FixMe: do not work with ppat_args *)
+ if is_included curpat ppat then
+ (* Length of previous pattern's signature *)
+ let lens = List.length ppat_sign in
+ (* Accumulated length of previous pattern's signatures *)
+ let len' = lens + len in
+ let acc =
+ ((* Jump over previous prevpat signs *)
+ lift_rel_context len ppat_sign @ sign,
+ len',
+ succ n, (* nth pattern *)
+ (papp evdref coq_eq_ind
+ [| lift (len' + liftsign) curpat_ty;
+ liftn (len + liftsign) (succ lens) ppat_c ;
+ lift len' curpat_c |]) ::
+ List.map (lift lens (* Jump over this prevpat signature *)) c)
+ in Some acc
+ else None)
+ (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)
+ [] prevpatterns
+ in match diffs with [] -> None
+ | _ -> Some (mk_coq_and diffs)
+
+let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity =
+ let i = ref 0 in
+ let (x, y, z) =
+ List.fold_left
+ (fun (branches, eqns, prevpatterns) eqn ->
+ let _, newpatterns, pats =
+ List.fold_left2
+ (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
+ in
+ let newpatterns = List.rev newpatterns and opats = List.rev pats in
+ let rhs_rels, pats, signlen =
+ List.fold_left
+ (fun (renv, pats, n) (sign,c, (s, args), p) ->
+ (* Recombine signatures and terms of all of the row's patterns *)
+ let sign' = lift_rel_context n sign in
+ let len = List.length sign' in
+ (sign' @ renv,
+ (* lift to get outside of previous pattern's signatures. *)
+ (sign', liftn n (succ len) c,
+ (s, List.map (liftn n (succ len)) args), p) :: pats,
+ len + n))
+ ([], [], 0) opats in
+ let pats, _ = List.fold_left
+ (* 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,
+ (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 _signenv = push_rel_context rhs_rels' env in
+ let arity =
+ let args, nargs =
+ List.fold_right (fun (sign, c, (_, args), _) (allargs,n) ->
+ (args @ c :: allargs, List.length args + succ n))
+ pats ([], 0)
+ in
+ let args = List.rev args in
+ substl args (liftn signlen (succ nargs) arity)
+ in
+ let rhs_rels', tycon =
+ let neqs_rels, arity =
+ match ineqs with
+ | None -> [], arity
+ | Some ineqs ->
+ [Anonymous, None, ineqs], lift 1 arity
+ in
+ let eqs_rels, arity = decompose_prod_n_assum neqs arity in
+ eqs_rels @ neqs_rels @ rhs_rels', arity
+ in
+ let rhs_env = push_rel_context rhs_rels' env in
+ let j = typing_fun (mk_tycon tycon) rhs_env eqn.rhs.it in
+ let bbody = it_mkLambda_or_LetIn j.uj_val rhs_rels'
+ and btype = it_mkProd_or_LetIn j.uj_type rhs_rels' in
+ let _btype = evd_comb1 (Typing.e_type_of env) evdref bbody in
+ let branch_name = Id.of_string ("program_branch_" ^ (string_of_int !i)) in
+ let branch_decl = (Name branch_name, Some (lift !i bbody), (lift !i btype)) in
+ let branch =
+ let bref = GVar (Loc.ghost, branch_name) in
+ match vars_of_ctx rhs_rels with
+ [] -> bref
+ | l -> GApp (Loc.ghost, bref, l)
+ in
+ let branch = match ineqs with
+ Some _ -> GApp (Loc.ghost, branch, [ hole ])
+ | None -> branch
+ in
+ incr i;
+ let rhs = { eqn.rhs with it = Some branch } in
+ (branch_decl :: branches,
+ { eqn with patterns = newpatterns; rhs = rhs } :: eqns,
+ opats :: prevpatterns))
+ ([], [], []) eqns
+ in x, y
+
+(* Builds the predicate. If the predicate is dependent, its context is
+ * made of 1+nrealargs assumptions for each matched term in an inductive
+ * type and 1 assumption for each term not _syntactically_ in an
+ * inductive type.
+
+ * Each matched terms are independently considered dependent or not.
+
+ * A type constraint but no annotation case: it is assumed non dependent.
+ *)
+
+let lift_ctx n ctx =
+ let ctx', _ =
+ List.fold_right (fun (c, t) (ctx, n') ->
+ (liftn n n' c, liftn_tomatch_type n n' t) :: ctx, succ n')
+ ctx ([], 0)
+ in ctx'
+
+(* Turn matched terms into variables. *)
+let abstract_tomatch env 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
+ 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
+ let name = next_ident_away (Id.of_string "filtered_var") names in
+ (mkRel 1, lift_tomatch_type (succ lenctx) t) :: lift_ctx 1 prev,
+ (Name name, Some (lift lenctx c), lift lenctx $ type_of_tomatch t) :: ctx,
+ name :: names, tycon)
+ ([], [], [], 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 pi1) 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
+ (fun (eqs, neqs, refl_args, slift, arsigns) (tm, ty) arsign ->
+ (* The accumulator:
+ previous eqs,
+ number of previous eqs,
+ lift to get outside eqs and in the introduced variables ('as' and 'in'),
+ new arity signatures
+ *)
+ match ty with
+ | IsInd (ty, IndType (indf, args), _) when List.length args > 0 ->
+ (* Build the arity signature following the names in matched terms
+ as much as possible *)
+ let argsign = List.tl arsign in (* arguments in inverse application order *)
+ let (appn, appb, appt) as _appsign = List.hd arsign in (* The matched argument *)
+ 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 (name, b, t) ->
+ let argt = Retyping.get_type_of env !evdref arg in
+ let eq, refl_arg =
+ if Reductionops.is_conv env !evdref argt t then
+ (mk_eq evdref (lift (nargeqs + slift) argt)
+ (mkRel (nargeqs + slift))
+ (lift (nargeqs + nar) arg),
+ mk_eq_refl evdref argt arg)
+ else
+ (mk_JMeq evdref (lift (nargeqs + slift) t)
+ (mkRel (nargeqs + slift))
+ (lift (nargeqs + nar) argt)
+ (lift (nargeqs + nar) arg),
+ mk_JMeq_refl evdref argt arg)
+ in
+ let previd, id =
+ let name =
+ match kind_of_term arg with
+ Rel n -> pi1 (lookup_rel n env)
+ | _ -> name
+ in
+ make_prime avoid name
+ in
+ (env, succ nargeqs,
+ (Name (eq_id avoid previd), None, eq) :: argeqs,
+ refl_arg :: refl_args,
+ pred slift,
+ (Name id, b, t) :: argsign'))
+ (env, neqs, [], [], slift, []) args argsign
+ in
+ let eq = mk_JMeq evdref
+ (lift (nargeqs + slift) appt)
+ (mkRel (nargeqs + slift))
+ (lift (nargeqs + nar) ty)
+ (lift (nargeqs + nar) tm)
+ in
+ let refl_eq = mk_JMeq_refl evdref ty tm in
+ let previd, id = make_prime avoid appn in
+ (((Name (eq_id avoid previd), None, eq) :: argeqs) :: eqs,
+ succ nargeqs,
+ refl_eq :: refl_args,
+ pred slift,
+ (((Name id, appb, appt) :: argsign') :: arsigns))
+
+ | _ -> (* Non dependent inductive or not inductive, just use a regular equality *)
+ let (name, b, typ) = match arsign with [x] -> x | _ -> assert(false) in
+ let previd, id = make_prime avoid name in
+ let arsign' = (Name id, b, typ) in
+ let tomatch_ty = type_of_tomatch ty in
+ let eq =
+ mk_eq evdref (lift nar tomatch_ty)
+ (mkRel slift) (lift nar tm)
+ in
+ ([(Name (eq_id avoid previd), None, eq)] :: eqs, succ neqs,
+ (mk_eq_refl evdref tomatch_ty tm) :: refl_args,
+ pred slift, (arsign' :: []) :: arsigns))
+ ([], 0, [], nar, []) tomatchs arsign
+ in
+ let arsign'' = List.rev arsign' in
+ assert(Int.equal slift 0); (* we must have folded over all elements of the arity signature *)
+ arsign'', allnames, nar, eqs, neqs, refls
+
+let context_of_arsign l =
+ let (x, _) = List.fold_right
+ (fun c (x, n) ->
+ (lift_rel_context n c @ x, List.length c + n))
+ l ([], 0)
+ in x
+
+let compile_program_cases loc style (typing_function, evdref) tycon env
+ (predopt, tomatchl, eqns) =
+ let typing_fun tycon env = function
+ | Some t -> typing_function tycon env evdref t
+ | None -> Evarutil.evd_comb0 use_unit_judge evdref in
+
+ (* We build the matrix of patterns and right-hand side *)
+ let matx = matx_of_eqns env eqns in
+
+ (* 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 tycon = valcon_of_tycon tycon in
+ let tomatchs, tomatchs_lets, tycon' = abstract_tomatch env 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
+ (* Build the dependent arity signature, the equalities which makes
+ the first part of the predicate and their instantiations. *)
+ let avoid = [] in
+ build_dependent_signature env evdref avoid tomatchs arsign
+
+ in
+ let tycon, arity =
+ match tycon' with
+ | None -> let ev = mkExistential env evdref in ev, ev
+ | Some t ->
+ let pred =
+ try
+ let pred = prepare_predicate_from_arsign_tycon loc tomatchs sign t in
+ (* The tycon may be ill-typed after abstraction. *)
+ let env' = push_rel_context (context_of_arsign sign) env in
+ ignore(Typing.sort_of env' evdref pred); pred
+ with e when Errors.noncritical e ->
+ let nar = List.fold_left (fun n sign -> List.length sign + n) 0 sign in
+ lift nar t
+ in Option.get tycon, pred
+ in
+ let neqs, arity =
+ let ctx = context_of_arsign eqs in
+ let neqs = List.length ctx in
+ neqs, it_mkProd_or_LetIn (lift neqs arity) ctx
+ in
+ let lets, matx =
+ (* Type the rhs under the assumption of equations *)
+ constrs_of_pats typing_fun env evdref matx tomatchs sign neqs arity
+ in
+ let matx = List.rev matx in
+ let _ = assert (Int.equal len (List.length lets)) in
+ let env = push_rel_context lets env in
+ let matx = List.map (fun eqn -> { eqn with rhs = { eqn.rhs with rhs_env = env } }) matx in
+ let tomatchs = List.map (fun (x, y) -> lift len x, lift_tomatch_type len y) tomatchs in
+ let args = List.rev_map (lift len) args in
+ let pred = liftn len (succ signlen) arity in
+ let nal, pred = build_initial_predicate sign pred in
+
+ (* We push the initial terms to match and push their alias to rhs' envs *)
+ (* names of aliases will be recovered from patterns (hence Anonymous here) *)
+
+ let out_tmt na = function NotInd (c,t) -> (na,c,t) | IsInd (typ,_,_) -> (na,None,typ) in
+ let typs = List.map2 (fun na (tm,tmt) -> (tm,out_tmt na tmt)) nal tomatchs in
+
+ let typs =
+ List.map (fun (c,d) -> (c,extract_inductive_data env !evdref d,d)) typs in
+
+ let dep_sign =
+ find_dependencies_signature
+ (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
+ ((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
+ | None -> evd_comb0 use_unit_judge evdref in
+
+ let pb =
+ { env = env;
+ evdref = evdref;
+ pred = pred;
+ tomatch = initial_pushed;
+ history = start_history (List.length initial_pushed);
+ mat = matx;
+ caseloc = loc;
+ casestyle= style;
+ typing_function = typing_function } in
+
+ 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 j =
+ { uj_val = it_mkLambda_or_LetIn body tomatchs_lets;
+ uj_type = nf_evar !evdref tycon; }
+ in j
+
(**************************************************************************)
(* Main entry of the matching compilation *)
let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, eqns) =
-
+ if predopt == None && Flags.is_program_mode () then
+ compile_program_cases loc style (typing_fun, evdref)
+ tycon env (predopt, tomatchl, eqns)
+ else
+
(* We build the matrix of patterns and right-hand side *)
- let matx = matx_of_eqns env tomatchl eqns in
+ let matx = matx_of_eqns env eqns in
(* 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
+
+
(* 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 !evdref env tomatchs arsign tycon predopt in
+ let preds = prepare_predicate loc typing_fun env !evdref 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 *)
@@ -1818,22 +2489,22 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e
let dep_sign =
find_dependencies_signature
- (list_make (List.length typs) true)
+ (List.make (List.length typs) true)
typs in
let typs' =
- list_map3
+ List.map3
(fun (tm,tmt) deps na ->
let deps = if not (isRel tm) then [] else deps in
((tm,tmt),deps,na))
tomatchs dep_sign nal in
- let initial_pushed = List.map (fun x -> Pushed x) typs' in
+ let initial_pushed = List.map (fun x -> Pushed (true,x)) typs' in
(* 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
- | None -> coq_unit_judge () in
+ | None -> evd_comb0 use_unit_judge evdref in
let myevdref = ref sigma in
@@ -1862,4 +2533,3 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e
(* We coerce to the tycon (if an elim predicate was provided) *)
inh_conv_coerce_to_tycon loc env evdref j tycon
-end
diff --git a/pretyping/cases.mli b/pretyping/cases.mli
index 826d68a4..c599766a 100644
--- a/pretyping/cases.mli
+++ b/pretyping/cases.mli
@@ -1,14 +1,14 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Util
open Names
open Term
+open Context
open Evd
open Environ
open Inductiveops
@@ -23,37 +23,101 @@ type pattern_matching_error =
| BadConstructor of constructor * inductive
| WrongNumargConstructor of constructor * int
| WrongNumargInductive of inductive * int
- | WrongPredicateArity of constr * constr * constr
- | NeedsInversion of constr * constr
| UnusedClause of cases_pattern list
| NonExhaustive of cases_pattern list
| CannotInferPredicate of (constr * types) array
-exception PatternMatchingError of env * pattern_matching_error
+exception PatternMatchingError of env * evar_map * pattern_matching_error
-val raise_pattern_matching_error : (loc * env * pattern_matching_error) -> 'a
+val error_wrong_numarg_constructor_loc : Loc.t -> env -> constructor -> int -> 'a
-val error_wrong_numarg_constructor_loc : loc -> env -> constructor -> int -> 'a
+val error_wrong_numarg_inductive_loc : Loc.t -> env -> inductive -> int -> 'a
-val error_wrong_numarg_inductive_loc : loc -> env -> inductive -> int -> 'a
+(** {6 Compilation primitive. } *)
-val error_bad_constructor_loc : loc -> constructor -> inductive -> 'a
+val compile_cases :
+ Loc.t -> case_style ->
+ (type_constraint -> env -> evar_map ref -> glob_constr -> unsafe_judgment) * evar_map ref ->
+ type_constraint ->
+ env -> glob_constr option * tomatch_tuples * cases_clauses ->
+ unsafe_judgment
-val error_bad_pattern_loc : loc -> constructor -> constr -> 'a
+val constr_of_pat :
+ Environ.env ->
+ Evd.evar_map ref ->
+ rel_declaration list ->
+ Glob_term.cases_pattern ->
+ Names.Id.t list ->
+ Glob_term.cases_pattern *
+ (rel_declaration list * Term.constr *
+ (Term.types * Term.constr list) * Glob_term.cases_pattern) *
+ Names.Id.t list
-val error_wrong_predicate_arity_loc : loc -> env -> constr -> constr -> constr -> 'a
+type 'a rhs =
+ { rhs_env : env;
+ rhs_vars : Id.t list;
+ avoid_ids : Id.t list;
+ it : 'a option}
-val error_needs_inversion : env -> constr -> types -> 'a
+type 'a equation =
+ { patterns : cases_pattern list;
+ rhs : 'a rhs;
+ alias_stack : Name.t list;
+ eqn_loc : Loc.t;
+ used : bool ref }
-(** {6 Compilation primitive. } *)
+type 'a matrix = 'a equation list
+
+(* 1st argument of IsInd is the original ind before extracting the summary *)
+type tomatch_type =
+ | IsInd of types * inductive_type * Name.t list
+ | NotInd of constr option * types
+
+(* spiwack: The first argument of [Pushed] is [true] for initial
+ Pushed and [false] otherwise. Used to decide whether the term being
+ matched on must be aliased in the variable case (only initial
+ Pushed need to be aliased). The first argument of [Alias] is [true]
+ if the alias was introduced by an initial pushed and [false]
+ otherwise.*)
+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 * rel_declaration
+
+type tomatch_stack = tomatch_status list
+
+(* We keep a constr for aliases and a cases_pattern for error message *)
+
+type pattern_history =
+ | Top
+ | MakeConstructor of constructor * pattern_continuation
+
+and pattern_continuation =
+ | Continuation of int * cases_pattern list * pattern_history
+ | Result of cases_pattern list
+
+type 'a pattern_matching_problem =
+ { env : env;
+ evdref : evar_map ref;
+ pred : constr;
+ tomatch : tomatch_stack;
+ history : pattern_continuation;
+ mat : 'a matrix;
+ caseloc : Loc.t;
+ casestyle : case_style;
+ typing_function: type_constraint -> env -> evar_map ref -> 'a option -> unsafe_judgment }
+
+
+val compile : 'a pattern_matching_problem -> Environ.unsafe_judgment
-module type S = sig
- val compile_cases :
- loc -> case_style ->
- (type_constraint -> env -> evar_map ref -> glob_constr -> unsafe_judgment) * evar_map ref ->
- type_constraint ->
- env -> glob_constr option * tomatch_tuples * cases_clauses ->
- unsafe_judgment
-end
+val prepare_predicate : Loc.t ->
+ (Evarutil.type_constraint ->
+ Environ.env -> Evd.evar_map ref -> 'a -> Environ.unsafe_judgment) ->
+ Environ.env ->
+ Evd.evar_map ->
+ (Term.types * tomatch_type) list ->
+ Context.rel_context list ->
+ Constr.constr option ->
+ 'a option -> (Evd.evar_map * Names.name list * Term.constr) list
-module Cases_F(C : Coercion.S) : S
diff --git a/pretyping/cbv.ml b/pretyping/cbv.ml
index 08ec25be..21bbede0 100644
--- a/pretyping/cbv.ml
+++ b/pretyping/cbv.ml
@@ -1,19 +1,15 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Util
-open Pp
-open Term
open Names
-open Environ
-open Univ
-open Evd
-open Conv_oracle
+open Term
+open Vars
open Closure
open Esubst
@@ -46,10 +42,10 @@ type cbv_value =
| VAL of int * constr
| STACK of int * cbv_value * cbv_stack
| CBN of constr * cbv_value subs
- | LAM of int * (name * constr) list * constr * cbv_value subs
+ | 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 * cbv_value array
+ | CONSTR of constructor 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
@@ -63,6 +59,8 @@ type cbv_value =
* the subs S, pat is information on the patterns of the Case
* (Weak reduction: we propagate the sub only when the selected branch
* is determined)
+ * PROJ(p,pb,stk) means the term is in a primitive projection p, itself in stk.
+ * pb is the associated projection body
*
* Important remark: the APPs should be collapsed:
* (APP (l,(APP ...))) forbidden
@@ -71,6 +69,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
(* les vars pourraient etre des constr,
cela permet de retarder les lift: utile ?? *)
@@ -90,7 +89,7 @@ let rec shift_value n = function
| CONSTR (c,args) ->
CONSTR (c, Array.map (shift_value n) args)
let shift_value n v =
- if n = 0 then v else shift_value n v
+ if Int.equal n 0 then v else shift_value n v
(* Contracts a fixpoint: given a fixpoint and a bindings,
* returns the corresponding fixpoint body, and the bindings in which
@@ -111,11 +110,11 @@ let contract_cofixp env (i,(_,_,bds as bodies)) =
let make_constr_ref n = function
| RelKey p -> mkRel (n+p)
| VarKey id -> mkVar id
- | ConstKey cst -> mkConst cst
+ | ConstKey cst -> mkConstU cst
(* Adds an application list. Collapse APPs! *)
let stack_app appl stack =
- if Array.length appl = 0 then stack else
+ if Int.equal (Array.length appl) 0 then stack else
match stack with
| APP(args,stk) -> APP(Array.append appl args,stk)
| _ -> APP(appl, stack)
@@ -125,6 +124,7 @@ let rec stack_concat stk1 stk2 =
TOP -> stk2
| APP(v,stk1') -> APP(v,stack_concat stk1' stk2)
| CASE(c,b,i,s,stk1') -> CASE(c,b,i,s,stack_concat stk1' stk2)
+ | PROJ (p,pinfo,stk1') -> PROJ (p,pinfo,stack_concat stk1' stk2)
(* merge stacks when there is no shifts in between *)
let mkSTACK = function
@@ -140,7 +140,7 @@ open RedFlags
let red_set_ref flags = function
| RelKey _ -> red_set flags fDELTA
| VarKey id -> red_set flags (fVAR id)
- | ConstKey sp -> red_set flags (fCONST sp)
+ | ConstKey (sp,_) -> red_set flags (fCONST sp)
(* Transfer application lists from a value to the stack
* useful because fixpoints may be totally applied in several times.
@@ -178,9 +178,9 @@ let cofixp_reducible flgs _ stk =
(* The main recursive functions
*
- * Go under applications and cases (pushed in the stack), expand head
- * constants or substitued de Bruijn, and try to make appear a
- * constructor, a lambda or a fixp in the head. If not, it is a value
+ * Go under applications and cases/projections (pushed in the stack),
+ * expand head constants or substitued de Bruijn, and try to a make a
+ * constructor, a lambda or a fixp appear in the head. If not, it is a value
* and is completely computed here. The head redexes are NOT reduced:
* the function returns the pair of a cbv_value and its stack. *
* Invariant: if the result of norm_head is CONSTR or (CO)FIXP, it last
@@ -197,7 +197,17 @@ let rec norm_head info env t stack =
norm_head info env head (stack_app nargs stack)
| Case (ci,p,c,v) -> norm_head info env c (CASE(p,v,ci,env,stack))
| Cast (ct,_,_) -> norm_head info env ct stack
-
+
+ | Proj (p, c) ->
+ let p' =
+ if red_set (info_flags info) (fCONST (Projection.constant p))
+ && red_set (info_flags info) fBETA
+ then Projection.unfold p
+ else p
+ in
+ let pinfo = Environ.lookup_projection p (info_env info) in
+ norm_head info env c (PROJ (p', pinfo, stack))
+
(* constants, axioms
* the first pattern is CRUCIAL, n=0 happens very often:
* when reducing closed terms, n is always 0 *)
@@ -222,10 +232,10 @@ let rec norm_head info env t stack =
let env' = subs_cons ([|cbv_stack_term info TOP env b|],env) in
norm_head info env' c stack
else
- (CBN(t,env), stack) (* Considérer une coupure commutative ? *)
+ (CBN(t,env), stack) (* Should we consider a commutative cut ? *)
| Evar ev ->
- (match evar_value info ev with
+ (match evar_value info.i_cache ev with
Some c -> norm_head info env c stack
| None -> (VAL(0, t), stack))
@@ -255,19 +265,21 @@ and norm_head_ref k info env stack normt =
* we build a value.
*)
and cbv_stack_term info stack env t =
- match norm_head info env t stack with
- (* a lambda meets an application -> BETA *)
- | (LAM (nlams,ctxt,b,env), APP (args, stk))
+ cbv_stack_value info env (norm_head info env t stack)
+
+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 ->
- let nargs = Array.length args in
- if nargs == nlams then
+ let nargs = Array.length args in
+ if nargs == nlams then
cbv_stack_term info stk (subs_cons(args,env)) b
else if nlams < nargs then
let env' = subs_cons(Array.sub args 0 nlams, env) in
let eargs = Array.sub args nlams (nargs-nlams) in
cbv_stack_term info (APP(eargs,stk)) env' b
else
- let ctxt' = list_skipn nargs ctxt in
+ let ctxt' = List.skipn nargs ctxt in
LAM(nlams-nargs,ctxt', b, subs_cons(args,env))
(* a Fix applied enough -> IOTA *)
@@ -283,22 +295,28 @@ and cbv_stack_term info stack env t =
cbv_stack_term info stk envf redfix
(* constructor in a Case -> IOTA *)
- | (CONSTR((sp,n),[||]), APP(args,CASE(_,br,ci,env,stk)))
+ | (CONSTR(((sp,n),u),[||]), APP(args,CASE(_,br,ci,env,stk)))
when red_set (info_flags info) fIOTA ->
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),[||]), CASE(_,br,_,env,stk))
+ | (CONSTR(((_,n),u),[||]), CASE(_,br,_,env,stk))
when red_set (info_flags info) fIOTA ->
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) fIOTA && Projection.unfolded p ->
+ let arg = args.(pi.Declarations.proj_npars + pi.Declarations.proj_arg) in
+ cbv_stack_value info env (arg, stk)
+
(* may be reduced later by application *)
| (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)
@@ -316,6 +334,8 @@ let rec apply_stack info t = function
(mkCase (ci, cbv_norm_term info env ty, t,
Array.map (cbv_norm_term info env) br))
st
+ | PROJ (p, pinfo, st) ->
+ apply_stack info (mkProj (p, t)) st
(* performs the reduction on a constr, and returns a constr *)
and cbv_norm_term info env t =
@@ -333,7 +353,7 @@ and cbv_norm_value info = function (* reduction under binders *)
map_constr_with_binders subs_lift (cbv_norm_term info) env t
| LAM (n,ctxt,b,env) ->
let nctxt =
- list_map_i (fun i (x,ty) ->
+ 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)
| FIXP ((lij,(names,lty,bds)),env,args) ->
@@ -352,7 +372,7 @@ and cbv_norm_value info = function (* reduction under binders *)
(subs_liftn (Array.length lty) env)) bds)),
Array.map (cbv_norm_value info) args)
| CONSTR (c,args) ->
- mkApp(mkConstruct c, Array.map (cbv_norm_value info) args)
+ mkApp(mkConstructU c, Array.map (cbv_norm_value info) args)
(* with profiling *)
let cbv_norm infos constr =
diff --git a/pretyping/cbv.mli b/pretyping/cbv.mli
index eb0abe97..bde85383 100644
--- a/pretyping/cbv.mli
+++ b/pretyping/cbv.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -27,15 +27,16 @@ type cbv_value =
| VAL of int * constr
| STACK of int * cbv_value * cbv_stack
| CBN of constr * cbv_value subs
- | LAM of int * (name * constr) list * constr * cbv_value subs
+ | 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 * cbv_value array
+ | CONSTR of constructor 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
val shift_value : int -> cbv_value -> cbv_value
diff --git a/pretyping/classops.ml b/pretyping/classops.ml
index 59f3c740..559f5fe6 100644
--- a/pretyping/classops.ml
+++ b/pretyping/classops.ml
@@ -1,24 +1,23 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Errors
open Util
open Pp
open Flags
open Names
open Libnames
+open Globnames
open Nametab
open Environ
open Libobject
-open Library
open Term
open Termops
-open Glob_term
-open Decl_kinds
open Mod_subst
(* usage qque peu general: utilise aussi dans record *)
@@ -32,6 +31,7 @@ type cl_typ =
| CL_SECVAR of variable
| CL_CONST of constant
| CL_IND of inductive
+ | CL_PROJ of constant
type cl_info_typ = {
cl_param : int
@@ -39,21 +39,42 @@ type cl_info_typ = {
type coe_typ = global_reference
+module CoeTypMap = Refmap_env
+
type coe_info_typ = {
coe_value : constr;
coe_type : types;
- coe_strength : locality;
+ coe_local : bool;
+ coe_context : Univ.universe_context_set;
coe_is_identity : bool;
+ coe_is_projection : bool;
coe_param : int }
let coe_info_typ_equal c1 c2 =
eq_constr c1.coe_value c2.coe_value &&
eq_constr c1.coe_type c2.coe_type &&
- c1.coe_strength = c2.coe_strength &&
- c1.coe_is_identity = c2.coe_is_identity &&
- c1.coe_param = c2.coe_param
+ c1.coe_local == c2.coe_local &&
+ c1.coe_is_identity == c2.coe_is_identity &&
+ c1.coe_is_projection == c2.coe_is_projection &&
+ Int.equal c1.coe_param c2.coe_param
+
+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_user_ord c1 c2
+ | CL_PROJ c1, CL_PROJ c2 -> con_user_ord c1 c2
+ | CL_IND i1, CL_IND i2 -> ind_user_ord i1 i2
+ | _ -> Pervasives.compare t1 t2 (** OK *)
+
+module ClTyp = struct
+ type t = cl_typ
+ let compare = cl_typ_ord
+end
+
+module ClTypMap = Map.Make(ClTyp)
-type cl_index = int
+module IntMap = Map.Make(Int)
+
+let cl_typ_eq t1 t2 = Int.equal (cl_typ_ord t1 t2) 0
type coe_index = coe_info_typ
@@ -61,31 +82,60 @@ type inheritance_path = coe_index list
(* table des classes, des coercions et graphe d'heritage *)
-module Bijint = struct
- type ('a,'b) t = { v : ('a * 'b) array; s : int; inv : ('a,int) Gmap.t }
- let empty = { v = [||]; s = 0; inv = Gmap.empty }
- let mem y b = Gmap.mem y b.inv
- let map x b = if 0 <= x & x < b.s then b.v.(x) else raise Not_found
- let revmap y b = let n = Gmap.find y b.inv in (n, snd (b.v.(n)))
+module Bijint :
+sig
+ module Index :
+ sig
+ type t
+ val compare : t -> t -> int
+ val equal : t -> t -> bool
+ val print : t -> std_ppcmds
+ end
+ type 'a t
+ val empty : 'a t
+ val mem : cl_typ -> 'a t -> bool
+ val map : Index.t -> 'a t -> cl_typ * 'a
+ val revmap : cl_typ -> 'a t -> Index.t * 'a
+ val add : cl_typ -> 'a -> 'a t -> 'a t
+ val dom : 'a t -> cl_typ list
+end
+=
+struct
+
+ module Index = struct include Int let print = Pp.int end
+
+ type 'a t = { v : (cl_typ * 'a) IntMap.t; s : int; inv : int ClTypMap.t }
+ let empty = { v = IntMap.empty; s = 0; inv = ClTypMap.empty }
+ let mem y b = ClTypMap.mem y b.inv
+ let map x b = IntMap.find x b.v
+ let revmap y b = let n = ClTypMap.find y b.inv in (n, snd (IntMap.find n b.v))
let add x y b =
- let v =
- if b.s = Array.length b.v then
- (let v = Array.make (b.s + 8) (x,y) in Array.blit b.v 0 v 0 b.s; v)
- else b.v in
- v.(b.s) <- (x,y); { v = v; s = b.s+1; inv = Gmap.add x b.s b.inv }
- let dom b = Gmap.dom b.inv
+ { v = IntMap.add b.s (x,y) b.v; s = b.s+1; inv = ClTypMap.add x b.s b.inv }
+ let dom b = List.rev (ClTypMap.fold (fun x _ acc -> x::acc) b.inv [])
end
+type cl_index = Bijint.Index.t
+
let class_tab =
- ref (Bijint.empty : (cl_typ, cl_info_typ) Bijint.t)
+ ref (Bijint.empty : cl_info_typ Bijint.t)
let coercion_tab =
- ref (Gmap.empty : (coe_typ, coe_info_typ) Gmap.t)
+ ref (CoeTypMap.empty : coe_info_typ CoeTypMap.t)
+
+module ClPairOrd =
+struct
+ type t = cl_index * cl_index
+ let compare (i1, j1) (i2, j2) =
+ let c = Bijint.Index.compare i1 i2 in
+ if Int.equal c 0 then Bijint.Index.compare j1 j2 else c
+end
+
+module ClPairMap = Map.Make(ClPairOrd)
let inheritance_graph =
- ref (Gmap.empty : (cl_index * cl_index, inheritance_path) Gmap.t)
+ ref (ClPairMap.empty : inheritance_path ClPairMap.t)
-let freeze () = (!class_tab, !coercion_tab, !inheritance_graph)
+let freeze _ = (!class_tab, !coercion_tab, !inheritance_graph)
let unfreeze (fcl,fco,fig) =
class_tab:=fcl;
@@ -99,17 +149,17 @@ let add_new_class cl s =
class_tab := Bijint.add cl s !class_tab
let add_new_coercion coe s =
- coercion_tab := Gmap.add coe s !coercion_tab
+ coercion_tab := CoeTypMap.add coe s !coercion_tab
let add_new_path x y =
- inheritance_graph := Gmap.add x y !inheritance_graph
+ inheritance_graph := ClPairMap.add x y !inheritance_graph
let init () =
class_tab:= Bijint.empty;
add_new_class CL_FUN { cl_param = 0 };
add_new_class CL_SORT { cl_param = 0 };
- coercion_tab:= Gmap.empty;
- inheritance_graph:= Gmap.empty
+ coercion_tab:= CoeTypMap.empty;
+ inheritance_graph:= ClPairMap.empty
let _ =
Summary.declare_summary "inh_graph"
@@ -135,20 +185,22 @@ let cl_sort_index = fst(class_info CL_SORT)
(* coercion_info : coe_typ -> coe_info_typ *)
-let coercion_info coe = Gmap.find coe !coercion_tab
+let coercion_info coe = CoeTypMap.find coe !coercion_tab
-let coercion_exists coe = Gmap.mem coe !coercion_tab
+let coercion_exists coe = CoeTypMap.mem coe !coercion_tab
-(* find_class_type : evar_map -> constr -> cl_typ * constr list *)
+(* find_class_type : evar_map -> constr -> cl_typ * universe_list * constr list *)
let find_class_type sigma t =
let t', args = Reductionops.whd_betaiotazeta_stack sigma t in
match kind_of_term t' with
- | Var id -> CL_SECVAR id, args
- | Const sp -> CL_CONST sp, args
- | Ind ind_sp -> CL_IND ind_sp, args
- | Prod (_,_,_) -> CL_FUN, []
- | Sort _ -> CL_SORT, []
+ | Var id -> CL_SECVAR id, Univ.Instance.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
+ | Ind (ind_sp,u) -> CL_IND ind_sp, u, args
+ | Prod (_,_,_) -> CL_FUN, Univ.Instance.empty, []
+ | Sort _ -> CL_SORT, Univ.Instance.empty, []
| _ -> raise Not_found
@@ -156,55 +208,57 @@ let subst_cl_typ subst ct = match ct with
CL_SORT
| CL_FUN
| CL_SECVAR _ -> ct
- | CL_CONST kn ->
- let kn',t = subst_con subst kn in
- if kn' == kn then ct else
- fst (find_class_type Evd.empty t)
- | CL_IND (kn,i) ->
- let kn' = subst_ind subst kn in
- if kn' == kn then ct else
- CL_IND (kn',i)
+ | CL_PROJ c ->
+ let c',t = subst_con_kn subst c in
+ if c' == c then ct else CL_PROJ c'
+ | 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)
+ | CL_IND i ->
+ let i' = subst_ind subst i in
+ if i' == i then ct else CL_IND i'
(*CSC: here we should change the datatype for coercions: it should be possible
to declare any term as a coercion *)
-let subst_coe_typ subst t = fst (subst_global subst t)
+let subst_coe_typ subst t = subst_global_reference subst t
(* class_of : Term.constr -> int *)
let class_of env sigma t =
- let (t, n1, i, args) =
+ let (t, n1, i, u, args) =
try
- let (cl,args) = find_class_type sigma t in
+ let (cl, u, args) = find_class_type sigma t in
let (i, { cl_param = n1 } ) = class_info cl in
- (t, n1, i, args)
+ (t, n1, i, u, args)
with Not_found ->
let t = Tacred.hnf_constr env sigma t in
- let (cl, args) = find_class_type sigma t in
+ let (cl, u, args) = find_class_type sigma t in
let (i, { cl_param = n1 } ) = class_info cl in
- (t, n1, i, args)
+ (t, n1, i, u, args)
in
- if List.length args = n1 then t, i else raise Not_found
+ if Int.equal (List.length args) n1 then t, i else raise Not_found
let inductive_class_of ind = fst (class_info (CL_IND ind))
-let class_args_of env sigma c = snd (find_class_type sigma c)
+let class_args_of env sigma c = pi3 (find_class_type sigma c)
let string_of_class = function
| CL_FUN -> "Funclass"
| CL_SORT -> "Sortclass"
- | CL_CONST sp ->
- string_of_qualid (shortest_qualid_of_global Idset.empty (ConstRef sp))
+ | CL_CONST sp | CL_PROJ sp ->
+ string_of_qualid (shortest_qualid_of_global Id.Set.empty (ConstRef sp))
| CL_IND sp ->
- string_of_qualid (shortest_qualid_of_global Idset.empty (IndRef sp))
+ string_of_qualid (shortest_qualid_of_global Id.Set.empty (IndRef sp))
| CL_SECVAR sp ->
- string_of_qualid (shortest_qualid_of_global Idset.empty (VarRef sp))
+ string_of_qualid (shortest_qualid_of_global Id.Set.empty (VarRef sp))
let pr_class x = str (string_of_class x)
(* lookup paths *)
let lookup_path_between_class (s,t) =
- Gmap.find (s,t) !inheritance_graph
+ ClPairMap.find (s,t) !inheritance_graph
let lookup_path_to_fun_from_class s =
lookup_path_between_class (s,cl_fun_index)
@@ -216,16 +270,16 @@ let lookup_path_to_sort_from_class s =
let apply_on_class_of env sigma t cont =
try
- let (cl,args) = find_class_type sigma t in
+ let (cl,u,args) = find_class_type sigma t in
let (i, { cl_param = n1 } ) = class_info cl in
- if List.length args <> n1 then raise Not_found;
+ if not (Int.equal (List.length args) n1) then raise Not_found;
t, cont i
with Not_found ->
(* Is it worth to be more incremental on the delta steps? *)
let t = Tacred.hnf_constr env sigma t in
- let (cl, args) = find_class_type sigma t in
+ let (cl, u, args) = find_class_type sigma t in
let (i, { cl_param = n1 } ) = class_info cl in
- if List.length args <> n1 then raise Not_found;
+ if not (Int.equal (List.length args) n1) then raise Not_found;
t, cont i
let lookup_path_between env sigma (s,t) =
@@ -241,31 +295,35 @@ let lookup_path_to_fun_from env sigma s =
let lookup_path_to_sort_from env sigma s =
apply_on_class_of env sigma s lookup_path_to_sort_from_class
-let get_coercion_constructor coe =
+let get_coercion_constructor env coe =
let c, _ =
- Reductionops.whd_betadeltaiota_stack (Global.env()) Evd.empty coe.coe_value
+ Reductionops.whd_betadeltaiota_stack env Evd.empty coe.coe_value
in
match kind_of_term c with
- | Construct cstr ->
- (cstr, Inductiveops.constructor_nrealargs (Global.env()) cstr -1)
+ | Construct (cstr,u) ->
+ (cstr, Inductiveops.constructor_nrealargs cstr -1)
| _ ->
raise Not_found
-let lookup_pattern_path_between (s,t) =
+let lookup_pattern_path_between env (s,t) =
let i = inductive_class_of s in
let j = inductive_class_of t in
- List.map get_coercion_constructor (Gmap.find (i,j) !inheritance_graph)
+ List.map (get_coercion_constructor env) (ClPairMap.find (i,j) !inheritance_graph)
(* coercion_value : coe_index -> unsafe_judgment * bool *)
-let coercion_value { coe_value = c; coe_type = t; coe_is_identity = b } =
- (make_judge c t, b)
+let coercion_value { coe_value = c; coe_type = t; coe_context = ctx;
+ coe_is_identity = b; coe_is_projection = b' } =
+ 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
(* pretty-print functions are now in Pretty *)
(* rajouter une coercion dans le graphe *)
let path_printer = ref (fun _ -> str "<a class path>"
- : (int * int) * inheritance_path -> std_ppcmds)
+ : (Bijint.Index.t * Bijint.Index.t) * inheritance_path -> std_ppcmds)
let install_path_printer f = path_printer := f
@@ -273,27 +331,33 @@ let print_path x = !path_printer x
let message_ambig l =
(str"Ambiguous paths:" ++ spc () ++
- prlist_with_sep pr_fnl (fun ijp -> print_path ijp) l)
+ prlist_with_sep fnl (fun ijp -> print_path ijp) l)
(* add_coercion_in_graph : coe_index * cl_index * cl_index -> unit
coercion,source,target *)
-let different_class_params i j =
- (snd (class_info_from_index i)).cl_param > 0
-
+let different_class_params i =
+ let ci = class_info_from_index i in
+ if (snd ci).cl_param > 0 then true
+ else
+ match fst ci with
+ | 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 old_inheritance_graph = !inheritance_graph in
let ambig_paths =
(ref [] : ((cl_index * cl_index) * inheritance_path) list ref) in
let try_add_new_path (i,j as ij) p =
try
- if i=j then begin
- if different_class_params i j then begin
+ if Bijint.Index.equal i j then begin
+ if different_class_params i then begin
let _ = lookup_path_between_class ij in
ambig_paths := (ij,p)::!ambig_paths
end
end else begin
- let _ = lookup_path_between_class (i,j) in
+ let _ = lookup_path_between_class ij in
ambig_paths := (ij,p)::!ambig_paths
end;
false
@@ -306,35 +370,50 @@ let add_coercion_in_graph (ic,source,target) =
let _ = try_add_new_path ij p in ()
in
if try_add_new_path (source,target) [ic] then begin
- Gmap.iter
+ ClPairMap.iter
(fun (s,t) p ->
- if s<>t then begin
- if t = source then begin
+ if not (Bijint.Index.equal s t) then begin
+ if Bijint.Index.equal t source then begin
try_add_new_path1 (s,target) (p@[ic]);
- Gmap.iter
+ ClPairMap.iter
(fun (u,v) q ->
- if u<>v & u = target && not (list_equal coe_info_typ_equal p q) then
+ if not (Bijint.Index.equal u v) && Bijint.Index.equal u target && not (List.equal coe_info_typ_equal p q) then
try_add_new_path1 (s,v) (p@[ic]@q))
old_inheritance_graph
end;
- if s = target then try_add_new_path1 (source,t) (ic::p)
+ if Bijint.Index.equal s target then try_add_new_path1 (source,t) (ic::p)
end)
old_inheritance_graph
end;
- if (!ambig_paths <> []) && is_verbose () then
- ppnl (message_ambig !ambig_paths)
-
-type coercion = coe_typ * locality * bool * cl_typ * cl_typ * int
+ let is_ambig = match !ambig_paths with [] -> false | _ -> true in
+ if is_ambig && is_verbose () then
+ msg_warning (message_ambig !ambig_paths)
+
+type coercion = {
+ coercion_type : coe_typ;
+ coercion_local : bool;
+ coercion_is_id : bool;
+ coercion_is_proj : bool;
+ coercion_source : cl_typ;
+ coercion_target : cl_typ;
+ coercion_params : int;
+}
-(* Calcul de l'arité d'une classe *)
+(* Computation of the class arity *)
let reference_arity_length ref =
- let t = Global.type_of_global ref in
+ let t = Universes.unsafe_type_of_global ref in
List.length (fst (Reductionops.splay_arity (Global.env()) Evd.empty t))
+let projection_arity_length p =
+ let len = reference_arity_length (ConstRef p) in
+ let pb = Environ.lookup_projection (Projection.make p false) (Global.env ()) in
+ len - pb.Declarations.proj_npars
+
let class_params = function
| CL_FUN | CL_SORT -> 0
| CL_CONST sp -> reference_arity_length (ConstRef sp)
+ | CL_PROJ sp -> projection_arity_length sp
| CL_SECVAR sp -> reference_arity_length (VarRef sp)
| CL_IND sp -> reference_arity_length (IndRef sp)
@@ -355,18 +434,22 @@ let _ =
optread = (fun () -> !automatically_import_coercions);
optwrite = (:=) automatically_import_coercions }
-let cache_coercion (_,(coe,stre,isid,cls,clt,ps)) =
- add_class cls;
- add_class clt;
- let is,_ = class_info cls in
- let it,_ = class_info clt in
+let cache_coercion (_, 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 xf =
- { coe_value = constr_of_global coe;
- coe_type = Global.type_of_global coe;
- coe_strength = stre;
- coe_is_identity = isid;
- coe_param = ps } in
- add_new_coercion coe xf;
+ { coe_value = value;
+ coe_type = typ;
+ coe_context = ctx;
+ coe_local = c.coercion_local;
+ coe_is_identity = c.coercion_is_id;
+ 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)
let load_coercion _ o =
@@ -376,40 +459,45 @@ let load_coercion _ o =
cache_coercion o
let open_coercion i o =
- if i = 1 && not
+ if Int.equal i 1 && not
(!automatically_import_coercions || Flags.version_less_or_equal Flags.V8_2)
then
cache_coercion o
-let subst_coercion (subst,(coe,stre,isid,cls,clt,ps as obj)) =
- let coe' = subst_coe_typ subst coe in
- let cls' = subst_cl_typ subst cls in
- let clt' = subst_cl_typ subst clt in
- if coe' == coe && cls' == cls & clt' == clt then obj else
- (coe',stre,isid,cls',clt',ps)
+let subst_coercion (subst, c) =
+ let coe = subst_coe_typ subst c.coercion_type in
+ let cls = subst_cl_typ subst c.coercion_source in
+ let clt = subst_cl_typ subst c.coercion_target in
+ if c.coercion_type == coe && c.coercion_source == cls && c.coercion_target == clt then c
+ else { c with coercion_type = coe; coercion_source = cls; coercion_target = clt }
let discharge_cl = function
| CL_CONST kn -> CL_CONST (Lib.discharge_con kn)
| CL_IND ind -> CL_IND (Lib.discharge_inductive ind)
+ | CL_PROJ p -> CL_PROJ (Lib.discharge_con p)
| cl -> cl
-let discharge_coercion (_,(coe,stre,isid,cls,clt,ps)) =
- if stre = Local then None else
- let n = try Array.length (Lib.section_instance coe) with Not_found -> 0 in
- Some (Lib.discharge_global coe,
- stre,
- isid,
- discharge_cl cls,
- discharge_cl clt,
- n + ps)
-
-let classify_coercion (coe,stre,isid,cls,clt,ps as obj) =
- if stre = Local then Dispose else Substitute obj
-
-type coercion_obj =
- coe_typ * Decl_kinds.locality * bool * cl_typ * cl_typ * int
-
-let inCoercion : coercion_obj -> obj =
+let discharge_coercion (_, c) =
+ if c.coercion_local then None
+ else
+ let n =
+ try
+ let ins = Lib.section_instance c.coercion_type in
+ Array.length (snd ins)
+ with Not_found -> 0
+ in
+ let nc = { c with
+ coercion_type = Lib.discharge_global c.coercion_type;
+ coercion_source = discharge_cl c.coercion_source;
+ coercion_target = discharge_cl c.coercion_target;
+ coercion_params = n + c.coercion_params;
+ } in
+ Some nc
+
+let classify_coercion obj =
+ if obj.coercion_local then Dispose else Substitute obj
+
+let inCoercion : coercion -> obj =
declare_object {(default_object "COERCION") with
open_function = open_coercion;
load_function = load_coercion;
@@ -418,31 +506,49 @@ let inCoercion : coercion_obj -> obj =
classify_function = classify_coercion;
discharge_function = discharge_coercion }
-let declare_coercion coef stre ~isid ~src:cls ~target:clt ~params:ps =
- Lib.add_anonymous_leaf (inCoercion (coef,stre,isid,cls,clt,ps))
+let declare_coercion coef ?(local = false) ~isid ~src:cls ~target:clt ~params:ps =
+ let isproj =
+ match coef with
+ | ConstRef c -> Environ.is_projection c (Global.env ())
+ | _ -> false
+ in
+ let c = {
+ coercion_type = coef;
+ coercion_local = local;
+ coercion_is_id = isid;
+ coercion_is_proj = isproj;
+ coercion_source = cls;
+ coercion_target = clt;
+ coercion_params = ps;
+ } in
+ Lib.add_anonymous_leaf (inCoercion c)
(* For printing purpose *)
let get_coercion_value v = v.coe_value
-let pr_cl_index n = int n
+let pr_cl_index = Bijint.Index.print
let classes () = Bijint.dom !class_tab
-let coercions () = Gmap.rng !coercion_tab
-let inheritance_graph () = Gmap.to_list !inheritance_graph
+let coercions () =
+ List.rev (CoeTypMap.fold (fun _ y acc -> y::acc) !coercion_tab [])
+
+let inheritance_graph () =
+ ClPairMap.bindings !inheritance_graph
let coercion_of_reference r =
let ref = Nametab.global r in
if not (coercion_exists ref) then
errorlabstrm "try_add_coercion"
- (Nametab.pr_global_env Idset.empty ref ++ str" is not a coercion.");
+ (Nametab.pr_global_env Id.Set.empty ref ++ str" is not a coercion.");
ref
module CoercionPrinting =
struct
type t = coe_typ
+ let compare = RefOrdered.compare
let encode = coercion_of_reference
let subst = subst_coe_typ
- let printer x = pr_global_env Idset.empty x
+ let printer x = pr_global_env Id.Set.empty x
let key = ["Printing";"Coercion"]
let title = "Explicitly printed coercions: "
let member_message x b =
diff --git a/pretyping/classops.mli b/pretyping/classops.mli
index 0136b90c..c421b450 100644
--- a/pretyping/classops.mli
+++ b/pretyping/classops.mli
@@ -1,17 +1,15 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Names
-open Decl_kinds
open Term
open Evd
open Environ
-open Nametab
open Mod_subst
(** {6 This is the type of class kinds } *)
@@ -21,6 +19,10 @@ type cl_typ =
| CL_SECVAR of variable
| CL_CONST of constant
| CL_IND of inductive
+ | CL_PROJ of constant
+
+(** Equality over [cl_typ] *)
+val cl_typ_eq : cl_typ -> cl_typ -> bool
val subst_cl_typ : substitution -> cl_typ -> cl_typ
@@ -29,7 +31,7 @@ type cl_info_typ = {
cl_param : int }
(** This is the type of coercion kinds *)
-type coe_typ = Libnames.global_reference
+type coe_typ = Globnames.global_reference
(** This is the type of infos for declared coercions *)
type coe_info_typ
@@ -44,13 +46,17 @@ type coe_index
type inheritance_path = coe_index list
(** {6 Access to classes infos } *)
-val class_info : cl_typ -> (cl_index * cl_info_typ)
+
val class_exists : cl_typ -> bool
+
+val class_info : cl_typ -> (cl_index * cl_info_typ)
+(** @raise Not_found if this type is not a class *)
+
val class_info_from_index : cl_index -> cl_typ * cl_info_typ
-(** [find_class_type env sigma c] returns the head reference of [c] and its
- arguments *)
-val find_class_type : evar_map -> types -> cl_typ * constr list
+(** [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
(** raises [Not_found] if not convertible to a class *)
val class_of : env -> evar_map -> types -> types * cl_index
@@ -62,16 +68,18 @@ val class_args_of : env -> evar_map -> types -> constr list
(** {6 [declare_coercion] adds a coercion in the graph of coercion paths } *)
val declare_coercion :
- coe_typ -> locality -> isid:bool ->
+ coe_typ -> ?local:bool -> isid:bool ->
src:cl_typ -> target:cl_typ -> params:int -> unit
(** {6 Access to coercions infos } *)
val coercion_exists : coe_typ -> bool
-val coercion_value : coe_index -> (unsafe_judgment * bool)
+val coercion_value : coe_index -> (unsafe_judgment * bool * bool) Univ.in_universe_context_set
(** {6 Lookup functions for coercion paths } *)
+
val lookup_path_between_class : cl_index * cl_index -> inheritance_path
+(** @raise Not_found when no such path exists *)
val lookup_path_between : env -> evar_map -> types * types ->
types * types * inheritance_path
@@ -80,7 +88,7 @@ val lookup_path_to_fun_from : env -> evar_map -> types ->
val lookup_path_to_sort_from : env -> evar_map -> types ->
types * inheritance_path
val lookup_pattern_path_between :
- inductive * inductive -> (constructor * int) list
+ env -> inductive * inductive -> (constructor * int) list
(**/**)
(* Crade *)
diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml
index 917856a2..8ebb8cd2 100644
--- a/pretyping/coercion.ml
+++ b/pretyping/coercion.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -9,279 +9,513 @@
(* Created by Hugo Herbelin for Coq V7 by isolating the coercion
mechanism out of the type inference algorithm in file trad.ml from
Coq V6.3, Nov 1999; The coercion mechanism was implemented in
- trad.ml by Amokrane Saïbi, May 1996 *)
+ trad.ml by Amokrane Saïbi, May 1996 *)
(* Addition of products and sorts in canonical structures by Pierre
Corbineau, Feb 2008 *)
(* Turned into an abstract compilation unit by Matthieu Sozeau, March 2006 *)
+open Errors
open Util
open Names
open Term
+open Vars
open Reductionops
open Environ
open Typeops
open Pretype_errors
open Classops
-open Recordops
open Evarutil
open Evarconv
-open Retyping
open Evd
open Termops
+open Globnames
-module type S = sig
- (*s Coercions. *)
-
- (* [inh_app_fun env evd j] coerces [j] to a function; i.e. it
- inserts a coercion into [j], if needed, in such a way it gets as
- type a product; it returns [j] if no coercion is applicable *)
- val inh_app_fun :
- bool -> env -> evar_map -> unsafe_judgment -> evar_map * unsafe_judgment
-
- (* [inh_coerce_to_sort env evd 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 ->
- env -> evar_map -> unsafe_judgment -> evar_map * unsafe_type_judgment
-
- (* [inh_coerce_to_base env evd 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 ->
- env -> evar_map -> unsafe_judgment -> evar_map * unsafe_judgment
-
- (* [inh_coerce_to_prod env evars t] coerces [t] to a product type *)
- val inh_coerce_to_prod : loc ->
- env -> evar_map -> type_constraint_type -> evar_map * type_constraint_type
-
- (* [inh_conv_coerce_to loc env evd j t] coerces [j] to an object of type
- [t]; i.e. it inserts a coercion into [j], if needed, in such a way [t] and
- [j.uj_type] are convertible; it fails if no coercion is applicable *)
- val inh_conv_coerce_to : bool -> loc ->
- env -> evar_map -> unsafe_judgment -> type_constraint_type -> evar_map * unsafe_judgment
-
- val inh_conv_coerce_rigid_to : bool -> loc ->
- env -> evar_map -> unsafe_judgment -> type_constraint_type -> evar_map * unsafe_judgment
-
- (* [inh_conv_coerces_to loc env evd 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 ->
- env -> evar_map -> types -> type_constraint_type -> evar_map
-
- (* [inh_pattern_coerce_to loc env evd 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 -> Glob_term.cases_pattern -> inductive -> inductive -> Glob_term.cases_pattern
-end
-
-module Default = struct
- (* Typing operations dealing with coercions *)
- exception NoCoercion
-
- (* Here, funj is a coercion therefore already typed in global context *)
- let apply_coercion_args env argl funj =
- let rec apply_rec acc typ = function
- | [] -> { uj_val = applist (j_val funj,argl);
- uj_type = typ }
- | h::restl ->
- (* On devrait pouvoir s'arranger pour qu'on n'ait pas à faire hnf_constr *)
- match kind_of_term (whd_betadeltaiota env Evd.empty typ) with
- | Prod (_,c1,c2) ->
- (* Typage garanti par l'appel à app_coercion*)
- apply_rec (h::acc) (subst1 h c2) restl
- | _ -> anomaly "apply_coercion_args"
+let use_typeclasses_for_conversion = ref true
+
+let _ =
+ Goptions.declare_bool_option
+ { Goptions.optsync = true;
+ 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
+exception NoCoercionNoUnifier of evar_map * unification_error
+
+(* Here, funj is a coercion therefore already typed in global context *)
+let apply_coercion_args env evd check isproj argl funj =
+ let evdref = ref evd in
+ let rec apply_rec acc typ = function
+ | [] ->
+ if isproj then
+ let cst = fst (destConst (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
+ let hd, tl = match args with hd :: tl -> hd, tl | [] -> assert false in
+ { uj_val = applist (mkProj (p, hd), tl);
+ uj_type = typ }
+ else
+ { 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_betadeltaiota env evd typ) with
+ | Prod (_,c1,c2) ->
+ if check && not (e_cumul env evdref (Retyping.get_type_of env evd h) c1) then
+ raise NoCoercion;
+ apply_rec (h::acc) (subst1 h c2) restl
+ | _ -> 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 =
+ 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))
+ pat p
+
+(* raise Not_found if no coercion found *)
+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
+
+(* Program coercions *)
+
+open Program
+
+let make_existential loc ?(opaque = Evar_kinds.Define true) env evdref c =
+ Evarutil.e_new_evar env evdref ~src:(loc, Evar_kinds.QuestionMark opaque) c
+
+let app_opt env evdref f t =
+ whd_betaiota !evdref (app_opt f t)
+
+let pair_of_array a = (a.(0), a.(1))
+
+let disc_subset x =
+ match kind_of_term x with
+ | App (c, l) ->
+ (match kind_of_term c with
+ Ind (i,_) ->
+ let len = Array.length l in
+ let sigty = delayed_force sig_typ in
+ if Int.equal len 2 && eq_ind i (Globnames.destIndRef sigty)
+ then
+ let (a, b) = pair_of_array l in
+ Some (a, b)
+ else None
+ | _ -> None)
+ | _ -> None
+
+exception NoSubtacCoercion
+
+let hnf env evd c = whd_betadeltaiota env evd c
+let hnf_nodelta env evd c = whd_betaiota evd c
+
+let lift_args n sign =
+ let rec liftrec k = function
+ | t::sign -> liftn n k t :: (liftrec (k-1) sign)
+ | [] -> []
+ in
+ liftrec (List.length sign) sign
+
+let mu env evdref t =
+ let rec aux v =
+ let v' = hnf env !evdref v in
+ match disc_subset v' with
+ Some (u, p) ->
+ let f, ct = aux u in
+ let p = hnf_nodelta env !evdref p in
+ (Some (fun x ->
+ app_opt env evdref
+ f (papp evdref sig_proj1 [| u; p; x |])),
+ ct)
+ | None -> (None, v)
+ in aux t
+
+and coerce loc env evdref (x : Term.constr) (y : Term.constr)
+ : (Term.constr -> Term.constr) option
+ =
+ let rec coerce_unify env x y =
+ let x = hnf env !evdref x and y = hnf env !evdref y in
+ try
+ 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 =
+ let subco () = subset_coerce env evdref x y in
+ let dest_prod c =
+ match Reductionops.splay_prod_n env ( !evdref) 1 c with
+ | [(na,b,t)], c -> (na,t), c
+ | _ -> raise NoSubtacCoercion
in
- apply_rec [] funj.uj_type argl
-
- (* appliquer le chemin de coercions de patterns 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_tabulate f (n+1), Anonymous))
- pat p
-
- (* raise Not_found if no coercion found *)
- let inh_pattern_coerce_to loc pat ind1 ind2 =
- let p = lookup_pattern_path_between (ind1,ind2) in
- apply_pattern_coercion loc pat p
-
- let saturate_evd env evd =
- Typeclasses.resolve_typeclasses
- ~filter:Typeclasses.no_goals ~split:true ~fail:false env evd
-
- (* appliquer le chemin de coercions p à hj *)
- let apply_coercion env sigma p hj typ_cl =
- try
- fst (List.fold_left
- (fun (ja,typ_cl) i ->
- let fv,isid = coercion_value i in
- let argl = (class_args_of env sigma typ_cl)@[ja.uj_val] in
- let jres = apply_coercion_args env argl fv in
- (if isid then
- { uj_val = ja.uj_val; uj_type = jres.uj_type }
- else
- jres),
- jres.uj_type)
- (hj,typ_cl) p)
- with e when Errors.noncritical e -> anomaly "apply_coercion"
-
- let inh_app_fun env evd j =
- let t = whd_betadeltaiota env evd j.uj_type in
- match kind_of_term t with
- | Prod (_,_,_) -> (evd,j)
- | Evar ev ->
- let (evd',t) = define_evar_as_product evd ev in
- (evd',{ uj_val = j.uj_val; uj_type = t })
- | _ ->
- let t,p =
- lookup_path_to_fun_from env evd j.uj_type in
- (evd,apply_coercion env evd p j t)
-
- let inh_app_fun resolve_tc env evd j =
- try inh_app_fun env evd j
- with
- | Not_found when not resolve_tc -> (evd, j)
- | Not_found ->
- try inh_app_fun env (saturate_evd env evd) j
- with Not_found -> (evd, 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 j1 = apply_coercion env evd p j t in
- let j2 = on_judgment_type (whd_evar evd) j1 in
- (evd,type_judgment env j2)
- with Not_found ->
- error_not_a_type_loc loc env evd j
-
- let inh_coerce_to_sort loc env evd j =
- let typ = whd_betadeltaiota 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_evar evd ev) ->
- let (evd',s) = define_evar_as_sort evd ev in
- (evd',{ utj_val = j.uj_val; utj_type = s })
- | _ ->
- inh_tosort_force loc env evd j
-
- let inh_coerce_to_base loc env evd j = (evd, j)
- let inh_coerce_to_prod loc env evd t = (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)
- then
- raise NoCoercion
- else
- let v', t' =
+ let coerce_application typ typ' c c' l l' =
+ let len = Array.length l in
+ let rec aux tele typ typ' i co =
+ if i < len then
+ let hdx = l.(i) and hdy = l'.(i) in
+ try evdref := the_conv_x_leq env hdx hdy !evdref;
+ let (n, eqT), restT = dest_prod typ in
+ let (n', eqT'), restT' = dest_prod typ' in
+ aux (hdx :: tele) (subst1 hdx restT) (subst1 hdy restT') (succ i) co
+ with UnableToUnify _ ->
+ let (n, eqT), restT = dest_prod typ in
+ let (n', eqT'), restT' = dest_prod typ' in
+ let _ =
+ try evdref := the_conv_x_leq env eqT eqT' !evdref
+ with UnableToUnify _ -> raise NoSubtacCoercion
+ in
+ (* Disallow equalities on arities *)
+ if Reduction.is_arity env 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 eq = papp evdref coq_eq_ind [| eqT; hdx; hdy |] in
+ let evar = make_existential loc env evdref eq in
+ let eq_app x = papp evdref coq_eq_rect
+ [| eqT; hdx; pred; x; hdy; evar|]
+ in
+ aux (hdy :: tele) (subst1 hdx restT)
+ (subst1 hdy restT') (succ i) (fun x -> eq_app (co x))
+ else Some (fun x ->
+ let term = co x in
+ Typing.solve_evars env evdref term)
+ in
+ if isEvar c || isEvar c' 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
+ | Sort s, Sort s' ->
+ (match s, 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))
+ in
+ let env' = push_rel (name', None, a') env in
+ let c1 = coerce_unify env' (lift 1 a') (lift 1 a) in
+ (* env, x : a' |- c1 : lift 1 a' > lift 1 a *)
+ let coec1 = app_opt env' evdref c1 (mkRel 1) in
+ (* env, x : a' |- c1[x] : lift 1 a *)
+ let c2 = coerce_unify env' (subst1 coec1 (liftn 1 2 b)) b' in
+ (* env, x : a' |- c2 : b[c1[x]/x]] > b' *)
+ (match c1, c2 with
+ | None, None -> None
+ | _, _ ->
+ Some
+ (fun f ->
+ mkLambda (name', a',
+ app_opt env' evdref c2
+ (mkApp (lift 1 f, [| coec1 |])))))
+
+ | App (c, l), App (c', l') ->
+ (match kind_of_term c, kind_of_term c' with
+ Ind (i, u), Ind (i', u') -> (* Inductive types *)
+ let len = Array.length l in
+ let sigT = delayed_force sigT_typ in
+ let prod = delayed_force prod_typ in
+ (* Sigma types *)
+ if Int.equal len (Array.length l') && Int.equal len 2 && eq_ind i i'
+ && (eq_ind i (destIndRef sigT) || eq_ind i (destIndRef prod))
+ then
+ if eq_ind i (destIndRef sigT)
+ then
+ begin
+ let (a, pb), (a', pb') =
+ pair_of_array l, pair_of_array l'
+ in
+ let c1 = coerce_unify env a a' in
+ let remove_head a c =
+ match kind_of_term c with
+ | Lambda (n, t, t') -> c, t'
+ (*| Prod (n, t, t') -> t'*)
+ | Evar (k, args) ->
+ let (evs, t) = Evarutil.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;
+ else ();
+ t, rng
+ | _ -> raise NoSubtacCoercion
+ in
+ let (pb, b), (pb', b') = remove_head a pb, remove_head a' pb' in
+ let env' = push_rel (Name Namegen.default_dependent_ident, None, a) env in
+ let c2 = coerce_unify env' b b' in
+ match c1, c2 with
+ | None, None -> None
+ | _, _ ->
+ Some
+ (fun x ->
+ let x, y =
+ app_opt env' evdref c1 (papp evdref sigT_proj1
+ [| a; pb; x |]),
+ app_opt env' evdref c2 (papp evdref sigT_proj2
+ [| a; pb; x |])
+ in
+ papp evdref sigT_intro [| a'; pb'; x ; y |])
+ end
+ else
+ begin
+ let (a, b), (a', b') =
+ pair_of_array l, pair_of_array l'
+ in
+ let c1 = coerce_unify env a a' in
+ let c2 = coerce_unify env b b' in
+ match c1, c2 with
+ None, None -> None
+ | _, _ ->
+ Some
+ (fun x ->
+ let x, y =
+ app_opt env evdref c1 (papp evdref prod_proj1
+ [| a; b; x |]),
+ app_opt env evdref c2 (papp evdref prod_proj2
+ [| a; b; x |])
+ in
+ papp evdref prod_intro [| a'; b'; x ; y |])
+ end
+ else
+ if eq_ind i i' && Int.equal len (Array.length l') then
+ let evm = !evdref in
+ (try subco ()
+ with NoSubtacCoercion ->
+ let typ = Typing.type_of env evm c in
+ let typ' = Typing.type_of env evm c' in
+ (* if not (is_arity env evm typ) then *)
+ coerce_application typ typ' c c' l l')
+ (* else subco () *)
+ else
+ subco ()
+ | x, y when Constr.equal c c' ->
+ if Int.equal (Array.length l) (Array.length l') then
+ let evm = !evdref in
+ let lam_type = Typing.type_of env evm c in
+ let lam_type' = Typing.type_of env evm c' in
+ (* if not (is_arity env evm lam_type) then ( *)
+ coerce_application lam_type lam_type' c c' l l'
+ (* ) else subco () *)
+ else subco ()
+ | _ -> subco ())
+ | _, _ -> subco ()
+
+ and subset_coerce env evdref x y =
+ match disc_subset 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
+ 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 |]))
+ in
+ (papp evdref sig_intro [| u; p; cx; evar |]))
+ | None ->
+ raise NoSubtacCoercion
+ in coerce_unify env x y
+
+let coerce_itf loc env evd v t c1 =
+ let evdref = ref evd in
+ let coercion = coerce loc env evdref t c1 in
+ let t = Option.map (app_opt env evdref coercion) v in
+ !evdref, t
+
+let saturate_evd env evd =
+ Typeclasses.resolve_typeclasses
+ ~filter:Typeclasses.no_goals ~split:false ~fail:false env evd
+
+(* appliquer le chemin de coercions p à hj *)
+let apply_coercion env sigma p hj typ_cl =
+ try
+ let j,t,evd =
+ List.fold_left
+ (fun (ja,typ_cl,sigma) i ->
+ let ((fv,isid,isproj),ctx) = coercion_value i in
+ let sigma = Evd.merge_context_set Evd.univ_flexible sigma ctx in
+ let argl = (class_args_of env sigma typ_cl)@[ja.uj_val] in
+ let sigma, jres =
+ apply_coercion_args env sigma true isproj argl fv
+ in
+ (if isid then
+ { uj_val = ja.uj_val; uj_type = jres.uj_type }
+ else
+ jres),
+ jres.uj_type,sigma)
+ (hj,typ_cl,sigma) p
+ in evd, j
+ with NoCoercion as e -> raise e
+ | e when Errors.noncritical e -> anomaly (Pp.str "apply_coercion")
+
+let inh_app_fun env evd j =
+ let t = whd_betadeltaiota env evd j.uj_type in
+ match kind_of_term t with
+ | Prod (_,_,_) -> (evd,j)
+ | Evar ev ->
+ let (evd',t) = define_evar_as_product evd ev in
+ (evd',{ uj_val = j.uj_val; uj_type = t })
+ | _ ->
+ try let t,p =
+ lookup_path_to_fun_from env evd j.uj_type in
+ apply_coercion env evd p j t
+ with Not_found | NoCoercion when Flags.is_program_mode () ->
+ try
+ let evdref = ref evd in
+ let coercef, t = mu env evdref t in
+ let res = { uj_val = app_opt env evdref coercef j.uj_val; uj_type = t } in
+ (!evdref, res)
+ with NoSubtacCoercion | NoCoercion ->
+ (evd,j)
+
+let inh_app_fun resolve_tc env evd j =
+ try inh_app_fun env evd j
+ with
+ | Not_found when not resolve_tc
+ || not !use_typeclasses_for_conversion -> (evd, j)
+ | Not_found ->
+ try inh_app_fun env (saturate_evd env evd) j
+ with Not_found -> (evd, 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)
+ with Not_found | NoCoercion ->
+ error_not_a_type_loc loc env evd j
+
+let inh_coerce_to_sort loc env evd j =
+ let typ = whd_betadeltaiota 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)) ->
+ let (evd',s) = define_evar_as_sort env evd ev in
+ (evd',{ utj_val = j.uj_val; utj_type = s })
+ | _ ->
+ inh_tosort_force 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_opt env evdref ct j.uj_val;
+ uj_type = typ' }
+ in !evdref, res
+ else (evd, j)
+
+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
+ !evdref, typ'
+ 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)
+ then
+ raise NoCoercion
+ else
+ let evd, v', t' =
try
let t2,t1,p = lookup_path_between env evd (t,c1) in
match v with
- Some v ->
- let j =
- apply_coercion env evd p
- {uj_val = v; uj_type = t} t2 in
- Some j.uj_val, j.uj_type
- | None -> None, t
+ | Some v ->
+ let evd,j =
+ apply_coercion env evd p
+ {uj_val = v; uj_type = t} t2 in
+ evd, Some j.uj_val, j.uj_type
+ | None -> evd, None, t
with Not_found -> raise NoCoercion
in
try (the_conv_x_leq env t' c1 evd, v')
- with Reduction.NotConvertible -> raise NoCoercion
+ with UnableToUnify _ -> raise NoCoercion
- 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 Reduction.NotConvertible ->
+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
+ match
kind_of_term (whd_betadeltaiota env evd t),
kind_of_term (whd_betadeltaiota 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 *)
- let name = match name with
- | Anonymous -> Name (id_of_string "x")
- | _ -> name in
- let env1 = push_rel (name,None,u1) env in
- let (evd', v1) =
- 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 t2 = match v2 with
- | None -> subst_term 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
- (evd'', Option.map (fun v2' -> mkLambda (name, u1, v2')) v2')
- | _ -> raise NoCoercion
-
- (* 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 (n, t) =
- match n with
- None ->
- let (evd', val') =
- try
- inh_conv_coerce_to_fail loc env evd rigidonly (Some cj.uj_val) cj.uj_type t
- with
- | NoCoercion when not resolve_tc -> error_actual_type_loc loc env evd cj t
- | NoCoercion ->
- let evd = saturate_evd env evd in
- try
- inh_conv_coerce_to_fail loc env evd rigidonly (Some cj.uj_val) cj.uj_type t
- with NoCoercion ->
- error_actual_type_loc loc env evd cj t
- in
- let val' = match val' with Some v -> v | None -> assert(false) in
- (evd',{ uj_val = val'; uj_type = t })
- | Some (init, cur) -> (evd, cj)
+ 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 *)
+ let name = match name with
+ | Anonymous -> Name Namegen.default_dependent_ident
+ | _ -> name in
+ let env1 = push_rel (name,None,u1) env in
+ let (evd', v1) =
+ 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 t2 = match v2 with
+ | None -> subst_term 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
+ (evd'', Option.map (fun v2' -> mkLambda (name, u1, v2')) v2')
+ | _ -> raise (NoCoercionNoUnifier (best_failed_evd,e))
- 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
+(* 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 (evd', val') =
+ try
+ 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
+ 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
+ | 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
+ else
+ inh_conv_coerce_to_fail loc env evd' rigidonly (Some cj.uj_val) cj.uj_type t
+ with NoCoercionNoUnifier (best_failed_evd,e) ->
+ error_actual_type_loc 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_coerces_to loc env (evd : evar_map) t (abs, t') =
- if abs = None then
- try
- fst (inh_conv_coerce_to_fail loc env evd true None t t')
- with NoCoercion ->
- evd (* Maybe not enough information to unify *)
- else
- evd
- (* Still problematic, as it changes unification
- let nabsinit, nabs =
- match abs with
- None -> 0, 0
- | Some (init, cur) -> init, cur
- in
- try
- let (rels, rng) =
- (* a little more effort to get products is needed *)
- try decompose_prod_n nabs t
- with _ ->
- if !Flags.debug then
- msg_warning (str "decompose_prod_n failed");
- raise (Invalid_argument "Coercion.inh_conv_coerces_to")
- in
- (* The final range free variables must have been replaced by evars, we accept only that evars
- in rng are applied to free vars. *)
- if noccur_with_meta 0 (succ nabsinit) rng then (
- let env', t, t' =
- let env' = List.fold_right (fun (n, t) env -> push_rel (n, None, t) env) rels env in
- env', rng, lift nabs t'
- in
- try
- pi1 (inh_conv_coerce_to_fail loc env' evd None t t')
- with NoCoercion ->
- evd) (* Maybe not enough information to unify *)
- (*let sigma = evd in
- error_cannot_coerce env' sigma (t, t'))*)
- else evd
- with Invalid_argument _ -> evd *)
-end
+let inh_conv_coerces_to loc env evd t t' =
+ try
+ 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 f06f58f4..f511f977 100644
--- a/pretyping/coercion.mli
+++ b/pretyping/coercion.mli
@@ -1,69 +1,62 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Util
open Evd
open Names
open Term
-open Sign
open Environ
-open Evarutil
open Glob_term
-module type S = sig
- (** {6 Coercions. } *)
-
- (** [inh_app_fun resolve_tc env isevars j] coerces [j] to a function; i.e. it
- inserts a coercion into [j], if needed, in such a way it gets as
- type a product; it returns [j] if no coercion is applicable.
- resolve_tc=false disables resolving type classes (as the last
- resort before failing) *)
- val inh_app_fun :
- bool -> env -> evar_map -> unsafe_judgment -> evar_map * unsafe_judgment
-
- (** [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 ->
- 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 ->
- 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 ->
- env -> evar_map -> type_constraint_type -> evar_map * type_constraint_type
-
- (** [inh_conv_coerce_to resolve_tc loc env isevars j t] coerces [j] to an object of type
- [t]; i.e. it inserts a coercion into [j], if needed, in such 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 ->
- env -> evar_map -> unsafe_judgment -> type_constraint_type -> evar_map * unsafe_judgment
-
- val inh_conv_coerce_rigid_to : bool -> loc ->
- env -> evar_map -> unsafe_judgment -> type_constraint_type -> 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 ->
- env -> evar_map -> types -> type_constraint_type -> 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 -> cases_pattern -> inductive -> inductive -> cases_pattern
-end
-
-module Default : S
+(** {6 Coercions. } *)
+
+(** [inh_app_fun resolve_tc env isevars j] coerces [j] to a function; i.e. it
+ inserts a coercion into [j], if needed, in such a way it gets as
+ type a product; it returns [j] if no coercion is applicable.
+ resolve_tc=false disables resolving type classes (as the last
+ resort before failing) *)
+val inh_app_fun : bool ->
+ env -> evar_map -> unsafe_judgment -> evar_map * unsafe_judgment
+
+(** [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 ->
+ 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 ->
+ 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 ->
+ env -> evar_map -> types -> evar_map * types
+
+(** [inh_conv_coerce_to resolve_tc Loc.t env isevars j t] coerces [j] to an
+ object of type [t]; i.e. it inserts a coercion into [j], if needed, in such
+ 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 ->
+ env -> evar_map -> unsafe_judgment -> types -> evar_map * unsafe_judgment
+
+val inh_conv_coerce_rigid_to : bool -> Loc.t ->
+ 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 ->
+ 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
diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml
new file mode 100644
index 00000000..a6e2bc19
--- /dev/null
+++ b/pretyping/constr_matching.ml
@@ -0,0 +1,494 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i*)
+open Pp
+open Errors
+open Util
+open Names
+open Globnames
+open Nameops
+open Termops
+open Reductionops
+open Term
+open Vars
+open Context
+open Pattern
+open Patternops
+open Misctypes
+(*i*)
+
+(* Given a term with second-order variables in it,
+ represented by Meta's, and possibly applied using [SOAPP] to
+ terms, this function will perform second-order, binding-preserving,
+ matching, in the case where the pattern is a pattern in the sense
+ of Dale Miller.
+
+ ALGORITHM:
+
+ Given a pattern, we decompose it, flattening Cast's and apply's,
+ recursing on all operators, and pushing the name of the binder each
+ time we descend a binder.
+
+ When we reach a first-order variable, we ask that the corresponding
+ term's free-rels all be higher than the depth of the current stack.
+
+ When we reach a second-order application, we ask that the
+ intersection of the free-rels of the term and the current stack be
+ contained in the arguments of the application, and in that case, we
+ construct a LAMBDA with the names on the stack.
+
+ *)
+
+type bound_ident_map = Id.t Id.Map.t
+
+exception PatternMatchingFailure
+
+let warn_bound_meta name =
+ msg_warning (str "Collision between bound variable " ++ pr_id name ++
+ str " and a metavariable of same name.")
+
+let warn_bound_bound name =
+ msg_warning (str "Collision between bound variables of name " ++ pr_id name)
+
+let warn_bound_again name =
+ msg_warning (str "Collision between bound variable " ++ pr_id name ++
+ str " and another bound variable of same name.")
+
+let constrain n (ids, m as x) (names, terms as subst) =
+ try
+ let (ids', m') = Id.Map.find n terms in
+ if List.equal Id.equal ids ids' && eq_constr m m' then subst
+ else raise PatternMatchingFailure
+ with Not_found ->
+ let () = if Id.Map.mem n names then warn_bound_meta n in
+ (names, Id.Map.add n x terms)
+
+let add_binders na1 na2 (names, terms as subst) = match na1, na2 with
+| Name id1, Name id2 ->
+ if Id.Map.mem id1 names then
+ let () = warn_bound_bound id1 in
+ (names, terms)
+ else
+ let names = Id.Map.add id1 id2 names in
+ let () = if Id.Map.mem id1 terms then warn_bound_again id1 in
+ (names, terms)
+| _ -> subst
+
+let rec build_lambda vars stk m = match vars with
+| [] ->
+ let len = List.length stk in
+ lift (-1 * len) m
+| n :: vars ->
+ (* change [ x1 ... xn y z1 ... zm |- t ] into
+ [ x1 ... xn z1 ... zm |- lam y. t ] *)
+ let len = List.length stk 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) stk in
+ 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
+
+let rec extract_bound_aux k accu frels stk = match stk with
+| [] -> accu
+| (na1, na2, _) :: stk ->
+ if Int.Set.mem k frels then
+ begin match na1 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 stk
+ | Anonymous -> raise PatternMatchingFailure
+ end
+ else extract_bound_aux (k + 1) accu frels stk
+
+let extract_bound_vars frels stk =
+ extract_bound_aux 1 Id.Set.empty frels stk
+
+let dummy_constr = mkProp
+
+let make_renaming ids = function
+| (Name id, Name _, _) ->
+ begin
+ try mkRel (List.index Id.equal id ids)
+ with Not_found -> dummy_constr
+ end
+| _ -> dummy_constr
+
+let merge_binding allow_bound_rels stk n cT subst =
+ let c = match stk with
+ | [] -> (* Optimization *)
+ ([], cT)
+ | _ ->
+ let frels = free_rels cT in
+ if allow_bound_rels then
+ let vars = extract_bound_vars frels stk in
+ let ordered_vars = Id.Set.elements vars in
+ let rename binding = make_renaming ordered_vars binding in
+ let renaming = List.map rename stk in
+ (ordered_vars, substl renaming cT)
+ else
+ let depth = List.length stk in
+ let min_elt = try Int.Set.min_elt frels with Not_found -> succ depth in
+ if depth < min_elt then
+ ([], lift (- depth) cT)
+ else raise PatternMatchingFailure
+ in
+ constrain n c subst
+
+let matches_core env sigma convert allow_partial_app allow_bound_rels pat c =
+ 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'
+ | 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)
+ in
+ let rec sorec stk env subst p t =
+ let cT = strip_outer_cast t in
+ match p,kind_of_term 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
+ (n :: ans, Int.Set.add n seen)
+ | _ -> error "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
+ if Int.Set.subset frels relset then
+ constrain n ([], build_lambda relargs stk cT) subst
+ else
+ raise PatternMatchingFailure
+
+ | PMeta (Some n), m -> merge_binding allow_bound_rels stk n cT subst
+
+ | PMeta None, m -> subst
+
+ | PRef (VarRef v1), Var v2 when Id.equal v1 v2 -> subst
+
+ | PVar v1, Var v2 when Id.equal v1 v2 -> subst
+
+ | PRef ref, _ when convref ref cT -> subst
+
+ | PRel n1, Rel n2 when Int.equal n1 n2 -> subst
+
+ | PSort GProp, Sort (Prop Null) -> subst
+
+ | PSort GSet, Sort (Prop Pos) -> subst
+
+ | PSort (GType _), Sort (Type _) -> subst
+
+ | PApp (p, [||]), _ -> sorec stk env subst p t
+
+ | PApp (PApp (h, a1), a2), _ ->
+ sorec stk env subst (PApp(h,Array.append a1 a2)) t
+
+ | PApp (PMeta meta,args1), App (c2,args2) when allow_partial_app ->
+ (let diff = Array.length args2 - Array.length args1 in
+ if diff >= 0 then
+ let args21, args22 = Array.chop diff args2 in
+ let c = mkApp(c2,args21) in
+ let subst =
+ match meta with
+ | None -> subst
+ | Some n -> merge_binding allow_bound_rels stk n c subst in
+ Array.fold_left2 (sorec stk env) subst args1 args22
+ else (* Might be a projection on the right *)
+ match kind_of_term 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 stk env subst p term
+ with Retyping.RetypeError _ -> raise PatternMatchingFailure)
+ | _ -> 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))
+ || Projection.unfolded pr ->
+ raise PatternMatchingFailure
+ | PProj (pr1,c1), Proj (pr,c) ->
+ if Projection.equal pr1 pr then
+ try Array.fold_left2 (sorec stk env) (sorec stk env subst c1 c) arg1 arg2
+ with Invalid_argument _ -> raise PatternMatchingFailure
+ else raise PatternMatchingFailure
+ | _, Proj (pr,c) when not (Projection.unfolded pr) ->
+ (try let term = Retyping.expand_projection env sigma pr c (Array.to_list arg2) in
+ sorec stk env subst p term
+ with Retyping.RetypeError _ -> raise PatternMatchingFailure)
+ | _, _ ->
+ try Array.fold_left2 (sorec stk env) (sorec stk env subst c1 c2) arg1 arg2
+ with Invalid_argument _ -> raise PatternMatchingFailure)
+
+ | PApp (PRef (ConstRef c1), _), Proj (pr, c2)
+ when Projection.unfolded pr || not (eq_constant c1 (Projection.constant pr)) ->
+ raise PatternMatchingFailure
+
+ | PApp (c, args), Proj (pr, c2) ->
+ (try let term = Retyping.expand_projection env sigma pr c2 [] in
+ sorec stk env subst p term
+ with Retyping.RetypeError _ -> raise PatternMatchingFailure)
+
+ | PProj (p1,c1), Proj (p2,c2) when Projection.equal p1 p2 ->
+ sorec stk env subst c1 c2
+
+ | PProd (na1,c1,d1), Prod(na2,c2,d2) ->
+ sorec ((na1,na2,c2)::stk) (Environ.push_rel (na2,None,c2) env)
+ (add_binders na1 na2 (sorec stk env subst c1 c2)) d1 d2
+
+ | PLambda (na1,c1,d1), Lambda(na2,c2,d2) ->
+ sorec ((na1,na2,c2)::stk) (Environ.push_rel (na2,None,c2) env)
+ (add_binders na1 na2 (sorec stk env subst c1 c2)) d1 d2
+
+ | PLetIn (na1,c1,d1), LetIn(na2,c2,t2,d2) ->
+ sorec ((na1,na2,t2)::stk) (Environ.push_rel (na2,Some c2,t2) env)
+ (add_binders na1 na2 (sorec stk env subst c1 c2)) d1 d2
+
+ | PIf (a1,b1,b1'), Case (ci,_,a2,[|b2;b2'|]) ->
+ let ctx,b2 = decompose_lam_n_assum ci.ci_cstr_ndecls.(0) b2 in
+ let ctx',b2' = decompose_lam_n_assum ci.ci_cstr_ndecls.(1) b2' in
+ let n = rel_context_length ctx in
+ let n' = rel_context_length ctx' in
+ if noccur_between 1 n b2 && noccur_between 1 n' b2' then
+ let s =
+ List.fold_left (fun l (na,_,t) -> (Anonymous,na,t)::l) stk ctx in
+ let s' =
+ List.fold_left (fun l (na,_,t) -> (Anonymous,na,t)::l) stk ctx' in
+ let b1 = lift_pattern n b1 and b1' = lift_pattern n' b1' in
+ sorec s' (Environ.push_rel_context ctx' env)
+ (sorec s (Environ.push_rel_context ctx env) (sorec stk env subst a1 a2) b1 b2) b1' b2'
+ else
+ raise PatternMatchingFailure
+
+ | PCase (ci1,p1,a1,br1), Case (ci2,p2,a2,br2) ->
+ let n2 = Array.length br2 in
+ let () = match ci1.cip_ind with
+ | None -> ()
+ | Some ind1 ->
+ (** ppedrot: Something spooky going here. The comparison used to be
+ the generic one, so I may have broken something. *)
+ if not (eq_ind ind1 ci2.ci_ind) then raise PatternMatchingFailure
+ in
+ let () =
+ if not ci1.cip_extensible && not (Int.equal (List.length br1) n2)
+ then raise PatternMatchingFailure
+ in
+ let chk_branch subst (j,n,c) =
+ (* (ind,j+1) is normally known to be a correct constructor
+ and br2 a correct match over the same inductive *)
+ assert (j < n2);
+ sorec stk env subst c br2.(j)
+ in
+ let chk_head = sorec stk env (sorec stk 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
+
+ in
+ sorec [] env (Id.Map.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 extended_matches env sigma = matches_core env sigma false true true
+
+let matches env sigma pat c = snd (matches_core_closed env sigma false true pat c)
+
+let special_meta = (-1)
+
+type matching_result =
+ { m_sub : bound_ident_map * patvar_map;
+ m_ctx : constr; }
+
+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 head =
+ match pat, kind_of_term c with
+ | PApp (c1,arg1), App (c2,arg2) ->
+ if isPMeta c1 then c else
+ let n1 = Array.length arg1 in
+ if n1 < Array.length arg2 then mkApp (c2,Array.sub arg2 0 n1) else c
+ | c1, App (c2,arg2) when not (isPMeta c1) -> c2
+ | _ -> c in
+ 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 next =
+ 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)
+ then next ()
+ else mkresult subst (mk_ctx (mkMeta special_meta)) next
+ with PatternMatchingFailure -> next ()
+
+(* Tries to match a subterm of [c] with [pat] *)
+let sub_match ?(partial_app=false) ?(closed=true) env sigma pat c =
+ let rec aux env c mk_ctx next =
+ match kind_of_term c with
+ | Cast (c1,k,c2) ->
+ let next_mk_ctx lc = mk_ctx (mkCast (List.hd lc, k,c2)) in
+ let next () = try_aux [env] [c1] next_mk_ctx next in
+ authorized_occ env sigma partial_app closed pat c mk_ctx next
+ | Lambda (x,c1,c2) ->
+ let next_mk_ctx lc = mk_ctx (mkLambda (x,List.hd lc,List.nth lc 1)) in
+ let next () =
+ let env' = Environ.push_rel (x,None,c1) env in
+ try_aux [env;env'] [c1; c2] next_mk_ctx next in
+ authorized_occ env sigma partial_app closed pat c mk_ctx next
+ | Prod (x,c1,c2) ->
+ let next_mk_ctx lc = mk_ctx (mkProd (x,List.hd lc,List.nth lc 1)) in
+ let next () =
+ let env' = Environ.push_rel (x,None,c1) env in
+ try_aux [env;env'] [c1;c2] next_mk_ctx next in
+ authorized_occ env sigma partial_app closed pat c mk_ctx next
+ | LetIn (x,c1,t,c2) ->
+ let next_mk_ctx = function
+ | [c1;c2] -> mkLetIn (x,c1,t,c2)
+ | _ -> assert false
+ in
+ let next () =
+ let env' = Environ.push_rel (x,Some c1,t) env in
+ try_aux [env;env'] [c1;c2] next_mk_ctx next in
+ authorized_occ env sigma partial_app closed pat c mk_ctx next
+ | App (c1,lc) ->
+ let next () =
+ 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;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
+ try_aux [env] (c1::Array.to_list lc) 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
+ try_aux [env] (c1::Array.to_list lc) mk_ctx next
+ in
+ authorized_occ env sigma partial_app closed pat c mk_ctx next
+ | Case (ci,hd,c1,lc) ->
+ let next_mk_ctx = function
+ | [] -> assert false
+ | c1 :: lc -> mk_ctx (mkCase (ci,hd,c1,Array.of_list lc))
+ in
+ let next () = try_aux [env] (c1 :: Array.to_list lc) next_mk_ctx next in
+ authorized_occ env sigma partial_app closed pat c mk_ctx next
+ | Fix (indx,(names,types,bodies)) ->
+ 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 next () =
+ try_aux
+ [env] ((Array.to_list types)@(Array.to_list bodies)) next_mk_ctx next in
+ authorized_occ env sigma partial_app closed pat c mk_ctx next
+ | CoFix (i,(names,types,bodies)) ->
+ 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 next () =
+ try_aux [env] ((Array.to_list types)@(Array.to_list bodies)) next_mk_ctx next in
+ authorized_occ env sigma partial_app closed pat c mk_ctx next
+ | Proj (p,c') ->
+ let next_mk_ctx le = mk_ctx (mkProj (p,List.hd le)) in
+ let next () =
+ if partial_app then
+ try
+ let term = Retyping.expand_projection env sigma p c' [] in
+ aux env term mk_ctx next
+ with Retyping.RetypeError _ -> raise PatternMatchingFailure
+ else
+ try_aux [env] [c'] next_mk_ctx next in
+ authorized_occ env sigma partial_app closed pat c mk_ctx next
+ | Construct _| Ind _|Evar _|Const _ | Rel _|Meta _|Var _|Sort _ ->
+ authorized_occ env sigma partial_app closed pat c mk_ctx next
+
+ (* Tries [sub_match] for all terms in the list *)
+ and try_aux lenv lc mk_ctx next =
+ let rec try_sub_match_rec lacc lenv lc =
+ match lenv, lc with
+ | _, [] -> next ()
+ | env :: tlenv, c::tl ->
+ let mk_ctx ce = mk_ctx (List.rev_append lacc (ce::tl)) in
+ let next () =
+ let env' = match tlenv with [] -> lenv | _ -> tlenv in
+ try_sub_match_rec (c::lacc) env' tl
+ in
+ aux env c mk_ctx next
+ | _ -> assert false in
+ try_sub_match_rec [] lenv lc in
+ let lempty () = IStream.Nil in
+ let result () = aux env c (fun x -> x) lempty in
+ IStream.thunk result
+
+let match_subterm env sigma pat c = sub_match env sigma pat c
+
+let match_appsubterm env sigma pat c =
+ sub_match ~partial_app:true env sigma pat c
+
+let match_subterm_gen env sigma app pat c =
+ sub_match ~partial_app:app env sigma pat c
+
+let is_matching env sigma pat c =
+ try let _ = matches env sigma pat c in true
+ with PatternMatchingFailure -> false
+
+let is_matching_head env sigma pat c =
+ try let _ = matches_head env sigma pat c in true
+ with PatternMatchingFailure -> false
+
+let is_matching_appsubterm ?(closed=true) env sigma pat c =
+ let results = sub_match ~partial_app:true ~closed env sigma pat c in
+ not (IStream.is_empty results)
+
+let matches_conv env sigma c p =
+ snd (matches_core_closed env sigma true false c p)
+
+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/matching.mli b/pretyping/constr_matching.mli
index 1a47b714..67854a89 100644
--- a/pretyping/matching.mli
+++ b/pretyping/constr_matching.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -12,7 +12,6 @@ open Names
open Term
open Environ
open Pattern
-open Termops
(** [PatternMatchingFailure] is the exception raised when pattern
matching fails *)
@@ -25,56 +24,60 @@ val special_meta : metavariable
(** [bound_ident_map] represents the result of matching binding
identifiers of the pattern with the binding identifiers of the term
matched *)
-type bound_ident_map = (identifier * identifier) list
+type bound_ident_map = Id.t Id.Map.t
(** [matches pat c] matches [c] against [pat] and returns the resulting
assignment of metavariables; it raises [PatternMatchingFailure] if
not matchable; bindings are given in increasing order based on the
numbers given in the pattern *)
-val matches : constr_pattern -> constr -> patvar_map
+val matches : env -> Evd.evar_map -> constr_pattern -> constr -> patvar_map
+
+(** [matches_head pat c] does the same as [matches pat c] but accepts
+ [pat] to match an applicative prefix of [c] *)
+val matches_head : env -> Evd.evar_map -> constr_pattern -> constr -> patvar_map
(** [extended_matches pat c] also returns the names of bound variables
in [c] that matches the bound variables in [pat]; if several bound
variables or metavariables have the same name, the metavariable,
or else the rightmost bound variable, takes precedence *)
val extended_matches :
- constr_pattern -> constr -> bound_ident_map * extended_patvar_map
+ env -> Evd.evar_map -> constr_pattern -> constr -> bound_ident_map * extended_patvar_map
(** [is_matching pat c] just tells if [c] matches against [pat] *)
-val is_matching : constr_pattern -> constr -> bool
+val is_matching : env -> Evd.evar_map -> constr_pattern -> constr -> bool
+
+(** [is_matching_head pat c] just tells if [c] or an applicative
+ 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
+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 with [special_meta]) + a continuation that
- either returns the next matching subterm or raise PatternMatchingFailure *)
-type subterm_matching_result =
- (bound_ident_map * patvar_map) * constr * (unit -> subterm_matching_result)
+ (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 the first **closed** subterm of [c] matching [pat], and
- a continuation that looks for the next matching subterm.
- It raises PatternMatchingFailure if no subterm matches the pattern *)
-val match_subterm : constr_pattern -> constr -> subterm_matching_result
+ corresponding to each **closed** subterm of [c] matching [pat]. *)
+val match_subterm : env -> Evd.evar_map -> constr_pattern -> constr -> matching_result IStream.t
(** [match_appsubterm pat c] returns the substitution and the context
- corresponding to the first **closed** subterm of [c] matching [pat],
- considering application contexts as well. It also returns a
- continuation that looks for the next matching subterm.
- It raises PatternMatchingFailure if no subterm matches the pattern *)
-val match_appsubterm : constr_pattern -> constr -> subterm_matching_result
+ 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 : bool (** true = with app context *) ->
- constr_pattern -> constr -> subterm_matching_result
+val match_subterm_gen : env -> Evd.evar_map -> bool (** true = with app context *) ->
+ 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 -> constr_pattern -> constr -> bool
+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 *)
diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml
index 9e808dd4..046ee0da 100644
--- a/pretyping/detyping.ml
+++ b/pretyping/detyping.ml
@@ -1,38 +1,51 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Pp
+open Errors
open Util
-open Univ
open Names
open Term
-open Declarations
-open Inductive
+open Vars
+open Context
open Inductiveops
open Environ
-open Sign
open Glob_term
-open Nameops
+open Glob_ops
open Termops
open Namegen
open Libnames
+open Globnames
open Nametab
-open Evd
open Mod_subst
+open Misctypes
+open Decl_kinds
-let dl = dummy_loc
+let dl = Loc.ghost
+
+(** Should we keep details of universes during detyping ? *)
+let print_universes = Flags.univ_print
+
+(** If true, prints local context of evars, whatever print_arguments *)
+let print_evar_arguments = ref false
+
+let add_name na b t (nenv, env) = add_name na nenv, push_rel (na, b, t) env
+let add_name_opt na b t (nenv, env) =
+ match t with
+ | None -> Termops.add_name na nenv, env
+ | Some t -> add_name na b t (nenv, env)
(****************************************************************************)
(* Tools for printing of Cases *)
let encode_inductive r =
let indsp = global_inductive r in
- let constr_lengths = mis_constr_nargs indsp in
+ let constr_lengths = constructors_nrealargs indsp in
(indsp,constr_lengths)
(* Parameterization of the translation from constr to ast *)
@@ -40,9 +53,9 @@ let encode_inductive r =
(* Tables for Cases printing under a "if" form, a "let" form, *)
let has_two_constructors lc =
- Array.length lc = 2 (* & lc.(0) = 0 & lc.(1) = 0 *)
+ Int.equal (Array.length lc) 2 (* & lc.(0) = 0 & lc.(1) = 0 *)
-let isomorphic_to_tuple lc = (Array.length lc = 1)
+let isomorphic_to_tuple lc = Int.equal (Array.length lc) 1
let encode_bool r =
let (x,lc) = encode_inductive r in
@@ -67,12 +80,10 @@ module PrintingInductiveMake =
end) ->
struct
type t = inductive
+ let compare = ind_ord
let encode = Test.encode
- let subst subst (kn, ints as obj) =
- let kn' = subst_ind subst kn in
- if kn' == kn then obj else
- kn', ints
- let printer ind = pr_global_env Idset.empty (IndRef ind)
+ let subst subst obj = subst_ind subst obj
+ let printer ind = pr_global_env Id.Set.empty (IndRef ind)
let key = ["Printing";Test.field]
let title = Test.title
let member_message x = Test.member_message (printer x)
@@ -83,7 +94,7 @@ module PrintingCasesIf =
PrintingInductiveMake (struct
let encode = encode_bool
let field = "If"
- let title = "Types leading to pretty-printing of Cases using a `if' form: "
+ let title = "Types leading to pretty-printing of Cases using a `if' form:"
let member_message s b =
str "Cases on elements of " ++ s ++
str
@@ -144,6 +155,17 @@ let _ = declare_bool_option
optread = reverse_matching;
optwrite = (:=) reverse_matching_value }
+let print_primproj_params_value = ref true
+let print_primproj_params () = !print_primproj_params_value
+
+let _ = declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "printing of primitive projection parameters";
+ optkey = ["Printing";"Primitive";"Projection";"Parameters"];
+ optread = print_primproj_params;
+ optwrite = (:=) print_primproj_params_value }
+
(* Auxiliary function for MutCase printing *)
(* [computable] tries to tell if the predicate typing the result is inferable*)
@@ -153,19 +175,19 @@ let computable p k =
works for normal eta-expanded term. For non eta-expanded or
non-normal terms, it may affirm the pred is synthetisable
because of an undetected ultimate dependent variable in the second
- clause, or else, it may affirms the pred non synthetisable
+ clause, or else, it may affirm the pred non synthetisable
because of a non normal term in the fourth clause.
A solution could be to store, in the MutCase, the eta-expanded
normal form of pred to decide if it depends on its variables
- Lorsque le prédicat est dépendant de manière certaine, on
- ne déclare pas le prédicat synthétisable (même si la
- variable dépendante ne l'est pas effectivement) parce que
- sinon on perd la réciprocité de la synthèse (qui, lui,
- engendrera un prédicat non dépendant) *)
+ Lorsque le prédicat est dépendant de manière certaine, on
+ ne déclare pas le prédicat synthétisable (même si la
+ variable dépendante ne l'est pas effectivement) parce que
+ 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
- (rel_context_length sign = k+1)
+ Int.equal (rel_context_length sign) (k + 1)
&&
noccur_between 1 (k+1) ccl
@@ -173,11 +195,11 @@ let lookup_name_as_displayed env t s =
let rec lookup avoid n c = match kind_of_term c with
| Prod (name,_,c') ->
(match compute_displayed_name_in RenamingForGoal avoid name c' with
- | (Name id,avoid') -> if id=s then Some n else lookup avoid' (n+1) c'
+ | (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
- | (Name id,avoid') -> if id=s then Some n else lookup avoid' (n+1) c'
+ | (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
@@ -189,9 +211,9 @@ let lookup_index_as_renamed env t n =
(match compute_displayed_name_in RenamingForGoal [] name c' with
(Name _,_) -> lookup n (d+1) c'
| (Anonymous,_) ->
- if n=0 then
+ if Int.equal n 0 then
Some (d-1)
- else if n=1 then
+ else if Int.equal n 1 then
Some d
else
lookup (n-1) (d+1) c')
@@ -199,55 +221,63 @@ let lookup_index_as_renamed env t n =
(match compute_displayed_name_in RenamingForGoal [] name c' with
| (Name _,_) -> lookup n (d+1) c'
| (Anonymous,_) ->
- if n=0 then
+ if Int.equal n 0 then
Some (d-1)
- else if n=1 then
+ else if Int.equal n 1 then
Some d
else
lookup (n-1) (d+1) c'
)
| Cast (c,_,_) -> lookup n d c
- | _ -> if n=0 then Some (d-1) else None
+ | _ -> if Int.equal n 0 then Some (d-1) else None
in lookup n 1 t
(**********************************************************************)
(* Fragile algorithm to reverse pattern-matching compilation *)
-let update_name na ((_,e),c) =
+let update_name na ((_,(e,_)),c) =
match na with
- | Name _ when force_wildcard () & noccurn (list_index na e) c ->
+ | Name _ when force_wildcard () && noccurn (List.index Name.equal na e) c ->
Anonymous
| _ ->
na
-let rec decomp_branch n nal b (avoid,env as e) c =
- let flag = if b then RenamingForGoal else RenamingForCasesPattern in
- if n=0 then (List.rev nal,(e,c))
- else
- let na,c,f =
- match kind_of_term (strip_outer_cast c) with
- | Lambda (na,_,c) -> na,c,compute_displayed_let_name_in
- | LetIn (na,_,_,c) -> na,c,compute_displayed_name_in
- | _ ->
- Name (id_of_string "x"),(applist (lift 1 c, [mkRel 1])),
- compute_displayed_name_in in
+let rec decomp_branch tags nal b (avoid,env as e) 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
+ | 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
+ | _, false ->
+ Name default_dependent_ident,(applist (lift 1 c, [mkRel 1])),
+ compute_displayed_name_in,None,None
+ | _, true ->
+ Anonymous,lift 1 c,compute_displayed_name_in,None,None
+ in
let na',avoid' = f flag avoid na c in
- decomp_branch (n-1) (na'::nal) b (avoid',add_name na' env) c
+ decomp_branch tags (na'::nal) b
+ (avoid', add_name_opt na' body t env) 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 cnl = ci.ci_cstr_ndecls in
+ let cnl = ci.ci_pp_info.cstr_tags in
+ let cna = ci.ci_cstr_nargs in
List.flatten
- (list_tabulate (fun i -> contract_branch isgoal e (cnl.(i),mkpat i,cl.(i)))
- (Array.length cl))
+ (List.init (Array.length cl)
+ (fun i -> contract_branch isgoal e (cnl.(i),cna.(i),mkpat i,cl.(i))))
and align_tree nal isgoal (e,c as rhs) = match nal with
| [] -> [[],rhs]
| na::nal ->
match kind_of_term c with
- | Case (ci,p,c,cl) when c = mkRel (list_index na (snd e))
- & (* don't contract if p dependent *)
- computable p (ci.ci_pp_info.ind_nargs) ->
+ | Case (ci,p,c,cl) when
+ eq_constr c (mkRel (List.index Name.equal na (fst (snd e))))
+ && (* 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
List.flatten
(List.map (fun (pat,rhs) ->
@@ -259,8 +289,8 @@ and align_tree nal isgoal (e,c as rhs) = match nal with
let mat = align_tree nal isgoal rhs in
List.map (fun (hd,rest) -> pat::hd,rest) mat
-and contract_branch isgoal e (cn,mkpat,b) =
- let nal,rhs = decomp_branch cn [] isgoal e b in
+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
@@ -268,48 +298,53 @@ and contract_branch isgoal e (cn,mkpat,b) =
(* Transform internal representation of pattern-matching into list of *)
(* clauses *)
-let is_nondep_branch c n =
+let is_nondep_branch c l =
try
- let sign,ccl = decompose_lam_n_assum n c in
+ (* FIXME: do better using tags from l *)
+ let sign,ccl = decompose_lam_n_assum (List.length l) c in
noccur_between 1 (rel_context_length sign) ccl
with e when Errors.noncritical e -> (* Not eta-expanded or not reduced *)
false
-let extract_nondep_branches test c b n =
- let rec strip n r = if n=0 then r else
- match r with
- | GLambda (_,_,_,_,t) -> strip (n-1) t
- | GLetIn (_,_,_,t) -> strip (n-1) t
- | _ -> assert false in
- if test c n then Some (strip n b) else None
-
-let it_destRLambda_or_LetIn_names n c =
- let rec aux n nal c =
- if n=0 then (List.rev nal,c) else match c with
- | GLambda (_,na,_,_,c) -> aux (n-1) (na::nal) c
- | GLetIn (_,na,_,c) -> aux (n-1) (na::nal) c
- | _ ->
+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
+ | GLetIn (_,_,_,t), true::l -> strip l t
+ (* FIXME: do we need adjustment? *)
+ | _,_ -> assert false in
+ if test c l then Some (strip l b) else None
+
+let it_destRLambda_or_LetIn_names l c =
+ let rec aux l nal c =
+ match 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
+ | _, true::l -> (* let-expansion *) aux l (Anonymous :: nal) c
+ | _, false::l ->
(* eta-expansion *)
- let rec next l =
- let x = next_ident_away (id_of_string "x") l in
+ let next l =
+ let x = next_ident_away default_dependent_ident l in
(* Not efficient but unusual and no function to get free glob_vars *)
(* if occur_glob_constr x c then next (x::l) else x in *)
x
in
let x = next (free_glob_vars c) in
let a = GVar (dl,x) in
- aux (n-1) (Name x :: nal)
+ aux l (Name x :: nal)
(match c with
| GApp (loc,p,l) -> GApp (loc,p,l@[a])
| _ -> (GApp (dl,c,[a])))
- in aux n [] c
+ in aux l [] c
let detype_case computable detype detype_eqns testdep avoid data p c bl =
- let (indsp,st,nparams,consnargsl,k) = data in
+ let (indsp,st,constagsl,k) = data in
let synth_type = synthetize_type () in
let tomatch = detype c in
let alias, aliastyp, pred=
- if (not !Flags.raw_print) & synth_type & computable & Array.length bl<>0
+ if (not !Flags.raw_print) && synth_type && computable && not (Int.equal (Array.length bl) 0)
then
Anonymous, None, None
else
@@ -321,8 +356,8 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl =
| GLambda (_,x,_,t,c) -> x, c
| _ -> Anonymous, typ in
let aliastyp =
- if List.for_all ((=) Anonymous) nl then None
- else Some (dl,indsp,nparams,nl) in
+ if List.for_all (Name.equal Anonymous) nl then None
+ else Some (dl,indsp,nl) in
n, aliastyp, Some typ
in
let constructs = Array.init (Array.length bl) (fun i -> (indsp,i+1)) in
@@ -330,7 +365,7 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl =
try
if !Flags.raw_print then
RegularStyle
- else if st = LetPatternStyle then
+ else if st == LetPatternStyle then
st
else if PrintingLet.active indsp then
LetStyle
@@ -340,117 +375,194 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl =
st
with Not_found -> st
in
- match tag with
- | LetStyle when aliastyp = None ->
+ match tag, aliastyp with
+ | LetStyle, None ->
let bl' = Array.map detype bl in
- let (nal,d) = it_destRLambda_or_LetIn_names consnargsl.(0) bl'.(0) in
+ let (nal,d) = it_destRLambda_or_LetIn_names constagsl.(0) bl'.(0) in
GLetTuple (dl,nal,(alias,pred),tomatch,d)
- | IfStyle when aliastyp = None ->
+ | IfStyle, None ->
let bl' = Array.map detype bl in
let nondepbrs =
- array_map3 (extract_nondep_branches testdep) bl bl' consnargsl in
- if array_for_all ((<>) None) nondepbrs then
+ Array.map3 (extract_nondep_branches testdep) bl bl' constagsl in
+ if Array.for_all ((!=) None) nondepbrs then
GIf (dl,tomatch,(alias,pred),
Option.get nondepbrs.(0),Option.get nondepbrs.(1))
else
- let eqnl = detype_eqns constructs consnargsl bl in
+ let eqnl = detype_eqns constructs constagsl bl in
GCases (dl,tag,pred,[tomatch,(alias,aliastyp)],eqnl)
| _ ->
- let eqnl = detype_eqns constructs consnargsl bl in
+ let eqnl = detype_eqns constructs constagsl bl in
GCases (dl,tag,pred,[tomatch,(alias,aliastyp)],eqnl)
-let detype_sort = function
- | Prop c -> GProp c
- | Type u -> GType (Some u)
+let detype_sort sigma = function
+ | Prop Null -> GProp
+ | Prop Pos -> GSet
+ | Type u ->
+ GType
+ (if !print_universes
+ then [Pp.string_of_ppcmds (Univ.Universe.pr_with (Evd.pr_evd_level sigma) u)]
+ else [])
type binder_kind = BProd | BLambda | BLetIn
(**********************************************************************)
(* Main detyping function *)
-let detype_anonymous = ref (fun loc n -> anomaly "detype: 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 rec detype (isgoal:bool) avoid env t =
+let detype_level sigma l =
+ GType (Some (Pp.string_of_ppcmds (Evd.pr_evd_level sigma l)))
+
+let detype_instance sigma l =
+ 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
| Rel n ->
- (try match lookup_name_of_rel n env with
+ (try match lookup_name_of_rel n (fst env) with
| Name id -> GVar (dl, id)
| Anonymous -> !detype_anonymous dl n
with Not_found ->
let s = "_UNBOUND_REL_"^(string_of_int n)
- in GVar (dl, id_of_string s))
+ in GVar (dl, Id.of_string s))
| Meta n ->
(* Meta in constr are not user-parsable and are mapped to Evar *)
- GEvar (dl, n, None)
+ (* using numbers to be unparsable *)
+ GEvar (dl, Id.of_string ("M" ^ string_of_int n), [])
| Var id ->
- (try
- let _ = Global.lookup_named id in GRef (dl, VarRef id)
- with e when Errors.noncritical e ->
- GVar (dl, id))
- | Sort s -> GSort (dl,detype_sort s)
+ (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)
| Cast (c1,REVERTcast,c2) when not !Flags.raw_print ->
- detype isgoal avoid env c1
+ detype flags avoid env sigma c1
| Cast (c1,k,c2) ->
- GCast(dl,detype isgoal avoid env c1, CastConv (k, detype isgoal avoid env c2))
- | Prod (na,ty,c) -> detype_binder isgoal BProd avoid env na ty c
- | Lambda (na,ty,c) -> detype_binder isgoal BLambda avoid env na ty c
- | LetIn (na,b,_,c) -> detype_binder isgoal BLetIn avoid env na b c
+ let d1 = detype flags avoid env sigma c1 in
+ let d2 = detype 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
| App (f,args) ->
- GApp (dl,detype isgoal avoid env f,
- array_map_to_list (detype isgoal avoid env) args)
- | Const sp -> GRef (dl, ConstRef sp)
- | Evar (ev,cl) ->
- GEvar (dl, ev,
- Some (List.map (detype isgoal avoid env) (Array.to_list cl)))
- | Ind ind_sp ->
- GRef (dl, IndRef ind_sp)
- | Construct cstr_sp ->
- GRef (dl, ConstructRef cstr_sp)
+ let mkapp f' args' =
+ match f' with
+ | GApp (dl',f',args'') ->
+ GApp (dl,f',args''@args')
+ | _ -> GApp (dl,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)
+ | 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 args = List.make pars hole in
+ GApp (dl, GRef (dl, ConstRef (Projection.constant p), None),
+ (args @ [detype 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])
+ else
+ if Projection.unfolded p then
+ (** Print the compatibility match version *)
+ let c' =
+ try
+ let pb = Environ.lookup_projection p (snd env) in
+ let body = pb.Declarations.proj_body in
+ 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
+ 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'
+ else
+ if print_primproj_params () then
+ try
+ let c = Retyping.expand_projection (snd env) sigma p c [] in
+ detype flags avoid env sigma c
+ with Retyping.RetypeError _ -> noparams ()
+ else noparams ()
+
+ | Evar (evk,cl) ->
+ let bound_to_itself id c =
+ try let n = List.index Name.equal (Name id) (fst env) in
+ isRelN n c
+ with Not_found -> isVarId id c in
+ let id,l =
+ try
+ let id = Evd.evar_ident evk sigma in
+ let l = Evd.evar_instance_array bound_to_itself (Evd.find sigma evk) cl in
+ let fvs,rels = List.fold_left (fun (fvs,rels) (_,c) -> (Id.Set.union fvs (collect_vars c), Int.Set.union rels (free_rels c))) (Id.Set.empty,Int.Set.empty) l in
+ let l = Evd.evar_instance_array (fun id c -> not !print_evar_arguments && (bound_to_itself id c && not (isRel c && Int.Set.mem (destRel c) rels || isVar c && (Id.Set.mem (destVar c) fvs)))) (Evd.find sigma evk) cl in
+ id,l
+ with Not_found ->
+ Id.of_string ("X" ^ string_of_int (Evar.repr evk)),
+ (Array.map_to_list (fun c -> (Id.of_string "A",c)) cl)
+ in
+ GEvar (dl,id,
+ List.map (on_snd (detype flags avoid env sigma)) l)
+ | Ind (ind_sp,u) ->
+ GRef (dl, IndRef ind_sp, detype_instance sigma u)
+ | Construct (cstr_sp,u) ->
+ GRef (dl, ConstructRef cstr_sp, detype_instance sigma u)
| Case (ci,p,c,bl) ->
- let comp = computable p (ci.ci_pp_info.ind_nargs) in
- detype_case comp (detype isgoal avoid env)
- (detype_eqns isgoal avoid env ci comp)
+ 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
- (ci.ci_ind,ci.ci_pp_info.style,ci.ci_npar,
- ci.ci_cstr_ndecls,ci.ci_pp_info.ind_nargs)
+ (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 isgoal avoid env nvn recdef
- | CoFix (n,recdef) -> detype_cofix isgoal avoid env n recdef
+ | Fix (nvn,recdef) -> detype_fix flags avoid env sigma nvn recdef
+ | CoFix (n,recdef) -> detype_cofix flags avoid env sigma n recdef
-and detype_fix isgoal avoid env (vn,_ as nvn) (names,tys,bodies) =
+and detype_fix flags avoid env sigma (vn,_ as nvn) (names,tys,bodies) =
let def_avoid, def_env, lfi =
- Array.fold_left
- (fun (avoid, env, l) na ->
+ Array.fold_left2
+ (fun (avoid, env, l) na ty ->
let id = next_name_away na avoid in
- (id::avoid, add_name (Name id) env, id::l))
- (avoid, env, []) names in
+ (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 isgoal (i+1) [] def_avoid def_env c (lift n t))
+ let v = Array.map3
+ (fun c t i -> share_names 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),
Array.map (fun (bl,_,_) -> bl) v,
Array.map (fun (_,_,ty) -> ty) v,
Array.map (fun (_,bd,_) -> bd) v)
-and detype_cofix isgoal avoid env n (names,tys,bodies) =
+and detype_cofix flags avoid env sigma n (names,tys,bodies) =
let def_avoid, def_env, lfi =
- Array.fold_left
- (fun (avoid, env, l) na ->
+ Array.fold_left2
+ (fun (avoid, env, l) na ty ->
let id = next_name_away na avoid in
- (id::avoid, add_name (Name id) env, id::l))
- (avoid, env, []) names in
+ (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 isgoal 0 [] def_avoid def_env c (lift ntys t))
+ let v = Array.map2
+ (fun c t -> share_names flags 0 [] def_avoid def_env sigma c (lift ntys t))
bodies tys in
GRec(dl,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 isgoal n l avoid env c t =
+and share_names flags n l avoid env sigma c t =
match kind_of_term c, kind_of_term t with
(* factorize even when not necessary to have better presentation *)
| Lambda (na,t,c), Prod (na',t',c') ->
@@ -458,93 +570,98 @@ and share_names isgoal n l avoid env c t =
Name _, _ -> na
| _, Name _ -> na'
| _ -> na in
- let t = detype isgoal avoid env t in
+ let t' = detype flags avoid env sigma t in
let id = next_name_away na avoid in
- let avoid = id::avoid and env = add_name (Name id) env in
- share_names isgoal (n-1) ((Name id,Explicit,None,t)::l) avoid env c c'
+ 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'
(* May occur for fix built interactively *)
| LetIn (na,b,t',c), _ when n > 0 ->
- let t' = detype isgoal avoid env t' in
- let b = detype isgoal avoid env b in
+ let t'' = detype flags avoid env sigma t' in
+ let b' = detype flags avoid env sigma b in
let id = next_name_away na avoid in
- let avoid = id::avoid and env = add_name (Name id) env in
- share_names isgoal n ((Name id,Explicit,Some b,t')::l) avoid env c (lift 1 t)
+ 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)
(* Only if built with the f/n notation or w/o let-expansion in types *)
| _, LetIn (_,b,_,t) when n > 0 ->
- share_names isgoal n l avoid env c (subst1 b t)
+ share_names 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 isgoal avoid env t' in
+ let t'' = detype flags avoid env sigma t' in
let id = next_name_away na' avoid in
- let avoid = id::avoid and env = add_name (Name id) env in
+ let avoid = id::avoid and env = add_name (Name id) None t' env in
let appc = mkApp (lift 1 c,[|mkRel 1|]) in
- share_names isgoal (n-1) ((Name id,Explicit,None,t')::l) avoid env appc c'
+ share_names 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 msg_warn "Detyping.detype: cannot factorize fix enough";
- let c = detype isgoal avoid env c in
- let t = detype isgoal avoid env t in
+ if n>0 then msg_warning (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
(List.rev l,c,t)
-and detype_eqns isgoal avoid env ci computable constructs consnargsl bl =
+and detype_eqns flags avoid env sigma ci computable constructs consnargsl bl =
try
- if !Flags.raw_print or not (reverse_matching ()) then raise Exit;
- let mat = build_tree Anonymous isgoal (avoid,env) ci bl in
- List.map (fun (pat,((avoid,env),c)) -> (dl,[],[pat],detype isgoal avoid env c))
+ 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))
mat
with e when Errors.noncritical e ->
Array.to_list
- (array_map3 (detype_eqn isgoal avoid env) constructs consnargsl bl)
+ (Array.map3 (detype_eqn flags avoid env sigma) constructs consnargsl bl)
-and detype_eqn isgoal avoid env constr construct_nargs branch =
- let make_pat x avoid env b ids =
- if force_wildcard () & noccurn 1 b then
- PatVar (dl,Anonymous),avoid,(add_name Anonymous env),ids
+and detype_eqn (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
else
- let id = next_name_away_in_cases_pattern x avoid in
- PatVar (dl,Name id),id::avoid,(add_name (Name id) env),id::ids
+ 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
in
- let rec buildrec ids patlist avoid env n b =
- if n=0 then
- (dl, ids,
- [PatCstr(dl, constr, List.rev patlist,Anonymous)],
- detype isgoal avoid env b)
- else
- match kind_of_term b with
- | Lambda (x,_,b) ->
- let pat,new_avoid,new_env,new_ids = make_pat x avoid env b ids in
- buildrec new_ids (pat::patlist) new_avoid new_env (n-1) b
-
- | LetIn (x,_,_,b) ->
- let pat,new_avoid,new_env,new_ids = make_pat x avoid env b ids in
- buildrec new_ids (pat::patlist) new_avoid new_env (n-1) b
-
- | Cast (c,_,_) -> (* Oui, il y a parfois des cast *)
- buildrec ids patlist avoid env n c
-
- | _ -> (* eta-expansion : n'arrivera plus lorsque tous les
- termes seront construits à partir de la syntaxe Cases *)
+ 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)
+ | 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
+
+ | LetIn (x,b,t,b'), true::l ->
+ let pat,new_avoid,new_env,new_ids = make_pat x avoid env b' (Some b) t ids in
+ buildrec new_ids (pat::patlist) new_avoid new_env l b'
+
+ | Cast (c,_,_), l -> (* Oui, il y a parfois des cast *)
+ buildrec ids patlist avoid env l c
+
+ | _, true::l ->
+ let pat = PatVar (dl,Anonymous) in
+ buildrec ids (pat::patlist) avoid env l b
+
+ | _, false::l ->
+ (* eta-expansion : n'arrivera plus lorsque tous les
+ termes seront construits à partir de la syntaxe Cases *)
(* nommage de la nouvelle variable *)
let new_b = applist (lift 1 b, [mkRel 1]) in
let pat,new_avoid,new_env,new_ids =
- make_pat Anonymous avoid env new_b ids in
- buildrec new_ids (pat::patlist) new_avoid new_env (n-1) new_b
+ make_pat Anonymous avoid env new_b None mkProp ids in
+ buildrec new_ids (pat::patlist) new_avoid new_env l new_b
in
- buildrec [] [] avoid env construct_nargs branch
-
-and detype_binder isgoal bk avoid env na ty c =
- let flag = if isgoal then RenamingForGoal else RenamingElsewhereFor (env,c) in
- let na',avoid' =
- if bk = BLetIn then compute_displayed_let_name_in flag avoid na c
- else compute_displayed_name_in flag avoid na c in
- let r = detype isgoal avoid' (add_name na' env) c 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 =
+ 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
match bk with
- | BProd -> GProd (dl, na',Explicit,detype false avoid env ty, r)
- | BLambda -> GLambda (dl, na',Explicit,detype false avoid env ty, r)
- | BLetIn -> GLetIn (dl, na',detype false avoid env ty, r)
+ | 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)
+ | BLetIn -> GLetIn (dl, na',detype (lax,false) avoid env sigma (Option.get body), r)
-let rec detype_rel_context where avoid env sign =
+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 rec aux avoid env = function
| [] -> []
@@ -553,17 +670,80 @@ let rec detype_rel_context where avoid env sign =
match where with
| None -> na,avoid
| Some c ->
- if b<>None then
+ if b != None then
compute_displayed_let_name_in
- (RenamingElsewhereFor (env,c)) avoid na c
+ (RenamingElsewhereFor (fst env,c)) avoid na c
else
compute_displayed_name_in
- (RenamingElsewhereFor (env,c)) avoid na c in
- let b = Option.map (detype false avoid env) b in
- let t = detype false avoid env t in
- (na',Explicit,b,t) :: aux avoid' (add_name na' env) rest
+ (RenamingElsewhereFor (fst env,c)) avoid na c in
+ let b' = Option.map (detype (lax,false) avoid env sigma) b in
+ let t' = detype (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
+
+let detype_closed_glob ?lax isgoal avoid env sigma t =
+ let convert_id cl id =
+ try Id.Map.find id cl.idents
+ with Not_found -> id
+ in
+ let convert_name cl = function
+ | Name id -> Name (convert_id cl id)
+ | Anonymous -> Anonymous
+ in
+ let rec detype_closed_glob cl = function
+ | GVar (loc,id) ->
+ (* if [id] is bound to a name. *)
+ begin try
+ GVar(loc,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 *)
+ let (b,c) = Id.Map.find id cl.typed in
+ (* 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
+ (* 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
+ (* Otherwise [id] stands for itself *)
+ with Not_found ->
+ GVar(loc,id)
+ end
+ | GLambda (loc,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) ->
+ 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) ->
+ 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) ->
+ 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) ->
+ 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)
+ | c ->
+ Glob_ops.map_glob_constr (detype_closed_glob cl) c
+ in
+ detype_closed_glob t.closure t.term
+
(**********************************************************************)
(* Module substitution: relies on detyping *)
@@ -571,17 +751,19 @@ let rec subst_cases_pattern subst pat =
match pat with
| PatVar _ -> pat
| PatCstr (loc,((kn,i),j),cpl,n) ->
- let kn' = subst_ind subst kn
- and cpl' = list_smartmap (subst_cases_pattern subst) cpl in
+ 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)
+let (f_subst_genarg, subst_genarg_hook) = Hook.make ()
+
let rec subst_glob_constr subst raw =
match raw with
- | GRef (loc,ref) ->
+ | GRef (loc,ref,u) ->
let ref',t = subst_global subst ref in
if ref' == ref then raw else
- detype false [] [] t
+ detype false [] (Global.env()) Evd.empty t
| GVar _ -> raw
| GEvar _ -> raw
@@ -589,7 +771,7 @@ let rec subst_glob_constr subst raw =
| GApp (loc,r,rl) ->
let r' = subst_glob_constr subst r
- and rl' = list_smartmap (subst_glob_constr subst) rl in
+ and rl' = List.smartmap (subst_glob_constr subst) rl in
if r' == r && rl' == rl then raw else
GApp(loc,r',rl')
@@ -610,18 +792,18 @@ let rec subst_glob_constr subst raw =
| GCases (loc,sty,rtno,rl,branches) ->
let rtno' = Option.smartmap (subst_glob_constr subst) rtno
- and rl' = list_smartmap (fun (a,x as y) ->
+ 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),x,y as t) ->
- let sp' = subst_ind subst sp in
- if sp == sp' then t else (loc,(sp',i),x,y)) topt in
+ (fun (loc,(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 a == a' && topt == topt' then y else (a',(n,topt'))) rl
- and branches' = list_smartmap
+ and branches' = List.smartmap
(fun (loc,idl,cpl,r as branch) ->
let cpl' =
- list_smartmap (subst_cases_pattern subst) 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'))
@@ -642,52 +824,51 @@ let rec subst_glob_constr subst raw =
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
+ if c' == c && po' == po && b1' == b1 && b2' == b2 then raw else
GIf (loc,c',(na,po'),b1',b2')
| GRec (loc,fix,ida,bl,ra1,ra2) ->
- let ra1' = array_smartmap (subst_glob_constr subst) ra1
- and ra2' = array_smartmap (subst_glob_constr subst) ra2 in
- let bl' = array_smartmap
- (list_smartmap (fun (na,k,obd,ty as dcl) ->
+ let ra1' = Array.smartmap (subst_glob_constr subst) ra1
+ and ra2' = Array.smartmap (subst_glob_constr subst) ra2 in
+ let bl' = Array.smartmap
+ (List.smartmap (fun (na,k,obd,ty as dcl) ->
let ty' = subst_glob_constr subst ty in
let obd' = Option.smartmap (subst_glob_constr subst) obd in
- if ty'==ty & obd'==obd then dcl else (na,k,obd',ty')))
+ 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
- | GHole (loc,ImplicitArg (ref,i,b)) ->
- let ref',_ = subst_global subst ref in
- if ref' == ref then raw else
- GHole (loc,InternalHole)
- | GHole (loc, (BinderType _ | QuestionMark _ | CasesType | InternalHole |
- TomatchTypeParameter _ | GoalEvar | ImpossibleCase | MatchingVar _)) ->
- raw
+ | GHole (loc, knd, naming, solve) ->
+ let nknd = match knd with
+ | Evar_kinds.ImplicitArg (ref, i, b) ->
+ let nref, _ = subst_global subst ref in
+ if nref == ref then knd else Evar_kinds.ImplicitArg (nref, i, b)
+ | _ -> knd
+ 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)
| GCast (loc,r1,k) ->
- (match k with
- CastConv (k,r2) ->
- let r1' = subst_glob_constr subst r1 and r2' = subst_glob_constr subst r2 in
- if r1' == r1 && r2' == r2 then raw else
- GCast (loc,r1', CastConv (k,r2'))
- | CastCoerce ->
- let r1' = subst_glob_constr subst r1 in
- if r1' == r1 then raw else GCast (loc,r1',k))
+ 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')
(* 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 (dummy_loc,na) in
- let p = PatCstr (dummy_loc,(ind,i+1),List.map mkPatVar nal,Anonymous) in
- let ids = map_succeed Nameops.out_name nal in
- (dummy_loc,ids,[p],c))
+ 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))
brs
-let return_type_of_predicate ind nparams nrealargs_ctxt pred =
- let nal,p = it_destRLambda_or_LetIn_names (nrealargs_ctxt+1) pred in
- (List.hd nal, Some (dummy_loc, ind, nparams, List.tl nal)), Some p
+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
diff --git a/pretyping/detyping.mli b/pretyping/detyping.mli
index bbd94cfe..eb158686 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-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Util
open Names
open Term
-open Sign
+open Context
open Environ
open Glob_term
open Termops
open Mod_subst
+open Misctypes
+open Evd
+
+(** Should we keep details of universes during detyping ? *)
+val print_universes : bool ref
+
+(** If true, prints full local context of evars *)
+val print_evar_arguments : bool ref
val subst_cases_pattern : substitution -> cases_pattern -> cases_pattern
@@ -24,36 +31,43 @@ val subst_glob_constr : substitution -> glob_constr -> glob_constr
[isgoal] tells if naming must avoid global-level synonyms as intro does
[ctx] gives the names of the free variables *)
-val detype : bool -> identifier list -> names_context -> constr -> glob_constr
+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_case :
- bool -> ('a -> glob_constr) ->
- (constructor array -> int array -> 'a array ->
- (loc * identifier list * cases_pattern list * glob_constr) list) ->
- ('a -> int -> bool) ->
- identifier list -> inductive * case_style * int * int array * int ->
- 'a option -> 'a -> 'a array -> glob_constr
+ 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_sort : sorts -> glob_sort
+val detype_sort : evar_map -> sorts -> glob_sort
-val detype_rel_context : constr option -> identifier list -> names_context ->
- rel_context -> glob_decl list
+val detype_rel_context : ?lax:bool -> constr option -> Id.t list -> (names_context * env) ->
+ evar_map -> rel_context -> glob_decl list
+
+val detype_closed_glob : ?lax:bool -> bool -> Id.t list -> 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 -> identifier -> int option
+val lookup_name_as_displayed : env -> constr -> Id.t -> int option
val lookup_index_as_renamed : env -> constr -> int -> int option
-val set_detype_anonymous : (loc -> int -> glob_constr) -> unit
+val set_detype_anonymous : (Loc.t -> int -> glob_constr) -> unit
val force_wildcard : unit -> bool
val synthetize_type : unit -> bool
(** Utilities to transform kernel cases to simple pattern-matching problem *)
-val it_destRLambda_or_LetIn_names : int -> glob_constr -> name list * glob_constr
+val it_destRLambda_or_LetIn_names : bool list -> glob_constr -> Name.t list * glob_constr
val simple_cases_matrix_of_branches :
- inductive -> (int * int * glob_constr) list -> cases_clauses
+ inductive -> (int * bool list * glob_constr) list -> cases_clauses
val return_type_of_predicate :
- inductive -> int -> int -> glob_constr -> predicate_pattern * glob_constr option
+ inductive -> bool list -> glob_constr -> predicate_pattern * glob_constr option
+
+val subst_genarg_hook :
+ (substitution -> Genarg.glob_generic_argument -> Genarg.glob_generic_argument) Hook.t
module PrintingInductiveMake :
functor (Test : sig
@@ -64,6 +78,7 @@ module PrintingInductiveMake :
end) ->
sig
type t = Names.inductive
+ val compare : t -> t -> int
val encode : Libnames.reference -> Names.inductive
val subst : substitution -> t -> t
val printer : t -> Pp.std_ppcmds
diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml
index 8b421ea3..a95af253 100644
--- a/pretyping/evarconv.ml
+++ b/pretyping/evarconv.ml
@@ -1,15 +1,16 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Pp
+open Errors
open Util
open Names
open Term
+open Vars
open Closure
open Reduction
open Reductionops
@@ -17,41 +18,40 @@ open Termops
open Environ
open Recordops
open Evarutil
-open Libnames
+open Evarsolve
+open Globnames
open Evd
+open Pretype_errors
+
+type unify_fun = transparent_state ->
+ env -> evar_map -> conv_pb -> constr -> constr -> Evarsolve.unification_result
let debug_unification = ref (false)
let _ = Goptions.declare_bool_option {
Goptions.optsync = true; Goptions.optdepr = false;
Goptions.optname =
- "Print states sended to Evarconv unification";
+ "Print states sent to Evarconv unification";
Goptions.optkey = ["Debug";"Unification"];
Goptions.optread = (fun () -> !debug_unification);
Goptions.optwrite = (fun a -> debug_unification:=a);
}
-
-type flex_kind_of_term =
- | Rigid of constr
- | PseudoRigid of constr (* approximated as rigid but not necessarily so *)
- | MaybeFlexible of constr (* approx'ed as reducible but not necessarily so *)
- | Flexible of existential
-
-let flex_kind_of_term c l =
+let unfold_projection env evd ts p c =
+ let cst = Projection.constant p in
+ if is_transparent_constant ts cst then
+ let c' = Some (mkProj (Projection.make cst true, c)) in
+ match ReductionBehaviour.get (Globnames.ConstRef cst) with
+ | None -> c'
+ | Some (recargs, nargs, flags) ->
+ if (List.mem `ReductionNeverUnfold flags) then None
+ else c'
+ else None
+
+let eval_flexible_term ts env evd c =
match kind_of_term c with
- | Rel _ | Const _ | Var _ -> MaybeFlexible c
- | Lambda _ when l<>[] -> MaybeFlexible c
- | LetIn _ -> MaybeFlexible c
- | Evar ev -> Flexible ev
- | Lambda _ | Prod _ | Sort _ | Ind _ | Construct _ | CoFix _ -> Rigid c
- | Meta _ | Case _ | Fix _ -> PseudoRigid c
- | Cast _ | App _ -> assert false
-
-let eval_flexible_term ts env c =
- match kind_of_term c with
- | Const c ->
+ | Const (c,u as cu) ->
if is_transparent_constant ts c
- then constant_opt_value env c
+ then constant_opt_value_in env cu
else None
| Rel n ->
(try let (_,v,_) = lookup_rel n env in Option.map (lift n) v
@@ -59,81 +59,126 @@ let eval_flexible_term ts env c =
| Var id ->
(try
if is_transparent_variable ts id then
- let (_,v,_) = lookup_named id env in v
- else None
+ let (_,v,_) = lookup_named id env in v
+ else None
with Not_found -> None)
| LetIn (_,b,_,c) -> Some (subst1 b c)
| Lambda _ -> Some c
+ | Proj (p, c) ->
+ if Projection.unfolded p then assert false
+ else unfold_projection env evd ts p c
| _ -> assert false
-let evar_apprec ts env evd stack c =
- let sigma = evd in
- let rec aux s =
- let (t,stack) = whd_betaiota_deltazeta_for_iota_state ts env sigma s in
- match kind_of_term t with
- | Evar (evk,_ as ev) when Evd.is_defined sigma evk ->
- aux (Evd.existential_value sigma ev, stack)
- | _ -> (t, list_of_stack stack)
- in aux (c, append_stack_list stack empty_stack)
+type flex_kind_of_term =
+ | Rigid
+ | MaybeFlexible of Constr.t (* reducible but not necessarily reduced *)
+ | Flexible of existential
+
+let flex_kind_of_term ts env evd c sk =
+ match kind_of_term 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
+ | Evar ev -> Flexible ev
+ | Lambda _ | Prod _ | Sort _ | Ind _ | Construct _ | CoFix _ -> Rigid
+ | Meta _ -> Rigid
+ | Fix _ -> Rigid (* happens when the fixpoint is partially applied *)
+ | Cast _ | App _ | Case _ -> assert false
let apprec_nohdbeta ts env evd c =
- match kind_of_term (fst (Reductionops.whd_stack evd c)) with
- | (Case _ | Fix _) -> applist (evar_apprec ts env evd [] c)
- | _ -> 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
+ ts env evd Cst_stack.empty appr))
+ else c
let position_problem l2r = function
| CONV -> None
| CUMUL -> Some l2r
-(* [check_conv_record (t1,l1) (t2,l2)] tries to decompose the problem
- (t1 l1) = (t2 l2) into a problem
-
- l1 = params1@c1::extra_args1
- l2 = us2@extra_args2
- (t1 params1 c1) = (proji params (c xs))
- (t2 us2) = (cstr us)
+let occur_rigidly ev evd t =
+ let (l, app) = decompose_app_vect t in
+ let rec aux t =
+ match kind_of_term (whd_evar 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)
+ | Evar (ev',_) -> if Evar.equal ev ev' then raise Occur else false
+ | Cast (p, _, _) -> aux p
+ | Lambda _ | LetIn _ -> false
+ | Const _ -> false
+ | Prod (_, b, t) -> ignore(aux b || aux t); true
+ | Rel _ | Var _ -> false
+ | Case _ -> false
+ in Array.exists (fun t -> try ignore(aux t); false with Occur -> true) app
+
+(* [check_conv_record env sigma (t1,stack1) (t2,stack2)] tries to decompose
+ the problem (t1 stack1) = (t2 stack2) into a problem
+
+ stack1 = params1@[c1]@extra_args1
+ stack2 = us2@extra_args2
+ t1 params1 c1 = proji params (c xs)
+ t2 us2 = head us
extra_args1 = extra_args2
by finding a record R and an object c := [xs:bs](Build_R params v1..vn)
- with vi = (cstr us), for which we know that the i-th projection proji
+ with vi = (head us), for which we know that the i-th projection proji
satisfies
- (proji params (c xs)) = (cstr us)
+ proji params (c xs) = head us
Rem: such objects, usable for conversion, are defined in the objdef
table; practically, it amounts to "canonically" equip t2 into a
object c in structure R (since, if c1 were not an evar, the
projection would have been reduced) *)
-let check_conv_record (t1,l1) (t2,l2) =
- try
- let proji = global_of_constr t1 in
- let canon_s,l2_effective =
- try
- match kind_of_term t2 with
- Prod (_,a,b) -> (* assert (l2=[]); *)
- if dependent (mkRel 1) b then raise Not_found
- else lookup_canonical_conversion (proji, Prod_cs),[a;pop b]
- | Sort s ->
- lookup_canonical_conversion
- (proji, Sort_cs (family_of_sort s)),[]
- | _ ->
- let c2 = global_of_constr t2 in
- lookup_canonical_conversion (proji, Const_cs c2),l2
- with Not_found ->
- lookup_canonical_conversion (proji,Default_cs),[]
- in
- let { o_DEF = c; o_INJ=n; o_TABS = bs;
- o_TPARAMS = params; o_NPARAMS = nparams; o_TCOMPS = us } = canon_s in
- let params1, c1, extra_args1 =
- match list_chop nparams l1 with
- | params1, c1::extra_args1 -> params1, c1, extra_args1
- | _ -> raise Not_found in
- let us2,extra_args2 = list_chop (List.length us) l2_effective in
- c,bs,(params,params1),(us,us2),(extra_args1,extra_args2),c1,
- (n,applist(t2,l2))
- with Failure _ | Not_found ->
- raise Not_found
+let check_conv_record env sigma (t1,sk1) (t2,sk2) =
+ let (proji, u), arg = Universes.global_app_of_constr t1 in
+ let canon_s,sk2_effective =
+ try
+ match kind_of_term t2 with
+ Prod (_,a,b) -> (* assert (l2=[]); *)
+ if dependent (mkRel 1) b then raise Not_found
+ else lookup_canonical_conversion (proji, Prod_cs),
+ (Stack.append_app [|a;pop b|] Stack.empty)
+ | Sort s ->
+ lookup_canonical_conversion
+ (proji, Sort_cs (family_of_sort s)),[]
+ | _ ->
+ let c2 = global_of_constr t2 in
+ lookup_canonical_conversion (proji, Const_cs c2),sk2
+ with Not_found ->
+ let (c, cs) = lookup_canonical_conversion (proji,Default_cs) in
+ (c,cs),[]
+ 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 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
+ with _ -> raise Not_found
+ in Stack.append_app_list ind_args Stack.empty, c, sk1
+ | None ->
+ match Stack.strip_n_app nparams sk1 with
+ | Some (params1, c1, extra_args1) -> params1, c1, extra_args1
+ | _ -> raise Not_found in
+ let us2,extra_args2 =
+ let l_us = List.length us in
+ if Int.equal l_us 0 then Stack.empty,sk2_effective
+ 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 c' = subst_univs_level_constr subst c 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
+ 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))
(* Precondition: one of the terms of the pb is an uninstantiated evar,
* possibly applied to arguments. *)
@@ -142,40 +187,135 @@ let rec ise_try evd = function
[] -> assert false
| [f] -> f evd
| f1::l ->
- let (evd',b) = f1 evd in
- if b then (evd',b) else ise_try evd l
+ match f1 evd with
+ | Success _ as x -> x
+ | UnifFailure _ -> ise_try evd l
let ise_and evd l =
let rec ise_and i = function
[] -> assert false
| [f] -> f i
| f1::l ->
- let (i',b) = f1 i in
- if b then ise_and i' l else (evd,false) in
+ match f1 i with
+ | Success i' -> ise_and i' l
+ | UnifFailure _ as x -> x in
ise_and evd l
-let ise_list2 evd f l1 l2 =
- let rec ise_list2 i l1 l2 =
- match l1,l2 with
- [], [] -> (i, true)
- | [x], [y] -> f i x y
- | x::l1, y::l2 ->
- let (i',b) = f i x y in
- if b then ise_list2 i' l1 l2 else (evd,false)
- | _ -> (evd, false) in
- ise_list2 evd l1 l2
+(* This function requires to get the outermost arguments first. It is
+ a fold_right for backward compatibility.
+
+ It tries to unify the suffix of 2 lists element by element and if
+ it reaches the end of a list, it returns the remaining elements in
+ the other list if there are some.
+*)
+let ise_exact ise x1 x2 =
+ match ise x1 x2 with
+ | None, out -> out
+ | _, (UnifFailure _ as out) -> out
+ | Some _, Success i -> UnifFailure (i,NotSameArgSize)
let ise_array2 evd f v1 v2 =
let rec allrec i = function
- | -1 -> (i,true)
+ | -1 -> Success i
| n ->
- let (i',b) = f i v1.(n) v2.(n) in
- if b then allrec i' (n-1) else (evd,false)
- in
+ match f i v1.(n) v2.(n) with
+ | Success i' -> allrec i' (n-1)
+ | UnifFailure _ as x -> x in
let lv1 = Array.length v1 in
- if lv1 = Array.length v2 then allrec evd (pred lv1)
- else (evd,false)
-
+ if Int.equal lv1 (Array.length v2) then allrec evd (pred lv1)
+ else UnifFailure (evd,NotSameArgSize)
+
+(* Applicative node of stack are read from the outermost to the innermost
+ but are unified the other way. *)
+let rec ise_app_stack2 env f evd sk1 sk2 =
+ match sk1,sk2 with
+ | Stack.App node1 :: q1, Stack.App node2 :: q2 ->
+ let (t1,l1) = Stack.decomp_node_last node1 q1 in
+ let (t2,l2) = Stack.decomp_node_last node2 q2 in
+ begin match ise_app_stack2 env f evd l1 l2 with
+ |(_,UnifFailure _) as x -> x
+ |x,Success i' -> x,f env i' CONV t1 t2
+ end
+ | _, _ -> (sk1,sk2), Success evd
+
+(* This function tries to unify 2 stacks element by element. It works
+ from the end to the beginning. If it unifies a non empty suffix of
+ stacks but not the entire stacks, the first part of the answer is
+ Some(the remaining prefixes to tackle)) *)
+let ise_stack2 no_app env evd f sk1 sk2 =
+ let rec ise_stack2 deep i sk1 sk2 =
+ let fail x = if deep then Some (List.rev sk1, List.rev sk2), Success i
+ else None, x in
+ match sk1, sk2 with
+ | [], [] -> None, Success i
+ | Stack.Case (_,t1,c1,_)::q1, Stack.Case (_,t2,c2,_)::q2 ->
+ (match f env i CONV t1 t2 with
+ | Success i' ->
+ (match ise_array2 i' (fun ii -> f env ii CONV) c1 c2 with
+ | Success i'' -> ise_stack2 true i'' q1 q2
+ | 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)
+ then ise_stack2 true i q1 q2
+ else fail (UnifFailure (i, NotSameHead))
+ | Stack.Fix (((li1, i1),(_,tys1,bds1 as recdef1)),a1,_)::q1,
+ Stack.Fix (((li2, i2),(_,tys2,bds2)),a2,_)::q2 ->
+ if Int.equal i1 i2 && Array.equal Int.equal li1 li2 then
+ match ise_and i [
+ (fun i -> ise_array2 i (fun ii -> f env ii CONV) tys1 tys2);
+ (fun i -> ise_array2 i (fun ii -> f (push_rec_types recdef1 env) ii CONV) bds1 bds2);
+ (fun i -> ise_exact (ise_stack2 false i) a1 a2)] with
+ | 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
+ |_,(UnifFailure _ as x) -> fail x
+ |(l1, l2), Success i' -> ise_stack2 true i' l1 l2
+ end
+ |_, _ -> fail (UnifFailure (i,(* Maybe improve: *) NotSameHead))
+ in ise_stack2 false evd (List.rev sk1) (List.rev sk2)
+
+(* Make sure that the matching suffix is the all stack *)
+let exact_ise_stack2 env evd f sk1 sk2 =
+ let rec ise_stack2 i sk1 sk2 =
+ match sk1, sk2 with
+ | [], [] -> Success i
+ | Stack.Case (_,t1,c1,_)::q1, Stack.Case (_,t2,c2,_)::q2 ->
+ ise_and i [
+ (fun i -> ise_stack2 i q1 q2);
+ (fun i -> ise_array2 i (fun ii -> f env ii CONV) c1 c2);
+ (fun i -> f env i CONV t1 t2)]
+ | Stack.Fix (((li1, i1),(_,tys1,bds1 as recdef1)),a1,_)::q1,
+ Stack.Fix (((li2, i2),(_,tys2,bds2)),a2,_)::q2 ->
+ if Int.equal i1 i2 && Array.equal Int.equal li1 li2 then
+ ise_and i [
+ (fun i -> ise_stack2 i q1 q2);
+ (fun i -> ise_array2 i (fun ii -> f env ii CONV) tys1 tys2);
+ (fun i -> ise_array2 i (fun ii -> f (push_rec_types recdef1 env) ii CONV) bds1 bds2);
+ (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)
+ 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
+ |(l1, l2), Success i' -> ise_stack2 i' l1 l2
+ end
+ |_, _ -> UnifFailure (i,(* Maybe improve: *) NotSameHead)
+ in
+ if Reductionops.Stack.compare_shape sk1 sk2 then
+ ise_stack2 evd (List.rev sk1) (List.rev sk2)
+ else UnifFailure (evd, (* Dummy *) NotSameHead)
+
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
@@ -183,426 +323,578 @@ let rec evar_conv_x ts env evd pbty term1 term2 =
could have found, we do it only if the terms are free of evar.
Note: incomplete heuristic... *)
let ground_test =
- if is_ground_term evd term1 && is_ground_term evd term2 then
- if is_trans_fconv pbty ts env evd term1 term2 then
- Some true
- else if is_ground_env evd env then Some false
- else None
- else None in
+ if is_ground_term evd term1 && is_ground_term evd term2 then (
+ let evd, b =
+ try infer_conv ~pb:pbty ~ts:(fst ts) env evd term1 term2
+ with Univ.UniverseInconsistency _ -> evd, false
+ in
+ if b then Some (evd, true)
+ else if is_ground_env evd env then Some (evd, false)
+ else None)
+ else None
+ in
match ground_test with
- Some b -> (evd,b)
+ | Some (evd, true) -> Success evd
+ | Some (evd, false) -> UnifFailure (evd,ConversionFailed (env,term1,term2))
| None ->
(* 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 ts env evd term1 in
- let term2 = apprec_nohdbeta ts env evd term2 in
- if is_undefined_evar evd term1 then
- solve_simple_eqn (evar_conv_x ts) env evd
- (position_problem true pbty,destEvar term1,term2)
- else if is_undefined_evar evd term2 then
- solve_simple_eqn (evar_conv_x ts) env evd
- (position_problem false pbty,destEvar term2,term1)
- else
+ let term1 = apprec_nohdbeta (fst ts) env evd term1 in
+ let term2 = apprec_nohdbeta (fst ts) env evd term2 in
+ let default () =
evar_eqappr_x ts env evd pbty
- (decompose_app term1) (decompose_app term2)
-
-and evar_eqappr_x ?(rhs_is_already_stuck = false)
- ts env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
-
- let eta env evd onleft term l term' l' =
- assert (l = []);
- let (na,c,body) = destLambda term in
- let c = nf_evar evd c in
+ (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
+ | 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
+ | 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
+ | UnifFailure (_, OccurCheck _) ->
+ (* Eta-expansion might apply *) default ()
+ | x -> x)
+ | _ -> default ()
+ end
+
+and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
+ ((term1,sk1 as appr1),csts1) ((term2,sk2 as appr2),csts2) =
+ let default_fail i = (* costly *)
+ UnifFailure (i,ConversionFailed (env, Stack.zip appr1, Stack.zip appr2)) in
+ let quick_fail i = (* not costly, loses info *)
+ UnifFailure (i, NotSameHead)
+ in
+ let miller_pfenning on_left fallback ev lF tM evd =
+ 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
+ solve_simple_eqn (evar_conv_x ts) env evd
+ (position_problem on_left pbty,ev,t2)
+ in
+ let consume_stack on_left (termF,skF) (termO,skO) evd =
+ let switch f a b = if on_left then f a b else f b a in
+ 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))
+ |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))
+ |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 c = nf_evar evd c1 in
let env' = push_rel (na,None,c) env in
- let appr1 = evar_apprec ts env' evd [] body in
- let appr2 = (lift 1 term', List.map (lift 1) l' @ [mkRel 1]) in
- if onleft then evar_eqappr_x ts env' evd CONV appr1 appr2
- else evar_eqappr_x ts env' evd CONV appr2 appr1
+ 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),
+ 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)
+ 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
+ let not_only_app = Stack.not_purely_applicative skM in
+ let f1 i =
+ match Stack.list_of_app_stack skF with
+ | None -> default_fail evd
+ | Some lF ->
+ let tM = Stack.zip 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
+ else quick_fail i)
+ ev lF tM i
+ and consume (termF,skF as apprF) (termM,skM as apprM) i =
+ if not (Stack.is_empty skF && Stack.is_empty skM) then
+ consume_stack on_left apprF apprM i
+ else quick_fail i
+ and delta i =
+ switch (evar_eqappr_x ts env i pbty) (apprF,cstsF)
+ (whd_betaiota_deltazeta_for_iota_state (fst ts) env i cstsM (vM,skM))
+ in
+ let default i = ise_try i [f1; consume apprF apprM; delta]
+ in
+ match kind_of_term 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. *)
+ let nargsF = Stack.args_size skF and nargsM = Stack.args_size skM in
+ begin
+ (* ?X argsF' ~= (p.c ..) argsM' -> ?X ~= (p.c ..), no need to expand *)
+ if nargsF <= nargsM then default evd
+ else
+ let f =
+ try
+ let termM' = Retyping.expand_projection env evd p c [] in
+ let apprM', cstsM' =
+ whd_betaiota_deltazeta_for_iota_state (fst ts) env evd cstsM (termM',skM)
+ in
+ let delta' i =
+ switch (evar_eqappr_x ts env i pbty) (apprF,cstsF) (apprM',cstsM')
+ in
+ fun i -> ise_try i [f1; consume apprF apprM'; delta']
+ with Retyping.RetypeError _ ->
+ (* Happens thanks to w_unify building ill-typed terms *)
+ default
+ in f evd
+ end
+ | _ -> default evd
+ in
+ 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
+ | Lambda _ -> eta env evd false skR termR skF termF
+ | Construct u -> eta_constructor ts env evd skR u skF termF
+ | _ -> UnifFailure (evd,NotSameHead)
+ in
+ match Stack.list_of_app_stack skF with
+ | None ->
+ ise_try evd [consume_stack on_left apprF apprR; eta]
+ | Some lF ->
+ let tR = Stack.zip apprR in
+ miller_pfenning on_left
+ (fun () ->
+ ise_try evd
+ [eta;(* Postpone the use of an heuristic *)
+ (fun i ->
+ if not (occur_rigidly (fst ev) i tR) then
+ let i,tF =
+ if isRel tR || isVar 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
+ switch (fun x y -> Success (add_conv_pb (pbty,env,x,y) i))
+ tF tR
+ else
+ UnifFailure (evd,OccurCheck (fst ev,tR)))])
+ ev lF tR evd
+ in
+ let app_empty = match sk1, sk2 with [], [] -> true | _ -> false in
(* Evar must be undefined since we have flushed evars *)
let () = if !debug_unification then
let open Pp in
- let pr_state (tm,l) =
- h 0 (Termops.print_constr tm ++ str "|" ++ cut ()
- ++ prlist_with_sep pr_semicolon
- (fun x -> hov 1 (Termops.print_constr x)) l) in
- pp (v 0 (pr_state appr1 ++ cut () ++ pr_state appr2 ++ cut ()) ++ fnl ()) in
- match (flex_kind_of_term term1 l1, flex_kind_of_term term2 l2) with
+ pp (v 0 (pr_state appr1 ++ cut () ++ pr_state appr2 ++ cut ())
+ ++ fnl ()) 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) ->
let f1 i =
- if List.length l1 > List.length l2 then
- let (deb1,rest1) = list_chop (List.length l1-List.length l2) l1 in
- ise_and i
- [(fun i -> solve_simple_eqn (evar_conv_x ts) env i
- (position_problem false pbty,ev2,applist(term1,deb1)));
- (fun i -> ise_list2 i
- (fun i -> evar_conv_x ts env i CONV) rest1 l2)]
- else
- let (deb2,rest2) = list_chop (List.length l2-List.length l1) l2 in
- ise_and i
- [(fun i -> solve_simple_eqn (evar_conv_x ts) env i
- (position_problem true pbty,ev1,applist(term2,deb2)));
- (fun i -> ise_list2 i
- (fun i -> evar_conv_x ts env i CONV) l1 rest2)]
+ match ise_stack2 false env i (evar_conv_x ts) sk1 sk2 with
+ |None, Success i' ->
+ (* Evar can be defined in i' *)
+ let ev1' = whd_evar i' (mkEvar ev1) in
+ if isEvar ev1' then
+ solve_simple_eqn (evar_conv_x ts) env i'
+ (position_problem true pbty,destEvar ev1',term2)
+ else
+ evar_eqappr_x ts env evd pbty
+ ((ev1', sk1), csts1) ((term2, sk2), csts2)
+ |Some (r,[]), Success i' ->
+ let ev2' = whd_evar i' (mkEvar ev2) in
+ if isEvar ev2' then
+ solve_simple_eqn (evar_conv_x ts) env i'
+ (position_problem false pbty,destEvar ev2',Stack.zip(term1,r))
+ else
+ evar_eqappr_x ts env evd pbty
+ ((ev2', sk1), csts1) ((term2, sk2), csts2)
+
+ |Some ([],r), Success i' ->
+ let ev1' = whd_evar i' (mkEvar ev1) in
+ if isEvar ev1' then
+ solve_simple_eqn (evar_conv_x ts) env i'
+ (position_problem true pbty,destEvar ev1',Stack.zip(term2,r))
+ else evar_eqappr_x ts env evd pbty
+ ((ev1', sk1), csts1) ((term2, sk2), csts2)
+ |_, (UnifFailure _ as x) -> x
+ |Some _, _ -> UnifFailure (i,NotSameArgSize)
and f2 i =
- if sp1 = sp2 then
- ise_and i
- [(fun i -> ise_list2 i
- (fun i -> evar_conv_x ts env i CONV) l1 l2);
- (fun i -> solve_refl (evar_conv_x ts) env i sp1 al1 al2,
- true)]
- else (i,false)
+ if Evar.equal sp1 sp2 then
+ match ise_stack2 false env i (evar_conv_x ts) sk1 sk2 with
+ |None, Success i' ->
+ Success (solve_refl (fun env i pbty a1 a2 ->
+ is_success (evar_conv_x ts env i pbty a1 a2))
+ env i' (position_problem true pbty) sp1 al1 al2)
+ |_, (UnifFailure _ as x) -> x
+ |Some _, _ -> UnifFailure (i,NotSameArgSize)
+ else UnifFailure (i,NotSameHead)
in
ise_try evd [f1; f2]
- | Flexible ev1, MaybeFlexible flex2 ->
- let f1 i =
- match is_unification_pattern_evar env evd ev1 l1 (applist appr2) with
- | Some l1' ->
- (* Miller-Pfenning's patterns unification *)
- (* Preserve generality (except that CCI has no eta-conversion) *)
- let t2 = nf_evar evd (applist appr2) in
- let t2 = solve_pattern_eqn env l1' t2 in
- solve_simple_eqn (evar_conv_x ts) env evd
- (position_problem true pbty,ev1,t2)
- | None -> (i,false)
- and f2 i =
- if
- List.length l1 <= List.length l2
- then
- (* Try first-order unification *)
- (* (heuristic that gives acceptable results in practice) *)
- let (deb2,rest2) =
- list_chop (List.length l2-List.length l1) l2 in
- ise_and i
- (* First compare extra args for better failure message *)
- [(fun i -> ise_list2 i
- (fun i -> evar_conv_x ts env i CONV) l1 rest2);
- (fun i -> evar_conv_x ts env i pbty term1 (applist(term2,deb2)))]
- else (i,false)
- and f3 i =
- match eval_flexible_term ts env flex2 with
- | Some v2 ->
- evar_eqappr_x ts env i pbty appr1 (evar_apprec ts env i l2 v2)
- | None -> (i,false)
- in
- ise_try evd [f1; f2; f3]
+ | Flexible ev1, MaybeFlexible v2 ->
+ flex_maybeflex true ev1 (appr1,csts1) (appr2,csts2) v2
- | MaybeFlexible flex1, Flexible ev2 ->
- let f1 i =
- match is_unification_pattern_evar env evd ev2 l2 (applist appr1) with
- | Some l1' ->
- (* Miller-Pfenning's patterns unification *)
- (* Preserve generality (except that CCI has no eta-conversion) *)
- let t1 = nf_evar evd (applist appr1) in
- let t1 = solve_pattern_eqn env l2 t1 in
- solve_simple_eqn (evar_conv_x ts) env evd
- (position_problem false pbty,ev2,t1)
- | None -> (i,false)
- and f2 i =
- if
- List.length l2 <= List.length l1
- then
- (* Try first-order unification *)
- (* (heuristic that gives acceptable results in practice) *)
- let (deb1,rest1) = list_chop (List.length l1-List.length l2) l1 in
- ise_and i
- (* First compare extra args for better failure message *)
- [(fun i -> ise_list2 i
- (fun i -> evar_conv_x ts env i CONV) rest1 l2);
- (fun i -> evar_conv_x ts env i pbty (applist(term1,deb1)) term2)]
- else (i,false)
- and f3 i =
- match eval_flexible_term ts env flex1 with
- | Some v1 ->
- evar_eqappr_x ts env i pbty (evar_apprec ts env i l1 v1) appr2
- | None -> (i,false)
- in
- ise_try evd [f1; f2; f3]
+ | MaybeFlexible v1, Flexible ev2 ->
+ flex_maybeflex false ev2 (appr2,csts2) (appr1,csts1) v1
- | MaybeFlexible flex1, MaybeFlexible flex2 -> begin
- match kind_of_term flex1, kind_of_term flex2 with
- | LetIn (na,b1,t1,c'1), LetIn (_,b2,_,c'2) ->
+ | MaybeFlexible v1, MaybeFlexible v2 -> begin
+ match kind_of_term term1, kind_of_term term2 with
+ | LetIn (na,b1,t1,c'1), LetIn (_,b2,t2,c'2) ->
let f1 i =
ise_and i
- [(fun i -> evar_conv_x ts env i CONV b1 b2);
+ [(fun i -> evar_conv_x ts env i CONV t1 t2);
+ (fun i -> evar_conv_x ts env i CONV b1 b2);
(fun i ->
let b = nf_evar i b1 in
let t = nf_evar i t1 in
evar_conv_x ts (push_rel (na,Some b,t) env) i pbty c'1 c'2);
- (fun i -> ise_list2 i (fun i -> evar_conv_x ts env i CONV) l1 l2)]
+ (fun i -> exact_ise_stack2 env i (evar_conv_x ts) sk1 sk2)]
and f2 i =
- let appr1 = evar_apprec ts env i l1 (subst1 b1 c'1)
- and appr2 = evar_apprec ts env i l2 (subst1 b2 c'2)
- in evar_eqappr_x ts env i pbty appr1 appr2
+ let out1 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts1 (v1,sk1)
+ and out2 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts2 (v2,sk2)
+ in evar_eqappr_x ts env i pbty out1 out2
in
ise_try evd [f1; f2]
+ | Proj (p, c), Proj (p', c')
+ when Constant.equal (Projection.constant p) (Projection.constant p') ->
+ let f1 i =
+ ise_and i
+ [(fun i -> evar_conv_x ts env i CONV c c');
+ (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)
+ and out2 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts2 (v2,sk2)
+ in evar_eqappr_x ts env i pbty out1 out2
+ in
+ 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' ->
+ let res =
+ try Some (destApp (Retyping.expand_projection env evd p c []))
+ with Retyping.RetypeError _ -> None
+ in
+ (match res with
+ | Some (f1,args1) ->
+ evar_eqappr_x ts env evd pbty ((f1,Stack.append_app args1 sk1),csts1)
+ (appr2,csts2)
+ | None -> UnifFailure (evd,NotSameHead))
+
+ | Const (p,u), Proj (p',c') when eq_constant p (Projection.constant p') ->
+ let res =
+ try Some (destApp (Retyping.expand_projection env evd p' c' []))
+ with Retyping.RetypeError _ -> None
+ in
+ (match res with
+ | Some (f2,args2) ->
+ evar_eqappr_x ts env evd pbty (appr1,csts1) ((f2,Stack.append_app args2 sk2),csts2)
+ | None -> UnifFailure (evd,NotSameHead))
+
| _, _ ->
- let f1 i =
- if eq_constr flex1 flex2 then
- ise_list2 i (fun i -> evar_conv_x ts env i CONV) l1 l2
- else
- (i,false)
+ let f1 i =
+ (* Gather the universe constraints that would make term1 and term2 equal.
+ If these only involve unifications of flexible universes to other universes,
+ 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
+ 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)
and f2 i =
- (try conv_record ts env i
- (try check_conv_record appr1 appr2
- with Not_found -> check_conv_record appr2 appr1)
- with Not_found -> (i,false))
+ (try
+ if not (snd ts) then raise Not_found
+ else conv_record ts env i
+ (try check_conv_record env i appr1 appr2
+ with Not_found -> check_conv_record env i appr2 appr1)
+ with Not_found -> UnifFailure (i,NoCanonicalStructure))
and f3 i =
(* heuristic: unfold second argument first, exception made
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
- | (Var _|Construct _|Ind _|Const _|Prod _|Sort _) -> false
- | (Case _|Fix _|CoFix _|Meta _|Rel _)-> true
- | Evar _ -> false (* immediate solution without Canon Struct *)
- | Lambda _ -> assert(args = []); true
- | LetIn (_,b,_,c) ->
- is_unnamed (evar_apprec ts env i args (subst1 b c))
- | App _| Cast _ -> assert false in
+ | (Var _|Construct _|Ind _|Const _|Prod _|Sort _) ->
+ Stack.not_purely_applicative args
+ | (CoFix _|Meta _|Rel _)-> true
+ | Evar _ -> Stack.not_purely_applicative args
+ (* false (* immediate solution without Canon Struct *)*)
+ | Lambda _ -> assert (match args with [] -> true | _ -> false); true
+ | LetIn (_,b,_,c) -> is_unnamed
+ (fst (whd_betaiota_deltazeta_for_iota_state
+ (fst ts) env i Cst_stack.empty (subst1 b c, args)))
+ | Fix _ -> true (* Partially applied fix can be the result of a whd call *)
+ | Proj (p, _) -> Projection.unfolded p || Stack.not_purely_applicative args
+ | Case _ | App _| Cast _ -> assert false in
let rhs_is_stuck_and_unnamed () =
- match eval_flexible_term ts env flex2 with
- | None -> false
- | Some v2 -> is_unnamed (evar_apprec ts env i l2 v2) in
+ let applicative_stack = fst (Stack.strip_app sk2) in
+ is_unnamed
+ (fst (whd_betaiota_deltazeta_for_iota_state
+ (fst ts) env i Cst_stack.empty (v2, applicative_stack))) in
let rhs_is_already_stuck =
rhs_is_already_stuck || rhs_is_stuck_and_unnamed () in
- if isLambda flex1 || rhs_is_already_stuck then
- match eval_flexible_term ts env flex1 with
- | Some v1 ->
- evar_eqappr_x ~rhs_is_already_stuck
- ts env i pbty (evar_apprec ts env i l1 v1) appr2
- | None ->
- match eval_flexible_term ts env flex2 with
- | Some v2 ->
- evar_eqappr_x ts env i pbty appr1 (evar_apprec ts env i l2 v2)
- | None -> (i,false)
+
+ if (isLambda 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
+ (fst ts) env i (Cst_stack.add_cst term1 csts1) (v1,sk1))
+ (appr2,csts2)
else
- match eval_flexible_term ts env flex2 with
- | Some v2 ->
- evar_eqappr_x ts env i pbty appr1 (evar_apprec ts env i l2 v2)
- | None ->
- match eval_flexible_term ts env flex1 with
- | Some v1 ->
- evar_eqappr_x ts env i pbty (evar_apprec ts env i l1 v1) appr2
- | None -> (i,false)
+ evar_eqappr_x ts env i pbty (appr1,csts1)
+ (whd_betaiota_deltazeta_for_iota_state
+ (fst ts) env i (Cst_stack.add_cst term2 csts2) (v2,sk2))
in
ise_try evd [f1; f2; f3]
- end
+ end
- | Rigid c1, Rigid c2 when isLambda c1 & isLambda c2 ->
- let (na,c1,c'1) = destLambda c1 in
- let (_,c2,c'2) = destLambda c2 in
- assert (l1=[] & l2=[]);
+ | Rigid, Rigid when isLambda term1 && isLambda term2 ->
+ let (na,c1,c'1) = destLambda term1 in
+ let (_,c2,c'2) = destLambda 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
evar_conv_x ts (push_rel (na,None,c) env) i CONV c'1 c'2)]
- | Flexible ev1, (Rigid _ | PseudoRigid _) ->
- (match is_unification_pattern_evar env evd ev1 l1 (applist appr2) with
- | Some l1 ->
- (* Miller-Pfenning's pattern unification *)
- (* Preserve generality thanks to eta-conversion) *)
- let t2 = nf_evar evd (applist appr2) in
- let t2 = solve_pattern_eqn env l1 t2 in
- solve_simple_eqn (evar_conv_x ts) env evd
- (position_problem true pbty,ev1,t2)
- | None ->
- if isLambda term2 then
- eta env evd false term2 l2 term1 l1
- else
- (* Postpone the use of an heuristic *)
- add_conv_pb (pbty,env,applist appr1,applist appr2) evd,
- true)
-
- | (Rigid _ | PseudoRigid _), Flexible ev2 ->
- (match is_unification_pattern_evar env evd ev2 l2 (applist appr1) with
- | Some l2 ->
- (* Miller-Pfenning's pattern unification *)
- (* Preserve generality thanks to eta-conversion) *)
- let t1 = nf_evar evd (applist appr1) in
- let t1 = solve_pattern_eqn env l2 t1 in
- solve_simple_eqn (evar_conv_x ts) env evd
- (position_problem false pbty,ev2,t1)
- | None ->
- if isLambda term1 then
- eta env evd true term1 l1 term2 l2
- else
- (* Postpone the use of an heuristic *)
- add_conv_pb (pbty,env,applist appr1,applist appr2) evd,
- true)
-
- | MaybeFlexible flex1, (Rigid _ | PseudoRigid _) ->
+ | Flexible ev1, Rigid -> flex_rigid true ev1 appr1 appr2
+ | Rigid, Flexible ev2 -> flex_rigid false ev2 appr2 appr1
+
+ | MaybeFlexible v1, Rigid ->
let f3 i =
- (try conv_record ts env i (check_conv_record appr1 appr2)
- with Not_found -> (i,false))
+ (try
+ if not (snd ts) then raise Not_found
+ else conv_record ts env i (check_conv_record env i appr1 appr2)
+ with Not_found -> UnifFailure (i,NoCanonicalStructure))
and f4 i =
- match eval_flexible_term ts env flex1 with
- | Some v1 ->
- evar_eqappr_x ts env i pbty (evar_apprec ts env i l1 v1) appr2
- | None ->
- if isLambda term2 then
- eta env i false term2 l2 term1 l1
- else
- (i,false)
+ evar_eqappr_x ts env i pbty
+ (whd_betaiota_deltazeta_for_iota_state
+ (fst ts) env i (Cst_stack.add_cst term1 csts1) (v1,sk1))
+ (appr2,csts2)
in
- ise_try evd [f3; f4]
+ ise_try evd [f3; f4]
- | (Rigid _ | PseudoRigid _), MaybeFlexible flex2 ->
+ | Rigid, MaybeFlexible v2 ->
let f3 i =
- (try conv_record ts env i (check_conv_record appr2 appr1)
- with Not_found -> (i,false))
+ (try
+ if not (snd ts) then raise Not_found
+ else conv_record ts env i (check_conv_record env i appr2 appr1)
+ with Not_found -> UnifFailure (i,NoCanonicalStructure))
and f4 i =
- match eval_flexible_term ts env flex2 with
- | Some v2 ->
- evar_eqappr_x ts env i pbty appr1 (evar_apprec ts env i l2 v2)
- | None ->
- if isLambda term1 then
- eta env i true term1 l1 term2 l2
- else
- (i,false)
+ evar_eqappr_x ts env i pbty (appr1,csts1)
+ (whd_betaiota_deltazeta_for_iota_state
+ (fst ts) env i (Cst_stack.add_cst term2 csts2) (v2,sk2))
in
- ise_try evd [f3; f4]
+ ise_try evd [f3; f4]
(* Eta-expansion *)
- | Rigid c1, _ when isLambda c1 ->
- eta env evd true term1 l1 term2 l2
-
- | _, Rigid c2 when isLambda c2 ->
- eta env evd false term2 l2 term1 l1
-
- | Rigid c1, Rigid c2 -> begin
- match kind_of_term c1, kind_of_term c2 with
-
- | Sort s1, Sort s2 when l1=[] & l2=[] ->
- (try
- let evd' =
- if pbty = CONV
- then Evd.set_eq_sort evd s1 s2
- else Evd.set_leq_sort evd s1 s2
- in (evd', true)
- with Univ.UniverseInconsistency _ -> (evd, false)
- | e when Errors.noncritical e -> (evd, false))
-
- | Prod (n,c1,c'1), Prod (_,c2,c'2) when l1=[] & l2=[] ->
+ | Rigid, _ when isLambda term1 ->
+ eta env evd true sk1 term1 sk2 term2
+
+ | _, Rigid when isLambda term2 ->
+ eta env evd false sk2 term2 sk1 term1
+
+ | Rigid, Rigid -> begin
+ match kind_of_term term1, kind_of_term term2 with
+
+ | Sort s1, Sort s2 when app_empty ->
+ (try
+ let evd' =
+ if pbty == CONV
+ then Evd.set_eq_sort env evd s1 s2
+ else Evd.set_leq_sort env evd s1 s2
+ in Success evd'
+ with Univ.UniverseInconsistency p ->
+ UnifFailure (evd,UnifUnivInconsistency p)
+ | e when Errors.noncritical e -> UnifFailure (evd,NotSameHead))
+
+ | Prod (n,c1,c'1), Prod (_,c2,c'2) when 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
evar_conv_x ts (push_rel (n,None,c) env) i pbty c'1 c'2)]
- | Ind sp1, Ind sp2 ->
- if eq_ind sp1 sp2 then
- ise_list2 evd (fun i -> evar_conv_x ts env i CONV) l1 l2
- else (evd, false)
-
- | Construct sp1, Construct sp2 ->
- if eq_constructor sp1 sp2 then
- ise_list2 evd (fun i -> evar_conv_x ts env i CONV) l1 l2
- else (evd, false)
+ | Rel x1, Rel x2 ->
+ if Int.equal x1 x2 then
+ exact_ise_stack2 env evd (evar_conv_x ts) sk1 sk2
+ else UnifFailure (evd,NotSameHead)
+
+ | Var var1, Var var2 ->
+ if Id.equal var1 var2 then
+ exact_ise_stack2 env evd (evar_conv_x ts) sk1 sk2
+ else UnifFailure (evd,NotSameHead)
+
+ | Const _, Const _
+ | Ind _, Ind _
+ | Construct _, Construct _ ->
+ rigids env evd sk1 term1 sk2 term2
+
+ | Construct u, _ ->
+ eta_constructor ts env evd sk1 u sk2 term2
+
+ | _, Construct u ->
+ eta_constructor ts env evd sk2 u sk1 term1
+
+ | Fix ((li1, i1),(_,tys1,bds1 as recdef1)), Fix ((li2, i2),(_,tys2,bds2)) -> (* Partially applied fixs *)
+ if Int.equal i1 i2 && Array.equal Int.equal li1 li2 then
+ ise_and evd [
+ (fun i -> ise_array2 i (fun i' -> evar_conv_x ts env i' CONV) tys1 tys2);
+ (fun i -> ise_array2 i (fun i' -> evar_conv_x ts (push_rec_types recdef1 env) i' CONV) bds1 bds2);
+ (fun i -> exact_ise_stack2 env i (evar_conv_x ts) sk1 sk2)]
+ else UnifFailure (evd, NotSameHead)
| CoFix (i1,(_,tys1,bds1 as recdef1)), CoFix (i2,(_,tys2,bds2)) ->
- if i1=i2 then
+ if Int.equal i1 i2 then
ise_and evd
[(fun i -> ise_array2 i
(fun i -> evar_conv_x ts env i CONV) tys1 tys2);
(fun i -> ise_array2 i
(fun i -> evar_conv_x ts (push_rec_types recdef1 env) i CONV)
bds1 bds2);
- (fun i -> ise_list2 i
- (fun i -> evar_conv_x ts env i CONV) l1 l2)]
- else (evd,false)
-
- | (Ind _ | Construct _ | Sort _ | Prod _ | CoFix _), _ -> (evd,false)
- | _, (Ind _ | Construct _ | Sort _ | Prod _ | CoFix _) -> (evd,false)
-
- | (App _ | Meta _ | Cast _ | Case _ | Fix _), _ -> assert false
- | (LetIn _ | Rel _ | Var _ | Const _ | Evar _), _ -> assert false
+ (fun i -> exact_ise_stack2 env i
+ (evar_conv_x ts) sk1 sk2)]
+ else UnifFailure (evd,NotSameHead)
+
+ | (Meta _, _) | (_, Meta _) ->
+ 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'))
+ end
+
+ | (Ind _ | Sort _ | Prod _ | CoFix _ | Fix _ | Rel _ | Var _ | Const _), _ ->
+ UnifFailure (evd,NotSameHead)
+ | _, (Ind _ | Sort _ | Prod _ | CoFix _ | Fix _ | Rel _ | Var _ | Const _) ->
+ UnifFailure (evd,NotSameHead)
+
+ | (App _ | Cast _ | Case _ | Proj _), _ -> assert false
+ | (LetIn _| Evar _), _ -> assert false
| (Lambda _), _ -> assert false
end
- | PseudoRigid c1, PseudoRigid c2 -> begin
- match kind_of_term c1, kind_of_term c2 with
+and conv_record trs env evd (ctx,(h,h2),c,bs,(params,params1),(us,us2),(sk1,sk2),c1,(n,t2)) =
+ (* Tries to unify the states
- | Case (_,p1,c1,cl1), Case (_,p2,c2,cl2) ->
- ise_and evd
- [(fun i -> evar_conv_x ts env i CONV p1 p2);
- (fun i -> evar_conv_x ts env i CONV c1 c2);
- (fun i -> ise_array2 i
- (fun i -> evar_conv_x ts env i CONV) cl1 cl2);
- (fun i -> ise_list2 i (fun i -> evar_conv_x ts env i CONV) l1 l2)]
-
- | Fix (li1,(_,tys1,bds1 as recdef1)), Fix (li2,(_,tys2,bds2)) ->
- if li1=li2 then
- ise_and evd
- [(fun i -> ise_array2 i
- (fun i -> evar_conv_x ts env i CONV) tys1 tys2);
- (fun i -> ise_array2 i
- (fun i -> evar_conv_x ts (push_rec_types recdef1 env) i CONV)
- bds1 bds2);
- (fun i -> ise_list2 i
- (fun i -> evar_conv_x ts env i CONV) l1 l2)]
- else (evd,false)
+ (proji params1 c1 | sk1) = (proji params2 (c (?xs:bs)) | sk2)
- | (Meta _ | Case _ | Fix _ | CoFix _),
- (Meta _ | Case _ | Fix _ | CoFix _) -> (evd,false)
+ and the terms
- | (App _ | Ind _ | Construct _ | Sort _ | Prod _), _ -> assert false
- | _, (App _ | Ind _ | Construct _ | Sort _ | Prod _) -> assert false
+ h us = h2 us2
- | (LetIn _ | Cast _), _ -> assert false
- | _, (LetIn _ | Cast _) -> assert false
+ where
- | (Lambda _ | Rel _ | Var _ | Const _ | Evar _), _ -> assert false
- | _, (Lambda _ | Rel _ | Var _ | Const _ | Evar _) -> assert false
- end
+ c = the constant for the canonical structure (i.e. some term of the form
+ fun (xs:bs) => Build_R params v1 .. vi-1 (h us) vi+1 .. vn)
+ bs = the types of the parameters of the canonical structure
+ c1 = the main argument of the canonical projection
+ sk1, sk2 = the surrounding stacks of the conversion problem
+ params1, params2 = the params of the projection (empty if a primitive proj)
- | PseudoRigid _, Rigid _ -> (evd,false)
+ knowing that
- | Rigid _, PseudoRigid _ -> (evd,false)
+ (proji params1 c1 | sk1) = (h2 us2 | sk2)
-and conv_record trs env evd (c,bs,(params,params1),(us,us2),(ts,ts1),c1,(n,t2)) =
- let (evd',ks,_) =
- List.fold_left
- (fun (i,ks,m) b ->
- if m=n then (i,t2::ks, m-1) else
- let dloc = (dummy_loc,InternalHole) in
- let (i',ev) = new_evar i env ~src:dloc (substl ks b) in
- (i', ev :: ks, m - 1))
- (evd,[],List.length bs - 1) bs
- in
- ise_and evd'
- [(fun i ->
- ise_list2 i
- (fun i x1 x -> evar_conv_x trs env i CONV x1 (substl ks x))
- params1 params);
- (fun i ->
- ise_list2 i
- (fun i u1 u -> evar_conv_x trs env i CONV u1 (substl ks u))
- us2 us);
- (fun i -> evar_conv_x trs env i CONV c1 (applist (c,(List.rev ks))));
- (fun i -> ise_list2 i (fun i -> evar_conv_x trs env i CONV) ts ts1)]
+ had to be initially resolved
+ *)
+ let evd = Evd.merge_context_set Evd.univ_flexible evd ctx in
+ if Reductionops.Stack.compare_shape sk1 sk2 then
+ let (evd',ks,_,test) =
+ List.fold_left
+ (fun (i,ks,m,test) b ->
+ if match n with Some n -> Int.equal m n | None -> false then
+ let ty = Retyping.get_type_of env i t2 in
+ 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',ev) = 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
+ let app = mkApp (c, Array.rev_of_list ks) in
+ ise_and evd'
+ [(fun i ->
+ exact_ise_stack2 env i
+ (fun env' i' cpb x1 x -> evar_conv_x trs env' i' cpb x1 (substl ks x))
+ params1 params);
+ (fun i ->
+ exact_ise_stack2 env i
+ (fun env' i' cpb u1 u -> evar_conv_x trs env' i' cpb u1 (substl ks u))
+ us2 us);
+ (fun i -> evar_conv_x trs env i CONV c1 app);
+ (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))))]
+ 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.CoFinite ->
+ 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)
+ in
+ exact_ise_stack2 env evd (evar_conv_x (fst ts, false)) l1'
+ (Stack.append_app_list l2' Stack.empty)
+ with
+ | Invalid_argument _ ->
+ (* Stack.tail: partially applied constructor *)
+ UnifFailure(evd,NotSameHead))
+ | _ -> UnifFailure (evd,NotSameHead)
+
+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
+ else evar_conv_x
+
+let evar_conv_hook_get, evar_conv_hook_set = Hook.make ~default:evar_conv_x ()
+
+let evar_conv_x ts = Hook.get evar_conv_hook_get ts
+
+let set_evar_conv f = Hook.set evar_conv_hook_set f
-(* getting rid of the optional argument rhs_is_already_stuck *)
-let evar_eqappr_x ts env evd pbty appr1 appr2 =
- evar_eqappr_x ts env evd pbty appr1 appr2
(* We assume here |l1| <= |l2| *)
let first_order_unification ts env evd (ev1,l1) (term2,l2) =
- let (deb2,rest2) = list_chop (List.length l2-List.length l1) l2 in
+ let (deb2,rest2) = Array.chop (Array.length l2-Array.length l1) l2 in
ise_and evd
(* First compare extra args for better failure message *)
- [(fun i -> ise_list2 i (fun i -> evar_conv_x ts env i CONV) rest2 l1);
+ [(fun i -> ise_array2 i (fun i -> evar_conv_x ts env i CONV) rest2 l1);
(fun i ->
(* Then instantiate evar unless already done by unifying args *)
- let t2 = applist(term2,deb2) in
- if is_defined_evar i ev1 then
+ let t2 = mkApp(term2,deb2) in
+ if is_defined i (fst ev1) then
evar_conv_x ts env i CONV t2 (mkEvar ev1)
else
solve_simple_eqn ~choose:true (evar_conv_x ts) env i (None,ev1,t2))]
@@ -610,46 +902,50 @@ 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) -> eq_constr c term) subst in
- if subst' = [] then evd, false else
- Evd.define evk (mkVar (fst (List.hd subst'))) evd, true
+ let subst' = List.filter (fun (id,c) -> Term.eq_constr c term) subst in
+ match subst' with
+ | [] -> None
+ | (id, _) :: _ -> Some (Evd.define evk (mkVar id) evd)
-let apply_on_subterm evdref f c t =
- let rec applyrec (k,c as kc) t =
+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 eq_constr c t then f k
+ if e_eq_constr_univs evdref c t then f k
else
match kind_of_term t with
| Evar (evk,args) when Evd.is_undefined !evdref evk ->
let ctx = evar_filtered_context (Evd.find_undefined !evdref evk) in
- let g (_,b,_) a = if b = None then applyrec kc a else a in
+ let g (_,b,_) a = if Option.is_empty b 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 (fun d (k,c) -> (k+1,lift 1 c))
- applyrec kc t
+ map_constr_with_binders_left_to_right
+ (fun d (env,(k,c)) -> (push_rel d env, (k+1,lift 1 c)))
+ applyrec acc t
in
- applyrec (0,c) t
+ applyrec (env,(0,c)) t
let filter_possible_projections c ty ctxt args =
- let fv1 = free_rels c in
- let fv2 = collect_vars c in
+ (* 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
+ let len = Array.length args in
let tyvars = collect_vars ty in
- List.map2 (fun (id,b,_) a ->
- b <> None ||
+ List.map_i (fun i (id,b,_) ->
+ let () = assert (i < len) in
+ let a = Array.unsafe_get args i in
+ (match b with None -> false | Some c -> not (isRel c || isVar 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 && Intset.mem (destRel a) fv1 ||
- isVar a && Idset.mem (destVar a) fv2 ||
- Idset.mem id tyvars)
- ctxt args
-
-let initial_evar_data evi =
- let ids = List.map pi1 (evar_context evi) in
- (evar_filter evi, List.map mkVar ids)
+ isRel a && Int.Set.mem (destRel a) fv1 ||
+ isVar a && Id.Set.mem (destVar a) fv2 ||
+ Id.Set.mem id tyvars)
+ 0 ctxt
let solve_evars = ref (fun _ -> failwith "solve_evars not installed")
let set_solve_evars f = solve_evars := f
@@ -671,48 +967,49 @@ let set_solve_evars f = solve_evars := f
* proposition from Dan Grayson]
*)
+exception TypingFailed of evar_map
+
let second_order_matching ts env_rhs evd (evk,args) argoccs rhs =
try
- let args = Array.to_list args in
let evi = Evd.find_undefined evd evk in
- let env_evar = evar_env evi in
+ let env_evar = evar_filtered_env evi in
let sign = named_context_val env_evar in
let ctxt = evar_filtered_context evi in
- let filter = evar_filter evi in
let instance = List.map mkVar (List.map pi1 ctxt) in
let rec make_subst = function
| (id,_,t)::ctxt', c::l, occs::occsl when isVarId id c ->
- if occs<>None then
+ begin match occs with
+ | Some _ ->
error "Cannot force abstraction on identity instance."
- else
+ | None ->
make_subst (ctxt',l,occsl)
+ end
| (id,_,t)::ctxt', c::l, occs::occsl ->
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 = List.map2 (&&) filter filter' in
- (id,t,c,ty,evs,filter,occs) :: make_subst (ctxt',l,occsl)
- | [], [], [] -> []
- | _ -> anomaly "Signature, instance and occurrences list do not match" 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
let rec set_holes evdref rhs = function
| (id,_,c,cty,evsref,filter,occs)::subst ->
let set_var k =
match occs with
- | Some (false,[]) -> mkVar id
+ | Some Locus.AllOccurrences -> mkVar id
| Some _ -> error "Selection of specific occurrences not supported"
| None ->
let evty = set_holes evdref cty subst in
- let instance = snd (list_filter2 (fun b c -> b) (filter,instance)) in
+ let instance = Filter.filter_list filter instance in
let evd,ev = new_evar_instance sign !evdref evty ~filter instance in
evdref := evd;
evsref := (fst (destEvar ev),evty)::!evsref;
ev in
- set_holes evdref (apply_on_subterm evdref set_var c rhs) subst
+ set_holes evdref (apply_on_subterm env_rhs evdref set_var c rhs) subst
| [] -> rhs in
- let subst = make_subst (ctxt,args,argoccs) in
+ let subst = make_subst (ctxt,Array.to_list args,argoccs) in
let evdref = ref evd in
let rhs = set_holes evdref rhs subst in
@@ -720,10 +1017,11 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs =
(* We instantiate the evars of which the value is forced by typing *)
let evd,rhs =
- try !solve_evars env_evar evd rhs
+ let evdref = ref evd in
+ try let c = !solve_evars env_evar evdref rhs in !evdref,c
with e when Pretype_errors.precatchable_exception e ->
(* Could not revert all subterms *)
- raise Exit in
+ raise (TypingFailed !evdref) in
let rec abstract_free_holes evd = function
| (id,idty,c,_,evsref,_,_)::l ->
@@ -735,10 +1033,12 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs =
(* and we use typing to propagate this instantiation *)
(* This is an arbitrary choice *)
let evd = Evd.define evk (mkVar id) evd in
- let evd,b = evar_conv_x ts env_evar evd CUMUL idty evty in
- if not b then error "Cannot find an instance";
- let evd,b = reconsider_conv_pbs (evar_conv_x ts) evd in
- if not b then error "Cannot find an instance";
+ match evar_conv_x ts env_evar evd CUMUL idty evty with
+ | UnifFailure _ -> error "Cannot find an instance"
+ | Success evd ->
+ match reconsider_conv_pbs (evar_conv_x ts) evd with
+ | UnifFailure _ -> error "Cannot find an instance"
+ | Success evd ->
evd
else
evd
@@ -749,63 +1049,78 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs =
in
force_instantiation evd !evsref
| [] ->
- Evd.define evk rhs evd in
-
+ let evd =
+ 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
+ in
abstract_free_holes evd subst, true
- with Exit -> evd, false
+ with TypingFailed evd -> evd, false
let second_order_matching_with_args ts env evd ev l t =
(*
let evd,ev = evar_absorb_arguments env evd ev l in
- let argoccs = array_map_to_list (fun _ -> None) (snd ev) in
- second_order_matching ts env evd ev argoccs t
+ let argoccs = Array.map_to_list (fun _ -> None) (snd ev) in
+ let evd, b = second_order_matching ts env evd ev argoccs t in
+ if b then Success evd
+ else UnifFailure (evd, ConversionFailed (env,mkApp(mkEvar ev,l),t))
+ if b then Success evd else
*)
- (evd,false)
+ UnifFailure (evd, ConversionFailed (env,mkApp(mkEvar ev,l),t))
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) = decompose_app t1 in
- let (term2,l2 as appr2) = decompose_app 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 app_empty = Array.is_empty l1 && Array.is_empty l2 in
match kind_of_term term1, kind_of_term term2 with
- | Evar (evk1,args1), (Rel _|Var _) when l1 = [] & l2 = []
- & List.for_all (fun a -> eq_constr a term2 or isEvar a)
- (remove_instance_local_defs evd evk1 (Array.to_list args1)) ->
+ | Evar (evk1,args1), (Rel _|Var _) when app_empty
+ && List.for_all (fun a -> Term.eq_constr a term2 || isEvar a)
+ (remove_instance_local_defs evd evk1 args1) ->
(* The typical kind of constraint coming from pattern-matching return
type inference *)
- choose_less_dependent_instance evk1 evd term2 args1
- | (Rel _|Var _), Evar (evk2,args2) when l1 = [] & l2 = []
- & List.for_all (fun a -> eq_constr a term1 or isEvar a)
- (remove_instance_local_defs evd evk2 (Array.to_list args2)) ->
+ (match choose_less_dependent_instance evk1 evd term2 args1 with
+ | Some evd -> Success evd
+ | None -> UnifFailure (evd, ConversionFailed (env,term1,term2)))
+ | (Rel _|Var _), Evar (evk2,args2) when app_empty
+ && List.for_all (fun a -> Term.eq_constr a term1 || isEvar a)
+ (remove_instance_local_defs evd evk2 args2) ->
(* The typical kind of constraint coming from pattern-matching return
type inference *)
- choose_less_dependent_instance evk2 evd term1 args2
- | Evar (evk1,args1), Evar (evk2,args2) when evk1 = evk2 ->
- let f env evd pbty x y = (evd,is_trans_fconv pbty ts env evd x y) in
- solve_refl ~can_drop:true f env evd evk1 args1 args2, true
+ (match choose_less_dependent_instance evk2 evd term1 args2 with
+ | Some evd -> Success evd
+ | None -> UnifFailure (evd, ConversionFailed (env,term1,term2)))
+ | Evar (evk1,args1), Evar (evk2,args2) when Evar.equal evk1 evk2 ->
+ let f env evd pbty x y = is_trans_fconv pbty ts env evd x y in
+ Success (solve_refl ~can_drop:true f env evd
+ (position_problem true pbty) evk1 args1 args2)
| Evar ev1, Evar ev2 ->
- solve_evar_evar ~force:true
- (evar_define (evar_conv_x ts)) (evar_conv_x ts) env evd ev1 ev2, true
- | Evar ev1,_ when List.length l1 <= List.length l2 ->
+ Success (solve_evar_evar ~force:true
+ (evar_define (evar_conv_x ts) ~choose:true) (evar_conv_x ts) env evd
+ (position_problem true pbty) ev1 ev2)
+ | Evar ev1,_ when Array.length l1 <= Array.length l2 ->
(* On "?n t1 .. tn = u u1 .. u(n+p)", try first-order unification *)
(* and otherwise second-order matching *)
ise_try evd
[(fun evd -> first_order_unification ts env evd (ev1,l1) appr2);
(fun evd ->
- second_order_matching_with_args ts env evd ev1 l1 (applist appr2))]
- | _,Evar ev2 when List.length l2 <= List.length l1 ->
+ second_order_matching_with_args ts env evd ev1 l1 t2)]
+ | _,Evar ev2 when Array.length l2 <= Array.length l1 ->
(* On "u u1 .. u(n+p) = ?n t1 .. tn", try first-order unification *)
(* and otherwise second-order matching *)
ise_try evd
[(fun evd -> first_order_unification ts env evd (ev2,l2) appr1);
(fun evd ->
- second_order_matching_with_args ts env evd ev2 l2 (applist appr1))]
+ second_order_matching_with_args ts env evd ev2 l2 t1)]
| Evar ev1,_ ->
(* Try second-order pattern-matching *)
- second_order_matching_with_args ts env evd ev1 l1 (applist appr2)
+ second_order_matching_with_args ts env evd ev1 l1 t2
| _,Evar ev2 ->
(* Try second-order pattern-matching *)
- second_order_matching_with_args ts env evd ev2 l2 (applist appr1)
+ second_order_matching_with_args ts env evd ev2 l2 t1
| _ ->
(* Some head evar have been instantiated, or unknown kind of problem *)
evar_conv_x ts env evd pbty t1 t2
@@ -828,9 +1143,9 @@ let max_undefined_with_candidates evd =
| Some l -> (evk,ev_info,l)::evars) evd [] in
match l with
| [] -> None
- | a::l -> Some (list_last (a::l))
+ | a::l -> Some (List.last (a::l))
-let rec solve_unconstrained_evars_with_canditates evd =
+let rec solve_unconstrained_evars_with_candidates ts evd =
(* max_undefined is supposed to return the most recent, hence
possibly most dependent evar *)
match max_undefined_with_candidates evd with
@@ -840,70 +1155,87 @@ let rec solve_unconstrained_evars_with_canditates evd =
| [] -> error "Unsolvable existential variables."
| a::l ->
try
- let conv_algo = evar_conv_x full_transparent_state in
+ let conv_algo = evar_conv_x ts in
let evd = check_evar_instance evd evk a conv_algo in
let evd = Evd.define evk a evd in
- let evd,b = reconsider_conv_pbs conv_algo evd in
- if b then solve_unconstrained_evars_with_canditates evd
- else aux l
- with e when Pretype_errors.precatchable_exception e ->
- aux l in
+ match reconsider_conv_pbs conv_algo evd with
+ | Success evd -> solve_unconstrained_evars_with_candidates ts evd
+ | UnifFailure _ -> aux l
+ with
+ | IllTypedInstance _ -> aux l
+ | e when Pretype_errors.precatchable_exception e -> aux l in
(* List.rev is there to favor most dependent solutions *)
(* and favor progress when used with the refine tactics *)
let evd = aux (List.rev l) in
- solve_unconstrained_evars_with_canditates evd
+ solve_unconstrained_evars_with_candidates ts evd
-let solve_unconstrained_impossible_cases evd =
+let solve_unconstrained_impossible_cases env evd =
Evd.fold_undefined (fun evk ev_info evd' ->
match ev_info.evar_source with
- | _,ImpossibleCase -> Evd.define evk (j_type (coq_unit_judge ())) evd'
+ | _,Evar_kinds.ImpossibleCase ->
+ let j, ctx = coq_unit_judge () in
+ let evd' = Evd.merge_context_set Evd.univ_flexible_alg 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') evd evd
-
-let consider_remaining_unif_problems ?(ts=full_transparent_state) env evd =
- let evd = solve_unconstrained_evars_with_canditates evd in
+let consider_remaining_unif_problems env
+ ?(ts=Conv_oracle.get_transp_state (Environ.oracle env)) evd =
+ let evd = solve_unconstrained_evars_with_candidates ts evd in
let rec aux evd pbs progress stuck =
match pbs with
| (pbty,env,t1,t2 as pb) :: pbs ->
- let evd', b = apply_conversion_problem_heuristic ts env evd pbty t1 t2 in
- if b then
- let (evd', rest) = extract_all_conv_pbs evd' in
- if rest = [] then aux evd' pbs true stuck
- else (* Unification got actually stuck, postpone *)
+ (match apply_conversion_problem_heuristic ts env evd pbty t1 t2 with
+ | Success evd' ->
+ let (evd', rest) = extract_all_conv_pbs evd' in
+ begin match rest with
+ | [] -> aux evd' pbs true stuck
+ | _ -> (* Unification got actually stuck, postpone *)
aux evd pbs progress (pb :: stuck)
- else Pretype_errors.error_cannot_unify env evd (t1, t2)
+ end
+ | UnifFailure (evd,reason) ->
+ Pretype_errors.error_cannot_unify_loc (loc_of_conv_pb evd pb)
+ env evd ~reason (t1, t2))
| _ ->
if progress then aux evd stuck false []
else
match stuck with
| [] -> (* We're finished *) evd
- | (pbty,env,t1,t2) :: _ ->
+ | (pbty,env,t1,t2 as pb) :: _ ->
(* There remains stuck problems *)
- Pretype_errors.error_cannot_unify env evd (t1, t2)
+ Pretype_errors.error_cannot_unify_loc (loc_of_conv_pb evd pb)
+ env evd (t1, t2)
in
let (evd,pbs) = extract_all_conv_pbs evd in
let heuristic_solved_evd = aux evd pbs false [] in
check_problems_are_solved env heuristic_solved_evd;
- solve_unconstrained_impossible_cases heuristic_solved_evd
+ solve_unconstrained_impossible_cases env heuristic_solved_evd
(* Main entry points *)
-let the_conv_x ?(ts=full_transparent_state) env t1 t2 evd =
+exception UnableToUnify of evar_map * unification_error
+
+let default_transparent_state env = full_transparent_state
+(* Conv_oracle.get_transp_state (Environ.oracle env) *)
+
+let the_conv_x env ?(ts=default_transparent_state env) t1 t2 evd =
match evar_conv_x ts env evd CONV t1 t2 with
- (evd',true) -> evd'
- | _ -> raise Reduction.NotConvertible
+ | Success evd' -> evd'
+ | UnifFailure (evd',e) -> raise (UnableToUnify (evd',e))
-let the_conv_x_leq ?(ts=full_transparent_state) env t1 t2 evd =
+let the_conv_x_leq env ?(ts=default_transparent_state env) t1 t2 evd =
match evar_conv_x ts env evd CUMUL t1 t2 with
- (evd', true) -> evd'
- | _ -> raise Reduction.NotConvertible
+ | Success evd' -> evd'
+ | UnifFailure (evd',e) -> raise (UnableToUnify (evd',e))
-let e_conv ?(ts=full_transparent_state) env evdref t1 t2 =
+let e_conv env ?(ts=default_transparent_state env) evdref t1 t2 =
match evar_conv_x ts env !evdref CONV t1 t2 with
- (evd',true) -> evdref := evd'; true
- | _ -> false
+ | Success evd' -> evdref := evd'; true
+ | _ -> false
-let e_cumul ?(ts=full_transparent_state) env evdref t1 t2 =
+let e_cumul env ?(ts=default_transparent_state env) evdref t1 t2 =
match evar_conv_x ts env !evdref CUMUL t1 t2 with
- (evd',true) -> evdref := evd'; true
- | _ -> false
+ | Success evd' -> evdref := evd'; true
+ | _ -> false
diff --git a/pretyping/evarconv.mli b/pretyping/evarconv.mli
index 3a352d13..8bc30a71 100644
--- a/pretyping/evarconv.mli
+++ b/pretyping/evarconv.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,40 +8,72 @@
open Names
open Term
-open Sign
open Environ
-open Termops
open Reductionops
open Evd
+open Locus
-(** returns exception Reduction.NotConvertible if not unifiable *)
-val the_conv_x : ?ts:transparent_state -> env -> constr -> constr -> evar_map -> evar_map
-val the_conv_x_leq : ?ts:transparent_state -> env -> constr -> constr -> evar_map -> evar_map
+(** {4 Unification for type inference. } *)
+
+exception UnableToUnify of evar_map * Pretype_errors.unification_error
+
+(** {6 Main unification algorithm for type inference. } *)
+
+(** returns exception NotUnifiable with best known evar_map if not unifiable *)
+val the_conv_x : env -> ?ts:transparent_state -> constr -> constr -> evar_map -> evar_map
+val the_conv_x_leq : env -> ?ts:transparent_state -> constr -> constr -> evar_map -> evar_map
(** The same function resolving evars by side-effect and
catching the exception *)
-val e_conv : ?ts:transparent_state -> env -> evar_map ref -> constr -> constr -> bool
-val e_cumul : ?ts:transparent_state -> env -> evar_map ref -> constr -> constr -> bool
+val e_conv : env -> ?ts:transparent_state -> evar_map ref -> constr -> constr -> bool
+val e_cumul : env -> ?ts:transparent_state -> evar_map ref -> constr -> constr -> bool
-(**/**)
-(* For debugging *)
-val evar_conv_x : transparent_state ->
- env -> evar_map -> conv_pb -> constr -> constr -> evar_map * bool
-val evar_eqappr_x : transparent_state ->
- env -> evar_map ->
- conv_pb -> constr * constr list -> constr * constr list ->
- evar_map * bool
-(**/**)
+(** {6 Unification heuristics. } *)
+
+(** Try heuristics to solve pending unification problems and to solve
+ evars with candidates *)
+
+val consider_remaining_unif_problems : env -> ?ts:transparent_state -> evar_map -> evar_map
+
+(** Check all pending unification problems are solved and raise an
+ error otherwise *)
+
+val check_problems_are_solved : env -> evar_map -> unit
-val consider_remaining_unif_problems : ?ts:transparent_state -> env -> evar_map -> evar_map
+(** Check if a canonical structure is applicable *)
-val check_conv_record : constr * types list -> constr * types list ->
- constr * constr list * (constr list * constr list) *
- (constr list * types list) *
- (constr list * types list) * constr *
- (int * constr)
+val check_conv_record : env -> evar_map ->
+ constr * types Stack.t -> constr * types Stack.t ->
+ Univ.universe_context_set * (constr * constr)
+ * constr * constr list * (constr Stack.t * constr Stack.t) *
+ (constr Stack.t * types Stack.t) *
+ (constr Stack.t * types Stack.t) * constr *
+ (int option * constr)
-val set_solve_evars : (env -> evar_map -> constr -> evar_map * constr) -> unit
+(** 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 ->
+val second_order_matching : transparent_state -> env -> evar_map ->
existential -> occurrences option list -> constr -> evar_map * bool
+
+(** Declare function to enforce evars resolution by using typing constraints *)
+
+val set_solve_evars : (env -> evar_map ref -> constr -> constr) -> unit
+
+type unify_fun = transparent_state ->
+ env -> evar_map -> conv_pb -> constr -> constr -> Evarsolve.unification_result
+
+(** Override default [evar_conv_x] algorithm. *)
+val set_evar_conv : unify_fun -> unit
+
+(** The default unification algorithm with evars and universes. *)
+val evar_conv_x : unify_fun
+
+(**/**)
+(* For debugging *)
+val evar_eqappr_x : ?rhs_is_already_stuck:bool -> transparent_state * bool ->
+ env -> evar_map ->
+ conv_pb -> state * Cst_stack.t -> state * Cst_stack.t ->
+ Evarsolve.unification_result
+(**/**)
+
diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml
new file mode 100644
index 00000000..5aa72c90
--- /dev/null
+++ b/pretyping/evarsolve.ml
@@ -0,0 +1,1548 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Util
+open Errors
+open Names
+open Term
+open Vars
+open Context
+open Environ
+open Termops
+open Evd
+open Namegen
+open Retyping
+open Reductionops
+open Evarutil
+open Pretype_errors
+
+let normalize_evar evd ev =
+ match kind_of_term (whd_evar evd (mkEvar ev)) with
+ | Evar (evk,args) -> (evk,args)
+ | _ -> assert false
+
+let get_polymorphic_positions f =
+ let open Declarations in
+ match kind_of_term 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
+
+(**
+ forall A (l : list A) -> typeof A = Type i <= Datatypes.j -> i not refreshed
+ hd ?A (l : list t) -> A = t
+
+*)
+let refresh_universes ?(inferred=false) ?(onlyalg=false) pbty env evd t =
+ let evdref = ref evd in
+ let modified = ref false in
+ let rec refresh dir t =
+ match kind_of_term t with
+ | Sort (Type u as s) when
+ (match Univ.universe_level u with
+ | None -> true
+ | Some l -> not onlyalg && Option.is_empty (Evd.is_sort_variable evd s)) ->
+ let status = if inferred then Evd.univ_flexible_alg else Evd.univ_flexible in
+ let s' = evd_comb0 (new_sort_variable status) evdref in
+ let evd =
+ if dir then set_leq_sort env !evdref s' s
+ else set_leq_sort env !evdref s s'
+ in
+ modified := true; evdref := evd; mkSort s'
+ | Prod (na,u,v) ->
+ mkProd (na,u,refresh dir v)
+ | _ -> t
+ (** Refresh the types of evars under template polymorphic references *)
+ and refresh_term_evars onevars t =
+ match kind_of_term t with
+ | App (f, args) when is_template_polymorphic env f ->
+ let pos = get_polymorphic_positions f in
+ refresh_polymorphic_positions args pos
+ | Evar (ev, a) when onevars ->
+ let evi = Evd.find !evdref ev in
+ let ty' = refresh true evi.evar_concl in
+ if !modified then
+ evdref := Evd.add !evdref ev {evi with evar_concl = ty'}
+ else ()
+ | _ -> iter_constr (refresh_term_evars onevars) t
+ and refresh_polymorphic_positions args pos =
+ let rec aux i = function
+ | Some l :: ls ->
+ if i < Array.length args then
+ ignore(refresh_term_evars true args.(i));
+ aux (succ i) ls
+ | None :: ls ->
+ if i < Array.length args then
+ ignore(refresh_term_evars false args.(i));
+ aux (succ i) ls
+ | [] -> ()
+ in aux 0 pos
+ in
+ let t' =
+ if isArity t then
+ (match pbty with
+ | None -> t
+ | Some dir -> refresh dir t)
+ else (refresh_term_evars false t; t)
+ in
+ if !modified then !evdref, t' else !evdref, t
+
+let get_type_of_refresh ?(polyprop=true) ?(lax=false) env sigma c =
+ let ty = Retyping.get_type_of ~polyprop ~lax env sigma c in
+ refresh_universes (Some false) env sigma ty
+
+(************************)
+(* Unification results *)
+(************************)
+
+type unification_result =
+ | Success of evar_map
+ | UnifFailure of evar_map * unification_error
+
+let is_success = function Success _ -> true | UnifFailure _ -> false
+
+let test_success conv_algo env evd c c' rhs =
+ is_success (conv_algo env evd c c' rhs)
+
+let add_conv_oriented_pb (pbty,env,t1,t2) evd =
+ match pbty with
+ | Some true -> add_conv_pb (Reduction.CUMUL,env,t1,t2) evd
+ | Some false -> add_conv_pb (Reduction.CUMUL,env,t2,t1) evd
+ | None -> add_conv_pb (Reduction.CONV,env,t1,t2) evd
+
+(*------------------------------------*
+ * Restricting existing evars *
+ *------------------------------------*)
+
+type 'a update =
+| UpdateWith of 'a
+| NoUpdate
+
+let inst_of_vars sign = Array.map_of_list (fun (id,_,_) -> mkVar id) sign
+
+let restrict_evar_key evd evk filter candidates =
+ match filter, candidates with
+ | None, NoUpdate -> evd, evk
+ | _ ->
+ let evi = Evd.find_undefined evd evk in
+ let oldfilter = evar_filter evi in
+ begin match filter, candidates with
+ | Some filter, NoUpdate when Filter.equal oldfilter filter ->
+ evd, evk
+ | _ ->
+ let filter = match filter with
+ | None -> evar_filter evi
+ | Some filter -> filter in
+ let candidates = match candidates with
+ | NoUpdate -> evi.evar_candidates
+ | UpdateWith c -> Some c in
+ restrict_evar evd evk filter candidates
+ end
+
+(* Restrict an applied evar and returns its restriction in the same context *)
+let restrict_applied_evar evd (evk,argsv) filter candidates =
+ let evd,newevk = restrict_evar_key evd evk filter candidates in
+ let newargsv = match filter with
+ | None -> (* optim *) argsv
+ | Some filter ->
+ let evi = Evd.find evd evk in
+ let subfilter = Filter.compose (evar_filter evi) filter in
+ Filter.filter_array subfilter argsv in
+ evd,(newevk,newargsv)
+
+(* Restrict an evar in the current evar_map *)
+let restrict_evar evd evk filter candidates =
+ fst (restrict_evar_key evd evk filter candidates)
+
+(* Restrict an evar in the current evar_map *)
+let restrict_instance evd evk filter argsv =
+ match filter with None -> argsv | Some filter ->
+ let evi = Evd.find evd evk in
+ Filter.filter_array (Filter.compose (evar_filter evi) filter) argsv
+
+let noccur_evar env evd evk c =
+ let rec occur_rec k c = match kind_of_term c with
+ | Evar (evk',args' as ev') ->
+ (match safe_evar_value evd ev' with
+ | Some c -> occur_rec k c
+ | None ->
+ if Evar.equal evk evk' then raise Occur
+ else Array.iter (occur_rec k) args')
+ | Rel i when i > k ->
+ (match pi2 (Environ.lookup_rel (i-k) env) with
+ | None -> ()
+ | Some b -> occur_rec k (lift i b))
+ | _ -> iter_constr_with_binders succ occur_rec k c
+ in
+ try occur_rec 0 c; true with Occur -> false
+
+(***************************************)
+(* Managing chains of local definitons *)
+(***************************************)
+
+(* 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 =
+ List.fold_right (fun (id,b,c) aliases ->
+ match b with
+ | Some t ->
+ (match kind_of_term 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
+ | _ ->
+ Id.Map.add id [t] aliases)
+ | None -> aliases)
+ sign Id.Map.empty
+
+let compute_rel_aliases var_aliases rels =
+ snd (List.fold_right (fun (_,b,t) (n,aliases) ->
+ (n-1,
+ match b with
+ | Some t ->
+ (match kind_of_term 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
+ | 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
+ | _ ->
+ Int.Map.add n [lift n t] aliases)
+ | None -> aliases))
+ rels (List.length rels,Int.Map.empty))
+
+let make_alias_map env =
+ (* 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 lift_aliases n (var_aliases,rel_aliases as 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
+ | Some a -> a
+ | None -> x
+
+let normalize_alias_var var_aliases id =
+ destVar (normalize_alias (var_aliases,Int.Map.empty) (mkVar id))
+
+let extend_alias (_,b,_) (var_aliases,rel_aliases) =
+ let rel_aliases =
+ Int.Map.fold (fun n l -> Int.Map.add (n+1) (List.map (lift 1) l))
+ rel_aliases Int.Map.empty in
+ let rel_aliases =
+ match b with
+ | Some t ->
+ (match kind_of_term 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
+ | 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
+ | _ ->
+ Int.Map.add 1 [lift 1 t] rel_aliases)
+ | None -> 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
+ | _ ->
+ map_constr_with_full_binders
+ extend_alias expand_vars_in_term_using aliases t
+
+let expand_vars_in_term env = expand_vars_in_term_using (make_alias_map env)
+
+let free_vars_and_rels_up_alias_expansion aliases c =
+ let acc1 = ref Int.Set.empty and acc2 = 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
+ 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
+ let rec frec (aliases,depth) c =
+ match kind_of_term c with
+ | Rel _ | Var _ as ck ->
+ if is_in_cache depth ck then () else begin
+ put_in_cache depth ck;
+ let c = expansion_of_var aliases c in
+ match kind_of_term 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
+ | _ ->
+ iter_constr_with_full_binders
+ (fun d (aliases,depth) -> (extend_alias d aliases,depth+1))
+ frec (aliases,depth) c
+ in
+ frec (aliases,0) c;
+ (!acc1,!acc2)
+
+(********************************)
+(* 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
+ (* 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 remove_instance_local_defs evd evk args =
+ let evi = Evd.find evd evk in
+ let len = Array.length args in
+ let rec aux sign i = match sign with
+ | [] ->
+ let () = assert (i = len) in []
+ | (_, None, _) :: sign ->
+ let () = assert (i < len) in
+ (Array.unsafe_get args i) :: aux sign (succ i)
+ | (_, Some _, _) :: sign ->
+ aux sign (succ i)
+ in
+ aux (evar_filtered_context evi) 0
+
+(* 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 is_unification_pattern_meta env 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
+ | _ -> None
+ else
+ 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
+ let args = remove_instance_local_defs evd evk args in
+ let n = List.length args in
+ match find_unification_pattern_args env (args @ l) t with
+ | Some l -> Some (List.skipn n l)
+ | _ -> None
+ else 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
+ match is_ev with
+ | None -> false
+ | 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
+ | Evar ev -> is_unification_pattern_evar env evd ev l t
+ | _ -> None
+
+(* From a unification problem "?X l = c", build "\x1...xn.(term1 l2)"
+ (pattern unification). It is assumed that l is made of rel's that
+ are distinct and not bound to aliases. *)
+(* It is also assumed that c does not contain metas because metas
+ *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 c' = List.fold_right (fun a c ->
+ let c' = subst_term (lift 1 a) (lift 1 c) in
+ match kind_of_term a with
+ (* Rem: if [a] links to a let-in, do as if it were an assumption *)
+ | Rel n ->
+ let d = map_rel_declaration (lift n) (lookup_rel n env) in
+ mkLambda_or_LetIn d c'
+ | Var 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 *)
+ whd_eta c'
+
+(*****************************************)
+(* Refining/solving unification problems *)
+(*****************************************)
+
+(* Knowing that [Gamma |- ev : T] and that [ev] is applied to [args],
+ * [make_projectable_subst ev args] builds the substitution [Gamma:=args].
+ * If a variable and an alias of it are bound to the same instance, we skip
+ * the alias (we just use eq_constr -- instead of conv --, since anyway,
+ * only instances that are variables -- or evars -- are later considered;
+ * morever, we can bet that similar instances came at some time from
+ * the very same substitution. The removal of aliased duplicates is
+ * useful to ensure the uniqueness of a projection.
+*)
+
+let make_projectable_subst aliases sigma evi args =
+ let sign = evar_filtered_context evi in
+ let evar_aliases = compute_var_aliases sign in
+ let (_,full_subst,cstr_subst) =
+ List.fold_right
+ (fun (id,b,c) (args,all,cstrs) ->
+ match b,args with
+ | None, 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
+ | 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)
+ | Some c, a::rest ->
+ let a = whd_evar sigma a in
+ (match kind_of_term c with
+ | Var id' ->
+ let idc = normalize_alias_var 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
+ (rest,all,cstrs)
+ else
+ (rest,
+ Id.Map.add idc ((a,normalize_alias_opt 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"))
+ sign (Array.rev_to_list args,Id.Map.empty,Constrmap.empty) in
+ (full_subst,cstr_subst)
+
+(*------------------------------------*
+ * operations on the evar constraints *
+ *------------------------------------*)
+
+(* We have a unification problem Σ; Γ |- ?e[u1..uq] = t : s where ?e is not yet
+ * declared in Σ but yet known to be declarable in some context x1:T1..xq:Tq.
+ * [define_evar_from_virtual_equation ... Γ Σ t (x1:T1..xq:Tq) .. (u1..uq) (x1..xq)]
+ * declares x1:T1..xq:Tq |- ?e : s such that ?e[u1..uq] = t holds.
+ *)
+
+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,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 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
+ (evd,whd_evar evd evar_in_sign)
+
+(* We have x1..xq |- ?e1 : Ï„ and had to solve something like
+ * Σ; Γ |- ?e1[u1..uq] = (...\y1 ... \yk ... c), where c is typically some
+ * ?e2[v1..vn], hence flexible. We had to go through k binders and now
+ * virtually have x1..xq, y1'..yk' | ?e1' : Ï„' and the equation
+ * Γ, y1..yk |- ?e1'[u1..uq y1..yk] = c.
+ * [materialize_evar Γ evd k (?e1[u1..uq]) τ'] extends Σ with the declaration
+ * of ?e1' and returns both its instance ?e1'[x1..xq y1..yk] in an extension
+ * of the context of e1 so that e1 can be instantiated by
+ * (...\y1' ... \yk' ... ?e1'[x1..xq y1'..yk']),
+ * and the instance ?e1'[u1..uq y1..yk] so that the remaining equation
+ * ?e1'[u1..uq y1..yk] = c can be registered
+ *
+ * Note that, because invert_definition does not check types, we need to
+ * guess the types of y1'..yn' by inverting the types of y1..yn along the
+ * substitution u1..uq.
+ *)
+
+let materialize_evar define_fun env evd k (evk1,args1) ty_in_env =
+ 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 pi1 (named_context_of_val sign1) in
+ let inst_in_sign = List.map mkVar (Filter.filter_list filter1 ids1) in
+ let (sign2,filter2,inst2_in_env,inst2_in_sign,_,evd,_) =
+ List.fold_right (fun (na,b,t_in_env as d) (sign,filter,inst_in_env,inst_in_sign,env,evd,avoid) ->
+ let id = next_name_away na avoid in
+ let evd,t_in_sign =
+ let s = Retyping.get_sort_of env evd t_in_env in
+ let evd,ty_t_in_sign = refresh_universes ~inferred:true (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 b with
+ | None -> evd,None
+ | Some 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 (id,b_in_sign,t_in_sign) 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))
+ rel_sign
+ (sign1,filter1,Array.to_list args1,inst_in_sign,env1,evd,ids1)
+ in
+ let evd,ev2ty_in_sign =
+ let s = Retyping.get_sort_of env evd ty_in_env in
+ let evd,ty_t_in_sign = refresh_universes ~inferred:true (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,ev2_in_sign =
+ new_evar_instance sign2 evd ev2ty_in_sign ~filter:filter2 ~src inst2_in_sign in
+ let ev2_in_env = (fst (destEvar 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 =
+ let oldfullfilter = evar_filter (Evd.find_undefined evd evk) in
+ let len = Array.length args in
+ Filter.restrict_upon oldfullfilter len (fun i -> p (Array.unsafe_get args i))
+
+(***************)
+(* Unification *)
+
+(* Inverting constructors in instances (common when inferring type of match) *)
+
+let find_projectable_constructor env evd cstr k args cstr_subst =
+ try
+ let l = Constrmap.find cstr cstr_subst in
+ let args = Array.map (lift (-k)) args in
+ let l =
+ 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
+ List.map snd l
+ with Not_found ->
+ []
+
+(* [find_projectable_vars env sigma y subst] finds all vars of [subst]
+ * that project on [y]. It is able to find solutions to the following
+ * two kinds of problems:
+ *
+ * - ?n[...;x:=y;...] = y
+ * - ?n[...;x:=?m[args];...] = y with ?m[args] = y recursively solvable
+ *
+ * (see test-suite/success/Fixpoint.v for an example of application of
+ * the second kind of problem).
+ *
+ * The seek for [y] is up to variable aliasing. In case of solutions that
+ * differ only up to aliasing, the binding that requires the less
+ * steps of alias reduction is kept. At the end, only one solution up
+ * to aliasing is kept.
+ *
+ * [find_projectable_vars] also unifies against evars that themselves mention
+ * [y] and recursively.
+ *
+ * In short, the following situations give the following solutions:
+ *
+ * problem evar ctxt soluce remark
+ * z1; z2:=z1 |- ?ev[z1;z2] = z1 y1:A; y2:=y1 y1 \ thanks to defs kept in
+ * z1; z2:=z1 |- ?ev[z1;z2] = z2 y1:A; y2:=y1 y2 / subst and preferring =
+ * z1; z2:=z1 |- ?ev[z1] = z2 y1:A y1 thanks to expand_var
+ * z1; z2:=z1 |- ?ev[z2] = z1 y1:A y1 thanks to expand_var
+ * z3 |- ?ev[z3;z3] = z3 y1:A; y2:=y1 y2 see make_projectable_subst
+ *
+ * Remark: [find_projectable_vars] assumes that identical instances of
+ * variables in the same set of aliased variables are already removed (see
+ * [make_projectable_subst])
+ *)
+
+type evar_projection =
+| ProjectVar
+| ProjectEvar of existential * evar_info * Id.t * evar_projection
+
+exception NotUnique
+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
+ 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
+
+let rec find_projectable_vars with_evars aliases sigma y subst =
+ let yc = normalize_alias aliases y in
+ let is_projectable idc idcl subst' =
+ (* First test if some [id] aliased to [idc] is bound to [y] in [subst] *)
+ try
+ let id = assoc_up_to_alias sigma aliases y yc idcl in
+ (id,ProjectVar)::subst'
+ with Not_found ->
+ (* Then test if [idc] is (indirectly) bound in [subst] to some evar *)
+ (* projectable on [y] *)
+ if with_evars then
+ let idcl' = List.filter (fun (c,_,id) -> isEvar c) idcl in
+ match idcl' with
+ | [c,_,id] ->
+ begin
+ let (evk,argsv as t) = destEvar 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
+ match l with
+ | [id',p] -> (id,ProjectEvar (t,evi,id',p))::subst'
+ | _ -> subst'
+ end
+ | [] -> subst'
+ | _ -> anomaly (Pp.str "More than one non var in aliases class of evar instance")
+ else
+ subst' in
+ Id.Map.fold is_projectable subst []
+
+(* [filter_solution] checks if one and only one possible projection exists
+ * among a set of solutions to a projection problem *)
+
+let filter_solution = function
+ | [] -> raise Not_found
+ | (id,p)::_::_ -> raise NotUnique
+ | [id,p] -> (mkVar id, p)
+
+let project_with_effects aliases sigma effects t subst =
+ let c, p =
+ filter_solution (find_projectable_vars false aliases sigma t subst) in
+ effects := p :: !effects;
+ c
+
+let rec find_solution_type evarenv = function
+ | (id,ProjectVar)::l -> pi3 (lookup_named id evarenv)
+ | [id,ProjectEvar _] -> (* bugged *) pi3 (lookup_named id evarenv)
+ | (id,ProjectEvar _)::l -> find_solution_type evarenv l
+ | [] -> 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
+ * type is an evar too.
+ *
+ * Note: typing creates new evar problems, which induces a recursive dependency
+ * with [define]. To avoid a too large set of recursive functions, we
+ * pass [define] to [do_projection_effects] as a parameter.
+ *)
+
+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
+ (* TODO: simplify constraints involving evk *)
+ let evd = do_projection_effects define_fun env ty evd p in
+ let ty = whd_betadeltaiota env evd (Lazy.force ty) in
+ if not (isSort 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
+ else
+ evd
+
+(* Assuming Σ; Γ, y1..yk |- c, [invert_arg_from_subst Γ k Σ [x1:=u1..xn:=un] c]
+ * tries to return φ(x1..xn) such that equation φ(u1..un) = c is valid.
+ * The strategy is to imitate the structure of c and then to invert
+ * the variables of c (i.e. rels or vars of Γ) using the algorithm
+ * implemented by project_with_effects/find_projectable_vars.
+ * It returns either a unique solution or says whether 0 or more than
+ * 1 solutions is found.
+ *
+ * Precondition: Σ; Γ, y1..yk |- c /\ Σ; Γ |- u1..un
+ * Postcondition: if φ(x1..xn) is returned then
+ * Σ; Γ, y1..yk |- φ(u1..un) = c /\ x1..xn |- φ(x1..xn)
+ *
+ * The effects correspond to evars instantiated while trying to project.
+ *
+ * [invert_arg_from_subst] is used on instances of evars. Since the
+ * evars are flexible, these instances are potentially erasable. This
+ * is why we don't investigate whether evars in the instances of evars
+ * are unifiable, to the contrary of [invert_definition].
+ *)
+
+type projectibility_kind =
+ | NoUniqueProjection
+ | UniqueProjection of constr * evar_projection list
+
+type projectibility_status =
+ | CannotInvert
+ | Invertible of projectibility_kind
+
+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
+ 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
+ | None -> raise Not_found
+ | Some c -> aux k c in
+ try
+ let c = aux 0 c_in_env_extended_with_k_binders in
+ Invertible (UniqueProjection (c,!effects))
+ with
+ | Not_found -> CannotInvert
+ | NotUnique -> Invertible NoUniqueProjection
+
+let invert_arg fullenv evd aliases k evk subst_in_env_extended_with_k_binders c_in_env_extended_with_k_binders =
+ let res = invert_arg_from_subst evd aliases k subst_in_env_extended_with_k_binders c_in_env_extended_with_k_binders in
+ match res with
+ | Invertible (UniqueProjection (c,_)) when not (noccur_evar fullenv evd evk c)
+ ->
+ CannotInvert
+ | _ ->
+ res
+
+exception NotEnoughInformationToInvert
+
+let extract_unique_projection = function
+| Invertible (UniqueProjection (c,_)) -> c
+| _ ->
+ (* For instance, there are evars with non-invertible arguments and *)
+ (* we cannot arbitrarily restrict these evars before knowing if there *)
+ (* will really be used; it can also be due to some argument *)
+ (* (typically a rel) that is not inversible and that cannot be *)
+ (* inverted either because it is needed for typing the conclusion *)
+ (* of the evar to project *)
+ raise NotEnoughInformationToInvert
+
+let extract_candidates sols =
+ try
+ UpdateWith
+ (List.map (function (id,ProjectVar) -> mkVar id | _ -> raise Exit) sols)
+ with Exit ->
+ NoUpdate
+
+let invert_invertible_arg fullenv evd aliases k (evk,argsv) args' =
+ let evi = Evd.find_undefined evd evk in
+ let subst,_ = make_projectable_subst aliases evd evi argsv in
+ let invert arg =
+ let p = invert_arg fullenv evd aliases k evk subst arg in
+ extract_unique_projection p
+ in
+ Array.map invert args'
+
+(* Redefines an evar with a smaller context (i.e. it may depend on less
+ * variables) such that c becomes closed.
+ * Example: in "fun (x:?1) (y:list ?2[x]) => x = y :> ?3[x,y] /\ x = nil bool"
+ * ?3 <-- ?1 no pb: env of ?3 is larger than ?1's
+ * ?1 <-- list ?2 pb: ?2 may depend on x, but not ?1.
+ * What we do is that ?2 is defined by a new evar ?4 whose context will be
+ * a prefix of ?2's env, included in ?1's env.
+ *
+ * If "hyps |- ?e : T" and "filter" selects a subset hyps' of hyps then
+ * [do_restrict_hyps evd ?e filter] sets ?e:=?e'[hyps'] and returns ?e'
+ * such that "hyps' |- ?e : T"
+ *)
+
+let set_of_evctx l =
+ List.fold_left (fun s (id,_,_) -> Id.Set.add id s) Id.Set.empty l
+
+let filter_effective_candidates 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
+
+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
+ | UpdateWith c -> Some c
+ in
+ match candidates with
+ | None -> NoUpdate
+ | Some l ->
+ let l' = filter_effective_candidates evi filter l in
+ if List.length l = List.length l' && candidates_update = NoUpdate then
+ NoUpdate
+ else
+ UpdateWith l'
+
+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 (id,c,_) = b || Idset.mem id vars || match c with None -> false | Some c -> not (isRel c || isVar c) in
+ let newfilter = Filter.map_along test filter (evar_context evi) in
+ if Filter.equal newfilter (evar_filter evi) then None else Some newfilter
+
+let restrict_hyps evd evk filter candidates =
+ (* What to do with dependencies?
+ Assume we have x:A, y:B(x), z:C(x,y) |- ?e:T(x,y,z) and restrict on y.
+ - If y is in a non-erasable position in C(x,y) (i.e. it is not below an
+ occurrence of x in the hnf of C), then z should be removed too.
+ - If y is in a non-erasable position in T(x,y,z) then the problem is
+ unsolvable.
+ Computing whether y is erasable or not may be costly and the
+ interest for this early detection in practice is not obvious. We let
+ it for future work. In any case, thanks to the use of filters, the whole
+ (unrestricted) context remains consistent. *)
+ let candidates = filter_candidates evd evk (Some filter) candidates in
+ let typablefilter = closure_of_filter evd evk (Some filter) in
+ (typablefilter,candidates)
+
+exception EvarSolvedWhileRestricting of evar_map * 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 [nc],_ ->
+ let evd = Evd.define evk nc evd in
+ raise (EvarSolvedWhileRestricting (evd,whd_evar evd (mkEvar ev)))
+ | NoUpdate, None -> evd,ev
+ | _ -> restrict_applied_evar evd ev filter candidates
+
+(* [postpone_non_unique_projection] postpones equation of the form ?e[?] = c *)
+(* ?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 filter =
+ restrict_upon_filter evd evk
+ (* Keep only variables that occur in rhs *)
+ (* This is not safe: is the variable is a local def, its body *)
+ (* may contain references to variables that are removed, leading to *)
+ (* an ill-formed context. We would actually need a notion of filter *)
+ (* 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)
+ argsv in
+ let filter = closure_of_filter evd evk filter in
+ let candidates = extract_candidates sols in
+ match candidates with
+ | NoUpdate ->
+ (* We made an approximation by not expanding a local definition *)
+ let evd,ev = restrict_applied_evar evd ev filter NoUpdate in
+ let pb = (pbty,env,mkEvar ev,rhs) in
+ add_conv_oriented_pb pb evd
+ | UpdateWith c ->
+ restrict_evar evd evk filter (UpdateWith c)
+
+(* [solve_evar_evar f Γ Σ ?e1[u1..un] ?e2[v1..vp]] applies an heuristic
+ * to solve the equation Σ; Γ ⊢ ?e1[u1..un] = ?e2[v1..vp]:
+ * - if there are at most one φj for each vj s.t. vj = φj(u1..un),
+ * we first restrict ?e2 to the subset v_k1..v_kq of the vj that are
+ * inversible and we set ?e1[x1..xn] := ?e2[φk1(x1..xn)..φkp(x1..xn)]
+ * (this is a case of pattern-unification)
+ * - symmetrically if there are at most one ψj for each uj s.t.
+ * uj = ψj(v1..vp),
+ * - otherwise, each position i s.t. ui does not occur in v1..vp has to
+ * be restricted and similarly for the vi, and we leave the equation
+ * as an open equation (performed by [postpone_evar])
+ *
+ * Warning: the notion of unique φj is relative to some given class
+ * of unification problems
+ *
+ * Note: argument f is the function used to instantiate evars.
+ *)
+
+let are_canonical_instances args1 args2 env =
+ let n1 = Array.length args1 in
+ let n2 = Array.length args2 in
+ let rec aux n = function
+ | (id,_,c)::sign
+ when n < n1 && isVarId id args1.(n) && isVarId id args2.(n) ->
+ aux (n+1) sign
+ | [] ->
+ let rec aux2 n =
+ Int.equal n n1 ||
+ (isRelN (n1-n) args1.(n) && isRelN (n1-n) args2.(n) && aux2 (n+1))
+ in aux2 n
+ | _ -> false in
+ Int.equal n1 n2 && aux 0 (named_context env)
+
+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
+ | Success evd -> Some (c,evd)
+ | UnifFailure _ -> None
+
+(* [restrict_candidates ... filter ev1 ev2] restricts the candidates
+ of ev1, removing those not compatible with the filter, as well as
+ those not convertible to some candidate of ev2 *)
+
+exception DoesNotPreserveCandidateRestriction
+
+let restrict_candidates conv_algo env evd filter1 (evk1,argsv1) (evk2,argsv2) =
+ let evi1 = Evd.find evd evk1 in
+ let evi2 = Evd.find evd evk2 in
+ match evi1.evar_candidates, evi2.evar_candidates with
+ | _, 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.filter (fun c1 ->
+ let c1' = instantiate_evar_array evi1 c1 argsv1 in
+ let filter c2 =
+ let compatibility = filter_compatible_candidates conv_algo env evd evi2 argsv2 c1' c2 in
+ match compatibility with
+ | None -> false
+ | Some _ -> true
+ in
+ let filtered = List.filter filter l2 in
+ match filtered with [] -> false | _ -> true) l1 in
+ if Int.equal (List.length l1) (List.length l1') then NoUpdate
+ else UpdateWith l1'
+
+exception CannotProject of evar_map * 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
+ FV(u[x1:=t1..xn:=tn]) are in the set U?
+ - If ti is a variable, it has to be in U.
+ - If ti is a constructor, its parameters cannot be erased even if u
+ matches on it, so we have to discard ti if the parameters
+ contain variables not in U.
+ - If ti is rigid, we have to discard it if it contains variables in U.
+
+ Note: when restricting as part of an equation ?n[x1:=t1..xn:=tn] = ?m[...]
+ then, occurrences of ?m in the ti can be seen, like variables, as occurrences
+ of subterms to eventually discard so as to be allowed to keep ti.
+*)
+
+let rec is_constrainable_in top k (ev,(fv_rels,fv_ids) as g) t =
+ let f,args = decompose_app_vect t in
+ match kind_of_term 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 *)
+ else
+ let params = fst (Array.chop n args) in
+ Array.for_all (is_constrainable_in false k g) params
+ | Ind _ -> Array.for_all (is_constrainable_in false k g) args
+ | Prod (_,t1,t2) -> is_constrainable_in false k g t1 && is_constrainable_in false k g t2
+ | Evar (ev',_) -> top || not (Evar.equal ev' ev) (*If ev' needed, one may also try to restrict it*)
+ | Var id -> Id.Set.mem id fv_ids
+ | Rel n -> n <= k || Int.Set.mem n fv_rels
+ | Sort _ -> true
+ | _ -> (* We don't try to be more clever *) true
+
+let has_constrainable_free_vars evd aliases k ev (fv_rels,fv_ids as fvs) t =
+ let t = expansion_of_var aliases t in
+ match kind_of_term t with
+ | Var id -> Id.Set.mem id fv_ids
+ | Rel n -> n <= k || Int.Set.mem n fv_rels
+ | _ -> is_constrainable_in true k (ev,fvs) t
+
+let ensure_evar_independent g env evd (evk1,argsv1 as ev1) (evk2,argsv2 as ev2)=
+ let filter1 =
+ restrict_upon_filter evd evk1 (noccur_evar env evd evk2) argsv1
+ in
+ let candidates1 = restrict_candidates g env evd filter1 ev1 ev2 in
+ let evd,(evk1,_ as ev1) = do_restrict_hyps evd ev1 filter1 candidates1 in
+ let filter2 =
+ restrict_upon_filter evd evk2 (noccur_evar env evd evk1) argsv2
+ in
+ let candidates2 = restrict_candidates g env evd filter2 ev2 ev1 in
+ let evd,ev2 = do_restrict_hyps evd ev2 filter2 candidates2 in
+ evd,ev1,ev2
+
+exception EvarSolvedOnTheFly of evar_map * 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 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 filter1 = restrict_upon_filter evd evk1
+ (has_constrainable_free_vars evd aliases k2 evk2 fvs2)
+ argsv1 in
+ let candidates1 =
+ try restrict_candidates g env evd filter1 ev1 ev2
+ with DoesNotPreserveCandidateRestriction ->
+ let evd,ev1' = do_restrict_hyps evd ev1 filter1 NoUpdate in
+ raise (CannotProject (evd,ev1')) in
+ let evd,(evk1',args1 as ev1') =
+ try do_restrict_hyps evd ev1 filter1 candidates1
+ with EvarSolvedWhileRestricting (evd,ev1) ->
+ raise (EvarSolvedOnTheFly (evd,ev1)) in
+ (* Only try pruning on variable substitutions, postpone otherwise. *)
+ (* Rules out non-linear instances. *)
+ if Option.is_empty pbty && is_unification_pattern_pure_evar env evd ev2 (mkEvar ev1) then
+ try
+ evd,mkEvar (evk1',invert_invertible_arg env evd aliases k2 ev2 args1)
+ with NotEnoughInformationToInvert ->
+ raise (CannotProject (evd,ev1'))
+ else
+ raise (CannotProject (evd,ev1'))
+
+exception IllTypedInstance of env * types * types
+
+let check_evar_instance evd evk1 body conv_algo =
+ let evi = Evd.find evd evk1 in
+ let evenv = evar_env evi in
+ (* FIXME: The body might be ill-typed when this is called from w_merge *)
+ (* 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"
+ in
+ match conv_algo evenv evd Reduction.CUMUL ty evi.evar_concl with
+ | Success evd -> evd
+ | UnifFailure _ -> raise (IllTypedInstance (evenv,ty,evi.evar_concl))
+
+let solve_evar_evar_l2r f g env evd aliases pbty ev1 (evk2,_ as ev2) =
+ try
+ let evd,body = project_evar_on_evar g env evd aliases 0 pbty ev1 ev2 in
+ let evd' = Evd.define evk2 body evd in
+ check_evar_instance evd' evk2 body g
+ with EvarSolvedOnTheFly (evd,c) ->
+ f env evd pbty ev2 c
+
+let opp_problem = function None -> None | Some b -> Some (not b)
+
+let preferred_orientation evd evk1 evk2 =
+ let _,src1 = (Evd.find_undefined evd evk1).evar_source in
+ let _,src2 = (Evd.find_undefined evd evk2).evar_source in
+ (* This is a heuristic useful for program to work *)
+ match src1,src2 with
+ | Evar_kinds.QuestionMark _, _ -> true
+ | _,Evar_kinds.QuestionMark _ -> false
+ | _ -> true
+
+let solve_evar_evar_aux f g env evd pbty (evk1,args1 as ev1) (evk2,args2 as ev2) =
+ let aliases = make_alias_map env in
+ if preferred_orientation evd evk1 evk2 then
+ try solve_evar_evar_l2r f g env evd aliases (opp_problem pbty) ev2 ev1
+ with CannotProject (evd,ev2) ->
+ try solve_evar_evar_l2r f g env evd aliases pbty ev1 ev2
+ with CannotProject (evd,ev1) ->
+ add_conv_oriented_pb (pbty,env,mkEvar ev1,mkEvar ev2) evd
+ else
+ try solve_evar_evar_l2r f g env evd aliases pbty ev1 ev2
+ with CannotProject (evd,ev1) ->
+ try solve_evar_evar_l2r f g env evd aliases (opp_problem pbty) ev2 ev1
+ with CannotProject (evd,ev2) ->
+ add_conv_oriented_pb (pbty,env,mkEvar ev1,mkEvar ev2) evd
+
+let solve_evar_evar ?(force=false) f g env evd pbty ev1 ev2 =
+ let (evd,(evk1,args1 as ev1),(evk2,args2 as ev2)),pbty =
+ (* If an evar occurs in the instance of the other evar and the
+ use of an heuristic is forced, we restrict *)
+ if force then ensure_evar_independent g env evd ev1 ev2, None
+ else (evd,ev1,ev2),pbty in
+ let evi = Evd.find evd evk1 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 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 ui, uj = univ_of_sort i, univ_of_sort j in
+ if i == j || Evd.check_eq evd ui uj
+ then (* Shortcut, i = j *)
+ evd
+ else if Evd.check_leq evd ui uj then
+ let t2 = it_mkProd_or_LetIn (mkSort i) ctx2 in
+ downcast evk2 t2 evd
+ else if Evd.check_leq evd uj ui then
+ let t1 = it_mkProd_or_LetIn (mkSort j) ctx1 in
+ downcast evk1 t1 evd
+ else
+ let evd, k = Evd.new_sort_variable univ_flexible_alg evd in
+ let t1 = it_mkProd_or_LetIn (mkSort k) ctx1 in
+ let t2 = it_mkProd_or_LetIn (mkSort k) ctx2 in
+ let evd = Evd.set_leq_sort env (Evd.set_leq_sort env evd k i) k j in
+ downcast evk2 t2 (downcast evk1 t1 evd)
+ with Reduction.NotArity ->
+ evd in
+ solve_evar_evar_aux f g env evd pbty ev1 ev2
+
+type conv_fun =
+ env -> evar_map -> conv_pb -> constr -> constr -> unification_result
+
+type conv_fun_bool =
+ env -> evar_map -> conv_pb -> constr -> 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
+ * that don't unify are discarded (i.e. ?e is redefined so that it does not
+ * depend on these args). *)
+
+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
+ (* Filter and restrict if needed *)
+ let args = Array.map2 (fun a1 a2 -> (a1, a2)) argsv1 argsv2 in
+ let untypedfilter =
+ restrict_upon_filter evd evk
+ (fun (a1,a2) -> conv_algo env evd Reduction.CONV a1 a2) args in
+ let candidates = filter_candidates evd evk untypedfilter NoUpdate in
+ let filter = closure_of_filter evd evk untypedfilter in
+ let evd,ev1 = restrict_applied_evar evd (evk,argsv1) filter candidates in
+ if Evar.equal (fst ev1) evk && can_drop then (* No refinement *) evd else
+ (* either progress, or not allowed to drop, e.g. to preserve possibly *)
+ (* informative equations such as ?e[x:=?y]=?e[x:=?y'] where we don't know *)
+ (* if e can depend on x until ?y is not resolved, or, conversely, we *)
+ (* don't know if ?y has to be unified with ?y, until e is resolved *)
+ let argsv2 = restrict_instance evd evk filter argsv2 in
+ let ev2 = (fst ev1,argsv2) in
+ (* Leave a unification problem *)
+ add_conv_oriented_pb (pbty,env,mkEvar ev1,mkEvar ev2) evd
+
+(* If the evar can be instantiated by a finite set of candidates known
+ in advance, we check which of them apply *)
+
+exception NoCandidates
+exception IncompatibleCandidates
+
+let solve_candidates conv_algo env evd (evk,argsv) rhs =
+ let evi = Evd.find evd evk in
+ match evi.evar_candidates with
+ | None -> raise NoCandidates
+ | Some l ->
+ let l' =
+ List.map_filter
+ (filter_compatible_candidates conv_algo env evd evi argsv rhs) 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 Evd.define evk c evd else evd
+ | l when List.length l < List.length l' ->
+ let candidates = List.map fst l in
+ restrict_evar evd evk None (UpdateWith candidates)
+ | l -> evd
+
+(* We try to instantiate the evar assuming the body won't depend
+ * on arguments that are not Rels or Vars, or appearing several times
+ * (i.e. we tackle a generalization of Miller-Pfenning patterns unification)
+ *
+ * 1) Let "env |- ?ev[hyps:=args] = rhs" be the unification problem
+ * 2) We limit it to a patterns unification problem "env |- ev[subst] = rhs"
+ * where only Rel's and Var's are relevant in subst
+ * 3) We recur on rhs, "imitating" the term, and failing if some Rel/Var is
+ * not in the scope of ?ev. For instance, the problem
+ * "y:nat |- ?x[] = y" where "|- ?1:nat" is not satisfiable because
+ * ?1 would be instantiated by y which is not in the scope of ?1.
+ * 4) We try to "project" the term if the process of imitation fails
+ * and that only one projection is possible
+ *
+ * Note: we don't assume rhs in normal form, it may fail while it would
+ * have succeeded after some reductions.
+ *
+ * This is the work of [invert_definition Γ Σ ?ev[hyps:=args] c]
+ * Precondition: Σ; Γ, y1..yk |- c /\ Σ; Γ |- u1..un
+ * Postcondition: if φ(x1..xn) is returned then
+ * Σ; Γ, y1..yk |- φ(u1..un) = c /\ x1..xn |- φ(x1..xn)
+ *)
+
+exception NotInvertibleUsingOurAlgorithm of constr
+exception NotEnoughInformationToProgress of (Id.t * evar_projection) list
+exception NotEnoughInformationEvarEvar of constr
+exception OccurCheckIn of evar_map * 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 evdref = ref evd in
+ let progress = ref false in
+ let evi = Evd.find evd evk in
+ let subst,cstr_subst = make_projectable_subst aliases evd evi argsv in
+
+ (* Projection *)
+ let project_variable t =
+ (* Evar/Var problem: unifiable iff variable projectable from ev subst *)
+ try
+ let sols = find_projectable_vars true aliases !evdref t subst in
+ 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)
+ in
+ let ty = lazy (Retyping.get_type_of env !evdref 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)
+ | NotUniqueInType sols ->
+ if not !progress then
+ raise (NotEnoughInformationToProgress sols);
+ (* No unique projection but still restrict to where it is possible *)
+ (* materializing is necessary, but is restricting useful? *)
+ let ty = find_solution_type (evar_filtered_env evi) sols in
+ 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 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
+ | UpdateWith _ ->
+ restrict_evar evd evk' filter candidates
+ in
+ evdref := evd;
+ evar in
+
+ let rec imitate (env',k as envk) t =
+ let t = whd_evar !evdref t in
+ match kind_of_term t with
+ | Rel i when i>k ->
+ (match pi2 (Environ.lookup_rel (i-k) env') with
+ | None -> project_variable (mkRel (i-k))
+ | Some b ->
+ try project_variable (mkRel (i-k))
+ with NotInvertibleUsingOurAlgorithm _ -> imitate envk (lift i b))
+ | Var id ->
+ (match pi2 (Environ.lookup_named id env') with
+ | None -> project_variable t
+ | Some b ->
+ try project_variable t
+ with NotInvertibleUsingOurAlgorithm _ -> imitate envk b)
+ | LetIn (na,b,u,c) ->
+ imitate envk (subst1 b c)
+ | Evar (evk',args' as ev') ->
+ if Evar.equal evk evk' then raise (OccurCheckIn (evd,rhs));
+ (* Evar/Evar problem (but left evar is virtual) *)
+ let aliases = lift_aliases k aliases in
+ (try
+ let ev = (evk,Array.map (lift k) argsv) in
+ let evd,body = project_evar_on_evar conv_algo env' !evdref aliases k None ev' ev in
+ evdref := evd;
+ body
+ with
+ | EvarSolvedOnTheFly (evd,t) -> evdref:=evd; imitate envk t
+ | CannotProject (evd,ev') ->
+ if not !progress then
+ raise (NotEnoughInformationEvarEvar t);
+ (* Make the virtual left evar real *)
+ let ty = get_type_of env' evd t in
+ let (evd,evar'',ev'') =
+ materialize_evar (evar_define conv_algo ~choose) env' evd k ev ty in
+ (* materialize_evar may instantiate ev' by another evar; adjust it *)
+ let (evk',args' as ev') = normalize_evar evd ev' in
+ let evd =
+ (* Try to project (a restriction of) the left evar ... *)
+ try
+ let evd,body = project_evar_on_evar conv_algo env' evd aliases 0 None ev'' ev' in
+ let evd = Evd.define evk' body evd in
+ check_evar_instance evd evk' body conv_algo
+ with
+ | EvarSolvedOnTheFly _ -> assert false (* ev has no candidates *)
+ | CannotProject (evd,ev'') ->
+ (* ... or postpone the problem *)
+ add_conv_oriented_pb (None,env',mkEvar ev'',mkEvar ev') evd in
+ evdref := evd;
+ evar'')
+ | _ ->
+ 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 ->
+ (* 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 *)
+ (* alternative inversion of the subterms of the constructor, etc)*)
+ (match find_projectable_constructor env evd cstr k args cstr_subst with
+ | _::_ as l -> Some (List.map mkVar l)
+ | _ -> None)
+ | _ -> None
+ with
+ | Some l ->
+ let ty = get_type_of env' !evdref t in
+ let candidates =
+ try
+ let t =
+ map_constr_with_full_binders (fun d (env,k) -> push_rel d env, k+1)
+ imitate envk t in
+ t::l
+ with e when Errors.noncritical e -> l in
+ (match candidates with
+ | [x] -> x
+ | _ ->
+ let (evd,evar'',ev'') =
+ materialize_evar (evar_define conv_algo ~choose) env' !evdref k ev ty in
+ evdref := restrict_evar evd (fst ev'') None (UpdateWith candidates);
+ 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)
+ imitate envk t in
+
+ let _fast rhs =
+ let filter_ctxt = evar_filtered_context evi in
+ let names = ref Idset.empty in
+ let rec is_id_subst ctxt s =
+ match ctxt, s with
+ | ((id, _, _) :: ctxt'), (c :: s') ->
+ names := Idset.add id !names;
+ isVarId 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 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 rec is_id_subst ctxt s =
+ match ctxt, s with
+ | ((id, _, _) :: ctxt'), (c :: s') ->
+ names := Idset.add id !names;
+ isVarId 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
+ in
+ let body =
+ if fast rhs then nf_evar evd rhs
+ else imitate (env,0) rhs
+ in (!evdref,body)
+
+(* [define] tries to solve the problem "?ev[args] = rhs" when "?ev" is
+ * an (uninstantiated) evar such that "hyps |- ?ev : typ". Otherwise said,
+ * [define] tries to find an instance lhs such that
+ * "lhs [hyps:=args]" unifies to rhs. The term "lhs" must be closed in
+ * context "hyps" and not referring to itself.
+ *)
+
+and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs =
+ match kind_of_term rhs with
+ | Evar (evk2,argsv2 as ev2) ->
+ if Evar.equal evk evk2 then
+ solve_refl ~can_drop:choose
+ (test_success conv_algo) env evd pbty evk argsv argsv2
+ else
+ solve_evar_evar ~force:choose
+ (evar_define conv_algo) conv_algo env evd pbty ev ev2
+ | _ ->
+ try solve_candidates conv_algo env evd 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;
+ (* invert_definition may have instantiate some evars of rhs with evk *)
+ (* so we recheck acyclicity *)
+ if occur_evar evk body then raise (OccurCheckIn (evd',body));
+ (* needed only if an inferred type *)
+ let evd', body = refresh_universes pbty env evd' body in
+(* Cannot strictly type instantiations since the unification algorithm
+ * does not unify applications from left to right.
+ * e.g problem f x == g y yields x==y and f==g (in that order)
+ * Another problem is that type variables are evars of type Type
+ let _ =
+ try
+ let env = evar_filtered_env evi in
+ let ty = evi.evar_concl in
+ Typing.check env evd' body ty
+ with e ->
+ msg_info
+ (str "Ill-typed evar instantiation: " ++ fnl() ++
+ pr_evar_map evd' ++ fnl() ++
+ str "----> " ++ int ev ++ str " := " ++
+ print_constr body);
+ raise e in*)
+ let evd' = check_evar_instance evd' evk body conv_algo in
+ Evd.define evk body evd'
+ with
+ | NotEnoughInformationToProgress sols ->
+ postpone_non_unique_projection env evd pbty ev sols rhs
+ | NotEnoughInformationEvarEvar t ->
+ add_conv_oriented_pb (pbty,env,mkEvar ev,t) evd
+ | NotInvertibleUsingOurAlgorithm _ | MetaOccurInBodyInternal as e ->
+ raise e
+ | OccurCheckIn (evd,rhs) ->
+ (* last chance: rhs actually reduces to ev *)
+ let c = whd_betadeltaiota env evd rhs in
+ match kind_of_term 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
+ | _ ->
+ raise (OccurCheckIn (evd,rhs))
+
+(* This code (i.e. solve_pb, etc.) takes a unification
+ * problem, and tries to solve it. If it solves it, then it removes
+ * all the conversion problems, and re-runs conversion on each one, in
+ * the hopes that the new solution will aid in solving them.
+ *
+ * The kinds of problems it knows how to solve are those in which
+ * the usable arguments of an existential var are all themselves
+ * universal variables.
+ * The solution to this problem is to do renaming for the Var's,
+ * to make them match up with the Var's which are found in the
+ * hyps of the existential, to do a "pop" for each Rel which is
+ * not an argument of the existential, and a subst1 for each which
+ * is, again, with the corresponding variable. This is done by
+ * define
+ *
+ * Thus, we take the arguments of the existential which we are about
+ * to assign, and zip them with the identifiers in the hypotheses.
+ * Then, we process all the Var's in the arguments, and sort the
+ * Rel's into ascending order. Then, we just march up, doing
+ * subst1's and pop's.
+ *
+ * NOTE: We can do this more efficiently for the relative arguments,
+ * by building a long substituend by hand, but this is a pain in the
+ * 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 reconsider_conv_pbs conv_algo evd =
+ let (evd,pbs) = extract_changed_conv_pbs evd status_changed 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
+ | Success _ as x -> x
+ | UnifFailure (i,e) -> UnifFailure (i,CannotSolveConstraint (x,e)))
+ | UnifFailure _ as x -> x)
+ (Success evd)
+ pbs
+
+(* Tries to solve problem t1 = t2.
+ * Precondition: t1 is an uninstantiated evar
+ * Returns an optional list of evars that were instantiated, or None
+ * if the problem couldn't be solved. *)
+
+(* Rq: uncomplete algorithm if pbty = CONV_X_LEQ ! *)
+let solve_simple_eqn conv_algo ?(choose=false) env evd (pbty,(evk1,args1 as ev1),t2) =
+ try
+ let t2 = whd_betaiota evd t2 in (* includes whd_evar *)
+ let evd = evar_define conv_algo ~choose env evd pbty ev1 t2 in
+ reconsider_conv_pbs conv_algo evd
+ with
+ | NotInvertibleUsingOurAlgorithm t ->
+ UnifFailure (evd,NotClean (ev1,env,t))
+ | OccurCheckIn (evd,rhs) ->
+ UnifFailure (evd,OccurCheck (evk1,rhs))
+ | MetaOccurInBodyInternal ->
+ UnifFailure (evd,MetaOccurInBody evk1)
+ | IllTypedInstance (env,t,u) ->
+ UnifFailure (evd,InstanceNotSameType (evk1,env,t,u))
+ | IncompatibleCandidates ->
+ UnifFailure (evd,ConversionFailed (env,mkEvar ev1,t2))
+
diff --git a/pretyping/evarsolve.mli b/pretyping/evarsolve.mli
new file mode 100644
index 00000000..21d97609
--- /dev/null
+++ b/pretyping/evarsolve.mli
@@ -0,0 +1,74 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Term
+open Evd
+open Environ
+
+type unification_result =
+ | Success of evar_map
+ | UnifFailure of evar_map * Pretype_errors.unification_error
+
+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
+
+(** [evar_define choose env ev c] try to instantiate [ev] with [c] (typed in [env]),
+ possibly solving related unification problems, possibly leaving open
+ some problems that cannot be solved in a unique way (except if choose is
+ true); fails if the instance is not valid for the given [ev] *)
+
+type conv_fun =
+ env -> evar_map -> conv_pb -> constr -> constr -> unification_result
+
+type conv_fun_bool =
+ env -> evar_map -> conv_pb -> constr -> constr -> bool
+
+val evar_define : conv_fun -> ?choose:bool -> env -> evar_map ->
+ bool option -> existential -> constr -> evar_map
+
+val refresh_universes : ?inferred:bool -> ?onlyalg:bool (* Only algebraic universes *) ->
+ bool option (* direction: true for levels lower than the existing levels *) ->
+ 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
+
+val solve_evar_evar : ?force:bool ->
+ (env -> evar_map -> bool option -> existential -> constr -> evar_map) ->
+ conv_fun ->
+ env -> evar_map -> bool option -> existential -> existential -> evar_map
+
+val solve_simple_eqn : conv_fun -> ?choose:bool -> env -> evar_map ->
+ bool option * existential * constr -> unification_result
+
+val reconsider_conv_pbs : conv_fun -> evar_map -> unification_result
+
+val is_unification_pattern_evar : env -> evar_map -> existential -> constr list ->
+ constr -> constr list option
+
+val is_unification_pattern : env * int -> evar_map -> constr -> constr list ->
+ constr -> constr list option
+
+val solve_pattern_eqn : env -> constr list -> constr -> constr
+
+val noccur_evar : env -> evar_map -> Evar.t -> constr -> bool
+
+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
+
+val remove_instance_local_defs :
+ evar_map -> existential_key -> constr array -> constr list
+
+val get_type_of_refresh :
+ ?polyprop:bool -> ?lax:bool -> env -> evar_map -> constr -> evar_map * types
diff --git a/pretyping/evarutil.ml b/pretyping/evarutil.ml
index 05b7e443..d286b98e 100644
--- a/pretyping/evarutil.ml
+++ b/pretyping/evarutil.ml
@@ -1,25 +1,48 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Errors
open Util
open Pp
open Names
-open Univ
open Term
+open Vars
+open Context
open Termops
open Namegen
-open Sign
open Pre_env
open Environ
open Evd
open Reductionops
open Pretype_errors
-open Retyping
+
+(** Combinators *)
+
+let evd_comb0 f evdref =
+ let (evd',x) = f !evdref in
+ evdref := evd';
+ x
+
+let evd_comb1 f evdref x =
+ let (evd',y) = f !evdref x in
+ evdref := evd';
+ y
+
+let evd_comb2 f evdref x y =
+ let (evd',z) = f !evdref x y in
+ evdref := evd';
+ z
+
+let e_new_global evdref x =
+ evd_comb1 (Evd.fresh_global (Global.env())) evdref x
+
+let new_global evd x =
+ Evd.fresh_global (Global.env()) evd x
(****************************************************)
(* Expanding/testing/exposing existential variables *)
@@ -37,48 +60,80 @@ let rec flush_and_check_evars sigma c =
| Some c -> flush_and_check_evars sigma c)
| _ -> map_constr (flush_and_check_evars sigma) c
-let nf_evar = Pretype_errors.nf_evar
-let j_nf_evar = Pretype_errors.j_nf_evar
-let jl_nf_evar = Pretype_errors.jl_nf_evar
-let jv_nf_evar = Pretype_errors.jv_nf_evar
-let tj_nf_evar = Pretype_errors.tj_nf_evar
+(* let nf_evar_key = Profile.declare_profile "nf_evar" *)
+(* let nf_evar = Profile.profile2 nf_evar_key Reductionops.nf_evar *)
+let nf_evar = Reductionops.nf_evar
+let j_nf_evar sigma j =
+ { uj_val = nf_evar sigma j.uj_val;
+ uj_type = nf_evar sigma j.uj_type }
+let j_nf_betaiotaevar sigma j =
+ { uj_val = nf_evar sigma j.uj_val;
+ uj_type = Reductionops.nf_betaiota sigma j.uj_type }
+let jl_nf_evar sigma jl = List.map (j_nf_evar sigma) jl
+let jv_nf_betaiotaevar sigma jl =
+ Array.map (j_nf_betaiotaevar sigma) jl
+let jv_nf_evar sigma = Array.map (j_nf_evar sigma)
+let tj_nf_evar sigma {utj_val=v;utj_type=t} =
+ {utj_val=nf_evar sigma v;utj_type=t}
+
+let env_nf_evar sigma env =
+ process_rel_context
+ (fun d e -> push_rel (map_rel_declaration (nf_evar sigma) d) e) env
+
+let env_nf_betaiotaevar sigma env =
+ process_rel_context
+ (fun d e ->
+ push_rel (map_rel_declaration (Reductionops.nf_betaiota sigma) d) e) env
+
+let nf_evars_universes evm =
+ Universes.nf_evars_and_universes_opt_subst (Reductionops.safe_evar_value evm)
+ (Evd.universe_subst evm)
+
+let nf_evars_and_universes evm =
+ let evm = Evd.nf_constraints evm in
+ evm, nf_evars_universes evm
+
+let e_nf_evars_and_universes evdref =
+ evdref := Evd.nf_constraints !evdref;
+ nf_evars_universes !evdref, Evd.universe_subst !evdref
+
+let nf_evar_map_universes evm =
+ let evm = Evd.nf_constraints evm in
+ let subst = Evd.universe_subst evm in
+ if Univ.LMap.is_empty subst then evm, nf_evar evm
+ else
+ let f = nf_evars_universes evm in
+ Evd.raw_map (fun _ -> map_evar_info f) evm, f
let nf_named_context_evar sigma ctx =
- Sign.map_named_context (Reductionops.nf_evar sigma) ctx
+ Context.map_named_context (nf_evar sigma) ctx
let nf_rel_context_evar sigma ctx =
- Sign.map_rel_context (Reductionops.nf_evar sigma) ctx
+ Context.map_rel_context (nf_evar sigma) ctx
let nf_env_evar sigma env =
let nc' = nf_named_context_evar sigma (Environ.named_context env) in
let rel' = nf_rel_context_evar sigma (Environ.rel_context env) in
push_rel_context rel' (reset_with_named_context (val_of_named_context nc') env)
-let nf_evar_info evc info =
- { info with
- evar_concl = Reductionops.nf_evar evc info.evar_concl;
- evar_hyps = map_named_val (Reductionops.nf_evar evc) info.evar_hyps;
- evar_body = match info.evar_body with
- | Evar_empty -> Evar_empty
- | Evar_defined c -> Evar_defined (Reductionops.nf_evar evc c) }
-let nf_evars evm =
- Evd.fold
- (fun ev evi evm' -> Evd.add evm' ev (nf_evar_info evm evi))
- evm Evd.empty
-
-let nf_evars_undefined evm =
- Evd.fold_undefined
- (fun ev evi evm' -> Evd.add evm' ev (nf_evar_info evm evi))
- evm (defined_evars evm)
-
-let nf_evar_map evd = Evd.evars_reset_evd (nf_evars evd) evd
-let nf_evar_map_undefined evd = Evd.evars_reset_evd (nf_evars_undefined evd) evd
+let nf_evar_info evc info = map_evar_info (nf_evar evc) info
+
+let nf_evar_map evm =
+ Evd.raw_map (fun _ evi -> nf_evar_info evm evi) evm
+
+let nf_evar_map_undefined evm =
+ Evd.raw_map_undefined (fun _ evi -> nf_evar_info evm evi) evm
(*-------------------*)
(* Auxiliary functions for the conversion algorithms modulo evars
*)
-let has_undefined_evars_or_sorts evd t =
+(* A probably faster though more approximative variant of
+ [has_undefined (nf_evar c)]: instances are not substituted and
+ maybe an evar occurs in an instance and it would disappear by
+ instantiation *)
+
+let has_undefined_evars evd t =
let rec has_ev t =
match kind_of_term t with
| Evar (ev,args) ->
@@ -87,13 +142,12 @@ let has_undefined_evars_or_sorts evd t =
has_ev c; Array.iter has_ev args
| Evar_empty ->
raise NotInstantiatedEvar)
- | Sort s when is_sort_variable evd s -> raise Not_found
| _ -> iter_constr has_ev t in
try let _ = has_ev t in false
with (Not_found | NotInstantiatedEvar) -> true
let is_ground_term evd t =
- not (has_undefined_evars_or_sorts evd t)
+ not (has_undefined_evars evd t)
let is_ground_env evd env =
let is_ground_decl = function
@@ -101,9 +155,16 @@ let is_ground_env evd env =
| _ -> true in
List.for_all is_ground_decl (rel_context env) &&
List.for_all is_ground_decl (named_context env)
+
(* Memoization is safe since evar_map and environ are applicative
structures *)
-let is_ground_env = memo1_2 is_ground_env
+let memo f =
+ let m = ref None in
+ fun x y -> match !m with
+ | Some (x', y', r) when x == x' && y == y' -> r
+ | _ -> let r = f x y in m := Some (x, y, r); r
+
+let is_ground_env = memo is_ground_env
(* Return the head evar if any *)
@@ -125,35 +186,27 @@ let head_evar =
let whd_head_evar_stack sigma c =
let rec whrec (c, l as s) =
match kind_of_term c with
- | Evar (evk,args as ev) when Evd.is_defined sigma evk
- -> whrec (existential_value sigma ev, l)
+ | Evar (evk,args as ev) ->
+ let v =
+ try Some (existential_value sigma ev)
+ with NotInstantiatedEvar | Not_found -> None in
+ begin match v with
+ | None -> s
+ | Some c -> whrec (c, l)
+ end
| Cast (c,_,_) -> whrec (c, l)
- | App (f,args) -> whrec (f, Array.fold_right (fun a l -> a::l) args l)
+ | App (f,args) -> whrec (f, args :: l)
| _ -> s
in
whrec (c, [])
-let whd_head_evar sigma c = applist (whd_head_evar_stack sigma c)
-
-let noccur_evar env evd evk c =
- let rec occur_rec k c = match kind_of_term c with
- | Evar (evk',args' as ev') ->
- (match safe_evar_value evd ev' with
- | Some c -> occur_rec k c
- | None ->
- if evk = evk' then raise Occur else Array.iter (occur_rec k) args')
- | Rel i when i > k ->
- (match pi2 (Environ.lookup_rel (i-k) env) with
- | None -> ()
- | Some b -> occur_rec k (lift i b))
- | _ -> iter_constr_with_binders succ occur_rec k c
- in
- try occur_rec 0 c; true with Occur -> false
-
-let normalize_evar evd ev =
- match kind_of_term (whd_evar evd (mkEvar ev)) with
- | Evar (evk,args) -> (evk,args)
- | _ -> assert false
+let whd_head_evar sigma c =
+ let (f, args) = whd_head_evar_stack sigma c in
+ (** optim: don't reallocate if empty/singleton *)
+ match args with
+ | [] -> f
+ | [arg] -> mkApp (f, arg)
+ | _ -> mkApp (f, Array.concat args)
(**********************)
(* Creating new metas *)
@@ -161,94 +214,28 @@ let normalize_evar evd ev =
(* Generator of metavariables *)
let new_meta =
- let meta_ctr = ref 0 in
- Summary.declare_summary "meta counter"
- { Summary.freeze_function = (fun () -> !meta_ctr);
- Summary.unfreeze_function = (fun n -> meta_ctr := n);
- Summary.init_function = (fun () -> meta_ctr := 0) };
+ let meta_ctr = Summary.ref 0 ~name:"meta counter" in
fun () -> incr meta_ctr; !meta_ctr
let mk_new_meta () = mkMeta(new_meta())
-let collect_evars emap c =
- let rec collrec acc c =
- match kind_of_term c with
- | Evar (evk,_) ->
- if Evd.is_undefined emap evk then evk::acc
- else (* No recursion on the evar instantiation *) acc
- | _ ->
- fold_constr collrec acc c in
- list_uniquize (collrec [] c)
-
-let push_dependent_evars sigma emap =
- Evd.fold_undefined (fun ev {evar_concl = ccl} (sigma',emap') ->
- List.fold_left
- (fun (sigma',emap') ev ->
- (Evd.add sigma' ev (Evd.find emap' ev),Evd.remove emap' ev))
- (sigma',emap') (collect_evars emap' ccl))
- emap (sigma,emap)
-
-let push_duplicated_evars sigma emap c =
- let rec collrec (one,(sigma,emap) as acc) c =
- match kind_of_term c with
- | Evar (evk,_) when not (Evd.mem sigma evk) ->
- if List.mem evk one then
- let sigma' = Evd.add sigma evk (Evd.find emap evk) in
- let emap' = Evd.remove emap evk in
- (one,(sigma',emap'))
- else
- (evk::one,(sigma,emap))
- | _ ->
- fold_constr collrec acc c
- in
- snd (collrec ([],(sigma,emap)) c)
-
-(* replaces a mapping of existentials into a mapping of metas.
- Problem if an evar appears in the type of another one (pops anomaly) *)
-let evars_to_metas sigma (emap, c) =
- let emap = nf_evar_map_undefined emap in
- let sigma',emap' = push_dependent_evars sigma emap in
- let sigma',emap' = push_duplicated_evars sigma' emap' c in
- (* if an evar has been instantiated in [emap] (as part of typing [c])
- then it is instantiated in [sigma]. *)
- let repair_evars sigma emap =
- fold_undefined begin fun ev _ sigma' ->
- try
- let info = find emap ev in
- match evar_body info with
- | Evar_empty -> sigma'
- | Evar_defined body -> define ev body sigma'
- with Not_found -> sigma'
- end sigma sigma
- in
- let sigma' = repair_evars sigma' emap in
- let change_exist evar =
- let ty = nf_betaiota emap (existential_type emap evar) in
- let n = new_meta() in
- mkCast (mkMeta n, DEFAULTcast, ty) in
- let rec replace c =
- match kind_of_term c with
- | Evar (evk,_ as ev) when Evd.mem emap' evk -> change_exist ev
- | _ -> map_constr replace c in
- (sigma', replace c)
-
(* The list of non-instantiated existential declarations (order is important) *)
let non_instantiated sigma =
- let listev = Evd.undefined_list sigma in
- List.map (fun (ev,evi) -> (ev,nf_evar_info sigma evi)) listev
+ let listev = Evd.undefined_map sigma in
+ Evar.Map.smartmap (fun evi -> nf_evar_info sigma evi) listev
(************************)
(* Manipulating filters *)
(************************)
-let apply_subfilter filter subfilter =
- fst (List.fold_right (fun oldb (l,filter) ->
- if oldb then List.hd filter::l,List.tl filter else (false::l,filter))
- filter ([], List.rev subfilter))
-
-let extract_subfilter initial_filter refined_filter =
- snd (list_filter2 (fun b1 b2 -> b1) (initial_filter,refined_filter))
+let make_pure_subst evi args =
+ snd (List.fold_right
+ (fun (id,b,c) (args,l) ->
+ match args with
+ | a::rest -> (rest, (id,a)::l)
+ | _ -> anomaly (Pp.str "Instance does not match its signature"))
+ (evar_filtered_context evi) (Array.rev_to_list args,[]))
(**********************)
(* Creating new evars *)
@@ -256,12 +243,8 @@ let extract_subfilter initial_filter refined_filter =
(* Generator of existential names *)
let new_untyped_evar =
- let evar_ctr = ref 0 in
- Summary.declare_summary "evar counter"
- { Summary.freeze_function = (fun () -> !evar_ctr);
- Summary.unfreeze_function = (fun n -> evar_ctr := n);
- Summary.init_function = (fun () -> evar_ctr := 0) };
- fun () -> incr evar_ctr; existential_of_int !evar_ctr
+ let evar_ctr = Summary.ref 0 ~name:"evar counter" in
+ fun () -> incr evar_ctr; Evar.unsafe_of_int !evar_ctr
(*------------------------------------*
* functional operations on evar sets *
@@ -296,109 +279,145 @@ let new_untyped_evar =
* we have the property that u and phi(t) are convertible in env.
*)
+let subst2 subst vsubst c =
+ substl subst (replace_vars vsubst c)
+
let push_rel_context_to_named_context env typ =
(* compute the instances relative to the named context and rel_context *)
let ids = List.map pi1 (named_context env) in
let inst_vars = List.map mkVar ids in
let inst_rels = List.rev (rel_list 0 (nb_rel env)) in
+ let replace_var_named_declaration id0 id (id',b,t) =
+ let id' = if Id.equal id0 id' then id else id' in
+ let vsubst = [id0 , mkVar id] in
+ let b = match b with
+ | None -> None
+ | Some c -> Some (replace_vars vsubst c)
+ in
+ id', b, replace_vars vsubst t
+ in
+ let replace_var_named_context id0 id env =
+ let nc = Environ.named_context env in
+ let nc' = List.map (replace_var_named_declaration id0 id) nc in
+ Environ.reset_with_named_context (val_of_named_context nc') env
+ in
+ let extract_if_neq id = function
+ | Anonymous -> None
+ | Name id' when id_ord id id' = 0 -> None
+ | Name id' -> Some id'
+ in
(* move the rel context to a named context and extend the named instance *)
(* with vars of the rel context *)
(* We do keep the instances corresponding to local definition (see above) *)
- let (subst, _, env) =
- Sign.fold_rel_context
- (fun (na,c,t) (subst, avoid, env) ->
- let id = next_name_away na avoid in
- let d = (id,Option.map (substl subst) c,substl subst t) in
- (mkVar id :: subst, id::avoid, push_named d env))
- (rel_context env) ~init:([], ids, env) in
- (named_context_val env, substl subst typ, inst_rels@inst_vars, subst)
+ let (subst, vsubst, _, env) =
+ Context.fold_rel_context
+ (fun (na,c,t) (subst, vsubst, avoid, env) ->
+ let id =
+ (* ppedrot: we want to infer nicer names for the refine tactic, but
+ keeping at the same time backward compatibility in other code
+ using this function. For now, we only attempt to preserve the
+ old behaviour of Program, but ultimately, one should do something
+ about this whole name generation problem. *)
+ if Flags.is_program_mode () then next_name_away na avoid
+ else next_ident_away (id_of_name_using_hdchar env t na) avoid
+ in
+ match extract_if_neq id na with
+ | Some id0 when not (is_section_variable id0) ->
+ (* spiwack: if [id<>id0], rather than introducing a new
+ binding named [id], we will keep [id0] (the name given
+ by the user) and rename [id0] into [id] in the named
+ context. Unless [id] is a section variable. *)
+ let subst = List.map (replace_vars [id0,mkVar id]) subst in
+ let vsubst = (id0,mkVar id)::vsubst in
+ let d = (id0, Option.map (subst2 subst vsubst) c, subst2 subst vsubst t) in
+ let env = replace_var_named_context id0 id env in
+ (mkVar id0 :: subst, vsubst, id::avoid, push_named d env)
+ | _ ->
+ (* spiwack: if [id0] is a section variable renaming it is
+ incorrect. We revert to a less robust behaviour where
+ the new binder has name [id]. Which amounts to the same
+ behaviour than when [id=id0]. *)
+ let d = (id,Option.map (subst2 subst vsubst) c,subst2 subst vsubst t) in
+ (mkVar id :: subst, vsubst, id::avoid, push_named d env)
+ )
+ (rel_context env) ~init:([], [], ids, env) in
+ (named_context_val env, subst2 subst vsubst typ, inst_rels@inst_vars, subst, vsubst)
(*------------------------------------*
* Entry points to define new evars *
*------------------------------------*)
-let default_source = (dummy_loc,InternalHole)
+let default_source = (Loc.ghost,Evar_kinds.InternalHole)
-let new_pure_evar evd sign ?(src=default_source) ?filter ?candidates typ =
+let restrict_evar evd evk filter candidates =
+ let evk' = new_untyped_evar () in
+ let evd = Evd.restrict evk evk' filter ?candidates evd in
+ Evd.declare_future_goal evk' evd, evk'
+
+let new_pure_evar_full evd evi =
+ let evk = new_untyped_evar () in
+ let evd = Evd.add evd evk evi in
+ let evd = Evd.declare_future_goal evk evd in
+ (evd, evk)
+
+let new_pure_evar sign evd ?(src=default_source) ?filter ?candidates ?store ?naming ?(principal=false) typ =
+ let default_naming =
+ if principal then
+ (* waiting for a more principled approach
+ (unnamed evars, private names?) *)
+ Misctypes.IntroFresh (Names.Id.of_string "tmp_goal")
+ else
+ Misctypes.IntroAnonymous
+ in
+ let naming = Option.default default_naming naming in
let newevk = new_untyped_evar() in
- let evd = evar_declare sign newevk typ ~src ?filter ?candidates evd in
+ let evd = evar_declare sign newevk typ ~src ?filter ?candidates ?store ~naming evd in
+ let evd =
+ if principal then Evd.declare_principal_goal newevk evd
+ else Evd.declare_future_goal newevk evd
+ in
(evd,newevk)
-let new_evar_instance sign evd typ ?src ?filter ?candidates instance =
+let new_evar_instance sign evd typ ?src ?filter ?candidates ?store ?naming ?principal instance =
assert (not !Flags.debug ||
- list_distinct (ids_of_named_context (named_context_of_val sign)));
- let evd,newevk = new_pure_evar evd sign ?src ?filter ?candidates typ in
+ List.distinct (ids_of_named_context (named_context_of_val sign)));
+ let evd,newevk = new_pure_evar sign evd ?src ?filter ?candidates ?store ?naming ?principal typ in
(evd,mkEvar (newevk,Array.of_list instance))
(* [new_evar] declares a new existential in an env env with type typ *)
(* Converting the env into the sign of the evar to define *)
-
-let new_evar evd env ?src ?filter ?candidates typ =
- let sign,typ',instance,subst = push_rel_context_to_named_context env typ in
- let candidates = Option.map (List.map (substl subst)) candidates in
+let new_evar env evd ?src ?filter ?candidates ?store ?naming ?principal typ =
+ let sign,typ',instance,subst,vsubst = push_rel_context_to_named_context env typ in
+ let candidates = Option.map (List.map (subst2 subst vsubst)) candidates in
let instance =
match filter with
| None -> instance
- | Some filter -> list_filter_with filter instance in
- new_evar_instance sign evd typ' ?src ?filter ?candidates instance
+ | Some filter -> Filter.filter_list filter instance in
+ new_evar_instance sign evd typ' ?src ?filter ?candidates ?store ?naming ?principal instance
-let new_type_evar ?src ?filter evd env =
- let evd', s = new_sort_variable evd in
- new_evar evd' env ?src ?filter (mkSort s)
+let new_type_evar env evd ?src ?filter ?naming ?principal rigid =
+ let evd', s = new_sort_variable rigid evd in
+ let evd', e = new_evar env evd' ?src ?filter ?naming ?principal (mkSort s) in
+ evd', (e, s)
- (* The same using side-effect *)
-let e_new_evar evdref env ?(src=(dummy_loc,InternalHole)) ?filter ?candidates ty =
- let (evd',ev) = new_evar !evdref env ~src:src ?filter ?candidates ty in
- evdref := evd';
- ev
+let e_new_type_evar env evdref ?src ?filter ?naming ?principal rigid =
+ let evd', c = new_type_evar env !evdref ?src ?filter ?naming ?principal rigid in
+ evdref := evd';
+ c
-(*------------------------------------*
- * Restricting existing evars *
- *------------------------------------*)
+let new_Type ?(rigid=Evd.univ_flexible) env evd =
+ let evd', s = new_sort_variable rigid evd in
+ evd', mkSort s
-let restrict_evar_key evd evk filter candidates =
- if filter = None && candidates = None then
- evd,evk
- else
- let evi = Evd.find_undefined evd evk in
- let oldfilter = evar_filter evi in
- if filter = Some oldfilter && candidates = None then
- evd,evk
- else
- let filter =
- match filter with
- | None -> evar_filter evi
- | Some filter -> filter in
- let candidates =
- match candidates with None -> evi.evar_candidates | _ -> candidates in
- let ccl = evi.evar_concl in
- let sign = evar_hyps evi in
- let src = evi.evar_source in
- let evd,newevk = new_pure_evar evd sign ccl ~src ~filter ?candidates in
- let ctxt = snd (list_filter2 (fun b c -> b) (filter,evar_context evi)) in
- let id_inst = Array.of_list (List.map (fun (id,_,_) -> mkVar id) ctxt) in
- Evd.define evk (mkEvar(newevk,id_inst)) evd,newevk
-
-(* Restrict an applied evar and returns its restriction in the same context *)
-let restrict_applied_evar evd (evk,argsv) filter candidates =
- let evd,newevk = restrict_evar_key evd evk filter candidates in
- let newargsv = match filter with
- | None -> (* optim *) argsv
- | Some filter ->
- let evi = Evd.find evd evk in
- let subfilter = extract_subfilter (evar_filter evi) filter in
- array_filter_with subfilter argsv in
- evd,(newevk,newargsv)
-
-(* Restrict an evar in the current evar_map *)
-let restrict_evar evd evk filter candidates =
- fst (restrict_evar_key evd evk filter candidates)
+let e_new_Type ?(rigid=Evd.univ_flexible) env evdref =
+ let evd', s = new_sort_variable rigid !evdref in
+ evdref := evd'; mkSort s
-(* Restrict an evar in the current evar_map *)
-let restrict_instance evd evk filter argsv =
- match filter with None -> argsv | Some filter ->
- let evi = Evd.find evd evk in
- array_filter_with (extract_subfilter (evar_filter evi) filter) argsv
+ (* The same using side-effect *)
+let e_new_evar env evdref ?(src=default_source) ?filter ?candidates ?store ?naming ?principal ty =
+ let (evd',ev) = new_evar env !evdref ~src:src ?filter ?candidates ?store ?naming ?principal ty in
+ evdref := evd';
+ ev
(* This assumes an evar with identity instance and generalizes it over only
the De Bruijn part of the context *)
@@ -410,176 +429,26 @@ let generalize_evar_over_rels sigma (ev,args) =
if isRel a then (mkNamedProd_or_LetIn d c,a::inst) else x)
(evi.evar_concl,[]) (Array.to_list args) sign
-(***************************************)
-(* Managing chains of local definitons *)
-(***************************************)
-
-(* 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 =
- List.fold_right (fun (id,b,c) aliases ->
- match b with
- | Some t ->
- (match kind_of_term t with
- | Var id' ->
- let aliases_of_id =
- try Idmap.find id' aliases with Not_found -> [] in
- Idmap.add id (aliases_of_id@[t]) aliases
- | _ ->
- Idmap.add id [t] aliases)
- | None -> aliases)
- sign Idmap.empty
-
-let compute_rel_aliases var_aliases rels =
- snd (List.fold_right (fun (_,b,t) (n,aliases) ->
- (n-1,
- match b with
- | Some t ->
- (match kind_of_term t with
- | Var id' ->
- let aliases_of_n =
- try Idmap.find id' var_aliases with Not_found -> [] in
- Intmap.add n (aliases_of_n@[t]) aliases
- | Rel p ->
- let aliases_of_n =
- try Intmap.find (p+n) aliases with Not_found -> [] in
- Intmap.add n (aliases_of_n@[mkRel (p+n)]) aliases
- | _ ->
- Intmap.add n [lift n t] aliases)
- | None -> aliases))
- rels (List.length rels,Intmap.empty))
-
-let make_alias_map env =
- (* 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 lift_aliases n (var_aliases,rel_aliases as aliases) =
- if n = 0 then aliases else
- (var_aliases,
- Intmap.fold (fun p l -> Intmap.add (p+n) (List.map (lift n) l))
- rel_aliases Intmap.empty)
-
-let get_alias_chain_of aliases x = match kind_of_term x with
- | Rel n -> (try Intmap.find n (snd aliases) with Not_found -> [])
- | Var id -> (try Idmap.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 or isVar a -> Some a
- | [_] -> None
- | _::a::_ -> Some a
-
-let normalize_alias aliases x =
- match normalize_alias_opt aliases x with
- | Some a -> a
- | None -> x
-
-let normalize_alias_var var_aliases id =
- destVar (normalize_alias (var_aliases,Intmap.empty) (mkVar id))
-
-let extend_alias (_,b,_) (var_aliases,rel_aliases) =
- let rel_aliases =
- Intmap.fold (fun n l -> Intmap.add (n+1) (List.map (lift 1) l))
- rel_aliases Intmap.empty in
- let rel_aliases =
- match b with
- | Some t ->
- (match kind_of_term t with
- | Var id' ->
- let aliases_of_binder =
- try Idmap.find id' var_aliases with Not_found -> [] in
- Intmap.add 1 (aliases_of_binder@[t]) rel_aliases
- | Rel p ->
- let aliases_of_binder =
- try Intmap.find (p+1) rel_aliases with Not_found -> [] in
- Intmap.add 1 (aliases_of_binder@[mkRel (p+1)]) rel_aliases
- | _ ->
- Intmap.add 1 [lift 1 t] rel_aliases)
- | None -> 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 rec 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
- | _ ->
- map_constr_with_full_binders
- extend_alias expand_vars_in_term_using aliases t
-
-let expand_vars_in_term env = expand_vars_in_term_using (make_alias_map env)
-
-let free_vars_and_rels_up_alias_expansion aliases c =
- let acc1 = ref Intset.empty and acc2 = ref Idset.empty in
- let cache_rel = ref Intset.empty and cache_var = ref Idset.empty in
- let is_in_cache depth = function
- | Rel n -> Intset.mem (n-depth) !cache_rel
- | Var s -> Idset.mem s !cache_var
- | _ -> false in
- let put_in_cache depth = function
- | Rel n -> cache_rel := Intset.add (n-depth) !cache_rel
- | Var s -> cache_var := Idset.add s !cache_var
- | _ -> () in
- let rec frec (aliases,depth) c =
- match kind_of_term c with
- | Rel _ | Var _ as ck ->
- if is_in_cache depth ck then () else begin
- put_in_cache depth ck;
- let c = expansion_of_var aliases c in
- match kind_of_term c with
- | Var id -> acc2 := Idset.add id !acc2
- | Rel n -> if n >= depth+1 then acc1 := Intset.add (n-depth) !acc1
- | _ -> frec (aliases,depth) c end
- | Const _ | Ind _ | Construct _ ->
- acc2 := List.fold_right Idset.add (vars_of_global (Global.env()) c) !acc2
- | _ ->
- iter_constr_with_full_binders
- (fun d (aliases,depth) -> (extend_alias d aliases,depth+1))
- frec (aliases,depth) c
- in
- frec (aliases,0) c;
- (!acc1,!acc2)
-
(************************************)
(* Removing a dependency in an evar *)
(************************************)
type clear_dependency_error =
-| OccurHypInSimpleClause of identifier option
+| OccurHypInSimpleClause of Id.t option
| EvarTypingBreak of existential
-exception ClearDependencyError of identifier * clear_dependency_error
-
-open Store.Field
+exception ClearDependencyError of Id.t * clear_dependency_error
let cleared = Store.field ()
-let rec check_and_clear_in_constr evdref err ids c =
+exception Depends of Id.t
+
+let rec check_and_clear_in_constr env evdref err ids c =
(* returns a new constr where all the evars have been 'cleaned'
(ie the hypotheses ids have been removed from the contexts of
evars) *)
let check id' =
- if List.mem id' ids then
+ if Id.Set.mem id' ids then
raise (ClearDependencyError (id',err))
in
match kind_of_term c with
@@ -587,14 +456,14 @@ let rec check_and_clear_in_constr evdref err ids c =
check id'; c
| ( Const _ | Ind _ | Construct _ ) ->
- let vars = Environ.vars_of_global (Global.env()) c in
- List.iter check vars; c
+ let vars = Environ.vars_of_global env c in
+ Id.Set.iter check vars; c
| Evar (evk,l as ev) ->
if Evd.is_defined !evdref evk then
(* If evk is already defined we replace it by its definition *)
let nc = whd_evar !evdref c in
- (check_and_clear_in_constr evdref err ids nc)
+ (check_and_clear_in_constr env evdref err ids nc)
else
(* We check for dependencies to elements of ids in the
evar_info corresponding to e and in the instance of
@@ -603,1217 +472,93 @@ let rec check_and_clear_in_constr evdref err ids c =
removed *)
let evi = Evd.find_undefined !evdref evk in
let ctxt = Evd.evar_filtered_context evi in
- let (nhyps,nargs,rids) =
- List.fold_right2
- (fun (rid,ob,c as h) a (hy,ar,ri) ->
- (* Check if some id to clear occurs in the instance
- a of rid in ev and remember the dependency *)
- match
- List.filter (fun id -> List.mem id ids) (Idset.elements (collect_vars a))
- with
- | id :: _ -> (hy,ar,(rid,id)::ri)
- | _ ->
- (* Check if some rid to clear in the context of ev
- has dependencies in another hyp of the context of ev
- and transitively remember the dependency *)
- match List.filter (fun (id,_) -> occur_var_in_decl (Global.env()) id h) ri with
- | (_,id') :: _ -> (hy,ar,(rid,id')::ri)
- | _ ->
- (* No dependency at all, we can keep this ev's context hyp *)
- (h::hy,a::ar,ri))
- ctxt (Array.to_list l) ([],[],[]) in
+ let (rids,filter) =
+ List.fold_right2
+ (fun (rid, ob,c as h) a (ri,filter) ->
+ try
+ (* Check if some id to clear occurs in the instance
+ a of rid in ev and remember the dependency *)
+ let check id = if Id.Set.mem id ids then raise (Depends id) in
+ let () = Id.Set.iter check (collect_vars a) in
+ (* Check if some rid to clear in the context of ev
+ has dependencies in another hyp of the context of ev
+ and transitively remember the dependency *)
+ let check id _ =
+ if occur_var_in_decl (Global.env ()) id h
+ then raise (Depends id)
+ in
+ let () = Id.Map.iter check ri in
+ (* No dependency at all, we can keep this ev's context hyp *)
+ (ri, true::filter)
+ with Depends id -> (Id.Map.add rid id ri, false::filter))
+ ctxt (Array.to_list l) (Id.Map.empty,[]) in
(* Check if some rid to clear in the context of ev has dependencies
in the type of ev and adjust the source of the dependency *)
- let nconcl =
- try check_and_clear_in_constr evdref (EvarTypingBreak ev)
- (List.map fst rids) (evar_concl evi)
+ let _nconcl =
+ try
+ let nids = Id.Map.domain rids in
+ check_and_clear_in_constr env evdref (EvarTypingBreak ev) nids (evar_concl evi)
with ClearDependencyError (rid,err) ->
- raise (ClearDependencyError (List.assoc rid rids,err)) in
+ raise (ClearDependencyError (Id.Map.find rid rids,err)) in
- if rids = [] then c else begin
- let env = Sign.fold_named_context push_named nhyps ~init:(empty_env) in
- let ev'= e_new_evar evdref env ~src:(evar_source evk !evdref) nconcl in
- evdref := Evd.define evk ev' !evdref;
- let (evk',_) = destEvar ev' in
+ if Id.Map.is_empty rids then c
+ else
+ let origfilter = Evd.evar_filter evi in
+ let filter = Evd.Filter.apply_subfilter origfilter filter in
+ let evd,_ = restrict_evar !evdref evk filter None in
+ evdref := evd;
(* spiwack: hacking session to mark the old [evk] as having been "cleared" *)
let evi = Evd.find !evdref evk in
let extra = evi.evar_extra in
- let extra' = cleared.set true extra in
+ let extra' = Store.set extra cleared true in
let evi' = { evi with evar_extra = extra' } in
evdref := Evd.add !evdref evk evi' ;
(* spiwack: /hacking session *)
- mkEvar(evk', Array.of_list nargs)
- end
+ whd_evar !evdref c
- | _ -> map_constr (check_and_clear_in_constr evdref err ids) c
+ | _ -> map_constr (check_and_clear_in_constr env evdref err ids) c
-let clear_hyps_in_evi evdref hyps concl ids =
+let clear_hyps_in_evi_main env evdref hyps terms ids =
(* clear_hyps_in_evi erases hypotheses ids in hyps, checking if some
hypothesis does not depend on a element of ids, and erases ids in
the contexts of the evars occuring in evi *)
- let nconcl =
- check_and_clear_in_constr evdref (OccurHypInSimpleClause None) ids concl in
+ let terms =
+ List.map (check_and_clear_in_constr env evdref (OccurHypInSimpleClause None) ids) terms in
let nhyps =
- let check_context (id,ob,c) =
+ let check_context ((id,ob,c) as decl) =
let err = OccurHypInSimpleClause (Some id) in
- (id, Option.map (check_and_clear_in_constr evdref err ids) ob,
- check_and_clear_in_constr evdref err ids c)
+ let ob' = Option.smartmap (fun c -> check_and_clear_in_constr env evdref err ids c) ob in
+ let c' = check_and_clear_in_constr env evdref err ids c in
+ if ob == ob' && c == c' then decl else (id, ob', c')
in
- let check_value vk =
- match !vk with
- | VKnone -> vk
- | VKvalue (v,d) ->
- if (List.for_all (fun e -> not (Idset.mem e d)) ids) then
- (* v does depend on any of ids, it's ok *)
- vk
- else
- (* v depends on one of the cleared hyps: we forget the computed value *)
- ref VKnone
+ let check_value vk = match force_lazy_val vk with
+ | None -> vk
+ | Some (_, d) ->
+ if (Id.Set.for_all (fun e -> not (Id.Set.mem e d)) ids) then
+ (* v does depend on any of ids, it's ok *)
+ vk
+ else
+ (* v depends on one of the cleared hyps:
+ we forget the computed value *)
+ dummy_lazy_val ()
in
remove_hyps ids check_context check_value hyps
in
- (nhyps,nconcl)
-
-(********************************)
-(* Managing pattern-unification *)
-(********************************)
-
-let rec expand_and_check_vars aliases = function
- | [] -> []
- | a::l when isRel a or isVar a ->
- let a = expansion_of_var aliases a in
- if isRel a or 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 = eq_constr
- let hash = hash_constr
- end)
-
-let rec 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
- (* 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 -> Idset.mem id fv_ids
- | Rel n -> Intset.mem n fv_rels
- | _ -> assert false) l
-
-let remove_instance_local_defs evd evk args =
- let evi = Evd.find evd evk in
- let rec aux = function
- | (_,Some _,_)::sign, a::args -> aux (sign,args)
- | (_,None,_)::sign, a::args -> a::aux (sign,args)
- | [], [] -> []
- | _ -> assert false in
- aux (evar_filtered_context evi, 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 is_unification_pattern_meta env 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
- | _ -> None
- else
- 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
- let args = remove_instance_local_defs evd evk (Array.to_list args) in
- let n = List.length args in
- match find_unification_pattern_args env (args @ l) t with
- | Some l -> Some (list_skipn n l)
- | _ -> None
- else
- None
-
-let is_unification_pattern_pure_evar env evd (evk,args) t =
- is_unification_pattern_evar env evd (evk,args) [] t <> None
-
-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
- | Evar ev -> is_unification_pattern_evar env evd ev l t
- | _ -> None
-
-(* From a unification problem "?X l = c", build "\x1...xn.(term1 l2)"
- (pattern unification). It is assumed that l is made of rel's that
- are distinct and not bound to aliases. *)
-(* It is also assumed that c does not contain metas because metas
- *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 c' = List.fold_right (fun a c ->
- let c' = subst_term (lift 1 a) (lift 1 c) in
- match kind_of_term a with
- (* Rem: if [a] links to a let-in, do as if it were an assumption *)
- | Rel n ->
- let d = map_rel_declaration (lift n) (lookup_rel n env) in
- mkLambda_or_LetIn d c'
- | Var 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 *)
- whd_eta c'
-
-(*****************************************)
-(* Refining/solving unification problems *)
-(*****************************************)
-
-(* Knowing that [Gamma |- ev : T] and that [ev] is applied to [args],
- * [make_projectable_subst ev args] builds the substitution [Gamma:=args].
- * If a variable and an alias of it are bound to the same instance, we skip
- * the alias (we just use eq_constr -- instead of conv --, since anyway,
- * only instances that are variables -- or evars -- are later considered;
- * morever, we can bet that similar instances came at some time from
- * the very same substitution. The removal of aliased duplicates is
- * useful to ensure the uniqueness of a projection.
-*)
-
-let make_projectable_subst aliases sigma evi args =
- let sign = evar_filtered_context evi in
- let evar_aliases = compute_var_aliases sign in
- let (_,full_subst,cstr_subst) =
- List.fold_right
- (fun (id,b,c) (args,all,cstrs) ->
- match b,args with
- | None, 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
- | Construct cstr ->
- let l = try Constrmap.find cstr cstrs with Not_found -> [] in
- Constrmap.add cstr ((args,id)::l) cstrs
- | _ -> cstrs in
- (rest,Idmap.add id [a,normalize_alias_opt aliases a,id] all,cstrs)
- | Some c, a::rest ->
- let a = whd_evar sigma a in
- (match kind_of_term c with
- | Var id' ->
- let idc = normalize_alias_var evar_aliases id' in
- let sub = try Idmap.find idc all with Not_found -> [] in
- if List.exists (fun (c,_,_) -> eq_constr a c) sub then
- (rest,all,cstrs)
- else
- (rest,
- Idmap.add idc ((a,normalize_alias_opt aliases a,id)::sub) all,
- cstrs)
- | _ ->
- (rest,Idmap.add id [a,normalize_alias_opt aliases a,id] all,cstrs))
- | _ -> anomaly "Instance does not match its signature")
- sign (array_rev_to_list args,Idmap.empty,Constrmap.empty) in
- (full_subst,cstr_subst)
-
-let make_pure_subst evi args =
- snd (List.fold_right
- (fun (id,b,c) (args,l) ->
- match args with
- | a::rest -> (rest, (id,a)::l)
- | _ -> anomaly "Instance does not match its signature")
- (evar_filtered_context evi) (array_rev_to_list args,[]))
+ (nhyps,terms)
-(*------------------------------------*
- * operations on the evar constraints *
- *------------------------------------*)
-
-(* We have a unification problem Σ; Γ |- ?e[u1..uq] = t : s where ?e is not yet
- * declared in Σ but yet known to be declarable in some context x1:T1..xq:Tq.
- * [define_evar_from_virtual_equation ... Γ Σ t (x1:T1..xq:Tq) .. (u1..uq) (x1..xq)]
- * declares x1:T1..xq:Tq |- ?e : s such that ?e[u1..uq] = t holds.
- *)
-
-let define_evar_from_virtual_equation define_fun env evd t_in_env sign filter inst_in_env =
- let ty_t_in_env = Retyping.get_type_of env evd t_in_env in
- let evd,evar_in_env = new_evar_instance sign evd ty_t_in_env ~filter inst_in_env in
- let t_in_env = whd_evar evd t_in_env in
- let evd = define_fun env evd (destEvar evar_in_env) t_in_env in
- let ids = List.map pi1 (named_context_of_val sign) in
- let inst_in_sign = List.map mkVar (list_filter_with filter ids) in
- let evar_in_sign = mkEvar (fst (destEvar evar_in_env), Array.of_list inst_in_sign) in
- (evd,whd_evar evd evar_in_sign)
-
-(* We have x1..xq |- ?e1 : Ï„ and had to solve something like
- * Σ; Γ |- ?e1[u1..uq] = (...\y1 ... \yk ... c), where c is typically some
- * ?e2[v1..vn], hence flexible. We had to go through k binders and now
- * virtually have x1..xq, y1'..yk' | ?e1' : Ï„' and the equation
- * Γ, y1..yk |- ?e1'[u1..uq y1..yk] = c.
- * [materialize_evar Γ evd k (?e1[u1..uq]) τ'] extends Σ with the declaration
- * of ?e1' and returns both its instance ?e1'[x1..xq y1..yk] in an extension
- * of the context of e1 so that e1 can be instantiated by
- * (...\y1' ... \yk' ... ?e1'[x1..xq y1'..yk']),
- * and the instance ?e1'[u1..uq y1..yk] so that the remaining equation
- * ?e1'[u1..uq y1..yk] = c can be registered
- *
- * Note that, because invert_definition does not check types, we need to
- * guess the types of y1'..yn' by inverting the types of y1..yn along the
- * substitution u1..uq.
- *)
-
-let materialize_evar define_fun env evd k (evk1,args1) ty_in_env =
- 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 ids1 = List.map pi1 (named_context_of_val sign1) in
- let inst_in_sign = List.map mkVar (list_filter_with filter1 ids1) in
- let (sign2,filter2,inst2_in_env,inst2_in_sign,_,evd,_) =
- List.fold_right (fun (na,b,t_in_env as d) (sign,filter,inst_in_env,inst_in_sign,env,evd,avoid) ->
- let id = next_name_away na avoid in
- let evd,t_in_sign =
- define_evar_from_virtual_equation define_fun env evd t_in_env
- sign filter inst_in_env in
- let evd,b_in_sign = match b with
- | None -> evd,None
- | Some b ->
- let evd,b = define_evar_from_virtual_equation define_fun env evd b
- sign filter inst_in_env in
- evd,Some b in
- (push_named_context_val (id,b_in_sign,t_in_sign) sign,true::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))
- rel_sign
- (sign1,filter1,Array.to_list args1,inst_in_sign,env1,evd,ids1)
- in
- let evd,ev2ty_in_sign =
- define_evar_from_virtual_equation define_fun env evd ty_in_env
- sign2 filter2 inst2_in_env in
- let evd,ev2_in_sign =
- new_evar_instance sign2 evd ev2ty_in_sign ~filter:filter2 inst2_in_sign in
- let ev2_in_env = (fst (destEvar 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 =
- let newfilter = List.map p args in
- if List.for_all (fun id -> id) newfilter then
- None
- else
- let oldfullfilter = evar_filter (Evd.find_undefined evd evk) in
- Some (apply_subfilter oldfullfilter newfilter)
-
-(* Inverting constructors in instances (common when inferring type of match) *)
-
-let find_projectable_constructor env evd cstr k args cstr_subst =
- try
- let l = Constrmap.find cstr cstr_subst in
- let args = Array.map (lift (-k)) args in
- let l =
- 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
- List.map snd l
- with Not_found ->
- []
-
-(* [find_projectable_vars env sigma y subst] finds all vars of [subst]
- * that project on [y]. It is able to find solutions to the following
- * two kinds of problems:
- *
- * - ?n[...;x:=y;...] = y
- * - ?n[...;x:=?m[args];...] = y with ?m[args] = y recursively solvable
- *
- * (see test-suite/success/Fixpoint.v for an example of application of
- * the second kind of problem).
- *
- * The seek for [y] is up to variable aliasing. In case of solutions that
- * differ only up to aliasing, the binding that requires the less
- * steps of alias reduction is kept. At the end, only one solution up
- * to aliasing is kept.
- *
- * [find_projectable_vars] also unifies against evars that themselves mention
- * [y] and recursively.
- *
- * In short, the following situations give the following solutions:
- *
- * problem evar ctxt soluce remark
- * z1; z2:=z1 |- ?ev[z1;z2] = z1 y1:A; y2:=y1 y1 \ thanks to defs kept in
- * z1; z2:=z1 |- ?ev[z1;z2] = z2 y1:A; y2:=y1 y2 / subst and preferring =
- * z1; z2:=z1 |- ?ev[z1] = z2 y1:A y1 thanks to expand_var
- * z1; z2:=z1 |- ?ev[z2] = z1 y1:A y1 thanks to expand_var
- * z3 |- ?ev[z3;z3] = z3 y1:A; y2:=y1 y2 see make_projectable_subst
- *
- * Remark: [find_projectable_vars] assumes that identical instances of
- * variables in the same set of aliased variables are already removed (see
- * [make_projectable_subst])
- *)
-
-type evar_projection =
-| ProjectVar
-| ProjectEvar of existential * evar_info * identifier * evar_projection
-
-exception NotUnique
-exception NotUniqueInType of (identifier * 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 eq_constr y c' then id
- else
- if l <> [] then assoc_up_to_alias sigma aliases y yc l
- else
- (* Last chance, we reason up to alias conversion *)
- match (if c == c' then cc else normalize_alias_opt aliases c') with
- | Some cc when eq_constr yc cc -> id
- | _ -> if eq_constr yc c 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 is_projectable idc idcl subst' =
- (* First test if some [id] aliased to [idc] is bound to [y] in [subst] *)
- try
- let id = assoc_up_to_alias sigma aliases y yc idcl in
- (id,ProjectVar)::subst'
- with Not_found ->
- (* Then test if [idc] is (indirectly) bound in [subst] to some evar *)
- (* projectable on [y] *)
- if with_evars then
- let idcl' = List.filter (fun (c,_,id) -> isEvar c) idcl in
- match idcl' with
- | [c,_,id] ->
- begin
- let (evk,argsv as t) = destEvar 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
- match l with
- | [id',p] -> (id,ProjectEvar (t,evi,id',p))::subst'
- | _ -> subst'
- end
- | [] -> subst'
- | _ -> anomaly "More than one non var in aliases class of evar instance"
- else
- subst' in
- Idmap.fold is_projectable subst []
-
-(* [filter_solution] checks if one and only one possible projection exists
- * among a set of solutions to a projection problem *)
-
-let filter_solution = function
- | [] -> raise Not_found
- | (id,p)::_::_ -> raise NotUnique
- | [id,p] -> (mkVar id, p)
-
-let project_with_effects aliases sigma effects t subst =
- let c, p =
- filter_solution (find_projectable_vars false aliases sigma t subst) in
- effects := p :: !effects;
- c
-
-let rec find_solution_type evarenv = function
- | (id,ProjectVar)::l -> pi3 (lookup_named id evarenv)
- | [id,ProjectEvar _] -> (* bugged *) pi3 (lookup_named id evarenv)
- | (id,ProjectEvar _)::l -> find_solution_type evarenv l
- | [] -> 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
- * type is an evar too.
- *
- * Note: typing creates new evar problems, which induces a recursive dependency
- * with [define]. To avoid a too large set of recursive functions, we
- * pass [define] to [do_projection_effects] as a parameter.
- *)
-
-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
- (* TODO: simplify constraints involving evk *)
- let evd = do_projection_effects define_fun env ty evd p in
- let ty = whd_betadeltaiota env evd (Lazy.force ty) in
- if not (isSort 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 (destEvar ty') ty else evd
- else
- evd
-
-(* Assuming Σ; Γ, y1..yk |- c, [invert_arg_from_subst Γ k Σ [x1:=u1..xn:=un] c]
- * tries to return φ(x1..xn) such that equation φ(u1..un) = c is valid.
- * The strategy is to imitate the structure of c and then to invert
- * the variables of c (i.e. rels or vars of Γ) using the algorithm
- * implemented by project_with_effects/find_projectable_vars.
- * It returns either a unique solution or says whether 0 or more than
- * 1 solutions is found.
- *
- * Precondition: Σ; Γ, y1..yk |- c /\ Σ; Γ |- u1..un
- * Postcondition: if φ(x1..xn) is returned then
- * Σ; Γ, y1..yk |- φ(u1..un) = c /\ x1..xn |- φ(x1..xn)
- *
- * The effects correspond to evars instantiated while trying to project.
- *
- * [invert_arg_from_subst] is used on instances of evars. Since the
- * evars are flexible, these instances are potentially erasable. This
- * is why we don't investigate whether evars in the instances of evars
- * are unifiable, to the contrary of [invert_definition].
- *)
-
-type projectibility_kind =
- | NoUniqueProjection
- | UniqueProjection of constr * evar_projection list
-
-type projectibility_status =
- | CannotInvert
- | Invertible of projectibility_kind
-
-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
- 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
- | None -> raise Not_found
- | Some c -> aux k c in
- try
- let c = aux 0 c_in_env_extended_with_k_binders in
- Invertible (UniqueProjection (c,!effects))
- with
- | Not_found -> CannotInvert
- | NotUnique -> Invertible NoUniqueProjection
-
-let invert_arg fullenv evd aliases k evk subst_in_env_extended_with_k_binders c_in_env_extended_with_k_binders =
- let res = invert_arg_from_subst evd aliases k subst_in_env_extended_with_k_binders c_in_env_extended_with_k_binders in
- match res with
- | Invertible (UniqueProjection (c,_)) when not (noccur_evar fullenv evd evk c)
- ->
- CannotInvert
- | _ ->
- res
-
-let effective_projections =
- map_succeed (function Invertible c -> c | _ -> failwith"")
-
-let instance_of_projection f env t evd projs =
- let ty = lazy (nf_evar evd (Retyping.get_type_of env evd t)) in
- match projs with
- | NoUniqueProjection -> raise NotUnique
- | UniqueProjection (c,effects) ->
- (List.fold_left (do_projection_effects f env ty) evd effects, c)
-
-exception NotEnoughInformationToInvert
-
-let extract_unique_projections projs =
- List.map (function
- | Invertible (UniqueProjection (c,_)) -> c
- | _ ->
- (* For instance, there are evars with non-invertible arguments and *)
- (* we cannot arbitrarily restrict these evars before knowing if there *)
- (* will really be used; it can also be due to some argument *)
- (* (typically a rel) that is not inversible and that cannot be *)
- (* inverted either because it is needed for typing the conclusion *)
- (* of the evar to project *)
- raise NotEnoughInformationToInvert) projs
-
-let extract_candidates sols =
- try
- Some
- (List.map (function (id,ProjectVar) -> mkVar id | _ -> raise Exit) sols)
- with Exit ->
- None
-
-let filter_of_projection = function Invertible _ -> true | _ -> false
-
-let invert_invertible_arg fullenv evd aliases k (evk,argsv) args' =
- let evi = Evd.find_undefined evd evk in
- let subst,_ = make_projectable_subst aliases evd evi argsv in
- let projs =
- array_map_to_list (invert_arg fullenv evd aliases k evk subst) args' in
- Array.of_list (extract_unique_projections projs)
-
-(* Redefines an evar with a smaller context (i.e. it may depend on less
- * variables) such that c becomes closed.
- * Example: in "fun (x:?1) (y:list ?2[x]) => x = y :> ?3[x,y] /\ x = nil bool"
- * ?3 <-- ?1 no pb: env of ?3 is larger than ?1's
- * ?1 <-- list ?2 pb: ?2 may depend on x, but not ?1.
- * What we do is that ?2 is defined by a new evar ?4 whose context will be
- * a prefix of ?2's env, included in ?1's env.
- *
- * If "hyps |- ?e : T" and "filter" selects a subset hyps' of hyps then
- * [do_restrict_hyps evd ?e filter] sets ?e:=?e'[hyps'] and returns ?e'
- * such that "hyps' |- ?e : T"
- *)
-
-let filter_effective_candidates evi filter candidates =
- match filter with
- | None -> candidates
- | Some filter ->
- let ids = List.map pi1 (list_filter_with filter (evar_context evi)) in
- List.filter (fun a -> list_subset (Idset.elements (collect_vars 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
- | None -> evi.evar_candidates
- | Some _ -> candidates_update
- in
- match candidates with
- | None -> None
- | Some l ->
- let l' = filter_effective_candidates evi filter l in
- if List.length l = List.length l' && candidates_update = None then
- None
- else
- Some l'
-
-let closure_of_filter evd evk filter =
- let evi = Evd.find_undefined evd evk in
- let vars = collect_vars (nf_evar evd (evar_concl evi)) in
- let test (id,c,_) b = b || Idset.mem id vars || c <> None in
- let newfilter = List.map2 test (evar_context evi) filter in
- if newfilter = evar_filter evi then None else Some newfilter
-
-let restrict_hyps evd evk filter candidates =
- (* What to do with dependencies?
- Assume we have x:A, y:B(x), z:C(x,y) |- ?e:T(x,y,z) and restrict on y.
- - If y is in a non-erasable position in C(x,y) (i.e. it is not below an
- occurrence of x in the hnf of C), then z should be removed too.
- - If y is in a non-erasable position in T(x,y,z) then the problem is
- unsolvable.
- Computing whether y is erasable or not may be costly and the
- interest for this early detection in practice is not obvious. We let
- it for future work. In any case, thanks to the use of filters, the whole
- (unrestricted) context remains consistent. *)
- let candidates = filter_candidates evd evk (Some filter) candidates in
- let typablefilter = closure_of_filter evd evk filter in
- (typablefilter,candidates)
-
-exception EvarSolvedWhileRestricting of evar_map * 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
- | Some [], _ -> error "Not solvable."
- | Some [nc],_ ->
- let evd = Evd.define evk nc evd in
- raise (EvarSolvedWhileRestricting (evd,whd_evar evd (mkEvar ev)))
- | None, None -> evd,ev
- | _ -> restrict_applied_evar evd ev filter candidates
-
-(* [postpone_non_unique_projection] postpones equation of the form ?e[?] = c *)
-(* ?e is assumed to have no candidates *)
-
-let postpone_non_unique_projection env evd (evk,argsv as ev) sols rhs =
- let rhs = expand_vars_in_term env rhs in
- let filter =
- restrict_upon_filter evd evk
- (* Keep only variables that occur in rhs *)
- (* This is not safe: is the variable is a local def, its body *)
- (* may contain references to variables that are removed, leading to *)
- (* a ill-formed context. We would actually need a notion of filter *)
- (* 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)
- (Array.to_list argsv) in
- let filter = match filter with
- | None -> None
- | Some filter -> closure_of_filter evd evk filter in
- let candidates = extract_candidates sols in
- if candidates <> None then
- restrict_evar evd evk filter candidates
- else
- (* We made an approximation by not expanding a local definition *)
- let evd,ev = restrict_applied_evar evd ev filter None in
- let pb = (Reduction.CONV,env,mkEvar ev,rhs) in
- Evd.add_conv_pb pb evd
-
-(* [postpone_evar_evar] postpones an equation of the form ?e1[?1] = ?e2[?2] *)
-
-let postpone_evar_evar f env evd filter1 ev1 filter2 ev2 =
- (* Leave an equation between (restrictions of) ev1 andv ev2 *)
- try
- let evd,ev1' = do_restrict_hyps evd ev1 filter1 None in
- try
- let evd,ev2' = do_restrict_hyps evd ev2 filter2 None in
- add_conv_pb (Reduction.CONV,env,mkEvar ev1',mkEvar ev2') evd
- with EvarSolvedWhileRestricting (evd,ev2) ->
- (* ev2 solved on the fly *)
- f env evd ev1' ev2
- with EvarSolvedWhileRestricting (evd,ev1) ->
- (* ev1 solved on the fly *)
- f env evd ev2 ev1
-
-(* [solve_evar_evar f Γ Σ ?e1[u1..un] ?e2[v1..vp]] applies an heuristic
- * to solve the equation Σ; Γ ⊢ ?e1[u1..un] = ?e2[v1..vp]:
- * - if there are at most one φj for each vj s.t. vj = φj(u1..un),
- * we first restrict ?e2 to the subset v_k1..v_kq of the vj that are
- * inversible and we set ?e1[x1..xn] := ?e2[φk1(x1..xn)..φkp(x1..xn)]
- * (this is a case of pattern-unification)
- * - symmetrically if there are at most one ψj for each uj s.t.
- * uj = ψj(v1..vp),
- * - otherwise, each position i s.t. ui does not occur in v1..vp has to
- * be restricted and similarly for the vi, and we leave the equation
- * as an open equation (performed by [postpone_evar])
- *
- * Warning: the notion of unique φj is relative to some given class
- * of unification problems
- *
- * Note: argument f is the function used to instantiate evars.
- *)
-
-let are_canonical_instances args1 args2 env =
- let n1 = Array.length args1 in
- let n2 = Array.length args2 in
- let rec aux n = function
- | (id,_,c)::sign
- when n < n1 && isVarId id args1.(n) && isVarId id args2.(n) ->
- aux (n+1) sign
- | [] ->
- let rec aux2 n =
- n = n1 ||
- (isRelN (n1-n) args1.(n) && isRelN (n1-n) args2.(n) && aux2 (n+1))
- in aux2 n
- | _ -> false in
- n1 = n2 & aux 0 (named_context env)
-
-let filter_compatible_candidates conv_algo env evd evi args rhs c =
- let c' = instantiate_evar (evar_filtered_context evi) c args in
- let evd, b = conv_algo env evd Reduction.CONV rhs c' in
- if b then Some (c,evd) else None
-
-exception DoesNotPreserveCandidateRestriction
-
-let restrict_candidates conv_algo env evd filter1 (evk1,argsv1) (evk2,argsv2) =
- let evi1 = Evd.find evd evk1 in
- let evi2 = Evd.find evd evk2 in
- let cand1 = filter_candidates evd evk1 filter1 None in
- let cand2 = evi2.evar_candidates in
- match cand1, cand2 with
- | _, None -> cand1
- | None, Some _ -> raise DoesNotPreserveCandidateRestriction
- | Some l1, Some l2 ->
- let args1 = Array.to_list argsv1 in
- let args2 = Array.to_list argsv2 in
- let l1' = List.filter (fun c1 ->
- let c1' = instantiate_evar (evar_filtered_context evi1) c1 args1 in
- List.filter (fun c2 ->
- (filter_compatible_candidates conv_algo env evd evi2 args2 c1' c2
- <> None)) l2 <> []) l1 in
- if List.length l1 = List.length l1' then None else Some l1'
-
-exception CannotProject of bool list option
-
-(* 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
- FV(u[x1:=t1..xn:=tn]) are in the set U?
- - If ti is a variable, it has to be in U.
- - If ti is a constructor, its parameters cannot be erased even if u
- matches on it, so we have to discard ti if the parameters
- contain variables not in U.
- - If ti is rigid, we have to discard it if it contains variables in U.
-
- Note: when restricting as part of an equation ?n[x1:=t1..xn:=tn] = ?m[...]
- then, occurrences of ?m in the ti can be seen, like variables, as occurrences
- of subterms to eventually discard so as to be allowed to keep ti.
-*)
-
-let rec is_constrainable_in k (ev,(fv_rels,fv_ids) as g) t =
- let f,args = decompose_app_vect t in
- match kind_of_term f with
- | Construct (ind,_) ->
- let nparams =
- (fst (Global.lookup_inductive ind)).Declarations.mind_nparams
- in
- if nparams > Array.length args
- then true (* We don't try to be more clever *)
- else
- let params,_ = array_chop nparams args in
- array_for_all (is_constrainable_in k g) params
- | Ind _ -> array_for_all (is_constrainable_in k g) args
- | Prod (_,t1,t2) -> is_constrainable_in k g t1 && is_constrainable_in k g t2
- | Evar (ev',_) -> ev' <> ev (*If ev' needed, one may also try to restrict it*)
- | Var id -> Idset.mem id fv_ids
- | Rel n -> n <= k || Intset.mem n fv_rels
- | Sort _ -> true
- | _ -> (* We don't try to be more clever *) true
-
-let has_constrainable_free_vars evd aliases k ev (fv_rels,fv_ids as fvs) t =
- let t = expansion_of_var aliases t in
- match kind_of_term t with
- | Var id -> Idset.mem id fv_ids
- | Rel n -> n <= k || Intset.mem n fv_rels
- | _ -> is_constrainable_in k (ev,fvs) t
-
-let ensure_evar_independent g env evd (evk1,argsv1 as ev1) (evk2,argsv2 as ev2)=
- let filter1 =
- restrict_upon_filter evd evk1 (noccur_evar env evd evk2)
- (Array.to_list argsv1)
- in
- let candidates1 = restrict_candidates g env evd filter1 ev1 ev2 in
- let evd,(evk1,_ as ev1) = do_restrict_hyps evd ev1 filter1 candidates1 in
- let filter2 =
- restrict_upon_filter evd evk2 (noccur_evar env evd evk1)
- (Array.to_list argsv2)
- in
- let candidates2 = restrict_candidates g env evd filter2 ev2 ev1 in
- let evd,ev2 = do_restrict_hyps evd ev2 filter2 candidates2 in
- evd,ev1,ev2
-
-exception EvarSolvedOnTheFly of evar_map * constr
-
-let project_evar_on_evar g env evd aliases k2 (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 filter1 = restrict_upon_filter evd evk1
- (has_constrainable_free_vars evd aliases k2 evk2 fvs2)
- (Array.to_list argsv1) in
- (* Only try pruning on variable substitutions, postpone otherwise. *)
- (* Rules out non-linear instances. *)
- if is_unification_pattern_pure_evar env evd ev2 (mkEvar ev1) then
- try
- let candidates1 = restrict_candidates g env evd filter1 ev1 ev2 in
- let evd,(evk1',args1) = do_restrict_hyps evd ev1 filter1 candidates1 in
- evd,mkEvar (evk1',invert_invertible_arg env evd aliases k2 ev2 args1)
- with
- | EvarSolvedWhileRestricting (evd,ev1) ->
- raise (EvarSolvedOnTheFly (evd,ev1))
- | DoesNotPreserveCandidateRestriction | NotEnoughInformationToInvert ->
- raise (CannotProject filter1)
- else
- raise (CannotProject filter1)
-
-let solve_evar_evar_l2r f g env evd aliases ev1 (evk2,_ as ev2) =
- try
- let evd,body = project_evar_on_evar g env evd aliases 0 ev1 ev2 in
- Evd.define evk2 body evd
- with EvarSolvedOnTheFly (evd,c) ->
- f env evd ev2 c
-
-let solve_evar_evar ?(force=false) f g env evd (evk1,args1 as ev1) (evk2,args2 as ev2) =
- if are_canonical_instances args1 args2 env then
- (* If instances are canonical, we solve the problem in linear time *)
- let sign = evar_filtered_context (Evd.find evd evk2) in
- let id_inst = list_map_to_array (fun (id,_,_) -> mkVar id) sign in
- Evd.define evk2 (mkEvar(evk1,id_inst)) evd
- else
- let evd,ev1,ev2 =
- (* If an evar occurs in the instance of the other evar and the
- use of an heuristic is forced, we restrict *)
- if force then ensure_evar_independent g env evd ev1 ev2 else (evd,ev1,ev2) in
- let aliases = make_alias_map env in
- try solve_evar_evar_l2r f g env evd aliases ev1 ev2
- with CannotProject filter1 ->
- try solve_evar_evar_l2r f g env evd aliases ev2 ev1
- with CannotProject filter2 ->
- postpone_evar_evar f env evd filter1 ev1 filter2 ev2
-
-type conv_fun =
- env -> evar_map -> conv_pb -> constr -> constr -> evar_map * bool
-
-let check_evar_instance evd evk1 body conv_algo =
- let evi = Evd.find evd evk1 in
- let evenv = evar_unfiltered_env evi in
- (* FIXME: The body might be ill-typed when this is called from w_merge *)
- let ty =
- try Retyping.get_type_of evenv evd body
- with e when Errors.noncritical e -> error "Ill-typed evar instance"
- in
- let evd,b = conv_algo evenv evd Reduction.CUMUL ty evi.evar_concl in
- if b then evd else
- user_err_loc (fst (evar_source evk1 evd),"",
- str "Unable to find a well-typed instantiation")
-
-(* 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
- * that don't unify are discarded (i.e. ?e is redefined so that it does not
- * depend on these args). *)
-
-let solve_refl ?(can_drop=false) conv_algo env evd evk argsv1 argsv2 =
- if array_equal eq_constr argsv1 argsv2 then evd else
- (* Filter and restrict if needed *)
- let untypedfilter =
- restrict_upon_filter evd evk
- (fun (a1,a2) -> snd (conv_algo env evd Reduction.CONV a1 a2))
- (List.combine (Array.to_list argsv1) (Array.to_list argsv2)) in
- let candidates = filter_candidates evd evk untypedfilter None in
- let filter = match untypedfilter with
- | None -> None
- | Some filter -> closure_of_filter evd evk filter in
- let evd,ev1 = restrict_applied_evar evd (evk,argsv1) filter candidates in
- if fst ev1 = evk & can_drop then (* No refinement *) evd else
- (* either progress, or not allowed to drop, e.g. to preserve possibly *)
- (* informative equations such as ?e[x:=?y]=?e[x:=?y'] where we don't know *)
- (* if e can depend on x until ?y is not resolved, or, conversely, we *)
- (* don't know if ?y has to be unified with ?y, until e is resolved *)
- let argsv2 = restrict_instance evd evk filter argsv2 in
- let ev2 = (fst ev1,argsv2) in
- (* Leave a unification problem *)
- Evd.add_conv_pb (Reduction.CONV,env,mkEvar ev1,mkEvar ev2) evd
-
-(* If the evar can be instantiated by a finite set of candidates known
- in advance, we check which of them apply *)
-
-exception NoCandidates
-
-let solve_candidates conv_algo env evd (evk,argsv as ev) rhs =
- let evi = Evd.find evd evk in
- let args = Array.to_list argsv in
- match evi.evar_candidates with
- | None -> raise NoCandidates
- | Some l ->
- let l' =
- list_map_filter
- (filter_compatible_candidates conv_algo env evd evi args rhs) l in
- match l' with
- | [] -> error_cannot_unify env evd (mkEvar ev, rhs)
- | [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 Evd.define evk c evd else evd
- | l when List.length l < List.length l' ->
- let candidates = List.map fst l in
- restrict_evar evd evk None (Some candidates)
- | l -> evd
-
-(* We try to instantiate the evar assuming the body won't depend
- * on arguments that are not Rels or Vars, or appearing several times
- * (i.e. we tackle a generalization of Miller-Pfenning patterns unification)
- *
- * 1) Let "env |- ?ev[hyps:=args] = rhs" be the unification problem
- * 2) We limit it to a patterns unification problem "env |- ev[subst] = rhs"
- * where only Rel's and Var's are relevant in subst
- * 3) We recur on rhs, "imitating" the term, and failing if some Rel/Var is
- * not in the scope of ?ev. For instance, the problem
- * "y:nat |- ?x[] = y" where "|- ?1:nat" is not satisfiable because
- * ?1 would be instantiated by y which is not in the scope of ?1.
- * 4) We try to "project" the term if the process of imitation fails
- * and that only one projection is possible
- *
- * Note: we don't assume rhs in normal form, it may fail while it would
- * have succeeded after some reductions.
- *
- * This is the work of [invert_definition Γ Σ ?ev[hyps:=args] c]
- * Precondition: Σ; Γ, y1..yk |- c /\ Σ; Γ |- u1..un
- * Postcondition: if φ(x1..xn) is returned then
- * Σ; Γ, y1..yk |- φ(u1..un) = c /\ x1..xn |- φ(x1..xn)
- *)
-
-exception NotInvertibleUsingOurAlgorithm of constr
-exception NotEnoughInformationToProgress of (identifier * evar_projection) list
-exception NotEnoughInformationEvarEvar of constr
-exception OccurCheckIn of evar_map * constr
-
-let rec invert_definition conv_algo choose env evd (evk,argsv as ev) rhs =
- let aliases = make_alias_map env in
- let evdref = ref evd in
- let progress = ref false in
- let evi = Evd.find evd evk in
- let subst,cstr_subst = make_projectable_subst aliases evd evi argsv in
-
- (* Projection *)
- let project_variable t =
- (* Evar/Var problem: unifiable iff variable projectable from ev subst *)
- try
- let sols = find_projectable_vars true aliases !evdref t subst in
- 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)
- in
- let ty = lazy (Retyping.get_type_of env !evdref t) in
- let evd = do_projection_effects (evar_define conv_algo) env ty !evdref p in
- evdref := evd;
- c
- with
- | Not_found -> raise (NotInvertibleUsingOurAlgorithm t)
- | NotUniqueInType sols ->
- if not !progress then
- raise (NotEnoughInformationToProgress sols);
- (* No unique projection but still restrict to where it is possible *)
- (* materializing is necessary, but is restricting useful? *)
- let ty = find_solution_type (evar_env evi) sols in
- let sign = evar_filtered_context evi in
- let ty' = instantiate_evar sign ty (Array.to_list argsv) in
- let (evd,evar,(evk',argsv' as ev')) =
- materialize_evar (evar_define conv_algo) env !evdref 0 ev ty' in
- let ts = expansions_of_var aliases t in
- let test c = isEvar c or List.mem c ts in
- let filter = array_map_to_list test argsv' in
- let filter = apply_subfilter (evar_filter (Evd.find_undefined evd evk)) filter in
-
- let filter = closure_of_filter evd evk' filter in
- let candidates = extract_candidates sols in
- let evd =
- if candidates <> None then restrict_evar evd evk' filter candidates
- else
- let evd,ev'' = restrict_applied_evar evd ev' filter None in
- Evd.add_conv_pb (Reduction.CONV,env,mkEvar ev'',t) evd in
- evdref := evd;
- evar in
-
- let rec imitate (env',k as envk) t =
- let t = whd_evar !evdref t in
- match kind_of_term t with
- | Rel i when i>k ->
- (match pi2 (Environ.lookup_rel (i-k) env') with
- | None -> project_variable (mkRel (i-k))
- | Some b ->
- try project_variable (mkRel (i-k))
- with NotInvertibleUsingOurAlgorithm _ -> imitate envk (lift i b))
- | Var id ->
- (match pi2 (Environ.lookup_named id env') with
- | None -> project_variable t
- | Some b ->
- try project_variable t
- with NotInvertibleUsingOurAlgorithm _ -> imitate envk b)
- | Evar (evk',args' as ev') ->
- if evk = evk' then raise (OccurCheckIn (evd,rhs));
- (* Evar/Evar problem (but left evar is virtual) *)
- let aliases = lift_aliases k aliases in
- (try
- let ev = (evk,Array.map (lift k) argsv) in
- let evd,body = project_evar_on_evar conv_algo env' !evdref aliases k ev' ev in
- evdref := evd;
- body
- with
- | EvarSolvedOnTheFly (evd,t) -> evdref:=evd; imitate envk t
- | CannotProject filter' ->
- if not !progress then
- raise (NotEnoughInformationEvarEvar t);
- (* Make the virtual left evar real *)
- let ty = get_type_of env' !evdref t in
- let (evd,evar'',ev'') =
- materialize_evar (evar_define conv_algo) env' !evdref k ev ty in
- (* materialize_evar may instantiate ev' by another evar; adjust it *)
- let (evk',args' as ev') = normalize_evar evd ev' in
- let evd =
- (* Try to project (a restriction of) the left evar ... *)
- try
- let evd,body = project_evar_on_evar conv_algo env' evd aliases 0 ev'' ev' in
- Evd.define evk' body evd
- with
- | EvarSolvedOnTheFly _ -> assert false (* ev has no candidates *)
- | CannotProject filter'' ->
- (* ... or postpone the problem *)
- postpone_evar_evar (evar_define conv_algo) env' evd filter'' ev'' filter' ev' in
- evdref := evd;
- evar'')
- | _ ->
- progress := true;
- match
- let c,args = decompose_app_vect t in
- match kind_of_term c with
- | Construct cstr when noccur_between 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 *)
- (* alternative inversion of the subterms of the constructor, etc)*)
- (match find_projectable_constructor env evd cstr k args cstr_subst with
- | _::_ as l -> Some (List.map mkVar l)
- | _ -> None)
- | _ -> None
- with
- | Some l ->
- let ty = get_type_of env' !evdref t in
- let candidates =
- try
- let t =
- map_constr_with_full_binders (fun d (env,k) -> push_rel d env, k+1)
- imitate envk t in
- t::l
- with e when Errors.noncritical e -> l in
- (match candidates with
- | [x] -> x
- | _ ->
- let (evd,evar'',ev'') =
- materialize_evar (evar_define conv_algo) env' !evdref k ev ty in
- evdref := restrict_evar evd (fst ev'') None (Some candidates);
- 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)
- imitate envk t in
-
- let rhs = whd_beta evd rhs (* heuristic *) in
- let body = imitate (env,0) rhs in
- (!evdref,body)
-
-(* [define] tries to solve the problem "?ev[args] = rhs" when "?ev" is
- * an (uninstantiated) evar such that "hyps |- ?ev : typ". Otherwise said,
- * [define] tries to find an instance lhs such that
- * "lhs [hyps:=args]" unifies to rhs. The term "lhs" must be closed in
- * context "hyps" and not referring to itself.
- *)
-
-and evar_define conv_algo ?(choose=false) env evd (evk,argsv as ev) rhs =
- match kind_of_term rhs with
- | Evar (evk2,argsv2 as ev2) ->
- if evk = evk2 then
- solve_refl ~can_drop:choose conv_algo env evd evk argsv argsv2
- else
- solve_evar_evar ~force:choose
- (evar_define conv_algo) conv_algo env evd ev ev2
- | _ ->
- try solve_candidates conv_algo env evd ev rhs
- with NoCandidates ->
- try
- let (evd',body) = invert_definition conv_algo choose env evd ev rhs in
- if occur_meta body then error "Meta cannot occur in evar body.";
- (* invert_definition may have instantiate some evars of rhs with evk *)
- (* so we recheck acyclicity *)
- if occur_evar evk body then raise (OccurCheckIn (evd',body));
- (* needed only if an inferred type *)
- let body = refresh_universes body in
-(* Cannot strictly type instantiations since the unification algorithm
- * does not unify applications from left to right.
- * e.g problem f x == g y yields x==y and f==g (in that order)
- * Another problem is that type variables are evars of type Type
- let _ =
- try
- let env = evar_env evi in
- let ty = evi.evar_concl in
- Typing.check env evd' body ty
- with e ->
- pperrnl
- (str "Ill-typed evar instantiation: " ++ fnl() ++
- pr_evar_map evd' ++ fnl() ++
- str "----> " ++ int ev ++ str " := " ++
- print_constr body);
- raise e in*)
- let evd' = Evd.define evk body evd' in
- check_evar_instance evd' evk body conv_algo
- with
- | NotEnoughInformationToProgress sols ->
- postpone_non_unique_projection env evd ev sols rhs
- | NotEnoughInformationEvarEvar t ->
- add_conv_pb (Reduction.CONV,env,mkEvar ev,t) evd
- | NotInvertibleUsingOurAlgorithm t ->
- error_not_clean env evd evk t (evar_source evk evd)
- | OccurCheckIn (evd,rhs) ->
- (* last chance: rhs actually reduces to ev *)
- let c = whd_betadeltaiota env evd rhs in
- match kind_of_term c with
- | Evar (evk',argsv2) when evk = evk' ->
- solve_refl
- (fun env sigma pb c c' -> (evd,is_fconv pb env sigma c c'))
- env evd evk argsv argsv2
- | _ ->
- error_occur_check env evd evk rhs
-
-(* This code (i.e. solve_pb, etc.) takes a unification
- * problem, and tries to solve it. If it solves it, then it removes
- * all the conversion problems, and re-runs conversion on each one, in
- * the hopes that the new solution will aid in solving them.
- *
- * The kinds of problems it knows how to solve are those in which
- * the usable arguments of an existential var are all themselves
- * universal variables.
- * The solution to this problem is to do renaming for the Var's,
- * to make them match up with the Var's which are found in the
- * hyps of the existential, to do a "pop" for each Rel which is
- * not an argument of the existential, and a subst1 for each which
- * is, again, with the corresponding variable. This is done by
- * define
- *
- * Thus, we take the arguments of the existential which we are about
- * to assign, and zip them with the identifiers in the hypotheses.
- * Then, we process all the Var's in the arguments, and sort the
- * Rel's into ascending order. Then, we just march up, doing
- * subst1's and pop's.
- *
- * NOTE: We can do this more efficiently for the relative arguments,
- * by building a long substituend by hand, but this is a pain in the
- * ass.
- *)
+let clear_hyps_in_evi env evdref hyps concl ids =
+ match clear_hyps_in_evi_main env evdref hyps [concl] ids with
+ | (nhyps,[nconcl]) -> (nhyps,nconcl)
+ | _ -> assert false
-let status_changed lev (pbty,_,t1,t2) =
- (try ExistentialSet.mem (head_evar t1) lev with NoHeadEvar -> false) or
- (try ExistentialSet.mem (head_evar t2) lev with NoHeadEvar -> false)
-
-let reconsider_conv_pbs conv_algo evd =
- let (evd,pbs) = extract_changed_conv_pbs evd status_changed in
- List.fold_left
- (fun (evd,b as p) (pbty,env,t1,t2) ->
- if b then conv_algo env evd pbty t1 t2 else p) (evd,true)
- pbs
-
-(* Tries to solve problem t1 = t2.
- * Precondition: t1 is an uninstantiated evar
- * Returns an optional list of evars that were instantiated, or None
- * if the problem couldn't be solved. *)
-
-(* Rq: uncomplete algorithm if pbty = CONV_X_LEQ ! *)
-let solve_simple_eqn conv_algo ?(choose=false) env evd (pbty,(evk1,args1 as ev1),t2) =
- try
- let t2 = whd_betaiota evd t2 in (* includes whd_evar *)
- let evd =
- match pbty with
- | Some true when isEvar t2 ->
- add_conv_pb (Reduction.CUMUL,env,mkEvar ev1,t2) evd
- | Some false when isEvar t2 ->
- add_conv_pb (Reduction.CUMUL,env,t2,mkEvar ev1) evd
- | _ ->
- evar_define conv_algo ~choose env evd ev1 t2 in
- reconsider_conv_pbs conv_algo evd
- with e when precatchable_exception e ->
- (evd,false)
-
-(** The following functions return the set of evars immediately
- contained in the object, including defined evars *)
-
-let evars_of_term c =
- let rec evrec acc c =
- match kind_of_term c with
- | Evar (n, l) -> Intset.add n (Array.fold_left evrec acc l)
- | _ -> fold_constr evrec acc c
- in
- evrec Intset.empty c
+let clear_hyps2_in_evi env evdref hyps t concl ids =
+ match clear_hyps_in_evi_main env evdref hyps [t;concl] ids with
+ | (nhyps,[t;nconcl]) -> (nhyps,t,nconcl)
+ | _ -> assert false
(* spiwack: a few functions to gather evars on which goals depend. *)
let queue_set q is_dependent set =
- Intset.iter (fun a -> Queue.push (is_dependent,a) q) set
+ Evar.Set.iter (fun a -> Queue.push (is_dependent,a) q) set
let queue_term q is_dependent c =
queue_set q is_dependent (evars_of_term c)
@@ -1830,7 +575,7 @@ let process_dependent_evar q acc evm is_dependent e =
end (Environ.named_context_of_val evi.evar_hyps);
match evi.evar_body with
| Evar_empty ->
- if is_dependent then Intmap.add e None acc else acc
+ if is_dependent then Evar.Map.add e None acc else acc
| Evar_defined b ->
let subevars = evars_of_term b in
(* evars appearing in the definition of an evar [e] are marked
@@ -1838,14 +583,14 @@ let process_dependent_evar q acc evm is_dependent e =
non-dependent goal, then, unless they are reach from another
path, these evars are just other non-dependent goals. *)
queue_set q is_dependent subevars;
- if is_dependent then Intmap.add e (Some subevars) acc else acc
+ if is_dependent then Evar.Map.add e (Some subevars) acc else acc
let gather_dependent_evars q evm =
- let acc = ref Intmap.empty in
+ let acc = ref Evar.Map.empty in
while not (Queue.is_empty q) do
let (is_dependent,e) = Queue.pop q in
(* checks if [e] has already been added to [!acc] *)
- begin if not (Intmap.mem e !acc) then
+ begin if not (Evar.Map.mem e !acc) then
acc := process_dependent_evar q !acc evm is_dependent e
end
done;
@@ -1858,21 +603,6 @@ let gather_dependent_evars evm l =
(* /spiwack *)
-let evars_of_named_context nc =
- List.fold_right (fun (_, b, t) s ->
- Option.fold_left (fun s t ->
- Intset.union s (evars_of_term t))
- (Intset.union s (evars_of_term t)) b)
- nc Intset.empty
-
-let evars_of_evar_info evi =
- Intset.union (evars_of_term evi.evar_concl)
- (Intset.union
- (match evi.evar_body with
- | Evar_empty -> Intset.empty
- | Evar_defined b -> evars_of_term b)
- (evars_of_named_context (named_context_of_val evi.evar_hyps)))
-
(** The following functions return the set of undefined evars
contained in the object, the defined evars being traversed.
This is roughly a combination of the previous functions and
@@ -1884,25 +614,25 @@ let undefined_evars_of_term evd t =
| Evar (n, l) ->
let acc = Array.fold_left evrec acc l in
(try match (Evd.find evd n).evar_body with
- | Evar_empty -> Intset.add n acc
+ | Evar_empty -> Evar.Set.add n acc
| Evar_defined c -> evrec acc c
- with Not_found -> anomaly "undefined_evars_of_term: evar not found")
+ with Not_found -> anomaly ~label:"undefined_evars_of_term" (Pp.str "evar not found"))
| _ -> fold_constr evrec acc c
in
- evrec Intset.empty t
+ evrec Evar.Set.empty t
let undefined_evars_of_named_context evd nc =
List.fold_right (fun (_, b, t) s ->
Option.fold_left (fun s t ->
- Intset.union s (undefined_evars_of_term evd t))
- (Intset.union s (undefined_evars_of_term evd t)) b)
- nc Intset.empty
+ Evar.Set.union s (undefined_evars_of_term evd t))
+ (Evar.Set.union s (undefined_evars_of_term evd t)) b)
+ nc Evar.Set.empty
let undefined_evars_of_evar_info evd evi =
- Intset.union (undefined_evars_of_term evd evi.evar_concl)
- (Intset.union
+ Evar.Set.union (undefined_evars_of_term evd evi.evar_concl)
+ (Evar.Set.union
(match evi.evar_body with
- | Evar_empty -> Intset.empty
+ | Evar_empty -> Evar.Set.empty
| Evar_defined b -> undefined_evars_of_term evd b)
(undefined_evars_of_named_context evd
(named_context_of_val evi.evar_hyps)))
@@ -1919,21 +649,29 @@ let check_evars env initial_sigma sigma c =
if not (Evd.mem initial_sigma evk) then
let (loc,k) = evar_source evk sigma in
match k with
- | ImplicitArg (gr, (i, id), false) -> ()
- | _ ->
- let evi = nf_evar_info sigma (Evd.find_undefined sigma evk) in
- error_unsolvable_implicit loc env sigma evi k None)
+ | Evar_kinds.ImplicitArg (gr, (i, id), false) -> ()
+ | _ -> error_unsolvable_implicit loc env sigma evk None)
| _ -> iter_constr proc_rec c
in proc_rec c
-open Glob_term
+(* spiwack: this is a more complete version of
+ {!Termops.occur_evar}. The latter does not look recursively into an
+ [evar_map]. If unification only need to check superficially, tactics
+ do not have this luxury, and need the more complete version. *)
+let occur_evar_upto sigma n c =
+ let rec occur_rec c = match kind_of_term c with
+ | Evar (sp,_) when Evar.equal sp n -> raise Occur
+ | Evar e -> Option.iter occur_rec (existential_opt_value sigma e)
+ | _ -> iter_constr occur_rec c
+ in
+ try occur_rec c; false with Occur -> true
+
(****************************************)
(* Operations on value/type constraints *)
(****************************************)
-type type_constraint_type = (int * int) option * constr
-type type_constraint = type_constraint_type option
+type type_constraint = types option
type val_constraint = constr option
@@ -1954,14 +692,8 @@ type val_constraint = constr option
(* The empty type constraint *)
let empty_tycon = None
-let mk_tycon_type c = (None, c)
-let mk_abstr_tycon_type n c = (Some (n, n), c) (* First component is initial abstraction, second
- is current abstraction *)
-
(* Builds a type constraint *)
-let mk_tycon ty = Some (mk_tycon_type ty)
-
-let mk_abstr_tycon n ty = Some (mk_abstr_tycon_type n ty)
+let mk_tycon ty = Some ty
(* Constrains the value of a type *)
let empty_valcon = None
@@ -1969,24 +701,34 @@ let empty_valcon = None
(* Builds a value constraint *)
let mk_valcon c = Some c
-
-let idx = id_of_string "x"
+let idx = Namegen.default_dependent_ident
(* Refining an evar to a product *)
let define_pure_evar_as_product evd evk =
let evi = Evd.find_undefined evd evk in
- let evenv = evar_unfiltered_env evi in
+ let evenv = evar_env evi in
let id = next_ident_away idx (ids_of_named_context (evar_context evi)) in
- let evd1,dom = new_type_evar evd evenv ~filter:(evar_filter evi) in
+ let concl = whd_evar evd evi.evar_concl in
+ let s = destSort concl in
+ let evd1,(dom,u1) = new_type_evar evenv evd univ_flexible_alg ~filter:(evar_filter evi) in
let evd2,rng =
let newenv = push_named (id, None, dom) evenv in
let src = evar_source evk evd1 in
- let filter = true::evar_filter evi in
- new_type_evar evd1 newenv ~src ~filter in
+ let filter = Filter.extend 1 (evar_filter evi) in
+ if is_prop_sort s then
+ (* Impredicative product, conclusion must fall in [Prop]. *)
+ new_evar newenv evd1 concl ~src ~filter
+ else
+ let evd3, (rng, srng) =
+ new_type_evar newenv evd1 univ_flexible_alg ~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
+ evd3, rng
+ in
let prod = mkProd (Name id, dom, subst_var id rng) in
let evd3 = Evd.define evk prod evd2 in
- evd3,prod
+ evd3,prod
(* Refine an applied evar to a product and returns its instantiation *)
@@ -1995,7 +737,7 @@ let define_evar_as_product evd (evk,args) =
(* 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 evrngargs = array_cons (mkRel 1) (Array.map (lift 1) 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)
@@ -2010,19 +752,19 @@ let define_evar_as_product evd (evk,args) =
let define_pure_evar_as_lambda env evd evk =
let evi = Evd.find_undefined evd evk in
- let evenv = evar_unfiltered_env evi in
+ let evenv = evar_env evi in
let typ = whd_betadeltaiota env evd (evar_concl evi) in
let evd1,(na,dom,rng) = match kind_of_term 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 dummy_loc env evd typ in
+ | _ -> error_not_product_loc Loc.ghost env evd typ in
let avoid = ids_of_named_context (evar_context evi) in
let id =
next_name_away_with_default_using_types "x" na avoid (whd_evar evd dom) in
let newenv = push_named (id, None, dom) evenv in
- let filter = true::evar_filter evi in
+ let filter = Filter.extend 1 (evar_filter evi) in
let src = evar_source evk evd1 in
- let evd2,body = new_evar evd1 newenv ~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
@@ -2030,7 +772,7 @@ 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 evbodyargs = array_cons (mkRel 1) (Array.map (lift 1) args) 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)
@@ -2041,30 +783,29 @@ let rec evar_absorb_arguments env evd (evk,args as ev) = function
let evd,lam = define_pure_evar_as_lambda env evd evk in
let _,_,body = destLambda lam in
let evk = fst (destEvar body) in
- evar_absorb_arguments env evd (evk, array_cons a args) l
+ evar_absorb_arguments env evd (evk, Array.cons a args) l
(* Refining an evar to a sort *)
-let define_evar_as_sort evd (ev,args) =
- let evd, s = new_sort_variable evd in
- Evd.define ev (mkSort s) evd, s
+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 evd' = Evd.define ev (mkSort s) evd in
+ Evd.set_leq_sort env evd' (Type (Univ.super u)) (destSort evi.evar_concl), s
(* We don't try to guess in which sort the type should be defined, since
any type has type Type. May cause some trouble, but not so far... *)
let judge_of_new_Type evd =
- let evd', s = new_univ_variable evd in
- evd', Typeops.judge_of_type s
+ let evd', s = new_univ_variable univ_rigid evd in
+ evd', { uj_val = mkSort (Type s); uj_type = mkSort (Type (Univ.super 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 unlift_tycon init cur c =
- if cur = 1 then None, c
- else Some (init, pred cur), c
-
let split_tycon loc env evd tycon =
let rec real_split evd c =
let t = whd_betadeltaiota env evd c in
@@ -2081,35 +822,19 @@ let split_tycon loc env evd tycon =
in
match tycon with
| None -> evd,(Anonymous,None,None)
- | Some (abs, c) ->
- (match abs with
- None ->
- let evd', (n, dom, rng) = real_split evd c in
- evd', (n, mk_tycon dom, mk_tycon rng)
- | Some (init, cur) ->
- evd, (Anonymous, None, Some (unlift_tycon init cur c)))
-
-let valcon_of_tycon x =
- match x with
- | Some (None, t) -> Some t
- | _ -> None
-
-let lift_abstr_tycon_type n (abs, t) =
- match abs with
- None -> raise (Invalid_argument "lift_abstr_tycon_type: not an abstraction")
- | Some (init, abs) ->
- let abs' = abs + n in
- if abs' < 0 then raise (Invalid_argument "lift_abstr_tycon_type")
- else (Some (init, abs'), t)
-
-let lift_tycon_type n (abs, t) = (abs, lift n t)
-let lift_tycon n = Option.map (lift_tycon_type n)
-
-let pr_tycon_type env (abs, t) =
- match abs with
- None -> Termops.print_constr_env env t
- | Some (init, cur) -> str "Abstract (" ++ int init ++ str "," ++ int cur ++ str ") " ++ Termops.print_constr_env env t
+ | Some c ->
+ let evd', (n, dom, rng) = real_split evd c in
+ evd', (n, mk_tycon dom, mk_tycon rng)
+
+let valcon_of_tycon x = x
+let lift_tycon n = Option.map (lift n)
let pr_tycon env = function
None -> str "None"
- | Some t -> pr_tycon_type env t
+ | Some t -> Termops.print_constr_env env t
+
+let subterm_source evk (loc,k) =
+ let evk = match k with
+ | Evar_kinds.SubEvar (evk) -> evk
+ | _ -> evk in
+ (loc,Evar_kinds.SubEvar evk)
diff --git a/pretyping/evarutil.mli b/pretyping/evarutil.mli
index 9269f138..f89266a6 100644
--- a/pretyping/evarutil.mli
+++ b/pretyping/evarutil.mli
@@ -1,19 +1,16 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Util
open Names
-open Glob_term
open Term
-open Sign
+open Context
open Evd
open Environ
-open Reductionops
(** {5 This modules provides useful functions for unification modulo evars } *)
@@ -23,23 +20,49 @@ open Reductionops
val new_meta : unit -> metavariable
val mk_new_meta : unit -> constr
-(** [new_untyped_evar] is a generator of unique evar keys *)
-val new_untyped_evar : unit -> existential_key
-
(** {6 Creating a fresh evar given their type and context} *)
val new_evar :
- evar_map -> env -> ?src:loc * hole_kind -> ?filter:bool list ->
- ?candidates:constr list -> types -> evar_map * constr
+ env -> evar_map -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t ->
+ ?candidates:constr list -> ?store:Store.t ->
+ ?naming:Misctypes.intro_pattern_naming_expr ->
+ ?principal:bool -> types -> evar_map * constr
+
+val new_pure_evar :
+ named_context_val -> evar_map -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t ->
+ ?candidates:constr list -> ?store:Store.t ->
+ ?naming:Misctypes.intro_pattern_naming_expr ->
+ ?principal:bool -> types -> evar_map * evar
+
+val new_pure_evar_full : evar_map -> evar_info -> evar_map * evar
(** the same with side-effects *)
val e_new_evar :
- evar_map ref -> env -> ?src:loc * hole_kind -> ?filter:bool list ->
- ?candidates:constr list -> types -> constr
+ env -> evar_map ref -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t ->
+ ?candidates:constr list -> ?store:Store.t ->
+ ?naming:Misctypes.intro_pattern_naming_expr ->
+ ?principal:bool -> types -> constr
(** Create a new Type existential variable, as we keep track of
them during type-checking and unification. *)
val new_type_evar :
- ?src:loc * hole_kind -> ?filter:bool list -> evar_map -> env -> evar_map * constr
+ env -> evar_map -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t ->
+ ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> rigid ->
+ evar_map * (constr * sorts)
+
+val e_new_type_evar : env -> evar_map ref ->
+ ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t ->
+ ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> rigid -> constr * sorts
+
+val new_Type : ?rigid:rigid -> env -> evar_map -> evar_map * constr
+val e_new_Type : ?rigid:rigid -> env -> evar_map ref -> constr
+
+val restrict_evar : evar_map -> existential_key -> Filter.t ->
+ constr list option -> evar_map * existential_key
+
+(** Polymorphic constants *)
+
+val new_global : evar_map -> Globnames.global_reference -> evar_map * constr
+val e_new_global : evar_map ref -> Globnames.global_reference -> constr
(** Create a fresh evar in a context different from its definition context:
[new_evar_instance sign evd ty inst] creates a new evar of context
@@ -48,30 +71,17 @@ val new_type_evar :
of [inst] are typed in the occurrence context and their type (seen
as a telescope) is [sign] *)
val new_evar_instance :
- named_context_val -> evar_map -> types -> ?src:loc * hole_kind -> ?filter:bool list -> ?candidates:constr list -> constr list -> evar_map * constr
-
-val make_pure_subst : evar_info -> constr array -> (identifier * constr) list
+ named_context_val -> evar_map -> types ->
+ ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> ?candidates:constr list ->
+ ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr ->
+ ?principal:bool ->
+ constr list -> evar_map * constr
-(** {6 Instantiate evars} *)
-
-type conv_fun =
- env -> evar_map -> conv_pb -> constr -> constr -> evar_map * bool
-
-(** [evar_define choose env ev c] try to instantiate [ev] with [c] (typed in [env]),
- possibly solving related unification problems, possibly leaving open
- some problems that cannot be solved in a unique way (except if choose is
- true); fails if the instance is not valid for the given [ev] *)
-val evar_define : conv_fun -> ?choose:bool -> env -> evar_map ->
- existential -> constr -> evar_map
+val make_pure_subst : evar_info -> constr array -> (Id.t * constr) list
(** {6 Evars/Metas switching...} *)
-(** [evars_to_metas] generates new metavariables for each non dependent
- existential and performs the replacement in the given constr; it also
- returns the evar_map extended with dependent evars *)
-val evars_to_metas : evar_map -> open_constr -> (evar_map * constr)
-
-val non_instantiated : evar_map -> (evar * evar_info) list
+val non_instantiated : evar_map -> evar_info Evar.Map.t
(** {6 Unification utils} *)
@@ -82,46 +92,26 @@ val head_evar : constr -> existential_key (** may raise NoHeadEvar *)
(* Expand head evar if any *)
val whd_head_evar : evar_map -> constr -> constr
+(* An over-approximation of [has_undefined (nf_evars evd c)] *)
+val has_undefined_evars : evar_map -> constr -> bool
+
val is_ground_term : evar_map -> constr -> bool
val is_ground_env : evar_map -> env -> bool
-val solve_refl : ?can_drop:bool -> conv_fun -> env -> evar_map ->
- existential_key -> constr array -> constr array -> evar_map
-val solve_evar_evar : ?force:bool ->
- (env -> evar_map -> existential -> constr -> evar_map) -> conv_fun ->
- env -> evar_map -> existential -> existential -> evar_map
-
-val solve_simple_eqn : conv_fun -> ?choose:bool -> env -> evar_map ->
- bool option * existential * constr -> evar_map * bool
-val reconsider_conv_pbs : conv_fun -> evar_map -> evar_map * bool
-
(** [check_evars env initial_sigma extended_sigma c] fails if some
new unresolved evar remains in [c] *)
val check_evars : env -> evar_map -> evar_map -> constr -> unit
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 : evar_map -> existential -> evar_map * sorts
-
-val is_unification_pattern_evar : env -> evar_map -> existential -> constr list ->
- constr -> constr list option
-
-val is_unification_pattern : env * int -> evar_map -> constr -> constr list ->
- constr -> constr list option
+val define_evar_as_sort : env -> evar_map -> existential -> evar_map * sorts
+(** Instantiate an evar by as many lambda's as needed so that its arguments
+ are moved to the evar substitution (i.e. turn [?x[vars1:=args1] args] into
+ [?y[vars1:=args1,vars:=args]] with
+ [vars1 |- ?x:=\vars.?y[vars1:=vars1,vars:=vars]] *)
val evar_absorb_arguments : env -> evar_map -> existential -> constr list ->
evar_map * existential
-val solve_pattern_eqn : env -> constr list -> constr -> constr
-
-(** The following functions return the set of evars immediately
- contained in the object, including defined evars *)
-
-
-val evars_of_term : constr -> Intset.t
-
-val evars_of_named_context : named_context -> Intset.t
-val evars_of_evar_info : evar_info -> Intset.t
-
(** [gather_dependent_evars evm seeds] classifies the evars in [evm]
as dependent_evars and goals (these may overlap). A goal is an
evar in [seeds] or an evar appearing in the (partial) definition
@@ -131,43 +121,39 @@ val evars_of_evar_info : evar_info -> Intset.t
associating to each dependent evar [None] if it has no (partial)
definition or [Some s] if [s] is the list of evars appearing in
its (partial) definition. *)
-val gather_dependent_evars : evar_map -> evar list -> (Intset.t option) Intmap.t
+val gather_dependent_evars : evar_map -> evar list -> (Evar.Set.t option) Evar.Map.t
(** The following functions return the set of undefined evars
contained in the object, the defined evars being traversed.
This is roughly a combination of the previous functions and
[nf_evar]. *)
-val undefined_evars_of_term : evar_map -> constr -> Intset.t
-val undefined_evars_of_named_context : evar_map -> named_context -> Intset.t
-val undefined_evars_of_evar_info : evar_map -> evar_info -> Intset.t
+val undefined_evars_of_term : evar_map -> constr -> Evar.Set.t
+val undefined_evars_of_named_context : evar_map -> named_context -> Evar.Set.t
+val undefined_evars_of_evar_info : evar_map -> evar_info -> Evar.Set.t
+
+(** [occur_evar_upto sigma k c] returns [true] if [k] appears in
+ [c]. It looks up recursively in [sigma] for the value of existential
+ variables. *)
+val occur_evar_upto : evar_map -> Evar.t -> Constr.t -> bool
(** {6 Value/Type constraints} *)
val judge_of_new_Type : evar_map -> evar_map * unsafe_judgment
-type type_constraint_type = (int * int) option * constr
-type type_constraint = type_constraint_type option
-
+type type_constraint = types option
type val_constraint = constr option
val empty_tycon : type_constraint
-val mk_tycon_type : constr -> type_constraint_type
-val mk_abstr_tycon_type : int -> constr -> type_constraint_type
val mk_tycon : constr -> type_constraint
-val mk_abstr_tycon : int -> constr -> type_constraint
val empty_valcon : val_constraint
val mk_valcon : constr -> val_constraint
val split_tycon :
- loc -> env -> evar_map -> type_constraint ->
- evar_map * (name * type_constraint * type_constraint)
+ Loc.t -> env -> evar_map -> type_constraint ->
+ evar_map * (Name.t * type_constraint * type_constraint)
val valcon_of_tycon : type_constraint -> val_constraint
-
-val lift_abstr_tycon_type : int -> type_constraint_type -> type_constraint_type
-
-val lift_tycon_type : int -> type_constraint_type -> type_constraint_type
val lift_tycon : int -> type_constraint -> type_constraint
(***********************************************************)
@@ -192,17 +178,29 @@ val nf_evar_info : evar_map -> evar_info -> evar_info
val nf_evar_map : evar_map -> evar_map
val nf_evar_map_undefined : evar_map -> evar_map
+val env_nf_evar : evar_map -> env -> env
+val env_nf_betaiotaevar : evar_map -> env -> env
+
+val j_nf_betaiotaevar : evar_map -> unsafe_judgment -> unsafe_judgment
+val jv_nf_betaiotaevar :
+ evar_map -> unsafe_judgment array -> unsafe_judgment array
+(** Presenting terms without solved evars *)
+
+val nf_evars_universes : evar_map -> constr -> constr
+
+val nf_evars_and_universes : evar_map -> evar_map * (constr -> constr)
+val e_nf_evars_and_universes : evar_map ref -> (constr -> constr) * Universes.universe_opt_subst
+
+(** Normalize the evar map w.r.t. universes, after simplification of constraints.
+ Return the substitution function for constrs as well. *)
+val nf_evar_map_universes : evar_map -> evar_map * (constr -> constr)
+
(** Replacing all evars, possibly raising [Uninstantiated_evar] *)
exception Uninstantiated_evar of existential_key
val flush_and_check_evars : evar_map -> constr -> constr
-(** 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
-
(** {6 debug pretty-printer:} *)
-val pr_tycon_type : env -> type_constraint_type -> Pp.std_ppcmds
val pr_tycon : env -> type_constraint -> Pp.std_ppcmds
@@ -210,24 +208,31 @@ val pr_tycon : env -> type_constraint -> Pp.std_ppcmds
raise OccurHypInSimpleClause if the removal breaks dependencies *)
type clear_dependency_error =
-| OccurHypInSimpleClause of identifier option
+| OccurHypInSimpleClause of Id.t option
| EvarTypingBreak of existential
-exception ClearDependencyError of identifier * clear_dependency_error
+exception ClearDependencyError of Id.t * clear_dependency_error
(* spiwack: marks an evar that has been "defined" by clear.
used by [Goal] and (indirectly) [Proofview] to handle the clear tactic gracefully*)
-val cleared : bool Store.Field.t
+val cleared : bool Store.field
-val clear_hyps_in_evi : evar_map ref -> named_context_val -> types ->
- identifier list -> named_context_val * types
+val clear_hyps_in_evi : env -> evar_map ref -> named_context_val -> types ->
+ Id.Set.t -> named_context_val * types
+
+val clear_hyps2_in_evi : env -> evar_map ref -> named_context_val -> types -> types ->
+ Id.Set.t -> named_context_val * types * types
val push_rel_context_to_named_context : Environ.env -> types ->
- named_context_val * types * constr list * constr list
+ named_context_val * types * constr list * constr list * (identifier*constr) list
val generalize_evar_over_rels : evar_map -> existential -> types * constr list
-val check_evar_instance : evar_map -> existential_key -> constr -> conv_fun ->
- evar_map
+(** Evar combinators *)
+
+val evd_comb0 : (evar_map -> evar_map * 'a) -> evar_map ref -> 'a
+val evd_comb1 : (evar_map -> 'b -> evar_map * 'a) -> evar_map ref -> 'b -> 'a
+val evd_comb2 : (evar_map -> 'b -> 'c -> evar_map * 'a) -> evar_map ref -> 'b -> 'c -> 'a
-val remove_instance_local_defs : evar_map -> existential_key -> constr list -> constr list
+val subterm_source : existential_key -> Evar_kinds.t Loc.located ->
+ Evar_kinds.t Loc.located
diff --git a/pretyping/evd.ml b/pretyping/evd.ml
index 72dc27e6..ee72d314 100644
--- a/pretyping/evd.ml
+++ b/pretyping/evd.ml
@@ -1,43 +1,134 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Pp
+open Errors
open Util
open Names
open Nameops
open Term
+open Vars
open Termops
-open Sign
open Environ
-open Libnames
-open Mod_subst
+open Globnames
+
+(** Generic filters *)
+module Filter :
+sig
+ type t
+ val equal : t -> t -> bool
+ val identity : t
+ val filter_list : t -> 'a list -> 'a list
+ val filter_array : t -> 'a array -> 'a array
+ val extend : int -> t -> t
+ val compose : t -> t -> t
+ val apply_subfilter : t -> bool list -> t
+ val restrict_upon : t -> int -> (int -> bool) -> t option
+ val map_along : (bool -> 'a -> bool) -> t -> 'a list -> t
+ val make : bool list -> t
+ val repr : t -> bool list option
+end =
+struct
+ type t = bool list option
+ (** We guarantee through the interface that if a filter is [Some _] then it
+ contains at least one [false] somewhere. *)
+
+ let identity = None
+
+ let rec equal l1 l2 = match l1, l2 with
+ | [], [] -> true
+ | h1 :: l1, h2 :: l2 ->
+ (if h1 then h2 else not h2) && equal l1 l2
+ | _ -> false
+
+ let equal l1 l2 = match l1, l2 with
+ | None, None -> true
+ | Some _, None | None, Some _ -> false
+ | Some l1, Some l2 -> equal l1 l2
+
+ let rec is_identity = function
+ | [] -> true
+ | true :: l -> is_identity l
+ | false :: _ -> false
+
+ let normalize f = if is_identity f then None else Some f
+
+ let filter_list f l = match f with
+ | None -> l
+ | Some f -> CList.filter_with f l
+
+ let filter_array f v = match f with
+ | None -> v
+ | Some f -> CArray.filter_with f v
+
+ let rec extend n l =
+ if n = 0 then l
+ else extend (pred n) (true :: l)
+
+ let extend n = function
+ | None -> None
+ | Some f -> Some (extend n f)
+
+ let compose f1 f2 = match f1 with
+ | None -> f2
+ | Some f1 ->
+ match f2 with
+ | None -> None
+ | Some f2 -> normalize (CList.filter_with f1 f2)
+
+ let apply_subfilter_array filter subfilter =
+ (** In both cases we statically know that the argument will contain at
+ least one [false] *)
+ match filter with
+ | None -> Some (Array.to_list subfilter)
+ | Some f ->
+ let len = Array.length subfilter in
+ let fold b (i, ans) =
+ if b then
+ let () = assert (0 <= i) in
+ (pred i, Array.unsafe_get subfilter i :: ans)
+ else
+ (i, false :: ans)
+ in
+ Some (snd (List.fold_right fold f (pred len, [])))
+
+ let apply_subfilter filter subfilter =
+ apply_subfilter_array filter (Array.of_list subfilter)
+
+ let restrict_upon f len p =
+ let newfilter = Array.init len p in
+ if Array.for_all (fun id -> id) newfilter then None
+ else
+ Some (apply_subfilter_array f newfilter)
-(* The kinds of existential variable *)
+ let map_along f flt l =
+ let ans = match flt with
+ | None -> List.map (fun x -> f true x) l
+ | Some flt -> List.map2 f flt l
+ in
+ normalize ans
-type obligation_definition_status = Define of bool | Expand
+ let make l = normalize l
+
+ let repr f = f
+
+end
-type hole_kind =
- | ImplicitArg of global_reference * (int * identifier option) * bool
- | BinderType of name
- | QuestionMark of obligation_definition_status
- | CasesType
- | InternalHole
- | TomatchTypeParameter of inductive * int
- | GoalEvar
- | ImpossibleCase
- | MatchingVar of bool * identifier
+(* The kinds of existential variables are now defined in [Evar_kinds] *)
(* The type of mappings for existential variables *)
-type evar = existential_key
+module Dummy = struct end
+module Store = Store.Make(Dummy)
-let string_of_existential evk = "?" ^ string_of_int evk
-let existential_of_int evk = evk
+type evar = Term.existential_key
+
+let string_of_existential evk = "?X" ^ string_of_int (Evar.repr evk)
type evar_body =
| Evar_empty
@@ -47,8 +138,8 @@ type evar_info = {
evar_concl : constr;
evar_hyps : named_context_val;
evar_body : evar_body;
- evar_filter : bool list;
- evar_source : hole_kind located;
+ evar_filter : Filter.t;
+ evar_source : Evar_kinds.t Loc.located;
evar_candidates : constr list option; (* if not None, list of allowed instances *)
evar_extra : Store.t }
@@ -56,194 +147,311 @@ let make_evar hyps ccl = {
evar_concl = ccl;
evar_hyps = hyps;
evar_body = Evar_empty;
- evar_filter = List.map (fun _ -> true) (named_context_of_val hyps);
- evar_source = (dummy_loc,InternalHole);
+ evar_filter = Filter.identity;
+ evar_source = (Loc.ghost,Evar_kinds.InternalHole);
evar_candidates = None;
evar_extra = Store.empty
}
+let instance_mismatch () =
+ anomaly (Pp.str "Signature and its instance do not match")
+
let evar_concl evi = evi.evar_concl
-let evar_hyps evi = evi.evar_hyps
-let evar_context evi = named_context_of_val evi.evar_hyps
-let evar_body evi = evi.evar_body
+
let evar_filter evi = evi.evar_filter
+
+let evar_body evi = evi.evar_body
+
+let evar_context evi = named_context_of_val evi.evar_hyps
+
let evar_filtered_context evi =
- snd (list_filter2 (fun b c -> b) (evar_filter evi,evar_context evi))
-let evar_filtered_hyps evi =
- List.fold_right push_named_context_val (evar_filtered_context evi)
- empty_named_context_val
-let evar_unfiltered_env evi = Global.env_of_context evi.evar_hyps
-let evar_env evi =
- List.fold_right push_named (evar_filtered_context evi)
- (reset_context (Global.env()))
-
-let eq_evar_info ei1 ei2 =
- ei1 == ei2 ||
- eq_constr ei1.evar_concl ei2.evar_concl &&
- eq_named_context_val (ei1.evar_hyps) (ei2.evar_hyps) &&
- ei1.evar_body = ei2.evar_body
-
-(* spiwack: Revised hierarchy :
- - ExistentialMap ( Maps of existential_keys )
- - EvarInfoMap ( .t = evar_info ExistentialMap.t * evar_info ExistentialMap )
- - EvarMap ( .t = EvarInfoMap.t * sort_constraints )
- - evar_map (exported)
-*)
+ Filter.filter_list (evar_filter evi) (evar_context evi)
-module ExistentialMap = Intmap
-module ExistentialSet = Intset
+let evar_hyps evi = evi.evar_hyps
+
+let evar_filtered_hyps evi = match Filter.repr (evar_filter evi) with
+| None -> evar_hyps evi
+| Some filter ->
+ let rec make_hyps filter ctxt = match filter, ctxt with
+ | [], [] -> empty_named_context_val
+ | false :: filter, _ :: ctxt -> make_hyps filter ctxt
+ | true :: filter, decl :: ctxt ->
+ let hyps = make_hyps filter ctxt in
+ push_named_context_val decl hyps
+ | _ -> instance_mismatch ()
+ in
+ make_hyps filter (evar_context evi)
+
+let evar_env evi = Global.env_of_context evi.evar_hyps
+
+let evar_filtered_env evi = match Filter.repr (evar_filter evi) with
+| None -> evar_env evi
+| Some filter ->
+ let rec make_env filter ctxt = match filter, ctxt with
+ | [], [] -> reset_context (Global.env ())
+ | false :: filter, _ :: ctxt -> make_env filter ctxt
+ | true :: filter, decl :: ctxt ->
+ let env = make_env filter ctxt in
+ push_named decl env
+ | _ -> instance_mismatch ()
+ in
+ make_env filter (evar_context evi)
+
+let map_evar_body f = function
+ | Evar_empty -> Evar_empty
+ | Evar_defined d -> Evar_defined (f d)
+
+let map_evar_info f evi =
+ {evi with
+ evar_body = map_evar_body f evi.evar_body;
+ evar_hyps = map_named_val f evi.evar_hyps;
+ evar_concl = f evi.evar_concl;
+ evar_candidates = Option.map (List.map f) evi.evar_candidates }
+
+let evar_ident_info evi =
+ match evi.evar_source with
+ | _,Evar_kinds.ImplicitArg (c,(n,Some id),b) -> id
+ | _,Evar_kinds.VarInstance id -> id
+ | _,Evar_kinds.GoalEvar -> Id.of_string "Goal"
+ | _ ->
+ let env = reset_with_named_context evi.evar_hyps (Global.env()) in
+ Namegen.id_of_name_using_hdchar env evi.evar_concl Anonymous
(* This exception is raised by *.existential_value *)
exception NotInstantiatedEvar
(* Note: let-in contributes to the instance *)
-let make_evar_instance sign args =
- let rec instrec = function
- | (id,_,_) :: sign, c::args when isVarId id c -> instrec (sign,args)
- | (id,_,_) :: sign, c::args -> (id,c) :: instrec (sign,args)
- | [],[] -> []
- | [],_ | _,[] -> anomaly "Signature and its instance do not match"
- in
- instrec (sign,args)
-
-let instantiate_evar sign c args =
- let inst = make_evar_instance sign args in
- if inst = [] then c else replace_vars inst c
-
-module EvarInfoMap = struct
- type t = evar_info ExistentialMap.t * evar_info ExistentialMap.t
-
- let empty = ExistentialMap.empty, ExistentialMap.empty
-
- let is_empty (d,u) = ExistentialMap.is_empty d && ExistentialMap.is_empty u
-
- let has_undefined (_,u) = not (ExistentialMap.is_empty u)
-
- let to_list (def,undef) =
- (* Workaround for change in Map.fold behavior in ocaml 3.08.4 *)
- let l = ref [] in
- ExistentialMap.iter (fun evk x -> l := (evk,x)::!l) def;
- ExistentialMap.iter (fun evk x -> l := (evk,x)::!l) undef;
- !l
-
- let undefined_list (def,undef) =
- (* Order is important: needs ocaml >= 3.08.4 from which "fold" is a
- "fold_left" *)
- ExistentialMap.fold (fun evk evi l -> (evk,evi)::l) undef []
-
- let undefined_evars (def,undef) = (ExistentialMap.empty,undef)
- let defined_evars (def,undef) = (def,ExistentialMap.empty)
-
- let find (def,undef) k =
- try ExistentialMap.find k def
- with Not_found -> ExistentialMap.find k undef
- let find_undefined (def,undef) k = ExistentialMap.find k undef
- let remove (def,undef) k =
- (ExistentialMap.remove k def,ExistentialMap.remove k undef)
- let mem (def,undef) k =
- ExistentialMap.mem k def || ExistentialMap.mem k undef
- let fold (def,undef) f a =
- ExistentialMap.fold f def (ExistentialMap.fold f undef a)
- let fold_undefined (def,undef) f a =
- ExistentialMap.fold f undef a
- let exists_undefined (def,undef) f =
- ExistentialMap.fold (fun k v b -> b || f k v) undef false
-
- let add (def,undef) evk newinfo =
- if newinfo.evar_body = Evar_empty then
- (def,ExistentialMap.add evk newinfo undef)
- else
- (ExistentialMap.add evk newinfo def,undef)
-
- let add_undefined (def,undef) evk newinfo =
- assert (newinfo.evar_body = Evar_empty);
- (def,ExistentialMap.add evk newinfo undef)
-
- let map f (def,undef) = (ExistentialMap.map f def, ExistentialMap.map f undef)
-
- let define (def,undef) evk body =
- let oldinfo =
- try ExistentialMap.find evk undef
- with Not_found ->
- try ExistentialMap.find evk def
- with Not_found ->
- anomaly "Evd.define: cannot define undeclared evar" in
- let newinfo =
- { oldinfo with
- evar_body = Evar_defined body } in
- match oldinfo.evar_body with
- | Evar_empty ->
- (ExistentialMap.add evk newinfo def,ExistentialMap.remove evk undef)
- | _ ->
- anomaly "Evd.define: cannot define an evar twice"
-
- let is_evar = mem
-
- let is_defined (def,undef) evk = ExistentialMap.mem evk def
- let is_undefined (def,undef) evk = ExistentialMap.mem evk undef
-
- (*******************************************************************)
- (* Formerly Instantiate module *)
-
- (* Existentials. *)
-
- let existential_type sigma (n,args) =
- let info =
- try find sigma n
- with Not_found ->
- anomaly ("Evar "^(string_of_existential n)^" was not declared") in
- let hyps = evar_filtered_context info in
- instantiate_evar hyps info.evar_concl (Array.to_list args)
-
- let existential_value sigma (n,args) =
- let info = find sigma n in
- let hyps = evar_filtered_context info in
- match evar_body info with
- | Evar_defined c ->
- instantiate_evar hyps c (Array.to_list args)
- | Evar_empty ->
- raise NotInstantiatedEvar
-
- let existential_opt_value sigma ev =
- try Some (existential_value sigma ev)
- with NotInstantiatedEvar -> None
+let evar_instance_array test_id info args =
+ let len = Array.length args in
+ let rec instrec filter ctxt i = match filter, ctxt with
+ | [], [] ->
+ if Int.equal i len then []
+ else instance_mismatch ()
+ | false :: filter, _ :: ctxt ->
+ instrec filter ctxt i
+ | true :: filter, (id, _, _) :: ctxt ->
+ if i < len then
+ let c = Array.unsafe_get args i in
+ if test_id id c then instrec filter ctxt (succ i)
+ else (id, c) :: instrec filter ctxt (succ i)
+ else instance_mismatch ()
+ | _ -> instance_mismatch ()
+ in
+ match Filter.repr (evar_filter info) with
+ | None ->
+ let map i (id, _, _) =
+ if (i < len) then
+ let c = Array.unsafe_get args i in
+ if test_id id c then None else Some (id,c)
+ else instance_mismatch ()
+ in
+ List.map_filter_i map (evar_context info)
+ | Some filter ->
+ instrec filter (evar_context info) 0
+
+let make_evar_instance_array info args =
+ evar_instance_array isVarId info args
+
+let instantiate_evar_array info c args =
+ let inst = make_evar_instance_array info args in
+ match inst with
+ | [] -> c
+ | _ -> replace_vars inst c
+
+module StringOrd = struct type t = string let compare = String.compare end
+module UNameMap = struct
+
+ include Map.Make(StringOrd)
+
+ let union s t =
+ if s == t then s
+ else
+ merge (fun k l r ->
+ match l, r with
+ | Some _, _ -> l
+ | _, _ -> r) s t
end
-module EvarMap = struct
- type t = EvarInfoMap.t * (Univ.UniverseLSet.t * Univ.universes)
- let empty = EvarInfoMap.empty, (Univ.UniverseLSet.empty, Univ.initial_universes)
- let is_empty (sigma,_) = EvarInfoMap.is_empty sigma
- let has_undefined (sigma,_) = EvarInfoMap.has_undefined sigma
- let add (sigma,sm) k v = (EvarInfoMap.add sigma k v, sm)
- let add_undefined (sigma,sm) k v = (EvarInfoMap.add_undefined sigma k v, sm)
- let find (sigma,_) = EvarInfoMap.find sigma
- let find_undefined (sigma,_) = EvarInfoMap.find_undefined sigma
- let remove (sigma,sm) k = (EvarInfoMap.remove sigma k, sm)
- let mem (sigma,_) = EvarInfoMap.mem sigma
- let to_list (sigma,_) = EvarInfoMap.to_list sigma
- let undefined_list (sigma,_) = EvarInfoMap.undefined_list sigma
- let undefined_evars (sigma,sm) = (EvarInfoMap.undefined_evars sigma, sm)
- let defined_evars (sigma,sm) = (EvarInfoMap.defined_evars sigma, sm)
- let fold (sigma,_) = EvarInfoMap.fold sigma
- let fold_undefined (sigma,_) = EvarInfoMap.fold_undefined sigma
- let define (sigma,sm) k v = (EvarInfoMap.define sigma k v, sm)
- let is_evar (sigma,_) = EvarInfoMap.is_evar sigma
- let is_defined (sigma,_) = EvarInfoMap.is_defined sigma
- let is_undefined (sigma,_) = EvarInfoMap.is_undefined sigma
- let existential_value (sigma,_) = EvarInfoMap.existential_value sigma
- let existential_type (sigma,_) = EvarInfoMap.existential_type sigma
- let existential_opt_value (sigma,_) = EvarInfoMap.existential_opt_value sigma
- let progress_evar_map (sigma1,sm1 as x) (sigma2,sm2 as y) = not (x == y) &&
- (EvarInfoMap.exists_undefined sigma1
- (fun k v -> assert (v.evar_body = Evar_empty);
- EvarInfoMap.is_defined sigma2 k))
-
- let merge e e' = fold e' (fun n v sigma -> add sigma n v) e
- let add_constraints (sigma, (us, sm)) cstrs =
- (sigma, (us, Univ.merge_constraints cstrs sm))
-end
+(* 2nd part used to check consistency on the fly. *)
+type evar_universe_context =
+ { uctx_names : Univ.Level.t UNameMap.t * string Univ.LMap.t;
+ uctx_local : Univ.universe_context_set; (** The local context of variables *)
+ uctx_univ_variables : Universes.universe_opt_subst;
+ (** The local universes that are unification variables *)
+ uctx_univ_algebraic : Univ.universe_set;
+ (** The subset of unification variables that
+ can be instantiated with algebraic universes as they appear in types
+ and universe instances only. *)
+ uctx_universes : Univ.universes; (** The current graph extended with the local constraints *)
+ uctx_initial_universes : Univ.universes; (** The graph at the creation of the evar_map *)
+ }
+
+let empty_evar_universe_context =
+ { uctx_names = UNameMap.empty, Univ.LMap.empty;
+ uctx_local = Univ.ContextSet.empty;
+ uctx_univ_variables = Univ.LMap.empty;
+ uctx_univ_algebraic = Univ.LSet.empty;
+ uctx_universes = Univ.initial_universes;
+ uctx_initial_universes = Univ.initial_universes }
+
+let evar_universe_context_from e =
+ let u = universes e in
+ {empty_evar_universe_context with
+ uctx_universes = u; uctx_initial_universes = u}
+
+let is_empty_evar_universe_context ctx =
+ Univ.ContextSet.is_empty ctx.uctx_local &&
+ Univ.LMap.is_empty ctx.uctx_univ_variables
+
+let union_evar_universe_context ctx ctx' =
+ if ctx == ctx' then ctx
+ else if is_empty_evar_universe_context ctx' then ctx
+ else
+ let local = Univ.ContextSet.union ctx.uctx_local ctx'.uctx_local in
+ let names = UNameMap.union (fst ctx.uctx_names) (fst ctx'.uctx_names) in
+ let names_rev = Univ.LMap.union (snd ctx.uctx_names) (snd ctx'.uctx_names) in
+ { uctx_names = (names, names_rev);
+ uctx_local = local;
+ uctx_univ_variables =
+ Univ.LMap.subst_union ctx.uctx_univ_variables ctx'.uctx_univ_variables;
+ uctx_univ_algebraic =
+ Univ.LSet.union ctx.uctx_univ_algebraic ctx'.uctx_univ_algebraic;
+ uctx_initial_universes = ctx.uctx_initial_universes;
+ uctx_universes =
+ if local == ctx.uctx_local then ctx.uctx_universes
+ else
+ let cstrsr = Univ.ContextSet.constraints ctx'.uctx_local in
+ Univ.merge_constraints cstrsr ctx.uctx_universes }
+
+(* let union_evar_universe_context_key = Profile.declare_profile "union_evar_universe_context";; *)
+(* let union_evar_universe_context = *)
+(* Profile.profile2 union_evar_universe_context_key union_evar_universe_context;; *)
+
+type 'a in_evar_universe_context = 'a * evar_universe_context
+
+let evar_universe_context_set ctx = ctx.uctx_local
+let evar_universe_context_constraints ctx = snd ctx.uctx_local
+let evar_context_universe_context ctx = Univ.ContextSet.to_context ctx.uctx_local
+let evar_universe_context_of ctx = { empty_evar_universe_context with uctx_local = ctx }
+let evar_universe_context_subst ctx = ctx.uctx_univ_variables
+
+let instantiate_variable l b v =
+ (* let b = Univ.subst_large_constraint (Univ.Universe.make l) Univ.type0m_univ b in *)
+ (* if Univ.univ_depends (Univ.Universe.make l) b then *)
+ (* error ("Occur-check in universe variable instantiation") *)
+ (* else *) v := Univ.LMap.add l (Some b) !v
+
+exception UniversesDiffer
+
+let process_universe_constraints univs vars alg cstrs =
+ let vars = ref vars in
+ let normalize = Universes.normalize_universe_opt_subst vars in
+ let rec unify_universes fo l d r local =
+ let l = normalize l and r = normalize r in
+ if Univ.Universe.equal l r then local
+ else
+ let varinfo x =
+ match Univ.Universe.level x with
+ | None -> Inl x
+ | Some l -> Inr (l, Univ.LMap.mem l !vars, Univ.LSet.mem l alg)
+ in
+ if d == Universes.ULe then
+ if Univ.check_leq univs l r then
+ (** Keep Prop/Set <= var around if var might be instantiated by prop or set
+ later. *)
+ if Univ.Universe.is_level l then
+ match Univ.Universe.level r with
+ | Some r ->
+ Univ.Constraint.add (Option.get (Univ.Universe.level l),Univ.Le,r) local
+ | _ -> local
+ else local
+ else
+ match Univ.Universe.level r with
+ | None -> error ("Algebraic universe on the right")
+ | Some rl ->
+ if Univ.Level.is_small rl then
+ let levels = Univ.Universe.levels l in
+ Univ.LSet.fold (fun l local ->
+ if Univ.Level.is_small l || Univ.LMap.mem l !vars then
+ Univ.enforce_eq (Univ.Universe.make l) r local
+ else raise (Univ.UniverseInconsistency (Univ.Le, Univ.Universe.make l, r, None)))
+ levels local
+ else
+ Univ.enforce_leq l r local
+ else if d == Universes.ULub then
+ match varinfo l, varinfo r with
+ | (Inr (l, true, _), Inr (r, _, _))
+ | (Inr (r, _, _), Inr (l, true, _)) ->
+ instantiate_variable l (Univ.Universe.make r) vars;
+ Univ.enforce_eq_level l r local
+ | Inr (_, _, _), Inr (_, _, _) ->
+ unify_universes true l Universes.UEq r local
+ | _, _ -> assert false
+ else (* d = Universes.UEq *)
+ match varinfo l, varinfo r with
+ | Inr (l', lloc, _), Inr (r', rloc, _) ->
+ let () =
+ if lloc then
+ instantiate_variable l' r vars
+ else if rloc then
+ instantiate_variable r' l vars
+ else if not (Univ.check_eq univs l r) then
+ (* Two rigid/global levels, none of them being local,
+ one of them being Prop/Set, disallow *)
+ if Univ.Level.is_small l' || Univ.Level.is_small r' then
+ raise (Univ.UniverseInconsistency (Univ.Eq, l, r, None))
+ else
+ if fo then
+ raise UniversesDiffer
+ in
+ Univ.enforce_eq_level l' r' local
+ | _, _ (* One of the two is algebraic or global *) ->
+ if Univ.check_eq univs l r then local
+ else raise (Univ.UniverseInconsistency (Univ.Eq, l, r, None))
+ in
+ let local =
+ Universes.Constraints.fold (fun (l,d,r) local -> unify_universes false l d r local)
+ cstrs Univ.Constraint.empty
+ in
+ !vars, local
+
+let add_constraints_context ctx cstrs =
+ let univs, local = ctx.uctx_local in
+ let cstrs' = Univ.Constraint.fold (fun (l,d,r) acc ->
+ let l = Univ.Universe.make l and r = Univ.Universe.make r in
+ let cstr' =
+ if d == Univ.Lt then (Univ.Universe.super l, Universes.ULe, r)
+ else (l, (if d == Univ.Le then Universes.ULe else Universes.UEq), r)
+ in Universes.Constraints.add cstr' acc)
+ cstrs Universes.Constraints.empty
+ in
+ let vars, local' =
+ process_universe_constraints ctx.uctx_universes
+ ctx.uctx_univ_variables ctx.uctx_univ_algebraic
+ cstrs'
+ in
+ { ctx with uctx_local = (univs, Univ.Constraint.union local local');
+ uctx_univ_variables = vars;
+ uctx_universes = Univ.merge_constraints cstrs ctx.uctx_universes }
+
+(* let addconstrkey = Profile.declare_profile "add_constraints_context";; *)
+(* let add_constraints_context = Profile.profile2 addconstrkey add_constraints_context;; *)
+
+let add_universe_constraints_context ctx cstrs =
+ let univs, local = ctx.uctx_local in
+ let vars, local' =
+ process_universe_constraints ctx.uctx_universes
+ ctx.uctx_univ_variables ctx.uctx_univ_algebraic
+ cstrs
+ in
+ { ctx with uctx_local = (univs, Univ.Constraint.union local local');
+ uctx_univ_variables = vars;
+ uctx_universes = Univ.merge_constraints local' ctx.uctx_universes }
+(* let addunivconstrkey = Profile.declare_profile "add_universe_constraints_context";; *)
+(* let add_universe_constraints_context = *)
+(* Profile.profile2 addunivconstrkey add_universe_constraints_context;; *)
(*******************************************************************)
(* Metamaps *)
@@ -253,16 +461,16 @@ end
type 'a freelisted = {
rebus : 'a;
- freemetas : Intset.t }
+ freemetas : Int.Set.t }
(* Collects all metavars appearing in a constr *)
let metavars_of c =
let rec collrec acc c =
match kind_of_term c with
- | Meta mv -> Intset.add mv acc
+ | Meta mv -> Int.Set.add mv acc
| _ -> fold_constr collrec acc c
in
- collrec Intset.empty c
+ collrec Int.Set.empty c
let mk_freelisted c =
{ rebus = c; freemetas = metavars_of c }
@@ -278,6 +486,8 @@ let map_fl f cfl = { cfl with rebus=f cfl.rebus }
type instance_constraint = IsSuperType | IsSubType | Conv
+let eq_instance_constraint c1 c2 = c1 == c2
+
(* Status of the unification of the type of an instance against the type of
the meta it instantiates:
- CoerceToType means that the unification of types has not been done
@@ -301,8 +511,8 @@ type instance_status = instance_constraint * instance_typing_status
(* Clausal environments *)
type clbinding =
- | Cltyp of name * constr freelisted
- | Clval of name * (constr freelisted * instance_status) * constr freelisted
+ | Cltyp of Name.t * constr freelisted
+ | Clval of Name.t * (constr freelisted * instance_status) * constr freelisted
let map_clb f = function
| Cltyp (na,cfl) -> Cltyp (na,map_fl f cfl)
@@ -315,11 +525,9 @@ let clb_name = function
(***********************)
-module Metaset = Intset
+module Metaset = Int.Set
-let meta_exists p s = Metaset.fold (fun x b -> b || (p x)) s false
-
-module Metamap = Intmap
+module Metamap = Int.Map
let metamap_to_list m =
Metamap.fold (fun n v l -> (n,v)::l) m []
@@ -329,136 +537,322 @@ let metamap_to_list m =
type conv_pb = Reduction.conv_pb
type evar_constraint = conv_pb * Environ.env * constr * constr
-type evar_map =
- { evars : EvarMap.t;
- conv_pbs : evar_constraint list;
- last_mods : ExistentialSet.t;
- metas : clbinding Metamap.t }
-(*** Lifting primitive from EvarMap. ***)
+module EvMap = Evar.Map
+
+type evar_map = {
+ (** Existential variables *)
+ defn_evars : evar_info EvMap.t;
+ undf_evars : evar_info EvMap.t;
+ evar_names : Id.t EvMap.t * existential_key Idmap.t;
+ (** Universes *)
+ universes : evar_universe_context;
+ (** Conversion problems *)
+ conv_pbs : evar_constraint list;
+ last_mods : Evar.Set.t;
+ (** Metas *)
+ metas : clbinding Metamap.t;
+ (** Interactive proofs *)
+ effects : Declareops.side_effects;
+ future_goals : Evar.t list; (** list of newly created evars, to be
+ eventually turned into goals if not solved.*)
+ principal_future_goal : Evar.t option; (** if [Some e], [e] must be
+ contained
+ [future_goals]. The evar
+ [e] will inherit
+ properties (now: the
+ name) of the evar which
+ will be instantiated with
+ a term containing [e]. *)
+}
+
+(*** Lifting primitive from Evar.Map. ***)
(* HH: The progress tactical now uses this function. *)
let progress_evar_map d1 d2 =
- EvarMap.progress_evar_map d1.evars d2.evars
-
-(* spiwack: tentative. It might very well not be the semantics we want
- for merging evar_map *)
-let merge d1 d2 = {
- evars = EvarMap.merge d1.evars d2.evars ;
- conv_pbs = List.rev_append d1.conv_pbs d2.conv_pbs ;
- last_mods = ExistentialSet.union d1.last_mods d2.last_mods ;
- metas = Metamap.fold (fun k m r -> Metamap.add k m r) d2.metas d1.metas
-}
-let add d e i = { d with evars=EvarMap.add d.evars e i }
-let remove d e = { d with evars=EvarMap.remove d.evars e }
-let find d e = EvarMap.find d.evars e
-let find_undefined d e = EvarMap.find_undefined d.evars e
-let mem d e = EvarMap.mem d.evars e
+ let is_new k v =
+ assert (v.evar_body == Evar_empty);
+ EvMap.mem k d2.defn_evars
+ in
+ not (d1 == d2) && EvMap.exists is_new d1.undf_evars
+
+let add_name_newly_undefined naming evk evi (evtoid,idtoev) =
+ let id = match naming with
+ | Misctypes.IntroAnonymous ->
+ let id = evar_ident_info evi in
+ Namegen.next_ident_away_from id (fun id -> Idmap.mem id idtoev)
+ | Misctypes.IntroIdentifier id ->
+ let id' =
+ Namegen.next_ident_away_from id (fun id -> Idmap.mem id idtoev) in
+ if not (Names.Id.equal id id') then
+ user_err_loc
+ (Loc.ghost,"",str "Already an existential evar of name " ++ pr_id id);
+ id'
+ | Misctypes.IntroFresh id ->
+ Namegen.next_ident_away_from id (fun id -> Idmap.mem id idtoev) in
+ (EvMap.add evk id evtoid, Idmap.add id evk idtoev)
+
+let add_name_undefined naming evk evi (evtoid,idtoev as evar_names) =
+ if EvMap.mem evk evtoid then
+ evar_names
+ else
+ add_name_newly_undefined naming evk evi evar_names
+
+let remove_name_defined evk (evtoid,idtoev) =
+ let id = EvMap.find evk evtoid in
+ (EvMap.remove evk evtoid, Idmap.remove id idtoev)
+
+let remove_name_possibly_already_defined evk evar_names =
+ try remove_name_defined evk evar_names
+ with Not_found -> evar_names
+
+let rename evk id evd =
+ let (evtoid,idtoev) = evd.evar_names in
+ let id' = EvMap.find evk evtoid in
+ if Idmap.mem id idtoev then anomaly (str "Evar name already in use");
+ { evd with evar_names =
+ (EvMap.add evk id evtoid (* overwrite old name *),
+ Idmap.add id evk (Idmap.remove id' idtoev)) }
+
+let reassign_name_defined evk evk' (evtoid,idtoev) =
+ let id = EvMap.find evk evtoid in
+ (EvMap.add evk' id (EvMap.remove evk evtoid),
+ Idmap.add id evk' (Idmap.remove id idtoev))
+
+let add d e i = match i.evar_body with
+| Evar_empty ->
+ let evar_names = add_name_undefined Misctypes.IntroAnonymous e i d.evar_names in
+ { d with undf_evars = EvMap.add e i d.undf_evars; evar_names }
+| Evar_defined _ ->
+ let evar_names = remove_name_possibly_already_defined e d.evar_names in
+ { d with defn_evars = EvMap.add e i d.defn_evars; evar_names }
+
+let remove d e =
+ let undf_evars = EvMap.remove e d.undf_evars in
+ let defn_evars = EvMap.remove e d.defn_evars in
+ { d with undf_evars; defn_evars; }
+
+let find d e =
+ try EvMap.find e d.undf_evars
+ with Not_found -> EvMap.find e d.defn_evars
+
+let find_undefined d e = EvMap.find e d.undf_evars
+
+let mem d e = EvMap.mem e d.undf_evars || EvMap.mem e d.defn_evars
+
(* spiwack: this function loses information from the original evar_map
it might be an idea not to export it. *)
-let to_list d = EvarMap.to_list d.evars
-let undefined_list d = EvarMap.undefined_list d.evars
-let undefined_evars d = { d with evars=EvarMap.undefined_evars d.evars }
-let defined_evars d = { d with evars=EvarMap.defined_evars d.evars }
+let to_list d =
+ (* Workaround for change in Map.fold behavior in ocaml 3.08.4 *)
+ let l = ref [] in
+ EvMap.iter (fun evk x -> l := (evk,x)::!l) d.defn_evars;
+ EvMap.iter (fun evk x -> l := (evk,x)::!l) d.undf_evars;
+ !l
+
+let undefined_map d = d.undf_evars
+
+let drop_all_defined d = { d with defn_evars = EvMap.empty }
+
(* spiwack: not clear what folding over an evar_map, for now we shall
simply fold over the inner evar_map. *)
-let fold f d a = EvarMap.fold d.evars f a
-let fold_undefined f d a = EvarMap.fold_undefined d.evars f a
-let is_evar d e = EvarMap.is_evar d.evars e
-let is_defined d e = EvarMap.is_defined d.evars e
-let is_undefined d e = EvarMap.is_undefined d.evars e
+let fold f d a =
+ EvMap.fold f d.defn_evars (EvMap.fold f d.undf_evars a)
+
+let fold_undefined f d a = EvMap.fold f d.undf_evars a
+
+let raw_map f d =
+ let f evk info =
+ let ans = f evk info in
+ let () = match info.evar_body, ans.evar_body with
+ | Evar_defined _, Evar_empty
+ | Evar_empty, Evar_defined _ ->
+ anomaly (str "Unrespectful mapping function.")
+ | _ -> ()
+ in
+ ans
+ in
+ let defn_evars = EvMap.smartmapi f d.defn_evars in
+ let undf_evars = EvMap.smartmapi f d.undf_evars in
+ { d with defn_evars; undf_evars; }
+
+let raw_map_undefined f d =
+ let f evk info =
+ let ans = f evk info in
+ let () = match ans.evar_body with
+ | Evar_defined _ ->
+ anomaly (str "Unrespectful mapping function.")
+ | _ -> ()
+ in
+ ans
+ in
+ { d with undf_evars = EvMap.smartmapi f d.undf_evars; }
+
+let is_evar = mem
+
+let is_defined d e = EvMap.mem e d.defn_evars
+
+let is_undefined d e = EvMap.mem e d.undf_evars
+
+let existential_value d (n, args) =
+ let info = find d n in
+ match evar_body info with
+ | Evar_defined c ->
+ instantiate_evar_array info c args
+ | Evar_empty ->
+ raise NotInstantiatedEvar
+
+let existential_opt_value d ev =
+ try Some (existential_value d ev)
+ with NotInstantiatedEvar -> None
-let existential_value d e = EvarMap.existential_value d.evars e
-let existential_type d e = EvarMap.existential_type d.evars e
-let existential_opt_value d e = EvarMap.existential_opt_value d.evars e
+let existential_type d (n, args) =
+ let info =
+ try find d n
+ with Not_found ->
+ anomaly (str "Evar " ++ str (string_of_existential n) ++ str " was not declared") in
+ instantiate_evar_array info info.evar_concl args
-let add_constraints d e = {d with evars= EvarMap.add_constraints d.evars e}
+let add_constraints d c =
+ { d with universes = add_constraints_context d.universes c }
+
+let add_universe_constraints d c =
+ { d with universes = add_universe_constraints_context d.universes c }
(*** /Lifting... ***)
(* evar_map are considered empty disregarding histories *)
let is_empty d =
- EvarMap.is_empty d.evars &&
- d.conv_pbs = [] &&
+ EvMap.is_empty d.defn_evars &&
+ EvMap.is_empty d.undf_evars &&
+ List.is_empty d.conv_pbs &&
Metamap.is_empty d.metas
-let subst_named_context_val s = map_named_val (subst_mps s)
-
-let subst_evar_info s evi =
- let subst_evb = function Evar_empty -> Evar_empty
- | Evar_defined c -> Evar_defined (subst_mps s c) in
- { evi with
- evar_concl = subst_mps s evi.evar_concl;
- evar_hyps = subst_named_context_val s evi.evar_hyps;
- evar_body = subst_evb evi.evar_body }
-
-let subst_evar_defs_light sub evd =
- assert (Univ.is_initial_universes (snd (snd evd.evars)));
- assert (evd.conv_pbs = []);
+let cmap f evd =
{ evd with
- metas = Metamap.map (map_clb (subst_mps sub)) evd.metas;
- evars = EvarInfoMap.map (subst_evar_info sub) (fst evd.evars), (snd evd.evars)
+ metas = Metamap.map (map_clb f) evd.metas;
+ defn_evars = EvMap.map (map_evar_info f) evd.defn_evars;
+ undf_evars = EvMap.map (map_evar_info f) evd.defn_evars
}
-let subst_evar_map = subst_evar_defs_light
-
(* spiwack: deprecated *)
let create_evar_defs sigma = { sigma with
- conv_pbs=[]; last_mods=ExistentialSet.empty; metas=Metamap.empty }
+ conv_pbs=[]; last_mods=Evar.Set.empty; metas=Metamap.empty }
(* spiwack: tentatively deprecated *)
let create_goal_evar_defs sigma = { sigma with
- (* conv_pbs=[]; last_mods=ExistentialSet.empty; metas=Metamap.empty } *)
+ (* conv_pbs=[]; last_mods=Evar.Set.empty; metas=Metamap.empty } *)
metas=Metamap.empty }
-let empty = {
- evars=EvarMap.empty;
- conv_pbs=[];
- last_mods = ExistentialSet.empty;
- metas=Metamap.empty
+
+let empty = {
+ defn_evars = EvMap.empty;
+ undf_evars = EvMap.empty;
+ universes = empty_evar_universe_context;
+ conv_pbs = [];
+ last_mods = Evar.Set.empty;
+ metas = Metamap.empty;
+ effects = Declareops.no_seff;
+ evar_names = (EvMap.empty,Idmap.empty); (* id<->key for undefined evars *)
+ future_goals = [];
+ principal_future_goal = None;
}
-let has_undefined evd =
- EvarMap.has_undefined evd.evars
+let from_env ?ctx e =
+ match ctx with
+ | None -> { empty with universes = evar_universe_context_from e }
+ | Some ctx -> { empty with universes = ctx }
+
+let has_undefined evd = not (EvMap.is_empty evd.undf_evars)
+
+let evars_reset_evd ?(with_conv_pbs=false) ?(with_univs=true) evd d =
+ let conv_pbs = if with_conv_pbs then evd.conv_pbs else d.conv_pbs in
+ let last_mods = if with_conv_pbs then evd.last_mods else d.last_mods in
+ let universes =
+ if not with_univs then evd.universes
+ else union_evar_universe_context evd.universes d.universes
+ in
+ { evd with
+ metas = d.metas;
+ last_mods; conv_pbs; universes }
+
+let merge_universe_context evd uctx' =
+ { evd with universes = union_evar_universe_context evd.universes uctx' }
+
+let set_universe_context evd uctx' =
+ { evd with universes = uctx' }
-let evars_reset_evd ?(with_conv_pbs=false) evd d =
- {d with evars = evd.evars;
- conv_pbs = if with_conv_pbs then evd.conv_pbs else d.conv_pbs }
let add_conv_pb pb d = {d with conv_pbs = pb::d.conv_pbs}
-let evar_source evk d = (EvarMap.find d.evars evk).evar_source
+
+let evar_source evk d = (find d evk).evar_source
+
+let evar_ident evk evd =
+ try EvMap.find evk (fst evd.evar_names)
+ with Not_found ->
+ (* Unnamed (non-dependent) evar *)
+ add_suffix (Id.of_string "X") (string_of_int (Evar.repr evk))
+
+let evar_key id evd =
+ Idmap.find id (snd evd.evar_names)
+
+let define_aux def undef evk body =
+ let oldinfo =
+ try EvMap.find evk undef
+ with Not_found ->
+ if EvMap.mem evk def then
+ anomaly ~label:"Evd.define" (Pp.str "cannot define an evar twice")
+ else
+ anomaly ~label:"Evd.define" (Pp.str "cannot define undeclared evar")
+ in
+ let () = assert (oldinfo.evar_body == Evar_empty) in
+ let newinfo = { oldinfo with evar_body = Evar_defined body } in
+ EvMap.add evk newinfo def, EvMap.remove evk undef
(* define the existential of section path sp as the constr body *)
let define evk body evd =
- { evd with
- evars = EvarMap.define evd.evars evk body;
- last_mods =
- match evd.conv_pbs with
- | [] -> evd.last_mods
- | _ -> ExistentialSet.add evk evd.last_mods }
-
-let evar_declare hyps evk ty ?(src=(dummy_loc,InternalHole)) ?filter ?candidates evd =
- let filter =
- if filter = None then
- List.map (fun _ -> true) (named_context_of_val hyps)
- else
- (let filter = Option.get filter in
- assert (List.length filter = List.length (named_context_of_val hyps));
- filter)
+ let (defn_evars, undf_evars) = define_aux evd.defn_evars evd.undf_evars evk body in
+ let last_mods = match evd.conv_pbs with
+ | [] -> evd.last_mods
+ | _ -> Evar.Set.add evk evd.last_mods
in
- { evd with
- evars = EvarMap.add_undefined evd.evars evk
- {evar_hyps = hyps;
- evar_concl = ty;
- evar_body = Evar_empty;
- evar_filter = filter;
- evar_source = src;
- evar_candidates = candidates;
- evar_extra = Store.empty } }
-
-let is_defined_evar evd (evk,_) = EvarMap.is_defined evd.evars evk
-
-(* Does k corresponds to an (un)defined existential ? *)
-let is_undefined_evar evd c = match kind_of_term c with
- | Evar ev -> not (is_defined_evar evd ev)
- | _ -> false
+ let evar_names = remove_name_defined evk evd.evar_names in
+ { evd with defn_evars; undf_evars; last_mods; evar_names }
+
+let evar_declare hyps evk ty ?(src=(Loc.ghost,Evar_kinds.InternalHole))
+ ?(filter=Filter.identity) ?candidates ?(store=Store.empty)
+ ?(naming=Misctypes.IntroAnonymous) evd =
+ let () = match Filter.repr filter with
+ | None -> ()
+ | Some filter ->
+ assert (Int.equal (List.length filter) (List.length (named_context_of_val hyps)))
+ in
+ let evar_info = {
+ evar_hyps = hyps;
+ evar_concl = ty;
+ evar_body = Evar_empty;
+ evar_filter = filter;
+ evar_source = src;
+ evar_candidates = candidates;
+ evar_extra = store; }
+ in
+ let evar_names = add_name_newly_undefined naming evk evar_info evd.evar_names in
+ { evd with undf_evars = EvMap.add evk evar_info evd.undf_evars; evar_names }
+
+let restrict evk evk' filter ?candidates evd =
+ let evar_info = EvMap.find evk evd.undf_evars in
+ let evar_info' =
+ { evar_info with evar_filter = filter;
+ evar_candidates = candidates;
+ evar_extra = Store.empty } in
+ let evar_names = reassign_name_defined evk evk' evd.evar_names in
+ let ctxt = Filter.filter_list filter (evar_context evar_info) in
+ let id_inst = Array.map_of_list (fun (id,_,_) -> mkVar id) ctxt in
+ let body = mkEvar(evk',id_inst) in
+ let (defn_evars, undf_evars) = define_aux evd.defn_evars evd.undf_evars evk body in
+ { evd with undf_evars = EvMap.add evk' evar_info' undf_evars;
+ defn_evars; evar_names }
+
+let downcast evk ccl evd =
+ let evar_info = EvMap.find evk evd.undf_evars in
+ let evar_info' = { evar_info with evar_concl = ccl } in
+ { evd with undf_evars = EvMap.add evk evar_info' evd.undf_evars }
(* extracts conversion problems that satisfy predicate p *)
(* Note: conv_pbs not satisying p are stored back in reverse order *)
@@ -473,129 +867,495 @@ let extract_conv_pbs evd p =
([],[])
evd.conv_pbs
in
- {evd with conv_pbs = pbs1; last_mods = ExistentialSet.empty},
+ {evd with conv_pbs = pbs1; last_mods = Evar.Set.empty},
pbs
let extract_changed_conv_pbs evd p =
- extract_conv_pbs evd (p evd.last_mods)
+ extract_conv_pbs evd (fun pb -> p evd.last_mods pb)
let extract_all_conv_pbs evd =
extract_conv_pbs evd (fun _ -> true)
-(* spiwack: should it be replaced by Evd.merge? *)
-let evar_merge evd evars =
- { evd with evars = EvarMap.merge evd.evars evars.evars }
+let loc_of_conv_pb evd (pbty,env,t1,t2) =
+ match kind_of_term (fst (decompose_app t1)) with
+ | Evar (evk1,_) -> fst (evar_source evk1 evd)
+ | _ ->
+ match kind_of_term (fst (decompose_app t2)) with
+ | Evar (evk2,_) -> fst (evar_source evk2 evd)
+ | _ -> Loc.ghost
+
+(** The following functions return the set of evars immediately
+ contained in the object *)
-let evar_list evd c =
+(* excluding defined evars *)
+
+let evar_list c =
let rec evrec acc c =
match kind_of_term c with
- | Evar (evk, _ as ev) when mem evd evk -> ev :: acc
+ | Evar (evk, _ as ev) -> ev :: acc
| _ -> fold_constr evrec acc c in
evrec [] c
-let collect_evars c =
- let rec collrec acc c =
+let evars_of_term c =
+ let rec evrec acc c =
match kind_of_term c with
- | Evar (evk,_) -> ExistentialSet.add evk acc
- | _ -> fold_constr collrec acc c
+ | Evar (n, l) -> Evar.Set.add n (Array.fold_left evrec acc l)
+ | _ -> fold_constr evrec acc c
in
- collrec ExistentialSet.empty c
+ evrec Evar.Set.empty c
+
+let evars_of_named_context nc =
+ List.fold_right (fun (_, b, t) s ->
+ Option.fold_left (fun s t ->
+ Evar.Set.union s (evars_of_term t))
+ (Evar.Set.union s (evars_of_term t)) b)
+ nc Evar.Set.empty
+
+let evars_of_filtered_evar_info evi =
+ Evar.Set.union (evars_of_term evi.evar_concl)
+ (Evar.Set.union
+ (match evi.evar_body with
+ | Evar_empty -> Evar.Set.empty
+ | Evar_defined b -> evars_of_term b)
+ (evars_of_named_context (evar_filtered_context evi)))
+
+(**********************************************************)
+(* Side effects *)
+
+let emit_side_effects eff evd =
+ { evd with effects = Declareops.union_side_effects eff evd.effects; }
+
+let drop_side_effects evd =
+ { evd with effects = Declareops.no_seff; }
+
+let eval_side_effects evd = evd.effects
+
+(* Future goals *)
+let declare_future_goal evk evd =
+ { evd with future_goals = evk::evd.future_goals }
+
+let declare_principal_goal evk evd =
+ match evd.principal_future_goal with
+ | None -> { evd with
+ future_goals = evk::evd.future_goals;
+ principal_future_goal=Some evk; }
+ | Some _ -> Errors.error "Only one main subgoal per instantiation."
+
+let future_goals evd = evd.future_goals
+
+let principal_future_goal evd = evd.principal_future_goal
+
+let reset_future_goals evd =
+ { evd with future_goals = [] ; principal_future_goal=None }
+
+let restore_future_goals evd gls pgl =
+ { evd with future_goals = gls ; principal_future_goal = pgl }
(**********************************************************)
(* Sort variables *)
-let new_univ_variable ({ evars = (sigma,(us,sm)) } as d) =
- let u = Termops.new_univ_level () in
- let us' = Univ.UniverseLSet.add u us in
- ({d with evars = (sigma, (us', sm))}, Univ.make_universe u)
-
-let new_sort_variable d =
- let (d', u) = new_univ_variable d in
+type rigid =
+ | UnivRigid
+ | UnivFlexible of bool (** Is substitution by an algebraic ok? *)
+
+let univ_rigid = UnivRigid
+let univ_flexible = UnivFlexible false
+let univ_flexible_alg = UnivFlexible true
+
+let evar_universe_context d = d.universes
+
+let universe_context_set d = d.universes.uctx_local
+
+let universe_context evd =
+ Univ.ContextSet.to_context evd.universes.uctx_local
+
+let universe_subst evd =
+ evd.universes.uctx_univ_variables
+
+let merge_uctx rigid uctx ctx' =
+ let open Univ in
+ let uctx =
+ match rigid with
+ | UnivRigid -> uctx
+ | UnivFlexible b ->
+ let levels = ContextSet.levels ctx' in
+ let fold u accu =
+ if LMap.mem u accu then accu
+ else LMap.add u None accu
+ in
+ let uvars' = LSet.fold fold levels uctx.uctx_univ_variables in
+ if b then
+ { uctx with uctx_univ_variables = uvars';
+ uctx_univ_algebraic = LSet.union uctx.uctx_univ_algebraic levels }
+ else { uctx with uctx_univ_variables = uvars' }
+ in
+ let uctx_local = ContextSet.append ctx' uctx.uctx_local in
+ let uctx_universes = merge_constraints (ContextSet.constraints ctx') uctx.uctx_universes in
+ { uctx with uctx_local; uctx_universes }
+
+let merge_context_set rigid evd ctx' =
+ {evd with universes = merge_uctx rigid evd.universes ctx'}
+
+let merge_uctx_subst uctx s =
+ { uctx with uctx_univ_variables = Univ.LMap.subst_union uctx.uctx_univ_variables s }
+
+let merge_universe_subst evd subst =
+ {evd with universes = merge_uctx_subst evd.universes subst }
+
+let with_context_set rigid d (a, ctx) =
+ (merge_context_set rigid d ctx, a)
+
+let add_uctx_names s l (names, names_rev) =
+ (UNameMap.add s l names, Univ.LMap.add l s names_rev)
+
+let uctx_new_univ_variable rigid name
+ ({ uctx_local = ctx; uctx_univ_variables = uvars; uctx_univ_algebraic = avars} as uctx) =
+ let u = Universes.new_univ_level (Global.current_dirpath ()) in
+ let ctx' = Univ.ContextSet.add_universe u ctx in
+ let uctx' =
+ match rigid with
+ | UnivRigid -> uctx
+ | UnivFlexible b ->
+ let uvars' = Univ.LMap.add u None uvars in
+ if b then {uctx with uctx_univ_variables = uvars';
+ uctx_univ_algebraic = Univ.LSet.add u avars}
+ else {uctx with uctx_univ_variables = Univ.LMap.add u None uvars} in
+ let names =
+ match name with
+ | Some n -> add_uctx_names n u uctx.uctx_names
+ | None -> uctx.uctx_names
+ in
+ {uctx' with uctx_names = names; uctx_local = ctx';
+ uctx_universes = Univ.add_universe u uctx.uctx_universes}, u
+
+let new_univ_level_variable ?name rigid evd =
+ let uctx', u = uctx_new_univ_variable rigid name evd.universes in
+ ({evd with universes = uctx'}, u)
+
+let new_univ_variable ?name rigid evd =
+ let uctx', u = uctx_new_univ_variable rigid name evd.universes in
+ ({evd with universes = uctx'}, Univ.Universe.make u)
+
+let new_sort_variable ?name rigid d =
+ let (d', u) = new_univ_variable rigid ?name d in
(d', Type u)
-let is_sort_variable {evars=(_,(us,_))} s = match s with Type u -> true | _ -> false
-let whd_sort_variable {evars=(_,sm)} t = t
+let make_flexible_variable evd b u =
+ let {uctx_univ_variables = uvars; uctx_univ_algebraic = avars} as ctx = evd.universes in
+ let uvars' = Univ.LMap.add u None uvars in
+ let avars' =
+ if b then
+ let uu = Univ.Universe.make u in
+ let substu_not_alg u' v =
+ Option.cata (fun vu -> Univ.Universe.equal uu vu && not (Univ.LSet.mem u' avars)) false v
+ in
+ if not (Univ.LMap.exists substu_not_alg uvars)
+ then Univ.LSet.add u avars else avars
+ else avars
+ in
+ {evd with universes = {ctx with uctx_univ_variables = uvars';
+ uctx_univ_algebraic = avars'}}
+
+(****************************************)
+(* Operations on constants *)
+(****************************************)
+
+let fresh_sort_in_family ?(rigid=univ_flexible) env evd s =
+ with_context_set rigid evd (Universes.fresh_sort_in_family env s)
+
+let fresh_constant_instance env evd c =
+ with_context_set univ_flexible evd (Universes.fresh_constant_instance env c)
-let univ_of_sort = function
- | Type u -> u
- | Prop Pos -> Univ.type0_univ
- | Prop Null -> Univ.type0m_univ
+let fresh_inductive_instance env evd i =
+ with_context_set univ_flexible evd (Universes.fresh_inductive_instance env i)
+
+let fresh_constructor_instance env evd c =
+ with_context_set univ_flexible evd (Universes.fresh_constructor_instance env c)
+
+let fresh_global ?(rigid=univ_flexible) ?names env evd gr =
+ with_context_set rigid evd (Universes.fresh_global_instance ?names env gr)
+
+let whd_sort_variable evd t = t
+
+let is_sort_variable evd s =
+ match s with
+ | Type u ->
+ (match Univ.universe_level u with
+ | Some l as x ->
+ let uctx = evd.universes in
+ if Univ.LSet.mem l (Univ.ContextSet.levels uctx.uctx_local) then x
+ else None
+ | None -> None)
+ | _ -> None
+
+let is_flexible_level evd l =
+ let uctx = evd.universes in
+ Univ.LMap.mem l uctx.uctx_univ_variables
let is_eq_sort s1 s2 =
- if s1 = s2 then None
- else
- let u1 = univ_of_sort s1 and u2 = univ_of_sort s2 in
- if u1 = u2 then None
+ if Sorts.equal s1 s2 then None
+ else
+ let u1 = univ_of_sort s1
+ and u2 = univ_of_sort s2 in
+ if Univ.Universe.equal u1 u2 then None
else Some (u1, u2)
-let is_univ_var_or_set u =
- Univ.is_univ_variable u || u = Univ.type0_univ
-
-let set_leq_sort ({evars = (sigma, (us, sm))} as d) s1 s2 =
+let normalize_universe evd =
+ let vars = ref evd.universes.uctx_univ_variables in
+ let normalize = Universes.normalize_universe_opt_subst vars in
+ normalize
+
+let normalize_universe_instance evd l =
+ let vars = ref evd.universes.uctx_univ_variables in
+ let normalize = Univ.level_subst_of (Universes.normalize_univ_variable_opt_subst vars) in
+ Univ.Instance.subst_fn normalize l
+
+let normalize_sort evars s =
+ match s with
+ | Prop _ -> s
+ | Type u ->
+ let u' = normalize_universe evars u in
+ if u' == u then s else Type u'
+
+(* FIXME inefficient *)
+let set_eq_sort env d s1 s2 =
+ let s1 = normalize_sort d s1 and s2 = normalize_sort d s2 in
match is_eq_sort s1 s2 with
| None -> d
| Some (u1, u2) ->
- match s1, s2 with
- | Prop c, Prop c' ->
- if c = Null && c' = Pos then d
- else (raise (Univ.UniverseInconsistency (Univ.Le, u1, u2)))
- | Type u, Prop c ->
- if c = Pos then
- add_constraints d (Univ.enforce_geq Univ.type0_univ u Univ.empty_constraint)
- else raise (Univ.UniverseInconsistency (Univ.Le, u1, u2))
- | _, Type u ->
- if is_univ_var_or_set u then
- add_constraints d (Univ.enforce_geq u2 u1 Univ.empty_constraint)
- else raise (Univ.UniverseInconsistency (Univ.Le, u1, u2))
-
-let is_univ_level_var us u =
- match Univ.universe_level u with
- | Some u -> Univ.UniverseLSet.mem u us
- | None -> false
-
-let set_eq_sort ({evars = (sigma, (us, sm))} as d) s1 s2 =
+ if not (type_in_type env) then
+ add_universe_constraints d
+ (Universes.Constraints.singleton (u1,Universes.UEq,u2))
+ else
+ d
+
+let has_lub evd u1 u2 =
+ (* let normalize = Universes.normalize_universe_opt_subst (ref univs.uctx_univ_variables) in *)
+ (* (\* let dref, norm = memo_normalize_universe d in *\) *)
+ (* let u1 = normalize u1 and u2 = normalize u2 in *)
+ if Univ.Universe.equal u1 u2 then evd
+ else add_universe_constraints evd
+ (Universes.Constraints.singleton (u1,Universes.ULub,u2))
+
+let set_eq_level d u1 u2 =
+ add_constraints d (Univ.enforce_eq_level u1 u2 Univ.Constraint.empty)
+
+let set_leq_level d u1 u2 =
+ add_constraints d (Univ.enforce_leq_level u1 u2 Univ.Constraint.empty)
+
+let set_eq_instances ?(flex=false) d u1 u2 =
+ add_universe_constraints d
+ (Universes.enforce_eq_instances_univs flex u1 u2 Universes.Constraints.empty)
+
+let set_leq_sort env evd s1 s2 =
+ let s1 = normalize_sort evd s1
+ and s2 = normalize_sort evd s2 in
match is_eq_sort s1 s2 with
- | None -> d
+ | None -> evd
| Some (u1, u2) ->
- match s1, s2 with
- | Prop c, Type u when is_univ_level_var us u ->
- add_constraints d (Univ.enforce_eq u1 u2 Univ.empty_constraint)
- | Type u, Prop c when is_univ_level_var us u ->
- add_constraints d (Univ.enforce_eq u1 u2 Univ.empty_constraint)
- | Type u, Type v when (is_univ_level_var us u) || (is_univ_level_var us v) ->
- add_constraints d (Univ.enforce_eq u1 u2 Univ.empty_constraint)
- | Prop c, Type u when is_univ_var_or_set u &&
- Univ.check_eq sm u1 u2 -> d
- | Type u, Prop c when is_univ_var_or_set u && Univ.check_eq sm u1 u2 -> d
- | Type u, Type v when is_univ_var_or_set u && is_univ_var_or_set v ->
- add_constraints d (Univ.enforce_eq u1 u2 Univ.empty_constraint)
- | _, _ -> raise (Univ.UniverseInconsistency (Univ.Eq, u1, u2))
+ (* if Univ.is_type0_univ u2 then *)
+ (* if Univ.is_small_univ u1 then evd *)
+ (* else raise (Univ.UniverseInconsistency (Univ.Le, u1, u2, [])) *)
+ (* else if Univ.is_type0m_univ u2 then *)
+ (* raise (Univ.UniverseInconsistency (Univ.Le, u1, u2, [])) *)
+ (* else *)
+ if not (type_in_type env) then
+ add_universe_constraints evd (Universes.Constraints.singleton (u1,Universes.ULe,u2))
+ else evd
+let check_eq evd s s' =
+ Univ.check_eq evd.universes.uctx_universes s s'
+
+let check_leq evd s s' =
+ Univ.check_leq evd.universes.uctx_universes s s'
+
+let subst_univs_context_with_def def usubst (ctx, cst) =
+ (Univ.LSet.diff ctx def, Univ.subst_univs_constraints usubst cst)
+
+let normalize_evar_universe_context_variables uctx =
+ let normalized_variables, undef, def, subst =
+ Universes.normalize_univ_variables uctx.uctx_univ_variables
+ in
+ let ctx_local = subst_univs_context_with_def def (Univ.make_subst subst) uctx.uctx_local in
+ let ctx_local', univs = Universes.refresh_constraints uctx.uctx_initial_universes ctx_local in
+ subst, { uctx with uctx_local = ctx_local';
+ uctx_univ_variables = normalized_variables;
+ uctx_universes = univs }
+
+(* let normvarsconstrkey = Profile.declare_profile "normalize_evar_universe_context_variables";; *)
+(* let normalize_evar_universe_context_variables = *)
+(* Profile.profile1 normvarsconstrkey normalize_evar_universe_context_variables;; *)
+
+let abstract_undefined_variables uctx =
+ let vars' =
+ Univ.LMap.fold (fun u v acc ->
+ if v == None then Univ.LSet.remove u acc
+ else acc)
+ uctx.uctx_univ_variables uctx.uctx_univ_algebraic
+ in { uctx with uctx_local = Univ.ContextSet.empty;
+ uctx_univ_algebraic = vars' }
+
+
+let refresh_undefined_univ_variables uctx =
+ let subst, ctx' = Universes.fresh_universe_context_set_instance uctx.uctx_local in
+ let alg = Univ.LSet.fold (fun u acc -> Univ.LSet.add (Univ.subst_univs_level_level subst u) acc)
+ uctx.uctx_univ_algebraic Univ.LSet.empty
+ in
+ let vars =
+ Univ.LMap.fold
+ (fun u v acc ->
+ Univ.LMap.add (Univ.subst_univs_level_level subst u)
+ (Option.map (Univ.subst_univs_level_universe subst) v) acc)
+ uctx.uctx_univ_variables Univ.LMap.empty
+ in
+ let uctx' = {uctx_names = uctx.uctx_names;
+ uctx_local = ctx';
+ uctx_univ_variables = vars; uctx_univ_algebraic = alg;
+ uctx_universes = Univ.initial_universes;
+ uctx_initial_universes = uctx.uctx_initial_universes } in
+ uctx', subst
+
+let refresh_undefined_universes evd =
+ let uctx', subst = refresh_undefined_univ_variables evd.universes in
+ let evd' = cmap (subst_univs_level_constr subst) {evd with universes = uctx'} in
+ evd', subst
+
+let normalize_evar_universe_context uctx =
+ let rec fixpoint uctx =
+ let ((vars',algs'), us') =
+ Universes.normalize_context_set uctx.uctx_local uctx.uctx_univ_variables
+ uctx.uctx_univ_algebraic
+ in
+ if Univ.LSet.equal (fst us') (fst uctx.uctx_local) then
+ uctx
+ else
+ let us', universes = Universes.refresh_constraints uctx.uctx_initial_universes us' in
+ let uctx' =
+ { uctx_names = uctx.uctx_names;
+ uctx_local = us';
+ uctx_univ_variables = vars';
+ uctx_univ_algebraic = algs';
+ uctx_universes = universes;
+ uctx_initial_universes = uctx.uctx_initial_universes }
+ in fixpoint uctx'
+ in fixpoint uctx
+
+let nf_univ_variables evd =
+ let subst, uctx' = normalize_evar_universe_context_variables evd.universes in
+ let evd' = {evd with universes = uctx'} in
+ evd', subst
+
+let nf_constraints evd =
+ let subst, uctx' = normalize_evar_universe_context_variables evd.universes in
+ let uctx' = normalize_evar_universe_context uctx' in
+ {evd with universes = uctx'}
+
+let nf_constraints =
+ if Flags.profile then
+ let nfconstrkey = Profile.declare_profile "nf_constraints" in
+ Profile.profile1 nfconstrkey nf_constraints
+ else nf_constraints
+
+let universe_of_name evd s =
+ UNameMap.find s (fst evd.universes.uctx_names)
+
+let add_universe_name evd s l =
+ let names' = add_uctx_names s l evd.universes.uctx_names in
+ {evd with universes = {evd.universes with uctx_names = names'}}
+
+let universes evd = evd.universes.uctx_universes
+
+(* Conversion w.r.t. an evar map and its local universes. *)
+
+let conversion_gen env evd pb t u =
+ match pb with
+ | Reduction.CONV ->
+ Reduction.trans_conv_universes
+ full_transparent_state ~evars:(existential_opt_value evd) env
+ evd.universes.uctx_universes t u
+ | Reduction.CUMUL -> Reduction.trans_conv_leq_universes
+ full_transparent_state ~evars:(existential_opt_value evd) env
+ evd.universes.uctx_universes t u
+
+(* let conversion_gen_key = Profile.declare_profile "conversion_gen" *)
+(* let conversion_gen = Profile.profile5 conversion_gen_key conversion_gen *)
+
+let conversion env d pb t u =
+ conversion_gen env d pb t u; d
+
+let test_conversion env d pb t u =
+ try conversion_gen env d pb t u; true
+ with _ -> false
+
+let eq_constr_univs evd t u =
+ let b, c = Universes.eq_constr_univs_infer evd.universes.uctx_universes t u in
+ if b then
+ try let evd' = add_universe_constraints evd c in evd', b
+ with Univ.UniverseInconsistency _ | UniversesDiffer -> evd, false
+ else evd, b
+
+let e_eq_constr_univs evdref t u =
+ let evd, b = eq_constr_univs !evdref t u in
+ evdref := evd; b
+
+let eq_constr_univs_test evd t u =
+ snd (eq_constr_univs evd t u)
+
+let eq_named_context_val d ctx1 ctx2 =
+ ctx1 == ctx2 ||
+ let c1 = named_context_of_val ctx1 and c2 = named_context_of_val ctx2 in
+ let eq_named_declaration (i1, c1, t1) (i2, c2, t2) =
+ Id.equal i1 i2 && Option.equal (eq_constr_univs_test d) c1 c2
+ && (eq_constr_univs_test d) t1 t2
+ in List.equal eq_named_declaration c1 c2
+
+let eq_evar_body d b1 b2 = match b1, b2 with
+| Evar_empty, Evar_empty -> true
+| Evar_defined t1, Evar_defined t2 -> eq_constr_univs_test d t1 t2
+| _ -> false
+
+let eq_evar_info d ei1 ei2 =
+ ei1 == ei2 ||
+ eq_constr_univs_test d ei1.evar_concl ei2.evar_concl &&
+ eq_named_context_val d (ei1.evar_hyps) (ei2.evar_hyps) &&
+ eq_evar_body d ei1.evar_body ei2.evar_body
+ (** ppedrot: [eq_constr] may be a bit too permissive here *)
+
+
(**********************************************************)
(* Accessing metas *)
-let meta_list evd = metamap_to_list evd.metas
+(** We use this function to overcome OCaml compiler limitations and to prevent
+ the use of costly in-place modifications. *)
+let set_metas evd metas = {
+ defn_evars = evd.defn_evars;
+ undf_evars = evd.undf_evars;
+ universes = evd.universes;
+ conv_pbs = evd.conv_pbs;
+ last_mods = evd.last_mods;
+ metas;
+ effects = evd.effects;
+ evar_names = evd.evar_names;
+ future_goals = evd.future_goals;
+ principal_future_goal = evd.principal_future_goal;
+}
-let find_meta evd mv = Metamap.find mv evd.metas
+let meta_list evd = metamap_to_list evd.metas
let undefined_metas evd =
- List.sort Pervasives.compare (map_succeed (function
- | (n,Clval(_,_,typ)) -> failwith ""
- | (n,Cltyp (_,typ)) -> n)
- (meta_list evd))
-
-let metas_of evd =
- List.map (function
- | (n,Clval(_,_,typ)) -> (n,typ.rebus)
- | (n,Cltyp (_,typ)) -> (n,typ.rebus))
- (meta_list evd)
+ let filter = function
+ | (n,Clval(_,_,typ)) -> None
+ | (n,Cltyp (_,typ)) -> Some n
+ in
+ let m = List.map_filter filter (meta_list evd) in
+ List.sort Int.compare m
let map_metas_fvalue f evd =
- { evd with metas =
- Metamap.map
- (function
- | Clval(id,(c,s),typ) -> Clval(id,(mk_freelisted (f c.rebus),s),typ)
- | x -> x) evd.metas }
+ let map = function
+ | Clval(id,(c,s),typ) -> Clval(id,(mk_freelisted (f c.rebus),s),typ)
+ | x -> x
+ in
+ set_metas evd (Metamap.smartmap map evd.metas)
let meta_opt_fvalue evd mv =
match Metamap.find mv evd.metas with
@@ -614,7 +1374,7 @@ let try_meta_fvalue evd mv =
let meta_fvalue evd mv =
try try_meta_fvalue evd mv
- with Not_found -> anomaly "meta_fvalue: meta has no value"
+ with Not_found -> anomaly ~label:"meta_fvalue" (Pp.str "meta has no value")
let meta_value evd mv =
(fst (try_meta_fvalue evd mv)).rebus
@@ -627,39 +1387,55 @@ let meta_ftype evd mv =
let meta_type evd mv = (meta_ftype evd mv).rebus
let meta_declare mv v ?(name=Anonymous) evd =
- { evd with metas = Metamap.add mv (Cltyp(name,mk_freelisted v)) evd.metas }
+ let metas = Metamap.add mv (Cltyp(name,mk_freelisted v)) evd.metas in
+ set_metas evd metas
-let meta_assign mv (v,pb) evd =
- match Metamap.find mv evd.metas with
- | Cltyp(na,ty) ->
- { evd with
- metas = Metamap.add mv (Clval(na,(mk_freelisted v,pb),ty)) evd.metas }
- | _ -> anomaly "meta_assign: already defined"
+let meta_assign mv (v, pb) evd =
+ let modify _ = function
+ | Cltyp (na, ty) -> Clval (na, (mk_freelisted v, pb), ty)
+ | _ -> anomaly ~label:"meta_assign" (Pp.str "already defined")
+ in
+ let metas = Metamap.modify mv modify evd.metas in
+ set_metas evd metas
-let meta_reassign mv (v,pb) evd =
- match Metamap.find mv evd.metas with
- | Clval(na,_,ty) ->
- { evd with
- metas = Metamap.add mv (Clval(na,(mk_freelisted v,pb),ty)) evd.metas }
- | _ -> anomaly "meta_reassign: not yet defined"
+let meta_reassign mv (v, pb) evd =
+ let modify _ = function
+ | Clval(na, _, ty) -> Clval (na, (mk_freelisted v, pb), ty)
+ | _ -> anomaly ~label:"meta_reassign" (Pp.str "not yet defined")
+ in
+ let metas = Metamap.modify mv modify evd.metas in
+ set_metas evd metas
(* If the meta is defined then forget its name *)
let meta_name evd mv =
try fst (clb_name (Metamap.find mv evd.metas)) with Not_found -> Anonymous
+let explain_no_such_bound_variable evd id =
+ let mvl =
+ List.rev (Metamap.fold (fun n clb l ->
+ let na = fst (clb_name clb) in
+ if na != Anonymous then out_name na :: l else l)
+ evd.metas []) in
+ errorlabstrm "Evd.meta_with_name"
+ (str"No such bound variable " ++ pr_id id ++
+ (if mvl == [] then str " (no bound variables at all in the expression)."
+ else
+ (str" (possible name" ++
+ str (if List.length mvl == 1 then " is: " else "s are: ") ++
+ pr_enum pr_id mvl ++ str").")))
+
let meta_with_name evd id =
let na = Name id in
let (mvl,mvnodef) =
Metamap.fold
(fun n clb (l1,l2 as l) ->
let (na',def) = clb_name clb in
- if na = na' then if def then (n::l1,l2) else (n::l1,n::l2)
+ if Name.equal na na' then if def then (n::l1,l2) else (n::l1,n::l2)
else l)
evd.metas ([],[]) in
match mvnodef, mvl with
| _,[] ->
- errorlabstrm "Evd.meta_with_name"
- (str"No such bound variable " ++ pr_id id ++ str".")
+ explain_no_such_bound_variable evd id
| ([n],_|_,[n]) ->
n
| _ ->
@@ -667,36 +1443,55 @@ let meta_with_name evd id =
(str "Binder name \"" ++ pr_id id ++
strbrk "\" occurs more than once in clause.")
+let clear_metas evd = {evd with metas = Metamap.empty}
let meta_merge evd1 evd2 =
- {evd2 with
- metas = List.fold_left (fun m (n,v) -> Metamap.add n v m)
- evd2.metas (metamap_to_list evd1.metas) }
+ let metas = Metamap.fold Metamap.add evd1.metas evd2.metas in
+ let universes = union_evar_universe_context evd2.universes evd1.universes in
+ {evd2 with universes; metas; }
type metabinding = metavariable * constr * instance_status
let retract_coercible_metas evd =
- let mc,ml =
- Metamap.fold (fun n v (mc,ml) ->
- match v with
- | Clval (na,(b,(Conv,CoerceToType as s)),typ) ->
- (n,b.rebus,s)::mc, Metamap.add n (Cltyp (na,typ)) ml
- | v ->
- mc, Metamap.add n v ml)
- evd.metas ([],Metamap.empty) in
- mc, { evd with metas = ml }
-
-let rec list_assoc_in_triple x = function
- [] -> raise Not_found
- | (a,b,_)::l -> if compare a x = 0 then b else list_assoc_in_triple x l
-
-let subst_defined_metas bl c =
+ let mc = ref [] in
+ let map n v = match v with
+ | Clval (na, (b, (Conv, CoerceToType as s)), typ) ->
+ let () = mc := (n, b.rebus, s) :: !mc in
+ Cltyp (na, typ)
+ | v -> v
+ in
+ let metas = Metamap.smartmapi map evd.metas in
+ !mc, set_metas evd metas
+
+let subst_defined_metas_evars (bl,el) c =
let rec substrec c = match kind_of_term c with
- | Meta i -> substrec (list_assoc_snd_in_triple i bl)
+ | Meta i ->
+ let select (j,_,_) = Int.equal i j in
+ substrec (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))
+ with Not_found -> map_constr substrec c)
| _ -> map_constr substrec c
in try Some (substrec c) with Not_found -> None
+let evar_source_of_meta mv evd =
+ match meta_name evd mv with
+ | Anonymous -> (Loc.ghost,Evar_kinds.GoalEvar)
+ | Name id -> (Loc.ghost,Evar_kinds.VarInstance id)
+
+let dependent_evar_ident ev evd =
+ let evi = find evd ev in
+ match evi.evar_source with
+ | (_,Evar_kinds.VarInstance id) -> id
+ | _ -> anomaly (str "Not an evar resulting of a dependent binding")
+
(*******************************************************************)
+
+type pending = (* before: *) evar_map * (* after: *) evar_map
+
+type pending_constr = pending * constr
+
type open_constr = evar_map * constr
(*******************************************************************)
@@ -704,10 +1499,57 @@ type open_constr = evar_map * constr
type ['a] *)
type 'a sigma = {
it : 'a ;
- sigma : evar_map}
+ sigma : evar_map
+}
let sig_it x = x.it
let sig_sig x = x.sigma
+let on_sig s f =
+ let sigma', v = f s.sigma in
+ { s with sigma = sigma' }, v
+
+(*******************************************************************)
+(* The state monad with state an evar map. *)
+
+module MonadR =
+ Monad.Make (struct
+
+ type +'a t = evar_map -> evar_map * 'a
+
+ let return a = fun s -> (s,a)
+
+ let (>>=) x f = fun s ->
+ let (s',a) = x s in
+ f a s'
+
+ let (>>) x y = fun s ->
+ let (s',()) = x s in
+ y s'
+
+ let map f x = fun s ->
+ on_snd f (x s)
+
+ end)
+
+module Monad =
+ Monad.Make (struct
+
+ type +'a t = evar_map -> 'a * evar_map
+
+ let return a = fun s -> (a,s)
+
+ let (>>=) x f = fun s ->
+ let (a,s') = x s in
+ f a s'
+
+ let (>>) x y = fun s ->
+ let ((),s') = x s in
+ y s'
+
+ let map f x = fun s ->
+ on_fst f (x s)
+
+ end)
(**********************************************************)
(* Failure explanation *)
@@ -717,6 +1559,8 @@ type unsolvability_explanation = SeveralInstancesFound of int
(**********************************************************)
(* Pretty-printing *)
+let pr_existential_key sigma evk = str "?" ++ pr_id (evar_ident evk sigma)
+
let pr_instance_status (sc,typ) =
begin match sc with
| IsSubType -> str " [or a subtype of it]"
@@ -753,27 +1597,35 @@ let pr_decl ((id,b,_),ok) =
| Some c -> str (if ok then "(" else "{") ++ pr_id id ++ str ":=" ++
print_constr c ++ str (if ok then ")" else "}")
-let pr_evar_source = function
- | QuestionMark _ -> str "underscore"
- | CasesType -> str "pattern-matching return predicate"
- | BinderType (Name id) -> str "type of " ++ Nameops.pr_id id
- | BinderType Anonymous -> str "type of anonymous binder"
- | ImplicitArg (c,(n,ido),b) ->
+let rec pr_evar_source = function
+ | Evar_kinds.QuestionMark _ -> str "underscore"
+ | Evar_kinds.CasesType false -> str "pattern-matching return predicate"
+ | Evar_kinds.CasesType true ->
+ str "subterm of pattern-matching return predicate"
+ | Evar_kinds.BinderType (Name id) -> str "type of " ++ Nameops.pr_id id
+ | Evar_kinds.BinderType Anonymous -> str "type of anonymous binder"
+ | Evar_kinds.ImplicitArg (c,(n,ido),b) ->
let id = Option.get ido in
str "parameter " ++ pr_id id ++ spc () ++ str "of" ++
- spc () ++ print_constr (constr_of_global c)
- | InternalHole -> str "internal placeholder"
- | TomatchTypeParameter (ind,n) ->
- nth n ++ str " argument of type " ++ print_constr (mkInd ind)
- | GoalEvar -> str "goal evar"
- | ImpossibleCase -> str "type of impossible pattern-matching clause"
- | MatchingVar _ -> str "matching variable"
+ spc () ++ print_constr (printable_constr_of_global c)
+ | Evar_kinds.InternalHole -> str "internal placeholder"
+ | Evar_kinds.TomatchTypeParameter (ind,n) ->
+ pr_nth n ++ str " argument of type " ++ print_constr (mkInd ind)
+ | Evar_kinds.GoalEvar -> str "goal evar"
+ | Evar_kinds.ImpossibleCase -> str "type of impossible pattern-matching clause"
+ | Evar_kinds.MatchingVar _ -> str "matching variable"
+ | Evar_kinds.VarInstance id -> str "instance of " ++ pr_id id
+ | Evar_kinds.SubEvar evk ->
+ str "subterm of " ++ str (string_of_existential evk)
let pr_evar_info evi =
let phyps =
try
- let decls = List.combine (evar_context evi) (evar_filter evi) in
- prlist_with_sep pr_spc pr_decl (List.rev decls)
+ let decls = match Filter.repr (evar_filter evi) with
+ | None -> List.map (fun c -> (c, true)) (evar_context evi)
+ | Some filter -> List.combine (evar_context evi) filter
+ in
+ prlist_with_sep spc pr_decl (List.rev decls)
with Invalid_argument _ -> str "Ill-formed filtered context" in
let pty = print_constr evi.evar_concl in
let pb =
@@ -794,95 +1646,151 @@ let pr_evar_info evi =
(str"[" ++ phyps ++ spc () ++ str"|- " ++ pty ++ pb ++ str"]" ++
candidates ++ spc() ++ src)
-let compute_evar_dependency_graph (sigma:evar_map) =
+let compute_evar_dependency_graph (sigma : evar_map) =
(* Compute the map binding ev to the evars whose body depends on ev *)
- fold (fun evk evi acc ->
- let deps =
- match evar_body evi with
- | Evar_empty -> ExistentialSet.empty
- | Evar_defined c -> collect_evars c in
- ExistentialSet.fold (fun evk' acc ->
- let tab = try ExistentialMap.find evk' acc with Not_found -> [] in
- ExistentialMap.add evk' ((evk,evi)::tab) acc) deps acc)
- sigma ExistentialMap.empty
+ let fold evk evi acc =
+ let fold_ev evk' acc =
+ let tab =
+ try EvMap.find evk' acc
+ with Not_found -> Evar.Set.empty
+ in
+ EvMap.add evk' (Evar.Set.add evk tab) acc
+ in
+ match evar_body evi with
+ | Evar_empty -> assert false
+ | Evar_defined c -> Evar.Set.fold fold_ev (evars_of_term c) acc
+ in
+ EvMap.fold fold sigma.defn_evars EvMap.empty
let evar_dependency_closure n sigma =
+ (** Create the DAG of depth [n] representing the recursive dependencies of
+ undefined evars. *)
let graph = compute_evar_dependency_graph sigma in
- let order a b = fst a < fst b in
- let rec aux n l =
- if n=0 then l
- else
- let l' =
- list_map_append (fun (evk,_) ->
- try ExistentialMap.find evk graph with Not_found -> []) l in
- aux (n-1) (list_uniquize (Sort.list order (l@l'))) in
- aux n (undefined_list sigma)
-
-let pr_evar_map_t depth sigma =
- let (evars,(uvs,univs)) = sigma.evars in
- let pr_evar_list l =
- h 0 (prlist_with_sep pr_fnl
- (fun (ev,evi) ->
- h 0 (str(string_of_existential ev) ++
- str"==" ++ pr_evar_info evi)) l) in
- let evs =
- if EvarInfoMap.is_empty evars then mt ()
+ let rec aux n curr accu =
+ if Int.equal n 0 then Evar.Set.union curr accu
else
- match depth with
- | None ->
- (* Print all evars *)
- str"EVARS:"++brk(0,1)++pr_evar_list (to_list sigma)++fnl()
- | Some n ->
- (* Print all evars *)
- str"UNDEFINED EVARS"++
- (if n=0 then mt() else str" (+level "++int n++str" closure):")++
- brk(0,1)++
- pr_evar_list (evar_dependency_closure n sigma)++fnl()
- and svs =
- if Univ.UniverseLSet.is_empty uvs then mt ()
- else str"UNIVERSE VARIABLES:"++brk(0,1)++
- h 0 (prlist_with_sep pr_fnl
- (fun u -> Univ.pr_uni_level u) (Univ.UniverseLSet.elements uvs))++fnl()
- and cs =
- if Univ.is_initial_universes univs then mt ()
- else str"UNIVERSES:"++brk(0,1)++
- h 0 (Univ.pr_universes univs)++fnl()
- in evs ++ svs ++ cs
+ let fold evk accu =
+ try
+ let deps = EvMap.find evk graph in
+ Evar.Set.union deps accu
+ with Not_found -> accu
+ in
+ (** Consider only the newly added evars *)
+ let ncurr = Evar.Set.fold fold curr Evar.Set.empty in
+ (** Merge the others *)
+ let accu = Evar.Set.union curr accu in
+ aux (n - 1) ncurr accu
+ in
+ let undef = EvMap.domain (undefined_map sigma) in
+ aux n undef Evar.Set.empty
+
+let evar_dependency_closure n sigma =
+ let deps = evar_dependency_closure n sigma in
+ let map = EvMap.bind (fun ev -> find sigma ev) deps in
+ EvMap.bindings map
+
+let has_no_evar sigma =
+ EvMap.is_empty sigma.defn_evars && EvMap.is_empty sigma.undf_evars
+
+let pr_uctx_level uctx =
+ let map, map_rev = uctx.uctx_names in
+ fun l ->
+ try str(Univ.LMap.find l map_rev)
+ with Not_found ->
+ Universes.pr_with_global_universes l
+
+let pr_evd_level evd = pr_uctx_level evd.universes
+
+let pr_evar_universe_context ctx =
+ let prl = pr_uctx_level ctx in
+ if is_empty_evar_universe_context ctx then mt ()
+ else
+ (str"UNIVERSES:"++brk(0,1)++
+ h 0 (Univ.pr_universe_context_set prl ctx.uctx_local) ++ fnl () ++
+ str"ALGEBRAIC UNIVERSES:"++brk(0,1)++
+ h 0 (Univ.LSet.pr prl ctx.uctx_univ_algebraic) ++ fnl() ++
+ str"UNDEFINED UNIVERSES:"++brk(0,1)++
+ h 0 (Universes.pr_universe_opt_subst ctx.uctx_univ_variables) ++ fnl())
let print_env_short env =
- let pr_body n = function None -> pr_name n | Some b -> str "(" ++ pr_name n ++ str " := " ++ print_constr b ++ str ")" in
+ let pr_body n = function
+ | None -> pr_name n
+ | Some b -> str "(" ++ pr_name n ++ str " := " ++ print_constr b ++ str ")" in
let pr_named_decl (n, b, _) = pr_body (Name n) b in
let pr_rel_decl (n, b, _) = pr_body n b in
let nc = List.rev (named_context env) in
let rc = List.rev (rel_context env) in
- str "[" ++ prlist_with_sep pr_spc pr_named_decl nc ++ str "]" ++ spc () ++
- str "[" ++ prlist_with_sep pr_spc pr_rel_decl rc ++ str "]"
-
-let pr_constraints pbs =
- h 0
- (prlist_with_sep pr_fnl
- (fun (pbty,env,t1,t2) ->
- print_env_short env ++ spc () ++ str "|-" ++ spc () ++
- print_constr t1 ++ spc() ++
- str (match pbty with
- | Reduction.CONV -> "=="
- | Reduction.CUMUL -> "<=") ++
- spc() ++ print_constr t2) pbs)
-
-let pr_evar_map_constraints evd =
- if evd.conv_pbs = [] then mt()
- else pr_constraints evd.conv_pbs++fnl()
-
-let pr_evar_map allevars evd =
- let pp_evm =
- if EvarMap.is_empty evd.evars then mt() else
- pr_evar_map_t allevars evd++fnl() in
- let cstrs = if evd.conv_pbs = [] then mt() else
- str"CONSTRAINTS:"++brk(0,1)++pr_constraints evd.conv_pbs++fnl() in
- let pp_met =
- if Metamap.is_empty evd.metas then mt() else
- str"METAS:"++brk(0,1)++pr_meta_map evd.metas in
- v 0 (pp_evm ++ cstrs ++ pp_met)
+ str "[" ++ pr_sequence pr_named_decl nc ++ str "]" ++ spc () ++
+ str "[" ++ pr_sequence pr_rel_decl rc ++ str "]"
+
+let pr_evar_constraints pbs =
+ let pr_evconstr (pbty, env, t1, t2) =
+ print_env_short env ++ spc () ++ str "|-" ++ spc () ++
+ print_constr_env env t1 ++ spc () ++
+ str (match pbty with
+ | Reduction.CONV -> "=="
+ | Reduction.CUMUL -> "<=") ++
+ spc () ++ print_constr_env env t2
+ in
+ prlist_with_sep fnl pr_evconstr pbs
+
+let pr_evar_map_gen with_univs pr_evars sigma =
+ let { universes = uvs } = sigma in
+ let evs = if has_no_evar sigma then mt () else pr_evars sigma ++ fnl ()
+ and svs = if with_univs then pr_evar_universe_context uvs else mt ()
+ and cstrs =
+ if List.is_empty sigma.conv_pbs then mt ()
+ else
+ str "CONSTRAINTS:" ++ brk (0, 1) ++
+ pr_evar_constraints sigma.conv_pbs ++ fnl ()
+ and metas =
+ if Metamap.is_empty sigma.metas then mt ()
+ else
+ str "METAS:" ++ brk (0, 1) ++ pr_meta_map sigma.metas
+ in
+ evs ++ svs ++ cstrs ++ metas
+
+let pr_evar_list sigma l =
+ let pr (ev, evi) =
+ h 0 (str (string_of_existential ev) ++
+ str "==" ++ pr_evar_info evi ++
+ (if evi.evar_body == Evar_empty
+ then str " {" ++ pr_id (evar_ident ev sigma) ++ str "}"
+ else mt ()))
+ in
+ h 0 (prlist_with_sep fnl pr l)
+
+let pr_evar_by_depth depth sigma = match depth with
+| None ->
+ (* Print all evars *)
+ str"EVARS:"++brk(0,1)++pr_evar_list sigma (to_list sigma)++fnl()
+| Some n ->
+ (* Print all evars *)
+ str"UNDEFINED EVARS:"++
+ (if Int.equal n 0 then mt() else str" (+level "++int n++str" closure):")++
+ brk(0,1)++
+ pr_evar_list sigma (evar_dependency_closure n sigma)++fnl()
+
+let pr_evar_by_filter filter sigma =
+ let defined = Evar.Map.filter filter sigma.defn_evars in
+ let undefined = Evar.Map.filter filter sigma.undf_evars in
+ let prdef =
+ if Evar.Map.is_empty defined then mt ()
+ else str "DEFINED EVARS:" ++ brk (0, 1) ++
+ pr_evar_list sigma (Evar.Map.bindings defined)
+ in
+ let prundef =
+ if Evar.Map.is_empty undefined then mt ()
+ else str "UNDEFINED EVARS:" ++ brk (0, 1) ++
+ pr_evar_list sigma (Evar.Map.bindings undefined)
+ in
+ prdef ++ prundef
+
+let pr_evar_map ?(with_univs=true) depth sigma =
+ pr_evar_map_gen with_univs (fun sigma -> pr_evar_by_depth depth sigma) sigma
+
+let pr_evar_map_filter ?(with_univs=true) filter sigma =
+ pr_evar_map_gen with_univs (fun sigma -> pr_evar_by_filter filter sigma) sigma
let pr_metaset metas =
- str "[" ++ prlist_with_sep spc pr_meta (Metaset.elements metas) ++ str "]"
+ str "[" ++ pr_sequence pr_meta (Metaset.elements metas) ++ str "]"
diff --git a/pretyping/evd.mli b/pretyping/evd.mli
index d5bf2f68..53f8b0db 100644
--- a/pretyping/evd.mli
+++ b/pretyping/evd.mli
@@ -1,126 +1,108 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Util
+open Loc
open Names
open Term
-open Sign
+open Context
open Environ
-open Libnames
open Mod_subst
-open Termops
-(********************************************************************
- Meta map *)
-
-module Metamap : Map.S with type key = metavariable
-
-module Metaset : Set.S with type elt = metavariable
+(** {5 Existential variables and unification states}
-val meta_exists : (metavariable -> bool) -> Metaset.t -> bool
+ A unification state (of type [evar_map]) is primarily a finite mapping
+ from existential variables to records containing the type of the evar
+ ([evar_concl]), the context under which it was introduced ([evar_hyps])
+ and its definition ([evar_body]). [evar_extra] is used to add any other
+ kind of information.
-type 'a freelisted = {
- rebus : 'a;
- freemetas : Metaset.t }
+ It also contains conversion constraints, debugging information and
+ information about meta variables. *)
-val metavars_of : constr -> Metaset.t
-val mk_freelisted : constr -> constr freelisted
-val map_fl : ('a -> 'b) -> 'a freelisted -> 'b freelisted
+(** {6 Evars} *)
-(** Status of an instance found by unification wrt to the meta it solves:
- - a supertype of the meta (e.g. the solution to ?X <= T is a supertype of ?X)
- - a subtype of the meta (e.g. the solution to T <= ?X is a supertype of ?X)
- - a term that can be eta-expanded n times while still being a solution
- (e.g. the solution [P] to [?X u v = P u v] can be eta-expanded twice)
-*)
-
-type instance_constraint = IsSuperType | IsSubType | Conv
+type evar = existential_key
+(** Existential variables. TODO: Should be made opaque one day. *)
-(** Status of the unification of the type of an instance against the type of
- the meta it instantiates:
- - CoerceToType means that the unification of types has not been done
- and that a coercion can still be inserted: the meta should not be
- substituted freely (this happens for instance given via the
- "with" binding clause).
- - TypeProcessed means that the information obtainable from the
- unification of types has been extracted.
- - TypeNotProcessed means that the unification of types has not been
- done but it is known that no coercion may be inserted: the meta
- can be substituted freely.
-*)
+val string_of_existential : evar -> string
-type instance_typing_status =
- CoerceToType | TypeNotProcessed | TypeProcessed
+(** {6 Evar filters} *)
-(** Status of an instance together with the status of its type unification *)
+module Filter :
+sig
+ type t
+ (** Evar filters, seen as bitmasks. *)
-type instance_status = instance_constraint * instance_typing_status
+ val equal : t -> t -> bool
+ (** Equality over filters *)
-(** Clausal environments *)
+ val identity : t
+ (** The identity filter. *)
-type clbinding =
- | Cltyp of name * constr freelisted
- | Clval of name * (constr freelisted * instance_status) * constr freelisted
+ val filter_list : t -> 'a list -> 'a list
+ (** Filter a list. Sizes must coincide. *)
-val map_clb : (constr -> constr) -> clbinding -> clbinding
+ val filter_array : t -> 'a array -> 'a array
+ (** Filter an array. Sizes must coincide. *)
+ val extend : int -> t -> t
+ (** [extend n f] extends [f] on the left with [n]-th times [true]. *)
-(********************************************************************
- ** Kinds of existential variables ***)
+ val compose : t -> t -> t
+ (** Horizontal composition : [compose f1 f2] only keeps parts of [f2] where
+ [f1] is set. In particular, [f1] and [f2] must have the same length. *)
-(** Should the obligation be defined (opaque or transparent (default)) or
- defined transparent and expanded in the term? *)
+ val apply_subfilter : t -> bool list -> t
+ (** [apply_subfilter f1 f2] applies filter [f2] where [f1] is [true]. In
+ particular, the length of [f2] is the number of times [f1] is [true] *)
-type obligation_definition_status = Define of bool | Expand
+ val restrict_upon : t -> int -> (int -> bool) -> t option
+ (** Ad-hoc primitive. *)
-(** Evars *)
-type hole_kind =
- | ImplicitArg of global_reference * (int * identifier option) * bool (** Force inference *)
- | BinderType of name
- | QuestionMark of obligation_definition_status
- | CasesType
- | InternalHole
- | TomatchTypeParameter of inductive * int
- | GoalEvar
- | ImpossibleCase
- | MatchingVar of bool * identifier
+ val map_along : (bool -> 'a -> bool) -> t -> 'a list -> t
+ (** Apply the function on the filter and the list. Sizes must coincide. *)
-(********************************************************************
- ** Existential variables and unification states ***)
+ val make : bool list -> t
+ (** Create out of a list *)
-(** A unification state (of type [evar_map]) is primarily a finite mapping
- from existential variables to records containing the type of the evar
- ([evar_concl]), the context under which it was introduced ([evar_hyps])
- and its definition ([evar_body]). [evar_extra] is used to add any other
- kind of information.
- It also contains conversion constraints, debugging information and
- information about meta variables. *)
+ val repr : t -> bool list option
+ (** Observe as a bool list. *)
-(** Information about existential variables. *)
-type evar = existential_key
+end
-val string_of_existential : evar -> string
-val existential_of_int : int -> evar
+(** {6 Evar infos} *)
type evar_body =
| Evar_empty
| Evar_defined of constr
+
+module Store : Store.S
+(** Datatype used to store additional information in evar maps. *)
+
type evar_info = {
evar_concl : constr;
+ (** Type of the evar. *)
evar_hyps : named_context_val;
+ (** Context of the evar. *)
evar_body : evar_body;
- evar_filter : bool list;
- evar_source : hole_kind located;
+ (** Optional content of the evar. *)
+ evar_filter : Filter.t;
+ (** Boolean mask over {!evar_hyps}. Should have the same length.
+ TODO: document me more. *)
+ evar_source : Evar_kinds.t located;
+ (** Information about the evar. *)
evar_candidates : constr list option;
- evar_extra : Store.t }
-
-val eq_evar_info : evar_info -> evar_info -> bool
+ (** TODO: document this *)
+ evar_extra : Store.t
+ (** Extra store, used for clever hacks. *)
+}
val make_evar : named_context_val -> types -> evar_info
val evar_concl : evar_info -> constr
@@ -129,171 +111,498 @@ val evar_filtered_context : evar_info -> named_context
val evar_hyps : evar_info -> named_context_val
val evar_filtered_hyps : evar_info -> named_context_val
val evar_body : evar_info -> evar_body
-val evar_filter : evar_info -> bool list
-val evar_unfiltered_env : evar_info -> env
+val evar_filter : evar_info -> Filter.t
val evar_env : evar_info -> env
+val evar_filtered_env : evar_info -> env
-(*** Unification state ***)
-type evar_map
+val map_evar_body : (constr -> constr) -> evar_body -> evar_body
+val map_evar_info : (constr -> constr) -> evar_info -> evar_info
+
+(** {6 Unification state} **)
+
+type evar_universe_context
+(** The universe context associated to an evar map *)
-(** Unification state and existential variables *)
+type evar_map
+(** Type of unification state. Essentially a bunch of state-passing data needed
+ to handle incremental term construction. *)
+val progress_evar_map : evar_map -> evar_map -> bool
(** Assuming that the second map extends the first one, this says if
some existing evar has been refined *)
-val progress_evar_map : evar_map -> evar_map -> bool
val empty : evar_map
+(** The empty evar map. *)
+
+val from_env : ?ctx:evar_universe_context -> env -> evar_map
+(** The empty evar map with given universe context, taking its initial
+ universes from env. *)
+
val is_empty : evar_map -> bool
+(** Whether an evarmap is empty. *)
+
+val has_undefined : evar_map -> bool
(** [has_undefined sigma] is [true] if and only if
there are uninstantiated evars in [sigma]. *)
-val has_undefined : evar_map -> bool
+val add : evar_map -> evar -> evar_info -> evar_map
(** [add sigma ev info] adds [ev] with evar info [info] in sigma.
Precondition: ev must not preexist in [sigma]. *)
-val add : evar_map -> evar -> evar_info -> evar_map
val find : evar_map -> evar -> evar_info
+(** Recover the data associated to an evar. *)
+
val find_undefined : evar_map -> evar -> evar_info
+(** Same as {!find} but restricted to undefined evars. For efficiency
+ reasons. *)
+
val remove : evar_map -> evar -> evar_map
+(** Remove an evar from an evar map. Use with caution. *)
+
val mem : evar_map -> evar -> bool
-val undefined_list : evar_map -> (evar * evar_info) list
-val to_list : evar_map -> (evar * evar_info) list
+(** Whether an evar is present in an evarmap. *)
+
val fold : (evar -> evar_info -> 'a -> 'a) -> evar_map -> 'a -> 'a
+(** Apply a function to all evars and their associated info in an evarmap. *)
+
val fold_undefined : (evar -> evar_info -> 'a -> 'a) -> evar_map -> 'a -> 'a
-val merge : evar_map -> evar_map -> evar_map
+(** Same as {!fold}, but restricted to undefined evars. For efficiency
+ reasons. *)
+
+val raw_map : (evar -> evar_info -> evar_info) -> evar_map -> evar_map
+(** Apply the given function to all evars in the map. Beware: this function
+ expects the argument function to preserve the kind of [evar_body], i.e. it
+ must send [Evar_empty] to [Evar_empty] and [Evar_defined c] to some
+ [Evar_defined c']. *)
+
+val raw_map_undefined : (evar -> evar_info -> evar_info) -> evar_map -> evar_map
+(** Same as {!raw_map}, but restricted to undefined evars. For efficiency
+ reasons. *)
+
val define : evar -> constr -> evar_map -> evar_map
+(** Set the body of an evar to the given constr. It is expected that:
+ {ul
+ {- The evar is already present in the evarmap.}
+ {- The evar is not defined in the evarmap yet.}
+ {- All the evars present in the constr should be present in the evar map.}
+ } *)
+
+val cmap : (constr -> constr) -> evar_map -> evar_map
+(** Map the function on all terms in the evar map. *)
val is_evar : evar_map -> evar -> bool
+(** Alias for {!mem}. *)
val is_defined : evar_map -> evar -> bool
+(** Whether an evar is defined in an evarmap. *)
+
val is_undefined : evar_map -> evar -> bool
+(** Whether an evar is not defined in an evarmap. *)
val add_constraints : evar_map -> Univ.constraints -> evar_map
+(** Add universe constraints in an evar map. *)
-(** {6 ... } *)
-(** [existential_value sigma ev] raises [NotInstantiatedEvar] if [ev] has
- no body and [Not_found] if it does not exist in [sigma] *)
+val undefined_map : evar_map -> evar_info Evar.Map.t
+(** Access the undefined evar mapping directly. *)
+
+val eq_evar_info : evar_map -> evar_info -> evar_info -> bool
+(** Compare the evar_info's up to the universe constraints of the evar map. *)
+
+val drop_all_defined : evar_map -> evar_map
+
+(** {6 Instantiating partial terms} *)
exception NotInstantiatedEvar
+
val existential_value : evar_map -> existential -> constr
+(** [existential_value sigma ev] raises [NotInstantiatedEvar] if [ev] has
+ no body and [Not_found] if it does not exist in [sigma] *)
+
val existential_type : evar_map -> existential -> types
+
val existential_opt_value : evar_map -> existential -> constr option
+(** Same as {!existential_value} but returns an option instead of raising an
+ exception. *)
-val instantiate_evar : named_context -> constr -> constr list -> constr
+val evar_instance_array : (Id.t -> 'a -> bool) -> evar_info ->
+ 'a array -> (Id.t * 'a) list
-(** Assume empty universe constraints in [evar_map] and [conv_pbs] *)
-val subst_evar_defs_light : substitution -> evar_map -> evar_map
+val instantiate_evar_array : evar_info -> constr -> constr array -> constr
+val evars_reset_evd : ?with_conv_pbs:bool -> ?with_univs:bool ->
+ evar_map -> evar_map -> evar_map
(** spiwack: this function seems to somewhat break the abstraction. *)
-val evars_reset_evd : ?with_conv_pbs:bool -> evar_map -> evar_map -> evar_map
+(** {6 Misc} *)
-(* spiwack: [is_undefined_evar] should be considered a candidate
- for moving to evarutils *)
-val is_undefined_evar : evar_map -> constr -> bool
-val undefined_evars : evar_map -> evar_map
-val defined_evars : evar_map -> evar_map
-(* [fold_undefined f m] iterates ("folds") function [f] over the undefined
- evars (that is, whose value is [Evar_empty]) of map [m].
- It optimizes the call of {!Evd.fold} to [f] and [undefined_evars m] *)
-val fold_undefined : (evar -> evar_info -> 'a -> 'a) -> evar_map -> 'a -> 'a
val evar_declare :
- named_context_val -> evar -> types -> ?src:loc * hole_kind ->
- ?filter:bool list -> ?candidates:constr list -> evar_map -> evar_map
-val evar_source : existential_key -> evar_map -> hole_kind located
+ named_context_val -> evar -> types -> ?src:Loc.t * Evar_kinds.t ->
+ ?filter:Filter.t -> ?candidates:constr list -> ?store:Store.t ->
+ ?naming:Misctypes.intro_pattern_naming_expr -> evar_map -> evar_map
+(** Convenience function. Just a wrapper around {!add}. *)
+
+val restrict : evar -> evar -> Filter.t -> ?candidates:constr list ->
+ evar_map -> evar_map
+(** Restrict an undefined evar into a new evar by filtering context and
+ possibly limiting the instances to a set of candidates *)
+
+val downcast : evar -> types -> evar_map -> evar_map
+(** Change the type of an undefined evar to a new type assumed to be a
+ subtype of its current type; subtyping must be ensured by caller *)
+
+val evar_source : existential_key -> evar_map -> Evar_kinds.t located
+(** Convenience function. Wrapper around {!find} to recover the source of an
+ evar in a given evar map. *)
+
+val evar_ident : existential_key -> evar_map -> Id.t
+
+val rename : existential_key -> Id.t -> evar_map -> evar_map
+
+val evar_key : Id.t -> evar_map -> existential_key
+
+val evar_source_of_meta : metavariable -> evar_map -> Evar_kinds.t located
+
+val dependent_evar_ident : existential_key -> evar_map -> Id.t
+
+(** {5 Side-effects} *)
+
+val emit_side_effects : Declareops.side_effects -> evar_map -> evar_map
+(** Push a side-effect into the evar map. *)
+
+val eval_side_effects : evar_map -> Declareops.side_effects
+(** Return the effects contained in the evar map. *)
+
+val drop_side_effects : evar_map -> evar_map
+(** This should not be used. For hacking purposes. *)
-(* spiwack: this function seems to somewhat break the abstraction.
- [evar_merge evd ev1] extends the evars of [evd] with [evd1] *)
-val evar_merge : evar_map -> evar_map -> evar_map
+(** {5 Future goals} *)
+
+val declare_future_goal : Evar.t -> evar_map -> evar_map
+(** Adds an existential variable to the list of future goals. For
+ internal uses only. *)
+
+val declare_principal_goal : Evar.t -> evar_map -> evar_map
+(** Adds an existential variable to the list of future goals and make
+ it principal. Only one existential variable can be made principal, an
+ error is raised otherwise. For internal uses only. *)
+
+val future_goals : evar_map -> Evar.t list
+(** Retrieves the list of future goals. Used by the [refine] primitive
+ of the tactic engine. *)
+
+val principal_future_goal : evar_map -> Evar.t option
+(** Retrieves the name of the principal existential variable if there
+ is one. Used by the [refine] primitive of the tactic engine. *)
+
+val reset_future_goals : evar_map -> evar_map
+(** Clears the list of future goals (as well as the principal future
+ goal). Used by the [refine] primitive of the tactic engine. *)
+
+val restore_future_goals : evar_map -> Evar.t list -> Evar.t option -> evar_map
+(** Sets the future goals (including the principal future goal) to a
+ previous value. Intended to be used after a local list of future
+ goals has been consumed. Used by the [refine] primitive of the
+ tactic engine. *)
+
+(** {5 Sort variables}
+
+ Evar maps also keep track of the universe constraints defined at a given
+ point. This section defines the relevant manipulation functions. *)
+
+val whd_sort_variable : evar_map -> constr -> constr
+
+exception UniversesDiffer
+
+val add_universe_constraints : evar_map -> Universes.universe_constraints -> evar_map
+(** Add the given universe unification constraints to the evar map.
+ @raises UniversesDiffer in case a first-order unification fails.
+ @raises UniverseInconsistency
+*)
+(** {5 Enriching with evar maps} *)
+
+type 'a sigma = {
+ it : 'a ;
+ (** The base object. *)
+ sigma : evar_map
+ (** The added unification state. *)
+}
+(** The type constructor ['a sigma] adds an evar map to an object of type
+ ['a]. *)
+
+val sig_it : 'a sigma -> 'a
+val sig_sig : 'a sigma -> evar_map
+val on_sig : 'a sigma -> (evar_map -> evar_map * 'b) -> 'a sigma * 'b
+
+(** {5 The state monad with state an evar map} *)
+
+module MonadR : Monad.S with type +'a t = evar_map -> evar_map * 'a
+module Monad : Monad.S with type +'a t = evar_map -> 'a * evar_map
+
+
+(** {5 Meta machinery}
+
+ These functions are almost deprecated. They were used before the
+ introduction of the full-fledged evar calculus. In an ideal world, they
+ should be removed. Alas, some parts of the code still use them. Do not use
+ in newly-written code. *)
+
+module Metaset : Set.S with type elt = metavariable
+module Metamap : Map.ExtS with type key = metavariable and module Set := Metaset
+
+type 'a freelisted = {
+ rebus : 'a;
+ freemetas : Metaset.t }
+
+val metavars_of : constr -> Metaset.t
+val mk_freelisted : constr -> constr freelisted
+val map_fl : ('a -> 'b) -> 'a freelisted -> 'b freelisted
+
+(** Status of an instance found by unification wrt to the meta it solves:
+ - a supertype of the meta (e.g. the solution to ?X <= T is a supertype of ?X)
+ - a subtype of the meta (e.g. the solution to T <= ?X is a supertype of ?X)
+ - a term that can be eta-expanded n times while still being a solution
+ (e.g. the solution [P] to [?X u v = P u v] can be eta-expanded twice)
+*)
+
+type instance_constraint = IsSuperType | IsSubType | Conv
+
+val eq_instance_constraint :
+ instance_constraint -> instance_constraint -> bool
+
+(** Status of the unification of the type of an instance against the type of
+ the meta it instantiates:
+ - CoerceToType means that the unification of types has not been done
+ and that a coercion can still be inserted: the meta should not be
+ substituted freely (this happens for instance given via the
+ "with" binding clause).
+ - TypeProcessed means that the information obtainable from the
+ unification of types has been extracted.
+ - TypeNotProcessed means that the unification of types has not been
+ done but it is known that no coercion may be inserted: the meta
+ can be substituted freely.
+*)
+
+type instance_typing_status =
+ CoerceToType | TypeNotProcessed | TypeProcessed
+
+(** Status of an instance together with the status of its type unification *)
+
+type instance_status = instance_constraint * instance_typing_status
+
+(** Clausal environments *)
+
+type clbinding =
+ | Cltyp of Name.t * constr freelisted
+ | Clval of Name.t * (constr freelisted * instance_status) * constr freelisted
(** Unification constraints *)
type conv_pb = Reduction.conv_pb
type evar_constraint = conv_pb * env * constr * constr
val add_conv_pb : evar_constraint -> evar_map -> evar_map
-module ExistentialMap : Map.S with type key = existential_key
-module ExistentialSet : Set.S with type elt = existential_key
val extract_changed_conv_pbs : evar_map ->
- (ExistentialSet.t -> evar_constraint -> bool) ->
+ (Evar.Set.t -> evar_constraint -> bool) ->
evar_map * evar_constraint list
val extract_all_conv_pbs : evar_map -> evar_map * evar_constraint list
+val loc_of_conv_pb : evar_map -> evar_constraint -> Loc.t
+
+(** The following functions return the set of evars immediately
+ contained in the object; need the term to be evar-normal otherwise
+ defined evars are returned too. *)
+
+val evar_list : constr -> existential list
+ (** excluding evars in instances of evars and collected with
+ redundancies from right to left (used by tactic "instantiate") *)
-val evar_list : evar_map -> constr -> existential list
-val collect_evars : constr -> ExistentialSet.t
+val evars_of_term : constr -> Evar.Set.t
+ (** including evars in instances of evars *)
+
+val evars_of_named_context : named_context -> Evar.Set.t
+
+val evars_of_filtered_evar_info : evar_info -> Evar.Set.t
(** Metas *)
-val find_meta : evar_map -> metavariable -> clbinding
val meta_list : evar_map -> (metavariable * clbinding) list
val meta_defined : evar_map -> metavariable -> bool
+val meta_value : evar_map -> metavariable -> constr
(** [meta_fvalue] raises [Not_found] if meta not in map or [Anomaly] if
meta has no value *)
-val meta_value : evar_map -> metavariable -> constr
+
val meta_fvalue : evar_map -> metavariable -> constr freelisted * instance_status
val meta_opt_fvalue : evar_map -> metavariable -> (constr freelisted * instance_status) option
val meta_type : evar_map -> metavariable -> types
val meta_ftype : evar_map -> metavariable -> types freelisted
-val meta_name : evar_map -> metavariable -> name
-val meta_with_name : evar_map -> identifier -> metavariable
+val meta_name : evar_map -> metavariable -> Name.t
+val meta_with_name : evar_map -> Id.t -> metavariable
val meta_declare :
- metavariable -> types -> ?name:name -> evar_map -> evar_map
+ metavariable -> types -> ?name:Name.t -> evar_map -> evar_map
val meta_assign : metavariable -> constr * instance_status -> evar_map -> evar_map
val meta_reassign : metavariable -> constr * instance_status -> evar_map -> evar_map
+val clear_metas : evar_map -> evar_map
+
(** [meta_merge evd1 evd2] returns [evd2] extended with the metas of [evd1] *)
val meta_merge : evar_map -> evar_map -> evar_map
val undefined_metas : evar_map -> metavariable list
-val metas_of : evar_map -> meta_type_map
val map_metas_fvalue : (constr -> constr) -> evar_map -> evar_map
type metabinding = metavariable * constr * instance_status
val retract_coercible_metas : evar_map -> metabinding list * evar_map
-val subst_defined_metas : metabinding list -> constr -> constr option
+val subst_defined_metas_evars : metabinding list * ('a * existential * constr) list -> constr -> constr option
+
+(** {5 FIXME: Nothing to do here} *)
(*********************************************************
- Sort variables *)
+ Sort/universe variables *)
+
+(** Rigid or flexible universe variables *)
+
+type rigid =
+ | UnivRigid
+ | UnivFlexible of bool (** Is substitution by an algebraic ok? *)
+
+val univ_rigid : rigid
+val univ_flexible : rigid
+val univ_flexible_alg : rigid
+
+type 'a in_evar_universe_context = 'a * evar_universe_context
+
+val evar_universe_context_set : evar_universe_context -> Univ.universe_context_set
+val evar_universe_context_constraints : evar_universe_context -> Univ.constraints
+val evar_context_universe_context : evar_universe_context -> Univ.universe_context
+val evar_universe_context_of : Univ.universe_context_set -> evar_universe_context
+val empty_evar_universe_context : evar_universe_context
+val union_evar_universe_context : evar_universe_context -> evar_universe_context ->
+ evar_universe_context
+val evar_universe_context_subst : evar_universe_context -> Universes.universe_opt_subst
+
+(** Raises Not_found if not a name for a universe in this map. *)
+val universe_of_name : evar_map -> string -> Univ.universe_level
+val add_universe_name : evar_map -> string -> Univ.universe_level -> evar_map
+
+val universes : evar_map -> Univ.universes
+
+val add_constraints_context : evar_universe_context ->
+ Univ.constraints -> evar_universe_context
+
+
+val normalize_evar_universe_context_variables : evar_universe_context ->
+ Univ.universe_subst in_evar_universe_context
+
+val normalize_evar_universe_context : evar_universe_context ->
+ evar_universe_context
+
+val new_univ_level_variable : ?name:string -> rigid -> evar_map -> evar_map * Univ.universe_level
+val new_univ_variable : ?name:string -> rigid -> evar_map -> evar_map * Univ.universe
+val new_sort_variable : ?name:string -> rigid -> evar_map -> evar_map * sorts
+val make_flexible_variable : evar_map -> bool -> Univ.universe_level -> evar_map
+val is_sort_variable : evar_map -> sorts -> Univ.universe_level option
+(** [is_sort_variable evm s] returns [Some u] or [None] if [s] is
+ not a local sort variable declared in [evm] *)
+val is_flexible_level : evar_map -> Univ.Level.t -> bool
-val new_univ_variable : evar_map -> evar_map * Univ.universe
-val new_sort_variable : evar_map -> evar_map * sorts
-val is_sort_variable : evar_map -> sorts -> bool
val whd_sort_variable : evar_map -> constr -> constr
-val set_leq_sort : evar_map -> sorts -> sorts -> evar_map
-val set_eq_sort : evar_map -> sorts -> sorts -> evar_map
+(* val normalize_universe_level : evar_map -> Univ.universe_level -> Univ.universe_level *)
+val normalize_universe : evar_map -> Univ.universe -> Univ.universe
+val normalize_universe_instance : evar_map -> Univ.universe_instance -> Univ.universe_instance
-(********************************************************************
- constr with holes *)
-type open_constr = evar_map * constr
+val set_leq_sort : env -> evar_map -> sorts -> sorts -> evar_map
+val set_eq_sort : env -> evar_map -> sorts -> sorts -> evar_map
+val has_lub : evar_map -> Univ.universe -> Univ.universe -> evar_map
+val set_eq_level : evar_map -> Univ.universe_level -> Univ.universe_level -> evar_map
+val set_leq_level : evar_map -> Univ.universe_level -> Univ.universe_level -> evar_map
+val set_eq_instances : ?flex:bool ->
+ evar_map -> Univ.universe_instance -> Univ.universe_instance -> evar_map
+
+val check_eq : evar_map -> Univ.universe -> Univ.universe -> bool
+val check_leq : evar_map -> Univ.universe -> Univ.universe -> bool
+
+val evar_universe_context : evar_map -> evar_universe_context
+val universe_context_set : evar_map -> Univ.universe_context_set
+val universe_context : evar_map -> Univ.universe_context
+val universe_subst : evar_map -> Universes.universe_opt_subst
+val universes : evar_map -> Univ.universes
+
+
+val merge_universe_context : evar_map -> evar_universe_context -> evar_map
+val set_universe_context : evar_map -> evar_universe_context -> evar_map
+
+val merge_context_set : rigid -> evar_map -> Univ.universe_context_set -> evar_map
+val merge_universe_subst : evar_map -> Universes.universe_opt_subst -> evar_map
+
+val with_context_set : rigid -> evar_map -> 'a Univ.in_universe_context_set -> evar_map * 'a
+
+val nf_univ_variables : evar_map -> evar_map * Univ.universe_subst
+val abstract_undefined_variables : evar_universe_context -> evar_universe_context
+
+val refresh_undefined_universes : evar_map -> evar_map * Univ.universe_level_subst
+
+val nf_constraints : evar_map -> evar_map
+
+(** Polymorphic universes *)
+
+val fresh_sort_in_family : ?rigid:rigid -> env -> evar_map -> sorts_family -> evar_map * sorts
+val fresh_constant_instance : env -> evar_map -> constant -> evar_map * pconstant
+val fresh_inductive_instance : env -> evar_map -> inductive -> evar_map * pinductive
+val fresh_constructor_instance : env -> evar_map -> constructor -> evar_map * pconstructor
+
+val fresh_global : ?rigid:rigid -> ?names:Univ.Instance.t -> env -> evar_map ->
+ Globnames.global_reference -> evar_map * constr
(********************************************************************
- The type constructor ['a sigma] adds an evar map to an object of
- type ['a] *)
-type 'a sigma = {
- it : 'a ;
- sigma : evar_map}
+ Conversion w.r.t. an evar map: might generate universe unifications
+ that are kept in the evarmap.
+ Raises [NotConvertible]. *)
-val sig_it : 'a sigma -> 'a
-val sig_sig : 'a sigma -> evar_map
+val conversion : env -> evar_map -> conv_pb -> constr -> constr -> evar_map
-(*********************************************************
- Failure explanation *)
+val test_conversion : env -> evar_map -> conv_pb -> constr -> constr -> bool
+(** This one forgets about the assignemts of universes. *)
+
+val eq_constr_univs : evar_map -> constr -> constr -> evar_map * bool
+(** Syntactic equality up to universes, recording the associated constraints *)
+
+val e_eq_constr_univs : evar_map ref -> constr -> constr -> bool
+(** Syntactic equality up to universes. *)
+
+val eq_constr_univs_test : evar_map -> constr -> constr -> bool
+(** Syntactic equality up to universes, throwing away the (consistent) constraints
+ in case of success. *)
+
+(********************************************************************)
+(* constr with holes and pending resolution of classes, conversion *)
+(* problems, candidates, etc. *)
+
+type pending = (* before: *) evar_map * (* after: *) evar_map
+
+type pending_constr = pending * constr
+
+type open_constr = evar_map * constr (* Special case when before is empty *)
+
+(** Partially constructed constrs. *)
type unsolvability_explanation = SeveralInstancesFound of int
+(** Failure explanation. *)
-(********************************************************************
- debug pretty-printer: *)
+val pr_existential_key : evar_map -> evar -> Pp.std_ppcmds
+
+(** {5 Debug pretty-printers} *)
val pr_evar_info : evar_info -> Pp.std_ppcmds
-val pr_evar_map_constraints : evar_map -> Pp.std_ppcmds
-val pr_evar_map : int option -> evar_map -> Pp.std_ppcmds
+val pr_evar_constraints : evar_constraint list -> Pp.std_ppcmds
+val pr_evar_map : ?with_univs:bool -> int option -> evar_map -> Pp.std_ppcmds
+val pr_evar_map_filter : ?with_univs:bool -> (Evar.t -> evar_info -> bool) ->
+ evar_map -> Pp.std_ppcmds
val pr_metaset : Metaset.t -> Pp.std_ppcmds
+val pr_evar_universe_context : evar_universe_context -> Pp.std_ppcmds
+val pr_evd_level : evar_map -> Univ.Level.t -> Pp.std_ppcmds
+(** {5 Deprecated functions} *)
-(*** /!\Deprecated /!\ **
- create an [evar_map] with empty meta map: *)
val create_evar_defs : evar_map -> evar_map
+(** Create an [evar_map] with empty meta map: *)
+
val create_goal_evar_defs : evar_map -> evar_map
-val is_defined_evar : evar_map -> existential -> bool
-val subst_evar_map : substitution -> evar_map -> evar_map
-(*** /Deprecaded ***)
diff --git a/pretyping/find_subterm.ml b/pretyping/find_subterm.ml
new file mode 100644
index 00000000..7f7f4d76
--- /dev/null
+++ b/pretyping/find_subterm.ml
@@ -0,0 +1,179 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Pp
+open Util
+open Errors
+open Names
+open Locus
+open Context
+open Term
+open Nameops
+open Termops
+open Pretype_errors
+
+(** Processing occurrences *)
+
+type occurrence_error =
+ | InvalidOccurrence of int list
+ | IncorrectInValueOccurrence of Id.t
+
+let explain_invalid_occurrence l =
+ let l = List.sort_uniquize Int.compare l in
+ str ("Invalid occurrence " ^ String.plural (List.length l) "number" ^": ")
+ ++ prlist_with_sep spc int l ++ str "."
+
+let explain_incorrect_in_value_occurrence id =
+ pr_id 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)
+
+let error_invalid_occurrence occ =
+ error_occurrences_error (InvalidOccurrence occ)
+
+let check_used_occurrences nbocc (nowhere_except_in,locs) =
+ let rest = List.filter (fun o -> o >= nbocc) locs in
+ match rest with
+ | [] -> ()
+ | _ -> error_occurrences_error (InvalidOccurrence rest)
+
+let proceed_with_occurrences f occs x =
+ match occs with
+ | NoOccurrences -> x
+ | occs ->
+ let plocs = Locusops.convert_occs occs in
+ assert (List.for_all (fun x -> x >= 0) (snd plocs));
+ let (nbocc,x) = f 1 x in
+ check_used_occurrences nbocc plocs;
+ x
+
+(** Applying a function over a named_declaration with an hypothesis
+ location request *)
+
+let map_named_declaration_with_hyploc f hyploc acc (id,bodyopt,typ) =
+ let f = f (Some (id,hyploc)) in
+ match bodyopt,hyploc with
+ | None, InHypValueOnly ->
+ error_occurrences_error (IncorrectInValueOccurrence id)
+ | None, _ | Some _, InHypTypeOnly ->
+ let acc,typ = f acc typ in acc,(id,bodyopt,typ)
+ | Some body, InHypValueOnly ->
+ let acc,body = f acc body in acc,(id,Some body,typ)
+ | Some body, InHyp ->
+ let acc,body = f acc body in
+ let acc,typ = f acc typ in
+ acc,(id,Some body,typ)
+
+(** Finding a subterm up to some testing function *)
+
+exception SubtermUnificationError of subterm_unification_error
+
+exception NotUnifiable of (constr * constr * unification_error) option
+
+type 'a testing_function = {
+ match_fun : 'a -> constr -> 'a;
+ merge_fun : 'a -> 'a -> 'a;
+ mutable testing_state : 'a;
+ mutable last_found : position_reporting option
+}
+
+(* Find subterms using a testing function, but only at a list of
+ locations or excluding a list of locations; in the occurrences list
+ (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 (nowhere_except_in,locs) = Locusops.convert_occs occs in
+ let maxocc = List.fold_right max locs 0 in
+ let pos = ref occ in
+ let nested = ref false in
+ let add_subst t subst =
+ try
+ test.testing_state <- test.merge_fun subst test.testing_state;
+ test.last_found <- Some ((cl,!pos),t)
+ with NotUnifiable e when not like_first ->
+ let lastpos = Option.get test.last_found in
+ raise (SubtermUnificationError (!nested,((cl,!pos),t),lastpos,e)) in
+ let rec substrec k t =
+ if nowhere_except_in && !pos > maxocc then t else
+ if not (Vars.closed0 t) then subst_below k t else
+ try
+ let subst = test.match_fun test.testing_state t in
+ if Locusops.is_selected !pos occs then
+ (if !nested then begin
+ (* in case it is nested but not later detected as unconvertible,
+ as when matching "id _" in "id (id 0)" *)
+ let lastpos = Option.get test.last_found in
+ raise (SubtermUnificationError (!nested,((cl,!pos),t),lastpos,None))
+ end;
+ add_subst t subst; incr pos;
+ (* Check nested matching subterms *)
+ if occs != Locus.AllOccurrences && occs != Locus.NoOccurrences then
+ begin nested := true; ignore (subst_below k t); nested := false end;
+ (* Do the effective substitution *)
+ Vars.lift k (bywhat ()))
+ else
+ (incr pos; subst_below k 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
+ in
+ let t' = substrec 0 t in
+ (!pos, t')
+
+let replace_term_occ_modulo 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
+
+let replace_term_occ_decl_modulo 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)
+ hyploc)
+ plocs d
+
+(** Finding an exact subterm *)
+
+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
+ 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
+}
+
+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
+ t', test.testing_state
+
+let subst_closed_term_occ_decl env evd occs c d =
+ let test = make_eq_univs_test env evd c in
+ let (plocs,hyploc),like_first =
+ match occs with AtOccs occs -> occs,false | LikeFirst -> (AllOccurrences,InHyp),true in
+ 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)
+ hyploc) plocs d,
+ test.testing_state
diff --git a/pretyping/find_subterm.mli b/pretyping/find_subterm.mli
new file mode 100644
index 00000000..82330b84
--- /dev/null
+++ b/pretyping/find_subterm.mli
@@ -0,0 +1,69 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Locus
+open Context
+open Term
+open Evd
+open Pretype_errors
+open Environ
+
+(** Finding subterms, possibly up to some unification function,
+ possibly at some given occurrences *)
+
+exception NotUnifiable of (constr * constr * unification_error) option
+
+exception SubtermUnificationError of subterm_unification_error
+
+(** A testing function is typically a unification function returning a
+ substitution or failing with a NotUnifiable error, together with a
+ function to merge substitutions and an initial substitution;
+ last_found is used for error messages and it has to be initialized
+ with None. *)
+
+type 'a testing_function = {
+ match_fun : 'a -> constr -> 'a;
+ merge_fun : 'a -> 'a -> 'a;
+ mutable testing_state : 'a;
+ mutable last_found : position_reporting option
+}
+
+(** This is the basic testing function, looking for exact matches of a
+ closed term *)
+val make_eq_univs_test : env -> evar_map -> constr -> evar_map testing_function
+
+(** [replace_term_occ_modulo occl test mk c] looks in [c] for subterm
+ modulo a testing function [test] and replaces successfully
+ 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 ->
+ '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 :
+ (occurrences * hyp_location_flag) or_like_first ->
+ 'a testing_function -> (unit -> constr) ->
+ 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),
+ unifying universes which results in a set of constraints. *)
+val subst_closed_term_occ : env -> evar_map -> occurrences or_like_first ->
+ constr -> constr -> constr * evar_map
+
+(** [subst_closed_term_occ_decl evd occl c decl] replaces occurrences of
+ 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 -> named_declaration -> named_declaration * evar_map
+
+(** Miscellaneous *)
+val error_invalid_occurrence : int list -> 'a
diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml
new file mode 100644
index 00000000..454d64f0
--- /dev/null
+++ b/pretyping/glob_ops.ml
@@ -0,0 +1,434 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Util
+open Names
+open Globnames
+open Misctypes
+open Glob_term
+
+(* Untyped intermediate terms, after ASTs and before constr. *)
+
+let cases_pattern_loc = function
+ PatVar(loc,_) -> loc
+ | PatCstr(loc,_,_,_) -> loc
+
+let cases_predicate_names tml =
+ List.flatten (List.map (function
+ | (tm,(na,None)) -> [na]
+ | (tm,(na,Some (_,_,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 map_glob_decl_left_to_right f (na,k,obd,ty) =
+ let comp1 = Option.map f obd in
+ let comp2 = f ty in
+ (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
+
+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
+
+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) =
+ 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
+
+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 glob_decl_eq (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) ->
+ let comp1 = f g in
+ let comp2 = Util.List.map_left f args in
+ GApp (loc,comp1,comp2)
+ | GLambda (loc,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) ->
+ let comp1 = f ty in
+ let comp2 = f c in
+ GProd (loc,na,bk,comp1,comp2)
+ | GLetIn (loc,na,b,c) ->
+ let comp1 = f b in
+ let comp2 = f c in
+ GLetIn (loc,na,comp1,comp2)
+ | GCases (loc,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 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) ->
+ 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) ->
+ 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) ->
+ let comp1 = f c in
+ let comp2 = Miscops.map_cast_type f k in
+ GCast (loc,comp1,comp2)
+ | (GVar _ | GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) as x -> x
+
+let map_glob_constr = map_glob_constr_left_to_right
+
+let fold_glob_constr f acc =
+ let rec fold acc = function
+ | GVar _ -> acc
+ | GApp (_,c,args) -> List.fold_left fold (fold acc c) args
+ | GLambda (_,_,_,b,c) | GProd (_,_,_,b,c) | GLetIn (_,_,b,c) ->
+ fold (fold acc b) c
+ | GCases (_,_,rtntypopt,tml,pl) ->
+ List.fold_left fold_pattern
+ (List.fold_left fold (Option.fold_left fold acc rtntypopt) (List.map fst tml))
+ pl
+ | GLetTuple (_,_,rtntyp,b,c) ->
+ fold (fold (fold_return_type acc rtntyp) b) c
+ | GIf (_,c,rtntyp,b1,b2) ->
+ fold (fold (fold (fold_return_type acc rtntyp) c) b1) b2
+ | GRec (_,_,_,bl,tyl,bv) ->
+ let acc = Array.fold_left
+ (List.fold_left (fun acc (na,k,bbd,bty) ->
+ fold (Option.fold_left fold acc bbd) bty)) acc bl in
+ Array.fold_left fold (Array.fold_left fold acc tyl) bv
+ | GCast (_,c,k) ->
+ let r = match k with
+ | CastConv t | CastVM t | CastNative t -> fold acc t | CastCoerce -> acc
+ in
+ fold r c
+ | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> acc
+
+ and fold_pattern acc (_,idl,p,c) = fold acc c
+
+ and fold_return_type acc (na,tyopt) = Option.fold_left fold acc tyopt
+
+ in fold acc
+
+let iter_glob_constr f = fold_glob_constr (fun () -> f) ()
+
+let same_id na id = match na with
+| Anonymous -> false
+| Name id' -> Id.equal id id'
+
+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 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
+ fun rt ->
+ let vs = vars Id.Set.empty Id.Set.empty rt in
+ Id.Set.elements vs
+
+(** Mapping of names in binders *)
+
+(* spiwack: I used a smartmap-style kind of mapping here, because the
+ operation will be the identity almost all of the time (with any
+ term outside of Ltac to begin with). But to be honest, there would
+ 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 r = CList.smartmap f nal in
+ if r == nal then x
+ else 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 r = f na in
+ if r == na then x
+ else PatVar (loc,r)
+ | PatCstr (loc,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)
+
+let map_cases_branch_binders f ((loc,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
+
+let map_pattern_binders f tomatch branches =
+ CList.smartmap (fun tm -> map_tomatch_binders f tm) tomatch,
+ CList.smartmap (fun br -> map_cases_branch_binders f br) branches
+
+(** /mapping of names in binders *)
+
+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_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
+
+(**********************************************************************)
+(* Conversion from glob_constr to cases pattern, if possible *)
+
+let rec cases_pattern_of_glob_constr na = function
+ | GVar (loc,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)
+ 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)
+ | _ -> raise Not_found
+
+(* 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)
+ | _ -> raise Not_found
+
+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))
+ | _ ->
+ raise Not_found
diff --git a/pretyping/glob_ops.mli b/pretyping/glob_ops.mli
new file mode 100644
index 00000000..67f3cb41
--- /dev/null
+++ b/pretyping/glob_ops.mli
@@ -0,0 +1,61 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Glob_term
+
+(** Equalities *)
+
+val cases_pattern_eq : cases_pattern -> cases_pattern -> bool
+
+val glob_constr_eq : glob_constr -> glob_constr -> bool
+
+(** Operations on [glob_constr] *)
+
+val cases_pattern_loc : cases_pattern -> Loc.t
+
+val cases_predicate_names : tomatch_tuples -> Name.t list
+
+(** Apply one argument to a glob_constr *)
+val mkGApp : Loc.t -> glob_constr -> glob_constr -> glob_constr
+
+val map_glob_constr :
+ (glob_constr -> glob_constr) -> glob_constr -> glob_constr
+
+(** Ensure traversal from left to right *)
+val map_glob_constr_left_to_right :
+ (glob_constr -> glob_constr) -> glob_constr -> glob_constr
+
+val fold_glob_constr : ('a -> glob_constr -> 'a) -> 'a -> glob_constr -> 'a
+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 loc_of_glob_constr : glob_constr -> Loc.t
+
+(** [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) ->
+ tomatch_tuples -> cases_clauses -> (tomatch_tuples*cases_clauses)
+
+(** [map_pattern f m c] applies [f] to the return predicate and the
+ right-hand side of a pattern-matching expression
+ ({!Glob_term.GCases}) represented here by its relevant components
+ [m] and [c]. *)
+val map_pattern : (glob_constr -> glob_constr) ->
+ tomatch_tuples -> cases_clauses -> (tomatch_tuples*cases_clauses)
+
+(** Conversion from glob_constr to cases pattern, if possible
+
+ 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 glob_constr_of_closed_cases_pattern : cases_pattern -> Name.t * glob_constr
diff --git a/pretyping/glob_term.ml b/pretyping/glob_term.ml
deleted file mode 100644
index b7023db0..00000000
--- a/pretyping/glob_term.ml
+++ /dev/null
@@ -1,418 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(*i*)
-open Util
-open Names
-open Sign
-open Term
-open Libnames
-open Nametab
-open Evd
-(*i*)
-
-(* Untyped intermediate terms, after ASTs and before constr. *)
-
-(* locs here refers to the ident's location, not whole pat *)
-(* the last argument of PatCstr is a possible alias ident for the pattern *)
-type cases_pattern =
- | PatVar of loc * name
- | PatCstr of loc * constructor * cases_pattern list * name
-
-let cases_pattern_loc = function
- PatVar(loc,_) -> loc
- | PatCstr(loc,_,_,_) -> loc
-
-type patvar = identifier
-
-type glob_sort = GProp of Term.contents | GType of Univ.universe option
-
-type binding_kind = Lib.binding_kind = Explicit | Implicit
-
-type quantified_hypothesis = AnonHyp of int | NamedHyp of identifier
-
-type 'a explicit_bindings = (loc * quantified_hypothesis * 'a) list
-
-type 'a bindings =
- | ImplicitBindings of 'a list
- | ExplicitBindings of 'a explicit_bindings
- | NoBindings
-
-type 'a with_bindings = 'a * 'a bindings
-
-type 'a cast_type =
- | CastConv of cast_kind * 'a
- | CastCoerce (* Cast to a base type (eg, an underlying inductive type) *)
-
-type glob_constr =
- | GRef of (loc * global_reference)
- | GVar of (loc * identifier)
- | GEvar of loc * existential_key * glob_constr list option
- | GPatVar of loc * (bool * patvar) (* Used for patterns only *)
- | GApp of loc * glob_constr * glob_constr list
- | GLambda of loc * name * binding_kind * glob_constr * glob_constr
- | GProd of loc * name * binding_kind * glob_constr * glob_constr
- | GLetIn of loc * name * glob_constr * glob_constr
- | GCases of loc * case_style * glob_constr option * tomatch_tuples * cases_clauses
- | GLetTuple of loc * name list * (name * glob_constr option) *
- glob_constr * glob_constr
- | GIf of loc * glob_constr * (name * glob_constr option) * glob_constr * glob_constr
- | GRec of loc * fix_kind * identifier array * glob_decl list array *
- glob_constr array * glob_constr array
- | GSort of loc * glob_sort
- | GHole of (loc * hole_kind)
- | GCast of loc * glob_constr * glob_constr cast_type
-
-and glob_decl = name * binding_kind * glob_constr option * glob_constr
-
-and fix_recursion_order = GStructRec | GWfRec of glob_constr | GMeasureRec of glob_constr * glob_constr option
-
-and fix_kind =
- | GFix of ((int option * fix_recursion_order) array * int)
- | GCoFix of int
-
-and predicate_pattern =
- name * (loc * inductive * int * name list) option
-
-and tomatch_tuple = (glob_constr * predicate_pattern)
-
-and tomatch_tuples = tomatch_tuple list
-
-and cases_clause = (loc * identifier list * cases_pattern list * glob_constr)
-
-and cases_clauses = cases_clause list
-
-let cases_predicate_names tml =
- List.flatten (List.map (function
- | (tm,(na,None)) -> [na]
- | (tm,(na,Some (_,_,_,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 map_glob_decl_left_to_right f (na,k,obd,ty) =
- let comp1 = Option.map f obd in
- let comp2 = f ty in
- (na,k,comp1,comp2)
-
-let map_glob_constr_left_to_right f = function
- | GApp (loc,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) ->
- let comp1 = f ty in
- let comp2 = f c in
- GLambda (loc,na,bk,comp1,comp2)
- | GProd (loc,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) ->
- let comp1 = f b in
- let comp2 = f c in
- GLetIn (loc,na,comp1,comp2)
- | GCases (loc,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 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) ->
- 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) ->
- 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) ->
- let comp1 = f c in
- let comp2 = match k with CastConv (k,t) -> CastConv (k, f t) | x -> x in
- GCast (loc,comp1,comp2)
- | (GVar _ | GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) as x -> x
-
-let map_glob_constr = map_glob_constr_left_to_right
-
-(*
-let name_app f e = function
- | Name id -> let (id, e) = f id e in (Name id, e)
- | Anonymous -> Anonymous, e
-
-let fold_ident g idl e =
- let (idl,e) =
- Array.fold_right
- (fun id (idl,e) -> let id,e = g id e in (id::idl,e)) idl ([],e)
- in (Array.of_list idl,e)
-
-let map_glob_constr_with_binders_loc loc g f e = function
- | GVar (_,id) -> GVar (loc,id)
- | GApp (_,a,args) -> GApp (loc,f e a, List.map (f e) args)
- | GLambda (_,na,ty,c) ->
- let na,e = name_app g e na in GLambda (loc,na,f e ty,f e c)
- | GProd (_,na,ty,c) ->
- let na,e = name_app g e na in GProd (loc,na,f e ty,f e c)
- | GLetIn (_,na,b,c) ->
- let na,e = name_app g e na in GLetIn (loc,na,f e b,f e c)
- | GCases (_,tyopt,tml,pl) ->
- (* We don't modify pattern variable since we don't traverse patterns *)
- let g' id e = snd (g id e) in
- let h (_,idl,p,c) = (loc,idl,p,f (List.fold_right g' idl e) c) in
- GCases
- (loc,Option.map (f e) tyopt,List.map (f e) tml, List.map h pl)
- | GRec (_,fk,idl,tyl,bv) ->
- let idl',e' = fold_ident g idl e in
- GRec (loc,fk,idl',Array.map (f e) tyl,Array.map (f e') bv)
- | GCast (_,c,t) -> GCast (loc,f e c,f e t)
- | GSort (_,x) -> GSort (loc,x)
- | GHole (_,x) -> GHole (loc,x)
- | GRef (_,x) -> GRef (loc,x)
- | GEvar (_,x,l) -> GEvar (loc,x,l)
- | GPatVar (_,x) -> GPatVar (loc,x)
-*)
-
-let fold_glob_constr f acc =
- let rec fold acc = function
- | GVar _ -> acc
- | GApp (_,c,args) -> List.fold_left fold (fold acc c) args
- | GLambda (_,_,_,b,c) | GProd (_,_,_,b,c) | GLetIn (_,_,b,c) ->
- fold (fold acc b) c
- | GCases (_,_,rtntypopt,tml,pl) ->
- List.fold_left fold_pattern
- (List.fold_left fold (Option.fold_left fold acc rtntypopt) (List.map fst tml))
- pl
- | GLetTuple (_,_,rtntyp,b,c) ->
- fold (fold (fold_return_type acc rtntyp) b) c
- | GIf (_,c,rtntyp,b1,b2) ->
- fold (fold (fold (fold_return_type acc rtntyp) c) b1) b2
- | GRec (_,_,_,bl,tyl,bv) ->
- let acc = Array.fold_left
- (List.fold_left (fun acc (na,k,bbd,bty) ->
- fold (Option.fold_left fold acc bbd) bty)) acc bl in
- Array.fold_left fold (Array.fold_left fold acc tyl) bv
- | GCast (_,c,k) -> fold (match k with CastConv (_, t) -> fold acc t | CastCoerce -> acc) c
- | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> acc
-
- and fold_pattern acc (_,idl,p,c) = fold acc c
-
- and fold_return_type acc (na,tyopt) = Option.fold_left fold acc tyopt
-
- in fold acc
-
-let iter_glob_constr f = fold_glob_constr (fun () -> f) ()
-
-let occur_glob_constr id =
- let rec occur = function
- | GVar (loc,id') -> id = id'
- | GApp (loc,f,args) -> (occur f) or (List.exists occur args)
- | GLambda (loc,na,bk,ty,c) -> (occur ty) or ((na <> Name id) & (occur c))
- | GProd (loc,na,bk,ty,c) -> (occur ty) or ((na <> Name id) & (occur c))
- | GLetIn (loc,na,b,c) -> (occur b) or ((na <> Name id) & (occur c))
- | GCases (loc,sty,rtntypopt,tml,pl) ->
- (occur_option rtntypopt)
- or (List.exists (fun (tm,_) -> occur tm) tml)
- or (List.exists occur_pattern pl)
- | GLetTuple (loc,nal,rtntyp,b,c) ->
- occur_return_type rtntyp id
- or (occur b) or (not (List.mem (Name id) nal) & (occur c))
- | GIf (loc,c,rtntyp,b1,b2) ->
- occur_return_type rtntyp id or (occur c) or (occur b1) or (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) && (fid=id or not(occur bd))
- | (na,k,bbd,bty)::bl ->
- not (occur bty) &&
- (match bbd with
- Some bd -> not (occur bd)
- | _ -> true) &&
- (na=Name id or not(occur_fix bl)) in
- occur_fix bl)
- idl bl tyl bv)
- | GCast (loc,c,k) -> (occur c) or (match k with CastConv (_, t) -> occur t | CastCoerce -> false)
- | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> false
-
- and occur_pattern (loc,idl,p,c) = not (List.mem id idl) & (occur c)
-
- and occur_option = function None -> false | Some p -> occur p
-
- and occur_return_type (na,tyopt) id = na <> Name id & occur_option tyopt
-
- in occur
-
-
-let add_name_to_ids set na =
- match na with
- | Anonymous -> set
- | Name id -> Idset.add id set
-
-let free_glob_vars =
- let rec vars bounded vs = function
- | GVar (loc,id') -> if Idset.mem id' bounded then vs else Idset.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 Idset.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) -> 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 Idset.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
- fun rt ->
- let vs = vars Idset.empty Idset.empty rt in
- Idset.elements vs
-
-
-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
-
-(**********************************************************************)
-(* Conversion from glob_constr to cases pattern, if possible *)
-
-let rec cases_pattern_of_glob_constr na = function
- | GVar (loc,id) when na<>Anonymous ->
- (* Unable to manage the presence of both an alias and a variable *)
- raise Not_found
- | GVar (loc,id) -> PatVar (loc,Name id)
- | GHole (loc,_) -> PatVar (loc,na)
- | GRef (loc,ConstructRef cstr) ->
- PatCstr (loc,cstr,[],na)
- | GApp (loc,GRef (_,ConstructRef (ind,_ as cstr)),args) ->
- let mib,_ = Global.lookup_inductive ind in
- let nparams = mib.Declarations.mind_nparams in
- if nparams > List.length args then
- user_err_loc (loc,"",Pp.str "Invalid notation for pattern.");
- let params,args = list_chop nparams args in
- List.iter (function GHole _ -> ()
- | _ -> user_err_loc (loc,"",Pp.str"Invalid notation for pattern."))
- params;
- let args = List.map (cases_pattern_of_glob_constr Anonymous) args in
- PatCstr (loc,cstr,args,na)
- | _ -> raise Not_found
-
-(* 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)
- | PatCstr (loc,cstr,l,Anonymous) ->
- let ref = GRef (loc,ConstructRef cstr) in
- GApp (loc,ref, List.map glob_constr_of_closed_cases_pattern_aux l)
- | _ -> raise Not_found
-
-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))
- | _ ->
- raise Not_found
-
-(**********************************************************************)
-(* Reduction expressions *)
-
-type 'a glob_red_flag = {
- rBeta : bool;
- rIota : bool;
- rZeta : bool;
- rDelta : bool; (* true = delta all but rConst; false = delta only on rConst*)
- rConst : 'a list
-}
-
-let all_flags =
- {rBeta = true; rIota = true; rZeta = true; rDelta = true; rConst = []}
-
-type 'a or_var = ArgArg of 'a | ArgVar of identifier located
-
-type occurrences_expr = bool * int or_var list
-
-let all_occurrences_expr_but l = (false,l)
-let no_occurrences_expr_but l = (true,l)
-let all_occurrences_expr = (false,[])
-let no_occurrences_expr = (true,[])
-
-type 'a with_occurrences = occurrences_expr * 'a
-
-type ('a,'b,'c) red_expr_gen =
- | Red of bool
- | Hnf
- | Simpl of 'c with_occurrences option
- | Cbv of 'b glob_red_flag
- | Lazy of 'b glob_red_flag
- | Unfold of 'b with_occurrences list
- | Fold of 'a list
- | Pattern of 'a with_occurrences list
- | ExtraRedExpr of string
- | CbvVm
-
-type ('a,'b,'c) may_eval =
- | ConstrTerm of 'a
- | ConstrEval of ('a,'b,'c) red_expr_gen * 'a
- | ConstrContext of (loc * identifier) * 'a
- | ConstrTypeOf of 'a
diff --git a/pretyping/glob_term.mli b/pretyping/glob_term.mli
deleted file mode 100644
index 09dd9203..00000000
--- a/pretyping/glob_term.mli
+++ /dev/null
@@ -1,167 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(**Untyped intermediate terms, after constr_expr and before constr
-
- Resolution of names, insertion of implicit arguments placeholder,
- and notations are done, but coercions, inference of implicit
- arguments and pattern-matching compilation are not. *)
-
-open Util
-open Names
-open Sign
-open Term
-open Libnames
-open Nametab
-
-(** The kind of patterns that occurs in "match ... with ... end"
-
- locs here refers to the ident's location, not whole pat *)
-type cases_pattern =
- | PatVar of loc * name
- | PatCstr of loc * constructor * cases_pattern list * name
- (** [PatCstr(p,C,l,x)] = "|'C' 'l' as 'x'" *)
-
-val cases_pattern_loc : cases_pattern -> loc
-
-type patvar = identifier
-
-type glob_sort = GProp of Term.contents | GType of Univ.universe option
-
-type binding_kind = Lib.binding_kind = Explicit | Implicit
-
-type quantified_hypothesis = AnonHyp of int | NamedHyp of identifier
-
-type 'a explicit_bindings = (loc * quantified_hypothesis * 'a) list
-
-type 'a bindings =
- | ImplicitBindings of 'a list
- | ExplicitBindings of 'a explicit_bindings
- | NoBindings
-
-type 'a with_bindings = 'a * 'a bindings
-
-type 'a cast_type =
- | CastConv of cast_kind * 'a
- | CastCoerce (** Cast to a base type (eg, an underlying inductive type) *)
-
-type glob_constr =
- | GRef of (loc * global_reference)
- | GVar of (loc * identifier)
- | GEvar of loc * existential_key * glob_constr list option
- | GPatVar of loc * (bool * patvar) (** Used for patterns only *)
- | GApp of loc * glob_constr * glob_constr list
- | GLambda of loc * name * binding_kind * glob_constr * glob_constr
- | GProd of loc * name * binding_kind * glob_constr * glob_constr
- | GLetIn of loc * name * glob_constr * glob_constr
- | GCases of loc * case_style * glob_constr option * tomatch_tuples * cases_clauses
- (** [GCases(l,style,r,tur,cc)] = "match 'tur' return 'r' with 'cc'" (in
- [MatchStyle]) *)
-
- | GLetTuple of loc * name list * (name * glob_constr option) *
- glob_constr * glob_constr
- | GIf of loc * glob_constr * (name * glob_constr option) * glob_constr * glob_constr
- | GRec of loc * fix_kind * identifier array * glob_decl list array *
- glob_constr array * glob_constr array
- | GSort of loc * glob_sort
- | GHole of (loc * Evd.hole_kind)
- | GCast of loc * glob_constr * glob_constr cast_type
-
-and glob_decl = name * binding_kind * glob_constr option * glob_constr
-
-and fix_recursion_order = GStructRec | GWfRec of glob_constr | GMeasureRec of glob_constr * glob_constr option
-
-and fix_kind =
- | GFix of ((int option * fix_recursion_order) array * int)
- | GCoFix of int
-
-and predicate_pattern =
- name * (loc * inductive * int * name list) option
- (** [(na,id)] = "as 'na' in 'id'" where if [id] is [Some(l,I,k,args)], [k]
- is the number of parameter of [I]. *)
-
-and tomatch_tuple = (glob_constr * predicate_pattern)
-
-and tomatch_tuples = tomatch_tuple list
-
-and cases_clause = (loc * identifier list * cases_pattern list * glob_constr)
-(** [(p,il,cl,t)] = "|'cl' as 'il' => 't'" *)
-
-and cases_clauses = cases_clause list
-
-val cases_predicate_names : tomatch_tuples -> name list
-
-(* Apply one argument to a glob_constr *)
-val mkGApp : loc -> glob_constr -> glob_constr -> glob_constr
-
-val map_glob_constr : (glob_constr -> glob_constr) -> glob_constr -> glob_constr
-
-(* Ensure traversal from left to right *)
-val map_glob_constr_left_to_right :
- (glob_constr -> glob_constr) -> glob_constr -> glob_constr
-
-(*
-val map_glob_constr_with_binders_loc : loc ->
- (identifier -> 'a -> identifier * 'a) ->
- ('a -> glob_constr -> glob_constr) -> 'a -> glob_constr -> glob_constr
-*)
-
-val fold_glob_constr : ('a -> glob_constr -> 'a) -> 'a -> glob_constr -> 'a
-val iter_glob_constr : (glob_constr -> unit) -> glob_constr -> unit
-val occur_glob_constr : identifier -> glob_constr -> bool
-val free_glob_vars : glob_constr -> identifier list
-val loc_of_glob_constr : glob_constr -> loc
-
-(** Conversion from glob_constr to cases pattern, if possible
-
- Take the current alias as parameter,
- @raise Not_found if translation is impossible *)
-val cases_pattern_of_glob_constr : name -> glob_constr -> cases_pattern
-
-val glob_constr_of_closed_cases_pattern : cases_pattern -> name * glob_constr
-
-(** {6 Reduction expressions} *)
-
-type 'a glob_red_flag = {
- rBeta : bool;
- rIota : bool;
- rZeta : bool;
- rDelta : bool; (** true = delta all but rConst; false = delta only on rConst*)
- rConst : 'a list
-}
-
-val all_flags : 'a glob_red_flag
-
-type 'a or_var = ArgArg of 'a | ArgVar of identifier located
-
-type occurrences_expr = bool * int or_var list
-
-val all_occurrences_expr_but : int or_var list -> occurrences_expr
-val no_occurrences_expr_but : int or_var list -> occurrences_expr
-val all_occurrences_expr : occurrences_expr
-val no_occurrences_expr : occurrences_expr
-
-type 'a with_occurrences = occurrences_expr * 'a
-
-type ('a,'b,'c) red_expr_gen =
- | Red of bool
- | Hnf
- | Simpl of 'c with_occurrences option
- | Cbv of 'b glob_red_flag
- | Lazy of 'b glob_red_flag
- | Unfold of 'b with_occurrences list
- | Fold of 'a list
- | Pattern of 'a with_occurrences list
- | ExtraRedExpr of string
- | CbvVm
-
-type ('a,'b,'c) may_eval =
- | ConstrTerm of 'a
- | ConstrEval of ('a,'b,'c) red_expr_gen * 'a
- | ConstrContext of (loc * identifier) * 'a
- | ConstrTypeOf of 'a
diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml
index c7cdb302..54d47fbe 100644
--- a/pretyping/indrec.ml
+++ b/pretyping/indrec.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -11,66 +11,75 @@
(* This file builds various inductive schemes *)
open Pp
+open Errors
open Util
open Names
open Libnames
+open Globnames
open Nameops
open Term
+open Vars
+open Context
open Namegen
open Declarations
-open Entries
+open Declareops
open Inductive
open Inductiveops
open Environ
open Reductionops
-open Typeops
-open Type_errors
-open Safe_typing
open Nametab
-open Sign
type dep_flag = bool
(* Errors related to recursors building *)
type recursion_scheme_error =
- | NotAllowedCaseAnalysis of (*isrec:*) bool * sorts * inductive
+ | NotAllowedCaseAnalysis of (*isrec:*) bool * sorts * pinductive
| NotMutualInScheme of inductive * inductive
exception RecursionSchemeError of recursion_scheme_error
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)
+let mkLambda_string s t c = mkLambda (Name (Id.of_string s), t, c)
(*******************************************)
(* Building curryfied elimination *)
(*******************************************)
+let is_private mib =
+ match mib.mind_private with
+ | Some true -> true
+ | _ -> false
+
+let check_privacy_block mib =
+ if is_private mib then
+ errorlabstrm ""(str"case analysis on a private inductive type")
+
(**********************************************************************)
(* Building case analysis schemes *)
(* Christine Paulin, 1996 *)
-let mis_make_case_com dep env sigma ind (mib,mip as specif) kind =
- let lnamespar = List.map
- (fun (n, c, t) -> (n, c, Termops.refresh_universes t))
- mib.mind_params_ctxt
+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, Termops.extended_rel_list 0 lnamespar) in
+ let constrs = get_constructors env indf in
+ let projs = get_projections env indf in
+
+ let () = if Option.is_empty projs then check_privacy_block mib in
+ let () =
+ if not (Sorts.List.mem kind (elim_sorts specif)) then
+ raise
+ (RecursionSchemeError
+ (NotAllowedCaseAnalysis (false, fst (Universes.fresh_sort_in_family env kind), pind)))
in
+ let ndepar = mip.mind_nrealdecls + 1 in
- if not (List.mem kind (elim_sorts specif)) then
- raise
- (RecursionSchemeError
- (NotAllowedCaseAnalysis (false, Termops.new_sort_in_family kind, ind)));
-
- let ndepar = mip.mind_nrealargs_ctxt + 1 in
-
- (* Pas génant car env ne sert pas à typer mais juste à renommer les Anonym *)
- (* mais pas très joli ... (mais manque get_sort_of à ce niveau) *)
+ (* Pas génant car env ne sert pas à typer mais juste à renommer les Anonym *)
+ (* mais pas très joli ... (mais manque get_sort_of à ce niveau) *)
let env' = push_rel_context lnamespar env in
- let indf = make_ind_family(ind, Termops.extended_rel_list 0 lnamespar) in
- let constrs = get_constructors env indf in
let rec add_branch env k =
- if k = Array.length mip.mind_consnames then
+ if Int.equal k (Array.length mip.mind_consnames) then
let nbprod = k+1 in
let indf' = lift_inductive_family nbprod indf in
@@ -78,7 +87,7 @@ let mis_make_case_com dep env sigma ind (mib,mip as specif) kind =
let depind = build_dependent_inductive env indf' in
let deparsign = (Anonymous,None,depind)::arsign in
- let ci = make_case_info env ind RegularStyle in
+ let ci = make_case_info env (fst pind) RegularStyle in
let pbody =
appvect
(mkRel (ndepar + nbprod),
@@ -90,21 +99,34 @@ let mis_make_case_com dep env sigma ind (mib,mip as specif) kind =
(Anonymous,depind,pbody))
arsign
in
- it_mkLambda_or_LetIn_name env'
- (mkCase (ci, lift ndepar p,
- mkRel 1,
- Termops.rel_vect ndepar k))
- deparsign
+ let obj =
+ match projs with
+ | None -> mkCase (ci, lift ndepar p, mkRel 1,
+ Termops.rel_vect ndepar k)
+ | Some ps ->
+ let term =
+ mkApp (mkRel 2,
+ Array.map
+ (fun p -> mkProj (Projection.make p true, mkRel 1)) ps) in
+ if dep then
+ let ty = mkApp (mkRel 3, [| mkRel 1 |]) in
+ mkCast (term, DEFAULTcast, ty)
+ else term
+ in
+ 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
mkLambda_string "f" t
(add_branch (push_rel (Anonymous, None, t) env) (k+1))
in
- let typP = make_arity env' dep indf (Termops.new_sort_in_family kind) in
- it_mkLambda_or_LetIn_name env
+ let sigma, s = Evd.fresh_sort_in_family ~rigid:Evd.univ_flexible_alg env sigma kind in
+ let typP = make_arity env' dep indf s in
+ let c =
+ it_mkLambda_or_LetIn_name env
(mkLambda_string "P" typP
- (add_branch (push_rel (Anonymous,None,typP) env') 0)) lnamespar
+ (add_branch (push_rel (Anonymous,None,typP) env') 0)) lnamespar
+ in sigma, c
(* check if the type depends recursively on one of the inductive scheme *)
@@ -137,7 +159,7 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
let d = (n,Some b,t) in
mkLetIn (n,b,t,prec (push_rel d env) (i+1) (d::sign) c)
| Ind (_,_) ->
- let realargs = list_skipn nparams largs in
+ let realargs = List.skipn nparams largs in
let base = applist (lift i pk,realargs) in
if depK then
Reduction.beta_appvect
@@ -158,7 +180,7 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
(match dest_recarg ra with
| Mrec (_,j) when is_rec -> (depPvect.(j),rest)
| Imbr _ ->
- Flags.if_warn msg_warning (str "Ignoring recursive call");
+ msg_warning (strbrk "Ignoring recursive call");
(None,rest)
| _ -> (None, rest))
in
@@ -172,7 +194,7 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
let nP = lift (i+1+decP) p in
let env' = push_rel (n,None,t) env in
let t_0 = process_pos env' dep' nP (lift 1 t) in
- make_prod_dep (dep or dep') env
+ make_prod_dep (dep || dep') env
(n,t,
mkArrow t_0
(process_constr
@@ -186,9 +208,9 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
| _ -> assert false
else
if dep then
- let realargs = List.map (fun k -> mkRel (i-k)) (List.rev li) in
+ let realargs = List.rev_map (fun k -> mkRel (i-k)) li in
let params = List.map (lift i) vargs in
- let co = applist (mkConstruct cs.cs_cstr,params@realargs) in
+ let co = applist (mkConstructU cs.cs_cstr,params@realargs) in
Reduction.beta_appvect c [|co|]
else c
in
@@ -212,14 +234,14 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
let d = (n,Some b,t) in
mkLetIn (n,b,t,prec (push_rel d env) (i+1) (d::hyps) c)
| Ind _ ->
- let realargs = list_skipn nparrec largs
+ let realargs = List.skipn nparrec largs
and arg = appvect (mkRel (i+1), Termops.extended_rel_vect 0 hyps) in
applist(lift i fk,realargs@[arg])
| _ -> assert false
in
prec env 0 []
in
- (* ici, cstrprods est la liste des produits du constructeur instantié *)
+ (* ici, cstrprods est la liste des produits du constructeur instantié *)
let rec process_constr env i f = function
| (n,None,t as d)::cprest, recarg::rest ->
let optionpos =
@@ -248,7 +270,7 @@ 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 "process_constr"
+ | _,[] | [],_ -> anomaly (Pp.str "process_constr")
in
process_constr env 0 f (List.rev cstr.cs_args, recargs)
@@ -264,21 +286,21 @@ let context_chop k ctx =
| (_, []) -> failwith "context_chop"
in chop_aux [] (k,ctx)
-
(* Main function *)
-let mis_make_indrec env sigma listdepkind mib =
+let mis_make_indrec env sigma listdepkind mib u =
let nparams = mib.mind_nparams in
- let nparrec = mib. mind_nparams_rec in
+ let nparrec = mib.mind_nparams_rec in
+ let evdref = ref sigma in
let lnonparrec,lnamesparrec =
- context_chop (nparams-nparrec) mib.mind_params_ctxt in
+ context_chop (nparams-nparrec) (Vars.subst_instance_context u mib.mind_params_ctxt) in
let nrec = List.length listdepkind in
let depPvec =
- Array.create mib.mind_ntypes (None : (bool * constr) option) in
+ Array.make mib.mind_ntypes (None : (bool * constr) option) in
let _ =
let rec
assign k = function
| [] -> ()
- | (indi,mibi,mipi,dep,_)::rest ->
+ | ((indi,u),mibi,mipi,dep,_)::rest ->
(Array.set depPvec (snd indi) (Some(dep,mkRel k));
assign (k-1) rest)
in
@@ -287,12 +309,12 @@ let mis_make_indrec env sigma listdepkind mib =
Array.map (fun mip -> mip.mind_recargs) mib.mind_packets in
(* recarg information for non recursive parameters *)
let rec recargparn l n =
- if n = 0 then l else recargparn (mk_norec::l) (n-1) in
+ if Int.equal n 0 then l else recargparn (mk_norec::l) (n-1) in
let recargpar = recargparn [] (nparams-nparrec) in
let make_one_rec p =
let makefix nbconstruct =
let rec mrec i ln ltyp ldef = function
- | (indi,mibi,mipi,dep,_)::rest ->
+ | ((indi,u),mibi,mipi,dep,_)::rest ->
let tyi = snd indi in
let nctyi =
Array.length mipi.mind_consnames in (* nb constructeurs du type*)
@@ -300,7 +322,7 @@ let mis_make_indrec env sigma listdepkind mib =
(* arity in the context of the fixpoint, i.e.
P1..P_nrec f1..f_nbconstruct *)
let args = Termops.extended_rel_list (nrec+nbconstruct) lnamesparrec in
- let indf = make_ind_family(indi,args) in
+ let indf = make_ind_family((indi,u),args) in
let arsign,_ = get_arity env indf in
let depind = build_dependent_inductive env indf in
@@ -315,7 +337,7 @@ let mis_make_indrec env sigma listdepkind mib =
P1..P_nrec f1..f_nbconstruct F_1..F_nrec a_1..a_nar x:I *)
let args' = Termops.extended_rel_list (dect+nrec) lnamesparrec in
let args'' = Termops.extended_rel_list ndepar lnonparrec in
- let indf' = make_ind_family(indi,args'@args'') in
+ let indf' = make_ind_family((indi,u),args'@args'') in
let branches =
let constrs = get_constructors env indf' in
@@ -324,8 +346,8 @@ let mis_make_indrec env sigma listdepkind mib =
(fun f -> appvect (f, Termops.extended_rel_vect ndepar lnonparrec))
fi
in
- array_map3
- (make_rec_branch_arg env sigma
+ Array.map3
+ (make_rec_branch_arg env !evdref
(nparrec,depPvec,larsign))
vecfi constrs (dest_subterms recargsvec.(tyi))
in
@@ -359,10 +381,27 @@ let mis_make_indrec env sigma listdepkind mib =
(Anonymous,depind',concl))
arsign'
in
- it_mkLambda_or_LetIn_name env
- (mkCase (ci, pred,
- mkRel 1,
- branches))
+ let obj =
+ let projs = get_projections env indf in
+ match projs with
+ | None -> (mkCase (ci, pred,
+ mkRel 1,
+ branches))
+ | Some ps ->
+ let branch = branches.(0) in
+ let ctx, br = decompose_lam_assum branch in
+ let n, subst =
+ List.fold_right (fun (na,b,t) (i, subst) ->
+ if b == None then
+ let t = mkProj (Projection.make ps.(i) true, mkRel 1) in
+ (i + 1, t :: subst)
+ else (i, mkRel 0 :: subst))
+ ctx (0, [])
+ in
+ let term = substl subst br in
+ term
+ in
+ it_mkLambda_or_LetIn_name env obj
(Termops.lift_rel_context nrec deparsign)
in
@@ -383,26 +422,26 @@ let mis_make_indrec env sigma listdepkind mib =
let fixn = Array.of_list (List.rev ln) in
let fixtyi = Array.of_list (List.rev ltyp) in
let fixdef = Array.of_list (List.rev ldef) in
- let names = Array.create nrec (Name(id_of_string "F")) in
+ let names = Array.make nrec (Name(Id.of_string "F")) in
mkFix ((fixn,p),(names,fixtyi,fixdef))
in
mrec 0 [] [] []
in
let rec make_branch env i = function
- | (indi,mibi,mipi,dep,_)::rest ->
+ | ((indi,u),mibi,mipi,dep,_)::rest ->
let tyi = snd indi in
let nconstr = Array.length mipi.mind_consnames in
let rec onerec env j =
- if j = nconstr then
+ if Int.equal j nconstr then
make_branch env (i+j) rest
else
let recarg = (dest_subterms recargsvec.(tyi)).(j) in
let recarg = recargpar@recarg in
let vargs = Termops.extended_rel_list (nrec+i+j) lnamesparrec in
- let cs = get_constructor (indi,mibi,mipi,vargs) (j+1) in
+ let cs = get_constructor ((indi,u),mibi,mipi,vargs) (j+1) in
let p_0 =
type_rec_branch
- true dep env sigma (vargs,depPvec,i+j) tyi cs recarg
+ true dep env !evdref (vargs,depPvec,i+j) tyi cs recarg
in
mkLambda_string "f" p_0
(onerec (push_rel (Anonymous,None,p_0) env) (j+1))
@@ -411,9 +450,13 @@ let mis_make_indrec env sigma listdepkind mib =
makefix i listdepkind
in
let rec put_arity env i = function
- | (indi,_,_,dep,kinds)::rest ->
- let indf = make_ind_family (indi, Termops.extended_rel_list i lnamesparrec) in
- let typP = make_arity env dep indf (Termops.new_sort_in_family kinds) in
+ | ((indi,u),_,_,dep,kinds)::rest ->
+ let indf = make_ind_family ((indi,u), Termops.extended_rel_list 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
mkLambda_string "P" typP
(put_arity (push_rel (Anonymous,None,typP) env) (i+1) rest)
| [] ->
@@ -421,33 +464,38 @@ let mis_make_indrec env sigma listdepkind mib =
in
(* Body on make_one_rec *)
- let (indi,mibi,mipi,dep,kind) = List.nth listdepkind p in
+ let ((indi,u),mibi,mipi,dep,kind) = List.nth listdepkind p in
if (mis_is_recursive_subset
- (List.map (fun (indi,_,_,_,_) -> snd indi) listdepkind)
+ (List.map (fun ((indi,u),_,_,_,_) -> snd indi) listdepkind)
mipi.mind_recargs)
then
let env' = push_rel_context lnamesparrec env in
it_mkLambda_or_LetIn_name env (put_arity env' 0 listdepkind)
lnamesparrec
else
- mis_make_case_com dep env sigma indi (mibi,mipi) kind
+ let evd', c = mis_make_case_com dep env !evdref (indi,u) (mibi,mipi) kind in
+ evdref := evd'; c
in
(* Body of mis_make_indrec *)
- list_tabulate make_one_rec nrec
+ !evdref, List.init nrec make_one_rec
(**********************************************************************)
(* This builds elimination predicate for Case tactic *)
-let build_case_analysis_scheme env sigma ity dep kind =
- let (mib,mip) = lookup_mind_specif env ity in
- mis_make_case_com dep env sigma ity (mib,mip) kind
+let build_case_analysis_scheme env sigma pity dep kind =
+ let (mib,mip) = lookup_mind_specif env (fst pity) in
+ mis_make_case_com dep env sigma pity (mib,mip) kind
-let build_case_analysis_scheme_default env sigma ity kind =
- let (mib,mip) = lookup_mind_specif env ity in
- let dep = inductive_sort_family mip <> InProp in
- mis_make_case_com dep env sigma ity (mib,mip) kind
+let is_in_prop mip =
+ match inductive_sort_family mip with
+ | InProp -> true
+ | _ -> false
+let build_case_analysis_scheme_default env sigma pity kind =
+ let (mib,mip) = lookup_mind_specif env (fst pity) in
+ let dep = not (is_in_prop mip) in
+ mis_make_case_com dep env sigma pity (mib,mip) kind
(**********************************************************************)
(* [modify_sort_scheme s rec] replaces the sort of the scheme
@@ -456,87 +504,78 @@ let build_case_analysis_scheme_default env sigma ity kind =
let change_sort_arity sort =
let rec drec a = match kind_of_term a with
| Cast (c,_,_) -> drec c
- | Prod (n,t,c) -> mkProd (n, t, drec c)
- | LetIn (n,b,t,c) -> mkLetIn (n,b, t, drec c)
- | Sort _ -> mkSort sort
+ | 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')
+ | Sort s -> s, mkSort sort
| _ -> assert false
in
drec
-(* [npar] is the number of expected arguments (then excluding letin's) *)
-let modify_sort_scheme sort =
- let rec drec npar elim =
- match kind_of_term elim with
- | Lambda (n,t,c) ->
- if npar = 0 then
- mkLambda (n, change_sort_arity sort t, c)
- else
- mkLambda (n, t, drec (npar-1) c)
- | LetIn (n,b,t,c) -> mkLetIn (n,b,t,drec npar c)
- | _ -> anomaly "modify_sort_scheme: wrong elimination type"
- in
- drec
-
(* Change the sort in the type of an inductive definition, builds the
corresponding eta-expanded term *)
-let weaken_sort_scheme sort npars term =
+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
| Prod (n,t,c) ->
- if np = 0 then
- let t' = change_sort_arity sort t in
- mkProd (n, t', c),
- mkLambda (n, t', mkApp(term,Termops.rel_vect 0 (npars+1)))
+ if Int.equal np 0 then
+ let osort, t' = change_sort_arity sort t in
+ evdref := (if set then Evd.set_eq_sort else Evd.set_leq_sort) env !evdref sort osort;
+ mkProd (n, t', c),
+ mkLambda (n, t', mkApp(term,Termops.rel_vect 0 (npars+1)))
else
let c',term' = drec (np-1) c in
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 "weaken_sort_scheme: wrong elimination type"
+ | _ -> anomaly ~label:"weaken_sort_scheme" (Pp.str "wrong elimination type")
in
- drec npars
+ let ty, term = drec npars ty in
+ !evdref, ty, term
(**********************************************************************)
(* Interface to build complex Scheme *)
(* Check inductive types only occurs once
(otherwise we obtain a meaning less scheme) *)
-let check_arities listdepkind =
+let check_arities env listdepkind =
let _ = List.fold_left
- (fun ln ((_,ni as mind),mibi,mipi,dep,kind) ->
+ (fun ln (((_,ni as mind),u),mibi,mipi,dep,kind) ->
let kelim = elim_sorts (mibi,mipi) in
- if not (List.exists ((=) kind) kelim) then raise
+ if not (Sorts.List.mem kind kelim) then raise
(RecursionSchemeError
- (NotAllowedCaseAnalysis (true, Termops.new_sort_in_family kind,mind)))
- else if List.mem ni ln then raise
+ (NotAllowedCaseAnalysis (true, fst (Universes.fresh_sort_in_family env
+ kind),(mind,u))))
+ else if Int.List.mem ni ln then raise
(RecursionSchemeError (NotMutualInScheme (mind,mind)))
else ni::ln)
[] listdepkind
in true
let build_mutual_induction_scheme env sigma = function
- | (mind,dep,s)::lrecspec ->
- let (mib,mip) = Global.lookup_inductive mind in
+ | ((mind,u),dep,s)::lrecspec ->
+ let (mib,mip) = lookup_mind_specif env mind in
let (sp,tyi) = mind in
let listdepkind =
- (mind,mib,mip,dep,s)::
+ ((mind,u),mib,mip,dep,s)::
(List.map
- (function (mind',dep',s') ->
+ (function ((mind',u'),dep',s') ->
let (sp',_) = mind' in
- if sp=sp' then
+ if eq_mind sp sp' then
let (mibi',mipi') = lookup_mind_specif env mind' in
- (mind',mibi',mipi',dep',s')
+ ((mind',u'),mibi',mipi',dep',s')
else
raise (RecursionSchemeError (NotMutualInScheme (mind,mind'))))
lrecspec)
in
- let _ = check_arities listdepkind in
- mis_make_indrec env sigma listdepkind mib
- | _ -> anomaly "build_induction_scheme expects a non empty list of inductive types"
+ 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")
-let build_induction_scheme env sigma ind dep kind =
- let (mib,mip) = lookup_mind_specif env ind in
- List.hd (mis_make_indrec env sigma [(ind,mib,mip,dep,kind)] mib)
+let build_induction_scheme env sigma pind dep kind =
+ let (mib,mip) = lookup_mind_specif env (fst pind) in
+ let sigma, l = mis_make_indrec env sigma [(pind,mib,mip,dep,kind)] mib (snd pind) in
+ sigma, List.hd l
(*s Eliminations. *)
@@ -559,17 +598,17 @@ 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 (make_kn mp dp (Label.of_id id)) in
let _ = Global.lookup_constant cst in
- mkConst cst
+ ConstRef cst
with Not_found ->
(* Then try to get a user-defined eliminator in some other places *)
(* using short name (e.g. for "eq_rec") *)
- try constr_of_global (Nametab.locate (qualid_of_ident id))
+ try Nametab.locate (qualid_of_ident id)
with Not_found ->
errorlabstrm "default_elim"
(strbrk "Cannot find the elimination combinator " ++
pr_id id ++ strbrk ", the elimination of the inductive definition " ++
- pr_global_env Idset.empty (IndRef ind_sp) ++
+ 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 19c70dc0..f616c967 100644
--- a/pretyping/indrec.mli
+++ b/pretyping/indrec.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,15 +8,13 @@
open Names
open Term
-open Declarations
-open Inductiveops
open Environ
open Evd
(** Errors related to recursors building *)
type recursion_scheme_error =
- | NotAllowedCaseAnalysis of (*isrec:*) bool * sorts * inductive
+ | NotAllowedCaseAnalysis of (*isrec:*) bool * sorts * pinductive
| NotMutualInScheme of inductive * inductive
exception RecursionSchemeError of recursion_scheme_error
@@ -27,42 +25,39 @@ type dep_flag = bool
(** Build a case analysis elimination scheme in some sort family *)
-val build_case_analysis_scheme : env -> evar_map -> inductive ->
- dep_flag -> sorts_family -> constr
+val build_case_analysis_scheme : env -> evar_map -> pinductive ->
+ dep_flag -> sorts_family -> evar_map * constr
(** Build a dependent case elimination predicate unless type is in Prop *)
-val build_case_analysis_scheme_default : env -> evar_map -> inductive ->
- sorts_family -> constr
+val build_case_analysis_scheme_default : env -> evar_map -> pinductive ->
+ sorts_family -> evar_map * constr
(** Builds a recursive induction scheme (Peano-induction style) in the same
sort family as the inductive family; it is dependent if not in Prop *)
-val build_induction_scheme : env -> evar_map -> inductive ->
- dep_flag -> sorts_family -> constr
+val build_induction_scheme : env -> evar_map -> pinductive ->
+ dep_flag -> sorts_family -> evar_map * constr
(** Builds mutual (recursive) induction schemes *)
val build_mutual_induction_scheme :
- env -> evar_map -> (inductive * dep_flag * sorts_family) list -> constr list
+ env -> evar_map -> (pinductive * dep_flag * sorts_family) list -> evar_map * constr list
(** Scheme combinators *)
-(** [modify_sort_scheme s n c] modifies the quantification sort of
- scheme c whose predicate is abstracted at position [n] of [c] *)
+(** [weaken_sort_scheme env sigma eq s n c t] derives by subtyping from [c:t]
+ whose conclusion is quantified on [Type i] at position [n] of [t] a
+ scheme quantified on sort [s]. [set] asks for [s] be declared equal to [i],
+ otherwise just less or equal to [i]. *)
-val modify_sort_scheme : sorts -> int -> constr -> constr
-
-(** [weaken_sort_scheme s n c t] derives by subtyping from [c:t]
- whose conclusion is quantified on [Type] at position [n] of [t] a
- scheme quantified on sort [s] *)
-
-val weaken_sort_scheme : sorts -> int -> constr -> types -> constr * types
+val weaken_sort_scheme : env -> evar_map -> bool -> sorts -> int -> constr -> types ->
+ evar_map * types * constr
(** Recursor names utilities *)
-val lookup_eliminator : inductive -> sorts_family -> constr
+val lookup_eliminator : inductive -> sorts_family -> Globnames.global_reference
val elimination_suffix : sorts_family -> string
-val make_elimination_ident : identifier -> sorts_family -> identifier
+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 5bfc57bf..654f914b 100644
--- a/pretyping/inductiveops.ml
+++ b/pretyping/inductiveops.ml
@@ -1,47 +1,52 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Errors
open Util
open Names
open Univ
open Term
+open Vars
+open Context
open Termops
-open Namegen
-open Sign
open Declarations
+open Declareops
open Environ
open Reductionops
+open Inductive
(* The following three functions are similar to the ones defined in
Inductive, but they expect an env *)
-let type_of_inductive env ind =
- let specif = Inductive.lookup_mind_specif env ind in
- Inductive.type_of_inductive env specif
+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;
+ Inductive.type_of_inductive env (specif,u)
(* Return type as quoted by the user *)
-let type_of_constructor env cstr =
- let specif =
- Inductive.lookup_mind_specif env (inductive_of_constructor cstr) in
- Inductive.type_of_constructor cstr specif
+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;
+ Inductive.type_of_constructor (cstr,u) specif
(* Return constructor types in user form *)
-let type_of_constructors env ind =
+let type_of_constructors env (ind,u as indu) =
let specif = Inductive.lookup_mind_specif env ind in
- Inductive.type_of_constructors ind specif
+ Inductive.type_of_constructors indu specif
(* Return constructor types in normal form *)
-let arities_of_constructors env ind =
+let arities_of_constructors env (ind,u as indu) =
let specif = Inductive.lookup_mind_specif env ind in
- Inductive.arities_of_constructors ind specif
+ Inductive.arities_of_constructors indu specif
(* [inductive_family] = [inductive_instance] applied to global parameters *)
-type inductive_family = inductive * constr list
+type inductive_family = pinductive * constr list
let make_ind_family (mis, params) = (mis,params)
let dest_ind_family (mis,params) = (mis,params)
@@ -68,88 +73,224 @@ let lift_inductive_type n = liftn_inductive_type n 1
let substnl_ind_type l n = map_inductive_type (substnl l n)
let mkAppliedInd (IndType ((ind,params), realargs)) =
- applist (mkInd ind,params@realargs)
+ applist (mkIndU ind,params@realargs)
(* Does not consider imbricated or mutually recursive types *)
let mis_is_recursive_subset listind rarg =
- let rec one_is_rec rvec =
+ let one_is_rec rvec =
List.exists
(fun ra ->
match dest_recarg ra with
- | Mrec (_,i) -> List.mem i listind
+ | Mrec (_,i) -> Int.List.mem i listind
| _ -> false) rvec
in
- array_exists one_is_rec (dest_subterms rarg)
+ Array.exists one_is_rec (dest_subterms rarg)
let mis_is_recursive (ind,mib,mip) =
- mis_is_recursive_subset (interval 0 (mib.mind_ntypes-1))
+ mis_is_recursive_subset (List.interval 0 (mib.mind_ntypes - 1))
mip.mind_recargs
-let mis_nf_constructor_type (ind,mib,mip) j =
+let mis_nf_constructor_type ((ind,u),mib,mip) j =
let specif = mip.mind_nf_lc
and ntypes = mib.mind_ntypes
and nconstr = Array.length mip.mind_consnames in
- let make_Ik k = mkInd ((fst ind),ntypes-k-1) in
+ let make_Ik k = mkIndU (((fst ind),ntypes-k-1),u) in
if j > nconstr then error "Not enough constructors in the type.";
- substl (list_tabulate make_Ik ntypes) specif.(j-1)
+ substl (List.init ntypes make_Ik) (subst_instance_constr u specif.(j-1))
+
+(* Number of constructors *)
+
+let nconstructors ind =
+ let (_,mip) = Global.lookup_inductive ind in
+ Array.length mip.mind_consnames
+
+let nconstructors_env env ind =
+ let (_,mip) = Inductive.lookup_mind_specif env ind in
+ Array.length mip.mind_consnames
+
+(* Arity of constructors excluding parameters, excluding local defs *)
+
+let constructors_nrealargs ind =
+ let (_,mip) = Global.lookup_inductive ind in
+ mip.mind_consnrealargs
+
+let constructors_nrealargs_env env ind =
+ let (_,mip) = Inductive.lookup_mind_specif env ind in
+ mip.mind_consnrealargs
+
+(* Arity of constructors excluding parameters, including local defs *)
+
+let constructors_nrealdecls ind =
+ let (_,mip) = Global.lookup_inductive ind in
+ mip.mind_consnrealdecls
+
+let constructors_nrealdecls_env env ind =
+ let (_,mip) = Inductive.lookup_mind_specif env ind in
+ mip.mind_consnrealdecls
-(* Arity of constructors excluding parameters and local defs *)
+(* Arity of constructors including parameters, excluding local defs *)
-let mis_constr_nargs indsp =
+let constructor_nallargs (indsp,j) =
let (mib,mip) = Global.lookup_inductive indsp in
- let recargs = dest_subterms mip.mind_recargs in
- Array.map List.length recargs
+ mip.mind_consnrealargs.(j-1) + mib.mind_nparams
-let mis_constr_nargs_env env (kn,i) =
+let constructor_nallargs_env env ((kn,i),j) =
let mib = Environ.lookup_mind kn env in
let mip = mib.mind_packets.(i) in
- let recargs = dest_subterms mip.mind_recargs in
- Array.map List.length recargs
+ mip.mind_consnrealargs.(j-1) + mib.mind_nparams
-let mis_constructor_nargs_env env ((kn,i),j) =
+(* Arity of constructors including params, including local defs *)
+
+let constructor_nalldecls (indsp,j) = (* TOCHANGE en decls *)
+ let (mib,mip) = Global.lookup_inductive indsp in
+ mip.mind_consnrealdecls.(j-1) + rel_context_length (mib.mind_params_ctxt)
+
+let constructor_nalldecls_env env ((kn,i),j) = (* TOCHANGE en decls *)
let mib = Environ.lookup_mind kn env in
let mip = mib.mind_packets.(i) in
- recarg_length mip.mind_recargs j + mib.mind_nparams
+ mip.mind_consnrealdecls.(j-1) + rel_context_length (mib.mind_params_ctxt)
+
+(* Arity of constructors excluding params, excluding local defs *)
-let constructor_nrealargs env (ind,j) =
+let constructor_nrealargs (ind,j) =
+ let (_,mip) = Global.lookup_inductive ind in
+ mip.mind_consnrealargs.(j-1)
+
+let constructor_nrealargs_env env (ind,j) =
let (_,mip) = Inductive.lookup_mind_specif env ind in
- recarg_length mip.mind_recargs j
+ mip.mind_consnrealargs.(j-1)
-let constructor_nrealhyps env (ind,j) =
- let (mib,mip) = Inductive.lookup_mind_specif env ind in
+(* Arity of constructors excluding params, including local defs *)
+
+let constructor_nrealdecls (ind,j) = (* TOCHANGE en decls *)
+ let (_,mip) = Global.lookup_inductive ind in
mip.mind_consnrealdecls.(j-1)
-let get_full_arity_sign env ind =
- let (mib,mip) = Inductive.lookup_mind_specif env ind in
- mip.mind_arity_ctxt
+let constructor_nrealdecls_env env (ind,j) = (* TOCHANGE en decls *)
+ let (_,mip) = Inductive.lookup_mind_specif env ind in
+ mip.mind_consnrealdecls.(j-1)
-let nconstructors ind =
- let (mib,mip) = Inductive.lookup_mind_specif (Global.env()) ind in
- Array.length mip.mind_consnames
+(* Length of arity, excluding params, excluding local defs *)
+
+let inductive_nrealargs ind =
+ let (_,mip) = Global.lookup_inductive ind in
+ mip.mind_nrealargs
+
+let inductive_nrealargs_env env ind =
+ let (_,mip) = Inductive.lookup_mind_specif env ind in
+ mip.mind_nrealargs
+
+(* Length of arity, excluding params, including local defs *)
+
+let inductive_nrealdecls ind =
+ let (_,mip) = Global.lookup_inductive ind in
+ mip.mind_nrealdecls
+
+let inductive_nrealdecls_env env ind =
+ let (_,mip) = Inductive.lookup_mind_specif env ind in
+ mip.mind_nrealdecls
+
+(* Full length of arity (w/o local defs) *)
+
+let inductive_nallargs ind =
+ let (mib,mip) = Global.lookup_inductive ind in
+ mib.mind_nparams + mip.mind_nrealargs
+
+let inductive_nallargs_env env ind =
+ let (mib,mip) = Inductive.lookup_mind_specif env ind in
+ mib.mind_nparams + mip.mind_nrealargs
(* Length of arity (w/o local defs) *)
-let inductive_nargs env ind =
+let inductive_nparams ind =
+ let (mib,mip) = Global.lookup_inductive ind in
+ mib.mind_nparams
+
+let inductive_nparams_env env ind =
+ let (mib,mip) = Inductive.lookup_mind_specif env ind in
+ mib.mind_nparams
+
+(* Length of arity (with local defs) *)
+
+let inductive_nparamdecls ind =
+ let (mib,mip) = Global.lookup_inductive ind in
+ rel_context_length mib.mind_params_ctxt
+
+let inductive_nparamdecls_env env ind =
+ let (mib,mip) = Inductive.lookup_mind_specif env ind in
+ rel_context_length mib.mind_params_ctxt
+
+(* Full length of arity (with local defs) *)
+
+let inductive_nalldecls ind =
+ let (mib,mip) = Global.lookup_inductive ind in
+ rel_context_length (mib.mind_params_ctxt) + mip.mind_nrealdecls
+
+let inductive_nalldecls_env env ind =
let (mib,mip) = Inductive.lookup_mind_specif env ind in
- (rel_context_length (mib.mind_params_ctxt), mip.mind_nrealargs_ctxt)
+ rel_context_length (mib.mind_params_ctxt) + mip.mind_nrealdecls
+
+(* Others *)
+
+let inductive_paramdecls (ind,u) =
+ let (mib,mip) = Global.lookup_inductive ind in
+ Inductive.inductive_paramdecls (mib,u)
+
+let inductive_paramdecls_env env (ind,u) =
+ let (mib,mip) = Inductive.lookup_mind_specif env ind in
+ Inductive.inductive_paramdecls (mib,u)
+
+let inductive_alldecls (ind,u) =
+ let (mib,mip) = Global.lookup_inductive ind in
+ Vars.subst_instance_context u mip.mind_arity_ctxt
+
+let inductive_alldecls_env env (ind,u) =
+ let (mib,mip) = Inductive.lookup_mind_specif env ind in
+ Vars.subst_instance_context u mip.mind_arity_ctxt
+
+let constructor_has_local_defs (indsp,j) =
+ let (mib,mip) = Global.lookup_inductive indsp in
+ let l1 = mip.mind_consnrealdecls.(j-1) + rel_context_length (mib.mind_params_ctxt) in
+ let l2 = recarg_length mip.mind_recargs j + mib.mind_nparams in
+ not (Int.equal l1 l2)
+
+let inductive_has_local_defs ind =
+ let (mib,mip) = Global.lookup_inductive ind in
+ let l1 = rel_context_length (mib.mind_params_ctxt) + mip.mind_nrealdecls in
+ let l2 = mib.mind_nparams + mip.mind_nrealargs in
+ not (Int.equal l1 l2)
let allowed_sorts env (kn,i as ind) =
let (mib,mip) = Inductive.lookup_mind_specif env ind in
mip.mind_kelim
+let projection_nparams_env env p =
+ let pb = lookup_projection p env in
+ pb.proj_npars
+
+let projection_nparams p = projection_nparams_env (Global.env ()) p
+
(* Annotation for cases *)
let make_case_info env ind style =
let (mib,mip) = Inductive.lookup_mind_specif env ind in
- let print_info = { ind_nargs = mip.mind_nrealargs_ctxt; style = style } in
+ let ind_tags =
+ rel_context_tags (List.firstn mip.mind_nrealargs mip.mind_arity_ctxt) in
+ let cstr_tags =
+ Array.map2 (fun c n ->
+ let d,_ = decompose_prod_assum c in
+ rel_context_tags (List.firstn n d))
+ mip.mind_nf_lc mip.mind_consnrealdecls in
+ let print_info = { ind_tags; cstr_tags; style } in
{ ci_ind = ind;
ci_npar = mib.mind_nparams;
ci_cstr_ndecls = mip.mind_consnrealdecls;
+ ci_cstr_nargs = mip.mind_consnrealargs;
ci_pp_info = print_info }
(*s Useful functions *)
type constructor_summary = {
- cs_cstr : constructor;
+ cs_cstr : pconstructor;
cs_params : constr list;
cs_nargs : int;
cs_args : rel_context;
@@ -170,33 +311,39 @@ let instantiate_params t args sign =
| ((_,None,_)::ctxt,a::args) ->
(match kind_of_term t with
| Prod(_,_,t) -> inst (a::s) t (ctxt,args)
- | _ -> anomaly"instantiate_params: type, ctxt and args mismatch")
+ | _ -> anomaly ~label:"instantiate_params" (Pp.str "type, ctxt and args mismatch"))
| ((_,(Some b),_)::ctxt,args) ->
(match kind_of_term t with
| LetIn(_,_,_,t) -> inst ((substl s b)::s) t (ctxt,args)
- | _ -> anomaly"instantiate_params: type, ctxt and args mismatch")
+ | _ -> anomaly ~label:"instantiate_params" (Pp.str "type, ctxt and args mismatch"))
| _, [] -> substl s t
- | _ -> anomaly"instantiate_params: type, ctxt and args mismatch"
+ | _ -> anomaly ~label:"instantiate_params" (Pp.str "type, ctxt and args mismatch")
in inst [] t (List.rev sign,args)
-let get_constructor (ind,mib,mip,params) j =
+let get_constructor ((ind,u as indu),mib,mip,params) j =
assert (j <= Array.length mip.mind_consnames);
- let typi = mis_nf_constructor_type (ind,mib,mip) j in
+ let typi = mis_nf_constructor_type (indu,mib,mip) j in
let typi = instantiate_params typi params mib.mind_params_ctxt in
let (args,ccl) = decompose_prod_assum typi in
let (_,allargs) = decompose_app ccl in
- let vargs = list_skipn (List.length params) allargs in
- { cs_cstr = ith_constructor_of_inductive ind j;
+ let vargs = List.skipn (List.length params) allargs in
+ { cs_cstr = (ith_constructor_of_inductive ind j,u);
cs_params = params;
cs_nargs = rel_context_length args;
cs_args = args;
cs_concl_realargs = Array.of_list vargs }
let get_constructors env (ind,params) =
- let (mib,mip) = Inductive.lookup_mind_specif env ind in
+ let (mib,mip) = Inductive.lookup_mind_specif env (fst ind) in
Array.init (Array.length mip.mind_consnames)
(fun j -> get_constructor (ind,mib,mip,params) (j+1))
+let get_projections env (ind,params) =
+ let (mib,mip) = Inductive.lookup_mind_specif env (fst ind) in
+ match mib.mind_record with
+ | Some (Some (id, projs, pbs)) -> Some projs
+ | _ -> None
+
(* substitution in a signature *)
let substnl_rel_context subst n sign =
@@ -207,15 +354,15 @@ let substnl_rel_context subst n sign =
let substl_rel_context subst = substnl_rel_context subst 0
-let rec instantiate_context sign args =
+let instantiate_context sign args =
let rec aux subst = function
| (_,None,_)::sign, a::args -> aux (a::subst) (sign,args)
| (_,Some b,_)::sign, args -> aux (substl subst b::subst) (sign,args)
| [], [] -> subst
- | _ -> anomaly "Signature/instance mismatch in inductive family"
+ | _ -> anomaly (Pp.str "Signature/instance mismatch in inductive family")
in aux [] (List.rev sign,args)
-let get_arity env (ind,params) =
+let get_arity env ((ind,u),params) =
let (mib,mip) = Inductive.lookup_mind_specif env ind in
let parsign =
(* Dynamically detect if called with an instance of recursively
@@ -223,19 +370,21 @@ let get_arity env (ind,params) =
parameters *)
let parsign = mib.mind_params_ctxt in
let nnonrecparams = mib.mind_nparams - mib.mind_nparams_rec in
- if List.length params = rel_context_nhyps parsign - nnonrecparams then
- snd (list_chop nnonrecparams mib.mind_params_ctxt)
+ if Int.equal (List.length params) (rel_context_nhyps parsign - nnonrecparams) then
+ snd (List.chop nnonrecparams mib.mind_params_ctxt)
else
parsign in
+ let parsign = Vars.subst_instance_context u parsign in
let arproperlength = List.length mip.mind_arity_ctxt - List.length parsign in
- let arsign,_ = list_chop arproperlength mip.mind_arity_ctxt in
+ let arsign,_ = List.chop arproperlength mip.mind_arity_ctxt in
let subst = instantiate_context parsign params in
+ let arsign = Vars.subst_instance_context u arsign in
(substl_rel_context subst arsign, Inductive.inductive_sort_family mip)
(* Functions to build standard types related to inductive *)
let build_dependent_constructor cs =
applist
- (mkConstruct cs.cs_cstr,
+ (mkConstructU cs.cs_cstr,
(List.map (lift cs.cs_nargs) cs.cs_params)
@(extended_rel_list 0 cs.cs_args))
@@ -243,7 +392,7 @@ let build_dependent_inductive env ((ind, params) as indf) =
let arsign,_ = get_arity env indf in
let nrealargs = List.length arsign in
applist
- (mkInd ind,
+ (mkIndU ind,
(List.map (lift nrealargs) params)@(extended_rel_list 0 arsign))
(* builds the arity of an elimination predicate in sort [s] *)
@@ -252,7 +401,7 @@ let make_arity_signature env dep indf =
let (arsign,_) = get_arity env indf in
if dep then
(* We need names everywhere *)
- name_context env
+ Namegen.name_context env
((Anonymous,None,build_dependent_inductive env indf)::arsign)
(* Costly: would be better to name once for all at definition time *)
else
@@ -265,7 +414,7 @@ let make_arity env dep indf s = mkArity (make_arity_signature env dep indf, s)
let build_branch_type env dep p cs =
let base = appvect (lift cs.cs_nargs p, cs.cs_concl_realargs) in
if dep then
- it_mkProd_or_LetIn_name env
+ Namegen.it_mkProd_or_LetIn_name env
(applist (base,[build_dependent_constructor cs]))
cs.cs_args
else
@@ -288,18 +437,18 @@ let find_mrectype env sigma c =
let find_rectype env sigma c =
let (t, l) = decompose_app (whd_betadeltaiota env sigma c) in
match kind_of_term t with
- | Ind ind ->
+ | Ind (ind,u as indu) ->
let (mib,mip) = Inductive.lookup_mind_specif env ind in
if mib.mind_nparams > List.length l then raise Not_found;
- let (par,rargs) = list_chop mib.mind_nparams l in
- IndType((ind, par),rargs)
+ let (par,rargs) = List.chop mib.mind_nparams l in
+ IndType((indu, par),rargs)
| _ -> raise Not_found
let find_inductive env sigma c =
let (t, l) = decompose_app (whd_betadeltaiota env sigma c) in
match kind_of_term t with
| Ind ind
- when (fst (Inductive.lookup_mind_specif env ind)).mind_finite ->
+ when (fst (Inductive.lookup_mind_specif env (fst ind))).mind_finite <> Decl_kinds.CoFinite ->
(ind, l)
| _ -> raise Not_found
@@ -307,7 +456,7 @@ let find_coinductive env sigma c =
let (t, l) = decompose_app (whd_betadeltaiota env sigma c) in
match kind_of_term t with
| Ind ind
- when not (fst (Inductive.lookup_mind_specif env ind)).mind_finite ->
+ when (fst (Inductive.lookup_mind_specif env (fst ind))).mind_finite == Decl_kinds.CoFinite ->
(ind, l)
| _ -> raise Not_found
@@ -322,7 +471,7 @@ let is_predicate_explicitly_dep env pred arsign =
match kind_of_term pv', arsign with
| Lambda (na,t,b), (_,None,_)::arsign ->
srec (push_rel_assum (na,t) env) b arsign
- | Lambda (na,_,_), _ ->
+ | Lambda (na,_,t), _ ->
(* The following code has an impact on the introduction names
given by the tactics "case" and "inversion": when the
@@ -341,13 +490,15 @@ let is_predicate_explicitly_dep env pred arsign =
dependency status (of course, Anonymous implies non
dependent, but not conversely).
- At the end, this is only to preserve the compatibility: a
- check whether the predicate is actually dependent or not
- would indeed be more natural! *)
+ From Coq > 8.2, using or not the the effective dependency of
+ the predicate is parametrable! *)
- na <> Anonymous
+ begin match na with
+ | Anonymous -> false
+ | Name _ -> true
+ end
- | _ -> anomaly "Non eta-expanded dep-expanded \"match\" predicate"
+ | _ -> anomaly (Pp.str "Non eta-expanded dep-expanded \"match\" predicate")
in
srec env pred arsign
@@ -357,7 +508,7 @@ let is_elim_predicate_explicitly_dependent env pred indf =
let set_names env n brty =
let (ctxt,cl) = decompose_prod_n_assum n brty in
- it_mkProd_or_LetIn_name env cl ctxt
+ Namegen.it_mkProd_or_LetIn_name env cl ctxt
let set_pattern_names env ind brv =
let (mib,mip) = Inductive.lookup_mind_specif env ind in
@@ -367,19 +518,19 @@ let set_pattern_names env ind brv =
rel_context_length ((prod_assum c)) -
mib.mind_nparams)
mip.mind_nf_lc in
- array_map2 (set_names env) arities brv
+ Array.map2 (set_names env) arities brv
let type_case_branches_with_names env indspec p c =
let (ind,args) = indspec in
- let (mib,mip as specif) = Inductive.lookup_mind_specif env ind 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
+ let (params,realargs) = List.chop nparams args in
let lbrty = Inductive.build_branches_type ind specif params p in
(* Build case type *)
let conclty = Reduction.beta_appvect p (Array.of_list (realargs@[c])) in
(* Adjust names *)
if is_elim_predicate_explicitly_dependent env p (ind,params) then
- (set_pattern_names env ind lbrty, conclty)
+ (set_pattern_names env (fst ind) lbrty, conclty)
else (lbrty, conclty)
(* Type of Case predicates *)
@@ -393,40 +544,44 @@ let arity_of_case_predicate env (ind,params) dep k =
(* Inferring the sort of parameters of a polymorphic inductive type
knowing the sort of the conclusion *)
+
(* Compute the inductive argument types: replace the sorts
that appear in the type of the inductive by the sort of the
conclusion, and the other ones by fresh universes. *)
-let rec instantiate_universes env scl is = function
+let rec instantiate_universes env evdref scl is = function
| (_,Some _,_ as d)::sign, exp ->
- d :: instantiate_universes env scl is (sign, exp)
+ d :: instantiate_universes env evdref scl is (sign, exp)
| d::sign, None::exp ->
- d :: instantiate_universes env scl is (sign, exp)
- | (na,None,ty)::sign, Some u::exp ->
+ d :: instantiate_universes env evdref scl is (sign, exp)
+ | (na,None,ty)::sign, Some l::exp ->
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)
the sort of inductive [is] ? *)
- if univ_depends u is then
+ if univ_level_mem l is then
scl (* constrained sort: replace by scl *)
else
- (* unconstriained sort: replace by fresh universe *)
- new_Type_sort() in
- (na,None,mkArity(ctx,s)):: instantiate_universes env scl is (sign, exp)
+ (* 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
+ in
+ (na,None,mkArity(ctx,s)):: instantiate_universes env evdref scl is (sign, exp)
| sign, [] -> sign (* Uniform parameters are exhausted *)
| [], _ -> assert false
-(* Does not deal with universes, but only with Set/Type distinction *)
-let type_of_inductive_knowing_conclusion env mip conclty =
+let type_of_inductive_knowing_conclusion env sigma ((mib,mip),u) conclty =
match mip.mind_arity with
- | Monomorphic s ->
- s.mind_user_arity
- | Polymorphic ar ->
- let _,scl = Reduction.dest_arity env conclty in
- let ctx = List.rev mip.mind_arity_ctxt in
- let ctx =
- instantiate_universes
- env scl ar.poly_level (ctx,ar.poly_param_levels) in
- mkArity (List.rev ctx,scl)
+ | RegularArity s -> sigma, subst_instance_constr u s.mind_user_arity
+ | TemplateArity ar ->
+ let _,scl = Reduction.dest_arity env conclty 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)
(***********************************************)
(* Guard condition *)
@@ -447,7 +602,3 @@ let control_only_guard env c =
iter_constr_with_full_binders push_rel iter env c
in
iter env c
-
-let subst_inductive subst (kn,i as ind) =
- let kn' = Mod_subst.subst_ind subst kn in
- if kn == kn' then ind else (kn',i)
diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli
index 133d7013..af1783b7 100644
--- a/pretyping/inductiveops.mli
+++ b/pretyping/inductiveops.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,27 +8,27 @@
open Names
open Term
+open Context
open Declarations
open Environ
open Evd
-open Sign
(** The following three functions are similar to the ones defined in
Inductive, but they expect an env *)
-val type_of_inductive : env -> inductive -> types
+val type_of_inductive : env -> pinductive -> types
(** Return type as quoted by the user *)
-val type_of_constructor : env -> constructor -> types
-val type_of_constructors : env -> inductive -> types array
+val type_of_constructor : env -> pconstructor -> types
+val type_of_constructors : env -> pinductive -> types array
(** Return constructor types in normal form *)
-val arities_of_constructors : env -> inductive -> types array
+val arities_of_constructors : env -> pinductive -> types array
(** An inductive type with its parameters *)
type inductive_family
-val make_ind_family : inductive * constr list -> inductive_family
-val dest_ind_family : inductive_family -> inductive * constr list
+val make_ind_family : inductive puniverses * constr list -> inductive_family
+val dest_ind_family : inductive_family -> inductive 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
@@ -49,31 +49,86 @@ val mis_is_recursive_subset : int list -> wf_paths -> bool
val mis_is_recursive :
inductive * mutual_inductive_body * one_inductive_body -> bool
val mis_nf_constructor_type :
- inductive * mutual_inductive_body * one_inductive_body -> int -> constr
+ pinductive * mutual_inductive_body * one_inductive_body -> int -> constr
-(** Extract information from an inductive name *)
-
-(** Arity of constructors excluding parameters and local defs *)
-val mis_constr_nargs : inductive -> int array
-val mis_constr_nargs_env : env -> inductive -> int array
+(** {6 Extract information from an inductive name} *)
+(** @return number of constructors *)
val nconstructors : inductive -> int
+val nconstructors_env : env -> inductive -> int
+
+(** @return arity of constructors excluding parameters, excluding local defs *)
+val constructors_nrealargs : inductive -> int array
+val constructors_nrealargs_env : env -> inductive -> int array
+
+(** @return arity of constructors excluding parameters, including local defs *)
+val constructors_nrealdecls : inductive -> int array
+val constructors_nrealdecls_env : env -> inductive -> int array
+
+(** @return the arity, excluding params, excluding local defs *)
+val inductive_nrealargs : inductive -> int
+val inductive_nrealargs_env : env -> inductive -> int
+
+(** @return the arity, excluding params, including local defs *)
+val inductive_nrealdecls : inductive -> int
+val inductive_nrealdecls_env : env -> inductive -> int
+
+(** @return the arity, including params, excluding local defs *)
+val inductive_nallargs : inductive -> int
+val inductive_nallargs_env : env -> inductive -> int
+
+(** @return the arity, including params, including local defs *)
+val inductive_nalldecls : inductive -> int
+val inductive_nalldecls_env : env -> inductive -> int
+
+(** @return nb of params without local defs *)
+val inductive_nparams : inductive -> int
+val inductive_nparams_env : env -> inductive -> int
+
+(** @return nb of params with local defs *)
+val inductive_nparamdecls : inductive -> int
+val inductive_nparamdecls_env : env -> inductive -> int
-(** Return the lengths of parameters signature and real arguments signature *)
-val inductive_nargs : env -> inductive -> int * int
+(** @return params context *)
+val inductive_paramdecls : pinductive -> rel_context
+val inductive_paramdecls_env : env -> pinductive -> rel_context
-val mis_constructor_nargs_env : env -> constructor -> int
-val constructor_nrealargs : env -> constructor -> int
-val constructor_nrealhyps : env -> constructor -> int
+(** @return full arity context, hence with letin *)
+val inductive_alldecls : pinductive -> rel_context
+val inductive_alldecls_env : env -> pinductive -> rel_context
-val get_full_arity_sign : env -> inductive -> rel_context
+(** {7 Extract information from a constructor name} *)
+
+(** @return param + args without letin *)
+val constructor_nallargs : constructor -> int
+val constructor_nallargs_env : env -> constructor -> int
+
+(** @return param + args with letin *)
+val constructor_nalldecls : constructor -> int
+val constructor_nalldecls_env : env -> constructor -> int
+
+(** @return args without letin *)
+val constructor_nrealargs : constructor -> int
+val constructor_nrealargs_env : env -> constructor -> int
+
+(** @return args with letin *)
+val constructor_nrealdecls : constructor -> int
+val constructor_nrealdecls_env : env -> constructor -> int
+
+(** Is there local defs in params or args ? *)
+val constructor_has_local_defs : constructor -> bool
+val inductive_has_local_defs : inductive -> bool
val allowed_sorts : env -> inductive -> sorts_family list
+(** Primitive projections *)
+val projection_nparams : projection -> int
+val projection_nparams_env : env -> projection -> int
+
(** Extract information from an inductive family *)
type constructor_summary = {
- cs_cstr : constructor; (* internal name of the constructor *)
+ cs_cstr : pconstructor; (* internal name of the constructor plus universes *)
cs_params : constr list; (* parameters of the constructor in current ctx *)
cs_nargs : int; (* length of arguments signature (letin included) *)
cs_args : rel_context; (* signature of the arguments (letin included) *)
@@ -81,22 +136,24 @@ type constructor_summary = {
}
val lift_constructor : int -> constructor_summary -> constructor_summary
val get_constructor :
- inductive * mutual_inductive_body * one_inductive_body * constr list ->
+ pinductive * mutual_inductive_body * one_inductive_body * constr list ->
int -> constructor_summary
val get_arity : env -> inductive_family -> rel_context * sorts_family
val get_constructors : env -> inductive_family -> constructor_summary array
+val get_projections : env -> inductive_family -> constant array option
+
val build_dependent_constructor : constructor_summary -> constr
val build_dependent_inductive : env -> inductive_family -> constr
val make_arity_signature : env -> bool -> inductive_family -> rel_context
val make_arity : env -> bool -> inductive_family -> sorts -> types
val build_branch_type : env -> bool -> constr -> constructor_summary -> types
-(** Raise [Not_found] if not given an valid inductive type *)
-val extract_mrectype : constr -> inductive * constr list
-val find_mrectype : env -> evar_map -> types -> inductive * constr list
+(** 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_rectype : env -> evar_map -> types -> inductive_type
-val find_inductive : env -> evar_map -> types -> inductive * constr list
-val find_coinductive : env -> evar_map -> types -> inductive * constr list
+val find_inductive : env -> evar_map -> types -> pinductive * constr list
+val find_coinductive : env -> evar_map -> types -> pinductive * constr list
(********************)
@@ -105,8 +162,7 @@ val arity_of_case_predicate :
env -> inductive_family -> bool -> sorts -> types
val type_case_branches_with_names :
- env -> inductive * constr list -> constr -> constr ->
- types array * types
+ env -> pinductive * constr list -> constr -> constr -> types array * types
(** Annotation for cases *)
val make_case_info : env -> inductive -> case_style -> case_info
@@ -118,9 +174,7 @@ i*)
(********************)
val type_of_inductive_knowing_conclusion :
- env -> one_inductive_body -> types -> types
+ env -> evar_map -> Inductive.mind_specif puniverses -> types -> evar_map * types
(********************)
val control_only_guard : env -> types -> unit
-
-val subst_inductive : Mod_subst.substitution -> inductive -> inductive
diff --git a/pretyping/locusops.ml b/pretyping/locusops.ml
new file mode 100644
index 00000000..4a5e11f0
--- /dev/null
+++ b/pretyping/locusops.ml
@@ -0,0 +1,125 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Locus
+
+(** Utilities on occurrences *)
+
+let occurrences_map f = function
+ | OnlyOccurrences l ->
+ let l' = f l in
+ if l' = [] then NoOccurrences else OnlyOccurrences l'
+ | AllOccurrencesBut l ->
+ let l' = f l in
+ if l' = [] then AllOccurrences else AllOccurrencesBut l'
+ | (NoOccurrences|AllOccurrences) as o -> o
+
+let convert_occs = function
+ | AllOccurrences -> (false,[])
+ | AllOccurrencesBut l -> (false,l)
+ | NoOccurrences -> (true,[])
+ | OnlyOccurrences l -> (true,l)
+
+let is_selected occ = function
+ | AllOccurrences -> true
+ | AllOccurrencesBut l -> not (Int.List.mem occ l)
+ | OnlyOccurrences l -> Int.List.mem occ l
+ | NoOccurrences -> false
+
+(** Usual clauses *)
+
+let allHypsAndConcl = { onhyps=None; concl_occs=AllOccurrences }
+let allHyps = { onhyps=None; concl_occs=NoOccurrences }
+let onConcl = { onhyps=Some[]; concl_occs=AllOccurrences }
+let nowhere = { onhyps=Some[]; concl_occs=NoOccurrences }
+let onHyp h =
+ { onhyps=Some[(AllOccurrences,h),InHyp]; concl_occs=NoOccurrences }
+
+let is_nowhere = function
+| { onhyps=Some[]; concl_occs=NoOccurrences } -> true
+| _ -> false
+
+(** Clause conversion functions, parametrized by a hyp enumeration function *)
+
+(** From [clause] to [simple_clause] *)
+
+let simple_clause_of enum_hyps cl =
+ let error_occurrences () =
+ Errors.error "This tactic does not support occurrences selection" in
+ let error_body_selection () =
+ Errors.error "This tactic does not support body selection" in
+ let hyps =
+ match cl.onhyps with
+ | None ->
+ List.map Option.make (enum_hyps ())
+ | Some l ->
+ List.map (fun ((occs,id),w) ->
+ if occs <> AllOccurrences then error_occurrences ();
+ if w = InHypValueOnly then error_body_selection ();
+ Some id) l in
+ if cl.concl_occs = NoOccurrences then hyps
+ else
+ if cl.concl_occs <> AllOccurrences then error_occurrences ()
+ else None :: hyps
+
+(** From [clause] to [concrete_clause] *)
+
+let concrete_clause_of enum_hyps cl =
+ let hyps =
+ match cl.onhyps with
+ | None ->
+ let f id = OnHyp (id,AllOccurrences,InHyp) in
+ List.map f (enum_hyps ())
+ | Some l ->
+ List.map (fun ((occs,id),w) -> OnHyp (id,occs,w)) l in
+ if cl.concl_occs = NoOccurrences then hyps
+ else
+ OnConcl cl.concl_occs :: hyps
+
+(** Miscellaneous functions *)
+
+let out_arg = function
+ | Misctypes.ArgVar _ -> Errors.anomaly (Pp.str "Unevaluated or_var variable")
+ | Misctypes.ArgArg x -> x
+
+let occurrences_of_hyp id cls =
+ let rec hyp_occ = function
+ [] -> NoOccurrences, InHyp
+ | ((occs,id'),hl)::_ when Names.Id.equal id id' ->
+ occurrences_map (List.map out_arg) occs, hl
+ | _::l -> hyp_occ l in
+ match cls.onhyps with
+ None -> AllOccurrences,InHyp
+ | Some l -> hyp_occ l
+
+let occurrences_of_goal cls =
+ occurrences_map (List.map out_arg) cls.concl_occs
+
+let in_every_hyp cls = Option.is_empty cls.onhyps
+
+let clause_with_generic_occurrences cls =
+ let hyps = match cls.onhyps with
+ | None -> true
+ | Some hyps ->
+ List.for_all
+ (function ((AllOccurrences,_),_) -> true | _ -> false) hyps in
+ let concl = match cls.concl_occs with
+ | AllOccurrences | NoOccurrences -> true
+ | _ -> false in
+ hyps && concl
+
+let clause_with_generic_context_selection cls =
+ let hyps = match cls.onhyps with
+ | None -> true
+ | Some hyps ->
+ List.for_all
+ (function ((AllOccurrences,_),InHyp) -> true | _ -> false) hyps in
+ let concl = match cls.concl_occs with
+ | AllOccurrences | NoOccurrences -> true
+ | _ -> false in
+ hyps && concl
diff --git a/pretyping/locusops.mli b/pretyping/locusops.mli
new file mode 100644
index 00000000..79dc3734
--- /dev/null
+++ b/pretyping/locusops.mli
@@ -0,0 +1,46 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Locus
+
+(** Utilities on occurrences *)
+
+val occurrences_map :
+ ('a list -> 'b list) -> 'a occurrences_gen -> 'b occurrences_gen
+
+(** From occurrences to a list of positions (or complement of positions) *)
+val convert_occs : occurrences -> bool * int list
+
+val is_selected : int -> occurrences -> bool
+
+(** Usual clauses *)
+
+val allHypsAndConcl : 'a clause_expr
+val allHyps : 'a clause_expr
+val onConcl : 'a clause_expr
+val nowhere : 'a clause_expr
+val onHyp : 'a -> 'a clause_expr
+
+(** Tests *)
+
+val is_nowhere : 'a clause_expr -> bool
+
+(** Clause conversion functions, parametrized by a hyp enumeration function *)
+
+val simple_clause_of : (unit -> Id.t list) -> clause -> simple_clause
+val concrete_clause_of : (unit -> Id.t list) -> clause -> concrete_clause
+
+(** Miscellaneous functions *)
+
+val occurrences_of_hyp : Id.t -> clause -> (occurrences * hyp_location_flag)
+val occurrences_of_goal : clause -> occurrences
+val in_every_hyp : clause -> bool
+
+val clause_with_generic_occurrences : 'a clause_expr -> bool
+val clause_with_generic_context_selection : 'a clause_expr -> bool
diff --git a/pretyping/matching.ml b/pretyping/matching.ml
deleted file mode 100644
index b86f3e45..00000000
--- a/pretyping/matching.ml
+++ /dev/null
@@ -1,357 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(*i*)
-open Pp
-open Util
-open Names
-open Libnames
-open Nameops
-open Termops
-open Reductionops
-open Term
-open Glob_term
-open Sign
-open Environ
-open Pattern
-(*i*)
-
-(* Given a term with second-order variables in it,
- represented by Meta's, and possibly applied using [SOAPP] to
- terms, this function will perform second-order, binding-preserving,
- matching, in the case where the pattern is a pattern in the sense
- of Dale Miller.
-
- ALGORITHM:
-
- Given a pattern, we decompose it, flattening Cast's and apply's,
- recursing on all operators, and pushing the name of the binder each
- time we descend a binder.
-
- When we reach a first-order variable, we ask that the corresponding
- term's free-rels all be higher than the depth of the current stack.
-
- When we reach a second-order application, we ask that the
- intersection of the free-rels of the term and the current stack be
- contained in the arguments of the application, and in that case, we
- construct a LAMBDA with the names on the stack.
-
- *)
-
-type bound_ident_map = (identifier * identifier) list
-
-exception PatternMatchingFailure
-
-let constrain (n,(ids,m as x)) (names,terms as subst) =
- try
- let (ids',m') = List.assoc n terms in
- if ids = ids' && eq_constr m m' then subst
- else raise PatternMatchingFailure
- with
- Not_found ->
- if List.mem_assoc n names then
- Flags.if_warn Pp.msg_warning
- (str "Collision between bound variable " ++ pr_id n ++
- str " and a metavariable of same name.");
- (names,(n,x)::terms)
-
-let add_binders na1 na2 (names,terms as subst) =
- match na1, na2 with
- | Name id1, Name id2 ->
- if List.mem_assoc id1 names then
- (Flags.if_warn Pp.msg_warning
- (str "Collision between bound variables of name " ++ pr_id id1);
- (names,terms))
- else
- (if List.mem_assoc id1 terms then
- Flags.if_warn Pp.msg_warning
- (str "Collision between bound variable " ++ pr_id id1 ++
- str " and another bound variable of same name.");
- ((id1,id2)::names,terms));
- | _ -> subst
-
-let build_lambda toabstract stk (m : constr) =
- let rec buildrec m p_0 p_1 = match p_0,p_1 with
- | (_, []) -> m
- | (n, (_,na,t)::tl) ->
- if List.mem n toabstract then
- buildrec (mkLambda (na,t,m)) (n+1) tl
- else
- buildrec (lift (-1) m) (n+1) tl
- in
- buildrec m 1 stk
-
-let rec list_insert f a = function
- | [] -> [a]
- | b::l when f a b -> a::b::l
- | b::l when a = b -> raise PatternMatchingFailure
- | b::l -> b :: list_insert f a l
-
-let extract_bound_vars =
- let rec aux k = function
- | ([],_) -> []
- | (n::l,(na1,na2,_)::stk) when k = n ->
- begin match na1,na2 with
- | Name id1,Name _ -> list_insert (<) id1 (aux (k+1) (l,stk))
- | Name _,Anonymous -> anomaly "Unnamed bound variable"
- | Anonymous,_ -> raise PatternMatchingFailure
- end
- | (l,_::stk) -> aux (k+1) (l,stk)
- | (_,[]) -> assert false
- in aux 1
-
-let dummy_constr = mkProp
-
-let rec make_renaming ids = function
- | (Name id,Name _,_)::stk ->
- let renaming = make_renaming ids stk in
- (try mkRel (list_index id ids) :: renaming
- with Not_found -> dummy_constr :: renaming)
- | (_,_,_)::stk ->
- dummy_constr :: make_renaming ids stk
- | [] ->
- []
-
-let merge_binding allow_bound_rels stk n cT subst =
- let depth = List.length stk in
- let c =
- if depth = 0 then
- (* Optimization *)
- ([],cT)
- else
- let frels = Intset.elements (free_rels cT) in
- let frels = List.filter (fun i -> i <= depth) frels in
- if allow_bound_rels then
- let frels = Sort.list (<) frels in
- let canonically_ordered_vars = extract_bound_vars (frels,stk) in
- let renaming = make_renaming canonically_ordered_vars stk in
- (canonically_ordered_vars, substl renaming cT)
- else if frels = [] then
- ([],lift (-depth) cT)
- else
- raise PatternMatchingFailure in
- constrain (n,c) subst
-
-let matches_core convert allow_partial_app allow_bound_rels pat c =
- let conv = match convert with
- | None -> eq_constr
- | Some (env,sigma) -> is_conv env sigma in
- let rec sorec stk subst p t =
- let cT = strip_outer_cast t in
- match p,kind_of_term cT with
- | PSoApp (n,args),m ->
- let relargs =
- List.map
- (function
- | PRel n -> n
- | _ -> error "Only bound indices allowed in second order pattern matching.")
- args in
- let frels = Intset.elements (free_rels cT) in
- if list_subset frels relargs then
- constrain (n,([],build_lambda relargs stk cT)) subst
- else
- raise PatternMatchingFailure
-
- | PMeta (Some n), m -> merge_binding allow_bound_rels stk n cT subst
-
- | PMeta None, m -> subst
-
- | PRef (VarRef v1), Var v2 when v1 = v2 -> subst
-
- | PVar v1, Var v2 when v1 = v2 -> subst
-
- | PRef ref, _ when conv (constr_of_global ref) cT -> subst
-
- | PRel n1, Rel n2 when n1 = n2 -> subst
-
- | PSort (GProp c1), Sort (Prop c2) when c1 = c2 -> subst
-
- | PSort (GType _), Sort (Type _) -> subst
-
- | PApp (p, [||]), _ -> sorec stk subst p t
-
- | PApp (PApp (h, a1), a2), _ ->
- sorec stk subst (PApp(h,Array.append a1 a2)) t
-
- | PApp (PMeta meta,args1), App (c2,args2) when allow_partial_app ->
- let p = Array.length args2 - Array.length args1 in
- if p>=0 then
- let args21, args22 = array_chop p args2 in
- let c = mkApp(c2,args21) in
- let subst =
- match meta with
- | None -> subst
- | Some n -> merge_binding allow_bound_rels stk n c subst in
- array_fold_left2 (sorec stk) subst args1 args22
- else raise PatternMatchingFailure
-
- | PApp (c1,arg1), App (c2,arg2) ->
- (try array_fold_left2 (sorec stk) (sorec stk subst c1 c2) arg1 arg2
- with Invalid_argument _ -> raise PatternMatchingFailure)
-
- | PProd (na1,c1,d1), Prod(na2,c2,d2) ->
- sorec ((na1,na2,c2)::stk)
- (add_binders na1 na2 (sorec stk subst c1 c2)) d1 d2
-
- | PLambda (na1,c1,d1), Lambda(na2,c2,d2) ->
- sorec ((na1,na2,c2)::stk)
- (add_binders na1 na2 (sorec stk subst c1 c2)) d1 d2
-
- | PLetIn (na1,c1,d1), LetIn(na2,c2,t2,d2) ->
- sorec ((na1,na2,t2)::stk)
- (add_binders na1 na2 (sorec stk subst c1 c2)) d1 d2
-
- | PIf (a1,b1,b1'), Case (ci,_,a2,[|b2;b2'|]) ->
- let ctx,b2 = decompose_lam_n_assum ci.ci_cstr_ndecls.(0) b2 in
- let ctx',b2' = decompose_lam_n_assum ci.ci_cstr_ndecls.(1) b2' in
- let n = rel_context_length ctx in
- let n' = rel_context_length ctx' in
- if noccur_between 1 n b2 & noccur_between 1 n' b2' then
- let s =
- List.fold_left (fun l (na,_,t) -> (Anonymous,na,t)::l) stk ctx in
- let s' =
- List.fold_left (fun l (na,_,t) -> (Anonymous,na,t)::l) stk ctx' in
- let b1 = lift_pattern n b1 and b1' = lift_pattern n' b1' in
- sorec s' (sorec s (sorec stk subst a1 a2) b1 b2) b1' b2'
- else
- raise PatternMatchingFailure
-
- | PCase (ci1,p1,a1,br1), Case (ci2,p2,a2,br2) ->
- let n2 = Array.length br2 in
- if (ci1.cip_ind <> None && ci1.cip_ind <> Some ci2.ci_ind) ||
- (not ci1.cip_extensible && List.length br1 <> n2)
- then raise PatternMatchingFailure;
- let chk_branch subst (j,n,c) =
- (* (ind,j+1) is normally known to be a correct constructor
- and br2 a correct match over the same inductive *)
- assert (j < n2);
- sorec stk subst c br2.(j)
- in
- let chk_head = sorec stk (sorec stk 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
-
- in
- let names,terms = sorec [] ([],[]) pat c in
- (names,Sort.list (fun (a,_) (b,_) -> a<b) terms)
-
-let matches_core_closed convert allow_partial_app pat c =
- let names,subst = matches_core convert allow_partial_app false pat c in
- (names, List.map (fun (a,(_,b)) -> (a,b)) subst)
-
-let extended_matches = matches_core None true true
-
-let matches c p = snd (matches_core_closed None true c p)
-
-let special_meta = (-1)
-
-(* Tells if it is an authorized occurrence and if the instance is closed *)
-let authorized_occ partial_app closed pat c mk_ctx next =
- try
- let sigma = matches_core_closed None partial_app pat c in
- if closed && not (List.for_all (fun (_,c) -> closed0 c) (snd sigma))
- then next ()
- else sigma, mk_ctx (mkMeta special_meta), next
- with
- PatternMatchingFailure -> next ()
-
-(* Tries to match a subterm of [c] with [pat] *)
-let sub_match ?(partial_app=false) ?(closed=true) pat c =
- let rec aux c mk_ctx next =
- match kind_of_term c with
- | Cast (c1,k,c2) ->
- authorized_occ partial_app closed pat c mk_ctx (fun () ->
- let mk_ctx lc = mk_ctx (mkCast (List.hd lc, k,c2)) in
- try_aux [c1] mk_ctx next)
- | Lambda (x,c1,c2) ->
- authorized_occ partial_app closed pat c mk_ctx (fun () ->
- let mk_ctx lc = mk_ctx (mkLambda (x,List.hd lc,List.nth lc 1)) in
- try_aux [c1;c2] mk_ctx next)
- | Prod (x,c1,c2) ->
- authorized_occ partial_app closed pat c mk_ctx (fun () ->
- let mk_ctx lc = mk_ctx (mkProd (x,List.hd lc,List.nth lc 1)) in
- try_aux [c1;c2] mk_ctx next)
- | LetIn (x,c1,t,c2) ->
- authorized_occ partial_app closed pat c mk_ctx (fun () ->
- let mk_ctx = function [c1;c2] -> mkLetIn (x,c1,t,c2) | _ -> assert false
- in try_aux [c1;c2] mk_ctx next)
- | App (c1,lc) ->
- authorized_occ partial_app closed pat c mk_ctx (fun () ->
- 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 [app;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
- try_aux (c1::Array.to_list lc) 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 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
- try_aux (c1::Array.to_list lc) mk_ctx next)
- | Case (ci,hd,c1,lc) ->
- authorized_occ partial_app closed pat c mk_ctx (fun () ->
- let mk_ctx le =
- mk_ctx (mkCase (ci,hd,List.hd le,Array.of_list (List.tl le))) in
- try_aux (c1::Array.to_list lc) mk_ctx next)
- | Construct _ | Fix _ | Ind _|CoFix _ |Evar _|Const _
- | Rel _|Meta _|Var _|Sort _ ->
- authorized_occ partial_app closed pat c mk_ctx next
-
- (* Tries [sub_match] for all terms in the list *)
- and try_aux lc mk_ctx next =
- let rec try_sub_match_rec lacc = function
- | [] -> next ()
- | c::tl ->
- let mk_ctx ce = mk_ctx (List.rev_append lacc (ce::tl)) in
- let next () = try_sub_match_rec (c::lacc) tl in
- aux c mk_ctx next in
- try_sub_match_rec [] lc in
- aux c (fun x -> x) (fun () -> raise PatternMatchingFailure)
-
-type subterm_matching_result =
- (bound_ident_map * patvar_map) * constr * (unit -> subterm_matching_result)
-
-let match_subterm pat c = sub_match pat c
-
-let match_appsubterm pat c = sub_match ~partial_app:true pat c
-
-let match_subterm_gen app pat c = sub_match ~partial_app:app pat c
-
-let is_matching pat c =
- try let _ = matches pat c in true
- with PatternMatchingFailure -> false
-
-let is_matching_appsubterm ?(closed=true) pat c =
- try let _ = sub_match ~partial_app:true ~closed pat c in true
- with PatternMatchingFailure -> false
-
-let matches_conv env sigma c p =
- snd (matches_core_closed (Some (env,sigma)) false c p)
-
-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/miscops.ml b/pretyping/miscops.ml
new file mode 100644
index 00000000..a2c97d2c
--- /dev/null
+++ b/pretyping/miscops.ml
@@ -0,0 +1,60 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Util
+open Misctypes
+open Genredexpr
+
+(** Mapping [cast_type] *)
+
+let map_cast_type f = function
+ | CastConv a -> CastConv (f a)
+ | CastVM a -> CastVM (f a)
+ | CastCoerce -> CastCoerce
+ | CastNative a -> CastNative (f a)
+
+let smartmap_cast_type f c =
+ match c with
+ | CastConv a -> let a' = f a in if a' == a then c else CastConv a'
+ | CastVM a -> let a' = f a in if a' == a then c else CastVM a'
+ | CastCoerce -> CastCoerce
+ | CastNative a -> let a' = f a in if a' == a then c else CastNative a'
+
+(** Equalities on [glob_sort] *)
+
+let glob_sort_eq g1 g2 = match g1, g2 with
+| GProp, GProp -> true
+| GSet, GSet -> true
+| GType l1, GType l2 -> List.for_all2 CString.equal l1 l2
+| _ -> false
+
+let intro_pattern_naming_eq nam1 nam2 = match nam1, nam2 with
+| IntroAnonymous, IntroAnonymous -> true
+| IntroIdentifier id1, IntroIdentifier id2 -> Names.Id.equal id1 id2
+| IntroFresh id1, IntroFresh id2 -> Names.Id.equal id1 id2
+| _ -> false
+
+(** Mapping [red_expr_gen] *)
+
+let map_flags f flags =
+ { flags with rConst = List.map f flags.rConst }
+
+let map_occs f (occ,e) = (occ,f e)
+
+let map_red_expr_gen f g h = function
+ | Fold l -> Fold (List.map f l)
+ | Pattern occs_l -> Pattern (List.map (map_occs f) occs_l)
+ | Simpl (flags,occs_o) ->
+ Simpl (map_flags g flags, Option.map (map_occs (map_union g h)) occs_o)
+ | Unfold occs_l -> Unfold (List.map (map_occs g) occs_l)
+ | Cbv flags -> Cbv (map_flags g flags)
+ | Lazy flags -> Lazy (map_flags g flags)
+ | CbvVm occs_o -> CbvVm (Option.map (map_occs (map_union g h)) occs_o)
+ | 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
diff --git a/pretyping/miscops.mli b/pretyping/miscops.mli
new file mode 100644
index 00000000..453648d4
--- /dev/null
+++ b/pretyping/miscops.mli
@@ -0,0 +1,29 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Misctypes
+open Genredexpr
+
+(** Mapping [cast_type] *)
+
+val map_cast_type : ('a -> 'b) -> 'a cast_type -> 'b cast_type
+val smartmap_cast_type : ('a -> 'a) -> 'a cast_type -> 'a cast_type
+
+(** Equalities on [glob_sort] *)
+
+val glob_sort_eq : glob_sort -> glob_sort -> bool
+
+(** Equalities on [intro_pattern_naming] *)
+
+val intro_pattern_naming_eq :
+ intro_pattern_naming_expr -> intro_pattern_naming_expr -> bool
+
+(** Mapping [red_expr_gen] *)
+
+val map_red_expr_gen : ('a -> 'd) -> ('b -> 'e) -> ('c -> 'f) ->
+ ('a,'b,'c) red_expr_gen -> ('d,'e,'f) red_expr_gen
diff --git a/pretyping/namegen.ml b/pretyping/namegen.ml
index 76b9bd8a..5aca11ae 100644
--- a/pretyping/namegen.ml
+++ b/pretyping/namegen.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -15,25 +15,42 @@
open Util
open Names
open Term
+open Vars
open Nametab
open Nameops
open Libnames
+open Globnames
open Environ
open Termops
(**********************************************************************)
+(* Conventional names *)
+
+let default_prop_string = "H"
+let default_prop_ident = Id.of_string default_prop_string
+
+let default_small_string = "H"
+let default_small_ident = Id.of_string default_small_string
+
+let default_type_string = "X"
+let default_type_ident = Id.of_string default_type_string
+
+let default_non_dependent_string = "H"
+let default_non_dependent_ident = Id.of_string default_non_dependent_string
+
+let default_dependent_ident = Id.of_string "x"
+
+(**********************************************************************)
(* Globality of identifiers *)
-let rec is_imported_modpath mp =
- let current_mp,_ = Lib.current_prefix() in
- match mp with
- | MPfile dp ->
- let rec find_prefix = function
- |MPfile dp1 -> not (dp1=dp)
- |MPdot(mp,_) -> find_prefix mp
- |MPbound(_) -> false
- in find_prefix current_mp
- | p -> false
+let is_imported_modpath = function
+ | MPfile dp ->
+ let rec find_prefix = function
+ |MPfile dp1 -> not (DirPath.equal dp1 dp)
+ |MPdot(mp,_) -> find_prefix mp
+ |MPbound(_) -> false
+ in find_prefix (Lib.current_mp ())
+ | _ -> false
let is_imported_ref = function
| VarRef _ -> false
@@ -61,8 +78,22 @@ let is_constructor id =
(**********************************************************************)
(* Generating "intuitive" names from its type *)
+let head_name c = (* Find the head constant of a constr if any *)
+ let rec hdrec c =
+ match kind_of_term c with
+ | Prod (_,_,c) | Lambda (_,_,c) | LetIn (_,_,_,c)
+ | Cast (c,_,_) | App (c,_) -> hdrec c
+ | Proj (kn,_) -> Some (Label.to_id (con_label (Projection.constant kn)))
+ | Const _ | Ind _ | Construct _ | Var _ ->
+ Some (basename_of_global (global_of_constr c))
+ | Fix ((_,i),(lna,_,_)) | CoFix (i,(lna,_,_)) ->
+ Some (match lna.(i) with Name id -> id | _ -> assert false)
+ | Sort _ | Rel _ | Meta _|Evar _|Case (_, _, _, _) -> None
+ in
+ hdrec c
+
let lowercase_first_char id = (* First character of a constr *)
- lowercase_first_char_utf8 (string_of_id id)
+ Unicode.lowercase_first_char (Id.to_string id)
let sort_hdchar = function
| Prop(_) -> "P"
@@ -71,14 +102,12 @@ let sort_hdchar = function
let hdchar env c =
let rec hdrec k c =
match kind_of_term c with
- | Prod (_,_,c) -> hdrec (k+1) c
- | Lambda (_,_,c) -> hdrec (k+1) c
- | LetIn (_,_,_,c) -> hdrec (k+1) c
- | Cast (c,_,_) -> hdrec k c
- | App (f,l) -> hdrec k f
- | Const kn -> lowercase_first_char (id_of_label (con_label kn))
- | Ind x -> lowercase_first_char (basename_of_global (IndRef x))
- | Construct x -> lowercase_first_char (basename_of_global (ConstructRef x))
+ | Prod (_,_,c) | Lambda (_,_,c) | LetIn (_,_,_,c) -> hdrec (k+1) c
+ | Cast (c,_,_) | App (c,_) -> hdrec k c
+ | Proj (kn,_) -> lowercase_first_char (Label.to_id (con_label (Projection.constant kn)))
+ | Const (kn,_) -> lowercase_first_char (Label.to_id (con_label kn))
+ | Ind (x,_) -> lowercase_first_char (basename_of_global (IndRef x))
+ | Construct (x,_) -> lowercase_first_char (basename_of_global (ConstructRef x))
| Var id -> lowercase_first_char id
| Sort s -> sort_hdchar s
| Rel n ->
@@ -88,22 +117,20 @@ let hdchar env c =
| (Name id,_,_) -> lowercase_first_char id
| (Anonymous,_,t) -> hdrec 0 (lift (n-k) t)
with Not_found -> "y")
- | Fix ((_,i),(lna,_,_)) ->
- let id = match lna.(i) with Name id -> id | _ -> assert false in
- lowercase_first_char id
- | CoFix (i,(lna,_,_)) ->
+ | Fix ((_,i),(lna,_,_)) | CoFix (i,(lna,_,_)) ->
let id = match lna.(i) with Name id -> id | _ -> assert false in
lowercase_first_char id
- | Meta _|Evar _|Case (_, _, _, _) -> "y"
+ | Evar _ (* We could do better... *)
+ | Meta _ | Case (_, _, _, _) -> "y"
in
hdrec 0 c
let id_of_name_using_hdchar env a = function
- | Anonymous -> id_of_string (hdchar env a)
+ | Anonymous -> Id.of_string (hdchar env a)
| Name id -> id
let named_hd env a = function
- | Anonymous -> Name (id_of_string (hdchar env a))
+ | Anonymous -> Name (Id.of_string (hdchar env a))
| x -> x
let mkProd_name env (n,a,b) = mkProd (named_hd env a n, a, b)
@@ -138,8 +165,6 @@ let it_mkLambda_or_LetIn_name env b hyps =
(**********************************************************************)
(* Fresh names *)
-let default_x = id_of_string "x"
-
(* Looks for next "good" name by lifting subscript *)
let next_ident_away_from id bad =
@@ -151,17 +176,46 @@ let next_ident_away_from id bad =
let restart_subscript id =
if not (has_subscript id) then id else
- (* Ce serait sans doute mieux avec quelque chose inspiré de
- *** make_ident id (Some 0) *** mais ça brise la compatibilité... *)
+ (* It would probably be better with something in the spirit of
+ *** make_ident id (Some 0) *** but compatibility would be lost... *)
forget_subscript id
+let rec to_avoid id = function
+| [] -> false
+| id' :: avoid -> Id.equal id id' || to_avoid id avoid
+
+let occur_rel p env id =
+ try
+ let name = lookup_name_of_rel p env in
+ begin match name with
+ | Name id' -> Id.equal id' id
+ | Anonymous -> false
+ end
+ with Not_found -> false (* Unbound indice : may happen in debug *)
+
+let visibly_occur_id id (nenv,c) =
+ let rec occur n c = match kind_of_term c with
+ | Const _ | Ind _ | Construct _ | Var _
+ when
+ let short = shortest_qualid_of_global Id.Set.empty (global_of_constr c) in
+ qualid_eq short (qualid_of_ident id) ->
+ raise Occur
+ | Rel p when p>n && occur_rel (p-n) nenv id -> raise Occur
+ | _ -> iter_constr_with_binders succ occur n c
+ in
+ try occur 1 c; false
+ with Occur -> true
+ | Not_found -> false (* Happens when a global is not in the env *)
+
(* Now, there are different renaming strategies... *)
(* 1- Looks for a fresh name for printing in cases pattern *)
-let next_name_away_in_cases_pattern na avoid =
- let id = match na with Name id -> id | Anonymous -> default_x in
- next_ident_away_from id (fun id -> List.mem id avoid or is_constructor id)
+let next_name_away_in_cases_pattern env_t na avoid =
+ let id = match na with Name id -> id | Anonymous -> default_dependent_ident in
+ let bad id = to_avoid id avoid || is_constructor id
+ || visibly_occur_id id env_t in
+ next_ident_away_from id bad
(* 2- Looks for a fresh name for introduction in goal *)
@@ -173,12 +227,14 @@ let next_name_away_in_cases_pattern na avoid =
name is taken by finding a free subscript starting from 0 *)
let next_ident_away_in_goal id avoid =
- let id = if List.mem id avoid then restart_subscript id else id in
- let bad id = List.mem id avoid || (is_global id & not (is_section_variable id)) in
+ let id = if to_avoid id avoid then restart_subscript id else id in
+ let bad id = to_avoid id avoid || (is_global id && not (is_section_variable id)) in
next_ident_away_from id bad
let next_name_away_in_goal na avoid =
- let id = match na with Name id -> id | Anonymous -> id_of_string "H" in
+ let id = match na with
+ | Name id -> id
+ | Anonymous -> default_non_dependent_ident in
next_ident_away_in_goal id avoid
(* 3- Looks for next fresh name outside a list that is moreover valid
@@ -189,20 +245,20 @@ let next_name_away_in_goal na avoid =
beyond the current subscript *)
let next_global_ident_away id avoid =
- let id = if List.mem id avoid then restart_subscript id else id in
- let bad id = List.mem id avoid || is_global id in
+ let id = if to_avoid id avoid then restart_subscript id else id in
+ let bad id = to_avoid id avoid || is_global id in
next_ident_away_from id bad
(* 4- Looks for next fresh name outside a list; if name already used,
looks for same name with lower available subscript *)
let next_ident_away id avoid =
- if List.mem id avoid then
- next_ident_away_from (restart_subscript id) (fun id -> List.mem id avoid)
+ if to_avoid id avoid then
+ next_ident_away_from (restart_subscript id) (fun id -> to_avoid id avoid)
else id
let next_name_away_with_default default na avoid =
- let id = match na with Name id -> id | Anonymous -> id_of_string default in
+ let id = match na with Name id -> id | Anonymous -> Id.of_string default in
next_ident_away id avoid
let reserved_type_name = ref (fun t -> Anonymous)
@@ -213,10 +269,10 @@ let next_name_away_with_default_using_types default na avoid t =
| Name id -> id
| Anonymous -> match !reserved_type_name t with
| Name id -> id
- | Anonymous -> id_of_string default in
+ | Anonymous -> Id.of_string default in
next_ident_away id avoid
-let next_name_away = next_name_away_with_default "H"
+let next_name_away = next_name_away_with_default default_non_dependent_string
let make_all_name_different env =
let avoid = ref (ids_of_named_context (named_context env)) in
@@ -232,24 +288,8 @@ let make_all_name_different env =
looks for name of same base with lower available subscript beyond current
subscript *)
-let occur_rel p env id =
- try lookup_name_of_rel p env = Name id
- with Not_found -> false (* Unbound indice : may happen in debug *)
-
-let visibly_occur_id id (nenv,c) =
- let rec occur n c = match kind_of_term c with
- | Const _ | Ind _ | Construct _ | Var _
- when shortest_qualid_of_global Idset.empty (global_of_constr c)
- = qualid_of_ident id -> raise Occur
- | Rel p when p>n & occur_rel (p-n) nenv id -> raise Occur
- | _ -> iter_constr_with_binders succ occur n c
- in
- try occur 1 c; false
- with Occur -> true
- | Not_found -> false (* Happens when a global is not in the env *)
-
let next_ident_away_for_default_printing env_t id avoid =
- let bad id = List.mem id avoid or visibly_occur_id id env_t in
+ let bad id = to_avoid id avoid || visibly_occur_id id env_t in
next_ident_away_from id bad
let next_name_away_for_default_printing env_t na avoid =
@@ -259,7 +299,7 @@ let next_name_away_for_default_printing env_t na avoid =
(* In principle, an anonymous name is not dependent and will not be *)
(* taken into account by the function compute_displayed_name_in; *)
(* just in case, invent a valid name *)
- id_of_string "H" in
+ default_non_dependent_ident in
next_ident_away_for_default_printing env_t id avoid
(**********************************************************************)
@@ -281,29 +321,31 @@ let next_name_away_for_default_printing env_t na avoid =
*)
type renaming_flags =
- | RenamingForCasesPattern
+ | RenamingForCasesPattern of (Name.t list * constr)
| RenamingForGoal
- | RenamingElsewhereFor of (name list * constr)
+ | RenamingElsewhereFor of (Name.t list * constr)
let next_name_for_display flags =
match flags with
- | RenamingForCasesPattern -> next_name_away_in_cases_pattern
+ | RenamingForCasesPattern env_t -> next_name_away_in_cases_pattern env_t
| RenamingForGoal -> next_name_away_in_goal
| RenamingElsewhereFor env_t -> next_name_away_for_default_printing env_t
(* Remark: Anonymous var may be dependent in Evar's contexts *)
let compute_displayed_name_in flags avoid na c =
- if na = Anonymous & noccurn 1 c then
+ match na with
+ | Anonymous when noccurn 1 c ->
(Anonymous,avoid)
- else
+ | _ ->
let fresh_id = next_name_for_display flags na avoid in
let idopt = if noccurn 1 c then Anonymous else Name fresh_id in
(idopt, fresh_id::avoid)
let compute_and_force_displayed_name_in flags avoid na c =
- if na = Anonymous & noccurn 1 c then
+ match na with
+ | Anonymous when noccurn 1 c ->
(Anonymous,avoid)
- else
+ | _ ->
let fresh_id = next_name_for_display flags na avoid in
(Name fresh_id, fresh_id::avoid)
@@ -311,7 +353,7 @@ let compute_displayed_let_name_in flags avoid na c =
let fresh_id = next_name_for_display flags na avoid in
(Name fresh_id, fresh_id::avoid)
-let rec rename_bound_vars_as_displayed avoid env c =
+let rename_bound_vars_as_displayed avoid env c =
let rec rename avoid env c =
match kind_of_term c with
| Prod (na,c1,c2) ->
@@ -328,3 +370,25 @@ let rec rename_bound_vars_as_displayed avoid env c =
| _ -> c
in
rename avoid env c
+
+(**********************************************************************)
+(* "H"-based naming strategy introduced June 2014 for hypotheses in
+ Prop produced by case/elim/destruct/induction, in place of the
+ strategy that was using the first letter of the type, leading to
+ inelegant "n:~A", "e:t=u", etc. when eliminating sumbool or similar
+ types *)
+
+let h_based_elimination_names = ref false
+
+let use_h_based_elimination_names () =
+ !h_based_elimination_names && Flags.version_strictly_greater Flags.V8_4
+
+open Goptions
+
+let _ = declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "use of \"H\"-based proposition names in elimination tactics";
+ optkey = ["Standard";"Proposition";"Elimination";"Names"];
+ optread = (fun () -> !h_based_elimination_names);
+ optwrite = (:=) h_based_elimination_names }
diff --git a/pretyping/namegen.mli b/pretyping/namegen.mli
index db078026..f66bc6d8 100644
--- a/pretyping/namegen.mli
+++ b/pretyping/namegen.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,23 +8,34 @@
open Names
open Term
+open Context
open Environ
(*********************************************************************
+ Conventional default names *)
+
+val default_prop_ident : Id.t (* "H" *)
+val default_small_ident : Id.t (* "H" *)
+val default_type_ident : Id.t (* "X" *)
+val default_non_dependent_ident : Id.t (* "H" *)
+val default_dependent_ident : Id.t (* "x" *)
+
+(*********************************************************************
Generating "intuitive" names from their type *)
-val lowercase_first_char : identifier -> string
+val lowercase_first_char : Id.t -> string
val sort_hdchar : sorts -> string
val hdchar : env -> types -> string
-val id_of_name_using_hdchar : env -> types -> name -> identifier
-val named_hd : env -> types -> name -> name
+val id_of_name_using_hdchar : env -> types -> Name.t -> Id.t
+val named_hd : env -> types -> Name.t -> Name.t
+val head_name : types -> Id.t option
-val mkProd_name : env -> name * types * types -> types
-val mkLambda_name : env -> name * types * constr -> constr
+val mkProd_name : env -> Name.t * types * types -> types
+val mkLambda_name : env -> Name.t * types * constr -> constr
(** Deprecated synonyms of [mkProd_name] and [mkLambda_name] *)
-val prod_name : env -> name * types * types -> types
-val lambda_name : env -> name * types * constr -> constr
+val prod_name : env -> Name.t * types * types -> types
+val lambda_name : env -> Name.t * types * constr -> constr
val prod_create : env -> types * types -> constr
val lambda_create : env -> types * constr -> constr
@@ -40,45 +51,52 @@ val it_mkLambda_or_LetIn_name : env -> constr -> rel_context -> constr
Fresh names *)
(** Avoid clashing with a name satisfying some predicate *)
-val next_ident_away_from : identifier -> (identifier -> bool) -> identifier
+val next_ident_away_from : Id.t -> (Id.t -> bool) -> Id.t
(** Avoid clashing with a name of the given list *)
-val next_ident_away : identifier -> identifier list -> identifier
+val next_ident_away : Id.t -> Id.t list -> Id.t
(** Avoid clashing with a name already used in current module *)
-val next_ident_away_in_goal : identifier -> identifier list -> identifier
+val next_ident_away_in_goal : Id.t -> Id.t list -> Id.t
(** Avoid clashing with a name already used in current module
but tolerate overwriting section variables, as in goals *)
-val next_global_ident_away : identifier -> identifier list -> identifier
+val next_global_ident_away : Id.t -> Id.t list -> Id.t
(** Avoid clashing with a constructor name already used in current module *)
-val next_name_away_in_cases_pattern : name -> identifier list -> identifier
+val next_name_away_in_cases_pattern : (Termops.names_context * constr) -> Name.t -> Id.t list -> Id.t
-val next_name_away : name -> identifier list -> identifier (** default is "H" *)
-val next_name_away_with_default : string -> name -> identifier list ->
- identifier
+(** Default is [default_non_dependent_ident] *)
+val next_name_away : Name.t -> Id.t list -> Id.t
-val next_name_away_with_default_using_types : string -> name ->
- identifier list -> types -> identifier
+val next_name_away_with_default : string -> Name.t -> Id.t list ->
+ Id.t
-val set_reserved_typed_name : (types -> name) -> unit
+val next_name_away_with_default_using_types : string -> Name.t ->
+ Id.t list -> types -> Id.t
+
+val set_reserved_typed_name : (types -> Name.t) -> unit
(*********************************************************************
Making name distinct for displaying *)
type renaming_flags =
- | RenamingForCasesPattern (** avoid only global constructors *)
+ | RenamingForCasesPattern of (Name.t list * constr) (** avoid only global constructors *)
| RenamingForGoal (** avoid all globals (as in intro) *)
- | RenamingElsewhereFor of (name list * constr)
+ | RenamingElsewhereFor of (Name.t list * constr)
val make_all_name_different : env -> env
val compute_displayed_name_in :
- renaming_flags -> identifier list -> name -> constr -> name * identifier list
+ renaming_flags -> Id.t list -> Name.t -> constr -> Name.t * Id.t list
val compute_and_force_displayed_name_in :
- renaming_flags -> identifier list -> name -> constr -> name * identifier list
+ renaming_flags -> Id.t list -> Name.t -> constr -> Name.t * Id.t list
val compute_displayed_let_name_in :
- renaming_flags -> identifier list -> name -> constr -> name * identifier list
+ renaming_flags -> Id.t list -> Name.t -> constr -> Name.t * Id.t list
val rename_bound_vars_as_displayed :
- identifier list -> name list -> types -> types
+ Id.t list -> Name.t list -> types -> types
+
+(**********************************************************************)
+(* Naming strategy for arguments in Prop when eliminating inductive types *)
+
+val use_h_based_elimination_names : unit -> bool
diff --git a/pretyping/nativenorm.ml b/pretyping/nativenorm.ml
new file mode 100644
index 00000000..bd427ecd
--- /dev/null
+++ b/pretyping/nativenorm.ml
@@ -0,0 +1,404 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+open Pp
+open Errors
+open Term
+open Vars
+open Environ
+open Reduction
+open Univ
+open Declarations
+open Names
+open Inductive
+open Util
+open Nativecode
+open Nativevalues
+open Nativelambda
+
+(** This module implements normalization by evaluation to OCaml code *)
+
+let evars_of_evar_map evd =
+ { evars_val = Evd.existential_opt_value evd;
+ evars_typ = Evd.existential_type evd;
+ evars_metas = Evd.meta_type evd }
+
+exception Find_at of int
+
+let invert_tag cst tag reloc_tbl =
+ try
+ for j = 0 to Array.length reloc_tbl - 1 do
+ let tagj,arity = reloc_tbl.(j) in
+ if Int.equal tag tagj && (cst && Int.equal arity 0 || not(cst || Int.equal arity 0)) then
+ raise (Find_at j)
+ else ()
+ done;raise Not_found
+ with Find_at j -> (j+1)
+
+let decompose_prod env t =
+ let (name,dom,codom as res) = destProd (whd_betadeltaiota env t) in
+ match name with
+ | Anonymous -> (Name (id_of_string "x"),dom,codom)
+ | _ -> res
+
+let app_type env c =
+ let t = whd_betadeltaiota env c in
+ try destApp t with DestKO -> (t,[||])
+
+
+let find_rectype_a env c =
+ let (t, l) = app_type env c in
+ match kind_of_term t with
+ | Ind ind -> (ind, l)
+ | _ -> raise Not_found
+
+(* Instantiate inductives and parameters in constructor type *)
+
+let type_constructor mind mib typ params =
+ let s = ind_subst mind mib Univ.Instance.empty (* FIXME *)in
+ let ctyp = substl s typ in
+ let nparams = Array.length params in
+ if Int.equal nparams 0 then ctyp
+ else
+ let _,ctyp = decompose_prod_n nparams ctyp in
+ substl (List.rev (Array.to_list params)) ctyp
+
+let construct_of_constr_notnative const env tag (mind, _ as ind) u allargs =
+ let mib,mip = lookup_mind_specif env ind in
+ let nparams = mib.mind_nparams in
+ let params = Array.sub allargs 0 nparams in
+ try
+ if const then
+ let ctyp = type_constructor mind mib (mip.mind_nf_lc.(0)) params in
+ retroknowledge Retroknowledge.get_vm_decompile_constant_info env (mkInd ind) tag, ctyp
+ else
+ raise Not_found
+ with Not_found ->
+ let i = invert_tag const tag mip.mind_reloc_tbl in
+ let ctyp = type_constructor mind mib (mip.mind_nf_lc.(i-1)) params in
+ (mkApp(mkConstructU((ind,i),u), params), ctyp)
+
+
+let construct_of_constr const env tag typ =
+ let t, l = app_type env typ in
+ match kind_of_term t with
+ | Ind (ind,u) ->
+ construct_of_constr_notnative const env tag ind u l
+ | _ -> assert false
+
+let construct_of_constr_const env tag typ =
+ fst (construct_of_constr true env tag typ)
+
+let construct_of_constr_block = construct_of_constr false
+
+let build_branches_type env (mind,_ as _ind) mib mip 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 cty params in
+ let decl,indapp = Reductionops.splay_prod env Evd.empty typi 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
+ let carity = snd (rtbl.(i)) in
+ let crealargs = Array.sub cargs nparams (Array.length cargs - nparams) in
+ let codom =
+ let ndecl = List.length decl in
+ let papp = mkApp(lift ndecl p,crealargs) in
+ if dep then
+ let cstr = ith_constructor_of_inductive (fst ind) (i+1) in
+ let relargs = Array.init carity (fun i -> mkRel (carity-i)) in
+ let params = Array.map (lift ndecl) params in
+ let dep_cstr = mkApp(mkApp(mkConstructU (cstr,snd ind),params),relargs) in
+ mkApp(papp,[|dep_cstr|])
+ else papp
+ in
+ decl, decl_with_letin, codom
+ in Array.mapi build_one_branch mip.mind_nf_lc
+
+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 =
+ let (_,_,ty) = lookup_rel n env in
+ lift n ty
+
+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 =
+ try let (_,_,ty) = lookup_named id env in ty
+ 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) ->
+ begin match engagement env with
+ | Some ImpredicativeSet ->
+ (* Rule is (Type,Set,Set) in the Set-impredicative calculus *)
+ rangsort
+ | _ ->
+ (* Rule is (Type_i,Set,Type_i) in the Set-predicative calculus *)
+ Type (sup u1 type0_univ)
+ end
+ (* 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 =
+ let tbl = ans.asw_reloc in
+ let branch i =
+ let tag,arity = tbl.(i) in
+ let ci =
+ if Int.equal arity 0 then mk_const tag
+ else mk_block tag (mk_rels_accu lvl arity) in
+ bs ci in
+ Array.init (Array.length tbl) branch
+
+let rec nf_val env v typ =
+ match kind_of_value v with
+ | Vaccu accu -> nf_accu env accu
+ | Vfun f ->
+ let lvl = nb_rel env in
+ let name,dom,codom =
+ try decompose_prod env typ
+ with DestKO ->
+ Errors.anomaly
+ (Pp.strbrk "Returned a functional value in a type not recognized as a product type.")
+ in
+ let env = push_rel (name,None,dom) env in
+ let body = nf_val env (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
+ mkApp(capp,args)
+
+and nf_type env v =
+ match kind_of_value v with
+ | Vaccu accu -> nf_accu env accu
+ | _ -> assert false
+
+and nf_type_sort env v =
+ match kind_of_value v with
+ | Vaccu accu ->
+ let t,s = nf_accu_type env accu in
+ let s = try destSort s with DestKO -> assert false in
+ t, s
+ | _ -> assert false
+
+and nf_accu env accu =
+ let atom = atom_of_accu accu in
+ if Int.equal (accu_nargs accu) 0 then nf_atom env atom
+ else
+ let a,typ = nf_atom_type env atom in
+ let _, args = nf_args env accu typ in
+ mkApp(a,Array.of_list args)
+
+and nf_accu_type env accu =
+ let atom = atom_of_accu accu in
+ if Int.equal (accu_nargs accu) 0 then nf_atom_type env atom
+ else
+ let a,typ = nf_atom_type env atom in
+ let t, args = nf_args env accu typ in
+ mkApp(a,Array.of_list args), t
+
+and nf_args env accu t =
+ let aux arg (t,l) =
+ let _,dom,codom =
+ try decompose_prod env t with
+ DestKO ->
+ Errors.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
+ (subst1 c codom, c::l)
+ in
+ let t,l = List.fold_right aux (args_of_accu accu) (t,[]) in
+ t, List.rev l
+
+and nf_bargs env b t =
+ let t = ref t in
+ let len = block_size b in
+ Array.init len
+ (fun i ->
+ let _,dom,codom =
+ try decompose_prod env !t with
+ DestKO ->
+ Errors.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
+ t := subst1 c codom; c)
+
+and nf_atom env atom =
+ match atom with
+ | Arel i -> mkRel (nb_rel env - i)
+ | Aconstant cst -> mkConstU cst
+ | Aind ind -> mkIndU ind
+ | Asort s -> mkSort s
+ | Avar id -> mkVar id
+ | Aprod(n,dom,codom) ->
+ let dom = nf_type env dom in
+ let vn = mk_rel_accu (nb_rel env) in
+ let env = push_rel (n,None,dom) env in
+ let codom = nf_type env (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
+ mkProj(Projection.make p false,c)
+ | _ -> fst (nf_atom_type env atom)
+
+and nf_atom_type env atom =
+ match atom with
+ | Arel i ->
+ let n = (nb_rel env - i) in
+ mkRel n, type_of_rel 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
+ | Avar id ->
+ mkVar id, type_of_var env id
+ | Acase(ans,accu,p,bs) ->
+ let a,ta = nf_accu_type env 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
+ let params,realargs = Array.chop nparams allargs in
+ let pT =
+ hnf_prod_applist env
+ (Inductiveops.type_of_inductive env ind) (Array.to_list params) in
+ let pT = whd_betadeltaiota env pT in
+ let dep, p = nf_predicate env ind mip params p pT in
+ (* Calcul du type des branches *)
+ let btypes = build_branches_type env (fst ind) mib mip 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
+ 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 a in
+ 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 lvl = nb_rel env in
+ let nbfix = Array.length ft in
+ let fargs = mk_rels_accu lvl (Array.length ft) in
+ (* Third argument of the tuple is ignored by push_rec_types *)
+ 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 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 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
+ mkCoFix(s,(name,tt,ft)), tt.(s)
+ | Aprod(n,dom,codom) ->
+ let dom,s1 = nf_type_sort env dom in
+ let vn = mk_rel_accu (nb_rel env) in
+ let env = push_rel (n,None,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
+ | Ameta(mv,ty) ->
+ let ty = nf_type env ty in
+ mkMeta mv, ty
+ | Aproj(p,c) ->
+ let c,tc = nf_accu_type env 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
+ | Vfun f, Prod _ ->
+ let k = nb_rel env in
+ let vb = f (mk_rel_accu k) in
+ let name,dom,codom =
+ try decompose_prod env pT with
+ DestKO ->
+ Errors.anomaly
+ (Pp.strbrk "Returned a functional value in a type not recognized as a product type.")
+ in
+ let dep,body =
+ nf_predicate (push_rel (name,None,dom) env) 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 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 (name,None,dom) env) vb in
+ true, mkLambda(name,dom,body)
+ | _, _ -> false, nf_type env v
+
+let native_norm env sigma c ty =
+ if !Flags.no_native_compiler then
+ error "Native_compute reduction has been disabled"
+ else
+ let penv = Environ.pre_env env 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
+ | true, fn ->
+ if !Flags.debug then Pp.msg_debug (Pp.str "Running norm ...");
+ let t0 = Sys.time () in
+ Nativelib.call_linker ~fatal:true prefix fn (Some upd);
+ let t1 = Sys.time () in
+ let time_info = Format.sprintf "Evaluation done in %.5f@." (t1 -. t0) in
+ if !Flags.debug then Pp.msg_debug (Pp.str time_info);
+ let res = nf_val env !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 Pp.msg_debug (Pp.str time_info);
+ res
+ | _ -> anomaly (Pp.str "Compilation failure")
diff --git a/pretyping/nativenorm.mli b/pretyping/nativenorm.mli
new file mode 100644
index 00000000..c854e8c9
--- /dev/null
+++ b/pretyping/nativenorm.mli
@@ -0,0 +1,17 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+open Term
+open Environ
+open Evd
+open Nativelambda
+
+(** This module implements normalization by evaluation to OCaml code *)
+
+val evars_of_evar_map : evar_map -> evars
+
+val native_norm : env -> evars -> constr -> types -> constr
diff --git a/pretyping/pattern.mli b/pretyping/pattern.mli
deleted file mode 100644
index 7fb53133..00000000
--- a/pretyping/pattern.mli
+++ /dev/null
@@ -1,126 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(** This module defines the type of pattern used for pattern-matching
- terms in tatics *)
-
-open Pp
-open Names
-open Sign
-open Term
-open Environ
-open Libnames
-open Nametab
-open Glob_term
-open Mod_subst
-
-(** {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 = identifier list * constr
-
-(** Types of substitutions with or w/o bound variables *)
-
-type patvar_map = (patvar * constr) list
-type extended_patvar_map = (patvar * constr_under_binders) list
-
-(** {5 Patterns} *)
-
-type case_info_pattern =
- { cip_style : case_style;
- cip_ind : inductive option;
- cip_ind_args : (int * int) option; (** number of params and args *)
- cip_extensible : bool (** does this match end with _ => _ ? *) }
-
-type constr_pattern =
- | PRef of global_reference
- | PVar of identifier
- | PEvar of existential_key * constr_pattern array
- | PRel of int
- | PApp of constr_pattern * constr_pattern array
- | PSoApp of patvar * constr_pattern list
- | PLambda of name * constr_pattern * constr_pattern
- | PProd of name * constr_pattern * constr_pattern
- | PLetIn of name * constr_pattern * constr_pattern
- | PSort of glob_sort
- | PMeta of patvar option
- | PIf of constr_pattern * constr_pattern * constr_pattern
- | PCase of case_info_pattern * constr_pattern * constr_pattern *
- (int * int * constr_pattern) list (** index of constructor, nb of args *)
- | PFix of fixpoint
- | PCoFix of cofixpoint
-
-(** Nota : in a [PCase], the array of branches might be shorter than
- expected, denoting the use of a final "_ => _" branch *)
-
-(** {5 Functions on patterns} *)
-
-val occur_meta_pattern : constr_pattern -> bool
-
-val subst_pattern : substitution -> constr_pattern -> constr_pattern
-
-exception BoundPattern
-
-(** [head_pattern_bound t] extracts the head variable/constant of the
- type [t] or raises [BoundPattern] (even if a sort); it raises an anomaly
- if [t] is an abstraction *)
-
-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
-
-(** [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 : Evd.evar_map -> constr -> named_context * 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
- are bound *)
-
-val pattern_of_glob_constr : glob_constr ->
- patvar list * constr_pattern
-
-val instantiate_pattern :
- Evd.evar_map -> (identifier * (identifier list * constr)) list ->
- constr_pattern -> constr_pattern
-
-val lift_pattern : int -> constr_pattern -> constr_pattern
diff --git a/pretyping/pattern.ml b/pretyping/patternops.ml
index 553615c2..c49bec9a 100644
--- a/pretyping/pattern.ml
+++ b/pretyping/patternops.ml
@@ -1,67 +1,97 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Errors
open Util
open Names
-open Libnames
+open Globnames
open Nameops
open Term
+open Vars
open Glob_term
-open Environ
-open Nametab
+open Glob_ops
open Pp
open Mod_subst
+open Misctypes
+open Decl_kinds
+open Pattern
+open Evd
+open Environ
-(* Metavariables *)
+let case_info_pattern_eq i1 i2 =
+ i1.cip_style == i2.cip_style &&
+ Option.equal eq_ind i1.cip_ind i2.cip_ind &&
+ Option.equal (List.equal (==)) i1.cip_ind_tags i2.cip_ind_tags &&
+ i1.cip_extensible == i2.cip_extensible
-type constr_under_binders = identifier list * constr
+let rec constr_pattern_eq p1 p2 = match p1, p2 with
+| PRef r1, PRef r2 -> eq_gr r1 r2
+| PVar v1, PVar v2 -> Id.equal v1 v2
+| PEvar (ev1, ctx1), PEvar (ev2, ctx2) ->
+ Evar.equal ev1 ev2 && Array.equal constr_pattern_eq ctx1 ctx2
+| PRel i1, PRel i2 ->
+ Int.equal i1 i2
+| PApp (t1, arg1), PApp (t2, arg2) ->
+ constr_pattern_eq t1 t2 && Array.equal constr_pattern_eq arg1 arg2
+| PSoApp (id1, arg1), PSoApp (id2, arg2) ->
+ Id.equal id1 id2 && List.equal constr_pattern_eq arg1 arg2
+| PLambda (v1, t1, b1), PLambda (v2, t2, b2) ->
+ 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
+| 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) ->
+ constr_pattern_eq t1 t2 && constr_pattern_eq l1 l2 && constr_pattern_eq r1 r2
+| PCase (info1, p1, r1, l1), PCase (info2, p2, r2, l2) ->
+ case_info_pattern_eq info1 info2 &&
+ constr_pattern_eq p1 p2 &&
+ constr_pattern_eq r1 r2 &&
+ List.equal pattern_eq l1 l2
+| PFix f1, PFix f2 ->
+ fixpoint_eq f1 f2
+| PCoFix f1, PCoFix f2 ->
+ cofixpoint_eq f1 f2
+| _ -> false
+(** FIXME: fixpoint and cofixpoint should be relativized to pattern *)
-type patvar_map = (patvar * constr) list
-type extended_patvar_map = (patvar * constr_under_binders) list
+and pattern_eq (i1, j1, p1) (i2, j2, p2) =
+ Int.equal i1 i2 && List.equal (==) j1 j2 && constr_pattern_eq p1 p2
-(* Patterns *)
+and fixpoint_eq ((arg1, i1), r1) ((arg2, i2), r2) =
+ Int.equal i1 i2 &&
+ Array.equal Int.equal arg1 arg2 &&
+ rec_declaration_eq r1 r2
-type case_info_pattern =
- { cip_style : case_style;
- cip_ind : inductive option;
- cip_ind_args : (int * int) option; (** number of params and args *)
- cip_extensible : bool (** does this match end with _ => _ ? *) }
+and cofixpoint_eq (i1, r1) (i2, r2) =
+ Int.equal i1 i2 &&
+ rec_declaration_eq r1 r2
-type constr_pattern =
- | PRef of global_reference
- | PVar of identifier
- | PEvar of existential_key * constr_pattern array
- | PRel of int
- | PApp of constr_pattern * constr_pattern array
- | PSoApp of patvar * constr_pattern list
- | PLambda of name * constr_pattern * constr_pattern
- | PProd of name * constr_pattern * constr_pattern
- | PLetIn of name * constr_pattern * constr_pattern
- | PSort of glob_sort
- | PMeta of patvar option
- | PIf of constr_pattern * constr_pattern * constr_pattern
- | PCase of case_info_pattern * constr_pattern * constr_pattern *
- (int * int * constr_pattern) list (** constructor index, nb of args *)
- | PFix of fixpoint
- | PCoFix of cofixpoint
+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
let rec occur_meta_pattern = function
| PApp (f,args) ->
- (occur_meta_pattern f) or (array_exists occur_meta_pattern args)
- | PLambda (na,t,c) -> (occur_meta_pattern t) or (occur_meta_pattern c)
- | PProd (na,t,c) -> (occur_meta_pattern t) or (occur_meta_pattern c)
- | PLetIn (na,t,c) -> (occur_meta_pattern t) or (occur_meta_pattern c)
+ (occur_meta_pattern f) || (Array.exists occur_meta_pattern args)
+ | 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)
| PIf (c,c1,c2) ->
- (occur_meta_pattern c) or
- (occur_meta_pattern c1) or (occur_meta_pattern c2)
+ (occur_meta_pattern c) ||
+ (occur_meta_pattern c1) || (occur_meta_pattern c2)
| PCase(_,p,c,br) ->
- (occur_meta_pattern p) or
- (occur_meta_pattern c) or
+ (occur_meta_pattern p) ||
+ (occur_meta_pattern c) ||
(List.exists (fun (_,_,p) -> occur_meta_pattern p) br)
| PMeta _ | PSoApp _ -> true
| PEvar _ | PVar _ | PRef _ | PRel _ | PSort _ | PFix _ | PCoFix _ -> false
@@ -77,72 +107,77 @@ let rec head_pattern_bound t =
| PCase (_,p,c,br) -> head_pattern_bound c
| PRef r -> r
| PVar id -> VarRef id
+ | PProj (p,c) -> ConstRef (Projection.constant p)
| PEvar _ | PRel _ | PMeta _ | PSoApp _ | PSort _ | PFix _
-> raise BoundPattern
(* Perhaps they were arguments, but we don't beta-reduce *)
| PLambda _ -> raise BoundPattern
- | PCoFix _ -> anomaly "head_pattern_bound: 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
- | Const sp -> ConstRef sp
- | Construct sp -> ConstructRef sp
- | Ind sp -> IndRef sp
+ | Const (sp,_) -> ConstRef sp
+ | Construct (sp,_) -> ConstructRef sp
+ | Ind (sp,_) -> IndRef sp
| Var id -> VarRef id
- | _ -> anomaly "Not a rigid reference"
-
-open Evd
+ | _ -> anomaly (Pp.str "Not a rigid reference")
-let pattern_of_constr sigma t =
+let pattern_of_constr env sigma t =
let ctx = ref [] in
- let rec pattern_of_constr t =
+ let rec pattern_of_constr env t =
match kind_of_term t with
| Rel n -> PRel n
- | Meta n -> PMeta (Some (id_of_string ("META" ^ string_of_int n)))
+ | Meta n -> PMeta (Some (Id.of_string ("META" ^ string_of_int n)))
| Var id -> PVar id
- | Sort (Prop c) -> PSort (GProp c)
- | Sort (Type _) -> PSort (GType None)
- | Cast (c,_,_) -> pattern_of_constr c
- | LetIn (na,c,_,b) -> PLetIn (na,pattern_of_constr c,pattern_of_constr b)
- | Prod (na,c,b) -> PProd (na,pattern_of_constr c,pattern_of_constr b)
- | Lambda (na,c,b) -> PLambda (na,pattern_of_constr c,pattern_of_constr b)
+ | Sort (Prop Null) -> PSort GProp
+ | 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,
+ pattern_of_constr (push_rel (na,Some c,t) env) b)
+ | Prod (na,c,b) -> PProd (na,pattern_of_constr env c,
+ pattern_of_constr (push_rel (na, None, c) env) b)
+ | Lambda (na,c,b) -> PLambda (na,pattern_of_constr env c,
+ pattern_of_constr (push_rel (na, None, c) env) b)
| App (f,a) ->
- (match
+ (match
match kind_of_term f with
Evar (evk,args as ev) ->
(match snd (Evd.evar_source evk sigma) with
- MatchingVar (true,id) ->
- ctx := (id,None,existential_type sigma ev)::!ctx;
+ Evar_kinds.MatchingVar (true,id) ->
+ ctx := (id,None,Evarutil.nf_evar sigma (existential_type sigma ev))::!ctx;
Some id
| _ -> None)
| _ -> None
with
- | Some n -> PSoApp (n,Array.to_list (Array.map pattern_of_constr a))
- | None -> PApp (pattern_of_constr f,Array.map (pattern_of_constr) a))
- | Const sp -> PRef (ConstRef (constant_of_kn(canonical_con sp)))
- | Ind sp -> PRef (canonical_gr (IndRef sp))
- | Construct sp -> PRef (canonical_gr (ConstructRef sp))
+ | 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)))
+ | 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 [])
| Evar (evk,ctxt as ev) ->
(match snd (Evd.evar_source evk sigma) with
- | MatchingVar (b,id) ->
- ctx := (id,None,existential_type sigma ev)::!ctx;
+ | Evar_kinds.MatchingVar (b,id) ->
+ ctx := (id,None,Evarutil.nf_evar sigma (existential_type sigma ev))::!ctx;
assert (not b); PMeta (Some id)
- | GoalEvar -> PEvar (evk,Array.map pattern_of_constr ctxt)
+ | Evar_kinds.GoalEvar -> PEvar (evk,Array.map (pattern_of_constr env) ctxt)
| _ -> PMeta None)
| Case (ci,p,a,br) ->
let cip =
{ cip_style = ci.ci_pp_info.style;
cip_ind = Some ci.ci_ind;
- cip_ind_args = Some (ci.ci_npar, ci.ci_pp_info.ind_nargs);
+ cip_ind_tags = Some ci.ci_pp_info.ind_tags;
cip_extensible = false }
in
let branch_of_constr i c =
- (i, ci.ci_cstr_ndecls.(i), pattern_of_constr c)
+ (i, ci.ci_pp_info.cstr_tags.(i), pattern_of_constr env c)
in
- PCase (cip, pattern_of_constr p, pattern_of_constr a,
+ PCase (cip, pattern_of_constr env p, pattern_of_constr env a,
Array.to_list (Array.mapi branch_of_constr br))
| Fix f -> PFix f
| CoFix f -> PCoFix f in
- let p = pattern_of_constr t in
+ let p = pattern_of_constr env t in
(* side-effect *)
(* Warning: the order of dependencies in ctx is not ensured *)
(!ctx,p)
@@ -158,32 +193,39 @@ let map_pattern_with_binders g f l = function
| 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)
+ | PProj (p,pc) -> PProj (p, f l pc)
(* Non recursive *)
| (PVar _ | PEvar _ | PRel _ | PRef _ | PSort _ | PMeta _
(* Bound to terms *)
| PFix _ | PCoFix _ as x) -> x
let error_instantiate_pattern id l =
- let is = if List.length l = 1 then "is" else "are" in
+ let is = match l with
+ | [_] -> "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
++ strbrk " which " ++ str is ++ strbrk " not bound in the pattern.")
-let instantiate_pattern sigma lvar c =
+let instantiate_pattern env sigma lvar c =
let rec aux vars = function
| PVar id as x ->
(try
- let ctx,c = List.assoc id lvar in
+ let ctx,c = Id.Map.find id lvar in
try
let inst =
- List.map (fun id -> mkRel (list_index (Name id) vars)) ctx in
+ List.map
+ (fun id -> mkRel (List.index Name.equal (Name id) vars))
+ ctx
+ in
let c = substl inst c in
- snd (pattern_of_constr sigma c)
- with Not_found (* list_index failed *) ->
+ snd (pattern_of_constr env sigma 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 ctx vars)
- with Not_found (* List.assoc failed *) ->
+ 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.")
| c ->
@@ -203,17 +245,23 @@ let rec subst_pattern subst pat =
| PRef ref ->
let ref',t = subst_global subst ref in
if ref' == ref then pat else
- snd (pattern_of_constr Evd.empty t)
+ snd (pattern_of_constr (Global.env()) Evd.empty t)
| PVar _
| PEvar _
| PRel _ -> pat
+ | PProj (p,c) ->
+ let p' = Projection.map (fun p ->
+ destConstRef (fst (subst_global subst (ConstRef p)))) p in
+ let c' = subst_pattern subst c in
+ if p' == p && c' == c then pat else
+ PProj(p',c')
| PApp (f,args) ->
let f' = subst_pattern subst f in
- let args' = array_smartmap (subst_pattern subst) args in
+ let args' = Array.smartmap (subst_pattern subst) args in
if f' == f && args' == args then pat else
PApp (f',args')
| PSoApp (i,args) ->
- let args' = list_smartmap (subst_pattern subst) args in
+ let args' = List.smartmap (subst_pattern subst) args in
if args' == args then pat else
PSoApp (i,args')
| PLambda (name,c1,c2) ->
@@ -241,7 +289,7 @@ let rec subst_pattern subst pat =
PIf (c',c1',c2')
| PCase (cip,typ,c,branches) ->
let ind = cip.cip_ind in
- let ind' = Option.smartmap (Inductiveops.subst_inductive subst) ind in
+ let ind' = Option.smartmap (subst_ind subst) ind in
let cip' = if ind' == ind then cip else { cip with cip_ind = ind' } in
let typ' = subst_pattern subst typ in
let c' = subst_pattern subst c in
@@ -249,7 +297,7 @@ let rec subst_pattern subst pat =
let c' = subst_pattern subst c in
if c' == c then br else (i,n,c')
in
- let branches' = list_smartmap subst_branch branches in
+ let branches' = List.smartmap subst_branch branches in
if cip' == cip && typ' == typ && c' == c && branches' == branches
then pat
else PCase(cip', typ', c', branches')
@@ -271,13 +319,13 @@ let err loc pp = user_err_loc (loc,"pattern_of_glob_constr", pp)
let rec pat_of_raw metas vars = function
| GVar (_,id) ->
- (try PRel (list_index (Name id) vars)
+ (try PRel (List.index Name.equal (Name id) vars)
with Not_found -> PVar id)
| GPatVar (_,(false,n)) ->
metas := n::!metas; PMeta (Some n)
- | GRef (_,gr) ->
+ | GRef (_,gr,_) ->
PRef (canonical_gr gr)
- (* Hack pour ne pas réécrire une interprétation complète des patterns*)
+ (* Hack to avoid rewriting a complete interpretation of patterns *)
| GApp (_, GPatVar (_,(true,n)), cl) ->
metas := n::!metas; PSoApp (n, List.map (pat_of_raw metas vars) cl)
| GApp (_,c,cl) ->
@@ -300,43 +348,45 @@ let rec pat_of_raw metas vars = function
| GHole _ ->
PMeta None
| GCast (_,c,_) ->
- Flags.if_warn
- Pp.msg_warning (str "Cast not taken into account in constr pattern");
+ Pp.msg_warning (strbrk "Cast not taken into account in constr pattern");
pat_of_raw metas vars c
| 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,Evd.InternalHole),c) in
+ 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
let cip =
{ cip_style = LetStyle;
cip_ind = None;
- cip_ind_args = None;
+ cip_ind_tags = None;
cip_extensible = false }
in
+ 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,1,pat_of_raw metas vars c])
+ [0,tags,pat_of_raw metas vars c])
| GCases (loc,sty,p,[c,(na,indnames)],brs) ->
let get_ind = function
| (_,_,[PatCstr(_,(ind,_),_,_)],_)::_ -> Some ind
| _ -> None
in
- let ind_nargs,ind = match indnames with
- | Some (_,ind,n,nal) -> Some (n,List.length nal), Some ind
+ let ind_tags,ind = match indnames with
+ | Some (_,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) ->
- rev_it_mkPLambda nal (mkPLambda na (pat_of_raw metas vars p))
+ | Some p, Some (_,_,nal) ->
+ let nvars = na :: List.rev nal @ vars in
+ rev_it_mkPLambda nal (mkPLambda na (pat_of_raw metas nvars p))
| _ -> PMeta None
in
let info =
{ cip_style = sty;
cip_ind = ind;
- cip_ind_args = ind_nargs;
+ cip_ind_tags = None;
cip_extensible = ext }
in
(* Nota : when we have a non-trivial predicate,
@@ -355,21 +405,25 @@ and pats_of_glob_branches loc metas vars ind brs =
| [] -> false, []
| [(_,_,[PatVar(_,Anonymous)],GHole _)] -> true, [] (* ends with _ => _ *)
| (_,_,[PatCstr(_,(indsp,j),lv,_)],br) :: brs ->
- if ind <> None && ind <> Some indsp then
- err loc (Pp.str "All constructors must be in the same inductive type.");
- if Intset.mem (j-1) indexes then
+ 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 (Intset.add (j-1) indexes) brs in
- ext, ((j-1, List.length lv, pat) :: pats)
+ 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.")
in
- get_pat Intset.empty brs
+ get_pat Int.Set.empty brs
let pattern_of_glob_constr c =
let metas = ref [] in
- let p = pat_of_raw metas [] c in
+ let p = pat_of_raw metas [] c in
(!metas,p)
diff --git a/pretyping/patternops.mli b/pretyping/patternops.mli
new file mode 100644
index 00000000..cf02421c
--- /dev/null
+++ b/pretyping/patternops.mli
@@ -0,0 +1,55 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Context
+open Term
+open Globnames
+open Glob_term
+open Mod_subst
+open Misctypes
+open Pattern
+
+(** {5 Functions on patterns} *)
+
+val constr_pattern_eq : constr_pattern -> constr_pattern -> bool
+
+val occur_meta_pattern : constr_pattern -> bool
+
+val subst_pattern : substitution -> constr_pattern -> constr_pattern
+
+exception BoundPattern
+
+(** [head_pattern_bound t] extracts the head variable/constant of the
+ type [t] or raises [BoundPattern] (even if a sort); it raises an anomaly
+ if [t] is an abstraction *)
+
+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
+
+(** [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 -> named_context * 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
+ are bound *)
+
+val pattern_of_glob_constr : glob_constr ->
+ patvar list * constr_pattern
+
+val instantiate_pattern : Environ.env ->
+ Evd.evar_map -> extended_patvar_map ->
+ constr_pattern -> constr_pattern
+
+val lift_pattern : int -> constr_pattern -> constr_pattern
diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml
index 52122974..030b4a11 100644
--- a/pretyping/pretype_errors.ml
+++ b/pretyping/pretype_errors.ml
@@ -1,100 +1,68 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Compat
open Util
open Names
-open Sign
open Term
-open Termops
-open Namegen
open Environ
open Type_errors
-open Glob_term
-open Inductiveops
+
+type unification_error =
+ | OccurCheck of existential_key * 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
+ | UnifUnivInconsistency of Univ.univ_inconsistency
+ | CannotSolveConstraint of Evd.evar_constraint * unification_error
+
+type position = (Id.t * Locus.hyp_location_flag) option
+
+type position_reporting = (position * int) * constr
+
+type subterm_unification_error = bool * position_reporting * position_reporting * (constr * constr * unification_error) option
type pretype_error =
(* Old Case *)
| CantFindCaseType of constr
- (* Unification *)
- | OccurCheck of existential_key * constr
- | NotClean of existential_key * constr * Evd.hole_kind
- | UnsolvableImplicit of Evd.evar_info * Evd.hole_kind *
- Evd.unsolvability_explanation option
- | CannotUnify of constr * 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
+ | CannotUnify of constr * constr * unification_error option
| CannotUnifyLocal of constr * constr * constr
| CannotUnifyBindingType of constr * constr
| CannotGeneralize of constr
- | NoOccurrenceFound of constr * identifier option
- | CannotFindWellTypedAbstraction of constr * constr list
- | AbstractionOverMeta of name * name
- | NonLinearUnification of name * constr
+ | NoOccurrenceFound of constr * Id.t option
+ | CannotFindWellTypedAbstraction of constr * constr list * (env * type_error) option
+ | WrongAbstractionType of Name.t * constr * types * types
+ | AbstractionOverMeta of Name.t * Name.t
+ | NonLinearUnification of Name.t * constr
(* Pretyping *)
- | VarNotFound of identifier
+ | VarNotFound of Id.t
| UnexpectedType of constr * constr
| NotProduct of constr
| TypingError of type_error
+ | CannotUnifyOccurrences of subterm_unification_error
+ | UnsatisfiableConstraints of
+ (existential_key * Evar_kinds.t) option * Evar.Set.t option
exception PretypeError of env * Evd.evar_map * pretype_error
let precatchable_exception = function
- | Util.UserError _ | TypeError _ | PretypeError _
- | Loc.Exc_located(_,(Util.UserError _ | TypeError _ |
- Nametab.GlobalizationError _ | PretypeError _)) -> true
+ | Errors.UserError _ | TypeError _ | PretypeError _
+ | Nametab.GlobalizationError _ -> true
| _ -> false
-let nf_evar = Reductionops.nf_evar
-let j_nf_evar sigma j =
- { uj_val = nf_evar sigma j.uj_val;
- uj_type = nf_evar sigma j.uj_type }
-let j_nf_betaiotaevar sigma j =
- { uj_val = nf_evar sigma j.uj_val;
- uj_type = Reductionops.nf_betaiota sigma j.uj_type }
-let jl_nf_evar sigma jl = List.map (j_nf_evar sigma) jl
-let jv_nf_betaiotaevar sigma jl =
- Array.map (j_nf_betaiotaevar sigma) jl
-let jv_nf_evar sigma = Array.map (j_nf_evar sigma)
-let tj_nf_evar sigma {utj_val=v;utj_type=t} =
- {utj_val=nf_evar sigma v;utj_type=t}
-
-let env_nf_evar sigma env =
- process_rel_context
- (fun d e -> push_rel (map_rel_declaration (nf_evar sigma) d) e) env
-
-let env_nf_betaiotaevar sigma env =
- process_rel_context
- (fun d e ->
- push_rel (map_rel_declaration (Reductionops.nf_betaiota sigma) d) e) env
-
-(* This simplifies the typing context of Cases clauses *)
-(* hope it does not disturb other typing contexts *)
-let contract env lc =
- let l = ref [] in
- let contract_context (na,c,t) env =
- match c with
- | Some c' when isRel c' ->
- l := (substl !l c') :: !l;
- env
- | _ ->
- let t' = substl !l t in
- let c' = Option.map (substl !l) c in
- let na' = named_hd env t' na in
- l := (mkRel 1) :: List.map (lift 1) !l;
- push_rel (na',c',t') env in
- let env = process_rel_context contract_context env in
- (env, List.map (substl !l) lc)
-
-let contract2 env a b = match contract env [a;b] with
- | env, [a;b] -> env,a,b | _ -> assert false
-
-let contract3 env a b c = match contract env [a;b;c] with
- | env, [a;b;c] -> env,a,b,c | _ -> assert false
-
let raise_pretype_error (loc,env,sigma,te) =
Loc.raise loc (PretypeError(env,sigma,te))
@@ -102,10 +70,10 @@ let raise_located_type_error (loc,env,sigma,te) =
Loc.raise loc (PretypeError(env,sigma,TypingError te))
-let error_actual_type_loc loc env sigma {uj_val=c;uj_type=actty} expty =
- let env, c, actty, expty = contract3 env c actty expty in
+let error_actual_type_loc loc env sigma {uj_val=c;uj_type=actty} expty reason =
let j = {uj_val=c;uj_type=actty} in
- raise_located_type_error (loc, env, sigma, ActualType (j, expty))
+ raise_pretype_error
+ (loc, env, sigma, ActualTypeNotCoercible (j, expty, reason))
let error_cant_apply_not_functional_loc loc env sigma rator randl =
raise_located_type_error
@@ -137,26 +105,29 @@ let error_not_a_type_loc loc env sigma j =
a precise location. *)
let error_occur_check env sigma ev c =
- raise (PretypeError (env, sigma, OccurCheck (ev,c)))
-
-let error_not_clean env sigma ev c (loc,k) =
- Loc.raise loc (PretypeError (env, sigma, NotClean (ev,c,k)))
+ raise (PretypeError (env, sigma, UnifOccurCheck (ev,c)))
-let error_unsolvable_implicit loc env sigma evi e explain =
+let error_unsolvable_implicit loc env sigma evk explain =
Loc.raise loc
- (PretypeError (env, sigma, UnsolvableImplicit (evi, e, explain)))
+ (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 (m,n) =
- raise (PretypeError (env, sigma,CannotUnify (m,n)))
+let error_cannot_unify env sigma ?reason (m,n) =
+ raise (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)))
let error_cannot_coerce env sigma (m,n) =
- raise (PretypeError (env, sigma,CannotUnify (m,n)))
+ raise (PretypeError (env, sigma,CannotUnify (m,n,None)))
-let error_cannot_find_well_typed_abstraction env sigma p l =
- raise (PretypeError (env, sigma,CannotFindWellTypedAbstraction (p,l)))
+let error_cannot_find_well_typed_abstraction env sigma p l e =
+ raise (PretypeError (env, sigma,CannotFindWellTypedAbstraction (p,l,e)))
+
+let error_wrong_abstraction_type env sigma na a p l =
+ raise (PretypeError (env, sigma,WrongAbstractionType (na,a,p,l)))
let error_abstraction_over_meta env sigma hdmeta metaarg =
let m = Evd.meta_name sigma hdmeta and n = Evd.meta_name sigma metaarg in
@@ -174,7 +145,6 @@ let error_cant_find_case_type_loc loc env sigma expr =
(*s Pretyping errors *)
let error_unexpected_type_loc loc env sigma actty expty =
- let env, actty, expty = contract2 env actty expty in
raise_pretype_error (loc, env, sigma, UnexpectedType (actty, expty))
let error_not_product_loc loc env sigma c =
@@ -184,3 +154,20 @@ let error_not_product_loc loc env sigma c =
let error_var_not_found_loc loc s =
raise_pretype_error (loc, empty_env, Evd.empty, VarNotFound s)
+
+(*s Typeclass errors *)
+
+let unsatisfiable_constraints env evd ev comp =
+ match ev with
+ | None ->
+ let err = UnsatisfiableConstraints (None, comp) in
+ raise (PretypeError (env,evd,err))
+ | 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))
+
+let unsatisfiable_exception exn =
+ match exn with
+ | PretypeError (_, _, UnsatisfiableConstraints _) -> true
+ | _ -> false
diff --git a/pretyping/pretype_errors.mli b/pretyping/pretype_errors.mli
index 75962d11..8fcfb59b 100644
--- a/pretyping/pretype_errors.mli
+++ b/pretyping/pretype_errors.mli
@@ -1,91 +1,97 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Pp
-open Util
open Names
open Term
-open Sign
open Environ
-open Glob_term
-open Inductiveops
+open Type_errors
(** {6 The type of errors raised by the pretyper } *)
+type unification_error =
+ | OccurCheck of existential_key * 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
+ | UnifUnivInconsistency of Univ.univ_inconsistency
+ | CannotSolveConstraint of Evd.evar_constraint * unification_error
+
+type position = (Id.t * Locus.hyp_location_flag) option
+
+type position_reporting = (position * int) * constr
+
+type subterm_unification_error = bool * position_reporting * position_reporting * (constr * constr * unification_error) option
+
type pretype_error =
(** Old Case *)
| CantFindCaseType of constr
- (** Unification *)
- | OccurCheck of existential_key * constr
- | NotClean of existential_key * constr * Evd.hole_kind
- | UnsolvableImplicit of Evd.evar_info * Evd.hole_kind *
- Evd.unsolvability_explanation option
- | CannotUnify of constr * 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
+ | CannotUnify of constr * constr * unification_error option
| CannotUnifyLocal of constr * constr * constr
| CannotUnifyBindingType of constr * constr
| CannotGeneralize of constr
- | NoOccurrenceFound of constr * identifier option
- | CannotFindWellTypedAbstraction of constr * constr list
- | AbstractionOverMeta of name * name
- | NonLinearUnification of name * constr
+ | NoOccurrenceFound of constr * Id.t option
+ | CannotFindWellTypedAbstraction of constr * constr list * (env * type_error) option
+ | WrongAbstractionType of Name.t * constr * types * types
+ | AbstractionOverMeta of Name.t * Name.t
+ | NonLinearUnification of Name.t * constr
(** Pretyping *)
- | VarNotFound of identifier
+ | VarNotFound of Id.t
| UnexpectedType of constr * constr
| NotProduct of constr
- | TypingError of Type_errors.type_error
+ | TypingError of type_error
+ | CannotUnifyOccurrences of subterm_unification_error
+ | UnsatisfiableConstraints of
+ (existential_key * Evar_kinds.t) option * Evar.Set.t option
+ (** unresolvable evar, connex component *)
exception PretypeError of env * Evd.evar_map * pretype_error
val precatchable_exception : exn -> bool
-(** Presenting terms without solved evars *)
-val nf_evar : Evd.evar_map -> constr -> constr
-val j_nf_evar : Evd.evar_map -> unsafe_judgment -> unsafe_judgment
-val jl_nf_evar : Evd.evar_map -> unsafe_judgment list -> unsafe_judgment list
-val jv_nf_evar : Evd.evar_map -> unsafe_judgment array -> unsafe_judgment array
-val tj_nf_evar : Evd.evar_map -> unsafe_type_judgment -> unsafe_type_judgment
-
-val env_nf_evar : Evd.evar_map -> env -> env
-val env_nf_betaiotaevar : Evd.evar_map -> env -> env
-
-val j_nf_betaiotaevar : Evd.evar_map -> unsafe_judgment -> unsafe_judgment
-val jv_nf_betaiotaevar :
- Evd.evar_map -> unsafe_judgment array -> unsafe_judgment array
-
(** Raising errors *)
val error_actual_type_loc :
- loc -> env -> Evd.evar_map -> unsafe_judgment -> constr -> 'b
+ Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> constr ->
+ unification_error -> 'b
val error_cant_apply_not_functional_loc :
- loc -> env -> Evd.evar_map ->
+ Loc.t -> env -> Evd.evar_map ->
unsafe_judgment -> unsafe_judgment list -> 'b
val error_cant_apply_bad_type_loc :
- loc -> env -> Evd.evar_map -> int * constr * constr ->
+ Loc.t -> env -> Evd.evar_map -> int * constr * constr ->
unsafe_judgment -> unsafe_judgment list -> 'b
val error_case_not_inductive_loc :
- loc -> env -> Evd.evar_map -> unsafe_judgment -> 'b
+ Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> 'b
val error_ill_formed_branch_loc :
- loc -> env -> Evd.evar_map ->
- constr -> constructor -> constr -> constr -> 'b
+ Loc.t -> env -> Evd.evar_map ->
+ constr -> pconstructor -> constr -> constr -> 'b
val error_number_branches_loc :
- loc -> env -> Evd.evar_map ->
+ Loc.t -> env -> Evd.evar_map ->
unsafe_judgment -> int -> 'b
val error_ill_typed_rec_body_loc :
- loc -> env -> Evd.evar_map ->
- int -> name array -> unsafe_judgment array -> types array -> 'b
+ Loc.t -> env -> Evd.evar_map ->
+ int -> Name.t array -> unsafe_judgment array -> types array -> 'b
val error_not_a_type_loc :
- loc -> env -> Evd.evar_map -> unsafe_judgment -> 'b
+ Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> 'b
val error_cannot_coerce : env -> Evd.evar_map -> constr * constr -> 'b
@@ -93,19 +99,23 @@ val error_cannot_coerce : env -> Evd.evar_map -> constr * constr -> 'b
val error_occur_check : env -> Evd.evar_map -> existential_key -> constr -> 'b
-val error_not_clean :
- env -> Evd.evar_map -> existential_key -> constr -> loc * Evd.hole_kind -> 'b
-
val error_unsolvable_implicit :
- loc -> env -> Evd.evar_map -> Evd.evar_info -> Evd.hole_kind ->
+ Loc.t -> env -> Evd.evar_map -> existential_key ->
Evd.unsolvability_explanation option -> 'b
-val error_cannot_unify : env -> Evd.evar_map -> constr * constr -> 'b
+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 ->
- constr -> constr list -> 'b
+ constr -> constr list -> (env * type_error) option -> 'b
+
+val error_wrong_abstraction_type : env -> Evd.evar_map ->
+ Name.t -> constr -> types -> types -> 'b
val error_abstraction_over_meta : env -> Evd.evar_map ->
metavariable -> metavariable -> 'b
@@ -116,16 +126,24 @@ val error_non_linear_unification : env -> Evd.evar_map ->
(** {6 Ml Case errors } *)
val error_cant_find_case_type_loc :
- loc -> env -> Evd.evar_map -> constr -> 'b
+ Loc.t -> env -> Evd.evar_map -> constr -> 'b
(** {6 Pretyping errors } *)
val error_unexpected_type_loc :
- loc -> env -> Evd.evar_map -> constr -> constr -> 'b
+ Loc.t -> env -> Evd.evar_map -> constr -> constr -> 'b
val error_not_product_loc :
- loc -> env -> Evd.evar_map -> constr -> 'b
+ Loc.t -> env -> Evd.evar_map -> constr -> 'b
(** {6 Error in conversion from AST to glob_constr } *)
-val error_var_not_found_loc : loc -> identifier -> '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 ->
+ Evar.Set.t option -> 'a
+
+val unsatisfiable_exception : exn -> bool
+
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index 08e8df05..040792ef 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -21,39 +21,44 @@
(* Secondary maintenance: collective *)
-open Compat
open Pp
+open Errors
open Util
open Names
-open Sign
open Evd
open Term
+open Vars
+open Context
open Termops
open Reductionops
open Environ
open Type_errors
open Typeops
-open Libnames
+open Globnames
open Nameops
-open Classops
-open List
-open Recordops
open Evarutil
open Pretype_errors
open Glob_term
+open Glob_ops
open Evarconv
open Pattern
-
-type typing_constraint = OfType of types option | IsType
-type var_map = (identifier * constr_under_binders) list
-type unbound_ltac_var_map = (identifier * identifier option) list
-type ltac_var_map = var_map * unbound_ltac_var_map
+open Misctypes
+
+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 = Genarg.tlevel Genarg.generic_argument 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
(************************************************************************)
(* This concerns Cases *)
-open Declarations
open Inductive
open Inductiveops
@@ -66,12 +71,15 @@ exception Found of int array
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. *)
- if List.for_all (fun l->1=List.length l) possible_indexes then
+ let is_singleton = function [_] -> true | _ -> false in
+ if List.for_all is_singleton possible_indexes then
let indexes = Array.of_list (List.map List.hd possible_indexes) in
let fix = ((indexes, 0),fixdefs) in
(try check_fix env fix
- with e when Errors.noncritical e ->
- if loc = dummy_loc then raise e else Loc.raise loc e);
+ with reraise ->
+ let (e, info) = Errors.push reraise in
+ let info = Loc.add_loc info loc in
+ iraise (e, info));
indexes
else
(* we now search recursively amoungst all combinations *)
@@ -82,9 +90,8 @@ let search_guard loc env possible_indexes fixdefs =
let fix = ((indexes, 0),fixdefs) in
try check_fix env fix; raise (Found indexes)
with TypeError _ -> ())
- (list_combinations possible_indexes);
+ (List.combinations possible_indexes);
let errmsg = "Cannot guess decreasing argument of fix." in
- if loc = dummy_loc then error errmsg else
user_err_loc (loc,"search_guard", Pp.str errmsg)
with Found indexes -> indexes)
@@ -93,704 +100,947 @@ let ((constr_in : constr -> Dyn.t),
(constr_out : Dyn.t -> constr)) = Dyn.create "constr"
(** Miscellaneous interpretation functions *)
-
-let interp_sort = function
- | GProp c -> Prop c
- | GType _ -> new_Type_sort ()
+let interp_universe_level_name evd s =
+ let names, _ = Universes.global_universe_names () in
+ try
+ let id = try Id.of_string s with _ -> raise Not_found in
+ evd, Idmap.find id names
+ with Not_found ->
+ try let level = Evd.universe_of_name evd s in
+ evd, level
+ with Not_found ->
+ new_univ_level_variable ~name:s univ_rigid evd
+
+let interp_universe evd = function
+ | [] -> let evd, l = new_univ_level_variable 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)))
+ (evd, Univ.Universe.type0m) l
+
+let interp_universe_level evd = function
+ | None -> new_univ_level_variable univ_rigid evd
+ | Some s -> interp_universe_level_name evd s
+
+let interp_sort evd = function
+ | GProp -> evd, Prop Null
+ | GSet -> evd, Prop Pos
+ | GType n ->
+ let evd, u = interp_universe evd n in
+ evd, Type u
let interp_elimination_sort = function
- | GProp Null -> InProp
- | GProp Pos -> InSet
+ | GProp -> InProp
+ | GSet -> InSet
| GType _ -> InType
-let resolve_evars env evdref fail_evar resolve_classes =
- if resolve_classes then
- evdref := (Typeclasses.resolve_typeclasses ~filter:Typeclasses.no_goals
- ~split:true ~fail:fail_evar env !evdref);
+type inference_flags = {
+ use_typeclasses : bool;
+ use_unif_heuristics : bool;
+ use_hook : (env -> evar_map -> evar -> constr) option;
+ fail_evar : bool;
+ expand_evars : bool
+}
+
+let pending_holes (sigma, sigma') =
+ let fold evk _ accu =
+ if not (Evd.mem sigma evk) then Evar.Set.add evk accu else accu
+ in
+ Evd.fold_undefined fold sigma' Evar.Set.empty
+
+let apply_typeclasses env evdref pending fail_evar =
+ let filter_pending evk = Evar.Set.mem evk pending in
+ evdref := Typeclasses.resolve_typeclasses
+ ~filter:(if Flags.is_program_mode ()
+ then (fun evk evi -> Typeclasses.no_goals_or_obligations evk evi && filter_pending evk)
+ else (fun evk evi -> Typeclasses.no_goals evk evi && filter_pending evk))
+ ~split:true ~fail:fail_evar env !evdref;
+ if Flags.is_program_mode () then (* Try optionally solving the obligations *)
+ evdref := Typeclasses.resolve_typeclasses
+ ~filter:(fun evk evi -> Typeclasses.all_evars evk evi && filter_pending evk) ~split:true ~fail:false env !evdref
+
+let apply_inference_hook hook evdref pending =
+ evdref := Evar.Set.fold (fun evk sigma ->
+ if Evd.is_undefined sigma evk (* in particular not defined by side-effect *)
+ then
+ try
+ let c = hook sigma evk in
+ Evd.define evk c sigma
+ with Exit ->
+ sigma
+ else
+ sigma) pending !evdref
+
+let apply_heuristics env evdref fail_evar =
(* Resolve eagerly, potentially making wrong choices *)
- evdref := (try consider_remaining_unif_problems
- ~ts:(Typeclasses.classes_transparent_state ()) env !evdref
- with e when Errors.noncritical e ->
- if fail_evar then raise e else !evdref)
-
-let solve_remaining_evars fail_evar use_classes hook env initial_sigma (evd,c) =
- let evdref = ref evd in
- resolve_evars env evdref fail_evar use_classes;
- let rec proc_rec c =
- let c = Reductionops.whd_evar !evdref c in
- match kind_of_term c with
- | Evar (evk,args as ev) when not (Evd.mem initial_sigma evk) ->
- let sigma = !evdref in
- (try
- let c = hook env sigma ev in
- evdref := Evd.define evk c !evdref;
- c
- with Exit ->
- if fail_evar then
- let evi = Evd.find_undefined sigma evk in
- let (loc,src) = evar_source evk !evdref in
- Pretype_errors.error_unsolvable_implicit loc env sigma evi src None
- else
- c)
- | _ -> map_constr proc_rec c in
- let c = proc_rec c in
- (* Side-effect *)
- !evdref,c
-
-module type S =
-sig
-
- module Cases : Cases.S
-
- (* Allow references to syntaxically inexistent variables (i.e., if applied on an inductive) *)
- val allow_anonymous_refs : bool ref
-
- (* Generic call to the interpreter from glob_constr to open_constr, leaving
- unresolved holes as evars and returning the typing contexts of
- these evars. Work as [understand_gen] for the rest. *)
-
- val understand_tcc : ?resolve_classes:bool ->
- evar_map -> env -> ?expected_type:types -> glob_constr -> open_constr
-
- val understand_tcc_evars : ?fail_evar:bool -> ?resolve_classes:bool ->
- evar_map ref -> env -> typing_constraint -> glob_constr -> constr
-
- (* More general entry point with evars from ltac *)
-
- (* Generic call to the interpreter from glob_constr to constr, failing
- unresolved holes in the glob_constr cannot be instantiated.
-
- In [understand_ltac expand_evars sigma env ltac_env constraint c],
-
- resolve_classes : launch typeclass resolution after typechecking.
- expand_evars : expand inferred evars by their value if any
- sigma : initial set of existential variables (typically dependent subgoals)
- ltac_env : partial substitution of variables (used for the tactic language)
- constraint : tell if interpreted as a possibly constrained term or a type
- *)
-
- val understand_ltac : ?resolve_classes:bool ->
- bool -> evar_map -> env -> ltac_var_map ->
- typing_constraint -> glob_constr -> pure_open_constr
-
- (* Standard call to get a constr from a glob_constr, resolving implicit args *)
-
- val understand : evar_map -> env -> ?expected_type:Term.types ->
- glob_constr -> constr
-
- (* Idem but the glob_constr is intended to be a type *)
-
- val understand_type : evar_map -> env -> glob_constr -> constr
-
- (* A generalization of the two previous case *)
-
- val understand_gen : typing_constraint -> evar_map -> env ->
- glob_constr -> constr
-
- (* Idem but returns the judgment of the understood term *)
-
- val understand_judgment : evar_map -> env -> glob_constr -> unsafe_judgment
-
- (* Idem but do not fail on unresolved evars *)
-
- val understand_judgment_tcc : evar_map ref -> env -> glob_constr -> unsafe_judgment
-
- (*i*)
- (* Internal of Pretyping...
- * Unused outside, but useful for debugging
- *)
- val pretype :
- bool -> type_constraint -> env -> evar_map ref ->
- ltac_var_map -> glob_constr -> unsafe_judgment
-
- val pretype_type :
- bool -> val_constraint -> env -> evar_map ref ->
- ltac_var_map -> glob_constr -> unsafe_type_judgment
-
- val pretype_gen :
- bool -> bool -> bool -> evar_map ref -> env ->
- ltac_var_map -> typing_constraint -> glob_constr -> constr
-
- (*i*)
-end
-
-module Pretyping_F (Coercion : Coercion.S) = struct
-
- module Cases = Cases.Cases_F(Coercion)
-
- (* Allow references to syntaxically inexistent variables (i.e., if applied on an inductive) *)
- let allow_anonymous_refs = ref false
-
- let evd_comb0 f evdref =
- let (evd',x) = f !evdref in
- evdref := evd';
- x
-
- let evd_comb1 f evdref x =
- let (evd',y) = f !evdref x in
- evdref := evd';
- y
-
- let evd_comb2 f evdref x y =
- let (evd',z) = f !evdref x y in
- evdref := evd';
- z
-
- let evd_comb3 f evdref x y z =
- let (evd',t) = f !evdref x y z in
- evdref := evd';
- t
-
- let mt_evd = Evd.empty
-
- (* Utilisé pour inférer le prédicat des Cases *)
- (* Semble exagérement fort *)
- (* Faudra préférer une unification entre les types de toutes les clauses *)
- (* et autoriser des ? à rester dans le résultat de l'unification *)
-
- let evar_type_fixpoint loc env evdref lna lar vdefj =
- let lt = Array.length vdefj in
- if Array.length lar = lt then
- for i = 0 to lt-1 do
- if not (e_cumul env evdref (vdefj.(i)).uj_type
- (lift lt lar.(i))) then
- error_ill_typed_rec_body_loc loc env !evdref
- i lna vdefj lar
- done
-
- (* coerce to tycon if any *)
- let inh_conv_coerce_to_tycon resolve_tc loc env evdref j = function
- | None -> j
- | Some t -> evd_comb2 (Coercion.inh_conv_coerce_to resolve_tc loc env) evdref j t
-
- let push_rels vars env = List.fold_right push_rel vars env
-
- (* used to enforce a name in Lambda when the type constraints itself
- is named, hence possibly dependent *)
-
- let orelse_name name name' = match name with
- | Anonymous -> name'
- | _ -> name
-
- let invert_ltac_bound_name env id0 id =
- 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 ++
- str " which is not bound in current context.")
-
- let protected_get_type_of env sigma c =
- try Retyping.get_type_of env sigma c
- with Anomaly _ ->
- errorlabstrm "" (str "Cannot reinterpret " ++ quote (print_constr c) ++ str " in the current environment.")
-
- let pretype_id loc env sigma (lvar,unbndltacvars) id =
- (* Look for the binder of [id] *)
- try
- let (n,_,typ) = lookup_rel_id id (rel_context env) in
+ try evdref := consider_remaining_unif_problems
+ ~ts:(Typeclasses.classes_transparent_state ()) env !evdref
+ with e when Errors.noncritical e ->
+ let e = Errors.push e in if fail_evar then iraise e
+
+let check_typeclasses_instances_are_solved env current_sigma pending =
+ (* Naive way, call resolution again with failure flag *)
+ apply_typeclasses env (ref current_sigma) pending true
+
+let check_extra_evars_are_solved env current_sigma pending =
+ Evar.Set.iter
+ (fun evk ->
+ if not (Evd.is_defined current_sigma evk) then
+ let (loc,k) = evar_source evk current_sigma in
+ match k with
+ | Evar_kinds.ImplicitArg (gr, (i, id), false) -> ()
+ | _ ->
+ error_unsolvable_implicit loc env current_sigma evk None) pending
+
+let check_evars_are_solved env current_sigma pending =
+ check_typeclasses_instances_are_solved env current_sigma pending;
+ check_problems_are_solved env current_sigma;
+ check_extra_evars_are_solved env current_sigma pending
+
+(* Try typeclasses, hooks, unification heuristics ... *)
+
+let solve_remaining_evars flags env current_sigma pending =
+ let pending = pending_holes pending in
+ let evdref = ref current_sigma in
+ if flags.use_typeclasses then apply_typeclasses env evdref pending false;
+ if Option.has_some flags.use_hook then
+ apply_inference_hook (Option.get flags.use_hook env) evdref pending;
+ if flags.use_unif_heuristics then apply_heuristics env evdref false;
+ if flags.fail_evar then check_evars_are_solved env !evdref pending;
+ !evdref
+
+let check_evars_are_solved env current_sigma pending =
+ let pending = pending_holes pending in
+ check_evars_are_solved env current_sigma pending
+
+let process_inference_flags flags env initial_sigma (sigma,c) =
+ let sigma = solve_remaining_evars flags env sigma (initial_sigma, sigma) in
+ let c = if flags.expand_evars then nf_evar sigma c else c in
+ sigma,c
+
+(* Allow references to syntaxically inexistent variables (i.e., if applied on an inductive) *)
+let allow_anonymous_refs = ref false
+
+(* Utilisé pour inférer le prédicat des Cases *)
+(* Semble exagérement fort *)
+(* Faudra préférer une unification entre les types de toutes les clauses *)
+(* et autoriser des ? à rester dans le résultat de l'unification *)
+
+let evar_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 (e_cumul env evdref (vdefj.(i)).uj_type
+ (lift lt lar.(i))) then
+ error_ill_typed_rec_body_loc loc env !evdref
+ i lna vdefj lar
+ done
+
+(* coerce to tycon if any *)
+let inh_conv_coerce_to_tycon resolve_tc loc env evdref j = function
+ | None -> j
+ | Some t ->
+ evd_comb2 (Coercion.inh_conv_coerce_to resolve_tc loc 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.")
+ else
+ user_err_loc (loc,"",str "No such variable in the signature of the existential variable: " ++ pr_id id ++ str ".")
+
+(* used to enforce a name in Lambda when the type constraints itself
+ is named, hence possibly dependent *)
+
+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 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 ++
+ str " which is not bound in current context.")
+
+let protected_get_type_of env sigma c =
+ try Retyping.get_type_of ~lax:true env sigma c
+ with Retyping.RetypeError _ ->
+ errorlabstrm ""
+ (str "Cannot reinterpret " ++ quote (print_constr c) ++
+ str " in the current environment.")
+
+let pretype_id pretype loc env evdref lvar id =
+ let sigma = !evdref in
+ (* Look for the binder of [id] *)
+ try
+ let id =
+ try Id.Map.find id lvar.ltac_idents
+ with Not_found -> id
+ in
+ let (n,_,typ) = lookup_rel_id id (rel_context env) in
{ uj_val = mkRel n; uj_type = lift n typ }
- with Not_found ->
+ with Not_found ->
(* Check if [id] is an ltac variable *)
try
- let (ids,c) = List.assoc id lvar in
- let subst = List.map (invert_ltac_bound_name env id) ids in
+ let (ids,c) = Id.Map.find id lvar.ltac_constrs in
+ let subst = List.map (invert_ltac_bound_name lvar env id) ids in
let c = substl subst c in
- { uj_val = c; uj_type = protected_get_type_of env sigma c }
+ { uj_val = c; uj_type = protected_get_type_of env sigma c }
+ with Not_found -> try
+ let {closure;term} = Id.Map.find id lvar.ltac_uconstrs in
+ let lvar = {
+ ltac_constrs = closure.typed;
+ ltac_uconstrs = closure.untyped;
+ ltac_idents = closure.idents;
+ ltac_genargs = Id.Map.empty; }
+ in
+ (* spiwack: I'm catching [Not_found] potentially too eagerly
+ here, as the call to the main pretyping function is caught
+ inside the try but I want to avoid refactoring this function
+ too much for now. *)
+ pretype env evdref lvar term
with Not_found ->
- (* if [id] an ltac variable not bound to a term *)
- (* build a nice error message *)
- try
- match List.assoc id unbndltacvars with
- | None -> user_err_loc (loc,"",
- str "Variable " ++ pr_id id ++ str " should be bound to a term.")
- | Some id0 -> Pretype_errors.error_var_not_found_loc loc id0
- with Not_found ->
- (* Check if [id] is a section or goal variable *)
- try
- let (_,_,typ) = lookup_named id env in
- { uj_val = mkVar id; uj_type = typ }
- 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)
-
- (*************************************************************************)
- (* Main pretyping function *)
-
- let pretype_ref evdref env ref =
- let c = constr_of_global ref in
- make_judge c (Retyping.get_type_of env Evd.empty c)
- let pretype_ref loc evdref env = function
- | VarRef id ->
- (* Section variable *)
- (try let (_,_,ty) = lookup_named id env in make_judge (mkVar id) ty
- 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)
- | ref ->
- let c = constr_of_global ref in
- make_judge c (Retyping.get_type_of env Evd.empty c)
-
- let pretype_sort evdref = function
- | GProp c -> judge_of_prop_contents c
- | GType _ -> evd_comb0 judge_of_new_Type evdref
-
- exception Found of fixpoint
-
- let new_type_evar evdref env loc =
- evd_comb0 (fun evd -> Evarutil.new_type_evar evd env ~src:(loc,InternalHole)) evdref
-
- (* [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 resolve_tc (tycon : type_constraint) env evdref lvar t =
- let pretype = pretype resolve_tc in
- let pretype_type = pretype_type resolve_tc in
- let inh_conv_coerce_to_tycon = inh_conv_coerce_to_tycon resolve_tc in
- match t with
- | GRef (loc,ref) ->
- inh_conv_coerce_to_tycon loc env evdref
- (pretype_ref loc evdref env ref)
- tycon
-
- | GVar (loc, id) ->
- inh_conv_coerce_to_tycon loc env evdref
- (pretype_id loc env !evdref lvar id)
- tycon
-
- | GEvar (loc, evk, instopt) ->
- (* 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 hyps = evar_filtered_context (Evd.find !evdref evk) in
- let args = match instopt with
- | None -> instance_from_named_context hyps
- | Some inst -> failwith "Evar subtitutions not implemented" in
- let c = mkEvar (evk, args) in
- let j = (Retyping.get_judgment_of env !evdref c) in
- inh_conv_coerce_to_tycon loc env evdref j tycon
-
- | GPatVar (loc,(someta,n)) ->
- let ty =
- match tycon with
- | Some (None, ty) -> ty
- | None | Some _ -> new_type_evar evdref env loc in
- let k = MatchingVar (someta,n) in
- { uj_val = e_new_evar evdref env ~src:(loc,k) ty; uj_type = ty }
-
- | GHole (loc,k) ->
- let ty =
- match tycon with
- | Some (None, ty) -> ty
- | None | Some _ ->
- new_type_evar evdref env loc in
- { uj_val = e_new_evar evdref env ~src:(loc,k) ty; uj_type = ty }
-
- | GRec (loc,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 = (na,None,ty'.utj_val) in
- type_bl (push_rel dcl env) (add_rel_decl 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 ty in
- let dcl = (na,Some bd'.uj_val,ty'.utj_val) in
- type_bl (push_rel dcl env) (add_rel_decl dcl ctxt) bl in
- let ctxtv = Array.map (type_bl env empty_rel_context) bl in
- let larj =
- array_map2
- (fun e ar ->
- pretype_type empty_valcon (push_rel_context 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
- let nbfix = Array.length lar in
- let names = Array.map (fun id -> Name id) names in
- (* Note: bodies are not used by push_rec_types, so [||] is safe *)
- let newenv = push_rec_types (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 (rel_context_length ctxt)
- (lift nbfix ftys.(i)) in
- let nenv = push_rel_context 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
- evar_type_fixpoint loc 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
- let fixj = match fixkind with
- | GFix (vn,i) ->
+ (* Check if [id] is a ltac variable not bound to a term *)
+ (* and build a nice error message *)
+ if Id.Map.mem id lvar.ltac_genargs then
+ user_err_loc (loc,"",
+ str "Variable " ++ pr_id id ++ str " should be bound to a term.");
+ (* Check if [id] is a section or goal variable *)
+ try
+ let (_,_,typ) = lookup_named id env in
+ (* let _ = *)
+ (* try *)
+ (* let ctx = Decls.variable_context id in *)
+ (* evdref := Evd.merge_context_set univ_rigid !evdref ctx; *)
+ (* with Not_found -> () *)
+ (* in *)
+ { uj_val = mkVar id; uj_type = typ }
+ 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)
+
+(*************************************************************************)
+(* Main pretyping function *)
+
+let interp_universe_level_name evd = function
+ | GProp -> evd, Univ.Level.prop
+ | GSet -> evd, Univ.Level.set
+ | GType s -> interp_universe_level evd s
+
+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 evd l in
+ (evd, l :: univs)) (evd, []) l
+ in
+ evd, Some (Univ.Instance.of_array (Array.of_list (List.rev l')))
+ in
+ Evd.fresh_global ~rigid ?names:instance env evd gr
+
+let pretype_ref loc evdref env ref us =
+ match ref with
+ | VarRef id ->
+ (* Section variable *)
+ (try let (_,_,ty) = lookup_named id env in
+ make_judge (mkVar id) ty
+ 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)
+ | ref ->
+ let evd, c = pretype_global loc univ_flexible env !evdref ref us in
+ let () = evdref := evd in
+ let ty = Typing.type_of env evd c in
+ make_judge c ty
+
+let judge_of_Type evd s =
+ let evd, s = interp_universe evd s in
+ let judge =
+ { uj_val = mkSort (Type s); uj_type = mkSort (Type (Univ.super s)) }
+ in
+ evd, judge
+
+let pretype_sort evdref = function
+ | GProp -> judge_of_prop
+ | GSet -> judge_of_set
+ | GType s -> evd_comb1 judge_of_Type evdref s
+
+let new_type_evar env evdref loc =
+ let e, s =
+ evd_comb0 (fun evd -> Evarutil.new_type_evar env evd univ_flexible_alg ~src:(loc,Evar_kinds.InternalHole)) evdref
+ in e
+
+let get_projection env cst =
+ let cb = lookup_constant cst env in
+ match cb.Declarations.const_proj with
+ | Some {Declarations.proj_ind = mind; proj_npars = n; proj_arg = m; proj_type = ty} ->
+ (cst,mind,n,m,ty)
+ | None -> raise Not_found
+
+let (f_genarg_interp, genarg_interp_hook) = Hook.make ()
+
+(* [pretype tycon env evdref lvar lmeta cstr] attempts to type [cstr] *)
+(* in environment [env], with existential variables [evdref] and *)
+(* the type constraint tycon *)
+
+let is_GHole = function
+ | GHole _ -> true
+ | _ -> false
+
+let evars = ref Id.Map.empty
+
+let rec pretype resolve_tc (tycon : type_constraint) env evdref (lvar : ltac_var_map) t =
+ let inh_conv_coerce_to_tycon = inh_conv_coerce_to_tycon resolve_tc in
+ let pretype_type = pretype_type resolve_tc in
+ let pretype = pretype resolve_tc in
+ match t with
+ | GRef (loc,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
+ (pretype_id (fun e r l t -> pretype tycon e r l t) loc env evdref lvar id)
+ tycon
+
+ | GEvar (loc, 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
+ let hyps = evar_filtered_context (Evd.find !evdref evk) in
+ let args = pretype_instance resolve_tc env evdref lvar loc hyps evk inst in
+ let c = mkEvar (evk, args) in
+ let j = (Retyping.get_judgment_of env !evdref c) in
+ inh_conv_coerce_to_tycon loc env evdref j tycon
+
+ | GPatVar (loc,(someta,n)) ->
+ let ty =
+ match tycon with
+ | Some ty -> ty
+ | None -> new_type_evar env evdref loc in
+ let k = Evar_kinds.MatchingVar (someta,n) in
+ { uj_val = e_new_evar env evdref ~src:(loc,k) ty; uj_type = ty }
+
+ | GHole (loc, k, naming, None) ->
+ let ty =
+ match tycon with
+ | Some ty -> ty
+ | None ->
+ 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 ty =
+ match tycon with
+ | Some ty -> ty
+ | None ->
+ new_type_evar env evdref loc in
+ let ist = lvar.ltac_genargs in
+ let (c, sigma) = Hook.get f_genarg_interp ty env !evdref ist arg in
+ let () = evdref := sigma in
+ { uj_val = c; uj_type = ty }
+
+ | GRec (loc,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 = (na,None,ty'.utj_val) in
+ type_bl (push_rel dcl env) (add_rel_decl 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 = (na,Some bd'.uj_val,ty'.utj_val) in
+ type_bl (push_rel dcl env) (add_rel_decl dcl ctxt) bl in
+ let ctxtv = Array.map (type_bl env empty_rel_context) bl in
+ let larj =
+ Array.map2
+ (fun e ar ->
+ pretype_type empty_valcon (push_rel_context 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
+ let nbfix = Array.length lar in
+ let names = Array.map (fun id -> Name id) names in
+ let _ =
+ match tycon with
+ | Some t ->
+ let fixi = match fixkind with
+ | GFix (vn,i) -> i
+ | GCoFix i -> i
+ in e_conv env evdref ftys.(fixi) t
+ | 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 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 (rel_context_length ctxt)
+ (lift nbfix ftys.(i)) in
+ let nenv = push_rel_context 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
+ evar_type_fixpoint loc 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
+ let fixj = match fixkind with
+ | GFix (vn,i) ->
(* First, let's find the guard indexes. *)
(* If recursive argument was not given by user, we try all args.
An earlier approach was to look only for inductive arguments,
but doing it properly involves delta-reduction, and it finally
doesn't seem worth the effort (except for huge mutual
fixpoints ?) *)
- let possible_indexes = Array.to_list (Array.mapi
- (fun i (n,_) -> match n with
- | Some n -> [n]
- | None -> list_map_i (fun i _ -> i) 0 ctxtv.(i))
- vn)
- in
- let fixdecls = (names,ftys,fdefs) in
- let indexes = search_guard loc env possible_indexes fixdecls in
- make_judge (mkFix ((indexes,i),fixdecls)) ftys.(i)
- | GCoFix i ->
- let cofix = (i,(names,ftys,fdefs)) in
- (try check_cofix env cofix
- with e when Errors.noncritical e -> Loc.raise loc e);
- make_judge (mkCoFix cofix) ftys.(i) in
+ let possible_indexes =
+ Array.to_list (Array.mapi
+ (fun i (n,_) -> match n with
+ | Some n -> [n]
+ | None -> List.map_i (fun i _ -> i) 0 ctxtv.(i))
+ vn)
+ in
+ let fixdecls = (names,ftys,fdefs) in
+ let indexes = search_guard loc env possible_indexes fixdecls in
+ make_judge (mkFix ((indexes,i),fixdecls)) ftys.(i)
+ | GCoFix i ->
+ let cofix = (i,(names,ftys,fdefs)) in
+ (try check_cofix env cofix
+ with reraise ->
+ let (e, info) = Errors.push reraise in
+ let info = Loc.add_loc info loc in
+ iraise (e, info));
+ make_judge (mkCoFix cofix) ftys.(i)
+ in
inh_conv_coerce_to_tycon loc env evdref fixj tycon
- | GSort (loc,s) ->
- let j = pretype_sort evdref s in
- inh_conv_coerce_to_tycon loc env evdref j tycon
-
- | GApp (loc,f,args) ->
- let fj = pretype empty_tycon env evdref lvar f in
- let floc = loc_of_glob_constr f in
- let rec apply_rec env n resj = function
- | [] -> resj
- | c::rest ->
- let argloc = loc_of_glob_constr c in
- let resj = evd_comb1 (Coercion.inh_app_fun resolve_tc env) evdref resj in
- let resty = whd_betadeltaiota env !evdref resj.uj_type in
- match kind_of_term resty with
- | Prod (na,c1,c2) ->
- let hj = pretype (mk_tycon c1) env evdref lvar c in
- let value, typ = applist (j_val resj, [j_val hj]), subst1 hj.uj_val c2 in
- apply_rec env (n+1)
- { uj_val = value;
- uj_type = typ }
- rest
- | _ ->
- let hj = pretype empty_tycon env evdref lvar c in
- error_cant_apply_not_functional_loc
- (join_loc floc argloc) env !evdref
- resj [hj]
- in
- let resj = apply_rec env 1 fj args in
- let resj =
- match evar_kind_of_term !evdref resj.uj_val with
- | App (f,args) ->
- let f = whd_evar !evdref f in
- begin match kind_of_term f with
- | Ind _ | Const _
- when isInd f or has_polymorphic_type (destConst f)
- ->
- let sigma = !evdref in
- let c = mkApp (f,Array.map (whd_evar sigma) args) in
- let t = Retyping.get_type_of env sigma c in
- make_judge c (* use this for keeping evars: resj.uj_val *) t
- | _ -> resj end
- | _ -> resj in
- inh_conv_coerce_to_tycon loc env evdref resj tycon
-
- | GLambda(loc,name,bk,c1,c2) ->
- let (name',dom,rng) = evd_comb1 (split_tycon loc env) evdref tycon in
- let dom_valcon = valcon_of_tycon dom in
- let j = pretype_type dom_valcon env evdref lvar c1 in
- let var = (name,None,j.utj_val) in
- let j' = pretype rng (push_rel var env) evdref lvar c2 in
- judge_of_abstraction env (orelse_name name name') j j'
-
- | GProd(loc,name,bk,c1,c2) ->
- let j = pretype_type empty_valcon env evdref lvar c1 in
- let j' =
- if name = Anonymous then
- let j = pretype_type empty_valcon env evdref lvar c2 in
- { j with utj_val = lift 1 j.utj_val }
- else
- let var = (name,j.utj_val) in
- let env' = push_rel_assum var env in
- pretype_type empty_valcon env' evdref lvar c2
- in
- let resj =
- try judge_of_product env name j j'
- with TypeError _ as e -> Loc.raise loc e in
- inh_conv_coerce_to_tycon loc env evdref resj tycon
-
- | GLetIn(loc,name,c1,c2) ->
- let j =
- match c1 with
- | GCast (loc, c, CastConv (DEFAULTcast, 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
- let t = refresh_universes j.uj_type in
- let var = (name,Some j.uj_val,t) in
- let tycon = lift_tycon 1 tycon in
- let j' = pretype tycon (push_rel var env) evdref lvar c2 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) ->
- let cj = pretype empty_tycon env evdref lvar c in
- let (IndType (indf,realargs)) =
- try find_rectype env !evdref cj.uj_type
- with Not_found ->
- let cloc = loc_of_glob_constr c in
- error_case_not_inductive_loc cloc env !evdref cj
- in
- let cstrs = get_constructors env indf in
- if Array.length cstrs <> 1 then
- user_err_loc (loc,"",str "Destructing let is only for inductive types with one constructor.");
- let cs = cstrs.(0) in
- if List.length nal <> cs.cs_nargs then
- user_err_loc (loc,"", str "Destructing let on this type expects " ++ int cs.cs_nargs ++ str " variables.");
- let fsign = List.map2 (fun na (_,c,t) -> (na,c,t))
- (List.rev nal) cs.cs_args in
- let env_f = push_rels fsign env in
- (* Make dependencies from arity signature impossible *)
- let arsgn =
- let arsgn,_ = get_arity env indf in
- if not !allow_anonymous_refs then
- List.map (fun (_,b,t) -> (Anonymous,b,t)) arsgn
- else arsgn
+ | GSort (loc,s) ->
+ let j = pretype_sort evdref s in
+ inh_conv_coerce_to_tycon loc env evdref j tycon
+
+ | GApp (loc,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
+ let candargs =
+ (* 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
+ match tycon with
+ | None -> []
+ | Some ty ->
+ let ((ind, i), u) = destConstruct 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 !evdref ty in
+ let ((ind',u'),pars) = dest_ind_family indf in
+ if eq_ind ind ind' then 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
+ | Const (p, u) when Environ.is_projection p env ->
+ let p = Projection.make p false in
+ let pb = Environ.lookup_projection p env in
+ let npars = pb.Declarations.proj_npars in
+ fun n ->
+ if n == npars + 1 then fun _ v -> mkProj (p, v)
+ else fun f v -> applist (f, [v])
+ | _ -> fun _ f v -> applist (f, [v])
+ in
+ let rec apply_rec env n resj candargs = function
+ | [] -> resj
+ | c::rest ->
+ let argloc = loc_of_glob_constr c in
+ let resj = evd_comb1 (Coercion.inh_app_fun resolve_tc env) evdref resj in
+ let resty = whd_betadeltaiota env !evdref resj.uj_type in
+ match kind_of_term resty with
+ | Prod (na,c1,c2) ->
+ let tycon = Some c1 in
+ let hj = pretype tycon env evdref lvar c in
+ let candargs, ujval =
+ match candargs with
+ | [] -> [], j_val hj
+ | arg :: args ->
+ if e_conv env evdref (j_val hj) arg then
+ args, nf_evar !evdref (j_val hj)
+ else [], j_val hj
in
- let psign = (na,None,build_dependent_inductive env indf)::arsgn in
- let nar = List.length arsgn in
- (match po with
- | Some p ->
- let env_p = push_rels psign env in
- let pj = pretype_type empty_valcon env_p evdref lvar p in
- let ccl = nf_evar !evdref pj.utj_val in
- let psign = make_arity_signature env true indf in (* with names *)
- let p = it_mkLambda_or_LetIn ccl psign in
- let inst =
- (Array.to_list cs.cs_concl_realargs)
- @[build_dependent_constructor cs] in
- let lp = lift cs.cs_nargs p in
- let fty = hnf_lam_applist env !evdref lp inst in
- let fj = pretype (mk_tycon fty) env_f evdref lvar d in
- let f = it_mkLambda_or_LetIn fj.uj_val fsign in
- let v =
- let ind,_ = dest_ind_family indf in
- let ci = make_case_info env ind LetStyle in
- Typing.check_allowed_sort env !evdref ind cj.uj_val p;
- mkCase (ci, p, cj.uj_val,[|f|]) in
- { 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 f = it_mkLambda_or_LetIn fj.uj_val fsign in
- let ccl = nf_evar !evdref fj.uj_type in
- let ccl =
- if noccur_between 1 cs.cs_nargs ccl then
- lift (- cs.cs_nargs) ccl
- else
- error_cant_find_case_type_loc loc 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 v =
- let ind,_ = dest_ind_family indf in
- let ci = make_case_info env ind LetStyle in
- Typing.check_allowed_sort env !evdref ind cj.uj_val p;
- mkCase (ci, p, cj.uj_val,[|f|])
- in { uj_val = v; uj_type = ccl })
-
- | GIf (loc,c,(na,po),b1,b2) ->
- let cj = pretype empty_tycon env evdref lvar c in
- let (IndType (indf,realargs)) =
- try find_rectype env !evdref cj.uj_type
- with Not_found ->
- let cloc = loc_of_glob_constr c in
- error_case_not_inductive_loc cloc env !evdref cj in
- let cstrs = get_constructors env indf in
- if Array.length cstrs <> 2 then
- user_err_loc (loc,"",
- str "If is only for inductive types with two constructors.");
-
- let arsgn =
- let arsgn,_ = get_arity env indf in
- if not !allow_anonymous_refs then
- (* Make dependencies from arity signature impossible *)
- List.map (fun (_,b,t) -> (Anonymous,b,t)) arsgn
- else arsgn
- in
- let nar = List.length arsgn in
- let psign = (na,None,build_dependent_inductive env indf)::arsgn in
- let pred,p = match po with
- | Some p ->
- let env_p = push_rels psign env in
- let pj = pretype_type empty_valcon env_p evdref lvar 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
- pred, typ
- | None ->
- let p = match tycon with
- | Some (None, ty) -> ty
- | None | Some _ -> new_type_evar evdref env loc
- 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 = rel_context_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 csgn =
- if not !allow_anonymous_refs then
- List.map (fun (_,b,t) -> (Anonymous,b,t)) cs.cs_args
- else
- List.map
- (fun (n, b, t) ->
- match n with
- Name _ -> (n, b, t)
- | Anonymous -> (Name (id_of_string "H"), b, t))
- cs.cs_args
+ 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 !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
+ | App (f,args) ->
+ let f = whd_evar !evdref f in
+ if isInd f && is_template_polymorphic env 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 = evd_comb1 (Evarsolve.refresh_universes (Some true) env) evdref c in
+ let t = Retyping.get_type_of 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
+
+ | GLambda(loc,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 evd ty in
+ evd, Some ty')
+ evdref tycon
+ in
+ let (name',dom,rng) = evd_comb1 (split_tycon loc 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 name = ltac_interp_name lvar name in
+ let var = (name,None,j.utj_val) in
+ let j' = pretype rng (push_rel var env) evdref lvar c2 in
+ let resj = judge_of_abstraction env (orelse_name name name') j j' in
+ inh_conv_coerce_to_tycon loc env evdref resj tycon
+
+ | GProd(loc,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
+ looked up in the rel context. *)
+ let name = ltac_interp_name lvar name in
+ let j' = match name with
+ | Anonymous ->
+ let j = pretype_type empty_valcon env evdref lvar c2 in
+ { j with utj_val = lift 1 j.utj_val }
+ | Name _ ->
+ let var = (name,j.utj_val) in
+ let env' = push_rel_assum var env in
+ pretype_type empty_valcon env' evdref lvar c2
+ in
+ let resj =
+ try
+ judge_of_product env name j j'
+ with TypeError _ as e ->
+ let (e, info) = Errors.push e in
+ let info = Loc.add_loc info loc in
+ iraise (e, info) in
+ 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
+ let t = j.uj_type 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 name = ltac_interp_name lvar name in
+ let var = (name,Some j.uj_val,t) in
+ let tycon = lift_tycon 1 tycon in
+ let j' = pretype tycon (push_rel var env) evdref lvar c2 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) ->
+ let cj = pretype empty_tycon env evdref lvar c in
+ let (IndType (indf,realargs)) =
+ try find_rectype env !evdref cj.uj_type
+ with Not_found ->
+ let cloc = loc_of_glob_constr c in
+ error_case_not_inductive_loc cloc env !evdref cj
+ in
+ let cstrs = get_constructors env indf in
+ if not (Int.equal (Array.length cstrs) 1) then
+ user_err_loc (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 " ++
+ int cs.cs_nargs ++ str " variables.");
+ let nal = List.map (fun na -> ltac_interp_name lvar na) nal in
+ let na = ltac_interp_name lvar na in
+ let fsign, record =
+ match get_projections env indf with
+ | None -> List.map2 (fun na (_,c,t) -> (na,c,t))
+ (List.rev nal) cs.cs_args, false
+ | Some ps ->
+ let rec aux n k names l =
+ match names, l with
+ | na :: names, ((_, None, t) :: l) ->
+ let proj = Projection.make ps.(cs.cs_nargs - k) true in
+ (na, Some (lift (cs.cs_nargs - n) (mkProj (proj, cj.uj_val))), t)
+ :: aux (n+1) (k + 1) names l
+ | na :: names, ((_, c, t) :: l) ->
+ (na, c, t) :: aux (n+1) k names l
+ | [], [] -> []
+ | _ -> assert false
+ in aux 1 1 (List.rev nal) cs.cs_args, true
+ in
+ let obj ind p v f =
+ if not record then
+ let f = it_mkLambda_or_LetIn f fsign in
+ let ci = make_case_info env (fst ind) LetStyle in
+ mkCase (ci, p, cj.uj_val,[|f|])
+ else it_mkLambda_or_LetIn f fsign
+ in
+ let env_f = push_rel_context fsign env in
+ (* Make dependencies from arity signature impossible *)
+ let arsgn =
+ let arsgn,_ = get_arity env indf in
+ if not !allow_anonymous_refs then
+ List.map (fun (_,b,t) -> (Anonymous,b,t)) arsgn
+ else arsgn
+ in
+ let psign = (na,None,build_dependent_inductive env indf)::arsgn 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 ccl = nf_evar !evdref pj.utj_val in
+ let psign = make_arity_signature env true indf in (* with names *)
+ let p = it_mkLambda_or_LetIn ccl psign in
+ let inst =
+ (Array.to_list cs.cs_concl_realargs)
+ @[build_dependent_constructor cs] in
+ let lp = lift cs.cs_nargs p in
+ let fty = hnf_lam_applist env !evdref lp inst in
+ let fj = pretype (mk_tycon fty) env_f evdref lvar d in
+ let v =
+ let ind,_ = dest_ind_family indf in
+ Typing.check_allowed_sort env !evdref ind cj.uj_val p;
+ obj ind p cj.uj_val fj.uj_val
in
- let env_c = push_rels 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
- let b1 = f cstrs.(0) b1 in
- let b2 = f cstrs.(1) b2 in
- let v =
- let ind,_ = dest_ind_family indf in
- let ci = make_case_info env ind IfStyle in
- let pred = nf_evar !evdref pred in
- Typing.check_allowed_sort env !evdref ind cj.uj_val pred;
- mkCase (ci, pred, cj.uj_val, [|b1;b2|])
+ { 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 ccl = nf_evar !evdref fj.uj_type in
+ let ccl =
+ if noccur_between 1 cs.cs_nargs ccl then
+ lift (- cs.cs_nargs) ccl
+ else
+ error_cant_find_case_type_loc loc 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 v =
+ let ind,_ = dest_ind_family indf in
+ Typing.check_allowed_sort 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) ->
+ let cj = pretype empty_tycon env evdref lvar c in
+ let (IndType (indf,realargs)) =
+ try find_rectype env !evdref cj.uj_type
+ with Not_found ->
+ let cloc = loc_of_glob_constr c in
+ error_case_not_inductive_loc cloc env !evdref cj in
+ let cstrs = get_constructors 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.");
+
+ let arsgn =
+ let arsgn,_ = get_arity env indf in
+ if not !allow_anonymous_refs then
+ (* Make dependencies from arity signature impossible *)
+ List.map (fun (_,b,t) -> (Anonymous,b,t)) arsgn
+ else arsgn
+ in
+ let nar = List.length arsgn in
+ let psign = (na,None,build_dependent_inductive env indf)::arsgn 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 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
+ pred, typ
+ | None ->
+ let p = match tycon with
+ | Some ty -> ty
+ | None -> new_type_evar env evdref loc
in
- { uj_val = v; uj_type = p }
-
- | GCases (loc,sty,po,tml,eqns) ->
- Cases.compile_cases loc sty
- ((fun vtyc env evdref -> pretype vtyc env evdref lvar),evdref)
- tycon env (* loc *) (po,tml,eqns)
-
- | GCast (loc,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) evdref cj
- | CastConv (k,t) ->
- let tj = pretype_type empty_valcon env evdref lvar t in
- let cj = pretype empty_tycon env evdref lvar c in
- let cty = nf_evar !evdref cj.uj_type and tval = nf_evar !evdref tj.utj_val in
- let cj = match k with
- | VMcast ->
- if not (occur_existential cty || occur_existential tval) then
- begin
- try
- ignore (Reduction.vm_conv Reduction.CUMUL env cty tval); cj
- with Reduction.NotConvertible ->
- error_actual_type_loc loc env !evdref cj tval
- end
- else user_err_loc (loc,"",str "Cannot check cast with vm: unresolved arguments remain.")
- | _ -> inh_conv_coerce_to_tycon loc env evdref cj (mk_tycon tval)
- 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
-
- (* [pretype_type valcon env evdref lvar c] coerces [c] into a type *)
- and pretype_type resolve_tc valcon env evdref lvar = function
- | GHole loc ->
- (match valcon with
- | Some v ->
- let s =
- let sigma = !evdref in
- let t = Retyping.get_type_of env sigma v in
- match kind_of_term (whd_betadeltaiota env sigma t) with
- | Sort s -> s
- | Evar ev when is_Type (existential_type sigma ev) ->
- evd_comb1 (define_evar_as_sort) evdref ev
- | _ -> anomaly "Found a type constraint which is not a type"
- in
- { utj_val = v;
- utj_type = s }
- | None ->
- let s = evd_comb0 new_sort_variable evdref in
- { utj_val = e_new_evar evdref env ~src:loc (mkSort s);
- utj_type = s})
- | c ->
- let j = pretype 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) evdref j in
- match valcon with
- | None -> tj
- | Some v ->
- if e_cumul env evdref v tj.utj_val then tj
- else
- error_unexpected_type_loc
- (loc_of_glob_constr c) env !evdref tj.utj_val v
-
- let pretype_gen expand_evar fail_evar resolve_classes evdref env lvar kind c =
- let c' = match kind with
- | OfType exptyp ->
- let tycon = match exptyp with None -> empty_tycon | Some t -> mk_tycon t in
- (pretype resolve_classes tycon env evdref lvar c).uj_val
- | IsType ->
- (pretype_type resolve_classes empty_valcon env evdref lvar c).utj_val
+ 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 = rel_context_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 csgn =
+ if not !allow_anonymous_refs then
+ List.map (fun (_,b,t) -> (Anonymous,b,t)) cs.cs_args
+ else
+ List.map
+ (fun (n, b, t) ->
+ match n with
+ Name _ -> (n, b, t)
+ | Anonymous -> (Name Namegen.default_non_dependent_ident, b, t))
+ cs.cs_args
+ in
+ let env_c = push_rel_context 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
+ let b1 = f cstrs.(0) b1 in
+ let b2 = f cstrs.(1) b2 in
+ let v =
+ let ind,_ = dest_ind_family indf in
+ let ci = make_case_info env (fst ind) IfStyle in
+ let pred = nf_evar !evdref pred in
+ Typing.check_allowed_sort env !evdref ind cj.uj_val pred;
+ 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
+
+ | GCases (loc,sty,po,tml,eqns) ->
+ let (tml,eqns) =
+ Glob_ops.map_pattern_binders (fun na -> ltac_interp_name lvar na) tml eqns
in
- resolve_evars env evdref fail_evar resolve_classes;
- let c = if expand_evar then nf_evar !evdref c' else c' in
- if fail_evar then check_evars env Evd.empty !evdref c;
- c
-
- (* TODO: comment faire remonter l'information si le typage a resolu des
- variables du sigma original. il faudrait que la fonction de typage
- retourne aussi le nouveau sigma...
- *)
-
- let understand_judgment sigma env c =
- let evdref = ref sigma in
- let j = pretype true empty_tycon env evdref ([],[]) c in
- resolve_evars env evdref true true;
- let j = j_nf_evar !evdref j in
- check_evars env sigma !evdref (mkCast(j.uj_val,DEFAULTcast, j.uj_type));
- j
-
- let understand_judgment_tcc evdref env c =
- let j = pretype true empty_tycon env evdref ([],[]) c in
- resolve_evars env evdref false true;
- j_nf_evar !evdref j
-
- (* Raw calls to the unsafe inference machine: boolean says if we must
- fail on unresolved evars; the unsafe_judgment list allows us to
- extend env with some bindings *)
-
- let ise_pretype_gen expand_evar fail_evar resolve_classes sigma env lvar kind c =
- let evdref = ref sigma in
- let c = pretype_gen expand_evar fail_evar resolve_classes evdref env lvar kind c in
- !evdref, c
-
- (** Entry points of the high-level type synthesis algorithm *)
-
- let understand_gen kind sigma env c =
- snd (ise_pretype_gen true true true sigma env ([],[]) kind c)
-
- let understand sigma env ?expected_type:exptyp c =
- snd (ise_pretype_gen true true true sigma env ([],[]) (OfType exptyp) c)
-
- let understand_type sigma env c =
- snd (ise_pretype_gen true true true sigma env ([],[]) IsType c)
-
- let understand_ltac ?(resolve_classes=false) expand_evar sigma env lvar kind c =
- ise_pretype_gen expand_evar false resolve_classes sigma env lvar kind c
-
- let understand_tcc ?(resolve_classes=true) sigma env ?expected_type:exptyp c =
- ise_pretype_gen true false resolve_classes sigma env ([],[]) (OfType exptyp) c
-
- let understand_tcc_evars ?(fail_evar=false) ?(resolve_classes=true) evdref env kind c =
- pretype_gen true fail_evar resolve_classes evdref env ([],[]) kind c
-end
-
-module Default : S = Pretyping_F(Coercion.Default)
+ Cases.compile_cases loc sty
+ ((fun vtyc env evdref -> pretype vtyc env evdref lvar),evdref)
+ tycon env (* loc *) (po,tml,eqns)
+
+ | GCast (loc,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) 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
+ let tval = nf_evar !evdref tj.utj_val in
+ let cj = match k with
+ | VMcast ->
+ let cj = pretype empty_tycon env evdref lvar c in
+ let cty = nf_evar !evdref cj.uj_type and tval = nf_evar !evdref tj.utj_val in
+ if not (occur_existential cty || occur_existential tval) then
+ begin
+ try
+ ignore (Reduction.vm_conv Reduction.CUMUL env cty tval); cj
+ with Reduction.NotConvertible ->
+ error_actual_type_loc loc env !evdref cj tval
+ (ConversionFailed (env,cty,tval))
+ end
+ else user_err_loc (loc,"",str "Cannot check cast with vm: " ++
+ str "unresolved arguments remain.")
+ | NATIVEcast ->
+ let cj = pretype empty_tycon env evdref lvar c in
+ let cty = nf_evar !evdref cj.uj_type and tval = nf_evar !evdref tj.utj_val in
+ let evars = Nativenorm.evars_of_evar_map !evdref in
+ begin
+ try
+ ignore (Nativeconv.native_conv Reduction.CUMUL evars env cty tval); cj
+ with Reduction.NotConvertible ->
+ error_actual_type_loc loc env !evdref cj tval
+ (ConversionFailed (env,cty,tval))
+ end
+ | _ ->
+ pretype (mk_tycon tval) env evdref lvar c
+ 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
+
+and pretype_instance resolve_tc env evdref lvar loc hyps evk update =
+ let f (id,_,t) (subst,update) =
+ let t = replace_vars subst t in
+ let c, update =
+ try
+ let c = List.assoc id update in
+ let c = pretype resolve_tc (mk_tycon t) env evdref lvar c in
+ c.uj_val, List.remove_assoc id update
+ with Not_found ->
+ try
+ let (n,_,t') = lookup_rel_id id (rel_context env) in
+ if is_conv env !evdref t t' then mkRel n, update else raise Not_found
+ with Not_found ->
+ try
+ let (_,_,t') = lookup_named id env in
+ if is_conv env !evdref t t' then mkVar id, update else raise Not_found
+ with Not_found ->
+ user_err_loc (loc,"",str "Cannot interpret " ++
+ pr_existential_key !evdref evk ++
+ str " in current context: no binding for " ++ pr_id 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 resolve_tc valcon env evdref lvar = function
+ | GHole (loc, knd, naming, None) ->
+ (match valcon with
+ | Some v ->
+ let s =
+ let sigma = !evdref in
+ let t = Retyping.get_type_of env sigma v in
+ match kind_of_term (whd_betadeltaiota env sigma t) with
+ | Sort s -> s
+ | Evar ev when is_Type (existential_type sigma ev) ->
+ evd_comb1 (define_evar_as_sort env) evdref ev
+ | _ -> anomaly (Pp.str "Found a type constraint which is not a type")
+ in
+ { utj_val = v;
+ utj_type = s }
+ | None ->
+ 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 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) evdref j in
+ match valcon with
+ | None -> tj
+ | Some v ->
+ if e_cumul env evdref v tj.utj_val then tj
+ else
+ error_unexpected_type_loc
+ (loc_of_glob_constr c) env !evdref tj.utj_val v
+
+let ise_pretype_gen flags env sigma lvar kind c =
+ let evdref = ref sigma in
+ let c' = match kind with
+ | WithoutTypeConstraint ->
+ (pretype flags.use_typeclasses empty_tycon env evdref lvar c).uj_val
+ | OfType exptyp ->
+ (pretype flags.use_typeclasses (mk_tycon exptyp) env evdref lvar c).uj_val
+ | IsType ->
+ (pretype_type flags.use_typeclasses empty_valcon env evdref lvar c).utj_val
+ in
+ process_inference_flags flags env sigma (!evdref,c')
+
+let default_inference_flags fail = {
+ use_typeclasses = true;
+ use_unif_heuristics = true;
+ use_hook = None;
+ fail_evar = fail;
+ expand_evars = true }
+
+let no_classes_no_fail_inference_flags = {
+ use_typeclasses = false;
+ use_unif_heuristics = true;
+ use_hook = None;
+ fail_evar = false;
+ expand_evars = true }
+
+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 evdref = ref sigma in
+ let j = pretype 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 sigma (!evdref,c) in
+ evdref := evd; c) j
+ in j, Evd.evar_universe_context !evdref
+
+let understand_judgment_tcc env evdref c =
+ let j = pretype true empty_tycon env evdref empty_lvar c in
+ on_judgment (fun c ->
+ let (evd,c) = process_inference_flags all_no_fail_flags 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
+
+(** Entry points of the high-level type synthesis algorithm *)
+
+let understand
+ ?(flags=all_and_fail_flags)
+ ?(expected_type=WithoutTypeConstraint)
+ 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 =
+ 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_ltac flags env sigma lvar kind c =
+ ise_pretype_gen flags env sigma lvar kind c
diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli
index a785b432..7d1e0c9b 100644
--- a/pretyping/pretyping.mli
+++ b/pretyping/pretyping.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -13,113 +13,137 @@
implicit arguments. *)
open Names
-open Sign
open Term
open Environ
open Evd
open Glob_term
open Evarutil
+open Misctypes
(** An auxiliary function for searching for fixpoint guard indexes *)
val search_guard :
- Util.loc -> env -> int list list -> rec_declaration -> int array
+ Loc.t -> env -> int list list -> rec_declaration -> int array
-type typing_constraint = OfType of types option | IsType
+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 = Genarg.tlevel Genarg.generic_argument 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 var_map = (identifier * Pattern.constr_under_binders) list
-type unbound_ltac_var_map = (identifier * identifier option) list
-type ltac_var_map = var_map * unbound_ltac_var_map
type glob_constr_ltac_closure = ltac_var_map * glob_constr
type pure_open_constr = evar_map * constr
-module type S =
-sig
+type inference_flags = {
+ use_typeclasses : bool;
+ use_unif_heuristics : bool;
+ use_hook : (env -> evar_map -> evar -> constr) option;
+ fail_evar : bool;
+ expand_evars : bool
+}
- module Cases : Cases.S
+val default_inference_flags : bool -> inference_flags
- (** Allow references to syntaxically inexistent variables (i.e., if applied on an inductive) *)
- val allow_anonymous_refs : bool ref
+val no_classes_no_fail_inference_flags : inference_flags
- (** Generic call to the interpreter from glob_constr to open_constr, leaving
- unresolved holes as evars and returning the typing contexts of
- these evars. Work as [understand_gen] for the rest. *)
+val all_no_fail_flags : inference_flags
- val understand_tcc : ?resolve_classes:bool ->
- evar_map -> env -> ?expected_type:types -> glob_constr -> open_constr
+val all_and_fail_flags : inference_flags
- val understand_tcc_evars : ?fail_evar:bool -> ?resolve_classes:bool ->
- evar_map ref -> env -> typing_constraint -> glob_constr -> constr
+(** Allow references to syntaxically inexistent variables (i.e., if applied on an inductive) *)
+val allow_anonymous_refs : bool ref
+
+(** Generic call to the interpreter from glob_constr to open_constr, leaving
+ unresolved holes as evars and returning the typing contexts of
+ these evars. Work as [understand_gen] for the rest. *)
- (** More general entry point with evars from ltac *)
+val understand_tcc : ?flags:inference_flags -> env -> evar_map ->
+ ?expected_type:typing_constraint -> glob_constr -> open_constr
- (** Generic call to the interpreter from glob_constr to constr, failing
- unresolved holes in the glob_constr cannot be instantiated.
+val understand_tcc_evars : ?flags:inference_flags -> env -> evar_map ref ->
+ ?expected_type:typing_constraint -> glob_constr -> constr
- In [understand_ltac expand_evars sigma env ltac_env constraint c],
+(** More general entry point with evars from ltac *)
- resolve_classes : launch typeclass resolution after typechecking.
- expand_evars : expand inferred evars by their value if any
- sigma : initial set of existential variables (typically dependent subgoals)
- ltac_env : partial substitution of variables (used for the tactic language)
- constraint : tell if interpreted as a possibly constrained term or a type
- *)
+(** Generic call to the interpreter from glob_constr to constr
- val understand_ltac : ?resolve_classes:bool ->
- bool -> evar_map -> env -> ltac_var_map ->
- typing_constraint -> glob_constr -> pure_open_constr
+ In [understand_ltac flags sigma env ltac_env constraint c],
- (** Standard call to get a constr from a glob_constr, resolving implicit args *)
+ flags: tell how to manage evars
+ sigma: initial set of existential variables (typically current goals)
+ ltac_env: partial substitution of variables (used for the tactic language)
+ constraint: tell if interpreted as a possibly constrained term or a type
+*)
- val understand : evar_map -> env -> ?expected_type:Term.types ->
- glob_constr -> constr
+val understand_ltac : inference_flags ->
+ env -> evar_map -> ltac_var_map ->
+ typing_constraint -> glob_constr -> pure_open_constr
- (** Idem but the glob_constr is intended to be a type *)
+(** Standard call to get a constr from a glob_constr, resolving implicit args *)
- val understand_type : evar_map -> env -> glob_constr -> constr
+val understand : ?flags:inference_flags -> ?expected_type:typing_constraint ->
+ env -> evar_map -> glob_constr -> constr Evd.in_evar_universe_context
- (** A generalization of the two previous case *)
+(** Idem but returns the judgment of the understood term *)
- val understand_gen : typing_constraint -> evar_map -> env ->
- glob_constr -> constr
+val understand_judgment : env -> evar_map ->
+ glob_constr -> unsafe_judgment Evd.in_evar_universe_context
- (** Idem but returns the judgment of the understood term *)
+(** Idem but do not fail on unresolved evars *)
+val understand_judgment_tcc : env -> evar_map ref ->
+ glob_constr -> unsafe_judgment
- val understand_judgment : evar_map -> env -> glob_constr -> unsafe_judgment
+(** Trying to solve remaining evars and remaining conversion problems
+ with type classes, heuristics, and possibly an external solver *)
+(* For simplicity, it is assumed that current map has no other evars
+ with candidate and no other conversion problems that the one in
+ [pending], however, it can contain more evars than the pending ones. *)
- (** Idem but do not fail on unresolved evars *)
- val understand_judgment_tcc : evar_map ref -> env -> glob_constr -> unsafe_judgment
+val solve_remaining_evars : inference_flags ->
+ env -> (* initial map *) evar_map -> (* map to solve *) pending -> evar_map
- (**/**)
- (** Internal of Pretyping... *)
- val pretype :
- bool -> type_constraint -> env -> evar_map ref ->
- ltac_var_map -> glob_constr -> unsafe_judgment
+(** Checking evars are all solved and reporting an appropriate error message *)
- val pretype_type :
- bool -> val_constraint -> env -> evar_map ref ->
- ltac_var_map -> glob_constr -> unsafe_type_judgment
+val check_evars_are_solved :
+ env -> (* current map: *) evar_map -> (* map to check: *) pending -> unit
- val pretype_gen :
- bool -> bool -> bool -> evar_map ref -> env ->
- ltac_var_map -> typing_constraint -> glob_constr -> constr
+(**/**)
+(** Internal of Pretyping... *)
+val pretype :
+ bool -> type_constraint -> env -> evar_map ref ->
+ ltac_var_map -> glob_constr -> unsafe_judgment
- (**/**)
+val pretype_type :
+ bool -> val_constraint -> env -> evar_map ref ->
+ ltac_var_map -> glob_constr -> unsafe_type_judgment
-end
+val ise_pretype_gen :
+ inference_flags -> env -> evar_map ->
+ ltac_var_map -> typing_constraint -> glob_constr -> evar_map * constr
-module Pretyping_F (C : Coercion.S) : S
-module Default : S
+(**/**)
(** To embed constr in glob_constr *)
val constr_in : constr -> Dyn.t
val constr_out : Dyn.t -> constr
-val interp_sort : glob_sort -> sorts
+val interp_sort : evar_map -> glob_sort -> evar_map * sorts
val interp_elimination_sort : glob_sort -> sorts_family
-(** Last chance for solving evars, possibly using external solver *)
-val solve_remaining_evars : bool -> bool ->
- (env -> evar_map -> existential -> constr) ->
- env -> evar_map -> pure_open_constr -> pure_open_constr
+val genarg_interp_hook :
+ (types -> env -> evar_map -> Genarg.typed_generic_argument Id.Map.t ->
+ Genarg.glob_generic_argument -> constr * evar_map) Hook.t
diff --git a/pretyping/pretyping.mllib b/pretyping/pretyping.mllib
index 9eec9414..25d17c7c 100644
--- a/pretyping/pretyping.mllib
+++ b/pretyping/pretyping.mllib
@@ -1,29 +1,34 @@
+Locusops
Termops
+Namegen
Evd
Reductionops
Vnorm
-Namegen
Inductiveops
+Arguments_renaming
+Nativenorm
Retyping
Cbv
Pretype_errors
+Find_subterm
Evarutil
-Term_dnet
+Evarsolve
Recordops
Evarconv
-Arguments_renaming
Typing
-Glob_term
-Pattern
-Matching
+Miscops
+Glob_ops
+Redops
+Patternops
+Constr_matching
Tacred
Typeclasses_errors
Typeclasses
Classops
+Program
Coercion
-Unification
Detyping
Indrec
Cases
Pretyping
-
+Unification
diff --git a/pretyping/program.ml b/pretyping/program.ml
new file mode 100644
index 00000000..cac8a6a3
--- /dev/null
+++ b/pretyping/program.ml
@@ -0,0 +1,69 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Pp
+open Errors
+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 sp = Libnames.make_path (make_dir dir) (Id.of_string s) in
+ try Nametab.global_of_path sp
+ with Not_found ->
+ anomaly ~label:locstr (Pp.str "cannot find" ++ spc () ++ Libnames.pr_path sp)
+
+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 papp evdref r args =
+ let gr = delayed_force r in
+ mkApp (Evarutil.e_new_global evdref gr, args)
+
+let sig_typ = init_reference ["Init"; "Specif"] "sig"
+let sig_intro = init_reference ["Init"; "Specif"] "exist"
+let sig_proj1 = init_reference ["Init"; "Specif"] "proj1_sig"
+
+let sigT_typ = init_reference ["Init"; "Specif"] "sigT"
+let sigT_intro = init_reference ["Init"; "Specif"] "existT"
+let sigT_proj1 = init_reference ["Init"; "Specif"] "projT1"
+let sigT_proj2 = init_reference ["Init"; "Specif"] "projT2"
+
+let prod_typ = init_reference ["Init"; "Datatypes"] "prod"
+let prod_intro = init_reference ["Init"; "Datatypes"] "pair"
+let prod_proj1 = init_reference ["Init"; "Datatypes"] "fst"
+let prod_proj2 = init_reference ["Init"; "Datatypes"] "snd"
+
+let coq_eq_ind = init_reference ["Init"; "Logic"] "eq"
+let coq_eq_refl = init_reference ["Init"; "Logic"] "eq_refl"
+let coq_eq_refl_ref = init_reference ["Init"; "Logic"] "eq_refl"
+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 mk_coq_not x = mkApp (delayed_force coq_not, [| 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
+ (fun c conj ->
+ mkApp (and_typ, [| c ; conj |]))
+ l
diff --git a/pretyping/program.mli b/pretyping/program.mli
new file mode 100644
index 00000000..3844f375
--- /dev/null
+++ b/pretyping/program.mli
@@ -0,0 +1,39 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Term
+open Globnames
+
+(** A bunch of Coq constants used by Progam *)
+
+val sig_typ : unit -> global_reference
+val sig_intro : unit -> global_reference
+val sig_proj1 : unit -> global_reference
+val sigT_typ : unit -> global_reference
+val sigT_intro : unit -> global_reference
+val sigT_proj1 : unit -> global_reference
+val sigT_proj2 : unit -> global_reference
+
+val prod_typ : unit -> global_reference
+val prod_intro : unit -> global_reference
+val prod_proj1 : unit -> global_reference
+val prod_proj2 : unit -> global_reference
+
+val coq_eq_ind : unit -> global_reference
+val coq_eq_refl : unit -> global_reference
+val coq_eq_refl_ref : unit -> global_reference
+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
+
+(** 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 f9cd3501..6dc0d1f3 100644
--- a/pretyping/recordops.ml
+++ b/pretyping/recordops.ml
@@ -1,27 +1,26 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* Created by Amokrane Saïbi, Dec 1998 *)
+(* Created by Amokrane Saïbi, Dec 1998 *)
(* Addition of products and sorts in canonical structures by Pierre
Corbineau, Feb 2008 *)
(* This file registers properties of records: projections and
canonical structures *)
+open Errors
open Util
open Pp
open Names
-open Libnames
+open Globnames
open Nametab
open Term
-open Typeops
open Libobject
-open Library
open Mod_subst
open Reductionops
@@ -29,25 +28,27 @@ open Reductionops
constructor (the name of which defaults to Build_S) *)
(* Table des structures: le nom de la structure (un [inductive]) donne
- le nom du constructeur, le nombre de paramètres et pour chaque
- argument réel du constructeur, le nom de la projection
- correspondante, si valide, et un booléen disant si c'est une vraie
- projection ou bien une fonction constante (associée à un LetIn) *)
+ le nom du constructeur, le nombre de paramètres et pour chaque
+ argument réel du constructeur, le nom de la projection
+ correspondante, si valide, et un booléen disant si c'est une vraie
+ projection ou bien une fonction constante (associée à un LetIn) *)
type struc_typ = {
s_CONST : constructor;
s_EXPECTEDPARAM : int;
- s_PROJKIND : (name * bool) list;
+ s_PROJKIND : (Name.t * bool) list;
s_PROJ : constant option list }
-let structure_table = ref (Indmap.empty : struc_typ Indmap.t)
-let projection_table = ref Cmap.empty
+let structure_table =
+ Summary.ref (Indmap.empty : struc_typ Indmap.t) ~name:"record-structs"
+let projection_table =
+ Summary.ref Cmap.empty ~name:"record-projs"
(* TODO: could be unify struc_typ and struc_tuple ? in particular,
is the inductive always (fst constructor) ? It seems so... *)
type struc_tuple =
- inductive * constructor * (name * bool) list * constant option list
+ inductive * constructor * (Name.t * bool) list * constant option list
let load_structure i (_,(ind,id,kl,projs)) =
let n = (fst (Global.lookup_inductive ind)).Declarations.mind_nparams in
@@ -62,12 +63,12 @@ let cache_structure o =
load_structure 1 o
let subst_structure (subst,((kn,i),id,kl,projs as obj)) =
- let kn' = subst_ind subst kn in
+ let kn' = subst_mind subst kn in
let projs' =
(* invariant: struc.s_PROJ is an evaluable reference. Thus we can take *)
(* the first component of subst_con. *)
- list_smartmap
- (Option.smartmap (fun kn -> fst (subst_con subst kn)))
+ List.smartmap
+ (Option.smartmap (fun kn -> fst (subst_con_kn subst kn)))
projs
in
let id' = fst (subst_constructor subst id) in
@@ -104,56 +105,6 @@ let find_projection = function
| ConstRef cst -> Cmap.find cst !projection_table
| _ -> raise Not_found
-(* Management of a field store : each field + argument of the inferred
- * records are stored in a discrimination tree *)
-
-let subst_id s (gr,ev,evm) =
- (fst(subst_global s gr),ev,Evd.subst_evar_map s evm)
-
-module MethodsDnet : Term_dnet.S
- with type ident = global_reference * Evd.evar * Evd.evar_map
- = Term_dnet.Make
- (struct
- type t = global_reference * Evd.evar * Evd.evar_map
- let compare = Pervasives.compare
- let subst = subst_id
- let constr_of (_,ev,evm) = Evd.evar_concl (Evd.find evm ev)
- end)
- (struct
- let reduce c = Reductionops.head_unfold_under_prod
- Names.full_transparent_state (Global.env()) Evd.empty c
- let direction = true
- end)
-
-let meth_dnet = ref MethodsDnet.empty
-
-open Summary
-
-let _ =
- declare_summary "record-methods-state"
- { freeze_function = (fun () -> !meth_dnet);
- unfreeze_function = (fun m -> meth_dnet := m);
- init_function = (fun () -> meth_dnet := MethodsDnet.empty) }
-
-open Libobject
-
-let load_method (_,(ty,id)) =
- meth_dnet := MethodsDnet.add ty id !meth_dnet
-
-let in_method : constr * MethodsDnet.ident -> obj =
- declare_object
- { (default_object "RECMETHODS") with
- load_function = (fun _ -> load_method);
- cache_function = load_method;
- subst_function = (fun (s,(ty,id)) -> Mod_subst.subst_mps s ty,subst_id s id);
- classify_function = (fun x -> Substitute x)
- }
-
-let methods_matching c = MethodsDnet.search_pattern !meth_dnet c
-
-let declare_method cons ev sign =
- Lib.add_anonymous_leaf (in_method ((Evd.evar_concl (Evd.find sign ev)),(cons,ev,sign)))
-
(************************************************************************)
(*s A canonical structure declares "canonical" conversion hints between *)
(* the effective components of a structure and the projections of the *)
@@ -163,16 +114,18 @@ let declare_method cons ev sign =
c := [x1:B1]...[xk:Bk](Build_R a1...am t1...t_n)
- If ti has the form (ci ui1...uir) where ci is a global reference and
-if the corresponding projection Li of the structure R is defined, one
-declare a "conversion" between ci and Li
+ If ti has the form (ci ui1...uir) where ci is a global reference (or
+ a sort, or a product or a reference to a parameter) and if the
+ corresponding projection Li of the structure R is defined, one
+ declares a "conversion" between ci and Li.
x1:B1..xk:Bk |- (Li a1..am (c x1..xk)) =_conv (ci ui1...uir)
-that maps the pair (Li,ci) to the following data
+ that maps the pair (Li,ci) to the following data
o_DEF = c
o_TABS = B1...Bk
+ o_INJ = Some n (when ci is a reference to the parameter xi)
o_PARAMS = a1...am
o_NARAMS = m
o_TCOMP = ui1...uir
@@ -181,7 +134,8 @@ that maps the pair (Li,ci) to the following data
type obj_typ = {
o_DEF : constr;
- o_INJ : int; (* position of trivial argument (negative= none) *)
+ o_CTX : Univ.ContextSet.t;
+ o_INJ : int option; (* position of trivial argument if any *)
o_TABS : constr list; (* ordered *)
o_TPARAMS : constr list; (* ordered *)
o_NPARAMS : int;
@@ -193,42 +147,60 @@ type cs_pattern =
| Sort_cs of sorts_family
| Default_cs
-let object_table = ref (Refmap.empty : (cs_pattern * obj_typ) list Refmap.t)
+let eq_cs_pattern p1 p2 = match p1, p2 with
+| Const_cs gr1, Const_cs gr2 -> eq_gr gr1 gr2
+| Prod_cs, Prod_cs -> true
+| Sort_cs s1, Sort_cs s2 -> Sorts.family_equal s1 s2
+| Default_cs, Default_cs -> true
+| _ -> false
+
+let rec assoc_pat a = function
+ | ((pat, t), e) :: xs -> if eq_cs_pattern pat a then (t, e) else assoc_pat a xs
+ | [] -> raise Not_found
+
+
+let object_table =
+ Summary.ref (Refmap.empty : ((cs_pattern * constr) * obj_typ) list Refmap.t)
+ ~name:"record-canonical-structs"
let canonical_projections () =
- Refmap.fold (fun x -> List.fold_right (fun (y,c) acc -> ((x,y),c)::acc))
+ Refmap.fold (fun x -> List.fold_right (fun ((y,_),c) acc -> ((x,y),c)::acc))
!object_table []
let keep_true_projections projs kinds =
- map_succeed (function (p,(_,true)) -> p | _ -> failwith "")
- (List.combine 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
App (f,vargs) ->
begin
- try Const_cs (global_of_constr f) , -1, Array.to_list vargs
+ try Const_cs (global_of_constr f) , None, Array.to_list vargs
with e when Errors.noncritical e -> raise Not_found
end
- | Rel n -> Default_cs, pred n, []
- | Prod (_,a,b) when not (Termops.dependent (mkRel 1) b) -> Prod_cs, -1, [a; Termops.pop b]
- | Sort s -> Sort_cs (family_of_sort s), -1, []
+ | 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, []
| _ ->
begin
- try Const_cs (global_of_constr t) , -1, []
+ try Const_cs (global_of_constr t) , None, []
with e when Errors.noncritical e -> raise Not_found
end
(* Intended to always succeed *)
let compute_canonical_projections (con,ind) =
- let v = mkConst con in
- let c = Environ.constant_value (Global.env()) con in
- let lt,t = Reductionops.splay_lam (Global.env()) Evd.empty c in
- let lt = List.rev (List.map snd lt) in
+ let env = Global.env () in
+ let ctx = Environ.constant_context env con in
+ let u = Univ.UContext.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 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 params, projs = List.chop p args in
let lpj = keep_true_projections lpj kl in
let lps = List.combine lpj projs in
let comp =
@@ -239,48 +211,48 @@ let compute_canonical_projections (con,ind) =
begin
try
let patt, n , args = cs_pattern_of_constr t in
- ((ConstRef proji_sp, patt, n, args) :: l)
+ ((ConstRef proji_sp, patt, t, n, args) :: l)
with Not_found ->
if Flags.is_verbose () then
- (let con_pp = Nametab.pr_global_env Idset.empty (ConstRef con)
- and proji_sp_pp = Nametab.pr_global_env Idset.empty (ConstRef proji_sp) in
- msg_warning (str "No global reference exists for projection value"
- ++ Termops.print_constr t ++ str " in instance "
- ++ con_pp ++ str " of " ++ proji_sp_pp ++ str ", ignoring it."));
+ (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
+ msg_warning (strbrk "No global reference exists for projection value"
+ ++ Termops.print_constr t ++ strbrk " in instance "
+ ++ con_pp ++ str " of " ++ proji_sp_pp ++ strbrk ", ignoring it."));
l
end
| _ -> l)
[] lps in
- List.map (fun (refi,c,inj,argj) ->
- (refi,c),
- {o_DEF=v; o_INJ=inj; o_TABS=lt;
+ List.map (fun (refi,c,t,inj,argj) ->
+ (refi,(c,t)),
+ {o_DEF=v; o_CTX=ctx; o_INJ=inj; o_TABS=lt;
o_TPARAMS=params; o_NPARAMS=List.length params; o_TCOMPS=argj})
comp
let pr_cs_pattern = function
- Const_cs c -> Nametab.pr_global_env Idset.empty c
+ Const_cs c -> Nametab.pr_global_env Id.Set.empty c
| Prod_cs -> str "_ -> _"
| Default_cs -> str "_"
| Sort_cs s -> Termops.pr_sort_family s
let open_canonical_structure i (_,o) =
- if i=1 then
+ if Int.equal i 1 then
let lo = compute_canonical_projections o in
- List.iter (fun ((proj,cs_pat),s) ->
+ List.iter (fun ((proj,(cs_pat,_ as pat)),s) ->
let l = try Refmap.find proj !object_table with Not_found -> [] in
- let ocs = try Some (List.assoc cs_pat l)
+ let ocs = try Some (assoc_pat cs_pat l)
with Not_found -> None
in match ocs with
- | None -> object_table := Refmap.add proj ((cs_pat,s)::l) !object_table;
- | Some cs ->
+ | None -> object_table := Refmap.add proj ((pat,s)::l) !object_table;
+ | Some (c, cs) ->
if Flags.is_verbose () then
let old_can_s = (Termops.print_constr cs.o_DEF)
and new_can_s = (Termops.print_constr s.o_DEF) in
- let prj = (Nametab.pr_global_env Idset.empty proj)
+ let prj = (Nametab.pr_global_env Id.Set.empty proj)
and hd_val = (pr_cs_pattern cs_pat) in
- msg_warning (str "Ignoring canonical projection to " ++ hd_val
- ++ str " by " ++ prj ++ str " in "
- ++ new_can_s ++ str ": redundant with " ++ old_can_s)) lo
+ msg_warning (strbrk "Ignoring canonical projection to " ++ hd_val
+ ++ strbrk " by " ++ prj ++ strbrk " in "
+ ++ new_can_s ++ strbrk ": redundant with " ++ old_can_s)) lo
let cache_canonical_structure o =
open_canonical_structure 1 o
@@ -288,9 +260,9 @@ let cache_canonical_structure o =
let subst_canonical_structure (subst,(cst,ind as obj)) =
(* invariant: cst is an evaluable reference. Thus we can take *)
(* the first component of subst_con. *)
- let cst' = fst (subst_con subst cst) in
- let ind' = Inductiveops.subst_inductive subst ind in
- if cst' == cst & ind' == ind then obj else (cst',ind')
+ let cst' = subst_constant subst cst in
+ let ind' = subst_ind subst ind in
+ if cst' == cst && ind' == ind then obj else (cst',ind')
let discharge_canonical_structure (_,(cst,ind)) =
Some (Lib.discharge_con cst,Lib.discharge_inductive ind)
@@ -314,7 +286,9 @@ let error_not_structure ref =
let check_and_decompose_canonical_structure ref =
let sp = match ref with ConstRef sp -> sp | _ -> error_not_structure ref in
let env = Global.env () in
- let vc = match Environ.constant_opt_value env sp with
+ let ctx = Environ.constant_context env sp in
+ let u = Univ.UContext.instance ctx 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
@@ -322,7 +296,7 @@ let check_and_decompose_canonical_structure ref =
| App (f,args) -> f,args
| _ -> error_not_structure ref in
let indsp = match kind_of_term f with
- | Construct (indsp,1) -> indsp
+ | Construct ((indsp,1),u) -> indsp
| _ -> error_not_structure ref in
let s = try lookup_structure indsp with Not_found -> error_not_structure ref in
let ntrue_projs = List.length (List.filter (fun (_, x) -> x) s.s_PROJKIND) in
@@ -334,32 +308,17 @@ let declare_canonical_structure ref =
add_canonical_structure (check_and_decompose_canonical_structure ref)
let lookup_canonical_conversion (proj,pat) =
- List.assoc pat (Refmap.find proj !object_table)
+ assoc_pat pat (Refmap.find proj !object_table)
let is_open_canonical_projection env sigma (c,args) =
try
- let n = find_projection_nparams (global_of_constr c) in
+ 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
try
- let arg = whd_betadeltaiota env sigma (List.nth args n) in
+ let arg = whd_betadeltaiota env sigma (Stack.nth args n) in
let hd = match kind_of_term arg with App (hd, _) -> hd | _ -> arg in
- not (isConstruct hd)
+ not (isConstruct hd)
with Failure _ -> false
with Not_found -> false
-
-let freeze () =
- !structure_table, !projection_table, !object_table
-
-let unfreeze (s,p,o) =
- structure_table := s; projection_table := p; object_table := o
-
-let init () =
- structure_table := Indmap.empty; projection_table := Cmap.empty;
- object_table := Refmap.empty
-
-let _ = init()
-
-let _ =
- Summary.declare_summary "objdefs"
- { Summary.freeze_function = freeze;
- Summary.unfreeze_function = unfreeze;
- Summary.init_function = init }
diff --git a/pretyping/recordops.mli b/pretyping/recordops.mli
index 442e51db..37d5b4c2 100644
--- a/pretyping/recordops.mli
+++ b/pretyping/recordops.mli
@@ -1,17 +1,14 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Names
-open Nametab
open Term
-open Libnames
-open Libobject
-open Library
+open Globnames
(** Operations concerning records and canonical structures *)
@@ -22,11 +19,11 @@ open Library
type struc_typ = {
s_CONST : constructor;
s_EXPECTEDPARAM : int;
- s_PROJKIND : (name * bool) list;
+ s_PROJKIND : (Name.t * bool) list;
s_PROJ : constant option list }
type struc_tuple =
- inductive * constructor * (name * bool) list * constant option list
+ inductive * constructor * (Name.t * bool) list * constant option list
val declare_structure : struc_tuple -> unit
@@ -46,20 +43,12 @@ val find_projection_nparams : global_reference -> int
(** raise [Not_found] if not a projection *)
val find_projection : global_reference -> struc_typ
-(** we keep an index (dnet) of record's arguments + fields
- (=methods). Here is how to declare them: *)
-val declare_method :
- global_reference -> Evd.evar -> Evd.evar_map -> unit
- (** and here is how to search for methods matched by a given term: *)
-val methods_matching : constr ->
- ((global_reference*Evd.evar*Evd.evar_map) *
- (constr*existential_key)*Termops.subst) list
-
(** {6 Canonical structures } *)
(** A canonical structure declares "canonical" conversion hints between
the effective components of a structure and the projections of the
structure *)
+(** A cs_pattern characterizes the form of a component of canonical structure *)
type cs_pattern =
Const_cs of global_reference
| Prod_cs
@@ -68,18 +57,21 @@ type cs_pattern =
type obj_typ = {
o_DEF : constr;
- o_INJ : int; (** position of trivial argument *)
+ o_CTX : Univ.ContextSet.t;
+ o_INJ : int option; (** position of trivial argument *)
o_TABS : constr list; (** ordered *)
o_TPARAMS : constr list; (** ordered *)
o_NPARAMS : int;
o_TCOMPS : constr list } (** ordered *)
-val cs_pattern_of_constr : constr -> cs_pattern * int * constr list
+(** Return the form of the component of a canonical structure *)
+val cs_pattern_of_constr : constr -> cs_pattern * int option * constr list
+
val pr_cs_pattern : cs_pattern -> Pp.std_ppcmds
-val lookup_canonical_conversion : (global_reference * cs_pattern) -> obj_typ
+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 list) -> bool
+ Environ.env -> Evd.evar_map -> (constr * constr Reductionops.Stack.t) -> bool
val canonical_projections : unit ->
((global_reference * cs_pattern) * obj_typ) list
diff --git a/pretyping/redops.ml b/pretyping/redops.ml
new file mode 100644
index 00000000..92782737
--- /dev/null
+++ b/pretyping/redops.ml
@@ -0,0 +1,38 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Genredexpr
+
+let union_consts l1 l2 = Util.List.union Pervasives.(=) l1 l2 (* FIXME *)
+
+let make_red_flag l =
+ let rec add_flag red = function
+ | [] -> red
+ | FBeta :: lf -> add_flag { red with rBeta = true } lf
+ | FIota :: lf -> add_flag { red with rIota = true } lf
+ | FZeta :: lf -> add_flag { red with rZeta = true } lf
+ | FConst l :: lf ->
+ if red.rDelta then
+ Errors.error
+ "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
+ Errors.error
+ "Cannot set both constants to unfold and constants not to unfold";
+ add_flag
+ { red with rConst = union_consts red.rConst l; rDelta = true }
+ lf
+ in
+ add_flag
+ {rBeta = false; rIota = false; rZeta = false; rDelta = false; rConst = []}
+ l
+
+
+let all_flags =
+ {rBeta = true; rIota = true; rZeta = true; rDelta = true; rConst = []}
diff --git a/pretyping/redops.mli b/pretyping/redops.mli
new file mode 100644
index 00000000..89c68ff3
--- /dev/null
+++ b/pretyping/redops.mli
@@ -0,0 +1,13 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Genredexpr
+
+val make_red_flag : 'a red_atom list -> 'a glob_red_flag
+
+val all_flags : 'a glob_red_flag
diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml
index 7c348d97..a23963ab 100644
--- a/pretyping/reductionops.ml
+++ b/pretyping/reductionops.ml
@@ -1,107 +1,572 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Pp
+open Errors
open Util
open Names
open Term
+open Vars
+open Context
open Termops
open Univ
open Evd
-open Declarations
open Environ
-open Closure
-open Esubst
-open Reduction
exception Elimconst
+(** This module implements a call by name reduction used by (at
+ least) evarconv unification and cbn tactic.
-(**********************************************************************)
-(* The type of (machine) stacks (= lambda-bar-calculus' contexts) *)
-
-type 'a stack_member =
- | Zapp of 'a list
- | Zcase of case_info * 'a * 'a array
- | Zfix of 'a * 'a stack
- | Zshift of int
- | Zupdate of 'a
-
-and 'a stack = 'a stack_member list
-
-let empty_stack = []
-let append_stack_list l s =
- match (l,s) with
- | ([],s) -> s
- | (l1, Zapp l :: s) -> Zapp (l1@l) :: s
- | (l1, s) -> Zapp l1 :: s
-let append_stack v s = append_stack_list (Array.to_list v) s
-
-let rec stack_args_size = function
- | Zapp l::s -> List.length l + stack_args_size s
- | Zshift(_)::s -> stack_args_size s
- | Zupdate(_)::s -> stack_args_size s
- | _ -> 0
-
-(* When used as an argument stack (only Zapp can appear) *)
-let rec decomp_stack = function
- | Zapp[v]::s -> Some (v, s)
- | Zapp(v::l)::s -> Some (v, (Zapp l :: s))
- | Zapp [] :: s -> decomp_stack s
- | _ -> None
-let array_of_stack s =
- let rec stackrec = function
- | [] -> []
- | Zapp args :: s -> args :: (stackrec s)
- | _ -> assert false
- in Array.of_list (List.concat (stackrec s))
-let rec list_of_stack = function
- | [] -> []
- | Zapp args :: s -> args @ (list_of_stack s)
- | _ -> assert false
-let rec app_stack = function
- | f, [] -> f
- | f, (Zapp [] :: s) -> app_stack (f, s)
- | f, (Zapp args :: s) ->
- app_stack (applist (f, args), s)
- | _ -> assert false
-let rec stack_assign s p c = match s with
- | Zapp args :: s ->
- let q = List.length args in
- if p >= q then
- Zapp args :: stack_assign s (p-q) c
- else
- (match list_chop p args with
- (bef, _::aft) -> Zapp (bef@c::aft) :: s
- | _ -> assert false)
- | _ -> s
-let rec stack_tail p s =
- if p = 0 then s else
- match s with
- | Zapp args :: s ->
- let q = List.length args in
- if p >= q then stack_tail (p-q) s
- else Zapp (list_skipn p args) :: s
- | _ -> failwith "stack_tail"
-let rec stack_nth s p = match s with
- | Zapp args :: s ->
- let q = List.length args in
- if p >= q then stack_nth s (p-q)
- else List.nth args p
- | _ -> raise Not_found
-
-(**************************************************************)
-(* The type of (machine) states (= lambda-bar-calculus' cuts) *)
-type state = constr * constr stack
+ It has an ability to "refold" constants by storing constants and
+ their parameters in its stack.
+*)
+
+(** Machinery to custom the behavior of the reduction *)
+module ReductionBehaviour = struct
+ open Globnames
+ open Libobject
+
+ type t = {
+ b_nargs: int;
+ b_recargs: int list;
+ b_dont_expose_case: bool;
+ }
+
+ let table =
+ Summary.ref (Refmap.empty : t Refmap.t) ~name:"reductionbehaviour"
+
+ type flag = [ `ReductionDontExposeCase | `ReductionNeverUnfold ]
+ type req =
+ | ReqLocal
+ | ReqGlobal of global_reference * (int list * int * flag list)
+
+ let load _ (_,(_,(r, b))) =
+ table := Refmap.add r b !table
+
+ let cache o = load 1 o
+
+ let classify = function
+ | ReqLocal, _ -> Dispose
+ | ReqGlobal _, _ as o -> Substitute o
+
+ let subst (subst, (_, (r,o as orig))) =
+ ReqLocal,
+ let r' = fst (subst_global subst r) in if r==r' then orig else (r',o)
+
+ let discharge = function
+ | _,(ReqGlobal (ConstRef c, req), (_, b)) ->
+ let c' = pop_con c in
+ let vars, _subst, _ctx = Lib.section_segment_of_constant c in
+ let extra = List.length vars in
+ let nargs' = if b.b_nargs < 0 then b.b_nargs else b.b_nargs + extra in
+ let recargs' = List.map ((+) extra) b.b_recargs in
+ let b' = { b with b_nargs = nargs'; b_recargs = recargs' } in
+ Some (ReqGlobal (ConstRef c', req), (ConstRef c', b'))
+ | _ -> None
+
+ let rebuild = function
+ | req, (ConstRef c, _ as x) -> req, x
+ | _ -> assert false
+
+ let inRedBehaviour = declare_object {
+ (default_object "REDUCTIONBEHAVIOUR") with
+ load_function = load;
+ cache_function = cache;
+ classify_function = classify;
+ subst_function = subst;
+ discharge_function = discharge;
+ rebuild_function = rebuild;
+ }
+
+ let set local r (recargs, nargs, flags as req) =
+ let nargs = if List.mem `ReductionNeverUnfold flags then max_int else nargs in
+ let behaviour = {
+ b_nargs = nargs; b_recargs = recargs;
+ b_dont_expose_case = List.mem `ReductionDontExposeCase flags } in
+ let req = if local then ReqLocal else ReqGlobal (r, req) in
+ Lib.add_anonymous_leaf (inRedBehaviour (req, (r, behaviour)))
+ ;;
+
+ let get r =
+ try
+ let b = Refmap.find r !table in
+ let flags =
+ if Int.equal b.b_nargs max_int then [`ReductionNeverUnfold]
+ else if b.b_dont_expose_case then [`ReductionDontExposeCase] else [] in
+ Some (b.b_recargs, (if Int.equal b.b_nargs max_int then -1 else b.b_nargs), flags)
+ with Not_found -> None
+
+ let print ref =
+ let open Pp in
+ let pr_global = Nametab.pr_global_env Id.Set.empty in
+ match get ref with
+ | None -> mt ()
+ | Some (recargs, nargs, flags) ->
+ let never = List.mem `ReductionNeverUnfold flags in
+ let nomatch = List.mem `ReductionDontExposeCase flags in
+ let pp_nomatch = spc() ++ if nomatch then
+ str "but avoid exposing match constructs" else str"" in
+ let pp_recargs = spc() ++ str "when the " ++
+ pr_enum (fun x -> pr_nth (x+1)) recargs ++ str (String.plural (List.length recargs) " argument") ++
+ str (String.plural (if List.length recargs >= 2 then 1 else 2) " evaluate") ++
+ str " to a constructor" in
+ let pp_nargs =
+ spc() ++ str "when applied to " ++ int nargs ++
+ str (String.plural nargs " argument") in
+ hov 2 (str "The reduction tactics " ++
+ match recargs, nargs, never with
+ | _,_, true -> str "never unfold " ++ pr_global ref
+ | [], 0, _ -> str "always unfold " ++ pr_global ref
+ | _::_, n, _ when n < 0 ->
+ str "unfold " ++ pr_global ref ++ pp_recargs ++ pp_nomatch
+ | _::_, n, _ when n > List.fold_left max 0 recargs ->
+ str "unfold " ++ pr_global ref ++ pp_recargs ++
+ str " and" ++ pp_nargs ++ pp_nomatch
+ | _::_, _, _ ->
+ str "unfold " ++ pr_global ref ++ pp_recargs ++ pp_nomatch
+ | [], n, _ when n > 0 ->
+ str "unfold " ++ pr_global ref ++ pp_nargs ++ pp_nomatch
+ | _ -> str "unfold " ++ pr_global ref ++ pp_nomatch )
+end
+
+(** Machinery about stack of unfolded constants *)
+module Cst_stack = struct
+(** constant * params * args
+
+- constant applied to params = term in head applied to args
+- there is at most one arguments with an empty list of args, it must be the first.
+- in args, the int represents the indice of the first arg to consider *)
+ type t = (constr * constr list * (int * constr array) list) list
+
+ let empty = []
+ let is_empty = CList.is_empty
+
+ let sanity x y =
+ assert(Term.eq_constr x y)
+
+ let drop_useless = function
+ | _ :: ((_,_,[])::_ as q) -> q
+ | l -> l
+
+ let add_param h cst_l =
+ let append2cst = function
+ | (c,params,[]) -> (c, h::params, [])
+ | (c,params,((i,t)::q)) when i = pred (Array.length t) ->
+ let () = sanity h t.(i) in (c, params, q)
+ | (c,params,(i,t)::q) ->
+ let () = sanity h t.(i) in (c, params, (succ i,t)::q)
+ in
+ drop_useless (List.map append2cst cst_l)
+
+ let add_args cl =
+ List.map (fun (a,b,args) -> (a,b,(0,cl)::args))
+
+ let add_cst cst = function
+ | (_,_,[]) :: q as l -> l
+ | l -> (cst,[],[])::l
+
+ let best_cst = function
+ | (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))
+ | _ -> 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 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
+ (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
+ 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:" ++
+ pr_sequence (fun (i,el) -> prvect_with_sep spc p_c (Array.sub el i (Array.length el - i))) args ++
+ str ")")) l
+end
+
+
+(** The type of (machine) stacks (= lambda-bar-calculus' contexts) *)
+module Stack :
+sig
+ type 'a app_node
+ val pr_app_node : ('a -> Pp.std_ppcmds) -> 'a app_node -> Pp.std_ppcmds
+
+ type cst_member =
+ | Cst_const of pconstant
+ | Cst_proj of projection
+
+ 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
+ | 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
+ 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 compare_shape : 'a t -> 'a t -> bool
+ val map : (constr -> constr) -> constr t -> constr t
+ val fold2 : ('a -> constr -> constr -> 'a) -> 'a ->
+ constr t -> constr t -> 'a * int * int
+ 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
+ val not_purely_applicative : 'a t -> bool
+ val will_expose_iota : 'a t -> bool
+ val list_of_app_stack : constr t -> constr list option
+ val assign : 'a t -> int -> 'a -> 'a t
+ 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
+end =
+struct
+ type 'a app_node = int * 'a array * int
+ (* first releavnt position, arguments, last relevant position *)
+
+ (*
+ Invariant that this module must ensure :
+ (behare of direct access to app_node by the rest of Reductionops)
+ - in app_node (i,_,j) i <= j
+ - There is no array realocation (outside of debug printing)
+ *)
+
+ let pr_app_node pr (i,a,j) =
+ let open Pp in surround (
+ prvect_with_sep pr_comma pr (Array.sub a i (j - i + 1))
+ )
+
+
+ type cst_member =
+ | Cst_const of pconstant
+ | Cst_proj of projection
+
+ 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
+ | 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 =
+ let open Pp in
+ let pr_c x = hov 1 (pr_c x) in
+ match member with
+ | App app -> str "ZApp" ++ pr_app_node pr_c app
+ | Case (_,_,br,cst) ->
+ str "ZCase(" ++
+ prvect_with_sep (pr_bar) pr_c br
+ ++ str ")"
+ | Proj (n,m,p,cst) ->
+ str "ZProj(" ++ int n ++ pr_comma () ++ int m ++
+ pr_comma () ++ pr_con (Projection.constant p) ++ str ")"
+ | Fix (f,args,cst) ->
+ str "ZFix(" ++ Termops.pr_fix Termops.print_constr 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
+
+ and pr_cst_member pr_c c =
+ let open Pp in
+ match c with
+ | Cst_const (c, u) ->
+ if Univ.Instance.is_empty u then Constant.print c
+ else str"(" ++ Constant.print c ++ str ", " ++
+ Univ.Instance.pr Univ.Level.pr u ++ str")"
+ | Cst_proj p ->
+ str".(" ++ Constant.print (Projection.constant p) ++ str")"
+
+ let empty = []
+ let is_empty = CList.is_empty
+
+ let append_app v s =
+ let le = Array.length v in
+ if Int.equal le 0 then s else App (0,v,pred le) :: s
+
+ let decomp_node (i,l,j) sk =
+ if i < j then (l.(i), App (succ i,l,j) :: sk)
+ else (l.(i), sk)
+
+ let decomp = function
+ | App node::s -> Some (decomp_node node s)
+ | _ -> None
+
+ let decomp_node_last (i,l,j) sk =
+ if i < j then (l.(j), App (i,l,pred j) :: sk)
+ else (l.(j), sk)
+
+ let equal f f_fix sk1 sk2 =
+ let equal_cst_member x lft1 y lft2 =
+ match x, y with
+ | Cst_const (c1,u1), Cst_const (c2, u2) ->
+ Constant.equal c1 c2 && Univ.Instance.equal u1 u2
+ | Cst_proj p1, Cst_proj p2 ->
+ Constant.equal (Projection.constant p1) (Projection.constant p2)
+ | _, _ -> false
+ in
+ let rec equal_rec sk1 lft1 sk2 lft2 =
+ 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)
+ | 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
+ | 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
+ | (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
+ | 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
+ | 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
+
+ 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) ->
+ Int.equal bal 0 (* && c1.ci_ind = c2.ci_ind *) && compare_rec 0 s1 s2
+ | (Proj (n1,m1,p,_)::s1, Proj(n2,m2,p2,_)::s2) ->
+ Int.equal bal 0 && compare_rec 0 s1 s2
+ | (Fix(_,a1,_)::s1, Fix(_,a2,_)::s2) ->
+ Int.equal bal 0 && compare_rec 0 a1 a2 && compare_rec 0 s1 s2
+ | (Cst (_,_,_,p1,_)::s1, Cst (_,_,_,p2,_)::s2) ->
+ Int.equal bal 0 && compare_rec 0 p1 p2 && compare_rec 0 s1 s2
+ | (_,_) -> false in
+ compare_rec 0 stk1 stk2
+
+ 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
+ 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
+ | 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
+ | 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
+ | Proj (n1,m1,p1,_) :: q1, Proj (n2,m2,p2,_) :: q2 ->
+ aux o lft1 q1 lft2 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
+ | 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 rec map f x = List.map (function
+ | Update _ -> assert false
+ | (Proj (_,_,_,_) | Shift _) 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)
+ | Case (info,ty,br,alt) -> Case (info, f ty, Array.map f br, alt)
+ | Fix ((r,(na,ty,bo)),arg,alt) ->
+ Fix ((r,(na,Array.map f ty, Array.map f bo)),map f arg,alt)
+ | Cst (cst,curr,remains,params,alt) ->
+ Cst (cst,curr,remains,map f params,alt)
+ ) x
+
+ let append_app_list l s =
+ let a = Array.of_list l in
+ append_app a s
+
+ 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
+ | 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
+ aux (n - nb) (e::out) s
+ else
+ let p = i+n in
+ Some (CList.rev
+ (if Int.equal n 0 then out else App (i,a,p-1) :: out),
+ a.(p),
+ if j > p then App(succ p,a,j)::s else s)
+ | s -> None
+ in aux n [] s
+
+ let not_purely_applicative args =
+ List.exists (function (Fix _ | Case _ | Proj _ | Cst _) -> true | _ -> false) args
+ let will_expose_iota args =
+ List.exists
+ (function (Fix (_,_,l) | Case (_,_,_,l) |
+ Proj (_,_,_,l) | Cst (_,_,_,_,l)) when Cst_stack.is_empty l -> true | _ -> false)
+ args
+
+ 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
+ Option.init init out
+
+ let assign s p c =
+ match strip_n_app p s with
+ | Some (pre,_,sk) -> pre @ (App (0,[|c|],0)::sk)
+ | 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
+ match s with
+ | App (i,a,j) :: s ->
+ let nb = j - i + 1 in
+ if n >= nb then
+ aux lft (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
+
+ let nth s p =
+ match strip_n_app p s with
+ | Some (_,el,_) -> el
+ | None -> raise Not_found
+
+ (** This function breaks the abstraction of Cst_stack ! *)
+ let best_state (_,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) ->
+ if i = pred (Array.length t)
+ then aux sk' def (cst, params, q)
+ else aux sk' def (cst, params, (succ i,t)::q)
+ | _ -> def
+ in List.fold_left (aux sk) s l
+
+ let constr_of_cst_member f sk =
+ match f with
+ | Cst_const (c, u) -> mkConstU (c,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
+ | 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)
+ | 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)
+ | 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
+ (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)
+ | 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)
+ | 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
+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 local_reduction_function = evar_map -> constr -> constr
+type e_reduction_function = env -> evar_map -> constr -> evar_map * constr
type contextual_stack_reduction_function =
env -> evar_map -> constr -> constr * constr list
@@ -114,6 +579,10 @@ type contextual_state_reduction_function =
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)
+
(*************************************)
(*** Reduction Functions Operators ***)
(*************************************)
@@ -122,26 +591,10 @@ let safe_evar_value sigma ev =
try Some (Evd.existential_value sigma ev)
with NotInstantiatedEvar | Not_found -> None
-let rec whd_app_state sigma (x, stack as s) =
- match kind_of_term x with
- | App (f,cl) -> whd_app_state sigma (f, append_stack cl stack)
- | Cast (c,_,_) -> whd_app_state sigma (c, stack)
- | Evar ev ->
- (match safe_evar_value sigma ev with
- Some c -> whd_app_state sigma (c,stack)
- | _ -> s)
- | _ -> s
-
let safe_meta_value sigma ev =
try Some (Evd.meta_value sigma ev)
with Not_found -> None
-let appterm_of_stack (f,s) = (f,list_of_stack s)
-
-let whd_stack sigma x =
- appterm_of_stack (whd_app_state sigma (x, empty_stack))
-let whd_castapp_stack = whd_stack
-
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
@@ -161,70 +614,39 @@ let rec strong_prodspine redfun sigma c =
(*** Reduction using bindingss ***)
(*************************************)
-(* This signature is very similar to Closure.RedFlagsSig except there
- is eta but no per-constant unfolding *)
-
-module type RedFlagsSig = sig
- type flags
- type flag
- val fbeta : flag
- val fdelta : flag
- val feta : flag
- val fiota : flag
- val fzeta : flag
- val mkflags : flag list -> flags
- val red_beta : flags -> bool
- val red_delta : flags -> bool
- val red_eta : flags -> bool
- val red_iota : flags -> bool
- val red_zeta : flags -> bool
-end
-
-(* Compact Implementation *)
-module RedFlags = (struct
- type flag = int
- type flags = int
- let fbeta = 1
- let fdelta = 2
- let feta = 8
- let fiota = 16
- let fzeta = 32
- let mkflags = List.fold_left (lor) 0
- let red_beta f = f land fbeta <> 0
- let red_delta f = f land fdelta <> 0
- let red_eta f = f land feta <> 0
- let red_iota f = f land fiota <> 0
- let red_zeta f = f land fzeta <> 0
-end : RedFlagsSig)
-
-open RedFlags
-
(* Local *)
-let beta = mkflags [fbeta]
-let eta = mkflags [feta]
-let zeta = mkflags [fzeta]
-let betaiota = mkflags [fiota; fbeta]
-let betaiotazeta = mkflags [fiota; fbeta;fzeta]
+let nored = Closure.RedFlags.no_red
+let beta = Closure.beta
+let eta = Closure.RedFlags.mkflags [Closure.RedFlags.fETA]
+let zeta = Closure.RedFlags.mkflags [Closure.RedFlags.fZETA]
+let betaiota = Closure.betaiota
+let betaiotazeta = Closure.betaiotazeta
(* Contextual *)
-let delta = mkflags [fdelta]
-let betadelta = mkflags [fbeta;fdelta;fzeta]
-let betadeltaeta = mkflags [fbeta;fdelta;fzeta;feta]
-let betadeltaiota = mkflags [fbeta;fdelta;fzeta;fiota]
-let betadeltaiota_nolet = mkflags [fbeta;fdelta;fiota]
-let betadeltaiotaeta = mkflags [fbeta;fdelta;fzeta;fiota;feta]
-let betaetalet = mkflags [fbeta;feta;fzeta]
-let betalet = mkflags [fbeta;fzeta]
+let delta = Closure.RedFlags.mkflags [Closure.RedFlags.fDELTA]
+let betalet = Closure.RedFlags.mkflags [Closure.RedFlags.fBETA;Closure.RedFlags.fZETA]
+let betaetalet = Closure.RedFlags.red_add betalet Closure.RedFlags.fETA
+let betadelta = Closure.RedFlags.red_add betalet Closure.RedFlags.fDELTA
+let betadeltaeta = Closure.RedFlags.red_add betadelta Closure.RedFlags.fETA
+let betadeltaiota = Closure.RedFlags.red_add betadelta Closure.RedFlags.fIOTA
+let betadeltaiota_nolet = Closure.betadeltaiotanolet
+let betadeltaiotaeta = Closure.RedFlags.red_add betadeltaiota Closure.RedFlags.fETA
(* Beta Reduction tools *)
-let rec stacklam recfun env t stack =
- match (decomp_stack stack,kind_of_term t) with
- | Some (h,stacktl), Lambda (_,_,c) -> stacklam recfun (h::env) c stacktl
- | _ -> recfun (substl env t, stack)
+let apply_subst recfun env cst_l t stack =
+ let rec aux env cst_l t stack =
+ match (Stack.decomp stack,kind_of_term t) with
+ | Some (h,stacktl), Lambda (_,_,c) ->
+ aux (h::env) (Cst_stack.add_param h cst_l) c stacktl
+ | _ -> recfun cst_l (substl env t, stack)
+ in aux env cst_l t stack
+
+let stacklam recfun env t stack =
+ apply_subst (fun _ -> recfun) env Cst_stack.empty t stack
let beta_applist (c,l) =
- stacklam app_stack [] c (append_stack_list l empty_stack)
+ stacklam Stack.zip [] c (Stack.append_app_list l Stack.empty)
(* Iota reduction tools *)
@@ -239,16 +661,64 @@ let reducible_mind_case c = match kind_of_term c with
| Construct _ | CoFix _ -> true
| _ -> false
-let contract_cofix (bodynum,(types,names,bodies as typedbodies)) =
+(** @return c if there is a constant c whose body is bd
+ @return bd else.
+
+ It has only a meaning because internal representation of "Fixpoint f x
+ := t" is Definition f := fix f x => t
+
+ Even more fragile that we could hope because do Module M. Fixpoint
+ 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
+ | Name.Anonymous -> bd
+ | Name.Name id ->
+ 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
+ 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
+ with
+ | Not_found -> bd
+
+let contract_cofix ?env ?reference (bodynum,(names,types,bodies as typedbodies)) =
let nbodies = Array.length bodies in
- let make_Fi j = mkCoFix (nbodies-j-1,typedbodies) in
- substl (list_tabulate make_Fi nbodies) bodies.(bodynum)
+ let make_Fi j =
+ let ind = nbodies-j-1 in
+ if Int.equal bodynum ind then mkCoFix (ind,typedbodies)
+ else
+ let bd = mkCoFix (ind,typedbodies) in
+ match env with
+ | None -> bd
+ | Some e ->
+ match reference with
+ | None -> bd
+ | Some r -> magicaly_constant_of_fixbody e 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 cst_l cofix sk =
+ let raw_answer = contract_cofix ~env ?reference:(Cst_stack.reference cst_l) cofix in
+ apply_subst
+ (fun x (t,sk') -> recfun x (Cst_stack.best_replace (mkCoFix cofix) cst_l t,sk'))
+ [] Cst_stack.empty raw_answer sk
let reduce_mind_case mia =
match kind_of_term mia.mconstr with
- | Construct (ind_sp,i) ->
+ | 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
+ 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
@@ -258,164 +728,370 @@ let reduce_mind_case mia =
(* 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 ((recindices,bodynum),(types,names,bodies as typedbodies)) =
- let nbodies = Array.length recindices in
- let make_Fi j = mkFix ((recindices,nbodies-j-1),typedbodies) in
- substl (list_tabulate make_Fi nbodies) bodies.(bodynum)
+let contract_fix ?env ?reference ((recindices,bodynum),(names,types,bodies as typedbodies)) =
+ let nbodies = Array.length recindices in
+ let make_Fi j =
+ let ind = nbodies-j-1 in
+ if Int.equal bodynum ind then mkFix ((recindices,ind),typedbodies)
+ else
+ let bd = mkFix ((recindices,ind),typedbodies) in
+ match env with
+ | None -> bd
+ | Some e ->
+ match reference with
+ | None -> bd
+ | Some r -> magicaly_constant_of_fixbody e r bd names.(ind) in
+ let closure = List.init nbodies make_Fi in
+ substl closure bodies.(bodynum)
+
+(** First we substitute the Rel bodynum by the fixpoint and then we try to
+ 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 cst_l fix sk =
+ let raw_answer = contract_fix ~env ?reference:(Cst_stack.reference cst_l) fix in
+ apply_subst
+ (fun x (t,sk') -> recfun x (Cst_stack.best_replace (mkFix fix) cst_l t,sk'))
+ [] Cst_stack.empty raw_answer sk
let fix_recarg ((recindices,bodynum),_) stack =
- assert (0 <= bodynum & bodynum < Array.length recindices);
+ assert (0 <= bodynum && bodynum < Array.length recindices);
let recargnum = Array.get recindices bodynum in
try
- Some (recargnum, stack_nth stack recargnum)
+ Some (recargnum, Stack.nth stack recargnum)
with Not_found ->
None
-type fix_reduction_result = NotReducible | Reduced of state
+(** Generic reduction function with environment
-let reduce_fix whdfun sigma fix stack =
- match fix_recarg fix stack with
- | None -> NotReducible
- | Some (recargnum,recarg) ->
- let (recarg'hd,_ as recarg') = whdfun sigma (recarg, empty_stack) in
- let stack' = stack_assign stack recargnum (app_stack recarg') in
- (match kind_of_term recarg'hd with
- | Construct _ -> Reduced (contract_fix fix, stack')
- | _ -> NotReducible)
+ Here is where unfolded constant are stored in order to be
+ eventualy refolded.
-(* Generic reduction function *)
+ If tactic_mode is true, it uses ReductionBehaviour, prefers
+ refold constant instead of value and tries to infer constants
+ fix and cofix came from.
-(* Y avait un commentaire pour whd_betadeltaiota :
-
- NB : Cette fonction alloue peu c'est l'appel
- ``let (c,cargs) = whfun (recarg, empty_stack)''
- -------------------
- qui coute cher *)
+ It substitutes fix and cofix by the constant they come from in
+ contract_* in any case .
+*)
-let rec whd_state_gen flags ts env sigma =
- let rec whrec (x, stack as s) =
+let debug_RAKAM = ref (false)
+let _ = Goptions.declare_bool_option {
+ Goptions.optsync = true; Goptions.optdepr = false;
+ Goptions.optname =
+ "Print states of the Reductionops abstract machine";
+ Goptions.optkey = ["Debug";"RAKAM"];
+ Goptions.optread = (fun () -> !debug_RAKAM);
+ 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 rec whd_state_gen ?csts tactic_mode flags env sigma =
+ let rec whrec cst_l (x, stack as s) =
+ let () = if !debug_RAKAM then
+ let open Pp in
+ pp (h 0 (str "<<" ++ Termops.print_constr x ++
+ str "|" ++ cut () ++ Cst_stack.pr cst_l ++
+ str "|" ++ cut () ++ Stack.pr Termops.print_constr stack ++
+ str ">>") ++ fnl ())
+ in
+ let fold () =
+ let () = if !debug_RAKAM then
+ let open Pp in pp (str "<><><><><>" ++ fnl ()) in
+ if tactic_mode then (Stack.best_state s cst_l,Cst_stack.empty) else (s,cst_l)
+ in
match kind_of_term x with
- | Rel n when red_delta flags ->
- (match lookup_rel n env with
- | (_,Some body,_) -> whrec (lift n body, stack)
- | _ -> s)
- | Var id when red_delta flags ->
- (match lookup_named id env with
- | (_,Some body,_) -> whrec (body, stack)
- | _ -> s)
- | Evar ev ->
- (match safe_evar_value sigma ev with
- | Some body -> whrec (body, stack)
- | None -> s)
- | Meta ev ->
- (match safe_meta_value sigma ev with
- | Some body -> whrec (body, stack)
- | None -> s)
- | Const const when is_transparent_constant ts const ->
- (match constant_opt_value env const with
- | Some body -> whrec (body, stack)
- | None -> s)
- | LetIn (_,b,_,c) when red_zeta flags -> stacklam whrec [b] c stack
- | Cast (c,_,_) -> whrec (c, stack)
- | App (f,cl) -> whrec (f, append_stack cl stack)
- | Lambda (na,t,c) ->
- (match decomp_stack stack with
- | Some (a,m) when red_beta flags -> stacklam whrec [a] c m
- | None when red_eta flags ->
- let env' = push_rel (na,None,t) env in
- let whrec' = whd_state_gen flags ts env' sigma in
- (match kind_of_term (app_stack (whrec' (c, empty_stack))) with
- | App (f,cl) ->
- let napp = Array.length cl in
- if napp > 0 then
- let x', l' = whrec' (array_last cl, empty_stack) in
- match kind_of_term x', decomp_stack l' with
- | Rel 1, None ->
- let lc = Array.sub cl 0 (napp-1) in
- let u = if napp=1 then f else appvect (f,lc) in
- if noccurn 1 u then (pop u,empty_stack) else s
- | _ -> s
- else s
- | _ -> s)
- | _ -> s)
-
- | Case (ci,p,d,lf) when red_iota flags ->
- let (c,cargs) = whrec (d, empty_stack) in
- if reducible_mind_case c then
- whrec (reduce_mind_case
- {mP=p; mconstr=c; mcargs=list_of_stack cargs;
- mci=ci; mlf=lf}, stack)
- else
- (mkCase (ci, p, app_stack (c,cargs), lf), stack)
-
- | Fix fix when red_iota flags ->
- (match reduce_fix (fun _ -> whrec) sigma fix stack with
- | Reduced s' -> whrec s'
- | NotReducible -> s)
-
- | x -> s
+ | Rel n when Closure.RedFlags.red_set flags Closure.RedFlags.fDELTA ->
+ (match lookup_rel n env with
+ | (_,Some body,_) -> whrec Cst_stack.empty (lift n body, stack)
+ | _ -> fold ())
+ | Var id when Closure.RedFlags.red_set flags (Closure.RedFlags.fVAR id) ->
+ (match lookup_named id env with
+ | (_,Some body,_) -> whrec (Cst_stack.add_cst (mkVar id) cst_l) (body, stack)
+ | _ -> fold ())
+ | Evar ev ->
+ (match safe_evar_value sigma ev with
+ | Some body -> whrec cst_l (body, stack)
+ | None -> fold ())
+ | Meta ev ->
+ (match safe_meta_value sigma ev with
+ | Some body -> whrec cst_l (body, stack)
+ | None -> fold ())
+ | Const (c,u as const) when Closure.RedFlags.red_set flags (Closure.RedFlags.fCONST c) ->
+ (match constant_opt_value_in env const with
+ | None -> fold ()
+ | Some body ->
+ if not tactic_mode
+ then whrec (Cst_stack.add_cst (mkConstU const) cst_l) (body, stack)
+ else (* Looks for ReductionBehaviour *)
+ match ReductionBehaviour.get (Globnames.ConstRef c) with
+ | None -> whrec (Cst_stack.add_cst (mkConstU const) cst_l) (body, stack)
+ | Some (recargs, nargs, flags) ->
+ if (List.mem `ReductionNeverUnfold flags
+ || (nargs > 0 && Stack.args_size stack < nargs))
+ then fold ()
+ else (* maybe unfolds *)
+ if List.mem `ReductionDontExposeCase flags then
+ let app_sk,sk = Stack.strip_app stack in
+ let (tm',sk'),cst_l' =
+ whrec (Cst_stack.add_cst (mkConstU const) cst_l) (body, app_sk)
+ in
+ if equal_stacks (x, app_sk) (tm', sk') || Stack.will_expose_iota sk'
+ then fold ()
+ else whrec cst_l' (tm', sk' @ sk)
+ else match recargs with
+ |[] -> (* if nargs has been specified *)
+ (* CAUTION : the constant is NEVER refold
+ (even when it hides a (co)fix) *)
+ whrec cst_l (body, stack)
+ |curr::remains -> match Stack.strip_n_app curr stack with
+ | None -> fold ()
+ | Some (bef,arg,s') ->
+ whrec Cst_stack.empty
+ (arg,Stack.Cst(Stack.Cst_const const,curr,remains,bef,cst_l)::s')
+ )
+ | Proj (p, c) when Closure.RedFlags.red_projection flags p ->
+ (let pb = lookup_projection p env in
+ let kn = Projection.constant p in
+ let npars = pb.Declarations.proj_npars
+ and arg = pb.Declarations.proj_arg in
+ if not tactic_mode then
+ let stack' = (c, Stack.Proj (npars, arg, p, Cst_stack.empty (*cst_l*)) :: stack) in
+ whrec Cst_stack.empty stack'
+ else match ReductionBehaviour.get (Globnames.ConstRef kn) with
+ | 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 ()
+ else stack'', csts
+ | Some (recargs, nargs, flags) ->
+ if (List.mem `ReductionNeverUnfold flags
+ || (nargs > 0 && Stack.args_size stack < (nargs - (npars + 1))))
+ then fold ()
+ else
+ let recargs = List.map_filter (fun x ->
+ let idx = x - npars in
+ if idx < 0 then None else Some idx) recargs
+ in
+ match recargs with
+ |[] -> (* if nargs has been specified *)
+ (* CAUTION : the constant is NEVER refold
+ (even when it hides a (co)fix) *)
+ let stack' = (c, Stack.Proj (npars, arg, p, cst_l) :: stack) in
+ whrec Cst_stack.empty(* cst_l *) stack'
+ | curr::remains ->
+ if curr == 0 then (* Try to reduce the record argument *)
+ whrec Cst_stack.empty
+ (c, Stack.Cst(Stack.Cst_proj p,curr,remains,Stack.empty,cst_l)::stack)
+ else
+ match Stack.strip_n_app curr stack with
+ | None -> fold ()
+ | Some (bef,arg,s') ->
+ whrec Cst_stack.empty
+ (arg,Stack.Cst(Stack.Cst_proj p,curr,remains,
+ Stack.append_app [|c|] bef,cst_l)::s'))
+
+ | LetIn (_,b,_,c) when Closure.RedFlags.red_set flags Closure.RedFlags.fZETA ->
+ apply_subst whrec [b] cst_l c stack
+ | Cast (c,_,_) -> whrec cst_l (c, stack)
+ | App (f,cl) ->
+ whrec
+ (Cst_stack.add_args cl cst_l)
+ (f, Stack.append_app cl stack)
+ | Lambda (na,t,c) ->
+ (match Stack.decomp stack with
+ | Some _ when Closure.RedFlags.red_set flags Closure.RedFlags.fBETA ->
+ apply_subst whrec [] cst_l x stack
+ | None when Closure.RedFlags.red_set flags Closure.RedFlags.fETA ->
+ let env' = push_rel (na,None,t) env in
+ let whrec' = whd_state_gen tactic_mode flags env' sigma in
+ (match kind_of_term (Stack.zip ~refold:true (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
+ | 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 ()
+ | _ -> fold ()
+ else fold ()
+ | _ -> fold ())
+ | _ -> fold ())
+
+ | Case (ci,p,d,lf) ->
+ whrec Cst_stack.empty (d, Stack.Case (ci,p,lf,cst_l) :: stack)
+
+ | Fix ((ri,n),_ as f) ->
+ (match Stack.strip_n_app ri.(n) stack with
+ |None -> fold ()
+ |Some (bef,arg,s') ->
+ whrec Cst_stack.empty (arg, Stack.Fix(f,bef,cst_l)::s'))
+
+ | Construct ((ind,c),u) ->
+ if Closure.RedFlags.red_set flags Closure.RedFlags.fIOTA then
+ match Stack.strip_app stack with
+ |args, (Stack.Case(ci, _, lf,_)::s') ->
+ whrec Cst_stack.empty (lf.(c-1), (Stack.tail ci.ci_npar args) @ s')
+ |args, (Stack.Proj (n,m,p,_)::s') ->
+ whrec Cst_stack.empty (Stack.nth args (n+m), s')
+ |args, (Stack.Fix (f,s',cst_l)::s'') ->
+ let x' = Stack.zip(x,args) in
+ let out_sk = s' @ (Stack.append_app [|x'|] s'') in
+ reduce_and_refold_fix whrec env cst_l f out_sk
+ |args, (Stack.Cst (const,curr,remains,s',cst_l) :: s'') ->
+ let x' = Stack.zip(x,args) in
+ begin match remains with
+ | [] ->
+ (match const with
+ | Stack.Cst_const const ->
+ (match constant_opt_value_in env const with
+ | None -> fold ()
+ | Some body ->
+ whrec (Cst_stack.add_cst (mkConstU const) cst_l)
+ (body, s' @ (Stack.append_app [|x'|] s'')))
+ | Stack.Cst_proj p ->
+ let pb = lookup_projection p env in
+ let npars = pb.Declarations.proj_npars in
+ let narg = pb.Declarations.proj_arg in
+ let stack = s' @ (Stack.append_app [|x'|] s'') in
+ match Stack.strip_n_app 0 stack with
+ | None -> assert false
+ | Some (_,arg,s'') ->
+ whrec Cst_stack.empty (arg, Stack.Proj (npars,narg,p,cst_l) :: s''))
+ | next :: remains' -> match Stack.strip_n_app (next-curr-1) s'' with
+ | None -> fold ()
+ | Some (bef,arg,s''') ->
+ whrec Cst_stack.empty
+ (arg,
+ Stack.Cst (const,next,remains',s' @ (Stack.append_app [|x'|] bef),cst_l) :: s''')
+ end
+ |_, (Stack.App _|Stack.Update _|Stack.Shift _)::_ -> assert false
+ |_, [] -> fold ()
+ else fold ()
+
+ | CoFix cofix ->
+ if Closure.RedFlags.red_set flags Closure.RedFlags.fIOTA then
+ match Stack.strip_app stack with
+ |args, (Stack.Case(ci, _, lf,_)::s') ->
+ reduce_and_refold_cofix whrec env cst_l cofix stack
+ |_ -> fold ()
+ else fold ()
+
+ | Rel _ | Var _ | Const _ | LetIn _ | Proj _ -> fold ()
+ | Sort _ | Ind _ | Prod _ -> fold ()
in
- whrec
+ whrec (Option.default Cst_stack.empty csts)
+(** 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
- | LetIn (_,b,_,c) when red_zeta flags -> stacklam whrec [b] c stack
- | Cast (c,_,_) -> whrec (c, stack)
- | App (f,cl) -> whrec (f, append_stack cl stack)
- | Lambda (_,_,c) ->
- (match decomp_stack stack with
- | Some (a,m) when red_beta flags -> stacklam whrec [a] c m
- | None when red_eta flags ->
- (match kind_of_term (app_stack (whrec (c, empty_stack))) with
- | App (f,cl) ->
- let napp = Array.length cl in
- if napp > 0 then
- let x', l' = whrec (array_last cl, empty_stack) in
- match kind_of_term x', decomp_stack l' with
- | Rel 1, None ->
- let lc = Array.sub cl 0 (napp-1) in
- let u = if napp=1 then f else appvect (f,lc) in
- if noccurn 1 u then (pop u,empty_stack) else s
- | _ -> s
- else s
- | _ -> s)
- | _ -> s)
-
- | Case (ci,p,d,lf) when red_iota flags ->
- let (c,cargs) = whrec (d, empty_stack) in
- if reducible_mind_case c then
- whrec (reduce_mind_case
- {mP=p; mconstr=c; mcargs=list_of_stack cargs;
- mci=ci; mlf=lf}, stack)
- else
- (mkCase (ci, p, app_stack (c,cargs), lf), stack)
-
- | Fix fix when red_iota flags ->
- (match reduce_fix (fun _ ->whrec) sigma fix stack with
- | Reduced s' -> whrec s'
- | NotReducible -> s)
-
- | Evar ev ->
- (match safe_evar_value sigma ev with
- Some c -> whrec (c,stack)
- | None -> s)
-
- | Meta ev ->
- (match safe_meta_value sigma ev with
- Some c -> whrec (c,stack)
- | None -> s)
-
- | x -> s
+ | LetIn (_,b,_,c) when Closure.RedFlags.red_set flags Closure.RedFlags.fZETA ->
+ stacklam whrec [b] 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 Closure.RedFlags.red_set flags Closure.RedFlags.fBETA ->
+ stacklam whrec [a] c m
+ | None when Closure.RedFlags.red_set flags Closure.RedFlags.fETA ->
+ (match kind_of_term (Stack.zip (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
+ | 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
+ | _ -> s
+ else s
+ | _ -> s)
+ | _ -> s)
+
+ | Proj (p,c) when Closure.RedFlags.red_projection flags p ->
+ (let pb = lookup_projection p (Global.env ()) in
+ whrec (c, Stack.Proj (pb.Declarations.proj_npars, pb.Declarations.proj_arg,
+ p, Cst_stack.empty)
+ :: stack))
+
+ | Case (ci,p,d,lf) ->
+ whrec (d, Stack.Case (ci,p,lf,Cst_stack.empty) :: stack)
+
+ | Fix ((ri,n),_ as f) ->
+ (match Stack.strip_n_app ri.(n) stack with
+ |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)
+
+ | Meta ev ->
+ (match safe_meta_value sigma ev with
+ Some c -> whrec (c,stack)
+ | None -> s)
+
+ | Construct ((ind,c),u) ->
+ if Closure.RedFlags.red_set flags Closure.RedFlags.fIOTA then
+ match Stack.strip_app stack with
+ |args, (Stack.Case(ci, _, lf,_)::s') ->
+ whrec (lf.(c-1), (Stack.tail ci.ci_npar args) @ s')
+ |args, (Stack.Proj (n,m,p,_) :: s') ->
+ whrec (Stack.nth args (n+m), s')
+ |args, (Stack.Fix (f,s',cst)::s'') ->
+ 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
+ |_, [] -> s
+ else s
+
+ | CoFix cofix ->
+ if Closure.RedFlags.red_set flags Closure.RedFlags.fIOTA then
+ match Stack.strip_app stack with
+ |args, (Stack.Case(ci, _, lf,_)::s') ->
+ whrec (contract_cofix cofix, stack)
+ |_ -> s
+ else s
+
+ | x -> s
in
whrec
+let raw_whd_state_gen flags env =
+ let f sigma s = fst (whd_state_gen false flags env sigma s) in
+ f
-let stack_red_of_state_red f sigma x =
- appterm_of_stack (f sigma (x, empty_stack))
+let stack_red_of_state_red f =
+ let f sigma x = decompose_app (Stack.zip (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 flags env sigma (t,Stack.empty) in
+ let whd_sk = Stack.map aux sk in
+ Stack.zip ~refold (hd,whd_sk)
+ in aux s
let red_of_state_red f sigma x =
- app_stack (f sigma (x,empty_stack))
+ Stack.zip (f sigma (x,Stack.empty))
+
+(* 0. No Reduction Functions *)
+
+let whd_nored_state = local_whd_state_gen nored
+let whd_nored_stack = stack_red_of_state_red whd_nored_state
+let whd_nored = red_of_state_red whd_nored_state
(* 1. Beta Reduction Functions *)
@@ -434,19 +1110,18 @@ let whd_betalet = red_of_state_red whd_betalet_state
(* 2. Delta Reduction Functions *)
-let whd_delta_state e = whd_state_gen delta full_transparent_state e
+let whd_delta_state e = raw_whd_state_gen delta e
let whd_delta_stack env = stack_red_of_state_red (whd_delta_state env)
let whd_delta env = red_of_state_red (whd_delta_state env)
-let whd_betadelta_state e = whd_state_gen betadelta full_transparent_state e
+let whd_betadelta_state e = raw_whd_state_gen betadelta e
let whd_betadelta_stack env =
stack_red_of_state_red (whd_betadelta_state env)
let whd_betadelta env =
red_of_state_red (whd_betadelta_state env)
-let whd_betadeltaeta_state e =
- whd_state_gen betadeltaeta full_transparent_state e
+let whd_betadeltaeta_state e = raw_whd_state_gen betadeltaeta e
let whd_betadeltaeta_stack env =
stack_red_of_state_red (whd_betadeltaeta_state env)
let whd_betadeltaeta env =
@@ -462,29 +1137,19 @@ let whd_betaiotazeta_state = local_whd_state_gen betaiotazeta
let whd_betaiotazeta_stack = stack_red_of_state_red whd_betaiotazeta_state
let whd_betaiotazeta = red_of_state_red whd_betaiotazeta_state
-let whd_betadeltaiota_state env =
- whd_state_gen betadeltaiota full_transparent_state env
+let whd_betadeltaiota_state env = raw_whd_state_gen betadeltaiota env
let whd_betadeltaiota_stack env =
stack_red_of_state_red (whd_betadeltaiota_state env)
let whd_betadeltaiota env =
red_of_state_red (whd_betadeltaiota_state env)
-let whd_betadeltaiota_state_using ts env =
- whd_state_gen betadeltaiota ts env
-let whd_betadeltaiota_stack_using ts env =
- stack_red_of_state_red (whd_betadeltaiota_state_using ts env)
-let whd_betadeltaiota_using ts env =
- red_of_state_red (whd_betadeltaiota_state_using ts env)
-
-let whd_betadeltaiotaeta_state env =
- whd_state_gen betadeltaiotaeta full_transparent_state env
+let whd_betadeltaiotaeta_state env = raw_whd_state_gen betadeltaiotaeta env
let whd_betadeltaiotaeta_stack env =
stack_red_of_state_red (whd_betadeltaiotaeta_state env)
let whd_betadeltaiotaeta env =
red_of_state_red (whd_betadeltaiotaeta_state env)
-let whd_betadeltaiota_nolet_state env =
- whd_state_gen betadeltaiota_nolet full_transparent_state env
+let whd_betadeltaiota_nolet_state env = raw_whd_state_gen betadeltaiota_nolet env
let whd_betadeltaiota_nolet_stack env =
stack_red_of_state_red (whd_betadeltaiota_nolet_state env)
let whd_betadeltaiota_nolet env =
@@ -492,11 +1157,11 @@ let whd_betadeltaiota_nolet env =
(* 4. Eta reduction Functions *)
-let whd_eta c = app_stack (local_whd_state_gen eta Evd.empty (c,empty_stack))
+let whd_eta c = Stack.zip (local_whd_state_gen eta Evd.empty (c,Stack.empty))
(* 5. Zeta Reduction Functions *)
-let whd_zeta c = app_stack (local_whd_state_gen zeta Evd.empty (c,empty_stack))
+let whd_zeta c = Stack.zip (local_whd_state_gen zeta Evd.empty (c,Stack.empty))
(****************************************************************************)
(* Reduction Functions *)
@@ -506,10 +1171,23 @@ let whd_zeta c = app_stack (local_whd_state_gen zeta Evd.empty (c,empty_stack))
let rec whd_evar sigma c =
match kind_of_term c with
| Evar ev ->
- (match safe_evar_value sigma ev with
+ let (evk, args) = ev in
+ let args = Array.map (fun c -> whd_evar sigma c) args in
+ (match safe_evar_value sigma (evk, args) with
Some c -> whd_evar sigma c
| None -> c)
- | Sort s -> whd_sort_variable sigma c
+ | Sort (Type u) ->
+ let u' = Evd.normalize_universe sigma u in
+ if u' == u then c else mkSort (Sorts.sort_of_univ u')
+ | Const (c', u) when not (Univ.Instance.is_empty u) ->
+ let u' = Evd.normalize_universe_instance sigma u in
+ if u' == u then c else mkConstU (c', u')
+ | Ind (i, u) when not (Univ.Instance.is_empty u) ->
+ let u' = Evd.normalize_universe_instance sigma u in
+ if u' == u then c else mkIndU (i, u')
+ | Construct (co, u) when not (Univ.Instance.is_empty u) ->
+ let u' = Evd.normalize_universe_instance sigma u in
+ if u' == u then c else mkConstructU (co, u')
| _ -> c
let nf_evar =
@@ -520,66 +1198,19 @@ let nf_evar =
a [nf_evar] here *)
let clos_norm_flags flgs env sigma t =
try
- norm_val
- (create_clos_infos ~evars:(safe_evar_value sigma) flgs env)
- (inject t)
- with Anomaly _ -> error "Tried to normalized ill-typed term"
-
-let nf_beta = clos_norm_flags Closure.beta empty_env
-let nf_betaiota = clos_norm_flags Closure.betaiota empty_env
+ let evars ev = safe_evar_value sigma ev in
+ Closure.norm_val
+ (Closure.create_clos_infos ~evars flgs env)
+ (Closure.inject t)
+ with e when is_anomaly e -> error "Tried to normalize ill-typed term"
+
+let nf_beta = clos_norm_flags Closure.beta (Global.env ())
+let nf_betaiota = clos_norm_flags Closure.betaiota (Global.env ())
+let nf_betaiotazeta = clos_norm_flags Closure.betaiotazeta (Global.env ())
let nf_betadeltaiota env sigma =
clos_norm_flags Closure.betadeltaiota env sigma
-(* Attention reduire un beta-redexe avec un argument qui n'est pas
- une variable, peut changer enormement le temps de conversion lors
- du type checking :
- (fun x => x + x) M
-*)
-let rec whd_betaiota_preserving_vm_cast env sigma t =
- let rec stacklam_var subst t stack =
- match (decomp_stack stack,kind_of_term t) with
- | Some (h,stacktl), Lambda (_,_,c) ->
- begin match kind_of_term h with
- | Rel i when not (evaluable_rel i env) ->
- stacklam_var (h::subst) c stacktl
- | Var id when not (evaluable_named id env)->
- stacklam_var (h::subst) c stacktl
- | _ -> whrec (substl subst t, stack)
- end
- | _ -> whrec (substl subst t, stack)
- and whrec (x, stack as s) =
- match kind_of_term x with
- | Evar ev ->
- (match safe_evar_value sigma ev with
- | Some body -> whrec (body, stack)
- | None -> s)
- | Cast (c,VMcast,t) ->
- let c = app_stack (whrec (c,empty_stack)) in
- let t = app_stack (whrec (t,empty_stack)) in
- (mkCast(c,VMcast,t),stack)
- | Cast (c,DEFAULTcast,_) ->
- whrec (c, stack)
- | App (f,cl) -> whrec (f, append_stack cl stack)
- | Lambda (na,t,c) ->
- (match decomp_stack stack with
- | Some (a,m) -> stacklam_var [a] c m
- | _ -> s)
- | Case (ci,p,d,lf) ->
- let (c,cargs) = whrec (d, empty_stack) in
- if reducible_mind_case c then
- whrec (reduce_mind_case
- {mP=p; mconstr=c; mcargs=list_of_stack cargs;
- mci=ci; mlf=lf}, stack)
- else
- (mkCase (ci, p, app_stack (c,cargs), lf), stack)
- | x -> s
- in
- app_stack (whrec (t,empty_stack))
-
-let nf_betaiota_preserving_vm_cast =
- strong whd_betaiota_preserving_vm_cast
-
(********************************************************************)
(* Conversion *)
(********************************************************************)
@@ -591,40 +1222,88 @@ let fakey = Profile.declare_profile "fhnf_apply";;
let fhnf_apply info k h a = Profile.profile4 fakey fhnf_apply info k h a;;
*)
-let is_transparent k =
- Conv_oracle.get_strategy k <> Conv_oracle.Opaque
+let is_transparent e k =
+ match Conv_oracle.get_strategy (Environ.oracle e) k with
+ | Conv_oracle.Opaque -> false
+ | _ -> true
(* Conversion utility functions *)
type conversion_test = constraints -> constraints
-let pb_is_equal pb = pb = CONV
+let pb_is_equal pb = pb == Reduction.CONV
let pb_equal = function
- | CUMUL -> CONV
- | CONV -> CONV
-
-let sort_cmp = sort_cmp
+ | Reduction.CUMUL -> Reduction.CONV
+ | Reduction.CONV -> Reduction.CONV
-let test_conversion (f: ?l2r:bool-> ?evars:'a->'b) env sigma x y =
- try let _ =
- f ~evars:(safe_evar_value sigma) env x y in true
- with NotConvertible -> false
- | Anomaly _ -> error "Conversion test raised an anomaly"
-
-let is_conv env sigma = test_conversion Reduction.conv env sigma
-let is_conv_leq env sigma = test_conversion Reduction.conv_leq env sigma
-let is_fconv = function | CONV -> is_conv | CUMUL -> is_conv_leq
+let sort_cmp cv_pb s1 s2 u =
+ Reduction.check_sort_cmp_universes cv_pb s1 s2 u
let test_trans_conversion (f: ?l2r:bool-> ?evars:'a->'b) reds env sigma x y =
- try let _ = f ~evars:(safe_evar_value sigma) reds env x y in true
- with NotConvertible -> false
- | Anomaly _ -> error "Conversion test raised an anomaly"
-
-let is_trans_conv reds env sigma = test_trans_conversion Reduction.trans_conv reds env sigma
-let is_trans_conv_leq reds env sigma = test_trans_conversion Reduction.trans_conv_leq reds env sigma
-let is_trans_fconv = function | CONV -> is_trans_conv | CUMUL -> is_trans_conv_leq
-
+ try
+ let evars ev = safe_evar_value sigma ev in
+ let _ = f ~evars reds env (Evd.universes sigma) x y in
+ true
+ with Reduction.NotConvertible -> false
+ | e when is_anomaly e -> error "Conversion test raised an anomaly"
+
+let is_trans_conv reds env sigma = test_trans_conversion Reduction.trans_conv_universes reds env sigma
+let is_trans_conv_leq reds env sigma = test_trans_conversion Reduction.trans_conv_leq_universes reds env sigma
+let is_trans_fconv = function Reduction.CONV -> is_trans_conv | Reduction.CUMUL -> is_trans_conv_leq
+
+let is_conv = is_trans_conv full_transparent_state
+let is_conv_leq = is_trans_conv_leq full_transparent_state
+let is_fconv = function | Reduction.CONV -> is_conv | Reduction.CUMUL -> is_conv_leq
+
+let check_conv ?(pb=Reduction.CUMUL) ?(ts=full_transparent_state) env sigma x y =
+ let f = match pb with
+ | Reduction.CONV -> Reduction.trans_conv_universes
+ | Reduction.CUMUL -> Reduction.trans_conv_leq_universes
+ in
+ try f ~evars:(safe_evar_value sigma) ts env (Evd.universes sigma) x y; true
+ with Reduction.NotConvertible -> false
+ | Univ.UniverseInconsistency _ -> false
+ | e when is_anomaly e -> error "Conversion test raised an anomaly"
+
+let sigma_compare_sorts env pb s0 s1 sigma =
+ match pb with
+ | Reduction.CONV -> Evd.set_eq_sort env sigma s0 s1
+ | Reduction.CUMUL -> Evd.set_leq_sort env sigma s0 s1
+
+let sigma_compare_instances flex i0 i1 sigma =
+ try Evd.set_eq_instances ~flex sigma i0 i1
+ with Evd.UniversesDiffer -> raise Reduction.NotConvertible
+
+let sigma_univ_state =
+ { Reduction.compare = sigma_compare_sorts;
+ Reduction.compare_instances = sigma_compare_instances }
+
+let infer_conv ?(pb=Reduction.CUMUL) ?(ts=full_transparent_state) env sigma x y =
+ try
+ let b, sigma =
+ let b, cstrs =
+ if pb == Reduction.CUMUL then
+ Universes.leq_constr_univs_infer (Evd.universes sigma) x y
+ else
+ Universes.eq_constr_univs_infer (Evd.universes sigma) x y
+ in
+ if b then
+ try true, Evd.add_universe_constraints sigma cstrs
+ with Univ.UniverseInconsistency _ | Evd.UniversesDiffer -> false, sigma
+ else false, sigma
+ in
+ if b then sigma, true
+ else
+ let sigma' =
+ Reduction.generic_conv pb false (safe_evar_value sigma) ts
+ env (sigma, sigma_univ_state) x y in
+ sigma', true
+ with
+ | Reduction.NotConvertible -> sigma, false
+ | Univ.UniverseInconsistency _ -> sigma, false
+ | e when is_anomaly e -> error "Conversion test raised an anomaly"
+
(********************************************************************)
(* Special-Purpose Reduction *)
(********************************************************************)
@@ -633,33 +1312,36 @@ let whd_meta sigma c = match kind_of_term c with
| Meta p -> (try 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
- | Meta p -> (try lift n (List.assoc p s) with Not_found -> u)
+ | 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
- let l' = Array.map (irec n) l in
+ let l' = CArray.Fun1.smartmap irec n l in
(match kind_of_term 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 = List.assoc p s in
+ (try let g = Metamap.find p s in
match kind_of_term g with
| App _ ->
- let h = id_of_string "H" in
- mkLetIn (Name h,g,t,mkApp(mkRel 1,Array.map (lift 1) l'))
+ 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 (List.assoc (destMeta m) s) with Not_found -> u)
+ (try lift n (Metamap.find (destMeta m) s) with Not_found -> u)
| _ ->
map_constr_with_binders succ irec n u
in
- if s = [] then c else irec 0 c
+ if Metamap.is_empty s then c
+ else irec 0 c
(* [instance] is used for [res_pf]; the call to [local_strong whd_betaiota]
has (unfortunately) different subtle side effects:
@@ -708,7 +1390,7 @@ let instance sigma s c =
let hnf_prod_app env sigma t n =
match kind_of_term (whd_betadeltaiota env sigma t) with
| Prod (_,_,b) -> subst1 n b
- | _ -> anomaly "hnf_prod_app: 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
@@ -719,7 +1401,7 @@ let hnf_prod_applist env sigma t nl =
let hnf_lam_app env sigma t n =
match kind_of_term (whd_betadeltaiota env sigma t) with
| Lambda (_,_,b) -> subst1 n b
- | _ -> anomaly "hnf_lam_app: 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
@@ -760,7 +1442,10 @@ let splay_prod_assum env sigma =
prodec_rec (push_rel (x, Some b, t) env)
(add_rel_decl (x, Some b, t) l) c
| Cast (c,_,_) -> prodec_rec env l c
- | _ -> l,t
+ | _ ->
+ let t' = whd_betadeltaiota env sigma t in
+ if Term.eq_constr t t' then l,t
+ else prodec_rec env l t'
in
prodec_rec env empty_rel_context
@@ -773,7 +1458,7 @@ let splay_arity env sigma c =
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 m = 0 then (ln,c) else
+ let rec decrec env m ln c = if Int.equal m 0 then (ln,c) else
match kind_of_term (whd_betadeltaiota env sigma c) with
| Prod (n,a,c0) ->
decrec (push_rel (n,None,a) env)
@@ -783,7 +1468,7 @@ let splay_prod_n env sigma n =
decrec env n empty_rel_context
let splay_lam_n env sigma n =
- let rec decrec env m ln c = if m = 0 then (ln,c) else
+ let rec decrec env m ln c = if Int.equal m 0 then (ln,c) else
match kind_of_term (whd_betadeltaiota env sigma c) with
| Lambda (n,a,c0) ->
decrec (push_rel (n,None,a) env)
@@ -792,87 +1477,34 @@ let splay_lam_n env sigma n =
in
decrec env n empty_rel_context
-exception NotASort
-
-let decomp_sort env sigma t =
+let is_sort env sigma t =
match kind_of_term (whd_betadeltaiota env sigma t) with
- | Sort s -> s
- | _ -> raise NotASort
-
-let is_sort env sigma arity =
- try let _ = decomp_sort env sigma arity in true
- with NotASort -> false
+ | Sort s -> true
+ | _ -> false
(* reduction to head-normal-form allowing delta/zeta only in argument
of case/fix (heuristic used by evar_conv) *)
-let whd_betaiota_deltazeta_for_iota_state ts env sigma s =
- let rec whrec s =
- let (t, stack as s) = whd_betaiota_state sigma s in
- match kind_of_term t with
- | Case (ci,p,d,lf) ->
- let (cr,crargs) = whd_betadeltaiota_stack_using ts env sigma d in
- let rslt = mkCase (ci, p, applist (cr,crargs), lf) in
- if reducible_mind_case cr then
- whrec (rslt, stack)
- else
- s
- | Fix fix ->
- (match
- reduce_fix (whd_betadeltaiota_state_using ts env) sigma fix stack
- with
- | Reduced s -> whrec s
- | NotReducible -> s)
- | _ -> s
- in whrec s
-
-(* A reduction function like whd_betaiota but which keeps casts
- * and does not reduce redexes containing existential variables.
- * Used in Correctness.
- * Added by JCF, 29/1/98. *)
-
-let whd_programs_stack env sigma =
- let rec whrec (x, stack as s) =
- match kind_of_term x with
- | App (f,cl) ->
- let n = Array.length cl - 1 in
- let c = cl.(n) in
- if occur_existential c then
- s
- else
- whrec (mkApp (f, Array.sub cl 0 n), append_stack [|c|] stack)
- | LetIn (_,b,_,c) ->
- if occur_existential b then
- s
- else
- stacklam whrec [b] c stack
- | Lambda (_,_,c) ->
- (match decomp_stack stack with
- | None -> s
- | Some (a,m) -> stacklam whrec [a] c m)
- | Case (ci,p,d,lf) ->
- if occur_existential d then
- s
- else
- let (c,cargs) = whrec (d, empty_stack) in
- if reducible_mind_case c then
- whrec (reduce_mind_case
- {mP=p; mconstr=c; mcargs=list_of_stack cargs;
- mci=ci; mlf=lf}, stack)
- else
- (mkCase (ci, p, app_stack(c,cargs), lf), stack)
- | Fix fix ->
- (match reduce_fix (fun _ ->whrec) sigma fix stack with
- | Reduced s' -> whrec s'
- | NotReducible -> s)
- | _ -> s
- in
- whrec
-
-let whd_programs env sigma x =
- app_stack (whd_programs_stack env sigma (x, empty_stack))
-
-exception IsType
+let whd_betaiota_deltazeta_for_iota_state ts env sigma csts s =
+ let rec whrec csts s =
+ let (t, stack as s),csts' = whd_state_gen ~csts false betaiota env sigma s in
+ match Stack.strip_app stack with
+ |args, (Stack.Case _ :: _ as stack') ->
+ let (t_o,stack_o),csts_o = whd_state_gen ~csts:csts' false
+ (Closure.RedFlags.red_add_transparent betadeltaiota ts) env sigma (t,args) in
+ if reducible_mind_case 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' false
+ (Closure.RedFlags.red_add_transparent betadeltaiota ts) env sigma (t,args) in
+ if isConstruct 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' false
+ (Closure.RedFlags.red_add_transparent betadeltaiota ts) env sigma (t,args) in
+ if isConstruct 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'
+ in whrec csts s
let find_conclusion env sigma =
let rec decrec env c =
@@ -896,74 +1528,86 @@ let meta_value evd mv =
let rec valrec mv =
match meta_opt_fvalue evd mv with
| Some (b,_) ->
- instance evd
- (List.map (fun mv' -> (mv',valrec mv')) (Metaset.elements b.freemetas))
- b.rebus
+ let metas = Metamap.bind valrec b.freemetas in
+ instance evd metas b.rebus
| None -> mkMeta mv
in
valrec mv
let meta_instance sigma b =
- let c_sigma =
- List.map
- (fun mv -> (mv,meta_value sigma mv)) (Metaset.elements b.freemetas)
- in
- if c_sigma = [] then b.rebus else instance sigma c_sigma b.rebus
+ let fm = b.freemetas in
+ if Metaset.is_empty fm then b.rebus
+ else
+ 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)
(* Instantiate metas that create beta/iota redexes *)
let meta_reducible_instance evd b =
- let fm = Metaset.elements b.freemetas in
- let metas = List.fold_left (fun l mv ->
- match (try meta_opt_fvalue evd mv with Not_found -> None) with
- | Some (g,(_,s)) -> (mv,(g.rebus,s))::l
- | None -> l) [] fm in
+ let fm = b.freemetas in
+ let fold mv accu =
+ let fvalue = try meta_opt_fvalue evd mv with Not_found -> None in
+ match fvalue with
+ | None -> accu
+ | Some (g, (_, s)) -> Metamap.add mv (g.rebus, s) accu
+ 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 c or isCast c & isMeta (pi1 (destCast c)) ->
- let m =
- try destMeta c
- with e when Errors.noncritical e -> destMeta (pi1 (destCast c))
- in
+ | Case (ci,p,c,bl) when isMeta (strip_outer_cast c) ->
+ let m = destMeta (strip_outer_cast c) in
(match
try
- let g,s = List.assoc m metas in
- if isConstruct g or s <> CoerceToType then Some g else None
+ let g, s = Metamap.find m metas in
+ let is_coerce = match s with CoerceToType -> true | _ -> false in
+ if isConstruct 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 f or isCast f & isMeta (pi1 (destCast f)) ->
- let m =
- try destMeta f
- with e when Errors.noncritical e -> destMeta (pi1 (destCast f))
- in
+ | App (f,l) when isMeta (strip_outer_cast f) ->
+ let m = destMeta (strip_outer_cast f) in
(match
try
- let g,s = List.assoc m metas in
- if isLambda g or s <> CoerceToType then Some g else None
+ let g, s = Metamap.find m metas in
+ let is_coerce = match s with CoerceToType -> true | _ -> false in
+ if isLambda 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 = List.assoc m metas in if s<>CoerceToType then irec g else u
+ (try let g, s = Metamap.find m metas 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
+ (match
+ try
+ let g, s = Metamap.find m metas in
+ let is_coerce = match s with CoerceToType -> true | _ -> false in
+ if isConstruct g || not is_coerce then Some g else None
+ with Not_found -> None
+ with
+ | Some g -> irec (mkProj (p,g))
+ | None -> mkProj (p,c))
| _ -> map_constr irec u
in
- if fm = [] then (* nf_betaiota? *) b.rebus else irec b.rebus
+ 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 =
+ let unfold (cst,u as cstu) =
if Cpred.mem cst (snd ts) then
- match constant_opt_value env cst with
+ match constant_opt_value_in env cstu with
| Some c -> c
- | None -> mkConst cst
- else mkConst cst in
+ | None -> mkConstU cstu
+ else mkConstU cstu in
let rec aux c =
match kind_of_term c with
| Prod (n,t,c) -> mkProd (n,aux t, aux c)
diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli
index f508ea6c..7c61d4e1 100644
--- a/pretyping/reductionops.mli
+++ b/pretyping/reductionops.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,50 +8,109 @@
open Names
open Term
+open Context
open Univ
open Evd
open Environ
-open Closure
(** Reduction Functions. *)
exception Elimconst
-(***********************************************************************
- s A [stack] is a context of arguments, arguments are pushed by
- [append_stack] one array at a time but popped with [decomp_stack]
- one by one *)
-
-type 'a stack_member =
- | Zapp of 'a list
- | Zcase of case_info * 'a * 'a array
- | Zfix of 'a * 'a stack
- | Zshift of int
- | Zupdate of 'a
-
-and 'a stack = 'a stack_member list
-
-val empty_stack : 'a stack
-val append_stack : 'a array -> 'a stack -> 'a stack
-val append_stack_list : 'a list -> 'a stack -> 'a stack
-
-val decomp_stack : 'a stack -> ('a * 'a stack) option
-val list_of_stack : 'a stack -> 'a list
-val array_of_stack : 'a stack -> 'a array
-val stack_assign : 'a stack -> int -> 'a -> 'a stack
-val stack_args_size : 'a stack -> int
-val app_stack : constr * constr stack -> constr
-val stack_tail : int -> 'a stack -> 'a stack
-val stack_nth : 'a stack -> int -> 'a
+(** Machinery to customize the behavior of the reduction *)
+module ReductionBehaviour : sig
+ type flag = [ `ReductionDontExposeCase | `ReductionNeverUnfold ]
+
+(** [set is_local ref (recargs, nargs, flags)] *)
+ val set :
+ 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
+end
+
+(** {6 Machinery about a stack of unfolded constant }
+
+ cst applied to params must convertible to term of the state applied to args
+*)
+module Cst_stack : sig
+ type t
+ val empty : t
+ val add_param : constr -> t -> t
+ 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
+end
+
+
+module Stack : sig
+ type 'a app_node
+
+ val pr_app_node : ('a -> Pp.std_ppcmds) -> 'a app_node -> Pp.std_ppcmds
+
+ type cst_member =
+ | Cst_const of pconstant
+ | Cst_proj of projection
+
+ 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
+ | 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 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 compare_shape : 'a t -> 'a t -> bool
+ (** [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
+ 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 *)
+ 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
+
+ val not_purely_applicative : 'a t -> bool
+ val list_of_app_stack : constr t -> constr list option
+
+ val assign : 'a t -> int -> 'a -> 'a t
+ 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
+end
(************************************************************************)
-type state = constr * constr stack
+type state = constr * constr Stack.t
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 = 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
@@ -63,11 +122,7 @@ type contextual_state_reduction_function =
type state_reduction_function = contextual_state_reduction_function
type local_state_reduction_function = evar_map -> state -> state
-(** Removes cast and put into applicative form *)
-val whd_stack : local_stack_reduction_function
-
-(** For compatibility: alias for whd\_stack *)
-val whd_castapp_stack : local_stack_reduction_function
+val pr_state : state -> Pp.std_ppcmds
(** {6 Reduction Function Operators } *)
@@ -78,7 +133,13 @@ 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 -> 'a
+val stacklam : (state -> 'a) -> constr list -> constr -> constr Stack.t -> 'a
+
+val whd_state_gen : ?csts:Cst_stack.t -> bool -> Closure.RedFlags.reds ->
+ Environ.env -> Evd.evar_map -> state -> state * Cst_stack.t
+
+val iterate_whd_gen : bool -> Closure.RedFlags.reds ->
+ Environ.env -> Evd.evar_map -> Term.constr -> Term.constr
(** {6 Generic Optimized Reduction Function using Closures } *)
@@ -87,13 +148,14 @@ val clos_norm_flags : Closure.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_betadeltaiota : reduction_function
val nf_evar : evar_map -> constr -> constr
-val nf_betaiota_preserving_vm_cast : reduction_function
-
(** Lazy strategy, weak head reduction *)
+
val whd_evar : evar_map -> constr -> constr
+val whd_nored : local_reduction_function
val whd_beta : local_reduction_function
val whd_betaiota : local_reduction_function
val whd_betaiotazeta : local_reduction_function
@@ -102,6 +164,8 @@ val whd_betadeltaiota_nolet : contextual_reduction_function
val whd_betaetalet : local_reduction_function
val whd_betalet : local_reduction_function
+(** Removes cast and put into applicative form *)
+val whd_nored_stack : local_stack_reduction_function
val whd_beta_stack : local_stack_reduction_function
val whd_betaiota_stack : local_stack_reduction_function
val whd_betaiotazeta_stack : local_stack_reduction_function
@@ -110,6 +174,7 @@ val whd_betadeltaiota_nolet_stack : contextual_stack_reduction_function
val whd_betaetalet_stack : local_stack_reduction_function
val whd_betalet_stack : local_stack_reduction_function
+val whd_nored_state : local_state_reduction_function
val whd_beta_state : local_state_reduction_function
val whd_betaiota_state : local_state_reduction_function
val whd_betaiotazeta_state : local_state_reduction_function
@@ -149,15 +214,14 @@ val hnf_lam_app : env -> evar_map -> constr -> constr -> constr
val hnf_lam_appvect : env -> evar_map -> constr -> constr array -> constr
val hnf_lam_applist : env -> evar_map -> constr -> constr list -> constr
-val splay_prod : env -> evar_map -> constr -> (name * constr) list * constr
-val splay_lam : env -> evar_map -> constr -> (name * constr) list * constr
-val splay_arity : env -> evar_map -> constr -> (name * constr) list * sorts
+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 -> rel_context * constr
val splay_lam_n : env -> evar_map -> int -> constr -> rel_context * constr
val splay_prod_assum :
env -> evar_map -> constr -> rel_context * constr
-val decomp_sort : env -> evar_map -> types -> sorts
val is_sort : env -> evar_map -> types -> bool
type 'a miota_args = {
@@ -172,20 +236,13 @@ val reduce_mind_case : constr miota_args -> constr
val find_conclusion : env -> evar_map -> constr -> (constr,constr) kind_of_term
val is_arity : env -> evar_map -> constr -> bool
+val is_sort : env -> evar_map -> types -> bool
-val whd_programs : reduction_function
-
-(** [reduce_fix redfun fix stk] contracts [fix stk] if it is actually
- reducible; the structural argument is reduced by [redfun] *)
-
-type fix_reduction_result = NotReducible | Reduced of state
-
-val fix_recarg : fixpoint -> constr stack -> (int * constr) option
-val reduce_fix : local_state_reduction_function -> evar_map -> fixpoint
- -> constr stack -> fix_reduction_result
+val contract_fix : ?env:Environ.env -> ?reference:Constant.t -> fixpoint -> constr
+val fix_recarg : fixpoint -> constr Stack.t -> (int * constr) option
(** {6 Querying the kernel conversion oracle: opaque/transparent constants } *)
-val is_transparent : 'a tableKey -> bool
+val is_transparent : Environ.env -> constant tableKey -> bool
(** {6 Conversion Functions (uses closures, lazy strategy) } *)
@@ -194,7 +251,7 @@ type conversion_test = constraints -> constraints
val pb_is_equal : conv_pb -> bool
val pb_equal : conv_pb -> conv_pb
-val sort_cmp : conv_pb -> sorts -> sorts -> conversion_test
+val sort_cmp : env -> conv_pb -> sorts -> sorts -> universes -> unit
val is_conv : env -> evar_map -> constr -> constr -> bool
val is_conv_leq : env -> evar_map -> constr -> constr -> bool
@@ -204,17 +261,29 @@ val is_trans_conv : transparent_state -> env -> evar_map -> constr -> constr ->
val is_trans_conv_leq : transparent_state -> env -> evar_map -> constr -> constr -> bool
val is_trans_fconv : conv_pb -> transparent_state -> env -> evar_map -> constr -> constr -> bool
+(** [check_conv] Checks universe constraints only.
+ pb defaults to CUMUL and ts to a full transparent state.
+ *)
+val check_conv : ?pb:conv_pb -> ?ts:transparent_state -> env -> evar_map -> constr -> constr -> bool
+
+(** [infer_fconv] Adds necessary universe constraints to the evar map.
+ pb defaults to CUMUL and ts to a full transparent state.
+ *)
+val infer_conv : ?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 : (metavariable * constr) list -> constr -> constr
-val instance :evar_map -> (metavariable * constr) list -> constr -> constr
+val plain_instance : constr Metamap.t -> constr -> constr
+val instance : evar_map -> constr Metamap.t -> constr -> constr
val head_unfold_under_prod : transparent_state -> reduction_function
(** {6 Heuristic for Conversion with Evar } *)
val whd_betaiota_deltazeta_for_iota_state :
- transparent_state -> state_reduction_function
+ transparent_state -> Environ.env -> Evd.evar_map -> Cst_stack.t -> state ->
+ state * Cst_stack.t
(** {6 Meta-related reduction functions } *)
val meta_instance : evar_map -> constr freelisted -> constr
diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml
index 1e1960f5..cd52ba44 100644
--- a/pretyping/retyping.ml
+++ b/pretyping/retyping.ml
@@ -1,26 +1,54 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Pp
+open Errors
open Util
open Term
+open Vars
open Inductive
open Inductiveops
open Names
open Reductionops
open Environ
-open Typeops
-open Declarations
open Termops
+open Arguments_renaming
+
+type retype_error =
+ | NotASort
+ | NotAnArity
+ | NotAType
+ | BadVariable of Id.t
+ | BadMeta of int
+ | BadRecursiveType
+ | NonFunctionalConstruction
+
+let print_retype_error = function
+ | NotASort -> str "Not a sort"
+ | NotAnArity -> str "Not an arity"
+ | NotAType -> str "Not a type (1)"
+ | BadVariable id -> str "variable " ++ Id.print id ++ str " unbound"
+ | BadMeta n -> str "unknown meta " ++ int n
+ | BadRecursiveType -> str "Bad recursive type"
+ | NonFunctionalConstruction -> str "Non-functional construction"
+
+exception RetypeError of retype_error
+
+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)
let get_type_from_constraints env sigma t =
if isEvar (fst (decompose_app t)) then
match
- list_map_filter (fun (pbty,env,t1,t2) ->
+ 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
else None)
@@ -29,46 +57,48 @@ let get_type_from_constraints env sigma t =
| t::l -> t
| _ -> raise Not_found
else raise Not_found
-
+
let rec subst_type env sigma typ = function
| [] -> typ
| h::rest ->
match kind_of_term (whd_betadeltaiota env sigma typ) with
| Prod (na,c1,c2) -> subst_type env sigma (subst1 h c2) rest
- | _ -> anomaly "Non-functional construction"
+ | _ -> retype_error NonFunctionalConstruction
-(* Si ft est le type d'un terme f, lequel est appliqué à args, *)
-(* [sort_of_atomic_ty] calcule ft[args] qui doit être une sorte *)
-(* On suit une méthode paresseuse, en espèrant que ft est une arité *)
-(* et sinon on substitue *)
+(* If ft is the type of f which itself is applied to args, *)
+(* [sort_of_atomic_type] computes ft[args] which has to be a sort *)
let sort_of_atomic_type env sigma ft args =
- let rec concl_of_arity env ar args =
+ let rec concl_of_arity env n ar args =
match kind_of_term (whd_betadeltaiota env sigma ar), args with
- | Prod (na, t, b), h::l -> concl_of_arity (push_rel (na,Some h,t) env) b l
+ | Prod (na, t, b), h::l -> concl_of_arity (push_rel (na,Some (lift n h),t) env) (n + 1) b l
| Sort s, [] -> s
- | _ -> anomaly "Not a sort"
- in concl_of_arity env ft (Array.to_list args)
+ | _ -> retype_error NotASort
+ in concl_of_arity env 0 ft (Array.to_list args)
let type_of_var env id =
try let (_,_,ty) = lookup_named id env in ty
- with Not_found ->
- anomaly ("type_of: variable "^(string_of_id id)^" unbound")
+ with Not_found -> retype_error (BadVariable id)
+
+let decomp_sort env sigma t =
+ match kind_of_term (whd_betadeltaiota env sigma t) with
+ | Sort s -> s
+ | _ -> retype_error NotASort
let retype ?(polyprop=true) sigma =
let rec type_of env cstr=
match kind_of_term cstr with
| Meta n ->
(try strip_outer_cast (Evd.meta_ftype sigma n).Evd.rebus
- with Not_found -> anomaly ("type_of: unknown meta " ^ string_of_int n))
+ with Not_found -> retype_error (BadMeta n))
| Rel n ->
let (_,_,ty) = lookup_rel n env in
lift n ty
| Var id -> type_of_var env id
- | Const cst -> Typeops.type_of_constant env cst
+ | Const cst -> rename_type_of_constant env cst
| Evar ev -> Evd.existential_type sigma ev
- | Ind ind -> type_of_inductive env ind
- | Construct cstr -> type_of_constructor env cstr
+ | Ind ind -> rename_type_of_inductive env ind
+ | Construct cstr -> rename_type_of_constructor env cstr
| Case (_,p,c,lf) ->
let Inductiveops.IndType(_,realargs) =
let t = type_of env c in
@@ -77,7 +107,8 @@ let retype ?(polyprop=true) sigma =
try
let t = get_type_from_constraints env sigma t in
Inductiveops.find_rectype env sigma t
- with Not_found -> anomaly "type_of: Bad recursive type" in
+ with Not_found -> retype_error BadRecursiveType
+ in
let t = whd_beta sigma (applist (p, realargs)) in
(match kind_of_term (whd_betadeltaiota env sigma (type_of env t)) with
| Prod _ -> whd_beta sigma (applist (t, [c]))
@@ -88,12 +119,20 @@ let retype ?(polyprop=true) sigma =
subst1 b (type_of (push_rel (name,Some b,c1) env) c2)
| Fix ((_,i),(_,tys,_)) -> tys.(i)
| CoFix (i,(_,tys,_)) -> tys.(i)
- | App(f,args) when isGlobalRef f ->
+ | App(f,args) when is_template_polymorphic env 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))
| App(f,args) ->
strip_outer_cast
(subst_type env sigma (type_of env f) (Array.to_list args))
+ | Proj (p,c) ->
+ let Inductiveops.IndType(pars,realargs) =
+ let ty = type_of env c in
+ try Inductiveops.find_rectype env sigma ty
+ with Not_found -> retype_error BadRecursiveType
+ in
+ let (_,u), pars = dest_ind_family pars in
+ substl (c :: List.rev pars) (Typeops.type_of_projection env (p,u))
| Cast (c,_, t) -> t
| Sort _ | Prod _ -> mkSort (sort_of env cstr)
@@ -106,20 +145,16 @@ let retype ?(polyprop=true) sigma =
(match (sort_of env t, sort_of (push_rel (name,None,t) env) c2) with
| _, (Prop Null as s) -> s
| Prop _, (Prop Pos as s) -> s
- | Type _, (Prop Pos as s) when
- Environ.engagement env = Some ImpredicativeSet -> s
- | (Type _, _) | (_, Type _) -> new_Type_sort ()
-(*
+ | Type _, (Prop Pos as s) when is_impredicative_set env -> s
| Type u1, Prop Pos -> Type (Univ.sup u1 Univ.type0_univ)
| 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 isGlobalRef f ->
- let t = type_of_global_reference_knowing_parameters env f args in
+ | Type u1, Type u2 -> Type (Univ.sup u1 u2))
+ | App(f,args) when is_template_polymorphic env 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 _ ->
- anomaly "sort_of: Not a type (1)"
+ | Lambda _ | Fix _ | Construct _ -> retype_error NotAType
| _ -> decomp_sort env sigma (type_of env t)
and sort_family_of env t =
@@ -129,29 +164,29 @@ let retype ?(polyprop=true) sigma =
| Sort (Type u) -> InType
| Prod (name,t,c2) ->
let s2 = sort_family_of (push_rel (name,None,t) env) c2 in
- if Environ.engagement env <> Some ImpredicativeSet &&
- s2 = InSet & sort_family_of env t = InType then InType else s2
- | App(f,args) when isGlobalRef f ->
+ 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 ->
let t = type_of_global_reference_knowing_parameters env f args in
family_of_sort (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)
- | Lambda _ | Fix _ | Construct _ ->
- anomaly "sort_of: Not a type (1)"
- | _ -> family_of_sort (decomp_sort env sigma (type_of env t))
+ | Lambda _ | Fix _ | Construct _ -> retype_error NotAType
+ | _ ->
+ 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 -> nf_evar sigma (type_of env c)) args in
+ 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 ind in
+ let mip = lookup_mind_specif env (fst ind) in
(try Inductive.type_of_inductive_knowing_parameters
- ~polyprop env mip argtyps
- with Reduction.NotArity -> anomaly "type_of: Not an arity")
+ ~polyprop env (mip,snd ind) argtyps
+ with Reduction.NotArity -> retype_error NotAnArity)
| Const cst ->
- let t = constant_type env cst in
- (try Typeops.type_of_constant_knowing_parameters env t argtyps
- with Reduction.NotArity -> anomaly "type_of: Not an arity")
+ (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
@@ -160,36 +195,59 @@ let retype ?(polyprop=true) sigma =
type_of_global_reference_knowing_parameters
let get_sort_of ?(polyprop=true) env sigma t =
- let _,f,_,_ = retype ~polyprop sigma in f env 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 f env c
+ let _,_,f,_ = retype ~polyprop sigma in anomaly_on_error (f env) c
let type_of_global_reference_knowing_parameters env sigma c args =
- let _,_,_,f = retype sigma in 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
- | Ind ind ->
- let (_,mip) = Inductive.lookup_mind_specif env ind in
- type_of_inductive_knowing_conclusion env mip conclty
+ | 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 env cst in
+ let t = constant_type_in env cst in
(* TODO *)
- Typeops.type_of_constant_knowing_parameters env t [||]
- | Var id -> type_of_var env id
- | Construct cstr -> type_of_constructor env cstr
+ sigma, Typeops.type_of_constant_type_knowing_parameters env t [||]
+ | Var id -> sigma, type_of_var env id
+ | Construct cstr -> sigma, type_of_constructor env cstr
| _ -> assert false
-(* We are outside the kernel: we take fresh universes *)
-(* to avoid tactics and co to refresh universes themselves *)
-let get_type_of ?(polyprop=true) ?(refresh=true) env sigma c =
- let f,_,_,_ = retype ~polyprop sigma in
- let t = f env c in
- if refresh then refresh_universes t else t
+(* Profiling *)
+(* let get_type_of polyprop lax env sigma c = *)
+(* 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 *)
-(* Makes an assumption from a constr *)
-let get_assumption_of env evc c = c
+(* 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
+ if lax then f env c else anomaly_on_error (f env) c
(* Makes an unsafe judgment from a constr *)
let get_judgment_of env evc c = { uj_val = c; uj_type = get_type_of env evc c }
+(* Returns sorts of a context *)
+let sorts_of_context env evc ctxt =
+ let rec aux = function
+ | [] -> env,[]
+ | (_,_,t as d)::ctxt ->
+ let env,sorts = aux ctxt in
+ let s = get_sort_of env evc t in
+ (push_rel d env,s::sorts) in
+ snd (aux ctxt)
+
+let expand_projection env sigma pr c args =
+ let ty = get_type_of ~lax:true env sigma c in
+ let (i,u), ind_args =
+ try Inductiveops.find_mrectype env sigma ty
+ with Not_found -> retype_error BadRecursiveType
+ in
+ mkApp (mkConstU (Projection.constant pr,u),
+ Array.of_list (ind_args @ (c :: args)))
diff --git a/pretyping/retyping.mli b/pretyping/retyping.mli
index 4ef54c13..89ba46db 100644
--- a/pretyping/retyping.mli
+++ b/pretyping/retyping.mli
@@ -1,14 +1,14 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Names
open Term
open Evd
+open Context
open Environ
(** This family of functions assumes its constr argument is known to be
@@ -20,8 +20,14 @@ open Environ
(** The "polyprop" optional argument is used by the extraction to
disable "Prop-polymorphism", cf comment in [inductive.ml] *)
+(** The "lax" optional argument provides a relaxed version of
+ [get_type_of] that won't raise any anomaly but RetypeError instead *)
+
+type retype_error
+exception RetypeError of retype_error
+
val get_type_of :
- ?polyprop:bool -> ?refresh:bool -> env -> evar_map -> constr -> types
+ ?polyprop:bool -> ?lax:bool -> env -> evar_map -> constr -> types
val get_sort_of :
?polyprop:bool -> env -> evar_map -> types -> sorts
@@ -29,9 +35,6 @@ val get_sort_of :
val get_sort_family_of :
?polyprop:bool -> env -> evar_map -> types -> sorts_family
-(** Makes an assumption from a constr *)
-val get_assumption_of : env -> evar_map -> constr -> types
-
(** Makes an unsafe judgment from a constr *)
val get_judgment_of : env -> evar_map -> constr -> unsafe_judgment
@@ -39,4 +42,8 @@ val type_of_global_reference_knowing_parameters : env -> evar_map -> constr ->
constr array -> types
val type_of_global_reference_knowing_conclusion :
- env -> evar_map -> constr -> types -> types
+ env -> evar_map -> constr -> types -> evar_map * types
+
+val sorts_of_context : env -> evar_map -> rel_context -> sorts list
+
+val expand_projection : env -> evar_map -> Names.projection -> constr -> constr list -> constr
diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml
index b18314f7..b4e0459c 100644
--- a/pretyping/tacred.ml
+++ b/pretyping/tacred.ml
@@ -1,34 +1,34 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Pp
+open Errors
open Util
open Names
-open Nameops
open Term
+open Vars
open Libnames
+open Globnames
open Termops
+open Find_subterm
open Namegen
-open Declarations
-open Inductive
-open Libobject
open Environ
open Closure
open Reductionops
open Cbv
-open Glob_term
-open Pattern
-open Matching
+open Patternops
+open Locus
+open Pretype_errors
(* Errors *)
type reduction_tactic_error =
- InvalidAbstraction of env * constr * (env * Type_errors.type_error)
+ InvalidAbstraction of env * Evd.evar_map * constr * (env * Type_errors.type_error)
exception ReductionTacticError of reduction_tactic_error
@@ -39,27 +39,28 @@ exception Redelimination
let error_not_evaluable r =
errorlabstrm "error_not_evaluable"
- (str "Cannot coerce" ++ spc () ++ Nametab.pr_global_env Idset.empty r ++
+ (str "Cannot coerce" ++ spc () ++ Nametab.pr_global_env Id.Set.empty r ++
spc () ++ str "to an evaluable reference.")
let is_evaluable_const env cst =
- is_transparent (ConstKey cst) && evaluable_constant cst env
+ is_transparent env (ConstKey cst) &&
+ (evaluable_constant cst env || is_projection cst env)
let is_evaluable_var env id =
- is_transparent (VarKey id) && evaluable_named id env
+ is_transparent env (VarKey id) && evaluable_named id env
let is_evaluable env = function
| EvalConstRef cst -> is_evaluable_const env cst
| EvalVarRef id -> is_evaluable_var env id
-let value_of_evaluable_ref env = function
- | EvalConstRef con -> constant_value env con
+let value_of_evaluable_ref env evref u =
+ match evref with
+ | EvalConstRef con ->
+ (try constant_value_in env (con,u)
+ with NotEvaluableConst IsProj ->
+ raise (Invalid_argument "value_of_evaluable_ref"))
| EvalVarRef id -> Option.get (pi2 (lookup_named id env))
-let constr_of_evaluable_ref = function
- | EvalConstRef con -> mkConst con
- | EvalVarRef id -> mkVar id
-
let evaluable_of_global_reference env = function
| ConstRef cst when is_evaluable_const env cst -> EvalConstRef cst
| VarRef id when is_evaluable_var env id -> EvalVarRef id
@@ -71,31 +72,55 @@ let global_of_evaluable_reference = function
type evaluable_reference =
| EvalConst of constant
- | EvalVar of identifier
+ | EvalVar of Id.t
| EvalRel of int
| EvalEvar of existential
-let mkEvalRef = function
- | EvalConst cst -> mkConst cst
+let evaluable_reference_eq r1 r2 = match r1, r2 with
+| EvalConst c1, EvalConst c2 -> eq_constant 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
+| _ -> false
+
+let mkEvalRef ref u =
+ match ref with
+ | EvalConst cst -> mkConstU (cst,u)
| EvalVar id -> mkVar id
| EvalRel n -> mkRel n
| EvalEvar ev -> mkEvar ev
let isEvalRef env c = match kind_of_term c with
- | Const sp -> is_evaluable env (EvalConstRef sp)
+ | Const (sp,_) -> is_evaluable env (EvalConstRef sp)
| Var id -> is_evaluable env (EvalVarRef id)
| Rel _ | Evar _ -> true
| _ -> false
-let destEvalRef c = match kind_of_term c with
- | Const cst -> EvalConst cst
- | Var id -> EvalVar id
- | Rel n -> EvalRel n
- | Evar ev -> EvalEvar ev
- | _ -> anomaly "Not an unfoldable reference"
+let destEvalRefU c = match kind_of_term 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")
+
+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)
+ | _ -> None)
+ | EvalVar id ->
+ let (_,v,_) = lookup_named id env in
+ v
+ | EvalRel n ->
+ let (_,v,_) = lookup_rel n env in
+ Option.map (lift n) v
+ | EvalEvar ev -> Evd.existential_opt_value sigma ev
-let reference_opt_value sigma env = function
- | EvalConst cst -> constant_opt_value env cst
+let reference_opt_value env sigma eval u =
+ match eval with
+ | EvalConst cst -> constant_opt_value_in env (cst,u)
| EvalVar id ->
let (_,v,_) = lookup_named id env in
v
@@ -105,8 +130,8 @@ let reference_opt_value sigma env = function
| EvalEvar ev -> Evd.existential_opt_value sigma ev
exception NotEvaluable
-let reference_value sigma env c =
- match reference_opt_value sigma env c with
+let reference_value env sigma c u =
+ match reference_opt_value env sigma c u with
| None -> raise NotEvaluable
| Some d -> d
@@ -121,28 +146,14 @@ type constant_evaluation =
((int*evaluable_reference) option array *
(int * (int * constr) list * int))
| EliminationCases of int
+ | EliminationProj of int
| NotAnElimination
(* We use a cache registered as a global table *)
-let eval_table = ref Cmap_env.empty
-
-type frozen = (int * constant_evaluation) Cmap_env.t
-
-let init () =
- eval_table := Cmap_env.empty
+type frozen = constant_evaluation Cmap.t
-let freeze () =
- !eval_table
-
-let unfreeze ct =
- eval_table := ct
-
-let _ =
- Summary.declare_summary "evaluation"
- { Summary.freeze_function = freeze;
- Summary.unfreeze_function = unfreeze;
- Summary.init_function = init }
+let eval_table = Summary.ref (Cmap.empty : frozen) ~name:"evaluation"
(* [compute_consteval] determines whether c is an "elimination constant"
@@ -177,8 +188,8 @@ let check_fix_reversibility labs args ((lv,i),(_,tys,bds)) =
(function d -> match kind_of_term d with
| Rel k ->
if
- array_for_all (noccurn k) tys
- && array_for_all (noccurn (k+nbfix)) bds
+ Array.for_all (noccurn k) tys
+ && Array.for_all (noccurn (k+nbfix)) bds
then
(k, List.nth labs (k-1))
else
@@ -187,12 +198,14 @@ let check_fix_reversibility labs args ((lv,i),(_,tys,bds)) =
raise Elimconst) args
in
let reversible_rels = List.map fst li in
- if not (list_distinct reversible_rels) then
+ if not (List.distinct_f Int.compare reversible_rels) then
raise Elimconst;
- list_iter_i (fun i t_i ->
- if not (List.mem_assoc (i+1) li) then
- let fvs = List.map ((+) (i+1)) (Intset.elements (free_rels t_i)) in
- if list_intersect fvs reversible_rels <> [] then 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
+ match List.intersect Int.equal fvs reversible_rels with
+ | [] -> ()
+ | _ -> raise Elimconst)
labs;
let k = lv.(i) in
if k < nargs then
@@ -210,57 +223,64 @@ let check_fix_reversibility labs args ((lv,i),(_,tys,bds)) =
let invert_name labs l na0 env sigma ref = function
| Name id ->
let minfxargs = List.length l in
- if na0 <> Name id then
+ begin match na0 with
+ | Name id' when Id.equal id' id ->
+ Some (minfxargs,ref)
+ | _ ->
let refi = match ref with
| EvalRel _ | EvalEvar _ -> None
| EvalVar id' -> Some (EvalVar id)
| EvalConst kn ->
- Some (EvalConst (con_with_label kn (label_of_id id))) in
+ Some (EvalConst (con_with_label kn (Label.of_id id))) in
match refi with
| None -> None
| Some ref ->
- try match reference_opt_value sigma env ref with
+ try match unsafe_reference_opt_value env sigma ref with
| None -> None
| Some c ->
let labs',ccl = decompose_lam c in
let _, l' = whd_betalet_stack sigma ccl in
let labs' = List.map snd labs' in
- if labs' = labs & l = l' then Some (minfxargs,ref)
+ (** ppedrot: there used to be generic equality on terms here *)
+ if List.equal eq_constr labs' labs &&
+ List.equal eq_constr l l' then Some (minfxargs,ref)
else None
with Not_found (* Undefined ref *) -> None
- else Some (minfxargs,ref)
+ end
| Anonymous -> None (* Actually, should not occur *)
(* [compute_consteval_direct] expand all constant in a whole, but
[compute_consteval_mutual_fix] only one by one, until finding the
last one before the Fix if the latter is mutually defined *)
-let compute_consteval_direct sigma env ref =
- let rec srec env n labs c =
+let compute_consteval_direct env sigma ref =
+ let rec srec env n labs onlyproj c =
let c',l = whd_betadelta_stack env sigma c in
match kind_of_term c' with
- | Lambda (id,t,g) when l=[] ->
- srec (push_rel (id,None,t) env) (n+1) (t::labs) g
- | Fix fix ->
+ | Lambda (id,t,g) when List.is_empty l && not onlyproj ->
+ srec (push_rel (id,None,t) env) (n+1) (t::labs) onlyproj g
+ | Fix fix when not onlyproj ->
(try check_fix_reversibility labs l fix
with Elimconst -> NotAnElimination)
- | Case (_,_,d,_) when isRel d -> EliminationCases n
+ | Case (_,_,d,_) when isRel d && not onlyproj -> EliminationCases n
+ | Case (_,_,d,_) -> srec env n labs true d
+ | Proj (p, d) when isRel d -> EliminationProj n
| _ -> NotAnElimination
in
- match reference_opt_value sigma env ref with
+ match unsafe_reference_opt_value env sigma ref with
| None -> NotAnElimination
- | Some c -> srec env 0 [] c
+ | Some c -> srec env 0 [] false c
-let compute_consteval_mutual_fix sigma env ref =
+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
- | Lambda (na,t,g) when l=[] ->
+ | Lambda (na,t,g) when List.is_empty l ->
srec (push_rel (na,None,t) env) (minarg+1) (t::labs) ref g
| Fix ((lv,i),(names,_,_)) ->
(* Last known constant wrapping Fix is ref = [labs](Fix l) *)
- (match compute_consteval_direct sigma env ref with
+ (match compute_consteval_direct env sigma ref with
| NotAnElimination -> (*Above const was eliminable but this not!*)
NotAnElimination
| EliminationFix (minarg',minfxargs,infos) ->
@@ -272,32 +292,32 @@ let compute_consteval_mutual_fix sigma env ref =
| _ -> assert false)
| _ when isEvalRef env c' ->
(* Forget all \'s and args and do as if we had started with c' *)
- let ref = destEvalRef c' in
- (match reference_opt_value sigma env ref with
- | None -> anomaly "Should have been trapped by compute_direct"
+ let ref,_ = destEvalRefU c' in
+ (match unsafe_reference_opt_value env sigma ref with
+ | None -> anomaly (Pp.str "Should have been trapped by compute_direct")
| Some c -> srec env (minarg-nargs) [] ref c)
| _ -> (* Should not occur *) NotAnElimination
in
- match reference_opt_value sigma env ref with
+ match unsafe_reference_opt_value env sigma ref with
| None -> (* Should not occur *) NotAnElimination
| Some c -> srec env 0 [] ref c
-let compute_consteval sigma env ref =
- match compute_consteval_direct sigma env ref with
- | EliminationFix (_,_,(nbfix,_,_)) when nbfix <> 1 ->
- compute_consteval_mutual_fix sigma env ref
+let compute_consteval env sigma ref =
+ match compute_consteval_direct env sigma ref with
+ | EliminationFix (_,_,(nbfix,_,_)) when not (Int.equal nbfix 1) ->
+ compute_consteval_mutual_fix env sigma ref
| elim -> elim
-let reference_eval sigma env = function
+let reference_eval env sigma = function
| EvalConst cst as ref ->
(try
- Cmap_env.find cst !eval_table
+ Cmap.find cst !eval_table
with Not_found -> begin
- let v = compute_consteval sigma env ref in
- eval_table := Cmap_env.add cst v !eval_table;
+ let v = compute_consteval env sigma ref in
+ eval_table := Cmap.add cst v !eval_table;
v
end)
- | ref -> compute_consteval sigma env ref
+ | ref -> compute_consteval env sigma ref
(* If f is bound to EliminationFix (n',infos), then n' is the minimal
number of args for starting the reduction and infos is
@@ -320,16 +340,16 @@ let reference_eval sigma env = function
The type Tij' is Tij[yi(j-1)..y1 <- ai(j-1)..a1]
*)
-let x = Name (id_of_string "x")
+let x = Name default_dependent_ident
-let make_elim_fun (names,(nbfix,lv,n)) largs =
- let lu = list_firstn n (list_of_stack largs) in
+let make_elim_fun (names,(nbfix,lv,n)) u largs =
+ let lu = List.firstn n largs in
let p = List.length lv in
let lyi = List.map fst lv in
let la =
- list_map_i (fun q aq ->
+ List.map_i (fun q aq ->
(* k from the comment is q+1 *)
- try mkRel (p+1-(list_index (n-q) lyi))
+ try mkRel (p+1-(List.index Int.equal (n-q) lyi))
with Not_found -> aq)
0 (List.map (lift p) lu)
in
@@ -337,10 +357,10 @@ let make_elim_fun (names,(nbfix,lv,n)) largs =
match names.(i) with
| None -> None
| Some (minargs,ref) ->
- let body = applistc (mkEvalRef ref) la in
+ let body = applistc (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
+ 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
mkLambda (x,tij',c)) 1 body (List.rev lv)
in Some (minargs,g)
@@ -349,37 +369,32 @@ let make_elim_fun (names,(nbfix,lv,n)) largs =
do so that the reduction uses this extra information *)
let dummy = mkProp
-let vfx = id_of_string"_expanded_fix_"
-let vfun = id_of_string"_eliminator_function_"
+let vfx = Id.of_string "_expanded_fix_"
+let vfun = Id.of_string "_eliminator_function_"
+let venv = val_of_named_context [(vfx, None, dummy); (vfun, None, dummy)]
(* Mark every occurrence of substituted vars (associated to a function)
as a problem variable: an evar that can be instantiated either by
vfx (expanded fixpoint) or vfun (named function). *)
-let substl_with_function subst constr =
- let cnt = ref 0 in
- let evd = ref Evd.empty in
- let minargs = ref Intmap.empty in
+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
- Rel i when k<i ->
- if i <= k + Array.length v then
- match v.(i-k-1) with
- | (fx,Some(min,ref)) ->
- decr cnt;
- evd := Evd.add !evd !cnt
- (Evd.make_evar
- (val_of_named_context
- [(vfx,None,dummy);(vfun,None,dummy)])
- dummy);
- minargs := Intmap.add !cnt min !minargs;
- lift k (mkEvar(!cnt,[|fx;ref|]))
- | (fx,None) -> lift k fx
- else mkRel (i - Array.length v)
- | _ ->
- map_constr_with_binders succ subst_total k c in
+ let rec subst_total k c = match kind_of_term 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, evk) = Evarutil.new_pure_evar venv !evd dummy in
+ evd := sigma;
+ minargs := Evar.Map.add evk min !minargs;
+ lift k (mkEvar (evk, [|fx;ref|]))
+ | (fx, None) -> lift k fx
+ else mkRel (i - Array.length v)
+ | _ ->
+ map_constr_with_binders succ subst_total k c in
let c = subst_total 0 constr in
- (c,!evd,!minargs)
+ (c, !evd, !minargs)
exception Partial
@@ -392,15 +407,16 @@ let solve_arity_problem env sigma fxminargs c =
let c' = whd_betaiotazeta sigma c in
let (h,rcargs) = decompose_app c' in
match kind_of_term h with
- Evar(i,_) when Intmap.mem i fxminargs && not (Evd.is_defined !evm i) ->
- let minargs = Intmap.find i fxminargs in
+ 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 strict then set_fix i
else raise Partial;
List.iter (check strict) rcargs
| (Var _|Const _) when isEvalRef env h ->
- (match reference_opt_value sigma env (destEvalRef h) with
- Some h' ->
+ (let ev, u = destEvalRefU h in
+ match reference_opt_value env sigma ev u with
+ | Some h' ->
let bak = !evm in
(try List.iter (check false) rcargs
with Partial ->
@@ -411,42 +427,52 @@ let solve_arity_problem env sigma fxminargs c =
check true c;
!evm
-let substl_checking_arity env subst c =
+let substl_checking_arity env subst sigma c =
(* we initialize the problem: *)
- let body,sigma,minargs = substl_with_function subst c in
+ let body,sigma,minargs = substl_with_function subst sigma c in
(* we collect arity constraints *)
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 Intmap.mem i minargs ->
+ 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
nf_fix body
+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
+ | None -> NotReducible
+ | 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')
+ | _ -> NotReducible)
let contract_fix_use_function env sigma f
((recindices,bodynum),(_names,_types,bodies as typedbodies)) =
let nbodies = Array.length recindices in
let make_Fi j = (mkFix((recindices,j),typedbodies), f j) in
- let lbodies = list_tabulate make_Fi nbodies in
- substl_checking_arity env (List.rev lbodies) (nf_beta sigma bodies.(bodynum))
+ let lbodies = List.init nbodies make_Fi in
+ substl_checking_arity env (List.rev lbodies) sigma (nf_beta sigma bodies.(bodynum))
let reduce_fix_use_function env sigma f whfun fix stack =
- match fix_recarg fix stack with
+ 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
(* The recarg cannot be a local def, no worry about the right env *)
- (recarg, empty_stack)
+ (recarg, [])
else
- whfun (recarg, empty_stack) in
- let stack' = stack_assign stack recargnum (app_stack recarg') in
+ whfun recarg in
+ let stack' = List.assign stack recargnum (applist recarg') in
(match kind_of_term recarg'hd with
| Construct _ ->
Reduced (contract_fix_use_function env sigma f fix,stack')
@@ -456,21 +482,21 @@ let contract_cofix_use_function env sigma f
(bodynum,(_names,_,bodies as typedbodies)) =
let nbodies = Array.length bodies in
let make_Fi j = (mkCoFix(j,typedbodies), f j) in
- let subbodies = list_tabulate make_Fi nbodies in
+ let subbodies = List.init nbodies make_Fi in
substl_checking_arity env (List.rev subbodies)
- (nf_beta sigma bodies.(bodynum))
+ sigma (nf_beta sigma bodies.(bodynum))
let reduce_mind_case_use_function func env sigma mia =
match kind_of_term mia.mconstr with
- | Construct(ind_sp,i) ->
- let real_cargs = list_skipn mia.mci.ci_npar mia.mcargs in
+ | 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
let minargs = List.length mia.mcargs in
fun i ->
- if i = bodynum then Some (minargs,func)
+ if Int.equal i bodynum then Some (minargs,func)
else match names.(i) with
| Anonymous -> None
| Name id ->
@@ -478,12 +504,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 = con_with_label (destConst func) (label_of_id id)
+ let kn = map_puniverses (fun x -> con_with_label x (Label.of_id id))
+ (destConst func)
in
- try match constant_opt_value env kn with
+ try match constant_opt_value_in env kn with
| None -> None
(* TODO: check kn is correct *)
- | Some _ -> Some (minargs,mkConst kn)
+ | Some _ -> Some (minargs,mkConstU kn)
with Not_found -> None
else
fun _ -> None in
@@ -492,226 +519,277 @@ let reduce_mind_case_use_function func env sigma mia =
mkCase (mia.mci, mia.mP, applist(cofix_def,mia.mcargs), mia.mlf)
| _ -> assert false
-let special_red_case env sigma whfun (ci, p, c, lf) =
+
+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)
+ | _ -> 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))
+ | Var id when is_evaluable env (EvalVarRef id) ->
+ let (_,v,_) = lookup_named id env in v
+ | Rel n -> let (_,v,_) = lookup_rel n env in
+ Option.map (lift n) v
+ | Evar ev -> Evd.existential_opt_value sigma ev
+ | _ -> None
+
+let special_red_case env sigma whfun (ci, p, c, lf) =
let rec redrec s =
let (constr, cargs) = whfun s in
- if isEvalRef env constr then
- let ref = destEvalRef constr in
- match reference_opt_value sigma env ref with
- | None -> raise Redelimination
- | Some gvalue ->
- if reducible_mind_case gvalue then
- reduce_mind_case_use_function constr env sigma
- {mP=p; mconstr=gvalue; mcargs=list_of_stack cargs;
- mci=ci; mlf=lf}
- else
- redrec (gvalue, cargs)
- else
+ match match_eval_ref env constr with
+ | Some (ref, u) ->
+ (match reference_opt_value env sigma ref u with
+ | None -> raise Redelimination
+ | Some gvalue ->
+ if reducible_mind_case 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
- {mP=p; mconstr=constr; mcargs=list_of_stack cargs;
+ {mP=p; mconstr=constr; mcargs=cargs;
mci=ci; mlf=lf}
else
raise Redelimination
in
- redrec (c, empty_stack)
-
-(* data structure to hold the map kn -> rec_args for simpl *)
-
-type behaviour = {
- b_nargs: int;
- b_recargs: int list;
- b_dont_expose_case: bool;
-}
-
-let behaviour_table = ref (Refmap.empty : behaviour Refmap.t)
-
-let _ =
- Summary.declare_summary "simplbehaviour"
- { Summary.freeze_function = (fun () -> !behaviour_table);
- Summary.unfreeze_function = (fun x -> behaviour_table := x);
- Summary.init_function = (fun () -> behaviour_table := Refmap.empty) }
-
-type simpl_flag = [ `SimplDontExposeCase | `SimplNeverUnfold ]
-type req =
- | ReqLocal
- | ReqGlobal of global_reference * (int list * int * simpl_flag list)
-
-let load_simpl_behaviour _ (_,(_,(r, b))) =
- behaviour_table := Refmap.add r b !behaviour_table
-
-let cache_simpl_behaviour o = load_simpl_behaviour 1 o
-
-let classify_simpl_behaviour = function
- | ReqLocal, _ -> Dispose
- | ReqGlobal _, _ as o -> Substitute o
-
-let subst_simpl_behaviour (subst, (_, (r,o as orig))) =
- ReqLocal,
- let r' = fst (subst_global subst r) in if r==r' then orig else (r',o)
-
-let discharge_simpl_behaviour = function
- | _,(ReqGlobal (ConstRef c, req), (_, b)) ->
- let c' = pop_con c in
- let vars = Lib.section_segment_of_constant c in
- let extra = List.length vars in
- let nargs' = if b.b_nargs < 0 then b.b_nargs else b.b_nargs + extra in
- let recargs' = List.map ((+) extra) b.b_recargs in
- let b' = { b with b_nargs = nargs'; b_recargs = recargs' } in
- Some (ReqGlobal (ConstRef c', req), (ConstRef c', b'))
- | _ -> None
+ redrec c
-let rebuild_simpl_behaviour = function
- | req, (ConstRef c, _ as x) -> req, x
- | _ -> assert false
-
-let inSimplBehaviour = declare_object { (default_object "SIMPLBEHAVIOUR") with
- load_function = load_simpl_behaviour;
- cache_function = cache_simpl_behaviour;
- classify_function = classify_simpl_behaviour;
- subst_function = subst_simpl_behaviour;
- discharge_function = discharge_simpl_behaviour;
- rebuild_function = rebuild_simpl_behaviour;
-}
-
-let set_simpl_behaviour local r (recargs, nargs, flags as req) =
- let nargs = if List.mem `SimplNeverUnfold flags then max_int else nargs in
- let behaviour = {
- b_nargs = nargs; b_recargs = recargs;
- b_dont_expose_case = List.mem `SimplDontExposeCase flags } in
- let req = if local then ReqLocal else ReqGlobal (r, req) in
- Lib.add_anonymous_leaf (inSimplBehaviour (req, (r, behaviour)))
-;;
-
-let get_simpl_behaviour r =
- try
- let b = Refmap.find r !behaviour_table in
- let flags =
- if b.b_nargs = max_int then [`SimplNeverUnfold]
- else if b.b_dont_expose_case then [`SimplDontExposeCase] else [] in
- Some (b.b_recargs, (if b.b_nargs = max_int then -1 else b.b_nargs), flags)
- with Not_found -> None
+let recargs = function
+ | EvalVar _ | EvalRel _ | EvalEvar _ -> None
+ | EvalConst c -> ReductionBehaviour.get (ConstRef c)
-let get_behaviour = function
- | EvalVar _ | EvalRel _ | EvalEvar _ -> raise Not_found
- | EvalConst c -> Refmap.find (ConstRef c) !behaviour_table
+let reduce_projection env sigma pb (recarg'hd,stack') stack =
+ (match kind_of_term recarg'hd with
+ | Construct _ ->
+ let proj_narg =
+ pb.Declarations.proj_npars + pb.Declarations.proj_arg
+ in Reduced (List.nth stack' proj_narg, stack)
+ | _ -> NotReducible)
+
+let reduce_proj env sigma whfun whfun' c =
+ let rec redrec s =
+ match kind_of_term 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
+ | Construct _ ->
+ let proj_narg =
+ let pb = lookup_projection proj env in
+ pb.Declarations.proj_npars + pb.Declarations.proj_arg
+ in List.nth cargs proj_narg
+ | _ -> raise Redelimination)
+ | Case (n,p,c,brs) ->
+ let c' = redrec c in
+ let p = (n,p,c',brs) in
+ (try special_red_case env sigma whfun' p
+ with Redelimination -> mkCase p)
+ | _ -> raise Redelimination
+ in redrec c
-let recargs r =
- try let b = get_behaviour r in Some (b.b_recargs, b.b_nargs)
- with Not_found -> None
-let dont_expose_case r =
- try (get_behaviour r).b_dont_expose_case with Not_found -> false
+let dont_expose_case = function
+ | EvalVar _ | EvalRel _ | EvalEvar _ -> false
+ | EvalConst c ->
+ Option.cata (fun (_,_,z) -> List.mem `ReductionDontExposeCase z)
+ false (ReductionBehaviour.get (ConstRef c))
+
+let whd_nothing_for_iota env sigma s =
+ let rec whrec (x, stack as s) =
+ match kind_of_term x with
+ | Rel n ->
+ (match lookup_rel n env with
+ | (_,Some body,_) -> whrec (lift n body, stack)
+ | _ -> s)
+ | Var id ->
+ (match lookup_named id env with
+ | (_,Some body,_) -> whrec (body, stack)
+ | _ -> s)
+ | Evar ev ->
+ (try whrec (Evd.existential_value sigma ev, stack)
+ with Evd.NotInstantiatedEvar | Not_found -> s)
+ | Meta ev ->
+ (try whrec (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)
+ | None -> s)
+ | LetIn (_,b,_,c) -> stacklam whrec [b] 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
+ | _ -> s)
+
+ | x -> s
+ in
+ decompose_app (Stack.zip (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;
it fails if no redex is around *)
-let rec red_elim_const env sigma ref largs =
- let nargs = stack_args_size largs in
- let largs, unfold_anyway, unfold_nonelim =
+let rec red_elim_const env sigma ref u largs =
+ let nargs = List.length largs in
+ let largs, unfold_anyway, unfold_nonelim, nocase =
match recargs ref with
- | None -> largs, false, false
- | Some (_,n) when nargs < n -> raise Redelimination
- | Some (x::l,_) when nargs <= List.fold_left max x l -> raise Redelimination
- | Some (l,n) ->
- List.fold_left (fun stack i ->
- let arg = stack_nth stack i in
- let rarg = whd_construct_state env sigma (arg, empty_stack) in
- match kind_of_term (fst rarg) with
- | Construct _ -> stack_assign stack i (app_stack rarg)
- | _ -> raise Redelimination)
- largs l, n >= 0 && l = [] && nargs >= n,
- n >= 0 && l <> [] && nargs >= n in
- try match reference_eval sigma env ref with
+ | None -> largs, false, false, false
+ | Some (_,n,f) when nargs < n || List.mem `ReductionNeverUnfold f -> raise Redelimination
+ | Some (x::l,_,_) when nargs <= List.fold_left max x l -> raise Redelimination
+ | Some (l,n,f) ->
+ let is_empty = match l with [] -> true | _ -> false in
+ reduce_params env sigma largs l,
+ n >= 0 && is_empty && nargs >= n,
+ n >= 0 && not is_empty && nargs >= n,
+ List.mem `ReductionDontExposeCase f
+ in
+ try match reference_eval env sigma ref with
| EliminationCases n when nargs >= n ->
- let c = reference_value sigma env ref in
- let c', lrest = whd_betadelta_state env sigma (c,largs) in
- let whfun = whd_simpl_state env sigma in
- (special_red_case env sigma whfun (destCase c'), lrest)
+ 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
+ | 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
+ let whfun = whd_construct_stack env sigma in
+ let whfun' = whd_simpl_stack env sigma in
+ (reduce_proj env sigma whfun whfun' c', lrest), nocase
| EliminationFix (min,minfxargs,infos) when nargs >= min ->
- let c = reference_value sigma env ref in
- let d, lrest = whd_betadelta_state env sigma (c,largs) in
- let f = make_elim_fun ([|Some (minfxargs,ref)|],infos) largs in
- let whfun = whd_construct_state env sigma in
+ let c = reference_value env sigma ref u in
+ 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
| NotReducible -> raise Redelimination
- | Reduced (c,rest) -> (nf_beta sigma c, rest))
+ | Reduced (c,rest) -> (nf_beta sigma c, rest), nocase)
| EliminationMutualFix (min,refgoal,refinfos) when nargs >= min ->
- let rec descend ref args =
- let c = reference_value sigma env ref in
- if ref = refgoal then
+ let rec descend (ref,u) args =
+ let c = reference_value env sigma ref u in
+ if evaluable_reference_eq ref refgoal then
(c,args)
else
- let c', lrest = whd_betalet_state sigma (c,args) in
- descend (destEvalRef c') lrest in
- let (_, midargs as s) = descend ref largs in
- let d, lrest = whd_betadelta_state env sigma s in
- let f = make_elim_fun refinfos midargs in
- let whfun = whd_construct_state env sigma in
+ let c', lrest = whd_betalet_stack sigma (applist(c,args)) in
+ descend (destEvalRefU 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
| NotReducible -> raise Redelimination
- | Reduced (c,rest) -> (nf_beta sigma c, rest))
- | NotAnElimination when unfold_nonelim ->
- let c = reference_value sigma env ref in
- whd_betaiotazeta sigma (app_stack (c, largs)), empty_stack
+ | Reduced (c,rest) -> (nf_beta sigma c, rest), nocase)
+ | NotAnElimination when unfold_nonelim ->
+ let c = reference_value env sigma ref u in
+ (whd_betaiotazeta sigma (applist (c, largs)), []), nocase
| _ -> raise Redelimination
with Redelimination when unfold_anyway ->
- let c = reference_value sigma env ref in
- whd_betaiotazeta sigma (app_stack (c, largs)), empty_stack
+ let c = reference_value env sigma ref u in
+ (whd_betaiotazeta sigma (applist (c, largs)), []), nocase
+
+and reduce_params env sigma stack l =
+ let len = List.length stack in
+ List.fold_left (fun stack i ->
+ if len <= i then raise Redelimination
+ else
+ let arg = List.nth stack i in
+ let rarg = whd_construct_stack env sigma arg in
+ match kind_of_term (fst rarg) with
+ | Construct _ -> List.assign stack i (applist rarg)
+ | _ -> raise Redelimination)
+ stack l
+
(* reduce to whd normal form or to an applied constant that does not hide
a reducible iota/fix/cofix redex (the "simpl" tactic) *)
-and whd_simpl_state env sigma s =
- let rec redrec (x, stack as s) =
+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
| Lambda (na,t,c) ->
- (match decomp_stack stack with
- | None -> s
- | Some (a,rest) -> stacklam redrec [a] c rest)
- | LetIn (n,b,t,c) -> stacklam redrec [b] c stack
- | App (f,cl) -> redrec (f, append_stack cl stack)
- | Cast (c,_,_) -> redrec (c, stack)
+ (match stack with
+ | [] -> s'
+ | a :: rest -> redrec (beta_applist (x,stack)))
+ | LetIn (n,b,t,c) -> redrec (applist (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) ->
(try
- redrec (special_red_case env sigma redrec (ci,p,c,lf), stack)
+ redrec (applist(special_red_case env sigma redrec (ci,p,c,lf), stack))
with
- Redelimination -> s)
+ Redelimination -> s')
| Fix fix ->
- (try match reduce_fix (whd_construct_state env) sigma fix stack with
- | Reduced s' -> redrec s'
- | NotReducible -> s
- with Redelimination -> s)
- | _ when isEvalRef env x ->
- let ref = destEvalRef x in
+ (try match reduce_fix (whd_construct_stack env) sigma fix stack with
+ | Reduced s' -> redrec (applist s')
+ | NotReducible -> s'
+ with Redelimination -> s')
+
+ | Proj (p, c) ->
+ (try
+ let unf = Projection.unfolded p in
+ if unf || is_evaluable env (EvalConstRef (Projection.constant p)) then
+ let pb = lookup_projection p env in
+ (match unf, ReductionBehaviour.get (ConstRef (Projection.constant p)) with
+ | false, Some (l, n, f) when List.mem `ReductionNeverUnfold f ->
+ (* simpl never *) s'
+ | false, Some (l, n, f) when not (List.is_empty l) ->
+ let l' = List.map_filter (fun i ->
+ let idx = (i - (pb.Declarations.proj_npars + 1)) in
+ if idx < 0 then None else Some idx) l in
+ let stack = reduce_params env sigma stack l' in
+ (match reduce_projection env sigma pb
+ (whd_construct_stack env sigma c) stack
+ with
+ | Reduced s' -> redrec (applist s')
+ | NotReducible -> s')
+ | _ ->
+ match reduce_projection env sigma pb (whd_construct_stack env sigma c) stack with
+ | Reduced s' -> redrec (applist s')
+ | NotReducible -> s')
+ else s'
+ with Redelimination -> s')
+
+ | _ ->
+ match match_eval_ref env x with
+ | Some (ref, u) ->
(try
- let hd, _ as s' = redrec (red_elim_const env sigma ref stack) in
- let rec is_case x = match kind_of_term x with
- | Lambda (_,_, x) | LetIn (_,_,_, x) | Cast (x, _,_) -> is_case x
- | App (hd, _) -> is_case hd
- | Case _ -> true
- | _ -> false in
- if dont_expose_case ref && is_case hd then raise Redelimination
- else s'
- with Redelimination ->
- s)
- | _ -> s
+ 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
+ | Lambda (_,_, x) | LetIn (_,_,_, x) | Cast (x, _,_) -> is_case x
+ | App (hd, _) -> is_case hd
+ | Case _ -> true
+ | _ -> false in
+ if nocase && is_case hd then raise Redelimination
+ else s''
+ with Redelimination -> s')
+ | None -> s'
in
- redrec s
+ redrec
(* reduce until finding an applied constructor or fail *)
-and whd_construct_state env sigma s =
- let (constr, cargs as s') = whd_simpl_state env sigma s in
+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 if isEvalRef env constr then
- let ref = destEvalRef constr in
- match reference_opt_value sigma env ref with
- | None -> raise Redelimination
- | Some gvalue -> whd_construct_state env sigma (gvalue, cargs)
- else
- raise Redelimination
+ else match match_eval_ref env constr with
+ | Some (ref, u) ->
+ (match reference_opt_value env sigma ref u with
+ | None -> raise Redelimination
+ | Some gvalue -> whd_construct_stack env sigma (applist(gvalue, cargs)))
+ | _ -> raise Redelimination
(************************************************************************)
(* Special Purpose Reduction Strategies *)
@@ -724,35 +802,47 @@ and whd_construct_state env sigma s =
let try_red_product env sigma c =
let simpfun = clos_norm_flags betaiotazeta env sigma in
let rec redrec env x =
+ let x = whd_betaiota sigma x in
match kind_of_term x with
| App (f,l) ->
(match kind_of_term f with
| Fix fix ->
- let stack = append_stack l empty_stack in
+ let stack = Stack.append_app l Stack.empty in
(match fix_recarg fix stack with
| None -> raise Redelimination
| Some (recargnum,recarg) ->
let recarg' = redrec env recarg in
- let stack' = stack_assign stack recargnum recarg' in
- simpfun (app_stack (f,stack')))
+ let stack' = Stack.assign stack recargnum recarg' in
+ simpfun (Stack.zip (f,stack')))
| _ -> simpfun (appvect (redrec env f, l)))
| Cast (c,_,_) -> redrec env c
| Prod (x,a,b) -> mkProd (x, a, redrec (push_rel (x,None,a) env) b)
| LetIn (x,a,b,t) -> redrec env (subst1 a t)
| Case (ci,p,d,lf) -> simpfun (mkCase (ci,p,redrec env d,lf))
- | _ when isEvalRef env x ->
+ | Proj (p, c) ->
+ let c' =
+ match kind_of_term c with
+ | Construct _ -> c
+ | _ -> redrec env c
+ in
+ let pb = lookup_projection p env in
+ (match reduce_projection env sigma pb (whd_betaiotazeta_stack sigma c') [] with
+ | Reduced s -> simpfun (applist s)
+ | NotReducible -> raise Redelimination)
+ | _ ->
+ (match match_eval_ref env x with
+ | Some (ref, u) ->
(* TO DO: re-fold fixpoints after expansion *)
(* to get true one-step reductions *)
- let ref = destEvalRef x in
- (match reference_opt_value sigma env ref with
+ (match reference_opt_value env sigma ref u with
| None -> raise Redelimination
| Some c -> c)
- | _ -> raise Redelimination
+ | _ -> raise Redelimination)
in redrec env c
let red_product env sigma c =
try try_red_product env sigma c
- with Redelimination -> error "Not reducible."
+ with Redelimination -> error "No head constant to reduce."
(*
(* This old version of hnf uses betadeltaiota instead of itself (resp
@@ -798,7 +888,7 @@ let whd_simpl_orelse_delta_but_fix_old env sigma c =
(try
redrec (red_elim_const env sigma ref stack)
with Redelimination ->
- match reference_opt_value sigma env ref with
+ match reference_opt_value env sigma ref with
| Some c ->
(match kind_of_term (strip_lam c) with
| CoFix _ | Fix _ -> s
@@ -808,96 +898,160 @@ let whd_simpl_orelse_delta_but_fix_old env sigma c =
in app_stack (redrec (c, empty_stack))
*)
+let whd_simpl_stack =
+ if Flags.profile then
+ let key = Profile.declare_profile "whd_simpl_stack" in
+ Profile.profile3 key whd_simpl_stack
+ else whd_simpl_stack
+
(* Same as [whd_simpl] but also reduces constants that do not hide a
reducible fix, but does this reduction of constants only until it
immediately hides a non reducible fix or a cofix *)
let whd_simpl_orelse_delta_but_fix env sigma c =
let rec redrec s =
- let (constr, stack as s') = whd_simpl_state env sigma s in
- if isEvalRef env constr then
- match reference_opt_value sigma env (destEvalRef constr) with
- | Some c ->
- (match kind_of_term (strip_lam c) with
- | CoFix _ | Fix _ -> s'
- | _ -> redrec (c, stack))
- | None -> s'
- else s'
- in app_stack (redrec (c, empty_stack))
+ let (constr, stack as s') = whd_simpl_stack env sigma s in
+ match match_eval_ref_value env sigma constr with
+ | Some c ->
+ (match kind_of_term (strip_lam c) with
+ | CoFix _ | Fix _ -> s'
+ | Proj (p,t) when
+ (match kind_of_term constr with
+ | Const (c', _) -> eq_constant (Projection.constant p) c'
+ | _ -> false) ->
+ let pb = Environ.lookup_projection p env in
+ if List.length stack <= pb.Declarations.proj_npars then
+ (** Do not show the eta-expanded form *)
+ s'
+ else redrec (applist (c, stack))
+ | _ -> redrec (applist(c, stack)))
+ | None -> s'
+ in
+ let simpfun = clos_norm_flags betaiota env sigma in
+ simpfun (applist (redrec c))
let hnf_constr = whd_simpl_orelse_delta_but_fix
(* The "simpl" reduction tactic *)
let whd_simpl env sigma c =
- app_stack (whd_simpl_state env sigma (c, empty_stack))
+ applist (whd_simpl_stack env sigma c)
let simpl env sigma c = strong whd_simpl env sigma c
(* Reduction at specific subterms *)
-let matches_head c t =
+let matches_head env sigma c t =
match kind_of_term t with
- | App (f,_) -> matches c f
- | _ -> raise PatternMatchingFailure
+ | App (f,_) -> Constr_matching.matches env sigma c f
+ | Proj (p, _) -> Constr_matching.matches env sigma c (mkConst (Projection.constant p))
+ | _ -> raise Constr_matching.PatternMatchingFailure
+
+let is_pattern_meta = function Pattern.PMeta _ -> true | _ -> false
-let contextually byhead ((nowhere_except_in,locs),c) f env sigma t =
+(** FIXME: Specific function to handle projections: it ignores what happens on the
+ parameters. This is a temporary fix while rewrite etc... are not up to equivalence
+ 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
+ | 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 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
+ | 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
+
+let e_contextually byhead (occs,c) f 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 rec traverse (env,c as envc) t =
- if nowhere_except_in & (!pos > maxocc) then t
+ let evd = ref sigma in
+ let rec traverse nested (env,c as envc) t =
+ if nowhere_except_in && (!pos > maxocc) then (* Shortcut *) t
else
try
- let subst = if byhead then matches_head c t else matches c t in
+ let subst =
+ if byhead then matches_head env sigma c t
+ else Constr_matching.matches env sigma c t in
let ok =
- if nowhere_except_in then List.mem !pos locs
- else not (List.mem !pos locs) in
+ if nowhere_except_in then Int.List.mem !pos locs
+ else not (Int.List.mem !pos locs) in
incr pos;
- if ok then
- f subst env sigma t
- else if byhead then
- (* find other occurrences of c in t; TODO: ensure left-to-right *)
- let (f,l) = destApp t in
- mkApp (f, array_map_left (traverse envc) l)
+ 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 ".");
+ (* Skip inner occurrences for stable counting of occurrences *)
+ if locs != [] then
+ ignore (traverse_below (Some (!pos-1)) envc t);
+ let evm, t = f subst env !evd t in
+ (evd := evm; t)
+ end
else
- t
- with PatternMatchingFailure ->
- map_constr_with_binders_left_to_right
- (fun d (env,c) -> (push_rel d env,lift_pattern 1 c))
- traverse envc t
+ traverse_below nested envc t
+ with Constr_matching.PatternMatchingFailure ->
+ traverse_below nested envc 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
+ | 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)
+ | _ ->
+ change_map_constr_with_binders_left_to_right
+ (fun d (env,c) -> (push_rel d env,lift_pattern 1 c))
+ (traverse nested) envc sigma t
in
- let t' = traverse (env,c) t in
+ let t' = traverse None (env,c) t in
if List.exists (fun o -> o >= !pos) locs then error_invalid_occurrence locs;
- t'
+ !evd, t'
+
+let contextually byhead occs f env sigma t =
+ 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 occurence of name.
* ol is the occurence list to find. *)
-let substlin env evalref n (nowhere_except_in,locs) c =
+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
+ | _, _ -> None
+
+let substlin env sigma evalref n (nowhere_except_in,locs) c =
let maxocc = List.fold_right max locs 0 in
let pos = ref n in
assert (List.for_all (fun x -> x >= 0) locs);
- let value = value_of_evaluable_ref env evalref in
- let term = constr_of_evaluable_ref evalref in
+ let value u = value_of_evaluable_ref env evalref u in
let rec substrec () c =
- if nowhere_except_in & !pos > maxocc then c
- else if eq_constr c term then
- let ok =
- if nowhere_except_in then List.mem !pos locs
- else not (List.mem !pos locs) in
- incr pos;
- if ok then value else c
- else
- map_constr_with_binders_left_to_right
- (fun _ () -> ())
- substrec () c
+ if nowhere_except_in && !pos > maxocc then c
+ else
+ match match_constr_evaluable_ref sigma c evalref with
+ | Some u ->
+ let ok =
+ if nowhere_except_in then Int.List.mem !pos locs
+ else not (Int.List.mem !pos locs) in
+ incr pos;
+ if ok then value u else c
+ | None ->
+ map_constr_with_binders_left_to_right
+ (fun _ () -> ())
+ substrec () c
in
let t' = substrec () c in
(!pos, t')
let string_of_evaluable_ref env = function
- | EvalVarRef id -> string_of_id id
+ | EvalVarRef id -> Id.to_string id
| EvalConstRef kn ->
string_of_qualid
(Nametab.shortest_qualid_of_global (vars_of_env env) (ConstRef kn))
@@ -908,19 +1062,31 @@ let unfold env sigma name =
else
error (string_of_evaluable_ref env name^" is opaque.")
+let is_projection env = function
+ | EvalVarRef _ -> false
+ | EvalConstRef c -> Environ.is_projection c env
+
(* [unfoldoccs : (readable_constraints -> (int list * full_path) -> constr -> constr)]
* Unfolds the constant name in a term c following a list of occurrences occl.
* at the occurrences of occ_list. If occ_list is empty, unfold all occurences.
* Performs a betaiota reduction after unfolding. *)
-let unfoldoccs env sigma ((nowhere_except_in,locs as plocs),name) c =
- if locs = [] then if nowhere_except_in then c else unfold env sigma name c
- else
- let (nbocc,uc) = substlin env name 1 plocs c in
- if nbocc = 1 then
+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.");
let rest = List.filter (fun o -> o >= nbocc) locs in
- if rest <> [] then error_invalid_occurrence rest;
- nf_betaiota sigma uc
+ let () = match rest with
+ | [] -> ()
+ | _ -> error_invalid_occurrence rest
+ in
+ nf_betaiotazeta sigma uc
+ in
+ match occs with
+ | NoOccurrences -> c
+ | AllOccurrences -> unfold env sigma name c
+ | OnlyOccurrences l -> unfo true l
+ | AllOccurrencesBut l -> unfo false l
(* Unfold reduction tactic: *)
let unfoldn loccname env sigma c =
@@ -962,25 +1128,39 @@ 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 sigma (locc,a) c =
+let abstract_scheme env (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
- mkLambda (na,ta,c)
+ mkLambda (na,ta,c), sigma
else
- mkLambda (na,ta,subst_closed_term_occ locc a c)
+ let c', sigma' = subst_closed_term_occ env sigma (AtOccs locc) a c in
+ mkLambda (na,ta,c'), sigma'
let pattern_occs loccs_trm env sigma c =
- let abstr_trm = List.fold_right (abstract_scheme env sigma) loccs_trm c in
+ let abstr_trm, sigma = List.fold_right (abstract_scheme env) loccs_trm (c,sigma) in
try
let _ = Typing.type_of env sigma abstr_trm in
- applist(abstr_trm, List.map snd loccs_trm)
+ sigma, applist(abstr_trm, List.map snd loccs_trm)
with Type_errors.TypeError (env',t) ->
- raise (ReductionTacticError (InvalidAbstraction (env,abstr_trm,(env',t))))
+ raise (ReductionTacticError (InvalidAbstraction (env,sigma,abstr_trm,(env',t))))
(* 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.")
+ 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: " ++
+ str "use projections or let instead.")
+ else ind
+
(* put t as t'=(x1:A1)..(xn:An)B with B an inductive definition of name name
return name, B and t' *)
@@ -988,7 +1168,7 @@ 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
- | Ind ind-> (ind, it_mkProd_or_LetIn t l)
+ | Ind ind-> (check_privacy env ind, it_mkProd_or_LetIn t l)
| Prod (n,ty,t') ->
if allow_product then
elimrec (push_rel (n,None,ty) env) t' ((n,None,ty)::l)
@@ -999,7 +1179,7 @@ let reduce_to_ind_gen allow_product env sigma t =
was partially the case between V5.10 and V8.1 *)
let t' = whd_betadeltaiota env sigma t in
match kind_of_term (fst (decompose_app t')) with
- | Ind ind-> (ind, it_mkProd_or_LetIn t' l)
+ | Ind ind-> (check_privacy env ind, it_mkProd_or_LetIn t' l)
| _ -> errorlabstrm "" (str"Not an inductive product.")
in
elimrec env t []
@@ -1007,7 +1187,7 @@ let reduce_to_ind_gen allow_product env sigma 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 rec find_hnf_rectype env sigma t =
+let find_hnf_rectype env sigma t =
let ind,t = reduce_to_atomic_ind env sigma t in
ind, snd (decompose_app t)
@@ -1020,69 +1200,66 @@ let one_step_reduce env sigma c =
let rec redrec (x, stack) =
match kind_of_term x with
| Lambda (n,t,c) ->
- (match decomp_stack stack with
- | None -> raise NotStepReducible
- | Some (a,rest) -> (subst1 a c, rest))
- | App (f,cl) -> redrec (f, append_stack cl stack)
+ (match stack with
+ | [] -> raise NotStepReducible
+ | a :: rest -> (subst1 a c, rest))
+ | App (f,cl) -> redrec (f, (Array.to_list cl)@stack)
| LetIn (_,f,_,cl) -> (subst1 f cl,stack)
| Cast (c,_,_) -> redrec (c,stack)
| Case (ci,p,c,lf) ->
(try
- (special_red_case env sigma (whd_simpl_state env sigma)
+ (special_red_case env sigma (whd_simpl_stack env sigma)
(ci,p,c,lf), stack)
with Redelimination -> raise NotStepReducible)
| Fix fix ->
- (match reduce_fix (whd_construct_state env) sigma fix stack with
+ (match reduce_fix (whd_construct_stack env) sigma fix stack with
| Reduced s' -> s'
| NotReducible -> raise NotStepReducible)
| _ when isEvalRef env x ->
- let ref = destEvalRef x in
+ let ref,u = destEvalRefU x in
(try
- red_elim_const env sigma ref stack
+ fst (red_elim_const env sigma ref u stack)
with Redelimination ->
- match reference_opt_value sigma env ref with
- | Some d -> d, stack
+ match reference_opt_value env sigma ref u with
+ | Some d -> (d, stack)
| None -> raise NotStepReducible)
| _ -> raise NotStepReducible
in
- app_stack (redrec (c, empty_stack))
+ applist (redrec (c,[]))
-let isIndRef = function IndRef _ -> true | _ -> false
+let error_cannot_recognize ref =
+ errorlabstrm ""
+ (str "Cannot recognize a statement based on " ++
+ Nametab.pr_global_env Id.Set.empty ref ++ str".")
let reduce_to_ref_gen allow_product env sigma ref t =
if isIndRef ref then
- let (mind,t) = reduce_to_ind_gen allow_product env sigma t in
- if IndRef mind <> ref then
- errorlabstrm "" (str "Cannot recognize a statement based on " ++
- Nametab.pr_global_env Idset.empty ref ++ str".")
- else
- t
+ let ((mind,u),t) = reduce_to_ind_gen allow_product env sigma t in
+ begin match ref with
+ | IndRef mind' when eq_ind mind mind' -> t
+ | _ -> error_cannot_recognize ref
+ end
else
(* lazily reduces to match the head of [t] with the expected [ref] *)
let rec elimrec env t l =
- let c, _ = Reductionops.whd_stack sigma t in
+ let c, _ = decompose_appvect (Reductionops.whd_nored sigma t) in
match kind_of_term c with
| Prod (n,ty,t') ->
- if allow_product then
+ if allow_product then
elimrec (push_rel (n,None,t) env) t' ((n,None,ty)::l)
- else
- errorlabstrm ""
- (str "Cannot recognize an atomic statement based on " ++
- Nametab.pr_global_env Idset.empty ref ++ str".")
+ else
+ error_cannot_recognize ref
| _ ->
try
- if global_of_constr c = ref
+ if eq_gr (global_of_constr 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
elimrec env t' l
- with NotStepReducible ->
- errorlabstrm ""
- (str "Cannot recognize a statement based on " ++
- Nametab.pr_global_env Idset.empty ref ++ str".")
+ with NotStepReducible -> error_cannot_recognize ref
in
elimrec env t []
diff --git a/pretyping/tacred.mli b/pretyping/tacred.mli
index aec9da95..03c4cb41 100644
--- a/pretyping/tacred.mli
+++ b/pretyping/tacred.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -11,14 +11,13 @@ open Term
open Environ
open Evd
open Reductionops
-open Closure
-open Glob_term
-open Termops
open Pattern
-open Libnames
+open Globnames
+open Locus
+open Univ
type reduction_tactic_error =
- InvalidAbstraction of env * constr * (env * Type_errors.type_error)
+ InvalidAbstraction of env * evar_map * constr * (env * Type_errors.type_error)
exception ReductionTacticError of reduction_tactic_error
@@ -28,13 +27,13 @@ exception ReductionTacticError of reduction_tactic_error
val is_evaluable : Environ.env -> evaluable_global_reference -> bool
-val error_not_evaluable : Libnames.global_reference -> 'a
+val error_not_evaluable : Globnames.global_reference -> 'a
val evaluable_of_global_reference :
- Environ.env -> Libnames.global_reference -> evaluable_global_reference
+ Environ.env -> Globnames.global_reference -> evaluable_global_reference
val global_of_evaluable_reference :
- evaluable_global_reference -> Libnames.global_reference
+ evaluable_global_reference -> Globnames.global_reference
exception Redelimination
@@ -44,14 +43,6 @@ val red_product : reduction_function
(** Red (raise Redelimination if nothing reducible) *)
val try_red_product : reduction_function
-(** Tune the behaviour of simpl for the given constant name *)
-type simpl_flag = [ `SimplDontExposeCase | `SimplNeverUnfold ]
-
-val set_simpl_behaviour :
- bool -> global_reference -> (int list * int * simpl_flag list) -> unit
-val get_simpl_behaviour :
- global_reference -> (int list * int * simpl_flag list) option
-
(** Simpl *)
val simpl : reduction_function
@@ -70,7 +61,8 @@ val unfoldn :
val fold_commands : constr list -> reduction_function
(** Pattern *)
-val pattern_occs : (occurrences * constr) list -> reduction_function
+val pattern_occs : (occurrences * constr) list -> env -> evar_map -> constr ->
+ evar_map * constr
(** Rem: Lazy strategies are defined in Reduction *)
@@ -84,12 +76,12 @@ val cbv_norm_flags : Closure.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 -> inductive * types
+val reduce_to_atomic_ind : env -> evar_map -> types -> pinductive * types
(** [reduce_to_quantified_ind env sigma t] puts [t] in the form
[t'=(x1:A1)..(xn:An)(I args)] with [I] an inductive definition;
returns [I] and [t'] or fails with a user error *)
-val reduce_to_quantified_ind : env -> evar_map -> types -> inductive * types
+val reduce_to_quantified_ind : env -> evar_map -> types -> pinductive * 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 *)
@@ -100,7 +92,18 @@ val reduce_to_atomic_ref :
env -> evar_map -> global_reference -> types -> types
val find_hnf_rectype :
- env -> evar_map -> types -> inductive * constr list
+ env -> evar_map -> types -> pinductive * constr list
val contextually : bool -> occurrences * constr_pattern ->
(patvar_map -> reduction_function) -> reduction_function
+
+val e_contextually : bool -> occurrences * constr_pattern ->
+ (patvar_map -> e_reduction_function) -> e_reduction_function
+
+(** Returns the same inductive if it is allowed for pattern-matching
+ raises an error otherwise. **)
+val check_privacy : env -> inductive puniverses -> inductive puniverses
+
+(** Returns the same inductive if it is not a primitive record
+ raises an error otherwise. **)
+val check_not_primitive_record : env -> inductive puniverses -> inductive puniverses
diff --git a/pretyping/term_dnet.ml b/pretyping/term_dnet.ml
deleted file mode 100644
index 28a6c92c..00000000
--- a/pretyping/term_dnet.ml
+++ /dev/null
@@ -1,402 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(*i*)
-open Util
-open Term
-open Sign
-open Names
-open Libnames
-open Mod_subst
-open Pp (* debug *)
-(*i*)
-
-
-(* Representation/approximation of terms to use in the dnet:
- *
- * - no meta or evar (use ['a pattern] for that)
- *
- * - [Rel]s and [Sort]s are not taken into account (that's why we need
- * a second pass of linear filterin on the results - it's not a perfect
- * term indexing structure)
-
- * - Foralls and LetIns are represented by a context DCtx (a list of
- * generalization, similar to rel_context, and coded with DCons and
- * DNil). This allows for matching under an unfinished context
- *)
-
-module DTerm =
-struct
-
- type 't t =
- | DRel
- | DSort
- | DRef of global_reference
- | DCtx of 't * 't (* (binding list, subterm) = Prods and LetIns *)
- | DLambda of 't * 't
- | DApp of 't * 't (* binary app *)
- | DCase of case_info * 't * 't * 't array
- | DFix of int array * int * 't array * 't array
- | DCoFix of int * 't array * 't array
-
- (* special constructors only inside the left-hand side of DCtx or
- DApp. Used to encode lists of foralls/letins/apps as contexts *)
- | DCons of ('t * 't option) * 't
- | DNil
-
- type dconstr = dconstr t
-
- (* debug *)
- let rec pr_dconstr f : 'a t -> std_ppcmds = function
- | DRel -> str "*"
- | DSort -> str "Sort"
- | DRef _ -> str "Ref"
- | DCtx (ctx,t) -> f ctx ++ spc() ++ str "|-" ++ spc () ++ f t
- | DLambda (t1,t2) -> str "fun"++ spc() ++ f t1 ++ spc() ++ str"->" ++ spc() ++ f t2
- | DApp (t1,t2) -> f t1 ++ spc() ++ f t2
- | DCase (_,t1,t2,ta) -> str "case"
- | DFix _ -> str "fix"
- | DCoFix _ -> str "cofix"
- | DCons ((t,dopt),tl) -> f t ++ (match dopt with
- Some t' -> str ":=" ++ f t'
- | None -> str "") ++ spc() ++ str "::" ++ spc() ++ f tl
- | DNil -> str "[]"
-
- (*
- * Functional iterators for the t datatype
- * a.k.a boring and error-prone boilerplate code
- *)
-
- let map f = function
- | (DRel | DSort | DNil | DRef _) as c -> c
- | DCtx (ctx,c) -> DCtx (f ctx, f c)
- | DLambda (t,c) -> DLambda (f t, f c)
- | DApp (t,u) -> DApp (f t,f u)
- | DCase (ci,p,c,bl) -> DCase (ci, f p, f c, Array.map f bl)
- | DFix (ia,i,ta,ca) ->
- DFix (ia,i,Array.map f ta,Array.map f ca)
- | DCoFix(i,ta,ca) ->
- DCoFix (i,Array.map f ta,Array.map f ca)
- | DCons ((t,topt),u) -> DCons ((f t,Option.map f topt), f u)
-
- let compare x y =
- let make_name n =
- match n with
- | DRef(ConstRef con) ->
- DRef(ConstRef(constant_of_kn(canonical_con con)))
- | DRef(IndRef (kn,i)) ->
- DRef(IndRef(mind_of_kn(canonical_mind kn),i))
- | DRef(ConstructRef ((kn,i),j ))->
- DRef(ConstructRef((mind_of_kn(canonical_mind kn),i),j))
- | k -> k in
- Pervasives.compare (make_name x) (make_name y)
-
- let fold f acc = function
- | (DRel | DNil | DSort | DRef _) -> acc
- | DCtx (ctx,c) -> f (f acc ctx) c
- | DLambda (t,c) -> f (f acc t) c
- | DApp (t,u) -> f (f acc t) u
- | DCase (ci,p,c,bl) -> Array.fold_left f (f (f acc p) c) bl
- | DFix (ia,i,ta,ca) ->
- Array.fold_left f (Array.fold_left f acc ta) ca
- | DCoFix(i,ta,ca) ->
- Array.fold_left f (Array.fold_left f acc ta) ca
- | DCons ((t,topt),u) -> f (Option.fold_left f (f acc t) topt) u
-
- let choose f = function
- | (DRel | DSort | DNil | DRef _) -> invalid_arg "choose"
- | DCtx (ctx,c) -> f ctx
- | DLambda (t,c) -> f t
- | DApp (t,u) -> f u
- | DCase (ci,p,c,bl) -> f c
- | DFix (ia,i,ta,ca) -> f ta.(0)
- | DCoFix (i,ta,ca) -> f ta.(0)
- | DCons ((t,topt),u) -> f u
-
- let fold2 (f:'a -> 'b -> 'c -> 'a) (acc:'a) (c1:'b t) (c2:'c t) : 'a =
- let head w = map (fun _ -> ()) w in
- if compare (head c1) (head c2) <> 0
- then invalid_arg "fold2:compare" else
- match c1,c2 with
- | (DRel, DRel | DNil, DNil | DSort, DSort | DRef _, DRef _) -> acc
- | (DCtx (c1,t1), DCtx (c2,t2)
- | DApp (c1,t1), DApp (c2,t2)
- | DLambda (c1,t1), DLambda (c2,t2)) -> f (f acc c1 c2) t1 t2
- | DCase (ci,p1,c1,bl1),DCase (_,p2,c2,bl2) ->
- array_fold_left2 f (f (f acc p1 p2) c1 c2) bl1 bl2
- | DFix (ia,i,ta1,ca1), DFix (_,_,ta2,ca2) ->
- array_fold_left2 f (array_fold_left2 f acc ta1 ta2) ca1 ca2
- | DCoFix(i,ta1,ca1), DCoFix(_,ta2,ca2) ->
- array_fold_left2 f (array_fold_left2 f acc ta1 ta2) ca1 ca2
- | DCons ((t1,topt1),u1), DCons ((t2,topt2),u2) ->
- f (Option.fold_left2 f (f acc t1 t2) topt1 topt2) u1 u2
- | _ -> assert false
-
- let map2 (f:'a -> 'b -> 'c) (c1:'a t) (c2:'b t) : 'c t =
- let head w = map (fun _ -> ()) w in
- if compare (head c1) (head c2) <> 0
- then invalid_arg "map2_t:compare" else
- match c1,c2 with
- | (DRel, DRel | DSort, DSort | DNil, DNil | DRef _, DRef _) as cc ->
- let (c,_) = cc in c
- | DCtx (c1,t1), DCtx (c2,t2) -> DCtx (f c1 c2, f t1 t2)
- | DLambda (t1,c1), DLambda (t2,c2) -> DLambda (f t1 t2, f c1 c2)
- | DApp (t1,u1), DApp (t2,u2) -> DApp (f t1 t2,f u1 u2)
- | DCase (ci,p1,c1,bl1), DCase (_,p2,c2,bl2) ->
- DCase (ci, f p1 p2, f c1 c2, array_map2 f bl1 bl2)
- | DFix (ia,i,ta1,ca1), DFix (_,_,ta2,ca2) ->
- DFix (ia,i,array_map2 f ta1 ta2,array_map2 f ca1 ca2)
- | DCoFix (i,ta1,ca1), DCoFix (_,ta2,ca2) ->
- DCoFix (i,array_map2 f ta1 ta2,array_map2 f ca1 ca2)
- | DCons ((t1,topt1),u1), DCons ((t2,topt2),u2) ->
- DCons ((f t1 t2,Option.lift2 f topt1 topt2), f u1 u2)
- | _ -> assert false
-
- let terminal = function
- | (DRel | DSort | DNil | DRef _) -> true
- | _ -> false
-end
-
-(*
- * Terms discrimination nets
- * Uses the general dnet datatype on DTerm.t
- * (here you can restart reading)
- *)
-
-(*
- * Construction of the module
- *)
-
-module type IDENT =
-sig
- type t
- val compare : t -> t -> int
- val subst : substitution -> t -> t
- val constr_of : t -> constr
-end
-
-module type OPT =
-sig
- val reduce : constr -> constr
- val direction : bool
-end
-
-module Make =
- functor (Ident : IDENT) ->
- functor (Opt : OPT) ->
-struct
-
- module TDnet : Dnet.S with type ident=Ident.t
- and type 'a structure = 'a DTerm.t
- and type meta = metavariable
- = Dnet.Make(DTerm)(Ident)
- (struct
- type t = metavariable
- let compare = Pervasives.compare
- end)
-
- type t = TDnet.t
-
- type ident = TDnet.ident
-
- type 'a pattern = 'a TDnet.pattern
- type term_pattern = term_pattern DTerm.t pattern
-
- type idset = TDnet.Idset.t
-
- type result = ident * (constr*existential_key) * Termops.subst
-
- open DTerm
- open TDnet
-
- let rec pat_of_constr c : term_pattern =
- match kind_of_term c with
- | Rel _ -> Term DRel
- | Sort _ -> Term DSort
- | Var i -> Term (DRef (VarRef i))
- | Const c -> Term (DRef (ConstRef c))
- | Ind i -> Term (DRef (IndRef i))
- | Construct c -> Term (DRef (ConstructRef c))
- | Term.Meta _ -> assert false
- | Evar (i,_) -> Meta i
- | Case (ci,c1,c2,ca) ->
- Term(DCase(ci,pat_of_constr c1,pat_of_constr c2,Array.map pat_of_constr ca))
- | Fix ((ia,i),(_,ta,ca)) ->
- Term(DFix(ia,i,Array.map pat_of_constr ta, Array.map pat_of_constr ca))
- | CoFix (i,(_,ta,ca)) ->
- Term(DCoFix(i,Array.map pat_of_constr ta,Array.map pat_of_constr ca))
- | Cast (c,_,_) -> pat_of_constr c
- | Lambda (_,t,c) -> Term(DLambda (pat_of_constr t, pat_of_constr c))
- | (Prod (_,_,_) | LetIn(_,_,_,_)) ->
- let (ctx,c) = ctx_of_constr (Term DNil) c in Term (DCtx (ctx,c))
- | App (f,ca) ->
- Array.fold_left (fun c a -> Term (DApp (c,a)))
- (pat_of_constr f) (Array.map pat_of_constr ca)
-
- and ctx_of_constr ctx c : term_pattern * term_pattern =
- match kind_of_term c with
- | Prod (_,t,c) -> ctx_of_constr (Term(DCons((pat_of_constr t,None),ctx))) c
- | LetIn(_,d,t,c) -> ctx_of_constr (Term(DCons((pat_of_constr t, Some (pat_of_constr d)),ctx))) c
- | _ -> ctx,pat_of_constr c
-
- let empty_ctx : term_pattern -> term_pattern = function
- | Meta _ as c -> c
- | Term (DCtx(_,_)) as c -> c
- | c -> Term (DCtx (Term DNil, c))
-
- (*
- * Basic primitives
- *)
-
- let empty = TDnet.empty
-
- let subst s t =
- let sleaf id = Ident.subst s id in
- let snode = function
- | DTerm.DRef gr -> DTerm.DRef (fst (subst_global s gr))
- | n -> n in
- TDnet.map sleaf snode t
-
- let union = TDnet.union
-
- let add (c:constr) (id:Ident.t) (dn:t) =
- let c = Opt.reduce c in
- let c = empty_ctx (pat_of_constr c) in
- TDnet.add dn c id
-
- let new_meta_no =
- let ctr = ref 0 in
- fun () -> decr ctr; !ctr
-
- let new_meta_no = Evarutil.new_untyped_evar
-
- let neutral_meta = new_meta_no()
-
- let new_meta () = Meta (new_meta_no())
- let new_evar () = mkEvar(new_meta_no(),[||])
-
- let rec remove_cap : term_pattern -> term_pattern = function
- | Term (DCons (t,u)) -> Term (DCons (t,remove_cap u))
- | Term DNil -> new_meta()
- | Meta _ as m -> m
- | _ -> assert false
-
- let under_prod : term_pattern -> term_pattern = function
- | Term (DCtx (t,u)) -> Term (DCtx (remove_cap t,u))
- | Meta m -> Term (DCtx(new_meta(), Meta m))
- | _ -> assert false
-
- let init = let e = new_meta_no() in (mkEvar (e,[||]),e)
-
- let rec e_subst_evar i (t:unit->constr) c =
- match kind_of_term c with
- | Evar (j,_) when i=j -> t()
- | _ -> map_constr (e_subst_evar i t) c
-
- let subst_evar i c = e_subst_evar i (fun _ -> c)
-
- (* debug *)
- let rec pr_term_pattern p =
- (fun pr_t -> function
- | Term t -> pr_t t
- | Meta m -> str"["++Util.pr_int (Obj.magic m)++str"]"
- ) (pr_dconstr pr_term_pattern) p
-
- let search_pat cpat dpat dn (up,plug) =
- let whole_c = subst_evar plug cpat up in
- (* if we are at the root, add an empty context *)
- let dpat = if isEvar_or_Meta up then under_prod (empty_ctx dpat) else dpat in
- TDnet.Idset.fold
- (fun id acc ->
- let c_id = Opt.reduce (Ident.constr_of id) in
- let (ctx,wc) =
- try Termops.align_prod_letin whole_c c_id
- with Invalid_argument _ -> [],c_id in
- let up = it_mkProd_or_LetIn up ctx in
- let wc,whole_c = if Opt.direction then whole_c,wc else wc,whole_c in
- try (id,(up,plug),Termops.filtering ctx Reduction.CUMUL wc whole_c)::acc
- with Termops.CannotFilter -> (* msgnl(str"recon "++Termops.print_constr_env (Global.env()) wc); *) acc
- ) (TDnet.find_match dpat dn) []
-
- let fold_pattern_neutral f =
- fold_pattern (fun acc (mset,m,dn) -> if m=neutral_meta then acc else f m dn acc)
-
- let fold_pattern_nonlin f =
- let defined = ref Gmap.empty in
- fold_pattern_neutral
- ( fun m dn acc ->
- let dn = try TDnet.inter dn (Gmap.find m !defined) with Not_found -> dn in
- defined := Gmap.add m dn !defined;
- f m dn acc )
-
- let fold_pattern_up f acc dpat cpat dn (up,plug) =
- fold_pattern_nonlin
- ( fun m dn acc ->
- f dn (subst_evar plug (e_subst_evar neutral_meta new_evar cpat) up, m) acc
- ) acc dpat dn
-
- let possibly_under pat k dn (up,plug) =
- let rec aux fst dn (up,plug) acc =
- let cpat = pat() in
- let dpat = pat_of_constr cpat in
- let dpat = if fst then under_prod (empty_ctx dpat) else dpat in
- (k dn (up,plug)) @
- snd (fold_pattern_up (aux false) acc dpat cpat dn (up,plug)) in
- aux true dn (up,plug) []
-
- let eq_pat eq () = mkApp(eq,[|mkEvar(neutral_meta,[||]);new_evar();new_evar()|])
- let app_pat () = mkApp(new_evar(),[|mkEvar(neutral_meta,[||])|])
-
- (*
- * High-level primitives describing specific search problems
- *)
-
- let search_pattern dn pat =
- let pat = Opt.reduce pat in
- search_pat pat (empty_ctx (pat_of_constr pat)) dn init
-
- let search_concl dn pat =
- let pat = Opt.reduce pat in
- search_pat pat (under_prod (empty_ctx (pat_of_constr pat))) dn init
-
- let search_eq_concl dn eq pat =
- let pat = Opt.reduce pat in
- let eq_pat = eq_pat eq () in
- let eq_dpat = under_prod (empty_ctx (pat_of_constr eq_pat)) in
- snd (fold_pattern_up
- (fun dn up acc ->
- search_pat pat (pat_of_constr pat) dn up @ acc
- ) [] eq_dpat eq_pat dn init)
-
- let search_head_concl dn pat =
- let pat = Opt.reduce pat in
- possibly_under app_pat (search_pat pat (pat_of_constr pat)) dn init
-
- let find_all dn = Idset.elements (TDnet.find_all dn)
-
- let map f dn = TDnet.map f (fun x -> x) dn
-end
-
-module type S =
-sig
- type t
- type ident
-
- type result = ident * (constr*existential_key) * Termops.subst
-
- val empty : t
- val add : constr -> ident -> t -> t
- val union : t -> t -> t
- val subst : substitution -> t -> t
- val search_pattern : t -> constr -> result list
- val search_concl : t -> constr -> result list
- val search_head_concl : t -> constr -> result list
- val search_eq_concl : t -> constr -> constr -> result list
- val find_all : t -> ident list
- val map : (ident -> ident) -> t -> t
-end
diff --git a/pretyping/term_dnet.mli b/pretyping/term_dnet.mli
deleted file mode 100644
index 9e629dcd..00000000
--- a/pretyping/term_dnet.mli
+++ /dev/null
@@ -1,108 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-open Term
-open Sign
-open Libnames
-open Mod_subst
-
-(** Dnets on constr terms.
-
- An instantiation of Dnet on (an approximation of) constr. It
- associates a term (possibly with Evar) with an
- identifier. Identifiers must be unique (no two terms sharing the
- same ident), and there must be a way to recover the full term from
- the identifier (function constr_of).
-
- Optionally, a pre-treatment on terms can be performed before adding
- or searching (reduce). Practically, it is used to do some kind of
- delta-reduction on terms before indexing them.
-
- The results returned here are perfect, since post-filtering is done
- inside here.
-
- See lib/dnet.mli for more details.
-*)
-
-(** Identifiers to store (right hand side of the association) *)
-module type IDENT = sig
- type t
- val compare : t -> t -> int
-
- (** how to substitute them for storage *)
- val subst : substitution -> t -> t
-
- (** how to recover the term from the identifier *)
- val constr_of : t -> constr
-end
-
-(** Options : *)
-module type OPT = sig
-
- (** pre-treatment to terms before adding or searching *)
- val reduce : constr -> constr
-
- (** direction of post-filtering w.r.t sort subtyping :
- - true means query <= terms in the structure
- - false means terms <= query
- *)
- val direction : bool
-end
-
-module type S =
-sig
- type t
- type ident
-
- (** results of filtering : identifier, a context (term with Evar
- hole) and the substitution in that context*)
- type result = ident * (constr*existential_key) * Termops.subst
-
- val empty : t
-
- (** [add c i dn] adds the binding [(c,i)] to [dn]. [c] can be a
- closed term or a pattern (with untyped Evars). No Metas accepted *)
- val add : constr -> ident -> t -> t
-
- (** merge of dnets. Faster than re-adding all terms *)
- val union : t -> t -> t
-
- val subst : substitution -> t -> t
-
- (*
- * High-level primitives describing specific search problems
- *)
-
- (** [search_pattern dn c] returns all terms/patterns in dn
- matching/matched by c *)
- val search_pattern : t -> constr -> result list
-
- (** [search_concl dn c] returns all matches under products and
- letins, i.e. it finds subterms whose conclusion matches c. The
- complexity depends only on c ! *)
- val search_concl : t -> constr -> result list
-
- (** [search_head_concl dn c] matches under products and applications
- heads. Finds terms of the form [forall H_1...H_n, C t_1...t_n]
- where C matches c *)
- val search_head_concl : t -> constr -> result list
-
- (** [search_eq_concl dn eq c] searches terms of the form [forall
- H1...Hn, eq _ X1 X2] where either X1 or X2 matches c *)
- val search_eq_concl : t -> constr -> constr -> result list
-
- (** [find_all dn] returns all idents contained in dn *)
- val find_all : t -> ident list
-
- val map : (ident -> ident) -> t -> t
-end
-
-module Make :
- functor (Ident : IDENT) ->
- functor (Opt : OPT) ->
- S with type ident = Ident.t
diff --git a/pretyping/termops.ml b/pretyping/termops.ml
index 02661a93..5862a852 100644
--- a/pretyping/termops.ml
+++ b/pretyping/termops.ml
@@ -1,27 +1,28 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Pp
+open Errors
open Util
open Names
open Nameops
open Term
-open Sign
+open Vars
+open Context
open Environ
-open Libnames
-open Nametab
+open Locus
(* Sorts and sort family *)
let print_sort = function
| Prop Pos -> (str "Set")
| Prop Null -> (str "Prop")
- | Type u -> (str "Type(" ++ Univ.pr_uni u ++ str ")")
+ | Type u -> (str "Type(" ++ Univ.Universe.pr u ++ str ")")
let pr_sort_family = function
| InSet -> (str "Set")
@@ -34,6 +35,19 @@ let pr_name = function
let pr_con sp = str(string_of_con sp)
+let pr_fix pr_constr ((t,i),(lna,tl,bl)) =
+ let fixl = Array.mapi (fun i na -> (na,t.(i),tl.(i),bl.(i))) lna in
+ hov 1
+ (str"fix " ++ int i ++ spc() ++ str"{" ++
+ v 0 (prlist_with_sep spc (fun (na,i,ty,bd) ->
+ pr_name na ++ str"/" ++ int i ++ str":" ++ pr_constr ty ++
+ cut() ++ str":=" ++ pr_constr bd) (Array.to_list fixl)) ++
+ str"}")
+
+let pr_puniverses p u =
+ if Univ.Instance.is_empty u then p
+ else p ++ str"(*" ++ Univ.Instance.pr Universes.pr_with_global_universes u ++ str"*)"
+
let rec pr_constr c = match kind_of_term c with
| Rel n -> str "#"++int n
| Meta n -> str "Meta(" ++ int n ++ str ")"
@@ -59,25 +73,19 @@ let rec pr_constr c = match kind_of_term c with
(str"(" ++ pr_constr c ++ spc() ++
prlist_with_sep spc pr_constr (Array.to_list l) ++ str")")
| Evar (e,l) -> hov 1
- (str"Evar#" ++ int e ++ str"{" ++
+ (str"Evar#" ++ int (Evar.repr e) ++ str"{" ++
prlist_with_sep spc pr_constr (Array.to_list l) ++str"}")
- | Const c -> str"Cst(" ++ pr_con c ++ str")"
- | Ind (sp,i) -> str"Ind(" ++ pr_mind sp ++ str"," ++ int i ++ str")"
- | Construct ((sp,i),j) ->
- str"Constr(" ++ pr_mind sp ++ str"," ++ int i ++ str"," ++ int j ++ str")"
+ | Const (c,u) -> str"Cst(" ++ pr_puniverses (pr_con c) u ++ str")"
+ | Ind ((sp,i),u) -> str"Ind(" ++ pr_puniverses (pr_mind sp ++ str"," ++ int i) u ++ str")"
+ | Construct (((sp,i),j),u) ->
+ str"Constr(" ++ pr_puniverses (pr_mind sp ++ str"," ++ int i ++ str"," ++ int j) u ++ str")"
+ | Proj (p,c) -> str"Proj(" ++ pr_con (Projection.constant p) ++ str"," ++ bool (Projection.unfolded p) ++ pr_constr c ++ str")"
| Case (ci,p,c,bl) -> v 0
(hv 0 (str"<"++pr_constr p++str">"++ cut() ++ str"Case " ++
pr_constr c ++ str"of") ++ cut() ++
prlist_with_sep (fun _ -> brk(1,2)) pr_constr (Array.to_list bl) ++
cut() ++ str"end")
- | Fix ((t,i),(lna,tl,bl)) ->
- let fixl = Array.mapi (fun i na -> (na,t.(i),tl.(i),bl.(i))) lna in
- hov 1
- (str"fix " ++ int i ++ spc() ++ str"{" ++
- v 0 (prlist_with_sep spc (fun (na,i,ty,bd) ->
- pr_name na ++ str"/" ++ int i ++ str":" ++ pr_constr ty ++
- cut() ++ str":=" ++ pr_constr bd) (Array.to_list fixl)) ++
- str"}")
+ | Fix f -> pr_fix pr_constr f
| CoFix(i,(lna,tl,bl)) ->
let fixl = Array.mapi (fun i na -> (na,tl.(i),bl.(i))) lna in
hov 1
@@ -142,43 +150,6 @@ let print_env env =
in
(sign_env ++ db_env)
-(*let current_module = ref empty_dirpath
-
-let set_module m = current_module := m*)
-
-let new_univ_level =
- let univ_gen = ref 0 in
- (fun sp ->
- incr univ_gen;
- Univ.make_universe_level (Lib.library_dp(),!univ_gen))
-
-let new_univ () = Univ.make_universe (new_univ_level ())
-let new_Type () = mkType (new_univ ())
-let new_Type_sort () = Type (new_univ ())
-
-(* This refreshes universes in types; works only for inferred types (i.e. for
- types of the form (x1:A1)...(xn:An)B with B a sort or an atom in
- head normal form) *)
-let refresh_universes_gen strict t =
- let modified = ref false in
- let rec refresh t = match kind_of_term t with
- | Sort (Type u) when strict or u <> Univ.type0m_univ ->
- modified := true; new_Type ()
- | Prod (na,u,v) -> mkProd (na,u,refresh v)
- | _ -> t in
- let t' = refresh t in
- if !modified then t' else t
-
-let refresh_universes = refresh_universes_gen false
-let refresh_universes_strict = refresh_universes_gen true
-
-let new_sort_in_family = function
- | InProp -> prop_sort
- | InSet -> set_sort
- | InType -> Type (new_univ ())
-
-
-
(* [Rel (n+m);...;Rel(n+1)] *)
let rel_vect n m = Array.init m (fun i -> mkRel(n+m-i))
@@ -208,22 +179,23 @@ let push_rels_assum assums =
let push_named_rec_types (lna,typarray,_) env =
let ctxt =
- array_map2_i
+ Array.map2_i
(fun i na t ->
match na with
| Name id -> (id, None, lift i t)
- | Anonymous -> anomaly "Fix declarations must be named")
+ | Anonymous -> anomaly (Pp.str "Fix declarations must be named"))
lna typarray in
Array.fold_left
(fun e assum -> push_named assum e) env ctxt
-let rec lookup_rel_id id sign =
- let rec lookrec = function
- | (n, (Anonymous,_,_)::l) -> lookrec (n+1,l)
- | (n, (Name id',b,t)::l) -> if id' = id then (n,b,t) else lookrec (n+1,l)
- | (_, []) -> raise Not_found
+let lookup_rel_id id sign =
+ let rec lookrec n = function
+ | [] -> raise Not_found
+ | (Anonymous, _, _) :: l -> lookrec (n + 1) l
+ | (Name id', b, t) :: l ->
+ if Names.Id.equal id' id then (n, b, t) else lookrec (n + 1) l
in
- lookrec (1,sign)
+ lookrec 1 sign
(* Constructs either [forall x:t, c] or [let x:=b:t in c] *)
let mkProd_or_LetIn (na,body,t) c =
@@ -250,6 +222,13 @@ let it_mkNamedProd_or_LetIn init = it_named_context_quantifier mkNamedProd_or_Le
let it_mkNamedProd_wo_LetIn init = it_named_context_quantifier mkNamedProd_wo_LetIn ~init
let it_mkNamedLambda_or_LetIn init = it_named_context_quantifier mkNamedLambda_or_LetIn ~init
+let it_mkLambda_or_LetIn_from_no_LetIn c decls =
+ let rec aux k decls c = match decls with
+ | [] -> c
+ | (na,Some b,t)::decls -> mkLetIn (na,b,t,aux (k-1) decls (liftn 1 k c))
+ | (na,None,t)::decls -> mkLambda (na,t,aux (k-1) decls c)
+ in aux (List.length decls) (List.rev decls) c
+
(* *)
(* strips head casts and flattens head applications *)
@@ -258,7 +237,7 @@ let rec strip_head_cast c = match kind_of_term c with
let rec collapse_rec f cl2 = match kind_of_term f with
| App (g,cl1) -> collapse_rec g (Array.append cl1 cl2)
| Cast (c,_,_) -> collapse_rec c cl2
- | _ -> if Array.length cl2 = 0 then f else mkApp (f,cl2)
+ | _ -> if Int.equal (Array.length cl2) 0 then f else mkApp (f,cl2)
in
collapse_rec f cl
| Cast (c,_,_) -> strip_head_cast c
@@ -267,15 +246,15 @@ let rec strip_head_cast c = match kind_of_term c with
let rec drop_extra_implicit_args c = match kind_of_term c with
(* Removed trailing extra implicit arguments, what improves compatibility
for constants with recently added maximal implicit arguments *)
- | App (f,args) when isEvar (array_last args) ->
+ | App (f,args) when isEvar (Array.last args) ->
drop_extra_implicit_args
- (mkApp (f,fst (array_chop (Array.length args - 1) args)))
+ (mkApp (f,fst (Array.chop (Array.length args - 1) args)))
| _ -> c
(* Get the last arg of an application *)
let last_arg c = match kind_of_term c with
- | App (f,cl) -> array_last cl
- | _ -> anomaly "last_arg"
+ | App (f,cl) -> Array.last cl
+ | _ -> anomaly (Pp.str "last_arg")
(* Get the last arg of an application *)
let decompose_app_vect c =
@@ -285,22 +264,22 @@ let decompose_app_vect c =
let adjust_app_list_size f1 l1 f2 l2 =
let len1 = List.length l1 and len2 = List.length l2 in
- if len1 = len2 then (f1,l1,f2,l2)
+ if Int.equal len1 len2 then (f1,l1,f2,l2)
else if len1 < len2 then
- let extras,restl2 = list_chop (len2-len1) l2 in
+ let extras,restl2 = List.chop (len2-len1) l2 in
(f1, l1, applist (f2,extras), restl2)
else
- let extras,restl1 = list_chop (len1-len2) l1 in
+ let extras,restl1 = List.chop (len1-len2) l1 in
(applist (f1,extras), restl1, f2, l2)
let adjust_app_array_size f1 l1 f2 l2 =
let len1 = Array.length l1 and len2 = Array.length l2 in
- if len1 = len2 then (f1,l1,f2,l2)
+ if Int.equal len1 len2 then (f1,l1,f2,l2)
else if len1 < len2 then
- let extras,restl2 = array_chop (len2-len1) l2 in
+ let extras,restl2 = Array.chop (len2-len1) l2 in
(f1, l1, appvect (f2,extras), restl2)
else
- let extras,restl1 = array_chop (len1-len2) l1 in
+ let extras,restl1 = Array.chop (len1-len2) l1 in
(appvect (f1,extras), restl1, f2, l2)
(* [map_constr_with_named_binders g f l c] maps [f l] on the immediate
@@ -317,6 +296,7 @@ let map_constr_with_named_binders g f l c = match kind_of_term c with
| Lambda (na,t,c) -> mkLambda (na, f l t, f (g na l) c)
| LetIn (na,b,t,c) -> mkLetIn (na, f l b, f l t, f (g na l) c)
| App (c,al) -> mkApp (f l c, Array.map (f l) al)
+ | Proj (p,c) -> mkProj (p, f l c)
| Evar (e,al) -> mkEvar (e, Array.map (f l) al)
| Case (ci,p,c,bl) -> mkCase (ci, f l p, f l c, Array.map (f l) bl)
| Fix (ln,(lna,tl,bl)) ->
@@ -337,45 +317,81 @@ let map_constr_with_named_binders g f l c = match kind_of_term c with
(co-)fixpoint) *)
let fold_rec_types g (lna,typarray,_) e =
- let ctxt = array_map2_i (fun i na t -> (na, None, lift i t)) lna typarray in
+ let ctxt = Array.map2_i (fun i na t -> (na, None, lift i t)) lna typarray in
Array.fold_left (fun e assum -> g assum e) e ctxt
+let map_left2 f a g b =
+ let l = Array.length a in
+ if Int.equal l 0 then [||], [||] else begin
+ let r = Array.make l (f a.(0)) in
+ let s = Array.make l (g b.(0)) in
+ for i = 1 to l - 1 do
+ r.(i) <- f a.(i);
+ s.(i) <- g b.(i)
+ done;
+ r, s
+ end
let map_constr_with_binders_left_to_right g f l c = match kind_of_term c with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> c
- | Cast (c,k,t) -> let c' = f l c in mkCast (c',k,f l t)
- | Prod (na,t,c) ->
+ | Cast (b,k,t) ->
+ let b' = f l b in
+ let t' = f l t in
+ if b' == b && t' == t then c
+ else mkCast (b',k,t')
+ | Prod (na,t,b) ->
let t' = f l t in
- mkProd (na, t', f (g (na,None,t) l) c)
- | Lambda (na,t,c) ->
+ let b' = f (g (na,None,t) l) b in
+ if t' == t && b' == b then c
+ else mkProd (na, t', b')
+ | Lambda (na,t,b) ->
let t' = f l t in
- mkLambda (na, t', f (g (na,None,t) l) c)
- | LetIn (na,b,t,c) ->
- let b' = f l b in
+ let b' = f (g (na,None,t) l) b in
+ if t' == t && b' == b then c
+ else mkLambda (na, t', b')
+ | LetIn (na,bo,t,b) ->
+ let bo' = f l bo in
let t' = f l t in
- let c' = f (g (na,Some b,t) l) c in
- mkLetIn (na, b', t', c')
+ let b' = f (g (na,Some bo,t) l) b in
+ if bo' == bo && t' == t && b' == b then c
+ else mkLetIn (na, bo', t', b')
| App (c,[||]) -> assert false
- | App (c,al) ->
+ | App (t,al) ->
(*Special treatment to be able to recognize partially applied subterms*)
let a = al.(Array.length al - 1) in
- let hd = f l (mkApp (c, Array.sub al 0 (Array.length al - 1))) in
- mkApp (hd, [| f l a |])
- | Evar (e,al) -> mkEvar (e, array_map_left (f l) al)
- | Case (ci,p,c,bl) ->
+ let app = (mkApp (t, Array.sub al 0 (Array.length al - 1))) in
+ let app' = f l app in
+ let a' = f l a in
+ if app' == app && a' == a then c
+ else mkApp (app', [| a' |])
+ | Proj (p,b) ->
+ let b' = f l b in
+ if b' == b then c
+ else mkProj (p, b')
+ | Evar (e,al) ->
+ let al' = Array.map_left (f l) al in
+ if Array.for_all2 (==) al' al then c
+ else mkEvar (e, al')
+ | Case (ci,p,b,bl) ->
(* In v8 concrete syntax, predicate is after the term to match! *)
- let c' = f l c in
+ let b' = f l b in
let p' = f l p in
- mkCase (ci, p', c', array_map_left (f l) bl)
+ let bl' = Array.map_left (f l) bl in
+ if b' == b && p' == p && bl' == bl then c
+ else mkCase (ci, p', b', bl')
| Fix (ln,(lna,tl,bl as fx)) ->
let l' = fold_rec_types g fx l in
- let (tl',bl') = array_map_left_pair (f l) tl (f l') bl in
- mkFix (ln,(lna,tl',bl'))
+ let (tl', bl') = map_left2 (f l) tl (f l') bl in
+ if Array.for_all2 (==) tl tl' && Array.for_all2 (==) bl bl'
+ then c
+ else mkFix (ln,(lna,tl',bl'))
| CoFix(ln,(lna,tl,bl as fx)) ->
let l' = fold_rec_types g fx l in
- let (tl',bl') = array_map_left_pair (f l) tl (f l') bl in
- mkCoFix (ln,(lna,tl',bl'))
+ let (tl', bl') = map_left2 (f l) tl (f l') bl in
+ if Array.for_all2 (==) tl tl' && Array.for_all2 (==) bl bl'
+ then c
+ else mkCoFix (ln,(lna,tl',bl'))
(* strong *)
let map_constr_with_full_binders g f l cstr = match kind_of_term cstr with
@@ -401,30 +417,33 @@ let map_constr_with_full_binders g f l cstr = match kind_of_term cstr with
| App (c,al) ->
let c' = f l c in
let al' = Array.map (f l) al in
- if c==c' && array_for_all2 (==) al al' then cstr else mkApp (c', al')
+ if c==c' && Array.for_all2 (==) al al' then cstr else mkApp (c', al')
+ | Proj (p,c) ->
+ let c' = f l c in
+ if c' == c then cstr else mkProj (p, c')
| Evar (e,al) ->
let al' = Array.map (f l) al in
- if array_for_all2 (==) al al' then cstr else mkEvar (e, al')
+ if Array.for_all2 (==) al al' then cstr else mkEvar (e, al')
| Case (ci,p,c,bl) ->
let p' = f l p in
let c' = f l c in
let bl' = Array.map (f l) bl in
- if p==p' && c==c' && array_for_all2 (==) bl bl' then cstr else
+ if p==p' && c==c' && Array.for_all2 (==) bl bl' then cstr else
mkCase (ci, p', c', bl')
| Fix (ln,(lna,tl,bl)) ->
let tl' = Array.map (f l) tl in
let l' =
- array_fold_left2 (fun l na t -> g (na,None,t) l) l lna tl in
+ Array.fold_left2 (fun l na t -> g (na,None,t) l) l lna tl in
let bl' = Array.map (f l') bl in
- if array_for_all2 (==) tl tl' && array_for_all2 (==) bl bl'
+ if Array.for_all2 (==) tl tl' && Array.for_all2 (==) bl bl'
then cstr
else mkFix (ln,(lna,tl',bl'))
| CoFix(ln,(lna,tl,bl)) ->
let tl' = Array.map (f l) tl in
let l' =
- array_fold_left2 (fun l na t -> g (na,None,t) l) l lna tl in
+ Array.fold_left2 (fun l na t -> g (na,None,t) l) l lna tl in
let bl' = Array.map (f l') bl in
- if array_for_all2 (==) tl tl' && array_for_all2 (==) bl bl'
+ if Array.for_all2 (==) tl tl' && Array.for_all2 (==) bl bl'
then cstr
else mkCoFix (ln,(lna,tl',bl'))
@@ -443,15 +462,16 @@ let fold_constr_with_binders g f n acc c = match kind_of_term c with
| Lambda (_,t,c) -> f (g n) (f n acc t) c
| LetIn (_,b,t,c) -> f (g n) (f n (f n acc b) t) c
| App (c,l) -> Array.fold_left (f n) (f n acc c) l
+ | Proj (p,c) -> f n acc c
| Evar (_,l) -> Array.fold_left (f n) acc l
| Case (_,p,c,bl) -> Array.fold_left (f n) (f n (f n acc p) c) bl
| Fix (_,(lna,tl,bl)) ->
let n' = iterate g (Array.length tl) n in
- let fd = array_map2 (fun t b -> (t,b)) tl bl in
+ let fd = Array.map2 (fun t b -> (t,b)) tl bl in
Array.fold_left (fun acc (t,b) -> f n' (f n acc t) b) acc fd
| CoFix (_,(lna,tl,bl)) ->
let n' = iterate g (Array.length tl) n in
- let fd = array_map2 (fun t b -> (t,b)) tl bl in
+ let fd = Array.map2 (fun t b -> (t,b)) tl bl in
Array.fold_left (fun acc (t,b) -> f n' (f n acc t) b) acc fd
(* [iter_constr_with_full_binders g f acc c] iters [f acc] on the immediate
@@ -467,14 +487,15 @@ let iter_constr_with_full_binders g f l c = match kind_of_term c with
| Lambda (na,t,c) -> f l t; f (g (na,None,t) l) c
| LetIn (na,b,t,c) -> f l b; f l t; f (g (na,Some b,t) l) c
| App (c,args) -> f l c; Array.iter (f l) args
+ | Proj (p,c) -> f l c
| Evar (_,args) -> Array.iter (f l) args
| Case (_,p,c,bl) -> f l p; f l c; Array.iter (f l) bl
| Fix (_,(lna,tl,bl)) ->
- let l' = array_fold_left2 (fun l na t -> g (na,None,t) l) l lna tl in
+ let l' = Array.fold_left2 (fun l na t -> g (na,None,t) l) l lna tl in
Array.iter (f l) tl;
Array.iter (f l') bl
| CoFix (_,(lna,tl,bl)) ->
- let l' = array_fold_left2 (fun l na t -> g (na,None,t) l) l lna tl in
+ let l' = Array.fold_left2 (fun l na t -> g (na,None,t) l) l lna tl in
Array.iter (f l) tl;
Array.iter (f l') bl
@@ -503,23 +524,16 @@ let occur_meta_or_existential c =
| _ -> iter_constr occrec c
in try occrec c; false with Occur -> true
-let occur_const s c =
- let rec occur_rec c = match kind_of_term c with
- | Const sp when sp=s -> raise Occur
- | _ -> iter_constr occur_rec c
- in
- try occur_rec c; false with Occur -> true
-
let occur_evar n c =
let rec occur_rec c = match kind_of_term c with
- | Evar (sp,_) when sp=n -> raise Occur
+ | Evar (sp,_) when Evar.equal sp n -> raise Occur
| _ -> iter_constr occur_rec c
in
try occur_rec c; false with Occur -> true
let occur_in_global env id constr =
let vars = vars_of_global env constr in
- if List.mem id vars then raise Occur
+ if Id.Set.mem id vars then raise Occur
let occur_var env id c =
let rec occur_rec c =
@@ -540,10 +554,10 @@ let occur_var_in_decl env hyp (_,c,typ) =
let free_rels m =
let rec frec depth acc c = match kind_of_term c with
- | Rel n -> if n >= depth then Intset.add (n-depth+1) acc else acc
+ | Rel n -> if n >= depth then Int.Set.add (n-depth+1) acc else acc
| _ -> fold_constr_with_binders succ frec depth acc c
in
- frec 1 Intset.empty m
+ frec 1 Int.Set.empty m
(* collects all metavar occurences, in left-to-right order, preserving
* repetitions and all. *)
@@ -551,7 +565,7 @@ let free_rels m =
let collect_metas c =
let rec collrec acc c =
match kind_of_term c with
- | Meta mv -> list_add_set mv acc
+ | Meta mv -> List.add_set Int.equal mv acc
| _ -> fold_constr collrec acc c
in
List.rev (collrec [] c)
@@ -560,32 +574,41 @@ let collect_metas c =
all section variables; for the latter, use global_vars_set *)
let collect_vars c =
let rec aux vars c = match kind_of_term c with
- | Var id -> Idset.add id vars
+ | Var id -> Id.Set.add id vars
| _ -> fold_constr aux vars c in
- aux Idset.empty c
+ aux Id.Set.empty c
(* Tests whether [m] is a subterm of [t]:
[m] is appropriately lifted through abstractions of [t] *)
-let dependent_main noevar m t =
+let dependent_main noevar univs m t =
+ let eqc x y = if univs then fst (Universes.eq_constr_universes x y) else eq_constr_nounivs x y in
let rec deprec m t =
- if eq_constr m t then
+ if eqc m t then
raise Occur
else
match kind_of_term m, kind_of_term t with
| App (fm,lm), App (ft,lt) when Array.length lm < Array.length lt ->
deprec m (mkApp (ft,Array.sub lt 0 (Array.length lm)));
- Array.iter (deprec m)
+ CArray.Fun1.iter deprec m
(Array.sub lt
(Array.length lm) ((Array.length lt) - (Array.length lm)))
- | _, Cast (c,_,_) when noevar & isMeta c -> ()
+ | _, Cast (c,_,_) when noevar && isMeta c -> ()
| _, Evar _ when noevar -> ()
- | _ -> iter_constr_with_binders (lift 1) deprec m t
+ | _ -> iter_constr_with_binders (fun c -> lift 1 c) deprec m t
in
try deprec m t; false with Occur -> true
-let dependent = dependent_main false
-let dependent_no_evar = dependent_main true
+let dependent = dependent_main false false
+let dependent_no_evar = dependent_main true false
+
+let dependent_univs = dependent_main false true
+let dependent_univs_no_evar = dependent_main true true
+
+let dependent_in_decl a (_,c,t) =
+ match c with
+ | None -> dependent a t
+ | Some body -> dependent a body || dependent a t
let count_occurrences m t =
let n = ref 0 in
@@ -621,7 +644,7 @@ type meta_value_map = (metavariable * constr) list
let rec subst_meta bl c =
match kind_of_term c with
- | Meta i -> (try List.assoc i bl with Not_found -> c)
+ | Meta i -> (try Int.List.assoc i bl with Not_found -> c)
| _ -> map_constr (subst_meta bl) c
(* First utilities for avoiding telescope computation for subst_term *)
@@ -686,184 +709,42 @@ let replace_term_gen eq_fun c by_c in_t =
let replace_term = replace_term_gen eq_constr
-(* Substitute only at a list of locations or excluding a list of
- locations; in the occurrences list (b,l), b=true means no
- occurrence except the ones in l and b=false, means all occurrences
- except the ones in l *)
-
-type hyp_location_flag = (* To distinguish body and type of local defs *)
- | InHyp
- | InHypTypeOnly
- | InHypValueOnly
-
-type occurrences = bool * int list
-let all_occurrences = (false,[])
-let no_occurrences_in_set = (true,[])
-
-let error_invalid_occurrence l =
- let l = list_uniquize (List.sort Pervasives.compare l) in
- errorlabstrm ""
- (str ("Invalid occurrence " ^ plural (List.length l) "number" ^": ") ++
- prlist_with_sep spc int l ++ str ".")
-
-let pr_position (cl,pos) =
- let clpos = match cl with
- | None -> str " of the goal"
- | Some (id,InHyp) -> str " of hypothesis " ++ pr_id id
- | Some (id,InHypTypeOnly) -> str " of the type of hypothesis " ++ pr_id id
- | Some (id,InHypValueOnly) -> str " of the body of hypothesis " ++ pr_id id in
- int pos ++ clpos
-
-let error_cannot_unify_occurrences nested (cl2,pos2,t2) (cl1,pos1,t1) (nowhere_except_in,locs) =
- let s = if nested then "Found nested occurrences of the pattern"
- else "Found incompatible occurrences of the pattern" in
- errorlabstrm ""
- (str s ++ str ":" ++
- spc () ++ str "Matched term " ++ quote (print_constr t2) ++
- strbrk " at position " ++ pr_position (cl2,pos2) ++
- strbrk " is not compatible with matched term " ++
- quote (print_constr t1) ++ strbrk " at position " ++
- pr_position (cl1,pos1) ++ str ".")
-
-let is_selected pos (nowhere_except_in,locs) =
- nowhere_except_in && List.mem pos locs ||
- not nowhere_except_in && not (List.mem pos locs)
-
-exception NotUnifiable
-
-type 'a testing_function = {
- match_fun : constr -> 'a;
- merge_fun : 'a -> 'a -> 'a;
- mutable testing_state : 'a;
- mutable last_found : ((identifier * hyp_location_flag) option * int * constr) option
-}
-
-let subst_closed_term_occ_gen_modulo (nowhere_except_in,locs as plocs) test cl occ t =
- let maxocc = List.fold_right max locs 0 in
- let pos = ref occ in
- let nested = ref false in
- let add_subst t subst =
- try
- test.testing_state <- test.merge_fun subst test.testing_state;
- test.last_found <- Some (cl,!pos,t)
- with NotUnifiable ->
- let lastpos = Option.get test.last_found in
- error_cannot_unify_occurrences !nested (cl,!pos,t) lastpos plocs in
- let rec substrec k t =
- if nowhere_except_in & !pos > maxocc then t else
- try
- let subst = test.match_fun t in
- if is_selected !pos plocs then
- (add_subst t subst; incr pos;
- (* Check nested matching subterms *)
- nested := true; ignore (subst_below k t); nested := false;
- (* Do the effective substitution *)
- mkRel k)
- else
- (incr pos; subst_below k 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
- in
- let t' = substrec 1 t in
- (!pos, t')
-
-let is_nowhere (nowhere_except_in,locs) = nowhere_except_in && locs = []
-
-let check_used_occurrences nbocc (nowhere_except_in,locs) =
- let rest = List.filter (fun o -> o >= nbocc) locs in
- if rest <> [] then error_invalid_occurrence rest
-
-let proceed_with_occurrences f plocs x =
- if is_nowhere plocs then (* optimization *) x else
- begin
- assert (List.for_all (fun x -> x >= 0) (snd plocs));
- let (nbocc,x) = f 1 x in
- check_used_occurrences nbocc plocs;
- x
- end
-
-let make_eq_test c = {
- match_fun = (fun c' -> if eq_constr c c' then () else raise NotUnifiable);
- merge_fun = (fun () () -> ());
- testing_state = ();
- last_found = None
-}
-
-let subst_closed_term_occ_gen plocs pos c t =
- subst_closed_term_occ_gen_modulo plocs (make_eq_test c) None pos t
-
-let subst_closed_term_occ plocs c t =
- proceed_with_occurrences (fun occ -> subst_closed_term_occ_gen plocs occ c)
- plocs t
-
-let subst_closed_term_occ_modulo plocs test cl t =
- proceed_with_occurrences
- (subst_closed_term_occ_gen_modulo plocs test cl) plocs t
-
-let map_named_declaration_with_hyploc f hyploc acc (id,bodyopt,typ) =
- let f = f (Some (id,hyploc)) in
- match bodyopt,hyploc with
- | None, InHypValueOnly ->
- errorlabstrm "" (pr_id id ++ str " has no value.")
- | None, _ | Some _, InHypTypeOnly ->
- let acc,typ = f acc typ in acc,(id,bodyopt,typ)
- | Some body, InHypValueOnly ->
- let acc,body = f acc body in acc,(id,Some body,typ)
- | Some body, InHyp ->
- let acc,body = f acc body in
- let acc,typ = f acc typ in
- acc,(id,Some body,typ)
-
-let subst_closed_term_occ_decl (plocs,hyploc) c d =
- proceed_with_occurrences
- (map_named_declaration_with_hyploc
- (fun _ occ -> subst_closed_term_occ_gen plocs occ c) hyploc) plocs d
-
-let subst_closed_term_occ_decl_modulo (plocs,hyploc) test d =
- proceed_with_occurrences
- (map_named_declaration_with_hyploc
- (subst_closed_term_occ_gen_modulo plocs test)
- hyploc)
- plocs d
-
let vars_of_env env =
let s =
- Sign.fold_named_context (fun (id,_,_) s -> Idset.add id s)
- (named_context env) ~init:Idset.empty in
- Sign.fold_rel_context
- (fun (na,_,_) s -> match na with Name id -> Idset.add id s | _ -> s)
+ Context.fold_named_context (fun (id,_,_) s -> Id.Set.add id s)
+ (named_context env) ~init:Id.Set.empty in
+ Context.fold_rel_context
+ (fun (na,_,_) s -> match na with Name id -> Id.Set.add id s | _ -> s)
(rel_context env) ~init:s
let add_vname vars = function
- Name id -> Idset.add id vars
+ Name id -> Id.Set.add id vars
| _ -> vars
(*************************)
(* Names environments *)
(*************************)
-type names_context = name list
+type names_context = Name.t list
let add_name n nl = n::nl
let lookup_name_of_rel p names =
try List.nth names (p-1)
with Invalid_argument _ | Failure _ -> raise Not_found
-let rec lookup_rel_of_name id names =
+let lookup_rel_of_name id names =
let rec lookrec n = function
| Anonymous :: l -> lookrec (n+1) l
- | (Name id') :: l -> if id' = id then n else lookrec (n+1) l
+ | (Name id') :: l -> if Id.equal id' id then n else lookrec (n+1) l
| [] -> raise Not_found
in
lookrec 1 names
let empty_names_context = []
let ids_of_rel_context sign =
- Sign.fold_rel_context
+ Context.fold_rel_context
(fun (na,_,_) l -> match na with Name id -> id::l | Anonymous -> l)
sign ~init:[]
let ids_of_named_context sign =
- Sign.fold_named_context (fun (id,_,_) idl -> id::idl) sign ~init:[]
+ Context.fold_named_context (fun (id,_,_) idl -> id::idl) sign ~init:[]
let ids_of_context env =
(ids_of_rel_context (rel_context env))
@@ -882,15 +763,16 @@ let isGlobalRef c =
| Const _ | Ind _ | Construct _ | Var _ -> true
| _ -> false
-let has_polymorphic_type c =
- match (Global.lookup_constant c).Declarations.const_type with
- | Declarations.PolymorphicArity _ -> true
+let is_template_polymorphic env f =
+ match kind_of_term f with
+ | Ind ind -> Environ.template_polymorphic_pind ind env
+ | Const c -> Environ.template_polymorphic_pconstant c env
| _ -> false
let base_sort_cmp pb s0 s1 =
match (s0,s1) with
- | (Prop c1, Prop c2) -> c1 = Null or c2 = Pos (* Prop <= Set *)
- | (Prop c1, Type u) -> pb = Reduction.CUMUL
+ | (Prop c1, Prop c2) -> c1 == Null || c2 == Pos (* Prop <= Set *)
+ | (Prop c1, Type u) -> pb == Reduction.CUMUL
| (Type u1, Type u2) -> true
| _ -> false
@@ -899,45 +781,48 @@ let compare_constr_univ f cv_pb t1 t2 =
match kind_of_term t1, kind_of_term t2 with
Sort s1, Sort s2 -> base_sort_cmp cv_pb s1 s2
| Prod (_,t1,c1), Prod (_,t2,c2) ->
- f Reduction.CONV t1 t2 & f cv_pb c1 c2
- | _ -> compare_constr (f Reduction.CONV) t1 t2
+ f Reduction.CONV t1 t2 && f cv_pb c1 c2
+ | _ -> compare_constr (fun t1 t2 -> f Reduction.CONV t1 t2) t1 t2
let rec constr_cmp cv_pb t1 t2 = compare_constr_univ constr_cmp cv_pb t1 t2
-let eq_constr = constr_cmp Reduction.CONV
+let eq_constr t1 t2 = constr_cmp Reduction.CONV t1 t2
(* App(c,[t1,...tn]) -> ([c,t1,...,tn-1],tn)
App(c,[||]) -> ([],c) *)
let split_app c = match kind_of_term c with
App(c,l) ->
let len = Array.length l in
- if len=0 then ([],c) else
+ if Int.equal len 0 then ([],c) else
let last = Array.get l (len-1) in
let prev = Array.sub l 0 (len-1) in
c::(Array.to_list prev), last
| _ -> assert false
-let hdtl l = List.hd l, List.tl l
-
-type subst = (rel_context*constr) Intmap.t
+type subst = (rel_context*constr) Evar.Map.t
exception CannotFilter
let filtering env cv_pb c1 c2 =
- let evm = ref Intmap.empty in
+ let evm = ref Evar.Map.empty in
let define cv_pb e1 ev c1 =
- try let (e2,c2) = Intmap.find ev !evm in
+ try let (e2,c2) = Evar.Map.find ev !evm in
let shift = List.length e1 - List.length e2 in
if constr_cmp cv_pb c1 (lift shift c2) then () else raise CannotFilter
with Not_found ->
- evm := Intmap.add ev (e1,c1) !evm
+ evm := Evar.Map.add ev (e1,c1) !evm
in
let rec aux env cv_pb c1 c2 =
match kind_of_term c1, kind_of_term c2 with
| App _, App _ ->
- let ((p1,l1),(p2,l2)) = (split_app c1),(split_app c2) in
- aux env cv_pb l1 l2; if p1=[] & p2=[] then () else
- aux env cv_pb (applist (hdtl p1)) (applist (hdtl p2))
+ let ((p1,l1),(p2,l2)) = (split_app c1),(split_app c2) in
+ let () = aux env cv_pb l1 l2 in
+ begin match p1, p2 with
+ | [], [] -> ()
+ | (h1 :: p1), (h2 :: p2) ->
+ aux env cv_pb (applistc h1 p1) (applistc h2 p2)
+ | _ -> assert false
+ end
| Prod (n,t1,c1), Prod (_,t2,c2) ->
aux env cv_pb t1 t2;
aux ((n,None,t1)::env) cv_pb c1 c2
@@ -963,12 +848,12 @@ let align_prod_letin c a : rel_context * constr =
let (lc,_,_) = decompose_prod_letin c in
let (la,l,a) = decompose_prod_letin a in
if not (la >= lc) then invalid_arg "align_prod_letin";
- let (l1,l2) = Util.list_chop lc l in
+ let (l1,l2) = Util.List.chop lc l in
l2,it_mkProd_or_LetIn a l1
-(* On reduit une serie d'eta-redex de tete ou rien du tout *)
+(* We reduce a series of head eta-redex or nothing at all *)
(* [x1:c1;...;xn:cn]@(f;a1...an;x1;...;xn) --> @(f;a1...an) *)
-(* Remplace 2 versions précédentes buggées *)
+(* Remplace 2 earlier buggish versions *)
let rec eta_reduce_head c =
match kind_of_term c with
@@ -976,12 +861,12 @@ let rec eta_reduce_head c =
(match kind_of_term (eta_reduce_head c') with
| App (f,cl) ->
let lastn = (Array.length cl) - 1 in
- if lastn < 1 then anomaly "application without arguments"
+ if lastn < 0 then anomaly (Pp.str "application without arguments")
else
(match kind_of_term cl.(lastn) with
| Rel 1 ->
let c' =
- if lastn = 1 then f
+ if Int.equal lastn 0 then f
else mkApp (f, Array.sub cl 0 lastn)
in
if noccurn 1 c'
@@ -992,24 +877,15 @@ let rec eta_reduce_head c =
| _ -> c
-(* alpha-eta conversion : ignore print names and casts *)
-let eta_eq_constr =
- let rec aux t1 t2 =
- let t1 = eta_reduce_head (strip_head_cast t1)
- and t2 = eta_reduce_head (strip_head_cast t2) in
- t1=t2 or compare_constr aux t1 t2
- in aux
-
-
(* iterator on rel context *)
let process_rel_context f env =
let sign = named_context_val env in
let rels = rel_context env in
let env0 = reset_with_named_context sign env in
- Sign.fold_rel_context f rels ~init:env0
+ Context.fold_rel_context f rels ~init:env0
let assums_of_rel_context sign =
- Sign.fold_rel_context
+ Context.fold_rel_context
(fun (na,c,t) l ->
match c with
Some _ -> l
@@ -1054,37 +930,49 @@ let adjust_subst_to_rel_context sign l =
| (_,Some c,_)::sign', args' ->
aux (substl (List.rev subst) c :: subst) sign' args'
| [], [] -> List.rev subst
- | _ -> anomaly "Instance and signature do not match"
+ | _ -> anomaly (Pp.str "Instance and signature do not match")
in aux [] (List.rev sign) l
-let fold_named_context_both_sides f l ~init = list_fold_right_and_left f l init
+let fold_named_context_both_sides f l ~init = List.fold_right_and_left f l init
let rec mem_named_context id = function
- | (id',_,_) :: _ when id=id' -> true
+ | (id',_,_) :: _ when Id.equal id id' -> true
| _ :: sign -> mem_named_context id sign
| [] -> false
+let compact_named_context_reverse sign =
+ let compact l (i1,c1,t1) =
+ match l with
+ | [] -> [[i1],c1,t1]
+ | (l2,c2,t2)::q ->
+ if Option.equal Constr.equal c1 c2 && Constr.equal t1 t2
+ then (i1::l2,c2,t2)::q
+ else ([i1],c1,t1)::l
+ in Context.fold_named_context_reverse compact ~init:[] sign
+
+let compact_named_context sign = List.rev (compact_named_context_reverse sign)
+
let clear_named_body id env =
- let rec aux _ = function
- | (id',Some c,t) when id = id' -> push_named (id,None,t)
+ let aux _ = function
+ | (id',Some c,t) when Id.equal id id' -> push_named (id,None,t)
| d -> push_named d in
fold_named_context aux env ~init:(reset_context env)
-let global_vars env ids = Idset.elements (global_vars_set env ids)
+let global_vars env ids = Id.Set.elements (global_vars_set env ids)
let global_vars_set_of_decl env = function
| (_,None,t) -> global_vars_set env t
| (_,Some c,t) ->
- Idset.union (global_vars_set env t)
+ Id.Set.union (global_vars_set env t)
(global_vars_set env c)
let dependency_closure env sign hyps =
- if Idset.is_empty hyps then [] else
+ if Id.Set.is_empty hyps then [] else
let (_,lh) =
- Sign.fold_named_context_reverse
+ Context.fold_named_context_reverse
(fun (hs,hl) (x,_,_ as d) ->
- if Idset.mem x hs then
- (Idset.union (global_vars_set_of_decl env d) (Idset.remove x hs),
+ if Id.Set.mem x hs then
+ (Id.Set.union (global_vars_set_of_decl env d) (Id.Set.remove x hs),
x::hl)
else (hs,hl))
~init:(hyps,[])
@@ -1104,13 +992,13 @@ let context_chop k ctx =
| (0, l2) -> (List.rev acc, l2)
| (n, ((_,Some _,_ as h)::t)) -> chop_aux (h::acc) (n, t)
| (n, (h::t)) -> chop_aux (h::acc) (pred n, t)
- | (_, []) -> anomaly "context_chop"
+ | (_, []) -> anomaly (Pp.str "context_chop")
in chop_aux [] (k,ctx)
(* Do not skip let-in's *)
let env_rel_context_chop k env =
let rels = rel_context env in
- let ctx1,ctx2 = list_chop k rels in
+ let ctx1,ctx2 = List.chop k rels in
push_rel_context ctx2 (reset_with_named_context (named_context_val env) env),
ctx1
@@ -1122,13 +1010,15 @@ let impossible_default_case = ref None
let set_impossible_default_clause c = impossible_default_case := Some c
let coq_unit_judge =
- let na1 = Name (id_of_string "A") in
- let na2 = Name (id_of_string "H") 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) ->
- make_judge id type_of_id
+ | Some fn ->
+ let (id,type_of_id), ctx = fn () in
+ make_judge id type_of_id, ctx
| None ->
(* In case the constants id/ID are not defined *)
make_judge (mkLambda (na1,mkProp,mkLambda(na2,mkRel 1,mkRel 1)))
- (mkProd (na1,mkProp,mkArrow (mkRel 1) (mkRel 2)))
+ (mkProd (na1,mkProp,mkArrow (mkRel 1) (mkRel 2))),
+ Univ.ContextSet.empty
diff --git a/pretyping/termops.mli b/pretyping/termops.mli
index 81b23d7e..9f3efd72 100644
--- a/pretyping/termops.mli
+++ b/pretyping/termops.mli
@@ -1,30 +1,22 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Util
open Pp
open Names
open Term
-open Sign
+open Context
open Environ
-
-(** Universes *)
-val new_univ_level : unit -> Univ.universe_level
-val new_univ : unit -> Univ.universe
-val new_sort_in_family : sorts_family -> sorts
-val new_Type : unit -> types
-val new_Type_sort : unit -> sorts
-val refresh_universes : types -> types
-val refresh_universes_strict : types -> types
+open Locus
(** printers *)
val print_sort : sorts -> std_ppcmds
val pr_sort_family : sorts_family -> std_ppcmds
+val pr_fix : (constr -> std_ppcmds) -> fixpoint -> std_ppcmds
(** debug printer: do not use to display terms to the casual user... *)
val set_print_constr : (env -> constr -> std_ppcmds) -> unit
@@ -36,12 +28,19 @@ val print_rel_context : env -> std_ppcmds
val print_env : env -> std_ppcmds
(** about contexts *)
-val push_rel_assum : name * types -> env -> env
-val push_rels_assum : (name * types) list -> env -> env
-val push_named_rec_types : name array * types array * 'a -> env -> env
-val lookup_rel_id : identifier -> rel_context -> int * constr option * types
-
-(** builds argument lists matching a block of binders or a context *)
+val push_rel_assum : Name.t * types -> env -> env
+val push_rels_assum : (Name.t * types) list -> env -> env
+val push_named_rec_types : Name.t array * types array * 'a -> env -> env
+
+val lookup_rel_id : Id.t -> rel_context -> int * constr option * types
+(** Associates the contents of an identifier in a [rel_context]. Raise
+ [Not_found] if there is no such identifier. *)
+
+(** Functions that build argument lists matching a block of binders or a context.
+ [rel_vect n m] builds [|Rel (n+m);...;Rel(n+1)|]
+ [extended_rel_vect n ctx] extends the [ctx] context of length [m]
+ with [n] elements.
+*)
val rel_vect : int -> int -> constr array
val rel_list : int -> int -> constr list
val extended_rel_list : int -> rel_context -> constr list
@@ -50,8 +49,8 @@ val extended_rel_vect : int -> rel_context -> constr array
(** iterators/destructors on terms *)
val mkProd_or_LetIn : rel_declaration -> types -> types
val mkProd_wo_LetIn : rel_declaration -> types -> types
-val it_mkProd : types -> (name * types) list -> types
-val it_mkLambda : constr -> (name * types) list -> constr
+val it_mkProd : types -> (Name.t * types) list -> types
+val it_mkLambda : constr -> (Name.t * types) list -> constr
val it_mkProd_or_LetIn : types -> rel_context -> types
val it_mkProd_wo_LetIn : types -> rel_context -> types
val it_mkLambda_or_LetIn : constr -> rel_context -> constr
@@ -59,13 +58,14 @@ val it_mkNamedProd_or_LetIn : types -> named_context -> types
val it_mkNamedProd_wo_LetIn : types -> named_context -> types
val it_mkNamedLambda_or_LetIn : constr -> named_context -> constr
-val it_named_context_quantifier :
- (named_declaration -> 'a -> 'a) -> init:'a -> named_context -> 'a
+(* Ad hoc version reinserting letin, assuming the body is defined in
+ the context where the letins are expanded *)
+val it_mkLambda_or_LetIn_from_no_LetIn : constr -> rel_context -> constr
(** {6 Generic iterators on constr} *)
val map_constr_with_named_binders :
- (name -> 'a -> 'a) ->
+ (Name.t -> 'a -> 'a) ->
('a -> constr -> constr) -> 'a -> constr -> constr
val map_constr_with_binders_left_to_right :
(rel_declaration -> 'a -> 'a) ->
@@ -99,18 +99,22 @@ exception Occur
val occur_meta : types -> bool
val occur_existential : types -> bool
val occur_meta_or_existential : types -> bool
-val occur_const : constant -> types -> bool
val occur_evar : existential_key -> types -> bool
-val occur_var : env -> identifier -> types -> bool
+val occur_var : env -> Id.t -> types -> bool
val occur_var_in_decl :
env ->
- identifier -> 'a * types option * types -> bool
-val free_rels : constr -> Intset.t
+ Id.t -> 'a * types option * types -> bool
+val free_rels : constr -> Int.Set.t
+
+(** [dependent m t] tests whether [m] is a subterm of [t] *)
val dependent : constr -> constr -> bool
val dependent_no_evar : constr -> constr -> bool
+val dependent_univs : constr -> constr -> bool
+val dependent_univs_no_evar : constr -> constr -> bool
+val dependent_in_decl : constr -> named_declaration -> bool
val count_occurrences : constr -> constr -> int
val collect_metas : constr -> int list
-val collect_vars : constr -> Idset.t (** for visible vars only *)
+val collect_vars : constr -> Id.Set.t (** for visible vars only *)
val occur_term : constr -> constr -> bool (** Synonymous
of dependent
Substitution of metavariables *)
@@ -144,68 +148,14 @@ val subst_term : constr -> constr -> constr
(** [replace_term d e c] replaces [d] by [e] in [c] *)
val replace_term : constr -> constr -> constr -> constr
-(** In occurrences sets, false = everywhere except and true = nowhere except *)
-type occurrences = bool * int list
-val all_occurrences : occurrences
-val no_occurrences_in_set : occurrences
-
-(** [subst_closed_term_occ_gen occl n c d] replaces occurrences of closed [c] at
- positions [occl], counting from [n], by [Rel 1] in [d] *)
-val subst_closed_term_occ_gen :
- occurrences -> int -> constr -> types -> int * types
-
-(** [subst_closed_term_occ_modulo] looks for subterm modulo a
- testing function returning a substitution of type ['a] (or failing
- with NotUnifiable); a function for merging substitution (possibly
- failing with NotUnifiable) and an initial substitution are
- required too *)
-
-type hyp_location_flag = (** To distinguish body and type of local defs *)
- | InHyp
- | InHypTypeOnly
- | InHypValueOnly
-
-type 'a testing_function = {
- match_fun : constr -> 'a;
- merge_fun : 'a -> 'a -> 'a;
- mutable testing_state : 'a;
- mutable last_found : ((identifier * hyp_location_flag) option * int * constr) option
-}
-
-val make_eq_test : constr -> unit testing_function
-
-exception NotUnifiable
-
-val subst_closed_term_occ_modulo :
- occurrences -> 'a testing_function -> (identifier * hyp_location_flag) option
- -> constr -> types
-
-(** [subst_closed_term_occ occl c d] replaces occurrences of closed [c] at
- positions [occl] by [Rel 1] in [d] (see also Note OCC) *)
-val subst_closed_term_occ : occurrences -> constr -> constr -> constr
-
-(** [subst_closed_term_occ_decl occl c decl] replaces occurrences of closed [c]
- at positions [occl] by [Rel 1] in [decl] *)
-
-val subst_closed_term_occ_decl :
- occurrences * hyp_location_flag -> constr -> named_declaration ->
- named_declaration
-
-val subst_closed_term_occ_decl_modulo :
- occurrences * hyp_location_flag -> 'a testing_function ->
- named_declaration -> named_declaration
-
-val error_invalid_occurrence : int list -> 'a
-
(** Alternative term equalities *)
val base_sort_cmp : Reduction.conv_pb -> sorts -> sorts -> bool
val compare_constr_univ : (Reduction.conv_pb -> constr -> constr -> bool) ->
Reduction.conv_pb -> constr -> constr -> bool
val constr_cmp : Reduction.conv_pb -> constr -> constr -> bool
-val eq_constr : constr -> constr -> bool
+val eq_constr : constr -> constr -> bool (* FIXME rename: erases universes*)
val eta_reduce_head : constr -> constr
-val eta_eq_constr : constr -> constr -> bool
exception CannotFilter
@@ -215,7 +165,7 @@ exception CannotFilter
(context,term), or raises [CannotFilter].
Warning: Outer-kernel sort subtyping are taken into account: c1 has
to be smaller than c2 wrt. sorts. *)
-type subst = (rel_context*constr) Intmap.t
+type subst = (rel_context*constr) Evar.Map.t
val filtering : rel_context -> Reduction.conv_pb -> constr -> constr -> subst
val decompose_prod_letin : constr -> int * rel_context * constr
@@ -233,26 +183,26 @@ val adjust_app_array_size : constr -> constr array -> constr -> constr array ->
(constr * constr array * constr * constr array)
(** name contexts *)
-type names_context = name list
-val add_name : name -> names_context -> names_context
-val lookup_name_of_rel : int -> names_context -> name
-val lookup_rel_of_name : identifier -> names_context -> int
+type names_context = Name.t list
+val add_name : Name.t -> names_context -> names_context
+val lookup_name_of_rel : int -> names_context -> Name.t
+val lookup_rel_of_name : Id.t -> names_context -> int
val empty_names_context : names_context
-val ids_of_rel_context : rel_context -> identifier list
-val ids_of_named_context : named_context -> identifier list
-val ids_of_context : env -> identifier list
+val ids_of_rel_context : rel_context -> Id.t list
+val ids_of_named_context : named_context -> Id.t list
+val ids_of_context : env -> Id.t list
val names_of_rel_context : env -> names_context
val context_chop : int -> rel_context -> rel_context * rel_context
val env_rel_context_chop : int -> env -> env * rel_context
(** Set of local names *)
-val vars_of_env: env -> Idset.t
-val add_vname : Idset.t -> name -> Idset.t
+val vars_of_env: env -> Id.Set.t
+val add_vname : Id.Set.t -> Name.t -> Id.Set.t
(** other signature iterators *)
val process_rel_context : (rel_declaration -> env -> env) -> env -> env
-val assums_of_rel_context : rel_context -> (name * constr) list
+val assums_of_rel_context : rel_context -> (Name.t * constr) list
val lift_rel_context : int -> rel_context -> rel_context
val substl_rel_context : constr list -> rel_context -> rel_context
val smash_rel_context : rel_context -> rel_context (** expand lets in context *)
@@ -264,23 +214,25 @@ val map_rel_context_with_binders :
val fold_named_context_both_sides :
('a -> named_declaration -> named_declaration list -> 'a) ->
named_context -> init:'a -> 'a
-val mem_named_context : identifier -> named_context -> bool
+val mem_named_context : Id.t -> named_context -> bool
+val compact_named_context : named_context -> named_list_context
+val compact_named_context_reverse : named_context -> named_list_context
-val clear_named_body : identifier -> env -> env
+val clear_named_body : Id.t -> env -> env
-val global_vars : env -> constr -> identifier list
-val global_vars_set_of_decl : env -> named_declaration -> Idset.t
+val global_vars : env -> constr -> Id.t list
+val global_vars_set_of_decl : env -> named_declaration -> Id.Set.t
(** Gives an ordered list of hypotheses, closed by dependencies,
containing a given set *)
-val dependency_closure : env -> named_context -> Idset.t -> identifier list
+val dependency_closure : env -> named_context -> Id.Set.t -> Id.t list
(** Test if an identifier is the basename of a global reference *)
-val is_section_variable : identifier -> bool
+val is_section_variable : Id.t -> bool
val isGlobalRef : constr -> bool
-val has_polymorphic_type : constant -> bool
+val is_template_polymorphic : env -> constr -> bool
(** Combinators on judgments *)
@@ -289,5 +241,5 @@ val on_judgment_value : (types -> types) -> unsafe_judgment -> unsafe_judgment
val on_judgment_type : (types -> types) -> unsafe_judgment -> unsafe_judgment
(** {6 Functions to deal with impossible cases } *)
-val set_impossible_default_clause : constr * types -> unit
-val coq_unit_judge : unit -> unsafe_judgment
+val set_impossible_default_clause : (unit -> (constr * types) Univ.in_universe_context_set) -> unit
+val coq_unit_judge : unit -> unsafe_judgment Univ.in_universe_context_set
diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml
index 2b5b7fe2..817d6878 100644
--- a/pretyping/typeclasses.ml
+++ b/pretyping/typeclasses.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,45 +8,50 @@
(*i*)
open Names
-open Libnames
+open Globnames
open Decl_kinds
open Term
-open Sign
+open Vars
+open Context
open Evd
open Environ
-open Nametab
-open Mod_subst
open Util
open Typeclasses_errors
open Libobject
(*i*)
+let typeclasses_unique_solutions = ref false
+let set_typeclasses_unique_solutions d = (:=) typeclasses_unique_solutions d
+let get_typeclasses_unique_solutions () = !typeclasses_unique_solutions
-let add_instance_hint_ref = ref (fun id local pri -> assert false)
-let register_add_instance_hint =
- (:=) add_instance_hint_ref
-let add_instance_hint id = !add_instance_hint_ref id
+open Goptions
-let remove_instance_hint_ref = ref (fun id -> assert false)
-let register_remove_instance_hint =
- (:=) remove_instance_hint_ref
-let remove_instance_hint id = !remove_instance_hint_ref id
+let set_typeclasses_unique_solutions =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "check that typeclasses proof search returns unique solutions";
+ optkey = ["Typeclasses";"Unique";"Solutions"];
+ optread = get_typeclasses_unique_solutions;
+ optwrite = set_typeclasses_unique_solutions; }
-let set_typeclass_transparency_ref = ref (fun id local c -> assert false)
-let register_set_typeclass_transparency =
- (:=) set_typeclass_transparency_ref
-let set_typeclass_transparency gr local c = !set_typeclass_transparency_ref gr local c
+let (add_instance_hint, add_instance_hint_hook) = Hook.make ()
+let add_instance_hint id = Hook.get add_instance_hint id
-let classes_transparent_state_ref = ref (fun () -> assert false)
-let register_classes_transparent_state = (:=) classes_transparent_state_ref
-let classes_transparent_state () = !classes_transparent_state_ref ()
+let (remove_instance_hint, remove_instance_hint_hook) = Hook.make ()
+let remove_instance_hint id = Hook.get remove_instance_hint id
-let solve_instanciation_problem = ref (fun _ _ _ -> assert false)
+let (set_typeclass_transparency, set_typeclass_transparency_hook) = Hook.make ()
+let set_typeclass_transparency gr local c = Hook.get set_typeclass_transparency gr local c
-let resolve_one_typeclass env evm t =
- !solve_instanciation_problem env evm t
+let (classes_transparent_state, classes_transparent_state_hook) = Hook.make ()
+let classes_transparent_state () = Hook.get classes_transparent_state ()
+
+let solve_instantiation_problem = ref (fun _ _ _ _ -> assert false)
+
+let resolve_one_typeclass ?(unique=get_typeclasses_unique_solutions ()) env evm t =
+ !solve_instantiation_problem env evm t unique
-type rels = constr list
type direction = Forward | Backward
(* This module defines type-classes *)
@@ -61,12 +66,14 @@ type typeclass = {
cl_props : rel_context;
(* The method implementaions as projections. *)
- cl_projs : (name * (direction * int option) option * constant option) list;
-}
+ cl_projs : (Name.t * (direction * int option) option * constant option) list;
+
+ cl_strict : bool;
-module Gmap = Fmap.Make(RefOrdered)
+ cl_unique : bool;
+}
-type typeclasses = typeclass Gmap.t
+type typeclasses = typeclass Refmap.t
type instance = {
is_class: global_reference;
@@ -75,14 +82,17 @@ type instance = {
-1 for discard, 0 for none, mutable to avoid redeclarations
when multiple rebuild_object happen. *)
is_global: int;
+ is_poly: bool;
is_impl: global_reference;
}
-type instances = (instance Gmap.t) Gmap.t
+type instances = (instance Refmap.t) Refmap.t
let instance_impl is = is.is_impl
-let new_instance cl pri glob impl =
+let instance_priority is = is.is_pri
+
+let new_instance cl pri glob poly impl =
let global =
if glob then Lib.sections_depth ()
else -1
@@ -90,38 +100,45 @@ let new_instance cl pri glob impl =
{ is_class = cl.cl_impl;
is_pri = pri ;
is_global = global ;
+ is_poly = poly;
is_impl = impl }
(*
* states management
*)
-let classes : typeclasses ref = ref Gmap.empty
-
-let instances : instances ref = ref Gmap.empty
-
-let freeze () = !classes, !instances
-
-let unfreeze (cl,is) =
- classes:=cl;
- instances:=is
+let classes : typeclasses ref = Summary.ref Refmap.empty ~name:"classes"
+let instances : instances ref = Summary.ref Refmap.empty ~name:"instances"
-let init () =
- classes:= Gmap.empty;
- instances:= Gmap.empty
+open Declarations
-let _ =
- Summary.declare_summary "classes_and_instances"
- { Summary.freeze_function = freeze;
- Summary.unfreeze_function = unfreeze;
- Summary.init_function = init }
+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.map_rel_context (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 class_info c =
- try Gmap.find c !classes
- with Not_found -> not_a_class (Global.env()) (constr_of_global c)
+ try Refmap.find c !classes
+ with Not_found -> not_a_class (Global.env()) (printable_constr_of_global c)
let global_class_of_constr env c =
- try class_info (global_of_constr c)
+ try let gr, u = Universes.global_of_constr c in
+ class_info gr, u
with Not_found -> not_a_class env c
let dest_class_app env c =
@@ -129,19 +146,26 @@ let dest_class_app env c =
global_class_of_constr env cl, args
let dest_class_arity env c =
- let rels, c = Term.decompose_prod_assum c in
+ let rels, c = decompose_prod_assum c in
rels, dest_class_app env c
let class_of_constr c =
try Some (dest_class_arity (Global.env ()) c)
with e when Errors.noncritical e -> None
-let rec is_class_type evd c =
- match kind_of_term c with
- | Prod (_, _, t) -> is_class_type evd t
- | Evar (e, _) when is_defined evd e -> is_class_type evd (Evarutil.nf_evar evd c)
- | _ -> class_of_constr c <> None
+let is_class_constr c =
+ try let gr, u = Universes.global_of_constr 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
+ | 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
+
let is_class_evar evd evi =
is_class_type evd evi.Evd.evar_concl
@@ -150,25 +174,28 @@ let is_class_evar evd evi =
*)
let load_class (_, cl) =
- classes := Gmap.add cl.cl_impl cl !classes
+ classes := Refmap.add cl.cl_impl cl !classes
let cache_class = load_class
let subst_class (subst,cl) =
- let do_subst_con c = fst (Mod_subst.subst_con subst c)
+ 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 ctx = list_smartmap
+ let do_subst_ctx ctx = List.smartmap
(fun (na, b, t) -> (na, Option.smartmap do_subst b, do_subst t))
ctx in
let do_subst_context (grs,ctx) =
- list_smartmap (Option.smartmap (fun (gr,b) -> do_subst_gr gr, b)) grs,
+ List.smartmap (Option.smartmap (fun (gr,b) -> do_subst_gr gr, b)) 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
+ 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_context = do_subst_context cl.cl_context;
cl_props = do_subst_ctx cl.cl_props;
- cl_projs = do_subst_projs cl.cl_projs; }
+ cl_projs = do_subst_projs cl.cl_projs;
+ cl_strict = cl.cl_strict;
+ cl_unique = cl.cl_unique }
let discharge_class (_,cl) =
let repl = Lib.replacement_context () in
@@ -196,22 +223,26 @@ let discharge_class (_,cl) =
let newgrs = List.map (fun (_, _, t) ->
match class_of_constr t with
| None -> None
- | Some (_, (tc, _)) -> Some (tc.cl_impl, true))
+ | Some (_, ((tc,_), _)) -> Some (tc.cl_impl, true))
ctx'
in
- list_smartmap (Option.smartmap (fun (gr, b) -> Lib.discharge_global gr, b)) grs
+ List.smartmap (Option.smartmap (fun (gr, b) -> Lib.discharge_global gr, b)) 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 = abs_context cl in
+ let ctx, usubst, uctx = abs_context cl 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
- { cl_impl = cl_impl';
- cl_context = context;
- cl_props = props;
- cl_projs = list_smartmap (fun (x, y, z) -> x, y, Option.smartmap Lib.discharge_con z) cl.cl_projs }
+ let discharge_proj (x, y, z) = x, y, Option.smartmap Lib.discharge_con z in
+ { cl_impl = cl_impl';
+ cl_context = context;
+ cl_props = props;
+ cl_projs = List.smartmap discharge_proj cl.cl_projs;
+ cl_strict = cl.cl_strict;
+ cl_unique = cl.cl_unique
+ }
let rebuild_class cl =
try
@@ -239,25 +270,35 @@ let check_instance env sigma c =
try
let (evd, c) = resolve_one_typeclass env sigma
(Retyping.get_type_of env sigma c) in
- Evd.is_empty (Evd.undefined_evars evd)
+ not (Evd.has_undefined evd)
with e when Errors.noncritical e -> false
let build_subclasses ~check env sigma glob pri =
- let rec aux pri c =
- let ty = Evarutil.nf_evar sigma (Retyping.get_type_of env sigma c) in
+ let _id = Nametab.basename_of_global glob in
+ let _next_id =
+ let i = ref (-1) in
+ (fun () -> incr i;
+ 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 rec aux pri c ty path =
+ let ty = Evarutil.nf_evar sigma ty in
match class_of_constr ty with
| None -> []
- | Some (rels, (tc, args)) ->
- let instapp = Reductionops.whd_beta sigma (appvectc c (Termops.extended_rel_vect 0 rels)) in
+ | Some (rels, ((tc,u), args)) ->
+ let instapp =
+ Reductionops.whd_beta sigma (appvectc c (Termops.extended_rel_vect 0 rels))
+ in
let projargs = Array.of_list (args @ [instapp]) in
- let projs = list_map_filter
+ let projs = List.map_filter
(fun (n, b, proj) ->
match b with
| None -> None
| Some (Backward, _) -> None
| Some (Forward, pri') ->
let proj = Option.get proj in
- let body = it_mkLambda_or_LetIn (mkApp (mkConst proj, projargs)) rels in
+ let body = it_mkLambda_or_LetIn (mkApp (mkConstU (proj,u), projargs)) rels in
if check && check_instance env sigma body then None
else
let pri =
@@ -269,10 +310,16 @@ let build_subclasses ~check env sigma glob pri =
Some (ConstRef proj, pri, body)) tc.cl_projs
in
let declare_proj hints (cref, pri, body) =
- let rest = aux pri body in
- hints @ (pri, body) :: rest
+ let path' = cref :: path in
+ let ty = Retyping.get_type_of env sigma body in
+ let rest = aux pri body ty path' in
+ hints @ (path', pri, body) :: rest
in List.fold_left declare_proj [] projs
- in aux pri (constr_of_global glob)
+ in
+ let term = Universes.constr_of_global_univ (glob,Univ.UContext.instance ctx) in
+ (*FIXME subclasses should now get substituted for each particular instance of
+ the polymorphic superclass *)
+ aux pri term ty [glob]
(*
* instances persistent object
@@ -284,17 +331,17 @@ type instance_action =
let load_instance inst =
let insts =
- try Gmap.find inst.is_class !instances
- with Not_found -> Gmap.empty in
- let insts = Gmap.add inst.is_impl inst insts in
- instances := Gmap.add inst.is_class insts !instances
+ try Refmap.find inst.is_class !instances
+ with Not_found -> Refmap.empty in
+ let insts = Refmap.add inst.is_impl inst insts in
+ instances := Refmap.add inst.is_class insts !instances
let remove_instance inst =
let insts =
- try Gmap.find inst.is_class !instances
+ try Refmap.find inst.is_class !instances
with Not_found -> assert false in
- let insts = Gmap.remove inst.is_impl insts in
- instances := Gmap.add inst.is_class insts !instances
+ let insts = Refmap.remove inst.is_impl insts in
+ instances := Refmap.add inst.is_class insts !instances
let cache_instance (_, (action, i)) =
match action with
@@ -315,27 +362,28 @@ let discharge_instance (_, (action, inst)) =
is_impl = Lib.discharge_global inst.is_impl })
-let is_local i = i.is_global = -1
+let is_local i = Int.equal i.is_global (-1)
let add_instance check inst =
- add_instance_hint (constr_of_global inst.is_impl) (is_local inst) inst.is_pri;
- List.iter (fun (pri, c) -> add_instance_hint c (is_local inst) pri)
+ let poly = Global.is_polymorphic inst.is_impl in
+ add_instance_hint (IsGlobal inst.is_impl) [inst.is_impl] (is_local inst)
+ inst.is_pri poly;
+ List.iter (fun (path, pri, c) -> add_instance_hint (IsConstr c) path
+ (is_local inst) pri poly)
(build_subclasses ~check:(check && not (isVarRef inst.is_impl))
(Global.env ()) Evd.empty inst.is_impl inst.is_pri)
let rebuild_instance (action, inst) =
- if action = AddInstance then add_instance true inst;
+ let () = match action with
+ | AddInstance -> add_instance true inst
+ | _ -> ()
+ in
(action, inst)
let classify_instance (action, inst) =
if is_local inst then Dispose
else Substitute (action, inst)
-let load_instance (_, (action, inst) as ai) =
- cache_instance ai;
- if action = AddInstance then
- add_instance_hint (constr_of_global inst.is_impl) (is_local inst) inst.is_pri
-
let instance_input : instance_action * instance -> obj =
declare_object
{ (default_object "type classes instances state") with
@@ -356,11 +404,10 @@ let remove_instance i =
remove_instance_hint i.is_impl
let declare_instance pri local glob =
- let c = constr_of_global glob in
- let ty = Retyping.get_type_of (Global.env ()) Evd.empty c in
+ let ty = Global.type_of_global_unsafe glob in
match class_of_constr ty with
- | Some (rels, (tc, args) as _cl) ->
- add_instance (new_instance tc pri (not local) glob)
+ | Some (rels, ((tc,_), args) as _cl) ->
+ add_instance (new_instance tc pri (not local) (Flags.use_polymorphic_flag ()) glob)
(* let path, hints = build_subclasses (not local) (Global.env ()) Evd.empty glob in *)
(* let entries = List.map (fun (path, pri, c) -> (pri, local, path, c)) hints in *)
(* Auto.add_hints local [typeclasses_db] (Auto.HintsResolveEntry entries); *)
@@ -373,71 +420,57 @@ let add_class cl =
List.iter (fun (n, inst, body) ->
match inst with
| Some (Backward, pri) ->
- declare_instance pri false (ConstRef (Option.get body))
+ (match body with
+ | None -> Errors.error "Non-definable projection can not be declared as a subinstance"
+ | Some b -> declare_instance pri false (ConstRef b))
| _ -> ())
cl.cl_projs
open Declarations
-
-let add_constant_class cst =
- let ty = Typeops.type_of_constant (Global.env ()) cst in
- let ctx, arity = decompose_prod_assum ty in
- let tc =
- { cl_impl = ConstRef cst;
- cl_context = (List.map (const None) ctx, ctx);
- cl_props = [(Anonymous, None, arity)];
- cl_projs = []
- }
- in add_class tc;
- set_typeclass_transparency (EvalConstRef cst) false false
-
-let add_inductive_class ind =
- let mind, oneind = Global.lookup_inductive ind in
- let k =
- let ctx = oneind.mind_arity_ctxt in
- let ty = Inductive.type_of_inductive_knowing_parameters
- (push_rel_context ctx (Global.env ()))
- oneind (Termops.extended_rel_vect 0 ctx)
- in
- { cl_impl = IndRef ind;
- cl_context = List.map (const None) ctx, ctx;
- cl_props = [Anonymous, None, ty];
- cl_projs = [] }
- in add_class k
(*
* interface functions
*)
-let instance_constructor cl args =
- let lenpars = List.length (List.filter (fun (na, b, t) -> b = None) (snd cl.cl_context)) in
- let pars = fst (list_chop lenpars args) in
+let instance_constructor (cl,u) args =
+ let filter (_, b, _) = match b with
+ | None -> true
+ | Some _ -> false
+ in
+ let lenpars = List.length (List.filter filter (snd cl.cl_context)) in
+ let pars = fst (List.chop lenpars args) in
match cl.cl_impl with
- | IndRef ind -> Some (applistc (mkConstruct (ind, 1)) args),
- applistc (mkInd ind) pars
+ | IndRef ind ->
+ let ind = ind, u in
+ (Some (applistc (mkConstructUi (ind, 1)) args),
+ applistc (mkIndU ind) pars)
| ConstRef cst ->
- let term = if args = [] then None else Some (list_last args) in
- term, applistc (mkConst cst) pars
+ let cst = cst, u in
+ let term = match args with
+ | [] -> None
+ | _ -> Some (List.last args)
+ in
+ (term, applistc (mkConstU cst) pars)
| _ -> assert false
-let typeclasses () = Gmap.fold (fun _ l c -> l :: c) !classes []
+let typeclasses () = Refmap.fold (fun _ l c -> l :: c) !classes []
-let cmap_elements c = Gmap.fold (fun k v acc -> v :: acc) c []
+let cmap_elements c = Refmap.fold (fun k v acc -> v :: acc) c []
let instances_of c =
- try cmap_elements (Gmap.find c.cl_impl !instances) with Not_found -> []
+ try cmap_elements (Refmap.find c.cl_impl !instances) with Not_found -> []
let all_instances () =
- Gmap.fold (fun k v acc ->
- Gmap.fold (fun k v acc -> v :: acc) v acc)
+ Refmap.fold (fun k v acc ->
+ Refmap.fold (fun k v acc -> v :: acc) v acc)
!instances []
let instances r =
let cl = class_info r in instances_of cl
-
+
let is_class gr =
- Gmap.fold (fun k v acc -> acc || v.cl_impl = gr) !classes false
+ Refmap.exists (fun _ v -> eq_gr v.cl_impl gr) !classes
let is_instance = function
| ConstRef c ->
@@ -452,6 +485,14 @@ let is_instance = function
is_class (IndRef ind)
| _ -> false
+let is_implicit_arg = function
+| Evar_kinds.GoalEvar -> false
+| _ -> true
+ (* match k with *)
+ (* ImplicitArg (ref, (n, id), b) -> true *)
+ (* | InternalHole -> true *)
+ (* | _ -> false *)
+
(* To embed a boolean for resolvability status.
This is essentially a hack to mark which evars correspond to
@@ -463,42 +504,61 @@ let is_instance = function
*)
let resolvable = Store.field ()
-open Store.Field
+
+let set_resolvable s b =
+ Store.set s resolvable b
let is_resolvable evi =
- assert (evi.evar_body = Evar_empty);
- Option.default true (resolvable.get evi.evar_extra)
+ assert (match evi.evar_body with Evar_empty -> true | _ -> false);
+ Option.default true (Store.get evi.evar_extra resolvable)
let mark_resolvability_undef b evi =
- let t = resolvable.set b evi.evar_extra in
+ let t = Store.set evi.evar_extra resolvable b in
{ evi with evar_extra = t }
let mark_resolvability b evi =
- assert (evi.evar_body = Evar_empty);
+ assert (match evi.evar_body with Evar_empty -> true | _ -> false);
mark_resolvability_undef b evi
let mark_unresolvable evi = mark_resolvability false evi
let mark_resolvable evi = mark_resolvability true evi
-let mark_resolvability b sigma =
- Evd.fold_undefined (fun ev evi evs ->
- Evd.add evs ev (mark_resolvability_undef b evi))
- sigma (Evd.defined_evars sigma)
+open Evar_kinds
+type evar_filter = existential_key -> Evar_kinds.t -> bool
-let mark_unresolvables sigma = mark_resolvability false sigma
+let all_evars _ _ = true
+let all_goals _ = function VarInstance _ | GoalEvar -> true | _ -> false
+let no_goals ev evi = not (all_goals ev evi)
+let no_goals_or_obligations _ = function
+ | VarInstance _ | GoalEvar | QuestionMark _ -> false
+ | _ -> true
-let has_typeclasses evd =
- Evd.fold_undefined (fun ev evi has -> has ||
- (is_resolvable evi && is_class_evar evd evi))
- evd false
+let mark_resolvability filter b sigma =
+ let map ev evi =
+ if filter ev (snd evi.evar_source) then mark_resolvability_undef b evi
+ else evi
+ in
+ Evd.raw_map_undefined map sigma
+
+let mark_unresolvables ?(filter=all_evars) sigma = mark_resolvability filter false sigma
+let mark_resolvables ?(filter=all_evars) sigma = mark_resolvability filter true sigma
+
+let has_typeclasses filter evd =
+ let check ev evi =
+ filter ev (snd evi.evar_source) && is_resolvable evi && is_class_evar evd evi
+ in
+ Evar.Map.exists check (Evd.undefined_map evd)
-let solve_instanciations_problem = ref (fun _ _ _ _ _ -> assert false)
+let solve_instantiations_problem = ref (fun _ _ _ _ _ _ -> assert false)
-type evar_filter = hole_kind -> bool
+let solve_problem env evd filter unique split fail =
+ !solve_instantiations_problem env evd filter unique split fail
-let no_goals = function GoalEvar -> false | _ -> true
-let all_evars _ = true
+(** Profiling resolution of typeclasses *)
+(* let solve_classeskey = Profile.declare_profile "solve_typeclasses" *)
+(* let solve_problem = Profile.profile5 solve_classeskey solve_problem *)
-let resolve_typeclasses ?(filter=no_goals) ?(split=true) ?(fail=true) env evd =
- if not (has_typeclasses evd) then evd
- else !solve_instanciations_problem env evd filter split fail
+let resolve_typeclasses ?(filter=no_goals) ?(unique=get_typeclasses_unique_solutions ())
+ ?(split=true) ?(fail=true) env evd =
+ if not (has_typeclasses filter evd) then evd
+ else solve_problem env evd filter unique split fail
diff --git a/pretyping/typeclasses.mli b/pretyping/typeclasses.mli
index 91069b70..1a0b6696 100644
--- a/pretyping/typeclasses.mli
+++ b/pretyping/typeclasses.mli
@@ -1,22 +1,17 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Names
-open Libnames
-open Decl_kinds
+open Globnames
open Term
-open Sign
+open Context
open Evd
open Environ
-open Nametab
-open Mod_subst
-open Topconstr
-open Util
type direction = Forward | Backward
@@ -38,7 +33,14 @@ 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 * (direction * int option) option * constant option) list;
+ cl_projs : (Name.t * (direction * int option) option * constant option) list;
+
+ (** Whether we use matching or full unification during resolution *)
+ cl_strict : bool;
+
+ (** Whether we can assume that instances are unique, which allows
+ no backtracking and sharing of resolution. *)
+ cl_unique : bool;
}
type instance
@@ -49,64 +51,79 @@ val all_instances : unit -> instance list
val add_class : typeclass -> unit
-val add_constant_class : constant -> unit
-
-val add_inductive_class : inductive -> unit
-
-val new_instance : typeclass -> int option -> bool -> global_reference -> instance
+val new_instance : typeclass -> int option -> bool -> Decl_kinds.polymorphic ->
+ global_reference -> instance
val add_instance : instance -> unit
val remove_instance : instance -> unit
val class_info : global_reference -> typeclass (** raises a UserError if not a class *)
-(** These raise a UserError if not a class. *)
-val dest_class_app : env -> constr -> typeclass * constr list
+(** 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
+
+(** Get the instantiated typeclass structure for a given universe instance. *)
+val typeclass_univ_instance : typeclass puniverses -> typeclass puniverses
(** Just return None if not a class *)
-val class_of_constr : constr -> (rel_context * (typeclass * constr list)) option
+val class_of_constr : constr -> (rel_context * (typeclass puniverses * constr list)) option
val instance_impl : instance -> global_reference
+val instance_priority : instance -> int option
+
val is_class : global_reference -> bool
val is_instance : global_reference -> bool
+val is_implicit_arg : Evar_kinds.t -> bool
+
(** Returns the term and type for the given instance of the parameters and fields
of the type class. *)
-val instance_constructor : typeclass -> constr list -> constr option * types
+val instance_constructor : typeclass puniverses -> constr list ->
+ constr option * types
+
+(** Filter which evars to consider for resolution. *)
+type evar_filter = existential_key -> 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 could be marked or checked for resolvability. *)
+ Only undefined evars can be marked or checked for resolvability. *)
+val set_resolvable : Evd.Store.t -> bool -> Evd.Store.t
val is_resolvable : evar_info -> bool
val mark_unresolvable : evar_info -> evar_info
+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 mark_unresolvables : evar_map -> evar_map
val is_class_evar : evar_map -> evar_info -> bool
+val is_class_type : evar_map -> types -> bool
-(** Filter which evars to consider for resolution. *)
-type evar_filter = hole_kind -> bool
-val no_goals : evar_filter
-val all_evars : evar_filter
-
-val resolve_typeclasses : ?filter:evar_filter -> ?split:bool -> ?fail:bool ->
- env -> evar_map -> evar_map
-val resolve_one_typeclass : env -> evar_map -> types -> open_constr
+val resolve_typeclasses : ?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 register_set_typeclass_transparency : (evaluable_global_reference -> bool (*local?*) -> bool -> unit) -> unit
+val set_typeclass_transparency_hook : (evaluable_global_reference -> bool (*local?*) -> bool -> unit) Hook.t
val set_typeclass_transparency : evaluable_global_reference -> bool -> bool -> unit
-val register_classes_transparent_state : (unit -> transparent_state) -> unit
+val classes_transparent_state_hook : (unit -> transparent_state) Hook.t
val classes_transparent_state : unit -> transparent_state
-val register_add_instance_hint : (constr -> bool (* local? *) -> int option -> unit) -> unit
-val register_remove_instance_hint : (global_reference -> unit) -> unit
-val add_instance_hint : constr -> bool -> int option -> unit
+val add_instance_hint_hook :
+ (global_reference_or_constr -> global_reference list ->
+ bool (* local? *) -> int option -> Decl_kinds.polymorphic -> unit) Hook.t
+val remove_instance_hint_hook : (global_reference -> unit) Hook.t
+val add_instance_hint : global_reference_or_constr -> global_reference list ->
+ bool -> int option -> Decl_kinds.polymorphic -> unit
val remove_instance_hint : global_reference -> unit
-val solve_instanciations_problem : (env -> evar_map -> evar_filter -> bool -> bool -> evar_map) ref
-val solve_instanciation_problem : (env -> evar_map -> types -> open_constr) ref
+val solve_instantiations_problem : (env -> evar_map -> evar_filter -> bool -> bool -> bool -> evar_map) ref
+val solve_instantiation_problem : (env -> evar_map -> types -> bool -> open_constr) ref
val declare_instance : int option -> bool -> global_reference -> unit
@@ -116,4 +133,4 @@ val declare_instance : int option -> bool -> global_reference -> unit
subinstances and add only the missing ones. *)
val build_subclasses : check:bool -> env -> evar_map -> global_reference -> int option (* priority *) ->
- (int option * constr) list
+ (global_reference list * int option * constr) list
diff --git a/pretyping/typeclasses_errors.ml b/pretyping/typeclasses_errors.ml
index da5dc909..4f88dd86 100644
--- a/pretyping/typeclasses_errors.ml
+++ b/pretyping/typeclasses_errors.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,26 +8,19 @@
(*i*)
open Names
-open Decl_kinds
open Term
-open Sign
+open Context
open Evd
open Environ
-open Nametab
-open Mod_subst
-open Topconstr
-open Compat
-open Util
-open Libnames
+open Constrexpr
+open Globnames
(*i*)
type contexts = Parameters | Properties
type typeclass_error =
| NotAClass of constr
- | UnboundMethod of global_reference * identifier located (* Class name, method *)
- | NoInstance of identifier located * constr list
- | UnsatisfiableConstraints of evar_map * (existential_key * hole_kind) option
+ | UnboundMethod of global_reference * Id.t Loc.located (* Class name, method *)
| MismatchedContextInstance of contexts * constr_expr list * rel_context (* found, expected *)
exception TypeClassError of env * typeclass_error
@@ -38,21 +31,4 @@ let not_a_class env c = typeclass_error env (NotAClass c)
let unbound_method env cid id = typeclass_error env (UnboundMethod (cid, id))
-let no_instance env id args = typeclass_error env (NoInstance (id, args))
-
-let unsatisfiable_constraints env evd ev =
- match ev with
- | None ->
- raise (TypeClassError (env, UnsatisfiableConstraints (evd, None)))
- | Some ev ->
- let loc, kind = Evd.evar_source ev evd in
- raise (Loc.Exc_located (loc, TypeClassError
- (env, UnsatisfiableConstraints (evd, Some (ev, kind)))))
-
let mismatched_ctx_inst env c n m = typeclass_error env (MismatchedContextInstance (c, n, m))
-
-let rec unsatisfiable_exception exn =
- match exn with
- | TypeClassError (_, UnsatisfiableConstraints _) -> true
- | Loc.Exc_located(_, e) -> unsatisfiable_exception e
- | _ -> false
diff --git a/pretyping/typeclasses_errors.mli b/pretyping/typeclasses_errors.mli
index 6e0fc6ad..dd808771 100644
--- a/pretyping/typeclasses_errors.mli
+++ b/pretyping/typeclasses_errors.mli
@@ -1,42 +1,32 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Loc
open Names
-open Decl_kinds
open Term
-open Sign
+open Context
open Evd
open Environ
-open Nametab
-open Mod_subst
-open Topconstr
-open Util
-open Libnames
+open Constrexpr
+open Globnames
type contexts = Parameters | Properties
type typeclass_error =
| NotAClass of constr
- | UnboundMethod of global_reference * identifier located (** Class name, method *)
- | NoInstance of identifier located * constr list
- | UnsatisfiableConstraints of evar_map * (existential_key * hole_kind) option
+ | UnboundMethod of global_reference * Id.t located (** Class name, method *)
| MismatchedContextInstance of contexts * constr_expr list * rel_context (** found, expected *)
exception TypeClassError of env * typeclass_error
val not_a_class : env -> constr -> 'a
-val unbound_method : env -> global_reference -> identifier located -> 'a
-
-val no_instance : env -> identifier located -> constr list -> 'a
-
-val unsatisfiable_constraints : env -> evar_map -> evar option -> 'a
+val unbound_method : env -> global_reference -> Id.t located -> 'a
val mismatched_ctx_inst : env -> contexts -> constr_expr list -> rel_context -> 'a
-val unsatisfiable_exception : exn -> bool
diff --git a/pretyping/typing.ml b/pretyping/typing.ml
index 9fbfc197..c6209cc3 100644
--- a/pretyping/typing.ml
+++ b/pretyping/typing.ml
@@ -1,44 +1,44 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Pp
+open Errors
open Util
-open Names
open Term
+open Vars
open Environ
open Reductionops
open Type_errors
-open Pretype_errors
open Inductive
open Inductiveops
open Typeops
-open Evd
open Arguments_renaming
let meta_type evd mv =
let ty =
try Evd.meta_ftype evd mv
- with Not_found -> anomaly ("unknown meta ?"^Nameops.string_of_meta mv) in
+ with Not_found -> anomaly (str "unknown meta ?" ++ str (Nameops.string_of_meta mv)) in
meta_instance evd ty
let constant_type_knowing_parameters env cst jl =
- let paramstyp = Array.map (fun j -> j.uj_type) jl in
- type_of_constant_knowing_parameters env (constant_type env cst) paramstyp
+ 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 jl =
- let (mib,mip) = lookup_mind_specif env ind in
- let paramstyp = Array.map (fun j -> j.uj_type) jl in
- Inductive.type_of_inductive_knowing_parameters env mip paramstyp
+let inductive_type_knowing_parameters env (ind,u) jl =
+ let mspec = lookup_mind_specif env ind in
+ let paramstyp = Array.map (fun j -> lazy 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_betadeltaiota env !evdref j.uj_type) with
| Sort s -> {utj_val = j.uj_val; utj_type = s }
| Evar ev ->
- let (evd,s) = Evarutil.define_evar_as_sort !evdref ev in
+ let (evd,s) = Evarutil.define_evar_as_sort env !evdref ev in
evdref := evd; { utj_val = j.uj_val; utj_type = s }
| _ -> error_not_type env j
@@ -69,17 +69,17 @@ let e_judge_of_apply env evdref funj argjv =
in
apply_rec 1 funj.uj_type (Array.to_list argjv)
-let e_check_branch_types env evdref ind cj (lfj,explft) =
- if Array.length lfj <> Array.length explft then
+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);
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) lfj.(i).uj_type explft.(i)
+ error_ill_formed_branch env cj.uj_val ((ind,i+1),u) lfj.(i).uj_type explft.(i)
done
-let rec max_sort l =
- if List.mem InType l then InType else
- if List.mem InSet l then InSet else InProp
+let max_sort l =
+ if Sorts.List.mem InType l then InType else
+ 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
@@ -92,10 +92,11 @@ let e_is_correct_arity env evdref c pj ind specif params =
if not (Evarconv.e_cumul env evdref a1 a1') then error ();
srec (push_rel (na1,None,a1) env) t ar'
| Sort s, [] ->
- if not (List.mem (family_of_sort s) allowed_sorts) then error ()
+ if not (Sorts.List.mem (Sorts.family s) allowed_sorts)
+ then error ()
| Evar (ev,_), [] ->
- let s = Termops.new_sort_in_family (max_sort allowed_sorts) in
- evdref := Evd.define ev (mkSort s) !evdref
+ let evd, s = Evd.fresh_sort_in_family env !evdref (max_sort allowed_sorts) in
+ evdref := Evd.define ev (mkSort s) evd
| _, (_,Some _,_ as d)::ar' ->
srec (push_rel d env) (lift 1 pt') ar'
| _ ->
@@ -104,13 +105,13 @@ let e_is_correct_arity env evdref c pj ind specif params =
srec env pj.uj_type (List.rev arsign)
let e_type_case_branches env evdref (ind,largs) pj c =
- let specif = lookup_mind_specif env ind in
+ let specif = lookup_mind_specif env (fst ind) in
let nparams = inductive_params specif in
- let (params,realargs) = list_chop nparams largs in
+ let (params,realargs) = List.chop nparams largs in
let p = pj.uj_val in
let univ = e_is_correct_arity env evdref c pj ind specif params in
let lc = build_branches_type ind specif params p in
- let n = (snd specif).Declarations.mind_nrealargs_ctxt in
+ let n = (snd specif).Declarations.mind_nrealargs in
let ty =
whd_betaiota !evdref (Reduction.betazeta_appvect (n+1) p (Array.of_list (realargs@[c]))) in
(lc, ty, univ)
@@ -125,12 +126,13 @@ let e_judge_of_case env evdref ci pj cj lfj =
{ uj_val = mkCase (ci, pj.uj_val, cj.uj_val, Array.map j_val lfj);
uj_type = rslty }
+(* 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 specif = Global.lookup_inductive ind in
+ let specif = Global.lookup_inductive (fst ind) in
let sorts = elim_sorts specif in
- if not (List.exists ((=) ksort) sorts) then
+ 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))
@@ -195,22 +197,26 @@ let rec execute env evdref cstr =
judge_of_prop_contents c
| Sort (Type u) ->
- judge_of_type u
+ judge_of_type u
+
+ | Proj (p, c) ->
+ let cj = execute env evdref c in
+ judge_of_projection env p (Evarutil.j_nf_evar !evdref cj)
| App (f,args) ->
let jl = execute_array env evdref args in
let j =
match kind_of_term f with
- | Ind ind ->
+ | Ind ind when Environ.template_polymorphic_pind ind env ->
(* Sort-polymorphism of inductive types *)
make_judge f
(inductive_type_knowing_parameters env ind
- (jv_nf_evar !evdref jl))
- | Const cst ->
+ (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
- (jv_nf_evar !evdref jl))
+ (Evarutil.jv_nf_evar !evdref jl))
| _ ->
execute env evdref f
in
@@ -235,7 +241,7 @@ let rec execute env evdref cstr =
let j1 = execute env evdref c1 in
let j2 = execute env evdref c2 in
let j2 = e_type_judgment env evdref j2 in
- let _ = judge_of_cast env j1 DEFAULTcast j2 in
+ let _ = e_judge_of_cast env evdref j1 DEFAULTcast j2 in
let env1 = push_rel (name,Some j1.uj_val,j2.utj_val) env in
let j3 = execute env1 evdref c3 in
judge_of_letin env name j1 j2 j3
@@ -257,40 +263,37 @@ and execute_recdef env evdref (names,lar,vdef) =
and execute_array env evdref = Array.map (execute env evdref)
-let check env evd c t =
- let evdref = ref evd in
+let check env evdref c t =
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 evd t)
+ error_actual_type env j (nf_evar !evdref t)
(* Type of a constr *)
let type_of env evd c =
let j = execute env (ref evd) c in
- (* We are outside the kernel: we take fresh universes *)
- (* to avoid tactics and co to refresh universes themselves *)
- Termops.refresh_universes j.uj_type
+ j.uj_type
(* Sort of a type *)
-let sort_of env evd c =
- let evdref = ref evd in
+let sort_of env evdref c =
let j = execute env evdref c in
let a = e_type_judgment env evdref j in
a.utj_type
(* Try to solve the existential variables by typing *)
-let e_type_of env evd c =
+let e_type_of ?(refresh=false) env evd c =
let evdref = ref evd in
let j = execute env evdref c in
(* side-effect on evdref *)
- !evdref, Termops.refresh_universes j.uj_type
+ if refresh then
+ Evarsolve.refresh_universes ~onlyalg:true (Some false) env !evdref j.uj_type
+ else !evdref, j.uj_type
-let solve_evars env evd c =
- let evdref = ref evd in
+let solve_evars env evdref c =
let c = (execute env evdref c).uj_val in
(* side-effect on evdref *)
- !evdref, nf_evar !evdref c
+ nf_evar !evdref c
let _ = Evarconv.set_solve_evars solve_evars
diff --git a/pretyping/typing.mli b/pretyping/typing.mli
index 5aa0a2d4..c933106d 100644
--- a/pretyping/typing.mli
+++ b/pretyping/typing.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -12,27 +12,28 @@ open Environ
open Evd
(** This module provides the typing machine with existential variables
- (but without universes). *)
+ and universes. *)
(** Typecheck a term and return its type *)
val type_of : env -> evar_map -> constr -> types
-(** Typecheck a term and return its type + updated evars *)
-val e_type_of : env -> evar_map -> constr -> evar_map * types
+(** Typecheck a term and return its type + updated evars, optionally refreshing
+ universes *)
+val e_type_of : ?refresh:bool -> env -> evar_map -> constr -> evar_map * types
(** Typecheck a type and return its sort *)
-val sort_of : env -> evar_map -> types -> sorts
+val sort_of : env -> evar_map ref -> types -> sorts
(** Typecheck a term has a given type (assuming the type is OK) *)
-val check : env -> evar_map -> constr -> types -> unit
+val check : env -> evar_map ref -> constr -> types -> unit
(** Returns the instantiated type of a metavariable *)
val meta_type : evar_map -> metavariable -> types
(** Solve existential variables using typing *)
-val solve_evars : env -> evar_map -> constr -> evar_map * constr
+val solve_evars : env -> evar_map ref -> constr -> constr
(** Raise an error message if incorrect elimination for this inductive *)
(** (first constr is term to match, second is return predicate) *)
-val check_allowed_sort : env -> evar_map -> inductive -> constr -> constr ->
+val check_allowed_sort : env -> evar_map -> pinductive -> constr -> constr ->
unit
diff --git a/pretyping/unification.ml b/pretyping/unification.ml
index d6b1e2e4..203b1ec8 100644
--- a/pretyping/unification.ml
+++ b/pretyping/unification.ml
@@ -1,29 +1,51 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Errors
open Pp
open Util
open Names
-open Nameops
open Term
+open Vars
open Termops
open Namegen
-open Sign
open Environ
open Evd
open Reduction
open Reductionops
-open Glob_term
open Evarutil
+open Evarsolve
open Pretype_errors
open Retyping
-open Coercion.Default
+open Coercion
open Recordops
+open Locus
+open Locusops
+open Find_subterm
+
+let keyed_unification = ref (false)
+let _ = Goptions.declare_bool_option {
+ Goptions.optsync = true; Goptions.optdepr = false;
+ Goptions.optname = "Unification is keyed";
+ Goptions.optkey = ["Keyed";"Unification"];
+ Goptions.optread = (fun () -> !keyed_unification);
+ Goptions.optwrite = (fun a -> keyed_unification:=a);
+}
+
+let debug_unification = ref (false)
+let _ = Goptions.declare_bool_option {
+ Goptions.optsync = true; Goptions.optdepr = false;
+ Goptions.optname =
+ "Print states sent to tactic unification";
+ Goptions.optkey = ["Debug";"Tactic";"Unification"];
+ Goptions.optread = (fun () -> !debug_unification);
+ Goptions.optwrite = (fun a -> debug_unification:=a);
+}
let occur_meta_or_undefined_evar evd c =
let rec occrec c = match kind_of_term c with
@@ -33,7 +55,6 @@ let occur_meta_or_undefined_evar evd c =
| Evar_defined c ->
occrec c; Array.iter occrec args
| Evar_empty -> raise Occur)
- | Sort s when is_sort_variable evd s -> raise Occur
| _ -> iter_constr occrec c
in try occrec c; false with Occur | Not_found -> true
@@ -42,48 +63,60 @@ let occur_meta_evd sigma mv 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
- | Meta mv' when mv = mv' -> raise Occur
+ | Meta mv' when Int.equal mv mv' -> raise Occur
| _ -> iter_constr 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 c l lname_typ =
+let abstract_scheme env evd c l lname_typ =
List.fold_left2
- (fun t (locc,a) (na,_,ta) ->
+ (fun (t,evd) (locc,a) (na,_,ta) ->
let na = match kind_of_term 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)
- else mkLambda_name env (na,ta,subst_closed_term_occ locc a t))
- c
+ else *)
+ if occur_meta 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')
+ (c,evd)
(List.rev l)
lname_typ
+(* Precondition: resulting abstraction is expected to be of type [typ] *)
+
let abstract_list_all env evd typ c l =
let ctxt,_ = splay_prod_n env evd (List.length l) typ in
- let l_with_all_occs = List.map (function a -> (all_occurrences,a)) l in
- let p = abstract_scheme env c l_with_all_occs ctxt in
- try
- if is_conv_leq env evd (Typing.type_of env evd p) typ then p
- else error "abstract_list_all"
- with UserError _ | Type_errors.TypeError _ ->
- error_cannot_find_well_typed_abstraction env evd p l
+ let l_with_all_occs = List.map (function a -> (LikeFirst,a)) l in
+ let p,evd = abstract_scheme env evd c l_with_all_occs ctxt in
+ let evd,typp =
+ try Typing.e_type_of env evd p
+ with
+ | UserError _ ->
+ error_cannot_find_well_typed_abstraction env evd p l None
+ | Type_errors.TypeError (env',x) ->
+ error_cannot_find_well_typed_abstraction env evd p l (Some (env',x)) in
+ evd,(p,typp)
let set_occurrences_of_last_arg args =
- Some all_occurrences :: List.tl (array_map_to_list (fun _ -> None) 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,ev = new_evar evd env typ in
+ let evd,ev = new_evar env evd typ in
let evd,ev' = evar_absorb_arguments env evd (destEvar ev) l in
- let argoccs = set_occurrences_of_last_arg (snd ev') 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 nf_evar evd (existential_value evd (destEvar ev))
- else error "Cannot find a well-typed abstraction."
+ if b then
+ let p = nf_evar evd (existential_value evd (destEvar ev)) in
+ evd, p
+ else error_cannot_find_well_typed_abstraction env evd
+ (nf_evar evd c) l None
(**)
@@ -103,15 +136,15 @@ let extract_instance_status = function
| CUMUL -> add_type_status (IsSubType, IsSuperType)
| CONV -> add_type_status (Conv, Conv)
-let rec assoc_pair x = function
- [] -> raise Not_found
- | (a,b,_)::l -> if compare a x = 0 then b else assoc_pair x l
-
let rec subst_meta_instances bl c =
match kind_of_term c with
- | Meta i -> (try assoc_pair i bl with Not_found -> c)
+ | Meta i ->
+ let select (j,_,_) = Int.equal i j in
+ (try pi2 (List.find select bl) with Not_found -> c)
| _ -> map_constr (subst_meta_instances 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
@@ -120,8 +153,9 @@ let pose_all_metas_as_evars env evd t =
| Some ({rebus=c},_) -> c
| None ->
let {rebus=ty;freemetas=mvs} = Evd.meta_ftype evd mv in
- let ty = if mvs = Evd.Metaset.empty then ty else aux ty in
- let ev = Evarutil.e_new_evar evdref env ~src:(dummy_loc,GoalEvar) ty in
+ let ty = if Evd.Metaset.is_empty mvs then ty else aux ty 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;
ev)
| _ ->
@@ -133,14 +167,18 @@ let pose_all_metas_as_evars env evd t =
let solve_pattern_eqn_array (env,nb) f l c (sigma,metasubst,evarsubst) =
match kind_of_term f with
| Meta k ->
- let sigma,c = pose_all_metas_as_evars env sigma c in
+ (* 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 pb = (Conv,TypeNotProcessed) in
if noccur_between 1 nb c then
sigma,(k,lift (-nb) c,pb)::metasubst,evarsubst
else error_cannot_unify_local env sigma (applist (f, l),c,c)
| Evar ev ->
- let sigma,c = pose_all_metas_as_evars env sigma c in
+ 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
| _ -> assert false
@@ -166,18 +204,23 @@ let unify_r2l x = x
let sort_eqns = unify_r2l
*)
-(* Option introduced and activated in Coq 8.3 *)
-let global_evars_pattern_unification_flag = ref true
+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 = false;
+ optdepr = true;
optname = "pattern-unification for existential variables in tactics";
optkey = ["Tactic";"Evars";"Pattern";"Unification"];
- optread = (fun () -> !global_evars_pattern_unification_flag);
- optwrite = (:=) global_evars_pattern_unification_flag }
+ 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
@@ -185,10 +228,10 @@ let _ =
optdepr = false;
optname = "pattern-unification for existential variables in tactics";
optkey = ["Tactic";"Pattern";"Unification"];
- optread = (fun () -> !global_evars_pattern_unification_flag);
- optwrite = (:=) global_evars_pattern_unification_flag }
+ optread = (fun () -> !global_pattern_unification_flag);
+ optwrite = (:=) global_pattern_unification_flag }
-type unify_flags = {
+type core_unify_flags = {
modulo_conv_on_closed_terms : Names.transparent_state option;
(* What this flag controls was activated with all constants transparent, *)
(* even for auto, since Coq V5.10 *)
@@ -197,37 +240,33 @@ type unify_flags = {
(* This refinement of the conversion on closed terms is activable *)
(* (and activated for apply, rewrite but not auto since Feb 2008 for 8.2) *)
+ use_evars_eagerly_in_conv_on_closed_terms : bool;
+
modulo_delta : Names.transparent_state;
(* This controls which constants are unfoldable; this is on for apply *)
(* (but not simple apply) since Feb 2008 for 8.2 *)
modulo_delta_types : Names.transparent_state;
- modulo_delta_in_merge : Names.transparent_state option;
- (* This controls whether unfoldability is different when trying to unify *)
- (* several instances of the same metavariable *)
- (* Typical situation is when we give a pattern to be matched *)
- (* syntactically against a subterm but we want the metas of the *)
- (* pattern to be modulo convertibility *)
-
check_applied_meta_types : bool;
(* This controls whether meta's applied to arguments have their *)
(* type unified with the type of their instance *)
- resolve_evars : bool;
- (* This says if type classes instances resolution must be used to infer *)
- (* the remaining evars *)
-
use_pattern_unification : bool;
- (* This says if type classes instances resolution must be used to infer *)
- (* the remaining evars *)
+ (* 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" *)
use_meta_bound_pattern_unification : bool;
- (* This solves pattern "?n x1 ... xn = t" when the xi are distinct rels *)
- (* This allows for instance to unify "forall x:A, B(x)" with "A' -> B'" *)
+ (* This is implied by use_pattern_unification (though deactivated *)
+ (* by unsetting Tactic 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 *)
+ (* This allowed for instance to unify "forall x:?A, ?B x" with "A' -> B'" *)
+ (* when ?B is a Meta. *)
- frozen_evars : ExistentialSet.t;
+ frozen_evars : Evar.Set.t;
(* Evars of this set are considered axioms and never instantiated *)
(* Useful e.g. for autorewrite *)
@@ -240,43 +279,86 @@ type unify_flags = {
modulo_eta : bool;
(* Support eta in the reduction *)
+}
+
+type unify_flags = {
+ core_unify_flags : core_unify_flags;
+ (* Governs unification of problems of the form "t(?x) = u(?x)" in apply *)
+
+ merge_unify_flags : core_unify_flags;
+ (* These are the flags to be used when trying to unify *)
+ (* several instances of the same metavariable *)
+ (* Typical situation is when we give a pattern to be matched *)
+ (* syntactically against a subterm but we want the metas of the *)
+ (* pattern to be modulo convertibility *)
+
+ subterm_unify_flags : core_unify_flags;
+ (* Governs unification of problems of the form "?X a1..an = u" in apply, *)
+ (* hence in rewrite and elim *)
- allow_K_in_toplevel_higher_order_unification : bool
- (* This is used only in second/higher order unification when looking for *)
- (* subterms (rewrite and elim) *)
+ allow_K_in_toplevel_higher_order_unification : bool;
+ (* Tells in second-order abstraction over subterms which have not *)
+ (* been found in term are allowed (used for rewrite, elim, or *)
+ (* apply with a lemma whose type has the form "?X a1 ... an") *)
+
+ resolve_evars : bool
+ (* This says if type classes instances resolution must be used to infer *)
+ (* the remaining evars *)
}
(* Default flag for unifying a type against a type (e.g. apply) *)
(* We set all conversion flags (no flag should be modified anymore) *)
-let default_unify_flags = {
- modulo_conv_on_closed_terms = Some full_transparent_state;
+let default_core_unify_flags () =
+ let ts = Names.full_transparent_state in {
+ modulo_conv_on_closed_terms = Some ts;
use_metas_eagerly_in_conv_on_closed_terms = true;
- modulo_delta = full_transparent_state;
- modulo_delta_types = full_transparent_state;
- modulo_delta_in_merge = None;
+ use_evars_eagerly_in_conv_on_closed_terms = false;
+ modulo_delta = ts;
+ modulo_delta_types = ts;
check_applied_meta_types = true;
- resolve_evars = false;
use_pattern_unification = true;
use_meta_bound_pattern_unification = true;
- frozen_evars = ExistentialSet.empty;
+ frozen_evars = Evar.Set.empty;
restrict_conv_on_strict_subterms = false;
modulo_betaiota = true;
modulo_eta = true;
- allow_K_in_toplevel_higher_order_unification = false
- (* in fact useless when not used in w_unify_to_subterm_list *)
+ }
+
+(* Default flag for first-order or second-order unification of a type *)
+(* against another type (e.g. apply) *)
+(* We set all conversion flags (no flag should be modified anymore) *)
+let default_unify_flags () =
+ let flags = default_core_unify_flags () in {
+ core_unify_flags = flags;
+ merge_unify_flags = flags;
+ subterm_unify_flags = { flags with modulo_delta = var_full_transparent_state };
+ allow_K_in_toplevel_higher_order_unification = false; (* Why not? *)
+ resolve_evars = false
+}
+
+let set_no_delta_core_flags flags = { flags with
+ modulo_conv_on_closed_terms = None;
+ modulo_delta = empty_transparent_state;
+ check_applied_meta_types = false;
+ use_pattern_unification = false;
+ use_meta_bound_pattern_unification = true;
+ modulo_betaiota = false
}
-let set_merge_flags flags =
- match flags.modulo_delta_in_merge with
- | None -> flags
- | Some ts ->
- { flags with modulo_delta = ts; modulo_conv_on_closed_terms = Some ts }
+let set_no_delta_flags flags = {
+ core_unify_flags = set_no_delta_core_flags flags.core_unify_flags;
+ merge_unify_flags = set_no_delta_core_flags flags.merge_unify_flags;
+ subterm_unify_flags = set_no_delta_core_flags flags.subterm_unify_flags;
+ allow_K_in_toplevel_higher_order_unification =
+ flags.allow_K_in_toplevel_higher_order_unification;
+ resolve_evars = flags.resolve_evars
+}
(* Default flag for the "simple apply" version of unification of a *)
(* type against a type (e.g. apply) *)
-(* We set only the flags available at the time the new "apply" extends *)
+(* We set only the flags available at the time the new "apply" extended *)
(* out of "simple apply" *)
-let default_no_delta_unify_flags = { default_unify_flags with
+let default_no_delta_core_unify_flags () = { (default_core_unify_flags ()) with
modulo_delta = empty_transparent_state;
check_applied_meta_types = false;
use_pattern_unification = false;
@@ -284,56 +366,133 @@ let default_no_delta_unify_flags = { default_unify_flags with
modulo_betaiota = false;
}
+let default_no_delta_unify_flags () =
+ let flags = default_no_delta_core_unify_flags () in {
+ core_unify_flags = flags;
+ merge_unify_flags = flags;
+ subterm_unify_flags = flags;
+ allow_K_in_toplevel_higher_order_unification = false;
+ resolve_evars = false
+}
+
(* Default flags for looking for subterms in elimination tactics *)
(* Not used in practice at the current date, to the exception of *)
(* allow_K) because only closed terms are involved in *)
(* induction/destruct/case/elim and w_unify_to_subterm_list does not *)
(* call w_unify for induction/destruct/case/elim (13/6/2011) *)
-let elim_flags = { default_unify_flags with
- restrict_conv_on_strict_subterms = false; (* ? *)
+let elim_core_flags sigma = { (default_core_unify_flags ()) with
modulo_betaiota = false;
- allow_K_in_toplevel_higher_order_unification = true
+ frozen_evars =
+ fold_undefined (fun evk _ evars -> Evar.Set.add evk evars)
+ sigma Evar.Set.empty;
}
-let elim_no_delta_flags = { elim_flags with
+let elim_flags_evars sigma =
+ let flags = elim_core_flags sigma in {
+ core_unify_flags = flags;
+ merge_unify_flags = flags;
+ subterm_unify_flags = { flags with modulo_delta = empty_transparent_state };
+ allow_K_in_toplevel_higher_order_unification = true;
+ resolve_evars = false
+}
+
+let elim_flags () = elim_flags_evars Evd.empty
+
+let elim_no_delta_core_flags () = { (elim_core_flags Evd.empty) with
modulo_delta = empty_transparent_state;
check_applied_meta_types = false;
use_pattern_unification = false;
+ modulo_betaiota = false;
}
-let set_no_head_reduction flags =
- { flags with restrict_conv_on_strict_subterms = true }
+let elim_no_delta_flags () =
+ let flags = elim_no_delta_core_flags () in {
+ core_unify_flags = flags;
+ merge_unify_flags = flags;
+ subterm_unify_flags = flags;
+ allow_K_in_toplevel_higher_order_unification = true;
+ resolve_evars = false
+}
+
+(* On types, we don't restrict unification, but possibly for delta *)
+let set_flags_for_type flags = { flags with
+ modulo_delta = flags.modulo_delta_types;
+ modulo_conv_on_closed_terms = Some flags.modulo_delta_types;
+ use_pattern_unification = true;
+ modulo_betaiota = true;
+ modulo_eta = true;
+}
let use_evars_pattern_unification flags =
- !global_evars_pattern_unification_flag && flags.use_pattern_unification
+ !global_pattern_unification_flag && flags.use_pattern_unification
&& Flags.version_strictly_greater Flags.V8_2
let use_metas_pattern_unification flags nb l =
- !global_evars_pattern_unification_flag && flags.use_pattern_unification
+ !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 expand_key env = function
- | Some (ConstKey cst) -> constant_opt_value env cst
- | Some (VarKey id) -> (try named_body id env with Not_found -> None)
- | Some (RelKey _) -> None
+ Array.for_all (fun c -> isRel c && destRel c <= nb) l
+
+type key =
+ | IsKey of Closure.table_key
+ | IsProj of projection * constr
+
+let expand_table_key env = function
+ | ConstKey cst -> constant_opt_value_in env cst
+ | VarKey id -> (try named_body id env with Not_found -> None)
+ | RelKey _ -> None
+
+let unfold_projection env p stk =
+ (match try Some (lookup_projection p env) with Not_found -> None with
+ | Some pb ->
+ let s = Stack.Proj (pb.Declarations.proj_npars, pb.Declarations.proj_arg,
+ p, Cst_stack.empty) in
+ s :: stk
+ | None -> assert false)
+
+let expand_key ts env sigma = function
+ | Some (IsKey k) -> expand_table_key env k
+ | Some (IsProj (p, c)) ->
+ let red = Stack.zip (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
| None -> None
-let subterm_restriction is_subterm flags =
- not is_subterm && flags.restrict_conv_on_strict_subterms
+
+type unirec_flags = {
+ at_top: bool;
+ with_types: bool;
+ with_cs : bool;
+}
+
+let subterm_restriction opt flags =
+ not opt.at_top && flags.restrict_conv_on_strict_subterms
-let key_of b flags f =
+let key_of env b flags f =
if subterm_restriction b flags then None else
match kind_of_term f with
- | Const cst when is_transparent (ConstKey cst) &&
- Cpred.mem cst (snd flags.modulo_delta) ->
- Some (ConstKey cst)
- | Var id when is_transparent (VarKey id) &&
- Idpred.mem id (fst flags.modulo_delta) ->
- Some (VarKey id)
+ | Const (cst, u) when is_transparent env (ConstKey cst) &&
+ (Cpred.mem cst (snd flags.modulo_delta)
+ || Environ.is_projection cst env) ->
+ Some (IsKey (ConstKey (cst, u)))
+ | Var id when is_transparent env (VarKey id) &&
+ Id.Pred.mem id (fst flags.modulo_delta) ->
+ Some (IsKey (VarKey id))
+ | Proj (p, c) when Projection.unfolded p
+ || Cpred.mem (Projection.constant p) (snd flags.modulo_delta) ->
+ Some (IsProj (p, c))
| _ -> None
+
+let translate_key = function
+ | ConstKey (cst,u) -> ConstKey cst
+ | VarKey id -> VarKey id
+ | RelKey n -> RelKey n
+
+let translate_key = function
+ | IsKey k -> translate_key k
+ | IsProj (c, _) -> ConstKey (Projection.constant c)
+
let oracle_order env cf1 cf2 =
match cf1 with
| None ->
@@ -343,57 +502,151 @@ let oracle_order env cf1 cf2 =
| Some k1 ->
match cf2 with
| None -> Some true
- | Some k2 -> Some (Conv_oracle.oracle_order false k1 k2)
+ | Some k2 ->
+ match k1, k2 with
+ | IsProj (p, _), IsKey (ConstKey (p',_))
+ when eq_constant (Projection.constant p) p' ->
+ Some (not (Projection.unfolded p))
+ | IsKey (ConstKey (p,_)), IsProj (p', _)
+ when eq_constant 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
+ | 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
+ in
+ if b then
+ try Evd.add_universe_constraints sigma cstrs, b
+ with Univ.UniverseInconsistency _ -> sigma, false
+ | Evd.UniversesDiffer ->
+ if is_rigid_head flags t then
+ try Evd.add_universe_constraints sigma (force_eqs cstrs), b
+ with Univ.UniverseInconsistency _ -> sigma, false
+ else sigma, false
+ else sigma, b
+
let do_reduce ts (env, nb) sigma c =
- let (t, stack') = whd_betaiota_deltazeta_for_iota_state ts env sigma (c, empty_stack) in
- let l = list_of_stack stack' in
- applist (t, l)
+ Stack.zip (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
- | Evar (evk,_) -> not (ExistentialSet.mem evk flags.frozen_evars)
+ | Evar (evk,_) -> not (Evar.Set.mem evk flags.frozen_evars)
| _ -> false
-let check_compatibility env (sigma,metasubst,evarsubst) tyM tyN =
- match subst_defined_metas metasubst tyM with
- | None -> ()
+let check_compatibility env pbty flags (sigma,metasubst,evarsubst) tyM tyN =
+ match subst_defined_metas_evars (metasubst,[]) tyM with
+ | None -> sigma
| Some m ->
- match subst_defined_metas metasubst tyN with
- | None -> ()
+ match subst_defined_metas_evars (metasubst,[]) tyN with
+ | None -> sigma
| Some n ->
- if not (is_trans_fconv CONV full_transparent_state env sigma m n)
- && is_ground_term sigma m && is_ground_term sigma n
- then
- error_cannot_unify env sigma (m,n)
+ if is_ground_term sigma m && is_ground_term sigma n then
+ let sigma, b = infer_conv ~pb:pbty ~ts:flags.modulo_delta_types env sigma m n in
+ if b then sigma
+ else error_cannot_unify env sigma (m,n)
+ else sigma
+
+
+let rec is_neutral env ts t =
+ let (f, l) = decompose_appvect t in
+ match kind_of_term f with
+ | Const (c, u) ->
+ not (Environ.evaluable_constant c env) ||
+ not (is_transparent env (ConstKey c)) ||
+ not (Cpred.mem c (snd ts))
+ | Var id ->
+ not (Environ.evaluable_named id env) ||
+ not (is_transparent env (VarKey id)) ||
+ 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
+ | 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.CoFinite &&
+ Array.length projs == Array.length l1 - mib.Declarations.mind_nparams ->
+ (** Check that the other term is neutral *)
+ is_neutral env ts term
+ | _ -> false)
+ | _ -> false
-let 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 b wt ((sigma,metasubst,evarsubst) as substn) curm curn =
+let eta_constructor_app env f l1 term =
+ match kind_of_term f with
+ | Construct (((_, i as ind), j), u) ->
+ let mib = lookup_mind (fst ind) env in
+ (match mib.Declarations.mind_record with
+ | Some (Some (_, projs, _)) ->
+ let npars = mib.Declarations.mind_nparams in
+ let pars, l1' = Array.chop npars l1 in
+ let arg = Array.append pars [|term|] in
+ let l2 = Array.map (fun p -> mkApp (mkConstU (p,u), arg)) projs in
+ l1', l2
+ | _ -> 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 cM = Evarutil.whd_head_evar sigma curm
and cN = Evarutil.whd_head_evar sigma curn in
+ let () =
+ if !debug_unification then
+ msg_debug (Termops.print_constr_env curenv cM ++ str" ~= " ++ Termops.print_constr_env curenv cN)
+ in
match (kind_of_term cM,kind_of_term cN) with
| Meta k1, Meta k2 ->
- if k1 = k2 then substn else
+ if Int.equal k1 k2 then substn else
let stM,stN = extract_instance_status pb in
- if wt && flags.check_applied_meta_types then
- (let tyM = Typing.meta_type sigma k1 in
- let tyN = Typing.meta_type sigma k2 in
- check_compatibility curenv substn tyM tyN);
+ let sigma =
+ if opt.with_types && flags.check_applied_meta_types then
+ let tyM = Typing.meta_type sigma k1 in
+ let tyN = Typing.meta_type sigma k2 in
+ let l, r = if k2 < k1 then tyN, tyM else tyM, tyN in
+ check_compatibility curenv CUMUL flags substn l r
+ else sigma
+ in
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 *) ->
- if wt && flags.check_applied_meta_types then
- (try
- let tyM = Typing.meta_type sigma k in
- let tyN = get_type_of curenv sigma cN in
- check_compatibility curenv substn tyM tyN
- with Anomaly _ (* Hack *) ->
- (* Renounce, maybe metas/evars prevents typing *) ());
+ let sigma =
+ if opt.with_types && flags.check_applied_meta_types then
+ (try
+ let tyM = Typing.meta_type sigma k in
+ let tyN = get_type_of curenv ~lax:true sigma cN in
+ check_compatibility curenv CUMUL flags substn tyN tyM
+ with RetypeError _ ->
+ (* Renounce, maybe metas/evars prevents typing *) sigma)
+ else sigma
+ in
(* Here we check that [cN] does not contain any local variables *)
- if nb = 0 then
+ if Int.equal nb 0 then
sigma,(k,cN,snd (extract_instance_status pb))::metasubst,evarsubst
else if noccur_between 1 nb cN then
(sigma,
@@ -402,15 +655,18 @@ let unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb flag
else error_cannot_unify_local curenv sigma (m,n,cN)
| _, Meta k
when not (dependent cN cM) (* helps early trying alternatives *) ->
- if wt && flags.check_applied_meta_types then
+ let sigma =
+ if opt.with_types && flags.check_applied_meta_types then
(try
- let tyM = get_type_of curenv sigma cM in
+ let tyM = get_type_of curenv ~lax:true sigma cM in
let tyN = Typing.meta_type sigma k in
- check_compatibility curenv substn tyM tyN
- with Anomaly _ (* Hack *) ->
- (* Renounce, maybe metas/evars prevents typing *) ());
+ check_compatibility curenv CUMUL flags substn tyM tyN
+ with RetypeError _ ->
+ (* Renounce, maybe metas/evars prevents typing *) sigma)
+ else sigma
+ in
(* Here we check that [cM] does not contain any local variables *)
- if nb = 0 then
+ if Int.equal nb 0 then
(sigma,(k,cM,fst (extract_instance_status pb))::metasubst,evarsubst)
else if noccur_between 1 nb cM
then
@@ -418,125 +674,205 @@ let unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb flag
evarsubst)
else error_cannot_unify_local curenv sigma (m,n,cM)
| Evar (evk,_ as ev), _
- when not (ExistentialSet.mem evk flags.frozen_evars) ->
+ 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
- if Intset.subset cnvars cmvars then
+ 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 (ExistentialSet.mem evk flags.frozen_evars) ->
+ 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
- if Intset.subset cmvars cnvars then
+ 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 sigma' =
- if cv_pb = CUMUL
- then Evd.set_leq_sort sigma s1 s2
- else Evd.set_eq_sort sigma s1 s2
+ if pb == CUMUL
+ then Evd.set_leq_sort curenv sigma s1 s2
+ else Evd.set_eq_sort curenv sigma s1 s2
in (sigma', metasubst, evarsubst)
with e when Errors.noncritical e ->
error_cannot_unify curenv sigma (m,n))
| Lambda (na,t1,c1), Lambda (_,t2,c2) ->
- unirec_rec (push (na,t1) curenvnb) CONV true wt
- (unirec_rec curenvnb CONV true false substn t1 t2) c1 c2
+ unirec_rec (push (na,t1) curenvnb) CONV {opt with at_top = true}
+ (unirec_rec curenvnb CONV {opt with at_top = true; with_types = false} substn t1 t2) c1 c2
| Prod (na,t1,c1), Prod (_,t2,c2) ->
- unirec_rec (push (na,t1) curenvnb) pb true false
- (unirec_rec curenvnb CONV true false substn t1 t2) c1 c2
- | LetIn (_,a,_,c), _ -> unirec_rec curenvnb pb b wt substn (subst1 a c) cN
- | _, LetIn (_,a,_,c) -> unirec_rec curenvnb pb b wt substn cM (subst1 a c)
+ unirec_rec (push (na,t1) curenvnb) pb {opt with at_top = true}
+ (unirec_rec curenvnb CONV {opt with at_top = true; with_types = false} substn t1 t2) c1 c2
+ | LetIn (_,a,_,c), _ -> unirec_rec curenvnb pb opt substn (subst1 a c) cN
+ | _, 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
+ (Projection.constant p1) (Projection.constant p2) ->
+ (try unify_same_proj curenvnb cv_pb {opt with at_top = true}
+ substn c1 c2
+ with ex when precatchable_exception ex ->
+ unify_not_same_head curenvnb pb opt substn cM cN)
(* eta-expansion *)
| Lambda (na,t1,c1), _ when flags.modulo_eta ->
- unirec_rec (push (na,t1) curenvnb) CONV true wt substn
+ unirec_rec (push (na,t1) curenvnb) CONV {opt with at_top = true} substn
c1 (mkApp (lift 1 cN,[|mkRel 1|]))
| _, Lambda (na,t2,c2) when flags.modulo_eta ->
- unirec_rec (push (na,t2) curenvnb) CONV true wt substn
+ unirec_rec (push (na,t2) curenvnb) CONV {opt with at_top = true} substn
(mkApp (lift 1 cM,[|mkRel 1|])) c2
+ (* For records *)
+ | 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 ->
+ (try
+ let l1', l2' = eta_constructor_app curenv 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
+ | App(f2,l2) when
+ (isMeta f2 && use_metas_pattern_unification flags nb l2
+ || use_evars_pattern_unification flags && isAllowedEvar 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 ->
+ (try
+ let l2', l1' = eta_constructor_app curenv 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
+ | App(f1,l1) when
+ (isMeta f1 && use_metas_pattern_unification flags nb l1
+ || use_evars_pattern_unification flags && isAllowedEvar flags f1) ->
+ unify_app_pattern true curenvnb pb opt substn cM f1 l1 cN f2 l2
+ | _ -> raise ex)
+
| Case (_,p1,c1,cl1), Case (_,p2,c2,cl2) ->
(try
- array_fold_left2 (unirec_rec curenvnb CONV true wt)
- (unirec_rec curenvnb CONV true false
- (unirec_rec curenvnb CONV true false substn p1 p2) c1 c2)
+ let opt' = {opt with at_top = true; with_types = false} in
+ Array.fold_left2 (unirec_rec curenvnb CONV {opt with at_top = true})
+ (unirec_rec curenvnb CONV opt'
+ (unirec_rec curenvnb CONV opt' substn p1 p2) c1 c2)
cl1 cl2
with ex when precatchable_exception ex ->
- reduce curenvnb pb b wt substn cM cN)
+ 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) ->
- (match
- is_unification_pattern curenvnb sigma f1 (Array.to_list l1) cN
- with
- | None ->
- (match kind_of_term cN with
- | App (f2,l2) -> unify_app curenvnb pb b substn cM f1 l1 cN f2 l2
- | _ -> unify_not_same_head curenvnb pb b wt substn cM cN)
- | Some l ->
- solve_pattern_eqn_array curenvnb f1 l cN substn)
+ 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) ->
- (match
- is_unification_pattern curenvnb sigma f2 (Array.to_list l2) cM
- with
- | None ->
- (match kind_of_term cM with
- | App (f1,l1) -> unify_app curenvnb pb b substn cM f1 l1 cN f2 l2
- | _ -> unify_not_same_head curenvnb pb b wt substn cM cN)
- | Some l ->
- solve_pattern_eqn_array curenvnb f2 l cM substn)
+ unify_app_pattern false curenvnb pb opt substn cM cM [||] cN f2 l2
| App (f1,l1), App (f2,l2) ->
- unify_app curenvnb pb b substn cM f1 l1 cN f2 l2
+ unify_app curenvnb pb opt substn cM f1 l1 cN f2 l2
+
+ | App (f1,l1), Proj(p2,c2) ->
+ unify_app curenvnb pb opt substn cM f1 l1 cN cN [||]
- | _ ->
- unify_not_same_head curenvnb pb b wt substn cM cN
+ | Proj (p1,c1), App(f2,l2) ->
+ unify_app curenvnb pb opt substn cM cM [||] cN f2 l2
- and unify_app curenvnb pb b substn cM f1 l1 cN f2 l2 =
+ | _ ->
+ 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 =
+ 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
+ | 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
+ | Proj _ -> unify_app curenvnb pb opt substn cM f1 l1 cN f2 l2
+ | _ -> unify_not_same_head curenvnb pb opt substn cM cN)
+ | 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 =
try
+ let needs_expansion p c' =
+ match kind_of_term 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
+ | 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))
+ with RetypeError _ -> (** Unification can be called on ill-typed terms, due
+ to FO and eta in particular, fail gracefully in that case *)
+ (c, l))
+ | _ -> (c, l)
+ in
+ let f1, l1 = expand_proj f1 f2 l1 in
+ let f2, l2 = expand_proj f2 f1 l2 in
+ let opta = {opt with at_top = true; with_types = false} in
+ let optf = {opt with at_top = true; with_types = true} in
let (f1,l1,f2,l2) = adjust_app_array_size f1 l1 f2 l2 in
- array_fold_left2 (unirec_rec curenvnb CONV true false)
- (unirec_rec curenvnb CONV true true substn f1 f2) l1 l2
- with ex when precatchable_exception ex ->
- try reduce curenvnb pb b false substn cM cN
+ if Array.length l1 == 0 then error_cannot_unify (fst curenvnb) sigma (cM,cN)
+ else
+ Array.fold_left2 (unirec_rec curenvnb CONV opta)
+ (unirec_rec curenvnb CONV optf substn f1 f2) l1 l2
with ex when precatchable_exception ex ->
- try expand curenvnb pb b false substn cM f1 l1 cN f2 l2
+ try reduce curenvnb pb {opt with with_types = false} substn cM cN
with ex when precatchable_exception ex ->
- canonical_projections curenvnb pb b cM cN substn
-
- and unify_not_same_head curenvnb pb b wt substn cM cN =
- try canonical_projections curenvnb pb b cM cN substn
+ try canonical_projections curenvnb pb opt cM cN substn
with ex when precatchable_exception ex ->
- if constr_cmp cv_pb cM cN then substn else
- try reduce curenvnb pb b wt substn cM cN
+ expand curenvnb pb {opt with with_types = false} substn cM f1 l1 cN f2 l2
+
+ and unify_same_proj (curenv, nb as curenvnb) cv_pb opt substn c1 c2 =
+ let substn = unirec_rec curenvnb CONV opt substn c1 c2 in
+ try (* Force unification of the types to fill in parameters *)
+ let ty1 = get_type_of curenv ~lax:true sigma c1 in
+ let ty2 = get_type_of curenv ~lax:true sigma c2 in
+ unify_0_with_initial_metas substn true curenv cv_pb
+ { flags with modulo_conv_on_closed_terms = Some full_transparent_state;
+ modulo_delta = full_transparent_state;
+ modulo_eta = true;
+ modulo_betaiota = true }
+ ty1 ty2
+ with RetypeError _ -> substn
+
+ and unify_not_same_head curenvnb pb opt (sigma, metas, evars as substn) cM cN =
+ try canonical_projections curenvnb pb opt cM cN substn
with ex when precatchable_exception ex ->
- let (f1,l1) =
- match kind_of_term 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
- expand curenvnb pb b wt substn cM f1 l1 cN f2 l2
-
- and reduce curenvnb pb b wt (sigma, metas, evars as substn) cM cN =
- if use_full_betaiota flags && not (subterm_restriction b flags) then
+ let sigma', b = constr_cmp cv_pb 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
+ let (f2,l2) =
+ match kind_of_term 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
let cM' = do_reduce flags.modulo_delta curenvnb sigma cM in
- if not (eq_constr cM cM') then
- unirec_rec curenvnb pb b wt substn cM' cN
+ if not (Term.eq_constr 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 (eq_constr cN cN') then
- unirec_rec curenvnb pb b wt substn cM cN'
+ if not (Term.eq_constr 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)
+ else error_cannot_unify (fst curenvnb) sigma (cM,cN)
- and expand (curenv,_ as curenvnb) pb b wt (sigma,metasubst,_ as substn) cM f1 l1 cN f2 l2 =
-
- if
+ and expand (curenv,_ as curenvnb) pb opt (sigma,metasubst,evarsubst as substn) 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
heuristic was to apply conversion on meta-free (but not
@@ -549,117 +885,144 @@ let unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb flag
(it is used by apply and rewrite); it might now be redundant
with the support for delta-expansion (which is used
essentially for apply)... *)
- not (subterm_restriction b flags) &&
+ if subterm_restriction opt flags then None else
match flags.modulo_conv_on_closed_terms with
- | None -> false
+ | None -> None
| Some convflags ->
- let subst = if flags.use_metas_eagerly_in_conv_on_closed_terms then metasubst else ms in
- match subst_defined_metas subst cM with
- | None -> (* some undefined Metas in cM *) false
+ 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
+ | None -> (* some undefined Metas in cM *) None
| Some m1 ->
- match subst_defined_metas subst cN with
- | None -> (* some undefined Metas in cN *) false
+ match subst_defined_metas_evars subst cN with
+ | None -> (* some undefined Metas in cN *) None
| Some n1 ->
(* No subterm restriction there, too much incompatibilities *)
- if is_trans_fconv pb convflags env sigma m1 n1
- then true else
- if is_ground_term sigma m1 && is_ground_term sigma n1 then
- error_cannot_unify curenv sigma (cM,cN)
- else false
- then
- substn
- else
- let cf1 = key_of b flags f1 and cf2 = key_of b flags f2 in
+ let sigma, b = infer_conv ~pb ~ts:convflags curenv sigma m1 n1 in
+ if b then Some (sigma, metasubst, evarsubst)
+ else
+ if is_ground_term sigma m1 && is_ground_term sigma n1 then
+ error_cannot_unify curenv sigma (cM,cN)
+ else None
+ in
+ match res with
+ | Some substn -> substn
+ | None ->
+ let cf1 = key_of curenv opt flags f1 and cf2 = key_of curenv opt flags f2 in
match oracle_order curenv cf1 cf2 with
| None -> error_cannot_unify curenv sigma (cM,cN)
| Some true ->
- (match expand_key curenv cf1 with
+ (match expand_key flags.modulo_delta curenv sigma cf1 with
| Some c ->
- unirec_rec curenvnb pb b wt substn
+ unirec_rec curenvnb pb opt substn
(whd_betaiotazeta sigma (mkApp(c,l1))) cN
| None ->
- (match expand_key curenv cf2 with
+ (match expand_key flags.modulo_delta curenv sigma cf2 with
| Some c ->
- unirec_rec curenvnb pb b wt substn cM
+ unirec_rec curenvnb pb opt substn cM
(whd_betaiotazeta sigma (mkApp(c,l2)))
| None ->
error_cannot_unify curenv sigma (cM,cN)))
| Some false ->
- (match expand_key curenv cf2 with
+ (match expand_key flags.modulo_delta curenv sigma cf2 with
| Some c ->
- unirec_rec curenvnb pb b wt substn cM
+ unirec_rec curenvnb pb opt substn cM
(whd_betaiotazeta sigma (mkApp(c,l2)))
| None ->
- (match expand_key curenv cf1 with
+ (match expand_key flags.modulo_delta curenv sigma cf1 with
| Some c ->
- unirec_rec curenvnb pb b wt substn
+ unirec_rec curenvnb pb opt substn
(whd_betaiotazeta sigma (mkApp(c,l1))) cN
| None ->
error_cannot_unify curenv sigma (cM,cN)))
- and canonical_projections curenvnb pb b cM cN (sigma,_,_ as substn) =
+ and canonical_projections (curenv, _ as curenvnb) pb opt cM cN (sigma,_,_ as substn) =
let f1 () =
if isApp cM then
- let f1l1 = decompose_app cM in
- if is_open_canonical_projection env sigma f1l1 then
- let f2l2 = decompose_app cN in
- solve_canonical_projection curenvnb pb b cM f1l1 cN f2l2 substn
+ 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
+ solve_canonical_projection curenvnb pb opt cM f1l1 cN f2l2 substn
else error_cannot_unify (fst curenvnb) sigma (cM,cN)
else error_cannot_unify (fst curenvnb) sigma (cM,cN)
in
- if flags.modulo_conv_on_closed_terms = None ||
- subterm_restriction b flags then
+ if not opt.with_cs ||
+ begin match flags.modulo_conv_on_closed_terms with
+ | None -> true
+ | Some _ -> subterm_restriction opt flags
+ end then
error_cannot_unify (fst curenvnb) sigma (cM,cN)
else
try f1 () with e when precatchable_exception e ->
if isApp cN then
- let f2l2 = decompose_app cN in
- if is_open_canonical_projection env sigma f2l2 then
- let f1l1 = decompose_app cM in
- solve_canonical_projection curenvnb pb b cN f2l2 cM f1l1 substn
+ 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
+ solve_canonical_projection curenvnb pb opt cN f2l2 cM f1l1 substn
else error_cannot_unify (fst curenvnb) sigma (cM,cN)
else error_cannot_unify (fst curenvnb) sigma (cM,cN)
- and solve_canonical_projection curenvnb pb b cM f1l1 cN f2l2 (sigma,ms,es) =
- let (c,bs,(params,params1),(us,us2),(ts,ts1),c1,(n,t2)) =
- try Evarconv.check_conv_record f1l1 f2l2
+ and solve_canonical_projection curenvnb pb opt cM f1l1 cN f2l2 (sigma,ms,es) =
+ let (ctx,t,c,bs,(params,params1),(us,us2),(ts,ts1),c1,(n,t2)) =
+ try Evarconv.check_conv_record (fst curenvnb) sigma f1l1 f2l2
with Not_found -> error_cannot_unify (fst curenvnb) sigma (cM,cN)
in
- let (evd,ks,_) =
- List.fold_left
- (fun (evd,ks,m) b ->
- if m=n then (evd,t2::ks, m-1) else
- let mv = new_meta () in
- let evd' = meta_declare mv (substl ks b) evd in
+ if Reductionops.Stack.compare_shape ts ts1 then
+ let sigma = Evd.merge_context_set Evd.univ_flexible sigma ctx in
+ let (evd,ks,_) =
+ List.fold_left
+ (fun (evd,ks,m) b ->
+ if match n with Some n -> Int.equal m n | None -> false then
+ (evd,t2::ks, m-1)
+ else
+ let mv = new_meta () in
+ let evd' = meta_declare mv (substl ks b) evd in
(evd', mkMeta mv :: ks, m - 1))
- (sigma,[],List.length bs - 1) bs
- in
- let unilist2 f substn l l' =
- try List.fold_left2 f substn l l'
- with Invalid_argument "List.fold_left2" -> error_cannot_unify (fst curenvnb) sigma (cM,cN)
- in
- let substn = unilist2 (fun s u1 u -> unirec_rec curenvnb pb b false s u1 (substl ks u))
- (evd,ms,es) us2 us in
- let substn = unilist2 (fun s u1 u -> unirec_rec curenvnb pb b false s u1 (substl ks u))
- substn params1 params in
- let substn = unilist2 (unirec_rec curenvnb pb b false) substn ts ts1 in
- unirec_rec curenvnb pb b false substn c1 (applist (c,(List.rev ks)))
-
+ (sigma,[],List.length bs) bs
+ in
+ try
+ let opt' = {opt with with_types = false} in
+ 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
+ (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 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" ->
+ error_cannot_unify (fst curenvnb) sigma (cM,cN)
+ else error_cannot_unify (fst curenvnb) sigma (cM,cN)
in
- let evd = sigma in
- if (if occur_meta_or_undefined_evar evd m || occur_meta_or_undefined_evar evd n
- || subterm_restriction conv_at_top flags then false
- else if (match flags.modulo_conv_on_closed_terms with
- | Some convflags -> is_trans_fconv cv_pb convflags env sigma m n
- | _ -> constr_cmp cv_pb m n) then true
- else if (match flags.modulo_conv_on_closed_terms, flags.modulo_delta with
- | Some (cv_id, cv_k), (dl_id, dl_k) ->
- Idpred.subset dl_id cv_id && Cpred.subset dl_k cv_k
- | None,(dl_id, dl_k) ->
- Idpred.is_empty dl_id && Cpred.is_empty dl_k)
- then error_cannot_unify env sigma (m, n) else false)
- then subst
- else unirec_rec (env,0) cv_pb conv_at_top false subst m n
+
+ if !debug_unification then msg_debug (str "Starting unification");
+ let opt = { at_top = conv_at_top; with_types = false; with_cs = true } in
+ try
+ let res =
+ if occur_meta_or_undefined_evar sigma m || occur_meta_or_undefined_evar sigma n
+ || subterm_restriction opt flags then None
+ 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
+ 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) ->
+ Id.Pred.subset dl_id cv_id && Cpred.subset dl_k cv_k
+ | None,(dl_id, dl_k) ->
+ Id.Pred.is_empty dl_id && Cpred.is_empty dl_k)
+ then error_cannot_unify env sigma (m, n) else None
+ in
+ let a = match res with
+ | Some sigma -> sigma, ms, es
+ | None -> unirec_rec (env,0) cv_pb opt subst m n in
+ if !debug_unification then msg_debug (str "Leaving unification with success");
+ a
+ with e ->
+ if !debug_unification then msg_debug (str "Leaving unification with failure");
+ raise e
+
let unify_0 env sigma = unify_0_with_initial_metas (sigma,[],[]) true env
@@ -704,14 +1067,14 @@ let merge_instances env sigma flags st1 st2 c1 c2 =
| ((IsSubType | Conv as oppst1),
(IsSubType | Conv)) ->
let res = unify_0 env sigma CUMUL flags c2 c1 in
- if oppst1=st2 then (* arbitrary choice *) (left, st1, res)
- else if st2=IsSubType then (left, st1, res)
+ if eq_instance_constraint oppst1 st2 then (* arbitrary choice *) (left, st1, res)
+ else if eq_instance_constraint st2 IsSubType then (left, st1, res)
else (right, st2, res)
| ((IsSuperType | Conv as oppst1),
(IsSuperType | Conv)) ->
let res = unify_0 env sigma CUMUL flags c1 c2 in
- if oppst1=st2 then (* arbitrary choice *) (left, st1, res)
- else if st2=IsSuperType then (left, st1, res)
+ if eq_instance_constraint oppst1 st2 then (* arbitrary choice *) (left, st1, res)
+ else if eq_instance_constraint st2 IsSuperType then (left, st1, res)
else (right, st2, res)
| (IsSuperType,IsSubType) ->
(try (left, IsSubType, unify_0 env sigma CUMUL flags c2 c1)
@@ -773,13 +1136,13 @@ let merge_instances env sigma flags st1 st2 c1 c2 =
let applyHead env evd n c =
let rec apprec n c cty evd =
- if n = 0 then
+ if Int.equal n 0 then
(evd, c)
else
match kind_of_term (whd_betadeltaiota env evd cty) with
| Prod (_,c1,c2) ->
let (evd',evar) =
- Evarutil.new_evar evd env ~src:(dummy_loc,GoalEvar) c1 in
+ Evarutil.new_evar env evd ~src:(Loc.ghost,Evar_kinds.GoalEvar) c1 in
apprec (n-1) (mkApp(c,[|evar|])) (subst1 evar c2) evd'
| _ -> error "Apply_Head_Then"
in
@@ -787,21 +1150,20 @@ let applyHead env evd n c =
let is_mimick_head ts f =
match kind_of_term f with
- | Const c -> not (Closure.is_transparent_constant ts c)
+ | Const (c,u) -> not (Closure.is_transparent_constant ts c)
| Var id -> not (Closure.is_transparent_variable ts id)
| (Rel _|Construct _|Ind _) -> true
| _ -> false
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 dummy_loc env evd j tycon in
+ let (evd',j') = inh_conv_coerce_rigid_to true Loc.ghost env evd j tycon in
let evd' = Evarconv.consider_remaining_unif_problems env evd' in
let evd' = Evd.map_metas_fvalue (nf_evar evd') evd' in
(evd',j'.uj_val)
let w_coerce_to_type env evd c cty mvty =
- let evd,mvty = pose_all_metas_as_evars env evd mvty in
- let tycon = mk_tycon_type mvty in
+ let evd,tycon = pose_all_metas_as_evars env evd mvty in
try try_to_coerce env evd c cty tycon
with e when precatchable_exception e ->
(* inh_conv_coerce_rigid_to should have reasoned modulo reduction
@@ -816,8 +1178,8 @@ let w_coerce env evd mv c =
w_coerce_to_type env evd c cty mvty
let unify_to_type env sigma flags c status u =
- let c = refresh_universes c in
- let t = get_type_of env sigma c in
+ 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
unify_0 env sigma CUMUL flags t u
@@ -825,9 +1187,7 @@ let unify_type env sigma flags mv status c =
let mvty = Typing.meta_type sigma mv in
let mvty = nf_meta sigma mvty in
unify_to_type env sigma
- {flags with modulo_delta = flags.modulo_delta_types;
- modulo_conv_on_closed_terms = Some flags.modulo_delta_types;
- modulo_betaiota = true}
+ (set_flags_for_type flags)
c status mvty
(* Move metas that may need coercion at the end of the list of instances *)
@@ -835,17 +1195,20 @@ let unify_type env sigma flags mv status c =
let order_metas metas =
let rec order latemetas = function
| [] -> List.rev latemetas
- | (_,_,(status,to_type) as meta)::metas ->
- if to_type = CoerceToType then order (meta::latemetas) metas
- else meta :: order latemetas metas
+ | (_,_,(_,CoerceToType) as meta)::metas ->
+ order (meta::latemetas) metas
+ | (_,_,(_,_) as meta)::metas ->
+ meta :: order latemetas metas
in order [] metas
(* Solve an equation ?n[x1=u1..xn=un] = t where ?n is an evar *)
let solve_simple_evar_eqn ts env evd ev rhs =
- let evd,b = solve_simple_eqn (Evarconv.evar_conv_x ts) env evd (None,ev,rhs) in
- if not b then error_cannot_unify env evd (mkEvar ev,rhs);
- Evarconv.consider_remaining_unif_problems env evd
+ 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 ->
+ Evarconv.consider_remaining_unif_problems env evd
(* [w_merge env sigma b metas evars] merges common instances in metas
or in evars, possibly generating new unification problems; if [b]
@@ -860,7 +1223,7 @@ let w_merge env with_types flags (evd,metas,evars) =
if Evd.is_defined evd evk then
let v = Evd.existential_value evd ev in
let (evd,metas',evars'') =
- unify_0 curenv evd CONV (set_merge_flags flags) rhs v in
+ 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 *)
@@ -872,16 +1235,22 @@ let w_merge env with_types flags (evd,metas,evars) =
if is_mimick_head 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
+ (* let evd' = Evarconv.consider_remaining_unif_problems env evd' in *)
+ w_merge_rec evd' metas evars eqns
else
- let evd', rhs'' = pose_all_metas_as_evars curenv evd rhs' in
- w_merge_rec (solve_simple_evar_eqn flags.modulo_delta_types curenv evd' ev rhs'')
- metas evars' eqns
-
+ let evd' =
+ let evd', rhs'' = pose_all_metas_as_evars curenv evd rhs' in
+ try solve_simple_evar_eqn flags.modulo_delta_types curenv evd' ev rhs''
+ with Retyping.RetypeError _ ->
+ error_cannot_unify curenv evd' (mkEvar ev,rhs'')
+ in w_merge_rec evd' metas evars' eqns
| _ ->
let evd', rhs'' = pose_all_metas_as_evars curenv evd rhs' in
- w_merge_rec (solve_simple_evar_eqn flags.modulo_delta_types curenv evd' ev rhs'')
- metas evars' eqns
+ let evd' =
+ try solve_simple_evar_eqn flags.modulo_delta_types curenv evd' ev rhs''
+ with Retyping.RetypeError _ -> error_cannot_unify curenv evd' (mkEvar ev, rhs'')
+ in
+ w_merge_rec evd' metas evars' eqns
end
| [] ->
@@ -889,13 +1258,15 @@ let w_merge env with_types flags (evd,metas,evars) =
match metas with
| (mv,c,(status,to_type))::metas ->
let ((evd,c),(metas'',evars'')),eqns =
- if with_types & to_type <> TypeProcessed then
- if to_type = CoerceToType then
+ if with_types && to_type != TypeProcessed then
+ begin match to_type with
+ | CoerceToType ->
(* Some coercion may have to be inserted *)
(w_coerce env evd mv c,([],[])),eqns
- else
+ | _ ->
(* No coercion needed: delay the unification of types *)
((evd,c),([],[])),(mv,status,c)::eqns
+ end
else
((evd,c),([],[])),eqns
in
@@ -938,19 +1309,30 @@ let w_merge env with_types flags (evd,metas,evars) =
let sp_env = Global.env_of_context ev.evar_hyps in
let (evd', c) = applyHead sp_env evd nargs hdc in
let (evd'',mc,ec) =
- unify_0 sp_env evd' CUMUL (set_merge_flags flags)
+ unify_0 sp_env evd' CUMUL flags
(get_type_of sp_env evd' c) 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
- (* merge constraints *)
- w_merge_rec evd (order_metas metas) evars []
+ 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
+ | _ -> acc) [] metas
+ in w_merge_rec evd [] [] eqns
+ in
+ let res = (* merge constraints *)
+ w_merge_rec evd (order_metas metas) (List.rev evars) []
+ in
+ if with_types then check_types res
+ else res
-let w_unify_meta_types env ?(flags=default_unify_flags) 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 (evd,metas,[])
+ w_merge env true flags.merge_unify_flags (evd,metas,[])
(* [w_unify env evd M N]
performs a unification of M and N, generating a bunch of
@@ -962,49 +1344,49 @@ let w_unify_meta_types env ?(flags=default_unify_flags) evd =
[clenv_typed_unify M N clenv] expects in addition that expected
types of metavars are unifiable with the types of their instances *)
+let head_app sigma m =
+ fst (whd_nored_state sigma (m, Stack.empty))
+
let check_types env flags (sigma,_,_ as subst) m n =
- if isEvar_or_Meta (fst (whd_stack sigma m)) then
+ if isEvar_or_Meta (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 (fst (whd_stack sigma n)) then
+ else if isEvar_or_Meta (head_app sigma n) then
unify_0_with_initial_metas subst true env CUMUL
flags
(get_type_of env sigma m)
(get_type_of env sigma n)
else subst
-let try_resolve_typeclasses env evd flags m n =
- if flags.resolve_evars then
- try Typeclasses.resolve_typeclasses ~filter:Typeclasses.no_goals ~split:false
- ~fail:true env evd
- with e when Typeclasses_errors.unsatisfiable_exception e ->
- error_cannot_unify env evd (m, n)
+let try_resolve_typeclasses env evd flag m n =
+ if flag then
+ Typeclasses.resolve_typeclasses ~filter:Typeclasses.no_goals ~split:false
+ ~fail:true env evd
else evd
let w_unify_core_0 env evd with_types cv_pb flags m n =
let (mc1,evd') = retract_coercible_metas evd in
- let (sigma,ms,es) = check_types env flags (evd,mc1,[]) m n in
+ let (sigma,ms,es) = check_types env (set_flags_for_type flags.core_unify_flags) (evd',mc1,[]) m n in
let subst2 =
- unify_0_with_initial_metas (evd',ms,es) false env cv_pb flags m n
+ unify_0_with_initial_metas (sigma,ms,es) false env cv_pb
+ flags.core_unify_flags m n
in
- let evd = w_merge env with_types flags subst2 in
- try_resolve_typeclasses env evd flags m n
+ let evd = w_merge env with_types flags.merge_unify_flags subst2 in
+ try_resolve_typeclasses env evd flags.resolve_evars m n
-let w_unify_0 env evd = w_unify_core_0 env evd false
let w_typed_unify env evd = w_unify_core_0 env evd true
-let w_typed_unify_list env evd flags f1 l1 f2 l2 =
- let flags' = { flags with resolve_evars = false } in
- let f1,l1,f2,l2 = adjust_app_list_size f1 l1 f2 l2 in
+let w_typed_unify_array env evd flags f1 l1 f2 l2 =
+ let f1,l1,f2,l2 = adjust_app_array_size f1 l1 f2 l2 in
let (mc1,evd') = retract_coercible_metas evd in
- let subst =
- List.fold_left2 (fun subst m n ->
- unify_0_with_initial_metas subst true env CONV flags' m n) (evd',[],[])
- (f1::l1) (f2::l2) in
- let evd = w_merge env true flags subst in
- try_resolve_typeclasses env evd flags (applist(f1,l1)) (applist(f2,l2))
+ let fold_subst subst m n = unify_0_with_initial_metas subst true env CONV flags.core_unify_flags m n in
+ let subst = fold_subst (evd', [], []) f1 f2 in
+ let subst = Array.fold_left2 fold_subst subst l1 l2 in
+ let evd = w_merge env true flags.merge_unify_flags subst in
+ try_resolve_typeclasses env evd flags.resolve_evars
+ (mkApp(f1,l1)) (mkApp(f2,l2))
(* takes a substitution s, an open term op and a closed term cl
try to find a subterm of cl which matches op, if op is just a Meta
@@ -1013,21 +1395,242 @@ let w_typed_unify_list env evd flags f1 l1 f2 l2 =
let iter_fail f a =
let n = Array.length a in
let rec ffail i =
- if i = n then error "iter_fail"
+ if Int.equal i n then error "iter_fail"
else
try f a.(i)
with ex when precatchable_exception ex -> ffail (i+1)
in ffail 0
+(* 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) &&
+ (* 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 (id,_,_) -> dependent_in_decl (mkVar id) d) decls
+
+let indirect_dependency d decls =
+ pi1 (List.hd (List.filter (fun (id,_,_) -> dependent_in_decl (mkVar id) d) decls))
+
+let finish_evar_resolution ?(flags=Pretyping.all_and_fail_flags) env current_sigma (pending,c) =
+ let sigma = Pretyping.solve_remaining_evars flags env current_sigma pending in
+ let sigma, subst = nf_univ_variables sigma in
+ sigma, subst_univs_constr subst (nf_evar sigma c)
+
+let default_matching_core_flags sigma =
+ let ts = Names.full_transparent_state in {
+ modulo_conv_on_closed_terms = Some empty_transparent_state;
+ use_metas_eagerly_in_conv_on_closed_terms = false;
+ use_evars_eagerly_in_conv_on_closed_terms = false;
+ modulo_delta = empty_transparent_state;
+ modulo_delta_types = ts;
+ check_applied_meta_types = true;
+ use_pattern_unification = false;
+ use_meta_bound_pattern_unification = false;
+ frozen_evars =
+ fold_undefined (fun evk _ evars -> Evar.Set.add evk evars)
+ sigma Evar.Set.empty;
+ restrict_conv_on_strict_subterms = false;
+ modulo_betaiota = false;
+ modulo_eta = false;
+}
+
+let default_matching_merge_flags sigma =
+ let ts = Names.full_transparent_state in
+ let flags = default_matching_core_flags sigma in {
+ flags with
+ modulo_conv_on_closed_terms = Some ts;
+ modulo_delta = ts;
+ modulo_betaiota = true;
+ modulo_eta = true;
+ use_pattern_unification = true;
+}
+
+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;
+ subterm_unify_flags = flags; (* does not matter *)
+ resolve_evars = false;
+ allow_K_in_toplevel_higher_order_unification = false;
+}
+
+(* This supports search of occurrences of term from a pattern *)
+(* from_prefix is useful e.g. for subterms in an inductive type: we can say *)
+(* "destruct t" and it finds "t u" *)
+
+exception PatternNotFound
+
+let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) =
+ let flags =
+ if from_prefix_of_ind then
+ let flags = default_matching_flags pending in
+ { flags with core_unify_flags = { flags.core_unify_flags with
+ 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 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 l1,l2 =
+ try List.chop n l with Failure _ -> raise (NotUnifiable None) in
+ if not (List.for_all closed0 l2) then raise (NotUnifiable None)
+ else
+ applist (t,l1), l2
+ else t, [] in
+ let sigma = w_typed_unify env sigma Reduction.CONV flags c t' in
+ let ty = Retyping.get_type_of env sigma t in
+ if not (is_correct_type ty) then raise (NotUnifiable None);
+ Some(sigma, t, l2)
+ with
+ | PretypeError (_,_,CannotUnify (c1,c2,Some e)) ->
+ raise (NotUnifiable (Some (c1,c2,e)))
+ (** MS: This is pretty bad, it catches Not_found for example *)
+ | e when Errors.noncritical e -> raise (NotUnifiable None) in
+ let merge_fun c1 c2 =
+ match c1, c2 with
+ | Some (evd,c1,_) as x, Some (_,c2,_) ->
+ if is_conv env sigma c1 c2 then x else raise (NotUnifiable None)
+ | Some _, None -> c1
+ | None, Some _ -> c2
+ | None, None -> None in
+ { match_fun = matching_fun; merge_fun = merge_fun;
+ testing_state = None; last_found = None },
+ (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 univs, subst = nf_univ_variables sigma in
+ Some (sigma,subst_univs_constr subst c))
+
+let make_eq_test env evd c =
+ let out cstr =
+ match cstr.last_found with None -> None | _ -> Some (cstr.testing_state, c)
+ in
+ (make_eq_univs_test env evd c, out)
+
+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
+ if name == Anonymous then next_ident_away_in_goal x ids else
+ if mem_named_context x (named_context env) then
+ error ("The variable "^(Id.to_string x)^" is already declared.")
+ else
+ x
+ in
+ let likefirst = clause_with_generic_occurrences occs in
+ let mkvarid () = mkVar id in
+ let compute_dependency _ (hyp,_,_ as d) (sign,depdecls) =
+ match occurrences_of_hyp hyp occs with
+ | NoOccurrences, InHyp ->
+ if indirectly_dependent c d depdecls then
+ (* Told explicitly not to abstract over [d], but it is dependent *)
+ let id' = indirect_dependency d depdecls in
+ errorlabstrm "" (str "Cannot abstract over " ++ Nameops.pr_id id'
+ ++ str " without also abstracting or erasing " ++ Nameops.pr_id hyp
+ ++ str ".")
+ else
+ (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.eq_named_declaration d newdecl
+ && not (indirectly_dependent c d depdecls)
+ then
+ if check_occs && not (in_every_hyp occs)
+ then raise (PretypeError (env,sigma,NoOccurrenceFound (c,Some hyp)))
+ else (push_named_context_val d sign, depdecls)
+ else
+ (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
+ (push_named_context_val newdecl sign, newdecl :: depdecls) in
+ try
+ let sign,depdecls =
+ fold_named_context compute_dependency env
+ ~init:(empty_named_context_val,[]) in
+ let ccl = match occurrences_of_goal occs with
+ | NoOccurrences -> concl
+ | occ ->
+ let occ = if likefirst then LikeFirst else AtOccs occ in
+ replace_term_occ_modulo occ test mkvarid concl
+ in
+ let lastlhyp =
+ if List.is_empty depdecls then None else Some (pi1(List.last depdecls)) in
+ (id,sign,depdecls,lastlhyp,ccl,out test)
+ with
+ SubtermUnificationError e ->
+ raise (PretypeError (env,sigma,CannotUnifyOccurrences e))
+
+(** [make_abstraction] is the main entry point to abstract over a term
+ or pattern at some occurrences; it returns:
+ - the id used for the abstraction
+ - the type of the abstraction
+ - the declarations from the context which depend on the term or pattern
+ - the most recent hyp before which there is no dependency in the term of pattern
+ - the abstracted conclusion
+ - an evar universe context effect to apply on the goal
+ - the term or pattern to abstract fully instantiated
+*)
+
+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
+| AbstractExact of Name.t * constr * types option * clause * bool
+
+type abstraction_result =
+ Names.Id.t * named_context_val *
+ Context.named_declaration list * Names.Id.t option *
+ types * (Evd.evar_map * constr) option
+
+let make_abstraction env evd ccl abs =
+ match abs with
+ | AbstractPattern (from_prefix,check,name,c,occs,check_occs) ->
+ make_abstraction_core name
+ (make_pattern_test from_prefix check env evd c)
+ env evd (snd c) None occs check_occs ccl
+ | AbstractExact (name,c,ty,occs,check_occs) ->
+ make_abstraction_core name
+ (make_eq_test env evd c)
+ env evd c ty occs check_occs ccl
+
+let keyed_unify env evd kop =
+ if not !keyed_unification then fun cl -> true
+ else
+ match kop with
+ | None -> fun _ -> true
+ | Some kop ->
+ fun cl ->
+ let kc = Keys.constr_key cl in
+ match kc with
+ | None -> false
+ | Some kc -> Keys.equiv_keys kop kc
+
(* Tries to find an instance of term [cl] in term [op].
Unifies [cl] to every subterm of [op] until it finds a match.
Fails if no match is found *)
-let w_unify_to_subterm env evd ?(flags=default_unify_flags) (op,cl) =
+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 rec matchrec cl =
let cl = strip_outer_cast cl in
(try
- if closed0 cl && not (isEvar cl)
- then w_typed_unify env evd CONV flags op cl,cl
+ if closed0 cl && not (isEvar cl) && keyed_unify env evd kop cl then
+ (try 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"
with ex when precatchable_exception ex ->
(match kind_of_term cl with
@@ -1051,6 +1654,8 @@ let w_unify_to_subterm env evd ?(flags=default_unify_flags) (op,cl) =
with ex when precatchable_exception ex ->
matchrec c2)
+ | Proj (p,c) -> matchrec c
+
| Fix(_,(_,types,terms)) ->
(try
iter_fail matchrec types
@@ -1077,15 +1682,17 @@ let w_unify_to_subterm env evd ?(flags=default_unify_flags) (op,cl) =
in
try matchrec cl
with ex when precatchable_exception ex ->
- raise (PretypeError (env,evd,NoOccurrenceFound (op, None)))
+ match !bestexn with
+ | None -> raise (PretypeError (env,evd,NoOccurrenceFound (op, None)))
+ | Some e -> raise e
(* Tries to find all instances of term [cl] in term [op].
Unifies [cl] to every subterm of [op] and return all the matches.
Fails if no match is found *)
-let w_unify_to_subterm_all 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') -> eq_constr c c') b then b else a :: b
+ if List.exists (fun (evd',c') -> Term.eq_constr c c') b then b else a :: b
in
let fail str _ = error str in
let bind f g a =
@@ -1099,7 +1706,7 @@ let w_unify_to_subterm_all env evd ?(flags=default_unify_flags) (op,cl) =
let bind_iter f a =
let n = Array.length a in
let rec ffail i =
- if i = n then fun a -> a
+ if Int.equal i n then fun a -> a
else bind (f a.(i)) (ffail (i+1))
in ffail 0
in
@@ -1120,6 +1727,8 @@ let w_unify_to_subterm_all env evd ?(flags=default_unify_flags) (op,cl) =
| Case(_,_,c,lf) -> (* does not search in the predicate *)
bind (matchrec c) (bind_iter matchrec lf)
+ | Proj (p,c) -> matchrec c
+
| LetIn(_,c1,_,c2) ->
bind (matchrec c1) (matchrec c2)
@@ -1138,10 +1747,10 @@ let w_unify_to_subterm_all env evd ?(flags=default_unify_flags) (op,cl) =
| _ -> fail "Match_subterm"))
in
let res = matchrec cl [] in
- if res = [] then
+ match res with
+ | [] ->
raise (PretypeError (env,evd,NoOccurrenceFound (op, None)))
- else
- res
+ | _ -> res
let w_unify_to_subterm_list env evd flags hdmeta oplist t =
List.fold_right
@@ -1150,47 +1759,64 @@ let w_unify_to_subterm_list env evd flags hdmeta oplist t =
if isMeta 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 if occur_meta_or_existential op then
+ else
+ let allow_K = flags.allow_K_in_toplevel_higher_order_unification in
+ let flags =
+ if occur_meta_or_existential op || !keyed_unification then
+ flags
+ else
+ (* up to Nov 2014, unification was bypassed on evar/meta-free terms;
+ now it is called in a minimalistic way, at least to possibly
+ unify pre-existing non frozen evars of the goal or of the
+ pattern *)
+ set_no_delta_flags flags in
let (evd',cl) =
try
(* This is up to delta for subterms w/o metas ... *)
w_unify_to_subterm env evd ~flags (strip_outer_cast op,t)
- with PretypeError (env,_,NoOccurrenceFound _) when
- flags.allow_K_in_toplevel_higher_order_unification -> (evd,op)
+ with PretypeError (env,_,NoOccurrenceFound _) when allow_K -> (evd,op)
in
- if not flags.allow_K_in_toplevel_higher_order_unification &&
+ if not allow_K &&
(* ensure we found a different instance *)
- List.exists (fun op -> eq_constr op cl) l
+ List.exists (fun op -> Term.eq_constr op cl) l
then error_non_linear_unification env evd hdmeta cl
- else (evd',cl::l)
- else if flags.allow_K_in_toplevel_higher_order_unification or dependent op t
- then
- (evd,op::l)
- else
- (* This is not up to delta ... *)
- raise (PretypeError (env,evd,NoOccurrenceFound (op, None))))
+ else (evd',cl::l))
oplist
(evd,[])
let secondOrderAbstraction env evd flags typ (p, oplist) =
(* Remove delta when looking for a subterm *)
- let flags = { flags with modulo_delta = (fst flags.modulo_delta, Cpred.empty) } in
+ let flags = { flags with core_unify_flags = flags.subterm_unify_flags } in
let (evd',cllist) = w_unify_to_subterm_list env evd flags p oplist typ in
let typp = Typing.meta_type evd' p in
- let pred = abstract_list_all env evd' typp typ cllist in
- w_merge env false flags (evd',[p,pred,(Conv,TypeProcessed)],[])
+ let evd',(pred,predtyp) = abstract_list_all env evd' typp typ cllist in
+ let evd', b = infer_conv ~pb:CUMUL env evd' predtyp typp in
+ if not b then
+ error_wrong_abstraction_type env evd'
+ (Evd.meta_name evd p) pred typp predtyp;
+ w_merge env false flags.merge_unify_flags
+ (evd',[p,pred,(Conv,TypeProcessed)],[])
+
+ (* let evd',metas,evars = *)
+ (* try unify_0 env evd' CUMUL flags predtyp typp *)
+ (* with NotConvertible -> *)
+ (* error_wrong_abstraction_type env evd *)
+ (* (Evd.meta_name evd p) pred typp predtyp *)
+ (* in *)
+ (* w_merge env false flags (evd',(p,pred,(Conv,TypeProcessed))::metas,evars) *)
let secondOrderDependentAbstraction env evd flags typ (p, oplist) =
let typp = Typing.meta_type evd p in
- let pred = abstract_list_all_with_dependencies env evd typp typ oplist in
- w_merge env false flags (evd,[p,pred,(Conv,TypeProcessed)],[])
+ let evd, pred = abstract_list_all_with_dependencies env evd typp typ oplist in
+ w_merge env false flags.merge_unify_flags
+ (evd,[p,pred,(Conv,TypeProcessed)],[])
let secondOrderAbstractionAlgo dep =
if dep then secondOrderDependentAbstraction else secondOrderAbstraction
let w_unify2 env evd flags dep cv_pb ty1 ty2 =
- let c1, oplist1 = whd_stack evd ty1 in
- let c2, oplist2 = whd_stack evd ty2 in
+ 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
| Meta p1, _ ->
(* Find the predicate *)
@@ -1220,15 +1846,17 @@ let w_unify2 env evd flags dep cv_pb ty1 ty2 =
Before, second-order was used if the type of Meta(1) and [x:A]t was
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 = whd_stack evd ty1 in
- let hd2,l2 = whd_stack evd ty2 in
- match kind_of_term hd1, l1<>[], kind_of_term hd2, l2<>[] with
+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 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
(* Pattern case *)
| (Meta _, true, Lambda _, _ | Lambda _, _, Meta _, true)
- when List.length l1 = List.length l2 ->
+ when Int.equal (Array.length l1) (Array.length l2) ->
(try
- w_typed_unify_list env evd flags hd1 l1 hd2 l2
+ w_typed_unify_array env evd flags hd1 l1 hd2 l2
with ex when precatchable_exception ex ->
try
w_unify2 env evd flags false cv_pb ty1 ty2
@@ -1241,7 +1869,7 @@ let w_unify env evd cv_pb ?(flags=default_unify_flags) ty1 ty2 =
with PretypeError (env,_,NoOccurrenceFound _) as e -> raise e
| ex when precatchable_exception ex ->
try
- w_typed_unify_list env evd flags hd1 l1 hd2 l2
+ w_typed_unify_array env evd flags hd1 l1 hd2 l2
with ex' when precatchable_exception ex' ->
(* Last chance, use pattern-matching with typed
dependencies (done late for compatibility) *)
@@ -1252,3 +1880,17 @@ let w_unify env evd cv_pb ?(flags=default_unify_flags) ty1 ty2 =
(* General case: try first order *)
| _ -> w_typed_unify env evd cv_pb flags ty1 ty2
+
+(* Profiling *)
+
+let w_unify env evd cv_pb flags ty1 ty2 =
+ w_unify env evd cv_pb ~flags:flags ty1 ty2
+
+let w_unify =
+ if Flags.profile then
+ let wunifkey = Profile.declare_profile "w_unify" in
+ Profile.profile6 wunifkey w_unify
+ else w_unify
+
+let w_unify env evd cv_pb ?(flags=default_unify_flags ()) ty1 ty2 =
+ w_unify env evd cv_pb flags ty1 ty2
diff --git a/pretyping/unification.mli b/pretyping/unification.mli
index 75bddc1f..119b1a75 100644
--- a/pretyping/unification.mli
+++ b/pretyping/unification.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -10,34 +10,43 @@ open Term
open Environ
open Evd
-type unify_flags = {
+type core_unify_flags = {
modulo_conv_on_closed_terms : Names.transparent_state option;
use_metas_eagerly_in_conv_on_closed_terms : bool;
+ use_evars_eagerly_in_conv_on_closed_terms : bool;
modulo_delta : Names.transparent_state;
modulo_delta_types : Names.transparent_state;
- modulo_delta_in_merge : Names.transparent_state option;
check_applied_meta_types : bool;
- resolve_evars : bool;
use_pattern_unification : bool;
use_meta_bound_pattern_unification : bool;
- frozen_evars : ExistentialSet.t;
+ frozen_evars : Evar.Set.t;
restrict_conv_on_strict_subterms : bool;
modulo_betaiota : bool;
modulo_eta : bool;
- allow_K_in_toplevel_higher_order_unification : bool
}
-val default_unify_flags : unify_flags
-val default_no_delta_unify_flags : unify_flags
+type unify_flags = {
+ core_unify_flags : core_unify_flags;
+ merge_unify_flags : core_unify_flags;
+ subterm_unify_flags : core_unify_flags;
+ allow_K_in_toplevel_higher_order_unification : bool;
+ resolve_evars : bool
+}
+
+val default_core_unify_flags : unit -> core_unify_flags
+val default_no_delta_core_unify_flags : unit -> core_unify_flags
+
+val default_unify_flags : unit -> unify_flags
+val default_no_delta_unify_flags : unit -> unify_flags
-val elim_flags : unify_flags
-val elim_no_delta_flags : unify_flags
+val elim_flags : unit -> unify_flags
+val elim_no_delta_flags : unit -> unify_flags
(** The "unique" unification fonction *)
val w_unify :
env -> evar_map -> conv_pb -> ?flags:unify_flags -> constr -> constr -> evar_map
-(** [w_unify_to_subterm env (c,t) m] performs unification of [c] with a
+(** [w_unify_to_subterm env m (c,t)] performs unification of [c] with a
subterm of [t]. Constraints are added to [m] and the matched
subterm of [t] is also returned. *)
val w_unify_to_subterm :
@@ -53,25 +62,59 @@ val w_unify_meta_types : env -> ?flags:unify_flags -> evar_map -> evar_map
val w_coerce_to_type : env -> evar_map -> constr -> types -> types ->
evar_map * constr
+(* Looking for subterms in contexts at some occurrences, possibly with pattern*)
+
+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
+| AbstractExact of Names.Name.t * constr * types option * Locus.clause * bool
+
+val finish_evar_resolution : ?flags:Pretyping.inference_flags ->
+ env -> Evd.evar_map -> pending_constr -> Evd.evar_map * constr
+
+type abstraction_result =
+ Names.Id.t * named_context_val *
+ Context.named_declaration list * Names.Id.t option *
+ types * (Evd.evar_map * constr) option
+
+val make_abstraction : env -> Evd.evar_map -> constr ->
+ abstraction_request -> abstraction_result
+
+val pose_all_metas_as_evars : env -> evar_map -> constr -> evar_map * constr
+
(*i This should be in another module i*)
-(** [abstract_list_all env evd t c l]
- abstracts the terms in l over c to get a term of type t
+(** [abstract_list_all env evd t c l]
+ abstracts the terms in l over c to get a term of type t
(exported for inv.ml) *)
val abstract_list_all :
- env -> evar_map -> constr -> constr -> constr list -> constr
-
+ env -> evar_map -> constr -> constr -> constr list -> evar_map * (constr * types)
(* For tracing *)
-val w_merge : env -> bool -> unify_flags -> evar_map *
+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
val unify_0 : Environ.env ->
Evd.evar_map ->
Evd.conv_pb ->
- unify_flags ->
+ core_unify_flags ->
+ Term.types ->
+ Term.types ->
+ Evd.evar_map * Evd.metabinding list *
+ (Environ.env * Term.types Term.pexistential * Term.constr) list
+
+val unify_0_with_initial_metas :
+ Evd.evar_map * Evd.metabinding list *
+ (Environ.env * Term.types Term.pexistential * Term.constr) list ->
+ bool ->
+ Environ.env ->
+ Evd.conv_pb ->
+ core_unify_flags ->
Term.types ->
Term.types ->
Evd.evar_map * Evd.metabinding list *
diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml
index 047971d5..19613c4e 100644
--- a/pretyping/vnorm.ml
+++ b/pretyping/vnorm.ml
@@ -1,14 +1,16 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Util
open Names
open Declarations
open Term
+open Vars
open Environ
open Inductive
open Reduction
@@ -22,8 +24,9 @@ let crazy_type = mkSet
let decompose_prod env t =
let (name,dom,codom as res) = destProd (whd_betadeltaiota env t) in
- if name = Anonymous then (Name (id_of_string "x"),dom,codom)
- else res
+ match name with
+ | Anonymous -> (Name (Id.of_string "x"), dom, codom)
+ | Name _ -> res
exception Find_at of int
@@ -34,7 +37,8 @@ let invert_tag cst tag reloc_tbl =
try
for j = 0 to Array.length reloc_tbl - 1 do
let tagj,arity = reloc_tbl.(j) in
- if tag = tagj && (cst && arity = 0 || not(cst || arity = 0)) then
+ let no_arity = Int.equal arity 0 in
+ if Int.equal tag tagj && (cst && no_arity || not (cst || no_arity)) then
raise (Find_at j)
else ()
done;raise Not_found
@@ -44,30 +48,31 @@ let invert_tag cst tag reloc_tbl =
let find_rectype_a env c =
let (t, l) =
let t = whd_betadeltaiota env c in
- try destApp t with e when Errors.noncritical e -> (t,[||]) in
+ try destApp t with DestKO -> (t,[||]) in
match kind_of_term t with
| Ind ind -> (ind, l)
| _ -> raise Not_found
(* Instantiate inductives and parameters in constructor type *)
-let type_constructor mind mib typ params =
- let s = ind_subst mind mib in
+let type_constructor mind mib u typ params =
+ let s = ind_subst mind mib u in
let ctyp = substl s typ in
+ let ctyp = subst_instance_constr u ctyp in
let nparams = Array.length params in
- if nparams = 0 then ctyp
+ if Int.equal nparams 0 then ctyp
else
let _,ctyp = decompose_prod_n nparams ctyp in
- substl (List.rev (Array.to_list params)) ctyp
+ substl (Array.rev_to_list params) ctyp
let construct_of_constr const env tag typ =
- let (mind,_ as ind), allargs = find_rectype_a env typ in
+ let ((mind,_ as ind), u) as indu, allargs = find_rectype_a env typ in
(* spiwack : here be a branch for specific decompilation handled by retroknowledge *)
try
if const then
- ((retroknowledge Retroknowledge.get_vm_decompile_constant_info env (Ind ind) tag),
+ ((retroknowledge Retroknowledge.get_vm_decompile_constant_info env (mkIndU indu) tag),
typ) (*spiwack: this may need to be changed in case there are parameters in the
type which may cause a constant value to have an arity.
(type_constructor seems to be all about parameters actually)
@@ -80,8 +85,8 @@ let construct_of_constr const env tag typ =
let nparams = mib.mind_nparams in
let i = invert_tag const tag mip.mind_reloc_tbl in
let params = Array.sub allargs 0 nparams in
- let ctyp = type_constructor mind mib (mip.mind_nf_lc.(i-1)) params in
- (mkApp(mkConstruct(ind,i), params), ctyp)
+ let ctyp = type_constructor mind mib u (mip.mind_nf_lc.(i-1)) params in
+ (mkApp(mkConstructUi(indu,i), params), ctyp)
let construct_of_constr_const env tag typ =
fst (construct_of_constr true env tag typ)
@@ -91,7 +96,8 @@ let construct_of_constr_block = construct_of_constr false
let constr_type_of_idkey env idkey =
match idkey with
| ConstKey cst ->
- mkConst cst, Typeops.type_of_constant env cst
+ let const_type = Typeops.type_of_constant_in env cst in
+ mkConstU cst, const_type
| VarKey id ->
let (_,_,ty) = lookup_named id env in
mkVar id, ty
@@ -100,30 +106,33 @@ let constr_type_of_idkey env idkey =
let (_,_,ty) = lookup_rel n env in
mkRel n, lift n ty
-let type_of_ind env ind =
- type_of_inductive env (Inductive.lookup_mind_specif env ind)
+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 params dep p =
+let build_branches_type env (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 cty params in
- let decl,indapp = decompose_prod_assum typi in
- let ind,cargs = find_rectype_a env indapp in
+ let typi = type_constructor mind mib u cty params in
+ let decl,indapp = Reductionops.splay_prod env Evd.empty typi 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
let carity = snd (rtbl.(i)) in
let crealargs = Array.sub cargs nparams (Array.length cargs - nparams) in
let codom =
- let papp = mkApp(lift (List.length decl) p,crealargs) in
+ let ndecl = List.length decl in
+ let papp = mkApp(lift ndecl p,crealargs) in
if dep then
let cstr = ith_constructor_of_inductive ind (i+1) in
let relargs = Array.init carity (fun i -> mkRel (carity-i)) in
- let dep_cstr = mkApp(mkApp(mkConstruct cstr,params),relargs) in
+ let params = Array.map (lift ndecl) params in
+ let dep_cstr = mkApp(mkApp(mkConstructU (cstr,u),params),relargs) in
mkApp(papp,[|dep_cstr|])
else papp
in
- decl, codom
+ decl, decl_with_letin, codom
in Array.mapi build_one_branch mip.mind_nf_lc
let build_case_type dep p realargs c =
@@ -141,7 +150,7 @@ and nf_whd env whd typ =
| Vsort s -> mkSort s
| Vprod p ->
let dom = nf_vtype env (dom p) in
- let name = Name (id_of_string "x") 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 (name,None,dom) env) vc in
mkProd(name,dom,codom)
@@ -166,7 +175,7 @@ and nf_whd env whd typ =
| Vatom_stk(Aiddef(idkey,v), stk) ->
nf_whd env (whd_stack v stk) typ
| Vatom_stk(Aind ind, stk) ->
- nf_stk env (mkInd ind) (type_of_ind env ind) stk
+ nf_stk env (mkIndU ind) (type_of_ind env ind) stk
and nf_stk env c t stk =
match stk with
@@ -176,28 +185,25 @@ and nf_stk env c t stk =
nf_stk env (mkApp(c,args)) t stk
| Zfix (f,vargs) :: stk ->
let fa, typ = nf_fix_app env f vargs in
- let _,_,codom =
- try decompose_prod env typ
- with e when Errors.noncritical e -> exit 120
- in
+ let _,_,codom = decompose_prod env typ in
nf_stk env (mkApp(fa,[|c|])) (subst1 c codom) stk
| Zswitch sw :: stk ->
- let (mind,_ as ind),allargs = find_rectype_a env t in
+ 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 params,realargs = Util.Array.chop nparams allargs in
let pT =
- hnf_prod_applist env (type_of_ind env ind) (Array.to_list params) in
+ hnf_prod_applist env (type_of_ind env (ind,u)) (Array.to_list params) in
let pT = whd_betadeltaiota env pT in
- let dep, p = nf_predicate env ind mip params (type_of_switch sw) pT in
+ let dep, p = nf_predicate env (ind,u) mip params (type_of_switch sw) pT in
(* Calcul du type des branches *)
- let btypes = build_branches_type env ind mib mip params dep p in
+ let btypes = build_branches_type env 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,codom = btypes.(i) in
- let b = nf_val (push_rel_context decl env) v codom in
- it_mkLambda_or_LetIn b decl
+ let decl,decl_with_letin,codom = btypes.(i) in
+ let b = nf_val (Termops.push_rels_assum decl env) 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
@@ -209,21 +215,18 @@ and nf_predicate env ind mip params v pT =
| Vfun f, Prod _ ->
let k = nb_rel env in
let vb = body_of_vfun k f in
- let name,dom,codom =
- try decompose_prod env pT
- with e when Errors.noncritical e -> exit 121
- in
+ let name,dom,codom = decompose_prod env pT in
let dep,body =
nf_predicate (push_rel (name,None,dom) env) 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 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 n=0 then params else Array.map (lift n) params in
- let dom = mkApp(mkInd ind,Array.append params rargs) 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 (name,None,dom) env) vb in
true, mkLambda(name,dom,body)
| _, _ -> false, nf_val env v crazy_type
@@ -234,10 +237,7 @@ and nf_args env vargs t =
let args =
Array.init len
(fun i ->
- let _,dom,codom =
- try decompose_prod env !t
- with e when Errors.noncritical e -> exit 123
- in
+ let _,dom,codom = decompose_prod env !t in
let c = nf_val env (arg vargs i) dom in
t := subst1 c codom; c) in
!t,args
@@ -248,10 +248,7 @@ and nf_bargs env b t =
let args =
Array.init len
(fun i ->
- let _,dom,codom =
- try decompose_prod env !t
- with e when Errors.noncritical e -> exit 124
- in
+ let _,dom,codom = decompose_prod env !t in
let c = nf_val env (bfield b i) dom in
t := subst1 c codom; c) in
args
@@ -261,8 +258,10 @@ and nf_fun env f typ =
let vb = body_of_vfun k f in
let name,dom,codom =
try decompose_prod env typ
- with e when Errors.noncritical e ->
- raise (Type_errors.TypeError(env,Type_errors.ReferenceVariables typ))
+ with DestKO ->
+ (* 27/2/13: Turned this into an anomaly *)
+ Errors.anomaly
+ (Pp.strbrk "Returned a functional value in a type not recognized as a product type.")
in
let body = nf_val (push_rel (name,None,dom) env) vb codom in
mkLambda(name,dom,body)
@@ -274,9 +273,13 @@ and nf_fix env f =
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 name = Array.init ndef (fun _ -> (Name (id_of_string "Ffix"))) 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
- let fb = Util.array_map2 (fun v t -> nf_fun env v t) vb ft 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 fb = Util.Array.map2 norm_vb vb ft in
mkFix ((rec_args,init),(name,ft,fb))
and nf_fix_app env f vargs =
@@ -292,9 +295,9 @@ and nf_cofix env cf =
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 name = Array.init ndef (fun _ -> (Name (id_of_string "Fcofix"))) 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 v t) vb cft in
mkCoFix (init,(name,cft,cfb))
let cbv_vm env c t =
diff --git a/pretyping/vnorm.mli b/pretyping/vnorm.mli
index a0cf2372..7dabbc6c 100644
--- a/pretyping/vnorm.mli
+++ b/pretyping/vnorm.mli
@@ -1,15 +1,13 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Names
open Term
open Environ
-open Reduction
(** {6 Reduction functions } *)
val cbv_vm : env -> constr -> types -> constr