summaryrefslogtreecommitdiff
path: root/pretyping
diff options
context:
space:
mode:
authorGravatar Stephane Glondu <steph@glondu.net>2012-01-12 16:02:20 +0100
committerGravatar Stephane Glondu <steph@glondu.net>2012-01-12 16:02:20 +0100
commit97fefe1fcca363a1317e066e7f4b99b9c1e9987b (patch)
tree97ec6b7d831cc5fb66328b0c63a11db1cbb2f158 /pretyping
parent300293c119981054c95182a90c829058530a6b6f (diff)
Imported Upstream version 8.4~betaupstream/8.4_beta
Diffstat (limited to 'pretyping')
-rw-r--r--pretyping/arguments_renaming.ml118
-rw-r--r--pretyping/arguments_renaming.mli22
-rw-r--r--pretyping/cases.ml1109
-rw-r--r--pretyping/cases.mli30
-rw-r--r--pretyping/cbv.ml17
-rw-r--r--pretyping/cbv.mli21
-rw-r--r--pretyping/classops.ml37
-rw-r--r--pretyping/classops.mli45
-rw-r--r--pretyping/clenv.ml482
-rw-r--r--pretyping/clenv.mli148
-rw-r--r--pretyping/coercion.ml46
-rw-r--r--pretyping/coercion.mli24
-rw-r--r--pretyping/detyping.ml212
-rw-r--r--pretyping/detyping.mli60
-rw-r--r--pretyping/evarconv.ml614
-rw-r--r--pretyping/evarconv.mli36
-rw-r--r--pretyping/evarutil.ml1138
-rw-r--r--pretyping/evarutil.mli163
-rw-r--r--pretyping/evd.ml561
-rw-r--r--pretyping/evd.mli129
-rw-r--r--pretyping/glob_term.ml (renamed from pretyping/rawterm.ml)284
-rw-r--r--pretyping/glob_term.mli167
-rw-r--r--pretyping/indrec.ml68
-rw-r--r--pretyping/indrec.mli20
-rw-r--r--pretyping/inductiveops.ml19
-rw-r--r--pretyping/inductiveops.mli45
-rw-r--r--pretyping/matching.ml60
-rw-r--r--pretyping/matching.mli37
-rw-r--r--pretyping/namegen.ml19
-rw-r--r--pretyping/namegen.mli42
-rw-r--r--pretyping/pattern.ml205
-rw-r--r--pretyping/pattern.mli71
-rw-r--r--pretyping/pretype_errors.ml113
-rw-r--r--pretyping/pretype_errors.mli74
-rw-r--r--pretyping/pretyping.ml300
-rw-r--r--pretyping/pretyping.mli76
-rw-r--r--pretyping/pretyping.mllib4
-rw-r--r--pretyping/rawterm.mli167
-rw-r--r--pretyping/recordops.ml39
-rw-r--r--pretyping/recordops.mli46
-rw-r--r--pretyping/reductionops.ml83
-rw-r--r--pretyping/reductionops.mli57
-rw-r--r--pretyping/retyping.ml15
-rw-r--r--pretyping/retyping.mli12
-rw-r--r--pretyping/tacred.ml160
-rw-r--r--pretyping/tacred.mli57
-rw-r--r--pretyping/term_dnet.ml4
-rw-r--r--pretyping/term_dnet.mli36
-rw-r--r--pretyping/termops.ml298
-rw-r--r--pretyping/termops.mli147
-rw-r--r--pretyping/typeclasses.ml226
-rw-r--r--pretyping/typeclasses.mli63
-rw-r--r--pretyping/typeclasses_errors.ml9
-rw-r--r--pretyping/typeclasses_errors.mli10
-rw-r--r--pretyping/typing.ml108
-rw-r--r--pretyping/typing.mli31
-rw-r--r--pretyping/unification.ml764
-rw-r--r--pretyping/unification.mli56
-rw-r--r--pretyping/vnorm.ml4
-rw-r--r--pretyping/vnorm.mli6
60 files changed, 5047 insertions, 3967 deletions
diff --git a/pretyping/arguments_renaming.ml b/pretyping/arguments_renaming.ml
new file mode 100644
index 00000000..5e2284e1
--- /dev/null
+++ b/pretyping/arguments_renaming.ml
@@ -0,0 +1,118 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i*)
+open Names
+open Libnames
+open Decl_kinds
+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) }
+
+type req =
+ | ReqLocal
+ | ReqGlobal of global_reference * name list list
+
+let load_rename_args _ (_, (_, (r, names))) =
+ name_table := Refmap.add r names !name_table
+
+let cache_rename_args o = load_rename_args 1 o
+
+let classify_rename_args = function
+ | ReqLocal, _ -> Dispose
+ | ReqGlobal _, _ as o -> Substitute o
+
+let subst_rename_args (subst, (_, (r, names as orig))) =
+ ReqLocal,
+ let r' = fst (subst_global subst r) in
+ if r==r' then orig else (r', names)
+
+let section_segment_of_reference = function
+ | ConstRef con -> Lib.section_segment_of_constant con
+ | IndRef (kn,_) | ConstructRef ((kn,_),_) ->
+ Lib.section_segment_of_mutual_inductive kn
+ | _ -> []
+
+let discharge_rename_args = function
+ | _, (ReqGlobal (c, names), _) ->
+ let c' = pop_global_reference c in
+ let vars = section_segment_of_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
+ Some (ReqGlobal (c', names), (c', names'))
+ | _ -> None
+
+let rebuild_rename_args x = x
+
+let inRenameArgs = declare_object { (default_object "RENAME-ARGUMENTS" ) with
+ load_function = load_rename_args;
+ cache_function = cache_rename_args;
+ classify_function = classify_rename_args;
+ subst_function = subst_rename_args;
+ discharge_function = discharge_rename_args;
+ rebuild_function = rebuild_rename_args;
+}
+
+let rename_arguments local r names =
+ let req = if local then ReqLocal else ReqGlobal (r, names) in
+ Lib.add_anonymous_leaf (inRenameArgs (req, (r, names)))
+
+let arguments_names r = Refmap.find r !name_table
+
+let rec rename_prod c = function
+ | [] -> c
+ | (Name _ as n) :: tl ->
+ (match kind_of_type c with
+ | ProdType (_, s, t) -> mkProd (n, s, rename_prod t tl)
+ | _ -> c)
+ | _ :: tl ->
+ match kind_of_type c with
+ | ProdType (n, s, t) -> mkProd (n, s, rename_prod t tl)
+ | _ -> c
+
+let rename_type ty ref =
+ try rename_prod ty (List.hd (arguments_names 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 rename_type_of_inductive env ind =
+ let ty = Inductiveops.type_of_inductive env ind in
+ rename_type ty (IndRef ind)
+
+let rename_type_of_constructor env cstruct =
+ let ty = Inductiveops.type_of_constructor env cstruct in
+ rename_type ty (ConstructRef 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
+
diff --git a/pretyping/arguments_renaming.mli b/pretyping/arguments_renaming.mli
new file mode 100644
index 00000000..a484ecf7
--- /dev/null
+++ b/pretyping/arguments_renaming.mli
@@ -0,0 +1,22 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Libnames
+open Environ
+open Term
+
+val rename_arguments : bool -> global_reference -> name list list -> unit
+
+(** [Not_found] is raised is no names are defined for [r] *)
+val arguments_names : global_reference -> name 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_typing : env -> constr -> unsafe_judgment
diff --git a/pretyping/cases.ml b/pretyping/cases.ml
index 749101f7..8963ea5e 100644
--- a/pretyping/cases.ml
+++ b/pretyping/cases.ml
@@ -1,13 +1,12 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: cases.ml 14675 2011-11-17 22:19:34Z herbelin $ *)
-
+open Compat
open Util
open Names
open Nameops
@@ -21,7 +20,7 @@ open Sign
open Reductionops
open Typeops
open Type_errors
-open Rawterm
+open Glob_term
open Retyping
open Pretype_errors
open Evarutil
@@ -44,7 +43,7 @@ type pattern_matching_error =
exception PatternMatchingError of env * pattern_matching_error
let raise_pattern_matching_error (loc,ctx,te) =
- Stdpp.raise_with_loc loc (PatternMatchingError(ctx,te))
+ Loc.raise loc (PatternMatchingError(ctx,te))
let error_bad_pattern_loc loc cstr ind =
raise_pattern_matching_error (loc, Global.env(), BadPattern (cstr,ind))
@@ -64,33 +63,12 @@ let error_wrong_predicate_arity_loc loc env c n1 n2 =
let error_needs_inversion env x t =
raise (PatternMatchingError (env, NeedsInversion (x,t)))
-(**********************************************************************)
-(* Functions to deal with impossible cases *)
-
-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
- fun () ->
- match !impossible_default_case with
- | Some (id,type_of_id) ->
- make_judge id type_of_id
- | 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)))
-
-(**********************************************************************)
-
module type S = sig
val compile_cases :
loc -> case_style ->
- (type_constraint -> env -> evar_map ref -> rawconstr -> unsafe_judgment) * evar_map ref ->
+ (type_constraint -> env -> evar_map ref -> glob_constr -> unsafe_judgment) * evar_map ref ->
type_constraint ->
- env -> rawconstr option * tomatch_tuples * cases_clauses ->
+ env -> glob_constr option * tomatch_tuples * cases_clauses ->
unsafe_judgment
end
@@ -100,7 +78,7 @@ let rec list_try_compile f = function
| h::t ->
try f h
with UserError _ | TypeError _ | PretypeError _
- | Stdpp.Exc_located (_,(UserError _ | TypeError _ | PretypeError _)) ->
+ | Loc.Exc_located (_,(UserError _ | TypeError _ | PretypeError _)) ->
list_try_compile f t
let force_name =
@@ -128,37 +106,13 @@ let push_rels vars env = List.fold_right push_rel vars env
(* 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 regeneralize_rel i k j = if j = i+k then k+1 else j
-
-let rec regeneralize_index i k t = match kind_of_term t with
- | Rel j when j = i+k -> mkRel (k+1)
- | Rel j when j < i+k -> t
- | Rel j when j > i+k -> t
- | _ -> map_constr_with_binders succ (regeneralize_index i) k t
+let relocate_rel n1 n2 k j = if j = n1+k then n2+k else j
-type alias_constr =
- | DepAlias
- | NonDepAlias
-
-let mkSpecialLetInJudge j (na,(deppat,nondeppat,d,t)) =
- { uj_val =
- if
- isRel deppat or not (dependent (mkRel 1) j.uj_val) or
- d = NonDepAlias & not (dependent (mkRel 1) j.uj_type)
- then
- (* The body of pat is not needed to type j - see *)
- (* insert_aliases - and both deppat and nondeppat have the *)
- (* same type, then one can freely substitute one by the other. *)
- (* We use nondeppat only if it's a Rel to preserve sharing. *)
- if isRel nondeppat then
- subst1 nondeppat j.uj_val
- else subst1 deppat j.uj_val
- else
- (* The body of pat is not needed to type j but its value *)
- (* is dependent in the type of j; our choice is to *)
- (* enforce this dependency *)
- mkLetIn (na,deppat,t,j.uj_val);
- uj_type = subst1 deppat j.uj_type }
+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 j < n1+k -> t
+ | Rel j when j > n1+k -> t
+ | _ -> map_constr_with_binders succ (relocate_index n1 n2) k t
(**********************************************************************)
(* Structures used in compiling pattern-matching *)
@@ -178,29 +132,24 @@ type 'a equation =
type 'a matrix = 'a equation list
-type dep_status = KnownDep | KnownNotDep | DepUnknown
-
(* 1st argument of IsInd is the original ind before extracting the summary *)
type tomatch_type =
| IsInd of types * inductive_type * name list
| NotInd of constr option * types
type tomatch_status =
- | Pushed of ((constr * tomatch_type) * int list * (name * dep_status))
- | Alias of (constr * constr * alias_constr * constr)
- | Abstract of rel_declaration
+ | Pushed of ((constr * tomatch_type) * int list * name)
+ | Alias of (name * 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 alias_builder =
- | AliasLeaf
- | AliasConstructor of constructor
-
type pattern_history =
| Top
- | MakeAlias of alias_builder * pattern_continuation
+ | MakeConstructor of constructor * pattern_continuation
and pattern_continuation =
| Continuation of int * cases_pattern list * pattern_history
@@ -218,51 +167,47 @@ let feed_history arg = function
(* This is for non exhaustive error message *)
-let rec rawpattern_of_partial_history args2 = function
+let rec glob_pattern_of_partial_history args2 = function
| Continuation (n, args1, h) ->
let args3 = make_anonymous_patvars (n - (List.length args2)) in
- build_rawpattern (List.rev_append args1 (args2@args3)) h
+ build_glob_pattern (List.rev_append args1 (args2@args3)) h
| Result pl -> pl
-and build_rawpattern args = function
+and build_glob_pattern args = function
| Top -> args
- | MakeAlias (AliasLeaf, rh) ->
- assert (args = []);
- rawpattern_of_partial_history [PatVar (dummy_loc, Anonymous)] rh
- | MakeAlias (AliasConstructor pci, rh) ->
- rawpattern_of_partial_history
+ | MakeConstructor (pci, rh) ->
+ glob_pattern_of_partial_history
[PatCstr (dummy_loc, pci, args, Anonymous)] rh
-let complete_history = rawpattern_of_partial_history []
+let complete_history = glob_pattern_of_partial_history []
(* This is to build glued pattern-matching history and alias bodies *)
-let rec simplify_history = function
- | Continuation (0, l, Top) -> Result (List.rev l)
- | Continuation (0, l, MakeAlias (f, rh)) ->
- let pargs = List.rev l in
- let pat = match f with
- | AliasConstructor pci ->
- PatCstr (dummy_loc,pci,pargs,Anonymous)
- | AliasLeaf ->
- assert (l = []);
- PatVar (dummy_loc, Anonymous) in
- feed_history pat rh
- | h -> h
+let rec 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
+ | _ ->
+ anomaly "Constructor not yet filled with its arguments"
+
+let pop_history h =
+ feed_history (PatVar (dummy_loc, Anonymous)) h
(* Builds a continuation expecting [n] arguments and building [ci] applied
to this [n] arguments *)
-let push_history_pattern n current cont =
- Continuation (n, [], MakeAlias (current, cont))
+let push_history_pattern n pci cont =
+ Continuation (n, [], MakeConstructor (pci, cont))
(* A pattern-matching problem has the following form:
- env, evd |- <pred> Cases tomatch of mat end
+ env, evd |- match terms_to_tomatch return pred with mat end
- where tomatch is some sequence of "instructions" (t1 ... tn)
+ where terms_to_match is some sequence of "instructions" (t1 ... tp)
and mat is some matrix
+
(p11 ... p1n -> rhs1)
( ... )
(pm1 ... pmn -> rhsm)
@@ -270,16 +215,25 @@ let push_history_pattern n current cont =
Terms to match: there are 3 kinds of instructions
- "Pushed" terms to match are typed in [env]; these are usually just
- Rel(n) except for the initial terms given by user and typed in [env]
- - "Abstract" instructions means an abstraction has to be inserted in the
+ Rel(n) except for the initial terms given by user; in Pushed ((c,tm),deps,na),
+ [c] is the reference to the term (which is a Rel or an initial term), [tm] is
+ its type (telling whether we know if it is an inductive type or not),
+ [deps] is the list of terms to abstract before matching on [c] (these are
+ rels too)
+ - "Abstract" instructions mean that an abstraction has to be inserted in the
current branch to build (this means a pattern has been detected dependent
- in another one and generalisation is necessary to ensure well-typing)
- - "Alias" instructions means an alias has to be inserted (this alias
+ in another one and a generalization is necessary to ensure well-typing)
+ Abstract instructions extend the [env] in which the other instructions
+ are typed
+ - "Alias" instructions mean an alias has to be inserted (this alias
is usually removed at the end, except when its type is not the
same as the type of the matched term from which it comes -
typically because the inductive types are "real" parameters)
+ - "NonDepAlias" instructions mean the completion of a matching over
+ a term to match as for Alias but without inserting this alias
+ because there is no dependency in it
- Right-hand-sides:
+ Right-hand sides:
They consist of a raw term to type in an environment specific to the
clause they belong to: the names of declarations are those of the
@@ -335,7 +289,7 @@ let inductive_template evdref env tmloc ind =
let e = e_new_evar evdref env ~src:(hole_source n) ty' in
(e::subst,e::evarl,n+1)
| Some b ->
- (b::subst,evarl,n+1))
+ (substl subst b::subst,evarl,n+1))
arsign ([],[],1) in
applist (mkInd ind,List.rev evarl)
@@ -347,13 +301,24 @@ let try_find_ind env sigma typ realnames =
| 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 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)
+
let unify_tomatch_with_patterns evdref env loc typ pats realnames =
match find_row_ind pats with
| None -> NotInd (None,typ)
@@ -370,7 +335,7 @@ let find_tomatch_tycon evdref env loc = function
empty_tycon,None
let coerce_row typing_fun evdref env pats (tomatch,(_,indopt)) =
- let loc = Some (loc_of_rawconstr tomatch) in
+ 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 typ = nf_evar !evdref j.uj_type in
@@ -466,6 +431,15 @@ let remove_current_pattern eqn =
alias_stack = alias_of_pat pat :: eqn.alias_stack }
| [] -> anomaly "Empty list of patterns"
+let push_current_pattern (cur,ty) eqn =
+ match eqn.patterns with
+ | pat::pats ->
+ let rhs_env = push_rel (alias_of_pat pat,Some cur,ty) eqn.rhs.rhs_env in
+ { eqn with
+ rhs = { eqn.rhs with rhs_env = rhs_env };
+ patterns = pats }
+ | [] -> anomaly "Empty list of patterns"
+
let prepend_pattern tms eqn = {eqn with patterns = tms@eqn.patterns }
(**********************************************************************)
@@ -540,7 +514,7 @@ let is_dep_patt_in eqn = function
| PatVar (_,name) -> occur_in_rhs name eqn.rhs
| PatCstr _ -> true
-let mk_dep_patt_row (pats,eqn) =
+let mk_dep_patt_row (pats,_,eqn) =
List.map (is_dep_patt_in eqn) pats
let dependencies_in_pure_rhs nargs eqns =
@@ -554,8 +528,8 @@ let dependent_decl a = function
| (na,Some c,t) -> dependent a t || dependent a c
let rec dep_in_tomatch n = function
- | (Pushed _ | Alias _) :: l -> dep_in_tomatch n l
- | Abstract d :: l -> dependent_decl (mkRel n) d or dep_in_tomatch (n+1) l
+ | (Pushed _ | Alias _ | NonDepAlias) :: l -> dep_in_tomatch n l
+ | Abstract (_,d) :: l -> dependent_decl (mkRel n) d or dep_in_tomatch (n+1) l
| [] -> false
let dependencies_in_rhs nargs current tms eqns =
@@ -565,75 +539,90 @@ let dependencies_in_rhs nargs current tms eqns =
(* Computing the matrix of dependencies *)
-(* We are in context d1...dn |- and [find_dependencies k 1 nextlist]
- computes for declaration [k+1] in which of declarations in
- [nextlist] (which corresponds to d(k+2)...dn) it depends;
- declarations are expressed by index, e.g. in dependency list
- [n-2;1], [1] points to [dn] and [n-2] to [d3] *)
+(* [find_dependency_list tmi [d(i+1);...;dn]] computes in which
+ declarations [d(i+1);...;dn] the term [tmi] is dependent in.
-let rec find_dependency_list k n = function
+ [find_dependencies_signature (used1,...,usedn) ((tm1,d1),...,(tmn,dn))]
+ returns [(deps1,...,depsn)] where [depsi] is a subset of n,..,i+1
+ denoting in which of the d(i+1)...dn, the term tmi is dependent.
+ Dependencies are expressed by index, e.g. in dependency list
+ [n-2;1], [1] points to [dn] and [n-2] to [d3]
+*)
+
+let rec find_dependency_list tmblock = function
| [] -> []
| (used,tdeps,d)::rest ->
- let deps = find_dependency_list k (n+1) rest in
- if used && dependent_decl (mkRel n) d
+ 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)
else deps
-let find_dependencies is_dep_or_cstr_in_rhs d (k,nextlist) =
- let deps = find_dependency_list k 1 nextlist in
+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 <> []
- then (k-1,(true ,deps,d)::nextlist)
- else (k-1,(false,[] ,d)::nextlist)
+ then ((true ,deps,d)::nextlist)
+ else ((false,[] ,d)::nextlist)
let find_dependencies_signature deps_in_rhs typs =
- let k = List.length deps_in_rhs in
- let _,l = List.fold_right2 find_dependencies deps_in_rhs typs (k,[]) in
+ let l = List.fold_right2 find_dependencies deps_in_rhs typs [] in
List.map (fun (_,deps,_) -> deps) l
(* Assume we had terms t1..tq to match in a context xp:Tp,...,x1:T1 |-
and xn:Tn has just been regeneralized into x:Tn so that the terms
to match are now to be considered in the context xp:Tp,...,x1:T1,x:Tn |-.
- [regeneralize_index_tomatch n tomatch] updates t1..tq so that
- former references to xn are now references to x. Note that t1..tq
- are already adjusted to the context xp:Tp,...,x1:T1,x:Tn |-. *)
+ [relocate_index_tomatch n 1 tomatch] updates t1..tq so that
+ former references to xn1 are now references to x. Note that t1..tq
+ are already adjusted to the context xp:Tp,...,x1:T1,x:Tn |-.
+
+ [relocate_index_tomatch 1 n tomatch] will go the way back.
+ *)
-let regeneralize_index_tomatch n =
+let relocate_index_tomatch n1 n2 =
let rec genrec depth = function
| [] ->
[]
- | Pushed ((c,tm),l,dep) :: rest ->
- let c = regeneralize_index n depth c in
- let tm = map_tomatch_type (regeneralize_index n depth) tm in
- let l = List.map (regeneralize_rel n depth) l in
- Pushed ((c,tm),l,dep) :: genrec depth rest
- | Alias (c1,c2,d,t) :: rest ->
- Alias (regeneralize_index n depth c1,c2,d,t) :: genrec depth rest
- | Abstract d :: rest ->
- Abstract (map_rel_declaration (regeneralize_index n depth) d)
+ | Pushed ((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 ->
+ (* [c] is out of relocation scope *)
+ Alias (na,c,map_pair (relocate_index n1 n2 depth) d) :: genrec depth rest
+ | NonDepAlias :: rest ->
+ NonDepAlias :: genrec depth rest
+ | Abstract (i,d) :: rest ->
+ let i = relocate_rel n1 n2 depth i in
+ Abstract (i,map_rel_declaration (relocate_index n1 n2 depth) d)
:: genrec (depth+1) rest in
genrec 0
+(* [replace_tomatch n c tomatch] replaces [Rel n] by [c] in [tomatch] *)
+
let rec replace_term n c k t =
- if t = mkRel (n+k) then lift k c
+ if isRel t && 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 (dep,_) = function
- | NotInd _ -> if dep<>Anonymous then 1 else 0
- | IsInd (_,_,names) -> List.length names + if dep<>Anonymous then 1 else 0
+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 replace_tomatch n c =
let rec replrec depth = function
| [] -> []
- | Pushed ((b,tm),l,dep) :: rest ->
+ | Pushed ((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,dep) :: replrec depth rest
- | Alias (c1,c2,d,t) :: rest ->
- Alias (replace_term n c depth c1,c2,d,t) :: replrec depth rest
- | Abstract d :: rest ->
- Abstract (map_rel_declaration (replace_term n c depth) d)
+ Pushed ((b,tm),l,na) :: replrec depth rest
+ | Alias (na,b,d) :: rest ->
+ (* [b] is out of replacement scope *)
+ Alias (na,b,map_pair (replace_term n c depth) d) :: replrec depth rest
+ | NonDepAlias :: rest ->
+ NonDepAlias :: replrec depth rest
+ | Abstract (i,d) :: rest ->
+ Abstract (i,map_rel_declaration (replace_term n c depth) d)
:: replrec (depth+1) rest in
replrec 0
@@ -646,16 +635,19 @@ let replace_tomatch n c =
let rec liftn_tomatch_stack n depth = function
| [] -> []
- | Pushed ((c,tm),l,dep)::rest ->
+ | Pushed ((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,dep)::(liftn_tomatch_stack n depth rest)
- | Alias (c1,c2,d,t)::rest ->
- Alias (liftn n depth c1,liftn n depth c2,d,liftn n depth t)
+ 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)
::(liftn_tomatch_stack n depth rest)
- | Abstract d::rest ->
- Abstract (map_rel_declaration (liftn n depth) d)
+ | NonDepAlias :: rest ->
+ NonDepAlias :: liftn_tomatch_stack n depth rest
+ | Abstract (i,d)::rest ->
+ let i = if i<depth then i else i+n in
+ Abstract (i,map_rel_declaration (liftn n depth) d)
::(liftn_tomatch_stack n (depth+1) rest)
let lift_tomatch_stack n = liftn_tomatch_stack n 1
@@ -675,8 +667,14 @@ let lift_tomatch_stack n = liftn_tomatch_stack n 1
and [match y with (S (S n)) => n | n => n end] into
[match y with O => y | (S n0) => match n0 with O => y | (S n) => n end end]
- i.e. user names should be preserved and created names should not
- interfere with user names *)
+ i.e. user names should be preserved and created names should not
+ interfere with user names
+
+ The exact names here are not important for typing (because they are
+ put in pb.env and not in the rhs.rhs_env of branches. However,
+ whether a name is Anonymous or not may have an effect on whether a
+ generalization is done or not.
+ *)
let merge_name get_name obj = function
| Anonymous -> get_name obj
@@ -687,15 +685,18 @@ 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
(* If any, we prefer names used in pats, from top to bottom *)
- let names2 =
+ let names2,aliasname =
List.fold_right
- (fun (pats,eqn) names -> merge_names alias_of_pat pats names)
- eqns names1 in
+ (fun (pats,pat_alias,eqn) (names,aliasname) ->
+ (merge_names alias_of_pat pats names,
+ merge_name (fun x -> x) pat_alias aliasname))
+ eqns (names1,Anonymous) in
(* 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) [] eqns in
- let names4,_ =
+ List.fold_left (fun l (_,_,eqn) -> list_union l eqn.rhs.avoid_ids)
+ [] eqns in
+ let names3,_ =
List.fold_left2
(fun (l,avoid) d na ->
let na =
@@ -705,10 +706,18 @@ let get_names env sign eqns =
in
(na::l,(out_name na)::avoid))
([],allvars) (List.rev sign) names2 in
- names4
+ names3,aliasname
-(************************************************************************)
+(*****************************************************************)
(* Recovering names for variables pushed to the rhs' environment *)
+(* We just factorized a match over a matrix of equations *)
+(* "C xi1 .. xin as xi" as a single match over "C y1 .. yn as y" *)
+(* We now replace the names y1 .. yn y by the actual names *)
+(* xi1 .. xin xi to be found in the i-th clause of the matrix *)
+
+let set_declaration_name x (_,c,t) = (x,c,t)
+
+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))
@@ -716,68 +725,28 @@ let push_rels_eqn sign eqn =
{eqn with rhs = {eqn.rhs with rhs_env = push_rels sign eqn.rhs.rhs_env} }
let push_rels_eqn_with_names sign eqn =
- let pats = List.rev (list_firstn (List.length sign) eqn.patterns) in
- let sign = recover_alias_names alias_of_pat pats sign 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
-let build_aliases_context env sigma names allpats pats =
- (* pats is the list of bodies to push as an alias *)
- (* They all are defined in env and we turn them into a sign *)
- (* cuts in sign need to be done in allpats *)
- let rec insert env sign1 sign2 n newallpats oldallpats = function
- | (deppat,_,_,_)::pats, Anonymous::names when not (isRel deppat) ->
- (* Anonymous leaves must be considered named and treated in the *)
- (* next clause because they may occur in implicit arguments *)
- insert env sign1 sign2
- n newallpats (List.map List.tl oldallpats) (pats,names)
- | (deppat,nondeppat,d,t)::pats, na::names ->
- let nondeppat = lift n nondeppat in
- let deppat = lift n deppat in
- let newallpats =
- List.map2 (fun l1 l2 -> List.hd l2::l1) newallpats oldallpats in
- let oldallpats = List.map List.tl oldallpats in
- let decl = (na,Some deppat,t) in
- let a = (deppat,nondeppat,d,t) in
- insert (push_rel decl env) (decl::sign1) ((na,a)::sign2) (n+1)
- newallpats oldallpats (pats,names)
- | [], [] -> newallpats, sign1, sign2, env
- | _ -> anomaly "Inconsistent alias and name lists" in
- let allpats = List.map (fun x -> [x]) allpats
- in insert env [] [] 0 (List.map (fun _ -> []) allpats) allpats (pats, names)
-
-let insert_aliases_eqn sign eqnnames alias_rest eqn =
- let thissign = List.map2 (fun na (_,c,t) -> (na,c,t)) eqnnames sign in
- { eqn with
- alias_stack = alias_rest;
- rhs = {eqn.rhs with rhs_env = push_rels thissign eqn.rhs.rhs_env } }
-
-let insert_aliases env sigma alias eqns =
- (* Là, y a une faiblesse, si un alias est utilisé dans un cas par *)
- (* défaut présent mais inutile, ce qui est le cas général, l'alias *)
- (* est introduit même s'il n'est pas utilisé dans les cas réguliers *)
- let eqnsnames = List.map (fun eqn -> List.hd eqn.alias_stack) eqns in
- let alias_rests = List.map (fun eqn -> List.tl eqn.alias_stack) eqns in
- (* name2 takes the meet of all needed aliases *)
- let name2 =
- List.fold_right (merge_name (fun x -> x)) eqnsnames Anonymous in
- (* Only needed aliases are kept by build_aliases_context *)
- let eqnsnames, sign1, sign2, env =
- build_aliases_context env sigma [name2] eqnsnames [alias] in
- let eqns = list_map3 (insert_aliases_eqn sign1) eqnsnames alias_rests eqns in
- sign2, env, eqns
+let push_generalized_decl_eqn env n (na,c,t) eqn =
+ let na = match na with
+ | Anonymous -> Anonymous
+ | Name id -> pi1 (Environ.lookup_rel n eqn.rhs.rhs_env) in
+ push_rels_eqn [(na,c,t)] eqn
-(**********************************************************************)
-(* Functions to deal with elimination predicate *)
+let drop_alias_eqn eqn =
+ { eqn with alias_stack = List.tl eqn.alias_stack }
-exception Occur
-let noccur_between_without_evar n m term =
- let rec occur_rec n c = match kind_of_term c with
- | Rel p -> if n<=p && p<n+m then raise Occur
- | Evar (_,cl) -> ()
- | _ -> iter_constr_with_binders succ occur_rec n c
- in
- (m = 0) or (try occur_rec n term; true with Occur -> false)
+let push_alias_eqn alias eqn =
+ let aliasname = List.hd eqn.alias_stack in
+ let eqn = drop_alias_eqn eqn in
+ let alias = set_declaration_name aliasname alias in
+ push_rels_eqn [alias] eqn
+(**********************************************************************)
+(* Functions to deal with elimination predicate *)
(* Infering the predicate *)
(*
@@ -797,8 +766,8 @@ Assume the types in the branches are the following
Assume the type of the global case expression is Gamma |- T
-The predicate has the form phi = [y1..yq][z:I(y1..yq)]? and must satisfy
-the following n+1 equations:
+The predicate has the form phi = [y1..yq][z:I(y1..yq)]psi and it has to
+satisfy the following n+1 equations:
Gamma, x11...x1p1 |- (phi u11..u1q (C1 x11..x1p1)) = T1
...
@@ -808,11 +777,11 @@ the following n+1 equations:
Some hints:
- Clearly, if xij occurs in Ti, then, a "match z with (Ci xi1..xipi)
- => ..." or a "psi(yk)", with psi extracting xij from uik, should be
+ => ... end" or a "psi(yk)", with psi extracting xij from uik, should be
inserted somewhere in Ti.
-- If T is undefined, an easy solution is to insert a "match z with (Ci
- xi1..xipi) => ..." in front of each Ti
+- If T is undefined, an easy solution is to insert a "match z with
+ (Ci xi1..xipi) => ... end" in front of each Ti
- Otherwise, T1..Tn and T must be step by step unified, if some of them
diverge, then try to replace the diverging subterm by one of y1..yq or z.
@@ -825,10 +794,10 @@ Some hints:
let rec map_predicate f k ccl = function
| [] -> f k ccl
- | Pushed ((_,tm),_,dep) :: rest ->
- let k' = length_of_tomatch_type_sign dep tm in
+ | Pushed ((_,tm),_,na) :: rest ->
+ let k' = length_of_tomatch_type_sign na tm in
map_predicate f (k+k') ccl rest
- | Alias _ :: rest ->
+ | (Alias _ | NonDepAlias) :: rest ->
map_predicate f k ccl rest
| Abstract _ :: rest ->
map_predicate f (k+1) ccl rest
@@ -839,7 +808,7 @@ let liftn_predicate n = map_predicate (liftn n)
let lift_predicate n = liftn_predicate n 1
-let regeneralize_index_predicate n = map_predicate (regeneralize_index n) 0
+let regeneralize_index_predicate n = map_predicate (relocate_index n 1) 0
let substnl_predicate sigma = map_predicate (substnl sigma)
@@ -867,55 +836,74 @@ let specialize_predicate_var (cur,typ,dep) tms ccl =
(* We first need to lift t(x) s.t. it is typed in Gamma, X:=rargs, x' *)
(* then we have to replace x by x' in t(x) and y by y' in P *)
(*****************************************************************************)
-let generalize_predicate (names,(nadep,_)) ny d tms ccl =
- if nadep=Anonymous then anomaly "Undetected dependency";
+let generalize_predicate (names,na) ny d tms ccl =
+ if na=Anonymous then anomaly "Undetected dependency";
let p = List.length names + 1 in
let ccl = lift_predicate 1 ccl tms in
regeneralize_index_predicate (ny+p+1) ccl tms
+(*****************************************************************************)
+(* We just matched over cur:ind(realargs) in the following matching problem *)
+(* *)
+(* env |- match cur tms return ccl with ... end *)
+(* *)
+(* and we want to build the predicate corresponding to the individual *)
+(* matching over cur *)
+(* *)
+(* pred = fun X:realargstyps x:ind(X)] PI tms.ccl *)
+(* *)
+(* where pred is computed by abstract_predicate and PI tms.ccl by *)
+(* extract_predicate *)
+(*****************************************************************************)
let rec extract_predicate ccl = function
- | Alias (deppat,nondeppat,_,_)::tms ->
- (* substitution already done in build_branch *)
+ | (Alias _ | NonDepAlias)::tms ->
+ (* substitution already done in build_branch *)
extract_predicate ccl tms
- | Abstract d::tms ->
- mkProd_or_LetIn d (extract_predicate ccl tms)
- | Pushed ((cur,NotInd _),_,(dep,_))::tms ->
- let tms = if dep<>Anonymous then lift_tomatch_stack 1 tms else tms in
+ | 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 dep<>Anonymous then subst1 cur pred else pred
- | Pushed ((cur,IsInd (_,IndType(_,realargs),_)),_,(dep,_))::tms ->
+ if na<>Anonymous then subst1 cur pred else pred
+ | Pushed ((cur,IsInd (_,IndType(_,realargs),_)),_,na)::tms ->
let realargs = List.rev realargs in
- let k = if dep<>Anonymous then 1 else 0 in
+ let k = if na<>Anonymous then 1 else 0 in
let tms = lift_tomatch_stack (List.length realargs + k) tms in
let pred = extract_predicate ccl tms in
- substl (if dep<>Anonymous then cur::realargs else realargs) pred
+ substl (if na<>Anonymous then cur::realargs else realargs) pred
| [] ->
ccl
-let abstract_predicate env sigma indf cur (names,(nadep,_)) tms ccl =
+let abstract_predicate env sigma indf cur realargs (names,na) tms ccl =
let sign = make_arity_signature env true indf in
- (* n is the number of real args + 1 *)
+ (* n is the number of real args + 1 (+ possible let-ins in sign) *)
let n = List.length sign in
- let tms = lift_tomatch_stack n tms in
- let tms =
- match kind_of_term cur with
- | Rel i -> regeneralize_index_tomatch (i+n) tms
- | _ -> (* Initial case *) tms in
- let sign = List.map2 (fun na (_,c,t) -> (na,c,t)) (nadep::names) sign in
- let ccl = if nadep <> Anonymous then ccl else lift_predicate 1 ccl tms in
+ (* Before abstracting we generalize over cur and on those realargs *)
+ (* that are rels, consistently with the specialization made in *)
+ (* build_branch *)
+ let tms = List.fold_right2 (fun par arg tomatch ->
+ match kind_of_term par with
+ | Rel i -> relocate_index_tomatch (i+n) (destRel arg) tomatch
+ | _ -> tomatch) (realargs@[cur]) (extended_rel_list 0 sign)
+ (lift_tomatch_stack n tms) in
+ (* Pred is already dependent in the current term to match (if *)
+ (* (na<>Anonymous) and its realargs; we just need to adjust it to *)
+ (* 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 pred = extract_predicate ccl tms in
+ (* Build the predicate properly speaking *)
+ let sign = List.map2 set_declaration_name (na::names) sign in
it_mkLambda_or_LetIn_name env pred sign
-let known_dependent (_,dep) = (dep = KnownDep)
-
(* [expand_arg] is used by [specialize_predicate]
- it replaces gamma, x1...xn, x1...xk |- pred
- by gamma, x1...xn, x1...xk-1 |- [X=realargs,xk=xk]pred (if dep) or
- by gamma, x1...xn, x1...xk-1 |- [X=realargs]pred (if not dep) *)
+ if Yk denotes [Xk;xk] or [Xk],
+ it replaces gamma, x1...xn, x1...xk Yk+1...Yn |- pred
+ by gamma, x1...xn, x1...xk-1 [Xk;xk] Yk+1...Yn |- pred (if dep) or
+ by gamma, x1...xn, x1...xk-1 [Xk] Yk+1...Yn |- pred (if not dep) *)
-let expand_arg tms ccl ((_,t),_,na) =
+let expand_arg tms (p,ccl) ((_,t),_,na) =
let k = length_of_tomatch_type_sign na t in
- lift_predicate (k-1) ccl tms
+ (p+k,liftn_predicate (k-1) (p+1) ccl tms)
let adjust_impossible_cases pb pred tomatch submat =
if submat = [] then
@@ -926,7 +914,7 @@ let adjust_impossible_cases pb pred tomatch submat =
(* we add an "assert false" case *)
let pats = List.map (fun _ -> PatVar (dummy_loc,Anonymous)) tomatch in
let aliasnames =
- map_succeed (function Alias _ -> Anonymous | _ -> failwith"") tomatch
+ map_succeed (function Alias _ | NonDepAlias -> Anonymous | _ -> failwith"") tomatch
in
[ { patterns = pats;
rhs = { rhs_env = pb.env;
@@ -942,71 +930,132 @@ let adjust_impossible_cases pb pred tomatch submat =
submat
(*****************************************************************************)
-(* pred = [X:=realargs;x:=c]P types the following problem: *)
+(* Let pred = PI [X;x:I(X)]. PI tms. P be a typing predicate for the *)
+(* following pattern-matching problem: *)
(* *)
-(* Gamma |- match Pushed(c:I(realargs)) rest with...end: pred *)
+(* Gamma |- match Pushed(c:I(V)) as x in I(X), tms return pred with...end *)
(* *)
(* where the branch with constructor Ci:(x1:T1)...(xn:Tn)->I(realargsi) *)
-(* is considered. Assume each Ti is some Ii(argsi). *)
-(* We let e=Ci(x1,...,xn) and replace pred by *)
+(* is considered. Assume each Ti is some Ii(argsi) with Ti:PI Ui. sort_i *)
+(* We let subst = X:=realargsi;x:=Ci(x1,...,xn) and replace pred by *)
(* *)
-(* pred' = [X1:=rargs1,x1:=x1']...[Xn:=rargsn,xn:=xn'](P[X:=realargsi;x:=e]) *)
+(* pred' = PI [X1:Ui;x1:I1(X1)]...[Xn:Un;xn:In(Xn)]. (PI tms. P)[subst] *)
(* *)
-(* s.t Gamma,x1'..xn' |- match Pushed(x1')..Pushed(xn') rest with..end :pred'*)
+(* s.t. the following well-typed sub-pattern-matching problem is obtained *)
+(* *)
+(* Gamma,x'1..x'n |- *)
+(* match *)
+(* Pushed(x'1) as x1 in I(X1), *)
+(* .., *)
+(* Pushed(x'n) as xn in I(Xn), *)
+(* tms *)
+(* return pred' *)
+(* with .. end *)
(* *)
(*****************************************************************************)
-let specialize_predicate newtomatchs (names,(depna,_)) arsign cs tms ccl =
- (* Assume some gamma st: gamma, (X,x:=realargs,copt), tms |- ccl *)
+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
- (* We adjust pred st: gamma, x1..xn, (X,x:=realargs,copt), tms |- ccl' *)
+ (* 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 *)
let n = cs.cs_nargs in
let ccl' = liftn_predicate n (k+1) ccl tms in
- let argsi =
+ (* We prepare the substitution of X and x:I(X) *)
+ let realargsi =
if 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
- (* The substituends argsi, copti are all defined in gamma, x1...xn *)
- (* We need _parallel_ bindings to get gamma, x1...xn, tms |- ccl'' *)
+ let copti =
+ if depna<>Anonymous then Some (build_dependent_constructor cs) else None 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?) *)
let ccl'' =
- whd_betaiota Evd.empty (subst_predicate (argsi, copti) ccl' tms) in
- (* We adjust ccl st: gamma, x1..xn, x1..xn, tms |- ccl'' *)
+ whd_betaiota Evd.empty (subst_predicate (realargsi, copti) ccl' tms) in
+ (* We adjust ccl st: gamma, x'1..x'n, x1..xn, tms |- ccl'' *)
let ccl''' = liftn_predicate n (n+1) ccl'' tms in
- (* We finally get gamma,x1..xn |- [X1,x1:=R1,x1]..[Xn,xn:=Rn,xn]pred'''*)
- List.fold_left (expand_arg tms) ccl''' newtomatchs
+ (* We finally get gamma,x'1..x'n,x |- [X1;x1:I(X1)]..[Xn;xn:I(Xn)]pred'''*)
+ snd (List.fold_left (expand_arg tms) (1,ccl''') newtomatchs)
let find_predicate loc env evdref p current (IndType (indf,realargs)) dep tms =
- let pred= abstract_predicate env !evdref indf current dep tms p in
+ let pred = abstract_predicate env !evdref indf current realargs dep tms p in
(pred, whd_betaiota !evdref
- (applist (pred, realargs@[current])), new_Type ())
+ (applist (pred, realargs@[current])))
(* Take into account that a type has been discovered to be inductive, leading
to more dependencies in the predicate if the type has indices *)
let adjust_predicate_from_tomatch tomatch (current,typ as ct) pb =
- let ((_,oldtyp),deps,((nadep,_) as dep)) = tomatch in
+ let ((_,oldtyp),deps,na) = tomatch in
match typ, oldtyp with
| IsInd (_,_,names), NotInd _ ->
- let k = if nadep <> Anonymous then 2 else 1 in
+ let k = if na <> Anonymous then 2 else 1 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,dep)
+ (ct,List.map (fun i -> if i >= k then i+n else i) deps,na)
| _ ->
- pb, (ct,deps,dep)
+ pb, (ct,deps,na)
+
+(* Remove commutative cuts that turn out to be non-dependent after
+ some evars have been instantiated *)
+
+let rec ungeneralize n ng body =
+ match kind_of_term body with
+ | Lambda (_,_,c) when ng = 0 ->
+ subst1 (mkRel n) c
+ | Lambda (na,t,c) ->
+ (* We traverse an inner generalization *)
+ mkLambda (na,t,ungeneralize (n+1) (ng-1) c)
+ | LetIn (na,b,t,c) ->
+ (* We traverse an alias *)
+ mkLetIn (na,b,t,ungeneralize (n+1) ng c)
+ | Case (ci,p,c,brs) ->
+ (* We traverse a split *)
+ let p =
+ let sign,p = decompose_lam_assum p in
+ let sign2,p = decompose_prod_n_assum ng p in
+ let p = prod_applist p [mkRel (n+List.length sign+ng)] in
+ it_mkLambda_or_LetIn (it_mkProd_or_LetIn p sign2) sign in
+ mkCase (ci,p,c,array_map2 (fun q c ->
+ let sign,b = decompose_lam_n_assum q c in
+ it_mkLambda_or_LetIn (ungeneralize (n+q) ng b) sign)
+ ci.ci_cstr_ndecls brs)
+ | App (f,args) ->
+ (* We traverse an inner generalization *)
+ assert (isCase f);
+ mkApp (ungeneralize n (ng+Array.length args) f,args)
+ | _ -> assert false
+
+let ungeneralize_branch n k (sign,body) cs =
+ (sign,ungeneralize (n+cs.cs_nargs) k body)
+
+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
+ (* 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
+ brs, Abstract (i,d) :: tomatch, pred, inst
+ else
+ (* Finally, no dependency remains, so, we can replace the generalized *)
+ (* terms by its actual value in both the remaining terms to match and *)
+ (* the bodies of the Case *)
+ let pred = lift_predicate (-1) pred tomatch in
+ let tomatch = relocate_index_tomatch 1 (n+1) tomatch in
+ let tomatch = lift_tomatch_stack (-1) tomatch 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
(************************************************************************)
(* Sorting equations by constructor *)
-type inversion_problem =
- (* the discriminating arg in some Ind and its order in Ind *)
- | Incompatible of int * (int * int)
- | Constraints of (int * constr) list
-
-let solve_constraints constr_info indt =
- (* TODO *)
- Constraints []
-
let rec irrefutable env = function
| PatVar (_,name) -> true
| PatCstr (_,cstr,args,_) ->
@@ -1034,12 +1083,12 @@ let group_equations pb ind current cstrs mat =
(* This is a default clause that we expand *)
for i=1 to Array.length cstrs do
let args = make_anonymous_patvars cstrs.(i-1).cs_nargs in
- brs.(i-1) <- (args, rest) :: brs.(i-1)
+ brs.(i-1) <- (args, name, rest) :: brs.(i-1)
done
- | PatCstr (loc,((_,i)),args,_) ->
+ | PatCstr (loc,((_,i)),args,name) ->
(* This is a regular clause *)
only_default := false;
- brs.(i-1) <- (args,rest) :: brs.(i-1)) mat () in
+ brs.(i-1) <- (args, name, rest) :: brs.(i-1)) mat () in
(brs,!only_default)
(************************************************************************)
@@ -1047,16 +1096,18 @@ let group_equations pb ind current cstrs mat =
(* Abstracting over dependent subterms to match *)
let rec generalize_problem names pb = function
- | [] -> pb
+ | [] -> pb, []
| i::l ->
- let d = map_rel_declaration (lift i) (Environ.lookup_rel i pb.env) in
+ 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 pb' = generalize_problem names pb l in
let tomatch = lift_tomatch_stack 1 pb'.tomatch in
- let tomatch = regeneralize_index_tomatch (i+1) tomatch in
+ let tomatch = relocate_index_tomatch (i+1) 1 tomatch in
{ pb' with
- tomatch = Abstract d :: tomatch;
- pred = generalize_predicate names i d pb.tomatch pb'.pred }
+ tomatch = Abstract (i,d) :: tomatch;
+ pred = generalize_predicate names i d pb'.tomatch pb'.pred },
+ i::deps
(* No more patterns: typing the right-hand side of equations *)
let build_leaf pb =
@@ -1064,87 +1115,92 @@ let build_leaf pb =
let j = pb.typing_function (mk_tycon pb.pred) rhs.rhs_env pb.evdref rhs.it in
j_nf_evar !(pb.evdref) j
-(* Building the sub-problem when all patterns are variables *)
-let shift_problem ((current,t),_,(nadep,_)) pb =
- {pb with
- tomatch = Alias (current,current,NonDepAlias,type_of_tomatch t)::pb.tomatch;
- pred = specialize_predicate_var (current,t,nadep) pb.tomatch pb.pred;
- history = push_history_pattern 0 AliasLeaf pb.history;
- mat = List.map remove_current_pattern pb.mat }
-
-(* Building the sub-pattern-matching problem for a given branch *)
-let build_branch current deps (realnames,dep) pb arsign eqns const_info =
- (* We remember that we descend through a constructor *)
- let alias_type =
- if Array.length const_info.cs_concl_realargs = 0
- & not (known_dependent dep) & deps = []
- then
- NonDepAlias
- else
- DepAlias
- in
+(* 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 =
+ (* We remember that we descend through constructor C *)
let history =
- push_history_pattern const_info.cs_nargs
- (AliasConstructor const_info.cs_cstr)
- pb.history in
+ push_history_pattern const_info.cs_nargs const_info.cs_cstr pb.history in
- (* We find matching clauses *)
+ (* 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 *)
+ (* that had matched constructor C *)
let cs_args = const_info.cs_args in
- let names = get_names pb.env cs_args eqns in
- let submat = List.map (fun (tms,eqn) -> prepend_pattern tms eqn) eqns in
+ let names,aliasname = get_names pb.env cs_args eqns in
let typs = List.map2 (fun (_,c,t) na -> (na,c,t)) cs_args names in
+ (* We build the matrix obtained by expanding the matching on *)
+ (* "C x1..xn as x" followed by a residual matching on eqn into *)
+ (* a matching on "x1 .. xn eqn" *)
+ let submat = List.map (fun (tms,_,eqn) -> prepend_pattern tms eqn) eqns in
+
+ (* 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
+
+ let extenv = push_rels typs pb.env in
+
+ let typs' =
+ List.map (fun (c,d) ->
+ (c,extract_inductive_data extenv !(pb.evdref) d,d)) typs' in
+
+ (* We compute over which of x(i+1)..xn and x matching on xi will need a *)
+ (* generalization *)
let dep_sign =
find_dependencies_signature
(dependencies_in_rhs const_info.cs_nargs current pb.tomatch eqns)
- (List.rev typs) in
+ (List.rev typs') in
(* The dependent term to subst in the types of the remaining UnPushed
terms is relative to the current context enriched by topushs *)
let ci = build_dependent_constructor const_info in
- (* We replace [(mkRel 1)] by its expansion [ci] *)
- (* and context "Gamma = Gamma1, current, Gamma2" by "Gamma;typs;curalias" *)
- (* This is done in two steps : first from "Gamma |- tms" *)
- (* into "Gamma; typs; curalias |- tms" *)
- let tomatch = lift_tomatch_stack const_info.cs_nargs pb.tomatch in
+ (* Current context Gamma has the form Gamma1;cur:I(realargs);Gamma2 *)
+ (* We go from Gamma |- PI tms. pred to *)
+ (* Gamma;x1..xn;curalias:I(x1..xn) |- PI tms'. pred' *)
+ (* where, in tms and pred, those realargs that are vars are *)
+ (* replaced by the corresponding xi and cur replaced by curalias *)
+ let cirealargs = Array.to_list const_info.cs_concl_realargs in
- let tomatch = match kind_of_term current with
- | Rel i -> replace_tomatch (i+const_info.cs_nargs) ci tomatch
- | _ -> (* non-rel initial term *) tomatch in
+ (* Do the specialization for terms to match *)
+ let tomatch = List.fold_right2 (fun par arg tomatch ->
+ match kind_of_term par with
+ | Rel i -> replace_tomatch (i+const_info.cs_nargs) arg tomatch
+ | _ -> tomatch) (current::realargs) (ci::cirealargs)
+ (lift_tomatch_stack const_info.cs_nargs pb.tomatch) in
let pred_is_not_dep =
noccur_predicate_between 1 (List.length realnames + 1) pb.pred tomatch in
let typs' =
- list_map2_i
- (fun i d deps ->
- let (na,c,t) = map_rel_declaration (lift i) d in
- let dep = match dep with
- | Name _ as na',k -> (if na <> Anonymous then na else na'),k
- | Anonymous,KnownNotDep ->
- if deps = [] && pred_is_not_dep then
- (Anonymous,KnownNotDep)
- else
- (force_name na,KnownDep)
- | _,_ -> anomaly "Inconsistent dependency" in
- ((mkRel i, NotInd (c,t)),deps,dep))
- 1 typs (List.rev dep_sign) in
-
+ 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
+ ((tm,tmtyp),deps,na))
+ typs' (List.rev dep_sign) in
+
+ (* Do the specialization for the predicate *)
let pred =
- specialize_predicate typs' (realnames,dep) arsign const_info tomatch pb.pred in
+ specialize_predicate typs' (realnames,curname) arsign const_info tomatch pb.pred in
let currents = List.map (fun x -> Pushed x) typs' in
- let ind =
- appvect (
- applist (mkInd (inductive_of_constructor const_info.cs_cstr),
- List.map (lift const_info.cs_nargs) const_info.cs_params),
- const_info.cs_concl_realargs) in
+ let alias =
+ if aliasname = Anonymous then
+ NonDepAlias
+ else
+ let cur_alias = lift const_info.cs_nargs current in
+ let ind =
+ appvect (
+ applist (mkInd (inductive_of_constructor 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
- let cur_alias = lift (List.length typs) current in
- let currents = Alias (ci,cur_alias,alias_type,ind) :: currents in
- let tomatch = List.rev_append currents tomatch in
+ let tomatch = List.rev_append (alias :: currents) tomatch in
let submat = adjust_impossible_cases pb pred tomatch submat in
if submat = [] then
@@ -1153,7 +1209,7 @@ let build_branch current deps (realnames,dep) pb arsign eqns const_info =
typs,
{ pb with
- env = push_rels typs pb.env;
+ env = extenv;
tomatch = tomatch;
pred = pred;
history = history;
@@ -1162,11 +1218,11 @@ let build_branch current deps (realnames,dep) pb arsign eqns const_info =
(**********************************************************************
INVARIANT:
- pb = { env, subst, tomatch, mat, ...}
- tomatch = list of Pushed (c:T) or Abstract (na:T) or Alias (c:T)
+ pb = { env, pred, tomatch, mat, ...}
+ tomatch = list of Pushed (c:T), Abstract (na:T), Alias (c:T) or NonDepAlias
- "Pushed" terms and types are relative to env
- "Abstract" types are relative to env enriched by the previous terms to match
+ all terms and types in Pushed, Abstract and Alias are relative to env
+ enriched by the Abstract coming before
*)
@@ -1174,9 +1230,10 @@ let build_branch current deps (realnames,dep) pb arsign eqns const_info =
(* 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
- | (Abstract d)::rest -> compile_generalization pb d rest
+ | Pushed cur :: rest -> match_current { pb with tomatch = rest } cur
+ | Alias x :: rest -> compile_alias 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 *)
@@ -1187,63 +1244,108 @@ and match_current pb tomatch =
match typ with
| NotInd (_,typ) ->
check_all_variables typ pb.mat;
- compile (shift_problem tomatch pb)
+ shift_problem tomatch pb
| IsInd (_,(IndType(indf,realargs) as indt),names) ->
let mind,_ = dest_ind_family indf 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
- compile (shift_problem tomatch pb)
+ shift_problem tomatch pb
else
- let _constraints = Array.map (solve_constraints indt) cstrs in
-
(* We generalize over terms depending on current term to match *)
- let pb = generalize_problem (names,dep) pb deps in
+ let pb,deps = generalize_problem (names,dep) pb deps in
(* We compile branches *)
- let brs = array_map2 (compile_branch current (names,dep) deps pb arsign) eqns cstrs in
-
+ let brvals = array_map2 (compile_branch current realargs (names,dep) deps pb arsign) eqns cstrs in
(* We build the (elementary) case analysis *)
- let brvals = Array.map (fun (v,_) -> v) brs in
- let (pred,typ,s) =
+ let depstocheck = current::binding_vars_of_inductive typ in
+ let brvals,tomatch,pred,inst =
+ postprocess_dependencies !(pb.evdref) depstocheck
+ brvals pb.tomatch pb.pred deps cstrs in
+ let brvals = Array.map (fun (sign,body) ->
+ it_mkLambda_or_LetIn body sign) brvals in
+ let (pred,typ) =
find_predicate pb.caseloc pb.env pb.evdref
- pb.pred current indt (names,dep) pb.tomatch in
+ pred current indt (names,dep) tomatch in
let ci = make_case_info pb.env mind pb.casestyle in
- let case = mkCase (ci,nf_betaiota Evd.empty pred,current,brvals) in
- let inst = List.map mkRel deps in
+ let pred = nf_betaiota !(pb.evdref) pred in
+ let case = mkCase (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 }
-and compile_branch current names deps pb arsign eqn cstr =
- let sign, pb = build_branch current deps names pb arsign eqn cstr in
+(* Building the sub-problem when all patterns are variables *)
+and shift_problem ((current,t),_,na) pb =
+ let ty = type_of_tomatch t in
+ let tomatch = lift_tomatch_stack 1 pb.tomatch in
+ let pred = specialize_predicate_var (current,t,na) pb.tomatch pb.pred in
+ let pb =
+ { pb with
+ env = push_rel (na,Some current,ty) pb.env;
+ tomatch = tomatch;
+ pred = lift_predicate 1 pred tomatch;
+ history = pop_history pb.history;
+ mat = List.map (push_current_pattern (current,ty)) pb.mat } in
let j = compile pb in
- (it_mkLambda_or_LetIn j.uj_val sign, j.uj_type)
+ { uj_val = subst1 current j.uj_val;
+ uj_type = subst1 current j.uj_type }
-and compile_generalization pb d rest =
+(* 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
+ sign, (compile pb).uj_val
+
+(* Abstract over a declaration before continuing splitting *)
+and compile_generalization pb i d rest =
let pb =
{ pb with
env = push_rel d pb.env;
tomatch = rest;
- mat = List.map (push_rels_eqn [d]) pb.mat } in
+ mat = List.map (push_generalized_decl_eqn pb.env i d) pb.mat } in
let j = compile pb in
{ uj_val = mkLambda_or_LetIn d j.uj_val;
- uj_type = mkProd_or_LetIn d j.uj_type }
+ uj_type = mkProd_wo_LetIn d j.uj_type }
-and compile_alias pb aliases rest =
- let history = simplify_history pb.history in
- let sign, newenv, mat = insert_aliases pb.env !(pb.evdref) aliases pb.mat in
- let n = List.length sign in
- let tomatch = lift_tomatch_stack n rest in
+and compile_alias pb (na,orig,(expanded,expanded_typ)) rest =
+ let f c t =
+ let alias = (na,Some c,t) in
+ let pb =
+ { pb with
+ env = push_rel alias pb.env;
+ tomatch = lift_tomatch_stack 1 rest;
+ pred = lift_predicate 1 pb.pred pb.tomatch;
+ history = pop_history_pattern pb.history;
+ mat = List.map (push_alias_eqn alias) pb.mat } in
+ let j = compile pb in
+ { uj_val =
+ if isRel c || isVar c || count_occurrences (mkRel 1) j.uj_val <= 1 then
+ subst1 c j.uj_val
+ else
+ mkLetIn (na,c,t,j.uj_val);
+ uj_type = subst1 c j.uj_type } in
+ if isRel orig or isVar orig then
+ (* Try to compile first using non expanded alias *)
+ try f orig (Retyping.get_type_of pb.env !(pb.evdref) orig)
+ with e when precatchable_exception e ->
+ (* Try then to compile using expanded alias *)
+ 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)
+
+
+(* Remember that a non-trivial pattern has been consumed *)
+and compile_non_dep_alias pb rest =
let pb =
- {pb with
- env = newenv;
- tomatch = tomatch;
- pred = lift_predicate n pb.pred tomatch;
- history = history;
- mat = mat } in
- let j = compile pb in
- List.fold_left mkSpecialLetInJudge j sign
+ { pb with
+ tomatch = rest;
+ history = pop_history_pattern pb.history;
+ mat = List.map drop_alias_eqn pb.mat } in
+ compile pb
(* pour les alias des initiaux, enrichir les env de ce qu'il faut et
substituer après par les initiaux *)
@@ -1260,7 +1362,7 @@ let matx_of_eqns env tomatchl eqns =
let initial_rhs = rhs in
let rhs =
{ rhs_env = env;
- rhs_vars = free_rawvars initial_rhs;
+ rhs_vars = free_glob_vars initial_rhs;
avoid_ids = ids@(ids_of_named_context (named_context env));
it = Some initial_rhs } in
{ patterns = initial_lpat;
@@ -1273,7 +1375,7 @@ let matx_of_eqns env tomatchl eqns =
(***************** Building an inversion predicate ************************)
(* Let "match t1 in I1 u11..u1n_1 ... tm in Im um1..umn_m with ... end : T"
- be a pattern-matching problem. We assume that the each uij can be
+ be a pattern-matching problem. We assume that each uij can be
decomposed under the form pij(vij1..vijq_ij) where pij(aij1..aijq_ij)
is a pattern depending on some variables aijk and the vijk are
instances of these variables. We also assume that each ti has the
@@ -1314,26 +1416,23 @@ let adjust_to_extended_env_and_remove_deps env extenv subst t =
(* \--env-/ (= x:ty) *)
(* \--------------extenv------------/ *)
let (p,_,_) = lookup_rel_id x (rel_context extenv) in
- let rec aux n (_,b,ty) =
- match b with
- | Some c ->
- assert (isRel c);
- let p = n + destRel c in aux p (lookup_rel p extenv)
- | None ->
- (n,ty) in
- let (p,ty) = aux p (lookup_rel p extenv) in
- if noccur_between_without_evar 1 (n'-p-n+1) ty
- then
- let u = lift (n'-n) u in
- (p,u,(expand_vars_in_term extenv u,lift p ty))
- else
- failwith "") subst 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 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
+
(* 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
* instances for variables aijk and bi.
@@ -1359,46 +1458,62 @@ let abstract_tycon loc env evdref subst _tycon extenv t =
by an evar that may depend (and only depend) on the corresponding
convertible subterms of the substitution *)
let rec aux (k,env,subst as x) t =
- if isRel t && pi2 (lookup_rel (destRel t) env) <> None then
- map_constr_with_full_binders push_binder aux x t
- else
+ 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
+ | Evar ev ->
+ let ty = get_type_of env sigma t in
+ let inst =
+ 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 good = List.filter (fun (_,u,_) -> is_conv_leq env sigma t u) subst in
if good <> [] then
- let (u,ty) = pi3 (List.hd good) in
+ let u = pi3 (List.hd good) in (* 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 depvl = free_rels ty in
let inst =
list_map_i
(fun i _ -> if List.mem i vl then u else mkRel i) 1
(rel_context extenv) in
let rel_filter =
- List.map (fun a -> not (isRel a) or dependent a u) inst in
+ List.map (fun a -> not (isRel a) || dependent a u
+ || Intset.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 candidates = u :: List.map mkRel vl in
let ev =
- e_new_evar evdref extenv ~src:(loc, CasesType) ~filter:filter ty in
- evdref := add_conv_pb (Reduction.CONV,extenv,substl inst ev,u) !evdref;
+ e_new_evar evdref extenv ~src:(loc, CasesType) ~filter ~candidates ty in
lift k ev
else
map_constr_with_full_binders push_binder aux x t in
aux (0,extenv,subst0) t0
let build_tycon loc env tycon_env subst tycon extenv evdref t =
- let t = match t with
+ let t,tt = match t with
| None ->
(* This is the situation we are building a return predicate and
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) (new_Type ()) in
- lift (n'-n) impossible_case_type
- | Some t -> abstract_tycon loc tycon_env evdref subst tycon extenv t in
- try get_judgment_of extenv !evdref t
- with Not_found | Anomaly _ ->
- (* Quick workaround to acknowledge failure to build a well-typed pred *)
- error "Unable to infer a well-typed return clause."
+ e_new_evar evdref env ~src:(loc,ImpossibleCase) tt in
+ (lift (n'-n) impossible_case_type, tt)
+ | 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
+ evdref := evd;
+ (t,tt) in
+ { uj_val = t; uj_type = tt }
(* For a multiple pattern-matching problem Xi on t1..tn with return
* type T, [build_inversion_problem Gamma Sigma (t1..tn) T] builds a return
@@ -1432,12 +1547,16 @@ let build_inversion_problem loc env sigma tms t =
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 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 ->
- aux n env acc_sign tms acc in
+ let pat,acc = make_patvar t acc in
+ let d = (alias_of_pat pat,None,t) 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
(* [patl] is a list of patterns revealing the substructure of
constructors present in the constraints on the type of the
@@ -1450,16 +1569,22 @@ let build_inversion_problem loc env sigma tms t =
problem have to be abstracted *)
let patl,sign,(subst,avoid) = aux 0 env [] tms ([],avoid0) in
let n = List.length sign in
- let (pb_env,_),sub_tms =
- list_fold_map (fun (env,i) (na,b,t as d) ->
- let typ =
- if b<>None then NotInd(None,t) else
- try try_find_ind env sigma t None
- with Not_found -> NotInd (None,t) in
- let ty = lift_tomatch_type (n-i) typ in
- let tm = Pushed ((mkRel (n-i),ty),[],(Anonymous,KnownNotDep)) in
- ((push_rel d env,i+1),tm))
- (env,0) (List.rev sign) in
+
+ let decls =
+ 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 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 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))
+ 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
serves as skeleton for the return type: [patl] is the
@@ -1508,26 +1633,14 @@ let build_inversion_problem loc env sigma tms t =
(* Here, [pred] is assumed to be in the context built from all *)
(* realargs and terms to match *)
-let build_initial_predicate knowndep allnames pred =
- let nar = List.fold_left (fun n names -> List.length names + n) 0 allnames in
- let rec buildrec n pred nal = function
- | [] -> List.rev nal,pred
- | names::lnames ->
- let names' = List.tl names in
- let n' = n + List.length names' in
- let pred, p =
- if dependent (mkRel (nar-n')) pred then pred, 1
- else liftn (-1) (nar-n') pred, 0 in
- let na =
- if p=1 then
- let na = List.hd names in
- ((if na = Anonymous then
- (* can happen if evars occur in the return clause *)
- Name (id_of_string "x") (*Hum*)
- else na),knowndep)
- else (Anonymous,KnownNotDep) in
- buildrec (n'+1) pred (na::nal) lnames
- in buildrec 0 pred [] allnames
+let build_initial_predicate arsign pred =
+ let rec buildrec n pred tmnames = function
+ | [] -> List.rev tmnames,pred
+ | ((na,c,t)::realdecls)::lnames ->
+ let n' = n + List.length realdecls in
+ buildrec (n'+1) pred (force_name na::tmnames) lnames
+ | _ -> assert false
+ in buildrec 0 pred [] (List.rev arsign)
let extract_arity_signature env0 tomatchl tmsign =
let get_one_sign n tm (na,t) =
@@ -1553,19 +1666,11 @@ let extract_arity_signature env0 tomatchl tmsign =
anomaly "Ill-formed 'in' clause in cases";
List.rev realnal
| None -> list_make nrealargs_ctxt Anonymous in
-(* let na = *)
-(* match na with *)
-(* | Name _ -> na *)
-(* | Anonymous -> *)
-(* match kind_of_term term with *)
-(* | Rel n -> pi1 (lookup_rel n (Environ.rel_context env0)) *)
-(* | _ -> 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
| [],[] -> []
- | (_,tm)::ltm, x::tmsign ->
+ | (_,tm)::ltm, (_,x)::tmsign ->
let l = get_one_sign n tm x in
l :: buildrec (n + List.length l) (ltm,tmsign)
| _ -> assert false
@@ -1581,7 +1686,7 @@ let inh_conv_coerce_to_tycon loc env evdref j tycon =
(* We put the tycon inside the arity signature, possibly discovering dependencies. *)
-let prepare_predicate_from_arsign_tycon loc tomatchs sign arsign c =
+let prepare_predicate_from_arsign_tycon loc tomatchs arsign c =
let nar = List.fold_left (fun n sign -> List.length sign + n) 0 arsign in
let subst, len =
List.fold_left2 (fun (subst, len) (tm, tmtype) sign ->
@@ -1635,9 +1740,7 @@ let prepare_predicate_from_arsign_tycon loc tomatchs sign arsign c =
* tycon to make the predicate if it is not closed.
*)
-let prepare_predicate loc typing_fun evdref env tomatchs sign tycon pred =
- let arsign = extract_arity_signature env tomatchs sign in
- let names = List.rev (List.map (List.map pi1) arsign) in
+let prepare_predicate loc typing_fun sigma env tomatchs arsign tycon pred =
let preds =
match pred, tycon with
(* No type annotation *)
@@ -1646,40 +1749,40 @@ let prepare_predicate loc typing_fun evdref env tomatchs sign tycon pred =
(* two different strategies *)
(* First strategy: we abstract the tycon wrt to the dependencies *)
let pred1 =
- prepare_predicate_from_arsign_tycon loc tomatchs sign arsign t in
+ prepare_predicate_from_arsign_tycon loc tomatchs arsign t in
(* Second strategy: we build an "inversion" predicate *)
- let sigma2,pred2 = build_inversion_problem loc env !evdref tomatchs t in
- [!evdref, KnownDep, pred1; sigma2, DepUnknown, pred2]
- | None, Some (None, t) ->
- (* First strategy: we build an "inversion" predicate *)
- let sigma1,pred = build_inversion_problem loc env !evdref tomatchs t in
- (* Second strategy: we abstract the tycon wrt to the dependencies *)
- let pred2 = lift (List.length names) t in
- [sigma1, DepUnknown, pred; !evdref, KnownNotDep, pred2]
- | None, _ ->
- (* No type constaints: we use two strategies *)
- let t = mkExistential env ~src:(loc, CasesType) evdref in
- (* First strategy: we build an inversion problem *)
- let sigma1,pred1 = build_inversion_problem loc env !evdref tomatchs t in
+ let sigma2,pred2 = build_inversion_problem loc env sigma tomatchs t in
+ [sigma, pred1; sigma2, pred2]
+ | None, _ ->
+ (* 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
+ (* First strategy: we build an "inversion" predicate *)
+ let sigma1,pred1 = build_inversion_problem loc env sigma tomatchs t in
(* Second strategy: we directly use the evar as a non dependent pred *)
- let pred2 = lift (List.length names) t in
- [sigma1, DepUnknown, pred1; !evdref, KnownNotDep, pred2]
+ let pred2 = lift (List.length arsign) t in
+ [sigma1, pred1; sigma, pred2]
(* Some type annotation *)
| Some rtntyp, _ ->
(* We extract the signature of the arity *)
- let env = List.fold_right push_rels arsign env in
- let predcclj = typing_fun (mk_tycon (new_Type ())) env evdref rtntyp in
- Option.iter (fun tycon ->
- let tycon = lift_tycon_type (List.length arsign) tycon in
- evdref :=
- Coercion.inh_conv_coerces_to loc env !evdref predcclj.uj_val tycon)
- tycon;
- let predccl = (j_nf_evar !evdref predcclj).uj_val in
- [!evdref, KnownDep, predccl]
+ let envar = List.fold_right push_rels arsign env in
+ let sigma, newt = new_sort_variable 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 predccl = (j_nf_evar sigma predcclj).uj_val in
+ [sigma, predccl]
in
List.map
- (fun (sigma,dep,pred) ->
- let (nal,pred) = build_initial_predicate dep names pred in
+ (fun (sigma,pred) ->
+ let (nal,pred) = build_initial_predicate arsign pred in
sigma,nal,pred)
preds
@@ -1698,15 +1801,33 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e
(* 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 sign = List.map snd tomatchl in
- let preds = prepare_predicate loc typing_fun evdref env tomatchs sign tycon predopt in
+ let arsign = extract_arity_signature env tomatchs tomatchl in
+ let preds = prepare_predicate loc typing_fun !evdref env 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 *)
(* names of aliases will be recovered from patterns (hence Anonymous *)
(* here) *)
- let initial_pushed = List.map2 (fun tm na -> Pushed(tm,[],na)) tomatchs nal in
+ 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 sigma 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 x) typs' in
(* A typing function that provides with a canonical term for absurd cases*)
let typing_fun tycon env evdref = function
diff --git a/pretyping/cases.mli b/pretyping/cases.mli
index 2711276a..f773da0e 100644
--- a/pretyping/cases.mli
+++ b/pretyping/cases.mli
@@ -1,24 +1,23 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: cases.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
open Util
open Names
open Term
open Evd
open Environ
open Inductiveops
-open Rawterm
+open Glob_term
open Evarutil
-(*i*)
+(** {5 Compilation of pattern-matching } *)
+
+(** {6 Pattern-matching errors } *)
type pattern_matching_error =
| BadPattern of constructor * constr
| BadConstructor of constructor * inductive
@@ -46,27 +45,14 @@ val error_wrong_predicate_arity_loc : loc -> env -> constr -> constr -> constr -
val error_needs_inversion : env -> constr -> types -> 'a
-val set_impossible_default_clause : constr * types -> unit
-(*s Compilation of pattern-matching. *)
-
-type alias_constr =
- | DepAlias
- | NonDepAlias
-type dep_status = KnownDep | KnownNotDep | DepUnknown
-type tomatch_type =
- | IsInd of types * inductive_type * name list
- | NotInd of constr option * types
-type tomatch_status =
- | Pushed of ((constr * tomatch_type) * int list * (name * dep_status))
- | Alias of (constr * constr * alias_constr * constr)
- | Abstract of rel_declaration
+(** {6 Compilation primitive. } *)
module type S = sig
val compile_cases :
loc -> case_style ->
- (type_constraint -> env -> evar_map ref -> rawconstr -> unsafe_judgment) * evar_map ref ->
+ (type_constraint -> env -> evar_map ref -> glob_constr -> unsafe_judgment) * evar_map ref ->
type_constraint ->
- env -> rawconstr option * tomatch_tuples * cases_clauses ->
+ env -> glob_constr option * tomatch_tuples * cases_clauses ->
unsafe_judgment
end
diff --git a/pretyping/cbv.ml b/pretyping/cbv.ml
index 802f9c58..ad33bae1 100644
--- a/pretyping/cbv.ml
+++ b/pretyping/cbv.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: cbv.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
open Util
open Pp
open Term
@@ -94,13 +92,6 @@ let rec shift_value n = function
let shift_value n v =
if n = 0 then v else shift_value n v
-let rec shift_stack n = function
- TOP -> TOP
- | APP(v,stk) -> APP(Array.map (shift_value n) v, shift_stack n stk)
- | CASE(c,b,i,s,stk) -> CASE(c,b,i,subs_shft(n,s), shift_stack n stk)
-let shift_stack n stk =
- if n = 0 then stk else shift_stack n stk
-
(* Contracts a fixpoint: given a fixpoint and a bindings,
* returns the corresponding fixpoint body, and the bindings in which
* it should be evaluated: its first variables are the fixpoint bodies
@@ -299,7 +290,7 @@ and cbv_stack_term info stack env t =
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),[||]), CASE(_,br,_,env,stk))
when red_set (info_flags info) fIOTA ->
cbv_stack_term info stk env br.(n-1)
@@ -365,14 +356,14 @@ and cbv_norm_value info = function (* reduction under binders *)
(* with profiling *)
let cbv_norm infos constr =
- with_stats (lazy (cbv_norm_term infos (ESID 0) constr))
+ with_stats (lazy (cbv_norm_term infos (subs_id 0) constr))
type cbv_infos = cbv_value infos
(* constant bodies are normalized at the first expansion *)
let create_cbv_infos flgs env sigma =
create
- (fun old_info c -> cbv_stack_term old_info TOP (ESID 0) c)
+ (fun old_info c -> cbv_stack_term old_info TOP (subs_id 0) c)
flgs
env
(Reductionops.safe_evar_value sigma)
diff --git a/pretyping/cbv.mli b/pretyping/cbv.mli
index e66e6dc6..acaf7e49 100644
--- a/pretyping/cbv.mli
+++ b/pretyping/cbv.mli
@@ -1,32 +1,28 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: cbv.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
open Names
open Term
open Environ
open Closure
open Esubst
-(*i*)
-(************************************************************************)
-(*s Call-by-value reduction *)
+(***********************************************************************
+ s Call-by-value reduction *)
-(* Entry point for cbv normalization of a constr *)
+(** Entry point for cbv normalization of a constr *)
type cbv_infos
val create_cbv_infos : RedFlags.reds -> env -> Evd.evar_map -> cbv_infos
val cbv_norm : cbv_infos -> constr -> constr
-(************************************************************************)
-(*i This is for cbv debug *)
+(***********************************************************************
+ i This is for cbv debug *)
type cbv_value =
| VAL of int * constr
| STACK of int * cbv_value * cbv_stack
@@ -46,7 +42,7 @@ val shift_value : int -> cbv_value -> cbv_value
val stack_app : cbv_value array -> cbv_stack -> cbv_stack
val strip_appl : cbv_value -> cbv_stack -> cbv_value * cbv_stack
-(* recursive functions... *)
+(** recursive functions... *)
val cbv_stack_term : cbv_infos ->
cbv_stack -> cbv_value subs -> constr -> cbv_value
val cbv_norm_term : cbv_infos -> cbv_value subs -> constr -> constr
@@ -54,4 +50,5 @@ val norm_head : cbv_infos ->
cbv_value subs -> constr -> cbv_stack -> cbv_value * cbv_stack
val apply_stack : cbv_infos -> constr -> cbv_stack -> constr
val cbv_norm_value : cbv_infos -> cbv_value -> constr
-(* End of cbv debug section i*)
+
+(** End of cbv debug section i*)
diff --git a/pretyping/classops.ml b/pretyping/classops.ml
index 504d01af..62d774bd 100644
--- a/pretyping/classops.ml
+++ b/pretyping/classops.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: classops.ml 14776 2011-12-07 17:54:28Z herbelin $ *)
-
open Util
open Pp
open Flags
@@ -19,7 +17,7 @@ open Libobject
open Library
open Term
open Termops
-open Rawterm
+open Glob_term
open Decl_kinds
open Mod_subst
@@ -48,6 +46,13 @@ type coe_info_typ = {
coe_is_identity : 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
+
type cl_index = int
type coe_index = coe_info_typ
@@ -134,9 +139,9 @@ let coercion_info coe = Gmap.find coe !coercion_tab
let coercion_exists coe = Gmap.mem coe !coercion_tab
-(* find_class_type : env -> evar_map -> constr -> cl_typ * int *)
+(* find_class_type : evar_map -> constr -> cl_typ * constr list *)
-let find_class_type env sigma t =
+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
@@ -154,7 +159,7 @@ let subst_cl_typ subst ct = match ct with
| CL_CONST kn ->
let kn',t = subst_con subst kn in
if kn' == kn then ct else
- fst (find_class_type (Global.env()) Evd.empty t)
+ fst (find_class_type Evd.empty t)
| CL_IND (kn,i) ->
let kn' = subst_ind subst kn in
if kn' == kn then ct else
@@ -169,12 +174,12 @@ let subst_coe_typ subst t = fst (subst_global subst t)
let class_of env sigma t =
let (t, n1, i, args) =
try
- let (cl,args) = find_class_type env sigma t in
+ let (cl,args) = find_class_type sigma t in
let (i, { cl_param = n1 } ) = class_info cl in
(t, n1, i, args)
with Not_found ->
let t = Tacred.hnf_constr env sigma t in
- let (cl, args) = find_class_type env sigma t in
+ let (cl, args) = find_class_type sigma t in
let (i, { cl_param = n1 } ) = class_info cl in
(t, n1, i, args)
in
@@ -182,7 +187,7 @@ let class_of env sigma t =
let inductive_class_of ind = fst (class_info (CL_IND ind))
-let class_args_of env sigma c = snd (find_class_type env sigma c)
+let class_args_of env sigma c = snd (find_class_type sigma c)
let string_of_class = function
| CL_FUN -> "Funclass"
@@ -211,14 +216,14 @@ let lookup_path_to_sort_from_class s =
let apply_on_class_of env sigma t cont =
try
- let (cl,args) = find_class_type env sigma t in
+ let (cl,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;
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 env sigma t in
+ let (cl, 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;
t, cont i
@@ -308,7 +313,7 @@ let add_coercion_in_graph (ic,source,target) =
try_add_new_path1 (s,target) (p@[ic]);
Gmap.iter
(fun (u,v) q ->
- if u<>v & (u = target) & (p <> q) then
+ if u<>v & 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;
@@ -344,6 +349,7 @@ open Goptions
let _ =
declare_bool_option
{ optsync = true;
+ optdepr = false;
optname = "automatic import of coercions";
optkey = ["Automatic";"Coercions";"Import"];
optread = (fun () -> !automatically_import_coercions);
@@ -400,7 +406,10 @@ let discharge_coercion (_,(coe,stre,isid,cls,clt,ps)) =
let classify_coercion (coe,stre,isid,cls,clt,ps as obj) =
if stre = Local then Dispose else Substitute obj
-let (inCoercion,_) =
+type coercion_obj =
+ coe_typ * Decl_kinds.locality * bool * cl_typ * cl_typ * int
+
+let inCoercion : coercion_obj -> obj =
declare_object {(default_object "COERCION") with
open_function = open_coercion;
load_function = load_coercion;
diff --git a/pretyping/classops.mli b/pretyping/classops.mli
index 964b44a0..66bb5c6c 100644
--- a/pretyping/classops.mli
+++ b/pretyping/classops.mli
@@ -1,14 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: classops.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
open Names
open Decl_kinds
open Term
@@ -16,9 +13,8 @@ open Evd
open Environ
open Nametab
open Mod_subst
-(*i*)
-(*s This is the type of class kinds *)
+(** {6 This is the type of class kinds } *)
type cl_typ =
| CL_SORT
| CL_FUN
@@ -28,53 +24,53 @@ type cl_typ =
val subst_cl_typ : substitution -> cl_typ -> cl_typ
-(* This is the type of infos for declared classes *)
+(** This is the type of infos for declared classes *)
type cl_info_typ = {
cl_param : int }
-(* This is the type of coercion kinds *)
+(** This is the type of coercion kinds *)
type coe_typ = Libnames.global_reference
-(* This is the type of infos for declared coercions *)
+(** This is the type of infos for declared coercions *)
type coe_info_typ
-(* [cl_index] is the type of class keys *)
+(** [cl_index] is the type of class keys *)
type cl_index
-(* [coe_index] is the type of coercion keys *)
+(** [coe_index] is the type of coercion keys *)
type coe_index
-(* This is the type of paths from a class to another *)
+(** This is the type of paths from a class to another *)
type inheritance_path = coe_index list
-(*s Access to classes infos *)
+(** {6 Access to classes infos } *)
val class_info : cl_typ -> (cl_index * cl_info_typ)
val class_exists : cl_typ -> bool
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
+(** [find_class_type env sigma c] returns the head reference of [c] and its
arguments *)
-val find_class_type : env -> evar_map -> types -> cl_typ * constr list
+val find_class_type : evar_map -> types -> cl_typ * constr list
-(* raises [Not_found] if not convertible to a class *)
+(** raises [Not_found] if not convertible to a class *)
val class_of : env -> evar_map -> types -> types * cl_index
-(* raises [Not_found] if not mapped to a class *)
+(** raises [Not_found] if not mapped to a class *)
val inductive_class_of : inductive -> cl_index
val class_args_of : env -> evar_map -> types -> constr list
-(*s [declare_coercion] adds a coercion in the graph of coercion paths *)
+(** {6 [declare_coercion] adds a coercion in the graph of coercion paths } *)
val declare_coercion :
coe_typ -> locality -> isid:bool ->
src:cl_typ -> target:cl_typ -> params:int -> unit
-(*s Access to coercions infos *)
+(** {6 Access to coercions infos } *)
val coercion_exists : coe_typ -> bool
val coercion_value : coe_index -> (unsafe_judgment * bool)
-(*s Lookup functions for coercion paths *)
+(** {6 Lookup functions for coercion paths } *)
val lookup_path_between_class : cl_index * cl_index -> inheritance_path
val lookup_path_between : env -> evar_map -> types * types ->
@@ -86,13 +82,14 @@ val lookup_path_to_sort_from : env -> evar_map -> types ->
val lookup_pattern_path_between :
inductive * inductive -> (constructor * int) list
-(*i Crade *)
+(**/**)
+(* Crade *)
open Pp
val install_path_printer :
((cl_index * cl_index) * inheritance_path -> std_ppcmds) -> unit
-(*i*)
+(**/**)
-(*s This is for printing purpose *)
+(** {6 This is for printing purpose } *)
val string_of_class : cl_typ -> string
val pr_class : cl_typ -> std_ppcmds
val pr_cl_index : cl_index -> std_ppcmds
@@ -101,6 +98,6 @@ val inheritance_graph : unit -> ((cl_index * cl_index) * inheritance_path) list
val classes : unit -> cl_typ list
val coercions : unit -> coe_index list
-(* [hide_coercion] returns the number of params to skip if the coercion must
+(** [hide_coercion] returns the number of params to skip if the coercion must
be hidden, [None] otherwise; it raises [Not_found] if not a coercion *)
val hide_coercion : coe_typ -> int option
diff --git a/pretyping/clenv.ml b/pretyping/clenv.ml
deleted file mode 100644
index abfef8d4..00000000
--- a/pretyping/clenv.ml
+++ /dev/null
@@ -1,482 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(* $Id: clenv.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
-open Pp
-open Util
-open Names
-open Nameops
-open Term
-open Termops
-open Namegen
-open Sign
-open Environ
-open Evd
-open Reduction
-open Reductionops
-open Rawterm
-open Pattern
-open Tacred
-open Pretype_errors
-open Evarutil
-open Unification
-open Mod_subst
-open Coercion.Default
-
-(* Abbreviations *)
-
-let pf_env gls = Global.env_of_context gls.it.evar_hyps
-let pf_hyps gls = named_context_of_val gls.it.evar_hyps
-let pf_type_of gls c = Typing.type_of (pf_env gls) gls.sigma c
-let pf_concl gl = gl.it.evar_concl
-
-(******************************************************************)
-(* Clausal environments *)
-
-type clausenv = {
- env : env;
- evd : evar_map;
- templval : constr freelisted;
- templtyp : constr freelisted }
-
-let cl_env ce = ce.env
-let cl_sigma ce = ce.evd
-
-let subst_clenv sub clenv =
- { templval = map_fl (subst_mps sub) clenv.templval;
- templtyp = map_fl (subst_mps sub) clenv.templtyp;
- evd = subst_evar_defs_light sub clenv.evd;
- env = clenv.env }
-
-let clenv_nf_meta clenv c = nf_meta clenv.evd c
-let clenv_term clenv c = meta_instance clenv.evd c
-let clenv_meta_type clenv mv = Typing.meta_type clenv.evd mv
-let clenv_value clenv = meta_instance clenv.evd clenv.templval
-let clenv_type clenv = meta_instance clenv.evd clenv.templtyp
-
-
-let clenv_hnf_constr ce t = hnf_constr (cl_env ce) (cl_sigma ce) t
-
-let clenv_get_type_of ce c = Retyping.get_type_of (cl_env ce) (cl_sigma ce) c
-
-exception NotExtensibleClause
-
-let clenv_push_prod cl =
- let typ = whd_betadeltaiota (cl_env cl) (cl_sigma cl) (clenv_type cl) in
- let rec clrec typ = match kind_of_term typ with
- | Cast (t,_,_) -> clrec t
- | Prod (na,t,u) ->
- let mv = new_meta () in
- let dep = dependent (mkRel 1) u in
- let na' = if dep then na else Anonymous in
- let e' = meta_declare mv t ~name:na' cl.evd in
- let concl = if dep then subst1 (mkMeta mv) u else u in
- let def = applist (cl.templval.rebus,[mkMeta mv]) in
- { templval = mk_freelisted def;
- templtyp = mk_freelisted concl;
- evd = e';
- env = cl.env }
- | _ -> raise NotExtensibleClause
- in clrec typ
-
-(* Instantiate the first [bound] products of [t] with metas (all products if
- [bound] is [None]; unfold local defs *)
-
-let clenv_environments evd bound t =
- let rec clrec (e,metas) n t =
- match n, kind_of_term t with
- | (Some 0, _) -> (e, List.rev metas, t)
- | (n, Cast (t,_,_)) -> clrec (e,metas) n t
- | (n, Prod (na,t1,t2)) ->
- let mv = new_meta () in
- let dep = dependent (mkRel 1) t2 in
- let na' = if dep then na else Anonymous in
- let e' = meta_declare mv t1 ~name:na' e in
- clrec (e', (mkMeta mv)::metas) (Option.map ((+) (-1)) n)
- (if dep then (subst1 (mkMeta mv) t2) else t2)
- | (n, LetIn (na,b,_,t)) -> clrec (e,metas) n (subst1 b t)
- | (n, _) -> (e, List.rev metas, t)
- in
- clrec (evd,[]) bound t
-
-(* Instantiate the first [bound] products of [t] with evars (all products if
- [bound] is [None]; unfold local defs *)
-
-let clenv_environments_evars env evd bound t =
- let rec clrec (e,ts) n t =
- match n, kind_of_term t with
- | (Some 0, _) -> (e, List.rev ts, t)
- | (n, Cast (t,_,_)) -> clrec (e,ts) n t
- | (n, Prod (na,t1,t2)) ->
- let e',constr = Evarutil.new_evar e env t1 in
- let dep = dependent (mkRel 1) t2 in
- clrec (e', constr::ts) (Option.map ((+) (-1)) n)
- (if dep then (subst1 constr t2) else t2)
- | (n, LetIn (na,b,_,t)) -> clrec (e,ts) n (subst1 b t)
- | (n, _) -> (e, List.rev ts, t)
- in
- clrec (evd,[]) bound t
-
-let clenv_conv_leq env sigma t c bound =
- let ty = Retyping.get_type_of env sigma c in
- let evd = Evd.create_goal_evar_defs sigma in
- let evars,args,_ = clenv_environments_evars env evd (Some bound) ty in
- let evars = Evarconv.the_conv_x_leq env t (applist (c,args)) evars in
- let evars = Evarconv.consider_remaining_unif_problems env evars in
- let args = List.map (whd_evar evars) args in
- check_evars env sigma evars (applist (c,args));
- args
-
-let mk_clenv_from_env environ sigma n (c,cty) =
- let evd = create_goal_evar_defs sigma in
- let (evd,args,concl) = clenv_environments evd n cty in
- { templval = mk_freelisted (match args with [] -> c | _ -> applist (c,args));
- templtyp = mk_freelisted concl;
- evd = evd;
- env = environ }
-
-let mk_clenv_from_n gls n (c,cty) =
- mk_clenv_from_env (Global.env_of_context gls.it.evar_hyps) gls.sigma n (c, cty)
-
-let mk_clenv_from gls = mk_clenv_from_n gls None
-
-let mk_clenv_rename_from_n gls n (c,t) =
- mk_clenv_from_n gls n (c,rename_bound_vars_as_displayed [] t)
-
-let mk_clenv_type_of gls t = mk_clenv_from gls (t,pf_type_of gls t)
-
-(******************************************************************)
-
-(* [mentions clenv mv0 mv1] is true if mv1 is defined and mentions
- * mv0, or if one of the free vars on mv1's freelist mentions
- * mv0 *)
-
-let mentions clenv mv0 =
- let rec menrec mv1 =
- mv0 = mv1 ||
- let mlist =
- try match meta_opt_fvalue clenv.evd mv1 with
- | Some (b,_) -> b.freemetas
- | None -> Metaset.empty
- with Not_found -> Metaset.empty in
- meta_exists menrec mlist
- in menrec
-
-let error_incompatible_inst clenv mv =
- let na = meta_name clenv.evd mv in
- match na with
- Name id ->
- errorlabstrm "clenv_assign"
- (str "An incompatible instantiation has already been found for " ++
- pr_id id)
- | _ ->
- anomaly "clenv_assign: non dependent metavar already assigned"
-
-(* TODO: replace by clenv_unify (mkMeta mv) rhs ? *)
-let clenv_assign mv rhs clenv =
- let rhs_fls = mk_freelisted rhs in
- if meta_exists (mentions clenv mv) rhs_fls.freemetas then
- error "clenv_assign: circularity in unification";
- try
- if meta_defined clenv.evd mv then
- if not (eq_constr (fst (meta_fvalue clenv.evd mv)).rebus rhs) then
- error_incompatible_inst clenv mv
- else
- clenv
- else
- let st = (ConvUpToEta 0,TypeNotProcessed) in
- {clenv with evd = meta_assign mv (rhs_fls.rebus,st) clenv.evd}
- with Not_found ->
- error "clenv_assign: undefined meta"
-
-
-
-(* [clenv_dependent hyps_only clenv]
- * returns a list of the metavars which appear in the template of clenv,
- * and which are dependent, This is computed by taking the metavars in cval,
- * in right-to-left order, and collecting the metavars which appear
- * in their types, and adding in all the metavars appearing in the
- * type of clenv.
- * If [hyps_only] then metavariables occurring in the type are _excluded_ *)
-
-(* [clenv_metavars clenv mv]
- * returns a list of the metavars which appear in the type of
- * the metavar mv. The list is unordered. *)
-
-let clenv_metavars evd mv =
- (mk_freelisted (meta_instance evd (meta_ftype evd mv))).freemetas
-
-let dependent_metas clenv mvs conclmetas =
- List.fold_right
- (fun mv deps ->
- Metaset.union deps (clenv_metavars clenv.evd mv))
- mvs conclmetas
-
-let duplicated_metas c =
- let rec collrec (one,more as acc) c =
- match kind_of_term c with
- | Meta mv -> if List.mem mv one then (one,mv::more) else (mv::one,more)
- | _ -> fold_constr collrec acc c
- in
- snd (collrec ([],[]) c)
-
-let clenv_dependent hyps_only clenv =
- let mvs = undefined_metas clenv.evd in
- let ctyp_mvs = (mk_freelisted (clenv_type clenv)).freemetas in
- let deps = dependent_metas clenv mvs ctyp_mvs in
- let nonlinear = duplicated_metas (clenv_value clenv) in
- (* Make the assumption that duplicated metas have internal dependencies *)
- List.filter
- (fun mv -> if Metaset.mem mv deps
- then not (hyps_only && Metaset.mem mv ctyp_mvs)
- else List.mem mv nonlinear)
- mvs
-
-let clenv_missing ce = clenv_dependent true ce
-
-(******************************************************************)
-
-let clenv_unify allow_K ?(flags=default_unify_flags) cv_pb t1 t2 clenv =
- { clenv with
- evd = w_unify allow_K ~flags:flags clenv.env cv_pb t1 t2 clenv.evd }
-
-let clenv_unify_meta_types ?(flags=default_unify_flags) clenv =
- { clenv with evd = w_unify_meta_types ~flags:flags clenv.env clenv.evd }
-
-let clenv_unique_resolver allow_K ?(flags=default_unify_flags) clenv gl =
- if isMeta (fst (whd_stack clenv.evd clenv.templtyp.rebus)) then
- clenv_unify allow_K CUMUL ~flags:flags (clenv_type clenv) (pf_concl gl)
- (clenv_unify_meta_types ~flags:flags clenv)
- else
- clenv_unify allow_K CUMUL ~flags:flags
- (meta_reducible_instance clenv.evd clenv.templtyp) (pf_concl gl) clenv
-
-(* [clenv_pose_metas_as_evars clenv dep_mvs]
- * For each dependent evar in the clause-env which does not have a value,
- * pose a value for it by constructing a fresh evar. We do this in
- * left-to-right order, so that every evar's type is always closed w.r.t.
- * metas.
-
- * Node added 14/4/08 [HH]: before this date, evars were collected in
- clenv_dependent by collect_metas in the fold_constr order which is
- (almost) the left-to-right order of dependencies in term. However,
- due to K-redexes, collect_metas was sometimes missing some metas.
- The call to collect_metas has been replaced by a call to
- undefined_metas, but then the order was the one of definition of
- the metas (numbers in increasing order) which is _not_ the
- dependency order when a clenv_fchain occurs (because clenv_fchain
- plugs a term with a list of consecutive metas in place of a - a priori -
- arbitrary metavariable belonging to another sequence of consecutive metas:
- e.g., clenv_fchain may plug (H ?1 ?2) at the position ?6 of
- (nat_ind ?3 ?4 ?5 ?6), leading to a dependency order 3<4<5<1<2).
- To ensure the dependency order, we check that the type of each meta
- to pose is already meta-free, otherwise we postpone the transformation,
- hoping that no cycle may happen.
-
- Another approach could have been to use decimal numbers for metas so that
- in the example above, (H ?1 ?2) would have been renumbered (H ?6.1 ?6.2)
- then making the numeric order match the dependency order.
-*)
-
-let clenv_pose_metas_as_evars clenv dep_mvs =
- let rec fold clenv = function
- | [] -> clenv
- | mv::mvs ->
- let ty = clenv_meta_type clenv mv in
- (* Postpone the evar-ization if dependent on another meta *)
- (* This assumes no cycle in the dependencies - is it correct ? *)
- if occur_meta ty then fold clenv (mvs@[mv])
- else
- let (evd,evar) =
- new_evar clenv.evd (cl_env clenv) ~src:(dummy_loc,GoalEvar) ty in
- let clenv = clenv_assign mv evar {clenv with evd=evd} in
- fold clenv mvs in
- fold clenv dep_mvs
-
-let evar_clenv_unique_resolver = clenv_unique_resolver
-
-(******************************************************************)
-
-let connect_clenv gls clenv =
- { clenv with
- evd = evars_reset_evd ~with_conv_pbs:true gls.sigma clenv.evd;
- env = Global.env_of_context gls.it.evar_hyps }
-
-(* [clenv_fchain mv clenv clenv']
- *
- * Resolves the value of "mv" (which must be undefined) in clenv to be
- * the template of clenv' be the value "c", applied to "n" fresh
- * metavars, whose types are chosen by destructing "clf", which should
- * be a clausale forme generated from the type of "c". The process of
- * resolution can cause unification of already-existing metavars, and
- * of the fresh ones which get created. This operation is a composite
- * of operations which pose new metavars, perform unification on
- * terms, and make bindings.
-
- Otherwise said, from
-
- [clenv] = [env;sigma;metas |- c:T]
- [clenv'] = [env';sigma';metas' |- d:U]
- [mv] = [mi] of type [Ti] in [metas]
-
- then, if the unification of [Ti] and [U] produces map [rho], the
- chaining is [env';sigma';rho'(metas),rho(metas') |- c:rho'(T)] for
- [rho'] being [rho;mi:=d].
-
- In particular, it assumes that [env'] and [sigma'] extend [env] and [sigma].
-*)
-let clenv_fchain ?(allow_K=true) ?(flags=default_unify_flags) mv clenv nextclenv =
- (* Add the metavars of [nextclenv] to [clenv], with their name-environment *)
- let clenv' =
- { templval = clenv.templval;
- templtyp = clenv.templtyp;
- evd = meta_merge nextclenv.evd clenv.evd;
- env = nextclenv.env } in
- (* unify the type of the template of [nextclenv] with the type of [mv] *)
- let clenv'' =
- clenv_unify allow_K ~flags:flags CUMUL
- (clenv_term clenv' nextclenv.templtyp)
- (clenv_meta_type clenv' mv)
- clenv' in
- (* assign the metavar *)
- let clenv''' =
- clenv_assign mv (clenv_term clenv' nextclenv.templval) clenv''
- in
- clenv'''
-
-(***************************************************************)
-(* Bindings *)
-
-type arg_bindings = constr explicit_bindings
-
-(* [clenv_independent clenv]
- * returns a list of metavariables which appear in the term cval,
- * and which are not dependent. That is, they do not appear in
- * the types of other metavars which are in cval, nor in the type
- * of cval, ctyp. *)
-
-let clenv_independent clenv =
- let mvs = collect_metas (clenv_value clenv) in
- let ctyp_mvs = (mk_freelisted (clenv_type clenv)).freemetas in
- let deps = dependent_metas clenv mvs ctyp_mvs in
- List.filter (fun mv -> not (Metaset.mem mv deps)) mvs
-
-let check_bindings bl =
- match list_duplicates (List.map pi2 bl) with
- | NamedHyp s :: _ ->
- errorlabstrm ""
- (str "The variable " ++ pr_id s ++
- str " occurs more than once in binding list.");
- | AnonHyp n :: _ ->
- errorlabstrm ""
- (str "The position " ++ int n ++
- str " occurs more than once in binding list.")
- | [] -> ()
-
-let meta_of_binder clause loc mvs = function
- | NamedHyp s -> meta_with_name clause.evd s
- | AnonHyp n ->
- try List.nth mvs (n-1)
- with (Failure _|Invalid_argument _) ->
- errorlabstrm "" (str "No such binder.")
-
-let error_already_defined b =
- match b with
- | NamedHyp id ->
- errorlabstrm ""
- (str "Binder name \"" ++ pr_id id ++
- str"\" already defined with incompatible value.")
- | AnonHyp n ->
- anomalylabstrm ""
- (str "Position " ++ int n ++ str" already defined.")
-
-let clenv_unify_binding_type clenv c t u =
- if isMeta (fst (whd_stack clenv.evd u)) then
- (* Not enough information to know if some subtyping is needed *)
- CoerceToType, clenv, c
- else
- (* Enough information so as to try a coercion now *)
- try
- let evd,c = w_coerce_to_type (cl_env clenv) clenv.evd c t u in
- TypeProcessed, { clenv with evd = evd }, c
- with
- | PretypeError (_,NotClean _) as e -> raise e
- | e when precatchable_exception e ->
- TypeNotProcessed, clenv, c
-
-let clenv_assign_binding clenv k c =
- let k_typ = clenv_hnf_constr clenv (clenv_meta_type clenv k) in
- let c_typ = nf_betaiota clenv.evd (clenv_get_type_of clenv c) in
- let status,clenv',c = clenv_unify_binding_type clenv c c_typ k_typ in
- { clenv' with evd = meta_assign k (c,(UserGiven,status)) clenv'.evd }
-
-let clenv_match_args bl clenv =
- if bl = [] then
- clenv
- else
- let mvs = clenv_independent clenv in
- check_bindings bl;
- List.fold_left
- (fun clenv (loc,b,c) ->
- let k = meta_of_binder clenv loc mvs b in
- if meta_defined clenv.evd k then
- if eq_constr (fst (meta_fvalue clenv.evd k)).rebus c then clenv
- else error_already_defined b
- else
- clenv_assign_binding clenv k c)
- clenv bl
-
-exception NoSuchBinding
-
-let clenv_constrain_last_binding c clenv =
- let all_mvs = collect_metas clenv.templval.rebus in
- let k = try list_last all_mvs with Failure _ -> raise NoSuchBinding in
- clenv_assign_binding clenv k c
-
-let clenv_constrain_dep_args hyps_only bl clenv =
- if bl = [] then
- clenv
- else
- let occlist = clenv_dependent hyps_only clenv in
- if List.length occlist = List.length bl then
- List.fold_left2 clenv_assign_binding clenv occlist bl
- else
- errorlabstrm ""
- (strbrk "Not the right number of missing arguments (expected " ++
- int (List.length occlist) ++ str ").")
-
-(****************************************************************)
-(* Clausal environment for an application *)
-
-
-let make_clenv_binding_gen hyps_only n env sigma (c,t) = function
- | ImplicitBindings largs ->
- let clause = mk_clenv_from_env env sigma n (c,t) in
- clenv_constrain_dep_args hyps_only largs clause
- | ExplicitBindings lbind ->
- let clause = mk_clenv_from_env env sigma n
- (c,rename_bound_vars_as_displayed [] t)
- in clenv_match_args lbind clause
- | NoBindings ->
- mk_clenv_from_env env sigma n (c,t)
-
-let make_clenv_binding_env_apply env sigma n =
- make_clenv_binding_gen true n env sigma
-
-let make_clenv_binding_apply gls n = make_clenv_binding_gen true n (pf_env gls) gls.sigma
-let make_clenv_binding gls = make_clenv_binding_gen false None (pf_env gls) gls.sigma
-
-(****************************************************************)
-(* Pretty-print *)
-
-let pr_clenv clenv =
- h 0
- (str"TEMPL: " ++ print_constr clenv.templval.rebus ++
- str" : " ++ print_constr clenv.templtyp.rebus ++ fnl () ++
- pr_evar_map clenv.evd)
diff --git a/pretyping/clenv.mli b/pretyping/clenv.mli
deleted file mode 100644
index 6fbb668b..00000000
--- a/pretyping/clenv.mli
+++ /dev/null
@@ -1,148 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(*i $Id: clenv.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
-open Util
-open Names
-open Term
-open Sign
-open Environ
-open Evd
-open Evarutil
-open Mod_subst
-open Rawterm
-open Unification
-(*i*)
-
-(***************************************************************)
-(* The Type of Constructions clausale environments. *)
-
-(* [env] is the typing context
- * [evd] is the mapping from metavar and evar numbers to their types
- * and values.
- * [templval] is the template which we are trying to fill out.
- * [templtyp] is its type.
- *)
-type clausenv = {
- env : env;
- evd : evar_map;
- templval : constr freelisted;
- templtyp : constr freelisted }
-
-(* Substitution is not applied on templenv (because [subst_clenv] is
- applied only on hints which typing env is overwritten by the
- goal env) *)
-val subst_clenv : substitution -> clausenv -> clausenv
-
-(* subject of clenv (instantiated) *)
-val clenv_value : clausenv -> constr
-(* type of clenv (instantiated) *)
-val clenv_type : clausenv -> types
-(* substitute resolved metas *)
-val clenv_nf_meta : clausenv -> constr -> constr
-(* type of a meta in clenv context *)
-val clenv_meta_type : clausenv -> metavariable -> types
-
-val mk_clenv_from : evar_info sigma -> constr * types -> clausenv
-val mk_clenv_from_n :
- evar_info sigma -> int option -> constr * types -> clausenv
-val mk_clenv_type_of : evar_info sigma -> constr -> clausenv
-val mk_clenv_from_env : env -> evar_map -> int option -> constr * types -> clausenv
-
-(***************************************************************)
-(* linking of clenvs *)
-
-val connect_clenv : evar_info sigma -> clausenv -> clausenv
-val clenv_fchain :
- ?allow_K:bool -> ?flags:unify_flags -> metavariable -> clausenv -> clausenv -> clausenv
-
-(***************************************************************)
-(* Unification with clenvs *)
-
-(* Unifies two terms in a clenv. The boolean is [allow_K] (see [Unification]) *)
-val clenv_unify :
- bool -> ?flags:unify_flags -> conv_pb -> constr -> constr -> clausenv -> clausenv
-
-(* unifies the concl of the goal with the type of the clenv *)
-val clenv_unique_resolver :
- bool -> ?flags:unify_flags -> clausenv -> evar_info sigma -> clausenv
-
-(* same as above ([allow_K=false]) but replaces remaining metas
- with fresh evars if [evars_flag] is [true] *)
-val evar_clenv_unique_resolver :
- bool -> ?flags:unify_flags -> clausenv -> evar_info sigma -> clausenv
-
-val clenv_dependent : bool -> clausenv -> metavariable list
-
-val clenv_pose_metas_as_evars : clausenv -> metavariable list -> clausenv
-
-(***************************************************************)
-(* Bindings *)
-
-type arg_bindings = constr explicit_bindings
-
-(* bindings where the key is the position in the template of the
- clenv (dependent or not). Positions can be negative meaning to
- start from the rightmost argument of the template. *)
-val clenv_independent : clausenv -> metavariable list
-val clenv_missing : clausenv -> metavariable list
-
-(** for the purpose of inversion tactics *)
-exception NoSuchBinding
-val clenv_constrain_last_binding : constr -> clausenv -> clausenv
-
-(* defines metas corresponding to the name of the bindings *)
-val clenv_match_args : arg_bindings -> clausenv -> clausenv
-
-val clenv_unify_meta_types : ?flags:unify_flags -> clausenv -> clausenv
-
-(* start with a clenv to refine with a given term with bindings *)
-
-(* the arity of the lemma is fixed *)
-(* the optional int tells how many prods of the lemma have to be used *)
-(* use all of them if None *)
-val make_clenv_binding_apply :
- evar_info sigma -> int option -> constr * constr -> constr bindings ->
- clausenv
-val make_clenv_binding_env_apply :
- env -> evar_map -> int option -> constr * constr -> constr bindings ->
- clausenv
-val make_clenv_binding :
- evar_info sigma -> constr * constr -> constr bindings -> clausenv
-
-(* [clenv_environments sigma n t] returns [sigma',lmeta,ccl] where
- [lmetas] is a list of metas to be applied to a proof of [t] so that
- it produces the unification pattern [ccl]; [sigma'] is [sigma]
- extended with [lmetas]; if [n] is defined, it limits the size of
- the list even if [ccl] is still a product; otherwise, it stops when
- [ccl] is not a product; example: if [t] is [forall x y, x=y -> y=x]
- and [n] is [None], then [lmetas] is [Meta n1;Meta n2;Meta n3] and
- [ccl] is [Meta n1=Meta n2]; if [n] is [Some 1], [lmetas] is [Meta n1]
- and [ccl] is [forall y, Meta n1=y -> y=Meta n1] *)
-val clenv_environments :
- evar_map -> int option -> types -> evar_map * constr list * types
-
-(* [clenv_environments_evars env sigma n t] does the same but returns
- a list of Evar's defined in [env] and extends [sigma] accordingly *)
-val clenv_environments_evars :
- env -> evar_map -> int option -> types -> evar_map * constr list * types
-
-(* [clenv_conv_leq env sigma t c n] looks for c1...cn s.t. [t <= c c1...cn] *)
-val clenv_conv_leq :
- env -> evar_map -> types -> constr -> int -> constr list
-
-(* if the clause is a product, add an extra meta for this product *)
-exception NotExtensibleClause
-val clenv_push_prod : clausenv -> clausenv
-
-(***************************************************************)
-(* Pretty-print *)
-val pr_clenv : clausenv -> Pp.std_ppcmds
-
diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml
index d0f20615..553c9127 100644
--- a/pretyping/coercion.ml
+++ b/pretyping/coercion.ml
@@ -1,11 +1,18 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: coercion.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
+
+(* 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 *)
+(* 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 Util
open Names
@@ -66,7 +73,7 @@ module type S = sig
pattern [pat] typed in [ind1] into a pattern typed in [ind2];
raises [Not_found] if no coercion found *)
val inh_pattern_coerce_to :
- loc -> Rawterm.cases_pattern -> inductive -> inductive -> Rawterm.cases_pattern
+ loc -> Glob_term.cases_pattern -> inductive -> inductive -> Glob_term.cases_pattern
end
module Default = struct
@@ -92,8 +99,8 @@ module Default = struct
let apply_pattern_coercion loc pat p =
List.fold_left
(fun pat (co,n) ->
- let f i = if i<n then Rawterm.PatVar (loc, Anonymous) else pat in
- Rawterm.PatCstr (loc, co, list_tabulate f (n+1), Anonymous))
+ 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 *)
@@ -130,8 +137,8 @@ module Default = struct
(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)
+ lookup_path_to_fun_from env evd j.uj_type in
+ (evd,apply_coercion env evd p j t)
let inh_app_fun env evd j =
try inh_app_fun env evd j
@@ -141,15 +148,15 @@ module Default = struct
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
+ 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
+ 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
+ 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) ->
@@ -195,6 +202,8 @@ module Default = struct
(* 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
@@ -204,7 +213,9 @@ module Default = struct
(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 = subst_term v1 t2 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
@@ -231,7 +242,14 @@ module Default = struct
let inh_conv_coerce_rigid_to = inh_conv_coerce_to_gen true
- let inh_conv_coerces_to loc env (evd : evar_map) t (abs, t') = evd
+ 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
diff --git a/pretyping/coercion.mli b/pretyping/coercion.mli
index 6ad1023e..5d814b29 100644
--- a/pretyping/coercion.mli
+++ b/pretyping/coercion.mli
@@ -1,14 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: coercion.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
open Util
open Evd
open Names
@@ -16,35 +13,34 @@ open Term
open Sign
open Environ
open Evarutil
-open Rawterm
-(*i*)
+open Glob_term
module type S = sig
- (*s Coercions. *)
+ (** {6 Coercions. } *)
- (* [inh_app_fun env isevars j] coerces [j] to a function; i.e. it
+ (** [inh_app_fun 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 *)
val inh_app_fun :
env -> evar_map -> unsafe_judgment -> evar_map * unsafe_judgment
- (* [inh_coerce_to_sort env isevars j] coerces [j] to a type; i.e. it
+ (** [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
+ (** [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 *)
+ (** [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 loc env isevars j t] coerces [j] to an object of type
+ (** [inh_conv_coerce_to 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 *)
val inh_conv_coerce_to : loc ->
@@ -53,13 +49,13 @@ module type S = sig
val inh_conv_coerce_rigid_to : 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]
+ (** [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
+ (** [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 :
diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml
index 14e72807..c194a0f2 100644
--- a/pretyping/detyping.ml
+++ b/pretyping/detyping.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: detyping.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
open Pp
open Util
open Univ
@@ -18,7 +16,7 @@ open Inductive
open Inductiveops
open Environ
open Sign
-open Rawterm
+open Glob_term
open Nameops
open Termops
open Namegen
@@ -47,34 +45,34 @@ let has_two_constructors lc =
let isomorphic_to_tuple lc = (Array.length lc = 1)
let encode_bool r =
- let (_,lc as x) = encode_inductive r in
+ let (x,lc) = encode_inductive r in
if not (has_two_constructors lc) then
user_err_loc (loc_of_reference r,"encode_if",
str "This type has not exactly two constructors.");
x
let encode_tuple r =
- let (_,lc as x) = encode_inductive r in
+ let (x,lc) = encode_inductive r in
if not (isomorphic_to_tuple lc) then
user_err_loc (loc_of_reference r,"encode_tuple",
str "This type cannot be seen as a tuple type.");
x
-module PrintingCasesMake =
+module PrintingInductiveMake =
functor (Test : sig
- val encode : reference -> inductive * int array
+ val encode : reference -> inductive
val member_message : std_ppcmds -> bool -> std_ppcmds
val field : string
val title : string
end) ->
struct
- type t = inductive * int array
+ type t = inductive
let encode = Test.encode
- let subst subst ((kn,i), ints as obj) =
+ let subst subst (kn, ints as obj) =
let kn' = subst_ind subst kn in
if kn' == kn then obj else
- (kn',i), ints
- let printer (ind,_) = pr_global_env Idset.empty (IndRef ind)
+ kn', ints
+ let printer ind = pr_global_env Idset.empty (IndRef ind)
let key = ["Printing";Test.field]
let title = Test.title
let member_message x = Test.member_message (printer x)
@@ -82,7 +80,7 @@ module PrintingCasesMake =
end
module PrintingCasesIf =
- PrintingCasesMake (struct
+ PrintingInductiveMake (struct
let encode = encode_bool
let field = "If"
let title = "Types leading to pretty-printing of Cases using a `if' form: "
@@ -94,7 +92,7 @@ module PrintingCasesIf =
end)
module PrintingCasesLet =
- PrintingCasesMake (struct
+ PrintingInductiveMake (struct
let encode = encode_tuple
let field = "Let"
let title =
@@ -118,6 +116,7 @@ let force_wildcard () = !wildcard_value
let _ = declare_bool_option
{ optsync = true;
+ optdepr = false;
optname = "forced wildcard";
optkey = ["Printing";"Wildcard"];
optread = force_wildcard;
@@ -128,6 +127,7 @@ let synthetize_type () = !synth_type_value
let _ = declare_bool_option
{ optsync = true;
+ optdepr = false;
optname = "pattern matching return type synthesizability";
optkey = ["Printing";"Synth"];
optread = synthetize_type;
@@ -138,6 +138,7 @@ let reverse_matching () = !reverse_matching_value
let _ = declare_bool_option
{ optsync = true;
+ optdepr = false;
optname = "pattern-matching reversibility";
optkey = ["Printing";"Matching"];
optread = reverse_matching;
@@ -168,8 +169,6 @@ let computable p k =
&&
noccur_between 1 (k+1) ccl
-let avoid_flag isgoal = if isgoal then Some true else None
-
let lookup_name_as_displayed env t s =
let rec lookup avoid n c = match kind_of_term c with
| Prod (name,_,c') ->
@@ -237,7 +236,7 @@ let rec decomp_branch n nal b (avoid,env as e) 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_nargs in
+ let cnl = ci.ci_cstr_ndecls in
List.flatten
(list_tabulate (fun i -> contract_branch isgoal e (cnl.(i),mkpat i,cl.(i)))
(Array.length cl))
@@ -279,30 +278,30 @@ let is_nondep_branch c n =
let extract_nondep_branches test c b n =
let rec strip n r = if n=0 then r else
match r with
- | RLambda (_,_,_,_,t) -> strip (n-1) t
- | RLetIn (_,_,_,t) -> strip (n-1) t
+ | 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
- | RLambda (_,na,_,_,c) -> aux (n-1) (na::nal) c
- | RLetIn (_,na,_,c) -> aux (n-1) (na::nal) c
+ | GLambda (_,na,_,_,c) -> aux (n-1) (na::nal) c
+ | GLetIn (_,na,_,c) -> aux (n-1) (na::nal) c
| _ ->
(* eta-expansion *)
let rec next l =
let x = next_ident_away (id_of_string "x") l in
- (* Not efficient but unusual and no function to get free rawvars *)
-(* if occur_rawconstr x c then next (x::l) else x 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_rawvars c) in
- let a = RVar (dl,x) in
+ let x = next (free_glob_vars c) in
+ let a = GVar (dl,x) in
aux (n-1) (Name x :: nal)
(match c with
- | RApp (loc,p,l) -> RApp (loc,c,l@[a])
- | _ -> (RApp (dl,c,[a])))
+ | GApp (loc,p,l) -> GApp (loc,c,l@[a])
+ | _ -> (GApp (dl,c,[a])))
in aux n [] c
let detype_case computable detype detype_eqns testdep avoid data p c bl =
@@ -319,7 +318,7 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl =
| Some p ->
let nl,typ = it_destRLambda_or_LetIn_names k p in
let n,typ = match typ with
- | RLambda (_,x,_,t,c) -> x, c
+ | GLambda (_,x,_,t,c) -> x, c
| _ -> Anonymous, typ in
let aliastyp =
if List.for_all ((=) Anonymous) nl then None
@@ -333,9 +332,9 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl =
RegularStyle
else if st = LetPatternStyle then
st
- else if PrintingLet.active (indsp,consnargsl) then
+ else if PrintingLet.active indsp then
LetStyle
- else if PrintingIf.active (indsp,consnargsl) then
+ else if PrintingIf.active indsp then
IfStyle
else
st
@@ -345,24 +344,24 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl =
| LetStyle when aliastyp = None ->
let bl' = Array.map detype bl in
let (nal,d) = it_destRLambda_or_LetIn_names consnargsl.(0) bl'.(0) in
- RLetTuple (dl,nal,(alias,pred),tomatch,d)
+ GLetTuple (dl,nal,(alias,pred),tomatch,d)
| IfStyle when aliastyp = 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
- RIf (dl,tomatch,(alias,pred),
+ GIf (dl,tomatch,(alias,pred),
Option.get nondepbrs.(0),Option.get nondepbrs.(1))
else
let eqnl = detype_eqns constructs consnargsl bl in
- RCases (dl,tag,pred,[tomatch,(alias,aliastyp)],eqnl)
+ GCases (dl,tag,pred,[tomatch,(alias,aliastyp)],eqnl)
| _ ->
let eqnl = detype_eqns constructs consnargsl bl in
- RCases (dl,tag,pred,[tomatch,(alias,aliastyp)],eqnl)
+ GCases (dl,tag,pred,[tomatch,(alias,aliastyp)],eqnl)
let detype_sort = function
- | Prop c -> RProp c
- | Type u -> RType (Some u)
+ | Prop c -> GProp c
+ | Type u -> GType (Some u)
type binder_kind = BProd | BLambda | BLetIn
@@ -376,43 +375,43 @@ let rec detype (isgoal:bool) avoid env t =
match kind_of_term (collapse_appl t) with
| Rel n ->
(try match lookup_name_of_rel n env with
- | Name id -> RVar (dl, id)
+ | Name id -> GVar (dl, id)
| Anonymous -> !detype_anonymous dl n
with Not_found ->
let s = "_UNBOUND_REL_"^(string_of_int n)
- in RVar (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 *)
- REvar (dl, n, None)
+ GEvar (dl, n, None)
| Var id ->
(try
- let _ = Global.lookup_named id in RRef (dl, VarRef id)
+ let _ = Global.lookup_named id in GRef (dl, VarRef id)
with _ ->
- RVar (dl, id))
- | Sort s -> RSort (dl,detype_sort s)
+ GVar (dl, id))
+ | Sort s -> GSort (dl,detype_sort s)
| Cast (c1,k,c2) ->
- RCast(dl,detype isgoal avoid env c1, CastConv (k, detype isgoal avoid env 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
| App (f,args) ->
- RApp (dl,detype isgoal avoid env f,
+ GApp (dl,detype isgoal avoid env f,
array_map_to_list (detype isgoal avoid env) args)
- | Const sp -> RRef (dl, ConstRef sp)
+ | Const sp -> GRef (dl, ConstRef sp)
| Evar (ev,cl) ->
- REvar (dl, ev,
+ GEvar (dl, ev,
Some (List.map (detype isgoal avoid env) (Array.to_list cl)))
| Ind ind_sp ->
- RRef (dl, IndRef ind_sp)
+ GRef (dl, IndRef ind_sp)
| Construct cstr_sp ->
- RRef (dl, ConstructRef cstr_sp)
+ GRef (dl, ConstructRef cstr_sp)
| 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)
is_nondep_branch avoid
(ci.ci_ind,ci.ci_pp_info.style,ci.ci_npar,
- ci.ci_cstr_nargs,ci.ci_pp_info.ind_nargs)
+ ci.ci_cstr_ndecls,ci.ci_pp_info.ind_nargs)
(Some p) c bl
| Fix (nvn,recdef) -> detype_fix isgoal avoid env nvn recdef
| CoFix (n,recdef) -> detype_cofix isgoal avoid env n recdef
@@ -428,7 +427,7 @@ and detype_fix isgoal avoid env (vn,_ as nvn) (names,tys,bodies) =
let v = array_map3
(fun c t i -> share_names isgoal (i+1) [] def_avoid def_env c (lift n t))
bodies tys vn in
- RRec(dl,RFix (Array.map (fun i -> Some i, RStructRec) (fst nvn), snd nvn),Array.of_list (List.rev lfi),
+ 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)
@@ -444,7 +443,7 @@ and detype_cofix isgoal avoid env n (names,tys,bodies) =
let v = array_map2
(fun c t -> share_names isgoal 0 [] def_avoid def_env c (lift ntys t))
bodies tys in
- RRec(dl,RCoFix n,Array.of_list (List.rev lfi),
+ 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)
@@ -539,9 +538,9 @@ and detype_binder isgoal bk avoid env na ty c =
else compute_displayed_name_in flag avoid na c in
let r = detype isgoal avoid' (add_name na' env) c in
match bk with
- | BProd -> RProd (dl, na',Explicit,detype false avoid env ty, r)
- | BLambda -> RLambda (dl, na',Explicit,detype false avoid env ty, r)
- | BLetIn -> RLetIn (dl, na',detype false avoid env ty, r)
+ | 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)
let rec detype_rel_context where avoid env sign =
let where = Option.map (fun c -> it_mkLambda_or_LetIn c sign) where in
@@ -573,42 +572,42 @@ let rec subst_cases_pattern subst pat =
if kn' == kn && cpl' == cpl then pat else
PatCstr (loc,((kn',i),j),cpl',n)
-let rec subst_rawconstr subst raw =
+let rec subst_glob_constr subst raw =
match raw with
- | RRef (loc,ref) ->
+ | GRef (loc,ref) ->
let ref',t = subst_global subst ref in
if ref' == ref then raw else
detype false [] [] t
- | RVar _ -> raw
- | REvar _ -> raw
- | RPatVar _ -> raw
+ | GVar _ -> raw
+ | GEvar _ -> raw
+ | GPatVar _ -> raw
- | RApp (loc,r,rl) ->
- let r' = subst_rawconstr subst r
- and rl' = list_smartmap (subst_rawconstr subst) rl in
+ | GApp (loc,r,rl) ->
+ let r' = subst_glob_constr subst r
+ and rl' = list_smartmap (subst_glob_constr subst) rl in
if r' == r && rl' == rl then raw else
- RApp(loc,r',rl')
+ GApp(loc,r',rl')
- | RLambda (loc,n,bk,r1,r2) ->
- let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
+ | GLambda (loc,n,bk,r1,r2) ->
+ let r1' = subst_glob_constr subst r1 and r2' = subst_glob_constr subst r2 in
if r1' == r1 && r2' == r2 then raw else
- RLambda (loc,n,bk,r1',r2')
+ GLambda (loc,n,bk,r1',r2')
- | RProd (loc,n,bk,r1,r2) ->
- let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
+ | GProd (loc,n,bk,r1,r2) ->
+ let r1' = subst_glob_constr subst r1 and r2' = subst_glob_constr subst r2 in
if r1' == r1 && r2' == r2 then raw else
- RProd (loc,n,bk,r1',r2')
+ GProd (loc,n,bk,r1',r2')
- | RLetIn (loc,n,r1,r2) ->
- let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
+ | GLetIn (loc,n,r1,r2) ->
+ let r1' = subst_glob_constr subst r1 and r2' = subst_glob_constr subst r2 in
if r1' == r1 && r2' == r2 then raw else
- RLetIn (loc,n,r1',r2')
+ GLetIn (loc,n,r1',r2')
- | RCases (loc,sty,rtno,rl,branches) ->
- let rtno' = Option.smartmap (subst_rawconstr subst) rtno
+ | GCases (loc,sty,rtno,rl,branches) ->
+ let rtno' = Option.smartmap (subst_glob_constr subst) rtno
and rl' = list_smartmap (fun (a,x as y) ->
- let a' = subst_rawconstr subst a in
+ 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) ->
@@ -619,72 +618,71 @@ let rec subst_rawconstr subst raw =
(fun (loc,idl,cpl,r as branch) ->
let cpl' =
list_smartmap (subst_cases_pattern subst) cpl
- and r' = subst_rawconstr subst r in
+ and r' = subst_glob_constr subst r in
if cpl' == cpl && r' == r then branch else
(loc,idl,cpl',r'))
branches
in
if rtno' == rtno && rl' == rl && branches' == branches then raw else
- RCases (loc,sty,rtno',rl',branches')
+ GCases (loc,sty,rtno',rl',branches')
- | RLetTuple (loc,nal,(na,po),b,c) ->
- let po' = Option.smartmap (subst_rawconstr subst) po
- and b' = subst_rawconstr subst b
- and c' = subst_rawconstr subst c in
+ | GLetTuple (loc,nal,(na,po),b,c) ->
+ let po' = Option.smartmap (subst_glob_constr subst) po
+ and b' = subst_glob_constr subst b
+ and c' = subst_glob_constr subst c in
if po' == po && b' == b && c' == c then raw else
- RLetTuple (loc,nal,(na,po'),b',c')
+ GLetTuple (loc,nal,(na,po'),b',c')
- | RIf (loc,c,(na,po),b1,b2) ->
- let po' = Option.smartmap (subst_rawconstr subst) po
- and b1' = subst_rawconstr subst b1
- and b2' = subst_rawconstr subst b2
- and c' = subst_rawconstr subst c in
+ | GIf (loc,c,(na,po),b1,b2) ->
+ let po' = Option.smartmap (subst_glob_constr subst) po
+ and b1' = subst_glob_constr subst b1
+ and b2' = subst_glob_constr subst b2
+ and c' = subst_glob_constr subst c in
if c' == c & po' == po && b1' == b1 && b2' == b2 then raw else
- RIf (loc,c',(na,po'),b1',b2')
+ GIf (loc,c',(na,po'),b1',b2')
- | RRec (loc,fix,ida,bl,ra1,ra2) ->
- let ra1' = array_smartmap (subst_rawconstr subst) ra1
- and ra2' = array_smartmap (subst_rawconstr subst) ra2 in
+ | 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 ty' = subst_rawconstr subst ty in
- let obd' = Option.smartmap (subst_rawconstr subst) obd in
+ 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')))
bl in
if ra1' == ra1 && ra2' == ra2 && bl'==bl then raw else
- RRec (loc,fix,ida,bl',ra1',ra2')
+ GRec (loc,fix,ida,bl',ra1',ra2')
- | RSort _ -> raw
+ | GSort _ -> raw
- | RHole (loc,ImplicitArg (ref,i,b)) ->
+ | GHole (loc,ImplicitArg (ref,i,b)) ->
let ref',_ = subst_global subst ref in
if ref' == ref then raw else
- RHole (loc,InternalHole)
- | RHole (loc, (BinderType _ | QuestionMark _ | CasesType | InternalHole |
+ GHole (loc,InternalHole)
+ | GHole (loc, (BinderType _ | QuestionMark _ | CasesType | InternalHole |
TomatchTypeParameter _ | GoalEvar | ImpossibleCase | MatchingVar _)) ->
raw
- | RCast (loc,r1,k) ->
+ | GCast (loc,r1,k) ->
(match k with
CastConv (k,r2) ->
- let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
+ let r1' = subst_glob_constr subst r1 and r2' = subst_glob_constr subst r2 in
if r1' == r1 && r2' == r2 then raw else
- RCast (loc,r1', CastConv (k,r2'))
+ GCast (loc,r1', CastConv (k,r2'))
| CastCoerce ->
- let r1' = subst_rawconstr subst r1 in
- if r1' == r1 then raw else RCast (loc,r1',k))
- | RDynamic _ -> raw
+ let r1' = subst_glob_constr subst r1 in
+ if r1' == r1 then raw else GCast (loc,r1',k))
(* Utilities to transform kernel cases to simple pattern-matching problem *)
-let simple_cases_matrix_of_branches ind brns brs =
- list_map2_i (fun i n b ->
+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))
- 0 brns brs
+ brs
let return_type_of_predicate ind nparams nrealargs_ctxt pred =
let nal,p = it_destRLambda_or_LetIn_names (nrealargs_ctxt+1) pred in
diff --git a/pretyping/detyping.mli b/pretyping/detyping.mli
index 556b2477..40e3d0f8 100644
--- a/pretyping/detyping.mli
+++ b/pretyping/detyping.mli
@@ -1,60 +1,74 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: detyping.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
open Util
open Names
open Term
open Sign
open Environ
-open Rawterm
+open Glob_term
open Termops
open Mod_subst
-(*i*)
val subst_cases_pattern : substitution -> cases_pattern -> cases_pattern
-val subst_rawconstr : substitution -> rawconstr -> rawconstr
+val subst_glob_constr : substitution -> glob_constr -> glob_constr
-(* [detype isgoal avoid ctx c] turns a closed [c], into a rawconstr *)
-(* de Bruijn indexes are turned to bound names, avoiding names in [avoid] *)
-(* [isgoal] tells if naming must avoid global-level synonyms as intro does *)
-(* [ctx] gives the names of the free variables *)
+(** [detype isgoal avoid ctx c] turns a closed [c], into a glob_constr
+ de Bruijn indexes are turned to bound names, avoiding names in [avoid]
+ [isgoal] tells if naming must avoid global-level synonyms as intro does
+ [ctx] gives the names of the free variables *)
-val detype : bool -> identifier list -> names_context -> constr -> rawconstr
+val detype : bool -> identifier list -> names_context -> constr -> glob_constr
val detype_case :
- bool -> ('a -> rawconstr) ->
+ bool -> ('a -> glob_constr) ->
(constructor array -> int array -> 'a array ->
- (loc * identifier list * cases_pattern list * rawconstr) list) ->
+ (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 -> rawconstr
+ 'a option -> 'a -> 'a array -> glob_constr
-val detype_sort : sorts -> rawsort
+val detype_sort : sorts -> glob_sort
val detype_rel_context : constr option -> identifier list -> names_context ->
- rel_context -> rawdecl list
+ rel_context -> glob_decl list
-(* look for the index of a named var or a nondep var as it is renamed *)
+(** 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_index_as_renamed : env -> constr -> int -> int option
-val set_detype_anonymous : (loc -> int -> rawconstr) -> unit
+val set_detype_anonymous : (loc -> int -> glob_constr) -> unit
val force_wildcard : unit -> bool
val synthetize_type : unit -> bool
-(* Utilities to transform kernel cases to simple pattern-matching problem *)
+(** Utilities to transform kernel cases to simple pattern-matching problem *)
-val it_destRLambda_or_LetIn_names : int -> rawconstr -> name list * rawconstr
+val it_destRLambda_or_LetIn_names : int -> glob_constr -> name list * glob_constr
val simple_cases_matrix_of_branches :
- inductive -> int list -> rawconstr list -> cases_clauses
+ inductive -> (int * int * glob_constr) list -> cases_clauses
val return_type_of_predicate :
- inductive -> int -> int -> rawconstr -> predicate_pattern * rawconstr option
+ inductive -> int -> int -> glob_constr -> predicate_pattern * glob_constr option
+
+module PrintingInductiveMake :
+ functor (Test : sig
+ val encode : Libnames.reference -> Names.inductive
+ val member_message : Pp.std_ppcmds -> bool -> Pp.std_ppcmds
+ val field : string
+ val title : string
+ end) ->
+ sig
+ type t = Names.inductive
+ val encode : Libnames.reference -> Names.inductive
+ val subst : substitution -> t -> t
+ val printer : t -> Pp.std_ppcmds
+ val key : Goptions.option_name
+ val title : string
+ val member_message : t -> bool -> Pp.std_ppcmds
+ val synchronous : bool
+ end
diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml
index 6d080016..04f86e70 100644
--- a/pretyping/evarconv.ml
+++ b/pretyping/evarconv.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: evarconv.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
open Pp
open Util
open Names
@@ -24,44 +22,52 @@ open Evd
type flex_kind_of_term =
| Rigid of constr
- | MaybeFlexible 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 =
match kind_of_term c with
- | Const _ -> MaybeFlexible c
- | Rel n -> MaybeFlexible c
- | Var id -> MaybeFlexible c
+ | Rel _ | Const _ | Var _ -> MaybeFlexible c
| Lambda _ when l<>[] -> MaybeFlexible c
| LetIn _ -> MaybeFlexible c
| Evar ev -> Flexible ev
- | _ -> Rigid c
+ | Lambda _ | Prod _ | Sort _ | Ind _ | Construct _ | CoFix _ -> Rigid c
+ | Meta _ | Case _ | Fix _ -> PseudoRigid c
+ | Cast _ | App _ -> assert false
-let eval_flexible_term env c =
+let eval_flexible_term ts env c =
match kind_of_term c with
- | Const c -> constant_opt_value env c
+ | Const c ->
+ if is_transparent_constant ts c
+ then constant_opt_value env c
+ else None
| Rel n ->
(try let (_,v,_) = lookup_rel n env in Option.map (lift n) v
with Not_found -> None)
| Var id ->
- (try let (_,v,_) = lookup_named id env in v with Not_found -> None)
+ (try
+ if is_transparent_variable ts id then
+ let (_,v,_) = lookup_named id env in v
+ else None
+ with Not_found -> None)
| LetIn (_,b,_,c) -> Some (subst1 b c)
| Lambda _ -> Some c
| _ -> assert false
-let evar_apprec env evd stack c =
+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 env sigma s in
+ 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)
-let apprec_nohdbeta env evd c =
+let apprec_nohdbeta ts env evd c =
match kind_of_term (fst (Reductionops.whd_stack evd c)) with
- | (Case _ | Fix _) -> applist (evar_apprec env evd [] c)
+ | (Case _ | Fix _) -> applist (evar_apprec ts env evd [] c)
| _ -> c
let position_problem l2r = function
@@ -159,16 +165,15 @@ let ise_array2 evd f v1 v2 =
if lv1 = Array.length v2 then allrec evd (pred lv1)
else (evd,false)
-let rec evar_conv_x env evd pbty term1 term2 =
- let sigma = evd in
- let term1 = whd_head_evar sigma term1 in
- let term2 = whd_head_evar sigma term2 in
+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
(* Maybe convertible but since reducing can erase evars which [evar_apprec]
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_fconv pbty env evd term1 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
@@ -176,19 +181,22 @@ let rec evar_conv_x env evd pbty term1 term2 =
match ground_test with
Some b -> (evd,b)
| None ->
- let term1 = apprec_nohdbeta env evd term1 in
- let term2 = apprec_nohdbeta env evd term2 in
+ (* 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 env evd
+ 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 env evd
+ solve_simple_eqn (evar_conv_x ts) env evd
(position_problem false pbty,destEvar term2,term1)
else
- evar_eqappr_x env evd pbty
+ evar_eqappr_x ts env evd pbty
(decompose_app term1) (decompose_app term2)
-and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
+and evar_eqappr_x ?(rhs_is_stuck_proj = false)
+ ts env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
(* Evar must be undefined since we have flushed evars *)
match (flex_kind_of_term term1 l1, flex_kind_of_term term2 l2) with
| Flexible (sp1,al1 as ev1), Flexible (sp2,al2 as ev2) ->
@@ -196,23 +204,23 @@ and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
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 env 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 env i CONV) rest1 l2)]
+ (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 env 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 env i CONV) l1 rest2)]
+ (fun i -> evar_conv_x ts env i CONV) l1 rest2)]
and f2 i =
if sp1 = sp2 then
ise_and i
[(fun i -> ise_list2 i
- (fun i -> evar_conv_x env i CONV) l1 l2);
- (fun i -> solve_refl evar_conv_x env i sp1 al1 al2,
+ (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)
in
@@ -220,17 +228,17 @@ and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
| Flexible ev1, MaybeFlexible flex2 ->
let f1 i =
- if
- is_unification_pattern_evar env ev1 l1 (applist appr2) &
- not (occur_evar (fst ev1) (applist appr2))
- then
+ 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 env evd
+ let t2 = solve_pattern_eqn env l1' t2 in
+ solve_simple_eqn (evar_conv_x ts) env evd
(position_problem true pbty,ev1,t2)
- else if
+ | None -> (i,false)
+ and f2 i =
+ if
List.length l1 <= List.length l2
then
(* Try first-order unification *)
@@ -240,30 +248,30 @@ and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
ise_and i
(* First compare extra args for better failure message *)
[(fun i -> ise_list2 i
- (fun i -> evar_conv_x env i CONV) l1 rest2);
- (fun i -> evar_conv_x env i pbty term1 (applist(term2,deb2)))]
+ (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 f4 i =
- match eval_flexible_term env flex2 with
+ and f3 i =
+ match eval_flexible_term ts env flex2 with
| Some v2 ->
- evar_eqappr_x env i pbty appr1 (evar_apprec env i l2 v2)
+ evar_eqappr_x ts env i pbty appr1 (evar_apprec ts env i l2 v2)
| None -> (i,false)
in
- ise_try evd [f1; f4]
+ ise_try evd [f1; f2; f3]
| MaybeFlexible flex1, Flexible ev2 ->
let f1 i =
- if
- is_unification_pattern_evar env ev2 l2 (applist appr1) &
- not (occur_evar (fst ev2) (applist appr1))
- then
+ 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 env evd
+ solve_simple_eqn (evar_conv_x ts) env evd
(position_problem false pbty,ev2,t1)
- else if
+ | None -> (i,false)
+ and f2 i =
+ if
List.length l2 <= List.length l1
then
(* Try first-order unification *)
@@ -272,25 +280,43 @@ and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
ise_and i
(* First compare extra args for better failure message *)
[(fun i -> ise_list2 i
- (fun i -> evar_conv_x env i CONV) rest1 l2);
- (fun i -> evar_conv_x env i pbty (applist(term1,deb1)) term2)]
+ (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 f4 i =
- match eval_flexible_term env flex1 with
+ and f3 i =
+ match eval_flexible_term ts env flex1 with
| Some v1 ->
- evar_eqappr_x env i pbty (evar_apprec env i l1 v1) appr2
+ evar_eqappr_x ts env i pbty (evar_apprec ts env i l1 v1) appr2
| None -> (i,false)
in
- ise_try evd [f1; f4]
+ ise_try evd [f1; f2; f3]
+
+ | 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) ->
+ let f1 i =
+ ise_and i
+ [(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)]
+ 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
+ in
+ ise_try evd [f1; f2]
- | MaybeFlexible flex1, MaybeFlexible flex2 ->
+ | _, _ ->
let f1 i =
- if flex1 = flex2 then
- ise_list2 i (fun i -> evar_conv_x env i CONV) l1 l2
+ if eq_constr flex1 flex2 then
+ ise_list2 i (fun i -> evar_conv_x ts env i CONV) l1 l2
else
(i,false)
and f2 i =
- (try conv_record env 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))
@@ -299,186 +325,204 @@ and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
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 *)
- if isLambda flex1 or is_open_canonical_projection i appr2
- then
- match eval_flexible_term env flex1 with
+ let rhs_is_stuck_proj =
+ rhs_is_stuck_proj || is_open_canonical_projection env i appr2 in
+ if isLambda flex1 || rhs_is_stuck_proj then
+ match eval_flexible_term ts env flex1 with
| Some v1 ->
- evar_eqappr_x env i pbty (evar_apprec env i l1 v1) appr2
+ evar_eqappr_x ~rhs_is_stuck_proj
+ ts env i pbty (evar_apprec ts env i l1 v1) appr2
| None ->
- match eval_flexible_term env flex2 with
+ match eval_flexible_term ts env flex2 with
| Some v2 ->
- evar_eqappr_x env i pbty appr1 (evar_apprec env i l2 v2)
+ evar_eqappr_x ts env i pbty appr1 (evar_apprec ts env i l2 v2)
| None -> (i,false)
else
- match eval_flexible_term env flex2 with
+ match eval_flexible_term ts env flex2 with
| Some v2 ->
- evar_eqappr_x env i pbty appr1 (evar_apprec env i l2 v2)
+ evar_eqappr_x ts env i pbty appr1 (evar_apprec ts env i l2 v2)
| None ->
- match eval_flexible_term env flex1 with
+ match eval_flexible_term ts env flex1 with
| Some v1 ->
- evar_eqappr_x env i pbty (evar_apprec env i l1 v1) appr2
+ 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]
-
- | Flexible ev1, Rigid _ ->
- if
- is_unification_pattern_evar env ev1 l1 (applist appr2) &
- not (occur_evar (fst ev1) (applist appr2))
- then
- (* Miller-Pfenning's patterns unification *)
- (* Preserve generality (except that CCI has no eta-conversion) *)
+ 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=[]);
+ 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 env evd
+ solve_simple_eqn (evar_conv_x ts) env evd
(position_problem true pbty,ev1,t2)
- else
+ | None ->
(* Postpone the use of an heuristic *)
add_conv_pb (pbty,env,applist appr1,applist appr2) evd,
- true
-
- | Rigid _, Flexible ev2 ->
- if
- is_unification_pattern_evar env ev2 l2 (applist appr1) &
- not (occur_evar (fst ev2) (applist appr1))
- then
- (* Miller-Pfenning's patterns unification *)
- (* Preserve generality (except that CCI has no eta-conversion) *)
+ 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 env evd
+ solve_simple_eqn (evar_conv_x ts) env evd
(position_problem false pbty,ev2,t1)
- else
+ | None ->
(* Postpone the use of an heuristic *)
add_conv_pb (pbty,env,applist appr1,applist appr2) evd,
- true
+ true)
- | MaybeFlexible flex1, Rigid _ ->
+ | MaybeFlexible flex1, (Rigid _ | PseudoRigid _) ->
let f3 i =
- (try conv_record env i (check_conv_record appr1 appr2)
+ (try conv_record ts env i (check_conv_record appr1 appr2)
with Not_found -> (i,false))
and f4 i =
- match eval_flexible_term env flex1 with
+ match eval_flexible_term ts env flex1 with
| Some v1 ->
- evar_eqappr_x env i pbty (evar_apprec env i l1 v1) appr2
+ evar_eqappr_x ts env i pbty (evar_apprec ts env i l1 v1) appr2
| None -> (i,false)
in
ise_try evd [f3; f4]
- | Rigid _ , MaybeFlexible flex2 ->
+ | (Rigid _ | PseudoRigid _), MaybeFlexible flex2 ->
let f3 i =
- (try conv_record env i (check_conv_record appr2 appr1)
+ (try conv_record ts env i (check_conv_record appr2 appr1)
with Not_found -> (i,false))
and f4 i =
- match eval_flexible_term env flex2 with
+ match eval_flexible_term ts env flex2 with
| Some v2 ->
- evar_eqappr_x env i pbty appr1 (evar_apprec env i l2 v2)
+ evar_eqappr_x ts env i pbty appr1 (evar_apprec ts env i l2 v2)
| None -> (i,false)
in
ise_try evd [f3; f4]
- | Rigid c1, Rigid c2 -> match kind_of_term c1, kind_of_term c2 with
-
- | Cast (c1,_,_), _ -> evar_eqappr_x env evd pbty (c1,l1) appr2
-
- | _, Cast (c2,_,_) -> evar_eqappr_x env evd pbty appr1 (c2,l2)
+ (* Eta-expansion *)
+ | Rigid c1, _ when isLambda c1 ->
+ assert (l1 = []);
+ let (na,c1,c'1) = destLambda c1 in
+ let c = nf_evar evd c1 in
+ let env' = push_rel (na,None,c) env in
+ let appr1 = evar_apprec ts env' evd [] c'1 in
+ let appr2 = (lift 1 term2, List.map (lift 1) l2 @ [mkRel 1]) in
+ evar_eqappr_x ts env' evd CONV appr1 appr2
+
+ | _, Rigid c2 when isLambda c2 ->
+ assert (l2 = []);
+ let (na,c2,c'2) = destLambda c2 in
+ let c = nf_evar evd c2 in
+ let env' = push_rel (na,None,c) env in
+ let appr1 = (lift 1 term1, List.map (lift 1) l1 @ [mkRel 1]) in
+ let appr2 = evar_apprec ts env' evd [] c'2 in
+ evar_eqappr_x ts env' evd CONV appr1 appr2
+
+ | Rigid c1, Rigid c2 -> begin
+ match kind_of_term c1, kind_of_term c2 with
| Sort s1, Sort s2 when l1=[] & l2=[] ->
- (evd,base_sort_cmp pbty s1 s2)
-
- | Lambda (na,c1,c'1), Lambda (_,c2,c'2) when l1=[] & l2=[] ->
- ise_and evd
- [(fun i -> evar_conv_x env i CONV c1 c2);
- (fun i ->
- let c = nf_evar i c1 in
- evar_conv_x (push_rel (na,None,c) env) i CONV c'1 c'2)]
-
- | LetIn (na,b1,t1,c'1), LetIn (_,b2,_,c'2) ->
- let f1 i =
- ise_and i
- [(fun i -> evar_conv_x env i CONV b1 b2);
- (fun i ->
- let b = nf_evar i b1 in
- let t = nf_evar i t1 in
- evar_conv_x (push_rel (na,Some b,t) env) i pbty c'1 c'2);
- (fun i -> ise_list2 i
- (fun i -> evar_conv_x env i CONV) l1 l2)]
- and f2 i =
- let appr1 = evar_apprec env i l1 (subst1 b1 c'1)
- and appr2 = evar_apprec env i l2 (subst1 b2 c'2)
- in evar_eqappr_x env i pbty appr1 appr2
- in
- ise_try evd [f1; f2]
-
- | LetIn (_,b1,_,c'1), _ ->(* On fait commuter les args avec le Let *)
- let appr1 = evar_apprec env evd l1 (subst1 b1 c'1)
- in evar_eqappr_x env evd pbty appr1 appr2
-
- | _, LetIn (_,b2,_,c'2) ->
- let appr2 = evar_apprec env evd l2 (subst1 b2 c'2)
- in evar_eqappr_x env evd pbty appr1 appr2
+ (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)
+ | _ -> (evd, false))
| Prod (n,c1,c'1), Prod (_,c2,c'2) when l1=[] & l2=[] ->
ise_and evd
- [(fun i -> evar_conv_x env i CONV c1 c2);
+ [(fun i -> evar_conv_x ts env i CONV c1 c2);
(fun i ->
let c = nf_evar i c1 in
- evar_conv_x (push_rel (n,None,c) env) i pbty c'1 c'2)]
+ 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 env i CONV) l1 l2
+ 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 env i CONV) l1 l2
+ ise_list2 evd (fun i -> evar_conv_x ts env i CONV) l1 l2
else (evd, false)
+ | CoFix (i1,(_,tys1,bds1 as recdef1)), CoFix (i2,(_,tys2,bds2)) ->
+ if 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
+ | (Lambda _), _ -> assert false
+
+ end
+
+ | PseudoRigid c1, PseudoRigid c2 -> begin
+ match kind_of_term c1, kind_of_term c2 with
+
| Case (_,p1,c1,cl1), Case (_,p2,c2,cl2) ->
ise_and evd
- [(fun i -> evar_conv_x env i CONV p1 p2);
- (fun i -> evar_conv_x env i CONV c1 c2);
+ [(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 env i CONV) cl1 cl2);
- (fun i -> ise_list2 i (fun i -> evar_conv_x env i CONV) l1 l2)]
+ (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 env i CONV) tys1 tys2);
+ (fun i -> evar_conv_x ts env i CONV) tys1 tys2);
(fun i -> ise_array2 i
- (fun i -> evar_conv_x (push_rec_types recdef1 env) i CONV)
+ (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 env i CONV) l1 l2)]
+ (fun i -> evar_conv_x ts env i CONV) l1 l2)]
else (evd,false)
- | CoFix (i1,(_,tys1,bds1 as recdef1)), CoFix (i2,(_,tys2,bds2)) ->
- if i1=i2 then
- ise_and evd
- [(fun i -> ise_array2 i
- (fun i -> evar_conv_x env i CONV) tys1 tys2);
- (fun i -> ise_array2 i
- (fun i -> evar_conv_x (push_rec_types recdef1 env) i CONV)
- bds1 bds2);
- (fun i -> ise_list2 i
- (fun i -> evar_conv_x env i CONV) l1 l2)]
- else (evd,false)
- | (Meta _ | Lambda _), _ -> (evd,false)
- | _, (Meta _ | Lambda _) -> (evd,false)
+ | (Meta _ | Case _ | Fix _ | CoFix _),
+ (Meta _ | Case _ | Fix _ | CoFix _) -> (evd,false)
+
+ | (App _ | Ind _ | Construct _ | Sort _ | Prod _), _ -> assert false
+ | _, (App _ | Ind _ | Construct _ | Sort _ | Prod _) -> assert false
- | (Ind _ | Construct _ | Sort _ | Prod _), _ -> (evd,false)
- | _, (Ind _ | Construct _ | Sort _ | Prod _) -> (evd,false)
+ | (LetIn _ | Cast _), _ -> assert false
+ | _, (LetIn _ | Cast _) -> assert false
- | (App _ | Case _ | Fix _ | CoFix _),
- (App _ | Case _ | Fix _ | CoFix _) -> (evd,false)
+ | (Lambda _ | Rel _ | Var _ | Const _ | Evar _), _ -> assert false
+ | _, (Lambda _ | Rel _ | Var _ | Const _ | Evar _) -> assert false
+ end
- | (Rel _ | Var _ | Const _ | Evar _), _ -> assert false
- | _, (Rel _ | Var _ | Const _ | Evar _) -> assert false
+ | PseudoRigid _, Rigid _ -> (evd,false)
-and conv_record env evd (c,bs,(params,params1),(us,us2),(ts,ts1),c1,(n,t2)) =
+ | Rigid _, PseudoRigid _ -> (evd,false)
+
+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 ->
@@ -491,89 +535,259 @@ and conv_record env evd (c,bs,(params,params1),(us,us2),(ts,ts1),c1,(n,t2)) =
ise_and evd'
[(fun i ->
ise_list2 i
- (fun i x1 x -> evar_conv_x env i CONV x1 (substl ks x))
+ (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 env i CONV u1 (substl ks u))
+ (fun i u1 u -> evar_conv_x trs env i CONV u1 (substl ks u))
us2 us);
- (fun i -> evar_conv_x env i CONV c1 (applist (c,(List.rev ks))));
- (fun i -> ise_list2 i (fun i -> evar_conv_x env i CONV) ts ts1)]
+ (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)]
+
+(* getting rid of the optional argument rhs_is_stuck_proj *)
+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 env evd (ev1,l1) (term2,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
ise_and evd
(* First compare extra args for better failure message *)
- [(fun i -> ise_list2 i (fun i -> evar_conv_x env i CONV) rest2 l1);
+ [(fun i -> ise_list2 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
- evar_conv_x env i CONV t2 (mkEvar ev1)
+ evar_conv_x ts env i CONV t2 (mkEvar ev1)
else
- solve_simple_eqn ~choose:true evar_conv_x env i (None,ev1,t2))]
+ solve_simple_eqn ~choose:true (evar_conv_x ts) env i (None,ev1,t2))]
let choose_less_dependent_instance evk evd term args =
- let evi = Evd.find evd evk in
+ let evi = Evd.find_undefined evd evk in
let subst = make_pure_subst evi args in
- let subst' = List.filter (fun (id,c) -> c = term) subst in
+ let subst' = List.filter (fun (id,c) -> eq_constr c term) subst in
if subst' = [] then error "Too complex unification problem." else
Evd.define evk (mkVar (fst (List.hd subst'))) evd
-let apply_conversion_problem_heuristic env evd pbty t1 t2 =
- let t1 = apprec_nohdbeta env evd (whd_head_evar evd t1) in
- let t2 = apprec_nohdbeta env evd (whd_head_evar evd t2) in
+let apply_on_subterm f c t =
+ let rec applyrec (k,c as kc) 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
+ else
+ map_constr_with_binders_left_to_right (fun d (k,c) -> (k+1,lift 1 c))
+ applyrec kc t
+ in
+ applyrec (0,c) t
+
+let filter_possible_projections c args =
+ let fv1 = free_rels c in
+ let fv2 = collect_vars c in
+ List.map (fun a ->
+ 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)
+ args
+
+let initial_evar_data evi =
+ let ids = List.map pi1 (evar_context evi) in
+ (evar_filter evi, List.map mkVar ids)
+
+let solve_evars = ref (fun _ -> failwith "solve_evars not installed")
+let set_solve_evars f = solve_evars := f
+
+(* We solve the problem env_rhs |- ?e[u1..un] = rhs knowing
+ * x1:T1 .. xn:Tn |- ev : ty
+ * by looking for a maximal well-typed abtraction over u1..un in rhs
+ *
+ * We first build C[e11..e1p1,..,en1..enpn] obtained from rhs by replacing
+ * all occurrences of u1..un by evars eij of type Ti' where itself Ti' has
+ * been obtained from the type of ui by also replacing all occurrences of
+ * u1..ui-1 by evars.
+ *
+ * Then, we use typing to infer the relations between the different
+ * occurrences. If some occurrence is still unconstrained after typing,
+ * we instantiate successively the unresolved occurrences of un by xn,
+ * of un-1 by xn-1, etc [the idea comes from Chung-Kil Hur, that he
+ * used for his Heq plugin; extensions to several arguments based on a
+ * proposition from Dan Grayson]
+ *)
+
+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 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
+ error "Cannot force abstraction on identity instance."
+ else
+ make_subst (ctxt,l,occsl)
+ | (id,_,t)::ctxt, c::l, occs::occsl ->
+ let evs = ref [] in
+ let filter = List.map2 (&&) filter (filter_possible_projections c args) in
+ let ty = Retyping.get_type_of env_rhs evd c in
+ (id,t,c,ty,evs,filter,occs) :: make_subst (ctxt,l,occsl)
+ | [], [], [] -> []
+ | _ -> anomaly "Signature, instance and occurrences list do not match" 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 _ -> 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 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 set_var c rhs) subst
+ | [] -> rhs in
+
+ let subst = make_subst (ctxt,args,argoccs) in
+
+ let evdref = ref evd in
+ let rhs = set_holes evdref rhs subst in
+ let evd = !evdref in
+
+ (* We instantiate the evars of which the value is forced by typing *)
+ let evd,rhs =
+ try !solve_evars env_evar evd rhs
+ with e when Pretype_errors.precatchable_exception e ->
+ (* Could not revert all subterms *)
+ raise Exit in
+
+ let rec abstract_free_holes evd = function
+ | (id,idty,c,_,evsref,_,_)::l ->
+ let rec force_instantiation evd = function
+ | (evk,evty)::evs ->
+ let evd =
+ if is_undefined evd evk then
+ (* We force abstraction over this unconstrained occurrence *)
+ (* and we use typing to propagate this instantiation *)
+ (* This is an arbitrary choice *)
+ let evd = Evd.define evk (mkVar id) evd in
+ let evd,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";
+ evd
+ else
+ evd
+ in
+ force_instantiation evd evs
+ | [] ->
+ abstract_free_holes evd l
+ in
+ force_instantiation evd !evsref
+ | [] ->
+ Evd.define evk rhs evd in
+
+ abstract_free_holes evd subst, true
+ with Exit -> 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
+*)
+ (evd,false)
+
+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
match kind_of_term term1, kind_of_term term2 with
| Evar (evk1,args1), (Rel _|Var _) when l1 = [] & l2 = []
- & array_for_all (fun a -> a = term2 or isEvar a) args1 ->
+ & array_for_all (fun a -> eq_constr a term2 or isEvar a) args1 ->
(* The typical kind of constraint coming from pattern-matching return
type inference *)
choose_less_dependent_instance evk1 evd term2 args1, true
| (Rel _|Var _), Evar (evk2,args2) when l1 = [] & l2 = []
- & array_for_all (fun a -> a = term1 or isEvar a) args2 ->
+ & array_for_all (fun a -> eq_constr a term1 or isEvar a) args2 ->
(* The typical kind of constraint coming from pattern-matching return
type inference *)
choose_less_dependent_instance evk2 evd term1 args2, true
| Evar ev1,_ when List.length l1 <= List.length l2 ->
(* On "?n t1 .. tn = u u1 .. u(n+p)", try first-order unification *)
- first_order_unification env evd (ev1,l1) appr2
+ (* 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 ->
(* On "u u1 .. u(n+p) = ?n t1 .. tn", try first-order unification *)
- first_order_unification env evd (ev2,l2) appr1
+ (* 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))]
+ | Evar ev1,_ ->
+ (* Try second-order pattern-matching *)
+ second_order_matching_with_args ts env evd ev1 l1 (applist appr2)
+ | _,Evar ev2 ->
+ (* Try second-order pattern-matching *)
+ second_order_matching_with_args ts env evd ev2 l2 (applist appr1)
| _ ->
(* Some head evar have been instantiated, or unknown kind of problem *)
- evar_conv_x env evd pbty t1 t2
+ evar_conv_x ts env evd pbty t1 t2
+
+let check_problems_are_solved env evd =
+ match snd (extract_all_conv_pbs evd) with
+ | (pbty,env,t1,t2)::_ -> Pretype_errors.error_cannot_unify env evd (t1, t2)
+ | _ -> ()
-let consider_remaining_unif_problems env evd =
+let consider_remaining_unif_problems ?(ts=full_transparent_state) env evd =
let (evd,pbs) = extract_all_conv_pbs evd in
- List.fold_left
+ let heuristic_solved_evd = List.fold_left
(fun evd (pbty,env,t1,t2) ->
- let evd', b = apply_conversion_problem_heuristic env evd pbty t1 t2 in
+ let evd', b = apply_conversion_problem_heuristic ts env evd pbty t1 t2 in
if b then evd' else Pretype_errors.error_cannot_unify env evd (t1, t2))
- evd pbs
+ evd pbs in
+ check_problems_are_solved env heuristic_solved_evd;
+ Evd.fold_undefined (fun ev ev_info evd' -> match ev_info.evar_source with
+ |_,ImpossibleCase ->
+ Evd.define ev (j_type (coq_unit_judge ())) evd'
+ |_ ->
+ match ev_info.evar_candidates with
+ | Some (a::l) -> Evd.define ev a evd'
+ | Some [] -> error "Unsolvable existential variables"
+ | None -> evd') heuristic_solved_evd heuristic_solved_evd
(* Main entry points *)
-let the_conv_x env t1 t2 evd =
- match evar_conv_x env evd CONV t1 t2 with
+let the_conv_x ?(ts=full_transparent_state) env t1 t2 evd =
+ match evar_conv_x ts env evd CONV t1 t2 with
(evd',true) -> evd'
| _ -> raise Reduction.NotConvertible
-let the_conv_x_leq env t1 t2 evd =
- match evar_conv_x env evd CUMUL t1 t2 with
+let the_conv_x_leq ?(ts=full_transparent_state) env t1 t2 evd =
+ match evar_conv_x ts env evd CUMUL t1 t2 with
(evd', true) -> evd'
| _ -> raise Reduction.NotConvertible
-let e_conv env evd t1 t2 =
- match evar_conv_x env !evd CONV t1 t2 with
+let e_conv ?(ts=full_transparent_state) env evd t1 t2 =
+ match evar_conv_x ts env !evd CONV t1 t2 with
(evd',true) -> evd := evd'; true
| _ -> false
-let e_cumul env evd t1 t2 =
- match evar_conv_x env !evd CUMUL t1 t2 with
+let e_cumul ?(ts=full_transparent_state) env evd t1 t2 =
+ match evar_conv_x ts env !evd CUMUL t1 t2 with
(evd',true) -> evd := evd'; true
| _ -> false
diff --git a/pretyping/evarconv.mli b/pretyping/evarconv.mli
index 50228d4e..3e2ca7ae 100644
--- a/pretyping/evarconv.mli
+++ b/pretyping/evarconv.mli
@@ -1,43 +1,47 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: evarconv.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
+open Names
open Term
open Sign
open Environ
+open Termops
open Reductionops
open Evd
-(*i*)
-(* returns exception Reduction.NotConvertible if not unifiable *)
-val the_conv_x : env -> constr -> constr -> evar_map -> evar_map
-val the_conv_x_leq : env -> constr -> constr -> evar_map -> evar_map
+(** 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
-(* The same function resolving evars by side-effect and
+(** The same function resolving evars by side-effect and
catching the exception *)
-val e_conv : env -> evar_map ref -> constr -> constr -> bool
-val e_cumul : env -> evar_map ref -> constr -> constr -> bool
+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
-(*i For debugging *)
-val evar_conv_x :
+(**/**)
+(* For debugging *)
+val evar_conv_x : transparent_state ->
env -> evar_map -> conv_pb -> constr -> constr -> evar_map * bool
-val evar_eqappr_x :
+val evar_eqappr_x : transparent_state ->
env -> evar_map ->
conv_pb -> constr * constr list -> constr * constr list ->
evar_map * bool
-(*i*)
+(**/**)
-val consider_remaining_unif_problems : env -> evar_map -> evar_map
+val consider_remaining_unif_problems : ?ts:transparent_state -> env -> evar_map -> evar_map
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 set_solve_evars : (env -> evar_map -> constr -> evar_map * constr) -> unit
+
+val second_order_matching : transparent_state -> env -> evar_map ->
+ existential -> occurrences option list -> constr -> evar_map * bool
diff --git a/pretyping/evarutil.ml b/pretyping/evarutil.ml
index b95b50b3..7d66bee0 100644
--- a/pretyping/evarutil.ml
+++ b/pretyping/evarutil.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: evarutil.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
open Util
open Pp
open Names
@@ -23,9 +21,6 @@ open Reductionops
open Pretype_errors
open Retyping
-open Pretype_errors
-open Retyping
-
(* Expanding existential variables *)
(* 1- flush_and_check_evars fails if an existential is undefined *)
@@ -63,11 +58,18 @@ let nf_evar_info evc info =
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 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
(**********************)
(* Creating new metas *)
@@ -76,6 +78,10 @@ let nf_evar_map evd = Evd.evars_reset_evd (nf_evars evd) evd
(* 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) };
fun () -> incr meta_ctr; !meta_ctr
let mk_new_meta () = mkMeta(new_meta())
@@ -84,14 +90,14 @@ let collect_evars emap c =
let rec collrec acc c =
match kind_of_term c with
| Evar (evk,_) ->
- if Evd.mem emap evk & not (Evd.is_defined emap evk) then evk::acc
+ 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 (fun ev {evar_concl = ccl} (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))
@@ -116,7 +122,7 @@ let push_duplicated_evars 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_evars emap in
+ 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
let change_exist evar =
@@ -129,15 +135,11 @@ let evars_to_metas sigma (emap, c) =
| _ -> map_constr replace c in
(sigma', replace c)
-(* The list of non-instantiated existential declarations *)
+(* The list of non-instantiated existential declarations (order is important) *)
let non_instantiated sigma =
- let listev = to_list sigma in
- List.fold_left
- (fun l (ev,evi) ->
- if evi.evar_body = Evar_empty then
- ((ev,nf_evar_info sigma evi)::l) else l)
- [] listev
+ let listev = Evd.undefined_list sigma in
+ List.map (fun (ev,evi) -> (ev,nf_evar_info sigma evi)) listev
(**********************)
(* Creating new evars *)
@@ -146,96 +148,148 @@ let non_instantiated sigma =
(* 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
(*------------------------------------*
* functional operations on evar sets *
*------------------------------------*)
-let new_evar_instance sign evd typ ?(src=(dummy_loc,InternalHole)) ?filter instance =
- let instance =
- match filter with
- | None -> instance
- | Some filter -> snd (list_filter2 (fun b c -> b) (filter,instance)) in
+let new_evar_instance sign evd typ ?(src=(dummy_loc,InternalHole)) ?filter ?candidates instance =
assert
(let ctxt = named_context_of_val sign in
list_distinct (ids_of_named_context ctxt));
let newevk = new_untyped_evar() in
- let evd = evar_declare sign newevk typ ~src:src ?filter evd in
+ let evd = evar_declare sign newevk typ ~src:src ?filter ?candidates evd in
(evd,mkEvar (newevk,Array.of_list instance))
(* 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_aliases sign =
+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 id'' = try Idmap.find id' aliases with Not_found -> id' in
- Idmap.add id id'' aliases
- | _ -> aliases)
- | None -> aliases) sign Idmap.empty
-
-let alias_of_var id aliases = try Idmap.find id aliases with Not_found -> id
+ (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 =
- let var_aliases = compute_aliases (named_context env) in
- let rels = rel_context env in
- let rel_aliases =
- snd (List.fold_right (fun (_,b,t) (n,aliases) ->
- (n-1,
- match b with
- | Some t when isRel t or isVar t -> Intmap.add n (lift n t) aliases
- | _ -> aliases)) rels (List.length rels,Intmap.empty)) in
+ (* 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 expand_var_once aliases x = match kind_of_term x with
- | Rel n -> Intmap.find n (snd aliases)
- | Var id -> mkVar (Idmap.find id (fst aliases))
- | _ -> raise Not_found
+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 rec expand_var_at_least_once aliases x =
- let t = expand_var_once aliases x in
- try expand_var_at_least_once aliases t
- with Not_found -> t
+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 expand_var aliases x =
- try expand_var_at_least_once aliases x with Not_found -> x
+let normalize_alias aliases x =
+ match normalize_alias_opt aliases x with
+ | Some a -> a
+ | None -> x
-let expand_var_opt aliases x =
- try Some (expand_var_at_least_once aliases x) with Not_found -> None
+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 c -> Intmap.add (n+1) (lift 1 c))
+ 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 when isRel t or isVar t -> Intmap.add 1 (lift 1 t) rel_aliases
- | _ -> rel_aliases in
+ | 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 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 _ ->
- expand_var aliases t
+ 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 rec expansions_of_var aliases x =
- try
- let t = expand_var_once aliases x in
- t :: expansions_of_var aliases t
- with Not_found ->
- [x]
-
-let expand_full_opt aliases y =
- try Some (expand_var aliases y) with Not_found -> None
+let free_vars_and_rels_up_alias_expansion aliases c =
+ let acc1 = ref Intset.empty and acc2 = ref Idset.empty in
+ let rec frec (aliases,depth) c = match kind_of_term c with
+ | Rel _ | Var _ ->
+ 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
+ | _ ->
+ (* not optimal: would need sharing if alias occurs more than once *)
+ frec (aliases,depth) c)
+ | 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)
(* Knowing that [Gamma |- ev : T] and that [ev] is applied to [args],
* [make_projectable_subst ev args] builds the substitution [Gamma:=args].
@@ -249,26 +303,38 @@ let expand_full_opt aliases y =
let make_projectable_subst aliases sigma evi args =
let sign = evar_filtered_context evi in
- let evar_aliases = compute_aliases sign in
- snd (List.fold_right
- (fun (id,b,c) (args,l) ->
- match b,args with
+ 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
- (rest,Idmap.add id [a,expand_full_opt aliases a,id] l)
+ 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 = alias_of_var id' evar_aliases in
- let sub = try Idmap.find idc l with Not_found -> [] in
- if List.exists (fun (c,_,_) -> eq_constr a c) sub then (rest,l)
+ 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,expand_full_opt aliases a,id)::sub) l)
+ (rest,
+ Idmap.add idc ((a,normalize_alias_opt aliases a,id)::sub) all,
+ cstrs)
| _ ->
- (rest,Idmap.add id [a,expand_full_opt aliases a,id] l))
+ (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))
+ 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
@@ -322,127 +388,152 @@ let push_rel_context_to_named_context env typ =
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)
+ (named_context_val env, substl subst typ, inst_rels@inst_vars, subst)
(* [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=(dummy_loc,InternalHole)) ?filter typ =
- let sign,typ',instance = push_rel_context_to_named_context env typ in
- new_evar_instance sign evd typ' ~src:src ?filter instance
+let new_evar evd env ?(src=(dummy_loc,InternalHole)) ?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 instance =
+ match filter with
+ | None -> instance
+ | Some filter -> snd (list_filter2 (fun b c -> b) (filter,instance)) in
+ new_evar_instance sign evd typ' ~src:src ?filter ?candidates instance
(* The same using side-effect *)
-let e_new_evar evdref env ?(src=(dummy_loc,InternalHole)) ?filter ty =
- let (evd',ev) = new_evar !evdref env ~src:src ?filter ty in
+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
+(* This assumes an evar with identity instance and generalizes it over only
+ the de Bruijn part of the context *)
+let generalize_evar_over_rels sigma (ev,args) =
+ let evi = Evd.find sigma ev in
+ let sign = named_context_of_val evi.evar_hyps in
+ List.fold_left2
+ (fun (c,inst as x) a d ->
+ if isRel a then (mkNamedProd_or_LetIn d c,a::inst) else x)
+ (evi.evar_concl,[]) (Array.to_list args) sign
+
(*------------------------------------*
* operations on the evar constraints *
*------------------------------------*)
-(* Pb: defined Rels and Vars should not be considered as a pattern... *)
-(*
-let is_pattern inst =
- let rec is_hopat l = function
- [] -> true
- | t :: tl ->
- (isRel t or isVar t) && not (List.mem t l) && is_hopat (t::l) tl in
- is_hopat [] (Array.to_list inst)
-*)
+exception IllTypedFilter
+
+let check_restricted_occur evd refine env filter constr =
+ let filter = Array.of_list filter in
+ let rec aux k c =
+ let c = whd_evar evd c in
+ match kind_of_term c with
+ | Var id ->
+ let idx = list_try_find_i (fun i (id', _, _) -> if id' = id then i else raise (Failure "")) 0 env in
+ if not filter.(idx)
+ then if refine then
+ (filter.(idx) <- true; c)
+ else raise IllTypedFilter
+ else c
+ | _ -> map_constr_with_binders succ aux k c
+ in
+ let res = aux 0 constr in
+ Array.to_list filter, res
+
+(* 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
+ inst_in_sign =
+ 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 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
+(* 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
+ * virtually have x1..xq, y1'..yk' | ?e1' : Ï„' and the equation
* Γ, y1..yk |- ?e1'[u1..uq y1..yk] = c.
- * What we do is to formally introduce ?e1' in context x1..xq, Γ, y1..yk,
- * but forbidding it to use the variables of Γ (otherwise said,
- * Γ is here only for ensuring the correct typing of ?e1').
+ * [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
*
- * In fact, we optimize a little and try to compute a maximum
- * common subpart of x1..xq and Γ. This is done by detecting the
- * longest subcontext x1..xp such that Γ = x1'..xp' z1..zm and
- * u1..up = x1'..xp'.
- *
- * At the end, we return ?e1'[x1..xn z1..zm y1..yk] so that ?e1 can be
- * instantiated by (...\y1 ... \yk ... ?e1[x1..xn z1..zm y1..yk]) and the
- * new problem is Σ; Γ, y1..yk |- ?e1'[u1..un z1..zm y1..yk] = c,
- * making the z1..zm unavailable.
- *
- * This is what [extend_evar Γ evd k (?e1[u1..uq]) c] does.
+ * 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 shrink_context env subst ty =
- let rev_named_sign = List.rev (named_context env) in
- let rel_sign = rel_context env in
- (* We merge the contexts (optimization) *)
- let rec shrink_rel i subst rel_subst rev_rel_sign =
- match subst,rev_rel_sign with
- | (id,c)::subst,_::rev_rel_sign when c = mkRel i ->
- shrink_rel (i-1) subst (mkVar id::rel_subst) rev_rel_sign
- | _ ->
- substl_rel_context rel_subst (List.rev rev_rel_sign),
- substl rel_subst ty
- in
- let rec shrink_named subst named_subst rev_named_sign =
- match subst,rev_named_sign with
- | (id,c)::subst,(id',b',t')::rev_named_sign when c = mkVar id' ->
- shrink_named subst ((id',mkVar id)::named_subst) rev_named_sign
- | _::_, [] ->
- let nrel = List.length rel_sign in
- let rel_sign, ty = shrink_rel nrel subst [] (List.rev rel_sign) in
- [], map_rel_context (replace_vars named_subst) rel_sign,
- replace_vars named_subst ty
- | _ ->
- map_named_context (replace_vars named_subst) (List.rev rev_named_sign),
- rel_sign, ty
+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 (snd (list_filter2 (fun b id -> b) (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) ->
+ match b with
+ | None ->
+ 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 inst_in_sign in
+ (push_named_context_val (id,None,t_in_sign) sign,true::filter,
+ (mkRel 1)::(List.map (lift 1) inst_in_env),(mkVar id)::inst_in_sign,
+ push_rel d env,evd,id::avoid)
+ | Some b ->
+ (sign,filter,inst_in_env,inst_in_sign,
+ push_rel d env,evd,avoid))
+ rel_sign
+ (sign1,filter1,Array.to_list args1,inst_in_sign,env1,evd,ids1)
in
- shrink_named subst [] rev_named_sign
-
-let extend_evar env evdref k (evk1,args1) c =
- let ty = get_type_of env !evdref c in
- let overwrite_first v1 v2 =
- let v = Array.copy v1 in
- let n = Array.length v - Array.length v2 in
- for i = 0 to Array.length v2 - 1 do v.(n+i) <- v2.(i) done;
- v in
- let evi1 = Evd.find !evdref evk1 in
- let named_sign',rel_sign',ty =
- if k = 0 then [], [], ty
- else shrink_context env (List.rev (make_pure_subst evi1 args1)) ty in
- let extenv =
- List.fold_right push_rel rel_sign'
- (List.fold_right push_named named_sign' (evar_unfiltered_env evi1)) in
- let nb_to_hide = rel_context_length rel_sign' - k in
- let rel_filter = list_map_i (fun i _ -> i > nb_to_hide) 1 rel_sign' in
- let named_filter1 = List.map (fun _ -> true) (evar_context evi1) in
- let named_filter2 = List.map (fun _ -> false) named_sign' in
- let filter = rel_filter@named_filter2@named_filter1 in
- let evar1' = e_new_evar evdref extenv ~filter:filter ty in
- let evk1',args1'_in_env = destEvar evar1' in
- let args1'_in_extenv = Array.map (lift k) (overwrite_first args1'_in_env args1) in
- (evar1',(evk1',args1'_in_extenv))
-
-let subfilter p filter l =
- let (filter,_,l) =
- List.fold_left (fun (filter,l,newl) b ->
- if b then
- let a,l' = match l with a::args -> a,args | _ -> assert false in
- if p a then (true::filter,l',a::newl) else (false::filter,l',newl)
- else (false::filter,l,newl))
- ([],l,[]) filter in
- (List.rev filter,List.rev l)
-
-let restrict_upon_filter evd evi evk p args =
+ let evd,ev2ty_in_sign =
+ define_evar_from_virtual_equation define_fun env evd ty_in_env
+ sign2 filter2 inst2_in_env inst2_in_sign 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 subfilter env ccl filter newfilter args =
+ let vars = collect_vars ccl in
+ let (filter, _, _, newargs) =
+ List.fold_left2
+ (fun (filter, newl, args, newargs) oldf (n, _, _) ->
+ if oldf then
+ let a, oldargs = match args with hd :: tl -> hd, tl | _ -> assert false in
+ if Idset.mem n vars then
+ (oldf :: filter, List.tl newl, oldargs, a :: newargs)
+ else if List.hd newl then (true :: filter, List.tl newl, oldargs, a :: newargs)
+ else (false :: filter, List.tl newl, oldargs, newargs)
+ else (oldf :: filter, newl, args, newargs))
+ ([], newfilter, args, []) filter env
+ in List.rev filter, List.rev newargs
+
+let restrict_upon_filter ?(refine=false) evd evi evk p args =
let filter = evar_filter evi in
- let newfilter,newargs = subfilter p filter args in
+ let newfilter = List.map p args in
+ let env = evar_unfiltered_env evi in
+ let ccl = nf_evar evd evi.evar_concl in
+ let newfilter, newargs =
+ subfilter (named_context env) ccl filter newfilter args
+ in
if newfilter <> filter then
- let (evd,newev) = new_evar evd (evar_unfiltered_env evi) ~src:(evar_source evk evd)
- ~filter:newfilter evi.evar_concl in
+ let (evd,newev) = new_evar evd env ~src:(evar_source evk evd)
+ ~filter:newfilter ccl in
let evd = Evd.define evk newev evd in
- evd,fst (destEvar newev),newargs
+ evd,fst (destEvar newev), newargs
else
evd,evk,args
@@ -460,6 +551,10 @@ type clear_dependency_error =
exception ClearDependencyError of identifier * clear_dependency_error
+open Store.Field
+
+let cleared = Store.field ()
+
let rec check_and_clear_in_constr 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
@@ -477,7 +572,7 @@ let rec check_and_clear_in_constr evdref err ids c =
List.iter check vars; c
| Evar (evk,l as ev) ->
- if Evd.is_defined_evar !evdref ev then
+ 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)
@@ -487,7 +582,7 @@ let rec check_and_clear_in_constr evdref err ids c =
arguments. Concurrently, we build a new evar
corresponding to e where hypotheses of ids have been
removed *)
- let evi = Evd.find !evdref evk in
+ let evi = Evd.find_undefined !evdref evk in
let ctxt = Evd.evar_filtered_context evi in
let (nhyps,nargs,rids) =
List.fold_right2
@@ -521,6 +616,13 @@ let rec check_and_clear_in_constr evdref err ids c =
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
+ (* 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 evi' = { evi with evar_extra = extra' } in
+ evdref := Evd.add !evdref evk evi' ;
+ (* spiwack: /hacking session *)
mkEvar(evk', Array.of_list nargs)
end
@@ -553,6 +655,20 @@ let clear_hyps_in_evi evdref hyps concl ids =
in
(nhyps,nconcl)
+(* 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
@@ -597,17 +713,17 @@ 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 y = c' then id
+ 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 expand_full_opt aliases c') with
- | Some cc when yc = cc -> id
- | _ -> raise Not_found
+ 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 = expand_var aliases y in
+ 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
@@ -623,7 +739,7 @@ let rec find_projectable_vars with_evars aliases sigma y subst =
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 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'
@@ -680,7 +796,7 @@ let rec do_projection_effects define_fun env ty evd = function
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 (destEvar ty') ty evd else evd
+ if isEvar ty' then define_fun env evd (destEvar ty') ty else evd
else
evd
@@ -741,7 +857,7 @@ let effective_projections =
map_succeed (function Invertible c -> c | _ -> failwith"")
let instance_of_projection f env t evd projs =
- let ty = lazy (Retyping.get_type_of env evd t) in
+ let ty = lazy (nf_evar evd (Retyping.get_type_of env evd t)) in
match projs with
| NoUniqueProjection -> raise NotUnique
| UniqueProjection (c,effects) ->
@@ -767,7 +883,7 @@ let filter_along f projs v =
* such that "hyps' |- ?e : T"
*)
-let restrict_hyps evd evk filter =
+let restrict_hyps ?(refine=false) evd evk filter =
(* 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
@@ -783,8 +899,9 @@ let restrict_hyps evd evk filter =
let oldfilter = evar_filter evi in
let filter,_ = List.fold_right (fun oldb (l,filter) ->
if oldb then List.hd filter::l,List.tl filter else (false::l,filter))
- oldfilter ([],List.rev filter) in
- (env,evar_source evk evd,filter,evi.evar_concl)
+ oldfilter ([], List.rev filter) in
+ let filter, ccl = check_restricted_occur evd refine (named_context env) filter evi.evar_concl in
+ (env,evar_source evk evd,filter,ccl)
let do_restrict_hyps evd evk projs =
let filter = List.map filter_of_projection projs in
@@ -792,14 +909,18 @@ let do_restrict_hyps evd evk projs =
evd,evk
else
let env,src,filter,ccl = restrict_hyps evd evk filter in
- let evd,nc = new_evar evd env ~src ~filter ccl in
- let evd = Evd.define evk nc evd in
- let evk',_ = destEvar nc in
- evd,evk'
+ if List.for_all (fun x -> x) filter then
+ evd,evk
+ else
+ let evd,nc = new_evar evd env ~src ~filter ccl in
+ let evd = Evd.define evk nc evd in
+ let evk',_ = destEvar nc in
+ evd,evk'
(* [postpone_evar_term] postpones an equation of the form ?e[σ] = c *)
let postpone_evar_term env evd (evk,argsv) rhs =
+ assert (isVar rhs or isRel rhs);
let rhs = expand_vars_in_term env rhs in
let evi = Evd.find evd evk in
let evd,evk,args =
@@ -850,44 +971,71 @@ let are_canonical_instances args1 args2 env =
let n2 = Array.length args2 in
let rec aux n = function
| (id,_,c)::sign
- when n < n1 && args1.(n) = mkVar id && args1.(n) = args2.(n) ->
+ when n < n1 && isVar args1.(n) && destVar args1.(n) = id && eq_constr args1.(n) args2.(n) ->
aux (n+1) sign
| [] ->
let rec aux2 n =
n = n1 ||
- (args1.(n) = mkRel (n1-n) && args2.(n) = mkRel (n1-n) && aux2 (n+1))
+ (isRel args1.(n) && destRel args1.(n) = n1-n &&
+ isRel args2.(n) && destRel args2.(n) = n1-n && aux2 (n+1))
in aux2 n
| _ -> false in
n1 = n2 & aux 0 (named_context env)
exception CannotProject of projectibility_status list
+let is_variable_subst args =
+ array_for_all (fun c -> isRel c || isVar c) args
+
let solve_evar_evar_l2r f env evd (evk1,args1) (evk2,args2 as ev2) =
let aliases = make_alias_map env in
- let subst = make_projectable_subst aliases evd (Evd.find evd evk2) args2 in
+ let subst,_ = make_projectable_subst aliases evd (Evd.find evd evk2) args2 in
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 subst = List.map (fun (id,_,_) -> mkVar id) sign in
Evd.define evk2 (mkEvar(evk1,Array.of_list subst)) evd
else
- let proj1 = array_map_to_list (invert_arg aliases 0 evd evk2 subst) args1 in
- try
- (* Instantiate ev2 with (a restriction of) ev1 if uniquely projectable *)
- let proj1' = effective_projections proj1 in
- let evd,args1' =
- list_fold_map (instance_of_projection f env (mkEvar ev2)) evd proj1' in
- let evd,evk1' = do_restrict_hyps evd evk1 proj1 in
- Evd.define evk2 (mkEvar(evk1',Array.of_list args1')) evd
- with NotUnique ->
- raise (CannotProject proj1)
+ (* Only try pruning on variable substitutions, postpone otherwise. *)
+ if is_variable_subst args1 && is_variable_subst args2 then
+ let proj1 = array_map_to_list (invert_arg aliases 0 evd evk2 subst) args1 in
+ try
+ (* Instantiate ev2 with (a restriction of) ev1 if uniquely projectable.
+ Rules out non-linear instances. *)
+ let proj1' = effective_projections proj1 in
+ let evd,args1' =
+ list_fold_map (instance_of_projection f env (mkEvar ev2)) evd proj1' in
+ let evd,evk1' = do_restrict_hyps evd evk1 proj1 in
+ Evd.define evk2 (mkEvar(evk1',Array.of_list args1')) evd
+ with NotUnique -> raise (CannotProject proj1)
+ else raise IllTypedFilter
let solve_evar_evar f env evd ev1 ev2 =
- try solve_evar_evar_l2r f env evd ev1 ev2
- with CannotProject projs1 ->
- try solve_evar_evar_l2r f env evd ev2 ev1
- with CannotProject projs2 ->
- postpone_evar_evar env evd projs1 ev1 projs2 ev2
+ try
+ try solve_evar_evar_l2r f env evd ev1 ev2
+ with CannotProject projs1 ->
+ try solve_evar_evar_l2r f env evd ev2 ev1
+ with CannotProject projs2 ->
+ postpone_evar_evar env evd projs1 ev1 projs2 ev2
+ with IllTypedFilter ->
+ let pb = (Reduction.CONV,env,mkEvar(ev1),mkEvar (ev2)) in
+ add_conv_pb pb evd
+
+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 _ -> 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 (?i x1..xn) = (?i y1..yn) which arises often in fixpoint
* definitions. We try to unify the xi with the yi pairwise. The pairs
@@ -895,8 +1043,8 @@ let solve_evar_evar f env evd ev1 ev2 =
* depend on these args). *)
let solve_refl conv_algo env evd evk argsv1 argsv2 =
- if argsv1 = argsv2 then evd else
- let evi = Evd.find evd evk in
+ if array_equal eq_constr argsv1 argsv2 then evd else
+ let evi = Evd.find_undefined evd evk in
(* Filter and restrict if needed *)
let evd,evk,args =
restrict_upon_filter evd evi evk
@@ -906,7 +1054,37 @@ let solve_refl conv_algo env evd evk argsv1 argsv2 =
let args1,args2 = List.split args in
let argsv1 = Array.of_list args1 and argsv2 = Array.of_list args2 in
let pb = (Reduction.CONV,env,mkEvar(evk,argsv1),mkEvar(evk,argsv2)) in
- Evd.add_conv_pb pb evd
+ Evd.add_conv_pb pb 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 (fun c ->
+ let c' = instantiate_evar (evar_filtered_context evi) c args in
+ let evd, b = conv_algo env evd Reduction.CONV c' rhs in
+ if b then Some (c,evd) else None) l in
+ match l' with
+ | [] -> error_cannot_unify env evd (mkEvar ev, rhs)
+ | [c,evd] -> Evd.define evk c evd
+ | l when List.length l < List.length l' ->
+ let candidates = List.map fst l in
+ let filter = evar_filter evi in
+ let sign = evar_hyps evi in
+ let ids = List.map pi1 (named_context_of_val sign) in
+ let inst_in_sign =
+ List.map mkVar (snd (list_filter2 (fun b id -> b) (filter,ids))) in
+ let evd,evar = new_evar_instance (evar_hyps evi) evd (evar_concl evi)
+ ~filter ~candidates inst_in_sign in
+ Evd.define evk evar evd
+ | 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
@@ -925,7 +1103,7 @@ let solve_refl conv_algo env evd evk argsv1 argsv2 =
* 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]
+ * 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)
@@ -935,12 +1113,12 @@ exception NotInvertibleUsingOurAlgorithm of constr
exception NotEnoughInformationToProgress
exception OccurCheckIn of evar_map * constr
-let rec invert_definition choose env evd (evk,argsv as ev) rhs =
+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 = make_projectable_subst aliases evd evi argsv in
+ let subst,cstr_subst = make_projectable_subst aliases evd evi argsv in
(* Projection *)
let project_variable t =
@@ -955,7 +1133,7 @@ let rec invert_definition choose env evd (evk,argsv as ev) rhs =
else raise (NotUniqueInType(find_solution_type (evar_env evi) sols))
in
let ty = lazy (Retyping.get_type_of env !evdref t) in
- let evd = do_projection_effects evar_define env ty !evdref p in
+ let evd = do_projection_effects (evar_define conv_algo) env ty !evdref p in
evdref := evd;
c
with
@@ -963,11 +1141,16 @@ let rec invert_definition choose env evd (evk,argsv as ev) rhs =
| NotUniqueInType ty ->
if not !progress then raise NotEnoughInformationToProgress;
(* No unique projection but still restrict to where it is possible *)
+ (* materializing is necessary, but is restricting useful? *)
+ let sign = evar_filtered_context evi in
+ let ty' = instantiate_evar sign ty (Array.to_list argsv) in
+ let (evd,_,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 evarenv,src,filter,_ = restrict_hyps ~refine:true evd (fst ev') filter in
let args' = filter_along (fun x -> x) filter argsv in
- let evarenv,src,filter,_ = restrict_hyps !evdref evk filter in
let evd,evar = new_evar !evdref evarenv ~src ~filter ty in
let evk',_ = destEvar evar in
let pb = (Reduction.CONV,env,mkEvar(evk',args'),t) in
@@ -986,32 +1169,52 @@ let rec invert_definition choose env evd (evk,argsv as ev) rhs =
array_map_to_list
(invert_arg_from_subst aliases k !evdref subst) args'
in
- (try
- (* Try to project (a restriction of) the right evar *)
- let eprojs' = effective_projections projs' in
- let evd,args' =
- list_fold_map (instance_of_projection evar_define env' t)
- !evdref eprojs' in
- let evd,evk' = do_restrict_hyps evd evk' projs' in
- evdref := evd;
- mkEvar (evk',Array.of_list args')
- with NotUnique ->
- assert !progress;
- (* Make the virtual left evar real *)
- let (evar'',ev'') = extend_evar env' evdref k ev t in
- let evd =
- (* Try to project (a restriction of) the left evar ... *)
- try solve_evar_evar_l2r evar_define env' !evdref ev'' ev'
- with CannotProject projs'' ->
- (* ... or postpone the problem *)
- postpone_evar_evar env' !evdref projs'' ev'' projs' ev' in
- evdref := evd;
- evar'')
+ (try
+ (* Try to project (a restriction of) the right evar *)
+ let eprojs' = effective_projections projs' in
+ let evd,args' =
+ list_fold_map (instance_of_projection (evar_define conv_algo) env' t)
+ !evdref eprojs' in
+ let evd,evk' = do_restrict_hyps evd evk' projs' in
+ evdref := evd;
+ mkEvar (evk',Array.of_list args')
+ with NotUnique | IllTypedFilter ->
+ assert !progress;
+ (* 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
+ (try
+ let evd =
+ (* Try to project (a restriction of) the left evar ... *)
+ try solve_evar_evar_l2r (evar_define conv_algo) env' evd ev'' ev'
+ with CannotProject projs'' ->
+ (* ... or postpone the problem *)
+ postpone_evar_evar env' evd projs'' ev'' projs' ev'
+ in
+ evdref := evd;
+ evar''
+ with IllTypedFilter -> raise (NotInvertibleUsingOurAlgorithm t)))
| _ ->
- progress := true;
- (* 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
+ 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
+ | [id] -> Some (mkVar id)
+ | _ -> None)
+ | _ -> None
+ with
+ | Some c -> c
+ | None ->
+ progress := true;
+ (* 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
@@ -1024,15 +1227,16 @@ let rec invert_definition choose env evd (evk,argsv as ev) rhs =
* context "hyps" and not referring to itself.
*)
-and occur_existential evm c =
- let rec occrec c = match kind_of_term c with
- | Evar (e, _) -> if not (is_defined evm e) then raise Occur
- | _ -> iter_constr occrec c
- in try occrec c; false with Occur -> true
-
-and evar_define ?(choose=false) env (evk,argsv as ev) rhs evd =
+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 conv_algo env evd evk argsv argsv2
+ else solve_evar_evar (evar_define conv_algo) env evd ev ev2
+ | _ ->
+ try solve_candidates conv_algo env evd ev rhs
+ with NoCandidates ->
try
- let (evd',body) = invert_definition choose env evd ev rhs in
+ 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 *)
@@ -1055,7 +1259,8 @@ and evar_define ?(choose=false) env (evk,argsv as ev) rhs evd =
str "----> " ++ int ev ++ str " := " ++
print_constr body);
raise e in*)
- Evd.define evk body evd'
+ let evd' = Evd.define evk body evd' in
+ check_evar_instance evd' evk body conv_algo
with
| NotEnoughInformationToProgress ->
postpone_evar_term env evd ev rhs
@@ -1133,73 +1338,108 @@ let whd_head_evar_stack sigma c =
let whd_head_evar sigma c = applist (whd_head_evar_stack sigma c)
-(* Check if an applied evar "?X[args] l" is a Miller's pattern; note
- that we don't care whether args itself contains Rel's or even Rel's
- distinct from the ones in l *)
-
-let rec expand_and_check_vars env = function
+let rec expand_and_check_vars aliases = function
| [] -> []
- | a::l ->
- if isRel a or isVar a then
- let l = expand_and_check_vars env l in
- match expand_var_opt env a with
- | None -> a :: l
- | Some a' when isRel a' or isVar a' -> list_add_set a' l
- | _ -> raise Exit
- else
- raise Exit
+ | 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 is_unification_pattern_evar env (_,args) l t =
- List.for_all (fun x -> isRel x || isVar x) l (* common failure case *)
- &&
+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
- let l' = Array.to_list args @ l in
- let l'' = try Some (expand_and_check_vars aliases l') with Exit -> None in
- match l'' with
- | Some l ->
- let deps =
- 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 t = expand_vars_in_term_using aliases t in
- let fv_rels = free_rels t in
- let fv_ids = global_vars env t in
- List.filter (fun c ->
- match kind_of_term c with
- | Var id -> List.mem id fv_ids
- | Rel n -> Intset.mem n fv_rels
- | _ -> assert false) l in
- list_distinct deps
- | None -> false
-
-let is_unification_pattern (env,nb) f l t =
- match kind_of_term f with
- | Meta _ ->
- array_for_all (fun c -> isRel c && destRel c <= nb) l
- && array_distinct l
- | Evar ev ->
- is_unification_pattern_evar env ev (Array.to_list l) t
- | _ ->
- false
+ 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 (* common failure case *) 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 when not (occur_evar evk t) -> Some (list_skipn n l)
+ | _ -> None
+ else
+ None
-(* From a unification problem "?X l1 = term1 l2" such that l1 is made
- of distinct rel's, build "\x1...xn.(term1 l2)" (patterns unification) *)
-(* NB: does not work when (term1 l2) contains metas because metas
+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 l1 c =
- let l1 = List.map (expand_var (make_alias_map env)) l1 in
+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 (na,_,t) = lookup_rel n env in mkLambda (na,lift n t,c')
- | Var id -> let (id,_,t) = lookup_named id env in mkNamedLambda id t c'
+ | 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)
- l1 c in
+ l c in
(* Warning: we may miss some opportunity to eta-reduce more since c'
is not in normal form *)
whd_eta c'
@@ -1234,23 +1474,12 @@ 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)
-(* Util *)
-
-let check_instance_type conv_algo env evd ev1 t2 =
- let t2 = nf_evar evd t2 in
- if has_undefined_evars_or_sorts evd t2 then
- (* May contain larger constraints than needed: don't want to
- commit to an equal solution while only subtyping is requested *)
- evd
- else
- let typ2 = Retyping.get_type_of env evd (refresh_universes t2) in
- if isEvar typ2 then (* Don't want to commit too early too *) evd
- else
- let typ1 = existential_type evd ev1 in
- let evd,b = conv_algo env evd Reduction.CUMUL typ2 typ1 in
- if b then evd else
- user_err_loc (fst (evar_source (fst ev1) evd),"",
- str "Unable to find a well-typed instantiation")
+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
@@ -1261,51 +1490,69 @@ let check_instance_type conv_algo env evd ev1 t2 =
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 kind_of_term t2 with
- | Evar (evk2,args2 as ev2) ->
- if evk1 = evk2 then
- solve_refl conv_algo env evd evk1 args1 args2
- else
- if pbty = None
- then solve_evar_evar evar_define env evd ev1 ev2
- else if pbty = Some true then
- add_conv_pb (Reduction.CUMUL,env,mkEvar ev1,t2) evd
- else
- add_conv_pb (Reduction.CUMUL,env,t2,mkEvar ev1) evd
+ 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
| _ ->
- let evd =
- if pbty = Some false then
- check_instance_type conv_algo env evd ev1 t2
- else
- evd in
- let evd = evar_define ~choose env ev1 t2 evd in
- let evi = Evd.find evd evk1 in
- if occur_existential evd evi.evar_concl then
- let evenv = evar_unfiltered_env evi in
- let evc = nf_evar evd evi.evar_concl in
- match evi.evar_body with
- | Evar_defined body ->
- let ty = nf_evar evd (Retyping.get_type_of evenv evd body) in
- add_conv_pb (Reduction.CUMUL,evenv,ty,evc) evd
- | Evar_empty -> (* Resulted in a constraint *)
- evd
- else evd
- in
- 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
+ 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, _) -> Intset.add n acc
+ | Evar (n, l) -> Intset.add n (Array.fold_left evrec acc l)
| _ -> fold_constr evrec acc c
in
- evrec Intset.empty c
+ evrec Intset.empty c
+
+(* spiwack: a few functions to gather the existential variables
+ that occur in the types of goals present or past. *)
+let add_evars_of_evars_of_term acc evm c =
+ let evars = evars_of_term c in
+ Intset.fold begin fun e r ->
+ let body = (Evd.find evm e).evar_body in
+ let subevars =
+ match body with
+ | Evar_empty -> None
+ | Evar_defined c' -> Some (evars_of_term c')
+ in
+ Intmap.add e subevars r
+ end evars acc
+
+let evars_of_evars_of_term = add_evars_of_evars_of_term Intmap.empty
+
+let add_evars_of_evars_in_type acc evm e =
+ let evi = Evd.find evm e in
+ let acc_with_concl = add_evars_of_evars_of_term acc evm evi.evar_concl in
+ let hyps = Environ.named_context_of_val evi.evar_hyps in
+ List.fold_left begin fun r (_,b,t) ->
+ let r = add_evars_of_evars_of_term r evm t in
+ match b with
+ | None -> r
+ | Some b -> add_evars_of_evars_of_term r evm b
+ end acc_with_concl hyps
+
+let rec add_evars_of_evars_in_types_of_set acc evm s =
+ Intset.fold begin fun e r ->
+ let r = add_evars_of_evars_in_type r evm e in
+ match (Evd.find evm e).evar_body with
+ | Evar_empty -> r
+ | Evar_defined b -> add_evars_of_evars_in_types_of_set r evm (evars_of_term b)
+ end s acc
+
+let evars_of_evars_in_types_of_list evm l =
+ let set_of_l = List.fold_left (fun x y -> Intset.add y x) Intset.empty l in
+ add_evars_of_evars_in_types_of_set Intmap.empty evm set_of_l
+
+(* /spiwack *)
let evars_of_named_context nc =
List.fold_right (fun (_, b, t) s ->
@@ -1322,34 +1569,60 @@ let evars_of_evar_info evi =
| 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
+ [nf_evar]. *)
+
+let undefined_evars_of_term evd t =
+ let rec evrec acc c =
+ match kind_of_term c with
+ | 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_defined c -> evrec acc c
+ with Not_found -> anomaly "undefined_evars_of_term: evar not found")
+ | _ -> fold_constr evrec acc c
+ in
+ evrec Intset.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
+
+let undefined_evars_of_evar_info evd evi =
+ Intset.union (undefined_evars_of_term evd evi.evar_concl)
+ (Intset.union
+ (match evi.evar_body with
+ | Evar_empty -> Intset.empty
+ | Evar_defined b -> undefined_evars_of_term evd b)
+ (undefined_evars_of_named_context evd
+ (named_context_of_val evi.evar_hyps)))
+
(* [check_evars] fails if some unresolved evar remains *)
-(* it assumes that the defined existentials have already been substituted *)
-let check_evars env initial_sigma evd c =
- let sigma = evd in
- let c = nf_evar sigma c in
+let check_evars env initial_sigma sigma c =
let rec proc_rec c =
match kind_of_term c with
- | Evar (evk,args) ->
- assert (Evd.mem sigma evk);
+ | Evar (evk,_ as ev) ->
+ (match existential_opt_value sigma ev with
+ | Some c -> proc_rec c
+ | None ->
if not (Evd.mem initial_sigma evk) then
let (loc,k) = evar_source evk sigma in
- (match k with
+ match k with
| ImplicitArg (gr, (i, id), false) -> ()
| _ ->
- let evi = nf_evar_info sigma (Evd.find sigma evk) in
+ let evi = nf_evar_info sigma (Evd.find_undefined sigma evk) in
error_unsolvable_implicit loc env sigma evi k None)
| _ -> iter_constr proc_rec c
in proc_rec c
-(* This returns the evars of [sigma] that are not in [sigma0] and
- [sigma] minus these evars *)
-
-let subtract_evars sigma0 sigma =
- Evd.fold (fun evk ev (sigma,sigma' as acc) ->
- if Evd.mem sigma0 evk || Evd.mem sigma' evk then acc else
- (Evd.remove sigma evk,Evd.add sigma' evk ev))
- sigma (sigma,Evd.empty)
+open Glob_term
(* Operations on value/type constraints *)
@@ -1390,44 +1663,96 @@ let empty_valcon = None
(* Builds a value constraint *)
let mk_valcon c = Some c
-(* Refining an evar to a product or a sort *)
-(* Declaring any type to be in the sort Type shouldn't be harmful since
- cumulativity now includes Prop and Set in Type...
- It is, but that's not too bad *)
-let define_evar_as_abstraction abs evd (ev,args) =
- let evi = Evd.find evd ev in
+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 idx = id_of_string "x"
+
+(* 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 (evd1,dom) = new_evar evd evenv (new_Type()) ~filter:(evar_filter evi) in
- let nvar =
- next_ident_away (id_of_string "x")
- (ids_of_named_context (evar_context evi)) in
- let newenv = push_named (nvar, None, dom) evenv in
- let (evd2,rng) =
- new_evar evd1 newenv ~src:(evar_source ev evd1) (new_Type())
- ~filter:(true::evar_filter evi) in
- let prod = abs (Name nvar, dom, subst_var nvar rng) in
- let evd3 = Evd.define ev prod evd2 in
- let evdom = fst (destEvar dom), args in
- let evrng =
- fst (destEvar rng), array_cons (mkRel 1) (Array.map (lift 1) args) in
- let prod' = abs (Name nvar, mkEvar evdom, mkEvar evrng) in
- (evd3,prod')
-
-let define_evar_as_product evd (ev,args) =
- define_evar_as_abstraction (fun t -> mkProd t) evd (ev,args)
-
-let define_evar_as_lambda evd (ev,args) =
- define_evar_as_abstraction (fun t -> mkLambda t) evd (ev,args)
+ 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 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 prod = mkProd (Name id, dom, subst_var id rng) in
+ let evd3 = Evd.define evk prod evd2 in
+ evd3,prod
+
+(* Refine an applied evar to a product and returns its instantiation *)
+
+let define_evar_as_product evd (evk,args) =
+ let evd,prod = define_pure_evar_as_product evd evk in
+ (* Quick way to compute the instantiation of evk with args *)
+ let na,dom,rng = destProd prod in
+ let evdom = mkEvar (fst (destEvar dom), args) in
+ let evrngargs = array_cons (mkRel 1) (Array.map (lift 1) args) in
+ let evrng = mkEvar (fst (destEvar rng), evrngargs) in
+ evd,mkProd (na, evdom, evrng)
+
+(* Refine an evar with an abstraction
+
+ I.e., solve x1..xq |- ?e:T(x1..xq) with e:=λy:A.?e'[x1..xq,y] where:
+ - either T(x1..xq) = πy:A(x1..xq).B(x1..xq,y)
+ or T(x1..xq) = ?d[x1..xq] and we define ?d := πy:?A.?B
+ with x1..xq |- ?A:Type and x1..xq,y |- ?B:Type
+ - x1..xq,y:A |- ?e':B
+*)
+
+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 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
+ 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 src = evar_source evk evd1 in
+ let evd2,body = new_evar evd1 newenv ~src (subst1 (mkVar id) rng) ~filter in
+ let lam = mkLambda (Name id, dom, subst_var id body) in
+ Evd.define evk lam evd2, lam
+
+let define_evar_as_lambda env evd (evk,args) =
+ let evd,lam = define_pure_evar_as_lambda env evd evk in
+ (* Quick way to compute the instantiation of evk with args *)
+ let na,dom,body = destLambda lam in
+ let evbodyargs = array_cons (mkRel 1) (Array.map (lift 1) args) in
+ let evbody = mkEvar (fst (destEvar body), evbodyargs) in
+ evd,mkLambda (na, dom, evbody)
+
+let rec evar_absorb_arguments env evd (evk,args as ev) = function
+ | [] -> evd,ev
+ | a::l ->
+ (* TODO: optimize and avoid introducing intermediate evars *)
+ let evd,lam = define_pure_evar_as_lambda env evd evk in
+ let _,_,body = destLambda lam in
+ let evk = fst (destEvar body) in
+ 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 s = new_Type () in
- Evd.define ev s evd, destSort s
+ let evd, s = new_sort_variable evd in
+ Evd.define ev (mkSort s) evd, 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 () = Typeops.judge_of_type (new_univ ())
+let judge_of_new_Type evd =
+ let evd', s = new_univ_variable evd in
+ evd', Typeops.judge_of_type s
(* Propagation of constraints through application and abstraction:
Given a type constraint on a functional term, returns the type
@@ -1443,10 +1768,13 @@ let split_tycon loc env evd tycon =
let t = whd_betadeltaiota env evd c in
match kind_of_term t with
| Prod (na,dom,rng) -> evd, (na, dom, rng)
- | Evar ev when not (Evd.is_defined_evar evd ev) ->
+ | Evar ev (* ev is undefined because of whd_betadeltaiota *) ->
let (evd',prod) = define_evar_as_product evd ev in
let (_,dom,rng) = destProd prod in
evd',(Anonymous, dom, rng)
+ | App (c,args) when isEvar c ->
+ let (evd',lam) = define_evar_as_lambda env evd (destEvar c) in
+ real_split evd' (mkApp (lam,args))
| _ -> error_not_product_loc loc env evd c
in
match tycon with
diff --git a/pretyping/evarutil.mli b/pretyping/evarutil.mli
index e29effc2..61f503c7 100644
--- a/pretyping/evarutil.mli
+++ b/pretyping/evarutil.mli
@@ -1,117 +1,148 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: evarutil.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
open Util
open Names
-open Rawterm
+open Glob_term
open Term
open Sign
open Evd
open Environ
open Reductionops
-(*i*)
-(*s This modules provides useful functions for unification modulo evars *)
+(** {5 This modules provides useful functions for unification modulo evars } *)
-(***********************************************************)
-(* Metas *)
+(** {6 Metas} *)
-(* [new_meta] is a generator of unique meta variables *)
+(** [new_meta] is a generator of unique meta variables *)
val new_meta : unit -> metavariable
val mk_new_meta : unit -> constr
-(* [new_untyped_evar] is a generator of unique evar keys *)
+(** [new_untyped_evar] is a generator of unique evar keys *)
val new_untyped_evar : unit -> existential_key
-(***********************************************************)
-(* Creating a fresh evar given their type and context *)
+(** {6 Creating a fresh evar given their type and context} *)
val new_evar :
- evar_map -> env -> ?src:loc * hole_kind -> ?filter:bool list -> types -> evar_map * constr
-(* the same with side-effects *)
+ evar_map -> env -> ?src:loc * hole_kind -> ?filter:bool list ->
+ ?candidates:constr list -> types -> evar_map * constr
+
+(** the same with side-effects *)
val e_new_evar :
- evar_map ref -> env -> ?src:loc * hole_kind -> ?filter:bool list -> types -> constr
+ evar_map ref -> env -> ?src:loc * hole_kind -> ?filter:bool list ->
+ ?candidates:constr list -> types -> constr
-(* Create a fresh evar in a context different from its definition context:
+(** 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
+
+(** 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
[sign] and type [ty], [inst] is a mapping of the evar context to
the context where the evar should occur. This means that the terms
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 -> constr list -> evar_map * constr
+ 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
-(***********************************************************)
-(* Instantiate evars *)
+(** {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]),
+(** [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 : ?choose:bool -> env -> existential -> constr -> evar_map -> evar_map
+val evar_define : conv_fun -> ?choose:bool -> env -> evar_map ->
+ existential -> constr -> evar_map
-(***********************************************************)
-(* Evars/Metas switching... *)
+(** {6 Evars/Metas switching...} *)
-(* [evars_to_metas] generates new metavariables for each non dependent
+(** [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
-(***********************************************************)
-(* Unification utils *)
+(** {6 Unification utils} *)
+(** [head_evar c] returns the head evar of [c] if any *)
exception NoHeadEvar
-val head_evar : constr -> existential_key (* may raise NoHeadEvar *)
+val head_evar : constr -> existential_key (** may raise NoHeadEvar *)
(* Expand head evar if any *)
val whd_head_evar : evar_map -> constr -> constr
val is_ground_term : evar_map -> constr -> bool
val is_ground_env : evar_map -> env -> bool
-val solve_refl :
- (env -> evar_map -> conv_pb -> constr -> constr -> evar_map * bool)
- -> env -> evar_map -> existential_key -> constr array -> constr array ->
- evar_map
-val solve_simple_eqn :
- (env -> evar_map -> conv_pb -> constr -> constr -> evar_map * bool)
- -> ?choose:bool -> env -> evar_map -> bool option * existential * constr ->
- evar_map * bool
-
-(* [check_evars env initial_sigma extended_sigma c] fails if some
+val solve_refl : conv_fun -> env -> evar_map ->
+ existential_key -> constr array -> constr array -> 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 subtract_evars : evar_map -> evar_map -> evar_map * evar_map
val define_evar_as_product : evar_map -> existential -> evar_map * types
-val define_evar_as_lambda : 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 -> existential -> constr list ->
- constr -> bool
-val is_unification_pattern : env * int -> constr -> constr array ->
- constr -> bool
+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 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
+
+(** returns the evars contained in the term associated with
+ the evars they contain themselves in their body, if any.
+ If the evar has no body, [None] is associated to it. *)
+val evars_of_evars_of_term : evar_map -> constr -> (Intset.t option) Intmap.t
val evars_of_named_context : named_context -> Intset.t
val evars_of_evar_info : evar_info -> Intset.t
-(***********************************************************)
-(* Value/Type constraints *)
+(** returns the evars which can be found in the typing context of the argument evars,
+ in the same format as {!evars_of_evars_of_term}.
+ It explores recursively the evars in the body of the argument evars -- but does
+ not return them. *)
+(* spiwack: tongue in cheek: it should have been called
+ [evars_of_evars_in_types_of_list_and_recursively_in_bodies] *)
+val evars_of_evars_in_types_of_list : evar_map -> evar list -> (Intset.t option) Intmap.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
+
+(** {6 Value/Type constraints} *)
-val judge_of_new_Type : unit -> unsafe_judgment
+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
@@ -139,8 +170,8 @@ val lift_tycon : int -> type_constraint -> type_constraint
(***********************************************************)
-(* [flush_and_check_evars] raise [Uninstantiated_evar] if an evar remains *)
-(* uninstantiated; [nf_evar] leave uninstantiated evars as is *)
+(** [flush_and_check_evars] raise [Uninstantiated_evar] if an evar remains
+ uninstantiated; [nf_evar] leaves uninstantiated evars as is *)
val nf_evar : evar_map -> constr -> constr
val j_nf_evar : evar_map -> unsafe_judgment -> unsafe_judgment
@@ -151,34 +182,30 @@ val jv_nf_evar :
val tj_nf_evar :
evar_map -> unsafe_type_judgment -> unsafe_type_judgment
-val nf_evar_info : evar_map -> evar_info -> evar_info
-val nf_evars : evar_map -> evar_map
-
val nf_named_context_evar : evar_map -> named_context -> named_context
val nf_rel_context_evar : evar_map -> rel_context -> rel_context
val nf_env_evar : evar_map -> env -> env
+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
-(* Replacing all evars, possibly raising [Uninstantiated_evar] *)
-(* exception Uninstantiated_evar of existential_key *)
+(** 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 *)
+(** 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
-(*********************************************************************)
-(* debug pretty-printer: *)
+(** {6 debug pretty-printer:} *)
val pr_tycon_type : env -> type_constraint_type -> Pp.std_ppcmds
val pr_tycon : env -> type_constraint -> Pp.std_ppcmds
-(*********************************************************************)
-(* Removing hyps in evars'context; *)
-(* raise OccurHypInSimpleClause if the removal breaks dependencies *)
+(** {6 Removing hyps in evars'context}
+raise OccurHypInSimpleClause if the removal breaks dependencies *)
type clear_dependency_error =
| OccurHypInSimpleClause of identifier option
@@ -186,8 +213,18 @@ type clear_dependency_error =
exception ClearDependencyError of identifier * 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 clear_hyps_in_evi : evar_map ref -> named_context_val -> types ->
identifier list -> named_context_val * types
-val push_rel_context_to_named_context : Environ.env -> types ->
- named_context_val * types * constr list
+val push_rel_context_to_named_context : Environ.env -> types ->
+ named_context_val * types * constr list * 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
+
diff --git a/pretyping/evd.ml b/pretyping/evd.ml
index 2db77837..5d6ca2ca 100644
--- a/pretyping/evd.ml
+++ b/pretyping/evd.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: evd.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
open Pp
open Util
open Names
@@ -51,7 +49,8 @@ type evar_info = {
evar_body : evar_body;
evar_filter : bool list;
evar_source : hole_kind located;
- evar_extra : Dyn.t option}
+ evar_candidates : constr list option; (* if not None, list of allowed instances *)
+ evar_extra : Store.t }
let make_evar hyps ccl = {
evar_concl = ccl;
@@ -59,7 +58,8 @@ let make_evar hyps ccl = {
evar_body = Evar_empty;
evar_filter = List.map (fun _ -> true) (named_context_of_val hyps);
evar_source = (dummy_loc,InternalHole);
- evar_extra = None
+ evar_candidates = None;
+ evar_extra = Store.empty
}
let evar_concl evi = evi.evar_concl
@@ -81,10 +81,10 @@ let eq_evar_info ei1 ei2 =
ei1.evar_body = ei2.evar_body
(* spiwack: Revised hierarchy :
- - ExistentialMap ( Maps of existential_keys )
- - EvarInfoMap ( .t = evar_info ExistentialMap.t )
- - EvarMap ( .t = EvarInfoMap.t * sort_constraints )
- - evar_map (exported)
+ - ExistentialMap ( Maps of existential_keys )
+ - EvarInfoMap ( .t = evar_info ExistentialMap.t * evar_info ExistentialMap )
+ - EvarMap ( .t = EvarInfoMap.t * sort_constraints )
+ - evar_map (exported)
*)
module ExistentialMap = Intmap
@@ -93,72 +93,95 @@ module ExistentialSet = Intset
(* 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
+ type t = evar_info ExistentialMap.t * evar_info ExistentialMap.t
- let empty = ExistentialMap.empty
+ let empty = ExistentialMap.empty, ExistentialMap.empty
- let to_list evc = (* Workaround for change in Map.fold behavior *)
+ 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) evc;
+ ExistentialMap.iter (fun evk x -> l := (evk,x)::!l) def;
+ ExistentialMap.iter (fun evk x -> l := (evk,x)::!l) undef;
!l
- let dom evc = ExistentialMap.fold (fun evk _ acc -> evk::acc) evc []
- let find evc k = ExistentialMap.find k evc
- let remove evc k = ExistentialMap.remove k evc
- let mem evc k = ExistentialMap.mem k evc
- let fold = ExistentialMap.fold
- let exists evc f = ExistentialMap.fold (fun k v b -> b || f k v) evc false
+ 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 evd evk newinfo = ExistentialMap.add evk newinfo evd
+ let add_undefined (def,undef) evk newinfo =
+ assert (newinfo.evar_body = Evar_empty);
+ (def,ExistentialMap.add evk newinfo undef)
- let equal = ExistentialMap.equal
+ let map f (def,undef) = (ExistentialMap.map f def, ExistentialMap.map f undef)
- let define evd evk body =
+ let define (def,undef) evk body =
let oldinfo =
- try find evd evk
- with Not_found -> error "Evd.define: cannot define undeclared evar" in
+ 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 evd
- | _ -> anomaly "Evd.define: cannot define an evar twice"
-
- let is_evar sigma evk = mem sigma evk
+ | Evar_empty ->
+ (ExistentialMap.add evk newinfo def,ExistentialMap.remove evk undef)
+ | _ ->
+ anomaly "Evd.define: cannot define an evar twice"
- let is_defined sigma evk =
- let info = find sigma evk in
- not (info.evar_body = Evar_empty)
+ 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 *)
- let is_id_inst inst =
- let is_id (id,c) = match kind_of_term c with
- | Var id' -> id = id'
- | _ -> false
- in
- List.for_all is_id inst
-
- (* Vérifier que les instances des let-in sont compatibles ?? *)
- let instantiate_sign_including_let sign args =
- let rec instrec = function
- | ((id,b,_) :: 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 = instantiate_sign_including_let sign args in
- if is_id_inst inst then
- c
- else
- replace_vars inst c
-
(* Existentials. *)
let existential_type sigma (n,args) =
@@ -184,160 +207,38 @@ module EvarInfoMap = struct
end
-(*******************************************************************)
-(* Constraints for sort variables *)
-(*******************************************************************)
-
-type sort_var = Univ.universe
-
-type sort_constraint =
- | DefinedSort of sorts (* instantiated sort var *)
- | SortVar of sort_var list * sort_var list (* (leq,geq) *)
- | EqSort of sort_var
-
-module UniverseMap =
- Map.Make (struct type t = Univ.universe let compare = compare end)
-
-type sort_constraints = sort_constraint UniverseMap.t
-
-let rec canonical_find u scstr =
- match UniverseMap.find u scstr with
- EqSort u' -> canonical_find u' scstr
- | c -> (u,c)
-
-let whd_sort_var scstr t =
- match kind_of_term t with
- Sort(Type u) ->
- (try
- match canonical_find u scstr with
- _, DefinedSort s -> mkSort s
- | _ -> t
- with Not_found -> t)
- | _ -> t
-
-let rec set_impredicative u s scstr =
- match UniverseMap.find u scstr with
- | DefinedSort s' ->
- if family_of_sort s = family_of_sort s' then scstr
- else failwith "sort constraint inconsistency"
- | EqSort u' ->
- UniverseMap.add u (DefinedSort s) (set_impredicative u' s scstr)
- | SortVar(_,ul) ->
- (* also set sorts lower than u as impredicative *)
- UniverseMap.add u (DefinedSort s)
- (List.fold_left (fun g u' -> set_impredicative u' s g) scstr ul)
-
-let rec set_predicative u s scstr =
- match UniverseMap.find u scstr with
- | DefinedSort s' ->
- if family_of_sort s = family_of_sort s' then scstr
- else failwith "sort constraint inconsistency"
- | EqSort u' ->
- UniverseMap.add u (DefinedSort s) (set_predicative u' s scstr)
- | SortVar(ul,_) ->
- UniverseMap.add u (DefinedSort s)
- (List.fold_left (fun g u' -> set_impredicative u' s g) scstr ul)
-
-let var_of_sort = function
- Type u -> u
- | _ -> assert false
-
-let is_sort_var s scstr =
- match s with
- Type u ->
- (try
- match canonical_find u scstr with
- _, DefinedSort _ -> false
- | _ -> true
- with Not_found -> false)
- | _ -> false
-
-let new_sort_var cstr =
- let u = Termops.new_univ() in
- (u, UniverseMap.add u (SortVar([],[])) cstr)
-
-
-let set_leq_sort (u1,(leq1,geq1)) (u2,(leq2,geq2)) scstr =
- let rec search_rec (is_b, betw, not_betw) u1 =
- if List.mem u1 betw then (true, betw, not_betw)
- else if List.mem u1 not_betw then (is_b, betw, not_betw)
- else if u1 = u2 then (true, u1::betw,not_betw) else
- match UniverseMap.find u1 scstr with
- EqSort u1' -> search_rec (is_b,betw,not_betw) u1'
- | SortVar(leq,_) ->
- let (is_b',betw',not_betw') =
- List.fold_left search_rec (false,betw,not_betw) leq in
- if is_b' then (true, u1::betw', not_betw')
- else (false, betw', not_betw')
- | DefinedSort _ -> (false,betw,u1::not_betw) in
- let (is_betw,betw,_) = search_rec (false, [], []) u1 in
- if is_betw then
- UniverseMap.add u1 (SortVar(leq1@leq2,geq1@geq2))
- (List.fold_left
- (fun g u -> UniverseMap.add u (EqSort u1) g) scstr betw)
- else
- UniverseMap.add u1 (SortVar(u2::leq1,geq1))
- (UniverseMap.add u2 (SortVar(leq2, u1::geq2)) scstr)
-
-let set_leq s1 s2 scstr =
- let u1 = var_of_sort s1 in
- let u2 = var_of_sort s2 in
- let (cu1,c1) = canonical_find u1 scstr in
- let (cu2,c2) = canonical_find u2 scstr in
- if cu1=cu2 then scstr
- else
- match c1,c2 with
- (EqSort _, _ | _, EqSort _) -> assert false
- | SortVar(leq1,geq1), SortVar(leq2,geq2) ->
- set_leq_sort (cu1,(leq1,geq1)) (cu2,(leq2,geq2)) scstr
- | _, DefinedSort(Prop _ as s) -> set_impredicative u1 s scstr
- | _, DefinedSort(Type _) -> scstr
- | DefinedSort(Type _ as s), _ -> set_predicative u2 s scstr
- | DefinedSort(Prop _), _ -> scstr
-
-let set_sort_variable s1 s2 scstr =
- let u = var_of_sort s1 in
- match s2 with
- Prop _ -> set_impredicative u s2 scstr
- | Type _ -> set_predicative u s2 scstr
-
-let pr_sort_cstrs g =
- let l = UniverseMap.fold (fun u c l -> (u,c)::l) g [] in
- str "SORT CONSTRAINTS:" ++ fnl() ++
- prlist_with_sep fnl (fun (u,c) ->
- match c with
- EqSort u' -> Univ.pr_uni u ++ str" == " ++ Univ.pr_uni u'
- | DefinedSort s -> Univ.pr_uni u ++ str " := " ++ print_sort s
- | SortVar(leq,geq) ->
- str"[" ++ hov 0 (prlist_with_sep spc Univ.pr_uni geq) ++
- str"] <= "++ Univ.pr_uni u ++ brk(0,0) ++ str"<= [" ++
- hov 0 (prlist_with_sep spc Univ.pr_uni leq) ++ str"]")
- l
-
module EvarMap = struct
- type t = EvarInfoMap.t * sort_constraints
- let empty = EvarInfoMap.empty, UniverseMap.empty
+ 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 dom (sigma,_) = EvarInfoMap.dom sigma
+ 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 fold f (sigma,_) = EvarInfoMap.fold f 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 sigma1
- (fun k v -> v.evar_body = Evar_empty &&
- (EvarInfoMap.find sigma2 k).evar_body <> Evar_empty)
- || not (UniverseMap.equal (=) sm1 sm2))
-
- let merge e e' = fold (fun n v sigma -> add sigma n v) e' e
+ (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
(*******************************************************************)
@@ -372,8 +273,7 @@ let map_fl f cfl = { cfl with rebus=f cfl.rebus }
(e.g. the solution [P] to [?X u v = P u v] can be eta-expanded twice)
*)
-type instance_constraint =
- IsSuperType | IsSubType | ConvUpToEta of int | UserGiven
+type instance_constraint = IsSuperType | IsSubType | Conv
(* Status of the unification of the type of an instance against the type of
the meta it instantiates:
@@ -441,7 +341,6 @@ let progress_evar_map d1 d2 =
(* spiwack: tentative. It might very well not be the semantics we want
for merging evar_map *)
let merge d1 d2 = {
-(* d1 with evars = EvarMap.merge d1.evars d2.evars*)
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 ;
@@ -449,27 +348,34 @@ let merge d1 d2 = {
}
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 dom d = EvarMap.dom d.evars
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
(* 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 }
(* 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 f d.evars a
+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 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 add_constraints d e = {d with evars= EvarMap.add_constraints d.evars e}
+
(*** /Lifting... ***)
(* evar_map are considered empty disregarding histories *)
let is_empty d =
- d.evars = EvarMap.empty &&
+ EvarMap.is_empty d.evars &&
d.conv_pbs = [] &&
Metamap.is_empty d.metas
@@ -484,11 +390,11 @@ let subst_evar_info s evi =
evar_body = subst_evb evi.evar_body }
let subst_evar_defs_light sub evd =
- assert (UniverseMap.is_empty (snd evd.evars));
+ assert (Univ.is_initial_universes (snd (snd evd.evars)));
assert (evd.conv_pbs = []);
{ evd with
metas = Metamap.map (map_clb (subst_mps sub)) evd.metas;
- evars = ExistentialMap.map (subst_evar_info sub) (fst evd.evars), snd evd.evars
+ evars = EvarInfoMap.map (subst_evar_info sub) (fst evd.evars), (snd evd.evars)
}
let subst_evar_map = subst_evar_defs_light
@@ -507,9 +413,12 @@ let empty = {
metas=Metamap.empty
}
-let evars_reset_evd ?(with_conv_pbs=false) evd d =
+let has_undefined evd =
+ EvarMap.has_undefined evd.evars
+
+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}
+ 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
@@ -522,7 +431,7 @@ let define evk body evd =
| [] -> evd.last_mods
| _ -> ExistentialSet.add evk evd.last_mods }
-let evar_declare hyps evk ty ?(src=(dummy_loc,InternalHole)) ?filter evd =
+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)
@@ -532,13 +441,14 @@ let evar_declare hyps evk ty ?(src=(dummy_loc,InternalHole)) ?filter evd =
filter)
in
{ evd with
- evars = EvarMap.add evd.evars evk
+ evars = EvarMap.add_undefined evd.evars evk
{evar_hyps = hyps;
evar_concl = ty;
evar_body = Evar_empty;
evar_filter = filter;
evar_source = src;
- evar_extra = None} }
+ evar_candidates = candidates;
+ evar_extra = Store.empty } }
let is_defined_evar evd (evk,_) = EvarMap.is_defined evd.evars evk
@@ -547,14 +457,6 @@ let is_undefined_evar evd c = match kind_of_term c with
| Evar ev -> not (is_defined_evar evd ev)
| _ -> false
-let undefined_evars evd =
- let evars =
- EvarMap.fold (fun evk evi sigma -> if evi.evar_body = Evar_empty then
- EvarMap.add sigma evk evi else sigma)
- evd.evars EvarMap.empty
- in
- { evd with evars = evars }
-
(* extracts conversion problems that satisfy predicate p *)
(* Note: conv_pbs not satisying p are stored back in reverse order *)
let extract_conv_pbs evd p =
@@ -588,21 +490,84 @@ let evar_list evd c =
| _ -> fold_constr evrec acc c in
evrec [] c
+let collect_evars c =
+ let rec collrec acc c =
+ match kind_of_term c with
+ | Evar (evk,_) -> ExistentialSet.add evk acc
+ | _ -> fold_constr collrec acc c
+ in
+ collrec ExistentialSet.empty c
+
(**********************************************************)
(* Sort variables *)
-let new_sort_variable ({ evars = (sigma,sm) } as d)=
- let (u,scstr) = new_sort_var sm in
- (Type u,{ d with evars = (sigma,scstr) } )
-let is_sort_variable {evars=(_,sm)} s =
- is_sort_var s sm
-let whd_sort_variable {evars=(_,sm)} t = whd_sort_var sm t
-let set_leq_sort_variable ({evars=(sigma,sm)}as d) u1 u2 =
- { d with evars = (sigma, set_leq u1 u2 sm) }
-let define_sort_variable ({evars=(sigma,sm)}as d) u s =
- { d with evars = (sigma, set_sort_variable u s sm) }
-let pr_sort_constraints {evars=(_,sm)} = pr_sort_cstrs sm
-
+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
+ (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 univ_of_sort = function
+ | Type u -> u
+ | Prop Pos -> Univ.type0_univ
+ | Prop Null -> Univ.type0m_univ
+
+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
+ 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 =
+ 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 =
+ match is_eq_sort s1 s2 with
+ | None -> d
+ | 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))
+
(**********************************************************)
(* Accessing metas *)
@@ -700,7 +665,6 @@ let meta_with_name evd id =
strbrk "\" occurs more than once in clause.")
-(* spiwack: we should try and replace this List.fold_left by a Metamap.fold. *)
let meta_merge evd1 evd2 =
{evd2 with
metas = List.fold_left (fun m (n,v) -> Metamap.add n v m)
@@ -712,7 +676,7 @@ let retract_coercible_metas evd =
let mc,ml =
Metamap.fold (fun n v (mc,ml) ->
match v with
- | Clval (na,(b,(UserGiven,CoerceToType as s)),typ) ->
+ | 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)
@@ -725,7 +689,7 @@ let rec list_assoc_in_triple x = function
let subst_defined_metas bl c =
let rec substrec c = match kind_of_term c with
- | Meta i -> substrec (list_assoc_in_triple i bl)
+ | Meta i -> substrec (list_assoc_snd_in_triple i bl)
| _ -> map_constr substrec c
in try Some (substrec c) with Not_found -> None
@@ -754,9 +718,7 @@ let pr_instance_status (sc,typ) =
begin match sc with
| IsSubType -> str " [or a subtype of it]"
| IsSuperType -> str " [or a supertype of it]"
- | ConvUpToEta 0 -> mt ()
- | UserGiven -> mt ()
- | ConvUpToEta n -> str" [or an eta-expansion up to " ++ int n ++ str" of it]"
+ | Conv -> mt ()
end ++
begin match typ with
| CoerceToType -> str " [up to coercion]"
@@ -788,49 +750,126 @@ 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 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"
+
let pr_evar_info evi =
- let decls = List.combine (evar_context evi) (evar_filter evi) in
- let phyps = prlist_with_sep pr_spc pr_decl (List.rev decls) in
+ let phyps =
+ try
+ let decls = List.combine (evar_context evi) (evar_filter evi) in
+ prlist_with_sep pr_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 =
match evi.evar_body with
| Evar_empty -> mt ()
| Evar_defined c -> spc() ++ str"=> " ++ print_constr c
in
- hov 2 (str"[" ++ phyps ++ spc () ++ str"|- " ++ pty ++ pb ++ str"]")
-
-let pr_evar_map_t (evars,cstrs as sigma) =
+ let src = str "(" ++ pr_evar_source (snd evi.evar_source) ++ str ")" in
+ hov 2
+ (str"[" ++ phyps ++ spc () ++ str"|- " ++ pty ++ pb ++ str"]" ++
+ spc() ++ src)
+
+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 evar_dependency_closure n sigma =
+ 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 evars = EvarInfoMap.empty then mt ()
+ if EvarInfoMap.is_empty evars then mt ()
else
- str"EVARS:"++brk(0,1)++
- h 0 (prlist_with_sep pr_fnl
- (fun (ev,evi) ->
- h 0 (str(string_of_existential ev)++str"=="++ pr_evar_info evi))
- (EvarMap.to_list sigma))++fnl()
+ 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 cstrs = UniverseMap.empty then mt ()
- else pr_sort_cstrs cstrs++fnl()
- in evs ++ 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 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_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,_,t1,t2) ->
- print_constr t1 ++ spc() ++
- str (match pbty with
- | Reduction.CONV -> "=="
- | Reduction.CUMUL -> "<=") ++
- spc() ++ print_constr t2) pbs)
-
-let pr_evar_map evd =
+ (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 evd.evars = EvarMap.empty then mt() else
- pr_evar_map_t evd.evars++fnl() in
- let cstrs =
- if evd.conv_pbs = [] then mt() else
+ 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 evd.metas = Metamap.empty then mt() else
+ 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)
diff --git a/pretyping/evd.mli b/pretyping/evd.mli
index 8a903f1b..55c54f2c 100644
--- a/pretyping/evd.mli
+++ b/pretyping/evd.mli
@@ -1,14 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: evd.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
open Util
open Names
open Term
@@ -17,10 +14,9 @@ open Environ
open Libnames
open Mod_subst
open Termops
-(*i*)
-(*********************************************************************)
-(* Meta map *)
+(********************************************************************
+ Meta map *)
module Metamap : Map.S with type key = metavariable
@@ -36,17 +32,16 @@ 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:
+(** 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 | ConvUpToEta of int | UserGiven
+type instance_constraint = IsSuperType | IsSubType | Conv
-(* Status of the unification of the type of an instance against the type of
+(** 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
@@ -62,11 +57,11 @@ type instance_constraint =
type instance_typing_status =
CoerceToType | TypeNotProcessed | TypeProcessed
-(* Status of an instance together with the status of its type unification *)
+(** Status of an instance together with the status of its type unification *)
type instance_status = instance_constraint * instance_typing_status
-(* Clausal environments *)
+(** Clausal environments *)
type clbinding =
| Cltyp of name * constr freelisted
@@ -75,17 +70,17 @@ type clbinding =
val map_clb : (constr -> constr) -> clbinding -> clbinding
-(*********************************************************************)
-(*** Kinds of existential variables ***)
+(********************************************************************
+ ** Kinds of existential variables ***)
-(* Should the obligation be defined (opaque or transparent (default)) or
+(** Should the obligation be defined (opaque or transparent (default)) or
defined transparent and expanded in the term? *)
type obligation_definition_status = Define of bool | Expand
-(* Evars *)
+(** Evars *)
type hole_kind =
- | ImplicitArg of global_reference * (int * identifier option) * bool (* Force inference *)
+ | ImplicitArg of global_reference * (int * identifier option) * bool (** Force inference *)
| BinderType of name
| QuestionMark of obligation_definition_status
| CasesType
@@ -95,10 +90,10 @@ type hole_kind =
| ImpossibleCase
| MatchingVar of bool * identifier
-(*********************************************************************)
-(*** Existential variables and unification states ***)
+(********************************************************************
+ ** Existential variables and unification states ***)
-(* A unification state (of type [evar_map]) is primarily a finite mapping
+(** 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
@@ -106,7 +101,7 @@ type hole_kind =
It also contains conversion constraints, debugging information and
information about meta variables. *)
-(* Information about existential variables. *)
+(** Information about existential variables. *)
type evar = existential_key
val string_of_existential : evar -> string
@@ -122,7 +117,8 @@ type evar_info = {
evar_body : evar_body;
evar_filter : bool list;
evar_source : hole_kind located;
- evar_extra : Dyn.t option}
+ evar_candidates : constr list option;
+ evar_extra : Store.t }
val eq_evar_info : evar_info -> evar_info -> bool
@@ -139,33 +135,40 @@ val evar_env : evar_info -> env
(*** Unification state ***)
type evar_map
-(* Unification state and existential variables *)
+(** Unification state and existential variables *)
-(* Assuming that the second map extends the first one, this says if
+(** 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
val is_empty : 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
-val dom : evar_map -> evar list
val find : evar_map -> evar -> evar_info
+val find_undefined : evar_map -> evar -> evar_info
val remove : evar_map -> evar -> evar_map
val mem : evar_map -> evar -> bool
+val undefined_list : evar_map -> (evar * evar_info) list
val to_list : evar_map -> (evar * evar_info) list
val fold : (evar -> evar_info -> 'a -> 'a) -> evar_map -> 'a -> 'a
-
+val fold_undefined : (evar -> evar_info -> 'a -> 'a) -> evar_map -> 'a -> 'a
val merge : evar_map -> evar_map -> evar_map
-
val define : evar -> constr -> evar_map -> evar_map
val is_evar : evar_map -> evar -> bool
val is_defined : evar_map -> evar -> bool
+val is_undefined : evar_map -> evar -> bool
-(*s [existential_value sigma ev] raises [NotInstantiatedEvar] if [ev] has
+val add_constraints : evar_map -> Univ.constraints -> evar_map
+
+(** {6 ... } *)
+(** [existential_value sigma ev] raises [NotInstantiatedEvar] if [ev] has
no body and [Not_found] if it does not exist in [sigma] *)
exception NotInstantiatedEvar
@@ -173,10 +176,12 @@ val existential_value : evar_map -> existential -> constr
val existential_type : evar_map -> existential -> types
val existential_opt_value : evar_map -> existential -> constr option
-(* Assume empty universe constraints in [evar_map] and [conv_pbs] *)
+val instantiate_evar : named_context -> constr -> constr list -> constr
+
+(** Assume empty universe constraints in [evar_map] and [conv_pbs] *)
val subst_evar_defs_light : substitution -> evar_map -> evar_map
-(* spiwack: this function seems to somewhat break the abstraction. *)
+(** spiwack: this function seems to somewhat break the abstraction. *)
val evars_reset_evd : ?with_conv_pbs:bool -> evar_map -> evar_map -> evar_map
@@ -184,34 +189,41 @@ val evars_reset_evd : ?with_conv_pbs:bool -> evar_map -> evar_map -> evar_map
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 -> evar_map -> evar_map
-val evar_source : existential_key -> evar_map -> loc * hole_kind
+ ?filter:bool list -> ?candidates:constr list -> evar_map -> evar_map
+val evar_source : existential_key -> evar_map -> hole_kind located
-(* spiwack: this function seems to somewhat break the abstraction. *)
-(* [evar_merge evd ev1] extends the evars of [evd] with [evd1] *)
+(* 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
-val evar_list : evar_map -> constr -> existential list
-
-(* Unification constraints *)
+(** 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_map * evar_constraint list
val extract_all_conv_pbs : evar_map -> evar_map * evar_constraint list
+val evar_list : evar_map -> constr -> existential list
+val collect_evars : constr -> ExistentialSet.t
-(* Metas *)
+(** Metas *)
val find_meta : evar_map -> metavariable -> clbinding
val meta_list : evar_map -> (metavariable * clbinding) list
val meta_defined : evar_map -> metavariable -> bool
-(* [meta_fvalue] raises [Not_found] if meta not in map or [Anomaly] if
+
+(** [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
@@ -225,7 +237,7 @@ val meta_declare :
val meta_assign : metavariable -> constr * instance_status -> evar_map -> evar_map
val meta_reassign : metavariable -> constr * instance_status -> evar_map -> evar_map
-(* [meta_merge evd1 evd2] returns [evd2] extended with the metas of [evd1] *)
+(** [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
@@ -237,21 +249,22 @@ 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
-(**********************************************************)
-(* Sort variables *)
+(*********************************************************
+ Sort variables *)
-val new_sort_variable : evar_map -> sorts * evar_map
+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_variable : evar_map -> sorts -> sorts -> evar_map
-val define_sort_variable : evar_map -> sorts -> sorts -> evar_map
+val set_leq_sort : evar_map -> sorts -> sorts -> evar_map
+val set_eq_sort : evar_map -> sorts -> sorts -> evar_map
-(*********************************************************************)
-(* constr with holes *)
+(********************************************************************
+ constr with holes *)
type open_constr = evar_map * constr
-(*********************************************************************)
-(* The type constructor ['a sigma] adds an evar map to an object of
+(********************************************************************
+ The type constructor ['a sigma] adds an evar map to an object of
type ['a] *)
type 'a sigma = {
it : 'a ;
@@ -260,22 +273,22 @@ type 'a sigma = {
val sig_it : 'a sigma -> 'a
val sig_sig : 'a sigma -> evar_map
-(**********************************************************)
-(* Failure explanation *)
+(*********************************************************
+ Failure explanation *)
type unsolvability_explanation = SeveralInstancesFound of int
-(*********************************************************************)
-(* debug pretty-printer: *)
+(********************************************************************
+ debug pretty-printer: *)
val pr_evar_info : evar_info -> Pp.std_ppcmds
-val pr_evar_map : evar_map -> Pp.std_ppcmds
-val pr_sort_constraints : evar_map -> 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_metaset : Metaset.t -> Pp.std_ppcmds
-(*** /!\Deprecated /!\ ***)
-(* create an [evar_map] with empty meta map: *)
+(*** /!\Deprecated /!\ **
+ create an [evar_map] with empty meta map: *)
val create_evar_defs : evar_map -> evar_map
val create_goal_evar_defs : evar_map -> evar_map
val is_defined_evar : evar_map -> existential -> bool
diff --git a/pretyping/rawterm.ml b/pretyping/glob_term.ml
index 978dbeef..a4113671 100644
--- a/pretyping/rawterm.ml
+++ b/pretyping/glob_term.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: rawterm.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
(*i*)
open Util
open Names
@@ -32,7 +30,7 @@ let cases_pattern_loc = function
type patvar = identifier
-type rawsort = RProp of Term.contents | RType of Univ.universe option
+type glob_sort = GProp of Term.contents | GType of Univ.universe option
type binding_kind = Lib.binding_kind = Explicit | Implicit
@@ -51,42 +49,41 @@ type 'a cast_type =
| CastConv of cast_kind * 'a
| CastCoerce (* Cast to a base type (eg, an underlying inductive type) *)
-type rawconstr =
- | RRef of (loc * global_reference)
- | RVar of (loc * identifier)
- | REvar of loc * existential_key * rawconstr list option
- | RPatVar of loc * (bool * patvar) (* Used for patterns only *)
- | RApp of loc * rawconstr * rawconstr list
- | RLambda of loc * name * binding_kind * rawconstr * rawconstr
- | RProd of loc * name * binding_kind * rawconstr * rawconstr
- | RLetIn of loc * name * rawconstr * rawconstr
- | RCases of loc * case_style * rawconstr option * tomatch_tuples * cases_clauses
- | RLetTuple of loc * name list * (name * rawconstr option) *
- rawconstr * rawconstr
- | RIf of loc * rawconstr * (name * rawconstr option) * rawconstr * rawconstr
- | RRec of loc * fix_kind * identifier array * rawdecl list array *
- rawconstr array * rawconstr array
- | RSort of loc * rawsort
- | RHole of (loc * hole_kind)
- | RCast of loc * rawconstr * rawconstr cast_type
- | RDynamic of loc * Dyn.t
-
-and rawdecl = name * binding_kind * rawconstr option * rawconstr
-
-and fix_recursion_order = RStructRec | RWfRec of rawconstr | RMeasureRec of rawconstr * rawconstr option
+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 =
- | RFix of ((int option * fix_recursion_order) array * int)
- | RCoFix of int
+ | 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 = (rawconstr * predicate_pattern)
+and tomatch_tuple = (glob_constr * predicate_pattern)
and tomatch_tuples = tomatch_tuple list
-and cases_clause = (loc * identifier list * cases_pattern list * rawconstr)
+and cases_clause = (loc * identifier list * cases_pattern list * glob_constr)
and cases_clauses = cases_clause list
@@ -95,55 +92,60 @@ let cases_predicate_names tml =
| (tm,(na,None)) -> [na]
| (tm,(na,Some (_,_,_,nal))) -> na::nal) tml)
-let map_rawdecl_left_to_right f (na,k,obd,ty) =
+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_rawconstr_left_to_right f = function
- | RApp (loc,g,args) ->
+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
- RApp (loc,comp1,comp2)
- | RLambda (loc,na,bk,ty,c) ->
+ GApp (loc,comp1,comp2)
+ | GLambda (loc,na,bk,ty,c) ->
let comp1 = f ty in
let comp2 = f c in
- RLambda (loc,na,bk,comp1,comp2)
- | RProd (loc,na,bk,ty,c) ->
+ GLambda (loc,na,bk,comp1,comp2)
+ | GProd (loc,na,bk,ty,c) ->
let comp1 = f ty in
let comp2 = f c in
- RProd (loc,na,bk,comp1,comp2)
- | RLetIn (loc,na,b,c) ->
+ GProd (loc,na,bk,comp1,comp2)
+ | GLetIn (loc,na,b,c) ->
let comp1 = f b in
let comp2 = f c in
- RLetIn (loc,na,comp1,comp2)
- | RCases (loc,sty,rtntypopt,tml,pl) ->
+ 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
- RCases (loc,sty,comp1,comp2,comp3)
- | RLetTuple (loc,nal,(na,po),b,c) ->
+ 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
- RLetTuple (loc,nal,(na,comp1),comp2,comp3)
- | RIf (loc,c,(na,po),b1,b2) ->
+ 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
- RIf (loc,f c,(na,comp1),comp2,comp3)
- | RRec (loc,fk,idl,bl,tyl,bv) ->
- let comp1 = Array.map (Util.list_map_left (map_rawdecl_left_to_right f)) bl 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
- RRec (loc,fk,idl,comp1,comp2,comp3)
- | RCast (loc,c,k) ->
+ 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
- RCast (loc,comp1,comp2)
- | (RVar _ | RSort _ | RHole _ | RRef _ | REvar _ | RPatVar _ | RDynamic _) as x -> x
+ GCast (loc,comp1,comp2)
+ | (GVar _ | GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) as x -> x
-let map_rawconstr = map_rawconstr_left_to_right
+let map_glob_constr = map_glob_constr_left_to_right
(*
let name_app f e = function
@@ -156,54 +158,53 @@ let fold_ident g idl e =
(fun id (idl,e) -> let id,e = g id e in (id::idl,e)) idl ([],e)
in (Array.of_list idl,e)
-let map_rawconstr_with_binders_loc loc g f e = function
- | RVar (_,id) -> RVar (loc,id)
- | RApp (_,a,args) -> RApp (loc,f e a, List.map (f e) args)
- | RLambda (_,na,ty,c) ->
- let na,e = name_app g e na in RLambda (loc,na,f e ty,f e c)
- | RProd (_,na,ty,c) ->
- let na,e = name_app g e na in RProd (loc,na,f e ty,f e c)
- | RLetIn (_,na,b,c) ->
- let na,e = name_app g e na in RLetIn (loc,na,f e b,f e c)
- | RCases (_,tyopt,tml,pl) ->
+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
- RCases
+ GCases
(loc,Option.map (f e) tyopt,List.map (f e) tml, List.map h pl)
- | RRec (_,fk,idl,tyl,bv) ->
+ | GRec (_,fk,idl,tyl,bv) ->
let idl',e' = fold_ident g idl e in
- RRec (loc,fk,idl',Array.map (f e) tyl,Array.map (f e') bv)
- | RCast (_,c,t) -> RCast (loc,f e c,f e t)
- | RSort (_,x) -> RSort (loc,x)
- | RHole (_,x) -> RHole (loc,x)
- | RRef (_,x) -> RRef (loc,x)
- | REvar (_,x,l) -> REvar (loc,x,l)
- | RPatVar (_,x) -> RPatVar (loc,x)
- | RDynamic (_,x) -> RDynamic (loc,x)
+ 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_rawconstr f acc =
+let fold_glob_constr f acc =
let rec fold acc = function
- | RVar _ -> acc
- | RApp (_,c,args) -> List.fold_left fold (fold acc c) args
- | RLambda (_,_,_,b,c) | RProd (_,_,_,b,c) | RLetIn (_,_,b,c) ->
+ | 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
- | RCases (_,_,rtntypopt,tml,pl) ->
+ | GCases (_,_,rtntypopt,tml,pl) ->
List.fold_left fold_pattern
(List.fold_left fold (Option.fold_left fold acc rtntypopt) (List.map fst tml))
pl
- | RLetTuple (_,_,rtntyp,b,c) ->
+ | GLetTuple (_,_,rtntyp,b,c) ->
fold (fold (fold_return_type acc rtntyp) b) c
- | RIf (_,c,rtntyp,b1,b2) ->
+ | GIf (_,c,rtntyp,b1,b2) ->
fold (fold (fold (fold_return_type acc rtntyp) c) b1) b2
- | RRec (_,_,_,bl,tyl,bv) ->
+ | 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
- | RCast (_,c,k) -> fold (match k with CastConv (_, t) -> fold acc t | CastCoerce -> acc) c
- | (RSort _ | RHole _ | RRef _ | REvar _ | RPatVar _ | RDynamic _) -> acc
+ | 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
@@ -211,25 +212,25 @@ let fold_rawconstr f acc =
in fold acc
-let iter_rawconstr f = fold_rawconstr (fun () -> f) ()
+let iter_glob_constr f = fold_glob_constr (fun () -> f) ()
-let occur_rawconstr id =
+let occur_glob_constr id =
let rec occur = function
- | RVar (loc,id') -> id = id'
- | RApp (loc,f,args) -> (occur f) or (List.exists occur args)
- | RLambda (loc,na,bk,ty,c) -> (occur ty) or ((na <> Name id) & (occur c))
- | RProd (loc,na,bk,ty,c) -> (occur ty) or ((na <> Name id) & (occur c))
- | RLetIn (loc,na,b,c) -> (occur b) or ((na <> Name id) & (occur c))
- | RCases (loc,sty,rtntypopt,tml,pl) ->
+ | 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)
- | RLetTuple (loc,nal,rtntyp,b,c) ->
+ | GLetTuple (loc,nal,rtntyp,b,c) ->
occur_return_type rtntyp id
or (occur b) or (not (List.mem (Name id) nal) & (occur c))
- | RIf (loc,c,rtntyp,b1,b2) ->
+ | GIf (loc,c,rtntyp,b1,b2) ->
occur_return_type rtntyp id or (occur c) or (occur b1) or (occur b2)
- | RRec (loc,fk,idl,bl,tyl,bv) ->
+ | 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))
@@ -241,8 +242,8 @@ let occur_rawconstr id =
(na=Name id or not(occur_fix bl)) in
occur_fix bl)
idl bl tyl bv)
- | RCast (loc,c,k) -> (occur c) or (match k with CastConv (_, t) -> occur t | CastCoerce -> false)
- | (RSort _ | RHole _ | RRef _ | REvar _ | RPatVar _ | RDynamic _) -> false
+ | 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)
@@ -258,29 +259,29 @@ let add_name_to_ids set na =
| Anonymous -> set
| Name id -> Idset.add id set
-let free_rawvars =
+let free_glob_vars =
let rec vars bounded vs = function
- | RVar (loc,id') -> if Idset.mem id' bounded then vs else Idset.add id' vs
- | RApp (loc,f,args) -> List.fold_left (vars bounded) vs (f::args)
- | RLambda (loc,na,_,ty,c) | RProd (loc,na,_,ty,c) | RLetIn (loc,na,ty,c) ->
+ | 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
- | RCases (loc,sty,rtntypopt,tml,pl) ->
+ | 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
- | RLetTuple (loc,nal,rtntyp,b,c) ->
+ | 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
- | RIf (loc,c,rtntyp,b1,b2) ->
+ | 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
- | RRec (loc,fk,idl,bl,tyl,bv) ->
+ | 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 =
@@ -298,9 +299,9 @@ let free_rawvars =
vars bounded1 vs2 bv.(i)
in
array_fold_left_i vars_fix vs idl
- | RCast (loc,c,k) -> let v = vars bounded vs c in
+ | GCast (loc,c,k) -> let v = vars bounded vs c in
(match k with CastConv (_,t) -> vars bounded v t | _ -> v)
- | (RSort _ | RHole _ | RRef _ | REvar _ | RPatVar _ | RDynamic _) -> vs
+ | (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
@@ -317,58 +318,57 @@ let free_rawvars =
Idset.elements vs
-let loc_of_rawconstr = function
- | RRef (loc,_) -> loc
- | RVar (loc,_) -> loc
- | REvar (loc,_,_) -> loc
- | RPatVar (loc,_) -> loc
- | RApp (loc,_,_) -> loc
- | RLambda (loc,_,_,_,_) -> loc
- | RProd (loc,_,_,_,_) -> loc
- | RLetIn (loc,_,_,_) -> loc
- | RCases (loc,_,_,_,_) -> loc
- | RLetTuple (loc,_,_,_,_) -> loc
- | RIf (loc,_,_,_,_) -> loc
- | RRec (loc,_,_,_,_,_) -> loc
- | RSort (loc,_) -> loc
- | RHole (loc,_) -> loc
- | RCast (loc,_,_) -> loc
- | RDynamic (loc,_) -> loc
+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 rawconstr to cases pattern, if possible *)
+(* Conversion from glob_constr to cases pattern, if possible *)
-let rec cases_pattern_of_rawconstr na = function
- | RVar (loc,id) when na<>Anonymous ->
+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
- | RVar (loc,id) -> PatVar (loc,Name id)
- | RHole (loc,_) -> PatVar (loc,na)
- | RRef (loc,ConstructRef cstr) ->
+ | GVar (loc,id) -> PatVar (loc,Name id)
+ | GHole (loc,_) -> PatVar (loc,na)
+ | GRef (loc,ConstructRef cstr) ->
PatCstr (loc,cstr,[],na)
- | RApp (loc,RRef (_,ConstructRef cstr),l) ->
- PatCstr (loc,cstr,List.map (cases_pattern_of_rawconstr Anonymous) l,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 rawconstr *)
-let rec rawconstr_of_closed_cases_pattern_aux = function
+(* Turn a closed cases pattern into a glob_constr *)
+let rec glob_constr_of_closed_cases_pattern_aux = function
| PatCstr (loc,cstr,[],Anonymous) ->
- RRef (loc,ConstructRef cstr)
+ GRef (loc,ConstructRef cstr)
| PatCstr (loc,cstr,l,Anonymous) ->
- let ref = RRef (loc,ConstructRef cstr) in
- RApp (loc,ref, List.map rawconstr_of_closed_cases_pattern_aux l)
+ let ref = GRef (loc,ConstructRef cstr) in
+ GApp (loc,ref, List.map glob_constr_of_closed_cases_pattern_aux l)
| _ -> raise Not_found
-let rawconstr_of_closed_cases_pattern = function
+let glob_constr_of_closed_cases_pattern = function
| PatCstr (loc,cstr,l,na) ->
- na,rawconstr_of_closed_cases_pattern_aux (PatCstr (loc,cstr,l,Anonymous))
+ na,glob_constr_of_closed_cases_pattern_aux (PatCstr (loc,cstr,l,Anonymous))
| _ ->
raise Not_found
(**********************************************************************)
(* Reduction expressions *)
-type 'a raw_red_flag = {
+type 'a glob_red_flag = {
rBeta : bool;
rIota : bool;
rZeta : bool;
@@ -394,8 +394,8 @@ type ('a,'b,'c) red_expr_gen =
| Red of bool
| Hnf
| Simpl of 'c with_occurrences option
- | Cbv of 'b raw_red_flag
- | Lazy of 'b raw_red_flag
+ | 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
diff --git a/pretyping/glob_term.mli b/pretyping/glob_term.mli
new file mode 100644
index 00000000..7fb995a8
--- /dev/null
+++ b/pretyping/glob_term.mli
@@ -0,0 +1,167 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
+(* \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 edbab0a7..a0976029 100644
--- a/pretyping/indrec.ml
+++ b/pretyping/indrec.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: indrec.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
(* File initially created by Christine Paulin, 1996 *)
(* This file builds various inductive schemes *)
@@ -18,7 +16,6 @@ open Names
open Libnames
open Nameops
open Term
-open Termops
open Namegen
open Declarations
open Entries
@@ -53,11 +50,15 @@ let mkLambda_string s t c = mkLambda (Name (id_of_string s), t, c)
(* Christine Paulin, 1996 *)
let mis_make_case_com dep env sigma ind (mib,mip as specif) kind =
- let lnamespar = mib.mind_params_ctxt in
+ let lnamespar = List.map
+ (fun (n, c, t) -> (n, c, Termops.refresh_universes t))
+ mib.mind_params_ctxt
+ in
+
if not (List.mem kind (elim_sorts specif)) then
raise
(RecursionSchemeError
- (NotAllowedCaseAnalysis (false,new_sort_in_family kind,ind)));
+ (NotAllowedCaseAnalysis (false, Termops.new_sort_in_family kind, ind)));
let ndepar = mip.mind_nrealargs_ctxt + 1 in
@@ -65,7 +66,7 @@ let mis_make_case_com dep env sigma ind (mib,mip as specif) kind =
(* 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, extended_rel_list 0 lnamespar) 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 =
@@ -81,8 +82,8 @@ let mis_make_case_com dep env sigma ind (mib,mip as specif) kind =
let pbody =
appvect
(mkRel (ndepar + nbprod),
- if dep then extended_rel_vect 0 deparsign
- else extended_rel_vect 1 arsign) in
+ if dep then Termops.extended_rel_vect 0 deparsign
+ else Termops.extended_rel_vect 1 arsign) in
let p =
it_mkLambda_or_LetIn_name env'
((if dep then mkLambda_name env' else mkLambda)
@@ -92,7 +93,7 @@ let mis_make_case_com dep env sigma ind (mib,mip as specif) kind =
it_mkLambda_or_LetIn_name env'
(mkCase (ci, lift ndepar p,
mkRel 1,
- rel_vect ndepar k))
+ Termops.rel_vect ndepar k))
deparsign
else
let cs = lift_constructor (k+1) constrs.(k) in
@@ -100,7 +101,7 @@ let mis_make_case_com dep env sigma ind (mib,mip as specif) kind =
mkLambda_string "f" t
(add_branch (push_rel (Anonymous, None, t) env) (k+1))
in
- let typP = make_arity env' dep indf (new_sort_in_family kind) in
+ let typP = make_arity env' dep indf (Termops.new_sort_in_family kind) in
it_mkLambda_or_LetIn_name env
(mkLambda_string "P" typP
(add_branch (push_rel (Anonymous,None,typP) env') 0)) lnamespar
@@ -140,7 +141,7 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
let base = applist (lift i pk,realargs) in
if depK then
Reduction.beta_appvect
- base [|applist (mkRel (i+1),extended_rel_list 0 sign)|]
+ base [|applist (mkRel (i+1), Termops.extended_rel_list 0 sign)|]
else
base
| _ -> assert false
@@ -155,9 +156,9 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
| [] -> None,[]
| ra::rest ->
(match dest_recarg ra with
- | Mrec j when is_rec -> (depPvect.(j),rest)
+ | Mrec (_,j) when is_rec -> (depPvect.(j),rest)
| Imbr _ ->
- Flags.if_verbose warning "Ignoring recursive call";
+ Flags.if_warn msg_warning (str "Ignoring recursive call");
(None,rest)
| _ -> (None, rest))
in
@@ -212,7 +213,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
mkLetIn (n,b,t,prec (push_rel d env) (i+1) (d::hyps) c)
| Ind _ ->
let realargs = list_skipn nparrec largs
- and arg = appvect (mkRel (i+1),extended_rel_vect 0 hyps) in
+ and arg = appvect (mkRel (i+1), Termops.extended_rel_vect 0 hyps) in
applist(lift i fk,realargs@[arg])
| _ -> assert false
in
@@ -225,7 +226,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
match dest_recarg recarg with
| Norec -> None
| Imbr _ -> None
- | Mrec i -> fvect.(i)
+ | Mrec (_,i) -> fvect.(i)
in
(match optionpos with
| None ->
@@ -298,7 +299,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 = extended_rel_list (nrec+nbconstruct) lnamesparrec in
+ let args = Termops.extended_rel_list (nrec+nbconstruct) lnamesparrec in
let indf = make_ind_family(indi,args) in
let arsign,_ = get_arity env indf in
@@ -312,15 +313,15 @@ let mis_make_indrec env sigma listdepkind mib =
(* constructors in context of the Cases expr, i.e.
P1..P_nrec f1..f_nbconstruct F_1..F_nrec a_1..a_nar x:I *)
- let args' = extended_rel_list (dect+nrec) lnamesparrec in
- let args'' = extended_rel_list ndepar lnonparrec in
+ 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 branches =
let constrs = get_constructors env indf' in
- let fi = rel_vect (dect-i-nctyi) nctyi in
+ let fi = Termops.rel_vect (dect-i-nctyi) nctyi in
let vecfi = Array.map
- (fun f -> appvect (f,extended_rel_vect ndepar lnonparrec))
+ (fun f -> appvect (f, Termops.extended_rel_vect ndepar lnonparrec))
fi
in
array_map3
@@ -341,9 +342,9 @@ let mis_make_indrec env sigma listdepkind mib =
let deparsign' = (Anonymous,None,depind')::arsign' in
let pargs =
- let nrpar = extended_rel_list (2*ndepar) lnonparrec
- and nrar = if dep then extended_rel_list 0 deparsign'
- else extended_rel_list 1 arsign'
+ let nrpar = Termops.extended_rel_list (2*ndepar) lnonparrec
+ and nrar = if dep then Termops.extended_rel_list 0 deparsign'
+ else Termops.extended_rel_list 1 arsign'
in nrpar@nrar
in
@@ -362,15 +363,15 @@ let mis_make_indrec env sigma listdepkind mib =
(mkCase (ci, pred,
mkRel 1,
branches))
- (lift_rel_context nrec deparsign)
+ (Termops.lift_rel_context nrec deparsign)
in
(* type of i-th component of the mutual fixpoint *)
let typtyi =
let concl =
- let pargs = if dep then extended_rel_vect 0 deparsign
- else extended_rel_vect 1 arsign
+ let pargs = if dep then Termops.extended_rel_vect 0 deparsign
+ else Termops.extended_rel_vect 1 arsign
in appvect (mkRel (nbconstruct+ndepar+nonrecpar+j),pargs)
in it_mkProd_or_LetIn_name env
concl
@@ -397,7 +398,7 @@ let mis_make_indrec env sigma listdepkind mib =
else
let recarg = (dest_subterms recargsvec.(tyi)).(j) in
let recarg = recargpar@recarg in
- let vargs = extended_rel_list (nrec+i+j) lnamesparrec in
+ let vargs = Termops.extended_rel_list (nrec+i+j) lnamesparrec in
let cs = get_constructor (indi,mibi,mipi,vargs) (j+1) in
let p_0 =
type_rec_branch
@@ -411,8 +412,8 @@ let mis_make_indrec env sigma listdepkind mib =
in
let rec put_arity env i = function
| (indi,_,_,dep,kinds)::rest ->
- let indf = make_ind_family (indi,extended_rel_list i lnamesparrec) in
- let typP = make_arity env dep indf (new_sort_in_family kinds) in
+ 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
mkLambda_string "P" typP
(put_arity (push_rel (Anonymous,None,typP) env) (i+1) rest)
| [] ->
@@ -506,7 +507,7 @@ let check_arities listdepkind =
let kelim = elim_sorts (mibi,mipi) in
if not (List.exists ((=) kind) kelim) then raise
(RecursionSchemeError
- (NotAllowedCaseAnalysis (true,new_sort_in_family kind,mind)))
+ (NotAllowedCaseAnalysis (true, Termops.new_sort_in_family kind,mind)))
else if List.mem ni ln then raise
(RecursionSchemeError (NotMutualInScheme (mind,mind)))
else ni::ln)
@@ -558,8 +559,7 @@ let lookup_eliminator ind_sp s =
(* Try first to get an eliminator defined in the same section as the *)
(* inductive type *)
try
- let cst =Global.constant_of_delta
- (make_con 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
with Not_found ->
@@ -571,5 +571,5 @@ let lookup_eliminator ind_sp s =
(strbrk "Cannot find the elimination combinator " ++
pr_id id ++ strbrk ", the elimination of the inductive definition " ++
pr_global_env Idset.empty (IndRef ind_sp) ++
- strbrk " on sort " ++ pr_sort_family s ++
+ strbrk " on sort " ++ Termops.pr_sort_family s ++
strbrk " is probably not allowed.")
diff --git a/pretyping/indrec.mli b/pretyping/indrec.mli
index b2375c8f..a421ae7d 100644
--- a/pretyping/indrec.mli
+++ b/pretyping/indrec.mli
@@ -1,23 +1,19 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: indrec.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
open Names
open Term
open Declarations
open Inductiveops
open Environ
open Evd
-(*i*)
-(* Errors related to recursors building *)
+(** Errors related to recursors building *)
type recursion_scheme_error =
| NotAllowedCaseAnalysis of (*isrec:*) bool * sorts * inductive
@@ -29,35 +25,35 @@ exception RecursionSchemeError of recursion_scheme_error
type dep_flag = bool
-(* Build a case analysis elimination scheme in some sort family *)
+(** Build a case analysis elimination scheme in some sort family *)
val build_case_analysis_scheme : env -> evar_map -> inductive ->
dep_flag -> sorts_family -> constr
-(* Build a dependent case elimination predicate unless type is in Prop *)
+(** Build a dependent case elimination predicate unless type is in Prop *)
val build_case_analysis_scheme_default : env -> evar_map -> inductive ->
sorts_family -> constr
-(* Builds a recursive induction scheme (Peano-induction style) in the same
+(** 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
-(* Builds mutual (recursive) induction schemes *)
+(** Builds mutual (recursive) induction schemes *)
val build_mutual_induction_scheme :
env -> evar_map -> (inductive * dep_flag * sorts_family) list -> constr list
(** Scheme combinators *)
-(* [modify_sort_scheme s n c] modifies the quantification sort of
+(** [modify_sort_scheme s n c] modifies the quantification sort of
scheme c whose predicate is abstracted at position [n] of [c] *)
val modify_sort_scheme : sorts -> int -> constr -> constr
-(* [weaken_sort_scheme s n c t] derives by subtyping from [c:t]
+(** [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] *)
diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml
index 6e54c170..15fd226f 100644
--- a/pretyping/inductiveops.ml
+++ b/pretyping/inductiveops.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: inductiveops.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
open Util
open Names
open Univ
@@ -78,7 +76,7 @@ let mis_is_recursive_subset listind rarg =
List.exists
(fun ra ->
match dest_recarg ra with
- | Mrec i -> List.mem i listind
+ | Mrec (_,i) -> List.mem i listind
| _ -> false) rvec
in
array_exists one_is_rec (dest_subterms rarg)
@@ -145,12 +143,9 @@ let make_case_info env ind style =
let print_info = { ind_nargs = mip.mind_nrealargs_ctxt; style = style } in
{ ci_ind = ind;
ci_npar = mib.mind_nparams;
- ci_cstr_nargs = mip.mind_consnrealdecls;
+ ci_cstr_ndecls = mip.mind_consnrealdecls;
ci_pp_info = print_info }
-let make_default_case_info env style ind =
- make_case_info env ind style
-
(*s Useful functions *)
type constructor_summary = {
@@ -202,12 +197,6 @@ let get_constructors env (ind,params) =
Array.init (Array.length mip.mind_consnames)
(fun j -> get_constructor (ind,mib,mip,params) (j+1))
-let rec instantiate args c = match kind_of_term c, args with
- | Prod (_,_,c), a::args -> instantiate args (subst1 a c)
- | LetIn (_,b,_,c), args -> instantiate args (subst1 b c)
- | _, [] -> c
- | _ -> anomaly "too short arity"
-
(* substitution in a signature *)
let substnl_rel_context subst n sign =
@@ -413,7 +402,7 @@ let is_constrained is u =
let _ =
merge_constraints
(enforce_geq u (super u')
- (enforce_geq u' is Constraint.empty))
+ (enforce_geq u' is empty_constraint))
initial_universes in
false
with UniverseInconsistency _ -> true
diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli
index e5759864..f361e8b8 100644
--- a/pretyping/inductiveops.mli
+++ b/pretyping/inductiveops.mli
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: inductiveops.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
open Names
open Term
open Declarations
@@ -15,19 +13,19 @@ open Environ
open Evd
open Sign
-(* The following three functions are similar to the ones defined in
+(** The following three functions are similar to the ones defined in
Inductive, but they expect an env *)
val type_of_inductive : env -> inductive -> types
-(* Return type as quoted by the user *)
+(** Return type as quoted by the user *)
val type_of_constructor : env -> constructor -> types
val type_of_constructors : env -> inductive -> types array
-(* Return constructor types in normal form *)
+(** Return constructor types in normal form *)
val arities_of_constructors : env -> inductive -> types array
-(* An inductive type with its parameters *)
+(** 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
@@ -37,7 +35,7 @@ val lift_inductive_family : int -> inductive_family -> inductive_family
val substnl_ind_family :
constr list -> int -> inductive_family -> inductive_family
-(* An inductive type with its parameters and real arguments *)
+(** An inductive type with its parameters and real arguments *)
type inductive_type = IndType of inductive_family * constr list
val make_ind_type : inductive_family * constr list -> inductive_type
val dest_ind_type : inductive_type -> inductive_family * constr list
@@ -53,14 +51,15 @@ val mis_is_recursive :
val mis_nf_constructor_type :
inductive * mutual_inductive_body * one_inductive_body -> int -> constr
-(* Extract information from an inductive name *)
+(** 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
val nconstructors : inductive -> int
-(* Return the lengths of parameters signature and real arguments signature *)
+(** Return the lengths of parameters signature and real arguments signature *)
val inductive_nargs : env -> inductive -> int * int
val mis_constructor_nargs_env : env -> constructor -> int
@@ -71,14 +70,14 @@ val get_full_arity_sign : env -> inductive -> rel_context
val allowed_sorts : env -> inductive -> sorts_family list
-(* Extract information from an inductive family *)
+(** Extract information from an inductive family *)
type constructor_summary = {
- cs_cstr : constructor;
- cs_params : constr list;
- cs_nargs : int;
- cs_args : rel_context;
- cs_concl_realargs : constr array;
+ cs_cstr : constructor; (* internal name of the constructor *)
+ 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) *)
+ cs_concl_realargs : constr array; (* actual realargs in the concl of cstr *)
}
val lift_constructor : int -> constructor_summary -> constructor_summary
val get_constructor :
@@ -92,22 +91,24 @@ 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 *)
+(** Raise [Not_found] if not given an valid inductive type *)
val extract_mrectype : constr -> inductive * constr list
-val find_mrectype : env -> evar_map -> constr -> inductive * constr list
-val find_rectype : env -> evar_map -> constr -> inductive_type
-val find_inductive : env -> evar_map -> constr -> inductive * constr list
-val find_coinductive : env -> evar_map -> constr -> inductive * constr list
+val find_mrectype : env -> evar_map -> types -> inductive * 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
(********************)
-(* Builds the case predicate arity (dependent or not) *)
+(** Builds the case predicate arity (dependent or not) *)
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
+
+(** Annotation for cases *)
val make_case_info : env -> inductive -> case_style -> case_info
(*i Compatibility
diff --git a/pretyping/matching.ml b/pretyping/matching.ml
index f0536f22..7a745118 100644
--- a/pretyping/matching.ml
+++ b/pretyping/matching.ml
@@ -1,14 +1,13 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: matching.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
(*i*)
+open Pp
open Util
open Names
open Libnames
@@ -16,7 +15,7 @@ open Nameops
open Termops
open Reductionops
open Term
-open Rawterm
+open Glob_term
open Sign
open Environ
open Pattern
@@ -56,23 +55,23 @@ let constrain (n,(ids,m as x)) (names,terms as subst) =
with
Not_found ->
if List.mem_assoc n names then
- Flags.if_verbose Pp.warning
- ("Collision between bound variable "^string_of_id n^
- " and a metavariable of same name.");
+ 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_verbose Pp.warning
- ("Collision between bound variables of name "^string_of_id id1);
+ (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_verbose Pp.warning
- ("Collision between bound variable "^string_of_id id1^
- " and another bound variable of same name.");
+ 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
@@ -87,18 +86,6 @@ let build_lambda toabstract stk (m : constr) =
in
buildrec m 1 stk
-let memb_metavars m n =
- match (m,n) with
- | (None, _) -> true
- | (Some mvs, n) -> List.mem n mvs
-
-let eq_context ctxt1 ctxt2 = array_for_all2 eq_constr ctxt1 ctxt2
-
-let same_case_structure (_,cs1,ind,_) ci2 br1 br2 =
- match ind with
- | Some ind -> ind = ci2.ci_ind
- | None -> cs1 = ci2.ci_cstr_nargs
-
let rec list_insert f a = function
| [] -> [a]
| b::l when f a b -> a::b::l
@@ -182,9 +169,9 @@ let matches_core convert allow_partial_app allow_bound_rels pat c =
| PRel n1, Rel n2 when n1 = n2 -> subst
- | PSort (RProp c1), Sort (Prop c2) when c1 = c2 -> subst
+ | PSort (GProp c1), Sort (Prop c2) when c1 = c2 -> subst
- | PSort (RType _), Sort (Type _) -> subst
+ | PSort (GType _), Sort (Type _) -> subst
| PApp (p, [||]), _ -> sorec stk subst p t
@@ -217,8 +204,8 @@ let matches_core convert allow_partial_app allow_bound_rels pat c =
(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_nargs.(0) b2 in
- let ctx',b2' = decompose_lam_n_assum ci.ci_cstr_nargs.(1) b2' in
+ 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
@@ -232,11 +219,18 @@ let matches_core convert allow_partial_app allow_bound_rels pat c =
raise PatternMatchingFailure
| PCase (ci1,p1,a1,br1), Case (ci2,p2,a2,br2) ->
- if same_case_structure ci1 ci2 br1 br2 then
- array_fold_left2 (sorec stk)
- (sorec stk (sorec stk subst a1 a2) p1 p2) br1 br2
- else
- raise PatternMatchingFailure
+ 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
diff --git a/pretyping/matching.mli b/pretyping/matching.mli
index fc0e3feb..49033286 100644
--- a/pretyping/matching.mli
+++ b/pretyping/matching.mli
@@ -1,79 +1,82 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: matching.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
+(** This module implements pattern-matching on terms *)
-(*i*)
open Names
open Term
open Environ
open Pattern
open Termops
-(*i*)
-
-(*s This modules implements pattern-matching on terms *)
+(** [PatternMatchingFailure] is the exception raised when pattern
+ matching fails *)
exception PatternMatchingFailure
+(** [special_meta] is the default name of the meta holding the
+ surrounding context in subterm matching *)
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
-(* [matches pat c] matches [c] against [pat] and returns the resulting
+(** [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
-(* [extended_matches pat c] also returns the names of bound variables
+(** [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
-(* [is_matching pat c] just tells if [c] matches against [pat] *)
+(** [is_matching pat c] just tells if [c] matches against [pat] *)
val is_matching : constr_pattern -> constr -> bool
-(* [matches_conv env sigma] matches up to conversion in environment
+(** [matches_conv env sigma] matches up to conversion in environment
[(env,sigma)] when constants in pattern are concerned; it raises
[PatternMatchingFailure] if not matchable; bindings are given in
increasing order based on the numbers given in the pattern *)
val matches_conv :env -> Evd.evar_map -> constr_pattern -> constr -> patvar_map
-(* The type of subterm matching results: a substitution + a context
+(** 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)
-(* [match_subterm n pat c] returns the substitution and the context
+(** [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
-(* [match_appsubterm pat c] returns the substitution and the context
+(** [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
-(* [match_subterm_gen] calls either [match_subterm] or [match_appsubterm] *)
-val match_subterm_gen : bool (* true = with app context *) ->
+(** [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
-(* [is_matching_appsubterm pat c] tells if a subterm of [c] matches
+(** [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
-(* [is_matching_conv env sigma pat c] tells if [c] matches against [pat]
+(** [is_matching_conv env sigma pat c] tells if [c] matches against [pat]
up to conversion for constants in patterns *)
val is_matching_conv :
env -> Evd.evar_map -> constr_pattern -> constr -> bool
diff --git a/pretyping/namegen.ml b/pretyping/namegen.ml
index 45060511..2ad2f351 100644
--- a/pretyping/namegen.ml
+++ b/pretyping/namegen.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: namegen.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
(* Created from contents that was formerly in termops.ml and
nameops.ml, Nov 2009 *)
@@ -133,9 +131,9 @@ let mkProd_or_LetIn_name env b d = mkProd_or_LetIn (name_assumption env d) b
let mkLambda_or_LetIn_name env b d = mkLambda_or_LetIn (name_assumption env d)b
let it_mkProd_or_LetIn_name env b hyps =
- it_mkProd_or_LetIn ~init:b (name_context env hyps)
+ it_mkProd_or_LetIn b (name_context env hyps)
let it_mkLambda_or_LetIn_name env b hyps =
- it_mkLambda_or_LetIn ~init:b (name_context env hyps)
+ it_mkLambda_or_LetIn b (name_context env hyps)
(**********************************************************************)
(* Fresh names *)
@@ -207,6 +205,17 @@ let next_name_away_with_default default na avoid =
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)
+let set_reserved_typed_name f = reserved_type_name := f
+
+let next_name_away_with_default_using_types default na avoid t =
+ let id = match na with
+ | Name id -> id
+ | Anonymous -> match !reserved_type_name t with
+ | Name id -> id
+ | Anonymous -> id_of_string default in
+ next_ident_away id avoid
+
let next_name_away = next_name_away_with_default "H"
let make_all_name_different env =
diff --git a/pretyping/namegen.mli b/pretyping/namegen.mli
index f99ee3f6..637cbf64 100644
--- a/pretyping/namegen.mli
+++ b/pretyping/namegen.mli
@@ -1,19 +1,17 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: namegen.mli 14641 2011-11-06 11:59:10Z herbelin $ *)
-
open Names
open Term
open Environ
-(**********************************************************************)
-(* Generating "intuitive" names from their type *)
+(*********************************************************************
+ Generating "intuitive" names from their type *)
val lowercase_first_char : identifier -> string
val sort_hdchar : sorts -> string
@@ -24,7 +22,7 @@ val named_hd : env -> types -> name -> name
val mkProd_name : env -> name * types * types -> types
val mkLambda_name : env -> name * types * constr -> constr
-(* Deprecated synonyms of [mkProd_name] and [mkLambda_name] *)
+(** Deprecated synonyms of [mkProd_name] and [mkLambda_name] *)
val prod_name : env -> name * types * types -> types
val lambda_name : env -> name * types * constr -> constr
@@ -38,32 +36,40 @@ val mkLambda_or_LetIn_name : env -> constr -> rel_declaration -> constr
val it_mkProd_or_LetIn_name : env -> types -> rel_context -> types
val it_mkLambda_or_LetIn_name : env -> constr -> rel_context -> constr
-(**********************************************************************)
-(* Fresh names *)
+(*********************************************************************
+ Fresh names *)
+
+(** Avoid clashing with a name satisfying some predicate *)
+val next_ident_away_from : identifier -> (identifier -> bool) -> identifier
-(* Avoid clashing with a name of the given list *)
+(** Avoid clashing with a name of the given list *)
val next_ident_away : identifier -> identifier list -> identifier
-(* Avoid clashing with a name already used in current module *)
+(** Avoid clashing with a name already used in current module *)
val next_ident_away_in_goal : identifier -> identifier list -> identifier
-(* Avoid clashing with a name already used in current module *)
-(* but tolerate overwriting section variables, as in goals *)
+(** 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
-(* Avoid clashing with a constructor name already used in current module *)
+(** 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 : name -> identifier list -> identifier (* default is "H" *)
+val next_name_away : name -> identifier list -> identifier (** default is "H" *)
val next_name_away_with_default : string -> name -> identifier list ->
identifier
-(**********************************************************************)
-(* Making name distinct for displaying *)
+val next_name_away_with_default_using_types : string -> name ->
+ identifier list -> types -> identifier
+
+val set_reserved_typed_name : (types -> name) -> unit
+
+(*********************************************************************
+ Making name distinct for displaying *)
type renaming_flags =
- | RenamingForCasesPattern (* avoid only global constructors *)
- | RenamingForGoal (* avoid all globals (as in intro) *)
+ | RenamingForCasesPattern (** avoid only global constructors *)
+ | RenamingForGoal (** avoid all globals (as in intro) *)
| RenamingElsewhereFor of constr
val make_all_name_different : env -> env
diff --git a/pretyping/pattern.ml b/pretyping/pattern.ml
index 83bfe9ea..65f342d8 100644
--- a/pretyping/pattern.ml
+++ b/pretyping/pattern.ml
@@ -1,19 +1,17 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: pattern.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
open Util
open Names
open Libnames
open Nameops
open Term
-open Rawterm
+open Glob_term
open Environ
open Nametab
open Pp
@@ -28,6 +26,12 @@ type extended_patvar_map = (patvar * constr_under_binders) list
(* 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
@@ -38,11 +42,11 @@ type constr_pattern =
| PLambda of name * constr_pattern * constr_pattern
| PProd of name * constr_pattern * constr_pattern
| PLetIn of name * constr_pattern * constr_pattern
- | PSort of rawsort
+ | PSort of glob_sort
| PMeta of patvar option
| PIf of constr_pattern * constr_pattern * constr_pattern
- | PCase of (case_style * int array * inductive option * (int * int) option)
- * constr_pattern * constr_pattern * constr_pattern array
+ | 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
@@ -57,16 +61,11 @@ let rec occur_meta_pattern = function
(occur_meta_pattern c1) or (occur_meta_pattern c2)
| PCase(_,p,c,br) ->
(occur_meta_pattern p) or
- (occur_meta_pattern c) or (array_exists occur_meta_pattern br)
+ (occur_meta_pattern c) or
+ (List.exists (fun (_,_,p) -> occur_meta_pattern p) br)
| PMeta _ | PSoApp _ -> true
| PEvar _ | PVar _ | PRef _ | PRel _ | PSort _ | PFix _ | PCoFix _ -> false
-type constr_label =
- | ConstNode of constant
- | IndNode of inductive
- | CstrNode of constructor
- | VarNode of identifier
-
exception BoundPattern;;
let rec head_pattern_bound t =
@@ -100,8 +99,8 @@ let pattern_of_constr sigma t =
| Rel n -> PRel n
| Meta n -> PMeta (Some (id_of_string ("META" ^ string_of_int n)))
| Var id -> PVar id
- | Sort (Prop c) -> PSort (RProp c)
- | Sort (Type _) -> PSort (RType None)
+ | 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)
@@ -130,11 +129,17 @@ let pattern_of_constr sigma t =
| GoalEvar -> PEvar (evk,Array.map pattern_of_constr ctxt)
| _ -> PMeta None)
| Case (ci,p,a,br) ->
- let cip = ci.ci_pp_info in
- let no = Some (ci.ci_npar,cip.ind_nargs) in
- PCase ((cip.style,ci.ci_cstr_nargs,Some ci.ci_ind,no),
- pattern_of_constr p,pattern_of_constr a,
- Array.map pattern_of_constr 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_extensible = false }
+ in
+ let branch_of_constr i c =
+ (i, ci.ci_cstr_ndecls.(i), pattern_of_constr c)
+ in
+ PCase (cip, pattern_of_constr p, pattern_of_constr 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
@@ -151,14 +156,13 @@ let map_pattern_with_binders g f l = function
| PProd (n,a,b) -> PProd (n,f l a,f (g n l) b)
| PLetIn (n,a,b) -> PLetIn (n,f l a,f (g n l) b)
| 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,Array.map (f l) pl)
+ | 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)
(* Non recursive *)
| (PVar _ | PEvar _ | PRel _ | PRef _ | PSort _ | PMeta _
(* Bound to terms *)
| PFix _ | PCoFix _ as x) -> x
-let map_pattern f = map_pattern_with_binders (fun _ () -> ()) (fun () -> f) ()
-
let error_instantiate_pattern id l =
let is = if List.length l = 1 then "is" else "are" in
errorlabstrm "" (str "Cannot substitute the term bound to " ++ pr_id id
@@ -235,14 +239,20 @@ let rec subst_pattern subst pat =
let c2' = subst_pattern subst c2 in
if c' == c && c1' == c1 && c2' == c2 then pat else
PIf (c',c1',c2')
- | PCase ((a,b,ind,n as cs),typ,c,branches) ->
+ | PCase (cip,typ,c,branches) ->
+ let ind = cip.cip_ind in
let ind' = Option.smartmap (Inductiveops.subst_inductive 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
- let branches' = array_smartmap (subst_pattern subst) branches in
- let cs' = if ind == ind' then cs else (a,b,ind',n) in
- if typ' == typ && c' == c && branches' == branches then pat else
- PCase(cs',typ', c', branches')
+ let subst_branch ((i,n,c) as br) =
+ 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
+ if cip' == cip && typ' == typ && c' == c && branches' == branches
+ then pat
+ else PCase(cip', typ', c', branches')
| PFix fixpoint ->
let cstr = mkFix fixpoint in
let fixpoint' = destFix (subst_mps subst cstr) in
@@ -257,94 +267,109 @@ let rec subst_pattern subst pat =
let mkPLambda na b = PLambda(na,PMeta None,b)
let rev_it_mkPLambda = List.fold_right mkPLambda
+let err loc pp = user_err_loc (loc,"pattern_of_glob_constr", pp)
+
let rec pat_of_raw metas vars = function
- | RVar (_,id) ->
+ | GVar (_,id) ->
(try PRel (list_index (Name id) vars)
with Not_found -> PVar id)
- | RPatVar (_,(false,n)) ->
+ | GPatVar (_,(false,n)) ->
metas := n::!metas; PMeta (Some n)
- | RRef (_,gr) ->
+ | GRef (_,gr) ->
PRef (canonical_gr gr)
(* Hack pour ne pas réécrire une interprétation complète des patterns*)
- | RApp (_, RPatVar (_,(true,n)), cl) ->
+ | GApp (_, GPatVar (_,(true,n)), cl) ->
metas := n::!metas; PSoApp (n, List.map (pat_of_raw metas vars) cl)
- | RApp (_,c,cl) ->
+ | GApp (_,c,cl) ->
PApp (pat_of_raw metas vars c,
Array.of_list (List.map (pat_of_raw metas vars) cl))
- | RLambda (_,na,bk,c1,c2) ->
+ | GLambda (_,na,bk,c1,c2) ->
name_iter (fun n -> metas := n::!metas) na;
PLambda (na, pat_of_raw metas vars c1,
pat_of_raw metas (na::vars) c2)
- | RProd (_,na,bk,c1,c2) ->
+ | GProd (_,na,bk,c1,c2) ->
name_iter (fun n -> metas := n::!metas) na;
PProd (na, pat_of_raw metas vars c1,
pat_of_raw metas (na::vars) c2)
- | RLetIn (_,na,c1,c2) ->
+ | GLetIn (_,na,c1,c2) ->
name_iter (fun n -> metas := n::!metas) na;
PLetIn (na, pat_of_raw metas vars c1,
pat_of_raw metas (na::vars) c2)
- | RSort (_,s) ->
+ | GSort (_,s) ->
PSort s
- | RHole _ ->
+ | GHole _ ->
PMeta None
- | RCast (_,c,_) ->
- Flags.if_verbose
- Pp.warning "Cast not taken into account in constr pattern";
+ | GCast (_,c,_) ->
+ Flags.if_warn
+ Pp.msg_warning (str "Cast not taken into account in constr pattern");
pat_of_raw metas vars c
- | RIf (_,c,(_,None),b1,b2) ->
+ | GIf (_,c,(_,None),b1,b2) ->
PIf (pat_of_raw metas vars c,
pat_of_raw metas vars b1,pat_of_raw metas vars b2)
- | RLetTuple (loc,nal,(_,None),b,c) ->
- let mkRLambda c na = RLambda (loc,na,Explicit,RHole (loc,Evd.InternalHole),c) in
- let c = List.fold_left mkRLambda c nal in
- PCase ((LetStyle,[|1|],None,None),PMeta None,pat_of_raw metas vars b,
- [|pat_of_raw metas vars c|])
- | RCases (loc,sty,p,[c,(na,indnames)],brs) ->
- let pred,ind_nargs, ind = match p,indnames with
- | Some p, Some (_,ind,n,nal) ->
- rev_it_mkPLambda nal (mkPLambda na (pat_of_raw metas vars p)),
- Some (n,List.length nal),Some ind
- | _ -> PMeta None, None, None in
- let ind = match ind with Some _ -> ind | None ->
- match brs with
- | (_,_,[PatCstr(_,(ind,_),_,_)],_)::_ -> Some ind
- | _ -> None in
- let cbrs =
- Array.init (List.length brs) (pat_of_raw_branch loc metas vars ind brs)
+ | GLetTuple (loc,nal,(_,None),b,c) ->
+ let mkGLambda c na = GLambda (loc,na,Explicit,GHole (loc,Evd.InternalHole),c) in
+ let c = List.fold_left mkGLambda c nal in
+ let cip =
+ { cip_style = LetStyle;
+ cip_ind = None;
+ cip_ind_args = None;
+ cip_extensible = false }
+ in
+ PCase (cip, PMeta None, pat_of_raw metas vars b,
+ [0,1,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
+ | 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))
+ | _ -> PMeta None
+ in
+ let info =
+ { cip_style = sty;
+ cip_ind = ind;
+ cip_ind_args = ind_nargs;
+ cip_extensible = ext }
in
- let cstr_nargs,brs = (Array.map fst cbrs, Array.map snd cbrs) in
- PCase ((sty,cstr_nargs,ind,ind_nargs), pred,
- pat_of_raw metas vars c, brs)
+ (* Nota : when we have a non-trivial predicate,
+ the inductive type is known. Same when we have at least
+ one non-trivial branch. These facts are used in [Constrextern]. *)
+ PCase (info, pred, pat_of_raw metas vars c, brs)
- | r ->
- let loc = loc_of_rawconstr r in
- user_err_loc (loc,"pattern_of_rawconstr", Pp.str"Non supported pattern.")
+ | r -> err (loc_of_glob_constr r) (Pp.str "Non supported pattern.")
-and pat_of_raw_branch loc metas vars ind brs i =
- let bri = List.filter
- (function
- (_,_,[PatCstr(_,c,lv,Anonymous)],_) -> snd c = i+1
- | (loc,_,_,_) ->
- user_err_loc (loc,"pattern_of_rawconstr",
- Pp.str "Non supported pattern.")) brs in
- match bri with
- | [(_,_,[PatCstr(_,(indsp,_),lv,_)],br)] ->
- if ind <> None & ind <> Some indsp then
- user_err_loc (loc,"pattern_of_rawconstr",
- Pp.str "All constructors must be in the same inductive type.");
- let lna =
- List.map
- (function PatVar(_,na) -> na
- | PatCstr(loc,_,_,_) ->
- user_err_loc (loc,"pattern_of_rawconstr",
- Pp.str "Non supported pattern.")) lv in
- let vars' = List.rev lna @ vars in
- List.length lv, rev_it_mkPLambda lna (pat_of_raw metas vars' br)
- | _ -> user_err_loc (loc,"pattern_of_rawconstr",
- str "No unique branch for " ++ int (i+1) ++
- str"-th constructor.")
+and pats_of_glob_branches loc metas vars ind brs =
+ let get_arg = function
+ | PatVar(_,na) -> na
+ | PatCstr(loc,_,_,_) -> err loc (Pp.str "Non supported pattern.")
+ in
+ let rec get_pat indexes = function
+ | [] -> 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
+ 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)
+ | (loc,_,_,_) :: _ -> err loc (Pp.str "Non supported pattern.")
+ in
+ get_pat Intset.empty brs
-let pattern_of_rawconstr c =
+let pattern_of_glob_constr c =
let metas = ref [] in
let p = pat_of_raw metas [] c in
(!metas,p)
diff --git a/pretyping/pattern.mli b/pretyping/pattern.mli
index 195fecfe..cde6d4dc 100644
--- a/pretyping/pattern.mli
+++ b/pretyping/pattern.mli
@@ -1,14 +1,14 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: pattern.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
+(** This module defines the type of pattern used for pattern-matching
+ terms in tatics *)
-(*i*)
open Pp
open Names
open Sign
@@ -16,17 +16,55 @@ open Term
open Environ
open Libnames
open Nametab
-open Rawterm
+open Glob_term
open Mod_subst
-(*i*)
-(* Types of substitutions with or w/o bound variables *)
+(** {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
-(* Patterns *)
+(** {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
@@ -38,14 +76,17 @@ type constr_pattern =
| PLambda of name * constr_pattern * constr_pattern
| PProd of name * constr_pattern * constr_pattern
| PLetIn of name * constr_pattern * constr_pattern
- | PSort of rawsort
+ | PSort of glob_sort
| PMeta of patvar option
| PIf of constr_pattern * constr_pattern * constr_pattern
- | PCase of (case_style * int array * inductive option * (int * int) option)
- * constr_pattern * constr_pattern * constr_pattern array
+ | 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
@@ -54,28 +95,28 @@ val subst_pattern : substitution -> constr_pattern -> constr_pattern
exception BoundPattern
-(* [head_pattern_bound t] extracts the head variable/constant of the
+(** [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
+(** [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
+(** [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_rawconstr l c] translates a term [c] with metavariables into
+(** [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_rawconstr : rawconstr ->
+val pattern_of_glob_constr : glob_constr ->
patvar list * constr_pattern
val instantiate_pattern :
diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml
index 688a25b9..2cf16791 100644
--- a/pretyping/pretype_errors.ml
+++ b/pretyping/pretype_errors.ml
@@ -1,15 +1,13 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: pretype_errors.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
+open Compat
open Util
-open Stdpp
open Names
open Sign
open Term
@@ -17,7 +15,7 @@ open Termops
open Namegen
open Environ
open Type_errors
-open Rawterm
+open Glob_term
open Inductiveops
type pretype_error =
@@ -40,12 +38,13 @@ type pretype_error =
| VarNotFound of identifier
| UnexpectedType of constr * constr
| NotProduct of constr
+ | TypingError of type_error
-exception PretypeError of env * pretype_error
+exception PretypeError of env * Evd.evar_map * pretype_error
let precatchable_exception = function
| Util.UserError _ | TypeError _ | PretypeError _
- | Stdpp.Exc_located(_,(Util.UserError _ | TypeError _ |
+ | Loc.Exc_located(_,(Util.UserError _ | TypeError _ |
Nametab.GlobalizationError _ | PretypeError _)) -> true
| _ -> false
@@ -57,25 +56,22 @@ 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 jl_nf_betaiotaevar sigma jl =
- List.map (j_nf_betaiotaevar 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_ise sigma env =
- let sign = named_context_val env in
- let ctxt = rel_context env in
- let env0 = reset_with_named_context sign env in
- Sign.fold_rel_context
- (fun (na,b,ty) e ->
- push_rel
- (na, Option.map (nf_evar sigma) b, nf_evar sigma ty)
- e)
- ctxt
- ~init:env0
-
-(* This simplify the typing context of Cases clauses *)
+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
@@ -99,111 +95,92 @@ let contract2 env a b = match contract env [a;b] with
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,ctx,sigma,te) =
- Stdpp.raise_with_loc loc (PretypeError(env_ise sigma ctx,te))
+let raise_pretype_error (loc,env,sigma,te) =
+ Loc.raise loc (PretypeError(env,sigma,te))
-let raise_located_type_error (loc,ctx,sigma,te) =
- Stdpp.raise_with_loc loc (TypeError(env_ise sigma ctx,te))
+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 j = j_nf_evar sigma {uj_val=c;uj_type=actty} in
- raise_located_type_error
- (loc, env, sigma, ActualType (j, nf_evar sigma expty))
+ let j = {uj_val=c;uj_type=actty} in
+ raise_located_type_error (loc, env, sigma, ActualType (j, expty))
let error_cant_apply_not_functional_loc loc env sigma rator randl =
- let ja = Array.of_list (jl_nf_evar sigma randl) in
raise_located_type_error
- (loc, env, sigma,
- CantApplyNonFunctional (j_nf_evar sigma rator, ja))
+ (loc, env, sigma, CantApplyNonFunctional (rator, Array.of_list randl))
let error_cant_apply_bad_type_loc loc env sigma (n,c,t) rator randl =
- let ja = Array.of_list (jl_nf_betaiotaevar sigma randl) in
raise_located_type_error
(loc, env, sigma,
- CantApplyBadType
- ((n,nf_evar sigma c, Reductionops.nf_betaiota sigma t),
- j_nf_evar sigma rator, ja))
+ CantApplyBadType ((n,c,t), rator, Array.of_list randl))
let error_ill_formed_branch_loc loc env sigma c i actty expty =
- let simp t = Reduction.nf_betaiota (nf_evar sigma t) in
raise_located_type_error
- (loc, env, sigma,
- IllFormedBranch (nf_evar sigma c,i,simp actty, simp expty))
+ (loc, env, sigma, IllFormedBranch (c, i, actty, expty))
let error_number_branches_loc loc env sigma cj expn =
- raise_located_type_error
- (loc, env, sigma,
- NumberBranches (j_nf_evar sigma cj, expn))
+ raise_located_type_error (loc, env, sigma, NumberBranches (cj, expn))
let error_case_not_inductive_loc loc env sigma cj =
- raise_located_type_error
- (loc, env, sigma, CaseNotInductive (j_nf_evar sigma cj))
+ raise_located_type_error (loc, env, sigma, CaseNotInductive cj)
let error_ill_typed_rec_body_loc loc env sigma i na jl tys =
raise_located_type_error
- (loc, env, sigma,
- IllTypedRecBody (i,na,jv_nf_evar sigma jl,
- Array.map (nf_evar sigma) tys))
+ (loc, env, sigma, IllTypedRecBody (i, na, jl, tys))
let error_not_a_type_loc loc env sigma j =
- raise_located_type_error (loc, env, sigma, NotAType (j_nf_evar sigma j))
+ raise_located_type_error (loc, env, sigma, NotAType j)
(*s Implicit arguments synthesis errors. It is hard to find
a precise location. *)
let error_occur_check env sigma ev c =
- let c = nf_evar sigma c in
- raise (PretypeError (env_ise sigma env, OccurCheck (ev,c)))
+ raise (PretypeError (env, sigma, OccurCheck (ev,c)))
let error_not_clean env sigma ev c (loc,k) =
- let c = nf_evar sigma c in
- raise_with_loc loc
- (PretypeError (env_ise sigma env, NotClean (ev,c,k)))
+ Loc.raise loc (PretypeError (env, sigma, NotClean (ev,c,k)))
let error_unsolvable_implicit loc env sigma evi e explain =
- raise_with_loc loc
- (PretypeError (env_ise sigma env, UnsolvableImplicit (evi, e, explain)))
+ Loc.raise loc
+ (PretypeError (env, sigma, UnsolvableImplicit (evi, e, explain)))
let error_cannot_unify env sigma (m,n) =
- raise (PretypeError (env_ise sigma env,CannotUnify (m,n)))
+ raise (PretypeError (env, sigma,CannotUnify (m,n)))
let error_cannot_unify_local env sigma (m,n,sn) =
- raise (PretypeError (env_ise sigma env,CannotUnifyLocal (m,n,sn)))
+ raise (PretypeError (env, sigma,CannotUnifyLocal (m,n,sn)))
let error_cannot_coerce env sigma (m,n) =
- raise (PretypeError (env_ise sigma env,CannotUnify (m,n)))
+ raise (PretypeError (env, sigma,CannotUnify (m,n)))
let error_cannot_find_well_typed_abstraction env sigma p l =
- raise (PretypeError (env_ise sigma env,CannotFindWellTypedAbstraction (p,l)))
+ raise (PretypeError (env, sigma,CannotFindWellTypedAbstraction (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
- raise (PretypeError (env_ise sigma env,AbstractionOverMeta (m,n)))
+ raise (PretypeError (env, sigma,AbstractionOverMeta (m,n)))
let error_non_linear_unification env sigma hdmeta t =
let m = Evd.meta_name sigma hdmeta in
- raise (PretypeError (env_ise sigma env,NonLinearUnification (m,t)))
+ raise (PretypeError (env, sigma,NonLinearUnification (m,t)))
(*s Ml Case errors *)
let error_cant_find_case_type_loc loc env sigma expr =
- raise_pretype_error
- (loc, env, sigma, CantFindCaseType (nf_evar sigma expr))
+ raise_pretype_error (loc, env, sigma, CantFindCaseType 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 (nf_evar sigma actty, nf_evar sigma expty))
+ raise_pretype_error (loc, env, sigma, UnexpectedType (actty, expty))
let error_not_product_loc loc env sigma c =
- raise_pretype_error (loc, env, sigma, NotProduct (nf_evar sigma c))
+ raise_pretype_error (loc, env, sigma, NotProduct c)
-(*s Error in conversion from AST to rawterms *)
+(*s Error in conversion from AST to glob_constr *)
let error_var_not_found_loc loc s =
raise_pretype_error (loc, empty_env, Evd.empty, VarNotFound s)
diff --git a/pretyping/pretype_errors.mli b/pretyping/pretype_errors.mli
index d3e6ffea..3a3dccb4 100644
--- a/pretyping/pretype_errors.mli
+++ b/pretyping/pretype_errors.mli
@@ -1,30 +1,26 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: pretype_errors.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
open Pp
open Util
open Names
open Term
open Sign
open Environ
-open Rawterm
+open Glob_term
open Inductiveops
-(*i*)
-(*s The type of errors raised by the pretyper *)
+(** {6 The type of errors raised by the pretyper } *)
type pretype_error =
- (* Old Case *)
+ (** Old Case *)
| CantFindCaseType of constr
- (* Unification *)
+ (** Unification *)
| OccurCheck of existential_key * constr
| NotClean of existential_key * constr * Evd.hole_kind
| UnsolvableImplicit of Evd.evar_info * Evd.hole_kind *
@@ -37,51 +33,55 @@ type pretype_error =
| CannotFindWellTypedAbstraction of constr * constr list
| AbstractionOverMeta of name * name
| NonLinearUnification of name * constr
- (* Pretyping *)
+ (** Pretyping *)
| VarNotFound of identifier
| UnexpectedType of constr * constr
| NotProduct of constr
+ | TypingError of Type_errors.type_error
-exception PretypeError of env * pretype_error
+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
+(** 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 *)
+(** Raising errors *)
val error_actual_type_loc :
- loc -> env -> Evd.evar_map -> unsafe_judgment -> constr -> 'b
+ loc -> env -> Evd.evar_map -> unsafe_judgment -> constr -> 'b
val error_cant_apply_not_functional_loc :
- loc -> env -> Evd.evar_map ->
+ loc -> 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 -> 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 -> env -> Evd.evar_map -> unsafe_judgment -> 'b
val error_ill_formed_branch_loc :
- loc -> env -> Evd.evar_map ->
- constr -> int -> constr -> constr -> 'b
+ loc -> env -> Evd.evar_map ->
+ constr -> constructor -> constr -> constr -> 'b
val error_number_branches_loc :
- loc -> env -> Evd.evar_map ->
+ loc -> env -> Evd.evar_map ->
unsafe_judgment -> int -> 'b
val error_ill_typed_rec_body_loc :
- loc -> env -> Evd.evar_map ->
+ loc -> env -> Evd.evar_map ->
int -> name array -> unsafe_judgment array -> types array -> 'b
val error_not_a_type_loc :
@@ -89,9 +89,9 @@ val error_not_a_type_loc :
val error_cannot_coerce : env -> Evd.evar_map -> constr * constr -> 'b
-(*s Implicit arguments synthesis errors *)
+(** {6 Implicit arguments synthesis errors } *)
-val error_occur_check : env -> Evd.evar_map -> existential_key -> constr -> 'b
+val error_occur_check : env -> Evd.evar_map -> existential_key -> constr -> 'b
val error_not_clean :
env -> Evd.evar_map -> existential_key -> constr -> loc * Evd.hole_kind -> 'b
@@ -113,19 +113,19 @@ val error_abstraction_over_meta : env -> Evd.evar_map ->
val error_non_linear_unification : env -> Evd.evar_map ->
metavariable -> constr -> 'b
-(*s Ml Case errors *)
+(** {6 Ml Case errors } *)
val error_cant_find_case_type_loc :
- loc -> env -> Evd.evar_map -> constr -> 'b
+ loc -> env -> Evd.evar_map -> constr -> 'b
-(*s Pretyping errors *)
+(** {6 Pretyping errors } *)
val error_unexpected_type_loc :
- loc -> env -> Evd.evar_map -> constr -> constr -> 'b
+ loc -> env -> Evd.evar_map -> constr -> constr -> 'b
val error_not_product_loc :
- loc -> env -> Evd.evar_map -> constr -> 'b
+ loc -> env -> Evd.evar_map -> constr -> 'b
-(*s Error in conversion from AST to rawterms *)
+(** {6 Error in conversion from AST to glob_constr } *)
val error_var_not_found_loc : loc -> identifier -> 'b
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index a3b1dd18..901936f3 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -1,13 +1,27 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: pretyping.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
+(* This file contains the syntax-directed part of the type inference
+ algorithm introduced by Murthy in Coq V5.10, 1995; the type
+ inference algorithm was initially developed in a file named trad.ml
+ which formerly contained a simple concrete-to-abstract syntax
+ translation function introduced in CoC V4.10 for implementing the
+ "exact" tactic, 1989 *)
+(* Support for typing term in Ltac environment by David Delahaye, 2000 *)
+(* Type inference algorithm made a functor of the coercion and
+ pattern-matching compilation by Matthieu Sozeau, March 2006 *)
+(* Fixpoint guard index computation by Pierre Letouzey, July 2007 *)
+(* Structural maintainer: Hugo Herbelin *)
+(* Secondary maintenance: collective *)
+
+
+open Compat
open Pp
open Util
open Names
@@ -26,7 +40,7 @@ open List
open Recordops
open Evarutil
open Pretype_errors
-open Rawterm
+open Glob_term
open Evarconv
open Pattern
@@ -34,7 +48,8 @@ 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
-type rawconstr_ltac_closure = ltac_var_map * rawconstr
+type glob_constr_ltac_closure = ltac_var_map * glob_constr
+type pure_open_constr = evar_map * constr
(************************************************************************)
(* This concerns Cases *)
@@ -55,7 +70,7 @@ let search_guard loc env possible_indexes fixdefs =
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 -> if loc = dummy_loc then raise e else Stdpp.raise_with_loc loc e);
+ | e -> if loc = dummy_loc then raise e else Loc.raise loc e);
indexes
else
(* we now search recursively amoungst all combinations *)
@@ -72,20 +87,53 @@ let search_guard loc env possible_indexes fixdefs =
user_err_loc (loc,"search_guard", Pp.str errmsg)
with Found indexes -> indexes)
-(* To embed constr in rawconstr *)
+(* To embed constr in glob_constr *)
let ((constr_in : constr -> Dyn.t),
(constr_out : Dyn.t -> constr)) = Dyn.create "constr"
(** Miscellaneous interpretation functions *)
let interp_sort = function
- | RProp c -> Prop c
- | RType _ -> new_Type_sort ()
+ | GProp c -> Prop c
+ | GType _ -> new_Type_sort ()
let interp_elimination_sort = function
- | RProp Null -> InProp
- | RProp Pos -> InSet
- | RType _ -> InType
+ | GProp Null -> InProp
+ | GProp Pos -> InSet
+ | GType _ -> InType
+
+let resolve_evars env evdref fail_evar resolve_classes =
+ if resolve_classes then
+ evdref := (Typeclasses.resolve_typeclasses ~onlyargs:false
+ ~split:true ~fail:fail_evar env !evdref);
+ (* Resolve eagerly, potentially making wrong choices *)
+ evdref := (try consider_remaining_unif_problems
+ ~ts:(Typeclasses.classes_transparent_state ()) env !evdref
+ with 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
@@ -95,20 +143,20 @@ sig
(* 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 rawconstr to open_constr, leaving
+ (* 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 -> rawconstr -> open_constr
+ 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 -> rawconstr -> constr
+ evar_map ref -> env -> typing_constraint -> glob_constr -> constr
(* More general entry point with evars from ltac *)
- (* Generic call to the interpreter from rawconstr to constr, failing
- unresolved holes in the rawterm cannot be instantiated.
+ (* 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],
@@ -120,29 +168,29 @@ sig
val understand_ltac :
bool -> evar_map -> env -> ltac_var_map ->
- typing_constraint -> rawconstr -> evar_map * constr
+ typing_constraint -> glob_constr -> pure_open_constr
- (* Standard call to get a constr from a rawconstr, resolving implicit args *)
+ (* Standard call to get a constr from a glob_constr, resolving implicit args *)
val understand : evar_map -> env -> ?expected_type:Term.types ->
- rawconstr -> constr
+ glob_constr -> constr
- (* Idem but the rawconstr is intended to be a type *)
+ (* Idem but the glob_constr is intended to be a type *)
- val understand_type : evar_map -> env -> rawconstr -> constr
+ val understand_type : evar_map -> env -> glob_constr -> constr
(* A generalization of the two previous case *)
val understand_gen : typing_constraint -> evar_map -> env ->
- rawconstr -> constr
+ glob_constr -> constr
(* Idem but returns the judgment of the understood term *)
- val understand_judgment : evar_map -> env -> rawconstr -> unsafe_judgment
+ 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 -> rawconstr -> unsafe_judgment
+ val understand_judgment_tcc : evar_map ref -> env -> glob_constr -> unsafe_judgment
(*i*)
(* Internal of Pretyping...
@@ -150,15 +198,15 @@ sig
*)
val pretype :
type_constraint -> env -> evar_map ref ->
- ltac_var_map -> rawconstr -> unsafe_judgment
+ ltac_var_map -> glob_constr -> unsafe_judgment
val pretype_type :
val_constraint -> env -> evar_map ref ->
- ltac_var_map -> rawconstr -> unsafe_type_judgment
+ ltac_var_map -> glob_constr -> unsafe_type_judgment
val pretype_gen :
bool -> bool -> bool -> evar_map ref -> env ->
- ltac_var_map -> typing_constraint -> rawconstr -> constr
+ ltac_var_map -> typing_constraint -> glob_constr -> constr
(*i*)
end
@@ -207,13 +255,6 @@ module Pretyping_F (Coercion : Coercion.S) = struct
i lna vdefj lar
done
- let check_branches_message loc env evdref c (explft,lft) =
- for i = 0 to Array.length explft - 1 do
- if not (e_cumul env evdref lft.(i) explft.(i)) then
- let sigma = !evdref in
- error_ill_formed_branch_loc loc env sigma c i lft.(i) explft.(i)
- done
-
(* coerce to tycon if any *)
let inh_conv_coerce_to_tycon loc env evdref j = function
| None -> j
@@ -241,45 +282,33 @@ module Pretyping_F (Coercion : Coercion.S) = struct
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
{ uj_val = mkRel n; uj_type = lift n typ }
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 c = substl subst c in
{ uj_val = c; uj_type = protected_get_type_of env sigma c }
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 ->
- try (* To build a nicer ltac error message *)
+ (* [id] not found, build nice error message if [id] yet known from ltac *)
+ 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 ->
+ (* [id] not found, standard error message *)
error_var_not_found_loc loc id
- (* make a dependent predicate from an undependent one *)
-
- let make_dep_of_undep env (IndType (indf,realargs)) pj =
- let n = List.length realargs in
- let rec decomp n p =
- if n=0 then p else
- match kind_of_term p with
- | Lambda (_,_,c) -> decomp (n-1) c
- | _ -> decomp (n-1) (applist (lift 1 p, [mkRel 1]))
- in
- let sign,s = decompose_prod_n n pj.uj_type in
- let ind = build_dependent_inductive env indf in
- let s' = mkProd (Anonymous, ind, s) in
- let ccl = lift 1 (decomp n pj.uj_val) in
- let ccl' = mkLambda (Anonymous, ind, ccl) in
- {uj_val=it_mkLambda ccl' sign; uj_type=it_mkProd s' sign}
-
let evar_kind_of_term sigma c =
kind_of_term (whd_evar sigma c)
@@ -290,27 +319,30 @@ module Pretyping_F (Coercion : Coercion.S) = struct
let c = constr_of_global ref in
make_judge c (Retyping.get_type_of env Evd.empty c)
- let pretype_sort = function
- | RProp c -> judge_of_prop_contents c
- | RType _ -> judge_of_new_Type ()
+ 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 (tycon : type_constraint) env evdref lvar = function
- | RRef (loc,ref) ->
+ | GRef (loc,ref) ->
inh_conv_coerce_to_tycon loc env evdref
(pretype_ref evdref env ref)
tycon
- | RVar (loc, id) ->
+ | GVar (loc, id) ->
inh_conv_coerce_to_tycon loc env evdref
(pretype_id loc env !evdref lvar id)
tycon
- | REvar (loc, evk, instopt) ->
+ | 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
@@ -321,24 +353,23 @@ module Pretyping_F (Coercion : Coercion.S) = struct
let j = (Retyping.get_judgment_of env !evdref c) in
inh_conv_coerce_to_tycon loc env evdref j tycon
- | RPatVar (loc,(someta,n)) ->
+ | GPatVar (loc,(someta,n)) ->
let ty =
match tycon with
| Some (None, ty) -> ty
- | None | Some _ ->
- e_new_evar evdref env ~src:(loc,InternalHole) (new_Type ()) in
+ | 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 }
- | RHole (loc,k) ->
+ | GHole (loc,k) ->
let ty =
match tycon with
| Some (None, ty) -> ty
| None | Some _ ->
- e_new_evar evdref env ~src:(loc,InternalHole) (new_Type ()) in
+ new_type_evar evdref env loc in
{ uj_val = e_new_evar evdref env ~src:(loc,k) ty; uj_type = ty }
- | RRec (loc,fixkind,names,bl,lar,vdef) ->
+ | GRec (loc,fixkind,names,bl,lar,vdef) ->
let rec type_bl env ctxt = function
[] -> ctxt
| (na,bk,None,ty)::bl ->
@@ -379,7 +410,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
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
- | RFix (vn,i) ->
+ | 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,
@@ -395,22 +426,23 @@ module Pretyping_F (Coercion : Coercion.S) = struct
let fixdecls = (names,ftys,fdefs) in
let indexes = search_guard loc env possible_indexes fixdecls in
make_judge (mkFix ((indexes,i),fixdecls)) ftys.(i)
- | RCoFix i ->
+ | GCoFix i ->
let cofix = (i,(names,ftys,fdefs)) in
- (try check_cofix env cofix with e -> Stdpp.raise_with_loc loc e);
+ (try check_cofix env cofix with e -> Loc.raise loc e);
make_judge (mkCoFix cofix) ftys.(i) in
inh_conv_coerce_to_tycon loc env evdref fixj tycon
- | RSort (loc,s) ->
- inh_conv_coerce_to_tycon loc env evdref (pretype_sort s) tycon
+ | GSort (loc,s) ->
+ let j = pretype_sort evdref s in
+ inh_conv_coerce_to_tycon loc env evdref j tycon
- | RApp (loc,f,args) ->
+ | GApp (loc,f,args) ->
let fj = pretype empty_tycon env evdref lvar f in
- let floc = loc_of_rawconstr 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_rawconstr c in
+ let argloc = loc_of_glob_constr c in
let resj = evd_comb1 (Coercion.inh_app_fun env) evdref resj in
let resty = whd_betadeltaiota env !evdref resj.uj_type in
match kind_of_term resty with
@@ -444,7 +476,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
| _ -> resj in
inh_conv_coerce_to_tycon loc env evdref resj tycon
- | RLambda(loc,name,bk,c1,c2) ->
+ | 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
@@ -452,7 +484,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
let j' = pretype rng (push_rel var env) evdref lvar c2 in
judge_of_abstraction env (orelse_name name name') j j'
- | RProd(loc,name,bk,c1,c2) ->
+ | GProd(loc,name,bk,c1,c2) ->
let j = pretype_type empty_valcon env evdref lvar c1 in
let j' =
if name = Anonymous then
@@ -465,13 +497,13 @@ module Pretyping_F (Coercion : Coercion.S) = struct
in
let resj =
try judge_of_product env name j j'
- with TypeError _ as e -> Stdpp.raise_with_loc loc e in
+ with TypeError _ as e -> Loc.raise loc e in
inh_conv_coerce_to_tycon loc env evdref resj tycon
- | RLetIn(loc,name,c1,c2) ->
+ | GLetIn(loc,name,c1,c2) ->
let j =
match c1 with
- | RCast (loc, c, CastConv (DEFAULTcast, t)) ->
+ | 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
@@ -483,12 +515,12 @@ module Pretyping_F (Coercion : Coercion.S) = struct
{ uj_val = mkLetIn (name, j.uj_val, t, j'.uj_val) ;
uj_type = subst1 j.uj_val j'.uj_type }
- | RLetTuple (loc,nal,(na,po),c,d) ->
+ | 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_rawconstr c in
+ 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
@@ -524,10 +556,11 @@ module Pretyping_F (Coercion : Coercion.S) = struct
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 mis,_ = dest_ind_family indf in
- let ci = make_case_info env mis LetStyle in
- mkCase (ci, p, cj.uj_val,[|f|]) in
- { uj_val = v; uj_type = substl (realargs@[cj.uj_val]) ccl }
+ 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
@@ -543,18 +576,18 @@ module Pretyping_F (Coercion : Coercion.S) = struct
let ccl = refresh_universes ccl in
let p = it_mkLambda_or_LetIn (lift (nar+1) ccl) psign in
let v =
- let mis,_ = dest_ind_family indf in
- let ci = make_case_info env mis LetStyle in
- mkCase (ci, p, cj.uj_val,[|f|] )
- in
- { uj_val = v; uj_type = ccl })
+ 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 })
- | RIf (loc,c,(na,po),b1,b2) ->
+ | 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_rawconstr c in
+ 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
@@ -577,15 +610,11 @@ module Pretyping_F (Coercion : Coercion.S) = struct
let ccl = nf_evar !evdref pj.utj_val in
let pred = it_mkLambda_or_LetIn ccl psign in
let typ = lift (- nar) (beta_applist (pred,[cj.uj_val])) in
- let jtyp = inh_conv_coerce_to_tycon loc env evdref {uj_val = pred;
- uj_type = typ} tycon
- in
- jtyp.uj_val, jtyp.uj_type
+ pred, typ
| None ->
let p = match tycon with
| Some (None, ty) -> ty
- | None | Some _ ->
- e_new_evar evdref env ~src:(loc,InternalHole) (new_Type ())
+ | 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
@@ -611,18 +640,20 @@ module Pretyping_F (Coercion : Coercion.S) = struct
let b1 = f cstrs.(0) b1 in
let b2 = f cstrs.(1) b2 in
let v =
- let mis,_ = dest_ind_family indf in
- let ci = make_case_info env mis IfStyle in
- mkCase (ci, pred, cj.uj_val, [|b1;b2|])
+ 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|])
in
{ uj_val = v; uj_type = p }
- | RCases (loc,sty,po,tml,eqns) ->
+ | 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)
- | RCast (loc,c,k) ->
+ | GCast (loc,c,k) ->
let cj =
match k with
CastCoerce ->
@@ -633,29 +664,24 @@ module Pretyping_F (Coercion : Coercion.S) = struct
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 when not (occur_existential cty || occur_existential tval) ->
- (try ignore (Reduction.vm_conv Reduction.CUMUL env cty tval); cj
- with Reduction.NotConvertible ->
- error_actual_type_loc loc env !evdref cj tval)
-
+ | 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
- | RDynamic (loc,d) ->
- if (Dyn.tag d) = "constr" then
- let c = constr_out d in
- let j = (Retyping.get_judgment_of env !evdref c) in
- j
- (*inh_conv_coerce_to_tycon loc env evdref j tycon*)
- else
- user_err_loc (loc,"pretype",(str "Not a constr tagged Dynamic."))
-
(* [pretype_type valcon env evdref lvar c] coerces [c] into a type *)
and pretype_type valcon env evdref lvar = function
- | RHole loc ->
+ | GHole loc ->
(match valcon with
| Some v ->
let s =
@@ -670,12 +696,12 @@ module Pretyping_F (Coercion : Coercion.S) = struct
{ utj_val = v;
utj_type = s }
| None ->
- let s = new_Type_sort () in
+ 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 empty_tycon env evdref lvar c in
- let loc = loc_of_rawconstr 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
@@ -683,7 +709,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
if e_cumul env evdref v tj.utj_val then tj
else
error_unexpected_type_loc
- (loc_of_rawconstr c) env !evdref tj.utj_val v
+ (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
@@ -691,18 +717,12 @@ module Pretyping_F (Coercion : Coercion.S) = struct
let tycon = match exptyp with None -> empty_tycon | Some t -> mk_tycon t in
(pretype tycon env evdref lvar c).uj_val
| IsType ->
- (pretype_type empty_valcon env evdref lvar c).utj_val in
- if resolve_classes then (
- evdref := Typeclasses.resolve_typeclasses ~onlyargs:false
- ~split:true ~fail:fail_evar env !evdref);
- evdref := (try consider_remaining_unif_problems env !evdref
- with e when not resolve_classes ->
- consider_remaining_unif_problems env
- (Typeclasses.resolve_typeclasses ~onlyargs:false
- ~split:true ~fail:fail_evar env !evdref));
- 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
+ (pretype_type empty_valcon env evdref lvar c).utj_val
+ 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
@@ -710,26 +730,24 @@ module Pretyping_F (Coercion : Coercion.S) = struct
*)
let understand_judgment sigma env c =
- let evdref = ref (create_evar_defs sigma) in
+ let evdref = ref sigma in
let j = pretype empty_tycon env evdref ([],[]) c in
- let evd = consider_remaining_unif_problems env !evdref in
- let evd = Typeclasses.resolve_typeclasses ~onlyargs:true ~split:false
- ~fail:true env evd
- in
- let j = j_nf_evar evd j in
- check_evars env sigma evd (mkCast(j.uj_val,DEFAULTcast, j.uj_type));
- j
+ 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 empty_tycon env evdref ([],[]) c in
- j_nf_evar !evdref j
+ 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 (Evd.create_evar_defs sigma) in
+ let evdref = ref sigma in
let c = pretype_gen expand_evar fail_evar resolve_classes evdref env lvar kind c in
!evdref, c
diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli
index b94f9789..47b3ec87 100644
--- a/pretyping/pretyping.mli
+++ b/pretyping/pretyping.mli
@@ -1,24 +1,26 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: pretyping.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
+(** This file implements type inference. It maps [glob_constr]
+ (i.e. untyped terms whose names are located) to [constr]. In
+ particular, it drives complex pattern-matching problems ("match")
+ into elementary ones, insertion of coercions and resolution of
+ implicit arguments. *)
-(*i*)
open Names
open Sign
open Term
open Environ
open Evd
-open Rawterm
+open Glob_term
open Evarutil
-(*i*)
-(* An auxiliary function for searching for fixpoint guard indexes *)
+(** An auxiliary function for searching for fixpoint guard indexes *)
val search_guard :
Util.loc -> env -> int list list -> rec_declaration -> int array
@@ -28,30 +30,31 @@ type typing_constraint = OfType of types option | IsType
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 rawconstr_ltac_closure = ltac_var_map * rawconstr
+type glob_constr_ltac_closure = ltac_var_map * glob_constr
+type pure_open_constr = evar_map * constr
module type S =
sig
module Cases : Cases.S
- (* Allow references to syntaxically inexistent variables (i.e., if applied on an inductive) *)
+ (** 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 rawconstr to open_constr, leaving
+ (** 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 -> rawconstr -> open_constr
+ 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 -> rawconstr -> constr
+ evar_map ref -> env -> typing_constraint -> glob_constr -> constr
- (* More general entry point with evars from ltac *)
+ (** More general entry point with evars from ltac *)
- (* Generic call to the interpreter from rawconstr to constr, failing
- unresolved holes in the rawterm cannot be instantiated.
+ (** 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],
@@ -63,56 +66,59 @@ sig
val understand_ltac :
bool -> evar_map -> env -> ltac_var_map ->
- typing_constraint -> rawconstr -> evar_map * constr
+ typing_constraint -> glob_constr -> pure_open_constr
- (* Standard call to get a constr from a rawconstr, resolving implicit args *)
+ (** Standard call to get a constr from a glob_constr, resolving implicit args *)
val understand : evar_map -> env -> ?expected_type:Term.types ->
- rawconstr -> constr
+ glob_constr -> constr
- (* Idem but the rawconstr is intended to be a type *)
+ (** Idem but the glob_constr is intended to be a type *)
- val understand_type : evar_map -> env -> rawconstr -> constr
+ val understand_type : evar_map -> env -> glob_constr -> constr
- (* A generalization of the two previous case *)
+ (** A generalization of the two previous case *)
val understand_gen : typing_constraint -> evar_map -> env ->
- rawconstr -> constr
+ glob_constr -> constr
- (* Idem but returns the judgment of the understood term *)
+ (** Idem but returns the judgment of the understood term *)
- val understand_judgment : evar_map -> env -> rawconstr -> unsafe_judgment
+ 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 -> rawconstr -> 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...
- *)
+ (**/**)
+ (** Internal of Pretyping... *)
val pretype :
type_constraint -> env -> evar_map ref ->
- ltac_var_map -> rawconstr -> unsafe_judgment
+ ltac_var_map -> glob_constr -> unsafe_judgment
val pretype_type :
val_constraint -> env -> evar_map ref ->
- ltac_var_map -> rawconstr -> unsafe_type_judgment
+ ltac_var_map -> glob_constr -> unsafe_type_judgment
val pretype_gen :
bool -> bool -> bool -> evar_map ref -> env ->
- ltac_var_map -> typing_constraint -> rawconstr -> constr
+ ltac_var_map -> typing_constraint -> glob_constr -> constr
- (*i*)
+ (**/**)
end
module Pretyping_F (C : Coercion.S) : S
module Default : S
-(* To embed constr in rawconstr *)
+(** To embed constr in glob_constr *)
val constr_in : constr -> Dyn.t
val constr_out : Dyn.t -> constr
-val interp_sort : rawsort -> sorts
-val interp_elimination_sort : rawsort -> sorts_family
+val interp_sort : glob_sort -> 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
diff --git a/pretyping/pretyping.mllib b/pretyping/pretyping.mllib
index cea33c1e..9eec9414 100644
--- a/pretyping/pretyping.mllib
+++ b/pretyping/pretyping.mllib
@@ -11,8 +11,9 @@ Evarutil
Term_dnet
Recordops
Evarconv
+Arguments_renaming
Typing
-Rawterm
+Glob_term
Pattern
Matching
Tacred
@@ -21,7 +22,6 @@ Typeclasses
Classops
Coercion
Unification
-Clenv
Detyping
Indrec
Cases
diff --git a/pretyping/rawterm.mli b/pretyping/rawterm.mli
deleted file mode 100644
index 10156cec..00000000
--- a/pretyping/rawterm.mli
+++ /dev/null
@@ -1,167 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(*i $Id: rawterm.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
-open Util
-open Names
-open Sign
-open Term
-open Libnames
-open Nametab
-(*i*)
-
-(**********************************************************************)
-(* The kind of patterns that occurs in "match ... with ... end" *)
-
-(* 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
-
-val cases_pattern_loc : cases_pattern -> loc
-
-(**********************************************************************)
-(* 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 *)
-
-type patvar = identifier
-
-type rawsort = RProp of Term.contents | RType 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 rawconstr =
- | RRef of (loc * global_reference)
- | RVar of (loc * identifier)
- | REvar of loc * existential_key * rawconstr list option
- | RPatVar of loc * (bool * patvar) (* Used for patterns only *)
- | RApp of loc * rawconstr * rawconstr list
- | RLambda of loc * name * binding_kind * rawconstr * rawconstr
- | RProd of loc * name * binding_kind * rawconstr * rawconstr
- | RLetIn of loc * name * rawconstr * rawconstr
- | RCases of loc * case_style * rawconstr option * tomatch_tuples * cases_clauses
- | RLetTuple of loc * name list * (name * rawconstr option) *
- rawconstr * rawconstr
- | RIf of loc * rawconstr * (name * rawconstr option) * rawconstr * rawconstr
- | RRec of loc * fix_kind * identifier array * rawdecl list array *
- rawconstr array * rawconstr array
- | RSort of loc * rawsort
- | RHole of (loc * Evd.hole_kind)
- | RCast of loc * rawconstr * rawconstr cast_type
- | RDynamic of loc * Dyn.t
-
-and rawdecl = name * binding_kind * rawconstr option * rawconstr
-
-and fix_recursion_order = RStructRec | RWfRec of rawconstr | RMeasureRec of rawconstr * rawconstr option
-
-and fix_kind =
- | RFix of ((int option * fix_recursion_order) array * int)
- | RCoFix of int
-
-and predicate_pattern =
- name * (loc * inductive * int * name list) option
-
-and tomatch_tuple = (rawconstr * predicate_pattern)
-
-and tomatch_tuples = tomatch_tuple list
-
-and cases_clause = (loc * identifier list * cases_pattern list * rawconstr)
-
-and cases_clauses = cases_clause list
-
-val cases_predicate_names : tomatch_tuples -> name list
-
-val map_rawconstr : (rawconstr -> rawconstr) -> rawconstr -> rawconstr
-
-(** Ensure traversal from left to right *)
-val map_rawconstr_left_to_right :
- (rawconstr -> rawconstr) -> rawconstr -> rawconstr
-
-(*i
-val map_rawconstr_with_binders_loc : loc ->
- (identifier -> 'a -> identifier * 'a) ->
- ('a -> rawconstr -> rawconstr) -> 'a -> rawconstr -> rawconstr
-i*)
-
-val fold_rawconstr : ('a -> rawconstr -> 'a) -> 'a -> rawconstr -> 'a
-val iter_rawconstr : (rawconstr -> unit) -> rawconstr -> unit
-val occur_rawconstr : identifier -> rawconstr -> bool
-val free_rawvars : rawconstr -> identifier list
-val loc_of_rawconstr : rawconstr -> loc
-
-(**********************************************************************)
-(* Conversion from rawconstr to cases pattern, if possible *)
-
-(* Take the current alias as parameter, raise Not_found if *)
-(* translation is impossible *)
-
-val cases_pattern_of_rawconstr : name -> rawconstr -> cases_pattern
-
-val rawconstr_of_closed_cases_pattern : cases_pattern -> name * rawconstr
-
-(**********************************************************************)
-(* Reduction expressions *)
-
-type 'a raw_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 raw_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 raw_red_flag
- | Lazy of 'b raw_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/recordops.ml b/pretyping/recordops.ml
index e8c6073e..b3be7afd 100644
--- a/pretyping/recordops.ml
+++ b/pretyping/recordops.ml
@@ -1,12 +1,17 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: recordops.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
+(* 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 Util
open Pp
@@ -14,7 +19,6 @@ open Names
open Libnames
open Nametab
open Term
-open Termops
open Typeops
open Libobject
open Library
@@ -39,6 +43,12 @@ type struc_typ = {
let structure_table = ref (Indmap.empty : struc_typ Indmap.t)
let projection_table = ref Cmap.empty
+(* 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
+
let load_structure i (_,(ind,id,kl,projs)) =
let n = (fst (Global.lookup_inductive ind)).Declarations.mind_nparams in
let struc =
@@ -71,7 +81,7 @@ let discharge_structure (_,(ind,id,kl,projs)) =
Some (Lib.discharge_inductive ind, discharge_constructor id, kl,
List.map (Option.map Lib.discharge_con) projs)
-let (inStruc,outStruc) =
+let inStruc : struc_tuple -> obj =
declare_object {(default_object "STRUCTURE") with
cache_function = cache_structure;
load_function = load_structure;
@@ -130,7 +140,7 @@ open Libobject
let load_method (_,(ty,id)) =
meth_dnet := MethodsDnet.add ty id !meth_dnet
-let (in_method,out_method) =
+let in_method : constr * MethodsDnet.ident -> obj =
declare_object
{ (default_object "RECMETHODS") with
load_function = (fun _ -> load_method);
@@ -201,7 +211,7 @@ let cs_pattern_of_constr t =
_ -> raise Not_found
end
| Rel n -> Default_cs, pred n, []
- | Prod (_,a,b) when not (dependent (mkRel 1) b) -> Prod_cs, -1, [a;pop b]
+ | 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, []
| _ ->
begin
@@ -235,7 +245,7 @@ let compute_canonical_projections (con,ind) =
(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"
- ++ print_constr t ++ str " in instance "
+ ++ Termops.print_constr t ++ str " in instance "
++ con_pp ++ str " of " ++ proji_sp_pp ++ str ", ignoring it."));
l
end
@@ -285,7 +295,7 @@ let subst_canonical_structure (subst,(cst,ind as obj)) =
let discharge_canonical_structure (_,(cst,ind)) =
Some (Lib.discharge_con cst,Lib.discharge_inductive ind)
-let (inCanonStruc,outCanonStruct) =
+let inCanonStruc : constant * inductive -> obj =
declare_object {(default_object "CANONICAL-STRUCTURE") with
open_function = open_canonical_structure;
cache_function = cache_canonical_structure;
@@ -323,16 +333,17 @@ let check_and_decompose_canonical_structure ref =
let declare_canonical_structure ref =
add_canonical_structure (check_and_decompose_canonical_structure ref)
-let outCanonicalStructure x = fst (outCanonStruct x)
-
let lookup_canonical_conversion (proj,pat) =
List.assoc pat (Refmap.find proj !object_table)
-let is_open_canonical_projection sigma (c,args) =
+let is_open_canonical_projection env sigma (c,args) =
try
- let l = Refmap.find (global_of_constr c) !object_table in
- let n = (snd (List.hd l)).o_NPARAMS in
- try isEvar_or_Meta (whd_evar sigma (List.nth args n)) with Failure _ -> false
+ let n = find_projection_nparams (global_of_constr c) in
+ try
+ let arg = whd_betadeltaiota env sigma (List.nth args n) in
+ let hd = match kind_of_term arg with App (hd, _) -> hd | _ -> arg in
+ not (isConstruct hd)
+ with Failure _ -> false
with Not_found -> false
let freeze () =
diff --git a/pretyping/recordops.mli b/pretyping/recordops.mli
index b71c4969..6165fac2 100644
--- a/pretyping/recordops.mli
+++ b/pretyping/recordops.mli
@@ -1,23 +1,22 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: recordops.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
open Names
open Nametab
open Term
open Libnames
open Libobject
open Library
-(*i*)
-(*s A structure S is a non recursive inductive type with a single
+(** Operations concerning records and canonical structures *)
+
+(** {6 Records } *)
+(** A structure S is a non recursive inductive type with a single
constructor (the name of which defaults to Build_S) *)
type struc_typ = {
@@ -26,37 +25,40 @@ type struc_typ = {
s_PROJKIND : (name * bool) list;
s_PROJ : constant option list }
-val declare_structure :
- inductive * constructor * (name * bool) list * constant option list -> unit
+type struc_tuple =
+ inductive * constructor * (name * bool) list * constant option list
+
+val declare_structure : struc_tuple -> unit
-(* [lookup_structure isp] returns the struc_typ associated to the
+(** [lookup_structure isp] returns the struc_typ associated to the
inductive path [isp] if it corresponds to a structure, otherwise
it fails with [Not_found] *)
val lookup_structure : inductive -> struc_typ
-(* [lookup_projections isp] returns the projections associated to the
+(** [lookup_projections isp] returns the projections associated to the
inductive path [isp] if it corresponds to a structure, otherwise
it fails with [Not_found] *)
val lookup_projections : inductive -> constant option list
-(* raise [Not_found] if not a projection *)
+(** raise [Not_found] if not a projection *)
val find_projection_nparams : global_reference -> int
-(* raise [Not_found] if not a projection *)
+(** raise [Not_found] if not a projection *)
val find_projection : global_reference -> struc_typ
-(* we keep an index (dnet) of record's arguments + fields
+(** 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: *)
+ (** 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
-(*s A canonical structure declares "canonical" conversion hints between *)
-(* the effective components of a structure and the projections of the *)
-(* structure *)
+(** {6 Canonical structures } *)
+(** A canonical structure declares "canonical" conversion hints between
+ the effective components of a structure and the projections of the
+ structure *)
type cs_pattern =
Const_cs of global_reference
@@ -66,11 +68,11 @@ type cs_pattern =
type obj_typ = {
o_DEF : constr;
- o_INJ : int; (* position of trivial argument *)
- o_TABS : constr list; (* ordered *)
- o_TPARAMS : constr list; (* ordered *)
+ o_INJ : int; (** position of trivial argument *)
+ o_TABS : constr list; (** ordered *)
+ o_TPARAMS : constr list; (** ordered *)
o_NPARAMS : int;
- o_TCOMPS : constr list } (* ordered *)
+ o_TCOMPS : constr list } (** ordered *)
val cs_pattern_of_constr : constr -> cs_pattern * int * constr list
val pr_cs_pattern : cs_pattern -> Pp.std_ppcmds
@@ -78,6 +80,6 @@ val pr_cs_pattern : cs_pattern -> Pp.std_ppcmds
val lookup_canonical_conversion : (global_reference * cs_pattern) -> obj_typ
val declare_canonical_structure : global_reference -> unit
val is_open_canonical_projection :
- Evd.evar_map -> (constr * constr list) -> bool
+ Environ.env -> Evd.evar_map -> (constr * constr list) -> bool
val canonical_projections : unit ->
((global_reference * cs_pattern) * obj_typ) list
diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml
index 5af20551..e8acd67c 100644
--- a/pretyping/reductionops.ml
+++ b/pretyping/reductionops.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: reductionops.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
open Pp
open Util
open Names
@@ -44,13 +42,6 @@ let append_stack_list l s =
| (l1, s) -> Zapp l1 :: s
let append_stack v s = append_stack_list (Array.to_list v) s
-(* Collapse the shifts in the stack *)
-let zshift n s =
- match (n,s) with
- (0,_) -> s
- | (_,Zshift(k)::s) -> Zshift(n+k)::s
- | _ -> Zshift(n)::s
-
let rec stack_args_size = function
| Zapp l::s -> List.length l + stack_args_size s
| Zshift(_)::s -> stack_args_size s
@@ -63,10 +54,6 @@ let rec decomp_stack = function
| Zapp(v::l)::s -> Some (v, (Zapp l :: s))
| Zapp [] :: s -> decomp_stack s
| _ -> None
-let rec decomp_stackn = function
- | Zapp [] :: s -> decomp_stackn s
- | Zapp l :: s -> (Array.of_list l, s)
- | _ -> assert false
let array_of_stack s =
let rec stackrec = function
| [] -> []
@@ -155,10 +142,6 @@ let whd_stack sigma x =
appterm_of_stack (whd_app_state sigma (x, empty_stack))
let whd_castapp_stack = whd_stack
-let stack_reduction_of_reduction red_fun env sigma s =
- let t = red_fun env sigma (app_stack s) in
- whd_stack t
-
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
@@ -309,7 +292,7 @@ let reduce_fix whdfun sigma fix stack =
-------------------
qui coute cher *)
-let rec whd_state_gen flags env sigma =
+let rec whd_state_gen flags ts env sigma =
let rec whrec (x, stack as s) =
match kind_of_term x with
| Rel n when red_delta flags ->
@@ -328,7 +311,7 @@ let rec whd_state_gen flags env sigma =
(match safe_meta_value sigma ev with
| Some body -> whrec (body, stack)
| None -> s)
- | Const const when red_delta flags ->
+ | Const const when is_transparent_constant ts const ->
(match constant_opt_value env const with
| Some body -> whrec (body, stack)
| None -> s)
@@ -340,7 +323,7 @@ let rec whd_state_gen flags env sigma =
| 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 env' sigma 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
@@ -451,18 +434,19 @@ let whd_betalet = red_of_state_red whd_betalet_state
(* 2. Delta Reduction Functions *)
-let whd_delta_state e = whd_state_gen delta e
+let whd_delta_state e = whd_state_gen delta full_transparent_state 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 e
+let whd_betadelta_state e = whd_state_gen betadelta full_transparent_state 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 e
+let whd_betadeltaeta_state e =
+ whd_state_gen betadeltaeta full_transparent_state e
let whd_betadeltaeta_stack env =
stack_red_of_state_red (whd_betadeltaeta_state env)
let whd_betadeltaeta env =
@@ -478,19 +462,29 @@ 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 e = whd_state_gen betadeltaiota e
+let whd_betadeltaiota_state env =
+ whd_state_gen betadeltaiota full_transparent_state 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_betadeltaiotaeta_state e = whd_state_gen betadeltaiotaeta e
+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_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 e = whd_state_gen betadeltaiota_nolet e
+let whd_betadeltaiota_nolet_state env =
+ whd_state_gen betadeltaiota_nolet full_transparent_state env
let whd_betadeltaiota_nolet_stack env =
stack_red_of_state_red (whd_betadeltaiota_nolet_state env)
let whd_betadeltaiota_nolet env =
@@ -586,14 +580,6 @@ let rec whd_betaiota_preserving_vm_cast env sigma t =
let nf_betaiota_preserving_vm_cast =
strong whd_betaiota_preserving_vm_cast
-(* lazy weak head reduction functions *)
-let whd_flags flgs env sigma t =
- try
- whd_val
- (create_clos_infos ~evars:(safe_evar_value sigma) flgs env)
- (inject t)
- with Anomaly _ -> error "Tried to normalized ill-typed term"
-
(********************************************************************)
(* Conversion *)
(********************************************************************)
@@ -620,7 +606,7 @@ let pb_equal = function
let sort_cmp = sort_cmp
-let test_conversion (f:?evars:'a->'b) env sigma x y =
+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
@@ -630,8 +616,8 @@ 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 test_trans_conversion f reds env sigma x y =
- try let _ = f reds env (nf_evar sigma x) (nf_evar sigma y) in true
+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"
@@ -784,7 +770,7 @@ let splay_arity env sigma c =
| Sort s -> l,s
| _ -> invalid_arg "splay_arity"
-let sort_of_arity env c = snd (splay_arity env Evd.empty 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
@@ -820,19 +806,21 @@ let is_sort env sigma arity =
(* 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 env sigma s =
+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 env sigma d in
+ 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 env) sigma fix stack with
+ (match
+ reduce_fix (whd_betadeltaiota_state_using ts env) sigma fix stack
+ with
| Reduced s -> whrec s
| NotReducible -> s)
| _ -> s
@@ -926,17 +914,6 @@ let nf_meta sigma c = meta_instance sigma (mk_freelisted c)
(* Instantiate metas that create beta/iota redexes *)
-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
- | None -> mkMeta mv
- in
- valrec mv
-
let meta_reducible_instance evd b =
let fm = Metaset.elements b.freemetas in
let metas = List.fold_left (fun l mv ->
diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli
index 95032bde..3ffc587e 100644
--- a/pretyping/reductionops.mli
+++ b/pretyping/reductionops.mli
@@ -1,28 +1,24 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: reductionops.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
open Names
open Term
open Univ
open Evd
open Environ
open Closure
-(*i*)
-(* Reduction Functions. *)
+(** Reduction Functions. *)
exception Elimconst
-(************************************************************************)
-(*s A [stack] is a context of arguments, arguments are pushed by
+(***********************************************************************
+ 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 *)
@@ -67,13 +63,13 @@ 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 *)
+(** Removes cast and put into applicative form *)
val whd_stack : local_stack_reduction_function
-(* For compatibility: alias for whd\_stack *)
+(** For compatibility: alias for whd\_stack *)
val whd_castapp_stack : local_stack_reduction_function
-(*s Reduction Function Operators *)
+(** {6 Reduction Function Operators } *)
val strong : reduction_function -> reduction_function
val local_strong : local_reduction_function -> local_reduction_function
@@ -84,17 +80,19 @@ val stack_reduction_of_reduction :
i*)
val stacklam : (state -> 'a) -> constr list -> constr -> constr stack -> 'a
-(*s Generic Optimized Reduction Function using Closures *)
+(** {6 Generic Optimized Reduction Function using Closures } *)
val clos_norm_flags : Closure.RedFlags.reds -> reduction_function
-(* Same as [(strong whd_beta[delta][iota])], but much faster on big terms *)
+
+(** 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_betadeltaiota : reduction_function
val nf_evar : evar_map -> constr -> constr
val nf_betaiota_preserving_vm_cast : reduction_function
-(* Lazy strategy, weak head reduction *)
+
+(** Lazy strategy, weak head reduction *)
val whd_evar : evar_map -> constr -> constr
val whd_beta : local_reduction_function
val whd_betaiota : local_reduction_function
@@ -120,7 +118,7 @@ val whd_betadeltaiota_nolet_state : contextual_state_reduction_function
val whd_betaetalet_state : local_state_reduction_function
val whd_betalet_state : local_state_reduction_function
-(*s Head normal forms *)
+(** {6 Head normal forms } *)
val whd_delta_stack : stack_reduction_function
val whd_delta_state : state_reduction_function
@@ -138,7 +136,7 @@ val whd_betadeltaiotaeta : reduction_function
val whd_eta : constr -> constr
val whd_zeta : constr -> constr
-(* Various reduction functions *)
+(** Various reduction functions *)
val safe_evar_value : evar_map -> existential -> constr option
@@ -154,7 +152,7 @@ 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 sort_of_arity : env -> constr -> 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 :
@@ -163,11 +161,11 @@ val decomp_sort : env -> evar_map -> types -> sorts
val is_sort : env -> evar_map -> types -> bool
type 'a miota_args = {
- mP : constr; (* the result type *)
- mconstr : constr; (* the constructor *)
- mci : case_info; (* special info to re-build pattern *)
- mcargs : 'a list; (* the constructor's arguments *)
- mlf : 'a array } (* the branch code vector *)
+ mP : constr; (** the result type *)
+ mconstr : constr; (** the constructor *)
+ mci : case_info; (** special info to re-build pattern *)
+ mcargs : 'a list; (** the constructor's arguments *)
+ mlf : 'a array } (** the branch code vector *)
val reducible_mind_case : constr -> bool
val reduce_mind_case : constr miota_args -> constr
@@ -177,7 +175,7 @@ val is_arity : env -> evar_map -> constr -> bool
val whd_programs : reduction_function
-(* [reduce_fix redfun fix stk] contracts [fix stk] if it is actually
+(** [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
@@ -186,10 +184,10 @@ val fix_recarg : fixpoint -> constr stack -> (int * constr) option
val reduce_fix : local_state_reduction_function -> evar_map -> fixpoint
-> constr stack -> fix_reduction_result
-(*s Querying the kernel conversion oracle: opaque/transparent constants *)
+(** {6 Querying the kernel conversion oracle: opaque/transparent constants } *)
val is_transparent : 'a tableKey -> bool
-(*s Conversion Functions (uses closures, lazy strategy) *)
+(** {6 Conversion Functions (uses closures, lazy strategy) } *)
type conversion_test = constraints -> constraints
@@ -206,18 +204,19 @@ 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
-(*s Special-Purpose Reduction Functions *)
+(** {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 head_unfold_under_prod : transparent_state -> reduction_function
-(*s Heuristic for Conversion with Evar *)
+(** {6 Heuristic for Conversion with Evar } *)
-val whd_betaiota_deltazeta_for_iota_state : state_reduction_function
+val whd_betaiota_deltazeta_for_iota_state :
+ transparent_state -> state_reduction_function
-(*s Meta-related reduction functions *)
+(** {6 Meta-related reduction functions } *)
val meta_instance : evar_map -> constr freelisted -> constr
val nf_meta : evar_map -> constr -> constr
val meta_reducible_instance : evar_map -> constr freelisted -> constr
diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml
index 7fb4f7ba..502e238b 100644
--- a/pretyping/retyping.ml
+++ b/pretyping/retyping.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: retyping.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
open Util
open Term
open Inductive
@@ -48,8 +46,8 @@ 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))
+ (try strip_outer_cast (Evd.meta_ftype sigma n).Evd.rebus
+ with Not_found -> anomaly ("type_of: unknown meta " ^ string_of_int n))
| Rel n ->
let (_,_,ty) = lookup_rel n env in
lift n ty
@@ -129,10 +127,13 @@ let retype ?(polyprop=true) sigma =
match kind_of_term c with
| Ind ind ->
let (_,mip) = lookup_mind_specif env ind in
- Inductive.type_of_inductive_knowing_parameters ~polyprop env mip argtyps
+ (try Inductive.type_of_inductive_knowing_parameters
+ ~polyprop env mip argtyps
+ with Reduction.NotArity -> anomaly "type_of: Not an arity")
| Const cst ->
let t = constant_type env cst in
- Typeops.type_of_constant_knowing_parameters env t argtyps
+ (try Typeops.type_of_constant_knowing_parameters env t argtyps
+ with Reduction.NotArity -> anomaly "type_of: Not an arity")
| Var id -> type_of_var env id
| Construct cstr -> type_of_constructor env cstr
| _ -> assert false
diff --git a/pretyping/retyping.mli b/pretyping/retyping.mli
index f2c030f9..445f623a 100644
--- a/pretyping/retyping.mli
+++ b/pretyping/retyping.mli
@@ -1,21 +1,17 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: retyping.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
open Names
open Term
open Evd
open Environ
-(*i*)
-(* This family of functions assumes its constr argument is known to be
+(** This family of functions assumes its constr argument is known to be
well-typable. It does not type-check, just recompute the type
without any costly verifications. On non well-typable terms, it
either produces a wrong result or raise an anomaly. Use with care.
@@ -33,10 +29,10 @@ val get_sort_of :
val get_sort_family_of :
?polyprop:bool -> env -> evar_map -> types -> sorts_family
-(* Makes an assumption from a constr *)
+(** Makes an assumption from a constr *)
val get_assumption_of : env -> evar_map -> constr -> types
-(* Makes an unsafe judgment from a constr *)
+(** Makes an unsafe judgment from a constr *)
val get_judgment_of : env -> evar_map -> constr -> unsafe_judgment
val type_of_global_reference_knowing_parameters : env -> evar_map -> constr ->
diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml
index bd1eed94..fc35e2d3 100644
--- a/pretyping/tacred.ml
+++ b/pretyping/tacred.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: tacred.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
open Pp
open Util
open Names
@@ -18,11 +16,12 @@ open Termops
open Namegen
open Declarations
open Inductive
+open Libobject
open Environ
open Closure
open Reductionops
open Cbv
-open Rawterm
+open Glob_term
open Pattern
open Matching
@@ -300,17 +299,6 @@ let reference_eval sigma env = function
end)
| ref -> compute_consteval sigma env ref
-let rev_firstn_liftn fn ln =
- let rec rfprec p res l =
- if p = 0 then
- res
- else
- match l with
- | [] -> invalid_arg "Reduction.rev_firstn_liftn"
- | a::rest -> rfprec (p-1) ((lift ln a)::res) rest
- in
- rfprec fn []
-
(* If f is bound to EliminationFix (n',infos), then n' is the minimal
number of args for starting the reduction and infos is
(nbfix,[(yi1,Ti1);...;(yip,Tip)],n) indicating that f converts
@@ -528,27 +516,127 @@ let special_red_case env sigma whfun (ci, p, c, lf) =
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
+
+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 nargs = List.fold_left max nargs recargs 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 get_behaviour = function
+ | EvalVar _ | EvalRel _ | EvalEvar _ -> raise Not_found
+ | EvalConst c -> Refmap.find (ConstRef c) !behaviour_table
+
+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
+
(* [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 =
- match reference_eval sigma env ref with
- | EliminationCases n when stack_args_size largs >= n ->
+ let nargs = stack_args_size largs in
+ let largs, unfold_anyway =
+ match recargs ref with
+ | None -> largs, false
+ | Some (_,n) when nargs < n -> 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 in
+ try match reference_eval sigma env 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)
- | EliminationFix (min,minfxargs,infos) when stack_args_size largs >=min ->
+ | 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
(match reduce_fix_use_function env sigma f whfun (destFix d) lrest with
| NotReducible -> raise Redelimination
- | Reduced (c,rest) -> (nf_beta sigma c, rest))
- | EliminationMutualFix (min,refgoal,refinfos)
- when stack_args_size largs >= min ->
+ | Reduced (c,rest) -> (nf_beta sigma c, rest))
+ | EliminationMutualFix (min,refgoal,refinfos) when nargs >= min ->
let rec descend ref args =
let c = reference_value sigma env ref in
if ref = refgoal then
@@ -564,6 +652,9 @@ let rec red_elim_const env sigma ref largs =
| NotReducible -> raise Redelimination
| Reduced (c,rest) -> (nf_beta sigma c, rest))
| _ -> raise Redelimination
+ with Redelimination when unfold_anyway ->
+ let c = reference_value sigma env ref in
+ whd_betaiotazeta sigma (app_stack (c, largs)), empty_stack
(* reduce to whd normal form or to an applied constant that does not hide
a reducible iota/fix/cofix redex (the "simpl" tactic) *)
@@ -591,7 +682,14 @@ and whd_simpl_state env sigma s =
| _ when isEvalRef env x ->
let ref = destEvalRef x in
(try
- redrec (red_elim_const env sigma ref stack)
+ 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
@@ -698,7 +796,7 @@ let whd_simpl_orelse_delta_but_fix_old env sigma c =
with Redelimination ->
match reference_opt_value sigma env ref with
| Some c ->
- (match kind_of_term ((strip_lam c)) with
+ (match kind_of_term (strip_lam c) with
| CoFix _ | Fix _ -> s
| _ -> redrec (c, stack))
| None -> s)
@@ -708,7 +806,7 @@ let whd_simpl_orelse_delta_but_fix_old env sigma c =
(* Same as [whd_simpl] but also reduces constants that do not hide a
reducible fix, but does this reduction of constants only until it
- it immediately hides a non reducible fix or a cofix *)
+ immediately hides a non reducible fix or a cofix *)
let whd_simpl_orelse_delta_but_fix env sigma c =
let rec redrec s =
@@ -716,7 +814,7 @@ let whd_simpl_orelse_delta_but_fix env sigma c =
if isEvalRef env constr then
match reference_opt_value sigma env (destEvalRef constr) with
| Some c ->
- (match kind_of_term ((strip_lam c)) with
+ (match kind_of_term (strip_lam c) with
| CoFix _ | Fix _ -> s'
| _ -> redrec (c, stack))
| None -> s'
@@ -780,7 +878,7 @@ let substlin env evalref n (nowhere_except_in,locs) c =
let term = constr_of_evaluable_ref evalref in
let rec substrec () c =
if nowhere_except_in & !pos > maxocc then c
- else if c = term then
+ 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
@@ -867,7 +965,7 @@ let abstract_scheme env sigma (locc,a) c =
if occur_meta a then
mkLambda (na,ta,c)
else
- mkLambda (na,ta,subst_term_occ locc a c)
+ mkLambda (na,ta,subst_closed_term_occ locc a c)
let pattern_occs loccs_trm env sigma c =
let abstr_trm = List.fold_right (abstract_scheme env sigma) loccs_trm c in
@@ -905,6 +1003,10 @@ 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 ind,t = reduce_to_atomic_ind env sigma t in
+ ind, snd (decompose_app t)
+
(* Reduce the weak-head redex [beta,iota/fix/cofix[all],cast,zeta,simpl/delta]
or raise [NotStepReducible] if not a weak-head redex *)
@@ -932,7 +1034,7 @@ let one_step_reduce env sigma c =
| _ when isEvalRef env x ->
let ref = destEvalRef x in
(try
- red_elim_const env sigma ref stack
+ red_elim_const env sigma ref stack
with Redelimination ->
match reference_opt_value sigma env ref with
| Some d -> d, stack
@@ -972,7 +1074,7 @@ let reduce_to_ref_gen allow_product env sigma ref t =
with Not_found ->
try
let t' = nf_betaiota sigma (one_step_reduce env sigma t) in
- elimrec env t' l
+ elimrec env t' l
with NotStepReducible ->
errorlabstrm ""
(str "Cannot recognize a statement based on " ++
diff --git a/pretyping/tacred.mli b/pretyping/tacred.mli
index 55c305cc..8fd14dcc 100644
--- a/pretyping/tacred.mli
+++ b/pretyping/tacred.mli
@@ -1,33 +1,30 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: tacred.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
open Names
open Term
open Environ
open Evd
open Reductionops
open Closure
-open Rawterm
+open Glob_term
open Termops
open Pattern
-(*i*)
+open Libnames
type reduction_tactic_error =
InvalidAbstraction of env * constr * (env * Type_errors.type_error)
exception ReductionTacticError of reduction_tactic_error
-(*s Reduction functions associated to tactics. \label{tacred} *)
+(** {6 Reduction functions associated to tactics. {% \label{%}tacred{% }%} } *)
-(* Evaluable global reference *)
+(** Evaluable global reference *)
val is_evaluable : Environ.env -> evaluable_global_reference -> bool
@@ -41,57 +38,69 @@ val global_of_evaluable_reference :
exception Redelimination
-(* Red (raise user error if nothing reducible) *)
+(** Red (raise user error if nothing reducible) *)
val red_product : reduction_function
-(* Red (raise Redelimination if nothing reducible) *)
+(** Red (raise Redelimination if nothing reducible) *)
val try_red_product : reduction_function
-(* Simpl *)
+(** 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
-(* Simpl only at the head *)
+(** Simpl only at the head *)
val whd_simpl : reduction_function
-(* Hnf: like whd_simpl but force delta-reduction of constants that do
+(** Hnf: like whd_simpl but force delta-reduction of constants that do
not immediately hide a non reducible fix or cofix *)
val hnf_constr : reduction_function
-(* Unfold *)
+(** Unfold *)
val unfoldn :
(occurrences * evaluable_global_reference) list -> reduction_function
-(* Fold *)
+(** Fold *)
val fold_commands : constr list -> reduction_function
-(* Pattern *)
+(** Pattern *)
val pattern_occs : (occurrences * constr) list -> reduction_function
-(* Rem: Lazy strategies are defined in Reduction *)
-(* Call by value strategy (uses Closures) *)
+(** Rem: Lazy strategies are defined in Reduction *)
+
+(** Call by value strategy (uses Closures) *)
val cbv_norm_flags : Closure.RedFlags.reds -> reduction_function
val cbv_beta : local_reduction_function
val cbv_betaiota : local_reduction_function
val cbv_betadeltaiota : reduction_function
- val compute : reduction_function (* = [cbv_betadeltaiota] *)
+ val compute : reduction_function (** = [cbv_betadeltaiota] *)
-(* [reduce_to_atomic_ind env sigma t] puts [t] in the form [t'=(I args)]
+(** [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
-(* [reduce_to_quantified_ind env sigma t] puts [t] in the form
+(** [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
-(* [reduce_to_quantified_ref env sigma ref t] try to put [t] in the form
+(** [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 *)
val reduce_to_quantified_ref :
- env -> evar_map -> Libnames.global_reference -> types -> types
+ env -> evar_map -> global_reference -> types -> types
val reduce_to_atomic_ref :
- env -> evar_map -> Libnames.global_reference -> types -> types
+ env -> evar_map -> global_reference -> types -> types
+
+val find_hnf_rectype :
+ env -> evar_map -> types -> inductive * constr list
val contextually : bool -> occurrences * constr_pattern ->
(patvar_map -> reduction_function) -> reduction_function
diff --git a/pretyping/term_dnet.ml b/pretyping/term_dnet.ml
index 9dda5143..98091087 100644
--- a/pretyping/term_dnet.ml
+++ b/pretyping/term_dnet.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id:$ *)
-
(*i*)
open Util
open Term
diff --git a/pretyping/term_dnet.mli b/pretyping/term_dnet.mli
index 9c4c9dbc..a2c535d1 100644
--- a/pretyping/term_dnet.mli
+++ b/pretyping/term_dnet.mli
@@ -1,21 +1,17 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id:$ *)
-
-(*i*)
open Term
open Sign
open Libnames
open Mod_subst
-(*i*)
-(* Dnets on constr terms.
+(** Dnets on constr terms.
An instantiation of Dnet on (an approximation of) constr. It
associates a term (possibly with Evar) with an
@@ -33,25 +29,25 @@ open Mod_subst
See lib/dnet.mli for more details.
*)
-(* Identifiers to store (right hand side of the association) *)
+(** 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 *)
+ (** how to substitute them for storage *)
val subst : substitution -> t -> t
- (* how to recover the term from the identifier *)
+ (** how to recover the term from the identifier *)
val constr_of : t -> constr
end
-(* Options : *)
+(** Options : *)
module type OPT = sig
- (* pre-treatment to terms before adding or searching *)
+ (** pre-treatment to terms before adding or searching *)
val reduce : constr -> constr
- (* direction of post-filtering w.r.t sort subtyping :
+ (** direction of post-filtering w.r.t sort subtyping :
- true means query <= terms in the structure
- false means terms <= query
*)
@@ -63,17 +59,17 @@ sig
type t
type ident
- (* results of filtering : identifier, a context (term with Evar
+ (** 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
+ (** [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 *)
+ (** merge of dnets. Faster than re-adding all terms *)
val union : t -> t -> t
val subst : substitution -> t -> t
@@ -82,25 +78,25 @@ sig
* High-level primitives describing specific search problems
*)
- (* [search_pattern dn c] returns all terms/patterns in dn
+ (** [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
+ (** [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
+ (** [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
+ (** [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 *)
+ (** [find_all dn] returns all idents contained in dn *)
val find_all : t -> ident list
val map : (ident -> ident) -> t -> t
diff --git a/pretyping/termops.ml b/pretyping/termops.ml
index 70f2279c..6371fd3a 100644
--- a/pretyping/termops.ml
+++ b/pretyping/termops.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: termops.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
open Pp
open Util
open Names
@@ -34,7 +32,6 @@ let pr_name = function
| Name id -> pr_id id
| Anonymous -> str "_"
-let pr_path sp = str(string_of_kn sp)
let pr_con sp = str(string_of_con sp)
let rec pr_constr c = match kind_of_term c with
@@ -149,11 +146,13 @@ let print_env env =
let set_module m = current_module := m*)
-let new_univ =
+let new_univ_level =
let univ_gen = ref 0 in
(fun sp ->
incr univ_gen;
- Univ.make_univ (Lib.library_dp(),!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 ())
@@ -238,18 +237,18 @@ let mkProd_wo_LetIn (na,body,t) c =
| None -> mkProd (na, t, c)
| Some b -> subst1 b c
-let it_mkProd ~init = List.fold_left (fun c (n,t) -> mkProd (n, t, c)) init
-let it_mkLambda ~init = List.fold_left (fun c (n,t) -> mkLambda (n, t, c)) init
+let it_mkProd init = List.fold_left (fun c (n,t) -> mkProd (n, t, c)) init
+let it_mkLambda init = List.fold_left (fun c (n,t) -> mkLambda (n, t, c)) init
let it_named_context_quantifier f ~init =
List.fold_left (fun c d -> f d c) init
-let it_mkProd_or_LetIn = it_named_context_quantifier mkProd_or_LetIn
-let it_mkProd_wo_LetIn = it_named_context_quantifier mkProd_wo_LetIn
-let it_mkLambda_or_LetIn = it_named_context_quantifier mkLambda_or_LetIn
-let it_mkNamedProd_or_LetIn = it_named_context_quantifier mkNamedProd_or_LetIn
-let it_mkNamedProd_wo_LetIn = it_named_context_quantifier mkNamedProd_wo_LetIn
-let it_mkNamedLambda_or_LetIn = it_named_context_quantifier mkNamedLambda_or_LetIn
+let it_mkProd_or_LetIn init = it_named_context_quantifier mkProd_or_LetIn ~init
+let it_mkProd_wo_LetIn init = it_named_context_quantifier mkProd_wo_LetIn ~init
+let it_mkLambda_or_LetIn init = it_named_context_quantifier mkLambda_or_LetIn ~init
+let it_mkNamedProd_or_LetIn init = it_named_context_quantifier mkNamedProd_or_LetIn ~init
+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
(* *)
@@ -265,11 +264,45 @@ let rec strip_head_cast c = match kind_of_term c with
| Cast (c,_,_) -> strip_head_cast c
| _ -> c
+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) ->
+ drop_extra_implicit_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"
+(* Get the last arg of an application *)
+let decompose_app_vect c =
+ match kind_of_term c with
+ | App (f,cl) -> (f, cl)
+ | _ -> (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)
+ else if len1 < len2 then
+ 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
+ (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)
+ else if len1 < len2 then
+ 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
+ (appvect (f1,extras), restl1, f2, l2)
+
(* [map_constr_with_named_binders g f l c] maps [f l] on the immediate
subterms of [c]; it carries an extra data [l] (typically a name
list) which is processed by [g na] (which typically cons [na] to
@@ -523,6 +556,14 @@ let collect_metas c =
in
List.rev (collrec [] c)
+(* collects all vars; warning: this is only visible vars, not dependencies in
+ 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
+ | _ -> fold_constr aux vars c in
+ aux Idset.empty c
+
(* Tests whether [m] is a subterm of [t]:
[m] is appropriately lifted through abstractions of [t] *)
@@ -546,6 +587,25 @@ let dependent_main noevar m t =
let dependent = dependent_main false
let dependent_no_evar = dependent_main true
+let count_occurrences m t =
+ let n = ref 0 in
+ let rec countrec m t =
+ if eq_constr m t then
+ incr n
+ else
+ match kind_of_term m, kind_of_term t with
+ | App (fm,lm), App (ft,lt) when Array.length lm < Array.length lt ->
+ countrec m (mkApp (ft,Array.sub lt 0 (Array.length lm)));
+ Array.iter (countrec m)
+ (Array.sub lt
+ (Array.length lm) ((Array.length lt) - (Array.length lm)))
+ | _, Cast (c,_,_) when isMeta c -> ()
+ | _, Evar _ -> ()
+ | _ -> iter_constr_with_binders (lift 1) countrec m t
+ in
+ countrec m t;
+ !n
+
(* Synonymous *)
let occur_term = dependent
@@ -592,10 +652,9 @@ let my_prefix_application eq_fun (k,c) (by_c : constr) (t : constr) =
None
| _ -> None
-(* Recognizing occurrences of a given (closed) subterm in a term for Pattern :
- [subst_term c t] substitutes [(Rel 1)] for all occurrences of (closed)
- term [c] in a term [t] *)
-(*i Bizarre : si on cherche un sous terme clos, pourquoi le lifter ? i*)
+(* Recognizing occurrences of a given subterm in a term: [subst_term c t]
+ substitutes [(Rel 1)] for all occurrences of term [c] in a term [t];
+ works if [c] has rels *)
let subst_term_gen eq_fun c t =
let rec substrec (k,c as kc) t =
@@ -608,10 +667,11 @@ let subst_term_gen eq_fun c t =
in
substrec (1,c) t
-(* Recognizing occurrences of a given (closed) subterm in a term :
- [replace_term c1 c2 t] substitutes [c2] for all occurrences of (closed)
- term [c1] in a term [t] *)
-(*i Meme remarque : a priori [c] n'est pas forcement clos i*)
+let subst_term = subst_term_gen eq_constr
+
+(* Recognizing occurrences of a given subterm in a term :
+ [replace_term c1 c2 t] substitutes [c2] for all occurrences of
+ term [c1] in a term [t]; works if [c1] and [c2] have rels *)
let replace_term_gen eq_fun c by_c in_t =
let rec substrec (k,c as kc) t =
@@ -624,8 +684,6 @@ let replace_term_gen eq_fun c by_c in_t =
in
substrec (0,c) in_t
-let subst_term = subst_term_gen eq_constr
-
let replace_term = replace_term_gen eq_constr
(* Substitute only at a list of locations or excluding a list of
@@ -633,6 +691,11 @@ let replace_term = replace_term_gen eq_constr
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,[])
@@ -643,57 +706,127 @@ let error_invalid_occurrence l =
(str ("Invalid occurrence " ^ plural (List.length l) "number" ^": ") ++
prlist_with_sep spc int l ++ str ".")
-let subst_term_occ_gen (nowhere_except_in,locs) occ c t =
+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
- assert (List.for_all (fun x -> x >= 0) locs);
- let rec substrec (k,c as kc) t =
- if nowhere_except_in & !pos > maxocc then t
- else
- if eq_constr c t then
- let r =
- if nowhere_except_in then
- if List.mem !pos locs then (mkRel k) else t
- else
- if List.mem !pos locs then t else (mkRel k)
- in incr pos; r
- else
- map_constr_with_binders_left_to_right
- (fun d (k,c) -> (k+1,lift 1 c))
- substrec kc t
+ 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,c) t in
+ let t' = substrec 1 t in
(!pos, t')
-let subst_term_occ (nowhere_except_in,locs as plocs) c t =
- if locs = [] then if nowhere_except_in then t else subst_term c t
- else
- let (nbocc,t') = subst_term_occ_gen plocs 1 c t in
- let rest = List.filter (fun o -> o >= nbocc) locs in
- if rest <> [] then error_invalid_occurrence rest;
- t'
-
-type hyp_location_flag = (* To distinguish body and type of local defs *)
- | InHyp
- | InHypTypeOnly
- | InHypValueOnly
-
-let subst_term_occ_decl ((nowhere_except_in,locs as plocs),hloc) c (id,bodyopt,typ as d) =
- match bodyopt,hloc with
- | None, InHypValueOnly -> errorlabstrm "" (pr_id id ++ str " has no value")
- | None, _ -> (id,None,subst_term_occ plocs c typ)
- | Some body, InHypTypeOnly -> (id,Some body,subst_term_occ plocs c typ)
- | Some body, InHypValueOnly -> (id,Some (subst_term_occ plocs c body),typ)
+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 ->
- if locs = [] then
- if nowhere_except_in then d
- else (id,Some (subst_term c body),subst_term c typ)
- else
- let (nbocc,body') = subst_term_occ_gen plocs 1 c body in
- let (nbocc',t') = subst_term_occ_gen plocs nbocc c typ in
- let rest = List.filter (fun o -> o >= nbocc') locs in
- if rest <> [] then error_invalid_occurrence rest;
- (id,Some body',t')
+ 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 =
@@ -965,7 +1098,7 @@ let on_judgment_value f j = { j with uj_val = f j.uj_val }
let on_judgment_type f j = { j with uj_type = f j.uj_type }
(* Cut a context ctx in 2 parts (ctx1,ctx2) with ctx1 containing k
- variables *)
+ variables; skips let-in's *)
let context_chop k ctx =
let rec chop_aux acc = function
| (0, l2) -> (List.rev acc, l2)
@@ -974,3 +1107,28 @@ let context_chop k ctx =
| (_, []) -> anomaly "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
+ push_rel_context ctx2 (reset_with_named_context (named_context_val env) env),
+ ctx1
+
+(*******************************************)
+(* Functions to deal with impossible cases *)
+(*******************************************)
+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
+ fun () ->
+ match !impossible_default_case with
+ | Some (id,type_of_id) ->
+ make_judge id type_of_id
+ | 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)))
diff --git a/pretyping/termops.mli b/pretyping/termops.mli
index 91c76564..5448d97c 100644
--- a/pretyping/termops.mli
+++ b/pretyping/termops.mli
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: termops.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
open Util
open Pp
open Names
@@ -15,7 +13,8 @@ open Term
open Sign
open Environ
-(* Universes *)
+(** 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
@@ -23,10 +22,11 @@ val new_Type_sort : unit -> sorts
val refresh_universes : types -> types
val refresh_universes_strict : types -> types
-(* printers *)
+(** printers *)
val print_sort : sorts -> std_ppcmds
val pr_sort_family : sorts_family -> std_ppcmds
-(* debug printer: do not use to display terms to the casual user... *)
+
+(** debug printer: do not use to display terms to the casual user... *)
val set_print_constr : (env -> constr -> std_ppcmds) -> unit
val print_constr : constr -> std_ppcmds
val print_constr_env : env -> constr -> std_ppcmds
@@ -35,35 +35,34 @@ val pr_rel_decl : env -> rel_declaration -> std_ppcmds
val print_rel_context : env -> std_ppcmds
val print_env : env -> std_ppcmds
-(* about contexts *)
+(** 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 *)
+(** builds argument lists matching a block of binders or a context *)
val rel_vect : int -> int -> constr array
val rel_list : int -> int -> constr list
val extended_rel_list : int -> rel_context -> constr list
val extended_rel_vect : int -> rel_context -> constr array
-(* iterators/destructors on terms *)
+(** iterators/destructors on terms *)
val mkProd_or_LetIn : rel_declaration -> types -> types
val mkProd_wo_LetIn : rel_declaration -> types -> types
-val it_mkProd : init:types -> (name * types) list -> types
-val it_mkLambda : init:constr -> (name * types) list -> constr
-val it_mkProd_or_LetIn : init:types -> rel_context -> types
-val it_mkProd_wo_LetIn : init:types -> rel_context -> types
-val it_mkLambda_or_LetIn : init:constr -> rel_context -> constr
-val it_mkNamedProd_or_LetIn : init:types -> named_context -> types
-val it_mkNamedProd_wo_LetIn : init:types -> named_context -> types
-val it_mkNamedLambda_or_LetIn : init:constr -> named_context -> constr
+val it_mkProd : types -> (name * types) list -> types
+val it_mkLambda : constr -> (name * 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
+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
-(**********************************************************************)
-(* Generic iterators on constr *)
+(** {6 Generic iterators on constr} *)
val map_constr_with_named_binders :
(name -> 'a -> 'a) ->
@@ -76,7 +75,7 @@ val map_constr_with_full_binders :
(rel_declaration -> 'a -> 'a) ->
('a -> constr -> constr) -> 'a -> constr -> constr
-(* [fold_constr_with_binders g f n acc c] folds [f n] on the immediate
+(** [fold_constr_with_binders g f n acc c] folds [f n] on the immediate
subterms of [c] starting from [acc] and proceeding from left to
right according to the usual representation of the constructions as
[fold_constr] but it carries an extra data [n] (typically a lift
@@ -93,8 +92,9 @@ val iter_constr_with_full_binders :
(**********************************************************************)
val strip_head_cast : constr -> constr
+val drop_extra_implicit_args : constr -> constr
-(* occur checks *)
+(** occur checks *)
exception Occur
val occur_meta : types -> bool
val occur_existential : types -> bool
@@ -108,68 +108,96 @@ val occur_var_in_decl :
val free_rels : constr -> Intset.t
val dependent : constr -> constr -> bool
val dependent_no_evar : constr -> constr -> bool
+val count_occurrences : constr -> constr -> int
val collect_metas : constr -> int list
-val occur_term : constr -> constr -> bool (* Synonymous
- of dependent *)
-(* Substitution of metavariables *)
+val collect_vars : constr -> Idset.t (** for visible vars only *)
+val occur_term : constr -> constr -> bool (** Synonymous
+ of dependent
+ Substitution of metavariables *)
type meta_value_map = (metavariable * constr) list
val subst_meta : meta_value_map -> constr -> constr
-(* Type assignment for metavariables *)
+(** Type assignment for metavariables *)
type meta_type_map = (metavariable * types) list
-(* [pop c] lifts by -1 the positive indexes in [c] *)
+(** [pop c] lifts by -1 the positive indexes in [c] *)
val pop : constr -> constr
-(*s Substitution of an arbitrary large term. Uses equality modulo
+(** {6 ... } *)
+(** Substitution of an arbitrary large term. Uses equality modulo
reduction of let *)
-(* [subst_term_gen eq d c] replaces [Rel 1] by [d] in [c] using [eq]
+(** [subst_term_gen eq d c] replaces [Rel 1] by [d] in [c] using [eq]
as equality *)
val subst_term_gen :
(constr -> constr -> bool) -> constr -> constr -> constr
-(* [replace_term_gen eq d e c] replaces [d] by [e] in [c] using [eq]
+(** [replace_term_gen eq d e c] replaces [d] by [e] in [c] using [eq]
as equality *)
val replace_term_gen :
(constr -> constr -> bool) ->
constr -> constr -> constr -> constr
-(* [subst_term d c] replaces [Rel 1] by [d] in [c] *)
+(** [subst_term d c] replaces [Rel 1] by [d] in [c] *)
val subst_term : constr -> constr -> constr
-(* [replace_term d e c] replaces [d] by [e] in [c] *)
+(** [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 *)
+(** 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_term_occ_gen occl n c d] replaces occurrences of [c] at
+(** [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_term_occ_gen :
+val subst_closed_term_occ_gen :
occurrences -> int -> constr -> types -> int * types
-(* [subst_term_occ occl c d] replaces occurrences of [c] at
- positions [occl] by [Rel 1] in [d] (see also Note OCC) *)
-val subst_term_occ : occurrences -> constr -> constr -> constr
-
-(* [subst_term_occ_decl occl c decl] replaces occurrences of [c] at
- positions [occl] by [Rel 1] in [decl] *)
+(** [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 *)
+type hyp_location_flag = (** To distinguish body and type of local defs *)
| InHyp
| InHypTypeOnly
| InHypValueOnly
-val subst_term_occ_decl :
+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 *)
+(** 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
@@ -181,7 +209,7 @@ val eta_eq_constr : constr -> constr -> bool
exception CannotFilter
-(* Lightweight first-order filtering procedure. Unification
+(** Lightweight first-order filtering procedure. Unification
variables ar represented by (untyped) Evars.
[filtering c1 c2] returns the substitution n'th evar ->
(context,term), or raises [CannotFilter].
@@ -193,10 +221,18 @@ val filtering : rel_context -> Reduction.conv_pb -> constr -> constr -> subst
val decompose_prod_letin : constr -> int * rel_context * constr
val align_prod_letin : constr -> constr -> rel_context * constr
-(* Get the last arg of a constr intended to be an application *)
+(** Get the last arg of a constr intended to be an application *)
val last_arg : constr -> constr
-(* name contexts *)
+(** Force the decomposition of a term as an applicative one *)
+val decompose_app_vect : constr -> constr * constr array
+
+val adjust_app_list_size : constr -> constr list -> constr -> constr list ->
+ (constr * constr list * constr * constr list)
+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
@@ -207,18 +243,19 @@ val ids_of_named_context : named_context -> identifier list
val ids_of_context : env -> identifier list
val names_of_rel_context : env -> names_context
-val context_chop : int -> rel_context -> (rel_context*rel_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 *)
+(** Set of local names *)
val vars_of_env: env -> Idset.t
val add_vname : Idset.t -> name -> Idset.t
-(* other signature iterators *)
+(** other signature iterators *)
val process_rel_context : (rel_declaration -> env -> env) -> env -> env
val assums_of_rel_context : rel_context -> (name * 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 *)
+val smash_rel_context : rel_context -> rel_context (** expand lets in context *)
val adjust_subst_to_rel_context : rel_context -> constr list -> constr list
val map_rel_context_in_env :
(env -> constr -> constr) -> env -> rel_context -> rel_context
@@ -234,19 +271,23 @@ val clear_named_body : identifier -> env -> env
val global_vars : env -> constr -> identifier list
val global_vars_set_of_decl : env -> named_declaration -> Idset.t
-(* Gives an ordered list of hypotheses, closed by dependencies,
+(** Gives an ordered list of hypotheses, closed by dependencies,
containing a given set *)
val dependency_closure : env -> named_context -> Idset.t -> identifier list
-(* Test if an identifier is the basename of a global reference *)
+(** Test if an identifier is the basename of a global reference *)
val is_section_variable : identifier -> bool
val isGlobalRef : constr -> bool
val has_polymorphic_type : constant -> bool
-(* Combinators on judgments *)
+(** Combinators on judgments *)
val on_judgment : (types -> types) -> unsafe_judgment -> unsafe_judgment
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
diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml
index 9e64143d..e85f174e 100644
--- a/pretyping/typeclasses.ml
+++ b/pretyping/typeclasses.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: typeclasses.ml 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
(*i*)
open Names
open Libnames
@@ -24,21 +22,32 @@ open Libobject
(*i*)
-let add_instance_hint_ref = ref (fun id pri -> assert false)
+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
-let set_typeclass_transparency_ref = ref (fun id pri -> assert false)
+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_typeclass_transparency_ref = ref (fun id local c -> assert false)
let register_set_typeclass_transparency =
(:=) set_typeclass_transparency_ref
-let set_typeclass_transparency gr c = !set_typeclass_transparency_ref gr c
+let set_typeclass_transparency gr local c = !set_typeclass_transparency_ref gr local c
+
+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 solve_instanciation_problem = ref (fun _ _ _ -> assert false)
-let mismatched_params env n m = mismatched_ctx_inst env Parameters n m
-(* let mismatched_defs env n m = mismatched_ctx_inst env Definitions n m *)
-let mismatched_props env n m = mismatched_ctx_inst env Properties n m
+let resolve_one_typeclass env evm t =
+ !solve_instanciation_problem env evm t
type rels = constr list
+type direction = Forward | Backward
(* This module defines type-classes *)
type typeclass = {
@@ -52,8 +61,9 @@ type typeclass = {
cl_props : rel_context;
(* The method implementaions as projections. *)
- cl_projs : (identifier * constant option) list;
+ cl_projs : (name * (direction * int option) option * constant option) list;
}
+
module Gmap = Fmap.Make(RefOrdered)
type typeclasses = typeclass Gmap.t
@@ -65,7 +75,7 @@ type instance = {
-1 for discard, 0 for none, mutable to avoid redeclarations
when multiple rebuild_object happen. *)
is_global: int;
- is_impl: global_reference;
+ is_impl: global_reference;
}
type instances = (instance Gmap.t) Gmap.t
@@ -118,7 +128,11 @@ let dest_class_app env c =
let cl, args = decompose_app c in
global_class_of_constr env cl, args
-let class_of_constr c = try Some (fst (dest_class_app (Global.env ()) c)) with _ -> None
+let dest_class_arity env c =
+ let rels, c = Term.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 _ -> None
let rec is_class_type evd c =
match kind_of_term c with
@@ -148,7 +162,7 @@ let subst_class (subst,cl) =
let do_subst_context (grs,ctx) =
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) -> (x, Option.smartmap do_subst_con y)) 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;
@@ -180,7 +194,7 @@ 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
@@ -195,11 +209,15 @@ let discharge_class (_,cl) =
{ cl_impl = cl_impl';
cl_context = context;
cl_props = props;
- cl_projs = list_smartmap (fun (x, y) -> x, Option.smartmap Lib.discharge_con y) cl.cl_projs }
+ cl_projs = list_smartmap (fun (x, y, z) -> x, y, Option.smartmap Lib.discharge_con z) cl.cl_projs }
-let rebuild_class cl = cl
+let rebuild_class cl =
+ try
+ let cst = Tacred.evaluable_of_global_reference (Global.env ()) cl.cl_impl in
+ set_typeclass_transparency cst false false; cl
+ with _ -> cl
-let (class_input,class_output) =
+let class_input : typeclass -> obj =
declare_object
{ (default_object "type classes state") with
cache_function = cache_class;
@@ -213,55 +231,150 @@ let (class_input,class_output) =
let add_class cl =
Lib.add_anonymous_leaf (class_input cl)
+(** Build the subinstances hints. *)
+
+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)
+ with _ -> 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
+ 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
+ let projargs = Array.of_list (args @ [instapp]) in
+ 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
+ if check && check_instance env sigma body then None
+ else
+ let pri =
+ match pri, pri' with
+ | Some p, Some p' -> Some (p + p')
+ | Some p, None -> Some (p + 1)
+ | _, _ -> None
+ in
+ 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
+ in List.fold_left declare_proj [] projs
+ in aux pri (constr_of_global glob)
(*
* instances persistent object
*)
-let load_instance (_,inst) =
+type instance_action =
+ | AddInstance
+ | RemoveInstance
+
+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
-let cache_instance = load_instance
+let remove_instance inst =
+ let insts =
+ try Gmap.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 cache_instance (_, (action, i)) =
+ match action with
+ | AddInstance -> load_instance i
+ | RemoveInstance -> remove_instance i
-let subst_instance (subst,inst) =
+let subst_instance (subst, (action, inst)) = action,
{ inst with
is_class = fst (subst_global subst inst.is_class);
is_impl = fst (subst_global subst inst.is_impl) }
-let discharge_instance (_,inst) =
+let discharge_instance (_, (action, inst)) =
if inst.is_global <= 0 then None
- else Some
+ else Some (action,
{ inst with
is_global = pred inst.is_global;
is_class = Lib.discharge_global inst.is_class;
- is_impl = Lib.discharge_global inst.is_impl }
+ is_impl = Lib.discharge_global inst.is_impl })
-let rebuild_instance inst =
- add_instance_hint inst.is_impl inst.is_pri;
- inst
-let classify_instance inst =
- if inst.is_global = -1 then Dispose
- else Substitute inst
+let is_local i = 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)
+ (build_subclasses ~check:(check && not (isVarRef inst.is_impl))
+ (Global.env ()) Evd.empty inst.is_impl inst.is_pri)
-let (instance_input,instance_output) =
+let rebuild_instance (action, inst) =
+ if action = AddInstance then add_instance true inst;
+ (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
cache_function = cache_instance;
- load_function = (fun _ -> load_instance);
- open_function = (fun _ -> load_instance);
+ load_function = (fun _ x -> cache_instance x);
+ open_function = (fun _ x -> cache_instance x);
classify_function = classify_instance;
discharge_function = discharge_instance;
rebuild_function = rebuild_instance;
subst_function = subst_instance }
let add_instance i =
- Lib.add_anonymous_leaf (instance_input i);
- add_instance_hint i.is_impl i.is_pri
+ Lib.add_anonymous_leaf (instance_input (AddInstance, i));
+ add_instance true i
+
+let remove_instance i =
+ Lib.add_anonymous_leaf (instance_input (RemoveInstance, 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
+ match class_of_constr ty with
+ | Some (rels, (tc, args) as _cl) ->
+ add_instance (new_instance tc pri (not local) 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); *)
+(* Auto.add_hints local [typeclasses_db] *)
+(* (Auto.HintsCutEntry (PathSeq (PathStar (PathAtom PathAny), path))) *)
+ | None -> ()
+
+let add_class cl =
+ add_class cl;
+ List.iter (fun (n, inst, body) ->
+ match inst with
+ | Some (Backward, pri) ->
+ declare_instance pri false (ConstRef (Option.get body))
+ | _ -> ())
+ cl.cl_projs
+
open Declarations
@@ -275,7 +388,7 @@ let add_constant_class cst =
cl_projs = []
}
in add_class tc;
- set_typeclass_transparency (EvalConstRef cst) false
+ set_typeclass_transparency (EvalConstRef cst) false false
let add_inductive_class ind =
let mind, oneind = Global.lookup_inductive ind in
@@ -308,13 +421,6 @@ let instance_constructor cl args =
let typeclasses () = Gmap.fold (fun _ l c -> l :: c) !classes []
-let cmapl_add x y m =
- try
- let l = Gmap.find x m in
- Gmap.add x (Gmap.add y.is_impl y l) m
- with Not_found ->
- Gmap.add x (Gmap.add y.is_impl y Gmap.empty) m
-
let cmap_elements c = Gmap.fold (fun k v acc -> v :: acc) c []
let instances_of c =
@@ -350,41 +456,43 @@ let is_implicit_arg k =
| InternalHole -> true
| _ -> false
+
(* To embed a boolean for resolvability status.
This is essentially a hack to mark which evars correspond to
goals and do not need to be resolved when we have nested [resolve_all_evars]
calls (e.g. when doing apply in an External hint in typeclass_instances).
- Would be solved by having real evars-as-goals. *)
+ Would be solved by having real evars-as-goals.
-let ((bool_in : bool -> Dyn.t),
- (bool_out : Dyn.t -> bool)) = Dyn.create "bool"
+ Nota: we will only check the resolvability status of undefined evars.
+ *)
-let bool_false = bool_in false
+let resolvable = Store.field ()
+open Store.Field
let is_resolvable evi =
- match evi.evar_extra with
- Some t -> if Dyn.tag t = "bool" then bool_out t else true
- | None -> true
+ assert (evi.evar_body = Evar_empty);
+ Option.default true (resolvable.get evi.evar_extra)
+
+let mark_unresolvable_undef evi =
+ let t = resolvable.set false evi.evar_extra in
+ { evi with evar_extra = t }
let mark_unresolvable evi =
- { evi with evar_extra = Some bool_false }
+ assert (evi.evar_body = Evar_empty);
+ mark_unresolvable_undef evi
let mark_unresolvables sigma =
- Evd.fold (fun ev evi evs ->
- Evd.add evs ev (mark_unresolvable evi))
- sigma Evd.empty
+ Evd.fold_undefined (fun ev evi evs ->
+ Evd.add evs ev (mark_unresolvable_undef evi))
+ sigma (Evd.defined_evars sigma)
let has_typeclasses evd =
- Evd.fold (fun ev evi has -> has ||
- (evi.evar_body = Evar_empty && is_class_evar evd evi && is_resolvable evi))
+ Evd.fold_undefined (fun ev evi has -> has ||
+ (is_class_evar evd evi && is_resolvable evi))
evd false
let solve_instanciations_problem = ref (fun _ _ _ _ _ -> assert false)
-let solve_instanciation_problem = ref (fun _ _ _ -> assert false)
let resolve_typeclasses ?(onlyargs=false) ?(split=true) ?(fail=true) env evd =
if not (has_typeclasses evd) then evd
else !solve_instanciations_problem env evd onlyargs split fail
-
-let resolve_one_typeclass env evm t =
- !solve_instanciation_problem env evm t
diff --git a/pretyping/typeclasses.mli b/pretyping/typeclasses.mli
index 83ae84a5..74ccaf83 100644
--- a/pretyping/typeclasses.mli
+++ b/pretyping/typeclasses.mli
@@ -1,14 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: typeclasses.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
open Names
open Libnames
open Decl_kinds
@@ -20,25 +17,28 @@ open Nametab
open Mod_subst
open Topconstr
open Util
-(*i*)
-(* This module defines type-classes *)
+type direction = Forward | Backward
+
+(** This module defines type-classes *)
type typeclass = {
- (* The class implementation: a record parameterized by the context with defs in it or a definition if
+ (** The class implementation: a record parameterized by the context with defs in it or a definition if
the class is a singleton. This acts as the class' global identifier. *)
cl_impl : global_reference;
- (* Context in which the definitions are typed. Includes both typeclass parameters and superclasses.
+ (** Context in which the definitions are typed. Includes both typeclass parameters and superclasses.
The boolean indicates if the typeclass argument is a direct superclass and the global reference
gives a direct link to the class itself. *)
cl_context : (global_reference * bool) option list * rel_context;
- (* Context of definitions and properties on defs, will not be shared *)
+ (** Context of definitions and properties on defs, will not be shared *)
cl_props : rel_context;
- (* The methods implementations of the typeclass as projections. Some may be undefinable due to
- sorting restrictions. *)
- cl_projs : (identifier * constant option) list;
+ (** The methods implementations of the typeclass as projections.
+ 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;
}
type instance
@@ -55,15 +55,16 @@ val add_inductive_class : inductive -> unit
val new_instance : typeclass -> int option -> bool -> 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 *)
+val class_info : global_reference -> typeclass (** raises a UserError if not a class *)
-(* These raise a UserError if not a class. *)
+(** These raise a UserError if not a class. *)
val dest_class_app : env -> constr -> typeclass * constr list
-(* Just return None if not a class *)
-val class_of_constr : constr -> typeclass option
+(** Just return None if not a class *)
+val class_of_constr : constr -> (rel_context * (typeclass * constr list)) option
val instance_impl : instance -> global_reference
@@ -72,14 +73,13 @@ val is_instance : global_reference -> bool
val is_implicit_arg : hole_kind -> bool
-(* Returns the term and type for the given instance of the parameters and fields
+(** 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
-(* Use evar_extra for marking resolvable evars. *)
-val bool_in : bool -> Dyn.t
-val bool_out : Dyn.t -> bool
+(** Resolvability.
+ Only undefined evars could be marked or checked for resolvability. *)
val is_resolvable : evar_info -> bool
val mark_unresolvable : evar_info -> evar_info
@@ -90,11 +90,26 @@ val resolve_typeclasses : ?onlyargs:bool -> ?split:bool -> ?fail:bool ->
env -> evar_map -> evar_map
val resolve_one_typeclass : env -> evar_map -> types -> open_constr
-val register_set_typeclass_transparency : (evaluable_global_reference -> bool -> unit) -> unit
-val set_typeclass_transparency : evaluable_global_reference -> bool -> unit
+val register_set_typeclass_transparency : (evaluable_global_reference -> bool (*local?*) -> bool -> unit) -> unit
+val set_typeclass_transparency : evaluable_global_reference -> bool -> bool -> unit
-val register_add_instance_hint : (global_reference -> int option -> unit) -> unit
-val add_instance_hint : global_reference -> int option -> unit
+val register_classes_transparent_state : (unit -> transparent_state) -> unit
+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 remove_instance_hint : global_reference -> unit
val solve_instanciations_problem : (env -> evar_map -> bool -> bool -> bool -> evar_map) ref
val solve_instanciation_problem : (env -> evar_map -> types -> open_constr) ref
+
+val declare_instance : int option -> bool -> global_reference -> unit
+
+
+(** Build the subinstances hints for a given typeclass object.
+ check tells if we should check for existence of the
+ subinstances and add only the missing ones. *)
+
+val build_subclasses : check:bool -> env -> evar_map -> global_reference -> int option (* priority *) ->
+ (int option * constr) list
diff --git a/pretyping/typeclasses_errors.ml b/pretyping/typeclasses_errors.ml
index 62941a76..ab7bb9dc 100644
--- a/pretyping/typeclasses_errors.ml
+++ b/pretyping/typeclasses_errors.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: typeclasses_errors.ml 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
(*i*)
open Names
open Decl_kinds
@@ -18,6 +16,7 @@ open Environ
open Nametab
open Mod_subst
open Topconstr
+open Compat
open Util
open Libnames
(*i*)
@@ -47,7 +46,7 @@ let unsatisfiable_constraints env evd ev =
raise (TypeClassError (env, UnsatisfiableConstraints (evd, None)))
| Some ev ->
let loc, kind = Evd.evar_source ev evd in
- raise (Stdpp.Exc_located (loc, TypeClassError
+ 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))
@@ -55,5 +54,5 @@ let mismatched_ctx_inst env c n m = typeclass_error env (MismatchedContextInstan
let rec unsatisfiable_exception exn =
match exn with
| TypeClassError (_, UnsatisfiableConstraints _) -> true
- | Stdpp.Exc_located(_, e) -> unsatisfiable_exception e
+ | Loc.Exc_located(_, e) -> unsatisfiable_exception e
| _ -> false
diff --git a/pretyping/typeclasses_errors.mli b/pretyping/typeclasses_errors.mli
index a763a80b..79f0678f 100644
--- a/pretyping/typeclasses_errors.mli
+++ b/pretyping/typeclasses_errors.mli
@@ -1,14 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: typeclasses_errors.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
open Names
open Decl_kinds
open Term
@@ -20,16 +17,15 @@ open Mod_subst
open Topconstr
open Util
open Libnames
-(*i*)
type contexts = Parameters | Properties
type typeclass_error =
| NotAClass of constr
- | UnboundMethod of global_reference * identifier located (* Class name, method *)
+ | UnboundMethod of global_reference * identifier located (** Class name, method *)
| NoInstance of identifier located * constr list
| UnsatisfiableConstraints of evar_map * (existential_key * hole_kind) option
- | MismatchedContextInstance of contexts * constr_expr list * rel_context (* found, expected *)
+ | MismatchedContextInstance of contexts * constr_expr list * rel_context (** found, expected *)
exception TypeClassError of env * typeclass_error
diff --git a/pretyping/typing.ml b/pretyping/typing.ml
index 80db26a4..efcdff9d 100644
--- a/pretyping/typing.ml
+++ b/pretyping/typing.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: typing.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
open Util
open Names
open Term
@@ -19,6 +17,7 @@ open Inductive
open Inductiveops
open Typeops
open Evd
+open Arguments_renaming
let meta_type evd mv =
let ty =
@@ -35,6 +34,14 @@ let inductive_type_knowing_parameters env ind jl =
let paramstyp = Array.map (fun j -> j.uj_type) jl in
Inductive.type_of_inductive_knowing_parameters env mip 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
+ evdref := evd; { utj_val = j.uj_val; utj_type = s }
+ | _ -> error_not_type env j
+
let e_judge_of_apply env evdref funj argjv =
let rec apply_rec n typ = function
| [] ->
@@ -47,29 +54,82 @@ let e_judge_of_apply env evdref funj argjv =
apply_rec (n+1) (subst1 hj.uj_val c2) restjl
else
error_cant_apply_bad_type env (n,c1, hj.uj_type) funj argjv
+ | Evar ev ->
+ let (evd',t) = Evarutil.define_evar_as_product !evdref ev in
+ evdref := evd';
+ let (_,_,c2) = destProd t in
+ apply_rec (n+1) (subst1 hj.uj_val c2) restjl
| _ ->
error_cant_apply_not_functional env funj argjv
in
apply_rec 1 funj.uj_type (Array.to_list argjv)
-let check_branch_types env evdref cj (lfj,explft) =
+let e_check_branch_types env evdref ind cj (lfj,explft) =
if 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 i lfj.(i).uj_type explft.(i)
+ error_ill_formed_branch env cj.uj_val (ind,i+1) 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 e_is_correct_arity env evdref c pj ind specif params =
+ let arsign = make_arity_signature env true (make_ind_family (ind,params)) in
+ let allowed_sorts = elim_sorts specif in
+ let error () = error_elim_arity env ind allowed_sorts c pj None in
+ let rec srec env pt ar =
+ let pt' = whd_betadeltaiota env !evdref pt in
+ match kind_of_term pt', ar with
+ | Prod (na1,a1,t), (_,None,a1')::ar' ->
+ 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 ()
+ | Evar (ev,_), [] ->
+ let s = Termops.new_sort_in_family (max_sort allowed_sorts) in
+ evdref := Evd.define ev (mkSort s) !evdref
+ | _, (_,Some _,_ as d)::ar' ->
+ srec (push_rel d env) (lift 1 pt') ar'
+ | _ ->
+ error ()
+ in
+ 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 nparams = inductive_params specif 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 ty =
+ whd_betaiota !evdref (Reduction.betazeta_appvect (n+1) p (Array.of_list (realargs@[c]))) in
+ (lc, ty, univ)
+
let e_judge_of_case env evdref ci pj cj lfj =
let indspec =
try find_mrectype env !evdref cj.uj_type
with Not_found -> error_case_not_inductive env cj in
let _ = check_case_info env (fst indspec) ci in
- let (bty,rslty,univ) = type_case_branches env indspec pj cj.uj_val in
- check_branch_types env evdref cj (lfj,bty);
+ let (bty,rslty,univ) = e_type_case_branches env evdref indspec pj cj.uj_val in
+ e_check_branch_types env evdref (fst indspec) cj (lfj,bty);
{ uj_val = mkCase (ci, pj.uj_val, cj.uj_val, Array.map j_val lfj);
uj_type = rslty }
+let check_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 sorts = elim_sorts specif in
+ if not (List.exists ((=) ksort) sorts) then
+ let s = inductive_sort_family (snd specif) in
+ error_elim_arity env ind sorts c pj
+ (Some(ksort,s,error_elim_explain ksort s))
+
let e_judge_of_cast env evdref cj k tj =
let expected_type = tj.utj_val in
if not (Evarconv.e_cumul env evdref cj.uj_type expected_type) then
@@ -100,13 +160,13 @@ let rec execute env evdref cstr =
judge_of_variable env id
| Const c ->
- make_judge cstr (type_of_constant env c)
+ make_judge cstr (rename_type_of_constant env c)
| Ind ind ->
- make_judge cstr (type_of_inductive env ind)
+ make_judge cstr (rename_type_of_inductive env ind)
| Construct cstruct ->
- make_judge cstr (type_of_constructor env cstruct)
+ make_judge cstr (rename_type_of_constructor env cstruct)
| Case (ci,p,c,lf) ->
let cj = execute env evdref c in
@@ -153,23 +213,23 @@ let rec execute env evdref cstr =
| Lambda (name,c1,c2) ->
let j = execute env evdref c1 in
- let var = type_judgment env j in
+ let var = e_type_judgment env evdref j in
let env1 = push_rel (name,None,var.utj_val) env in
let j' = execute env1 evdref c2 in
judge_of_abstraction env1 name var j'
| Prod (name,c1,c2) ->
let j = execute env evdref c1 in
- let varj = type_judgment env j in
+ let varj = e_type_judgment env evdref j in
let env1 = push_rel (name,None,varj.utj_val) env in
let j' = execute env1 evdref c2 in
- let varj' = type_judgment env1 j' in
+ let varj' = e_type_judgment env1 evdref j' in
judge_of_product env name varj varj'
| LetIn (name,c1,c2,c3) ->
let j1 = execute env evdref c1 in
let j2 = execute env evdref c2 in
- let j2 = type_judgment env j2 in
+ let j2 = e_type_judgment env evdref j2 in
let _ = judge_of_cast env 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
@@ -178,7 +238,7 @@ let rec execute env evdref cstr =
| Cast (c,k,t) ->
let cj = execute env evdref c in
let tj = execute env evdref t in
- let tj = type_judgment env tj in
+ let tj = e_type_judgment env evdref tj in
e_judge_of_cast env evdref cj k tj
and execute_recdef env evdref (names,lar,vdef) =
@@ -192,8 +252,6 @@ and execute_recdef env evdref (names,lar,vdef) =
and execute_array env evdref = Array.map (execute env evdref)
-and execute_list env evdref = List.map (execute env evdref)
-
let check env evd c t =
let evdref = ref evd in
let j = execute env evdref c in
@@ -211,15 +269,23 @@ let type_of env evd c =
(* Sort of a type *)
let sort_of env evd c =
- let j = execute env (ref evd) c in
- let a = type_judgment env j in
+ let evdref = ref evd in
+ 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 evdref = ref evd in
+ let j = execute env evdref c in
+ (* side-effect on evdref *)
+ !evdref, Termops.refresh_universes j.uj_type
+
let solve_evars env evd c =
let evdref = ref evd in
let c = (execute env evdref c).uj_val in
(* side-effect on evdref *)
- nf_evar !evdref c
-
+ !evdref, nf_evar !evdref c
+
+let _ = Evarconv.set_solve_evars solve_evars
diff --git a/pretyping/typing.mli b/pretyping/typing.mli
index 07cb7d59..224fd995 100644
--- a/pretyping/typing.mli
+++ b/pretyping/typing.mli
@@ -1,31 +1,38 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: typing.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
+open Names
open Term
open Environ
open Evd
-(*i*)
-(* This module provides the typing machine with existential variables
+(** This module provides the typing machine with existential variables
(but without universes). *)
-(* Typecheck a term and return its type *)
+(** Typecheck a term and return its type *)
val type_of : env -> evar_map -> constr -> types
-(* Typecheck a type and return its sort *)
+
+(** Typecheck a term and return its type + updated evars *)
+val e_type_of : env -> evar_map -> constr -> evar_map * types
+
+(** Typecheck a type and return its sort *)
val sort_of : env -> evar_map -> types -> sorts
-(* Typecheck a term has a given type (assuming the type is OK) *)
+
+(** Typecheck a term has a given type (assuming the type is OK) *)
val check : env -> evar_map -> constr -> types -> unit
-(* Returns the instantiated type of a metavariable *)
+(** 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 -> constr
+(** Solve existential variables using typing *)
+val solve_evars : env -> evar_map -> constr -> evar_map * 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 ->
+ unit
diff --git a/pretyping/unification.ml b/pretyping/unification.ml
index 7aa2272d..e6fa6eec 100644
--- a/pretyping/unification.ml
+++ b/pretyping/unification.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: unification.ml 14641 2011-11-06 11:59:10Z herbelin $ *)
-
open Pp
open Util
open Names
@@ -20,8 +18,7 @@ open Environ
open Evd
open Reduction
open Reductionops
-open Rawterm
-open Pattern
+open Glob_term
open Evarutil
open Pretype_errors
open Retyping
@@ -51,7 +48,7 @@ let abstract_scheme env c l lname_typ =
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_term_occ locc a t))
+ else mkLambda_name env (na,ta,subst_closed_term_occ locc a t))
c
(List.rev l)
lname_typ
@@ -66,6 +63,19 @@ let abstract_list_all env evd typ c l =
with UserError _ | Type_errors.TypeError _ ->
error_cannot_find_well_typed_abstraction env evd p l
+let set_occurrences_of_last_arg args =
+ Some all_occurrences :: 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' = evar_absorb_arguments env evd (destEvar ev) l in
+ let argoccs = set_occurrences_of_last_arg (snd ev') 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."
+
(**)
(* A refinement of [conv_pb]: the integers tells how many arguments
@@ -73,23 +83,16 @@ let abstract_list_all env evd typ c l =
is non zero, steps of eta-expansion will be allowed
*)
-type conv_pb_up_to_eta = Cumul | ConvUnderApp of int * int
-
-let topconv = ConvUnderApp (0,0)
-let of_conv_pb = function CONV -> topconv | CUMUL -> Cumul
-let conv_pb_of = function ConvUnderApp _ -> CONV | Cumul -> CUMUL
-let prod_pb = function ConvUnderApp _ -> topconv | pb -> pb
-
let opp_status = function
| IsSuperType -> IsSubType
| IsSubType -> IsSuperType
- | ConvUpToEta _ | UserGiven as x -> x
+ | Conv -> Conv
let add_type_status (x,y) = ((x,TypeNotProcessed),(y,TypeNotProcessed))
let extract_instance_status = function
- | Cumul -> add_type_status (IsSubType, IsSuperType)
- | ConvUnderApp (n1,n2) -> add_type_status (ConvUpToEta n1, ConvUpToEta n2)
+ | CUMUL -> add_type_status (IsSubType, IsSuperType)
+ | CONV -> add_type_status (Conv, Conv)
let rec assoc_pair x = function
[] -> raise Not_found
@@ -110,7 +113,7 @@ let pose_all_metas_as_evars env evd t =
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
- evdref := meta_assign mv (ev,(ConvUpToEta 0,TypeNotProcessed)) !evdref;
+ evdref := meta_assign mv (ev,(Conv,TypeNotProcessed)) !evdref;
ev)
| _ ->
map_constr aux t in
@@ -121,15 +124,15 @@ 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 c = solve_pattern_eqn env (Array.to_list l) c in
- let n = Array.length l - List.length (fst (decompose_lam c)) in
- let pb = (ConvUpToEta n,TypeNotProcessed) 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 (mkApp (f, l),c,c)
+ 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
- sigma,metasubst,(ev,solve_pattern_eqn env (Array.to_list l) c)::evarsubst
+ sigma,metasubst,(env,ev,solve_pattern_eqn env l c)::evarsubst
| _ -> assert false
let push d (env,n) = (push_rel_assum d env,n+1)
@@ -161,46 +164,172 @@ open Goptions
let _ =
declare_bool_option
{ optsync = true;
+ optdepr = false;
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 }
+let _ =
+ declare_bool_option
+ { optsync = true;
+ 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 }
+
type unify_flags = {
modulo_conv_on_closed_terms : Names.transparent_state option;
- use_metas_eagerly : bool;
+ (* What this flag controls was activated with all constants transparent, *)
+ (* even for auto, since Coq V5.10 *)
+
+ use_metas_eagerly_in_conv_on_closed_terms : bool;
+ (* This refinement of the conversion on closed terms is activable *)
+ (* (and activated for apply, rewrite but not auto since Feb 2008 for 8.2) *)
+
modulo_delta : Names.transparent_state;
+ (* This controls which constant are unfoldable; this is on for apply *)
+ (* (but not simple apply) since Feb 2008 for 8.2 *)
+
+ modulo_delta_types : Names.transparent_state;
+
+ 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;
- use_evars_pattern_unification : 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 *)
+
+ 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 was on for all tactics, including auto, since Sep 2006 for 8.1 *)
+
+ frozen_evars : ExistentialSet.t;
+ (* Evars of this set are considered axioms and never instantiated *)
+ (* Useful e.g. for autorewrite *)
+
+ restrict_conv_on_strict_subterms : bool;
+ (* No conversion at the root of the term; potentially useful for rewrite *)
+
+ modulo_betaiota : bool;
+ (* Support betaiota in the reduction *)
+ (* Note that zeta is always used *)
+
+ modulo_eta : bool;
+ (* Support eta in the reduction *)
+
+ allow_K_in_toplevel_higher_order_unification : bool
+ (* This is used only in second/higher order unification when looking for *)
+ (* subterms (rewrite and elim) *)
}
+(* Default flag for unifying a type against a type (e.g. apply) *)
+(* We set all conversion flags *)
let default_unify_flags = {
modulo_conv_on_closed_terms = Some full_transparent_state;
- use_metas_eagerly = true;
+ use_metas_eagerly_in_conv_on_closed_terms = true;
modulo_delta = full_transparent_state;
+ modulo_delta_types = full_transparent_state;
+ check_applied_meta_types = true;
resolve_evars = false;
- use_evars_pattern_unification = true;
+ use_pattern_unification = true;
+ use_meta_bound_pattern_unification = true;
+ frozen_evars = ExistentialSet.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 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 *)
+(* out of "simple apply" *)
let default_no_delta_unify_flags = {
modulo_conv_on_closed_terms = Some full_transparent_state;
- use_metas_eagerly = true;
+ use_metas_eagerly_in_conv_on_closed_terms = true;
+ modulo_delta = empty_transparent_state;
+ modulo_delta_types = full_transparent_state;
+ check_applied_meta_types = false;
+ resolve_evars = false;
+ use_pattern_unification = false;
+ use_meta_bound_pattern_unification = true;
+ frozen_evars = ExistentialSet.empty;
+ restrict_conv_on_strict_subterms = false;
+ modulo_betaiota = false;
+ modulo_eta = true;
+ allow_K_in_toplevel_higher_order_unification = 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 = {
+ modulo_conv_on_closed_terms = Some full_transparent_state;
+ use_metas_eagerly_in_conv_on_closed_terms = true;
+ modulo_delta = full_transparent_state;
+ modulo_delta_types = full_transparent_state;
+ check_applied_meta_types = true;
+ resolve_evars = false;
+ use_pattern_unification = true;
+ use_meta_bound_pattern_unification = true;
+ frozen_evars = ExistentialSet.empty;
+ restrict_conv_on_strict_subterms = false; (* ? *)
+ modulo_betaiota = false;
+ modulo_eta = true;
+ allow_K_in_toplevel_higher_order_unification = true
+}
+
+let elim_no_delta_flags = {
+ modulo_conv_on_closed_terms = Some full_transparent_state;
+ use_metas_eagerly_in_conv_on_closed_terms = true;
modulo_delta = empty_transparent_state;
+ modulo_delta_types = full_transparent_state;
+ check_applied_meta_types = false;
resolve_evars = false;
- use_evars_pattern_unification = false;
+ use_pattern_unification = false;
+ use_meta_bound_pattern_unification = true;
+ frozen_evars = ExistentialSet.empty;
+ restrict_conv_on_strict_subterms = false; (* ? *)
+ modulo_betaiota = false;
+ modulo_eta = true;
+ allow_K_in_toplevel_higher_order_unification = true
}
+let set_no_head_reduction flags =
+ { flags with restrict_conv_on_strict_subterms = true }
+
let use_evars_pattern_unification flags =
- !global_evars_pattern_unification_flag && flags.use_evars_pattern_unification
+ !global_evars_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
+ || (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
| None -> None
-let key_of flags f =
+let subterm_restriction is_subterm flags =
+ not is_subterm && flags.restrict_conv_on_strict_subterms
+
+let key_of 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) ->
@@ -219,20 +348,51 @@ let oracle_order env cf1 cf2 =
| Some k1 ->
match cf2 with
| None -> Some true
- | Some k2 -> Some (Conv_oracle.oracle_order k1 k2)
+ | Some k2 -> Some (Conv_oracle.oracle_order false k1 k2)
+
+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)
+
+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)
+ | _ -> false
+
+let check_compatibility env (sigma,metasubst,evarsubst) tyM tyN =
+ match subst_defined_metas metasubst tyM with
+ | None -> ()
+ | Some m ->
+ match subst_defined_metas metasubst tyN with
+ | None -> ()
+ | 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)
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 ((sigma,metasubst,evarsubst) as substn) curm curn =
+ let rec unirec_rec (curenv,nb as curenvnb) pb b wt ((sigma,metasubst,evarsubst) as substn) curm curn =
let cM = Evarutil.whd_head_evar sigma curm
and cN = Evarutil.whd_head_evar sigma curn in
match (kind_of_term cM,kind_of_term cN) with
| Meta k1, Meta k2 ->
+ if k1 = k2 then substn else
let stM,stN = extract_instance_status pb in
- if k2 < k1
- then sigma,(k1,cN,stN)::metasubst,evarsubst
- else if k1 = k2 then substn
+ 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);
+ if k2 < k1 then sigma,(k1,cN,stN)::metasubst,evarsubst
else sigma,(k2,cM,stM)::metasubst,evarsubst
| Meta k, _ when not (dependent cM cN) ->
+ if wt && flags.check_applied_meta_types then
+ (let tyM = Typing.meta_type sigma k in
+ let tyN = get_type_of curenv sigma cN in
+ check_compatibility curenv substn tyM tyN);
(* Here we check that [cN] does not contain any local variables *)
if nb = 0 then
sigma,(k,cN,snd (extract_instance_status pb))::metasubst,evarsubst
@@ -242,6 +402,10 @@ 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,cN)
| _, Meta k when not (dependent cN cM) ->
+ if wt && flags.check_applied_meta_types then
+ (let tyM = get_type_of curenv sigma cM in
+ let tyN = Typing.meta_type sigma k in
+ check_compatibility curenv substn tyM tyN);
(* Here we check that [cM] does not contain any local variables *)
if nb = 0 then
(sigma,(k,cM,fst (extract_instance_status pb))::metasubst,evarsubst)
@@ -250,65 +414,123 @@ let unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb flag
(sigma,(k,lift (-nb) cM,fst (extract_instance_status pb))::metasubst,
evarsubst)
else error_cannot_unify_local curenv sigma (m,n,cM)
- | Evar ev, _ -> sigma,metasubst,((ev,cN)::evarsubst)
- | _, Evar ev -> sigma,metasubst,((ev,cM)::evarsubst)
+ | Evar (evk,_ as ev), _
+ when not (ExistentialSet.mem evk flags.frozen_evars) ->
+ let cmvars = free_rels cM and cnvars = free_rels cN in
+ if Intset.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) ->
+ let cmvars = free_rels cM and cnvars = free_rels cN in
+ if Intset.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
+ in (sigma', metasubst, evarsubst)
+ with _ -> error_cannot_unify curenv sigma (m,n))
+
| Lambda (na,t1,c1), Lambda (_,t2,c2) ->
- unirec_rec (push (na,t1) curenvnb) topconv true
- (unirec_rec curenvnb topconv true substn t1 t2) c1 c2
+ unirec_rec (push (na,t1) curenvnb) CONV true wt
+ (unirec_rec curenvnb CONV true false substn t1 t2) c1 c2
| Prod (na,t1,c1), Prod (_,t2,c2) ->
- unirec_rec (push (na,t1) curenvnb) (prod_pb pb) true
- (unirec_rec curenvnb topconv true substn t1 t2) c1 c2
- | LetIn (_,a,_,c), _ -> unirec_rec curenvnb pb b substn (subst1 a c) cN
- | _, LetIn (_,a,_,c) -> unirec_rec curenvnb pb b substn cM (subst1 a c)
+ 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)
+
+ (* eta-expansion *)
+ | Lambda (na,t1,c1), _ when flags.modulo_eta ->
+ unirec_rec (push (na,t1) curenvnb) CONV true wt 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
+ (mkApp (lift 1 cM,[|mkRel 1|])) c2
| Case (_,p1,c1,cl1), Case (_,p2,c2,cl2) ->
- array_fold_left2 (unirec_rec curenvnb topconv true)
- (unirec_rec curenvnb topconv true
- (unirec_rec curenvnb topconv true substn p1 p2) c1 c2) cl1 cl2
-
- | App (f1,l1), _ when
- (isMeta f1 || use_evars_pattern_unification flags && isEvar f1) &
- is_unification_pattern curenvnb f1 l1 cN &
- not (dependent f1 cN) ->
- solve_pattern_eqn_array curenvnb f1 l1 cN substn
+ (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)
+ cl1 cl2
+ with ex when precatchable_exception ex ->
+ reduce curenvnb pb b wt 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)
| _, App (f2,l2) when
- (isMeta f2 || use_evars_pattern_unification flags && isEvar f2) &
- is_unification_pattern curenvnb f2 l2 cM &
- not (dependent f2 cM) ->
- solve_pattern_eqn_array curenvnb f2 l2 cM substn
+ (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)
| App (f1,l1), App (f2,l2) ->
- let len1 = Array.length l1
- and len2 = Array.length l2 in
- (try
- let (f1,l1,f2,l2) =
- if len1 = len2 then (f1,l1,f2,l2)
- else if len1 < len2 then
- 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
- (appvect (f1,extras), restl1, f2, l2) in
- let pb = ConvUnderApp (len1,len2) in
- array_fold_left2 (unirec_rec curenvnb topconv true)
- (unirec_rec curenvnb pb true substn f1 f2) l1 l2
- with ex when precatchable_exception ex ->
- try expand curenvnb pb b substn cM f1 l1 cN f2 l2
- with ex when precatchable_exception ex ->
- canonical_projections curenvnb pb b cM cN substn)
+ unify_app curenvnb pb b substn cM f1 l1 cN f2 l2
| _ ->
- try canonical_projections curenvnb pb b cM cN substn
- with ex when precatchable_exception ex ->
- if constr_cmp (conv_pb_of cv_pb) cM cN then substn else
- 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 substn cM f1 l1 cN f2 l2
+ unify_not_same_head curenvnb pb b wt substn cM cN
- and expand (curenv,_ as curenvnb) pb b (sigma,metasubst,_ as substn) cM f1 l1 cN f2 l2 =
+ and unify_app curenvnb pb b substn cM f1 l1 cN f2 l2 =
+ try
+ 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
+ with ex when precatchable_exception ex ->
+ try expand curenvnb pb b false substn cM f1 l1 cN f2 l2
+ 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
+ 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
+ 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 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
+ 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'
+ 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
(* Try full conversion on meta-free terms. *)
@@ -323,17 +545,19 @@ 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) &&
match flags.modulo_conv_on_closed_terms with
| None -> false
| Some convflags ->
- let subst = if flags.use_metas_eagerly then metasubst else ms in
+ 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
| Some m1 ->
match subst_defined_metas subst cN with
| None -> (* some undefined Metas in cN *) false
| Some n1 ->
- if is_trans_fconv (conv_pb_of pb) convflags env sigma m1 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)
@@ -341,55 +565,52 @@ let unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb flag
then
substn
else
- if b then
- (* Try delta-expansion if in subterms or if asked to conv at top *)
- let cf1 = key_of flags f1 and cf2 = key_of flags f2 in
+ let cf1 = key_of b flags f1 and cf2 = key_of b 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
| Some c ->
- unirec_rec curenvnb pb b substn
+ unirec_rec curenvnb pb b wt substn
(whd_betaiotazeta sigma (mkApp(c,l1))) cN
| None ->
(match expand_key curenv cf2 with
| Some c ->
- unirec_rec curenvnb pb b substn cM
+ unirec_rec curenvnb pb b wt substn cM
(whd_betaiotazeta sigma (mkApp(c,l2)))
| None ->
error_cannot_unify curenv sigma (cM,cN)))
| Some false ->
(match expand_key curenv cf2 with
| Some c ->
- unirec_rec curenvnb pb b substn cM
+ unirec_rec curenvnb pb b wt substn cM
(whd_betaiotazeta sigma (mkApp(c,l2)))
| None ->
(match expand_key curenv cf1 with
| Some c ->
- unirec_rec curenvnb pb b substn
+ unirec_rec curenvnb pb b wt substn
(whd_betaiotazeta sigma (mkApp(c,l1))) cN
| None ->
error_cannot_unify curenv sigma (cM,cN)))
- else
- error_cannot_unify curenv sigma (cM,cN)
and canonical_projections curenvnb pb b cM cN (sigma,_,_ as substn) =
let f1 () =
if isApp cM then
let f1l1 = decompose_app cM in
- if is_open_canonical_projection sigma f1l1 then
+ 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
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 then
+ if flags.modulo_conv_on_closed_terms = None ||
+ subterm_restriction b flags 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 sigma f2l2 then
+ 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
else error_cannot_unify (fst curenvnb) sigma (cM,cN)
@@ -413,20 +634,20 @@ let unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb flag
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 s u1 (substl ks u))
+ 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 s u1 (substl ks u))
+ 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) substn ts ts1 in
- unirec_rec curenvnb pb b substn c1 (applist (c,(List.rev ks)))
+ 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)))
in
- let evd = create_evar_defs sigma in
- if (if occur_meta_or_undefined_evar evd m || occur_meta_or_undefined_evar evd n then false
+ 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 flags ->
- is_trans_fconv (conv_pb_of cv_pb) flags env sigma m n
- | None -> constr_cmp (conv_pb_of cv_pb) m n) then true
+ | 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
@@ -434,74 +655,67 @@ let unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb flag
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 subst m n
+ else unirec_rec (env,0) cv_pb conv_at_top false subst m n
let unify_0 env sigma = unify_0_with_initial_metas (sigma,[],[]) true env
let left = true
let right = false
-let pop k = if k=0 then 0 else k-1
-
-let rec unify_with_eta keptside flags env sigma k1 k2 c1 c2 =
- (* Reason up to limited eta-expansion: ci is allowed to start with ki lam *)
- (* Question: try whd_betadeltaiota on ci if ki>0 ? *)
+let rec unify_with_eta keptside flags env sigma c1 c2 =
+(* Question: try whd_betadeltaiota on ci if not two lambdas? *)
match kind_of_term c1, kind_of_term c2 with
| (Lambda (na,t1,c1'), Lambda (_,t2,c2')) ->
- let env' = push_rel_assum (na,t1) env in
- let sigma,metas,evars = unify_0 env sigma topconv flags t1 t2 in
- let side,status,(sigma,metas',evars') =
- unify_with_eta keptside flags env' sigma (pop k1) (pop k2) c1' c2'
- in (side,status,(sigma,metas@metas',evars@evars'))
- | (Lambda (na,t,c1'),_) when k2 > 0 ->
- let env' = push_rel_assum (na,t) env in
- let side = left in (* expansion on the right: we keep the left side *)
- unify_with_eta side flags env' sigma (pop k1) (k2-1)
- c1' (mkApp (lift 1 c2,[|mkRel 1|]))
- | (_,Lambda (na,t,c2')) when k1 > 0 ->
- let env' = push_rel_assum (na,t) env in
- let side = right in (* expansion on the left: we keep the right side *)
- unify_with_eta side flags env' sigma (k1-1) (pop k2)
- (mkApp (lift 1 c1,[|mkRel 1|])) c2'
+ let env' = push_rel_assum (na,t1) env in
+ let sigma,metas,evars = unify_0 env sigma CONV flags t1 t2 in
+ let side,(sigma,metas',evars') =
+ unify_with_eta keptside flags env' sigma c1' c2'
+ in (side,(sigma,metas@metas',evars@evars'))
+ | (Lambda (na,t,c1'),_)->
+ let env' = push_rel_assum (na,t) env in
+ let side = left in (* expansion on the right: we keep the left side *)
+ unify_with_eta side flags env' sigma
+ c1' (mkApp (lift 1 c2,[|mkRel 1|]))
+ | (_,Lambda (na,t,c2')) ->
+ let env' = push_rel_assum (na,t) env in
+ let side = right in (* expansion on the left: we keep the right side *)
+ unify_with_eta side flags env' sigma
+ (mkApp (lift 1 c1,[|mkRel 1|])) c2'
| _ ->
- (keptside,ConvUpToEta(min k1 k2),
- unify_0 env sigma topconv flags c1 c2)
-
+ (keptside,unify_0 env sigma CONV flags c1 c2)
+
(* We solved problems [?n =_pb u] (i.e. [u =_(opp pb) ?n]) and [?n =_pb' u'],
we now compute the problem on [u =? u'] and decide which of u or u' is kept
Rem: the upper constraint is lost in case u <= ?n <= u' (and symmetrically
in the case u' <= ?n <= u)
*)
-
+
let merge_instances env sigma flags st1 st2 c1 c2 =
match (opp_status st1, st2) with
- | (UserGiven, ConvUpToEta n2) ->
- unify_with_eta left flags env sigma 0 n2 c1 c2
- | (ConvUpToEta n1, UserGiven) ->
- unify_with_eta right flags env sigma n1 0 c1 c2
- | (ConvUpToEta n1, ConvUpToEta n2) ->
+ | (Conv, Conv) ->
let side = left (* arbitrary choice, but agrees with compatibility *) in
- unify_with_eta side flags env sigma n1 n2 c1 c2
- | ((IsSubType | ConvUpToEta _ | UserGiven as oppst1),
- (IsSubType | ConvUpToEta _ | UserGiven)) ->
- let res = unify_0 env sigma Cumul flags c2 c1 in
+ let (side,res) = unify_with_eta side flags env sigma c1 c2 in
+ (side,Conv,res)
+ | ((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 or st1=UserGiven then (left, st1, res)
+ else if st2=IsSubType then (left, st1, res)
else (right, st2, res)
- | ((IsSuperType | ConvUpToEta _ | UserGiven as oppst1),
- (IsSuperType | ConvUpToEta _ | UserGiven)) ->
- let res = unify_0 env sigma Cumul flags c1 c2 in
+ | ((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 or st1=UserGiven then (left, st1, res)
+ else if st2=IsSuperType then (left, st1, res)
else (right, st2, res)
| (IsSuperType,IsSubType) ->
- (try (left, IsSubType, unify_0 env sigma Cumul flags c2 c1)
- with _ -> (right, IsSubType, unify_0 env sigma Cumul flags c1 c2))
+ (try (left, IsSubType, unify_0 env sigma CUMUL flags c2 c1)
+ with _ -> (right, IsSubType, unify_0 env sigma CUMUL flags c1 c2))
| (IsSubType,IsSuperType) ->
- (try (left, IsSuperType, unify_0 env sigma Cumul flags c1 c2)
- with _ -> (right, IsSuperType, unify_0 env sigma Cumul flags c2 c1))
-
+ (try (left, IsSuperType, unify_0 env sigma CUMUL flags c1 c2)
+ with _ -> (right, IsSuperType, unify_0 env sigma CUMUL flags c2 c1))
+
(* Unification
*
* Procedure:
@@ -557,18 +771,20 @@ let applyHead env evd n c =
(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
- apprec (n-1) (mkApp(c,[|evar|])) (subst1 evar c2) evd'
- | _ -> error "Apply_Head_Then"
+ | Prod (_,c1,c2) ->
+ let (evd',evar) =
+ Evarutil.new_evar evd env ~src:(dummy_loc,GoalEvar) c1 in
+ apprec (n-1) (mkApp(c,[|evar|])) (subst1 evar c2) evd'
+ | _ -> error "Apply_Head_Then"
in
- apprec n c (Typing.type_of env evd c) evd
-
-let is_mimick_head f =
+ apprec n c (Typing.type_of env evd c) evd
+
+let is_mimick_head ts f =
match kind_of_term f with
- (Const _|Var _|Rel _|Construct _|Ind _) -> true
- | _ -> false
+ | Const c -> 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
@@ -580,14 +796,14 @@ let try_to_coerce env evd c cty tycon =
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
- try try_to_coerce env evd c cty tycon
- with e when precatchable_exception e ->
+ 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
but there are cases where it though it was not rigid (like in
fst (nat,nat)) and stops while it could have seen that it is rigid *)
let cty = Tacred.hnf_constr env evd cty in
- try_to_coerce env evd c cty tycon
-
+ try_to_coerce env evd c cty tycon
+
let w_coerce env evd mv c =
let cty = get_type_of env evd c in
let mvty = Typing.meta_type evd mv in
@@ -596,25 +812,17 @@ let w_coerce env evd mv c =
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 t = Tacred.hnf_constr env sigma (nf_betaiota sigma (nf_meta sigma t)) in
- let u = Tacred.hnf_constr env sigma u in
- try
- if status = IsSuperType then
- unify_0 env sigma Cumul flags u t
- else if status = IsSubType then
- unify_0 env sigma Cumul flags t u
- else
- try unify_0 env sigma Cumul flags t u
- with e when precatchable_exception e ->
- unify_0 env sigma Cumul flags u t
- with e when precatchable_exception e ->
- (sigma,[],[])
+ let t = nf_betaiota sigma (nf_meta sigma t) in
+ unify_0 env sigma CUMUL flags t u
let unify_type env sigma flags mv status c =
let mvty = Typing.meta_type sigma mv in
- if occur_meta_or_existential mvty or is_arity env sigma mvty then
- unify_to_type env sigma flags c status mvty
- else (sigma,[],[])
+ 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}
+ c status mvty
(* Move metas that may need coercion at the end of the list of instances *)
@@ -628,8 +836,8 @@ let order_metas metas =
(* Solve an equation ?n[x1=u1..xn=un] = t where ?n is an evar *)
-let solve_simple_evar_eqn env evd ev rhs =
- let evd,b = solve_simple_eqn Evarconv.evar_conv_x env evd (None,ev,rhs) in
+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
@@ -642,28 +850,32 @@ let w_merge env with_types flags (evd,metas,evars) =
(* Process evars *)
match evars with
- | ((evn,_ as ev),rhs)::evars' ->
- if is_defined_evar evd ev then
+ | (curenv,(evk,_ as ev),rhs)::evars' ->
+ if Evd.is_defined evd evk then
let v = Evd.existential_value evd ev in
let (evd,metas',evars'') =
- unify_0 env evd topconv 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 *)
let rhs' = subst_meta_instances metas rhs in
match kind_of_term rhs with
| App (f,cl) when occur_meta rhs' ->
- if occur_evar evn rhs' then
- error_occur_check env evd evn rhs';
- if is_mimick_head f then
- let evd' = mimick_evar evd flags f (Array.length cl) evn in
- w_merge_rec evd' metas evars eqns
+ if occur_evar evk rhs' then
+ error_occur_check curenv evd evk rhs';
+ 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
else
- let evd', rhs'' = pose_all_metas_as_evars env evd rhs' in
- w_merge_rec (solve_simple_evar_eqn env evd' ev rhs'')
- 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
+
| _ ->
- w_merge_rec (solve_simple_evar_eqn env evd ev rhs')
- 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
end
| [] ->
@@ -679,7 +891,8 @@ let w_merge env with_types flags (evd,metas,evars) =
(* No coercion needed: delay the unification of types *)
((evd,c),([],[])),(mv,status,c)::eqns
else
- ((evd,c),([],[])),eqns in
+ ((evd,c),([],[])),eqns
+ in
if meta_defined evd mv then
let {rebus=c'},(status',_) = meta_fvalue evd mv in
let (take_left,st,(evd,metas',evars')) =
@@ -692,23 +905,30 @@ let w_merge env with_types flags (evd,metas,evars) =
w_merge_rec evd' (metas'@metas@metas'') (evars'@evars'') eqns
else
let evd' = meta_assign mv (c,(status,TypeProcessed)) evd in
- w_merge_rec evd' (metas@metas'') evars'' eqns
+ w_merge_rec evd' (metas''@metas) evars'' eqns
| [] ->
-
- (* Process type eqns *)
- match eqns with
- | (mv,status,c)::eqns ->
- let (evd,metas,evars) = unify_type env evd flags mv status c in
- w_merge_rec evd metas evars eqns
- | [] -> evd
-
- and mimick_evar evd flags hdc nargs sp =
- let ev = Evd.find evd sp in
+ (* Process type eqns *)
+ let rec process_eqns failures = function
+ | (mv,status,c)::eqns ->
+ (match (try Inl (unify_type env evd flags mv status c)
+ with e -> Inr e)
+ with
+ | Inr e -> process_eqns (((mv,status,c),e)::failures) eqns
+ | Inl (evd,metas,evars) ->
+ w_merge_rec evd metas evars (List.map fst failures @ eqns))
+ | [] ->
+ (match failures with
+ | [] -> evd
+ | ((mv,status,c),e)::_ -> raise e)
+ in process_eqns [] eqns
+
+ and mimick_undefined_evar evd flags hdc nargs sp =
+ let ev = Evd.find_undefined evd sp in
let sp_env = Global.env_of_context ev.evar_hyps in
let (evd', c) = applyHead sp_env evd nargs hdc in
let (evd'',mc,ec) =
- unify_0 sp_env evd' Cumul flags
- (Retyping.get_type_of sp_env evd' c) ev.evar_concl in
+ 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'''
@@ -733,33 +953,47 @@ let w_unify_meta_types env ?(flags=default_unify_flags) evd =
let check_types env flags (sigma,_,_ as subst) m n =
if isEvar_or_Meta (fst (whd_stack sigma m)) then
- unify_0_with_initial_metas subst true env Cumul
+ unify_0_with_initial_metas subst true env CUMUL
flags
- (Retyping.get_type_of env sigma n)
- (Retyping.get_type_of env sigma m)
+ (get_type_of env sigma n)
+ (get_type_of env sigma m)
else if isEvar_or_Meta (fst (whd_stack sigma n)) then
- unify_0_with_initial_metas subst true env Cumul
+ unify_0_with_initial_metas subst true env CUMUL
flags
- (Retyping.get_type_of env sigma m)
- (Retyping.get_type_of env sigma n)
+ (get_type_of env sigma m)
+ (get_type_of env sigma n)
else subst
-let w_unify_core_0 env with_types cv_pb flags m n evd =
+let try_resolve_typeclasses env evd flags m n =
+ if flags.resolve_evars then
+ try Typeclasses.resolve_typeclasses ~onlyargs:false ~split:false
+ ~fail:true env evd
+ with e when Typeclasses_errors.unsatisfiable_exception e ->
+ error_cannot_unify env evd (m, n)
+ 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 subst2 =
- unify_0_with_initial_metas (evd',ms,es) true env cv_pb flags m n
+ unify_0_with_initial_metas (evd',ms,es) false env cv_pb flags m n
in
let evd = w_merge env with_types flags subst2 in
- if flags.resolve_evars then
- try Typeclasses.resolve_typeclasses ~onlyargs:false ~split:false
- ~fail:true env evd
- with e when Typeclasses_errors.unsatisfiable_exception e ->
- error_cannot_unify env evd (m, n)
- else evd
+ try_resolve_typeclasses env evd flags m n
-let w_unify_0 env = w_unify_core_0 env false
-let w_typed_unify env = w_unify_core_0 env true
+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 (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))
(* 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
@@ -777,12 +1011,12 @@ let iter_fail f a =
(* 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 ?(flags=default_unify_flags) (op,cl) evd =
+let w_unify_to_subterm env evd ?(flags=default_unify_flags) (op,cl) =
let rec matchrec cl =
let cl = strip_outer_cast cl in
(try
if closed0 cl && not (isEvar cl)
- then w_typed_unify env topconv flags op cl evd,cl
+ then w_typed_unify env evd CONV flags op cl,cl
else error "Bound 1"
with ex when precatchable_exception ex ->
(match kind_of_term cl with
@@ -832,12 +1066,12 @@ let w_unify_to_subterm env ?(flags=default_unify_flags) (op,cl) evd =
in
try matchrec cl
with ex when precatchable_exception ex ->
- raise (PretypeError (env,NoOccurrenceFound (op, None)))
+ raise (PretypeError (env,evd,NoOccurrenceFound (op, None)))
(* 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 ?(flags=default_unify_flags) (op,cl) evd =
+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
@@ -862,7 +1096,7 @@ let w_unify_to_subterm_all env ?(flags=default_unify_flags) (op,cl) evd =
let cl = strip_outer_cast cl in
(bind
(if closed0 cl
- then return (fun () -> w_typed_unify env topconv flags op cl evd,cl)
+ then return (fun () -> w_typed_unify env evd CONV flags op cl,cl)
else fail "Bound 1")
(match kind_of_term cl with
| App (f,args) ->
@@ -894,61 +1128,65 @@ let w_unify_to_subterm_all env ?(flags=default_unify_flags) (op,cl) evd =
in
let res = matchrec cl [] in
if res = [] then
- raise (PretypeError (env,NoOccurrenceFound (op, None)))
+ raise (PretypeError (env,evd,NoOccurrenceFound (op, None)))
else
res
-let w_unify_to_subterm_list env flags allow_K hdmeta oplist t evd =
+let w_unify_to_subterm_list env evd flags hdmeta oplist t =
List.fold_right
(fun op (evd,l) ->
let op = whd_meta evd op in
if isMeta op then
- if allow_K then (evd,op::l)
+ 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
let (evd',cl) =
try
(* This is up to delta for subterms w/o metas ... *)
- w_unify_to_subterm env ~flags (strip_outer_cast op,t) evd
- with PretypeError (env,NoOccurrenceFound _) when allow_K -> (evd,op)
+ 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)
in
- if not allow_K && (* ensure we found a different instance *)
+ if not flags.allow_K_in_toplevel_higher_order_unification &&
+ (* ensure we found a different instance *)
List.exists (fun op -> eq_constr op cl) l
then error_non_linear_unification env evd hdmeta cl
else (evd',cl::l)
- else if allow_K or dependent op t then
+ 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,NoOccurrenceFound (op, None))))
+ raise (PretypeError (env,evd,NoOccurrenceFound (op, None))))
oplist
(evd,[])
-let secondOrderAbstraction env flags allow_K typ (p, 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 (evd',cllist) =
- w_unify_to_subterm_list env flags allow_K p oplist typ evd 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,(ConvUpToEta 0,TypeProcessed)],[])
+ w_merge env false flags (evd',[p,pred,(Conv,TypeProcessed)],[])
+
+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 secondOrderAbstractionAlgo dep =
+ if dep then secondOrderDependentAbstraction else secondOrderAbstraction
-let w_unify2 env flags allow_K cv_pb ty1 ty2 evd =
+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
match kind_of_term c1, kind_of_term c2 with
| Meta p1, _ ->
(* Find the predicate *)
- let evd' =
- secondOrderAbstraction env flags allow_K ty2 (p1,oplist1) evd in
- (* Resume first order unification *)
- w_unify_0 env cv_pb flags (nf_meta evd' ty1) ty2 evd'
+ secondOrderAbstractionAlgo dep env evd flags ty2 (p1,oplist1)
| _, Meta p2 ->
(* Find the predicate *)
- let evd' =
- secondOrderAbstraction env flags allow_K ty1 (p2, oplist2) evd in
- (* Resume first order unification *)
- w_unify_0 env cv_pb flags ty1 (nf_meta evd' ty2) evd'
+ secondOrderAbstractionAlgo dep env evd flags ty1 (p2, oplist2)
| _ -> error "w_unify2"
(* The unique unification algorithm works like this: If the pattern is
@@ -971,8 +1209,7 @@ let w_unify2 env flags allow_K cv_pb ty1 ty2 evd =
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 allow_K env cv_pb ?(flags=default_unify_flags) ty1 ty2 evd =
- let cv_pb = of_conv_pb cv_pb in
+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
@@ -980,22 +1217,27 @@ let w_unify allow_K env cv_pb ?(flags=default_unify_flags) ty1 ty2 evd =
| (Meta _, true, Lambda _, _ | Lambda _, _, Meta _, true)
when List.length l1 = List.length l2 ->
(try
- w_typed_unify env cv_pb flags ty1 ty2 evd
+ w_typed_unify_list env evd flags hd1 l1 hd2 l2
with ex when precatchable_exception ex ->
try
- w_unify2 env flags allow_K cv_pb ty1 ty2 evd
- with PretypeError (env,NoOccurrenceFound _) as e -> raise e)
+ w_unify2 env evd flags false cv_pb ty1 ty2
+ with PretypeError (env,_,NoOccurrenceFound _) as e -> raise e)
(* Second order case *)
| (Meta _, true, _, _ | _, _, Meta _, true) ->
(try
- w_unify2 env flags allow_K cv_pb ty1 ty2 evd
- with PretypeError (env,NoOccurrenceFound _) as e -> raise e
+ w_unify2 env evd flags false cv_pb ty1 ty2
+ with PretypeError (env,_,NoOccurrenceFound _) as e -> raise e
| ex when precatchable_exception ex ->
try
- w_typed_unify env cv_pb flags ty1 ty2 evd
+ w_typed_unify_list env evd flags hd1 l1 hd2 l2
with ex' when precatchable_exception ex' ->
- raise ex)
+ (* Last chance, use pattern-matching with typed
+ dependencies (done late for compatibility) *)
+ try
+ w_unify2 env evd flags true cv_pb ty1 ty2
+ with ex' when precatchable_exception ex' ->
+ raise ex)
(* General case: try first order *)
- | _ -> w_typed_unify env cv_pb flags ty1 ty2 evd
+ | _ -> w_typed_unify env evd cv_pb flags ty1 ty2
diff --git a/pretyping/unification.mli b/pretyping/unification.mli
index 2f48081a..e4bca4d3 100644
--- a/pretyping/unification.mli
+++ b/pretyping/unification.mli
@@ -1,54 +1,78 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: unification.mli 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(*i*)
open Term
open Environ
open Evd
-(*i*)
type unify_flags = {
modulo_conv_on_closed_terms : Names.transparent_state option;
- use_metas_eagerly : bool;
+ use_metas_eagerly_in_conv_on_closed_terms : bool;
modulo_delta : Names.transparent_state;
+ modulo_delta_types : Names.transparent_state;
+ check_applied_meta_types : bool;
resolve_evars : bool;
- use_evars_pattern_unification : bool
+ use_pattern_unification : bool;
+ use_meta_bound_pattern_unification : bool;
+ frozen_evars : ExistentialSet.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
-(* The "unique" unification fonction *)
+val elim_flags : unify_flags
+val elim_no_delta_flags : unify_flags
+
+(** The "unique" unification fonction *)
val w_unify :
- bool -> env -> conv_pb -> ?flags:unify_flags -> constr -> constr -> evar_map -> evar_map
+ 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 (c,t) m] 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 :
- env -> ?flags:unify_flags -> constr * constr -> evar_map -> evar_map * constr
+ env -> evar_map -> ?flags:unify_flags -> constr * constr -> evar_map * constr
val w_unify_to_subterm_all :
- env -> ?flags:unify_flags -> constr * constr -> evar_map -> (evar_map * constr) list
+ env -> evar_map -> ?flags:unify_flags -> constr * constr -> (evar_map * constr) list
val w_unify_meta_types : env -> ?flags:unify_flags -> evar_map -> evar_map
-(* [w_coerce_to_type env evd c ctyp typ] tries to coerce [c] of type
+(** [w_coerce_to_type env evd c ctyp typ] tries to coerce [c] of type
[ctyp] so that its gets type [typ]; [typ] may contain metavariables *)
val w_coerce_to_type : env -> evar_map -> constr -> types -> types ->
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 *)
-(* (exported for inv.ml) *)
+(** [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
+
+
+(* For tracing *)
+
+val w_merge : env -> bool -> 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 ->
+ Term.types ->
+ Term.types ->
+ Evd.evar_map * Evd.metabinding list *
+ (Environ.env * Term.types Term.pexistential * Term.constr) list
+
diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml
index 395f5c8d..fad2e6f0 100644
--- a/pretyping/vnorm.ml
+++ b/pretyping/vnorm.ml
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: vnorm.ml 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
open Names
open Declarations
open Term
diff --git a/pretyping/vnorm.mli b/pretyping/vnorm.mli
index 4cf99842..ee946e53 100644
--- a/pretyping/vnorm.mli
+++ b/pretyping/vnorm.mli
@@ -1,18 +1,16 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i*)
open Names
open Term
open Environ
open Reduction
-(*i*)
-(*s Reduction functions *)
+(** {6 Reduction functions } *)
val cbv_vm : env -> constr -> types -> constr