summaryrefslogtreecommitdiff
path: root/pretyping
diff options
context:
space:
mode:
authorGravatar Stephane Glondu <steph@glondu.net>2010-07-21 09:46:51 +0200
committerGravatar Stephane Glondu <steph@glondu.net>2010-07-21 09:46:51 +0200
commit5b7eafd0f00a16d78f99a27f5c7d5a0de77dc7e6 (patch)
tree631ad791a7685edafeb1fb2e8faeedc8379318ae /pretyping
parentda178a880e3ace820b41d38b191d3785b82991f5 (diff)
Imported Upstream snapshot 8.3~beta0+13298
Diffstat (limited to 'pretyping')
-rw-r--r--pretyping/cases.ml494
-rw-r--r--pretyping/cases.mli6
-rw-r--r--pretyping/cbv.ml170
-rw-r--r--pretyping/cbv.mli10
-rw-r--r--pretyping/classops.ml85
-rw-r--r--pretyping/classops.mli20
-rw-r--r--pretyping/clenv.ml103
-rw-r--r--pretyping/clenv.mli18
-rw-r--r--pretyping/coercion.ml135
-rw-r--r--pretyping/coercion.mli34
-rw-r--r--pretyping/detyping.ml230
-rw-r--r--pretyping/detyping.mli10
-rw-r--r--pretyping/evarconv.ml140
-rw-r--r--pretyping/evarconv.mli28
-rw-r--r--pretyping/evarutil.ml760
-rw-r--r--pretyping/evarutil.mli80
-rw-r--r--pretyping/evd.ml559
-rw-r--r--pretyping/evd.mli278
-rw-r--r--pretyping/indrec.ml365
-rw-r--r--pretyping/indrec.mli56
-rw-r--r--pretyping/inductiveops.ml67
-rw-r--r--pretyping/inductiveops.mli14
-rw-r--r--pretyping/matching.ml130
-rw-r--r--pretyping/matching.mli10
-rw-r--r--pretyping/namegen.ml312
-rw-r--r--pretyping/namegen.mli77
-rw-r--r--pretyping/pattern.ml120
-rw-r--r--pretyping/pattern.mli16
-rw-r--r--pretyping/pretype_errors.ml30
-rw-r--r--pretyping/pretype_errors.mli24
-rw-r--r--pretyping/pretyping.ml389
-rw-r--r--pretyping/pretyping.mli84
-rw-r--r--pretyping/pretyping.mllib29
-rw-r--r--pretyping/rawterm.ml156
-rw-r--r--pretyping/rawterm.mli35
-rw-r--r--pretyping/recordops.ml154
-rwxr-xr-xpretyping/recordops.mli27
-rw-r--r--pretyping/reductionops.ml332
-rw-r--r--pretyping/reductionops.mli32
-rw-r--r--pretyping/retyping.ml27
-rw-r--r--pretyping/retyping.mli11
-rw-r--r--pretyping/tacred.ml227
-rw-r--r--pretyping/tacred.mli32
-rw-r--r--pretyping/term_dnet.ml404
-rw-r--r--pretyping/term_dnet.mli112
-rw-r--r--pretyping/termops.ml521
-rw-r--r--pretyping/termops.mli116
-rw-r--r--pretyping/typeclasses.ml431
-rw-r--r--pretyping/typeclasses.mli39
-rw-r--r--pretyping/typeclasses_errors.ml32
-rw-r--r--pretyping/typeclasses_errors.mli18
-rw-r--r--pretyping/typing.ml194
-rw-r--r--pretyping/typing.mli19
-rw-r--r--pretyping/unification.ml675
-rw-r--r--pretyping/unification.mli23
-rw-r--r--pretyping/vnorm.ml98
56 files changed, 4965 insertions, 3633 deletions
diff --git a/pretyping/cases.ml b/pretyping/cases.ml
index abe4fcc1..d3813660 100644
--- a/pretyping/cases.ml
+++ b/pretyping/cases.ml
@@ -6,13 +6,14 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: cases.ml 13112 2010-06-10 19:58:23Z herbelin $ *)
+(* $Id$ *)
open Util
open Names
open Nameops
open Term
open Termops
+open Namegen
open Declarations
open Inductiveops
open Environ
@@ -73,11 +74,11 @@ 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 () ->
+ fun () ->
match !impossible_default_case with
| Some (id,type_of_id) ->
make_judge id type_of_id
- | None ->
+ | 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)))
@@ -87,8 +88,8 @@ let coq_unit_judge =
module type S = sig
val compile_cases :
loc -> case_style ->
- (type_constraint -> env -> evar_defs ref -> rawconstr -> unsafe_judgment) * evar_defs ref ->
- type_constraint ->
+ (type_constraint -> env -> evar_map ref -> rawconstr -> unsafe_judgment) * evar_map ref ->
+ type_constraint ->
env -> rawconstr option * tomatch_tuples * cases_clauses ->
unsafe_judgment
end
@@ -97,8 +98,8 @@ let rec list_try_compile f = function
| [a] -> f a
| [] -> anomaly "try_find_f"
| h::t ->
- try f h
- with UserError _ | TypeError _ | PretypeError _
+ try f h
+ with UserError _ | TypeError _ | PretypeError _
| Stdpp.Exc_located (_,(UserError _ | TypeError _ | PretypeError _)) ->
list_try_compile f t
@@ -119,14 +120,11 @@ let msg_may_need_inversion () =
(* Utils *)
let make_anonymous_patvars n =
- list_make n (PatVar (dummy_loc,Anonymous))
+ list_make n (PatVar (dummy_loc,Anonymous))
(* Environment management *)
let push_rels vars env = List.fold_right push_rel vars env
-let push_rel_defs =
- List.fold_right (fun (x,d,t) e -> push_rel (x,Some d,t) e)
-
(* 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'] *)
@@ -172,7 +170,7 @@ type 'a rhs =
it : 'a option}
type 'a equation =
- { patterns : cases_pattern list;
+ { patterns : cases_pattern list;
rhs : 'a rhs;
alias_stack : name list;
eqn_loc : loc;
@@ -210,14 +208,12 @@ and pattern_continuation =
let start_history n = Continuation (n, [], Top)
-let initial_history = function Continuation (_,[],Top) -> true | _ -> false
-
let feed_history arg = function
| Continuation (n, l, h) when n>=1 ->
Continuation (n-1, arg :: l, h)
| Continuation (n, _, _) ->
anomaly ("Bad number of expected remaining patterns: "^(string_of_int n))
- | Result _ ->
+ | Result _ ->
anomaly "Exhausted pattern history"
(* This is for non exhaustive error message *)
@@ -248,7 +244,7 @@ let rec simplify_history = function
let pat = match f with
| AliasConstructor pci ->
PatCstr (dummy_loc,pci,pargs,Anonymous)
- | AliasLeaf ->
+ | AliasLeaf ->
assert (l = []);
PatVar (dummy_loc, Anonymous) in
feed_history pat rh
@@ -266,7 +262,7 @@ let push_history_pattern n current cont =
where tomatch is some sequence of "instructions" (t1 ... tn)
- and mat is some matrix
+ and mat is some matrix
(p11 ... p1n -> rhs1)
( ... )
(pm1 ... pmn -> rhsm)
@@ -295,14 +291,14 @@ let push_history_pattern n current cont =
type 'a pattern_matching_problem =
{ env : env;
- evdref : evar_defs ref;
+ evdref : evar_map ref;
pred : constr;
tomatch : tomatch_stack;
history : pattern_continuation;
mat : 'a matrix;
caseloc : loc;
casestyle : case_style;
- typing_function: type_constraint -> env -> evar_defs ref -> 'a option -> unsafe_judgment }
+ typing_function: type_constraint -> env -> evar_map ref -> 'a option -> unsafe_judgment }
(*--------------------------------------------------------------------------*
* A few functions to infer the inductive type from the patterns instead of *
@@ -327,7 +323,7 @@ let rec find_row_ind = function
let inductive_template evdref env tmloc ind =
let arsign = get_full_arity_sign env ind in
- let hole_source = match tmloc with
+ let hole_source = match tmloc with
| Some loc -> fun i -> (loc, TomatchTypeParameter (ind,i))
| None -> fun _ -> (dummy_loc, InternalHole) in
let (_,evarl,_) =
@@ -337,7 +333,7 @@ let inductive_template evdref env tmloc ind =
| None ->
let ty' = substl subst ty in
let e = e_new_evar evdref env ~src:(hole_source n) ty' in
- (e::subst,e::evarl,n+1)
+ (e::subst,e::evarl,n+1)
| Some b ->
(b::subst,evarl,n+1))
arsign ([],[],1) in
@@ -354,7 +350,7 @@ let try_find_ind env sigma typ realnames =
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
+ (* 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 ()
@@ -363,23 +359,23 @@ let unify_tomatch_with_patterns evdref env loc typ pats realnames =
| None -> NotInd (None,typ)
| Some (_,(ind,_)) ->
inh_coerce_to_ind evdref env typ ind;
- try try_find_ind env (evars_of !evdref) typ realnames
+ try try_find_ind env !evdref typ realnames
with Not_found -> NotInd (None,typ)
let find_tomatch_tycon evdref env loc = function
(* Try if some 'in I ...' is present and can be used as a constraint *)
- | Some (_,ind,_,realnal) ->
+ | Some (_,ind,_,realnal) ->
mk_tycon (inductive_template evdref env loc ind),Some (List.rev realnal)
- | None ->
+ | None ->
empty_tycon,None
let coerce_row typing_fun evdref env pats (tomatch,(_,indopt)) =
let loc = Some (loc_of_rawconstr 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 (evars_of !evdref) j.uj_type in
+ let typ = nf_evar !evdref j.uj_type in
let t =
- try try_find_ind env (evars_of !evdref) typ realnames
+ try try_find_ind env !evdref typ realnames
with Not_found ->
unify_tomatch_with_patterns evdref env loc typ pats realnames in
(j.uj_val,t)
@@ -409,12 +405,12 @@ let adjust_tomatch_to_pattern pb ((current,typ),deps,dep) =
(* Ideally, we could find a common inductive type to which both the
term to match and the patterns coerce *)
(* In practice, we coerce the term to match if it is not already an
- inductive type and it is not dependent; moreover, we use only
+ inductive type and it is not dependent; moreover, we use only
the first pattern type and forget about the others *)
let typ,names =
match typ with IsInd(t,_,names) -> t,Some names | NotInd(_,t) -> t,None in
let typ =
- try try_find_ind pb.env (evars_of !(pb.evdref)) typ names
+ try try_find_ind pb.env !(pb.evdref) typ names
with Not_found -> NotInd (None,typ) in
let tomatch = ((current,typ),deps,dep) in
match typ with
@@ -432,7 +428,7 @@ let adjust_tomatch_to_pattern pb ((current,typ),deps,dep) =
else
(evd_comb2 (Coercion.inh_conv_coerce_to dummy_loc pb.env)
pb.evdref (make_judge current typ) (mk_tycon_type indt)).uj_val in
- let sigma = evars_of !(pb.evdref) in
+ let sigma = !(pb.evdref) in
let typ = try_find_ind pb.env sigma indt names in
((current,typ),deps,dep))
| _ -> tomatch
@@ -452,9 +448,6 @@ let map_tomatch_type f = function
let liftn_tomatch_type n depth = map_tomatch_type (liftn n depth)
let lift_tomatch_type n = liftn_tomatch_type n 1
-let lift_tomatch n ((current,typ),info) =
- ((lift n current,lift_tomatch_type n typ),info)
-
(**********************************************************************)
(* Utilities on patterns *)
@@ -467,12 +460,6 @@ let alias_of_pat = function
| PatVar (_,name) -> name
| PatCstr(_,_,_,name) -> name
-let unalias_pat = function
- | PatVar (c,name) as p ->
- if name = Anonymous then p else PatVar (c,Anonymous)
- | PatCstr(a,b,c,name) as p ->
- if name = Anonymous then p else PatCstr (a,b,c,Anonymous)
-
let remove_current_pattern eqn =
match eqn.patterns with
| pat::pats ->
@@ -497,18 +484,18 @@ let rec adjust_local_defs loc = function
| [], [] -> []
| _ -> raise NotAdjustable
-let check_and_adjust_constructor env ind cstrs = function
+let check_and_adjust_constructor env ind cstrs = function
| PatVar _ as pat -> pat
| PatCstr (loc,((_,i) as cstr),args,alias) as pat ->
(* Check it is constructor of the right type *)
let ind' = inductive_of_constructor cstr in
- if Closure.mind_equiv env ind' ind then
+ if eq_ind ind' ind then
(* Check the constructor has the right number of args *)
let ci = cstrs.(i-1) in
let nb_args_constr = ci.cs_nargs in
if List.length args = nb_args_constr then pat
else
- try
+ try
let args' = adjust_local_defs loc (args, List.rev ci.cs_args)
in PatCstr (loc, cstr, args', alias)
with NotAdjustable ->
@@ -518,7 +505,7 @@ let check_and_adjust_constructor env ind cstrs = function
(* Try to insert a coercion *)
try
Coercion.inh_pattern_coerce_to loc pat ind' ind
- with Not_found ->
+ with Not_found ->
error_bad_constructor_loc loc cstr ind
let check_all_variables typ mat =
@@ -530,14 +517,14 @@ let check_all_variables typ mat =
mat
let check_unused_pattern env eqn =
- if not !(eqn.used) then
+ if not !(eqn.used) then
raise_pattern_matching_error
(eqn.eqn_loc, env, UnusedClause eqn.patterns)
let set_used_pattern eqn = eqn.used := true
let extract_rhs pb =
- match pb.mat with
+ match pb.mat with
| [] -> errorlabstrm "build_leaf" (msg_may_need_inversion())
| eqn::_ ->
set_used_pattern eqn;
@@ -588,7 +575,7 @@ let dependencies_in_rhs nargs current tms eqns =
let rec find_dependency_list k n = function
| [] -> []
- | (used,tdeps,d)::rest ->
+ | (used,tdeps,d)::rest ->
let deps = find_dependency_list k (n+1) rest in
if used && dependent_decl (mkRel n) d
then list_add_set (List.length rest + 1) (list_union deps tdeps)
@@ -615,7 +602,7 @@ let find_dependencies_signature deps_in_rhs typs =
let regeneralize_index_tomatch n =
let rec genrec depth = function
- | [] ->
+ | [] ->
[]
| Pushed ((c,tm),l,dep) :: rest ->
let c = regeneralize_index n depth c in
@@ -629,7 +616,7 @@ let regeneralize_index_tomatch n =
:: genrec (depth+1) rest in
genrec 0
-let rec replace_term n c k t =
+let rec replace_term n c k t =
if t = mkRel (n+k) then lift k c
else map_constr_with_binders succ (replace_term n c) k t
@@ -652,9 +639,6 @@ let replace_tomatch n c =
:: replrec (depth+1) rest in
replrec 0
-let liftn_rel_declaration n k = map_rel_declaration (liftn n k)
-let substnl_rel_declaration sigma k = map_rel_declaration (substnl sigma k)
-
(* [liftn_tomatch_stack]: a term to match has just been substituted by
some constructor t = (ci x1...xn) and the terms x1 ... xn have been
added to match; all pushed terms to match must be lifted by n
@@ -690,7 +674,7 @@ let lift_tomatch_stack n = liftn_tomatch_stack n 1
[match y with (S (S x)) => x | x => x end] should be compiled into
[match y with O => y | (S n) => match n with O => y | (S x) => x end end]
- and [match y with (S (S n)) => n | n => n end] into
+ 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
@@ -705,7 +689,7 @@ 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 =
List.fold_right
(fun (pats,eqn) names -> merge_names alias_of_pat pats names)
eqns names1 in
@@ -719,7 +703,7 @@ let get_names env sign eqns =
let na =
merge_name
(fun (na,_,t) -> Name (next_name_away (named_hd env t na) avoid))
- d na
+ d na
in
(na::l,(out_name na)::avoid))
([],allvars) (List.rev sign) names2 in
@@ -756,7 +740,7 @@ let build_aliases_context env sigma names allpats pats =
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)
+ 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
@@ -776,7 +760,7 @@ let insert_aliases env sigma alias eqns =
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 =
+ 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 =
@@ -793,7 +777,7 @@ let noccur_between_without_evar n m term =
| Rel p -> if n<=p && p<n+m then raise Occur
| Evar (_,cl) -> ()
| _ -> iter_constr_with_binders succ occur_rec n c
- in
+ in
(m = 0) or (try occur_rec n term; true with Occur -> false)
@@ -870,7 +854,7 @@ let subst_predicate (args,copt) ccl tms =
let specialize_predicate_var (cur,typ,dep) tms ccl =
let c = if dep<>Anonymous then Some cur else None in
- let l =
+ let l =
match typ with
| IsInd (_,IndType(_,realargs),names) -> if names<>[] then realargs else []
| NotInd _ -> [] in
@@ -918,7 +902,7 @@ let abstract_predicate env sigma indf cur (names,(nadep,_)) tms ccl =
| 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
+ let ccl = if nadep <> Anonymous then ccl else lift_predicate 1 ccl tms in
let pred = extract_predicate [] ccl tms in
it_mkLambda_or_LetIn_name env pred sign
@@ -930,24 +914,24 @@ let known_dependent (_,dep) = (dep = KnownDep)
by gamma, x1...xn, x1...xk-1 |- [X=realargs]pred (if not dep) *)
let expand_arg tms ccl ((_,t),_,na) =
- let k = length_of_tomatch_type_sign na t in
+ let k = length_of_tomatch_type_sign na t in
lift_predicate (k-1) ccl tms
let adjust_impossible_cases pb pred tomatch submat =
if submat = [] then
- match kind_of_term (whd_evar (evars_of !(pb.evdref)) pred) with
+ match kind_of_term (whd_evar !(pb.evdref) pred) with
| Evar (evk,_) when snd (evar_source evk !(pb.evdref)) = ImpossibleCase ->
let default = (coq_unit_judge ()).uj_type in
- pb.evdref := Evd.evar_define evk default !(pb.evdref);
+ pb.evdref := Evd.define evk default !(pb.evdref);
(* we add an "assert false" case *)
let pats = List.map (fun _ -> PatVar (dummy_loc,Anonymous)) tomatch in
let aliasnames =
map_succeed (function Alias _ -> Anonymous | _ -> failwith"") tomatch
in
[ { patterns = pats;
- rhs = { rhs_env = pb.env;
- rhs_vars = [];
- avoid_ids = [];
+ rhs = { rhs_env = pb.env;
+ rhs_vars = [];
+ avoid_ids = [];
it = None };
alias_stack = Anonymous::aliasnames;
eqn_loc = dummy_loc;
@@ -971,14 +955,18 @@ let adjust_impossible_cases pb pred tomatch submat =
(* s.t Gamma,x1'..xn' |- match Pushed(x1')..Pushed(xn') rest with..end :pred'*)
(* *)
(*****************************************************************************)
-let specialize_predicate newtomatchs (names,(depna,_)) cs tms ccl =
+let specialize_predicate newtomatchs (names,(depna,_)) arsign cs tms ccl =
(* Assume some gamma st: gamma, (X,x:=realargs,copt), 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' *)
let n = cs.cs_nargs in
let ccl' = liftn_predicate n (k+1) ccl tms in
- let argsi = if nrealargs <> 0 then Array.to_list cs.cs_concl_realargs else [] in
+ let argsi =
+ 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'' *)
@@ -990,8 +978,8 @@ let specialize_predicate newtomatchs (names,(depna,_)) cs tms ccl =
List.fold_left (expand_arg tms) ccl''' newtomatchs
let find_predicate loc env evdref p current (IndType (indf,realargs)) dep tms =
- let pred= abstract_predicate env (evars_of !evdref) indf current dep tms p in
- (pred, whd_betaiota (evars_of !evdref)
+ let pred= abstract_predicate env !evdref indf current dep tms p in
+ (pred, whd_betaiota !evdref
(applist (pred, realargs@[current])), new_Type ())
let adjust_predicate_from_tomatch ((_,oldtyp),_,(nadep,_)) typ pb =
@@ -1037,8 +1025,8 @@ let group_equations pb ind current cstrs mat =
(fun eqn () ->
let rest = remove_current_pattern eqn in
let pat = current_pattern eqn in
- match check_and_adjust_constructor pb.env ind cstrs pat with
- | PatVar (_,name) ->
+ match check_and_adjust_constructor pb.env ind cstrs pat with
+ | PatVar (_,name) ->
(* This is a default clause that we expand *)
for i=1 to Array.length cstrs do
let args = make_anonymous_patvars cstrs.(i-1).cs_nargs in
@@ -1058,6 +1046,7 @@ let rec generalize_problem names pb = function
| [] -> pb
| i::l ->
let d = map_rel_declaration (lift i) (Environ.lookup_rel i pb.env) in
+ 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
@@ -1069,7 +1058,7 @@ let rec generalize_problem names pb = function
let build_leaf pb =
let rhs = extract_rhs pb in
let j = pb.typing_function (mk_tycon pb.pred) rhs.rhs_env pb.evdref rhs.it in
- j_nf_evar (evars_of !(pb.evdref)) j
+ j_nf_evar !(pb.evdref) j
(* Building the sub-problem when all patterns are variables *)
let shift_problem ((current,t),_,(nadep,_)) pb =
@@ -1080,17 +1069,17 @@ let shift_problem ((current,t),_,(nadep,_)) pb =
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 eqns const_info =
+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
+ else
DepAlias
in
- let history =
+ let history =
push_history_pattern const_info.cs_nargs
(AliasConstructor const_info.cs_cstr)
pb.history in
@@ -1109,10 +1098,10 @@ let build_branch current deps (realnames,dep) pb eqns const_info =
let dep_sign =
find_dependencies_signature
- (dependencies_in_rhs const_info.cs_nargs current pb.tomatch eqns)
+ (dependencies_in_rhs const_info.cs_nargs current pb.tomatch eqns)
(List.rev typs) in
- (* The dependent term to subst in the types of the remaining UnPushed
+ (* 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
@@ -1125,7 +1114,7 @@ let build_branch current deps (realnames,dep) pb eqns const_info =
let pred_is_not_dep =
noccur_predicate_between 1 (List.length realnames + 1) pb.pred tomatch in
- let typs'' =
+ let typs'' =
list_map2_i
(fun i (na,t) deps ->
let dep = match dep with
@@ -1139,8 +1128,8 @@ let build_branch current deps (realnames,dep) pb eqns const_info =
((mkRel i, lift_tomatch_type i t),deps,dep))
1 typs' (List.rev dep_sign) in
- let pred =
- specialize_predicate typs'' (realnames,dep) const_info tomatch pb.pred in
+ let pred =
+ specialize_predicate typs'' (realnames,dep) arsign const_info tomatch pb.pred in
let currents = List.map (fun x -> Pushed x) typs'' in
@@ -1199,6 +1188,7 @@ and match_current pb tomatch =
| 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 ct pb)
@@ -1209,12 +1199,12 @@ and match_current pb tomatch =
let pb = generalize_problem (names,dep) pb deps in
(* We compile branches *)
- let brs = array_map2 (compile_branch current (names,dep) deps pb) eqns cstrs in
+ let brs = array_map2 (compile_branch current (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) =
- find_predicate pb.caseloc pb.env pb.evdref
+ find_predicate pb.caseloc pb.env pb.evdref
pb.pred current indt (names,dep) pb.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
@@ -1222,8 +1212,8 @@ and match_current pb tomatch =
{ uj_val = applist (case, inst);
uj_type = substl inst typ }
-and compile_branch current names deps pb eqn cstr =
- let sign, pb = build_branch current deps names pb eqn cstr in
+and compile_branch current names deps pb arsign eqn cstr =
+ let sign, pb = build_branch current deps names pb arsign eqn cstr in
let j = compile pb in
(it_mkLambda_or_LetIn j.uj_val sign, j.uj_type)
@@ -1240,7 +1230,7 @@ and compile_generalization pb d rest =
and compile_alias pb (deppat,nondeppat,d,t) rest =
let history = simplify_history pb.history in
let sign, newenv, mat =
- insert_aliases pb.env (evars_of !(pb.evdref)) (deppat,nondeppat,d,t) pb.mat in
+ insert_aliases pb.env !(pb.evdref) (deppat,nondeppat,d,t) pb.mat in
let n = List.length sign in
(* We had Gamma1; x:current; Gamma2 |- tomatch(x) and we rebind x to get *)
@@ -1287,102 +1277,6 @@ let matx_of_eqns env tomatchl eqns =
rhs = rhs }
in List.map build_eqn eqns
-(************************************************************************)
-(* preparing the elimination predicate if any *)
-
-let build_expected_arity env evdref isdep tomatchl =
- let cook n = function
- | _,IsInd (_,IndType(indf,_),_) ->
- let indf' = lift_inductive_family n indf in
- Some (build_dependent_inductive env indf', fst (get_arity env indf'))
- | _,NotInd _ -> None
- in
- let rec buildrec n env = function
- | [] -> new_Type ()
- | tm::ltm ->
- match cook n tm with
- | None -> buildrec n env ltm
- | Some (ty1,aritysign) ->
- let rec follow n env = function
- | d::sign ->
- mkProd_or_LetIn_name env
- (follow (n+1) (push_rel d env) sign) d
- | [] ->
- if isdep then
- mkProd (Anonymous, ty1,
- buildrec (n+1)
- (push_rel_assum (Anonymous, ty1) env)
- ltm)
- else buildrec n env ltm
- in follow n env (List.rev aritysign)
- in buildrec 0 env tomatchl
-
-let extract_predicate_conclusion isdep tomatchl pred =
- let cook = function
- | _,IsInd (_,IndType(_,args),_) -> Some (List.length args)
- | _,NotInd _ -> None in
- let rec decomp_lam_force n l p =
- if n=0 then (l,p) else
- match kind_of_term p with
- | Lambda (na,_,c) -> decomp_lam_force (n-1) (na::l) c
- | _ -> (* eta-expansion *)
- let na = Name (id_of_string "x") in
- decomp_lam_force (n-1) (na::l) (applist (lift 1 p, [mkRel 1])) in
- let rec buildrec allnames p = function
- | [] -> (List.rev allnames,p)
- | tm::ltm ->
- match cook tm with
- | None ->
- let p =
- (* adjust to a sign containing the NotInd's *)
- if isdep then lift 1 p else p in
- let names = if isdep then [Anonymous] else [] in
- buildrec (names::allnames) p ltm
- | Some n ->
- let n = if isdep then n+1 else n in
- let names,p = decomp_lam_force n [] p in
- buildrec (names::allnames) p ltm
- in buildrec [] pred tomatchl
-
-let set_arity_signature dep n arsign tomatchl pred x =
- (* avoid is not exhaustive ! *)
- let rec decomp_lam_force n avoid l p =
- if n = 0 then (List.rev l,p,avoid) else
- match p with
- | RLambda (_,(Name id as na),_,_,c) ->
- decomp_lam_force (n-1) (id::avoid) (na::l) c
- | RLambda (_,(Anonymous as na),_,_,c) -> decomp_lam_force (n-1) avoid (na::l) c
- | _ ->
- let x = next_ident_away (id_of_string "x") avoid in
- decomp_lam_force (n-1) (x::avoid) (Name x :: l)
- (* eta-expansion *)
- (let a = RVar (dummy_loc,x) in
- match p with
- | RApp (loc,p,l) -> RApp (loc,p,l@[a])
- | _ -> (RApp (dummy_loc,p,[a]))) in
- let rec decomp_block avoid p = function
- | ([], _) -> x := Some p
- | ((_,IsInd (_,IndType(indf,realargs),_))::l),(y::l') ->
- let (ind,params) = dest_ind_family indf in
- let (nal,p,avoid') = decomp_lam_force (List.length realargs) avoid [] p
- in
- let na,p,avoid' =
- if dep then decomp_lam_force 1 avoid' [] p else [Anonymous],p,avoid'
- in
- y :=
- (List.hd na,
- if List.for_all ((=) Anonymous) nal then
- None
- else
- Some (dummy_loc, ind, (List.map (fun _ -> Anonymous) params)@nal));
- decomp_block avoid' p (l,l')
- | (_::l),(y::l') ->
- y := (Anonymous,None);
- decomp_block avoid p (l,l')
- | _ -> anomaly "set_arity_signature"
- in
- decomp_block [] pred (tomatchl,arsign)
-
(***************** Building an inversion predicate ************************)
(* Let "match t1 in I1 u11..u1n_1 ... tm in Im um1..umn_m with ... end : T"
@@ -1395,9 +1289,9 @@ let set_arity_signature dep n arsign tomatchl pred x =
variables (in practice, there is no reason that ti is already
constructed and the qi will be degenerated).
- We then look for a type U(..a1jk..b1 .. ..amjk..bm) so that
+ We then look for a type U(..a1jk..b1 .. ..amjk..bm) so that
T = U(..v1jk..t1 .. ..vmjk..tm). This a higher-order matching
- problem with a priori different solution (one of them if T itself!).
+ problem with a priori different solutions (one of them if T itself!).
We finally invert the uij and the ti and build the return clause
@@ -1414,27 +1308,27 @@ let set_arity_signature dep n arsign tomatchl pred x =
let adjust_to_extended_env_and_remove_deps env extenv subst t =
let n = rel_context_length (rel_context env) in
let n' = rel_context_length (rel_context extenv) in
- (* We first remove the bindings that are dependently typed (they are
+ (* We first remove the bindings that are dependently typed (they are
difficult to manage and it is not sure these are so useful in practice);
Notes:
- [subst] is made of pairs [(id,u)] where id is a name in [extenv] and
[u] a term typed in [env];
- [subst0] is made of items [(p,u,(u,ty))] where [ty] is the type of [u]
- and both are adjusted to [extenv] while [p] is the index of [id] in
+ and both are adjusted to [extenv] while [p] is the index of [id] in
[extenv] (after expansion of the aliases) *)
let subst0 = map_succeed (fun (x,u) ->
(* d1 ... dn dn+1 ... dn'-p+1 ... dn' *)
(* \--env-/ (= x:ty) *)
(* \--------------extenv------------/ *)
- let (p,_) = lookup_rel_id x (rel_context extenv) in
+ let (p,_,_) = lookup_rel_id x (rel_context extenv) in
let rec aux n (_,b,ty) =
match b with
| Some c ->
assert (isRel c);
- let p = n + destRel c in aux p (lookup_rel p (rel_context extenv))
+ let p = n + destRel c in aux p (lookup_rel p extenv)
| None ->
(n,ty) in
- let (p,ty) = aux p (lookup_rel p (rel_context extenv)) 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
@@ -1444,12 +1338,15 @@ let adjust_to_extended_env_and_remove_deps env extenv subst t =
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 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.
*
- * [abstract_tycon Gamma0 Sigma subst T Gamma] looks for U(..v1jk..t1 .. ..vmjk..tm)
- * defined in some extended context
+ * [abstract_tycon Gamma0 Sigma subst T Gamma] looks for U(..v1jk..t1 .. ..vmjk..tm)
+ * defined in some extended context
* "Gamma0, ..a1jk:V1jk.. b1:W1 .. ..amjk:Vmjk.. bm:Wm"
* such that env |- T = U(..v1jk..t1 .. ..vmjk..tm). To not commit to
* a particular solution, we replace each subterm t in T that unifies with
@@ -1460,7 +1357,7 @@ let adjust_to_extended_env_and_remove_deps env extenv subst t =
*)
let abstract_tycon loc env evdref subst _tycon extenv t =
- let sigma = evars_of !evdref in
+ let sigma = !evdref in
let t = nf_betaiota sigma t in (* it helps in some cases to remove K-redex *)
let subst0,t0 = adjust_to_extended_env_and_remove_deps env extenv subst t in
(* We traverse the type T of the original problem Xi looking for subterms
@@ -1469,15 +1366,18 @@ 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 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 vl = List.map pi1 good in
- let inst =
+ 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 =
+ let rel_filter =
List.map (fun a -> not (isRel a) or dependent a u) inst in
let named_filter =
List.map (fun (id,_,_) -> dependent (mkVar id) u)
@@ -1488,30 +1388,25 @@ let abstract_tycon loc env evdref subst _tycon extenv t =
evdref := add_conv_pb (Reduction.CONV,extenv,substl inst ev,u) !evdref;
lift k ev
else
- map_constr_with_full_binders
- (fun d (k,env,subst) ->
- k+1,
- push_rel d env,
- List.map (fun (na,u,d) -> (na,lift 1 u,d)) subst)
- aux x t in
+ 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
| None ->
- (* This is the situation we are building a return predicate and
+ (* 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 impossible_case_type =
+ 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
- get_judgment_of extenv (evars_of !evdref) t
+ get_judgment_of extenv !evdref t
(* 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
- * predicate for Xi that is itself made by an auxiliary
+ * predicate for Xi that is itself made by an auxiliary
* pattern-matching problem of which the first clause reveals the
* pattern structure of the constraints on the inductive types of the t1..tn,
* and the second clause is a wildcard clause for catching the
@@ -1519,8 +1414,7 @@ let build_tycon loc env tycon_env subst tycon extenv evdref t =
* further explanations
*)
-let build_inversion_problem loc env evdref tms t =
- let sigma = evars_of !evdref in
+let build_inversion_problem loc env sigma tms t =
let make_patvar t (subst,avoid) =
let id = next_name_away (named_hd env t Anonymous) avoid in
PatVar (dummy_loc,Name id), ((id,t)::subst, id::avoid) in
@@ -1596,12 +1490,13 @@ let build_inversion_problem loc env evdref tms t =
alias_stack = [];
eqn_loc = dummy_loc;
used = ref false;
- rhs = { rhs_env = pb_env;
- rhs_vars = [];
+ rhs = { rhs_env = pb_env;
+ rhs_vars = [];
avoid_ids = avoid0;
it = None } } in
- (* [pb] is the auxiliary pattern-matching serving as skeleton for the
+ (* [pb] is the auxiliary pattern-matching serving as skeleton for the
return type of the original problem Xi *)
+ let evdref = ref sigma in
let pb =
{ env = pb_env;
evdref = evdref;
@@ -1612,26 +1507,9 @@ let build_inversion_problem loc env evdref tms t =
caseloc = loc;
casestyle = RegularStyle;
typing_function = build_tycon loc env pb_env subst} in
- (compile pb).uj_val
+ let pred = (compile pb).uj_val in
+ (!evdref,pred)
-let prepare_predicate_from_tycon loc dep env evdref tomatchs sign c =
- let cook (n, l, env, signs) = function
- | c,IsInd (_,IndType(indf,realargs),_) ->
- let indf' = lift_inductive_family n indf in
- let sign = make_arity_signature env dep indf' in
- let p = List.length realargs in
- if dep then
- (n + p + 1, c::(List.rev realargs)@l, push_rels sign env,sign::signs)
- else
- (n + p, (List.rev realargs)@l, push_rels sign env,sign::signs)
- | c,NotInd (bo,typ) ->
- let sign = [Anonymous,Option.map (lift n) bo,lift n typ] in
- let sign = name_context env sign in
- (n + 1, c::l, push_rels sign env, sign::signs) in
- let n,allargs,env',signs = List.fold_left cook (0, [], env, []) tomatchs in
- let names = List.rev (List.map (List.map pi1) signs) in
- names, build_inversion_problem loc env evdref tomatchs c
-
(* Here, [pred] is assumed to be in the context built from all *)
(* realargs and terms to match *)
let build_initial_predicate knowndep allnames pred =
@@ -1658,27 +1536,27 @@ let build_initial_predicate knowndep allnames pred =
let extract_arity_signature env0 tomatchl tmsign =
let get_one_sign n tm (na,t) =
match tm with
- | NotInd (bo,typ) ->
+ | NotInd (bo,typ) ->
(match t with
| None -> [na,Option.map (lift n) bo,lift n typ]
- | Some (loc,_,_,_) ->
+ | Some (loc,_,_,_) ->
user_err_loc (loc,"",
str"Unexpected type annotation for a term of non inductive type."))
| IsInd (term,IndType(indf,realargs),_) ->
let indf' = lift_inductive_family n indf in
- let (ind,params) = dest_ind_family indf' in
- let nrealargs = List.length realargs in
+ let (ind,_) = dest_ind_family indf' in
+ let nparams_ctxt,nrealargs_ctxt = inductive_nargs env0 ind in
+ let arsign = fst (get_arity env0 indf') in
let realnal =
match t with
| Some (loc,ind',nparams,realnal) ->
if ind <> ind' then
user_err_loc (loc,"",str "Wrong inductive type.");
- if List.length params <> nparams
- or nrealargs <> List.length realnal then
+ if nparams_ctxt <> nparams
+ or nrealargs_ctxt <> List.length realnal then
anomaly "Ill-formed 'in' clause in cases";
List.rev realnal
- | None -> list_make nrealargs Anonymous in
- let arsign = fst (get_arity env0 indf') in
+ | None -> list_make nrealargs_ctxt Anonymous in
(* let na = *)
(* match na with *)
(* | Name _ -> na *)
@@ -1707,43 +1585,46 @@ 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 env tomatchs sign arsign c =
+let prepare_predicate_from_arsign_tycon loc tomatchs sign arsign c =
let nar = List.fold_left (fun n sign -> List.length sign + n) 0 arsign in
- let subst, len =
+ let subst, len =
List.fold_left2 (fun (subst, len) (tm, tmtype) sign ->
let signlen = List.length sign in
match kind_of_term tm with
- | Rel n when dependent tm c
+ | Rel n when dependent tm c
&& signlen = 1 (* The term to match is not of a dependent type itself *) ->
((n, len) :: subst, len - signlen)
- | Rel _ when not (dependent tm c)
- && signlen > 1 (* The term is of a dependent type but does not appear in
- the tycon, maybe some variable in its type does. *) ->
+ | Rel n when signlen > 1 (* The term is of a dependent type,
+ maybe some variable in its type appears in the tycon. *) ->
(match tmtype with
NotInd _ -> (* len - signlen, subst*) assert false (* signlen > 1 *)
| IsInd (_, IndType(indf,realargs),_) ->
- List.fold_left
- (fun (subst, len) arg ->
- match kind_of_term arg with
+ let subst =
+ if dependent tm c && List.for_all isRel realargs
+ then (n, 1) :: subst else subst
+ in
+ List.fold_left
+ (fun (subst, len) arg ->
+ match kind_of_term arg with
| Rel n when dependent arg c ->
((n, len) :: subst, pred len)
| _ -> (subst, pred len))
- (subst, len) realargs)
+ (subst, len) realargs)
| _ -> (subst, len - signlen))
([], nar) tomatchs arsign
in
let rec predicate lift c =
match kind_of_term c with
- | Rel n when n > lift ->
- (try
+ | Rel n when n > lift ->
+ (try
(* Make the predicate dependent on the matched variable *)
let idx = List.assoc (n - lift) subst in
mkRel (idx + lift)
- with Not_found ->
+ with Not_found ->
(* A variable that is not matched, lift over the arsign. *)
mkRel (n + nar))
| _ ->
- map_constr_with_binders succ predicate lift c
+ map_constr_with_binders succ predicate lift c
in predicate 0 c
@@ -1758,92 +1639,71 @@ let prepare_predicate_from_arsign_tycon loc env tomatchs sign arsign c =
* tycon to make the predicate if it is not closed.
*)
-let is_dependent_on_rel x t =
- match kind_of_term x with
- Rel n -> not (noccur_with_meta n n t)
- | _ -> false
-
let prepare_predicate loc typing_fun evdref env tomatchs sign tycon pred =
- match pred with
- (* No type annotation *)
- | None ->
- (match tycon with
- | Some (None, t) when not (noccur_with_meta 0 max_int t) ->
- (* If the tycon is not closed w.r.t real variables *)
- (* We try two different strategies *)
- let evdref2 = ref !evdref in
- let arsign = extract_arity_signature env tomatchs sign in
- let env' = List.fold_right push_rels arsign env in
- (* First strategy: we abstract the tycon wrt to the dependencies *)
- let names1 = List.rev (List.map (List.map pi1) arsign) in
- let pred1 = prepare_predicate_from_arsign_tycon loc env' tomatchs sign arsign t in
- let nal1,pred1 = build_initial_predicate KnownDep names1 pred1 in
- (* Second strategy: we build an "inversion" predicate *)
- let names2,pred2 =
- prepare_predicate_from_tycon loc true env evdref2 tomatchs sign t
- in
- let nal2,pred2 = build_initial_predicate DepUnknown names2 pred2 in
- [evdref, nal1, pred1; evdref2, nal2, pred2]
- | Some (None, t) ->
- (* Only one strategy: we build an "inversion" predicate *)
- let names,pred =
- prepare_predicate_from_tycon loc true env evdref tomatchs sign t
- in
- let nal,pred = build_initial_predicate DepUnknown names pred in
- [evdref, nal, pred]
- | _ ->
- (* No type constaints: we use two strategies *)
- let evdref2 = ref !evdref in
- let t1 = mkExistential env ~src:(loc, CasesType) evdref in
- (* First strategy: we pose a possibly dependent "inversion" evar *)
- let names1,pred1 =
- prepare_predicate_from_tycon loc true env evdref tomatchs sign t1
- in
- let nal1,pred1 = build_initial_predicate DepUnknown names1 pred1 in
- (* Second strategy: we pose a non dependent evar *)
- let t2 = mkExistential env ~src:(loc, CasesType) evdref2 in
- let arsign = extract_arity_signature env tomatchs sign in
- let names2 = List.rev (List.map (List.map pi1) arsign) in
- let pred2 = lift (List.length names2) t2 in
- let nal2,pred2 = build_initial_predicate KnownNotDep names2 pred2 in
- [evdref, nal1, pred1; evdref2, nal2, pred2])
-
- (* Some type annotation *)
- | Some rtntyp ->
+ let arsign = extract_arity_signature env tomatchs sign in
+ let names = List.rev (List.map (List.map pi1) arsign) in
+ let preds =
+ match pred, tycon with
+ (* No type annotation *)
+ | None, Some (None, t) when not (noccur_with_meta 0 max_int t) ->
+ (* If the tycon is not closed w.r.t real variables, we try *)
+ (* two different strategies *)
+ (* First strategy: we abstract the tycon wrt to the dependencies *)
+ let pred1 =
+ prepare_predicate_from_arsign_tycon loc tomatchs sign 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) ->
+ (* Only one strategy: we build an "inversion" predicate *)
+ let sigma,pred = build_inversion_problem loc env !evdref tomatchs t in
+ [sigma, DepUnknown, pred]
+ | 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
+ (* 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]
+ (* Some type annotation *)
+ | Some rtntyp, _ ->
(* We extract the signature of the arity *)
- let arsign = extract_arity_signature env tomatchs sign in
let env = List.fold_right push_rels arsign env in
- let allnames = List.rev (List.map (List.map pi1) arsign) in
let predcclj = typing_fun (mk_tycon (new_Type ())) env evdref rtntyp in
- let _ =
- Option.map (fun tycon ->
- evdref := Coercion.inh_conv_coerces_to loc env !evdref predcclj.uj_val
- (lift_tycon_type (List.length arsign) tycon))
- tycon
- in
- let predccl = (j_nf_isevar !evdref predcclj).uj_val in
- let nal,pred = build_initial_predicate KnownDep allnames predccl in
- [evdref, nal, pred]
+ 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]
+ in
+ List.map
+ (fun (sigma,dep,pred) ->
+ let (nal,pred) = build_initial_predicate dep names pred in
+ sigma,nal,pred)
+ preds
(**************************************************************************)
(* Main entry of the matching compilation *)
-
+
let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, eqns) =
(* We build the matrix of patterns and right-hand-side *)
let matx = matx_of_eqns env tomatchl eqns in
-
+
(* We build the vector of terms to match consistently with the *)
(* constructors found in patterns *)
let tomatchs = coerce_to_indtype typing_fun evdref env matx tomatchl in
-
+
(* If an elimination predicate is provided, we check it is compatible
with the type of arguments to match; if none is provided, we
build alternative possible predicates *)
let sign = List.map snd tomatchl in
let preds = prepare_predicate loc typing_fun evdref env tomatchs sign tycon predopt in
-
- let compile_for_one_predicate (myevdref,nal,pred) =
+
+ 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) *)
@@ -1854,6 +1714,8 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e
| Some t -> typing_fun tycon env evdref t
| None -> coq_unit_judge () in
+ let myevdref = ref sigma in
+
let pb =
{ env = env;
evdref = myevdref;
diff --git a/pretyping/cases.mli b/pretyping/cases.mli
index 98923b2a..66924031 100644
--- a/pretyping/cases.mli
+++ b/pretyping/cases.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: cases.mli 11014 2008-05-28 19:09:32Z herbelin $ i*)
+(*i $Id$ i*)
(*i*)
open Util
@@ -64,8 +64,8 @@ type tomatch_status =
module type S = sig
val compile_cases :
loc -> case_style ->
- (type_constraint -> env -> evar_defs ref -> rawconstr -> unsafe_judgment) * evar_defs ref ->
- type_constraint ->
+ (type_constraint -> env -> evar_map ref -> rawconstr -> unsafe_judgment) * evar_map ref ->
+ type_constraint ->
env -> rawconstr option * tomatch_tuples * cases_clauses ->
unsafe_judgment
end
diff --git a/pretyping/cbv.ml b/pretyping/cbv.ml
index f4c612a5..8c03d0df 100644
--- a/pretyping/cbv.ml
+++ b/pretyping/cbv.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: cbv.ml 11897 2009-02-09 19:28:02Z barras $ *)
+(* $Id$ *)
open Util
open Pp
@@ -27,7 +27,14 @@ open Esubst
* in normal form and neutral (i.e. not a lambda, a construct or a
* (co)fix, because they may produce redexes by applying them,
* or putting them in a case)
- * LAM(x,a,b,S) is the term [S]([x:a]b). the bindings is propagated
+ * STACK(k,v,stk) represents an irreductible value [v] in the stack [stk].
+ * [k] is a delayed shift to be applied to both the value and
+ * the stack.
+ * CBN(t,S) is the term [S]t. It is used to delay evaluation. For
+ * instance products are evaluated only when actually needed
+ * (CBN strategy).
+ * LAM(n,a,b,S) is the term [S]([x:a]b) where [a] is a list of bindings and
+ * [n] is the length of [a]. the environment [S] is propagated
* only when the abstraction is applied, and then we use the rule
* ([S]([x:a]b) c) --> [S.c]b
* This corresponds to the usual strategy of weak reduction
@@ -36,28 +43,47 @@ open Esubst
* weak reduction.
* CONSTR(c,args) is the constructor [c] applied to [args].
*
- * Note that any term has not an equivalent in cbv_value: for example,
- * a product (x:A)B must be in normal form because only VAL may
- * represent it, and the argument of VAL is always in normal
- * form. This remark precludes coding a head reduction with these
- * functions. Anyway, does it make sense to head reduce with a
- * call-by-value strategy ?
*)
type cbv_value =
| VAL of int * constr
+ | STACK of int * cbv_value * cbv_stack
+ | CBN of constr * cbv_value subs
| LAM of int * (name * constr) list * constr * cbv_value subs
| FIXP of fixpoint * cbv_value subs * cbv_value array
| COFIXP of cofixpoint * cbv_value subs * cbv_value array
| CONSTR of constructor * cbv_value array
+(* type of terms with a hole. This hole can appear only under App or Case.
+ * TOP means the term is considered without context
+ * APP(v,stk) means the term is applied to v, and then the context stk
+ * (v.0 is the first argument).
+ * this corresponds to the application stack of the KAM.
+ * The members of l are values: we evaluate arguments before
+ calling the function.
+ * CASE(t,br,pat,S,stk) means the term is in a case (which is himself in stk
+ * t is the type of the case and br are the branches, all of them under
+ * the subs S, pat is information on the patterns of the Case
+ * (Weak reduction: we propagate the sub only when the selected branch
+ * is determined)
+ *
+ * Important remark: the APPs should be collapsed:
+ * (APP (l,(APP ...))) forbidden
+ *)
+and cbv_stack =
+ | TOP
+ | APP of cbv_value array * cbv_stack
+ | CASE of constr * constr array * case_info * cbv_value subs * cbv_stack
+
(* les vars pourraient etre des constr,
- cela permet de retarder les lift: utile ?? *)
+ cela permet de retarder les lift: utile ?? *)
(* relocation of a value; used when a value stored in a context is expanded
* in a larger context. e.g. [%k (S.t)](k+1) --> [^k]t (t is shifted of k)
*)
let rec shift_value n = function
- | VAL (k,v) -> VAL ((k+n),v)
+ | VAL (k,t) -> VAL (k+n,t)
+ | STACK(k,v,stk) -> STACK(k+n,v,stk)
+ | CBN (t,s) -> CBN(t,subs_shft(n,s))
| LAM (nlams,ctxt,b,s) -> LAM (nlams,ctxt,b,subs_shft (n,s))
| FIXP (fix,s,args) ->
FIXP (fix,subs_shft (n,s), Array.map (shift_value n) args)
@@ -68,6 +94,13 @@ 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
@@ -89,29 +122,6 @@ let make_constr_ref n = function
| VarKey id -> mkVar id
| ConstKey cst -> mkConst cst
-
-(* type of terms with a hole. This hole can appear only under App or Case.
- * TOP means the term is considered without context
- * APP(v,stk) means the term is applied to v, and then the context stk
- * (v.0 is the first argument).
- * this corresponds to the application stack of the KAM.
- * The members of l are values: we evaluate arguments before
- calling the function.
- * CASE(t,br,pat,S,stk) means the term is in a case (which is himself in stk
- * t is the type of the case and br are the branches, all of them under
- * the subs S, pat is information on the patterns of the Case
- * (Weak reduction: we propagate the sub only when the selected branch
- * is determined)
- *
- * Important remark: the APPs should be collapsed:
- * (APP (l,(APP ...))) forbidden
- *)
-
-type cbv_stack =
- | TOP
- | APP of cbv_value array * cbv_stack
- | CASE of constr * constr array * case_info * cbv_value subs * cbv_stack
-
(* Adds an application list. Collapse APPs! *)
let stack_app appl stack =
if Array.length appl = 0 then stack else
@@ -119,6 +129,20 @@ let stack_app appl stack =
| APP(args,stk) -> APP(Array.append appl args,stk)
| _ -> APP(appl, stack)
+let rec stack_concat stk1 stk2 =
+ match stk1 with
+ TOP -> stk2
+ | APP(v,stk1') -> APP(v,stack_concat stk1' stk2)
+ | CASE(c,b,i,s,stk1') -> CASE(c,b,i,s,stack_concat stk1' stk2)
+
+(* merge stacks when there is no shifts in between *)
+let mkSTACK = function
+ v, TOP -> v
+ | STACK(0,v0,stk0), stk -> STACK(0,v0,stack_concat stk0 stk)
+ | v,stk -> STACK(0,v,stk)
+
+
+(* Change: zeta reduction cannot be avoided in CBV *)
open RedFlags
@@ -128,7 +152,8 @@ let red_set_ref flags = function
| ConstKey sp -> red_set flags (fCONST sp)
(* Transfer application lists from a value to the stack
- * useful because fixpoints may be totally applied in several times
+ * useful because fixpoints may be totally applied in several times.
+ * On the other hand, irreductible atoms absorb the full stack.
*)
let strip_appl head stack =
match head with
@@ -148,7 +173,7 @@ let fixp_reducible flgs ((reci,i),_) stk =
CONSTR _ -> true
| _ -> false)
| _ -> false
- else
+ else
false
let cofixp_reducible flgs _ stk =
@@ -156,7 +181,7 @@ let cofixp_reducible flgs _ stk =
match stk with
| (CASE _ | APP(_,CASE _)) -> true
| _ -> false
- else
+ else
false
@@ -196,27 +221,17 @@ let rec norm_head info env t stack =
| Const sp -> norm_head_ref 0 info env stack (ConstKey sp)
- | LetIn (x, b, t, c) ->
+ | LetIn (_, b, _, c) ->
(* zeta means letin are contracted; delta without zeta means we *)
(* allow bindings but leave let's in place *)
- let zeta = red_set (info_flags info) fZETA in
- let env' =
- if zeta
- (* New rule: for Cbv, Delta does not apply to locally bound variables
- or red_set (info_flags info) fDELTA
- *)
- then
- subs_cons ([|cbv_stack_term info TOP env b|],env)
- else
- subs_lift env in
- if zeta then
+ if red_set (info_flags info) fZETA then
+ (* New rule: for Cbv, Delta does not apply to locally bound variables
+ or red_set (info_flags info) fDELTA
+ *)
+ let env' = subs_cons ([|cbv_stack_term info TOP env b|],env) in
norm_head info env' c stack
else
- let normt =
- mkLetIn (x, cbv_norm_term info env b,
- cbv_norm_term info env t,
- cbv_norm_term info env' c) in
- (VAL(0,normt), stack) (* Considérer une coupure commutative ? *)
+ (CBN(t,env), stack) (* Considérer une coupure commutative ? *)
| Evar ev ->
(match evar_value info ev with
@@ -233,23 +248,20 @@ let rec norm_head info env t stack =
(* neutral cases *)
| (Sort _ | Meta _ | Ind _) -> (VAL(0, t), stack)
- | Prod (x,t,c) ->
- (VAL(0, mkProd (x, cbv_norm_term info env t,
- cbv_norm_term info (subs_lift env) c)),
- stack)
+ | Prod _ -> (CBN(t,env), stack)
and norm_head_ref k info env stack normt =
if red_set_ref (info_flags info) normt then
match ref_value_cache info normt with
| Some body -> strip_appl (shift_value k body) stack
- | None -> (VAL(0,make_constr_ref k normt), stack)
- else (VAL(0,make_constr_ref k normt), stack)
+ | None -> (VAL(0,make_constr_ref k normt),stack)
+ else (VAL(0,make_constr_ref k normt),stack)
(* cbv_stack_term performs weak reduction on constr t under the subs
* env, with context stack, i.e. ([env]t stack). First computes weak
* head normal form of t and checks if a redex appears with the stack.
* If so, recursive call to reach the real head normal form. If not,
- * we build a value.
+ * we build a value.
*)
and cbv_stack_term info stack env t =
match norm_head info env t stack with
@@ -268,47 +280,43 @@ and cbv_stack_term info stack env t =
LAM(nlams-nargs,ctxt', b, subs_cons(args,env))
(* a Fix applied enough -> IOTA *)
- | (FIXP(fix,env,_), stk)
+ | (FIXP(fix,env,[||]), stk)
when fixp_reducible (info_flags info) fix stk ->
let (envf,redfix) = contract_fixp env fix in
cbv_stack_term info stk envf redfix
(* constructor guard satisfied or Cofix in a Case -> IOTA *)
- | (COFIXP(cofix,env,_), stk)
+ | (COFIXP(cofix,env,[||]), stk)
when cofixp_reducible (info_flags info) cofix stk->
let (envf,redfix) = contract_cofixp env cofix in
cbv_stack_term info stk envf redfix
(* constructor in a Case -> IOTA *)
- | (CONSTR((sp,n),_), APP(args,CASE(_,br,ci,env,stk)))
+ | (CONSTR((sp,n),[||]), APP(args,CASE(_,br,ci,env,stk)))
when red_set (info_flags info) fIOTA ->
let cargs =
Array.sub args ci.ci_npar (Array.length args - ci.ci_npar) in
cbv_stack_term info (stack_app cargs stk) env br.(n-1)
-
+
(* constructor of arity 0 in a Case -> IOTA *)
| (CONSTR((_,n),_), CASE(_,br,_,env,stk))
when red_set (info_flags info) fIOTA ->
cbv_stack_term info stk env br.(n-1)
- (* may be reduced later by application *)
- | (head, TOP) -> head
- | (FIXP(fix,env,_), APP(appl,TOP)) -> FIXP(fix,env,appl)
- | (COFIXP(cofix,env,_), APP(appl,TOP)) -> COFIXP(cofix,env,appl)
- | (CONSTR(c,_), APP(appl,TOP)) -> CONSTR(c,appl)
-
- (* absurd cases (ill-typed) *)
- | (LAM _, CASE _) -> assert false
+ (* may be reduced later by application *)
+ | (FIXP(fix,env,[||]), APP(appl,TOP)) -> FIXP(fix,env,appl)
+ | (COFIXP(cofix,env,[||]), APP(appl,TOP)) -> COFIXP(cofix,env,appl)
+ | (CONSTR(c,[||]), APP(appl,TOP)) -> CONSTR(c,appl)
(* definitely a value *)
- | (head,stk) -> VAL(0,apply_stack info (cbv_norm_value info head) stk)
+ | (head,stk) -> mkSTACK(head, stk)
(* When we are sure t will never produce a redex with its stack, we
* normalize (even under binders) the applied terms and we build the
* final term
*)
-and apply_stack info t = function
+let rec apply_stack info t = function
| TOP -> t
| APP (args,st) ->
apply_stack info (mkApp(t,Array.map (cbv_norm_value info) args)) st
@@ -318,7 +326,6 @@ and apply_stack info t = function
Array.map (cbv_norm_term info env) br))
st
-
(* performs the reduction on a constr, and returns a constr *)
and cbv_norm_term info env t =
(* reduction under binders *)
@@ -326,7 +333,13 @@ and cbv_norm_term info env t =
(* reduction of a cbv_value to a constr *)
and cbv_norm_value info = function (* reduction under binders *)
- | VAL (n,v) -> lift n v
+ | VAL (n,t) -> lift n t
+ | STACK (0,v,stk) ->
+ apply_stack info (cbv_norm_value info v) stk
+ | STACK (n,v,stk) ->
+ lift n (apply_stack info (cbv_norm_value info v) stk)
+ | CBN(t,env) ->
+ map_constr_with_binders subs_lift (cbv_norm_term info) env t
| LAM (n,ctxt,b,env) ->
let nctxt =
list_map_i (fun i (x,ty) ->
@@ -337,14 +350,14 @@ and cbv_norm_value info = function (* reduction under binders *)
(mkFix (lij,
(names,
Array.map (cbv_norm_term info env) lty,
- Array.map (cbv_norm_term info
+ Array.map (cbv_norm_term info
(subs_liftn (Array.length lty) env)) bds)),
Array.map (cbv_norm_value info) args)
| COFIXP ((j,(names,lty,bds)),env,args) ->
mkApp
(mkCoFix (j,
(names,Array.map (cbv_norm_term info env) lty,
- Array.map (cbv_norm_term info
+ Array.map (cbv_norm_term info
(subs_liftn (Array.length lty) env)) bds)),
Array.map (cbv_norm_value info) args)
| CONSTR (c,args) ->
@@ -354,7 +367,6 @@ and cbv_norm_value info = function (* reduction under binders *)
let cbv_norm infos constr =
with_stats (lazy (cbv_norm_term infos (ESID 0) constr))
-
type cbv_infos = cbv_value infos
(* constant bodies are normalized at the first expansion *)
diff --git a/pretyping/cbv.mli b/pretyping/cbv.mli
index 9ab15886..de66d22b 100644
--- a/pretyping/cbv.mli
+++ b/pretyping/cbv.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: cbv.mli 11897 2009-02-09 19:28:02Z barras $ i*)
+(*i $Id$ i*)
(*i*)
open Names
@@ -29,18 +29,20 @@ val cbv_norm : cbv_infos -> constr -> constr
(*i This is for cbv debug *)
type cbv_value =
| VAL of int * constr
+ | STACK of int * cbv_value * cbv_stack
+ | CBN of constr * cbv_value subs
| LAM of int * (name * constr) list * constr * cbv_value subs
| FIXP of fixpoint * cbv_value subs * cbv_value array
| COFIXP of cofixpoint * cbv_value subs * cbv_value array
| CONSTR of constructor * cbv_value array
-val shift_value : int -> cbv_value -> cbv_value
-
-type cbv_stack =
+and cbv_stack =
| TOP
| APP of cbv_value array * cbv_stack
| CASE of constr * constr array * case_info * cbv_value subs * cbv_stack
+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
diff --git a/pretyping/classops.ml b/pretyping/classops.ml
index 27418b4d..23796325 100644
--- a/pretyping/classops.ml
+++ b/pretyping/classops.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: classops.ml 11897 2009-02-09 19:28:02Z barras $ *)
+(* $Id$ *)
open Util
open Pp
@@ -28,8 +28,8 @@ open Mod_subst
(* A class is a type constructor, its type is an arity whose number of
arguments is cl_param (0 for CL_SORT and CL_FUN) *)
-type cl_typ =
- | CL_SORT
+type cl_typ =
+ | CL_SORT
| CL_FUN
| CL_SECVAR of variable
| CL_CONST of constant
@@ -68,8 +68,6 @@ module Bijint = struct
(let v = Array.make (b.s + 8) (x,y) in Array.blit b.v 0 v 0 b.s; v)
else b.v in
v.(b.s) <- (x,y); { v = v; s = b.s+1; inv = Gmap.add x b.s b.inv }
- let replace n x y b =
- let v = Array.copy b.v in v.(n) <- (x,y); { b with v = v }
let dom b = Gmap.dom b.inv
end
@@ -84,7 +82,7 @@ let inheritance_graph =
let freeze () = (!class_tab, !coercion_tab, !inheritance_graph)
-let unfreeze (fcl,fco,fig) =
+let unfreeze (fcl,fco,fig) =
class_tab:=fcl;
coercion_tab:=fco;
inheritance_graph:=fig
@@ -95,26 +93,24 @@ let add_new_class cl s =
if not (Bijint.mem cl !class_tab) then
class_tab := Bijint.add cl s !class_tab
-let add_new_coercion coe s =
+let add_new_coercion coe s =
coercion_tab := Gmap.add coe s !coercion_tab
let add_new_path x y =
inheritance_graph := Gmap.add x y !inheritance_graph
let init () =
- class_tab:= Bijint.empty;
+ class_tab:= Bijint.empty;
add_new_class CL_FUN { cl_param = 0 };
add_new_class CL_SORT { cl_param = 0 };
coercion_tab:= Gmap.empty;
inheritance_graph:= Gmap.empty
-let _ =
+let _ =
Summary.declare_summary "inh_graph"
{ Summary.freeze_function = freeze;
Summary.unfreeze_function = unfreeze;
- Summary.init_function = init;
- Summary.survive_module = false;
- Summary.survive_section = false }
+ Summary.init_function = init }
let _ = init()
@@ -138,8 +134,6 @@ let coercion_info coe = Gmap.find coe !coercion_tab
let coercion_exists coe = Gmap.mem coe !coercion_tab
-let coercion_params coe_info = coe_info.coe_param
-
(* find_class_type : env -> evar_map -> constr -> cl_typ * int *)
let find_class_type env sigma t =
@@ -157,12 +151,12 @@ let subst_cl_typ subst ct = match ct with
CL_SORT
| CL_FUN
| CL_SECVAR _ -> ct
- | CL_CONST kn ->
- let kn',t = subst_con subst kn in
+ | 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)
| CL_IND (kn,i) ->
- let kn' = subst_kn subst kn in
+ let kn' = subst_ind subst kn in
if kn' == kn then ct else
CL_IND (kn',i)
@@ -172,15 +166,15 @@ let subst_coe_typ subst t = fst (subst_global subst t)
(* class_of : Term.constr -> int *)
-let class_of env sigma t =
- let (t, n1, i, args) =
+let class_of env sigma t =
+ let (t, n1, i, args) =
try
let (cl,args) = find_class_type env 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 env sigma t in
let (i, { cl_param = n1 } ) = class_info cl in
(t, n1, i, args)
in
@@ -224,7 +218,7 @@ let apply_on_class_of env sigma t cont =
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 env sigma t in
let (i, { cl_param = n1 } ) = class_info cl in
if List.length args <> n1 then raise Not_found;
t, cont i
@@ -239,7 +233,7 @@ let lookup_path_between env sigma (s,t) =
let lookup_path_to_fun_from env sigma s =
apply_on_class_of env sigma s lookup_path_to_fun_from_class
-let lookup_path_to_sort_from env sigma s =
+let lookup_path_to_sort_from env sigma s =
apply_on_class_of env sigma s lookup_path_to_sort_from_class
let get_coercion_constructor coe =
@@ -247,7 +241,7 @@ let get_coercion_constructor coe =
Reductionops.whd_betadeltaiota_stack (Global.env()) Evd.empty coe.coe_value
in
match kind_of_term c with
- | Construct cstr ->
+ | Construct cstr ->
(cstr, Inductiveops.constructor_nrealargs (Global.env()) cstr -1)
| _ ->
raise Not_found
@@ -269,14 +263,14 @@ let path_printer = ref (fun _ -> str "<a class path>"
: (int * int) * inheritance_path -> std_ppcmds)
let install_path_printer f = path_printer := f
-
+
let print_path x = !path_printer x
-let message_ambig l =
+let message_ambig l =
(str"Ambiguous paths:" ++ spc () ++
prlist_with_sep pr_fnl (fun ijp -> print_path ijp) l)
-(* add_coercion_in_graph : coe_index * cl_index * cl_index -> unit
+(* add_coercion_in_graph : coe_index * cl_index * cl_index -> unit
coercion,source,target *)
let different_class_params i j =
@@ -287,7 +281,7 @@ let add_coercion_in_graph (ic,source,target) =
let ambig_paths =
(ref [] : ((cl_index * cl_index) * inheritance_path) list ref) in
let try_add_new_path (i,j as ij) p =
- try
+ try
if i=j then begin
if different_class_params i j then begin
let _ = lookup_path_between_class ij in
@@ -303,26 +297,26 @@ let add_coercion_in_graph (ic,source,target) =
true
end
in
- let try_add_new_path1 ij p =
- let _ = try_add_new_path ij p in ()
+ let try_add_new_path1 ij p =
+ let _ = try_add_new_path ij p in ()
in
if try_add_new_path (source,target) [ic] then begin
- Gmap.iter
+ Gmap.iter
(fun (s,t) p ->
if s<>t then begin
if t = source then begin
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) & (p <> q) then
try_add_new_path1 (s,v) (p@[ic]@q))
old_inheritance_graph
end;
if s = target then try_add_new_path1 (source,t) (ic::p)
end)
- old_inheritance_graph
+ old_inheritance_graph
end;
- if (!ambig_paths <> []) && is_verbose () then
+ if (!ambig_paths <> []) && is_verbose () then
ppnl (message_ambig !ambig_paths)
type coercion = coe_typ * locality * bool * cl_typ * cl_typ * int
@@ -349,7 +343,7 @@ let load_coercion i (_,(coe,stre,isid,cls,clt,ps)) =
add_class clt;
let is,_ = class_info cls in
let it,_ = class_info clt in
- let xf =
+ let xf =
{ coe_value = constr_of_global coe;
coe_type = Global.type_of_global coe;
coe_strength = stre;
@@ -361,7 +355,7 @@ let load_coercion i (_,(coe,stre,isid,cls,clt,ps)) =
let cache_coercion o =
load_coercion 1 o
-let subst_coercion (_,subst,(coe,stre,isid,cls,clt,ps as obj)) =
+let subst_coercion (subst,(coe,stre,isid,cls,clt,ps as obj)) =
let coe' = subst_coe_typ subst coe in
let cls' = subst_cl_typ subst cls in
let clt' = subst_cl_typ subst clt in
@@ -374,7 +368,7 @@ let discharge_cl = function
| cl -> cl
let discharge_coercion (_,(coe,stre,isid,cls,clt,ps)) =
- if stre = Local then None else
+ if stre = Local then None else
let n = try Array.length (Lib.section_instance coe) with Not_found -> 0 in
Some (Lib.discharge_global coe,
stre,
@@ -383,21 +377,20 @@ let discharge_coercion (_,(coe,stre,isid,cls,clt,ps)) =
discharge_cl clt,
n + ps)
-let (inCoercion,outCoercion) =
- declare_object {(default_object "COERCION") with
+let classify_coercion (coe,stre,isid,cls,clt,ps as obj) =
+ if stre = Local then Dispose else Substitute obj
+
+let (inCoercion,_) =
+ declare_object {(default_object "COERCION") with
load_function = load_coercion;
cache_function = cache_coercion;
subst_function = subst_coercion;
- classify_function = (fun (_,x) -> Substitute x);
- discharge_function = discharge_coercion;
- export_function = (function x -> Some x) }
+ classify_function = classify_coercion;
+ discharge_function = discharge_coercion }
let declare_coercion coef stre ~isid ~src:cls ~target:clt ~params:ps =
Lib.add_anonymous_leaf (inCoercion (coef,stre,isid,cls,clt,ps))
-let coercion_strength v = v.coe_strength
-let coercion_identity v = v.coe_is_identity
-
(* For printing purpose *)
let get_coercion_value v = v.coe_value
@@ -410,7 +403,7 @@ let inheritance_graph () = Gmap.to_list !inheritance_graph
let coercion_of_reference r =
let ref = Nametab.global r in
if not (coercion_exists ref) then
- errorlabstrm "try_add_coercion"
+ errorlabstrm "try_add_coercion"
(Nametab.pr_global_env Idset.empty ref ++ str" is not a coercion.");
ref
@@ -420,7 +413,7 @@ module CoercionPrinting =
let encode = coercion_of_reference
let subst = subst_coe_typ
let printer x = pr_global_env Idset.empty x
- let key = Goptions.SecondaryTable ("Printing","Coercion")
+ let key = ["Printing";"Coercion"]
let title = "Explicitly printed coercions: "
let member_message x b =
str "Explicit printing of coercion " ++ printer x ++
diff --git a/pretyping/classops.mli b/pretyping/classops.mli
index a76fe75c..63d5b0a4 100644
--- a/pretyping/classops.mli
+++ b/pretyping/classops.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: classops.mli 11343 2008-09-01 20:55:13Z herbelin $ i*)
+(*i $Id$ i*)
(*i*)
open Names
@@ -19,9 +19,9 @@ open Mod_subst
(*i*)
(*s This is the type of class kinds *)
-type cl_typ =
- | CL_SORT
- | CL_FUN
+type cl_typ =
+ | CL_SORT
+ | CL_FUN
| CL_SECVAR of variable
| CL_CONST of constant
| CL_IND of inductive
@@ -36,7 +36,7 @@ type cl_info_typ = {
type coe_typ = Libnames.global_reference
(* This is the type of infos for declared coercions *)
-type coe_info_typ
+type coe_info_typ
(* [cl_index] is the type of class keys *)
type cl_index
@@ -65,7 +65,7 @@ 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 *)
-val declare_coercion :
+val declare_coercion :
coe_typ -> locality -> isid:bool ->
src:cl_typ -> target:cl_typ -> params:int -> unit
@@ -77,18 +77,18 @@ val coercion_value : coe_index -> (unsafe_judgment * bool)
(*s 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 ->
+val lookup_path_between : env -> evar_map -> types * types ->
types * types * inheritance_path
val lookup_path_to_fun_from : env -> evar_map -> types ->
types * inheritance_path
-val lookup_path_to_sort_from : env -> evar_map -> types ->
+val lookup_path_to_sort_from : env -> evar_map -> types ->
types * inheritance_path
-val lookup_pattern_path_between :
+val lookup_pattern_path_between :
inductive * inductive -> (constructor * int) list
(*i Crade *)
open Pp
-val install_path_printer :
+val install_path_printer :
((cl_index * cl_index) * inheritance_path -> std_ppcmds) -> unit
(*i*)
diff --git a/pretyping/clenv.ml b/pretyping/clenv.ml
index 0bfef04a..283d3f12 100644
--- a/pretyping/clenv.ml
+++ b/pretyping/clenv.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: clenv.ml 13126 2010-06-13 11:09:51Z herbelin $ *)
+(* $Id$ *)
open Pp
open Util
@@ -14,6 +14,7 @@ open Names
open Nameops
open Term
open Termops
+open Namegen
open Sign
open Environ
open Evd
@@ -31,8 +32,8 @@ 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_hnf_constr gls c = hnf_constr (pf_env gls) gls.sigma c
let pf_concl gl = gl.it.evar_concl
(******************************************************************)
@@ -40,14 +41,14 @@ let pf_concl gl = gl.it.evar_concl
type clausenv = {
env : env;
- evd : evar_defs;
+ evd : evar_map;
templval : constr freelisted;
templtyp : constr freelisted }
let cl_env ce = ce.env
-let cl_sigma ce = evars_of ce.evd
+let cl_sigma ce = ce.evd
-let subst_clenv sub clenv =
+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;
@@ -62,8 +63,7 @@ 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_with_meta (cl_env ce) (cl_sigma ce) (metas_of ce.evd) c
+let clenv_get_type_of ce c = Retyping.get_type_of (cl_env ce) (cl_sigma ce) c
exception NotExtensibleClause
@@ -102,7 +102,7 @@ let clenv_environments evd bound t =
(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
+ in
clrec (evd,[]) bound t
(* Instantiate the first [bound] products of [t] with evars (all products if
@@ -120,7 +120,7 @@ let clenv_environments_evars env evd bound t =
(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
+ in
clrec (evd,[]) bound t
let clenv_conv_leq env sigma t c bound =
@@ -129,16 +129,16 @@ let clenv_conv_leq env sigma t c bound =
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 (Evd.evars_of evars)) args 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 (env,args,concl) = clenv_environments evd n cty 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 = env;
+ evd = evd;
env = environ }
let mk_clenv_from_n gls n (c,cty) =
@@ -146,8 +146,8 @@ let mk_clenv_from_n gls 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_var (pf_env gls) [] t)
+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)
@@ -168,21 +168,19 @@ let mentions clenv mv0 =
meta_exists menrec mlist
in menrec
-let clenv_defined clenv mv = meta_defined clenv.evd mv
-
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 " ++
+ (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 ? *)
+(* TODO: replace by clenv_unify (mkMeta mv) rhs ? *)
let clenv_assign mv rhs clenv =
- let rhs_fls = mk_freelisted rhs in
+ let rhs_fls = mk_freelisted rhs in
if meta_exists (mentions clenv mv) rhs_fls.freemetas then
error "clenv_assign: circularity in unification";
try
@@ -191,15 +189,13 @@ let clenv_assign mv rhs clenv =
error_incompatible_inst clenv mv
else
clenv
- else
+ else
let st = (ConvUpToEta 0,TypeNotProcessed) in
{clenv with evd = meta_assign mv (rhs_fls.rebus,st) clenv.evd}
- with Not_found ->
+ with Not_found ->
error "clenv_assign: undefined meta"
-let clenv_wtactic f clenv = {clenv with evd = f clenv.evd }
-
(* [clenv_dependent hyps_only clenv]
* returns a list of the metavars which appear in the template of clenv,
@@ -222,7 +218,7 @@ let dependent_metas clenv mvs conclmetas =
Metaset.union deps (clenv_metavars clenv.evd mv))
mvs conclmetas
-let duplicated_metas c =
+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)
@@ -237,9 +233,9 @@ let clenv_dependent hyps_only clenv =
let nonlinear = duplicated_metas (clenv_value clenv) in
(* Make the assumption that duplicated metas have internal dependencies *)
List.filter
- (fun mv -> (Metaset.mem mv deps &&
- not (hyps_only && Metaset.mem mv ctyp_mvs))
- or List.mem mv nonlinear)
+ (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
@@ -254,7 +250,7 @@ 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 (evars_of clenv.evd) clenv.templtyp.rebus)) then
+ 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
@@ -265,7 +261,7 @@ let clenv_unique_resolver allow_K ?(flags=default_unify_flags) clenv gl =
* 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.
+ * 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
@@ -277,7 +273,7 @@ let clenv_unique_resolver allow_K ?(flags=default_unify_flags) clenv gl =
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
+ 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,
@@ -291,13 +287,13 @@ let clenv_unique_resolver allow_K ?(flags=default_unify_flags) clenv gl =
let clenv_pose_metas_as_evars clenv dep_mvs =
let rec fold clenv = function
| [] -> clenv
- | mv::mvs ->
+ | 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) =
+ 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
@@ -321,9 +317,9 @@ let connect_clenv gls clenv =
* 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.
+ * terms, and make bindings.
- Otherwise said, from
+ Otherwise said, from
[clenv] = [env;sigma;metas |- c:T]
[clenv'] = [env';sigma';metas' |- d:U]
@@ -340,8 +336,7 @@ let clenv_fchain ?(allow_K=true) ?(flags=default_unify_flags) mv clenv nextclenv
let clenv' =
{ templval = clenv.templval;
templtyp = clenv.templtyp;
- evd =
- evar_merge (meta_merge clenv.evd nextclenv.evd) (evars_of clenv.evd);
+ 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'' =
@@ -352,13 +347,13 @@ let clenv_fchain ?(allow_K=true) ?(flags=default_unify_flags) mv clenv nextclenv
(* assign the metavar *)
let clenv''' =
clenv_assign mv (clenv_term clenv' nextclenv.templval) clenv''
- in
+ in
clenv'''
(***************************************************************)
(* Bindings *)
-type arg_bindings = open_constr explicit_bindings
+type arg_bindings = constr explicit_bindings
(* [clenv_independent clenv]
* returns a list of metavariables which appear in the term cval,
@@ -374,9 +369,9 @@ let clenv_independent clenv =
let check_bindings bl =
match list_duplicates (List.map pi2 bl) with
- | NamedHyp s :: _ ->
+ | NamedHyp s :: _ ->
errorlabstrm ""
- (str "The variable " ++ pr_id s ++
+ (str "The variable " ++ pr_id s ++
str " occurs more than once in binding list.");
| AnonHyp n :: _ ->
errorlabstrm ""
@@ -402,7 +397,7 @@ let error_already_defined b =
(str "Position " ++ int n ++ str" already defined.")
let clenv_unify_binding_type clenv c t u =
- if isMeta (fst (whd_stack (evars_of clenv.evd) u)) then
+ if isMeta (fst (whd_stack clenv.evd u)) then
(* Not enough information to know if some subtyping is needed *)
CoerceToType, clenv, c
else
@@ -410,16 +405,16 @@ let clenv_unify_binding_type clenv c t u =
try
let evd,c = w_coerce_to_type (cl_env clenv) clenv.evd c t u in
TypeProcessed, { clenv with evd = evd }, c
- with
+ with
| PretypeError (_,NotClean _) as e -> raise e
- | e when precatchable_exception e -> TypeNotProcessed, clenv, c
+ | e when precatchable_exception e ->
+ TypeNotProcessed, clenv, c
-let clenv_assign_binding clenv k (sigma,c) =
+let clenv_assign_binding clenv k c =
let k_typ = clenv_hnf_constr clenv (clenv_meta_type clenv k) in
- let clenv' = { clenv with evd = evar_merge clenv.evd sigma} in
- let c_typ = nf_betaiota (evars_of 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 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
@@ -428,13 +423,13 @@ let clenv_match_args bl clenv =
let mvs = clenv_independent clenv in
check_bindings bl;
List.fold_left
- (fun clenv (loc,b,(sigma,c as sc)) ->
+ (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 sc)
+ clenv_assign_binding clenv k c)
clenv bl
exception NoSuchBinding
@@ -442,7 +437,7 @@ 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 (Evd.empty,c)
+ clenv_assign_binding clenv k c
let clenv_constrain_dep_args hyps_only bl clenv =
if bl = [] then
@@ -451,8 +446,8 @@ let clenv_constrain_dep_args hyps_only bl clenv =
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 ""
+ else
+ errorlabstrm ""
(strbrk "Not the right number of missing arguments (expected " ++
int (List.length occlist) ++ str ").")
@@ -479,4 +474,4 @@ let pr_clenv clenv =
h 0
(str"TEMPL: " ++ print_constr clenv.templval.rebus ++
str" : " ++ print_constr clenv.templtyp.rebus ++ fnl () ++
- pr_evar_defs clenv.evd)
+ pr_evar_map clenv.evd)
diff --git a/pretyping/clenv.mli b/pretyping/clenv.mli
index 4c5535b3..62dfa7b5 100644
--- a/pretyping/clenv.mli
+++ b/pretyping/clenv.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: clenv.mli 13126 2010-06-13 11:09:51Z herbelin $ i*)
+(*i $Id$ i*)
(*i*)
open Util
@@ -32,7 +32,7 @@ open Unification
*)
type clausenv = {
env : env;
- evd : evar_defs;
+ evd : evar_map;
templval : constr freelisted;
templtyp : constr freelisted }
@@ -60,14 +60,14 @@ val mk_clenv_from_env : env -> evar_map -> int option -> constr * types -> claus
(* linking of clenvs *)
val connect_clenv : evar_info sigma -> clausenv -> clausenv
-val clenv_fchain :
+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 :
+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 *)
@@ -86,7 +86,7 @@ val clenv_pose_metas_as_evars : clausenv -> metavariable list -> clausenv
(***************************************************************)
(* Bindings *)
-type arg_bindings = open_constr explicit_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
@@ -109,10 +109,10 @@ val clenv_unify_meta_types : ?flags:unify_flags -> clausenv -> clausenv
(* 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 -> open_constr bindings ->
+ evar_info sigma -> int option -> constr * constr -> constr bindings ->
clausenv
val make_clenv_binding :
- evar_info sigma -> constr * constr -> open_constr bindings -> clausenv
+ 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
@@ -124,12 +124,12 @@ val make_clenv_binding :
[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_defs -> int option -> types -> evar_defs * constr list * types
+ 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_defs -> int option -> types -> evar_defs * constr list * types
+ 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 :
diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml
index 73fcd0ea..9f0b4352 100644
--- a/pretyping/coercion.ml
+++ b/pretyping/coercion.ml
@@ -5,7 +5,7 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: coercion.ml 11343 2008-09-01 20:55:13Z herbelin $ *)
+(* $Id$ *)
open Util
open Names
@@ -24,43 +24,43 @@ open Termops
module type S = sig
(*s Coercions. *)
-
+
(* [inh_app_fun env evd j] coerces [j] to a function; i.e. it
inserts a coercion into [j], if needed, in such a way it gets as
type a product; it returns [j] if no coercion is applicable *)
val inh_app_fun :
- env -> evar_defs -> unsafe_judgment -> evar_defs * unsafe_judgment
-
+ env -> evar_map -> unsafe_judgment -> evar_map * unsafe_judgment
+
(* [inh_coerce_to_sort env evd j] coerces [j] to a type; i.e. it
inserts a coercion into [j], if needed, in such a way it gets as
type a sort; it fails if no coercion is applicable *)
val inh_coerce_to_sort : loc ->
- env -> evar_defs -> unsafe_judgment -> evar_defs * unsafe_type_judgment
+ env -> evar_map -> unsafe_judgment -> evar_map * unsafe_type_judgment
(* [inh_coerce_to_base env evd j] coerces [j] to its base type; i.e. it
inserts a coercion into [j], if needed, in such a way it gets as
type its base type (the notion depends on the coercion system) *)
val inh_coerce_to_base : loc ->
- env -> evar_defs -> unsafe_judgment -> evar_defs * unsafe_judgment
-
- (* [inh_coerce_to_prod env isevars t] coerces [t] to a product type *)
+ env -> evar_map -> unsafe_judgment -> evar_map * unsafe_judgment
+
+ (* [inh_coerce_to_prod env evars t] coerces [t] to a product type *)
val inh_coerce_to_prod : loc ->
- env -> evar_defs -> type_constraint_type -> evar_defs * type_constraint_type
+ env -> evar_map -> type_constraint_type -> evar_map * type_constraint_type
- (* [inh_conv_coerce_to loc env evd j t] coerces [j] to an object of type
+ (* [inh_conv_coerce_to loc env evd j t] coerces [j] to an object of type
[t]; i.e. it inserts a coercion into [j], if needed, in such a way [t] and
[j.uj_type] are convertible; it fails if no coercion is applicable *)
- val inh_conv_coerce_to : loc ->
- env -> evar_defs -> unsafe_judgment -> type_constraint_type -> evar_defs * unsafe_judgment
+ val inh_conv_coerce_to : loc ->
+ env -> evar_map -> unsafe_judgment -> type_constraint_type -> evar_map * unsafe_judgment
- val inh_conv_coerce_rigid_to : loc ->
- env -> evar_defs -> unsafe_judgment -> type_constraint_type -> evar_defs * unsafe_judgment
+ 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 evd t t'] checks if an object of type [t]
is coercible to an object of type [t'] adding evar constraints if needed;
it fails if no coercion exists *)
- val inh_conv_coerces_to : loc ->
- env -> evar_defs -> types -> type_constraint_type -> evar_defs
+ val inh_conv_coerces_to : loc ->
+ env -> evar_map -> types -> type_constraint_type -> evar_map
(* [inh_pattern_coerce_to loc env evd pat ind1 ind2] coerces the Cases
pattern [pat] typed in [ind1] into a pattern typed in [ind2];
@@ -73,24 +73,19 @@ module Default = struct
(* Typing operations dealing with coercions *)
exception NoCoercion
- let whd_app_evar sigma t =
- match kind_of_term t with
- | App (f,l) -> mkApp (whd_evar sigma f,l)
- | _ -> whd_evar sigma t
-
(* Here, funj is a coercion therefore already typed in global context *)
let apply_coercion_args env argl funj =
let rec apply_rec acc typ = function
| [] -> { uj_val = applist (j_val funj,argl);
uj_type = typ }
| h::restl ->
- (* On devrait pouvoir s'arranger pour qu'on n'ait pas à faire hnf_constr *)
+ (* On devrait pouvoir s'arranger pour qu'on n'ait pas à faire hnf_constr *)
match kind_of_term (whd_betadeltaiota env Evd.empty typ) with
- | Prod (_,c1,c2) ->
+ | Prod (_,c1,c2) ->
(* Typage garanti par l'appel à app_coercion*)
apply_rec (h::acc) (subst1 h c2) restl
| _ -> anomaly "apply_coercion_args"
- in
+ in
apply_rec [] funj.uj_type argl
(* appliquer le chemin de coercions de patterns p *)
@@ -106,47 +101,55 @@ module Default = struct
let p = lookup_pattern_path_between (ind1,ind2) in
apply_pattern_coercion loc pat p
+ let saturate_evd env evd =
+ Typeclasses.resolve_typeclasses
+ ~onlyargs:true ~split:true ~fail:false env evd
+
(* appliquer le chemin de coercions p à hj *)
let apply_coercion env sigma p hj typ_cl =
- try
+ try
fst (List.fold_left
- (fun (ja,typ_cl) i ->
+ (fun (ja,typ_cl) i ->
let fv,isid = coercion_value i in
let argl = (class_args_of env sigma typ_cl)@[ja.uj_val] in
let jres = apply_coercion_args env argl fv in
- (if isid then
+ (if isid then
{ uj_val = ja.uj_val; uj_type = jres.uj_type }
- else
+ else
jres),
jres.uj_type)
(hj,typ_cl) p)
with _ -> anomaly "apply_coercion"
- let inh_app_fun env evd j =
- let t = whd_betadeltaiota env (evars_of evd) j.uj_type in
+ let inh_app_fun env evd j =
+ let t = whd_betadeltaiota env evd j.uj_type in
match kind_of_term t with
| Prod (_,_,_) -> (evd,j)
| Evar ev ->
let (evd',t) = define_evar_as_product evd ev in
(evd',{ uj_val = j.uj_val; uj_type = t })
| _ ->
- (try
- let t,p =
- lookup_path_to_fun_from env (evars_of evd) j.uj_type in
- (evd,apply_coercion env (evars_of evd) p j t)
- with Not_found -> (evd,j))
+ let t,p =
+ lookup_path_to_fun_from env ( evd) j.uj_type in
+ (evd,apply_coercion env ( evd) p j t)
+
+ let inh_app_fun env evd j =
+ try inh_app_fun env evd j
+ with Not_found ->
+ try inh_app_fun env (saturate_evd env evd) j
+ with Not_found -> (evd, j)
let inh_tosort_force loc env evd j =
try
- let t,p = lookup_path_to_sort_from env (evars_of evd) j.uj_type in
- let j1 = apply_coercion env (evars_of evd) p j t in
- let j2 = on_judgment_type (whd_evar (evars_of 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 (evars_of 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 (evars_of 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) ->
@@ -156,7 +159,6 @@ module Default = struct
inh_tosort_force loc env evd j
let inh_coerce_to_base loc env evd j = (evd, j)
-
let inh_coerce_to_prod loc env evd t = (evd, t)
let inh_coerce_to_fail env evd rigidonly v t c1 =
@@ -165,16 +167,16 @@ module Default = struct
raise NoCoercion
else
let v', t' =
- try
- let t2,t1,p = lookup_path_between env (evars_of evd) (t,c1) in
+ try
+ let t2,t1,p = lookup_path_between env evd (t,c1) in
match v with
- Some v ->
+ Some v ->
let j =
- apply_coercion env (evars_of evd) p
+ apply_coercion env evd p
{uj_val = v; uj_type = t} t2 in
Some j.uj_val, j.uj_type
| None -> None, t
- with Not_found -> raise NoCoercion
+ with Not_found -> raise NoCoercion
in
try (the_conv_x_leq env t' c1 evd, v')
with Reduction.NotConvertible -> raise NoCoercion
@@ -185,15 +187,15 @@ module Default = struct
try inh_coerce_to_fail env evd rigidonly v t c1
with NoCoercion ->
match
- kind_of_term (whd_betadeltaiota env (evars_of evd) t),
- kind_of_term (whd_betadeltaiota env (evars_of evd) c1)
+ kind_of_term (whd_betadeltaiota env evd t),
+ kind_of_term (whd_betadeltaiota env evd c1)
with
- | Prod (name,t1,t2), Prod (_,u1,u2) ->
+ | Prod (name,t1,t2), Prod (_,u1,u2) ->
(* Conversion did not work, we may succeed with a coercion. *)
(* We eta-expand (hence possibly modifying the original term!) *)
(* and look for a coercion c:u1->t1 s.t. fun x:u1 => v' (c x)) *)
(* has type forall (x:u1), u2 (with v' recursively obtained) *)
- let name = match name with
+ let name = match name with
| Anonymous -> Name (id_of_string "x")
| _ -> name in
let env1 = push_rel (name,None,u1) env in
@@ -211,12 +213,15 @@ module Default = struct
let inh_conv_coerce_to_gen rigidonly loc env evd cj (n, t) =
match n with
None ->
- let (evd', val') =
- try
+ let (evd', val') =
+ try
inh_conv_coerce_to_fail loc env evd rigidonly (Some cj.uj_val) cj.uj_type t
with NoCoercion ->
- let sigma = evars_of evd in
- error_actual_type_loc loc env sigma cj t
+ let evd = saturate_evd env evd in
+ try
+ inh_conv_coerce_to_fail loc env evd rigidonly (Some cj.uj_val) cj.uj_type t
+ with NoCoercion ->
+ error_actual_type_loc loc env evd cj t
in
let val' = match val' with Some v -> v | None -> assert(false) in
(evd',{ uj_val = val'; uj_type = t })
@@ -225,19 +230,19 @@ module Default = struct
let inh_conv_coerce_to = inh_conv_coerce_to_gen false
let inh_conv_coerce_rigid_to = inh_conv_coerce_to_gen true
-
- let inh_conv_coerces_to loc env (evd : evar_defs) t (abs, t') = evd
- (* Still problematic, as it changes unification
- let nabsinit, nabs =
+
+ let inh_conv_coerces_to loc env (evd : evar_map) t (abs, t') = evd
+ (* Still problematic, as it changes unification
+ let nabsinit, nabs =
match abs with
None -> 0, 0
| Some (init, cur) -> init, cur
in
- try
- let (rels, rng) =
- (* a little more effort to get products is needed *)
+ try
+ let (rels, rng) =
+ (* a little more effort to get products is needed *)
try decompose_prod_n nabs t
- with _ ->
+ with _ ->
if !Flags.debug then
msg_warning (str "decompose_prod_n failed");
raise (Invalid_argument "Coercion.inh_conv_coerces_to")
@@ -245,15 +250,15 @@ module Default = struct
(* The final range free variables must have been replaced by evars, we accept only that evars
in rng are applied to free vars. *)
if noccur_with_meta 0 (succ nabsinit) rng then (
- let env', t, t' =
+ let env', t, t' =
let env' = List.fold_right (fun (n, t) env -> push_rel (n, None, t) env) rels env in
env', rng, lift nabs t'
in
- try
+ try
pi1 (inh_conv_coerce_to_fail loc env' evd None t t')
with NoCoercion ->
evd) (* Maybe not enough information to unify *)
- (*let sigma = evars_of evd in
+ (*let sigma = evd in
error_cannot_coerce env' sigma (t, t'))*)
else evd
with Invalid_argument _ -> evd *)
diff --git a/pretyping/coercion.mli b/pretyping/coercion.mli
index 8705080f..565cf0c4 100644
--- a/pretyping/coercion.mli
+++ b/pretyping/coercion.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: coercion.mli 11282 2008-07-28 11:51:53Z msozeau $ i*)
+(*i $Id$ i*)
(*i*)
open Util
@@ -21,44 +21,44 @@ open Rawterm
module type S = sig
(*s Coercions. *)
-
+
(* [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_defs -> unsafe_judgment -> evar_defs * unsafe_judgment
-
+ env -> evar_map -> unsafe_judgment -> evar_map * unsafe_judgment
+
(* [inh_coerce_to_sort env isevars j] coerces [j] to a type; i.e. it
inserts a coercion into [j], if needed, in such a way it gets as
type a sort; it fails if no coercion is applicable *)
val inh_coerce_to_sort : loc ->
- env -> evar_defs -> unsafe_judgment -> evar_defs * unsafe_type_judgment
+ env -> evar_map -> unsafe_judgment -> evar_map * unsafe_type_judgment
(* [inh_coerce_to_base env isevars j] coerces [j] to its base type; i.e. it
inserts a coercion into [j], if needed, in such a way it gets as
type its base type (the notion depends on the coercion system) *)
val inh_coerce_to_base : loc ->
- env -> evar_defs -> unsafe_judgment -> evar_defs * unsafe_judgment
+ env -> evar_map -> unsafe_judgment -> evar_map * unsafe_judgment
(* [inh_coerce_to_prod env isevars t] coerces [t] to a product type *)
val inh_coerce_to_prod : loc ->
- env -> evar_defs -> type_constraint_type -> evar_defs * type_constraint_type
-
- (* [inh_conv_coerce_to loc env isevars j t] coerces [j] to an object of type
+ 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
[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 ->
- env -> evar_defs -> unsafe_judgment -> type_constraint_type -> evar_defs * unsafe_judgment
+ val inh_conv_coerce_to : loc ->
+ env -> evar_map -> unsafe_judgment -> type_constraint_type -> evar_map * unsafe_judgment
+
+ val inh_conv_coerce_rigid_to : loc ->
+ env -> evar_map -> unsafe_judgment -> type_constraint_type -> evar_map * unsafe_judgment
- val inh_conv_coerce_rigid_to : loc ->
- env -> evar_defs -> unsafe_judgment -> type_constraint_type -> evar_defs * unsafe_judgment
-
(* [inh_conv_coerces_to loc env isevars t t'] checks if an object of type [t]
is coercible to an object of type [t'] adding evar constraints if needed;
it fails if no coercion exists *)
- val inh_conv_coerces_to : loc ->
- env -> evar_defs -> types -> type_constraint_type -> evar_defs
-
+ val inh_conv_coerces_to : loc ->
+ env -> evar_map -> types -> type_constraint_type -> evar_map
+
(* [inh_pattern_coerce_to loc env isevars pat ind1 ind2] coerces the Cases
pattern [pat] typed in [ind1] into a pattern typed in [ind2];
raises [Not_found] if no coercion found *)
diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml
index 20cbba94..9552fc24 100644
--- a/pretyping/detyping.ml
+++ b/pretyping/detyping.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: detyping.ml 12887 2010-03-27 15:57:02Z herbelin $ *)
+(* $Id$ *)
open Pp
open Util
@@ -21,6 +21,7 @@ open Sign
open Rawterm
open Nameops
open Termops
+open Namegen
open Libnames
open Nametab
open Evd
@@ -32,7 +33,7 @@ let dl = dummy_loc
(* Tools for printing of Cases *)
let encode_inductive r =
- let indsp = inductive_of_reference r in
+ let indsp = global_inductive r in
let constr_lengths = mis_constr_nargs indsp in
(indsp,constr_lengths)
@@ -60,7 +61,7 @@ let encode_tuple r =
x
module PrintingCasesMake =
- functor (Test : sig
+ functor (Test : sig
val encode : reference -> inductive * int array
val member_message : std_ppcmds -> bool -> std_ppcmds
val field : string
@@ -70,33 +71,33 @@ module PrintingCasesMake =
type t = inductive * int array
let encode = Test.encode
let subst subst ((kn,i), ints as obj) =
- let kn' = subst_kn subst kn in
+ 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)
- let key = Goptions.SecondaryTable ("Printing",Test.field)
+ let key = ["Printing";Test.field]
let title = Test.title
let member_message x = Test.member_message (printer x)
let synchronous = true
end
module PrintingCasesIf =
- PrintingCasesMake (struct
+ PrintingCasesMake (struct
let encode = encode_bool
let field = "If"
let title = "Types leading to pretty-printing of Cases using a `if' form: "
let member_message s b =
- str "Cases on elements of " ++ s ++
+ str "Cases on elements of " ++ s ++
str
(if b then " are printed using a `if' form"
else " are not printed using a `if' form")
end)
module PrintingCasesLet =
- PrintingCasesMake (struct
+ PrintingCasesMake (struct
let encode = encode_tuple
let field = "Let"
- let title =
+ let title =
"Types leading to a pretty-printing of Cases using a `let' form:"
let member_message s b =
str "Cases on elements of " ++ s ++
@@ -115,30 +116,30 @@ open Goptions
let wildcard_value = ref true
let force_wildcard () = !wildcard_value
-let _ = declare_bool_option
+let _ = declare_bool_option
{ optsync = true;
optname = "forced wildcard";
- optkey = SecondaryTable ("Printing","Wildcard");
+ optkey = ["Printing";"Wildcard"];
optread = force_wildcard;
optwrite = (:=) wildcard_value }
let synth_type_value = ref true
let synthetize_type () = !synth_type_value
-let _ = declare_bool_option
+let _ = declare_bool_option
{ optsync = true;
optname = "pattern matching return type synthesizability";
- optkey = SecondaryTable ("Printing","Synth");
+ optkey = ["Printing";"Synth"];
optread = synthetize_type;
optwrite = (:=) synth_type_value }
let reverse_matching_value = ref true
let reverse_matching () = !reverse_matching_value
-let _ = declare_bool_option
+let _ = declare_bool_option
{ optsync = true;
optname = "pattern-matching reversibility";
- optkey = SecondaryTable ("Printing","Matching");
+ optkey = ["Printing";"Matching"];
optread = reverse_matching;
optwrite = (:=) reverse_matching_value }
@@ -162,52 +163,48 @@ let computable p k =
sinon on perd la réciprocité de la synthèse (qui, lui,
engendrera un prédicat non dépendant) *)
- (nb_lam p = k+1)
+ let sign,ccl = decompose_lam_assum p in
+ (rel_context_length sign = k+1)
&&
- let _,ccl = decompose_lam p in
noccur_between 1 (k+1) ccl
let avoid_flag isgoal = if isgoal then Some true else None
-
-let lookup_name_as_renamed env t s =
- let rec lookup avoid env_names n c = match kind_of_term c with
+
+let lookup_name_as_displayed env t s =
+ let rec lookup avoid n c = match kind_of_term c with
| Prod (name,_,c') ->
- (match concrete_name (Some true) avoid env_names name c' with
- | (Name id,avoid') ->
- if id=s then (Some n)
- else lookup avoid' (add_name (Name id) env_names) (n+1) c'
- | (Anonymous,avoid') -> lookup avoid' env_names (n+1) (pop c'))
+ (match compute_displayed_name_in RenamingForGoal avoid name c' with
+ | (Name id,avoid') -> if id=s then Some n else lookup avoid' (n+1) c'
+ | (Anonymous,avoid') -> lookup avoid' (n+1) (pop c'))
| LetIn (name,_,_,c') ->
- (match concrete_name (Some true) avoid env_names name c' with
- | (Name id,avoid') ->
- if id=s then (Some n)
- else lookup avoid' (add_name (Name id) env_names) (n+1) c'
- | (Anonymous,avoid') -> lookup avoid' env_names (n+1) (pop c'))
- | Cast (c,_,_) -> lookup avoid env_names n c
+ (match compute_displayed_name_in RenamingForGoal avoid name c' with
+ | (Name id,avoid') -> if id=s then Some n else lookup avoid' (n+1) c'
+ | (Anonymous,avoid') -> lookup avoid' (n+1) (pop c'))
+ | Cast (c,_,_) -> lookup avoid n c
| _ -> None
- in lookup (ids_of_named_context (named_context env)) empty_names_context 1 t
+ in lookup (ids_of_named_context (named_context env)) 1 t
let lookup_index_as_renamed env t n =
let rec lookup n d c = match kind_of_term c with
| Prod (name,_,c') ->
- (match concrete_name (Some true) [] empty_names_context name c' with
+ (match compute_displayed_name_in RenamingForGoal [] name c' with
(Name _,_) -> lookup n (d+1) c'
- | (Anonymous,_) ->
+ | (Anonymous,_) ->
if n=0 then
Some (d-1)
- else if n=1 then
- Some d
- else
+ else if n=1 then
+ Some d
+ else
lookup (n-1) (d+1) c')
| LetIn (name,_,_,c') ->
- (match concrete_name (Some true) [] empty_names_context name c' with
+ (match compute_displayed_name_in RenamingForGoal [] name c' with
| (Name _,_) -> lookup n (d+1) c'
- | (Anonymous,_) ->
- if n=0 then
- Some (d-1)
- else if n=1 then
- Some d
- else
+ | (Anonymous,_) ->
+ if n=0 then
+ Some (d-1)
+ else if n=1 then
+ Some d
+ else
lookup (n-1) (d+1) c'
)
| Cast (c,_,_) -> lookup n d c
@@ -225,16 +222,17 @@ let update_name na ((_,e),c) =
na
let rec decomp_branch n nal b (avoid,env as e) c =
+ let flag = if b then RenamingForGoal else RenamingForCasesPattern in
if n=0 then (List.rev nal,(e,c))
else
let na,c,f =
match kind_of_term (strip_outer_cast c) with
- | Lambda (na,_,c) -> na,c,concrete_let_name
- | LetIn (na,_,_,c) -> na,c,concrete_name
- | _ ->
- Name (id_of_string "x"),(applist (lift 1 c, [mkRel 1])),
- concrete_name in
- let na',avoid' = f (Some b) avoid env na c in
+ | Lambda (na,_,c) -> na,c,compute_displayed_let_name_in
+ | LetIn (na,_,_,c) -> na,c,compute_displayed_name_in
+ | _ ->
+ Name (id_of_string "x"),(applist (lift 1 c, [mkRel 1])),
+ compute_displayed_name_in in
+ let na',avoid' = f flag avoid na c in
decomp_branch (n-1) (na'::nal) b (avoid',add_name na' env) c
let rec build_tree na isgoal e ci cl =
@@ -248,14 +246,14 @@ and align_tree nal isgoal (e,c as rhs) = match nal with
| [] -> [[],rhs]
| na::nal ->
match kind_of_term c with
- | Case (ci,p,c,cl) when c = mkRel (list_index na (snd e))
+ | Case (ci,p,c,cl) when c = mkRel (list_index na (snd e))
& (* don't contract if p dependent *)
computable p (ci.ci_pp_info.ind_nargs) ->
let clauses = build_tree na isgoal e ci cl in
List.flatten
(List.map (fun (pat,rhs) ->
let lines = align_tree nal isgoal rhs in
- List.map (fun (hd,rest) -> pat::hd,rest) lines)
+ List.map (fun (hd,rest) -> pat::hd,rest) lines)
clauses)
| _ ->
let pat = PatVar(dl,update_name na rhs) in
@@ -294,14 +292,14 @@ let it_destRLambda_or_LetIn_names n c =
| _ ->
(* eta-expansion *)
let rec next l =
- let x = Nameops.next_ident_away (id_of_string "x") l in
+ 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 *)
x
in
- let x = next (free_rawvars c) in
+ let x = next (free_rawvars c) in
let a = RVar (dl,x) in
- aux (n-1) (Name x :: nal)
+ aux (n-1) (Name x :: nal)
(match c with
| RApp (loc,p,l) -> RApp (loc,c,l@[a])
| _ -> (RApp (dl,c,[a])))
@@ -311,16 +309,16 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl =
let (indsp,st,nparams,consnargsl,k) = data in
let synth_type = synthetize_type () in
let tomatch = detype c in
- let alias, aliastyp, pred=
- if (not !Flags.raw_print) & synth_type & computable & Array.length bl<>0
- then
+ let alias, aliastyp, pred=
+ if (not !Flags.raw_print) & synth_type & computable & Array.length bl<>0
+ then
Anonymous, None, None
else
match Option.map detype p with
| None -> Anonymous, None, None
| Some p ->
let nl,typ = it_destRLambda_or_LetIn_names k p in
- let n,typ = match typ with
+ let n,typ = match typ with
| RLambda (_,x,_,t,c) -> x, c
| _ -> Anonymous, typ in
let aliastyp =
@@ -330,21 +328,21 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl =
in
let constructs = Array.init (Array.length bl) (fun i -> (indsp,i+1)) in
let tag =
- try
+ try
if !Flags.raw_print then
RegularStyle
- else if st = LetPatternStyle then
+ else if st = LetPatternStyle then
st
else if PrintingLet.active (indsp,consnargsl) then
LetStyle
- else if PrintingIf.active (indsp,consnargsl) then
+ else if PrintingIf.active (indsp,consnargsl) then
IfStyle
- else
+ else
st
with Not_found -> st
in
match tag with
- | LetStyle when aliastyp = None ->
+ | 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)
@@ -400,7 +398,7 @@ let rec detype (isgoal:bool) avoid env t =
array_map_to_list (detype isgoal avoid env) args)
| Const sp -> RRef (dl, ConstRef sp)
| Evar (ev,cl) ->
- REvar (dl, ev,
+ REvar (dl, ev,
Some (List.map (detype isgoal avoid env) (Array.to_list cl)))
| Ind ind_sp ->
RRef (dl, IndRef ind_sp)
@@ -410,7 +408,7 @@ let rec detype (isgoal:bool) avoid env t =
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
+ 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)
(Some p) c bl
@@ -421,7 +419,7 @@ and detype_fix isgoal avoid env (vn,_ as nvn) (names,tys,bodies) =
let def_avoid, def_env, lfi =
Array.fold_left
(fun (avoid, env, l) na ->
- let id = next_name_away na avoid in
+ let id = next_name_away na avoid in
(id::avoid, add_name (Name id) env, id::l))
(avoid, env, []) names in
let n = Array.length tys in
@@ -437,7 +435,7 @@ and detype_cofix isgoal avoid env n (names,tys,bodies) =
let def_avoid, def_env, lfi =
Array.fold_left
(fun (avoid, env, l) na ->
- let id = next_name_away na avoid in
+ let id = next_name_away na avoid in
(id::avoid, add_name (Name id) env, id::l))
(avoid, env, []) names in
let ntys = Array.length tys in
@@ -456,16 +454,16 @@ and share_names isgoal n l avoid env c t =
let na = match (na,na') with
Name _, _ -> na
| _, Name _ -> na'
- | _ -> na in
+ | _ -> na in
let t = detype isgoal avoid env t in
- let id = next_name_away na avoid in
+ let id = next_name_away na avoid in
let avoid = id::avoid and env = add_name (Name id) env in
share_names isgoal (n-1) ((Name id,Explicit,None,t)::l) avoid env c c'
(* May occur for fix built interactively *)
| LetIn (na,b,t',c), _ when n > 0 ->
let t' = detype isgoal avoid env t' in
let b = detype isgoal avoid env b in
- let id = next_name_away na avoid in
+ let id = next_name_away na avoid in
let avoid = id::avoid and env = add_name (Name id) env in
share_names isgoal n ((Name id,Explicit,Some b,t')::l) avoid env c t
(* Only if built with the f/n notation or w/o let-expansion in types *)
@@ -474,7 +472,7 @@ and share_names isgoal n l avoid env c t =
(* If it is an open proof: we cheat and eta-expand *)
| _, Prod (na',t',c') when n > 0 ->
let t' = detype isgoal avoid env t' in
- let id = next_name_away na' avoid in
+ let id = next_name_away na' avoid in
let avoid = id::avoid and env = add_name (Name id) env in
let appc = mkApp (lift 1 c,[|mkRel 1|]) in
share_names isgoal (n-1) ((Name id,Explicit,None,t')::l) avoid env appc c'
@@ -499,22 +497,22 @@ and detype_eqn isgoal avoid env constr construct_nargs branch =
let make_pat x avoid env b ids =
if force_wildcard () & noccurn 1 b then
PatVar (dl,Anonymous),avoid,(add_name Anonymous env),ids
- else
+ else
let id = next_name_away_in_cases_pattern x avoid in
PatVar (dl,Name id),id::avoid,(add_name (Name id) env),id::ids
in
let rec buildrec ids patlist avoid env n b =
if n=0 then
- (dl, ids,
+ (dl, ids,
[PatCstr(dl, constr, List.rev patlist,Anonymous)],
detype isgoal avoid env b)
else
match kind_of_term b with
- | Lambda (x,_,b) ->
+ | Lambda (x,_,b) ->
let pat,new_avoid,new_env,new_ids = make_pat x avoid env b ids in
buildrec new_ids (pat::patlist) new_avoid new_env (n-1) b
- | LetIn (x,_,_,b) ->
+ | LetIn (x,_,_,b) ->
let pat,new_avoid,new_env,new_ids = make_pat x avoid env b ids in
buildrec new_ids (pat::patlist) new_avoid new_env (n-1) b
@@ -528,21 +526,20 @@ and detype_eqn isgoal avoid env constr construct_nargs branch =
let pat,new_avoid,new_env,new_ids =
make_pat Anonymous avoid env new_b ids in
buildrec new_ids (pat::patlist) new_avoid new_env (n-1) new_b
-
- in
+
+ in
buildrec [] [] avoid env construct_nargs branch
and detype_binder isgoal bk avoid env na ty c =
+ let flag = if isgoal then RenamingForGoal else (RenamingElsewhereFor c) in
let na',avoid' =
- if bk = BLetIn then
- concrete_let_name (avoid_flag isgoal) avoid env na c
- else
- concrete_name (avoid_flag isgoal) avoid env na c in
+ if bk = BLetIn then compute_displayed_let_name_in flag avoid na c
+ else compute_displayed_name_in flag avoid na c in
let r = detype isgoal avoid' (add_name na' env) c in
match bk with
- | BProd -> RProd (dl, na',Explicit,detype isgoal avoid env ty, r)
- | BLambda -> RLambda (dl, na',Explicit,detype isgoal avoid env ty, r)
- | BLetIn -> RLetIn (dl, na',detype isgoal avoid env ty, r)
+ | 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)
let rec detype_rel_context where avoid env sign =
let where = Option.map (fun c -> it_mkLambda_or_LetIn c sign) where in
@@ -553,8 +550,10 @@ let rec detype_rel_context where avoid env sign =
match where with
| None -> na,avoid
| Some c ->
- if b<>None then concrete_let_name None avoid env na c
- else concrete_name None avoid env na c in
+ if b<>None then
+ compute_displayed_let_name_in (RenamingElsewhereFor c) avoid na c
+ else
+ compute_displayed_name_in (RenamingElsewhereFor c) avoid na c in
let b = Option.map (detype false avoid env) b in
let t = detype false avoid env t in
(na',Explicit,b,t) :: aux avoid' (add_name na' env) rest
@@ -563,19 +562,19 @@ let rec detype_rel_context where avoid env sign =
(**********************************************************************)
(* Module substitution: relies on detyping *)
-let rec subst_cases_pattern subst pat =
+let rec subst_cases_pattern subst pat =
match pat with
| PatVar _ -> pat
- | PatCstr (loc,((kn,i),j),cpl,n) ->
- let kn' = subst_kn subst kn
+ | PatCstr (loc,((kn,i),j),cpl,n) ->
+ let kn' = subst_ind subst kn
and cpl' = list_smartmap (subst_cases_pattern subst) cpl in
if kn' == kn && cpl' == cpl then pat else
PatCstr (loc,((kn',i),j),cpl',n)
-let rec subst_rawconstr subst raw =
+let rec subst_rawconstr subst raw =
match raw with
- | RRef (loc,ref) ->
- let ref',t = subst_global subst ref in
+ | RRef (loc,ref) ->
+ let ref',t = subst_global subst ref in
if ref' == ref then raw else
detype false [] [] t
@@ -583,38 +582,38 @@ let rec subst_rawconstr subst raw =
| REvar _ -> raw
| RPatVar _ -> raw
- | RApp (loc,r,rl) ->
- let r' = subst_rawconstr subst r
+ | RApp (loc,r,rl) ->
+ let r' = subst_rawconstr subst r
and rl' = list_smartmap (subst_rawconstr subst) rl in
if r' == r && rl' == rl then raw else
RApp(loc,r',rl')
- | RLambda (loc,n,bk,r1,r2) ->
+ | RLambda (loc,n,bk,r1,r2) ->
let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
if r1' == r1 && r2' == r2 then raw else
RLambda (loc,n,bk,r1',r2')
- | RProd (loc,n,bk,r1,r2) ->
+ | RProd (loc,n,bk,r1,r2) ->
let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
if r1' == r1 && r2' == r2 then raw else
RProd (loc,n,bk,r1',r2')
- | RLetIn (loc,n,r1,r2) ->
+ | RLetIn (loc,n,r1,r2) ->
let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
if r1' == r1 && r2' == r2 then raw else
RLetIn (loc,n,r1',r2')
- | RCases (loc,sty,rtno,rl,branches) ->
+ | RCases (loc,sty,rtno,rl,branches) ->
let rtno' = Option.smartmap (subst_rawconstr subst) rtno
and rl' = list_smartmap (fun (a,x as y) ->
let a' = subst_rawconstr subst a in
- let (n,topt) = x in
+ let (n,topt) = x in
let topt' = Option.smartmap
(fun (loc,(sp,i),x,y as t) ->
- let sp' = subst_kn subst sp in
+ let sp' = subst_ind subst sp in
if sp == sp' then t else (loc,(sp',i),x,y)) topt in
if a == a' && topt == topt' then y else (a',(n,topt'))) rl
- and branches' = list_smartmap
+ and branches' = list_smartmap
(fun (loc,idl,cpl,r as branch) ->
let cpl' =
list_smartmap (subst_cases_pattern subst) cpl
@@ -628,20 +627,20 @@ let rec subst_rawconstr subst raw =
| RLetTuple (loc,nal,(na,po),b,c) ->
let po' = Option.smartmap (subst_rawconstr subst) po
- and b' = subst_rawconstr subst b
+ and b' = subst_rawconstr subst b
and c' = subst_rawconstr subst c in
if po' == po && b' == b && c' == c then raw else
RLetTuple (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 b1' = subst_rawconstr subst b1
+ and b2' = subst_rawconstr subst b2
and c' = subst_rawconstr subst c in
if c' == c & po' == po && b1' == b1 && b2' == b2 then raw else
RIf (loc,c',(na,po'),b1',b2')
- | RRec (loc,fix,ida,bl,ra1,ra2) ->
+ | RRec (loc,fix,ida,bl,ra1,ra2) ->
let ra1' = array_smartmap (subst_rawconstr subst) ra1
and ra2' = array_smartmap (subst_rawconstr subst) ra2 in
let bl' = array_smartmap
@@ -655,20 +654,21 @@ let rec subst_rawconstr subst raw =
| RSort _ -> raw
- | RHole (loc,ImplicitArg (ref,i)) ->
- let ref',_ = subst_global subst ref in
+ | RHole (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 |
- TomatchTypeParameter _ | GoalEvar | ImpossibleCase)) -> raw
+ TomatchTypeParameter _ | GoalEvar | ImpossibleCase | MatchingVar _)) ->
+ raw
- | RCast (loc,r1,k) ->
- (match k with
+ | RCast (loc,r1,k) ->
+ (match k with
CastConv (k,r2) ->
let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
if r1' == r1 && r2' == r2 then raw else
RCast (loc,r1', CastConv (k,r2'))
- | CastCoerce ->
+ | CastCoerce ->
let r1' = subst_rawconstr subst r1 in
if r1' == r1 then raw else RCast (loc,r1',k))
| RDynamic _ -> raw
@@ -684,6 +684,6 @@ let simple_cases_matrix_of_branches ind brns brs =
(dummy_loc,ids,[p],c))
0 brns brs
-let return_type_of_predicate ind nparams n pred =
- let nal,p = it_destRLambda_or_LetIn_names (n+1) pred in
+let return_type_of_predicate ind nparams nrealargs_ctxt pred =
+ let nal,p = it_destRLambda_or_LetIn_names (nrealargs_ctxt+1) pred in
(List.hd nal, Some (dummy_loc, ind, nparams, List.tl nal)), Some p
diff --git a/pretyping/detyping.mli b/pretyping/detyping.mli
index 23090858..d7fb01ae 100644
--- a/pretyping/detyping.mli
+++ b/pretyping/detyping.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: detyping.mli 10410 2007-12-31 13:11:55Z msozeau $ i*)
+(*i $Id$ i*)
(*i*)
open Util
@@ -30,9 +30,9 @@ val subst_rawconstr : substitution -> rawconstr -> rawconstr
val detype : bool -> identifier list -> names_context -> constr -> rawconstr
-val detype_case :
+val detype_case :
bool -> ('a -> rawconstr) ->
- (constructor array -> int array -> 'a array ->
+ (constructor array -> int array -> 'a array ->
(loc * identifier list * cases_pattern list * rawconstr) list) ->
('a -> int -> bool) ->
identifier list -> inductive * case_style * int * int array * int ->
@@ -44,7 +44,7 @@ val detype_rel_context : constr option -> identifier list -> names_context ->
rel_context -> rawdecl list
(* look for the index of a named var or a nondep var as it is renamed *)
-val lookup_name_as_renamed : env -> constr -> identifier -> int option
+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
@@ -54,7 +54,7 @@ val synthetize_type : unit -> bool
(* Utilities to transform kernel cases to simple pattern-matching problem *)
val it_destRLambda_or_LetIn_names : int -> rawconstr -> name list * rawconstr
-val simple_cases_matrix_of_branches :
+val simple_cases_matrix_of_branches :
inductive -> int list -> rawconstr list -> cases_clauses
val return_type_of_predicate :
inductive -> int -> int -> rawconstr -> predicate_pattern * rawconstr option
diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml
index 18e79e85..c1922d5d 100644
--- a/pretyping/evarconv.ml
+++ b/pretyping/evarconv.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: evarconv.ml 12268 2009-08-11 09:02:16Z herbelin $ *)
+(* $Id$ *)
open Pp
open Util
@@ -17,15 +17,13 @@ open Reduction
open Reductionops
open Termops
open Environ
-open Typing
-open Classops
-open Recordops
+open Recordops
open Evarutil
open Libnames
open Evd
type flex_kind_of_term =
- | Rigid of constr
+ | Rigid of constr
| MaybeFlexible of constr
| Flexible of existential
@@ -52,7 +50,7 @@ let eval_flexible_term env c =
| _ -> assert false
let evar_apprec env evd stack c =
- let sigma = evars_of evd in
+ let sigma = evd in
let rec aux s =
let (t,stack) = whd_betaiota_deltazeta_for_iota_state env sigma s in
match kind_of_term t with
@@ -62,7 +60,7 @@ let evar_apprec env evd stack c =
in aux (c, append_stack_list stack empty_stack)
let apprec_nohdbeta env evd c =
- match kind_of_term (fst (Reductionops.whd_stack (evars_of evd) c)) with
+ match kind_of_term (fst (Reductionops.whd_stack evd c)) with
| (Case _ | Fix _) -> applist (evar_apprec env evd [] c)
| _ -> c
@@ -93,31 +91,31 @@ let position_problem l2r = function
let check_conv_record (t1,l1) (t2,l2) =
try
let proji = global_of_constr t1 in
- let canon_s,l2_effective =
+ let canon_s,l2_effective =
try
match kind_of_term t2 with
Prod (_,a,b) -> (* assert (l2=[]); *)
if dependent (mkRel 1) b then raise Not_found
else lookup_canonical_conversion (proji, Prod_cs),[a;pop b]
- | Sort s ->
- lookup_canonical_conversion
+ | Sort s ->
+ lookup_canonical_conversion
(proji, Sort_cs (family_of_sort s)),[]
- | _ ->
+ | _ ->
let c2 = global_of_constr t2 in
lookup_canonical_conversion (proji, Const_cs c2),l2
- with Not_found ->
+ with Not_found ->
lookup_canonical_conversion (proji,Default_cs),[]
in
- let { o_DEF = c; o_INJ=n; o_TABS = bs;
+ let { o_DEF = c; o_INJ=n; o_TABS = bs;
o_TPARAMS = params; o_NPARAMS = nparams; o_TCOMPS = us } = canon_s in
let params1, c1, extra_args1 =
- match list_chop nparams l1 with
+ match list_chop nparams l1 with
| params1, c1::extra_args1 -> params1, c1, extra_args1
| _ -> raise Not_found in
let us2,extra_args2 = list_chop (List.length us) l2_effective in
c,bs,(params,params1),(us,us2),(extra_args1,extra_args2),c1,
(n,applist(t2,l2))
- with Failure _ | Not_found ->
+ with Failure _ | Not_found ->
raise Not_found
(* Precondition: one of the terms of the pb is an uninstantiated evar,
@@ -156,21 +154,21 @@ let ise_array2 evd f v1 v2 =
| n ->
let (i',b) = f i v1.(n) v2.(n) in
if b then allrec i' (n-1) else (evd,false)
- in
+ in
let lv1 = Array.length v1 in
- if lv1 = Array.length v2 then allrec evd (pred lv1)
+ 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 = evars_of evd in
- let term1 = whd_castappevar sigma term1 in
- let term2 = whd_castappevar sigma term2 in
+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
(* 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 (evars_of evd) term1 term2 then
+ if is_fconv pbty env evd term1 term2 then
Some true
else if is_ground_env evd env then Some false
else None
@@ -191,11 +189,11 @@ let rec evar_conv_x env evd pbty term1 term2 =
(decompose_app term1) (decompose_app term2)
and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
- (* Evar must be undefined since we have whd_ised *)
+ (* 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) ->
let f1 i =
- if List.length l1 > List.length l2 then
+ 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
@@ -212,23 +210,23 @@ and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
and f2 i =
if sp1 = sp2 then
ise_and i
- [(fun i -> ise_list2 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,
true)]
else (i,false)
- in
+ in
ise_try evd [f1; f2]
| Flexible ev1, MaybeFlexible flex2 ->
let f1 i =
- if
- is_unification_pattern_evar env ev1 l1 (applist appr2) &
+ 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) *)
- let t2 = nf_evar (evars_of evd) (applist appr2) in
+ 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
(position_problem true pbty,ev1,t2)
@@ -250,18 +248,18 @@ and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
| Some v2 ->
evar_eqappr_x env i pbty appr1 (evar_apprec env i l2 v2)
| None -> (i,false)
- in
+ in
ise_try evd [f1; f4]
| MaybeFlexible flex1, Flexible ev2 ->
let f1 i =
- if
- is_unification_pattern_evar env ev2 l2 (applist appr1) &
+ 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) *)
- let t1 = nf_evar (evars_of evd) (applist appr1) in
+ 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
(position_problem false pbty,ev2,t1)
@@ -282,7 +280,7 @@ and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
| Some v1 ->
evar_eqappr_x env i pbty (evar_apprec env i l1 v1) appr2
| None -> (i,false)
- in
+ in
ise_try evd [f1; f4]
| MaybeFlexible flex1, MaybeFlexible flex2 ->
@@ -301,7 +299,7 @@ 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 (evars_of i) appr2
+ if isLambda flex1 or is_open_canonical_projection i appr2
then
match eval_flexible_term env flex1 with
| Some v1 ->
@@ -320,17 +318,17 @@ and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
| Some v1 ->
evar_eqappr_x env i pbty (evar_apprec env i l1 v1) appr2
| None -> (i,false)
- in
+ in
ise_try evd [f1; f2; f3]
| Flexible ev1, Rigid _ ->
- if
- is_unification_pattern_evar env ev1 l1 (applist appr2) &
+ 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) *)
- let t2 = nf_evar (evars_of evd) (applist appr2) in
+ 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
(position_problem true pbty,ev1,t2)
@@ -340,13 +338,13 @@ and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
true
| Rigid _, Flexible ev2 ->
- if
- is_unification_pattern_evar env ev2 l2 (applist appr1) &
+ 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) *)
- let t1 = nf_evar (evars_of evd) (applist appr1) in
+ 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
(position_problem false pbty,ev2,t1)
@@ -364,11 +362,11 @@ and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
| Some v1 ->
evar_eqappr_x env i pbty (evar_apprec env i l1 v1) appr2
| None -> (i,false)
- in
+ in
ise_try evd [f3; f4]
-
- | Rigid _ , MaybeFlexible flex2 ->
- let f3 i =
+
+ | Rigid _ , MaybeFlexible flex2 ->
+ let f3 i =
(try conv_record env i (check_conv_record appr2 appr1)
with Not_found -> (i,false))
and f4 i =
@@ -376,11 +374,11 @@ and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
| Some v2 ->
evar_eqappr_x env i pbty appr1 (evar_apprec env i l2 v2)
| None -> (i,false)
- in
+ 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)
@@ -388,11 +386,11 @@ and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
| 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=[] ->
+ | 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 (evars_of i) c1 in
+ 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) ->
@@ -400,8 +398,8 @@ and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
ise_and i
[(fun i -> evar_conv_x env i CONV b1 b2);
(fun i ->
- let b = nf_evar (evars_of i) b1 in
- let t = nf_evar (evars_of i) t1 in
+ 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)]
@@ -409,7 +407,7 @@ and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
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
+ in
ise_try evd [f1; f2]
| LetIn (_,b1,_,c'1), _ ->(* On fait commuter les args avec le Let *)
@@ -420,20 +418,20 @@ and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
let appr2 = evar_apprec env evd l2 (subst1 b2 c'2)
in evar_eqappr_x env evd pbty appr1 appr2
- | Prod (n,c1,c'1), Prod (_,c2,c'2) when l1=[] & l2=[] ->
+ | 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 ->
- let c = nf_evar (evars_of i) c1 in
+ let c = nf_evar i c1 in
evar_conv_x (push_rel (n,None,c) env) i pbty c'1 c'2)]
| Ind sp1, Ind sp2 ->
- if sp1=sp2 then
+ if eq_ind sp1 sp2 then
ise_list2 evd (fun i -> evar_conv_x env i CONV) l1 l2
else (evd, false)
| Construct sp1, Construct sp2 ->
- if sp1=sp2 then
+ if eq_constructor sp1 sp2 then
ise_list2 evd (fun i -> evar_conv_x env i CONV) l1 l2
else (evd, false)
@@ -474,13 +472,13 @@ and evar_eqappr_x env evd pbty (term1,l1 as appr1) (term2,l2 as appr2) =
| (Ind _ | Construct _ | Sort _ | Prod _), _ -> (evd,false)
| _, (Ind _ | Construct _ | Sort _ | Prod _) -> (evd,false)
- | (App _ | Case _ | Fix _ | CoFix _),
+ | (App _ | Case _ | Fix _ | CoFix _),
(App _ | Case _ | Fix _ | CoFix _) -> (evd,false)
| (Rel _ | Var _ | Const _ | Evar _), _ -> assert false
| _, (Rel _ | Var _ | Const _ | Evar _) -> assert false
-and conv_record env evd (c,bs,(params,params1),(us,us2),(ts,ts1),c1,(n,t2)) =
+and conv_record env evd (c,bs,(params,params1),(us,us2),(ts,ts1),c1,(n,t2)) =
let (evd',ks,_) =
List.fold_left
(fun (i,ks,m) b ->
@@ -492,15 +490,15 @@ and conv_record env evd (c,bs,(params,params1),(us,us2),(ts,ts1),c1,(n,t2)) =
in
ise_and evd'
[(fun i ->
- ise_list2 i
- (fun i u1 u -> evar_conv_x env i CONV u1 (substl ks u))
- us2 us);
- (fun i ->
ise_list2 i
(fun i x1 x -> evar_conv_x env i CONV x1 (substl ks x))
params1 params);
- (fun i -> ise_list2 i (fun i -> evar_conv_x env i CONV) ts ts1);
- (fun i -> evar_conv_x env i CONV c1 (applist (c,(List.rev ks))))]
+ (fun i ->
+ ise_list2 i
+ (fun i u1 u -> evar_conv_x 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)]
(* We assume here |l1| <= |l2| *)
@@ -518,15 +516,15 @@ let first_order_unification env evd (ev1,l1) (term2,l2) =
solve_simple_eqn ~choose:true evar_conv_x env i (None,ev1,t2))]
let choose_less_dependent_instance evk evd term args =
- let evi = Evd.find (evars_of evd) evk in
+ let evi = Evd.find evd evk in
let subst = make_pure_subst evi args in
let subst' = List.filter (fun (id,c) -> c = term) subst in
if subst' = [] then error "Too complex unification problem." else
- Evd.evar_define evk (mkVar (fst (List.hd subst'))) evd
+ 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_castappevar (evars_of evd) t1) in
- let t2 = apprec_nohdbeta env evd (whd_castappevar (evars_of evd) t2) in
+ 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 (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
@@ -535,7 +533,7 @@ let apply_conversion_problem_heuristic env evd pbty t1 t2 =
(* 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 = []
+ | (Rel _|Var _), Evar (evk2,args2) when l1 = [] & l2 = []
& array_for_all (fun a -> a = term1 or isEvar a) args2 ->
(* The typical kind of constraint coming from pattern-matching return
type inference *)
@@ -569,7 +567,7 @@ let the_conv_x_leq env t1 t2 evd =
match evar_conv_x 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
(evd',true) -> evd := evd'; true
diff --git a/pretyping/evarconv.mli b/pretyping/evarconv.mli
index f92a6fdb..9a4247bc 100644
--- a/pretyping/evarconv.mli
+++ b/pretyping/evarconv.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: evarconv.mli 9141 2006-09-15 10:07:01Z herbelin $ i*)
+(*i $Id$ i*)
(*i*)
open Term
@@ -17,21 +17,27 @@ open Evd
(*i*)
(* returns exception Reduction.NotConvertible if not unifiable *)
-val the_conv_x : env -> constr -> constr -> evar_defs -> evar_defs
-val the_conv_x_leq : env -> constr -> constr -> evar_defs -> evar_defs
+val the_conv_x : env -> constr -> constr -> evar_map -> evar_map
+val the_conv_x_leq : 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_defs ref -> constr -> constr -> bool
-val e_cumul : env -> evar_defs ref -> constr -> constr -> bool
+val e_conv : env -> evar_map ref -> constr -> constr -> bool
+val e_cumul : env -> evar_map ref -> constr -> constr -> bool
(*i For debugging *)
val evar_conv_x :
- env -> evar_defs -> conv_pb -> constr -> constr -> evar_defs * bool
-val evar_eqappr_x :
- env -> evar_defs ->
+ env -> evar_map -> conv_pb -> constr -> constr -> evar_map * bool
+val evar_eqappr_x :
+ env -> evar_map ->
conv_pb -> constr * constr list -> constr * constr list ->
- evar_defs * bool
+ evar_map * bool
(*i*)
-val consider_remaining_unif_problems : env -> evar_defs -> evar_defs * bool
+val consider_remaining_unif_problems : env -> evar_map -> evar_map * bool
+
+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)
diff --git a/pretyping/evarutil.ml b/pretyping/evarutil.ml
index fbaac79b..2b218da6 100644
--- a/pretyping/evarutil.ml
+++ b/pretyping/evarutil.ml
@@ -6,15 +6,15 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: evarutil.ml 13116 2010-06-12 15:18:07Z herbelin $ *)
+(* $Id$ *)
open Util
open Pp
open Names
-open Nameops
open Univ
open Term
open Termops
+open Namegen
open Sign
open Pre_env
open Environ
@@ -23,33 +23,21 @@ open Reductionops
open Pretype_errors
open Retyping
-(* Expanding existential variables (pretyping.ml) *)
-(* 1- whd_ise fails if an existential is undefined *)
+open Pretype_errors
+open Retyping
+
+(* Expanding existential variables *)
+(* 1- flush_and_check_evars fails if an existential is undefined *)
exception Uninstantiated_evar of existential_key
-let rec whd_ise sigma c =
+let rec flush_and_check_evars sigma c =
match kind_of_term c with
- | Evar (evk,args as ev) when Evd.mem sigma evk ->
- if Evd.is_defined sigma evk then
- whd_ise sigma (existential_value sigma ev)
- else raise (Uninstantiated_evar evk)
- | _ -> c
-
-
-(* Expand evars, possibly in the head of an application *)
-let whd_castappevar_stack sigma c =
- let rec whrec (c, l as s) =
- match kind_of_term c with
- | Evar (evk,args as ev) when Evd.mem sigma evk & Evd.is_defined sigma evk
- -> whrec (existential_value sigma ev, l)
- | Cast (c,_,_) -> whrec (c, l)
- | App (f,args) -> whrec (f, Array.fold_right (fun a l -> a::l) args l)
- | _ -> s
- in
- whrec (c, [])
-
-let whd_castappevar sigma c = applist (whd_castappevar_stack sigma c)
+ | Evar (evk,_ as ev) ->
+ (match existential_opt_value sigma ev with
+ | None -> raise (Uninstantiated_evar evk)
+ | Some c -> flush_and_check_evars sigma c)
+ | _ -> map_constr (flush_and_check_evars sigma) c
let nf_evar = Pretype_errors.nf_evar
let j_nf_evar = Pretype_errors.j_nf_evar
@@ -57,32 +45,29 @@ let jl_nf_evar = Pretype_errors.jl_nf_evar
let jv_nf_evar = Pretype_errors.jv_nf_evar
let tj_nf_evar = Pretype_errors.tj_nf_evar
-let nf_named_context_evar sigma ctx =
+let nf_named_context_evar sigma ctx =
Sign.map_named_context (Reductionops.nf_evar sigma) ctx
-let nf_rel_context_evar sigma ctx =
+let nf_rel_context_evar sigma ctx =
Sign.map_rel_context (Reductionops.nf_evar sigma) ctx
-
-let nf_env_evar sigma env =
+
+let nf_env_evar sigma env =
let nc' = nf_named_context_evar sigma (Environ.named_context env) in
let rel' = nf_rel_context_evar sigma (Environ.rel_context env) in
push_rel_context rel' (reset_with_named_context (val_of_named_context nc') env)
let nf_evar_info evc info =
- { info with
+ { info with
evar_concl = Reductionops.nf_evar evc info.evar_concl;
- evar_hyps = map_named_val (Reductionops.nf_evar evc) info.evar_hyps}
+ evar_hyps = map_named_val (Reductionops.nf_evar evc) info.evar_hyps;
+ evar_body = match info.evar_body with
+ | Evar_empty -> Evar_empty
+ | Evar_defined c -> Evar_defined (Reductionops.nf_evar evc c) }
let nf_evars evm = Evd.fold (fun ev evi evm' -> Evd.add evm' ev (nf_evar_info evm evi))
evm Evd.empty
-let nf_evar_defs evd = Evd.evars_reset_evd (nf_evars (Evd.evars_of evd)) evd
-
-let nf_isevar evd = nf_evar (Evd.evars_of evd)
-let j_nf_isevar evd = j_nf_evar (Evd.evars_of evd)
-let jl_nf_isevar evd = jl_nf_evar (Evd.evars_of evd)
-let jv_nf_isevar evd = jv_nf_evar (Evd.evars_of evd)
-let tj_nf_isevar evd = tj_nf_evar (Evd.evars_of evd)
+let nf_evar_map evd = Evd.evars_reset_evd (nf_evars evd) evd
(**********************)
(* Creating new metas *)
@@ -107,25 +92,25 @@ let collect_evars emap c =
let push_dependent_evars sigma emap =
Evd.fold (fun ev {evar_concl = ccl} (sigma',emap') ->
- List.fold_left
- (fun (sigma',emap') ev ->
+ List.fold_left
+ (fun (sigma',emap') ev ->
(Evd.add sigma' ev (Evd.find emap' ev),Evd.remove emap' ev))
(sigma',emap') (collect_evars emap' ccl))
emap (sigma,emap)
-let push_duplicated_evars sigma emap c =
+let push_duplicated_evars sigma emap c =
let rec collrec (one,(sigma,emap) as acc) c =
match kind_of_term c with
| Evar (evk,_) when not (Evd.mem sigma evk) ->
- if List.mem evk one then
- let sigma' = Evd.add sigma evk (Evd.find emap evk) in
- let emap' = Evd.remove emap evk in
- (one,(sigma',emap'))
- else
- (evk::one,(sigma,emap))
+ if List.mem evk one then
+ let sigma' = Evd.add sigma evk (Evd.find emap evk) in
+ let emap' = Evd.remove emap evk in
+ (one,(sigma',emap'))
+ else
+ (evk::one,(sigma,emap))
| _ ->
- fold_constr collrec acc c
- in
+ fold_constr collrec acc c
+ in
snd (collrec ([],(sigma,emap)) c)
(* replaces a mapping of existentials into a mapping of metas.
@@ -146,11 +131,11 @@ let evars_to_metas sigma (emap, c) =
(* The list of non-instantiated existential declarations *)
-let non_instantiated sigma =
+let non_instantiated sigma =
let listev = to_list sigma in
- List.fold_left
- (fun l (ev,evi) ->
- if evi.evar_body = Evar_empty then
+ List.fold_left
+ (fun l (ev,evi) ->
+ if evi.evar_body = Evar_empty then
((ev,nf_evar_info sigma evi)::l) else l)
[] listev
@@ -179,6 +164,79 @@ let new_evar_instance sign evd typ ?(src=(dummy_loc,InternalHole)) ?filter insta
let evd = evar_declare sign newevk typ ~src:src ?filter 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 =
+ 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
+
+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
+ (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 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 expand_var aliases x =
+ try expand_var_at_least_once aliases x with Not_found -> x
+
+let expand_var_opt aliases x =
+ try Some (expand_var_at_least_once aliases x) with Not_found -> None
+
+let extend_alias (_,b,_) (var_aliases,rel_aliases) =
+ let rel_aliases =
+ Intmap.fold (fun n c -> Intmap.add (n+1) (lift 1 c))
+ 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
+ (var_aliases, rel_aliases)
+
+let rec expand_vars_in_term_using aliases t = match kind_of_term t with
+ | Rel _ | Var _ ->
+ expand_var 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
+
(* Knowing that [Gamma |- ev : T] and that [ev] is applied to [args],
* [make_projectable_subst ev args] builds the substitution [Gamma:=args].
* If a variable and an alias of it are bound to the same instance, we skip
@@ -189,20 +247,28 @@ let new_evar_instance sign evd typ ?(src=(dummy_loc,InternalHole)) ?filter insta
* useful to ensure the uniqueness of a projection.
*)
-let make_projectable_subst sigma evi args =
+let make_projectable_subst aliases sigma evi args =
let sign = evar_filtered_context evi in
- let rec alias_of_var id =
- match pi2 (Sign.lookup_named id sign) with
- | Some t when isVar t -> alias_of_var (destVar t)
- | _ -> id in
+ let evar_aliases = compute_aliases sign in
snd (List.fold_right
(fun (id,b,c) (args,l) ->
match b,args with
- | Some c, a::rest when
- isVar c & (try eq_constr a (snd (List.assoc (destVar c) l)) with Not_found -> false) -> (rest,l)
- | _, a::rest -> (rest, (id, (alias_of_var id,whd_evar sigma a))::l)
+ | None, a::rest ->
+ let a = whd_evar sigma a in
+ (rest,Idmap.add id [a,expand_full_opt aliases a,id] l)
+ | 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)
+ else
+ (rest,Idmap.add idc ((a,expand_full_opt aliases a,id)::sub) l)
+ | _ ->
+ (rest,Idmap.add id [a,expand_full_opt aliases a,id] l))
| _ -> anomaly "Instance does not match its signature")
- sign (List.rev (Array.to_list args),[]))
+ sign (array_rev_to_list args,Idmap.empty))
let make_pure_subst evi args =
snd (List.fold_right
@@ -210,16 +276,16 @@ let make_pure_subst evi args =
match args with
| a::rest -> (rest, (id,a)::l)
| _ -> anomaly "Instance does not match its signature")
- (evar_filtered_context evi) (List.rev (Array.to_list args),[]))
+ (evar_filtered_context evi) (array_rev_to_list args,[]))
(* [push_rel_context_to_named_context] builds the defining context and the
* initial instance of an evar. If the evar is to be used in context
- *
+ *
* Gamma = a1 ... an xp ... x1
* \- named part -/ \- de Bruijn part -/
- *
+ *
* then the x1...xp are turned into variables so that the evar is declared in
- * context
+ * context
*
* a1 ... an xp ... x1
* \----------- named part ------------/
@@ -227,16 +293,17 @@ let make_pure_subst evi args =
* but used applied to the initial instance "a1 ... an Rel(p) ... Rel(1)"
* so that ev[a1:=a1 ... an:=an xp:=Rel(p) ... x1:=Rel(1)] is correctly typed
* in context Gamma.
- *
+ *
* Remark 1: The instance is reverted in practice (i.e. Rel(1) comes first)
* Remark 2: If some of the ai or xj are definitions, we keep them in the
* instance. This is necessary so that no unfolding of local definitions
* happens when inferring implicit arguments (consider e.g. the problem
- * "x:nat; x':=x; f:forall x, x=x -> Prop |- f _ (refl_equal x')"
- * we want the hole to be instantiated by x', not by x (which would have the
- * case in [invert_instance] if x' had disappear of the instance).
+ * "x:nat; x':=x; f:forall y, y=y -> Prop |- f _ (refl_equal x')" which
+ * produces the equation "?y[x,x']=?y[x,x']" =? "x'=x'": we want
+ * the hole to be instantiated by x', not by x (which would have been
+ * the case in [invert_definition] if x' had disappeared from the instance).
* Note that at any time, if, in some context env, the instance of
- * declaration x:A is t and the instance of definition x':=phi(x) is u, then
+ * declaration x:A is t and the instance of definition x':=phi(x) is u, then
* we have the property that u and phi(t) are convertible in env.
*)
@@ -256,7 +323,7 @@ let push_rel_context_to_named_context env typ =
(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)
-
+
(* [new_evar] declares a new existential in an env env with type typ *)
(* Converting the env into the sign of the evar to define *)
@@ -274,10 +341,6 @@ let e_new_evar evdref env ?(src=(dummy_loc,InternalHole)) ?filter ty =
* operations on the evar constraints *
*------------------------------------*)
-let is_pattern inst =
- array_for_all (fun a -> isRel a || isVar a) inst &&
- array_distinct inst
-
(* Pb: defined Rels and Vars should not be considered as a pattern... *)
(*
let is_pattern inst =
@@ -288,23 +351,10 @@ let is_pattern inst =
is_hopat [] (Array.to_list inst)
*)
-let evar_well_typed_body evd ev evi body =
- try
- let env = evar_unfiltered_env evi in
- let ty = evi.evar_concl in
- Typing.check env (evars_of evd) body ty;
- true
- with e ->
- pperrnl
- (str "Ill-typed evar instantiation: " ++ fnl() ++
- pr_evar_defs evd ++ fnl() ++
- str "----> " ++ int ev ++ str " := " ++
- print_constr body);
- false
-
-(* 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
+
+(* We have x1..xq |- ?e1 and had to solve something like
+ * Σ; Γ |- ?e1[u1..uq] = (...\y1 ... \yk ... c), where c is typically some
+ * ?e2[v1..vn], hence flexible. We had to go through k binders and now
* virtually have x1..xq, y1..yk | ?e1' and the equation
* Γ, y1..yk |- ?e1'[u1..uq y1..yk] = c.
* What we do is to formally introduce ?e1' in context x1..xq, Γ, y1..yk,
@@ -313,10 +363,10 @@ let evar_well_typed_body evd ev evi body =
*
* 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
+ * 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
+ * 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.
@@ -330,10 +380,10 @@ let shrink_context env subst ty =
(* 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 ->
+ | (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_context rel_subst (List.rev rev_rel_sign),
substl rel_subst ty
in
let rec shrink_named subst named_subst rev_named_sign =
@@ -352,13 +402,13 @@ let shrink_context env subst ty =
shrink_named subst [] rev_named_sign
let extend_evar env evdref k (evk1,args1) c =
- let ty = get_type_of env (evars_of !evdref) c in
+ 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 (evars_of !evdref) evk1 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
@@ -378,7 +428,7 @@ let extend_evar env evdref k (evk1,args1) c =
let subfilter p filter l =
let (filter,_,l) =
List.fold_left (fun (filter,l,newl) b ->
- if b then
+ 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))
@@ -391,7 +441,7 @@ let restrict_upon_filter evd evi evk p args =
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 = Evd.evar_define evk newev evd in
+ let evd = Evd.define evk newev evd in
evd,fst (destEvar newev),newargs
else
evd,evk,args
@@ -414,10 +464,10 @@ 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
evars) *)
- let check id' =
+ let check id' =
if List.mem id' ids then
raise (ClearDependencyError (id',err))
- in
+ in
match kind_of_term c with
| Var id' ->
check id'; c
@@ -426,25 +476,25 @@ let rec check_and_clear_in_constr evdref err ids c =
let vars = Environ.vars_of_global (Global.env()) c in
List.iter check vars; c
- | Evar (evk,l as ev) ->
+ | Evar (evk,l as ev) ->
if Evd.is_defined_evar !evdref ev then
(* If evk is already defined we replace it by its definition *)
- let nc = whd_evar (evars_of !evdref) c in
+ let nc = whd_evar !evdref c in
(check_and_clear_in_constr evdref err ids nc)
- else
+ else
(* We check for dependencies to elements of ids in the
evar_info corresponding to e and in the instance of
arguments. Concurrently, we build a new evar
corresponding to e where hypotheses of ids have been
removed *)
- let evi = Evd.find (evars_of !evdref) evk in
+ let evi = Evd.find !evdref evk in
let ctxt = Evd.evar_filtered_context evi in
let (nhyps,nargs,rids) =
- List.fold_right2
+ List.fold_right2
(fun (rid,ob,c as h) a (hy,ar,ri) ->
(* Check if some id to clear occurs in the instance
a of rid in ev and remember the dependency *)
- match
+ match
List.filter (fun id -> List.mem id ids) (collect_vars a)
with
| id :: _ -> (hy,ar,(rid,id)::ri)
@@ -462,15 +512,17 @@ let rec check_and_clear_in_constr evdref err ids c =
in the type of ev and adjust the source of the dependency *)
let nconcl =
try check_and_clear_in_constr evdref (EvarTypingBreak ev)
- (List.map fst rids) (evar_concl evi)
- with ClearDependencyError (rid,err) ->
+ (List.map fst rids) (evar_concl evi)
+ with ClearDependencyError (rid,err) ->
raise (ClearDependencyError (List.assoc rid rids,err)) in
- let env = Sign.fold_named_context push_named nhyps ~init:(empty_env) in
- let ev'= e_new_evar evdref env ~src:(evar_source evk !evdref) nconcl in
- evdref := Evd.evar_define evk ev' !evdref;
+ if rids = [] then c else begin
+ let env = Sign.fold_named_context push_named nhyps ~init:(empty_env) in
+ let ev'= e_new_evar evdref env ~src:(evar_source evk !evdref) nconcl in
+ evdref := Evd.define evk ev' !evdref;
let (evk',_) = destEvar ev' in
- mkEvar(evk', Array.of_list nargs)
+ mkEvar(evk', Array.of_list nargs)
+ end
| _ -> map_constr (check_and_clear_in_constr evdref err ids) c
@@ -480,7 +532,7 @@ let clear_hyps_in_evi evdref hyps concl ids =
the contexts of the evars occuring in evi *)
let nconcl =
check_and_clear_in_constr evdref (OccurHypInSimpleClause None) ids concl in
- let nhyps =
+ let nhyps =
let check_context (id,ob,c) =
let err = OccurHypInSimpleClause (Some id) in
(id, Option.map (check_and_clear_in_constr evdref err ids) ob,
@@ -502,53 +554,13 @@ let clear_hyps_in_evi evdref hyps concl ids =
(nhyps,nconcl)
-(* 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 expand_var_once env x = match kind_of_term x with
- | Rel n ->
- begin match pi2 (lookup_rel n env) with
- | Some t when isRel t or isVar t -> lift n t
- | _ -> raise Not_found
- end
- | Var id ->
- begin match pi2 (lookup_named id env) with
- | Some t when isVar t -> t
- | _ -> raise Not_found
- end
- | _ ->
- raise Not_found
-
-let rec expand_var_at_least_once env x =
- let t = expand_var_once env x in
- try expand_var_at_least_once env t
- with Not_found -> t
-
-let expand_var env x =
- try expand_var_at_least_once env x with Not_found -> x
-
-let expand_var_opt env x =
- try Some (expand_var_at_least_once env x) with Not_found -> None
-
-let rec expand_vars_in_term env t = match kind_of_term t with
- | Rel _ | Var _ -> expand_var env t
- | _ -> map_constr_with_full_binders push_rel expand_vars_in_term env t
-
-let rec expansions_of_var env x =
- try
- let t = expand_var_once env x in
- t :: expansions_of_var env t
- with Not_found ->
- [x]
-
(* [find_projectable_vars env sigma y subst] finds all vars of [subst]
* that project on [y]. It is able to find solutions to the following
* two kinds of problems:
*
* - ?n[...;x:=y;...] = y
* - ?n[...;x:=?m[args];...] = y with ?m[args] = y recursively solvable
- *
+ *
* (see test-suite/success/Fixpoint.v for an example of application of
* the second kind of problem).
*
@@ -577,29 +589,51 @@ let rec expansions_of_var env x =
exception NotUnique
exception NotUniqueInType of types
-type evar_projection =
-| ProjectVar
+type evar_projection =
+| ProjectVar
| ProjectEvar of existential * evar_info * identifier * evar_projection
-let rec find_projectable_vars with_evars env sigma y subst =
- let is_projectable (id,(idc,y')) =
- let y' = whd_evar sigma y' in
- if y = y' or expand_var env y = expand_var env y'
- then (idc,(y'=y,(id,ProjectVar)))
- else if with_evars & isEvar y' then
- (* TODO: infer conditions for being of unifiable types *)
- let (evk,argsv as t) = destEvar y' in
- let evi = Evd.find sigma evk in
- let subst = make_projectable_subst sigma evi argsv in
- let l = find_projectable_vars with_evars env sigma y subst in
- match l with
- | [id',p] -> (idc,(true,(id,ProjectEvar(t,evi,id',p))))
- | _ -> failwith ""
- else failwith "" in
- let l = map_succeed is_projectable subst in
- let l = list_partition_by (fun (idc,_) (idc',_) -> idc = idc') l in
- let l = List.map (List.map snd) l in
- List.map (fun l -> try List.assoc true l with Not_found -> snd (List.hd l)) l
+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
+ 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
+
+let rec find_projectable_vars with_evars aliases sigma y subst =
+ let yc = expand_var aliases y in
+ let is_projectable idc idcl subst' =
+ (* First test if some [id] aliased to [idc] is bound to [y] in [subst] *)
+ try
+ let id = assoc_up_to_alias sigma aliases y yc idcl in
+ (id,ProjectVar)::subst'
+ with Not_found ->
+ (* Then test if [idc] is (indirectly) bound in [subst] to some evar *)
+ (* projectable on [y] *)
+ if with_evars then
+ let idcl' = List.filter (fun (c,_,id) -> isEvar c) idcl in
+ match idcl' with
+ | [c,_,id] ->
+ begin
+ let (evk,argsv as t) = destEvar c in
+ let evi = Evd.find sigma evk in
+ let subst = make_projectable_subst aliases sigma evi argsv in
+ let l = find_projectable_vars with_evars aliases sigma y subst in
+ match l with
+ | [id',p] -> (id,ProjectEvar (t,evi,id',p))::subst'
+ | _ -> subst'
+ end
+ | [] -> subst'
+ | _ -> anomaly "More than one non var in aliases class of evar instance"
+ else
+ subst' in
+ Idmap.fold is_projectable subst []
(* [filter_solution] checks if one and only one possible projection exists
* among a set of solutions to a projection problem *)
@@ -609,8 +643,9 @@ let filter_solution = function
| (id,p)::_::_ -> raise NotUnique
| [id,p] -> (mkVar id, p)
-let project_with_effects env sigma effects t subst =
- let c, p = filter_solution (find_projectable_vars false env sigma t subst) in
+let project_with_effects aliases sigma effects t subst =
+ let c, p =
+ filter_solution (find_projectable_vars false aliases sigma t subst) in
effects := p :: !effects;
c
@@ -626,17 +661,17 @@ let rec find_solution_type evarenv = function
* type is an evar too.
*
* Note: typing creates new evar problems, which induces a recursive dependency
- * with [evar_define]. To avoid a too large set of recursive functions, we
- * pass [evar_define] to [do_projection_effects] as a parameter.
+ * with [define]. To avoid a too large set of recursive functions, we
+ * pass [define] to [do_projection_effects] as a parameter.
*)
let rec do_projection_effects define_fun env ty evd = function
| ProjectVar -> evd
| ProjectEvar ((evk,argsv),evi,id,p) ->
- let evd = Evd.evar_define evk (mkVar id) evd in
+ let evd = Evd.define evk (mkVar id) evd in
(* TODO: simplify constraints involving evk *)
let evd = do_projection_effects define_fun env ty evd p in
- let ty = whd_betadeltaiota env (evars_of evd) (Lazy.force ty) in
+ let ty = whd_betadeltaiota env evd (Lazy.force ty) in
if not (isSort ty) then
(* Don't try to instantiate if a sort because if evar_concl is an
evar it may commit to a univ level which is not the right
@@ -644,13 +679,13 @@ let rec do_projection_effects define_fun env ty evd = function
unif, we know that no coercion can be inserted) *)
let subst = make_pure_subst evi argsv in
let ty' = replace_vars subst evi.evar_concl in
- let ty' = whd_evar (evars_of evd) ty' in
+ let ty' = whd_evar evd ty' in
if isEvar ty' then define_fun env (destEvar ty') ty evd else evd
else
evd
-(* Assuming Σ; Γ, y1..yk |- c, [invert_subst Γ k Σ [x1:=u1;...;xn:=un] c]
- * tries to return φ(x1..xn) such that equation φ(u1..un) = c is valid.
+(* Assuming Σ; Γ, y1..yk |- c, [invert_arg_from_subst Γ k Σ [x1:=u1..xn:=un] c]
+ * tries to return φ(x1..xn) such that equation φ(u1..un) = c is valid.
* The strategy is to imitate the structure of c and then to invert
* the variables of c (i.e. rels or vars of Γ) using the algorithm
* implemented by project_with_effects/find_projectable_vars.
@@ -658,15 +693,15 @@ let rec do_projection_effects define_fun env ty evd = function
* 1 solutions is found.
*
* Precondition: Σ; Γ, y1..yk |- c /\ Σ; Γ |- u1..un
- * Postcondition: if φ(x1..xn) is returned then
+ * Postcondition: if φ(x1..xn) is returned then
* Σ; Γ, y1..yk |- φ(u1..un) = c /\ x1..xn |- φ(x1..xn)
*
* The effects correspond to evars instantiated while trying to project.
*
- * [invert_subst] is used on instances of evars. Since the evars are flexible,
- * these instances are potentially erasable. This is why we don't investigate
- * whether evars in the instances of evars are unifiable, to the contrary of
- * [invert_definition].
+ * [invert_arg_from_subst] is used on instances of evars. Since the
+ * evars are flexible, these instances are potentially erasable. This
+ * is why we don't investigate whether evars in the instances of evars
+ * are unifiable, to the contrary of [invert_definition].
*)
type projectibility_kind =
@@ -677,27 +712,26 @@ type projectibility_status =
| CannotInvert
| Invertible of projectibility_kind
-let invert_arg_from_subst env k sigma subst_in_env c_in_env_extended_with_k_binders =
+let invert_arg_from_subst aliases k sigma subst_in_env c_in_env_extended_with_k_binders =
let effects = ref [] in
let rec aux k t =
let t = whd_evar sigma t in
match kind_of_term t with
| Rel i when i>k ->
- project_with_effects env sigma effects (mkRel (i-k)) subst_in_env
+ project_with_effects aliases sigma effects (mkRel (i-k)) subst_in_env
| Var id ->
- project_with_effects env sigma effects t subst_in_env
+ project_with_effects aliases sigma effects t subst_in_env
| _ ->
map_constr_with_binders succ aux k t in
- try
+ try
let c = aux k c_in_env_extended_with_k_binders in
Invertible (UniqueProjection (c,!effects))
with
| Not_found -> CannotInvert
| NotUnique -> Invertible NoUniqueProjection
-let invert_arg env k sigma (evk,args_in_env) c_in_env_extended_with_k_binders =
- let subst_in_env = make_projectable_subst sigma (Evd.find sigma evk) args_in_env in
- let res = invert_arg_from_subst env k sigma subst_in_env c_in_env_extended_with_k_binders in
+let invert_arg aliases k sigma evk subst_in_env c_in_env_extended_with_k_binders =
+ let res = invert_arg_from_subst aliases k sigma subst_in_env c_in_env_extended_with_k_binders in
match res with
| Invertible (UniqueProjection (c,_)) when occur_evar evk c -> CannotInvert
| _ -> res
@@ -707,7 +741,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 (evars_of evd) t) in
+ let ty = lazy (Retyping.get_type_of env evd t) in
match projs with
| NoUniqueProjection -> raise NotUnique
| UniqueProjection (c,effects) ->
@@ -740,11 +774,11 @@ let restrict_hyps evd evk filter =
occurrence of x in the hnf of C), then z should be removed too.
- If y is in a non-erasable position in T(x,y,z) then the problem is
unsolvable.
- Computing whether y is erasable or not may be costly and the
+ Computing whether y is erasable or not may be costly and the
interest for this early detection in practice is not obvious. We let
it for future work. In any case, thanks to the use of filters, the whole
(unrestricted) context remains consistent. *)
- let evi = Evd.find (evars_of evd) evk in
+ let evi = Evd.find evd evk in
let env = evar_unfiltered_env evi in
let oldfilter = evar_filter evi in
let filter,_ = List.fold_right (fun oldb (l,filter) ->
@@ -759,7 +793,7 @@ let do_restrict_hyps evd evk projs =
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.evar_define evk nc evd in
+ let evd = Evd.define evk nc evd in
let evk',_ = destEvar nc in
evd,evk'
@@ -767,7 +801,7 @@ let do_restrict_hyps evd evk projs =
let postpone_evar_term env evd (evk,argsv) rhs =
let rhs = expand_vars_in_term env rhs in
- let evi = Evd.find (evars_of evd) evk in
+ let evi = Evd.find evd evk in
let evd,evk,args =
restrict_upon_filter evd evi evk
(* Keep only variables that depends in rhs *)
@@ -794,13 +828,13 @@ let postpone_evar_evar env evd projs1 (evk1,args1) projs2 (evk2,args2) =
let pb = (Reduction.CONV,env,mkEvar(evk1',args1'),mkEvar (evk2',args2')) in
add_conv_pb pb evd
-(* [solve_evar_evar f Γ Σ ?e1[u1..un] ?e2[v1..vp]] applies an heuristic
+(* [solve_evar_evar f Γ Σ ?e1[u1..un] ?e2[v1..vp]] applies an heuristic
* to solve the equation Σ; Γ ⊢ ?e1[u1..un] = ?e2[v1..vp]:
- * - if there are at most one φj for each vj s.t. vj = φj(u1..un),
- * we first restrict ?2 to the subset v_k1..v_kq of the vj that are
+ * - if there are at most one φj for each vj s.t. vj = φj(u1..un),
+ * we first restrict ?2 to the subset v_k1..v_kq of the vj that are
* inversible and we set ?1[x1..xn] := ?2[φk1(x1..xn)..φkp(x1..xn)]
- * - symmetrically if there are at most one ψj for each uj s.t.
- * uj = ψj(v1..vp),
+ * - symmetrically if there are at most one ψj for each uj s.t.
+ * uj = ψj(v1..vp),
* - otherwise, each position i s.t. ui does not occur in v1..vp has to
* be restricted and similarly for the vi, and we leave the equation
* as an open equation (performed by [postpone_evar])
@@ -811,17 +845,40 @@ let postpone_evar_evar env evd projs1 (evk1,args1) projs2 (evk2,args2) =
* Note: argument f is the function used to instantiate evars.
*)
+let are_canonical_instances args1 args2 env =
+ let n1 = Array.length args1 in
+ let n2 = Array.length args2 in
+ let rec aux n = function
+ | (id,_,c)::sign
+ when n < n1 && args1.(n) = mkVar id && 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))
+ in aux2 n
+ | _ -> false in
+ n1 = n2 & aux 0 (named_context env)
+
exception CannotProject of projectibility_status list
-let solve_evar_evar_l2r f env evd (evk1,args1) (evk2,_ as ev2) =
- let proj1 = array_map_to_list (invert_arg env 0 (evars_of evd) ev2) args1 in
+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
+ 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.evar_define evk2 (mkEvar(evk1',Array.of_list args1')) evd
+ Evd.define evk2 (mkEvar(evk1',Array.of_list args1')) evd
with NotUnique ->
raise (CannotProject proj1)
@@ -832,20 +889,33 @@ let solve_evar_evar f env evd ev1 ev2 =
with CannotProject projs2 ->
postpone_evar_evar env evd projs1 ev1 projs2 ev2
-let expand_rhs env sigma subst rhs =
- let d = (named_hd env rhs Anonymous,Some rhs,get_type_of env sigma rhs) in
- let rhs' = lift 1 rhs in
- let f (id,(idc,t)) = (id,(idc,replace_term rhs' (mkRel 1) (lift 1 t))) in
- push_rel d env, List.map f subst, mkRel 1
+(* 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
+ * that don't unify are discarded (i.e. ?i is redefined so that it does not
+ * 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
+ (* Filter and restrict if needed *)
+ let evd,evk,args =
+ restrict_upon_filter evd evi evk
+ (fun (a1,a2) -> snd (conv_algo env evd Reduction.CONV a1 a2))
+ (List.combine (Array.to_list argsv1) (Array.to_list argsv2)) in
+ (* Leave a unification problem *)
+ 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
(* We try to instantiate the evar assuming the body won't depend
* on arguments that are not Rels or Vars, or appearing several times
- * (i.e. we tackle a generalization of Miller-Pfenning patterns unification)
+ * (i.e. we tackle a generalization of Miller-Pfenning patterns unification)
*
* 1) Let "env |- ?ev[hyps:=args] = rhs" be the unification problem
* 2) We limit it to a patterns unification problem "env |- ev[subst] = rhs"
* where only Rel's and Var's are relevant in subst
- * 3) We recur on rhs, "imitating" the term, and failing if some Rel/Var is
+ * 3) We recur on rhs, "imitating" the term, and failing if some Rel/Var is
* not in the scope of ?ev. For instance, the problem
* "y:nat |- ?x[] = y" where "|- ?1:nat" is not satisfiable because
* ?1 would be instantiated by y which is not in the scope of ?1.
@@ -855,26 +925,28 @@ let expand_rhs env sigma subst rhs =
* 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]
* Precondition: Σ; Γ, y1..yk |- c /\ Σ; Γ |- u1..un
- * Postcondition: if φ(x1..xn) is returned then
+ * Postcondition: if φ(x1..xn) is returned then
* Σ; Γ, y1..yk |- φ(u1..un) = c /\ x1..xn |- φ(x1..xn)
*)
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 aliases = make_alias_map env in
let evdref = ref evd in
let progress = ref false in
- let evi = Evd.find (evars_of evd) evk in
- let subst = make_projectable_subst (evars_of evd) evi argsv in
+ let evi = Evd.find evd evk in
+ let subst = make_projectable_subst aliases evd evi argsv in
(* Projection *)
let project_variable t =
(* Evar/Var problem: unifiable iff variable projectable from ev subst *)
- try
- let sols = find_projectable_vars true env (evars_of !evdref) t subst in
+ try
+ let sols = find_projectable_vars true aliases !evdref t subst in
let c, p = match sols with
| [] -> raise Not_found
| [id,p] -> (mkVar id, p)
@@ -882,7 +954,7 @@ let rec invert_definition choose env evd (evk,argsv as ev) rhs =
if choose then (mkVar id, p)
else raise (NotUniqueInType(find_solution_type (evar_env evi) sols))
in
- let ty = lazy (Retyping.get_type_of env (evars_of !evdref) t) in
+ let ty = lazy (Retyping.get_type_of env !evdref t) in
let evd = do_projection_effects evar_define env ty !evdref p in
evdref := evd;
c
@@ -891,7 +963,7 @@ 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 *)
- let ts = expansions_of_var env t 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 args' = filter_along (fun x -> x) filter argsv in
@@ -903,21 +975,21 @@ let rec invert_definition choose env evd (evk,argsv as ev) rhs =
evar in
let rec imitate (env',k as envk) t =
- let t = whd_evar (evars_of !evdref) t in
+ let t = whd_evar !evdref t in
match kind_of_term t with
| Rel i when i>k -> project_variable (mkRel (i-k))
| Var id -> project_variable t
| Evar (evk',args' as ev') ->
- if evk = evk' then error_occur_check env (evars_of evd) evk rhs;
+ if evk = evk' then raise (OccurCheckIn (evd,rhs));
(* Evar/Evar problem (but left evar is virtual) *)
let projs' =
array_map_to_list
- (invert_arg_from_subst env k (evars_of !evdref) subst) args'
+ (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' =
+ 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
@@ -941,13 +1013,13 @@ let rec invert_definition choose env evd (evk,argsv as ev) rhs =
map_constr_with_full_binders (fun d (env,k) -> push_rel d env, k+1)
imitate envk t in
- let rhs = whd_beta (evars_of evd) rhs (* heuristic *) in
+ let rhs = whd_beta evd rhs (* heuristic *) in
let body = imitate (env,0) rhs in
(!evdref,body)
-(* [evar_define] tries to solve the problem "?ev[args] = rhs" when "?ev" is
+(* [define] tries to solve the problem "?ev[args] = rhs" when "?ev" is
* an (uninstantiated) evar such that "hyps |- ?ev : typ". Otherwise said,
- * [evar_define] tries to find an instance lhs such that
+ * [define] tries to find an instance lhs such that
* "lhs [hyps:=args]" unifies to rhs. The term "lhs" must be closed in
* context "hyps" and not referring to itself.
*)
@@ -958,58 +1030,68 @@ and occur_existential evm c =
| _ -> iter_constr occrec c
in try occrec c; false with Occur -> true
-and evar_define ?(choose=false) env (evk,_ as ev) rhs evd =
+and evar_define ?(choose=false) env (evk,argsv as ev) rhs evd =
try
let (evd',body) = invert_definition choose env evd ev rhs in
if occur_meta body then error "Meta cannot occur in evar body.";
(* invert_definition may have instantiate some evars of rhs with evk *)
(* so we recheck acyclicity *)
- if occur_evar evk body then error_occur_check env (evars_of evd) evk body;
+ if occur_evar evk body then raise (OccurCheckIn (evd',body));
(* needed only if an inferred type *)
let body = refresh_universes body in
(* Cannot strictly type instantiations since the unification algorithm
* does not unify applications from left to right.
- * e.g problem f x == g y yields x==y and f==g (in that order)
+ * e.g problem f x == g y yields x==y and f==g (in that order)
* Another problem is that type variables are evars of type Type
let _ =
try
let env = evar_env evi in
let ty = evi.evar_concl in
- Typing.check env (evars_of evd') body ty
+ Typing.check env evd' body ty
with e ->
pperrnl
(str "Ill-typed evar instantiation: " ++ fnl() ++
- pr_evar_defs evd' ++ fnl() ++
+ pr_evar_map evd' ++ fnl() ++
str "----> " ++ int ev ++ str " := " ++
print_constr body);
raise e in*)
- Evd.evar_define evk body evd'
+ Evd.define evk body evd'
with
| NotEnoughInformationToProgress ->
postpone_evar_term env evd ev rhs
- | NotInvertibleUsingOurAlgorithm t ->
- error_not_clean env (evars_of evd) evk t (evar_source evk evd)
+ | NotInvertibleUsingOurAlgorithm t ->
+ error_not_clean env evd evk t (evar_source evk evd)
+ | OccurCheckIn (evd,rhs) ->
+ (* last chance: rhs actually reduces to ev *)
+ let c = whd_betadeltaiota env evd rhs in
+ match kind_of_term c with
+ | Evar (evk',argsv2) when evk = evk' ->
+ solve_refl
+ (fun env sigma pb c c' -> (evd,is_fconv pb env sigma c c'))
+ env evd evk argsv argsv2
+ | _ ->
+ error_occur_check env evd evk rhs
(*-------------------*)
(* Auxiliary functions for the conversion algorithms modulo evars
*)
-let has_undefined_evars evd t =
- let evm = evars_of evd in
+let has_undefined_evars_or_sorts evd t =
let rec has_ev t =
match kind_of_term t with
- Evar (ev,args) ->
- (match evar_body (Evd.find evm ev) with
- | Evar_defined c ->
- has_ev c; Array.iter has_ev args
- | Evar_empty ->
- raise NotInstantiatedEvar)
- | _ -> iter_constr has_ev t in
+ | Evar (ev,args) ->
+ (match evar_body (Evd.find evd ev) with
+ | Evar_defined c ->
+ has_ev c; Array.iter has_ev args
+ | Evar_empty ->
+ raise NotInstantiatedEvar)
+ | Sort s when is_sort_variable evd s -> raise Not_found
+ | _ -> iter_constr has_ev t in
try let _ = has_ev t in false
with (Not_found | NotInstantiatedEvar) -> true
let is_ground_term evd t =
- not (has_undefined_evars evd t)
+ not (has_undefined_evars_or_sorts evd t)
let is_ground_env evd env =
let is_ground_decl = function
@@ -1021,15 +1103,35 @@ let is_ground_env evd env =
structures *)
let is_ground_env = memo1_2 is_ground_env
-let head_evar =
+(* Return the head evar if any *)
+
+exception NoHeadEvar
+
+let head_evar =
let rec hrec c = match kind_of_term c with
| Evar (evk,_) -> evk
| Case (_,_,c,_) -> hrec c
| App (c,_) -> hrec c
| Cast (c,_,_) -> hrec c
- | _ -> failwith "headconstant"
- in
- hrec
+ | _ -> raise NoHeadEvar
+ in
+ hrec
+
+(* Expand head evar if any (currently consider only applications but I
+ guess it should consider Case too) *)
+
+let whd_head_evar_stack sigma c =
+ let rec whrec (c, l as s) =
+ match kind_of_term c with
+ | Evar (evk,args as ev) when Evd.is_defined sigma evk
+ -> whrec (existential_value sigma ev, l)
+ | Cast (c,_,_) -> whrec (c, l)
+ | App (f,args) -> whrec (f, Array.fold_right (fun a l -> a::l) args l)
+ | _ -> s
+ in
+ whrec (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
@@ -1050,8 +1152,9 @@ let rec expand_and_check_vars env = function
let is_unification_pattern_evar env (_,args) l t =
List.for_all (fun x -> isRel x || isVar x) l (* common failure case *)
&&
+ let aliases = make_alias_map env in
let l' = Array.to_list args @ l in
- let l'' = try Some (expand_and_check_vars env l') with Exit -> None in
+ let l'' = try Some (expand_and_check_vars aliases l') with Exit -> None in
match l'' with
| Some l ->
let deps =
@@ -1060,7 +1163,7 @@ let is_unification_pattern_evar env (_,args) l t =
l
else
(* Probably strong restrictions coming from t being evar-closed *)
- let t = expand_vars_in_term env t in
+ 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 ->
@@ -1083,16 +1186,19 @@ let is_unification_pattern (env,nb) f l t =
(* 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
+ *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 env) l1 in
+ let l1 = List.map (expand_var (make_alias_map env)) l1 in
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'
- | _ -> assert false)
+ | _ -> assert false)
l1 c in
(* Warning: we may miss some opportunity to eta-reduce more since c'
is not in normal form *)
@@ -1111,7 +1217,7 @@ let solve_pattern_eqn env l1 c =
* hyps of the existential, to do a "pop" for each Rel which is
* not an argument of the existential, and a subst1 for each which
* is, again, with the corresponding variable. This is done by
- * evar_define
+ * define
*
* Thus, we take the arguments of the existential which we are about
* to assign, and zip them with the identifiers in the hypotheses.
@@ -1125,43 +1231,22 @@ let solve_pattern_eqn env l1 c =
*)
let status_changed lev (pbty,_,t1,t2) =
- try
- List.mem (head_evar t1) lev or List.mem (head_evar t2) lev
- with Failure _ ->
- try List.mem (head_evar t2) lev with Failure _ -> false
-
-(* 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
- * that don't unify are discarded (i.e. ?i is redefined so that it does not
- * depend on these args). *)
-
-let solve_refl conv_algo env evd evk argsv1 argsv2 =
- if argsv1 = argsv2 then evd else
- let evi = Evd.find (evars_of evd) evk in
- (* Filter and restrict if needed *)
- let evd,evk,args =
- restrict_upon_filter evd evi evk
- (fun (a1,a2) -> snd (conv_algo env evd Reduction.CONV a1 a2))
- (List.combine (Array.to_list argsv1) (Array.to_list argsv2)) in
- (* Leave a unification problem *)
- 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
+ (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 (evars_of evd) t2 in
- if has_undefined_evars evd t2 then
+ 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 (evars_of evd) (refresh_universes t2) in
+ 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 (evars_of evd) ev1 in
+ 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),"",
@@ -1175,7 +1260,7 @@ let check_instance_type conv_algo env evd ev1 t2 =
(* Rq: uncomplete algorithm if pbty = CONV_X_LEQ ! *)
let solve_simple_eqn conv_algo ?(choose=false) env evd (pbty,(evk1,args1 as ev1),t2) =
try
- let t2 = whd_evar (evars_of evd) t2 in
+ 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
@@ -1190,20 +1275,19 @@ let solve_simple_eqn conv_algo ?(choose=false) env evd (pbty,(evk1,args1 as ev1)
| _ ->
let evd =
if pbty = Some false then
- check_instance_type conv_algo env evd ev1 t2
+ check_instance_type conv_algo env evd ev1 t2
else
evd in
let evd = evar_define ~choose env ev1 t2 evd in
- let evm = evars_of evd in
- let evi = Evd.find evm evk1 in
- if occur_existential evm evi.evar_concl then
+ 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_isevar evd evi.evar_concl in
- match evi.evar_body with
- | Evar_defined body ->
- let ty = nf_isevar evd (Retyping.get_type_of_with_meta evenv evm (metas_of evd) body) 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 *)
+ | Evar_empty -> (* Resulted in a constraint *)
evd
else evd
in
@@ -1215,45 +1299,58 @@ let solve_simple_eqn conv_algo ?(choose=false) env evd (pbty,(evk1,args1 as ev1)
with e when precatchable_exception e ->
(evd,false)
-let evars_of_term c =
+let evars_of_term c =
let rec evrec acc c =
match kind_of_term c with
| Evar (n, _) -> Intset.add n acc
| _ -> fold_constr evrec acc c
- in
+ in
evrec Intset.empty c
let evars_of_named_context nc =
List.fold_right (fun (_, b, t) s ->
- Option.fold_left (fun s t ->
+ Option.fold_left (fun s t ->
Intset.union s (evars_of_term t))
- s b) nc Intset.empty
+ (Intset.union s (evars_of_term t)) b)
+ nc Intset.empty
let evars_of_evar_info evi =
Intset.union (evars_of_term evi.evar_concl)
- (Intset.union
- (match evi.evar_body with
+ (Intset.union
+ (match evi.evar_body with
| Evar_empty -> Intset.empty
| Evar_defined b -> evars_of_term b)
(evars_of_named_context (named_context_of_val evi.evar_hyps)))
-
+
(* [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 = evars_of evd in
+ let sigma = evd in
let c = nf_evar sigma c in
let rec proc_rec c =
match kind_of_term c with
| Evar (evk,args) ->
assert (Evd.mem sigma evk);
if not (Evd.mem initial_sigma evk) then
- let (loc,k) = evar_source evk evd in
- let evi = nf_evar_info sigma (Evd.find sigma evk) in
- error_unsolvable_implicit loc env sigma evi k None
+ let (loc,k) = evar_source evk sigma in
+ (match k with
+ | ImplicitArg (gr, (i, id), false) -> ()
+ | _ ->
+ let evi = nf_evar_info sigma (Evd.find 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)
+
(* Operations on value/type constraints *)
type type_constraint_type = (int * int) option * constr
@@ -1299,7 +1396,7 @@ let mk_valcon c = Some c
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 (evars_of evd) ev in
+ let evi = Evd.find evd ev in
let evenv = evar_unfiltered_env evi in
let (evd1,dom) = new_evar evd evenv (new_Type()) ~filter:(evar_filter evi) in
let nvar =
@@ -1307,16 +1404,16 @@ let define_evar_as_abstraction abs evd (ev,args) =
(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())
+ 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.evar_define ev prod evd2 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)
@@ -1325,7 +1422,7 @@ let define_evar_as_lambda evd (ev,args) =
let define_evar_as_sort evd (ev,args) =
let s = new_Type () in
- Evd.evar_define ev s evd, destSort s
+ Evd.define ev s evd, destSort 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... *)
@@ -1337,44 +1434,44 @@ let judge_of_new_Type () = Typeops.judge_of_type (new_univ ())
constraint on its domain and codomain. If the input constraint is
an evar instantiate it with the product of 2 new evars. *)
-let split_tycon loc env evd tycon =
- let rec real_split evd c =
- let t = whd_betadeltaiota env (Evd.evars_of evd) c in
+let split_tycon loc env evd tycon =
+ let rec real_split evd c =
+ 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) ->
let (evd',prod) = define_evar_as_product evd ev in
let (_,dom,rng) = destProd prod in
evd',(Anonymous, dom, rng)
- | _ -> error_not_product_loc loc env (Evd.evars_of evd) c
+ | _ -> error_not_product_loc loc env evd c
in
match tycon with
| None -> evd,(Anonymous,None,None)
| Some (abs, c) ->
(match abs with
- None ->
+ None ->
let evd', (n, dom, rng) = real_split evd c in
evd', (n, mk_tycon dom, mk_tycon rng)
| Some (init, cur) ->
- if cur = 0 then
+ if cur = 0 then
let evd', (x, dom, rng) = real_split evd c in
- evd, (Anonymous,
- Some (None, dom),
+ evd, (Anonymous,
+ Some (None, dom),
Some (None, rng))
else
- evd, (Anonymous, None,
+ evd, (Anonymous, None,
Some (if cur = 1 then None,c else Some (init, pred cur), c)))
-
-let valcon_of_tycon x =
+
+let valcon_of_tycon x =
match x with
| Some (None, t) -> Some t
| _ -> None
-
+
let lift_abstr_tycon_type n (abs, t) =
- match abs with
+ match abs with
None -> raise (Invalid_argument "lift_abstr_tycon_type: not an abstraction")
| Some (init, abs) ->
- let abs' = abs + n in
+ let abs' = abs + n in
if abs' < 0 then raise (Invalid_argument "lift_abstr_tycon_type")
else (Some (init, abs'), t)
@@ -1382,11 +1479,10 @@ let lift_tycon_type n (abs, t) = (abs, lift n t)
let lift_tycon n = Option.map (lift_tycon_type n)
let pr_tycon_type env (abs, t) =
- match abs with
+ match abs with
None -> Termops.print_constr_env env t
| Some (init, cur) -> str "Abstract (" ++ int init ++ str "," ++ int cur ++ str ") " ++ Termops.print_constr_env env t
-
+
let pr_tycon env = function
None -> str "None"
| Some t -> pr_tycon_type env t
-
diff --git a/pretyping/evarutil.mli b/pretyping/evarutil.mli
index eef41da3..283867e8 100644
--- a/pretyping/evarutil.mli
+++ b/pretyping/evarutil.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: evarutil.mli 12268 2009-08-11 09:02:16Z herbelin $ i*)
+(*i $Id$ i*)
(*i*)
open Util
@@ -34,19 +34,19 @@ val new_untyped_evar : unit -> existential_key
(***********************************************************)
(* Creating a fresh evar given their type and context *)
val new_evar :
- evar_defs -> env -> ?src:loc * hole_kind -> ?filter:bool list -> types -> evar_defs * constr
+ evar_map -> env -> ?src:loc * hole_kind -> ?filter:bool list -> types -> evar_map * constr
(* the same with side-effects *)
val e_new_evar :
- evar_defs ref -> env -> ?src:loc * hole_kind -> ?filter:bool list -> types -> constr
+ evar_map ref -> env -> ?src:loc * hole_kind -> ?filter:bool list -> types -> 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
+ 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_defs -> types -> ?src:loc * hole_kind -> ?filter:bool list -> constr list -> evar_defs * constr
+ named_context_val -> evar_map -> types -> ?src:loc * hole_kind -> ?filter:bool list -> constr list -> evar_map * constr
val make_pure_subst : evar_info -> constr array -> (identifier * constr) list
@@ -57,7 +57,7 @@ val make_pure_subst : evar_info -> constr array -> (identifier * constr) list
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_defs -> evar_defs
+val evar_define : ?choose:bool -> env -> existential -> constr -> evar_map -> evar_map
(***********************************************************)
(* Evars/Metas switching... *)
@@ -72,26 +72,33 @@ val non_instantiated : evar_map -> (evar * evar_info) list
(***********************************************************)
(* Unification utils *)
-val is_ground_term : evar_defs -> constr -> bool
-val is_ground_env : evar_defs -> env -> bool
-val solve_refl :
- (env -> evar_defs -> conv_pb -> constr -> constr -> evar_defs * bool)
- -> env -> evar_defs -> existential_key -> constr array -> constr array ->
- evar_defs
+exception 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_defs -> conv_pb -> constr -> constr -> evar_defs * bool)
- -> ?choose:bool -> env -> evar_defs -> bool option * existential * constr ->
- evar_defs * bool
+ (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
new unresolved evar remains in [c] *)
-val check_evars : env -> evar_map -> evar_defs -> constr -> unit
+val check_evars : env -> evar_map -> evar_map -> constr -> unit
-val define_evar_as_product : evar_defs -> existential -> evar_defs * types
-val define_evar_as_lambda : evar_defs -> existential -> evar_defs * types
-val define_evar_as_sort : evar_defs -> existential -> evar_defs * sorts
+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_sort : evar_map -> existential -> evar_map * sorts
-val is_unification_pattern_evar : env -> existential -> constr list ->
+val is_unification_pattern_evar : env -> existential -> constr list ->
constr -> bool
val is_unification_pattern : env * int -> constr -> constr array ->
constr -> bool
@@ -120,8 +127,8 @@ val empty_valcon : val_constraint
val mk_valcon : constr -> val_constraint
val split_tycon :
- loc -> env -> evar_defs -> type_constraint ->
- evar_defs * (name * type_constraint * type_constraint)
+ loc -> env -> evar_map -> type_constraint ->
+ evar_map * (name * type_constraint * type_constraint)
val valcon_of_tycon : type_constraint -> val_constraint
@@ -132,8 +139,8 @@ val lift_tycon : int -> type_constraint -> type_constraint
(***********************************************************)
-(* [whd_ise] raise [Uninstantiated_evar] if an evar remains uninstantiated; *)
-(* *[whd_evar]* and *[nf_evar]* leave uninstantiated evar as is *)
+(* [flush_and_check_evars] raise [Uninstantiated_evar] if an evar remains *)
+(* uninstantiated; [nf_evar] leave uninstantiated evars as is *)
val nf_evar : evar_map -> constr -> constr
val j_nf_evar : evar_map -> unsafe_judgment -> unsafe_judgment
@@ -151,26 +158,16 @@ 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
-(* Same for evar defs *)
-val nf_isevar : evar_defs -> constr -> constr
-val j_nf_isevar : evar_defs -> unsafe_judgment -> unsafe_judgment
-val jl_nf_isevar :
- evar_defs -> unsafe_judgment list -> unsafe_judgment list
-val jv_nf_isevar :
- evar_defs -> unsafe_judgment array -> unsafe_judgment array
-val tj_nf_isevar :
- evar_defs -> unsafe_type_judgment -> unsafe_type_judgment
+val nf_evar_map : evar_map -> evar_map
-val nf_evar_defs : evar_defs -> evar_defs
-
-(* Replacing all evars *)
+(* Replacing all evars, possibly raising [Uninstantiated_evar] *)
+(* exception Uninstantiated_evar of existential_key *)
exception Uninstantiated_evar of existential_key
-val whd_ise : evar_map -> constr -> constr
-val whd_castappevar : evar_map -> constr -> constr
+val flush_and_check_evars : evar_map -> constr -> constr
(* Replace the vars and rels that are aliases to other vars and rels by *)
(* their representative that is most ancient in the context *)
-val expand_vars_in_term : env -> constr -> constr
+val expand_vars_in_term : env -> constr -> constr
(*********************************************************************)
(* debug pretty-printer: *)
@@ -189,5 +186,8 @@ type clear_dependency_error =
exception ClearDependencyError of identifier * clear_dependency_error
-val clear_hyps_in_evi : evar_defs ref -> named_context_val -> types ->
+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
diff --git a/pretyping/evd.ml b/pretyping/evd.ml
index af070d7e..21753d3a 100644
--- a/pretyping/evd.ml
+++ b/pretyping/evd.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: evd.ml 11865 2009-01-28 17:34:30Z herbelin $ *)
+(* $Id$ *)
open Pp
open Util
@@ -19,12 +19,30 @@ open Environ
open Libnames
open Mod_subst
+(* The kinds of existential variable *)
+
+type obligation_definition_status = Define of bool | Expand
+
+type hole_kind =
+ | ImplicitArg of global_reference * (int * identifier option) * bool
+ | BinderType of name
+ | QuestionMark of obligation_definition_status
+ | CasesType
+ | InternalHole
+ | TomatchTypeParameter of inductive * int
+ | GoalEvar
+ | ImpossibleCase
+ | MatchingVar of bool * identifier
+
(* The type of mappings for existential variables *)
type evar = existential_key
+let string_of_existential evk = "?" ^ string_of_int evk
+let existential_of_int evk = evk
+
type evar_body =
- | Evar_empty
+ | Evar_empty
| Evar_defined of constr
type evar_info = {
@@ -32,6 +50,7 @@ type evar_info = {
evar_hyps : named_context_val;
evar_body : evar_body;
evar_filter : bool list;
+ evar_source : hole_kind located;
evar_extra : Dyn.t option}
let make_evar hyps ccl = {
@@ -39,6 +58,7 @@ let make_evar hyps ccl = {
evar_hyps = hyps;
evar_body = Evar_empty;
evar_filter = List.map (fun _ -> true) (named_context_of_val hyps);
+ evar_source = (dummy_loc,InternalHole);
evar_extra = None
}
@@ -48,109 +68,121 @@ let evar_context evi = named_context_of_val evi.evar_hyps
let evar_body evi = evi.evar_body
let evar_filter evi = evi.evar_filter
let evar_unfiltered_env evi = Global.env_of_context evi.evar_hyps
-let evar_filtered_context evi =
+let evar_filtered_context evi =
snd (list_filter2 (fun b c -> b) (evar_filter evi,evar_context evi))
-let evar_env evi =
+let evar_env evi =
List.fold_right push_named (evar_filtered_context evi)
(reset_context (Global.env()))
let eq_evar_info ei1 ei2 =
- ei1 == ei2 ||
- eq_constr ei1.evar_concl ei2.evar_concl &&
+ ei1 == ei2 ||
+ eq_constr ei1.evar_concl ei2.evar_concl &&
eq_named_context_val (ei1.evar_hyps) (ei2.evar_hyps) &&
ei1.evar_body = ei2.evar_body
-module Evarmap = Intmap
-
-type evar_map1 = evar_info Evarmap.t
-
-let empty = Evarmap.empty
-
-let to_list evc = (* Workaround for change in Map.fold behavior *)
- let l = ref [] in
- Evarmap.iter (fun evk x -> l := (evk,x)::!l) evc;
- !l
-
-let dom evc = Evarmap.fold (fun evk _ acc -> evk::acc) evc []
-let find evc k = Evarmap.find k evc
-let remove evc k = Evarmap.remove k evc
-let mem evc k = Evarmap.mem k evc
-let fold = Evarmap.fold
-
-let add evd evk newinfo = Evarmap.add evk newinfo evd
-
-let define evd evk body =
- let oldinfo =
- try find evd evk
- with Not_found -> error "Evd.define: cannot define undeclared evar" in
- let newinfo =
- { oldinfo with
- evar_body = Evar_defined body } in
- match oldinfo.evar_body with
- | Evar_empty -> Evarmap.add evk newinfo evd
- | _ -> anomaly "Evd.define: cannot define an evar twice"
-
-let is_evar sigma evk = mem sigma evk
-
-let is_defined sigma evk =
- let info = find sigma evk in
- not (info.evar_body = Evar_empty)
-
-let string_of_existential evk = "?" ^ string_of_int evk
-
-let existential_of_int evk = evk
-
-(*******************************************************************)
-(* 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. *)
+(* spiwack: Revised hierarchy :
+ - ExistentialMap ( Maps of existential_keys )
+ - EvarInfoMap ( .t = evar_info ExistentialMap.t )
+ - EvarMap ( .t = EvarInfoMap.t * sort_constraints )
+ - evar_map (exported)
+*)
-let existential_type sigma (n,args) =
- let info =
- try find sigma n
- with Not_found ->
- anomaly ("Evar "^(string_of_existential n)^" was not declared") in
- let hyps = evar_filtered_context info in
- instantiate_evar hyps info.evar_concl (Array.to_list args)
+module ExistentialMap = Intmap
+module ExistentialSet = Intset
+(* This exception is raised by *.existential_value *)
exception NotInstantiatedEvar
-let existential_value sigma (n,args) =
- let info = find sigma n in
- let hyps = evar_filtered_context info in
- match evar_body info with
- | Evar_defined c ->
- instantiate_evar hyps c (Array.to_list args)
- | Evar_empty ->
- raise NotInstantiatedEvar
+module EvarInfoMap = struct
+ type t = evar_info ExistentialMap.t
+
+ let empty = ExistentialMap.empty
+
+ let to_list evc = (* Workaround for change in Map.fold behavior *)
+ let l = ref [] in
+ ExistentialMap.iter (fun evk x -> l := (evk,x)::!l) evc;
+ !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 add evd evk newinfo = ExistentialMap.add evk newinfo evd
+
+ let equal = ExistentialMap.equal
+
+ let define evd evk body =
+ let oldinfo =
+ try find evd evk
+ with Not_found -> error "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
+
+ let is_defined sigma evk =
+ let info = find sigma evk in
+ not (info.evar_body = Evar_empty)
+
+
+ (*******************************************************************)
+ (* 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) =
+ let info =
+ try find sigma n
+ with Not_found ->
+ anomaly ("Evar "^(string_of_existential n)^" was not declared") in
+ let hyps = evar_filtered_context info in
+ instantiate_evar hyps info.evar_concl (Array.to_list args)
+
+ let existential_value sigma (n,args) =
+ let info = find sigma n in
+ let hyps = evar_filtered_context info in
+ match evar_body info with
+ | Evar_defined c ->
+ instantiate_evar hyps c (Array.to_list args)
+ | Evar_empty ->
+ raise NotInstantiatedEvar
+
+ let existential_opt_value sigma ev =
+ try Some (existential_value sigma ev)
+ with NotInstantiatedEvar -> None
-let existential_opt_value sigma ev =
- try Some (existential_value sigma ev)
- with NotInstantiatedEvar -> None
+end
(*******************************************************************)
(* Constraints for sort variables *)
@@ -163,11 +195,8 @@ type sort_constraint =
| SortVar of sort_var list * sort_var list (* (leq,geq) *)
| EqSort of sort_var
-module UniverseOrdered = struct
- type t = Univ.universe
- let compare = Pervasives.compare
-end
-module UniverseMap = Map.Make(UniverseOrdered)
+module UniverseMap =
+ Map.Make (struct type t = Univ.universe let compare = compare end)
type sort_constraints = sort_constraint UniverseMap.t
@@ -236,7 +265,7 @@ let set_leq_sort (u1,(leq1,geq1)) (u2,(leq2,geq2)) scstr =
match UniverseMap.find u1 scstr with
EqSort u1' -> search_rec (is_b,betw,not_betw) u1'
| SortVar(leq,_) ->
- let (is_b',betw',not_betw') =
+ 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')
@@ -285,41 +314,33 @@ let pr_sort_cstrs g =
hov 0 (prlist_with_sep spc Univ.pr_uni leq) ++ str"]")
l
-type evar_map = evar_map1 * sort_constraints
-let empty = empty, UniverseMap.empty
-let add (sigma,sm) k v = (add sigma k v, sm)
-let dom (sigma,_) = dom sigma
-let find (sigma,_) = find sigma
-let remove (sigma,sm) k = (remove sigma k, sm)
-let mem (sigma,_) = mem sigma
-let to_list (sigma,_) = to_list sigma
-let fold f (sigma,_) = fold f sigma
-let define (sigma,sm) k v = (define sigma k v, sm)
-let is_evar (sigma,_) = is_evar sigma
-let is_defined (sigma,_) = is_defined sigma
-let existential_value (sigma,_) = existential_value sigma
-let existential_type (sigma,_) = existential_type sigma
-let existential_opt_value (sigma,_) = existential_opt_value sigma
-let eq_evar_map x y = x == y ||
- (Evarmap.equal eq_evar_info (fst x) (fst y) &&
- UniverseMap.equal (=) (snd x) (snd y))
-
-let merge e e' = fold (fun n v sigma -> add sigma n v) e' e
+module EvarMap = struct
+ type t = EvarInfoMap.t * sort_constraints
+ let empty = EvarInfoMap.empty, UniverseMap.empty
+ let add (sigma,sm) k v = (EvarInfoMap.add sigma k v, sm)
+ let dom (sigma,_) = EvarInfoMap.dom sigma
+ let find (sigma,_) = EvarInfoMap.find 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 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 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
-(*******************************************************************)
-type open_constr = evar_map * constr
+end
(*******************************************************************)
-(* The type constructor ['a sigma] adds an evar map to an object of
- type ['a] *)
-type 'a sigma = {
- it : 'a ;
- sigma : evar_map}
-
-let sig_it x = x.it
-let sig_sig x = x.sigma
-
-(*******************************************************************)
(* Metamaps *)
(*******************************************************************)
@@ -390,65 +411,113 @@ let clb_name = function
| Clval (na,_,_) -> (na,true)
(***********************)
-
+
module Metaset = Intset
-
+
let meta_exists p s = Metaset.fold (fun x b -> b || (p x)) s false
module Metamap = Intmap
let metamap_to_list m =
Metamap.fold (fun n v l -> (n,v)::l) m []
-
+
(*************************)
(* Unification state *)
-type obligation_definition_status = Define of bool | Expand
-
-type hole_kind =
- | ImplicitArg of global_reference * (int * identifier option)
- | BinderType of name
- | QuestionMark of obligation_definition_status
- | CasesType
- | InternalHole
- | TomatchTypeParameter of inductive * int
- | GoalEvar
- | ImpossibleCase
-
type conv_pb = Reduction.conv_pb
type evar_constraint = conv_pb * Environ.env * constr * constr
-type evar_defs =
- { evars : evar_map;
+type evar_map =
+ { evars : EvarMap.t;
conv_pbs : evar_constraint list;
- last_mods : existential_key list;
- history : (existential_key * (loc * hole_kind)) list;
+ last_mods : ExistentialSet.t;
metas : clbinding Metamap.t }
+(*** Lifting primitive from EvarMap. ***)
+
+(* HH: The progress tactical now uses this function. *)
+let progress_evar_map d1 d2 =
+ EvarMap.progress_evar_map d1.evars d2.evars
+
+(* spiwack: tentative. It might very well not be the semantics we want
+ for merging evar_map *)
+let merge d1 d2 = {
+(* 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 ;
+ metas = Metamap.fold (fun k m r -> Metamap.add k m r) d2.metas d1.metas
+}
+let add d e i = { d with evars=EvarMap.add d.evars e i }
+let remove d e = { d with evars=EvarMap.remove d.evars e }
+let dom d = EvarMap.dom d.evars
+let find d e = EvarMap.find 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
+(* 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 is_evar d e = EvarMap.is_evar d.evars e
+let is_defined d e = EvarMap.is_defined 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
+
+(*** /Lifting... ***)
+
+(* evar_map are considered empty disregarding histories *)
+let is_empty d =
+ d.evars = EvarMap.empty &&
+ d.conv_pbs = [] &&
+ Metamap.is_empty d.metas
+
+let subst_named_context_val s = map_named_val (subst_mps s)
+
+let subst_evar_info s evi =
+ let subst_evb = function Evar_empty -> Evar_empty
+ | Evar_defined c -> Evar_defined (subst_mps s c) in
+ { evi with
+ evar_concl = subst_mps s evi.evar_concl;
+ evar_hyps = subst_named_context_val s evi.evar_hyps;
+ evar_body = subst_evb evi.evar_body }
+
let subst_evar_defs_light sub evd =
- assert (evd.evars = (Evarmap.empty,UniverseMap.empty));
+ assert (UniverseMap.is_empty (snd evd.evars));
assert (evd.conv_pbs = []);
{ evd with
- metas = Metamap.map (map_clb (subst_mps sub)) evd.metas }
-
-let create_evar_defs sigma =
- { evars=sigma; conv_pbs=[]; last_mods = []; history=[]; metas=Metamap.empty }
-let create_goal_evar_defs sigma =
- let h = fold (fun mv _ l -> (mv,(dummy_loc,GoalEvar))::l) sigma [] in
- { evars=sigma; conv_pbs=[]; last_mods = []; history=h; metas=Metamap.empty }
-let empty_evar_defs = create_evar_defs empty
-let evars_of d = d.evars
-let evars_reset_evd evd d = {d with evars = evd}
-let reset_evd (sigma,mmap) d = {d with evars = sigma; metas=mmap}
+ metas = Metamap.map (map_clb (subst_mps sub)) evd.metas;
+ evars = ExistentialMap.map (subst_evar_info sub) (fst evd.evars), snd evd.evars
+ }
+
+let subst_evar_map = subst_evar_defs_light
+
+(* spiwack: deprecated *)
+let create_evar_defs sigma = { sigma with
+ conv_pbs=[]; last_mods=ExistentialSet.empty; metas=Metamap.empty }
+(* spiwack: tentatively deprecated *)
+let create_goal_evar_defs sigma = { sigma with
+ conv_pbs=[]; last_mods=ExistentialSet.empty; metas=Metamap.empty }
+let empty = {
+ evars=EvarMap.empty;
+ conv_pbs=[];
+ last_mods = ExistentialSet.empty;
+ metas=Metamap.empty
+}
+
+let evars_reset_evd evd d = {d with evars = evd.evars}
let add_conv_pb pb d = {d with conv_pbs = pb::d.conv_pbs}
-let evar_source evk d =
- try List.assoc evk d.history
- with Not_found -> (dummy_loc, InternalHole)
+let evar_source evk d = (EvarMap.find d.evars evk).evar_source
(* define the existential of section path sp as the constr body *)
-let evar_define evk body evd =
+let define evk body evd =
{ evd with
- evars = define evd.evars evk body;
- last_mods = evk :: evd.last_mods }
+ evars = EvarMap.define evd.evars evk body;
+ last_mods =
+ match evd.conv_pbs with
+ | [] -> evd.last_mods
+ | _ -> ExistentialSet.add evk evd.last_mods }
let evar_declare hyps evk ty ?(src=(dummy_loc,InternalHole)) ?filter evd =
let filter =
@@ -460,43 +529,43 @@ let evar_declare hyps evk ty ?(src=(dummy_loc,InternalHole)) ?filter evd =
filter)
in
{ evd with
- evars = add evd.evars evk
+ evars = EvarMap.add evd.evars evk
{evar_hyps = hyps;
evar_concl = ty;
evar_body = Evar_empty;
evar_filter = filter;
- evar_extra = None};
- history = (evk,src)::evd.history }
+ evar_source = src;
+ evar_extra = None} }
-let is_defined_evar evd (evk,_) = is_defined evd.evars evk
+let is_defined_evar evd (evk,_) = EvarMap.is_defined evd.evars evk
(* Does k corresponds to an (un)defined existential ? *)
let is_undefined_evar evd c = match kind_of_term c with
| Evar ev -> not (is_defined_evar evd ev)
| _ -> false
-let undefined_evars evd =
- let evars =
- fold (fun evk evi sigma -> if evi.evar_body = Evar_empty then
- add sigma evk evi else sigma)
- evd.evars empty
- in
+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 =
- let (pbs,pbs1) =
+ let (pbs,pbs1) =
List.fold_left
(fun (pbs,pbs1) pb ->
- if p pb then
+ if p pb then
(pb::pbs,pbs1)
- else
+ else
(pbs,pb::pbs1))
([],[])
evd.conv_pbs
in
- {evd with conv_pbs = pbs1; last_mods = []},
+ {evd with conv_pbs = pbs1; last_mods = ExistentialSet.empty},
pbs
let extract_changed_conv_pbs evd p =
@@ -505,23 +574,31 @@ let extract_changed_conv_pbs evd p =
let extract_all_conv_pbs evd =
extract_conv_pbs evd (fun _ -> true)
+(* spiwack: should it be replaced by Evd.merge? *)
let evar_merge evd evars =
- { evd with evars = merge evd.evars evars }
+ { evd with evars = EvarMap.merge evd.evars evars.evars }
+
+let evar_list evd c =
+ let rec evrec acc c =
+ match kind_of_term c with
+ | Evar (evk, _ as ev) when mem evd evk -> ev :: acc
+ | _ -> fold_constr evrec acc c in
+ evrec [] c
(**********************************************************)
(* Sort variables *)
-let new_sort_variable (sigma,sm) =
+let new_sort_variable ({ evars = (sigma,sm) } as d)=
let (u,scstr) = new_sort_var sm in
- (Type u,(sigma,scstr))
-let is_sort_variable (_,sm) s =
+ (Type u,{ d with evars = (sigma,scstr) } )
+let is_sort_variable {evars=(_,sm)} s =
is_sort_var s sm
-let whd_sort_variable (_,sm) t = whd_sort_var sm t
-let set_leq_sort_variable (sigma,sm) u1 u2 =
- (sigma, set_leq u1 u2 sm)
-let define_sort_variable (sigma,sm) u s =
- (sigma, set_sort_variable u s sm)
-let pr_sort_constraints (_,sm) = pr_sort_cstrs 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
(**********************************************************)
(* Accessing metas *)
@@ -536,7 +613,7 @@ let undefined_metas evd =
| (n,Cltyp (_,typ)) -> n)
(meta_list evd))
-let metas_of evd =
+let metas_of evd =
List.map (function
| (n,Clval(_,_,typ)) -> (n,typ.rebus)
| (n,Cltyp (_,typ)) -> (n,typ.rebus))
@@ -544,8 +621,8 @@ let metas_of evd =
let map_metas_fvalue f evd =
{ evd with metas =
- Metamap.map
- (function
+ Metamap.map
+ (function
| Clval(id,(c,s),typ) -> Clval(id,(mk_freelisted (f c.rebus),s),typ)
| x -> x) evd.metas }
@@ -559,19 +636,28 @@ let meta_defined evd mv =
| Clval _ -> true
| Cltyp _ -> false
-let meta_fvalue evd mv =
+let try_meta_fvalue evd mv =
match Metamap.find mv evd.metas with
| Clval(_,b,_) -> b
- | Cltyp _ -> anomaly "meta_fvalue: meta has no value"
-
+ | Cltyp _ -> raise Not_found
+
+let meta_fvalue evd mv =
+ try try_meta_fvalue evd mv
+ with Not_found -> anomaly "meta_fvalue: meta has no value"
+
+let meta_value evd mv =
+ (fst (try_meta_fvalue evd mv)).rebus
+
let meta_ftype evd mv =
match Metamap.find mv evd.metas with
| Cltyp (_,b) -> b
| Clval(_,_,b) -> b
-
+
+let meta_type evd mv = (meta_ftype evd mv).rebus
+
let meta_declare mv v ?(name=Anonymous) evd =
{ evd with metas = Metamap.add mv (Cltyp(name,mk_freelisted v)) evd.metas }
-
+
let meta_assign mv (v,pb) evd =
match Metamap.find mv evd.metas with
| Cltyp(na,ty) ->
@@ -588,10 +674,7 @@ let meta_reassign mv (v,pb) evd =
(* If the meta is defined then forget its name *)
let meta_name evd mv =
- try
- let (na,def) = clb_name (Metamap.find mv evd.metas) in
- if def then Anonymous else na
- with Not_found -> Anonymous
+ try fst (clb_name (Metamap.find mv evd.metas)) with Not_found -> Anonymous
let meta_with_name evd id =
let na = Name id in
@@ -603,27 +686,28 @@ let meta_with_name evd id =
else l)
evd.metas ([],[]) in
match mvnodef, mvl with
- | _,[] ->
+ | _,[] ->
errorlabstrm "Evd.meta_with_name"
(str"No such bound variable " ++ pr_id id ++ str".")
- | ([n],_|_,[n]) ->
+ | ([n],_|_,[n]) ->
n
- | _ ->
+ | _ ->
errorlabstrm "Evd.meta_with_name"
(str "Binder name \"" ++ pr_id 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)
+ metas = List.fold_left (fun m (n,v) -> Metamap.add n v m)
evd2.metas (metamap_to_list evd1.metas) }
type metabinding = metavariable * constr * instance_status
let retract_coercible_metas evd =
- let mc,ml =
- Metamap.fold (fun n v (mc,ml) ->
+ let mc,ml =
+ Metamap.fold (fun n v (mc,ml) ->
match v with
| Clval (na,(b,(UserGiven,CoerceToType as s)),typ) ->
(n,b.rebus,s)::mc, Metamap.add n (Cltyp (na,typ)) ml
@@ -636,12 +720,25 @@ let rec list_assoc_in_triple x = function
[] -> raise Not_found
| (a,b,_)::l -> if compare a x = 0 then b else list_assoc_in_triple x l
-let subst_defined_metas bl c =
+let subst_defined_metas bl c =
let rec substrec c = match kind_of_term c with
| Meta i -> substrec (list_assoc_in_triple i bl)
| _ -> map_constr substrec c
in try Some (substrec c) with Not_found -> None
+(*******************************************************************)
+type open_constr = evar_map * constr
+
+(*******************************************************************)
+(* The type constructor ['a sigma] adds an evar map to an object of
+ type ['a] *)
+type 'a sigma = {
+ it : 'a ;
+ sigma : evar_map}
+
+let sig_it x = x.it
+let sig_sig x = x.sigma
+
(**********************************************************)
(* Failure explanation *)
@@ -670,20 +767,22 @@ let pr_meta_map mmap =
| _ -> mt() in
let pr_meta_binding = function
| (mv,Cltyp (na,b)) ->
- hov 0
+ hov 0
(pr_meta mv ++ pr_name na ++ str " : " ++
print_constr b.rebus ++ fnl ())
- | (mv,Clval(na,(b,s),_)) ->
- hov 0
+ | (mv,Clval(na,(b,s),t)) ->
+ hov 0
(pr_meta mv ++ pr_name na ++ str " := " ++
- print_constr b.rebus ++ spc () ++ pr_instance_status s ++ fnl ())
+ print_constr b.rebus ++
+ str " : " ++ print_constr t.rebus ++
+ spc () ++ pr_instance_status s ++ fnl ())
in
prlist pr_meta_binding (metamap_to_list mmap)
let pr_decl ((id,b,_),ok) =
match b with
| None -> if ok then pr_id id else (str "{" ++ pr_id id ++ str "}")
- | Some c -> str (if ok then "(" else "{") ++ pr_id id ++ str ":=" ++
+ | Some c -> str (if ok then "(" else "{") ++ pr_id id ++ str ":=" ++
print_constr c ++ str (if ok then ")" else "}")
let pr_evar_info evi =
@@ -697,12 +796,19 @@ let pr_evar_info evi =
in
hov 2 (str"[" ++ phyps ++ spc () ++ str"|- " ++ pty ++ pb ++ str"]")
-let pr_evar_map sigma =
- h 0
- (prlist_with_sep pr_fnl
- (fun (ev,evi) ->
- h 0 (str(string_of_existential ev)++str"=="++ pr_evar_info evi))
- (to_list sigma))
+let pr_evar_map_t (evars,cstrs as sigma) =
+ let evs =
+ if evars = EvarInfoMap.empty 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()
+ and cs =
+ if cstrs = UniverseMap.empty then mt ()
+ else pr_sort_cstrs cstrs++fnl()
+ in evs ++ cs
let pr_constraints pbs =
h 0
@@ -710,19 +816,20 @@ let pr_constraints pbs =
print_constr t1 ++ spc() ++
str (match pbty with
| Reduction.CONV -> "=="
- | Reduction.CUMUL -> "<=") ++
+ | Reduction.CUMUL -> "<=") ++
spc() ++ print_constr t2) pbs)
-let pr_evar_defs evd =
+let pr_evar_map evd =
let pp_evm =
- if evd.evars = empty then mt() else
- str"EVARS:"++brk(0,1)++pr_evar_map evd.evars++fnl() in
+ 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
str"CONSTRAINTS:"++brk(0,1)++pr_constraints evd.conv_pbs++fnl() in
let pp_met =
if evd.metas = Metamap.empty then mt() else
str"METAS:"++brk(0,1)++pr_meta_map evd.metas in
v 0 (pp_evm ++ cstrs ++ pp_met)
-let pr_metaset metas =
+let pr_metaset metas =
str "[" ++ prlist_with_sep spc pr_meta (Metaset.elements metas) ++ str "]"
diff --git a/pretyping/evd.mli b/pretyping/evd.mli
index b9cb2142..46f13d5f 100644
--- a/pretyping/evd.mli
+++ b/pretyping/evd.mli
@@ -7,7 +7,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: evd.mli 11865 2009-01-28 17:34:30Z herbelin $ i*)
+(*i $Id$ i*)
(*i*)
open Util
@@ -20,84 +20,6 @@ open Mod_subst
open Termops
(*i*)
-(* The type of mappings for existential variables.
- The keys are integers and the associated information is a record
- containing the type of the evar ([evar_concl]), the context under which
- it was introduced ([evar_hyps]) and its definition ([evar_body]).
- [evar_info] is used to add any other kind of information. *)
-
-type evar = existential_key
-
-type evar_body =
- | Evar_empty
- | Evar_defined of constr
-
-type evar_info = {
- evar_concl : constr;
- evar_hyps : named_context_val;
- evar_body : evar_body;
- evar_filter : bool list;
- evar_extra : Dyn.t option}
-
-val eq_evar_info : evar_info -> evar_info -> bool
-
-val make_evar : named_context_val -> types -> evar_info
-val evar_concl : evar_info -> constr
-val evar_context : evar_info -> named_context
-val evar_filtered_context : evar_info -> named_context
-val evar_hyps : evar_info -> named_context_val
-val evar_body : evar_info -> evar_body
-val evar_filter : evar_info -> bool list
-val evar_unfiltered_env : evar_info -> env
-val evar_env : evar_info -> env
-
-type evar_map
-val eq_evar_map : evar_map -> evar_map -> bool
-
-val empty : evar_map
-
-val add : evar_map -> evar -> evar_info -> evar_map
-
-val dom : evar_map -> evar list
-val find : evar_map -> evar -> evar_info
-val remove : evar_map -> evar -> evar_map
-val mem : evar_map -> evar -> bool
-val to_list : evar_map -> (evar * evar_info) list
-val fold : (evar -> evar_info -> 'a -> 'a) -> evar_map -> 'a -> 'a
-
-val merge : evar_map -> evar_map -> evar_map
-
-val define : evar_map -> evar -> constr -> evar_map
-
-val is_evar : evar_map -> evar -> bool
-
-val is_defined : evar_map -> evar -> bool
-
-val string_of_existential : evar -> string
-val existential_of_int : int -> evar
-
-(*s [existential_value sigma ev] raises [NotInstantiatedEvar] if [ev] has
- no body and [Not_found] if it does not exist in [sigma] *)
-
-exception NotInstantiatedEvar
-val existential_value : evar_map -> existential -> constr
-val existential_type : evar_map -> existential -> types
-val existential_opt_value : evar_map -> existential -> constr option
-
-(*********************************************************************)
-(* constr with holes *)
-type open_constr = evar_map * constr
-
-(*********************************************************************)
-(* The type constructor ['a sigma] adds an evar map to an object of
- type ['a] *)
-type 'a sigma = {
- it : 'a ;
- sigma : evar_map}
-
-val sig_it : 'a sigma -> 'a
-val sig_sig : 'a sigma -> evar_map
-
(*********************************************************************)
(* Meta map *)
@@ -122,7 +44,7 @@ val map_fl : ('a -> 'b) -> 'a freelisted -> 'b freelisted
(e.g. the solution [P] to [?X u v = P u v] can be eta-expanded twice)
*)
-type instance_constraint =
+type instance_constraint =
IsSuperType | IsSubType | ConvUpToEta of int | UserGiven
(* Status of the unification of the type of an instance against the type of
@@ -153,19 +75,9 @@ type clbinding =
val map_clb : (constr -> constr) -> clbinding -> clbinding
-(*********************************************************************)
-(* Unification state *)
-type evar_defs
-
-(* Assume empty [evar_map] and [conv_pbs] *)
-val subst_evar_defs_light : substitution -> evar_defs -> evar_defs
-(* create an [evar_defs] with empty meta map: *)
-val create_evar_defs : evar_map -> evar_defs
-val create_goal_evar_defs : evar_map -> evar_defs
-val empty_evar_defs : evar_defs
-val evars_of : evar_defs -> evar_map
-val evars_reset_evd : evar_map -> evar_defs -> evar_defs
+(*********************************************************************)
+(*** Kinds of existential variables ***)
(* Should the obligation be defined (opaque or transparent (default)) or
defined transparent and expanded in the term? *)
@@ -174,7 +86,7 @@ type obligation_definition_status = Define of bool | Expand
(* Evars *)
type hole_kind =
- | ImplicitArg of global_reference * (int * identifier option)
+ | ImplicitArg of global_reference * (int * identifier option) * bool (* Force inference *)
| BinderType of name
| QuestionMark of obligation_definition_status
| CasesType
@@ -182,54 +94,148 @@ type hole_kind =
| TomatchTypeParameter of inductive * int
| GoalEvar
| ImpossibleCase
-val is_defined_evar : evar_defs -> existential -> bool
-val is_undefined_evar : evar_defs -> constr -> bool
-val undefined_evars : evar_defs -> evar_defs
+ | MatchingVar of bool * identifier
+
+(*********************************************************************)
+(*** Existential variables and unification states ***)
+
+(* A unification state (of type [evar_map]) is primarily a finite mapping
+ from existential variables to records containing the type of the evar
+ ([evar_concl]), the context under which it was introduced ([evar_hyps])
+ and its definition ([evar_body]). [evar_extra] is used to add any other
+ kind of information.
+ It also contains conversion constraints, debugging information and
+ information about meta variables. *)
+
+(* Information about existential variables. *)
+type evar = existential_key
+
+val string_of_existential : evar -> string
+val existential_of_int : int -> evar
+
+type evar_body =
+ | Evar_empty
+ | Evar_defined of constr
+
+type evar_info = {
+ evar_concl : constr;
+ evar_hyps : named_context_val;
+ evar_body : evar_body;
+ evar_filter : bool list;
+ evar_source : hole_kind located;
+ evar_extra : Dyn.t option}
+
+val eq_evar_info : evar_info -> evar_info -> bool
+
+val make_evar : named_context_val -> types -> evar_info
+val evar_concl : evar_info -> constr
+val evar_context : evar_info -> named_context
+val evar_filtered_context : evar_info -> named_context
+val evar_hyps : evar_info -> named_context_val
+val evar_body : evar_info -> evar_body
+val evar_filter : evar_info -> bool list
+val evar_unfiltered_env : evar_info -> env
+val evar_env : evar_info -> env
+
+(*** Unification state ***)
+type evar_map
+
+(* Unification state and existential variables *)
+
+(* 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
+
+val add : evar_map -> evar -> evar_info -> evar_map
+
+val dom : evar_map -> evar list
+val find : evar_map -> evar -> evar_info
+val remove : evar_map -> evar -> evar_map
+val mem : evar_map -> evar -> bool
+val to_list : evar_map -> (evar * evar_info) list
+val fold : (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
+
+(*s [existential_value sigma ev] raises [NotInstantiatedEvar] if [ev] has
+ no body and [Not_found] if it does not exist in [sigma] *)
+
+exception NotInstantiatedEvar
+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 subst_evar_defs_light : substitution -> evar_map -> evar_map
+
+(* spiwack: this function seems to somewhat break the abstraction. *)
+val evars_reset_evd : evar_map -> evar_map -> evar_map
+
+
+(* spiwack: [is_undefined_evar] should be considered a candidate
+ for moving to evarutils *)
+val is_undefined_evar : evar_map -> constr -> bool
+val undefined_evars : evar_map -> evar_map
val evar_declare :
named_context_val -> evar -> types -> ?src:loc * hole_kind ->
- ?filter:bool list -> evar_defs -> evar_defs
-val evar_define : evar -> constr -> evar_defs -> evar_defs
-val evar_source : existential_key -> evar_defs -> loc * hole_kind
+ ?filter:bool list -> evar_map -> evar_map
+val evar_source : existential_key -> evar_map -> loc * hole_kind
+
+(* 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
-(* [evar_merge evd evars] extends the evars of [evd] with [evars] *)
-val evar_merge : evar_defs -> evar_map -> evar_defs
+val evar_list : evar_map -> constr -> existential list
(* Unification constraints *)
type conv_pb = Reduction.conv_pb
type evar_constraint = conv_pb * env * constr * constr
-val add_conv_pb : evar_constraint -> evar_defs -> evar_defs
-val extract_changed_conv_pbs : evar_defs ->
- (existential_key list -> evar_constraint -> bool) ->
- evar_defs * evar_constraint list
-val extract_all_conv_pbs : evar_defs -> evar_defs * evar_constraint list
+val add_conv_pb : evar_constraint -> evar_map -> evar_map
+
+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
(* Metas *)
-val find_meta : evar_defs -> metavariable -> clbinding
-val meta_list : evar_defs -> (metavariable * clbinding) list
-val meta_defined : evar_defs -> metavariable -> bool
+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 has no value *)
-val meta_fvalue : evar_defs -> metavariable -> constr freelisted * instance_status
-val meta_opt_fvalue : evar_defs -> metavariable -> (constr freelisted * instance_status) option
-val meta_ftype : evar_defs -> metavariable -> constr freelisted
-val meta_name : evar_defs -> metavariable -> name
-val meta_with_name : evar_defs -> identifier -> metavariable
+ meta has no value *)
+val meta_value : evar_map -> metavariable -> constr
+val meta_fvalue : evar_map -> metavariable -> constr freelisted * instance_status
+val meta_opt_fvalue : evar_map -> metavariable -> (constr freelisted * instance_status) option
+val meta_type : evar_map -> metavariable -> types
+val meta_ftype : evar_map -> metavariable -> types freelisted
+val meta_name : evar_map -> metavariable -> name
+val meta_with_name : evar_map -> identifier -> metavariable
val meta_declare :
- metavariable -> types -> ?name:name -> evar_defs -> evar_defs
-val meta_assign : metavariable -> constr * instance_status -> evar_defs -> evar_defs
-val meta_reassign : metavariable -> constr * instance_status -> evar_defs -> evar_defs
+ metavariable -> types -> ?name:name -> evar_map -> evar_map
+val meta_assign : metavariable -> constr * instance_status -> evar_map -> evar_map
+val meta_reassign : metavariable -> constr * instance_status -> evar_map -> evar_map
(* [meta_merge evd1 evd2] returns [evd2] extended with the metas of [evd1] *)
-val meta_merge : evar_defs -> evar_defs -> evar_defs
+val meta_merge : evar_map -> evar_map -> evar_map
-val undefined_metas : evar_defs -> metavariable list
-val metas_of : evar_defs -> metamap
-val map_metas_fvalue : (constr -> constr) -> evar_defs -> evar_defs
+val undefined_metas : evar_map -> metavariable list
+val metas_of : evar_map -> meta_type_map
+val map_metas_fvalue : (constr -> constr) -> evar_map -> evar_map
type metabinding = metavariable * constr * instance_status
-val retract_coercible_metas : evar_defs -> metabinding list * evar_defs
+val retract_coercible_metas : evar_map -> metabinding list * evar_map
val subst_defined_metas : metabinding list -> constr -> constr option
(**********************************************************)
@@ -241,6 +247,20 @@ 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
+(*********************************************************************)
+(* constr with holes *)
+type open_constr = evar_map * constr
+
+(*********************************************************************)
+(* The type constructor ['a sigma] adds an evar map to an object of
+ type ['a] *)
+type 'a sigma = {
+ it : 'a ;
+ sigma : evar_map}
+
+val sig_it : 'a sigma -> 'a
+val sig_sig : 'a sigma -> evar_map
+
(**********************************************************)
(* Failure explanation *)
@@ -250,7 +270,15 @@ type unsolvability_explanation = SeveralInstancesFound of int
(* debug pretty-printer: *)
val pr_evar_info : evar_info -> Pp.std_ppcmds
-val pr_evar_map : evar_map -> Pp.std_ppcmds
-val pr_evar_defs : evar_defs -> Pp.std_ppcmds
+val pr_evar_map : evar_map -> Pp.std_ppcmds
val pr_sort_constraints : evar_map -> Pp.std_ppcmds
val pr_metaset : Metaset.t -> Pp.std_ppcmds
+
+
+(*** /!\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
+val subst_evar_map : substitution -> evar_map -> evar_map
+(*** /Deprecaded ***)
diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml
index 88a0c2a6..1352b383 100644
--- a/pretyping/indrec.ml
+++ b/pretyping/indrec.ml
@@ -6,7 +6,11 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: indrec.ml 11897 2009-02-09 19:28:02Z barras $ *)
+(* $Id$ *)
+
+(* File initially created by Christine Paulin, 1996 *)
+
+(* This file builds various inductive schemes *)
open Pp
open Util
@@ -15,6 +19,7 @@ open Libnames
open Nameops
open Term
open Termops
+open Namegen
open Declarations
open Entries
open Inductive
@@ -27,6 +32,8 @@ open Safe_typing
open Nametab
open Sign
+type dep_flag = bool
+
(* Errors related to recursors building *)
type recursion_scheme_error =
| NotAllowedCaseAnalysis of (*isrec:*) bool * sorts * inductive
@@ -34,7 +41,7 @@ type recursion_scheme_error =
exception RecursionSchemeError of recursion_scheme_error
-let make_prod_dep dep env = if dep then prod_name env else mkProd
+let make_prod_dep dep env = if dep then mkProd_name env else mkProd
let mkLambda_string s t c = mkLambda (Name (id_of_string s), t, c)
(*******************************************)
@@ -43,22 +50,16 @@ let mkLambda_string s t c = mkLambda (Name (id_of_string s), t, c)
(**********************************************************************)
(* Building case analysis schemes *)
-(* Nouvelle version, plus concise mais plus coûteuse à cause de
- lift_constructor et lift_inductive_family qui ne se contentent pas de
- lifter les paramètres globaux *)
+(* Christine Paulin, 1996 *)
-let mis_make_case_com depopt env sigma ind (mib,mip as specif) kind =
+let mis_make_case_com dep env sigma ind (mib,mip as specif) kind =
let lnamespar = mib.mind_params_ctxt in
- let dep = match depopt with
- | None -> inductive_sort_family mip <> InProp
- | Some d -> d
- in
if not (List.mem kind (elim_sorts specif)) then
raise
(RecursionSchemeError
(NotAllowedCaseAnalysis (false,new_sort_in_family kind,ind)));
- let ndepar = mip.mind_nrealargs + 1 in
+ let ndepar = mip.mind_nrealargs_ctxt + 1 in
(* Pas génant car env ne sert pas à typer mais juste à renommer les Anonym *)
(* mais pas très joli ... (mais manque get_sort_of à ce niveau) *)
@@ -67,7 +68,7 @@ let mis_make_case_com depopt env sigma ind (mib,mip as specif) kind =
let indf = make_ind_family(ind, extended_rel_list 0 lnamespar) in
let constrs = get_constructors env indf in
- let rec add_branch env k =
+ let rec add_branch env k =
if k = Array.length mip.mind_consnames then
let nbprod = k+1 in
@@ -82,7 +83,7 @@ let mis_make_case_com depopt env sigma ind (mib,mip as specif) kind =
(mkRel (ndepar + nbprod),
if dep then extended_rel_vect 0 deparsign
else extended_rel_vect 1 arsign) in
- let p =
+ let p =
it_mkLambda_or_LetIn_name env'
((if dep then mkLambda_name env' else mkLambda)
(Anonymous,depind,pbody))
@@ -100,27 +101,28 @@ let mis_make_case_com depopt env sigma ind (mib,mip as specif) kind =
(add_branch (push_rel (Anonymous, None, t) env) (k+1))
in
let typP = make_arity env' dep indf (new_sort_in_family kind) in
- it_mkLambda_or_LetIn_name env
+ it_mkLambda_or_LetIn_name env
(mkLambda_string "P" typP
(add_branch (push_rel (Anonymous,None,typP) env') 0)) lnamespar
-
+
(* check if the type depends recursively on one of the inductive scheme *)
(**********************************************************************)
(* Building the recursive elimination *)
+(* Christine Paulin, 1996 *)
(*
- * t is the type of the constructor co and recargs is the information on
+ * t is the type of the constructor co and recargs is the information on
* the recursive calls. (It is assumed to be in form given by the user).
* build the type of the corresponding branch of the recurrence principle
- * assuming f has this type, branch_rec gives also the term
- * [x1]..[xk](f xi (F xi) ...) to be put in the corresponding branch of
+ * assuming f has this type, branch_rec gives also the term
+ * [x1]..[xk](f xi (F xi) ...) to be put in the corresponding branch of
* the case operation
- * FPvect gives for each inductive definition if we want an elimination
- * on it with which predicate and which recursive function.
+ * FPvect gives for each inductive definition if we want an elimination
+ * on it with which predicate and which recursive function.
*)
-let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
+let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
let make_prod = make_prod_dep dep in
let nparams = List.length vargs in
let process_pos env depK pk =
@@ -136,39 +138,39 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
| Ind (_,_) ->
let realargs = list_skipn nparams largs in
let base = applist (lift i pk,realargs) in
- if depK then
+ if depK then
Reduction.beta_appvect
base [|applist (mkRel (i+1),extended_rel_list 0 sign)|]
- else
+ else
base
- | _ -> assert false
+ | _ -> assert false
in
prec env 0 []
in
let rec process_constr env i c recargs nhyps li =
- if nhyps > 0 then match kind_of_term c with
+ if nhyps > 0 then match kind_of_term c with
| Prod (n,t,c_0) ->
- let (optionpos,rest) =
- match recargs with
+ let (optionpos,rest) =
+ match recargs with
| [] -> None,[]
| ra::rest ->
- (match dest_recarg ra with
+ (match dest_recarg ra with
| Mrec j when is_rec -> (depPvect.(j),rest)
- | Imbr _ ->
- Flags.if_verbose warning "Ignoring recursive call";
- (None,rest)
+ | Imbr _ ->
+ Flags.if_verbose warning "Ignoring recursive call";
+ (None,rest)
| _ -> (None, rest))
- in
- (match optionpos with
- | None ->
+ in
+ (match optionpos with
+ | None ->
make_prod env
(n,t,
process_constr (push_rel (n,None,t) env) (i+1) c_0 rest
(nhyps-1) (i::li))
- | Some(dep',p) ->
+ | Some(dep',p) ->
let nP = lift (i+1+decP) p in
let env' = push_rel (n,None,t) env in
- let t_0 = process_pos env' dep' nP (lift 1 t) in
+ let t_0 = process_pos env' dep' nP (lift 1 t) in
make_prod_dep (dep or dep') env
(n,t,
mkArrow t_0
@@ -190,27 +192,27 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
else c
in
let nhyps = List.length cs.cs_args in
- let nP = match depPvect.(tyi) with
+ let nP = match depPvect.(tyi) with
| Some(_,p) -> lift (nhyps+decP) p
| _ -> assert false in
let base = appvect (nP,cs.cs_concl_realargs) in
let c = it_mkProd_or_LetIn base cs.cs_args in
process_constr env 0 c recargs nhyps []
-let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
+let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
let process_pos env fk =
let rec prec env i hyps p =
let p',largs = whd_betadeltaiota_nolet_stack env sigma p in
match kind_of_term p' with
| Prod (n,t,c) ->
let d = (n,None,t) in
- lambda_name env (n,t,prec (push_rel d env) (i+1) (d::hyps) c)
+ mkLambda_name env (n,t,prec (push_rel d env) (i+1) (d::hyps) c)
| LetIn (n,b,t,c) ->
let d = (n,Some b,t) in
mkLetIn (n,b,t,prec (push_rel d env) (i+1) (d::hyps) c)
- | Ind _ ->
+ | 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),extended_rel_vect 0 hyps) in
applist(lift i fk,realargs@[arg])
| _ -> assert false
in
@@ -218,24 +220,24 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
in
(* ici, cstrprods est la liste des produits du constructeur instantié *)
let rec process_constr env i f = function
- | (n,None,t as d)::cprest, recarg::rest ->
- let optionpos =
- match dest_recarg recarg with
+ | (n,None,t as d)::cprest, recarg::rest ->
+ let optionpos =
+ match dest_recarg recarg with
| Norec -> None
| Imbr _ -> None
| Mrec i -> fvect.(i)
- in
- (match optionpos with
+ in
+ (match optionpos with
| None ->
- lambda_name env
+ mkLambda_name env
(n,t,process_constr (push_rel d env) (i+1)
(whd_beta Evd.empty (applist (lift 1 f, [(mkRel 1)])))
(cprest,rest))
- | Some(_,f_0) ->
+ | Some(_,f_0) ->
let nF = lift (i+1+decF) f_0 in
let env' = push_rel d env in
- let arg = process_pos env' nF (lift 1 t) in
- lambda_name env
+ let arg = process_pos env' nF (lift 1 t) in
+ mkLambda_name env
(n,t,process_constr env' (i+1)
(whd_beta Evd.empty (applist (lift 1 f, [(mkRel 1); arg])))
(cprest,rest)))
@@ -251,9 +253,9 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
process_constr env 0 f (List.rev cstr.cs_args, recargs)
-(* Cut a context ctx in 2 parts (ctx1,ctx2) with ctx1 containing k
+(* Cut a context ctx in 2 parts (ctx1,ctx2) with ctx1 containing k
variables *)
-let context_chop k ctx =
+let context_chop k ctx =
let rec chop_aux acc = function
| (0, l2) -> (List.rev acc, l2)
| (n, ((_,Some _,_ as h)::t)) -> chop_aux (h::acc) (n, t)
@@ -266,24 +268,24 @@ let context_chop k ctx =
let mis_make_indrec env sigma listdepkind mib =
let nparams = mib.mind_nparams in
let nparrec = mib. mind_nparams_rec in
- let lnonparrec,lnamesparrec =
+ let lnonparrec,lnamesparrec =
context_chop (nparams-nparrec) mib.mind_params_ctxt in
let nrec = List.length listdepkind in
let depPvec =
- Array.create mib.mind_ntypes (None : (bool * constr) option) in
- let _ =
- let rec
- assign k = function
+ Array.create mib.mind_ntypes (None : (bool * constr) option) in
+ let _ =
+ let rec
+ assign k = function
| [] -> ()
- | (indi,mibi,mipi,dep,_)::rest ->
+ | (indi,mibi,mipi,dep,_)::rest ->
(Array.set depPvec (snd indi) (Some(dep,mkRel k));
assign (k-1) rest)
- in
- assign nrec listdepkind in
+ in
+ assign nrec listdepkind in
let recargsvec =
Array.map (fun mip -> mip.mind_recargs) mib.mind_packets in
(* recarg information for non recursive parameters *)
- let rec recargparn l n =
+ let rec recargparn l n =
if n = 0 then l else recargparn (mk_norec::l) (n-1) in
let recargpar = recargparn [] (nparams-nparrec) in
let make_one_rec p =
@@ -293,80 +295,80 @@ let mis_make_indrec env sigma listdepkind mib =
let tyi = snd indi in
let nctyi =
Array.length mipi.mind_consnames in (* nb constructeurs du type*)
-
+
(* 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 indf = make_ind_family(indi,args) in
-
+
let arsign,_ = get_arity env indf in
let depind = build_dependent_inductive env indf in
let deparsign = (Anonymous,None,depind)::arsign in
-
+
let nonrecpar = rel_context_length lnonparrec in
let larsign = rel_context_length deparsign in
let ndepar = larsign - nonrecpar in
let dect = larsign+nrec+nbconstruct in
-
+
(* 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 indf' = make_ind_family(indi,args'@args'') in
-
- let branches =
+
+ let branches =
let constrs = get_constructors env indf' in
let fi = rel_vect (dect-i-nctyi) nctyi in
- let vecfi = Array.map
+ let vecfi = Array.map
(fun f -> appvect (f,extended_rel_vect ndepar lnonparrec))
- fi
+ fi
in
array_map3
- (make_rec_branch_arg env sigma
+ (make_rec_branch_arg env sigma
(nparrec,depPvec,larsign))
- vecfi constrs (dest_subterms recargsvec.(tyi))
+ vecfi constrs (dest_subterms recargsvec.(tyi))
in
-
- let j = (match depPvec.(tyi) with
- | Some (_,c) when isRel c -> destRel c
- | _ -> assert false)
+
+ let j = (match depPvec.(tyi) with
+ | Some (_,c) when isRel c -> destRel c
+ | _ -> assert false)
in
-
+
(* Predicate in the context of the case *)
-
+
let depind' = build_dependent_inductive env indf' in
let arsign',_ = get_arity env indf' in
let deparsign' = (Anonymous,None,depind')::arsign' in
-
+
let pargs =
- let nrpar = extended_rel_list (2*ndepar) lnonparrec
+ let nrpar = extended_rel_list (2*ndepar) lnonparrec
and nrar = if dep then extended_rel_list 0 deparsign'
else extended_rel_list 1 arsign'
in nrpar@nrar
-
+
in
(* body of i-th component of the mutual fixpoint *)
- let deftyi =
+ let deftyi =
let ci = make_case_info env indi RegularStyle in
- let concl = applist (mkRel (dect+j+ndepar),pargs) in
+ let concl = applist (mkRel (dect+j+ndepar),pargs) in
let pred =
- it_mkLambda_or_LetIn_name env
+ it_mkLambda_or_LetIn_name env
((if dep then mkLambda_name env else mkLambda)
(Anonymous,depind',concl))
arsign'
in
it_mkLambda_or_LetIn_name env
- (mkCase (ci, pred,
+ (mkCase (ci, pred,
mkRel 1,
branches))
(lift_rel_context nrec deparsign)
in
-
+
(* type of i-th component of the mutual fixpoint *)
-
+
let typtyi =
- let concl =
+ let concl =
let pargs = if dep then extended_rel_vect 0 deparsign
else extended_rel_vect 1 arsign
in appvect (mkRel (nbconstruct+ndepar+nonrecpar+j),pargs)
@@ -374,25 +376,25 @@ let mis_make_indrec env sigma listdepkind mib =
concl
deparsign
in
- mrec (i+nctyi) (rel_context_nhyps arsign ::ln) (typtyi::ltyp)
+ mrec (i+nctyi) (rel_context_nhyps arsign ::ln) (typtyi::ltyp)
(deftyi::ldef) rest
- | [] ->
+ | [] ->
let fixn = Array.of_list (List.rev ln) in
let fixtyi = Array.of_list (List.rev ltyp) in
- let fixdef = Array.of_list (List.rev ldef) in
+ let fixdef = Array.of_list (List.rev ldef) in
let names = Array.create nrec (Name(id_of_string "F")) in
mkFix ((fixn,p),(names,fixtyi,fixdef))
- in
- mrec 0 [] [] []
- in
- let rec make_branch env i = function
+ in
+ mrec 0 [] [] []
+ in
+ let rec make_branch env i = function
| (indi,mibi,mipi,dep,_)::rest ->
let tyi = snd indi in
let nconstr = Array.length mipi.mind_consnames in
- let rec onerec env j =
- if j = nconstr then
- make_branch env (i+j) rest
- else
+ let rec onerec env j =
+ if j = nconstr then
+ make_branch env (i+j) rest
+ else
let recarg = (dest_subterms recargsvec.(tyi)).(j) in
let recarg = recargpar@recarg in
let vargs = extended_rel_list (nrec+i+j) lnamesparrec in
@@ -400,106 +402,107 @@ let mis_make_indrec env sigma listdepkind mib =
let p_0 =
type_rec_branch
true dep env sigma (vargs,depPvec,i+j) tyi cs recarg
- in
+ in
mkLambda_string "f" p_0
(onerec (push_rel (Anonymous,None,p_0) env) (j+1))
in onerec env 0
- | [] ->
+ | [] ->
makefix i listdepkind
in
- let rec put_arity env i = function
- | (indi,_,_,dep,kinds)::rest ->
+ 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
mkLambda_string "P" typP
(put_arity (push_rel (Anonymous,None,typP) env) (i+1) rest)
- | [] ->
- make_branch env 0 listdepkind
+ | [] ->
+ make_branch env 0 listdepkind
in
-
+
(* Body on make_one_rec *)
let (indi,mibi,mipi,dep,kind) = List.nth listdepkind p in
-
+
if (mis_is_recursive_subset
(List.map (fun (indi,_,_,_,_) -> snd indi) listdepkind)
- mipi.mind_recargs)
- then
+ mipi.mind_recargs)
+ then
let env' = push_rel_context lnamesparrec env in
- it_mkLambda_or_LetIn_name env (put_arity env' 0 listdepkind)
+ it_mkLambda_or_LetIn_name env (put_arity env' 0 listdepkind)
lnamesparrec
- else
- mis_make_case_com (Some dep) env sigma indi (mibi,mipi) kind
- in
+ else
+ mis_make_case_com dep env sigma indi (mibi,mipi) kind
+ in
(* Body of mis_make_indrec *)
list_tabulate make_one_rec nrec
(**********************************************************************)
(* This builds elimination predicate for Case tactic *)
-let make_case_com depopt env sigma ity kind =
- let (mib,mip) = lookup_mind_specif env ity in
- mis_make_case_com depopt env sigma ity (mib,mip) kind
+let build_case_analysis_scheme env sigma ity dep kind =
+ let (mib,mip) = lookup_mind_specif env ity in
+ mis_make_case_com dep env sigma ity (mib,mip) kind
-let make_case_dep env = make_case_com (Some true) env
-let make_case_nodep env = make_case_com (Some false) env
-let make_case_gen env = make_case_com None env
+let build_case_analysis_scheme_default env sigma ity kind =
+ let (mib,mip) = lookup_mind_specif env ity in
+ let dep = inductive_sort_family mip <> InProp in
+ mis_make_case_com dep env sigma ity (mib,mip) kind
(**********************************************************************)
-(* [instantiate_indrec_scheme s rec] replace the sort of the scheme
+(* [modify_sort_scheme s rec] replaces the sort of the scheme
[rec] by [s] *)
-let change_sort_arity sort =
+let change_sort_arity sort =
let rec drec a = match kind_of_term a with
- | Cast (c,_,_) -> drec c
+ | Cast (c,_,_) -> drec c
| Prod (n,t,c) -> mkProd (n, t, drec c)
| LetIn (n,b,t,c) -> mkLetIn (n,b, t, drec c)
| Sort _ -> mkSort sort
| _ -> assert false
- in
- drec
+ in
+ drec
(* [npar] is the number of expected arguments (then excluding letin's) *)
-let instantiate_indrec_scheme sort =
+let modify_sort_scheme sort =
let rec drec npar elim =
match kind_of_term elim with
- | Lambda (n,t,c) ->
- if npar = 0 then
+ | Lambda (n,t,c) ->
+ if npar = 0 then
mkLambda (n, change_sort_arity sort t, c)
- else
+ else
mkLambda (n, t, drec (npar-1) c)
| LetIn (n,b,t,c) -> mkLetIn (n,b,t,drec npar c)
- | _ -> anomaly "instantiate_indrec_scheme: wrong elimination type"
+ | _ -> anomaly "modify_sort_scheme: wrong elimination type"
in
drec
(* Change the sort in the type of an inductive definition, builds the
corresponding eta-expanded term *)
-let instantiate_type_indrec_scheme sort npars term =
+let weaken_sort_scheme sort npars term =
let rec drec np elim =
match kind_of_term elim with
- | Prod (n,t,c) ->
- if np = 0 then
+ | Prod (n,t,c) ->
+ if np = 0 then
let t' = change_sort_arity sort t in
mkProd (n, t', c),
mkLambda (n, t', mkApp(term,Termops.rel_vect 0 (npars+1)))
- else
+ else
let c',term' = drec (np-1) c in
mkProd (n, t, c'), mkLambda (n, t, term')
| LetIn (n,b,t,c) -> let c',term' = drec np c in
- mkLetIn (n,b,t,c'), mkLetIn (n,b,t,term')
- | _ -> anomaly "instantiate_type_indrec_scheme: wrong elimination type"
+ mkLetIn (n,b,t,c'), mkLetIn (n,b,t,term')
+ | _ -> anomaly "weaken_sort_scheme: wrong elimination type"
in
drec npars
(**********************************************************************)
(* Interface to build complex Scheme *)
-(* Check inductive types only occurs once
+(* Check inductive types only occurs once
(otherwise we obtain a meaning less scheme) *)
-let check_arities listdepkind =
+let check_arities listdepkind =
let _ = List.fold_left
- (fun ln ((_,ni as mind),mibi,mipi,dep,kind) ->
+ (fun ln ((_,ni as mind),mibi,mipi,dep,kind) ->
let kelim = elim_sorts (mibi,mipi) in
if not (List.exists ((=) kind) kelim) then raise
(RecursionSchemeError
@@ -510,56 +513,29 @@ let check_arities listdepkind =
[] listdepkind
in true
-let build_mutual_indrec env sigma = function
- | (mind,mib,mip,dep,s)::lrecspec ->
+let build_mutual_induction_scheme env sigma = function
+ | (mind,dep,s)::lrecspec ->
+ let (mib,mip) = Global.lookup_inductive mind in
let (sp,tyi) = mind in
- let listdepkind =
- (mind,mib,mip, dep,s)::
+ let listdepkind =
+ (mind,mib,mip,dep,s)::
(List.map
- (function (mind',mibi',mipi',dep',s') ->
+ (function (mind',dep',s') ->
let (sp',_) = mind' in
if sp=sp' then
let (mibi',mipi') = lookup_mind_specif env mind' in
(mind',mibi',mipi',dep',s')
else
- raise (RecursionSchemeError (NotMutualInScheme (mind,mind'))))
+ raise (RecursionSchemeError (NotMutualInScheme (mind,mind'))))
lrecspec)
in
- let _ = check_arities listdepkind in
+ let _ = check_arities listdepkind in
mis_make_indrec env sigma listdepkind mib
- | _ -> anomaly "build_indrec expects a non empty list of inductive types"
+ | _ -> anomaly "build_induction_scheme expects a non empty list of inductive types"
-let build_indrec env sigma ind =
+let build_induction_scheme env sigma ind dep kind =
let (mib,mip) = lookup_mind_specif env ind in
- let kind = inductive_sort_family mip in
- let dep = kind <> InProp in
- List.hd (mis_make_indrec env sigma [(ind,mib,mip,dep,kind)] mib)
-
-(**********************************************************************)
-(* To handle old Case/Match syntax in Pretyping *)
-
-(*****************************************)
-(* To interpret Case and Match operators *)
-(* Expects a dependent predicate *)
-
-let type_rec_branches recursive env sigma indt p c =
- let IndType (indf,realargs) = indt in
- let (ind,params) = dest_ind_family indf in
- let (mib,mip) = lookup_mind_specif env ind in
- let recargs = mip.mind_recargs in
- let tyi = snd ind in
- let init_depPvec i = if i = tyi then Some(true,p) else None in
- let depPvec = Array.init mib.mind_ntypes init_depPvec in
- let constructors = get_constructors env indf in
- let lft =
- array_map2
- (type_rec_branch recursive true env sigma (params,depPvec,0) tyi)
- constructors (dest_subterms recargs) in
- (lft,Reduction.beta_appvect p (Array.of_list (realargs@[c])))
-(* Non recursive case. Pb: does not deal with unification
- let (p,ra,_) = type_case_branches env (ind,params@realargs) pj c in
- (p,ra)
-*)
+ List.hd (mis_make_indrec env sigma [(ind,mib,mip,dep,kind)] mib)
(*s Eliminations. *)
@@ -568,51 +544,32 @@ let elimination_suffix = function
| InSet -> "_rec"
| InType -> "_rect"
+let case_suffix = "_case"
+
let make_elimination_ident id s = add_suffix id (elimination_suffix s)
(* Look up function for the default elimination constant *)
let lookup_eliminator ind_sp s =
let kn,i = ind_sp in
- let mp,dp,l = repr_kn kn in
+ let mp,dp,l = repr_mind kn in
let ind_id = (Global.lookup_mind kn).mind_packets.(i).mind_typename in
let id = add_suffix ind_id (elimination_suffix s) in
(* Try first to get an eliminator defined in the same section as the *)
(* inductive type *)
- let ref = ConstRef (make_con mp dp (label_of_id id)) in
- try
- let _ = sp_of_global ref in
- constr_of_global ref
+ try
+ let cst =Global.constant_of_delta
+ (make_con mp dp (label_of_id id)) in
+ let _ = Global.lookup_constant cst in
+ mkConst cst
with Not_found ->
(* Then try to get a user-defined eliminator in some other places *)
(* using short name (e.g. for "eq_rec") *)
- try constr_of_global (Nametab.locate (make_short_qualid id))
+ try constr_of_global (Nametab.locate (qualid_of_ident id))
with Not_found ->
errorlabstrm "default_elim"
(strbrk "Cannot find the elimination combinator " ++
pr_id id ++ strbrk ", the elimination of the inductive definition " ++
- pr_global_env Idset.empty (IndRef ind_sp) ++
+ pr_global_env Idset.empty (IndRef ind_sp) ++
strbrk " on sort " ++ pr_sort_family s ++
strbrk " is probably not allowed.")
-
-
-(* let env = Global.env() in
- let path = sp_of_global None (IndRef ind_sp) in
- let dir, base = repr_path path in
- let id = add_suffix base (elimination_suffix s) in
- (* Try first to get an eliminator defined in the same section as the *)
- (* inductive type *)
- try construct_absolute_reference (Names.make_path dir id)
- with Not_found ->
- (* Then try to get a user-defined eliminator in some other places *)
- (* using short name (e.g. for "eq_rec") *)
- try constr_of_global (Nametab.locate (make_short_qualid id))
- with Not_found ->
- errorlabstrm "default_elim"
- (str "Cannot find the elimination combinator " ++
- pr_id id ++ spc () ++
- str "The elimination of the inductive definition " ++
- pr_id base ++ spc () ++ str "on sort " ++
- spc () ++ pr_sort_family s ++
- str " is probably not allowed")
-*)
diff --git a/pretyping/indrec.mli b/pretyping/indrec.mli
index 102c7c7f..91d559e1 100644
--- a/pretyping/indrec.mli
+++ b/pretyping/indrec.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: indrec.mli 11562 2008-11-09 11:30:10Z herbelin $ i*)
+(*i $Id$ i*)
(*i*)
open Names
@@ -27,36 +27,41 @@ exception RecursionSchemeError of recursion_scheme_error
(** Eliminations *)
-(* These functions build elimination predicate for Case tactic *)
+type dep_flag = bool
-val make_case_dep : env -> evar_map -> inductive -> sorts_family -> constr
-val make_case_nodep : env -> evar_map -> inductive -> sorts_family -> constr
-val make_case_gen : env -> evar_map -> inductive -> sorts_family -> constr
+(* Build a case analysis elimination scheme in some sort family *)
-(* This builds an elimination scheme associated (using the own arity
- of the inductive) *)
+val build_case_analysis_scheme : env -> evar_map -> inductive ->
+ dep_flag -> sorts_family -> constr
-val build_indrec : env -> evar_map -> inductive -> constr
-val instantiate_indrec_scheme : sorts -> int -> constr -> constr
-val instantiate_type_indrec_scheme : sorts -> int -> constr -> types ->
- constr * types
+(* Build a dependent case elimination predicate unless type is in Prop *)
-(** Complex recursion schemes [Scheme] *)
+val build_case_analysis_scheme_default : env -> evar_map -> inductive ->
+ sorts_family -> constr
-val build_mutual_indrec :
- env -> evar_map ->
- (inductive * mutual_inductive_body * one_inductive_body
- * bool * sorts_family) list
- -> constr list
+(* Builds a recursive induction scheme (Peano-induction style) in the same
+ sort family as the inductive family; it is dependent if not in Prop *)
-(** Old Case/Match typing *)
+val build_induction_scheme : env -> evar_map -> inductive ->
+ dep_flag -> sorts_family -> constr
-val type_rec_branches : bool -> env -> evar_map -> inductive_type
- -> constr -> constr -> constr array * constr
-val make_rec_branch_arg :
- env -> evar_map ->
- int * ('b * constr) option array * int ->
- constr -> constructor_summary -> wf_paths list -> constr
+(* 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
+ 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]
+ whose conclusion is quantified on [Type] at position [n] of [t] a
+ scheme quantified on sort [s] *)
+
+val weaken_sort_scheme : sorts -> int -> constr -> types -> constr * types
(** Recursor names utilities *)
@@ -64,5 +69,4 @@ val lookup_eliminator : inductive -> sorts_family -> constr
val elimination_suffix : sorts_family -> string
val make_elimination_ident : identifier -> sorts_family -> identifier
-
-
+val case_suffix : string
diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml
index 9f8c06da..636f8622 100644
--- a/pretyping/inductiveops.ml
+++ b/pretyping/inductiveops.ml
@@ -6,13 +6,14 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: inductiveops.ml 11436 2008-10-07 13:56:55Z barras $ *)
+(* $Id$ *)
open Util
open Names
open Univ
open Term
open Termops
+open Namegen
open Sign
open Declarations
open Environ
@@ -71,16 +72,15 @@ let substnl_ind_type l n = map_inductive_type (substnl l n)
let mkAppliedInd (IndType ((ind,params), realargs)) =
applist (mkInd ind,params@realargs)
-
-(* Does not consider imbricated or mutually recursive types *)
-let mis_is_recursive_subset listind rarg =
- let rec one_is_rec rvec =
+(* Does not consider imbricated or mutually recursive types *)
+let mis_is_recursive_subset listind rarg =
+ let rec one_is_rec rvec =
List.exists
(fun ra ->
match dest_recarg ra with
- | Mrec i -> List.mem i listind
+ | Mrec i -> List.mem i listind
| _ -> false) rvec
- in
+ in
array_exists one_is_rec (dest_subterms rarg)
let mis_is_recursive (ind,mib,mip) =
@@ -91,7 +91,7 @@ let mis_nf_constructor_type (ind,mib,mip) j =
let specif = mip.mind_nf_lc
and ntypes = mib.mind_ntypes
and nconstr = Array.length mip.mind_consnames in
- let make_Ik k = mkInd ((fst ind),ntypes-k-1) in
+ let make_Ik k = mkInd ((fst ind),ntypes-k-1) in
if j > nconstr then error "Not enough constructors in the type.";
substl (list_tabulate make_Ik ntypes) specif.(j-1)
@@ -102,15 +102,15 @@ let mis_constr_nargs indsp =
let recargs = dest_subterms mip.mind_recargs in
Array.map List.length recargs
-let mis_constr_nargs_env env (kn,i) =
+let mis_constr_nargs_env env (kn,i) =
let mib = Environ.lookup_mind kn env in
- let mip = mib.mind_packets.(i) in
+ let mip = mib.mind_packets.(i) in
let recargs = dest_subterms mip.mind_recargs in
Array.map List.length recargs
let mis_constructor_nargs_env env ((kn,i),j) =
let mib = Environ.lookup_mind kn env in
- let mip = mib.mind_packets.(i) in
+ let mip = mib.mind_packets.(i) in
recarg_length mip.mind_recargs j + mib.mind_nparams
let constructor_nrealargs env (ind,j) =
@@ -125,11 +125,15 @@ let get_full_arity_sign env ind =
let (mib,mip) = Inductive.lookup_mind_specif env ind in
mip.mind_arity_ctxt
+let nconstructors ind =
+ let (mib,mip) = Inductive.lookup_mind_specif (Global.env()) ind in
+ Array.length mip.mind_consnames
+
(* Length of arity (w/o local defs) *)
let inductive_nargs env ind =
let (mib,mip) = Inductive.lookup_mind_specif env ind in
- mib.mind_nparams, mip.mind_nrealargs
+ (rel_context_length (mib.mind_params_ctxt), mip.mind_nrealargs_ctxt)
let allowed_sorts env (kn,i as ind) =
let (mib,mip) = Inductive.lookup_mind_specif env ind in
@@ -138,7 +142,7 @@ let allowed_sorts env (kn,i as ind) =
(* Annotation for cases *)
let make_case_info env ind style =
let (mib,mip) = Inductive.lookup_mind_specif env ind in
- let print_info = { ind_nargs = mip.mind_nrealargs; style = style } in
+ 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;
@@ -172,7 +176,7 @@ let instantiate_params t args sign =
(match kind_of_term t with
| Prod(_,_,t) -> inst (a::s) t (ctxt,args)
| _ -> anomaly"instantiate_params: type, ctxt and args mismatch")
- | ((_,(Some b),_)::ctxt,args) ->
+ | ((_,(Some b),_)::ctxt,args) ->
(match kind_of_term t with
| LetIn(_,_,_,t) -> inst ((substl s b)::s) t (ctxt,args)
| _ -> anomaly"instantiate_params: type, ctxt and args mismatch")
@@ -249,7 +253,7 @@ let build_dependent_constructor cs =
let build_dependent_inductive env ((ind, params) as indf) =
let arsign,_ = get_arity env indf in
let nrealargs = List.length arsign in
- applist
+ applist
(mkInd ind,
(List.map (lift nrealargs) params)@(extended_rel_list 0 arsign))
@@ -322,7 +326,7 @@ let find_coinductive env sigma c =
(* find appropriate names for pattern variables. Useful in the Case
and Inversion (case_then_using et case_nodep_then_using) tactics. *)
-let is_predicate_explicitly_dep env pred arsign =
+let is_predicate_explicitly_dep env pred arsign =
let rec srec env pval arsign =
let pv' = whd_betadeltaiota env Evd.empty pval in
match kind_of_term pv', arsign with
@@ -330,7 +334,7 @@ let is_predicate_explicitly_dep env pred arsign =
srec (push_rel_assum (na,t) env) b arsign
| Lambda (na,_,_), _ ->
- (* The following code has impact on the introduction names
+ (* The following code has an impact on the introduction names
given by the tactics "case" and "inversion": when the
elimination is not dependent, "case" uses Anonymous for
inductive types in Prop and names created by mkProd_name for
@@ -370,18 +374,21 @@ let set_pattern_names env ind brv =
let arities =
Array.map
(fun c ->
- rel_context_length (fst (decompose_prod_assum c)) -
+ rel_context_length ((prod_assum c)) -
mib.mind_nparams)
mip.mind_nf_lc in
array_map2 (set_names env) arities brv
-
-let type_case_branches_with_names env indspec pj c =
+let type_case_branches_with_names env indspec p c =
let (ind,args) = indspec in
- let (lbrty,conclty,_) = Inductive.type_case_branches env indspec pj c in
- let (mib,mip) = Inductive.lookup_mind_specif env ind in
- let params = list_firstn mib.mind_nparams args in
- if is_elim_predicate_explicitly_dependent env pj.uj_val (ind,params) then
+ let (mib,mip as specif) = Inductive.lookup_mind_specif env ind in
+ let nparams = mib.mind_nparams in
+ let (params,realargs) = list_chop nparams args in
+ let lbrty = Inductive.build_branches_type ind specif params p in
+ (* Build case type *)
+ let conclty = Reduction.beta_appvect p (Array.of_list (realargs@[c])) in
+ (* Adjust names *)
+ if is_elim_predicate_explicitly_dependent env p (ind,params) then
(set_pattern_names env ind lbrty, conclty)
else (lbrty, conclty)
@@ -399,7 +406,7 @@ let arity_of_case_predicate env (ind,params) dep k =
(* Check if u (sort of a parameter) appears in the sort of the
inductive (is). This is done by trying to enforce u > u' >= is
in the empty univ graph. If an inconsistency appears, then
- is depends on u. *)
+ is depends on u. *)
let is_constrained is u =
try
let u' = fresh_local_univ() in
@@ -450,7 +457,7 @@ let type_of_inductive_knowing_conclusion env mip conclty =
(* A function which checks that a term well typed verifies both
syntactic conditions *)
-let control_only_guard env c =
+let control_only_guard env c =
let check_fix_cofix e c = match kind_of_term c with
| CoFix (_,(_,_,_) as cofix) ->
Inductive.check_cofix e cofix
@@ -458,12 +465,12 @@ let control_only_guard env c =
Inductive.check_fix e fix
| _ -> ()
in
- let rec iter env c =
- check_fix_cofix env c;
+ let rec iter env c =
+ check_fix_cofix env c;
iter_constr_with_full_binders push_rel iter env c
in
iter env c
-let subst_inductive subst (kn,i as ind) =
- let kn' = Mod_subst.subst_kn subst kn in
+let subst_inductive subst (kn,i as ind) =
+ let kn' = Mod_subst.subst_ind subst kn in
if kn == kn' then ind else (kn',i)
diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli
index 1cf940cb..a9a51d9a 100644
--- a/pretyping/inductiveops.mli
+++ b/pretyping/inductiveops.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: inductiveops.mli 11436 2008-10-07 13:56:55Z barras $ i*)
+(*i $Id$ i*)
open Names
open Term
@@ -58,7 +58,9 @@ val mis_nf_constructor_type :
val mis_constr_nargs : inductive -> int array
val mis_constr_nargs_env : env -> inductive -> int array
-(* Return number of expected parameters and of expected real arguments *)
+val nconstructors : inductive -> int
+
+(* 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
@@ -75,7 +77,7 @@ type constructor_summary = {
cs_cstr : constructor;
cs_params : constr list;
cs_nargs : int;
- cs_args : Sign.rel_context;
+ cs_args : rel_context;
cs_concl_realargs : constr array;
}
val lift_constructor : int -> constructor_summary -> constructor_summary
@@ -86,7 +88,7 @@ val get_arity : env -> inductive_family -> rel_context * sorts_family
val get_constructors : env -> inductive_family -> constructor_summary array
val build_dependent_constructor : constructor_summary -> constr
val build_dependent_inductive : env -> inductive_family -> constr
-val make_arity_signature : env -> bool -> inductive_family -> Sign.rel_context
+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
@@ -104,11 +106,11 @@ val arity_of_case_predicate :
env -> inductive_family -> bool -> sorts -> types
val type_case_branches_with_names :
- env -> inductive * constr list -> unsafe_judgment -> constr ->
+ env -> inductive * constr list -> constr -> constr ->
types array * types
val make_case_info : env -> inductive -> case_style -> case_info
-(*i Compatibility
+(*i Compatibility
val make_default_case_info : env -> case_style -> inductive -> case_info
i*)
diff --git a/pretyping/matching.ml b/pretyping/matching.ml
index 93bac98e..45432ec0 100644
--- a/pretyping/matching.ml
+++ b/pretyping/matching.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: matching.ml 11735 2009-01-02 17:22:31Z herbelin $ *)
+(* $Id$ *)
(*i*)
open Util
@@ -48,9 +48,10 @@ type bound_ident_map = (identifier * identifier) list
exception PatternMatchingFailure
-let constrain (n,m) (names,terms as subst) =
+let constrain (n,(ids,m as x)) (names,terms as subst) =
try
- if eq_constr m (List.assoc n terms) then subst
+ let (ids',m') = List.assoc n terms in
+ if ids = ids' && eq_constr m m' then subst
else raise PatternMatchingFailure
with
Not_found ->
@@ -58,14 +59,14 @@ let constrain (n,m) (names,terms as subst) =
Flags.if_verbose Pp.warning
("Collision between bound variable "^string_of_id n^
" and a metavariable of same name.");
- (names,(n,m)::terms)
+ (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);
+ ("Collision between bound variables of name "^string_of_id id1);
(names,terms))
else
(if List.mem_assoc id1 terms then
@@ -75,15 +76,15 @@ let add_binders na1 na2 (names,terms as subst) =
((id1,id2)::names,terms));
| _ -> subst
-let build_lambda toabstract stk (m : constr) =
- let rec buildrec m p_0 p_1 = match p_0,p_1 with
+let build_lambda toabstract stk (m : constr) =
+ let rec buildrec m p_0 p_1 = match p_0,p_1 with
| (_, []) -> m
- | (n, (na,t)::tl) ->
+ | (n, (_,na,t)::tl) ->
if List.mem n toabstract then
buildrec (mkLambda (na,t,m)) (n+1) tl
- else
+ else
buildrec (lift (-1) m) (n+1) tl
- in
+ in
buildrec m 1 stk
let memb_metavars m n =
@@ -98,7 +99,58 @@ let same_case_structure (_,cs1,ind,_) ci2 br1 br2 =
| Some ind -> ind = ci2.ci_ind
| None -> cs1 = ci2.ci_cstr_nargs
-let matches_core convert allow_partial_app pat c =
+let rec list_insert f a = function
+ | [] -> [a]
+ | b::l when f a b -> a::b::l
+ | b::l when a = b -> raise PatternMatchingFailure
+ | b::l -> b :: list_insert f a l
+
+let extract_bound_vars =
+ let rec aux k = function
+ | ([],_) -> []
+ | (n::l,(na1,na2,_)::stk) when k = n ->
+ begin match na1,na2 with
+ | Name id1,Name _ -> list_insert (<) id1 (aux (k+1) (l,stk))
+ | Name _,Anonymous -> anomaly "Unnamed bound variable"
+ | Anonymous,_ -> raise PatternMatchingFailure
+ end
+ | (l,_::stk) -> aux (k+1) (l,stk)
+ | (_,[]) -> assert false
+ in aux 1
+
+let dummy_constr = mkProp
+
+let rec make_renaming ids = function
+ | (Name id,Name _,_)::stk ->
+ let renaming = make_renaming ids stk in
+ (try mkRel (list_index id ids) :: renaming
+ with Not_found -> dummy_constr :: renaming)
+ | (_,_,_)::stk ->
+ dummy_constr :: make_renaming ids stk
+ | [] ->
+ []
+
+let merge_binding allow_bound_rels stk n cT subst =
+ let depth = List.length stk in
+ let c =
+ if depth = 0 then
+ (* Optimization *)
+ ([],cT)
+ else
+ let frels = Intset.elements (free_rels cT) in
+ let frels = List.filter (fun i -> i <= depth) frels in
+ if allow_bound_rels then
+ let frels = Sort.list (<) frels in
+ let canonically_ordered_vars = extract_bound_vars (frels,stk) in
+ let renaming = make_renaming canonically_ordered_vars stk in
+ (canonically_ordered_vars, substl renaming cT)
+ else if frels = [] then
+ ([],lift (-depth) cT)
+ else
+ raise PatternMatchingFailure in
+ constrain (n,c) subst
+
+let matches_core convert allow_partial_app allow_bound_rels pat c =
let conv = match convert with
| None -> eq_constr
| Some (env,sigma) -> is_conv env sigma in
@@ -114,21 +166,11 @@ let matches_core convert allow_partial_app pat c =
args in
let frels = Intset.elements (free_rels cT) in
if list_subset frels relargs then
- constrain (n,build_lambda relargs stk cT) subst
+ constrain (n,([],build_lambda relargs stk cT)) subst
else
raise PatternMatchingFailure
- | PMeta (Some n), m ->
- let depth = List.length stk in
- if depth = 0 then
- (* Optimisation *)
- constrain (n,cT) subst
- else
- let frels = Intset.elements (free_rels cT) in
- if List.for_all (fun i -> i > depth) frels then
- constrain (n,lift (-depth) cT) subst
- else
- raise PatternMatchingFailure
+ | PMeta (Some n), m -> merge_binding allow_bound_rels stk n cT subst
| PMeta None, m -> subst
@@ -153,14 +195,8 @@ let matches_core convert allow_partial_app pat c =
let p = Array.length args2 - Array.length args1 in
if p>=0 then
let args21, args22 = array_chop p args2 in
- let subst =
- let depth = List.length stk in
- let c = mkApp(c2,args21) in
- let frels = Intset.elements (free_rels c) in
- if List.for_all (fun i -> i > depth) frels then
- constrain (n,lift (-depth) c) subst
- else
- raise PatternMatchingFailure in
+ let c = mkApp(c2,args21) in
+ let subst = merge_binding allow_bound_rels stk n c subst in
array_fold_left2 (sorec stk) subst args1 args22
else raise PatternMatchingFailure
@@ -169,15 +205,15 @@ let matches_core convert allow_partial_app pat c =
with Invalid_argument _ -> raise PatternMatchingFailure)
| PProd (na1,c1,d1), Prod(na2,c2,d2) ->
- sorec ((na2,c2)::stk)
+ sorec ((na1,na2,c2)::stk)
(add_binders na1 na2 (sorec stk subst c1 c2)) d1 d2
| PLambda (na1,c1,d1), Lambda(na2,c2,d2) ->
- sorec ((na2,c2)::stk)
+ sorec ((na1,na2,c2)::stk)
(add_binders na1 na2 (sorec stk subst c1 c2)) d1 d2
| PLetIn (na1,c1,d1), LetIn(na2,c2,t2,d2) ->
- sorec ((na2,t2)::stk)
+ sorec ((na1,na2,t2)::stk)
(add_binders na1 na2 (sorec stk subst c1 c2)) d1 d2
| PIf (a1,b1,b1'), Case (ci,_,a2,[|b2;b2'|]) ->
@@ -186,8 +222,10 @@ let matches_core convert allow_partial_app pat c =
let n = rel_context_length ctx in
let n' = rel_context_length ctx' in
if noccur_between 1 n b2 & noccur_between 1 n' b2' then
- let s = List.fold_left (fun l (na,_,t) -> (na,t)::l) stk ctx in
- let s' = List.fold_left (fun l (na,_,t) -> (na,t)::l) stk ctx' in
+ let s =
+ List.fold_left (fun l (na,_,t) -> (Anonymous,na,t)::l) stk ctx in
+ let s' =
+ List.fold_left (fun l (na,_,t) -> (Anonymous,na,t)::l) stk ctx' in
let b1 = lift_pattern n b1 and b1' = lift_pattern n' b1' in
sorec s' (sorec s (sorec stk subst a1 a2) b1 b2) b1' b2'
else
@@ -195,7 +233,7 @@ let matches_core convert allow_partial_app pat c =
| PCase (ci1,p1,a1,br1), Case (ci2,p2,a2,br2) ->
if same_case_structure ci1 ci2 br1 br2 then
- array_fold_left2 (sorec stk)
+ array_fold_left2 (sorec stk)
(sorec stk (sorec stk subst a1 a2) p1 p2) br1 br2
else
raise PatternMatchingFailure
@@ -208,16 +246,20 @@ let matches_core convert allow_partial_app pat c =
let names,terms = sorec [] ([],[]) pat c in
(names,Sort.list (fun (a,_) (b,_) -> a<b) terms)
-let extended_matches = matches_core None true
+let matches_core_closed convert allow_partial_app pat c =
+ let names,subst = matches_core convert allow_partial_app false pat c in
+ (names, List.map (fun (a,(_,b)) -> (a,b)) subst)
+
+let extended_matches = matches_core None true true
-let matches c p = snd (matches_core None true c p)
+let matches c p = snd (matches_core_closed None true c p)
let special_meta = (-1)
(* Tells if it is an authorized occurrence and if the instance is closed *)
let authorized_occ partial_app closed pat c mk_ctx next =
- try
- let sigma = matches_core None partial_app pat c in
+ try
+ let sigma = matches_core_closed None partial_app pat c in
if closed && not (List.for_all (fun (_,c) -> closed0 c) (snd sigma))
then next ()
else sigma, mk_ctx (mkMeta special_meta), next
@@ -251,7 +293,7 @@ let sub_match ?(partial_app=false) ?(closed=true) pat c =
if topdown then
let lc1 = Array.sub lc 0 (Array.length lc - 1) in
let app = mkApp (c1,lc1) in
- let mk_ctx = function
+ let mk_ctx = function
| [app';c] -> mk_ctx (mkApp (app',[|c|]))
| _ -> assert false in
try_aux [app;array_last lc] mk_ctx next
@@ -274,7 +316,7 @@ let sub_match ?(partial_app=false) ?(closed=true) pat c =
try_aux (c1::Array.to_list lc) mk_ctx next)
| Case (ci,hd,c1,lc) ->
authorized_occ partial_app closed pat c mk_ctx (fun () ->
- let mk_ctx le =
+ let mk_ctx le =
mk_ctx (mkCase (ci,hd,List.hd le,Array.of_list (List.tl le))) in
try_aux (c1::Array.to_list lc) mk_ctx next)
| Construct _ | Fix _ | Ind _|CoFix _ |Evar _|Const _
@@ -310,7 +352,7 @@ let is_matching_appsubterm ?(closed=true) pat c =
with PatternMatchingFailure -> false
let matches_conv env sigma c p =
- snd (matches_core (Some (env,sigma)) false c p)
+ snd (matches_core_closed (Some (env,sigma)) false c p)
let is_matching_conv env sigma pat n =
try let _ = matches_conv env sigma pat n in true
diff --git a/pretyping/matching.mli b/pretyping/matching.mli
index b54a17b7..e8f23b1f 100644
--- a/pretyping/matching.mli
+++ b/pretyping/matching.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: matching.mli 11739 2009-01-02 19:33:19Z herbelin $ i*)
+(*i $Id$ i*)
(*i*)
open Names
@@ -34,8 +34,8 @@ val matches : constr_pattern -> constr -> patvar_map
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 * patvar_map
+val extended_matches :
+ constr_pattern -> constr -> bound_ident_map * extended_patvar_map
(* [is_matching pat c] just tells if [c] matches against [pat] *)
val is_matching : constr_pattern -> constr -> bool
@@ -59,14 +59,14 @@ type subterm_matching_result =
val match_subterm : constr_pattern -> constr -> subterm_matching_result
(* [match_appsubterm pat c] returns the substitution and the context
- corresponding to the first **closed** subterm of [c] matching [pat],
+ 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 *) ->
+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
diff --git a/pretyping/namegen.ml b/pretyping/namegen.ml
new file mode 100644
index 00000000..7d141faf
--- /dev/null
+++ b/pretyping/namegen.ml
@@ -0,0 +1,312 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* $Id$ *)
+
+(* Created from contents that was formerly in termops.ml and
+ nameops.ml, Nov 2009 *)
+
+(* This file is about generating new or fresh names and dealing with
+ alpha-renaming *)
+
+open Util
+open Names
+open Term
+open Nametab
+open Nameops
+open Libnames
+open Environ
+open Termops
+
+(**********************************************************************)
+(* Globality of identifiers *)
+
+let rec is_imported_modpath mp =
+ let current_mp,_ = Lib.current_prefix() in
+ match mp with
+ | MPfile dp ->
+ let rec find_prefix = function
+ |MPfile dp1 -> not (dp1=dp)
+ |MPdot(mp,_) -> find_prefix mp
+ |MPbound(_) -> false
+ in find_prefix current_mp
+ | p -> false
+
+let is_imported_ref = function
+ | VarRef _ -> false
+ | IndRef (kn,_)
+ | ConstructRef ((kn,_),_) ->
+ let (mp,_,_) = repr_mind kn in is_imported_modpath mp
+ | ConstRef kn ->
+ let (mp,_,_) = repr_con kn in is_imported_modpath mp
+
+let is_global id =
+ try
+ let ref = locate (qualid_of_ident id) in
+ not (is_imported_ref ref)
+ with Not_found ->
+ false
+
+let is_constructor id =
+ try
+ match locate (qualid_of_ident id) with
+ | ConstructRef _ as ref -> not (is_imported_ref ref)
+ | _ -> false
+ with Not_found ->
+ false
+
+(**********************************************************************)
+(* Generating "intuitive" names from its type *)
+
+let lowercase_first_char id = (* First character of a constr *)
+ lowercase_first_char_utf8 (string_of_id id)
+
+let sort_hdchar = function
+ | Prop(_) -> "P"
+ | Type(_) -> "T"
+
+let hdchar env c =
+ let rec hdrec k c =
+ match kind_of_term c with
+ | Prod (_,_,c) -> hdrec (k+1) c
+ | Lambda (_,_,c) -> hdrec (k+1) c
+ | LetIn (_,_,_,c) -> hdrec (k+1) c
+ | Cast (c,_,_) -> hdrec k c
+ | App (f,l) -> hdrec k f
+ | Const kn -> lowercase_first_char (id_of_label (con_label kn))
+ | Ind x -> lowercase_first_char (basename_of_global (IndRef x))
+ | Construct x -> lowercase_first_char (basename_of_global (ConstructRef x))
+ | Var id -> lowercase_first_char id
+ | Sort s -> sort_hdchar s
+ | Rel n ->
+ (if n<=k then "p" (* the initial term is flexible product/function *)
+ else
+ try match Environ.lookup_rel (n-k) env with
+ | (Name id,_,_) -> lowercase_first_char id
+ | (Anonymous,_,t) -> hdrec 0 (lift (n-k) t)
+ with Not_found -> "y")
+ | Fix ((_,i),(lna,_,_)) ->
+ let id = match lna.(i) with Name id -> id | _ -> assert false in
+ lowercase_first_char id
+ | CoFix (i,(lna,_,_)) ->
+ let id = match lna.(i) with Name id -> id | _ -> assert false in
+ lowercase_first_char id
+ | Meta _|Evar _|Case (_, _, _, _) -> "y"
+ in
+ hdrec 0 c
+
+let id_of_name_using_hdchar env a = function
+ | Anonymous -> id_of_string (hdchar env a)
+ | Name id -> id
+
+let named_hd env a = function
+ | Anonymous -> Name (id_of_string (hdchar env a))
+ | x -> x
+
+let mkProd_name env (n,a,b) = mkProd (named_hd env a n, a, b)
+let mkLambda_name env (n,a,b) = mkLambda (named_hd env a n, a, b)
+
+let lambda_name = mkLambda_name
+let prod_name = mkProd_name
+
+let prod_create env (a,b) = mkProd (named_hd env a Anonymous, a, b)
+let lambda_create env (a,b) = mkLambda (named_hd env a Anonymous, a, b)
+
+let name_assumption env (na,c,t) =
+ match c with
+ | None -> (named_hd env t na, None, t)
+ | Some body -> (named_hd env body na, c, t)
+
+let name_context env hyps =
+ snd
+ (List.fold_left
+ (fun (env,hyps) d ->
+ let d' = name_assumption env d in (push_rel d' env, d' :: hyps))
+ (env,[]) (List.rev hyps))
+
+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)
+let it_mkLambda_or_LetIn_name env b hyps =
+ it_mkLambda_or_LetIn ~init:b (name_context env hyps)
+
+(**********************************************************************)
+(* Fresh names *)
+
+let default_x = id_of_string "x"
+
+(* Looks for next "good" name by lifting subscript *)
+
+let next_ident_away_from id bad =
+ let rec name_rec id = if bad id then name_rec (lift_subscript id) else id in
+ name_rec id
+
+(* Restart subscript from x0 if name starts with xN, or x00 if name
+ starts with x0N, etc *)
+
+let restart_subscript id =
+ if not (has_subscript id) then id else
+ (* Ce serait sans doute mieux avec quelque chose inspiré de
+ *** make_ident id (Some 0) *** mais ça brise la compatibilité... *)
+ forget_subscript id
+
+(* Now, there are different renaming strategies... *)
+
+(* 1- Looks for a fresh name for printing in cases pattern *)
+
+let next_name_away_in_cases_pattern na avoid =
+ let id = match na with Name id -> id | Anonymous -> default_x in
+ next_ident_away_from id (fun id -> List.mem id avoid or is_constructor id)
+
+(* 2- Looks for a fresh name for introduction in goal *)
+
+(* The legacy strategy for renaming introduction variables is not very uniform:
+ - if the name to use is fresh in the context but used as a global
+ name, then a fresh name is taken by finding a free subscript
+ starting from the current subscript;
+ - but if the name to use is not fresh in the current context, the fresh
+ name is taken by finding a free subscript starting from 0 *)
+
+let next_ident_away_in_goal id avoid =
+ let id = if List.mem id avoid then restart_subscript id else id in
+ let bad id = List.mem id avoid || (is_global id & not (is_section_variable id)) in
+ next_ident_away_from id bad
+
+let next_name_away_in_goal na avoid =
+ let id = match na with Name id -> id | Anonymous -> id_of_string "H" in
+ next_ident_away_in_goal id avoid
+
+(* 3- Looks for next fresh name outside a list that is moreover valid
+ as a global identifier; the legacy algorithm is that if the name is
+ already used in the list, one looks for a name of same base with
+ lower available subscript; if the name is not in the list but is
+ used globally, one looks for a name of same base with lower subscript
+ beyond the current subscript *)
+
+let next_global_ident_away id avoid =
+ let id = if List.mem id avoid then restart_subscript id else id in
+ let bad id = List.mem id avoid || is_global id in
+ next_ident_away_from id bad
+
+(* 4- Looks for next fresh name outside a list; if name already used,
+ looks for same name with lower available subscript *)
+
+let next_ident_away id avoid =
+ if List.mem id avoid then
+ next_ident_away_from (restart_subscript id) (fun id -> List.mem id avoid)
+ else id
+
+let next_name_away_with_default default na avoid =
+ let id = match na with Name id -> id | Anonymous -> id_of_string default in
+ next_ident_away id avoid
+
+let next_name_away = next_name_away_with_default "H"
+
+let make_all_name_different env =
+ let avoid = ref (ids_of_named_context (named_context env)) in
+ process_rel_context
+ (fun (na,c,t) newenv ->
+ let id = next_name_away na !avoid in
+ avoid := id::!avoid;
+ push_rel (Name id,c,t) newenv)
+ env
+
+(* 5- Looks for next fresh name outside a list; avoids also to use names that
+ would clash with short name of global references; if name is already used,
+ looks for name of same base with lower available subscript beyond current
+ subscript *)
+
+let visibly_occur_id id c =
+ let rec occur c = match kind_of_term c with
+ | Const _ | Ind _ | Construct _ | Var _
+ when shortest_qualid_of_global Idset.empty (global_of_constr c)
+ = qualid_of_ident id -> raise Occur
+ | _ -> iter_constr occur c
+ in
+ try occur c; false
+ with Occur -> true
+ | Not_found -> false (* Happens when a global is not in the env *)
+
+let next_ident_away_for_default_printing t id avoid =
+ let bad id = List.mem id avoid or visibly_occur_id id t in
+ next_ident_away_from id bad
+
+let next_name_away_for_default_printing t na avoid =
+ let id = match na with
+ | Name id -> id
+ | Anonymous ->
+ (* In principle, an anonymous name is not dependent and will not be *)
+ (* taken into account by the function compute_displayed_name_in; *)
+ (* just in case, invent a valid name *)
+ id_of_string "H" in
+ next_ident_away_for_default_printing t id avoid
+
+(**********************************************************************)
+(* Displaying terms avoiding bound variables clashes *)
+
+(* Renaming strategy introduced in December 1998:
+
+ - Rule number 1: all names, even if unbound and not displayed, contribute
+ to the list of names to avoid
+ - Rule number 2: only the dependency status is used for deciding if
+ a name is displayed or not
+
+ Example:
+ bool_ind: "forall (P:bool->Prop)(f:(P true))(f:(P false))(b:bool), P b" is
+ displayed "forall P:bool->Prop, P true -> P false -> forall b:bool, P b"
+ but f and f0 contribute to the list of variables to avoid (knowing
+ that f and f0 are how the f's would be named if introduced, assuming
+ no other f and f0 are already used).
+*)
+
+type renaming_flags =
+ | RenamingForCasesPattern
+ | RenamingForGoal
+ | RenamingElsewhereFor of constr
+
+let next_name_for_display flags =
+ match flags with
+ | RenamingForCasesPattern -> next_name_away_in_cases_pattern
+ | RenamingForGoal -> next_name_away_in_goal
+ | RenamingElsewhereFor t -> next_name_away_for_default_printing t
+
+(* Remark: Anonymous var may be dependent in Evar's contexts *)
+let compute_displayed_name_in flags avoid na c =
+ if na = Anonymous & noccurn 1 c then
+ (Anonymous,avoid)
+ else
+ let fresh_id = next_name_for_display flags na avoid in
+ let idopt = if noccurn 1 c then Anonymous else Name fresh_id in
+ (idopt, fresh_id::avoid)
+
+let compute_and_force_displayed_name_in flags avoid na c =
+ if na = Anonymous & noccurn 1 c then
+ (Anonymous,avoid)
+ else
+ let fresh_id = next_name_for_display flags na avoid in
+ (Name fresh_id, fresh_id::avoid)
+
+let compute_displayed_let_name_in flags avoid na c =
+ let fresh_id = next_name_for_display flags na avoid in
+ (Name fresh_id, fresh_id::avoid)
+
+let rec rename_bound_vars_as_displayed avoid c =
+ let rec rename avoid c =
+ match kind_of_term c with
+ | Prod (na,c1,c2) ->
+ let na',avoid' = compute_displayed_name_in (RenamingElsewhereFor c2) avoid na c2 in
+ mkProd (na', c1, rename avoid' c2)
+ | LetIn (na,c1,t,c2) ->
+ let na',avoid' = compute_displayed_let_name_in (RenamingElsewhereFor c2) avoid na c2 in
+ mkLetIn (na',c1,t, rename avoid' c2)
+ | Cast (c,k,t) -> mkCast (rename avoid c, k,t)
+ | _ -> c
+ in
+ rename avoid c
diff --git a/pretyping/namegen.mli b/pretyping/namegen.mli
new file mode 100644
index 00000000..096c09b5
--- /dev/null
+++ b/pretyping/namegen.mli
@@ -0,0 +1,77 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* $Id$ *)
+
+open Names
+open Term
+open Environ
+
+(**********************************************************************)
+(* Generating "intuitive" names from their type *)
+
+val lowercase_first_char : identifier -> string
+val sort_hdchar : sorts -> string
+val hdchar : env -> types -> string
+val id_of_name_using_hdchar : env -> types -> name -> identifier
+val named_hd : env -> types -> name -> name
+
+val mkProd_name : env -> name * types * types -> types
+val mkLambda_name : env -> name * types * constr -> constr
+
+(* Deprecated synonyms of [mkProd_name] and [mkLambda_name] *)
+val prod_name : env -> name * types * types -> types
+val lambda_name : env -> name * types * constr -> constr
+
+val prod_create : env -> types * types -> constr
+val lambda_create : env -> types * constr -> constr
+val name_assumption : env -> rel_declaration -> rel_declaration
+val name_context : env -> rel_context -> rel_context
+
+val mkProd_or_LetIn_name : env -> types -> rel_declaration -> types
+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 *)
+
+(* 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 *)
+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 *)
+val next_global_ident_away : identifier -> identifier list -> identifier
+
+(* 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_with_default : string -> name -> identifier list ->
+ identifier
+
+(**********************************************************************)
+(* Making name distinct for displaying *)
+
+type renaming_flags =
+ | RenamingForCasesPattern (* avoid only global constructors *)
+ | RenamingForGoal (* avoid all globals (as in intro) *)
+ | RenamingElsewhereFor of constr
+
+val make_all_name_different : env -> env
+
+val compute_displayed_name_in :
+ renaming_flags -> identifier list -> name -> constr -> name * identifier list
+val compute_and_force_displayed_name_in :
+ renaming_flags -> identifier list -> name -> constr -> name * identifier list
+val compute_displayed_let_name_in :
+ renaming_flags -> identifier list -> name -> constr -> name * identifier list
+val rename_bound_vars_as_displayed : identifier list -> types -> types
diff --git a/pretyping/pattern.ml b/pretyping/pattern.ml
index 057f9d1c..38da0a71 100644
--- a/pretyping/pattern.ml
+++ b/pretyping/pattern.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: pattern.ml 11309 2008-08-06 10:30:35Z herbelin $ *)
+(* $Id$ *)
open Util
open Names
@@ -21,8 +21,10 @@ open Mod_subst
(* Metavariables *)
+type constr_under_binders = identifier list * constr
+
type patvar_map = (patvar * constr) list
-let pr_patvar = pr_id
+type extended_patvar_map = (patvar * constr_under_binders) list
(* Patterns *)
@@ -69,8 +71,8 @@ exception BoundPattern;;
let rec head_pattern_bound t =
match t with
- | PProd (_,_,b) -> head_pattern_bound b
- | PLetIn (_,_,b) -> head_pattern_bound b
+ | PProd (_,_,b) -> head_pattern_bound b
+ | PLetIn (_,_,b) -> head_pattern_bound b
| PApp (c,args) -> head_pattern_bound c
| PIf (c,_,_) -> head_pattern_bound c
| PCase (_,p,c,br) -> head_pattern_bound c
@@ -89,7 +91,11 @@ let head_of_constr_reference c = match kind_of_term c with
| Var id -> VarRef id
| _ -> anomaly "Not a rigid reference"
-let rec pattern_of_constr t =
+open Evd
+
+let pattern_of_constr sigma t =
+ let ctx = ref [] in
+ let rec pattern_of_constr t =
match kind_of_term t with
| Rel n -> PRel n
| Meta n -> PMeta (Some (id_of_string ("META" ^ string_of_int n)))
@@ -100,11 +106,29 @@ let rec pattern_of_constr t =
| LetIn (na,c,_,b) -> PLetIn (na,pattern_of_constr c,pattern_of_constr b)
| Prod (na,c,b) -> PProd (na,pattern_of_constr c,pattern_of_constr b)
| Lambda (na,c,b) -> PLambda (na,pattern_of_constr c,pattern_of_constr b)
- | App (f,a) -> PApp (pattern_of_constr f,Array.map pattern_of_constr a)
- | Const sp -> PRef (ConstRef sp)
- | Ind sp -> PRef (IndRef sp)
- | Construct sp -> PRef (ConstructRef sp)
- | Evar (n,ctxt) -> PEvar (n,Array.map pattern_of_constr ctxt)
+ | App (f,a) ->
+ (match
+ match kind_of_term f with
+ Evar (evk,args as ev) ->
+ (match snd (Evd.evar_source evk sigma) with
+ MatchingVar (true,id) ->
+ ctx := (id,None,existential_type sigma ev)::!ctx;
+ Some id
+ | _ -> None)
+ | _ -> None
+ with
+ | Some n -> PSoApp (n,Array.to_list (Array.map pattern_of_constr a))
+ | None -> PApp (pattern_of_constr f,Array.map (pattern_of_constr) a))
+ | Const sp -> PRef (ConstRef (constant_of_kn(canonical_con sp)))
+ | Ind sp -> PRef (canonical_gr (IndRef sp))
+ | Construct sp -> PRef (canonical_gr (ConstructRef sp))
+ | Evar (evk,ctxt as ev) ->
+ (match snd (Evd.evar_source evk sigma) with
+ | MatchingVar (b,id) ->
+ ctx := (id,None,existential_type sigma ev)::!ctx;
+ assert (not b); PMeta (Some id)
+ | 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
@@ -112,16 +136,20 @@ let rec pattern_of_constr t =
pattern_of_constr p,pattern_of_constr a,
Array.map pattern_of_constr br)
| Fix f -> PFix f
- | CoFix f -> PCoFix f
+ | CoFix f -> PCoFix f in
+ let p = pattern_of_constr t in
+ (* side-effect *)
+ (* Warning: the order of dependencies in ctx is not ensured *)
+ (!ctx,p)
(* To process patterns, we need a translation without typing at all. *)
let map_pattern_with_binders g f l = function
| PApp (p,pl) -> PApp (f l p, Array.map (f l) pl)
| PSoApp (n,pl) -> PSoApp (n, List.map (f l) pl)
- | PLambda (n,a,b) -> PLambda (n,f l a,f (g l) b)
- | PProd (n,a,b) -> PProd (n,f l a,f (g l) b)
- | PLetIn (n,a,b) -> PLetIn (n,f l a,f (g l) b)
+ | PLambda (n,a,b) -> PLambda (n,f l a,f (g n l) b)
+ | PProd (n,a,b) -> PProd (n,f l a,f (g n l) b)
+ | PLetIn (n,a,b) -> PLetIn (n,f l a,f (g n l) b)
| 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)
(* Non recursive *)
@@ -129,31 +157,54 @@ let map_pattern_with_binders g f l = function
(* Bound to terms *)
| PFix _ | PCoFix _ as x) -> x
-let map_pattern f = map_pattern_with_binders (fun () -> ()) (fun () -> f) ()
+let map_pattern f = map_pattern_with_binders (fun _ () -> ()) (fun () -> f) ()
-let rec instantiate_pattern lvar = function
- | PVar id as x -> (try Lazy.force(List.assoc id lvar) with Not_found -> x)
+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
+ ++ strbrk " in pattern because the term refers to " ++ pr_enum pr_id l
+ ++ strbrk " which " ++ str is ++ strbrk " not bound in the pattern.")
+
+let instantiate_pattern sigma lvar c =
+ let rec aux vars = function
+ | PVar id as x ->
+ (try
+ let ctx,c = List.assoc id lvar in
+ try
+ let inst =
+ List.map (fun id -> mkRel (list_index (Name id) vars)) ctx in
+ let c = substl inst c in
+ snd (pattern_of_constr sigma c)
+ with Not_found (* list_index failed *) ->
+ let vars =
+ list_map_filter (function Name id -> Some id | _ -> None) vars in
+ error_instantiate_pattern id (list_subtract ctx vars)
+ with Not_found (* List.assoc failed *) ->
+ x)
| (PFix _ | PCoFix _) -> error ("Non instantiable pattern.")
- | c -> map_pattern (instantiate_pattern lvar) c
+ | c ->
+ map_pattern_with_binders (fun id vars -> id::vars) aux vars c in
+ aux [] c
let rec liftn_pattern k n = function
| PRel i as x -> if i >= n then PRel (i+k) else x
| PFix x -> PFix (destFix (liftn k n (mkFix x)))
| PCoFix x -> PCoFix (destCoFix (liftn k n (mkCoFix x)))
- | c -> map_pattern_with_binders succ (liftn_pattern k) n c
+ | c -> map_pattern_with_binders (fun _ -> succ) (liftn_pattern k) n c
let lift_pattern k = liftn_pattern k 1
-let rec subst_pattern subst pat = match pat with
+let rec subst_pattern subst pat =
+ match pat with
| PRef ref ->
let ref',t = subst_global subst ref in
if ref' == ref then pat else
- pattern_of_constr t
- | PVar _
+ snd (pattern_of_constr Evd.empty t)
+ | PVar _
| PEvar _
| PRel _ -> pat
| PApp (f,args) ->
- let f' = subst_pattern subst f in
+ let f' = subst_pattern subst f in
let args' = array_smartmap (subst_pattern subst) args in
if f' == f && args' == args then pat else
PApp (f',args')
@@ -176,7 +227,7 @@ let rec subst_pattern subst pat = match pat with
let c2' = subst_pattern subst c2 in
if c1' == c1 && c2' == c2 then pat else
PLetIn (name,c1',c2')
- | PSort _
+ | PSort _
| PMeta _ -> pat
| PIf (c,c1,c2) ->
let c' = subst_pattern subst c in
@@ -186,12 +237,12 @@ let rec subst_pattern subst pat = match pat with
PIf (c',c1',c2')
| PCase ((a,b,ind,n as cs),typ,c,branches) ->
let ind' = Option.smartmap (Inductiveops.subst_inductive subst) ind in
- let typ' = subst_pattern subst typ 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')
+ PCase(cs',typ', c', branches')
| PFix fixpoint ->
let cstr = mkFix fixpoint in
let fixpoint' = destFix (subst_mps subst cstr) in
@@ -204,7 +255,7 @@ let rec subst_pattern subst pat = match pat with
PCoFix cofixpoint'
let mkPLambda na b = PLambda(na,PMeta None,b)
-let rev_it_mkPLambda = List.fold_right mkPLambda
+let rev_it_mkPLambda = List.fold_right mkPLambda
let rec pat_of_raw metas vars = function
| RVar (_,id) ->
@@ -212,21 +263,24 @@ let rec pat_of_raw metas vars = function
with Not_found -> PVar id)
| RPatVar (_,(false,n)) ->
metas := n::!metas; PMeta (Some n)
- | RRef (_,r) ->
- PRef r
+ | RRef (_,gr) ->
+ PRef (canonical_gr gr)
(* Hack pour ne pas réécrire une interprétation complète des patterns*)
| RApp (_, RPatVar (_,(true,n)), cl) ->
metas := n::!metas; PSoApp (n, List.map (pat_of_raw metas vars) cl)
- | RApp (_,c,cl) ->
+ | RApp (_,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) ->
+ 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) ->
+ 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) ->
+ 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) ->
@@ -261,7 +315,7 @@ let rec pat_of_raw metas vars = function
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)
-
+
| r ->
let loc = loc_of_rawconstr r in
user_err_loc (loc,"pattern_of_rawconstr", Pp.str"Non supported pattern.")
@@ -284,7 +338,7 @@ and pat_of_raw_branch loc metas vars ind brs i =
| PatCstr(loc,_,_,_) ->
user_err_loc (loc,"pattern_of_rawconstr",
Pp.str "Non supported pattern.")) lv in
- let vars' = List.rev lna @ vars 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) ++
@@ -292,5 +346,5 @@ and pat_of_raw_branch loc metas vars ind brs i =
let pattern_of_rawconstr c =
let metas = ref [] in
- let p = pat_of_raw metas [] c in
+ let p = pat_of_raw metas [] c in
(!metas,p)
diff --git a/pretyping/pattern.mli b/pretyping/pattern.mli
index 4102db9e..6beb996b 100644
--- a/pretyping/pattern.mli
+++ b/pretyping/pattern.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: pattern.mli 8963 2006-06-19 18:54:49Z barras $ i*)
+(*i $Id$ i*)
(*i*)
open Pp
@@ -20,10 +20,11 @@ open Rawterm
open Mod_subst
(*i*)
-(* Pattern variables *)
+(* Types of substitutions with or w/o bound variables *)
+type constr_under_binders = identifier list * constr
type patvar_map = (patvar * constr) list
-val pr_patvar : patvar -> std_ppcmds
+type extended_patvar_map = (patvar * constr_under_binders) list
(* Patterns *)
@@ -45,6 +46,8 @@ type constr_pattern =
| PFix of fixpoint
| PCoFix of cofixpoint
+(** {5 Functions on patterns} *)
+
val occur_meta_pattern : constr_pattern -> bool
val subst_pattern : substitution -> constr_pattern -> constr_pattern
@@ -66,16 +69,17 @@ val head_of_constr_reference : Term.constr -> global_reference
a pattern; currently, no destructor (Cases, Fix, Cofix) and no
existential variable are allowed in [c] *)
-val pattern_of_constr : constr -> constr_pattern
+val pattern_of_constr : Evd.evar_map -> constr -> named_context * constr_pattern
(* [pattern_of_rawconstr 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_rawconstr : rawconstr ->
patvar list * constr_pattern
val instantiate_pattern :
- (identifier * constr_pattern Lazy.t) list -> constr_pattern -> constr_pattern
+ Evd.evar_map -> (identifier * (identifier list * constr)) list ->
+ constr_pattern -> constr_pattern
val lift_pattern : int -> constr_pattern -> constr_pattern
diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml
index a513d558..02d0a11f 100644
--- a/pretyping/pretype_errors.ml
+++ b/pretyping/pretype_errors.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: pretype_errors.ml 10860 2008-04-27 21:39:08Z herbelin $ *)
+(* $Id$ *)
open Util
open Stdpp
@@ -14,6 +14,7 @@ open Names
open Sign
open Term
open Termops
+open Namegen
open Environ
open Type_errors
open Rawterm
@@ -25,7 +26,7 @@ type pretype_error =
(* Unification *)
| OccurCheck of existential_key * constr
| NotClean of existential_key * constr * Evd.hole_kind
- | UnsolvableImplicit of Evd.evar_info * Evd.hole_kind *
+ | UnsolvableImplicit of Evd.evar_info * Evd.hole_kind *
Evd.unsolvability_explanation option
| CannotUnify of constr * constr
| CannotUnifyLocal of constr * constr * constr
@@ -33,6 +34,8 @@ type pretype_error =
| CannotGeneralize of constr
| NoOccurrenceFound of constr * identifier option
| CannotFindWellTypedAbstraction of constr * constr list
+ | AbstractionOverMeta of name * name
+ | NonLinearUnification of name * constr
(* Pretyping *)
| VarNotFound of identifier
| UnexpectedType of constr * constr
@@ -47,10 +50,15 @@ let precatchable_exception = function
| _ -> false
let nf_evar = Reductionops.nf_evar
-let j_nf_evar sigma j =
+let j_nf_evar sigma j =
{ uj_val = nf_evar sigma j.uj_val;
uj_type = nf_evar sigma j.uj_type }
+let j_nf_betaiotaevar sigma j =
+ { uj_val = nf_evar sigma j.uj_val;
+ uj_type = Reductionops.nf_betaiota sigma j.uj_type }
let jl_nf_evar sigma jl = List.map (j_nf_evar sigma) jl
+let jl_nf_betaiotaevar sigma jl =
+ List.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}
@@ -76,7 +84,7 @@ let contract env lc =
| Some c' when isRel c' ->
l := (substl !l c') :: !l;
env
- | _ ->
+ | _ ->
let t' = substl !l t in
let c' = Option.map (substl !l) c in
let na' = named_hd env t' na in
@@ -111,11 +119,11 @@ let error_cant_apply_not_functional_loc loc env sigma rator randl =
CantApplyNonFunctional (j_nf_evar sigma rator, ja))
let error_cant_apply_bad_type_loc loc env sigma (n,c,t) rator randl =
- let ja = Array.of_list (jl_nf_evar sigma randl) in
+ let ja = Array.of_list (jl_nf_betaiotaevar sigma randl) in
raise_located_type_error
(loc, env, sigma,
CantApplyBadType
- ((n,nf_evar sigma c, nf_evar sigma t),
+ ((n,nf_evar sigma c, Reductionops.nf_betaiota sigma t),
j_nf_evar sigma rator, ja))
let error_ill_formed_branch_loc loc env sigma c i actty expty =
@@ -161,7 +169,7 @@ let error_unsolvable_implicit loc env sigma evi e explain =
let error_cannot_unify env sigma (m,n) =
raise (PretypeError (env_ise sigma env,CannotUnify (m,n)))
-let error_cannot_unify_local env sigma (m,n,sn) =
+let error_cannot_unify_local env sigma (m,n,sn) =
raise (PretypeError (env_ise sigma env,CannotUnifyLocal (m,n,sn)))
let error_cannot_coerce env sigma (m,n) =
@@ -170,6 +178,14 @@ let error_cannot_coerce env sigma (m,n) =
let error_cannot_find_well_typed_abstraction env sigma p l =
raise (PretypeError (env_ise sigma env,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)))
+
+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)))
+
(*s Ml Case errors *)
let error_cant_find_case_type_loc loc env sigma expr =
diff --git a/pretyping/pretype_errors.mli b/pretyping/pretype_errors.mli
index 6ad2793f..537bc386 100644
--- a/pretyping/pretype_errors.mli
+++ b/pretyping/pretype_errors.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: pretype_errors.mli 10860 2008-04-27 21:39:08Z herbelin $ i*)
+(*i $Id$ i*)
(*i*)
open Pp
@@ -27,7 +27,7 @@ type pretype_error =
(* Unification *)
| OccurCheck of existential_key * constr
| NotClean of existential_key * constr * Evd.hole_kind
- | UnsolvableImplicit of Evd.evar_info * Evd.hole_kind *
+ | UnsolvableImplicit of Evd.evar_info * Evd.hole_kind *
Evd.unsolvability_explanation option
| CannotUnify of constr * constr
| CannotUnifyLocal of constr * constr * constr
@@ -35,6 +35,8 @@ type pretype_error =
| CannotGeneralize of constr
| NoOccurrenceFound of constr * identifier option
| CannotFindWellTypedAbstraction of constr * constr list
+ | AbstractionOverMeta of name * name
+ | NonLinearUnification of name * constr
(* Pretyping *)
| VarNotFound of identifier
| UnexpectedType of constr * constr
@@ -59,22 +61,22 @@ val tj_nf_evar :
val error_actual_type_loc :
loc -> env -> Evd.evar_map -> unsafe_judgment -> constr -> 'b
-val error_cant_apply_not_functional_loc :
+val error_cant_apply_not_functional_loc :
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 ->
+val error_cant_apply_bad_type_loc :
+ 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
-val error_ill_formed_branch_loc :
+val error_ill_formed_branch_loc :
loc -> env -> Evd.evar_map ->
constr -> int -> constr -> constr -> 'b
-val error_number_branches_loc :
+val error_number_branches_loc :
loc -> env -> Evd.evar_map ->
unsafe_judgment -> int -> 'b
@@ -95,7 +97,7 @@ val error_not_clean :
env -> Evd.evar_map -> existential_key -> constr -> loc * Evd.hole_kind -> 'b
val error_unsolvable_implicit :
- loc -> env -> Evd.evar_map -> Evd.evar_info -> Evd.hole_kind ->
+ loc -> env -> Evd.evar_map -> Evd.evar_info -> Evd.hole_kind ->
Evd.unsolvability_explanation option -> 'b
val error_cannot_unify : env -> Evd.evar_map -> constr * constr -> 'b
@@ -105,6 +107,12 @@ val error_cannot_unify_local : env -> Evd.evar_map -> constr * constr * constr -
val error_cannot_find_well_typed_abstraction : env -> Evd.evar_map ->
constr -> constr list -> 'b
+val error_abstraction_over_meta : env -> Evd.evar_map ->
+ metavariable -> metavariable -> 'b
+
+val error_non_linear_unification : env -> Evd.evar_map ->
+ metavariable -> constr -> 'b
+
(*s Ml Case errors *)
val error_cant_find_case_type_loc :
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index c0f820a2..e11811ec 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: pretyping.ml 12053 2009-04-06 16:20:42Z msozeau $ *)
+(* $Id$ *)
open Pp
open Util
@@ -23,17 +23,18 @@ open Libnames
open Nameops
open Classops
open List
-open Recordops
+open Recordops
open Evarutil
open Pretype_errors
open Rawterm
open Evarconv
open Pattern
-open Dyn
type typing_constraint = OfType of types option | IsType
-type var_map = (identifier * unsafe_judgment) list
+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
(************************************************************************)
(* This concerns Cases *)
@@ -47,119 +48,117 @@ open Inductiveops
exception Found of int array
-let search_guard loc env possible_indexes fixdefs =
+let search_guard loc env possible_indexes fixdefs =
(* Standard situation with only one possibility for each fix. *)
(* We treat it separately in order to get proper error msg. *)
- if List.for_all (fun l->1=List.length l) possible_indexes then
- let indexes = Array.of_list (List.map List.hd possible_indexes) in
+ if List.for_all (fun l->1=List.length l) possible_indexes then
+ let indexes = Array.of_list (List.map List.hd possible_indexes) in
let fix = ((indexes, 0),fixdefs) in
- (try check_fix env fix with
+ (try check_fix env fix with
| e -> if loc = dummy_loc then raise e else Stdpp.raise_with_loc loc e);
indexes
else
(* we now search recursively amoungst all combinations *)
- (try
- List.iter
- (fun l ->
- let indexes = Array.of_list l in
+ (try
+ List.iter
+ (fun l ->
+ let indexes = Array.of_list l in
let fix = ((indexes, 0),fixdefs) in
- try check_fix env fix; raise (Found indexes)
+ try check_fix env fix; raise (Found indexes)
with TypeError _ -> ())
- (list_combinations possible_indexes);
- let errmsg = "Cannot guess decreasing argument of fix." in
- if loc = dummy_loc then error errmsg else
+ (list_combinations possible_indexes);
+ let errmsg = "Cannot guess decreasing argument of fix." in
+ if loc = dummy_loc then error errmsg else
user_err_loc (loc,"search_guard", Pp.str errmsg)
with Found indexes -> indexes)
(* To embed constr in rawconstr *)
let ((constr_in : constr -> Dyn.t),
- (constr_out : Dyn.t -> constr)) = create "constr"
+ (constr_out : Dyn.t -> constr)) = Dyn.create "constr"
(** Miscellaneous interpretation functions *)
-
+
let interp_sort = function
| RProp c -> Prop c
| RType _ -> new_Type_sort ()
-
+
let interp_elimination_sort = function
| RProp Null -> InProp
| RProp Pos -> InSet
| RType _ -> InType
-module type S =
+module type S =
sig
module Cases : Cases.S
-
+
(* Allow references to syntaxically inexistent variables (i.e., if applied on an inductive) *)
val allow_anonymous_refs : bool ref
(* Generic call to the interpreter from rawconstr 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
- val understand_tcc_evars :
- evar_defs ref -> env -> typing_constraint -> rawconstr -> constr
-
+ val understand_tcc_evars : ?fail_evar:bool -> ?resolve_classes:bool ->
+ evar_map ref -> env -> typing_constraint -> rawconstr -> 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.
-
- In [understand_ltac sigma env ltac_env constraint c],
-
+
+ In [understand_ltac expand_evars sigma env ltac_env constraint c],
+
+ expand_evars : expand inferred evars by their value if any
sigma : initial set of existential variables (typically dependent subgoals)
ltac_env : partial substitution of variables (used for the tactic language)
- constraint : tell if interpreted as a possibly constrained term or a type
+ constraint : tell if interpreted as a possibly constrained term or a type
*)
-
+
val understand_ltac :
- evar_map -> env -> var_map * unbound_ltac_var_map ->
- typing_constraint -> rawconstr -> evar_defs * constr
-
+ bool -> evar_map -> env -> ltac_var_map ->
+ typing_constraint -> rawconstr -> evar_map * constr
+
(* Standard call to get a constr from a rawconstr, resolving implicit args *)
-
+
val understand : evar_map -> env -> ?expected_type:Term.types ->
rawconstr -> constr
-
+
(* Idem but the rawconstr is intended to be a type *)
-
+
val understand_type : evar_map -> env -> rawconstr -> constr
-
+
(* A generalization of the two previous case *)
-
- val understand_gen : typing_constraint -> evar_map -> env ->
+
+ val understand_gen : typing_constraint -> evar_map -> env ->
rawconstr -> constr
-
+
(* Idem but returns the judgment of the understood term *)
-
+
val understand_judgment : evar_map -> env -> rawconstr -> unsafe_judgment
(* Idem but do not fail on unresolved evars *)
- val understand_judgment_tcc : evar_defs ref -> env -> rawconstr -> unsafe_judgment
+ val understand_judgment_tcc : evar_map ref -> env -> rawconstr -> unsafe_judgment
(*i*)
(* Internal of Pretyping...
* Unused outside, but useful for debugging
*)
- val pretype :
- type_constraint -> env -> evar_defs ref ->
- var_map * (identifier * identifier option) list ->
- rawconstr -> unsafe_judgment
-
- val pretype_type :
- val_constraint -> env -> evar_defs ref ->
- var_map * (identifier * identifier option) list ->
- rawconstr -> unsafe_type_judgment
+ val pretype :
+ type_constraint -> env -> evar_map ref ->
+ ltac_var_map -> rawconstr -> unsafe_judgment
+
+ val pretype_type :
+ val_constraint -> env -> evar_map ref ->
+ ltac_var_map -> rawconstr -> unsafe_type_judgment
val pretype_gen :
- evar_defs ref -> env ->
- var_map * (identifier * identifier option) list ->
- typing_constraint -> rawconstr -> constr
+ bool -> bool -> bool -> evar_map ref -> env ->
+ ltac_var_map -> typing_constraint -> rawconstr -> constr
(*i*)
end
@@ -190,28 +189,28 @@ module Pretyping_F (Coercion : Coercion.S) = struct
let (evd',t) = f !evdref x y z in
evdref := evd';
t
-
+
let mt_evd = Evd.empty
-
+
(* Utilisé pour inférer le prédicat des Cases *)
(* Semble exagérement fort *)
(* Faudra préférer une unification entre les types de toutes les clauses *)
(* et autoriser des ? à rester dans le résultat de l'unification *)
-
+
let evar_type_fixpoint loc env evdref lna lar vdefj =
- let lt = Array.length vdefj in
- if Array.length lar = lt then
- for i = 0 to lt-1 do
+ let lt = Array.length vdefj in
+ if Array.length lar = lt then
+ for i = 0 to lt-1 do
if not (e_cumul env evdref (vdefj.(i)).uj_type
(lift lt lar.(i))) then
- error_ill_typed_rec_body_loc loc env (evars_of !evdref)
+ error_ill_typed_rec_body_loc loc env !evdref
i lna vdefj lar
done
- let check_branches_message loc env evdref c (explft,lft) =
+ 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 = evars_of !evdref in
+ 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
@@ -229,13 +228,23 @@ module Pretyping_F (Coercion : Coercion.S) = struct
| Anonymous -> name'
| _ -> name
- let pretype_id loc env (lvar,unbndltacvars) id =
+ let invert_ltac_bound_name env id0 id =
+ try mkRel (pi1 (lookup_rel_id id (rel_context env)))
+ with Not_found ->
+ errorlabstrm "" (str "Ltac variable " ++ pr_id id0 ++
+ str " depends on pattern variable name " ++ pr_id id ++
+ str " which is not bound in current context.")
+
+ let pretype_id loc env sigma (lvar,unbndltacvars) id =
try
- let (n,typ) = lookup_rel_id id (rel_context env) in
+ let (n,_,typ) = lookup_rel_id id (rel_context env) in
{ uj_val = mkRel n; uj_type = lift n typ }
with Not_found ->
try
- List.assoc id lvar
+ 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 = Retyping.get_type_of env sigma c }
with Not_found ->
try
let (_,_,typ) = lookup_named id env in
@@ -257,22 +266,22 @@ module Pretyping_F (Coercion : Coercion.S) = struct
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]))
+ | _ -> 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=lam_it ccl' sign; uj_type=prod_it s' sign}
+ {uj_val=it_mkLambda ccl' sign; uj_type=it_mkProd s' sign}
let evar_kind_of_term sigma c =
- kind_of_term (whd_evar (Evd.evars_of sigma) c)
+ kind_of_term (whd_evar sigma c)
(*************************************************************************)
(* Main pretyping function *)
- let pretype_ref evdref env ref =
+ let pretype_ref evdref env ref =
let c = constr_of_global ref in
make_judge c (Retyping.get_type_of env Evd.empty c)
@@ -293,26 +302,32 @@ module Pretyping_F (Coercion : Coercion.S) = struct
| RVar (loc, id) ->
inh_conv_coerce_to_tycon loc env evdref
- (pretype_id loc env lvar id)
+ (pretype_id loc env !evdref lvar id)
tycon
| REvar (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 (evars_of !evdref) evk) in
+ let hyps = evar_filtered_context (Evd.find !evdref evk) in
let args = match instopt with
| None -> instance_from_named_context hyps
| Some inst -> failwith "Evar subtitutions not implemented" in
let c = mkEvar (evk, args) in
- let j = (Retyping.get_judgment_of env (evars_of !evdref) c) in
+ let j = (Retyping.get_judgment_of env !evdref c) in
inh_conv_coerce_to_tycon loc env evdref j tycon
- | RPatVar (loc,(someta,n)) ->
- anomaly "Found a pattern variable in a rawterm to type"
-
+ | RPatVar (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
+ let k = MatchingVar (someta,n) in
+ { uj_val = e_new_evar evdref env ~src:(loc,k) ty; uj_type = ty }
+
| RHole (loc,k) ->
let ty =
- match tycon with
+ match tycon with
| Some (None, ty) -> ty
| None | Some _ ->
e_new_evar evdref env ~src:(loc,InternalHole) (new_Type ()) in
@@ -343,7 +358,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
(* Note: bodies are not used by push_rec_types, so [||] is safe *)
let newenv = push_rec_types (names,ftys,[||]) env in
let vdefj =
- array_map2_i
+ array_map2_i
(fun i ctxt def ->
(* we lift nbfix times the type in tycon, because of
* the nbfix variables pushed to newenv *)
@@ -356,24 +371,24 @@ module Pretyping_F (Coercion : Coercion.S) = struct
uj_type = it_mkProd_or_LetIn j.uj_type ctxt })
ctxtv vdef in
evar_type_fixpoint loc env evdref names ftys vdefj;
- let ftys = Array.map (nf_evar (evars_of !evdref)) ftys in
- let fdefs = Array.map (fun x -> nf_evar (evars_of !evdref) (j_val x)) vdefj in
+ 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) ->
(* First, let's find the guard indexes. *)
(* If recursive argument was not given by user, we try all args.
An earlier approach was to look only for inductive arguments,
- but doing it properly involves delta-reduction, and it finally
- doesn't seem worth the effort (except for huge mutual
+ but doing it properly involves delta-reduction, and it finally
+ doesn't seem worth the effort (except for huge mutual
fixpoints ?) *)
- let possible_indexes = Array.to_list (Array.mapi
- (fun i (n,_) -> match n with
+ let possible_indexes = Array.to_list (Array.mapi
+ (fun i (n,_) -> match n with
| Some n -> [n]
| None -> list_map_i (fun i _ -> i) 0 ctxtv.(i))
vn)
- in
- let fixdecls = (names,ftys,fdefs) in
- let indexes = search_guard loc env possible_indexes fixdecls in
+ in
+ let fixdecls = (names,ftys,fdefs) in
+ let indexes = search_guard loc env possible_indexes fixdecls in
make_judge (mkFix ((indexes,i),fixdecls)) ftys.(i)
| RCoFix i ->
let cofix = (i,(names,ftys,fdefs)) in
@@ -384,7 +399,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
| RSort (loc,s) ->
inh_conv_coerce_to_tycon loc env evdref (pretype_sort s) tycon
- | RApp (loc,f,args) ->
+ | RApp (loc,f,args) ->
let fj = pretype empty_tycon env evdref lvar f in
let floc = loc_of_rawconstr f in
let rec apply_rec env n resj = function
@@ -392,34 +407,34 @@ module Pretyping_F (Coercion : Coercion.S) = struct
| c::rest ->
let argloc = loc_of_rawconstr c in
let resj = evd_comb1 (Coercion.inh_app_fun env) evdref resj in
- let resty = whd_betadeltaiota env (evars_of !evdref) resj.uj_type in
+ let resty = whd_betadeltaiota env !evdref resj.uj_type in
match kind_of_term resty with
| Prod (na,c1,c2) ->
let hj = pretype (mk_tycon c1) env evdref lvar c in
let value, typ = applist (j_val resj, [j_val hj]), subst1 hj.uj_val c2 in
- apply_rec env (n+1)
+ apply_rec env (n+1)
{ uj_val = value;
uj_type = typ }
rest
| _ ->
let hj = pretype empty_tycon env evdref lvar c in
- error_cant_apply_not_functional_loc
- (join_loc floc argloc) env (evars_of !evdref)
+ error_cant_apply_not_functional_loc
+ (join_loc floc argloc) env !evdref
resj [hj]
in
let resj = apply_rec env 1 fj args in
let resj =
match evar_kind_of_term !evdref resj.uj_val with
| App (f,args) ->
- let f = whd_evar (Evd.evars_of !evdref) f in
+ let f = whd_evar !evdref f in
begin match kind_of_term f with
| Ind _ | Const _
when isInd f or has_polymorphic_type (destConst f)
->
- let sigma = evars_of !evdref in
+ let sigma = !evdref in
let c = mkApp (f,Array.map (whd_evar sigma) args) in
let t = Retyping.get_type_of env sigma c in
- make_judge c t
+ make_judge c (* use this for keeping evars: resj.uj_val *) t
| _ -> resj end
| _ -> resj in
inh_conv_coerce_to_tycon loc env evdref resj tycon
@@ -429,24 +444,30 @@ module Pretyping_F (Coercion : Coercion.S) = struct
let dom_valcon = valcon_of_tycon dom in
let j = pretype_type dom_valcon env evdref lvar c1 in
let var = (name,None,j.utj_val) in
- let j' = pretype rng (push_rel var env) evdref lvar c2 in
+ 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) ->
let j = pretype_type empty_valcon env evdref lvar c1 in
- let var = (name,j.utj_val) in
- let env' = push_rel_assum var env in
- let j' = pretype_type empty_valcon env' evdref lvar c2 in
+ let j' =
+ if name = Anonymous then
+ let j = pretype_type empty_valcon env evdref lvar c2 in
+ { j with utj_val = lift 1 j.utj_val }
+ else
+ let var = (name,j.utj_val) in
+ let env' = push_rel_assum var env in
+ pretype_type empty_valcon env' evdref lvar c2
+ in
let resj =
try judge_of_product env name j j'
with TypeError _ as e -> Stdpp.raise_with_loc loc e in
inh_conv_coerce_to_tycon loc env evdref resj tycon
-
+
| RLetIn(loc,name,c1,c2) ->
- let j =
+ let j =
match c1 with
| RCast (loc, c, CastConv (DEFAULTcast, t)) ->
- let tj = pretype_type empty_valcon env evdref lvar t in
+ let tj = pretype_type empty_valcon env evdref lvar t in
pretype (mk_tycon tj.utj_val) env evdref lvar c
| _ -> pretype empty_tycon env evdref lvar c1
in
@@ -459,11 +480,11 @@ module Pretyping_F (Coercion : Coercion.S) = struct
| RLetTuple (loc,nal,(na,po),c,d) ->
let cj = pretype empty_tycon env evdref lvar c in
- let (IndType (indf,realargs)) =
- try find_rectype env (evars_of !evdref) cj.uj_type
+ let (IndType (indf,realargs)) =
+ try find_rectype env !evdref cj.uj_type
with Not_found ->
let cloc = loc_of_rawconstr c in
- error_case_not_inductive_loc cloc env (evars_of !evdref) cj
+ error_case_not_inductive_loc cloc env !evdref cj
in
let cstrs = get_constructors env indf in
if Array.length cstrs <> 1 then
@@ -487,59 +508,59 @@ module Pretyping_F (Coercion : Coercion.S) = struct
| Some p ->
let env_p = push_rels psign env in
let pj = pretype_type empty_valcon env_p evdref lvar p in
- let ccl = nf_isevar !evdref pj.utj_val in
+ let ccl = nf_evar !evdref pj.utj_val in
let psign = make_arity_signature env true indf in (* with names *)
let p = it_mkLambda_or_LetIn ccl psign in
- let inst =
+ let inst =
(Array.to_list cs.cs_concl_realargs)
@[build_dependent_constructor cs] in
let lp = lift cs.cs_nargs p in
- let fty = hnf_lam_applist env (evars_of !evdref) lp inst in
+ let fty = hnf_lam_applist env !evdref lp inst in
let fj = pretype (mk_tycon fty) env_f evdref lvar d in
let f = it_mkLambda_or_LetIn fj.uj_val fsign in
let v =
let mis,_ = dest_ind_family indf in
let ci = make_case_info env mis LetStyle in
- mkCase (ci, p, cj.uj_val,[|f|]) in
+ mkCase (ci, p, cj.uj_val,[|f|]) in
{ uj_val = v; uj_type = substl (realargs@[cj.uj_val]) ccl }
- | None ->
+ | None ->
let tycon = lift_tycon cs.cs_nargs tycon in
let fj = pretype tycon env_f evdref lvar d in
let f = it_mkLambda_or_LetIn fj.uj_val fsign in
- let ccl = nf_isevar !evdref fj.uj_type in
+ let ccl = nf_evar !evdref fj.uj_type in
let ccl =
if noccur_between 1 cs.cs_nargs ccl then
- lift (- cs.cs_nargs) ccl
+ lift (- cs.cs_nargs) ccl
else
- error_cant_find_case_type_loc loc env (evars_of !evdref)
+ error_cant_find_case_type_loc loc env !evdref
cj.uj_val in
let ccl = refresh_universes ccl in
let p = it_mkLambda_or_LetIn (lift (nar+1) ccl) psign in
let v =
let mis,_ = dest_ind_family indf in
let ci = make_case_info env mis LetStyle in
- mkCase (ci, p, cj.uj_val,[|f|] )
+ mkCase (ci, p, cj.uj_val,[|f|] )
in
{ uj_val = v; uj_type = ccl })
| RIf (loc,c,(na,po),b1,b2) ->
let cj = pretype empty_tycon env evdref lvar c in
- let (IndType (indf,realargs)) =
- try find_rectype env (evars_of !evdref) cj.uj_type
+ let (IndType (indf,realargs)) =
+ try find_rectype env !evdref cj.uj_type
with Not_found ->
let cloc = loc_of_rawconstr c in
- error_case_not_inductive_loc cloc env (evars_of !evdref) cj in
- let cstrs = get_constructors env indf in
+ error_case_not_inductive_loc cloc env !evdref cj in
+ let cstrs = get_constructors env indf in
if Array.length cstrs <> 2 then
user_err_loc (loc,"",
str "If is only for inductive types with two constructors.");
- let arsgn =
+ let arsgn =
let arsgn,_ = get_arity env indf in
if not !allow_anonymous_refs then
(* Make dependencies from arity signature impossible *)
- List.map (fun (_,b,t) -> (Anonymous,b,t)) arsgn
+ List.map (fun (_,b,t) -> (Anonymous,b,t)) arsgn
else arsgn
in
let nar = List.length arsgn in
@@ -548,40 +569,38 @@ module Pretyping_F (Coercion : Coercion.S) = struct
| Some p ->
let env_p = push_rels psign env in
let pj = pretype_type empty_valcon env_p evdref lvar p in
- let ccl = nf_evar (evars_of !evdref) pj.utj_val in
+ let ccl = nf_evar !evdref pj.utj_val in
let pred = it_mkLambda_or_LetIn ccl psign in
let typ = lift (- nar) (beta_applist (pred,[cj.uj_val])) in
let jtyp = inh_conv_coerce_to_tycon loc env evdref {uj_val = pred;
- uj_type = typ} tycon
+ uj_type = typ} tycon
in
jtyp.uj_val, jtyp.uj_type
- | None ->
+ | None ->
let p = match tycon with
| Some (None, ty) -> ty
| None | Some _ ->
e_new_evar evdref env ~src:(loc,InternalHole) (new_Type ())
in
it_mkLambda_or_LetIn (lift (nar+1) p) psign, p in
- let pred = nf_evar (evars_of !evdref) pred in
- let p = nf_evar (evars_of !evdref) p in
- (* msgnl (str "Pred is: " ++ Termops.print_constr_env env pred);*)
+ let pred = nf_evar !evdref pred in
+ let p = nf_evar !evdref p in
let f cs b =
let n = rel_context_length cs.cs_args in
let pi = lift n pred in (* liftn n 2 pred ? *)
let pi = beta_applist (pi, [build_dependent_constructor cs]) in
- let csgn =
+ let csgn =
if not !allow_anonymous_refs then
- List.map (fun (_,b,t) -> (Anonymous,b,t)) cs.cs_args
- else
- List.map
+ List.map (fun (_,b,t) -> (Anonymous,b,t)) cs.cs_args
+ else
+ List.map
(fun (n, b, t) ->
match n with
Name _ -> (n, b, t)
| Anonymous -> (Name (id_of_string "H"), b, t))
cs.cs_args
in
- let env_c = push_rels csgn env in
-(* msgnl (str "Pi is: " ++ Termops.print_constr_env env_c pi); *)
+ let env_c = push_rels csgn env in
let bj = pretype (mk_tycon pi) env_c evdref lvar b in
it_mkLambda_or_LetIn bj.uj_val cs.cs_args in
let b1 = f cstrs.(0) b1 in
@@ -592,7 +611,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
mkCase (ci, pred, cj.uj_val, [|b1;b2|])
in
{ uj_val = v; uj_type = p }
-
+
| RCases (loc,sty,po,tml,eqns) ->
Cases.compile_cases loc sty
((fun vtyc env evdref -> pretype vtyc env evdref lvar),evdref)
@@ -606,17 +625,21 @@ module Pretyping_F (Coercion : Coercion.S) = struct
evd_comb1 (Coercion.inh_coerce_to_base loc env) evdref cj
| CastConv (k,t) ->
let tj = pretype_type empty_valcon env evdref lvar t in
- let cj = pretype (mk_tycon tj.utj_val) env evdref lvar c in
- (* User Casts are for helping pretyping, experimentally not to be kept*)
- (* ... except for Correctness *)
- let v = mkCast (cj.uj_val, k, tj.utj_val) in
- { uj_val = v; uj_type = tj.utj_val }
+ 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) ->
+ ignore (Reduction.vm_conv Reduction.CUMUL env cty tval); cj
+ | _ -> 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 (tag d) = "constr" then
+ if (Dyn.tag d) = "constr" then
let c = constr_out d in
- let j = (Retyping.get_judgment_of env (evars_of !evdref) c) in
+ let j = (Retyping.get_judgment_of env !evdref c) in
j
(*inh_conv_coerce_to_tycon loc env evdref j tycon*)
else
@@ -628,11 +651,11 @@ module Pretyping_F (Coercion : Coercion.S) = struct
(match valcon with
| Some v ->
let s =
- let sigma = evars_of !evdref in
+ let sigma = !evdref in
let t = Retyping.get_type_of env sigma v in
match kind_of_term (whd_betadeltaiota env sigma t) with
| Sort s -> s
- | Evar ev when is_Type (existential_type sigma ev) ->
+ | Evar ev when is_Type (existential_type sigma ev) ->
evd_comb1 (define_evar_as_sort) evdref ev
| _ -> anomaly "Found a type constraint which is not a type"
in
@@ -652,24 +675,24 @@ 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 (evars_of !evdref) tj.utj_val v
+ (loc_of_rawconstr c) env !evdref tj.utj_val v
- let pretype_gen_aux evdref env lvar kind c =
+ let pretype_gen expand_evar fail_evar resolve_classes evdref env lvar kind c =
let c' = match kind with
| OfType exptyp ->
let tycon = match exptyp with None -> empty_tycon | Some t -> mk_tycon t in
- (pretype tycon env evdref lvar c).uj_val
+ (pretype tycon env evdref lvar c).uj_val
| IsType ->
(pretype_type empty_valcon env evdref lvar c).utj_val in
- let evd,_ = consider_remaining_unif_problems env !evdref in
- evdref := evd;
- nf_isevar !evdref c'
-
- let pretype_gen evdref env lvar kind c =
- let c = pretype_gen_aux evdref env lvar kind c in
- evdref := Typeclasses.resolve_typeclasses ~onlyargs:true ~fail:false env !evdref;
- nf_isevar !evdref c
-
+ evdref := fst (consider_remaining_unif_problems env !evdref);
+ if resolve_classes then
+ evdref :=
+ 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
+
(* TODO: comment faire remonter l'information si le typage a resolu des
variables du sigma original. il faudrait que la fonction de typage
retourne aussi le nouveau sigma...
@@ -679,59 +702,45 @@ module Pretyping_F (Coercion : Coercion.S) = struct
let evdref = ref (create_evar_defs sigma) in
let j = pretype empty_tycon env evdref ([],[]) c in
let evd,_ = consider_remaining_unif_problems env !evdref in
- let j = j_nf_evar (evars_of evd) j in
- let evd = Typeclasses.resolve_typeclasses ~onlyargs:true ~fail:true env evd in
- let j = j_nf_evar (evars_of evd) j 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
let understand_judgment_tcc evdref env c =
let j = pretype empty_tycon env evdref ([],[]) c in
- let sigma = evars_of !evdref in
- let j = j_nf_evar sigma j in
- j
+ 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 fail_evar sigma env lvar kind c =
+ 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 c = pretype_gen_aux evdref env lvar kind c in
- if fail_evar then
- let evd = Typeclasses.resolve_typeclasses ~onlyargs:false ~fail:true env !evdref in
- let c = Evarutil.nf_isevar evd c in
- check_evars env Evd.empty evd c;
- evd, c
- else !evdref, c
+ let c = pretype_gen expand_evar fail_evar resolve_classes evdref env lvar kind c in
+ !evdref, c
(** Entry points of the high-level type synthesis algorithm *)
let understand_gen kind sigma env c =
- snd (ise_pretype_gen true sigma env ([],[]) kind c)
+ snd (ise_pretype_gen true true true sigma env ([],[]) kind c)
let understand sigma env ?expected_type:exptyp c =
- snd (ise_pretype_gen true sigma env ([],[]) (OfType exptyp) c)
+ snd (ise_pretype_gen true true true sigma env ([],[]) (OfType exptyp) c)
let understand_type sigma env c =
- snd (ise_pretype_gen true sigma env ([],[]) IsType c)
-
- let understand_ltac sigma env lvar kind c =
- ise_pretype_gen false sigma env lvar kind c
-
- let understand_tcc_evars evdref env kind c =
- pretype_gen evdref env ([],[]) kind c
-
+ snd (ise_pretype_gen true true true sigma env ([],[]) IsType c)
+
+ let understand_ltac expand_evar sigma env lvar kind c =
+ ise_pretype_gen expand_evar false false sigma env lvar kind c
+
let understand_tcc ?(resolve_classes=true) sigma env ?expected_type:exptyp c =
- let evd, t =
- let evdref = ref (Evd.create_evar_defs sigma) in
- let c =
- if resolve_classes then
- pretype_gen evdref env ([],[]) (OfType exptyp) c
- else
- pretype_gen_aux evdref env ([],[]) (OfType exptyp) c
- in !evdref, c
- in Evd.evars_of evd, t
+ ise_pretype_gen true false resolve_classes sigma env ([],[]) (OfType exptyp) c
+
+ let understand_tcc_evars ?(fail_evar=false) ?(resolve_classes=true) evdref env kind c =
+ pretype_gen true fail_evar resolve_classes evdref env ([],[]) kind c
end
-
+
module Default : S = Pretyping_F(Coercion.Default)
diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli
index 4f116053..d28b076d 100644
--- a/pretyping/pretyping.mli
+++ b/pretyping/pretyping.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: pretyping.mli 11047 2008-06-03 23:08:00Z msozeau $ i*)
+(*i $Id$ i*)
(*i*)
open Names
@@ -20,99 +20,99 @@ open Evarutil
(* An auxiliary function for searching for fixpoint guard indexes *)
-val search_guard :
+val search_guard :
Util.loc -> env -> int list list -> rec_declaration -> int array
type typing_constraint = OfType of types option | IsType
-type var_map = (identifier * unsafe_judgment) list
+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
-module type S =
+module type S =
sig
module Cases : Cases.S
-
+
(* Allow references to syntaxically inexistent variables (i.e., if applied on an inductive) *)
val allow_anonymous_refs : bool ref
(* Generic call to the interpreter from rawconstr 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
-
- val understand_tcc_evars :
- evar_defs ref -> env -> typing_constraint -> rawconstr -> constr
+
+ val understand_tcc_evars : ?fail_evar:bool -> ?resolve_classes:bool ->
+ evar_map ref -> env -> typing_constraint -> rawconstr -> 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.
-
- In [understand_ltac sigma env ltac_env constraint c],
-
+
+ In [understand_ltac expand_evars sigma env ltac_env constraint c],
+
+ expand_evars : expand inferred evars by their value if any
sigma : initial set of existential variables (typically dependent subgoals)
ltac_env : partial substitution of variables (used for the tactic language)
- constraint : tell if interpreted as a possibly constrained term or a type
+ constraint : tell if interpreted as a possibly constrained term or a type
*)
-
+
val understand_ltac :
- evar_map -> env -> var_map * unbound_ltac_var_map ->
- typing_constraint -> rawconstr -> evar_defs * constr
-
+ bool -> evar_map -> env -> ltac_var_map ->
+ typing_constraint -> rawconstr -> evar_map * constr
+
(* Standard call to get a constr from a rawconstr, resolving implicit args *)
-
+
val understand : evar_map -> env -> ?expected_type:Term.types ->
rawconstr -> constr
-
+
(* Idem but the rawconstr is intended to be a type *)
-
+
val understand_type : evar_map -> env -> rawconstr -> constr
-
+
(* A generalization of the two previous case *)
-
- val understand_gen : typing_constraint -> evar_map -> env ->
+
+ val understand_gen : typing_constraint -> evar_map -> env ->
rawconstr -> constr
-
+
(* Idem but returns the judgment of the understood term *)
-
+
val understand_judgment : evar_map -> env -> rawconstr -> unsafe_judgment
(* Idem but do not fail on unresolved evars *)
- val understand_judgment_tcc : evar_defs ref -> env -> rawconstr -> unsafe_judgment
+ val understand_judgment_tcc : evar_map ref -> env -> rawconstr -> unsafe_judgment
(*i*)
(* Internal of Pretyping...
*)
- val pretype :
- type_constraint -> env -> evar_defs ref ->
- var_map * (identifier * identifier option) list ->
- rawconstr -> unsafe_judgment
-
- val pretype_type :
- val_constraint -> env -> evar_defs ref ->
- var_map * (identifier * identifier option) list ->
- rawconstr -> unsafe_type_judgment
+ val pretype :
+ type_constraint -> env -> evar_map ref ->
+ ltac_var_map -> rawconstr -> unsafe_judgment
+
+ val pretype_type :
+ val_constraint -> env -> evar_map ref ->
+ ltac_var_map -> rawconstr -> unsafe_type_judgment
val pretype_gen :
- evar_defs ref -> env ->
- var_map * (identifier * identifier option) list ->
- typing_constraint -> rawconstr -> constr
+ bool -> bool -> bool -> evar_map ref -> env ->
+ ltac_var_map -> typing_constraint -> rawconstr -> constr
(*i*)
-
+
end
module Pretyping_F (C : Coercion.S) : S
module Default : S
(* To embed constr in rawconstr *)
-
+
val constr_in : constr -> Dyn.t
val constr_out : Dyn.t -> constr
-val interp_sort : rawsort -> sorts
+val interp_sort : rawsort -> sorts
val interp_elimination_sort : rawsort -> sorts_family
diff --git a/pretyping/pretyping.mllib b/pretyping/pretyping.mllib
new file mode 100644
index 00000000..cea33c1e
--- /dev/null
+++ b/pretyping/pretyping.mllib
@@ -0,0 +1,29 @@
+Termops
+Evd
+Reductionops
+Vnorm
+Namegen
+Inductiveops
+Retyping
+Cbv
+Pretype_errors
+Evarutil
+Term_dnet
+Recordops
+Evarconv
+Typing
+Rawterm
+Pattern
+Matching
+Tacred
+Typeclasses_errors
+Typeclasses
+Classops
+Coercion
+Unification
+Clenv
+Detyping
+Indrec
+Cases
+Pretyping
+
diff --git a/pretyping/rawterm.ml b/pretyping/rawterm.ml
index 30b62ea8..ba523402 100644
--- a/pretyping/rawterm.ml
+++ b/pretyping/rawterm.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: rawterm.ml 11576 2008-11-10 19:13:15Z msozeau $ *)
+(* $Id$ *)
(*i*)
open Util
@@ -42,7 +42,7 @@ type quantified_hypothesis = AnonHyp of int | NamedHyp of identifier
type 'a explicit_bindings = (loc * quantified_hypothesis * 'a) list
-type 'a bindings =
+type 'a bindings =
| ImplicitBindings of 'a list
| ExplicitBindings of 'a explicit_bindings
| NoBindings
@@ -53,7 +53,7 @@ type 'a cast_type =
| CastConv of cast_kind * 'a
| CastCoerce (* Cast to a base type (eg, an underlying inductive type) *)
-type rawconstr =
+type rawconstr =
| RRef of (loc * global_reference)
| RVar of (loc * identifier)
| REvar of loc * existential_key * rawconstr list option
@@ -63,7 +63,7 @@ type 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) *
+ | 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 *
@@ -75,7 +75,7 @@ type rawconstr =
and rawdecl = name * binding_kind * rawconstr option * rawconstr
-and fix_recursion_order = RStructRec | RWfRec of rawconstr | RMeasureRec of 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)
@@ -97,37 +97,55 @@ let cases_predicate_names tml =
| (tm,(na,None)) -> [na]
| (tm,(na,Some (_,_,_,nal))) -> na::nal) tml)
-(*i - if PRec (_, names, arities, bodies) is in env then arities are
- typed in env too and bodies are typed in env enriched by the
- arities incrementally lifted
-
- [On pourrait plutot mettre les arités aves le type qu'elles auront
- dans le contexte servant à typer les body ???]
-
- - boolean in POldCase means it is recursive
-i*)
-let map_rawdecl f (na,k,obd,ty) = (na,k,Option.map f obd,f ty)
-
-let map_rawconstr f = function
- | RVar (loc,id) -> RVar (loc,id)
- | RApp (loc,g,args) -> RApp (loc,f g, List.map f args)
- | RLambda (loc,na,bk,ty,c) -> RLambda (loc,na,bk,f ty,f c)
- | RProd (loc,na,bk,ty,c) -> RProd (loc,na,bk,f ty,f c)
- | RLetIn (loc,na,b,c) -> RLetIn (loc,na,f b,f c)
+let map_rawdecl_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 comp1 = f g in
+ let comp2 = Util.list_map_left f args in
+ RApp (loc,comp1,comp2)
+ | RLambda (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) ->
+ let comp1 = f ty in
+ let comp2 = f c in
+ RProd (loc,na,bk,comp1,comp2)
+ | RLetIn (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) ->
- RCases (loc,sty,Option.map f rtntypopt,
- List.map (fun (tm,x) -> (f tm,x)) tml,
- List.map (fun (loc,idl,p,c) -> (loc,idl,p,f c)) 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) ->
- RLetTuple (loc,nal,(na,Option.map f po),f b,f 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) ->
- RIf (loc,f c,(na,Option.map f po),f b1,f 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) ->
- RRec (loc,fk,idl,Array.map (List.map (map_rawdecl f)) bl,
- Array.map f tyl,Array.map f bv)
- | RCast (loc,c,k) -> RCast (loc,f c, match k with CastConv (k,t) -> CastConv (k, f t) | x -> x)
- | (RSort _ | RHole _ | RRef _ | REvar _ | RPatVar _ | RDynamic _) as x -> x
-
+ let comp1 = Array.map (Util.list_map_left (map_rawdecl_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) ->
+ 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
+
+let map_rawconstr = map_rawconstr_left_to_right
(*
let name_app f e = function
@@ -178,10 +196,10 @@ let occur_rawconstr id =
(occur_option rtntypopt)
or (List.exists (fun (tm,_) -> occur tm) tml)
or (List.exists occur_pattern pl)
- | RLetTuple (loc,nal,rtntyp,b,c) ->
+ | RLetTuple (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) ->
+ | RIf (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) ->
not (array_for_all4 (fun fid bl ty bd ->
@@ -207,67 +225,67 @@ let occur_rawconstr id =
in occur
-let add_name_to_ids set na =
- match na with
- | Anonymous -> set
- | Name id -> Idset.add id set
+let add_name_to_ids set na =
+ match na with
+ | Anonymous -> set
+ | Name id -> Idset.add id set
let free_rawvars =
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) ->
- let vs' = vars bounded vs ty in
- let bounded' = add_name_to_ids bounded na in
+ | RLambda (loc,na,_,ty,c) | RProd (loc,na,_,ty,c) | RLetIn (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) ->
- let vs1 = vars_option bounded vs rtntypopt in
- let vs2 = List.fold_left (fun vs (tm,_) -> vars bounded vs tm) vs1 tml in
+ 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) ->
- let vs1 = vars_return_type bounded vs rtntyp in
- let vs2 = vars bounded vs1 b in
+ 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) ->
- let vs1 = vars_return_type bounded vs rtntyp in
- let vs2 = vars bounded vs1 c in
- let vs3 = vars bounded vs2 b1 in
+ | RIf (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) ->
- let bounded' = Array.fold_right Idset.add idl bounded in
- let vars_fix i vs fid =
- let vs1,bounded1 =
- List.fold_left
- (fun (vs,bounded) (na,k,bbd,bty) ->
- let vs' = vars_option bounded vs bbd in
+ let bounded' = Array.fold_right Idset.add idl bounded in
+ let vars_fix i vs fid =
+ let vs1,bounded1 =
+ List.fold_left
+ (fun (vs,bounded) (na,k,bbd,bty) ->
+ let vs' = vars_option bounded vs bbd in
let vs'' = vars bounded vs' bty in
- let bounded' = add_name_to_ids bounded na in
+ let bounded' = add_name_to_ids bounded na in
(vs'',bounded')
)
(vs,bounded')
bl.(i)
in
- let vs2 = vars bounded1 vs1 tyl.(i) in
+ let vs2 = vars bounded1 vs1 tyl.(i) in
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
+ | RCast (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
- and vars_pattern bounded vs (loc,idl,p,c) =
- let bounded' = List.fold_right Idset.add idl bounded in
+ and vars_pattern bounded vs (loc,idl,p,c) =
+ let bounded' = List.fold_right Idset.add idl bounded in
vars bounded' vs c
and vars_option bounded vs = function None -> vs | Some p -> vars bounded vs p
- and vars_return_type bounded vs (na,tyopt) =
- let bounded' = add_name_to_ids bounded na in
+ and vars_return_type bounded vs (na,tyopt) =
+ let bounded' = add_name_to_ids bounded na in
vars_option bounded' vs tyopt
- in
- fun rt ->
- let vs = vars Idset.empty Idset.empty rt in
+ in
+ fun rt ->
+ let vs = vars Idset.empty Idset.empty rt in
Idset.elements vs
@@ -344,10 +362,10 @@ let no_occurrences_expr = (true,[])
type 'a with_occurrences = occurrences_expr * 'a
-type ('a,'b) red_expr_gen =
+type ('a,'b,'c) red_expr_gen =
| Red of bool
| Hnf
- | Simpl of 'a with_occurrences option
+ | Simpl of 'c with_occurrences option
| Cbv of 'b raw_red_flag
| Lazy of 'b raw_red_flag
| Unfold of 'b with_occurrences list
@@ -356,8 +374,8 @@ type ('a,'b) red_expr_gen =
| ExtraRedExpr of string
| CbvVm
-type ('a,'b) may_eval =
+type ('a,'b,'c) may_eval =
| ConstrTerm of 'a
- | ConstrEval of ('a,'b) red_expr_gen * 'a
+ | ConstrEval of ('a,'b,'c) red_expr_gen * 'a
| ConstrContext of (loc * identifier) * 'a
| ConstrTypeOf of 'a
diff --git a/pretyping/rawterm.mli b/pretyping/rawterm.mli
index eecee8b0..b2b70bc9 100644
--- a/pretyping/rawterm.mli
+++ b/pretyping/rawterm.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: rawterm.mli 11282 2008-07-28 11:51:53Z msozeau $ i*)
+(*i $Id$ i*)
(*i*)
open Util
@@ -46,7 +46,7 @@ type quantified_hypothesis = AnonHyp of int | NamedHyp of identifier
type 'a explicit_bindings = (loc * quantified_hypothesis * 'a) list
-type 'a bindings =
+type 'a bindings =
| ImplicitBindings of 'a list
| ExplicitBindings of 'a explicit_bindings
| NoBindings
@@ -57,7 +57,7 @@ type 'a cast_type =
| CastConv of cast_kind * 'a
| CastCoerce (* Cast to a base type (eg, an underlying inductive type) *)
-type rawconstr =
+type rawconstr =
| RRef of (loc * global_reference)
| RVar of (loc * identifier)
| REvar of loc * existential_key * rawconstr list option
@@ -67,7 +67,7 @@ type 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) *
+ | 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 *
@@ -79,7 +79,7 @@ type rawconstr =
and rawdecl = name * binding_kind * rawconstr option * rawconstr
-and fix_recursion_order = RStructRec | RWfRec of rawconstr | RMeasureRec of 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)
@@ -98,21 +98,14 @@ and cases_clauses = cases_clause list
val cases_predicate_names : tomatch_tuples -> name list
-(*i - if PRec (_, names, arities, bodies) is in env then arities are
- typed in env too and bodies are typed in env enriched by the
- arities incrementally lifted
-
- [On pourrait plutot mettre les arités aves le type qu'elles auront
- dans le contexte servant à typer les body ???]
-
- - boolean in POldCase means it is recursive
- - option in PHole tell if the "?" was apparent or has been implicitely added
-i*)
-
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 ->
+val map_rawconstr_with_binders_loc : loc ->
(identifier -> 'a -> identifier * 'a) ->
('a -> rawconstr -> rawconstr) -> 'a -> rawconstr -> rawconstr
i*)
@@ -155,10 +148,10 @@ val no_occurrences_expr : occurrences_expr
type 'a with_occurrences = occurrences_expr * 'a
-type ('a,'b) red_expr_gen =
+type ('a,'b,'c) red_expr_gen =
| Red of bool
| Hnf
- | Simpl of 'a with_occurrences option
+ | Simpl of 'c with_occurrences option
| Cbv of 'b raw_red_flag
| Lazy of 'b raw_red_flag
| Unfold of 'b with_occurrences list
@@ -167,8 +160,8 @@ type ('a,'b) red_expr_gen =
| ExtraRedExpr of string
| CbvVm
-type ('a,'b) may_eval =
+type ('a,'b,'c) may_eval =
| ConstrTerm of 'a
- | ConstrEval of ('a,'b) red_expr_gen * '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 711f332e..6c903238 100644
--- a/pretyping/recordops.ml
+++ b/pretyping/recordops.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: recordops.ml 12187 2009-06-13 19:36:59Z msozeau $ *)
+(* $Id$ *)
open Util
open Pp
@@ -18,7 +18,6 @@ open Termops
open Typeops
open Libobject
open Library
-open Classops
open Mod_subst
open Reductionops
@@ -32,7 +31,7 @@ open Reductionops
projection ou bien une fonction constante (associée à un LetIn) *)
type struc_typ = {
- s_CONST : constructor;
+ s_CONST : constructor;
s_EXPECTEDPARAM : int;
s_PROJKIND : (name * bool) list;
s_PROJ : constant option list }
@@ -45,19 +44,19 @@ let load_structure i (_,(ind,id,kl,projs)) =
let struc =
{ s_CONST = id; s_EXPECTEDPARAM = n; s_PROJ = projs; s_PROJKIND = kl } in
structure_table := Indmap.add ind struc !structure_table;
- projection_table :=
+ projection_table :=
List.fold_right (Option.fold_right (fun proj -> Cmap.add proj struc))
projs !projection_table
let cache_structure o =
load_structure 1 o
-let subst_structure (_,subst,((kn,i),id,kl,projs as obj)) =
- let kn' = subst_kn subst kn in
+let subst_structure (subst,((kn,i),id,kl,projs as obj)) =
+ let kn' = subst_ind subst kn in
let projs' =
(* invariant: struc.s_PROJ is an evaluable reference. Thus we can take *)
(* the first component of subst_con. *)
- list_smartmap
+ list_smartmap
(Option.smartmap (fun kn -> fst (subst_con subst kn)))
projs
in
@@ -65,7 +64,7 @@ let subst_structure (_,subst,((kn,i),id,kl,projs as obj)) =
if projs' == projs && kn' == kn && id' == id then obj else
((kn',i),id',kl,projs')
-let discharge_constructor (ind, n) =
+let discharge_constructor (ind, n) =
(Lib.discharge_inductive ind, n)
let discharge_structure (_,(ind,id,kl,projs)) =
@@ -73,15 +72,14 @@ let discharge_structure (_,(ind,id,kl,projs)) =
List.map (Option.map Lib.discharge_con) projs)
let (inStruc,outStruc) =
- declare_object {(default_object "STRUCTURE") with
+ declare_object {(default_object "STRUCTURE") with
cache_function = cache_structure;
load_function = load_structure;
subst_function = subst_structure;
- classify_function = (fun (_,x) -> Substitute x);
- discharge_function = discharge_structure;
- export_function = (function x -> Some x) }
+ classify_function = (fun x -> Substitute x);
+ discharge_function = discharge_structure }
-let declare_structure (s,c,kl,pl) =
+let declare_structure (s,c,kl,pl) =
Lib.add_anonymous_leaf (inStruc (s,c,kl,pl))
let lookup_structure indsp = Indmap.find indsp !structure_table
@@ -96,6 +94,55 @@ let find_projection = function
| ConstRef cst -> Cmap.find cst !projection_table
| _ -> raise Not_found
+(* Management of a field store : each field + argument of the inferred
+ * records are stored in a discrimination tree *)
+
+let subst_id s (gr,ev,evm) =
+ (fst(subst_global s gr),ev,Evd.subst_evar_map s evm)
+
+module MethodsDnet : Term_dnet.S
+ with type ident = global_reference * Evd.evar * Evd.evar_map
+ = Term_dnet.Make
+ (struct
+ type t = global_reference * Evd.evar * Evd.evar_map
+ let compare = Pervasives.compare
+ let subst = subst_id
+ let constr_of (_,ev,evm) = Evd.evar_concl (Evd.find evm ev)
+ end)
+ (struct
+ let reduce c = Reductionops.head_unfold_under_prod
+ Names.full_transparent_state (Global.env()) Evd.empty c
+ let direction = true
+ end)
+
+let meth_dnet = ref MethodsDnet.empty
+
+open Summary
+
+let _ =
+ declare_summary "record-methods-state"
+ { freeze_function = (fun () -> !meth_dnet);
+ unfreeze_function = (fun m -> meth_dnet := m);
+ init_function = (fun () -> meth_dnet := MethodsDnet.empty) }
+
+open Libobject
+
+let load_method (_,(ty,id)) =
+ meth_dnet := MethodsDnet.add ty id !meth_dnet
+
+let (in_method,out_method) =
+ declare_object
+ { (default_object "RECMETHODS") with
+ load_function = (fun _ -> load_method);
+ cache_function = load_method;
+ subst_function = (fun (s,(ty,id)) -> Mod_subst.subst_mps s ty,subst_id s id);
+ classify_function = (fun x -> Substitute x)
+ }
+
+let methods_matching c = MethodsDnet.search_pattern !meth_dnet c
+
+let declare_method cons ev sign =
+ Lib.add_anonymous_leaf (in_method ((Evd.evar_concl (Evd.find sign ev)),(cons,ev,sign)))
(************************************************************************)
(*s A canonical structure declares "canonical" conversion hints between *)
@@ -138,7 +185,7 @@ type cs_pattern =
let object_table = ref (Refmap.empty : (cs_pattern * obj_typ) list Refmap.t)
-let canonical_projections () =
+let canonical_projections () =
Refmap.fold (fun x -> List.fold_right (fun (y,c) acc -> ((x,y),c)::acc))
!object_table []
@@ -148,28 +195,28 @@ let keep_true_projections projs kinds =
let cs_pattern_of_constr t =
match kind_of_term t with
- App (f,vargs) ->
- begin
+ App (f,vargs) ->
+ begin
try Const_cs (global_of_constr f) , -1, Array.to_list vargs with
- _ -> raise Not_found
- end
+ _ -> 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]
| Sort s -> Sort_cs (family_of_sort s), -1, []
- | _ ->
- begin
+ | _ ->
+ begin
try Const_cs (global_of_constr t) , -1, [] with
- _ -> raise Not_found
- end
+ _ -> raise Not_found
+ end
(* Intended to always succeed *)
let compute_canonical_projections (con,ind) =
let v = mkConst con in
let c = Environ.constant_value (Global.env()) con in
- let lt,t = Reductionops.splay_lambda (Global.env()) Evd.empty c in
+ let lt,t = Reductionops.splay_lam (Global.env()) Evd.empty c in
let lt = List.rev (List.map snd lt) in
let args = snd (decompose_app t) in
- let { s_EXPECTEDPARAM = p; s_PROJ = lpj; s_PROJKIND = kl } =
+ let { s_EXPECTEDPARAM = p; s_PROJ = lpj; s_PROJKIND = kl } =
lookup_structure ind in
let params, projs = list_chop p args in
let lpj = keep_true_projections lpj kl in
@@ -180,31 +227,55 @@ let compute_canonical_projections (con,ind) =
match spopt with
| Some proji_sp ->
begin
- try
+ try
let patt, n , args = cs_pattern_of_constr t in
((ConstRef proji_sp, patt, n, args) :: l)
- with Not_found -> l
+ with Not_found ->
+ if Flags.is_verbose () then
+ (let con_pp = Nametab.pr_global_env Idset.empty (ConstRef con)
+ and proji_sp_pp = Nametab.pr_global_env Idset.empty (ConstRef proji_sp) in
+ msg_warning (str "No global reference exists for projection value"
+ ++ print_constr t ++ str " in instance "
+ ++ con_pp ++ str " of " ++ proji_sp_pp ++ str ", ignoring it."));
+ l
end
| _ -> l)
[] lps in
List.map (fun (refi,c,inj,argj) ->
(refi,c),
- {o_DEF=v; o_INJ=inj; o_TABS=lt;
+ {o_DEF=v; o_INJ=inj; o_TABS=lt;
o_TPARAMS=params; o_NPARAMS=List.length params; o_TCOMPS=argj})
comp
+let pr_cs_pattern = function
+ Const_cs c -> Nametab.pr_global_env Idset.empty c
+ | Prod_cs -> str "_ -> _"
+ | Default_cs -> str "_"
+ | Sort_cs s -> Termops.pr_sort_family s
+
let open_canonical_structure i (_,o) =
if i=1 then
let lo = compute_canonical_projections o in
List.iter (fun ((proj,cs_pat),s) ->
let l = try Refmap.find proj !object_table with Not_found -> [] in
- if not (List.mem_assoc cs_pat l) then
- object_table := Refmap.add proj ((cs_pat,s)::l) !object_table) lo
+ let ocs = try Some (List.assoc cs_pat l)
+ with Not_found -> None
+ in match ocs with
+ | None -> object_table := Refmap.add proj ((cs_pat,s)::l) !object_table;
+ | Some cs ->
+ if Flags.is_verbose () then
+ let old_can_s = (Termops.print_constr cs.o_DEF)
+ and new_can_s = (Termops.print_constr s.o_DEF) in
+ let prj = (Nametab.pr_global_env Idset.empty proj)
+ and hd_val = (pr_cs_pattern cs_pat) in
+ msg_warning (str "Ignoring canonical projection to " ++ hd_val
+ ++ str " by " ++ prj ++ str " in "
+ ++ new_can_s ++ str ": redundant with " ++ old_can_s)) lo
let cache_canonical_structure o =
open_canonical_structure 1 o
-let subst_canonical_structure (_,subst,(cst,ind as obj)) =
+let subst_canonical_structure (subst,(cst,ind as obj)) =
(* invariant: cst is an evaluable reference. Thus we can take *)
(* the first component of subst_con. *)
let cst' = fst (subst_con subst cst) in
@@ -215,13 +286,12 @@ let discharge_canonical_structure (_,(cst,ind)) =
Some (Lib.discharge_con cst,Lib.discharge_inductive ind)
let (inCanonStruc,outCanonStruct) =
- declare_object {(default_object "CANONICAL-STRUCTURE") with
+ declare_object {(default_object "CANONICAL-STRUCTURE") with
open_function = open_canonical_structure;
cache_function = cache_canonical_structure;
subst_function = subst_canonical_structure;
- classify_function = (fun (_,x) -> Substitute x);
- discharge_function = discharge_canonical_structure;
- export_function = (function x -> Some x) }
+ classify_function = (fun x -> Substitute x);
+ discharge_function = discharge_canonical_structure }
let add_canonical_structure x = Lib.add_anonymous_leaf (inCanonStruc x)
@@ -229,7 +299,7 @@ let add_canonical_structure x = Lib.add_anonymous_leaf (inCanonStruc x)
let error_not_structure ref =
errorlabstrm "object_declare"
- (Nameops.pr_id (id_of_global ref) ++ str" is not a structure object.")
+ (Nameops.pr_id (basename_of_global ref) ++ str" is not a structure object.")
let check_and_decompose_canonical_structure ref =
let sp = match ref with ConstRef sp -> sp | _ -> error_not_structure ref in
@@ -237,7 +307,7 @@ let check_and_decompose_canonical_structure ref =
let vc = match Environ.constant_opt_value env sp with
| Some vc -> vc
| None -> error_not_structure ref in
- let body = snd (splay_lambda (Global.env()) Evd.empty vc) in
+ let body = snd (splay_lam (Global.env()) Evd.empty vc) in
let f,args = match kind_of_term body with
| App (f,args) -> f,args
| _ -> error_not_structure ref in
@@ -259,16 +329,16 @@ let lookup_canonical_conversion (proj,pat) =
List.assoc pat (Refmap.find proj !object_table)
let is_open_canonical_projection sigma (c,args) =
- try
+ try
let l = Refmap.find (global_of_constr c) !object_table in
let n = (snd (List.hd l)).o_NPARAMS in
- try isEvar (whd_evar sigma (List.nth args n)) with Failure _ -> false
+ try isEvar_or_Meta (whd_evar sigma (List.nth args n)) with Failure _ -> false
with Not_found -> false
let freeze () =
!structure_table, !projection_table, !object_table
-let unfreeze (s,p,o) =
+let unfreeze (s,p,o) =
structure_table := s; projection_table := p; object_table := o
let init () =
@@ -277,10 +347,8 @@ let init () =
let _ = init()
-let _ =
+let _ =
Summary.declare_summary "objdefs"
{ Summary.freeze_function = freeze;
Summary.unfreeze_function = unfreeze;
- Summary.init_function = init;
- Summary.survive_module = false;
- Summary.survive_section = false }
+ Summary.init_function = init }
diff --git a/pretyping/recordops.mli b/pretyping/recordops.mli
index ea960aa9..8fc430ae 100755
--- a/pretyping/recordops.mli
+++ b/pretyping/recordops.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: recordops.mli 11576 2008-11-10 19:13:15Z msozeau $ i*)
+(*i $Id$ i*)
(*i*)
open Names
@@ -22,14 +22,19 @@ open Library
constructor (the name of which defaults to Build_S) *)
type struc_typ = {
- s_CONST : constructor;
+ s_CONST : constructor;
s_EXPECTEDPARAM : int;
s_PROJKIND : (name * bool) list;
s_PROJ : constant option list }
-val declare_structure :
+val declare_structure :
inductive * constructor * (name * bool) list * constant option list -> unit
+(* [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
inductive path [isp] if it corresponds to a structure, otherwise
it fails with [Not_found] *)
@@ -41,13 +46,22 @@ val find_projection_nparams : global_reference -> int
(* raise [Not_found] if not a projection *)
val find_projection : global_reference -> struc_typ
+(* we keep an index (dnet) of record's arguments + fields
+ (=methods). Here is how to declare them: *)
+val declare_method :
+ global_reference -> Evd.evar -> Evd.evar_map -> unit
+ (* and here is how to search for methods matched by a given term: *)
+val methods_matching : constr ->
+ ((global_reference*Evd.evar*Evd.evar_map) *
+ (constr*existential_key)*Termops.subst) list
+
(*s 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
- | Prod_cs
+ | Prod_cs
| Sort_cs of sorts_family
| Default_cs
@@ -60,10 +74,11 @@ type obj_typ = {
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
+
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
-val canonical_projections : unit ->
+val canonical_projections : unit ->
((global_reference * cs_pattern) * obj_typ) list
diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml
index 33928f67..1a69b633 100644
--- a/pretyping/reductionops.ml
+++ b/pretyping/reductionops.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: reductionops.ml 11897 2009-02-09 19:28:02Z barras $ *)
+(* $Id$ *)
open Pp
open Util
@@ -25,7 +25,7 @@ exception Elimconst
(**********************************************************************)
-(* The type of (machine) stacks (= lambda-bar-calculus' contexts) *)
+(* The type of (machine) stacks (= lambda-bar-calculus' contexts) *)
type 'a stack_member =
| Zapp of 'a list
@@ -80,12 +80,12 @@ let rec list_of_stack = function
let rec app_stack = function
| f, [] -> f
| f, (Zapp [] :: s) -> app_stack (f, s)
- | f, (Zapp args :: s) ->
+ | f, (Zapp args :: s) ->
app_stack (applist (f, args), s)
| _ -> assert false
let rec stack_assign s p c = match s with
| Zapp args :: s ->
- let q = List.length args in
+ let q = List.length args in
if p >= q then
Zapp args :: stack_assign s (p-q) c
else
@@ -109,20 +109,20 @@ let rec stack_nth s p = match s with
| _ -> raise Not_found
(**************************************************************)
-(* The type of (machine) states (= lambda-bar-calculus' cuts) *)
+(* The type of (machine) states (= lambda-bar-calculus' cuts) *)
type state = constr * constr stack
type contextual_reduction_function = env -> evar_map -> constr -> constr
type reduction_function = contextual_reduction_function
type local_reduction_function = evar_map -> constr -> constr
-type contextual_stack_reduction_function =
+type contextual_stack_reduction_function =
env -> evar_map -> constr -> constr * constr list
type stack_reduction_function = contextual_stack_reduction_function
type local_stack_reduction_function =
evar_map -> constr -> constr * constr list
-type contextual_state_reduction_function =
+type contextual_state_reduction_function =
env -> evar_map -> state -> state
type state_reduction_function = contextual_state_reduction_function
type local_state_reduction_function = evar_map -> state -> state
@@ -135,36 +135,40 @@ let safe_evar_value sigma ev =
try Some (Evd.existential_value sigma ev)
with NotInstantiatedEvar | Not_found -> None
-let rec whd_state sigma (x, stack as s) =
+let rec whd_app_state sigma (x, stack as s) =
match kind_of_term x with
- | App (f,cl) -> whd_state sigma (f, append_stack cl stack)
- | Cast (c,_,_) -> whd_state sigma (c, stack)
+ | App (f,cl) -> whd_app_state sigma (f, append_stack cl stack)
+ | Cast (c,_,_) -> whd_app_state sigma (c, stack)
| Evar ev ->
(match safe_evar_value sigma ev with
- Some c -> whd_state sigma (c,stack)
+ Some c -> whd_app_state sigma (c,stack)
| _ -> s)
| _ -> s
+let safe_meta_value sigma ev =
+ try Some (Evd.meta_value sigma ev)
+ with Not_found -> None
+
let appterm_of_stack (f,s) = (f,list_of_stack s)
let whd_stack sigma x =
- appterm_of_stack (whd_state sigma (x, empty_stack))
+ 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 strong whdfun env sigma t =
let rec strongrec env t =
map_constr_with_full_binders push_rel strongrec env (whdfun env sigma t) in
strongrec env t
-let local_strong whdfun sigma =
+let local_strong whdfun sigma =
let rec strongrec t = map_constr strongrec (whdfun sigma t) in
strongrec
-let rec strong_prodspine redfun sigma c =
+let rec strong_prodspine redfun sigma c =
let x = redfun sigma c in
match kind_of_term x with
| Prod (na,a,b) -> mkProd (na,a,strong_prodspine redfun sigma b)
@@ -193,33 +197,13 @@ module type RedFlagsSig = sig
val red_zeta : flags -> bool
end
-(* Naive Implementation
-module RedFlags = (struct
- type flag = BETA | DELTA | EVAR | IOTA | ZETA | ETA
- type flags = flag list
- let fbeta = BETA
- let fdelta = DELTA
- let fevar = EVAR
- let fiota = IOTA
- let fzeta = ZETA
- let feta = ETA
- let mkflags l = l
- let red_beta = List.mem BETA
- let red_delta = List.mem DELTA
- let red_evar = List.mem EVAR
- let red_eta = List.mem ETA
- let red_iota = List.mem IOTA
- let red_zeta = List.mem ZETA
-end : RedFlagsSig)
-*)
-
(* Compact Implementation *)
module RedFlags = (struct
type flag = int
type flags = int
let fbeta = 1
let fdelta = 2
- let feta = 8
+ let feta = 8
let fiota = 16
let fzeta = 32
let mkflags = List.fold_left (lor) 0
@@ -235,6 +219,7 @@ open RedFlags
(* Local *)
let beta = mkflags [fbeta]
let eta = mkflags [feta]
+let zeta = mkflags [fzeta]
let betaiota = mkflags [fiota; fbeta]
let betaiotazeta = mkflags [fiota; fbeta;fzeta]
@@ -298,7 +283,7 @@ let contract_fix ((recindices,bodynum),(types,names,bodies as typedbodies)) =
let fix_recarg ((recindices,bodynum),_) stack =
assert (0 <= bodynum & bodynum < Array.length recindices);
let recargnum = Array.get recindices bodynum in
- try
+ try
Some (recargnum, stack_nth stack recargnum)
with Not_found ->
None
@@ -319,12 +304,12 @@ let reduce_fix whdfun sigma fix stack =
(* Y avait un commentaire pour whd_betadeltaiota :
- NB : Cette fonction alloue peu c'est l'appel
+ NB : Cette fonction alloue peu c'est l'appel
``let (c,cargs) = whfun (recarg, empty_stack)''
-------------------
qui coute cher *)
-let rec whd_state_gen flags env sigma =
+let rec whd_state_gen flags env sigma =
let rec whrec (x, stack as s) =
match kind_of_term x with
| Rel n when red_delta flags ->
@@ -339,6 +324,10 @@ let rec whd_state_gen flags env sigma =
(match safe_evar_value sigma ev with
| Some body -> whrec (body, stack)
| None -> s)
+ | Meta ev ->
+ (match safe_meta_value sigma ev with
+ | Some body -> whrec (body, stack)
+ | None -> s)
| Const const when red_delta flags ->
(match constant_opt_value env const with
| Some body -> whrec (body, stack)
@@ -373,19 +362,19 @@ let rec whd_state_gen flags env sigma =
whrec (reduce_mind_case
{mP=p; mconstr=c; mcargs=list_of_stack cargs;
mci=ci; mlf=lf}, stack)
- else
+ else
(mkCase (ci, p, app_stack (c,cargs), lf), stack)
-
+
| Fix fix when red_iota flags ->
(match reduce_fix (fun _ -> whrec) sigma fix stack with
| Reduced s' -> whrec s'
| NotReducible -> s)
| x -> s
- in
+ in
whrec
-let local_whd_state_gen flags sigma =
+let local_whd_state_gen flags sigma =
let rec whrec (x, stack as s) =
match kind_of_term x with
| LetIn (_,b,_,c) when red_zeta flags -> stacklam whrec [b] c stack
@@ -395,7 +384,7 @@ let local_whd_state_gen flags sigma =
(match decomp_stack stack with
| Some (a,m) when red_beta flags -> stacklam whrec [a] c m
| None when red_eta flags ->
- (match kind_of_term (app_stack (whrec (c, empty_stack))) with
+ (match kind_of_term (app_stack (whrec (c, empty_stack))) with
| App (f,cl) ->
let napp = Array.length cl in
if napp > 0 then
@@ -416,9 +405,9 @@ let local_whd_state_gen flags sigma =
whrec (reduce_mind_case
{mP=p; mconstr=c; mcargs=list_of_stack cargs;
mci=ci; mlf=lf}, stack)
- else
+ else
(mkCase (ci, p, app_stack (c,cargs), lf), stack)
-
+
| Fix fix when red_iota flags ->
(match reduce_fix (fun _ ->whrec) sigma fix stack with
| Reduced s' -> whrec s'
@@ -429,8 +418,13 @@ let local_whd_state_gen flags sigma =
Some c -> whrec (c,stack)
| None -> s)
+ | Meta ev ->
+ (match safe_meta_value sigma ev with
+ Some c -> whrec (c,stack)
+ | None -> s)
+
| x -> s
- in
+ in
whrec
@@ -471,7 +465,7 @@ let whd_betadelta env =
let whd_betadeltaeta_state e = whd_state_gen betadeltaeta e
let whd_betadeltaeta_stack env =
stack_red_of_state_red (whd_betadeltaeta_state env)
-let whd_betadeltaeta env =
+let whd_betadeltaeta env =
red_of_state_red (whd_betadeltaeta_state env)
(* 3. Iota reduction Functions *)
@@ -487,25 +481,29 @@ let whd_betaiotazeta = red_of_state_red whd_betaiotazeta_state
let whd_betadeltaiota_state e = whd_state_gen betadeltaiota e
let whd_betadeltaiota_stack env =
stack_red_of_state_red (whd_betadeltaiota_state env)
-let whd_betadeltaiota 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_betadeltaiotaeta_stack env =
stack_red_of_state_red (whd_betadeltaiotaeta_state env)
-let whd_betadeltaiotaeta 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_stack env =
stack_red_of_state_red (whd_betadeltaiota_nolet_state env)
-let whd_betadeltaiota_nolet env =
+let whd_betadeltaiota_nolet env =
red_of_state_red (whd_betadeltaiota_nolet_state env)
-(* 3. Eta reduction Functions *)
+(* 4. Eta reduction Functions *)
let whd_eta c = app_stack (local_whd_state_gen eta Evd.empty (c,empty_stack))
+(* 5. Zeta Reduction Functions *)
+
+let whd_zeta c = app_stack (local_whd_state_gen zeta Evd.empty (c,empty_stack))
+
(****************************************************************************)
(* Reduction Functions *)
(****************************************************************************)
@@ -517,8 +515,8 @@ let rec whd_evar sigma c =
(match safe_evar_value sigma ev with
Some c -> whd_evar sigma c
| None -> c)
- | Sort s when is_sort_variable sigma s -> whd_sort_variable sigma c
- | _ -> collapse_appl c
+ | Sort s -> whd_sort_variable sigma c
+ | _ -> c
let nf_evar =
local_strong whd_evar
@@ -537,53 +535,53 @@ let nf_betadeltaiota env sigma =
clos_norm_flags Closure.betadeltaiota env sigma
-(* Attention reduire un beta-redexe avec un argument qui n'est pas
+(* Attention reduire un beta-redexe avec un argument qui n'est pas
une variable, peut changer enormement le temps de conversion lors
du type checking :
(fun x => x + x) M
*)
-let rec whd_betaiota_preserving_vm_cast env sigma t =
- let rec stacklam_var subst t stack =
- match (decomp_stack stack,kind_of_term t) with
- | Some (h,stacktl), Lambda (_,_,c) ->
- begin match kind_of_term h with
- | Rel i when not (evaluable_rel i env) ->
- stacklam_var (h::subst) c stacktl
- | Var id when not (evaluable_named id env)->
- stacklam_var (h::subst) c stacktl
- | _ -> whrec (substl subst t, stack)
- end
- | _ -> whrec (substl subst t, stack)
- and whrec (x, stack as s) =
- match kind_of_term x with
- | Evar ev ->
- (match safe_evar_value sigma ev with
- | Some body -> whrec (body, stack)
- | None -> s)
- | Cast (c,VMcast,t) ->
- let c = app_stack (whrec (c,empty_stack)) in
- let t = app_stack (whrec (t,empty_stack)) in
- (mkCast(c,VMcast,t),stack)
- | Cast (c,DEFAULTcast,_) ->
- whrec (c, stack)
- | App (f,cl) -> whrec (f, append_stack cl stack)
- | Lambda (na,t,c) ->
- (match decomp_stack stack with
- | Some (a,m) -> stacklam_var [a] c m
- | _ -> s)
- | Case (ci,p,d,lf) ->
- let (c,cargs) = whrec (d, empty_stack) in
- if reducible_mind_case c then
- whrec (reduce_mind_case
- {mP=p; mconstr=c; mcargs=list_of_stack cargs;
- mci=ci; mlf=lf}, stack)
- else
- (mkCase (ci, p, app_stack (c,cargs), lf), stack)
- | x -> s
- in
+let rec whd_betaiota_preserving_vm_cast env sigma t =
+ let rec stacklam_var subst t stack =
+ match (decomp_stack stack,kind_of_term t) with
+ | Some (h,stacktl), Lambda (_,_,c) ->
+ begin match kind_of_term h with
+ | Rel i when not (evaluable_rel i env) ->
+ stacklam_var (h::subst) c stacktl
+ | Var id when not (evaluable_named id env)->
+ stacklam_var (h::subst) c stacktl
+ | _ -> whrec (substl subst t, stack)
+ end
+ | _ -> whrec (substl subst t, stack)
+ and whrec (x, stack as s) =
+ match kind_of_term x with
+ | Evar ev ->
+ (match safe_evar_value sigma ev with
+ | Some body -> whrec (body, stack)
+ | None -> s)
+ | Cast (c,VMcast,t) ->
+ let c = app_stack (whrec (c,empty_stack)) in
+ let t = app_stack (whrec (t,empty_stack)) in
+ (mkCast(c,VMcast,t),stack)
+ | Cast (c,DEFAULTcast,_) ->
+ whrec (c, stack)
+ | App (f,cl) -> whrec (f, append_stack cl stack)
+ | Lambda (na,t,c) ->
+ (match decomp_stack stack with
+ | Some (a,m) -> stacklam_var [a] c m
+ | _ -> s)
+ | Case (ci,p,d,lf) ->
+ let (c,cargs) = whrec (d, empty_stack) in
+ if reducible_mind_case c then
+ whrec (reduce_mind_case
+ {mP=p; mconstr=c; mcargs=list_of_stack cargs;
+ mci=ci; mlf=lf}, stack)
+ else
+ (mkCase (ci, p, app_stack (c,cargs), lf), stack)
+ | x -> s
+ in
app_stack (whrec (t,empty_stack))
-let nf_betaiota_preserving_vm_cast =
+let nf_betaiota_preserving_vm_cast =
strong whd_betaiota_preserving_vm_cast
(* lazy weak head reduction functions *)
@@ -631,26 +629,26 @@ 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
with NotConvertible -> false
-let is_trans_conv env sigma = test_trans_conversion Reduction.trans_conv env sigma
-let is_trans_conv_leq env sigma = test_trans_conversion Reduction.trans_conv_leq env sigma
+let is_trans_conv reds env sigma = test_trans_conversion Reduction.trans_conv reds env sigma
+let is_trans_conv_leq reds env sigma = test_trans_conversion Reduction.trans_conv_leq reds env sigma
let is_trans_fconv = function | CONV -> is_trans_conv | CUMUL -> is_trans_conv_leq
(********************************************************************)
(* Special-Purpose Reduction *)
(********************************************************************)
-let whd_meta metamap c = match kind_of_term c with
- | Meta p -> (try List.assoc p metamap with Not_found -> c)
+let whd_meta sigma c = match kind_of_term c with
+ | Meta p -> (try meta_value sigma p with Not_found -> c)
| _ -> c
(* Try to replace all metas. Does not replace metas in the metas' values
* Differs from (strong whd_meta). *)
-let plain_instance s c =
+let plain_instance s c =
let rec irec n u = match kind_of_term u with
| Meta p -> (try lift n (List.assoc p s) with Not_found -> u)
| App (f,l) when isCast f ->
let (f,_,t) = destCast f in
- let l' = Array.map (irec n) l in
+ let l' = Array.map (irec n) l in
(match kind_of_term f with
| Meta p ->
(* Don't flatten application nodes: this is used to extract a
@@ -658,21 +656,21 @@ let plain_instance s c =
of the proof-tree *)
(try let g = List.assoc p s in
match kind_of_term g with
- | App _ ->
+ | App _ ->
let h = id_of_string "H" in
mkLetIn (Name h,g,t,mkApp(mkRel 1,Array.map (lift 1) l'))
| _ -> mkApp (g,l')
with Not_found -> mkApp (f,l'))
- | _ -> mkApp (irec n f,l'))
+ | _ -> mkApp (irec n f,l'))
| Cast (m,_,_) when isMeta m ->
(try lift n (List.assoc (destMeta m) s) with Not_found -> u)
| _ ->
map_constr_with_binders succ irec n u
- in
+ in
if s = [] then c else irec 0 c
(* [instance] is used for [res_pf]; the call to [local_strong whd_betaiota]
- has (unfortunately) different subtle side effects:
+ has (unfortunately) different subtle side effects:
- ** Order of subgoals **
If the lemma is a case analysis with parameters, it will move the
@@ -689,7 +687,7 @@ let plain_instance s c =
been contracted). A goal to rewrite may then fail or succeed
differently.
- - ** Naming of hypotheses **
+ - ** Naming of hypotheses **
If a lemma is a function of the form "fun H:(forall a:A, P a)
=> .. F H .." where the expected type of H is "forall b:A, P b",
then, without reduction, the application of the lemma will
@@ -705,9 +703,9 @@ let plain_instance s c =
empty map).
*)
-let instance s c =
+let instance sigma s c =
(* if s = [] then c else *)
- local_strong whd_betaiota Evd.empty (plain_instance s c)
+ local_strong whd_betaiota sigma (plain_instance s c)
(* pseudo-reduction rule:
* [hnf_prod_app env s (Prod(_,B)) N --> B[N]
@@ -720,24 +718,24 @@ let hnf_prod_app env sigma t n =
| Prod (_,_,b) -> subst1 n b
| _ -> anomaly "hnf_prod_app: Need a product"
-let hnf_prod_appvect env sigma t nl =
+let hnf_prod_appvect env sigma t nl =
Array.fold_left (hnf_prod_app env sigma) t nl
-let hnf_prod_applist env sigma t nl =
+let hnf_prod_applist env sigma t nl =
List.fold_left (hnf_prod_app env sigma) t nl
-
+
let hnf_lam_app env sigma t n =
match kind_of_term (whd_betadeltaiota env sigma t) with
| Lambda (_,_,b) -> subst1 n b
| _ -> anomaly "hnf_lam_app: Need an abstraction"
-let hnf_lam_appvect env sigma t nl =
+let hnf_lam_appvect env sigma t nl =
Array.fold_left (hnf_lam_app env sigma) t nl
-let hnf_lam_applist env sigma t nl =
+let hnf_lam_applist env sigma t nl =
List.fold_left (hnf_lam_app env sigma) t nl
-let splay_prod env sigma =
+let splay_prod env sigma =
let rec decrec env m c =
let t = whd_betadeltaiota env sigma c in
match kind_of_term t with
@@ -745,10 +743,10 @@ let splay_prod env sigma =
decrec (push_rel (n,None,a) env)
((n,a)::m) c0
| _ -> m,t
- in
+ in
decrec env []
-let splay_lambda env sigma =
+let splay_lam env sigma =
let rec decrec env m c =
let t = whd_betadeltaiota env sigma c in
match kind_of_term t with
@@ -756,23 +754,23 @@ let splay_lambda env sigma =
decrec (push_rel (n,None,a) env)
((n,a)::m) c0
| _ -> m,t
- in
+ in
decrec env []
-let splay_prod_assum env sigma =
+let splay_prod_assum env sigma =
let rec prodec_rec env l c =
let t = whd_betadeltaiota_nolet env sigma c in
match kind_of_term t with
| Prod (x,t,c) ->
prodec_rec (push_rel (x,None,t) env)
- (Sign.add_rel_decl (x, None, t) l) c
+ (add_rel_decl (x, None, t) l) c
| LetIn (x,b,t,c) ->
prodec_rec (push_rel (x, Some b, t) env)
- (Sign.add_rel_decl (x, Some b, t) l) c
+ (add_rel_decl (x, Some b, t) l) c
| Cast (c,_,_) -> prodec_rec env l c
| _ -> l,t
in
- prodec_rec env Sign.empty_rel_context
+ prodec_rec env empty_rel_context
let splay_arity env sigma c =
let l, c = splay_prod env sigma c in
@@ -782,15 +780,25 @@ let splay_arity env sigma c =
let sort_of_arity env c = snd (splay_arity env Evd.empty c)
-let decomp_n_prod env sigma n =
- let rec decrec env m ln c = if m = 0 then (ln,c) else
+let splay_prod_n env sigma n =
+ let rec decrec env m ln c = if m = 0 then (ln,c) else
match kind_of_term (whd_betadeltaiota env sigma c) with
| Prod (n,a,c0) ->
decrec (push_rel (n,None,a) env)
- (m-1) (Sign.add_rel_decl (n,None,a) ln) c0
- | _ -> invalid_arg "decomp_n_prod"
- in
- decrec env n Sign.empty_rel_context
+ (m-1) (add_rel_decl (n,None,a) ln) c0
+ | _ -> invalid_arg "splay_prod_n"
+ in
+ decrec env n empty_rel_context
+
+let splay_lam_n env sigma n =
+ let rec decrec env m ln c = if m = 0 then (ln,c) else
+ match kind_of_term (whd_betadeltaiota env sigma c) with
+ | Lambda (n,a,c0) ->
+ decrec (push_rel (n,None,a) env)
+ (m-1) (add_rel_decl (n,None,a) ln) c0
+ | _ -> invalid_arg "splay_lam_n"
+ in
+ decrec env n empty_rel_context
exception NotASort
@@ -800,22 +808,22 @@ let decomp_sort env sigma t =
| _ -> raise NotASort
let is_sort env sigma arity =
- try let _ = decomp_sort env sigma arity in true
+ try let _ = decomp_sort env sigma arity in true
with NotASort -> false
(* reduction to head-normal-form allowing delta/zeta only in argument
of case/fix (heuristic used by evar_conv) *)
let whd_betaiota_deltazeta_for_iota_state env sigma s =
- let rec whrec 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 rslt = mkCase (ci, p, applist (cr,crargs), lf) in
- if reducible_mind_case cr then
+ if reducible_mind_case cr then
whrec (rslt, stack)
- else
+ else
s
| Fix fix ->
(match reduce_fix (whd_betadeltaiota_state env) sigma fix stack with
@@ -829,15 +837,15 @@ let whd_betaiota_deltazeta_for_iota_state env sigma s =
* Used in Correctness.
* Added by JCF, 29/1/98. *)
-let whd_programs_stack env sigma =
+let whd_programs_stack env sigma =
let rec whrec (x, stack as s) =
match kind_of_term x with
| App (f,cl) ->
let n = Array.length cl - 1 in
let c = cl.(n) in
- if occur_existential c then
- s
- else
+ if occur_existential c then
+ s
+ else
whrec (mkApp (f, Array.sub cl 0 n), append_stack [|c|] stack)
| LetIn (_,b,_,c) ->
if occur_existential b then
@@ -864,7 +872,7 @@ let whd_programs_stack env sigma =
| Reduced s' -> whrec s'
| NotReducible -> s)
| _ -> s
- in
+ in
whrec
let whd_programs env sigma x =
@@ -879,7 +887,7 @@ let find_conclusion env sigma =
| Prod (x,t,c0) -> decrec (push_rel (x,None,t) env) c0
| Lambda (x,t,c0) -> decrec (push_rel (x,None,t) env) c0
| t -> t
- in
+ in
decrec env
let is_arity env sigma c =
@@ -890,44 +898,44 @@ let is_arity env sigma c =
(*************************************)
(* Metas *)
-let meta_value evd mv =
+let meta_value evd mv =
let rec valrec mv =
match meta_opt_fvalue evd mv with
- | Some (b,_) ->
- instance
+ | Some (b,_) ->
+ instance evd
(List.map (fun mv' -> (mv',valrec mv')) (Metaset.elements b.freemetas))
b.rebus
| None -> mkMeta mv
- in
+ in
valrec mv
-let meta_instance env b =
+let meta_instance sigma b =
let c_sigma =
- List.map
- (fun mv -> (mv,meta_value env mv)) (Metaset.elements b.freemetas)
- in
- if c_sigma = [] then b.rebus else instance c_sigma b.rebus
+ List.map
+ (fun mv -> (mv,meta_value sigma mv)) (Metaset.elements b.freemetas)
+ in
+ if c_sigma = [] then b.rebus else instance sigma c_sigma b.rebus
-let nf_meta env c = meta_instance env (mk_freelisted c)
+let nf_meta sigma c = meta_instance sigma (mk_freelisted c)
(* Instantiate metas that create beta/iota redexes *)
-let meta_value evd mv =
+let meta_value evd mv =
let rec valrec mv =
match meta_opt_fvalue evd mv with
| Some (b,_) ->
- instance
+ instance evd
(List.map (fun mv' -> (mv',valrec mv')) (Metaset.elements b.freemetas))
b.rebus
| None -> mkMeta mv
- in
+ in
valrec mv
let meta_reducible_instance evd b =
let fm = Metaset.elements b.freemetas in
- let metas = List.fold_left (fun l mv ->
+ let metas = List.fold_left (fun l mv ->
match (try meta_opt_fvalue evd mv with Not_found -> None) with
- | Some (g,(_,s)) -> (mv,(g.rebus,s))::l
+ | Some (g,(_,s)) -> (mv,(g.rebus,s))::l
| None -> l) [] fm in
let rec irec u =
let u = whd_betaiota Evd.empty u in
@@ -956,5 +964,23 @@ let meta_reducible_instance evd b =
(try let g,s = List.assoc m metas in if s<>CoerceToType then irec g else u
with Not_found -> u)
| _ -> map_constr irec u
- in
+ in
if fm = [] then (* nf_betaiota? *) b.rebus else irec b.rebus
+
+
+let head_unfold_under_prod ts env _ c =
+ let unfold cst =
+ if Cpred.mem cst (snd ts) then
+ match constant_opt_value env cst with
+ | Some c -> c
+ | None -> mkConst cst
+ else mkConst cst in
+ let rec aux c =
+ match kind_of_term c with
+ | Prod (n,t,c) -> mkProd (n,aux t, aux c)
+ | _ ->
+ let (h,l) = decompose_app c in
+ match kind_of_term h with
+ | Const cst -> beta_applist (unfold cst,l)
+ | _ -> c in
+ aux c
diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli
index 8b3657c7..127dbe6b 100644
--- a/pretyping/reductionops.mli
+++ b/pretyping/reductionops.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: reductionops.mli 11897 2009-02-09 19:28:02Z barras $ i*)
+(*i $Id$ i*)
(*i*)
open Names
@@ -56,13 +56,13 @@ type contextual_reduction_function = env -> evar_map -> constr -> constr
type reduction_function = contextual_reduction_function
type local_reduction_function = evar_map -> constr -> constr
-type contextual_stack_reduction_function =
+type contextual_stack_reduction_function =
env -> evar_map -> constr -> constr * constr list
type stack_reduction_function = contextual_stack_reduction_function
type local_stack_reduction_function =
evar_map -> constr -> constr * constr list
-type contextual_state_reduction_function =
+type contextual_state_reduction_function =
env -> evar_map -> state -> state
type state_reduction_function = contextual_state_reduction_function
type local_state_reduction_function = evar_map -> state -> state
@@ -79,15 +79,15 @@ val strong : reduction_function -> reduction_function
val local_strong : local_reduction_function -> local_reduction_function
val strong_prodspine : local_reduction_function -> local_reduction_function
(*i
-val stack_reduction_of_reduction :
+val stack_reduction_of_reduction :
'a reduction_function -> 'a state_reduction_function
i*)
-val stacklam : (state -> 'a) -> constr list -> constr -> constr stack -> 'a
+val stacklam : (state -> 'a) -> constr list -> constr -> constr stack -> 'a
(*s 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
@@ -136,6 +136,7 @@ val whd_betadeltaiotaeta_state : state_reduction_function
val whd_betadeltaiotaeta : reduction_function
val whd_eta : constr -> constr
+val whd_zeta : constr -> constr
(* Various reduction functions *)
@@ -151,13 +152,13 @@ val hnf_lam_appvect : env -> evar_map -> constr -> constr array -> constr
val hnf_lam_applist : env -> evar_map -> constr -> constr list -> constr
val splay_prod : env -> evar_map -> constr -> (name * constr) list * constr
-val splay_lambda : 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 decomp_n_prod :
- env -> evar_map -> int -> constr -> Sign.rel_context * constr
+val splay_prod_n : env -> evar_map -> int -> constr -> rel_context * constr
+val splay_lam_n : env -> evar_map -> int -> constr -> rel_context * constr
val splay_prod_assum :
- env -> evar_map -> constr -> Sign.rel_context * constr
+ env -> evar_map -> constr -> rel_context * constr
val decomp_sort : env -> evar_map -> types -> sorts
val is_sort : env -> evar_map -> types -> bool
@@ -207,15 +208,16 @@ val is_trans_fconv : conv_pb -> transparent_state -> env -> evar_map -> constr
(*s Special-Purpose Reduction Functions *)
-val whd_meta : (metavariable * constr) list -> constr -> constr
+val whd_meta : evar_map -> constr -> constr
val plain_instance : (metavariable * constr) list -> constr -> constr
-val 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 *)
val whd_betaiota_deltazeta_for_iota_state : state_reduction_function
(*s Meta-related reduction functions *)
-val meta_instance : evar_defs -> constr freelisted -> constr
-val nf_meta : evar_defs -> constr -> constr
-val meta_reducible_instance : evar_defs -> constr freelisted -> constr
+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 19e46a47..1e0649da 100644
--- a/pretyping/retyping.ml
+++ b/pretyping/retyping.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: retyping.ml 11897 2009-02-09 19:28:02Z barras $ *)
+(* $Id$ *)
open Util
open Term
@@ -44,11 +44,11 @@ let type_of_var env id =
with Not_found ->
anomaly ("type_of: variable "^(string_of_id id)^" unbound")
-let retype sigma metamap =
+let retype sigma =
let rec type_of env cstr=
match kind_of_term cstr with
| Meta n ->
- (try strip_outer_cast (List.assoc n metamap)
+ (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
@@ -81,7 +81,7 @@ let retype sigma metamap =
| Cast (c,_, t) -> t
| Sort _ | Prod _ -> mkSort (sort_of env cstr)
- and sort_of env t =
+ and sort_of env t =
match kind_of_term t with
| Cast (c,_, s) when isSort s -> destSort s
| Sort (Prop c) -> type1_sort
@@ -111,14 +111,14 @@ let retype sigma metamap =
| Cast (c,_, s) when isSort s -> family_of_sort (destSort s)
| Sort (Prop c) -> InType
| Sort (Type u) -> InType
- | Prod (name,t,c2) ->
+ | Prod (name,t,c2) ->
let s2 = sort_family_of (push_rel (name,None,t) env) c2 in
if Environ.engagement env <> Some ImpredicativeSet &&
s2 = InSet & sort_family_of env t = InType then InType else s2
| App(f,args) when isGlobalRef f ->
let t = type_of_global_reference_knowing_parameters env f args in
family_of_sort (sort_of_atomic_type env sigma t args)
- | App(f,args) ->
+ | App(f,args) ->
family_of_sort (sort_of_atomic_type env sigma (type_of env f) args)
| Lambda _ | Fix _ | Construct _ ->
anomaly "sort_of: Not a type (1)"
@@ -140,10 +140,10 @@ let retype sigma metamap =
in type_of, sort_of, sort_family_of,
type_of_global_reference_knowing_parameters
-let get_sort_of env sigma t = let _,f,_,_ = retype sigma [] in f env t
-let get_sort_family_of env sigma c = let _,_,f,_ = retype sigma [] in f env c
+let get_sort_of env sigma t = let _,f,_,_ = retype sigma in f env t
+let get_sort_family_of env sigma c = let _,_,f,_ = retype sigma in f env c
let type_of_global_reference_knowing_parameters env sigma c args =
- let _,_,_,f = retype sigma [] in f env c args
+ let _,_,_,f = retype sigma in f env c args
let type_of_global_reference_knowing_conclusion env sigma c conclty =
let conclty = nf_evar sigma conclty in
@@ -161,11 +161,10 @@ let type_of_global_reference_knowing_conclusion env sigma c conclty =
(* We are outside the kernel: we take fresh universes *)
(* to avoid tactics and co to refresh universes themselves *)
-let get_type_of env sigma c =
- let f,_,_,_ = retype sigma [] in refresh_universes (f env c)
-
-let get_type_of_with_meta env sigma metamap c =
- let f,_,_,_ = retype sigma metamap in refresh_universes (f env c)
+let get_type_of ?(refresh=true) env sigma c =
+ let f,_,_,_ = retype sigma in
+ let t = f env c in
+ if refresh then refresh_universes t else t
(* Makes an assumption from a constr *)
let get_assumption_of env evc c = c
diff --git a/pretyping/retyping.mli b/pretyping/retyping.mli
index ec1fc827..8576d5ba 100644
--- a/pretyping/retyping.mli
+++ b/pretyping/retyping.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: retyping.mli 11436 2008-10-07 13:56:55Z barras $ i*)
+(*i $Id$ i*)
(*i*)
open Names
@@ -21,21 +21,18 @@ open Environ
either produces a wrong result or raise an anomaly. Use with care.
It doesn't handle predicative universes too. *)
-val get_type_of : env -> evar_map -> constr -> types
+val get_type_of : ?refresh:bool -> env -> evar_map -> constr -> types
val get_sort_of : env -> evar_map -> types -> sorts
val get_sort_family_of : env -> evar_map -> types -> sorts_family
-val get_type_of_with_meta :
- env -> evar_map -> Termops.metamap -> constr -> types
-
(* Makes an assumption from a constr *)
val get_assumption_of : env -> evar_map -> constr -> types
(* Makes an unsafe judgment from a constr *)
val get_judgment_of : env -> evar_map -> constr -> unsafe_judgment
-val type_of_global_reference_knowing_parameters : env -> evar_map -> constr ->
+val type_of_global_reference_knowing_parameters : env -> evar_map -> constr ->
constr array -> types
-
+
val type_of_global_reference_knowing_conclusion :
env -> evar_map -> constr -> types -> types
diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml
index f579afa6..a103c49b 100644
--- a/pretyping/tacred.ml
+++ b/pretyping/tacred.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: tacred.ml 12233 2009-07-08 22:50:56Z herbelin $ *)
+(* $Id$ *)
open Pp
open Util
@@ -15,6 +15,7 @@ open Nameops
open Term
open Libnames
open Termops
+open Namegen
open Declarations
open Inductive
open Environ
@@ -22,10 +23,12 @@ open Closure
open Reductionops
open Cbv
open Rawterm
+open Pattern
+open Matching
(* Errors *)
-type reduction_tactic_error =
+type reduction_tactic_error =
InvalidAbstraction of env * constr * (env * Type_errors.type_error)
exception ReductionTacticError of reduction_tactic_error
@@ -35,15 +38,20 @@ exception ReductionTacticError of reduction_tactic_error
exception Elimconst
exception Redelimination
+let error_not_evaluable r =
+ errorlabstrm "error_not_evaluable"
+ (str "Cannot coerce" ++ spc () ++ Nametab.pr_global_env Idset.empty r ++
+ spc () ++ str "to an evaluable reference.")
+
+let is_evaluable_const env cst =
+ is_transparent (ConstKey cst) && evaluable_constant cst env
+
+let is_evaluable_var env id =
+ is_transparent (VarKey id) && evaluable_named id env
+
let is_evaluable env = function
- | EvalConstRef kn ->
- is_transparent (ConstKey kn) &&
- let cb = Environ.lookup_constant kn env in
- cb.const_body <> None & not cb.const_opaque
- | EvalVarRef id ->
- is_transparent (VarKey id) &&
- let (_,value,_) = Environ.lookup_named id env in
- value <> None
+ | EvalConstRef cst -> is_evaluable_const env cst
+ | EvalVarRef id -> is_evaluable_var env id
let value_of_evaluable_ref env = function
| EvalConstRef con -> constant_value env con
@@ -53,6 +61,15 @@ let constr_of_evaluable_ref = function
| EvalConstRef con -> mkConst con
| EvalVarRef id -> mkVar id
+let evaluable_of_global_reference env = function
+ | ConstRef cst when is_evaluable_const env cst -> EvalConstRef cst
+ | VarRef id when is_evaluable_var env id -> EvalVarRef id
+ | r -> error_not_evaluable r
+
+let global_of_evaluable_reference = function
+ | EvalConstRef cst -> ConstRef cst
+ | EvalVarRef id -> VarRef id
+
type evaluable_reference =
| EvalConst of constant
| EvalVar of identifier
@@ -98,7 +115,7 @@ let reference_value sigma env c =
(* Reduction of constants hiding a fixpoint (e.g. for "simpl" tactic). *)
(* One reuses the name of the function after reduction of the fixpoint *)
-type constant_evaluation =
+type constant_evaluation =
| EliminationFix of int * int * (int * (int * constr) list * int)
| EliminationMutualFix of
int * evaluable_reference *
@@ -109,19 +126,12 @@ type constant_evaluation =
(* We use a cache registered as a global table *)
-module CstOrdered =
- struct
- type t = constant
- let compare = Pervasives.compare
- end
-module Cstmap = Map.Make(CstOrdered)
-
-let eval_table = ref Cstmap.empty
+let eval_table = ref Cmap.empty
-type frozen = (int * constant_evaluation) Cstmap.t
+type frozen = (int * constant_evaluation) Cmap.t
let init () =
- eval_table := Cstmap.empty
+ eval_table := Cmap.empty
let freeze () =
!eval_table
@@ -129,22 +139,20 @@ let freeze () =
let unfreeze ct =
eval_table := ct
-let _ =
+let _ =
Summary.declare_summary "evaluation"
{ Summary.freeze_function = freeze;
Summary.unfreeze_function = unfreeze;
- Summary.init_function = init;
- Summary.survive_module = false;
- Summary.survive_section = false }
+ Summary.init_function = init }
(* [compute_consteval] determines whether c is an "elimination constant"
either [yn:Tn]..[y1:T1](match yi with f1..fk end g1 ..gp)
or [yn:Tn]..[y1:T1](Fix(f|t) yi1..yip)
- with yi1..yip distinct variables among the yi, not occurring in t
+ with yi1..yip distinct variables among the yi, not occurring in t
- In the second case, [check_fix_reversibility [T1;...;Tn] args fix]
+ In the second case, [check_fix_reversibility [T1;...;Tn] args fix]
checks that [args] is a subset of disjoint variables in y1..yn (a necessary
condition for reversibility). It also returns the relevant
information ([i1,Ti1;..;ip,Tip],n) in order to compute an
@@ -153,7 +161,7 @@ let _ =
g := [xp:Tip']..[x1:Ti1'](f a1..an)
== [xp:Tip']..[x1:Ti1'](Fix(f|t) yi1..yip)
- with a_k:=y_k if k<>i_j, a_k:=args_k otherwise, and
+ with a_k:=y_k if k<>i_j, a_k:=args_k otherwise, and
Tij':=Tij[x1..xi(j-1) <- a1..ai(j-1)]
Note that the types Tk, when no i_j=k, must not be dependent on
@@ -172,15 +180,15 @@ let check_fix_reversibility labs args ((lv,i),(_,tys,bds)) =
if
array_for_all (noccurn k) tys
&& array_for_all (noccurn (k+nbfix)) bds
- then
- (k, List.nth labs (k-1))
- else
+ then
+ (k, List.nth labs (k-1))
+ else
raise Elimconst
- | _ ->
+ | _ ->
raise Elimconst) args
in
let reversible_rels = List.map fst li in
- if not (list_distinct reversible_rels) then
+ if not (list_distinct reversible_rels) then
raise Elimconst;
list_iter_i (fun i t_i ->
if not (List.mem_assoc (i+1) li) then
@@ -189,8 +197,8 @@ let check_fix_reversibility labs args ((lv,i),(_,tys,bds)) =
labs;
let k = lv.(i) in
if k < nargs then
-(* Such an optimisation would need eta-expansion
- let p = destRel (List.nth args k) in
+(* Such an optimisation would need eta-expansion
+ let p = destRel (List.nth args k) in
EliminationFix (n-p+1,(nbfix,li,n))
*)
EliminationFix (n,nargs,(nbfix,li,n))
@@ -201,7 +209,7 @@ let check_fix_reversibility labs args ((lv,i),(_,tys,bds)) =
components of a mutual fixpoint *)
let invert_name labs l na0 env sigma ref = function
- | Name id ->
+ | Name id ->
let minfxargs = List.length l in
if na0 <> Name id then
let refi = match ref with
@@ -215,7 +223,7 @@ let invert_name labs l na0 env sigma ref = function
| Some ref ->
try match reference_opt_value sigma env ref with
| None -> None
- | Some c ->
+ | Some c ->
let labs',ccl = decompose_lam c in
let _, l' = whd_betalet_stack sigma ccl in
let labs' = List.map snd labs' in
@@ -236,11 +244,11 @@ let compute_consteval_direct sigma env ref =
| Lambda (id,t,g) when l=[] ->
srec (push_rel (id,None,t) env) (n+1) (t::labs) g
| Fix fix ->
- (try check_fix_reversibility labs l fix
+ (try check_fix_reversibility labs l fix
with Elimconst -> NotAnElimination)
| Case (_,_,d,_) when isRel d -> EliminationCases n
| _ -> NotAnElimination
- in
+ in
match reference_opt_value sigma env ref with
| None -> NotAnElimination
| Some c -> srec env 0 [] c
@@ -271,7 +279,7 @@ let compute_consteval_mutual_fix sigma env ref =
| None -> anomaly "Should have been trapped by compute_direct"
| Some c -> srec env (minarg-nargs) [] ref c)
| _ -> (* Should not occur *) NotAnElimination
- in
+ in
match reference_opt_value sigma env ref with
| None -> (* Should not occur *) NotAnElimination
| Some c -> srec env 0 [] ref c
@@ -281,27 +289,27 @@ let compute_consteval sigma env ref =
| EliminationFix (_,_,(nbfix,_,_)) when nbfix <> 1 ->
compute_consteval_mutual_fix sigma env ref
| elim -> elim
-
+
let reference_eval sigma env = function
- | EvalConst cst as ref ->
+ | EvalConst cst as ref ->
(try
- Cstmap.find cst !eval_table
+ Cmap.find cst !eval_table
with Not_found -> begin
let v = compute_consteval sigma env ref in
- eval_table := Cstmap.add cst v !eval_table;
+ eval_table := Cmap.add cst v !eval_table;
v
end)
| ref -> compute_consteval sigma env ref
-let rev_firstn_liftn fn ln =
- let rec rfprec p res l =
- if p = 0 then
- res
+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
+ in
rfprec fn []
(* If f is bound to EliminationFix (n',infos), then n' is the minimal
@@ -318,7 +326,7 @@ let rev_firstn_liftn fn ln =
s.t. (g u1 ... up) reduces to (Fix(..) u1 ... up)
- This is made possible by setting
+ This is made possible by setting
a_k:=x_j if k=i_j for some j
a_k:=arg_k otherwise
@@ -332,30 +340,30 @@ let make_elim_fun (names,(nbfix,lv,n)) largs =
let p = List.length lv in
let lyi = List.map fst lv in
let la =
- list_map_i (fun q aq ->
- (* k from the comment is q+1 *)
+ list_map_i (fun q aq ->
+ (* k from the comment is q+1 *)
try mkRel (p+1-(list_index (n-q) lyi))
with Not_found -> aq)
- 0 (List.map (lift p) lu)
- in
+ 0 (List.map (lift p) lu)
+ in
fun i ->
match names.(i) with
| None -> None
| Some (minargs,ref) ->
let body = applistc (mkEvalRef ref) la in
- let g =
+ let g =
list_fold_left_i (fun q (* j = n+1-q *) c (ij,tij) ->
let subst = List.map (lift (-q)) (list_firstn (n-ij) la) in
let tij' = substl (List.rev subst) tij in
mkLambda (x,tij',c)) 1 body (List.rev lv)
in Some (minargs,g)
-(* [f] is convertible to [Fix(recindices,bodynum),bodyvect)]:
+(* [f] is convertible to [Fix(recindices,bodynum),bodyvect)]:
do so that the reduction uses this extra information *)
let dummy = mkProp
let vfx = id_of_string"_expanded_fix_"
-let vfun = id_of_string"_elimminator_function_"
+let vfun = id_of_string"_eliminator_function_"
(* Mark every occurrence of substituted vars (associated to a function)
as a problem variable: an evar that can be instantiated either by
@@ -392,7 +400,7 @@ exception Partial
reduction is solved by the expanded fix term. *)
let solve_arity_problem env sigma fxminargs c =
let evm = ref sigma in
- let set_fix i = evm := Evd.define !evm i (mkVar vfx) in
+ let set_fix i = evm := Evd.define i (mkVar vfx) !evm in
let rec check strict c =
let c' = whd_betaiotazeta sigma c in
let (h,rcargs) = decompose_app c' in
@@ -448,7 +456,7 @@ let reduce_fix_use_function env sigma f whfun fix stack =
let (recarg'hd,_ as recarg') =
if isRel recarg then
(* The recarg cannot be a local def, no worry about the right env *)
- (recarg, empty_stack)
+ (recarg, empty_stack)
else
whfun (recarg, empty_stack) in
let stack' = stack_assign stack recargnum (app_stack recarg') in
@@ -466,7 +474,7 @@ let contract_cofix_use_function env sigma f
(nf_beta sigma bodies.(bodynum))
let reduce_mind_case_use_function func env sigma mia =
- match kind_of_term mia.mconstr with
+ match kind_of_term mia.mconstr with
| Construct(ind_sp,i) ->
let real_cargs = list_skipn mia.mci.ci_npar mia.mcargs in
applist (mia.mlf.(i-1), real_cargs)
@@ -480,9 +488,9 @@ let reduce_mind_case_use_function func env sigma mia =
else match names.(i) with
| Anonymous -> None
| Name id ->
- (* In case of a call to another component of a block of
+ (* In case of a call to another component of a block of
mutual inductive, try to reuse the global name if
- the block was indeed initially built as a global
+ the block was indeed initially built as a global
definition *)
let kn = make_con mp dp (label_of_id id) in
try match constant_opt_value env kn with
@@ -498,8 +506,8 @@ let reduce_mind_case_use_function func env sigma mia =
| _ -> assert false
let special_red_case env sigma whfun (ci, p, c, lf) =
- let rec redrec s =
- let (constr, cargs) = whfun s in
+ let rec redrec s =
+ let (constr, cargs) = whfun s in
if isEvalRef env constr then
let ref = destEvalRef constr in
match reference_opt_value sigma env ref with
@@ -516,9 +524,9 @@ let special_red_case env sigma whfun (ci, p, c, lf) =
reduce_mind_case
{mP=p; mconstr=constr; mcargs=list_of_stack cargs;
mci=ci; mlf=lf}
- else
+ else
raise Redelimination
- in
+ in
redrec (c, empty_stack)
(* [red_elim_const] contracts iota/fix/cofix redexes hidden behind
@@ -565,14 +573,14 @@ and whd_simpl_state env sigma s =
let rec redrec (x, stack as s) =
match kind_of_term x with
| Lambda (na,t,c) ->
- (match decomp_stack stack with
+ (match decomp_stack stack with
| None -> s
| Some (a,rest) -> stacklam redrec [a] c rest)
| LetIn (n,b,t,c) -> stacklam redrec [b] c stack
| App (f,cl) -> redrec (f, append_stack cl stack)
| Cast (c,_,_) -> redrec (c, stack)
| Case (ci,p,c,lf) ->
- (try
+ (try
redrec (special_red_case env sigma redrec (ci,p,c,lf), stack)
with
Redelimination -> s)
@@ -588,13 +596,13 @@ and whd_simpl_state env sigma s =
with Redelimination ->
s)
| _ -> s
- in
+ in
redrec s
(* reduce until finding an applied constructor or fail *)
and whd_construct_state env sigma s =
- let (constr, cargs as s') = whd_simpl_state env sigma s in
+ let (constr, cargs as s') = whd_simpl_state env sigma s in
if reducible_mind_case constr then s'
else if isEvalRef env constr then
let ref = destEvalRef constr in
@@ -612,11 +620,11 @@ and whd_construct_state env sigma s =
sequence of products; fails if no delta redex is around
*)
-let try_red_product env sigma c =
+let try_red_product env sigma c =
let simpfun = clos_norm_flags betaiotazeta env sigma in
let rec redrec env x =
match kind_of_term x with
- | App (f,l) ->
+ | App (f,l) ->
(match kind_of_term f with
| Fix fix ->
let stack = append_stack l empty_stack in
@@ -631,7 +639,7 @@ let try_red_product env sigma c =
| Prod (x,a,b) -> mkProd (x, a, redrec (push_rel (x,None,a) env) b)
| LetIn (x,a,b,t) -> redrec env (subst1 a t)
| Case (ci,p,d,lf) -> simpfun (mkCase (ci,p,redrec env d,lf))
- | _ when isEvalRef env x ->
+ | _ when isEvalRef env x ->
(* TO DO: re-fold fixpoints after expansion *)
(* to get true one-step reductions *)
let ref = destEvalRef x in
@@ -641,17 +649,17 @@ let try_red_product env sigma c =
| _ -> raise Redelimination
in redrec env c
-let red_product env sigma c =
+let red_product env sigma c =
try try_red_product env sigma c
with Redelimination -> error "Not reducible."
(*
-(* This old version of hnf uses betadeltaiota instead of itself (resp
+(* This old version of hnf uses betadeltaiota instead of itself (resp
whd_construct_state) to reduce the argument of Case (resp Fix);
The new version uses the "simpl" strategy instead. For instance,
Variable n:nat.
- Eval hnf in match (plus (S n) O) with S n => n | _ => O end.
+ Eval hnf in match (plus (S n) O) with S n => n | _ => O end.
returned
@@ -678,7 +686,7 @@ let whd_simpl_orelse_delta_but_fix_old env sigma c =
| Case (ci,p,d,lf) ->
(try
redrec (special_red_case env sigma whd_all (ci,p,d,lf), stack)
- with Redelimination ->
+ with Redelimination ->
s)
| Fix fix ->
(match reduce_fix whd_all fix stack with
@@ -691,7 +699,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 (snd (decompose_lam c)) with
+ (match kind_of_term ((strip_lam c)) with
| CoFix _ | Fix _ -> s
| _ -> redrec (c, stack))
| None -> s)
@@ -705,11 +713,11 @@ let whd_simpl_orelse_delta_but_fix_old env sigma c =
let whd_simpl_orelse_delta_but_fix env sigma c =
let rec redrec s =
- let (constr, stack as s') = whd_simpl_state env sigma s in
+ let (constr, stack as s') = whd_simpl_state env sigma s in
if isEvalRef env constr then
match reference_opt_value sigma env (destEvalRef constr) with
| Some c ->
- (match kind_of_term (snd (decompose_lam c)) with
+ (match kind_of_term ((strip_lam c)) with
| CoFix _ | Fix _ -> s'
| _ -> redrec (c, stack))
| None -> s'
@@ -725,14 +733,12 @@ let whd_simpl env sigma c =
let simpl env sigma c = strong whd_simpl env sigma c
-let nf = simpl (* Compatibility *)
-
(* Reduction at specific subterms *)
-let is_head c t =
+let matches_head c t =
match kind_of_term t with
- | App (f,_) -> f = c
- | _ -> false
+ | App (f,_) -> matches c f
+ | _ -> raise PatternMatchingFailure
let contextually byhead ((nowhere_except_in,locs),c) f env sigma t =
let maxocc = List.fold_right max locs 0 in
@@ -740,22 +746,23 @@ let contextually byhead ((nowhere_except_in,locs),c) f env sigma t =
let rec traverse (env,c as envc) t =
if nowhere_except_in & (!pos > maxocc) then t
else
- if (not byhead & eq_constr c t) or (byhead & is_head c t) then
- let ok =
+ try
+ let subst = if byhead then matches_head c t else matches c t in
+ let ok =
if nowhere_except_in then List.mem !pos locs
else not (List.mem !pos locs) in
incr pos;
if ok then
- f env sigma t
+ f subst env sigma t
else if byhead then
(* find other occurrences of c in t; TODO: ensure left-to-right *)
let (f,l) = destApp t in
mkApp (f, array_map_left (traverse envc) l)
else
t
- else
+ with PatternMatchingFailure ->
map_constr_with_binders_left_to_right
- (fun d (env,c) -> (push_rel d env,lift 1 c))
+ (fun d (env,c) -> (push_rel d env,lift_pattern 1 c))
traverse envc t
in
let t' = traverse (env,c) t in
@@ -775,7 +782,7 @@ let substlin env evalref n (nowhere_except_in,locs) c =
let rec substrec () c =
if nowhere_except_in & !pos > maxocc then c
else if c = term then
- let ok =
+ let ok =
if nowhere_except_in then List.mem !pos locs
else not (List.mem !pos locs) in
incr pos;
@@ -791,7 +798,7 @@ let substlin env evalref n (nowhere_except_in,locs) c =
let string_of_evaluable_ref env = function
| EvalVarRef id -> string_of_id id
| EvalConstRef kn ->
- string_of_qualid
+ string_of_qualid
(Nametab.shortest_qualid_of_global (vars_of_env env) (ConstRef kn))
let unfold env sigma name =
@@ -800,7 +807,7 @@ let unfold env sigma name =
else
error (string_of_evaluable_ref env name^" is opaque.")
-(* [unfoldoccs : (readable_constraints -> (int list * section_path) -> constr -> constr)]
+(* [unfoldoccs : (readable_constraints -> (int list * full_path) -> constr -> constr)]
* Unfolds the constant name in a term c following a list of occurrences occl.
* at the occurrences of occ_list. If occ_list is empty, unfold all occurences.
* Performs a betaiota reduction after unfolding. *)
@@ -808,14 +815,14 @@ let unfoldoccs env sigma ((nowhere_except_in,locs as plocs),name) c =
if locs = [] then if nowhere_except_in then c else unfold env sigma name c
else
let (nbocc,uc) = substlin env name 1 plocs c in
- if nbocc = 1 then
+ if nbocc = 1 then
error ((string_of_evaluable_ref env name)^" does not occur.");
let rest = List.filter (fun o -> o >= nbocc) locs in
if rest <> [] then error_invalid_occurrence rest;
nf_betaiota sigma uc
(* Unfold reduction tactic: *)
-let unfoldn loccname env sigma c =
+let unfoldn loccname env sigma c =
List.fold_left (fun c occname -> unfoldoccs env sigma occname c) c loccname
(* Re-folding constants tactics: refold com in term c *)
@@ -858,9 +865,9 @@ let abstract_scheme env sigma (locc,a) c =
let ta = Retyping.get_type_of env sigma a in
let na = named_hd env ta Anonymous in
if occur_meta ta then error "Cannot find a type for the generalisation.";
- if occur_meta a then
+ if occur_meta a then
mkLambda (na,ta,c)
- else
+ else
mkLambda (na,ta,subst_term_occ locc a c)
let pattern_occs loccs_trm env sigma c =
@@ -876,7 +883,7 @@ let pattern_occs loccs_trm env sigma c =
(* put t as t'=(x1:A1)..(xn:An)B with B an inductive definition of name name
return name, B and t' *)
-let reduce_to_ind_gen allow_product env sigma t =
+let reduce_to_ind_gen allow_product env sigma t =
let rec elimrec env t l =
let t = hnf_constr env sigma t in
match kind_of_term (fst (decompose_app t)) with
@@ -904,7 +911,7 @@ let reduce_to_atomic_ind x = reduce_to_ind_gen false x
exception NotStepReducible
-let one_step_reduce env sigma c =
+let one_step_reduce env sigma c =
let rec redrec (x, stack) =
match kind_of_term x with
| Lambda (n,t,c) ->
@@ -933,7 +940,7 @@ let one_step_reduce env sigma c =
| None -> raise NotStepReducible)
| _ -> raise NotStepReducible
- in
+ in
app_stack (redrec (c, empty_stack))
let isIndRef = function IndRef _ -> true | _ -> false
@@ -942,34 +949,34 @@ let reduce_to_ref_gen allow_product env sigma ref t =
if isIndRef ref then
let (mind,t) = reduce_to_ind_gen allow_product env sigma t in
if IndRef mind <> ref then
- errorlabstrm "" (str "Cannot recognize a statement based on " ++
+ errorlabstrm "" (str "Cannot recognize a statement based on " ++
Nametab.pr_global_env Idset.empty ref ++ str".")
else
t
else
(* lazily reduces to match the head of [t] with the expected [ref] *)
- let rec elimrec env t l =
+ let rec elimrec env t l =
let c, _ = Reductionops.whd_stack sigma t in
match kind_of_term c with
| Prod (n,ty,t') ->
if allow_product then
elimrec (push_rel (n,None,t) env) t' ((n,None,ty)::l)
- else
- errorlabstrm ""
- (str "Cannot recognize an atomic statement based on " ++
+ else
+ errorlabstrm ""
+ (str "Cannot recognize an atomic statement based on " ++
Nametab.pr_global_env Idset.empty ref ++ str".")
| _ ->
- try
- if global_of_constr c = ref
+ try
+ if global_of_constr c = ref
then it_mkProd_or_LetIn t l
else raise Not_found
with Not_found ->
- try
- let t' = nf_betaiota sigma (one_step_reduce env sigma t) in
+ try
+ let t' = nf_betaiota sigma (one_step_reduce env sigma t) in
elimrec env t' l
with NotStepReducible ->
errorlabstrm ""
- (str "Cannot recognize a statement based on " ++
+ (str "Cannot recognize a statement based on " ++
Nametab.pr_global_env Idset.empty ref ++ str".")
in
elimrec env t []
diff --git a/pretyping/tacred.mli b/pretyping/tacred.mli
index e12d6ad2..2bba87a7 100644
--- a/pretyping/tacred.mli
+++ b/pretyping/tacred.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: tacred.mli 11094 2008-06-10 19:35:23Z herbelin $ i*)
+(*i $Id$ i*)
(*i*)
open Names
@@ -17,16 +17,27 @@ open Reductionops
open Closure
open Rawterm
open Termops
+open Pattern
(*i*)
-type reduction_tactic_error =
+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} *)
-val is_evaluable : env -> evaluable_global_reference -> bool
+(* Evaluable global reference *)
+
+val is_evaluable : Environ.env -> evaluable_global_reference -> bool
+
+val error_not_evaluable : Libnames.global_reference -> 'a
+
+val evaluable_of_global_reference :
+ Environ.env -> Libnames.global_reference -> evaluable_global_reference
+
+val global_of_evaluable_reference :
+ evaluable_global_reference -> Libnames.global_reference
exception Redelimination
@@ -37,7 +48,7 @@ val red_product : reduction_function
val try_red_product : reduction_function
(* Simpl *)
-val simpl : reduction_function
+val simpl : reduction_function
(* Simpl only at the head *)
val whd_simpl : reduction_function
@@ -47,7 +58,7 @@ val whd_simpl : reduction_function
val hnf_constr : reduction_function
(* Unfold *)
-val unfoldn :
+val unfoldn :
(occurrences * evaluable_global_reference) list -> reduction_function
(* Fold *)
@@ -75,17 +86,12 @@ val reduce_to_atomic_ind : env -> evar_map -> types -> inductive * types
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
- [t'=(x1:A1)..(xn:An)(ref args)] and raise [Not_found] if not possible *)
+ [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
val reduce_to_atomic_ref :
env -> evar_map -> Libnames.global_reference -> types -> types
-val contextually : bool -> occurrences * constr -> reduction_function
- -> reduction_function
-
-(* Compatibility *)
-(* use [simpl] instead of [nf] *)
-val nf : reduction_function
-
+val contextually : bool -> occurrences * constr_pattern ->
+ (patvar_map -> reduction_function) -> reduction_function
diff --git a/pretyping/term_dnet.ml b/pretyping/term_dnet.ml
new file mode 100644
index 00000000..04e328cb
--- /dev/null
+++ b/pretyping/term_dnet.ml
@@ -0,0 +1,404 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* $Id:$ *)
+
+(*i*)
+open Util
+open Term
+open Sign
+open Names
+open Libnames
+open Mod_subst
+open Pp (* debug *)
+(*i*)
+
+
+(* Representation/approximation of terms to use in the dnet:
+ *
+ * - no meta or evar (use ['a pattern] for that)
+ *
+ * - [Rel]s and [Sort]s are not taken into account (that's why we need
+ * a second pass of linear filterin on the results - it's not a perfect
+ * term indexing structure)
+
+ * - Foralls and LetIns are represented by a context DCtx (a list of
+ * generalization, similar to rel_context, and coded with DCons and
+ * DNil). This allows for matching under an unfinished context
+ *)
+
+module DTerm =
+struct
+
+ type 't t =
+ | DRel
+ | DSort
+ | DRef of global_reference
+ | DCtx of 't * 't (* (binding list, subterm) = Prods and LetIns *)
+ | DLambda of 't * 't
+ | DApp of 't * 't (* binary app *)
+ | DCase of case_info * 't * 't * 't array
+ | DFix of int array * int * 't array * 't array
+ | DCoFix of int * 't array * 't array
+
+ (* special constructors only inside the left-hand side of DCtx or
+ DApp. Used to encode lists of foralls/letins/apps as contexts *)
+ | DCons of ('t * 't option) * 't
+ | DNil
+
+ type dconstr = dconstr t
+
+ (* debug *)
+ let rec pr_dconstr f : 'a t -> std_ppcmds = function
+ | DRel -> str "*"
+ | DSort -> str "Sort"
+ | DRef _ -> str "Ref"
+ | DCtx (ctx,t) -> f ctx ++ spc() ++ str "|-" ++ spc () ++ f t
+ | DLambda (t1,t2) -> str "fun"++ spc() ++ f t1 ++ spc() ++ str"->" ++ spc() ++ f t2
+ | DApp (t1,t2) -> f t1 ++ spc() ++ f t2
+ | DCase (_,t1,t2,ta) -> str "case"
+ | DFix _ -> str "fix"
+ | DCoFix _ -> str "cofix"
+ | DCons ((t,dopt),tl) -> f t ++ (match dopt with
+ Some t' -> str ":=" ++ f t'
+ | None -> str "") ++ spc() ++ str "::" ++ spc() ++ f tl
+ | DNil -> str "[]"
+
+ (*
+ * Functional iterators for the t datatype
+ * a.k.a boring and error-prone boilerplate code
+ *)
+
+ let map f = function
+ | (DRel | DSort | DNil | DRef _) as c -> c
+ | DCtx (ctx,c) -> DCtx (f ctx, f c)
+ | DLambda (t,c) -> DLambda (f t, f c)
+ | DApp (t,u) -> DApp (f t,f u)
+ | DCase (ci,p,c,bl) -> DCase (ci, f p, f c, Array.map f bl)
+ | DFix (ia,i,ta,ca) ->
+ DFix (ia,i,Array.map f ta,Array.map f ca)
+ | DCoFix(i,ta,ca) ->
+ DCoFix (i,Array.map f ta,Array.map f ca)
+ | DCons ((t,topt),u) -> DCons ((f t,Option.map f topt), f u)
+
+ let compare x y =
+ let make_name n =
+ match n with
+ | DRef(ConstRef con) ->
+ DRef(ConstRef(constant_of_kn(canonical_con con)))
+ | DRef(IndRef (kn,i)) ->
+ DRef(IndRef(mind_of_kn(canonical_mind kn),i))
+ | DRef(ConstructRef ((kn,i),j ))->
+ DRef(ConstructRef((mind_of_kn(canonical_mind kn),i),j))
+ | k -> k in
+ Pervasives.compare (make_name x) (make_name y)
+
+ let fold f acc = function
+ | (DRel | DNil | DSort | DRef _) -> acc
+ | DCtx (ctx,c) -> f (f acc ctx) c
+ | DLambda (t,c) -> f (f acc t) c
+ | DApp (t,u) -> f (f acc t) u
+ | DCase (ci,p,c,bl) -> Array.fold_left f (f (f acc p) c) bl
+ | DFix (ia,i,ta,ca) ->
+ Array.fold_left f (Array.fold_left f acc ta) ca
+ | DCoFix(i,ta,ca) ->
+ Array.fold_left f (Array.fold_left f acc ta) ca
+ | DCons ((t,topt),u) -> f (Option.fold_left f (f acc t) topt) u
+
+ let choose f = function
+ | (DRel | DSort | DNil | DRef _) -> invalid_arg "choose"
+ | DCtx (ctx,c) -> f ctx
+ | DLambda (t,c) -> f t
+ | DApp (t,u) -> f u
+ | DCase (ci,p,c,bl) -> f c
+ | DFix (ia,i,ta,ca) -> f ta.(0)
+ | DCoFix (i,ta,ca) -> f ta.(0)
+ | DCons ((t,topt),u) -> f u
+
+ let fold2 (f:'a -> 'b -> 'c -> 'a) (acc:'a) (c1:'b t) (c2:'c t) : 'a =
+ let head w = map (fun _ -> ()) w in
+ if compare (head c1) (head c2) <> 0
+ then invalid_arg "fold2:compare" else
+ match c1,c2 with
+ | (DRel, DRel | DNil, DNil | DSort, DSort | DRef _, DRef _) -> acc
+ | (DCtx (c1,t1), DCtx (c2,t2)
+ | DApp (c1,t1), DApp (c2,t2)
+ | DLambda (c1,t1), DLambda (c2,t2)) -> f (f acc c1 c2) t1 t2
+ | DCase (ci,p1,c1,bl1),DCase (_,p2,c2,bl2) ->
+ array_fold_left2 f (f (f acc p1 p2) c1 c2) bl1 bl2
+ | DFix (ia,i,ta1,ca1), DFix (_,_,ta2,ca2) ->
+ array_fold_left2 f (array_fold_left2 f acc ta1 ta2) ca1 ca2
+ | DCoFix(i,ta1,ca1), DCoFix(_,ta2,ca2) ->
+ array_fold_left2 f (array_fold_left2 f acc ta1 ta2) ca1 ca2
+ | DCons ((t1,topt1),u1), DCons ((t2,topt2),u2) ->
+ f (Option.fold_left2 f (f acc t1 t2) topt1 topt2) u1 u2
+ | _ -> assert false
+
+ let map2 (f:'a -> 'b -> 'c) (c1:'a t) (c2:'b t) : 'c t =
+ let head w = map (fun _ -> ()) w in
+ if compare (head c1) (head c2) <> 0
+ then invalid_arg "map2_t:compare" else
+ match c1,c2 with
+ | (DRel, DRel | DSort, DSort | DNil, DNil | DRef _, DRef _) as cc ->
+ let (c,_) = cc in c
+ | DCtx (c1,t1), DCtx (c2,t2) -> DCtx (f c1 c2, f t1 t2)
+ | DLambda (t1,c1), DLambda (t2,c2) -> DLambda (f t1 t2, f c1 c2)
+ | DApp (t1,u1), DApp (t2,u2) -> DApp (f t1 t2,f u1 u2)
+ | DCase (ci,p1,c1,bl1), DCase (_,p2,c2,bl2) ->
+ DCase (ci, f p1 p2, f c1 c2, array_map2 f bl1 bl2)
+ | DFix (ia,i,ta1,ca1), DFix (_,_,ta2,ca2) ->
+ DFix (ia,i,array_map2 f ta1 ta2,array_map2 f ca1 ca2)
+ | DCoFix (i,ta1,ca1), DCoFix (_,ta2,ca2) ->
+ DCoFix (i,array_map2 f ta1 ta2,array_map2 f ca1 ca2)
+ | DCons ((t1,topt1),u1), DCons ((t2,topt2),u2) ->
+ DCons ((f t1 t2,Option.lift2 f topt1 topt2), f u1 u2)
+ | _ -> assert false
+
+ let terminal = function
+ | (DRel | DSort | DNil | DRef _) -> true
+ | _ -> false
+end
+
+(*
+ * Terms discrimination nets
+ * Uses the general dnet datatype on DTerm.t
+ * (here you can restart reading)
+ *)
+
+(*
+ * Construction of the module
+ *)
+
+module type IDENT =
+sig
+ type t
+ val compare : t -> t -> int
+ val subst : substitution -> t -> t
+ val constr_of : t -> constr
+end
+
+module type OPT =
+sig
+ val reduce : constr -> constr
+ val direction : bool
+end
+
+module Make =
+ functor (Ident : IDENT) ->
+ functor (Opt : OPT) ->
+struct
+
+ module TDnet : Dnet.S with type ident=Ident.t
+ and type 'a structure = 'a DTerm.t
+ and type meta = metavariable
+ = Dnet.Make(DTerm)(Ident)
+ (struct
+ type t = metavariable
+ let compare = Pervasives.compare
+ end)
+
+ type t = TDnet.t
+
+ type ident = TDnet.ident
+
+ type 'a pattern = 'a TDnet.pattern
+ type term_pattern = term_pattern DTerm.t pattern
+
+ type idset = TDnet.Idset.t
+
+ type result = ident * (constr*existential_key) * Termops.subst
+
+ open DTerm
+ open TDnet
+
+ let rec pat_of_constr c : term_pattern =
+ match kind_of_term c with
+ | Rel _ -> Term DRel
+ | Sort _ -> Term DSort
+ | Var i -> Term (DRef (VarRef i))
+ | Const c -> Term (DRef (ConstRef c))
+ | Ind i -> Term (DRef (IndRef i))
+ | Construct c -> Term (DRef (ConstructRef c))
+ | Term.Meta _ -> assert false
+ | Evar (i,_) -> Meta i
+ | Case (ci,c1,c2,ca) ->
+ Term(DCase(ci,pat_of_constr c1,pat_of_constr c2,Array.map pat_of_constr ca))
+ | Fix ((ia,i),(_,ta,ca)) ->
+ Term(DFix(ia,i,Array.map pat_of_constr ta, Array.map pat_of_constr ca))
+ | CoFix (i,(_,ta,ca)) ->
+ Term(DCoFix(i,Array.map pat_of_constr ta,Array.map pat_of_constr ca))
+ | Cast (c,_,_) -> pat_of_constr c
+ | Lambda (_,t,c) -> Term(DLambda (pat_of_constr t, pat_of_constr c))
+ | (Prod (_,_,_) | LetIn(_,_,_,_)) ->
+ let (ctx,c) = ctx_of_constr (Term DNil) c in Term (DCtx (ctx,c))
+ | App (f,ca) ->
+ Array.fold_left (fun c a -> Term (DApp (c,a)))
+ (pat_of_constr f) (Array.map pat_of_constr ca)
+
+ and ctx_of_constr ctx c : term_pattern * term_pattern =
+ match kind_of_term c with
+ | Prod (_,t,c) -> ctx_of_constr (Term(DCons((pat_of_constr t,None),ctx))) c
+ | LetIn(_,d,t,c) -> ctx_of_constr (Term(DCons((pat_of_constr t, Some (pat_of_constr d)),ctx))) c
+ | _ -> ctx,pat_of_constr c
+
+ let empty_ctx : term_pattern -> term_pattern = function
+ | Meta _ as c -> c
+ | Term (DCtx(_,_)) as c -> c
+ | c -> Term (DCtx (Term DNil, c))
+
+ (*
+ * Basic primitives
+ *)
+
+ let empty = TDnet.empty
+
+ let subst s t =
+ let sleaf id = Ident.subst s id in
+ let snode = function
+ | DTerm.DRef gr -> DTerm.DRef (fst (subst_global s gr))
+ | n -> n in
+ TDnet.map sleaf snode t
+
+ let union = TDnet.union
+
+ let add (c:constr) (id:Ident.t) (dn:t) =
+ let c = Opt.reduce c in
+ let c = empty_ctx (pat_of_constr c) in
+ TDnet.add dn c id
+
+ let new_meta_no =
+ let ctr = ref 0 in
+ fun () -> decr ctr; !ctr
+
+ let new_meta_no = Evarutil.new_untyped_evar
+
+ let neutral_meta = new_meta_no()
+
+ let new_meta () = Meta (new_meta_no())
+ let new_evar () = mkEvar(new_meta_no(),[||])
+
+ let rec remove_cap : term_pattern -> term_pattern = function
+ | Term (DCons (t,u)) -> Term (DCons (t,remove_cap u))
+ | Term DNil -> new_meta()
+ | Meta _ as m -> m
+ | _ -> assert false
+
+ let under_prod : term_pattern -> term_pattern = function
+ | Term (DCtx (t,u)) -> Term (DCtx (remove_cap t,u))
+ | Meta m -> Term (DCtx(new_meta(), Meta m))
+ | _ -> assert false
+
+ let init = let e = new_meta_no() in (mkEvar (e,[||]),e)
+
+ let rec e_subst_evar i (t:unit->constr) c =
+ match kind_of_term c with
+ | Evar (j,_) when i=j -> t()
+ | _ -> map_constr (e_subst_evar i t) c
+
+ let subst_evar i c = e_subst_evar i (fun _ -> c)
+
+ (* debug *)
+ let rec pr_term_pattern p =
+ (fun pr_t -> function
+ | Term t -> pr_t t
+ | Meta m -> str"["++Util.pr_int (Obj.magic m)++str"]"
+ ) (pr_dconstr pr_term_pattern) p
+
+ let search_pat cpat dpat dn (up,plug) =
+ let whole_c = subst_evar plug cpat up in
+ (* if we are at the root, add an empty context *)
+ let dpat = if isEvar_or_Meta up then under_prod (empty_ctx dpat) else dpat in
+ TDnet.Idset.fold
+ (fun id acc ->
+ let c_id = Opt.reduce (Ident.constr_of id) in
+ let (ctx,wc) =
+ try Termops.align_prod_letin whole_c c_id
+ with Invalid_argument _ -> [],c_id in
+ let up = it_mkProd_or_LetIn up ctx in
+ let wc,whole_c = if Opt.direction then whole_c,wc else wc,whole_c in
+ try (id,(up,plug),Termops.filtering ctx Reduction.CUMUL wc whole_c)::acc
+ with Termops.CannotFilter -> (* msgnl(str"recon "++Termops.print_constr_env (Global.env()) wc); *) acc
+ ) (TDnet.find_match dpat dn) []
+
+ let fold_pattern_neutral f =
+ fold_pattern (fun acc (mset,m,dn) -> if m=neutral_meta then acc else f m dn acc)
+
+ let fold_pattern_nonlin f =
+ let defined = ref Gmap.empty in
+ fold_pattern_neutral
+ ( fun m dn acc ->
+ let dn = try TDnet.inter dn (Gmap.find m !defined) with Not_found -> dn in
+ defined := Gmap.add m dn !defined;
+ f m dn acc )
+
+ let fold_pattern_up f acc dpat cpat dn (up,plug) =
+ fold_pattern_nonlin
+ ( fun m dn acc ->
+ f dn (subst_evar plug (e_subst_evar neutral_meta new_evar cpat) up, m) acc
+ ) acc dpat dn
+
+ let possibly_under pat k dn (up,plug) =
+ let rec aux fst dn (up,plug) acc =
+ let cpat = pat() in
+ let dpat = pat_of_constr cpat in
+ let dpat = if fst then under_prod (empty_ctx dpat) else dpat in
+ (k dn (up,plug)) @
+ snd (fold_pattern_up (aux false) acc dpat cpat dn (up,plug)) in
+ aux true dn (up,plug) []
+
+ let eq_pat eq () = mkApp(eq,[|mkEvar(neutral_meta,[||]);new_evar();new_evar()|])
+ let app_pat () = mkApp(new_evar(),[|mkEvar(neutral_meta,[||])|])
+
+ (*
+ * High-level primitives describing specific search problems
+ *)
+
+ let search_pattern dn pat =
+ let pat = Opt.reduce pat in
+ search_pat pat (empty_ctx (pat_of_constr pat)) dn init
+
+ let search_concl dn pat =
+ let pat = Opt.reduce pat in
+ search_pat pat (under_prod (empty_ctx (pat_of_constr pat))) dn init
+
+ let search_eq_concl dn eq pat =
+ let pat = Opt.reduce pat in
+ let eq_pat = eq_pat eq () in
+ let eq_dpat = under_prod (empty_ctx (pat_of_constr eq_pat)) in
+ snd (fold_pattern_up
+ (fun dn up acc ->
+ search_pat pat (pat_of_constr pat) dn up @ acc
+ ) [] eq_dpat eq_pat dn init)
+
+ let search_head_concl dn pat =
+ let pat = Opt.reduce pat in
+ possibly_under app_pat (search_pat pat (pat_of_constr pat)) dn init
+
+ let find_all dn = Idset.elements (TDnet.find_all dn)
+
+ let map f dn = TDnet.map f (fun x -> x) dn
+end
+
+module type S =
+sig
+ type t
+ type ident
+
+ type result = ident * (constr*existential_key) * Termops.subst
+
+ val empty : t
+ val add : constr -> ident -> t -> t
+ val union : t -> t -> t
+ val subst : substitution -> t -> t
+ val search_pattern : t -> constr -> result list
+ val search_concl : t -> constr -> result list
+ val search_head_concl : t -> constr -> result list
+ val search_eq_concl : t -> constr -> constr -> result list
+ val find_all : t -> ident list
+ val map : (ident -> ident) -> t -> t
+end
diff --git a/pretyping/term_dnet.mli b/pretyping/term_dnet.mli
new file mode 100644
index 00000000..0e7fdb82
--- /dev/null
+++ b/pretyping/term_dnet.mli
@@ -0,0 +1,112 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \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.
+
+ An instantiation of Dnet on (an approximation of) constr. It
+ associates a term (possibly with Evar) with an
+ identifier. Identifiers must be unique (no two terms sharing the
+ same ident), and there must be a way to recover the full term from
+ the identifier (function constr_of).
+
+ Optionally, a pre-treatment on terms can be performed before adding
+ or searching (reduce). Practically, it is used to do some kind of
+ delta-reduction on terms before indexing them.
+
+ The results returned here are perfect, since post-filtering is done
+ inside here.
+
+ See lib/dnet.mli for more details.
+*)
+
+(* Identifiers to store (right hand side of the association) *)
+module type IDENT = sig
+ type t
+ val compare : t -> t -> int
+
+ (* how to substitute them for storage *)
+ val subst : substitution -> t -> t
+
+ (* how to recover the term from the identifier *)
+ val constr_of : t -> constr
+end
+
+(* Options : *)
+module type OPT = sig
+
+ (* pre-treatment to terms before adding or searching *)
+ val reduce : constr -> constr
+
+ (* direction of post-filtering w.r.t sort subtyping :
+ - true means query <= terms in the structure
+ - false means terms <= query
+ *)
+ val direction : bool
+end
+
+module type S =
+sig
+ type t
+ type ident
+
+ (* results of filtering : identifier, a context (term with Evar
+ hole) and the substitution in that context*)
+ type result = ident * (constr*existential_key) * Termops.subst
+
+ val empty : t
+
+ (* [add c i dn] adds the binding [(c,i)] to [dn]. [c] can be a
+ closed term or a pattern (with untyped Evars). No Metas accepted *)
+ val add : constr -> ident -> t -> t
+
+ (* merge of dnets. Faster than re-adding all terms *)
+ val union : t -> t -> t
+
+ val subst : substitution -> t -> t
+
+ (*
+ * High-level primitives describing specific search problems
+ *)
+
+ (* [search_pattern dn c] returns all terms/patterns in dn
+ matching/matched by c *)
+ val search_pattern : t -> constr -> result list
+
+ (* [search_concl dn c] returns all matches under products and
+ letins, i.e. it finds subterms whose conclusion matches c. The
+ complexity depends only on c ! *)
+ val search_concl : t -> constr -> result list
+
+ (* [search_head_concl dn c] matches under products and applications
+ heads. Finds terms of the form [forall H_1...H_n, C t_1...t_n]
+ where C matches c *)
+ val search_head_concl : t -> constr -> result list
+
+ (* [search_eq_concl dn eq c] searches terms of the form [forall
+ H1...Hn, eq _ X1 X2] where either X1 or X2 matches c *)
+ val search_eq_concl : t -> constr -> constr -> result list
+
+ (* [find_all dn] returns all idents contained in dn *)
+ val find_all : t -> ident list
+
+ val map : (ident -> ident) -> t -> t
+end
+
+module Make :
+ functor (Ident : IDENT) ->
+ functor (Opt : OPT) ->
+ S with type ident = Ident.t
diff --git a/pretyping/termops.ml b/pretyping/termops.ml
index 4f38fbb3..3e4c5ae5 100644
--- a/pretyping/termops.ml
+++ b/pretyping/termops.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: termops.ml 12058 2009-04-08 10:54:59Z herbelin $ *)
+(* $Id$ *)
open Pp
open Util
@@ -34,7 +34,7 @@ let pr_name = function
| Name id -> pr_id id
| Anonymous -> str "_"
-let pr_sp sp = str(string_of_kn sp)
+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
@@ -42,7 +42,7 @@ let rec pr_constr c = match kind_of_term c with
| Meta n -> str "Meta(" ++ int n ++ str ")"
| Var id -> pr_id id
| Sort s -> print_sort s
- | Cast (c,_, t) -> hov 1
+ | Cast (c,_, t) -> hov 1
(str"(" ++ pr_constr c ++ cut() ++
str":" ++ pr_constr t ++ str")")
| Prod (Name(id),t,c) -> hov 1
@@ -65,9 +65,9 @@ let rec pr_constr c = match kind_of_term c with
(str"Evar#" ++ int e ++ str"{" ++
prlist_with_sep spc pr_constr (Array.to_list l) ++str"}")
| Const c -> str"Cst(" ++ pr_con c ++ str")"
- | Ind (sp,i) -> str"Ind(" ++ pr_sp sp ++ str"," ++ int i ++ str")"
+ | Ind (sp,i) -> str"Ind(" ++ pr_mind sp ++ str"," ++ int i ++ str")"
| Construct ((sp,i),j) ->
- str"Constr(" ++ pr_sp sp ++ str"," ++ int i ++ str"," ++ int j ++ str")"
+ str"Constr(" ++ pr_mind sp ++ str"," ++ int i ++ str"," ++ int j ++ str")"
| Case (ci,p,c,bl) -> v 0
(hv 0 (str"<"++pr_constr p++str">"++ cut() ++ str"Case " ++
pr_constr c ++ str"of") ++ cut() ++
@@ -99,7 +99,7 @@ let pr_var_decl env (id,c,typ) =
let pbody = match c with
| None -> (mt ())
| Some c ->
- (* Force evaluation *)
+ (* Force evaluation *)
let pb = print_constr_env env c in
(str" := " ++ pb ++ cut () ) in
let pt = print_constr_env env typ in
@@ -110,7 +110,7 @@ let pr_rel_decl env (na,c,typ) =
let pbody = match c with
| None -> mt ()
| Some c ->
- (* Force evaluation *)
+ (* Force evaluation *)
let pb = print_constr_env env c in
(str":=" ++ spc () ++ pb ++ spc ()) in
let ptyp = print_constr_env env typ in
@@ -120,39 +120,39 @@ let pr_rel_decl env (na,c,typ) =
let print_named_context env =
hv 0 (fold_named_context
- (fun env d pps ->
+ (fun env d pps ->
pps ++ ws 2 ++ pr_var_decl env d)
env ~init:(mt ()))
-let print_rel_context env =
+let print_rel_context env =
hv 0 (fold_rel_context
(fun env d pps -> pps ++ ws 2 ++ pr_rel_decl env d)
env ~init:(mt ()))
-
+
let print_env env =
let sign_env =
fold_named_context
(fun env d pps ->
let pidt = pr_var_decl env d in
(pps ++ fnl () ++ pidt))
- env ~init:(mt ())
+ env ~init:(mt ())
in
let db_env =
fold_rel_context
(fun env d pps ->
let pnat = pr_rel_decl env d in (pps ++ fnl () ++ pnat))
env ~init:(mt ())
- in
+ in
(sign_env ++ db_env)
-
+
(*let current_module = ref empty_dirpath
let set_module m = current_module := m*)
-let new_univ =
+let new_univ =
let univ_gen = ref 0 in
(fun sp ->
- incr univ_gen;
+ incr univ_gen;
Univ.make_univ (Lib.library_dp(),!univ_gen))
let new_Type () = mkType (new_univ ())
let new_Type_sort () = Type (new_univ ())
@@ -173,26 +173,20 @@ let refresh_universes_gen strict t =
let refresh_universes = refresh_universes_gen false
let refresh_universes_strict = refresh_universes_gen true
-let new_sort_in_family = function
+let new_sort_in_family = function
| InProp -> prop_sort
| InSet -> set_sort
| InType -> Type (new_univ ())
-(* prod_it b [xn:Tn;..;x1:T1] = (x1:T1)..(xn:Tn)b *)
-let prod_it ~init = List.fold_left (fun c (n,t) -> mkProd (n, t, c)) init
-
-(* lam_it b [xn:Tn;..;x1:T1] = [x1:T1]..[xn:Tn]b *)
-let lam_it ~init = List.fold_left (fun c (n,t) -> mkLambda (n, t, c)) init
-
(* [Rel (n+m);...;Rel(n+1)] *)
let rel_vect n m = Array.init m (fun i -> mkRel(n+m-i))
-let rel_list n m =
- let rec reln l p =
+let rel_list n m =
+ let rec reln l p =
if p>m then l else reln (mkRel(n+p)::l) (p+1)
- in
+ in
reln [] 1
(* Same as [rel_list] but takes a context as argument and skips let-ins *)
@@ -201,7 +195,7 @@ let extended_rel_list n hyps =
| (_,None,_) :: hyps -> reln (mkRel (n+p) :: l) (p+1) hyps
| (_,Some _,_) :: hyps -> reln l (p+1) hyps
| [] -> l
- in
+ in
reln [] 1 hyps
let extended_rel_vect n hyps = Array.of_list (extended_rel_list n hyps)
@@ -224,53 +218,49 @@ let push_named_rec_types (lna,typarray,_) env =
Array.fold_left
(fun e assum -> push_named assum e) env ctxt
-let rec lookup_rel_id id sign =
+let rec lookup_rel_id id sign =
let rec lookrec = function
| (n, (Anonymous,_,_)::l) -> lookrec (n+1,l)
- | (n, (Name id',_,t)::l) -> if id' = id then (n,t) else lookrec (n+1,l)
+ | (n, (Name id',b,t)::l) -> if id' = id then (n,b,t) else lookrec (n+1,l)
| (_, []) -> raise Not_found
- in
+ in
lookrec (1,sign)
-(* Constructs either [(x:t)c] or [[x=b:t]c] *)
+(* Constructs either [forall x:t, c] or [let x:=b:t in c] *)
let mkProd_or_LetIn (na,body,t) c =
match body with
| None -> mkProd (na, t, c)
| Some b -> mkLetIn (na, b, t, c)
-(* Constructs either [(x:t)c] or [c] where [x] is replaced by [b] *)
+(* Constructs either [forall x:t, c] or [c] in which [x] is replaced by [b] *)
let mkProd_wo_LetIn (na,body,t) c =
match body with
| None -> mkProd (na, t, c)
| Some b -> subst1 b c
-let it_mkProd_wo_LetIn ~init =
- List.fold_left (fun c d -> mkProd_wo_LetIn d c) init
-
-let it_mkProd_or_LetIn ~init =
- List.fold_left (fun c d -> mkProd_or_LetIn d c) init
-
-let it_mkLambda_or_LetIn ~init =
- List.fold_left (fun c d -> mkLambda_or_LetIn d 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 =
+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_mkNamedLambda_or_LetIn = it_named_context_quantifier mkNamedLambda_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
(* *)
(* strips head casts and flattens head applications *)
let rec strip_head_cast c = match kind_of_term c with
- | App (f,cl) ->
+ | App (f,cl) ->
let rec collapse_rec f cl2 = match kind_of_term f with
| App (g,cl1) -> collapse_rec g (Array.append cl1 cl2)
| Cast (c,_,_) -> collapse_rec c cl2
| _ -> if Array.length cl2 = 0 then f else mkApp (f,cl2)
- in
+ in
collapse_rec f cl
| Cast (c,_,_) -> strip_head_cast c
| _ -> c
@@ -358,7 +348,7 @@ let map_constr_with_binders_left_to_right g f l c = match kind_of_term c with
let map_constr_with_full_binders g f l cstr = match kind_of_term cstr with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> cstr
- | Cast (c,k, t) ->
+ | Cast (c,k, t) ->
let c' = f l c in
let t' = f l t in
if c==c' && t==t' then cstr else mkCast (c', k, t')
@@ -422,7 +412,7 @@ let fold_constr_with_binders g f n acc c = match kind_of_term c with
| App (c,l) -> Array.fold_left (f n) (f n acc c) l
| Evar (_,l) -> Array.fold_left (f n) acc l
| Case (_,p,c,bl) -> Array.fold_left (f n) (f n (f n acc p) c) bl
- | Fix (_,(lna,tl,bl)) ->
+ | Fix (_,(lna,tl,bl)) ->
let n' = iterate g (Array.length tl) n in
let fd = array_map2 (fun t b -> (t,b)) tl bl in
Array.fold_left (fun acc (t,b) -> f n' (f n acc t) b) acc fd
@@ -446,7 +436,7 @@ let iter_constr_with_full_binders g f l c = match kind_of_term c with
| App (c,args) -> f l c; Array.iter (f l) args
| Evar (_,args) -> Array.iter (f l) args
| Case (_,p,c,bl) -> f l p; f l c; Array.iter (f l) bl
- | Fix (_,(lna,tl,bl)) ->
+ | Fix (_,(lna,tl,bl)) ->
let l' = array_fold_left2 (fun l na t -> g (na,None,t) l) l lna tl in
Array.iter (f l) tl;
Array.iter (f l') bl
@@ -456,7 +446,7 @@ let iter_constr_with_full_binders g f l c = match kind_of_term c with
Array.iter (f l') bl
(***************************)
-(* occurs check functions *)
+(* occurs check functions *)
(***************************)
exception Occur
@@ -467,42 +457,43 @@ let occur_meta c =
| _ -> iter_constr occrec c
in try occrec c; false with Occur -> true
-let occur_existential c =
+let occur_existential c =
let rec occrec c = match kind_of_term c with
| Evar _ -> raise Occur
| _ -> iter_constr occrec c
in try occrec c; false with Occur -> true
-let occur_meta_or_existential c =
+let occur_meta_or_existential c =
let rec occrec c = match kind_of_term c with
| Evar _ -> raise Occur
| Meta _ -> raise Occur
| _ -> iter_constr occrec c
in try occrec c; false with Occur -> true
-let occur_const s c =
+let occur_const s c =
let rec occur_rec c = match kind_of_term c with
| Const sp when sp=s -> raise Occur
| _ -> iter_constr occur_rec c
- in
+ in
try occur_rec c; false with Occur -> true
-let occur_evar n c =
+let occur_evar n c =
let rec occur_rec c = match kind_of_term c with
| Evar (sp,_) when sp=n -> raise Occur
| _ -> iter_constr occur_rec c
- in
+ in
try occur_rec c; false with Occur -> true
let occur_in_global env id constr =
let vars = vars_of_global env constr in
if List.mem id vars then raise Occur
-let occur_var env s c =
+let occur_var env id c =
let rec occur_rec c =
- occur_in_global env s c;
- iter_constr occur_rec c
- in
+ match kind_of_term c with
+ | Var _ | Const _ | Ind _ | Construct _ -> occur_in_global env id c
+ | _ -> iter_constr occur_rec c
+ in
try occur_rec c; false with Occur -> true
let occur_var_in_decl env hyp (_,c,typ) =
@@ -512,25 +503,19 @@ let occur_var_in_decl env hyp (_,c,typ) =
occur_var env hyp typ ||
occur_var env hyp body
-(* Tests that t is a subterm of c *)
-let occur_term t c =
- let eq_constr_fail c = if eq_constr t c then raise Occur
- in let rec occur_rec c = eq_constr_fail c; iter_constr occur_rec c
- in try occur_rec c; false with Occur -> true
-
(* returns the list of free debruijn indices in a term *)
-let free_rels m =
+let free_rels m =
let rec frec depth acc c = match kind_of_term c with
| Rel n -> if n >= depth then Intset.add (n-depth+1) acc else acc
| _ -> fold_constr_with_binders succ frec depth acc c
- in
+ in
frec 1 Intset.empty m
(* collects all metavar occurences, in left-to-right order, preserving
* repetitions and all. *)
-let collect_metas c =
+let collect_metas c =
let rec collrec acc c =
match kind_of_term c with
| Meta mv -> list_add_set mv acc
@@ -538,10 +523,10 @@ let collect_metas c =
in
List.rev (collrec [] c)
-(* (dependent M N) is true iff M is eq_term with a subterm of N
- M is appropriately lifted through abstractions of N *)
+(* Tests whether [m] is a subterm of [t]:
+ [m] is appropriately lifted through abstractions of [t] *)
-let dependent m t =
+let dependent_main noevar m t =
let rec deprec m t =
if eq_constr m t then
raise Occur
@@ -550,28 +535,38 @@ let dependent m t =
| App (fm,lm), App (ft,lt) when Array.length lm < Array.length lt ->
deprec m (mkApp (ft,Array.sub lt 0 (Array.length lm)));
Array.iter (deprec m)
- (Array.sub lt
+ (Array.sub lt
(Array.length lm) ((Array.length lt) - (Array.length lm)))
+ | _, Cast (c,_,_) when noevar & isMeta c -> ()
+ | _, Evar _ when noevar -> ()
| _ -> iter_constr_with_binders (lift 1) deprec m t
- in
+ in
try deprec m t; false with Occur -> true
+let dependent = dependent_main false
+let dependent_no_evar = dependent_main true
+
+(* Synonymous *)
+let occur_term = dependent
+
let pop t = lift (-1) t
(***************************)
-(* bindings functions *)
+(* bindings functions *)
(***************************)
-type metamap = (metavariable * constr) list
+type meta_type_map = (metavariable * types) list
+
+type meta_value_map = (metavariable * constr) list
-let rec subst_meta bl c =
+let rec subst_meta bl c =
match kind_of_term c with
| Meta i -> (try List.assoc i bl with Not_found -> c)
| _ -> map_constr (subst_meta bl) c
(* First utilities for avoiding telescope computation for subst_term *)
-let prefix_application eq_fun (k,c) (t : constr) =
+let prefix_application eq_fun (k,c) (t : constr) =
let c' = collapse_appl c and t' = collapse_appl t in
match kind_of_term c', kind_of_term t' with
| App (f1,cl1), App (f2,cl2) ->
@@ -580,11 +575,11 @@ let prefix_application eq_fun (k,c) (t : constr) =
if l1 <= l2
&& eq_fun c' (mkApp (f2, Array.sub cl2 0 l1)) then
Some (mkApp (mkRel k, Array.sub cl2 l1 (l2 - l1)))
- else
+ else
None
| _ -> None
-let my_prefix_application eq_fun (k,c) (by_c : constr) (t : constr) =
+let my_prefix_application eq_fun (k,c) (by_c : constr) (t : constr) =
let c' = collapse_appl c and t' = collapse_appl t in
match kind_of_term c', kind_of_term t' with
| App (f1,cl1), App (f2,cl2) ->
@@ -593,7 +588,7 @@ let my_prefix_application eq_fun (k,c) (by_c : constr) (t : constr) =
if l1 <= l2
&& eq_fun c' (mkApp (f2, Array.sub cl2 0 l1)) then
Some (mkApp ((lift k by_c), Array.sub cl2 l1 (l2 - l1)))
- else
+ else
None
| _ -> None
@@ -602,7 +597,7 @@ let my_prefix_application eq_fun (k,c) (by_c : constr) (t : constr) =
term [c] in a term [t] *)
(*i Bizarre : si on cherche un sous terme clos, pourquoi le lifter ? i*)
-let subst_term_gen eq_fun c t =
+let subst_term_gen eq_fun c t =
let rec substrec (k,c as kc) t =
match prefix_application eq_fun kc t with
| Some x -> x
@@ -610,7 +605,7 @@ let subst_term_gen eq_fun c t =
if eq_fun c t then mkRel k
else
map_constr_with_binders (fun (k,c) -> (k+1,lift 1 c)) substrec kc t
- in
+ in
substrec (1,c) t
(* Recognizing occurrences of a given (closed) subterm in a term :
@@ -618,7 +613,7 @@ let subst_term_gen eq_fun c t =
term [c1] in a term [t] *)
(*i Meme remarque : a priori [c] n'est pas forcement clos i*)
-let replace_term_gen eq_fun c by_c in_t =
+let replace_term_gen eq_fun c by_c in_t =
let rec substrec (k,c as kc) t =
match my_prefix_application eq_fun kc by_c t with
| Some x -> x
@@ -626,7 +621,7 @@ let replace_term_gen eq_fun c by_c in_t =
(if eq_fun c t then (lift k by_c) else
map_constr_with_binders (fun (k,c) -> (k+1,lift 1 c))
substrec kc t)
- in
+ in
substrec (0,c) in_t
let subst_term = subst_term_gen eq_constr
@@ -645,7 +640,7 @@ let no_occurrences_in_set = (true,[])
let error_invalid_occurrence l =
let l = list_uniquize (List.sort Pervasives.compare l) in
errorlabstrm ""
- (str ("Invalid occurrence " ^ plural (List.length l) "number" ^": ") ++
+ (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 =
@@ -656,10 +651,10 @@ let subst_term_occ_gen (nowhere_except_in,locs) occ c t =
if nowhere_except_in & !pos > maxocc then t
else
if eq_constr c t then
- let r =
+ let r =
if nowhere_except_in then
if List.mem !pos locs then (mkRel k) else t
- else
+ else
if List.mem !pos locs then t else (mkRel k)
in incr pos; r
else
@@ -670,9 +665,9 @@ let subst_term_occ_gen (nowhere_except_in,locs) occ c t =
let t' = substrec (1,c) t in
(!pos, t')
-let subst_term_occ (nowhere_except_in,locs as plocs) c 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
+ 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;
@@ -693,20 +688,15 @@ let subst_term_occ_decl ((nowhere_except_in,locs as plocs),hloc) c (id,bodyopt,t
if locs = [] then
if nowhere_except_in then d
else (id,Some (subst_term c body),subst_term c typ)
- else
+ 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')
-(* First character of a constr *)
-
-let lowercase_first_char id =
- lowercase_first_char_utf8 (string_of_id id)
-
let vars_of_env env =
- let s =
+ let s =
Sign.fold_named_context (fun (id,_,_) s -> Idset.add id s)
(named_context env) ~init:Idset.empty in
Sign.fold_rel_context
@@ -717,85 +707,6 @@ let add_vname vars = function
Name id -> Idset.add id vars
| _ -> vars
-let id_of_global = Nametab.id_of_global
-
-let sort_hdchar = function
- | Prop(_) -> "P"
- | Type(_) -> "T"
-
-let hdchar env c =
- let rec hdrec k c =
- match kind_of_term c with
- | Prod (_,_,c) -> hdrec (k+1) c
- | Lambda (_,_,c) -> hdrec (k+1) c
- | LetIn (_,_,_,c) -> hdrec (k+1) c
- | Cast (c,_,_) -> hdrec k c
- | App (f,l) -> hdrec k f
- | Const kn ->
- lowercase_first_char (id_of_label (con_label kn))
- | Ind ((kn,i) as x) ->
- if i=0 then
- lowercase_first_char (id_of_label (label kn))
- else
- lowercase_first_char (id_of_global (IndRef x))
- | Construct ((sp,i) as x) ->
- lowercase_first_char (id_of_global (ConstructRef x))
- | Var id -> lowercase_first_char id
- | Sort s -> sort_hdchar s
- | Rel n ->
- (if n<=k then "p" (* the initial term is flexible product/function *)
- else
- try match Environ.lookup_rel (n-k) env with
- | (Name id,_,_) -> lowercase_first_char id
- | (Anonymous,_,t) -> hdrec 0 (lift (n-k) t)
- with Not_found -> "y")
- | Fix ((_,i),(lna,_,_)) ->
- let id = match lna.(i) with Name id -> id | _ -> assert false in
- lowercase_first_char id
- | CoFix (i,(lna,_,_)) ->
- let id = match lna.(i) with Name id -> id | _ -> assert false in
- lowercase_first_char id
- | Meta _|Evar _|Case (_, _, _, _) -> "y"
- in
- hdrec 0 c
-
-let id_of_name_using_hdchar env a = function
- | Anonymous -> id_of_string (hdchar env a)
- | Name id -> id
-
-let named_hd env a = function
- | Anonymous -> Name (id_of_string (hdchar env a))
- | x -> x
-
-let mkProd_name env (n,a,b) = mkProd (named_hd env a n, a, b)
-let mkLambda_name env (n,a,b) = mkLambda (named_hd env a n, a, b)
-
-let lambda_name = mkLambda_name
-let prod_name = mkProd_name
-
-let prod_create env (a,b) = mkProd (named_hd env a Anonymous, a, b)
-let lambda_create env (a,b) = mkLambda (named_hd env a Anonymous, a, b)
-
-let name_assumption env (na,c,t) =
- match c with
- | None -> (named_hd env t na, None, t)
- | Some body -> (named_hd env body na, c, t)
-
-let name_context env hyps =
- snd
- (List.fold_left
- (fun (env,hyps) d ->
- let d' = name_assumption env d in (push_rel d' env, d' :: hyps))
- (env,[]) (List.rev hyps))
-
-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 b (name_context env hyps)
-let it_mkLambda_or_LetIn_name env b hyps =
- it_mkLambda_or_LetIn b (name_context env hyps)
-
(*************************)
(* Names environments *)
(*************************)
@@ -804,12 +715,12 @@ let add_name n nl = n::nl
let lookup_name_of_rel p names =
try List.nth names (p-1)
with Invalid_argument _ | Failure _ -> raise Not_found
-let rec lookup_rel_of_name id names =
+let rec lookup_rel_of_name id names =
let rec lookrec n = function
| Anonymous :: l -> lookrec (n+1) l
| (Name id') :: l -> if id' = id then n else lookrec (n+1) l
| [] -> raise Not_found
- in
+ in
lookrec 1 names
let empty_names_context = []
@@ -821,7 +732,7 @@ let ids_of_rel_context sign =
let ids_of_named_context sign =
Sign.fold_named_context (fun (id,_,_) idl -> id::idl) sign ~init:[]
-let ids_of_context env =
+let ids_of_context env =
(ids_of_rel_context (rel_context env))
@ (ids_of_named_context (named_context env))
@@ -829,57 +740,11 @@ let ids_of_context env =
let names_of_rel_context env =
List.map (fun (na,_,_) -> na) (rel_context env)
-(**** Globality of identifiers *)
-
-let rec is_imported_modpath = function
- | MPfile dp -> true
- | p -> false
-
-let is_imported_ref = function
- | VarRef _ -> false
- | IndRef (kn,_)
- | ConstructRef ((kn,_),_) ->
- let (mp,_,_) = repr_kn kn in is_imported_modpath mp
- | ConstRef kn ->
- let (mp,_,_) = repr_con kn in is_imported_modpath mp
-
-let is_global id =
- try
- let ref = locate (make_short_qualid id) in
- not (is_imported_ref ref)
- with Not_found ->
- false
-
-let is_constructor id =
- try
- match locate (make_short_qualid id) with
- | ConstructRef _ as ref -> not (is_imported_ref ref)
- | _ -> false
- with Not_found ->
- false
-
let is_section_variable id =
try let _ = Global.lookup_named id in true
with Not_found -> false
-let next_global_ident_from allow_secvar id avoid =
- let rec next_rec id =
- let id = next_ident_away_from id avoid in
- if (allow_secvar && is_section_variable id) || not (is_global id) then
- id
- else
- next_rec (lift_ident id)
- in
- next_rec id
-
-let next_global_ident_away allow_secvar id avoid =
- let id = next_ident_away id avoid in
- if (allow_secvar && is_section_variable id) || not (is_global id) then
- id
- else
- next_global_ident_from allow_secvar (lift_ident id) avoid
-
-let isGlobalRef c =
+let isGlobalRef c =
match kind_of_term c with
| Const _ | Ind _ | Construct _ | Var _ -> true
| _ -> false
@@ -889,68 +754,6 @@ let has_polymorphic_type c =
| Declarations.PolymorphicArity _ -> true
| _ -> false
-(* nouvelle version de renommage des variables (DEC 98) *)
-(* This is the algorithm to display distinct bound variables
-
- - Règle 1 : un nom non anonyme, même non affiché, contribue à la liste
- des noms à éviter
- - Règle 2 : c'est la dépendance qui décide si on affiche ou pas
-
- Exemple :
- si bool_ind = (P:bool->Prop)(f:(P true))(f:(P false))(b:bool)(P b), alors
- il est affiché (P:bool->Prop)(P true)->(P false)->(b:bool)(P b)
- mais f et f0 contribue à la liste des variables à éviter (en supposant
- que les noms f et f0 ne sont pas déjà pris)
- Intérêt : noms homogènes dans un but avant et après Intro
-*)
-
-type used_idents = identifier list
-
-let occur_rel p env id =
- try lookup_name_of_rel p env = Name id
- with Not_found -> false (* Unbound indice : may happen in debug *)
-
-let occur_id nenv id0 c =
- let rec occur n c = match kind_of_term c with
- | Var id when id=id0 -> raise Occur
- | Const kn when id_of_global (ConstRef kn) = id0 -> raise Occur
- | Ind ind_sp
- when id_of_global (IndRef ind_sp) = id0 ->
- raise Occur
- | Construct cstr_sp
- when id_of_global (ConstructRef cstr_sp) = id0 ->
- raise Occur
- | Rel p when p>n & occur_rel (p-n) nenv id0 -> raise Occur
- | _ -> iter_constr_with_binders succ occur n c
- in
- try occur 1 c; false
- with Occur -> true
- | Not_found -> false (* Case when a global is not in the env *)
-
-type avoid_flags = bool option
-
-let next_name_not_occuring avoid_flags name l env_names t =
- let rec next id =
- if List.mem id l or occur_id env_names id t or
- (* Further restrictions ? *)
- match avoid_flags with None -> false | Some not_only_cstr ->
- (if not_only_cstr then
- (* To be consistent with the intro mechanism *)
- is_global id & not (is_section_variable id)
- else
- (* To avoid constructors in pattern-matchings *)
- is_constructor id)
- then next (lift_ident id)
- else id
- in
- match name with
- | Name id -> next id
- | Anonymous ->
- (* Normally, an anonymous name is not dependent and will not be *)
- (* taken into account by the function concrete_name; just in case *)
- (* invent a valid name *)
- next (id_of_string "H")
-
let base_sort_cmp pb s0 s1 =
match (s0,s1) with
| (Prop c1, Prop c2) -> c1 = Null or c2 = Pos (* Prop <= Set *)
@@ -959,14 +762,77 @@ let base_sort_cmp pb s0 s1 =
| _ -> false
(* eq_constr extended with universe erasure *)
-let rec constr_cmp cv_pb t1 t2 =
- (match kind_of_term t1, kind_of_term t2 with
+let compare_constr_univ f cv_pb t1 t2 =
+ match kind_of_term t1, kind_of_term t2 with
Sort s1, Sort s2 -> base_sort_cmp cv_pb s1 s2
- | _ -> false)
- || compare_constr (constr_cmp cv_pb) t1 t2
+ | Prod (_,t1,c1), Prod (_,t2,c2) ->
+ f Reduction.CONV t1 t2 & f cv_pb c1 c2
+ | _ -> compare_constr (f Reduction.CONV) t1 t2
+
+let rec constr_cmp cv_pb t1 t2 = compare_constr_univ constr_cmp cv_pb t1 t2
let eq_constr = constr_cmp Reduction.CONV
+(* App(c,[t1,...tn]) -> ([c,t1,...,tn-1],tn)
+ App(c,[||]) -> ([],c) *)
+let split_app c = match kind_of_term c with
+ App(c,l) ->
+ let len = Array.length l in
+ if len=0 then ([],c) else
+ let last = Array.get l (len-1) in
+ let prev = Array.sub l 0 (len-1) in
+ c::(Array.to_list prev), last
+ | _ -> assert false
+
+let hdtl l = List.hd l, List.tl l
+
+type subst = (rel_context*constr) Intmap.t
+
+exception CannotFilter
+
+let filtering env cv_pb c1 c2 =
+ let evm = ref Intmap.empty in
+ let define cv_pb e1 ev c1 =
+ try let (e2,c2) = Intmap.find ev !evm in
+ let shift = List.length e1 - List.length e2 in
+ if constr_cmp cv_pb c1 (lift shift c2) then () else raise CannotFilter
+ with Not_found ->
+ evm := Intmap.add ev (e1,c1) !evm
+ in
+ let rec aux env cv_pb c1 c2 =
+ match kind_of_term c1, kind_of_term c2 with
+ | App _, App _ ->
+ let ((p1,l1),(p2,l2)) = (split_app c1),(split_app c2) in
+ aux env cv_pb l1 l2; if p1=[] & p2=[] then () else
+ aux env cv_pb (applist (hdtl p1)) (applist (hdtl p2))
+ | Prod (n,t1,c1), Prod (_,t2,c2) ->
+ aux env cv_pb t1 t2;
+ aux ((n,None,t1)::env) cv_pb c1 c2
+ | _, Evar (ev,_) -> define cv_pb env ev c1
+ | Evar (ev,_), _ -> define cv_pb env ev c2
+ | _ ->
+ if compare_constr_univ
+ (fun pb c1 c2 -> aux env pb c1 c2; true) cv_pb c1 c2 then ()
+ else raise CannotFilter
+ (* TODO: le reste des binders *)
+ in
+ aux env cv_pb c1 c2; !evm
+
+let decompose_prod_letin : constr -> int * rel_context * constr =
+ let rec prodec_rec i l c = match kind_of_term c with
+ | Prod (n,t,c) -> prodec_rec (succ i) ((n,None,t)::l) c
+ | LetIn (n,d,t,c) -> prodec_rec (succ i) ((n,Some d,t)::l) c
+ | Cast (c,_,_) -> prodec_rec i l c
+ | _ -> i,l,c in
+ prodec_rec 0 []
+
+let align_prod_letin c a : rel_context * constr =
+ let (lc,_,_) = decompose_prod_letin c in
+ let (la,l,a) = decompose_prod_letin a in
+ if not (la >= lc) then invalid_arg "align_prod_letin";
+ let (l1,l2) = Util.list_chop lc l in
+ l2,it_mkProd_or_LetIn a l1
+
(* On reduit une serie d'eta-redex de tete ou rien du tout *)
(* [x1:c1;...;xn:cn]@(f;a1...an;x1;...;xn) --> @(f;a1...an) *)
(* Remplace 2 versions précédentes buggées *)
@@ -976,7 +842,7 @@ let rec eta_reduce_head c =
| Lambda (_,c1,c') ->
(match kind_of_term (eta_reduce_head c') with
| App (f,cl) ->
- let lastn = (Array.length cl) - 1 in
+ let lastn = (Array.length cl) - 1 in
if lastn < 1 then anomaly "application without arguments"
else
(match kind_of_term cl.(lastn) with
@@ -1017,7 +883,7 @@ let assums_of_rel_context sign =
| None -> (na, t)::l)
sign ~init:[]
-let fold_map_rel_context f env sign =
+let map_rel_context_in_env f env sign =
let rec aux env acc = function
| d::sign ->
aux (push_rel d env) (map_rel_declaration (f env) d :: acc) sign
@@ -1039,6 +905,25 @@ let substl_rel_context l =
let lift_rel_context n =
map_rel_context_with_binders (liftn n)
+let smash_rel_context sign =
+ let rec aux acc = function
+ | [] -> acc
+ | (_,None,_ as d) :: l -> aux (d::acc) l
+ | (_,Some b,_) :: l ->
+ (* Quadratic in the number of let but there are probably a few of them *)
+ aux (List.rev (substl_rel_context [b] (List.rev acc))) l
+ in List.rev (aux [] sign)
+
+let adjust_subst_to_rel_context sign l =
+ let rec aux subst sign l =
+ match sign, l with
+ | (_,None,_)::sign', a::args' -> aux (a::subst) sign' args'
+ | (_,Some c,_)::sign', args' ->
+ aux (substl (List.rev subst) c :: subst) sign' args'
+ | [], [] -> List.rev subst
+ | _ -> anomaly "Instance and signature do not match"
+ in aux [] (List.rev sign) l
+
let fold_named_context_both_sides f l ~init = list_fold_right_and_left f l init
let rec mem_named_context id = function
@@ -1046,14 +931,11 @@ let rec mem_named_context id = function
| _ :: sign -> mem_named_context id sign
| [] -> false
-let make_all_name_different env =
- let avoid = ref (ids_of_named_context (named_context env)) in
- process_rel_context
- (fun (na,c,t) newenv ->
- let id = next_name_away na !avoid in
- avoid := id::!avoid;
- push_rel (Name id,c,t) newenv)
- env
+let clear_named_body id env =
+ let rec aux _ = function
+ | (id',Some c,t) when id = id' -> push_named (id,None,t)
+ | d -> push_named d in
+ fold_named_context aux env ~init:(reset_context env)
let global_vars env ids = Idset.elements (global_vars_set env ids)
@@ -1076,50 +958,13 @@ let dependency_closure env sign hyps =
sign in
List.rev lh
-let default_x = id_of_string "x"
-
-let rec next_name_away_in_cases_pattern id avoid =
- let id = match id with Name id -> id | Anonymous -> default_x in
- let rec next id =
- if List.mem id avoid or is_constructor id then next (lift_ident id)
- else id in
- next id
-
-(* Remark: Anonymous var may be dependent in Evar's contexts *)
-let concrete_name avoid_flags l env_names n c =
- if n = Anonymous & noccurn 1 c then
- (Anonymous,l)
- else
- let fresh_id = next_name_not_occuring avoid_flags n l env_names c in
- let idopt = if noccurn 1 c then Anonymous else Name fresh_id in
- (idopt, fresh_id::l)
-
-let concrete_let_name avoid_flags l env_names n c =
- let fresh_id = next_name_not_occuring avoid_flags n l env_names c in
- (Name fresh_id, fresh_id::l)
-
-let rec rename_bound_var env avoid c =
- let env_names = names_of_rel_context env in
- let rec rename avoid c =
- match kind_of_term c with
- | Prod (na,c1,c2) ->
- let na',avoid' = concrete_name None avoid env_names na c2 in
- mkProd (na', c1, rename avoid' c2)
- | LetIn (na,c1,t,c2) ->
- let na',avoid' = concrete_let_name None avoid env_names na c2 in
- mkLetIn (na',c1,t, rename avoid' c2)
- | Cast (c,k,t) -> mkCast (rename avoid c, k,t)
- | _ -> c
- in
- rename avoid c
-
(* Combinators on judgments *)
let on_judgment f j = { uj_val = f j.uj_val; uj_type = f j.uj_type }
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
+(* Cut a context ctx in 2 parts (ctx1,ctx2) with ctx1 containing k
variables *)
let context_chop k ctx =
let rec chop_aux acc = function
diff --git a/pretyping/termops.mli b/pretyping/termops.mli
index e0bbe7b5..f9ea7b22 100644
--- a/pretyping/termops.mli
+++ b/pretyping/termops.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: termops.mli 12058 2009-04-08 10:54:59Z herbelin $ i*)
+(*i $Id$ i*)
open Util
open Pp
@@ -35,27 +35,32 @@ val pr_rel_decl : env -> rel_declaration -> std_ppcmds
val print_rel_context : env -> std_ppcmds
val print_env : env -> std_ppcmds
-(* iterators on terms *)
-val prod_it : init:types -> (name * types) list -> types
-val lam_it : init:constr -> (name * types) list -> constr
+(* about contexts *)
+val push_rel_assum : name * types -> env -> env
+val push_rels_assum : (name * types) list -> env -> env
+val push_named_rec_types : name array * types array * 'a -> env -> env
+val lookup_rel_id : identifier -> rel_context -> int * constr option * types
+
+(* builds argument lists matching a block of binders or a context *)
val 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
-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 * types
+
+(* iterators/destructors on terms *)
val mkProd_or_LetIn : rel_declaration -> types -> types
val mkProd_wo_LetIn : rel_declaration -> types -> types
-val it_mkProd_wo_LetIn : init:types -> rel_context -> 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_named_context_quantifier :
- (named_declaration -> 'a -> 'a) -> init:'a -> named_context -> 'a
val it_mkNamedProd_or_LetIn : init:types -> named_context -> types
-val it_mkNamedLambda_or_LetIn : init:constr -> named_context -> constr
val it_mkNamedProd_wo_LetIn : init:types -> named_context -> types
+val it_mkNamedLambda_or_LetIn : init:constr -> named_context -> constr
+
+val it_named_context_quantifier :
+ (named_declaration -> 'a -> 'a) -> init:'a -> named_context -> 'a
(**********************************************************************)
(* Generic iterators on constr *)
@@ -64,7 +69,7 @@ val map_constr_with_named_binders :
(name -> 'a -> 'a) ->
('a -> constr -> constr) -> 'a -> constr -> constr
val map_constr_with_binders_left_to_right :
- (rel_declaration -> 'a -> 'a) ->
+ (rel_declaration -> 'a -> 'a) ->
('a -> constr -> constr) ->
'a -> constr -> constr
val map_constr_with_full_binders :
@@ -82,7 +87,7 @@ val fold_constr_with_binders :
('a -> 'a) -> ('a -> 'b -> constr -> 'b) -> 'a -> 'b -> constr -> 'b
val iter_constr_with_full_binders :
- (rel_declaration -> 'a -> 'a) -> ('a -> constr -> unit) -> 'a ->
+ (rel_declaration -> 'a -> 'a) -> ('a -> constr -> unit) -> 'a ->
constr -> unit
(**********************************************************************)
@@ -96,18 +101,22 @@ val occur_existential : types -> bool
val occur_meta_or_existential : types -> bool
val occur_const : constant -> types -> bool
val occur_evar : existential_key -> types -> bool
-val occur_in_global : env -> identifier -> constr -> unit
val occur_var : env -> identifier -> types -> bool
val occur_var_in_decl :
env ->
identifier -> 'a * types option * types -> bool
-val occur_term : constr -> constr -> bool
val free_rels : constr -> Intset.t
val dependent : constr -> constr -> bool
+val dependent_no_evar : constr -> constr -> bool
val collect_metas : constr -> int list
+val occur_term : constr -> constr -> bool (* Synonymous
+ of dependent *)
(* Substitution of metavariables *)
-type metamap = (metavariable * constr) list
-val subst_meta : metamap -> constr -> constr
+type meta_value_map = (metavariable * constr) list
+val subst_meta : meta_value_map -> constr -> constr
+
+(* Type assignment for metavariables *)
+type meta_type_map = (metavariable * types) list
(* [pop c] lifts by -1 the positive indexes in [c] *)
val pop : constr -> constr
@@ -139,7 +148,7 @@ val no_occurrences_in_set : occurrences
(* [subst_term_occ_gen occl n c d] replaces occurrences of [c] at
positions [occl], counting from [n], by [Rel 1] in [d] *)
-val subst_term_occ_gen :
+val subst_term_occ_gen :
occurrences -> int -> constr -> types -> int * types
(* [subst_term_occ occl c d] replaces occurrences of [c] at
@@ -155,43 +164,34 @@ type hyp_location_flag = (* To distinguish body and type of local defs *)
| InHypValueOnly
val subst_term_occ_decl :
- occurrences * hyp_location_flag -> constr -> named_declaration ->
+ occurrences * hyp_location_flag -> constr -> named_declaration ->
named_declaration
val error_invalid_occurrence : int list -> 'a
(* Alternative term equalities *)
val base_sort_cmp : Reduction.conv_pb -> sorts -> sorts -> bool
+val compare_constr_univ : (Reduction.conv_pb -> constr -> constr -> bool) ->
+ Reduction.conv_pb -> constr -> constr -> bool
val constr_cmp : Reduction.conv_pb -> constr -> constr -> bool
val eq_constr : constr -> constr -> bool
val eta_reduce_head : constr -> constr
val eta_eq_constr : constr -> constr -> bool
-(* finding "intuitive" names to hypotheses *)
-val lowercase_first_char : identifier -> string
-val sort_hdchar : sorts -> string
-val hdchar : env -> types -> string
-val id_of_name_using_hdchar :
- env -> types -> name -> identifier
-val named_hd : env -> types -> name -> name
-
-val mkProd_name : env -> name * types * types -> types
-val mkLambda_name : env -> name * types * constr -> constr
-
-(* Deprecated synonyms of [mkProd_name] and [mkLambda_name] *)
-val prod_name : env -> name * types * types -> types
-val lambda_name : env -> name * types * constr -> constr
+exception CannotFilter
-val prod_create : env -> types * types -> constr
-val lambda_create : env -> types * constr -> constr
-val name_assumption : env -> rel_declaration -> rel_declaration
-val name_context : env -> rel_context -> rel_context
+(* 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].
+ Warning: Outer-kernel sort subtyping are taken into account: c1 has
+ to be smaller than c2 wrt. sorts. *)
+type subst = (rel_context*constr) Intmap.t
+val filtering : rel_context -> Reduction.conv_pb -> constr -> constr -> subst
-val mkProd_or_LetIn_name : env -> types -> rel_declaration -> types
-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
+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 *)
val last_arg : constr -> constr
@@ -213,43 +213,23 @@ val context_chop : int -> rel_context -> (rel_context*rel_context)
val vars_of_env: env -> Idset.t
val add_vname : Idset.t -> name -> Idset.t
-(* sets of free identifiers *)
-type used_idents = identifier list
-val occur_rel : int -> name list -> identifier -> bool
-val occur_id : name list -> identifier -> constr -> bool
-
-type avoid_flags = bool option
- (* Some true = avoid all globals (as in intro);
- Some false = avoid only global constructors; None = don't avoid globals *)
-
-val next_name_away_in_cases_pattern :
- name -> identifier list -> identifier
-val next_global_ident_away :
- (*allow section vars:*) bool -> identifier -> identifier list -> identifier
-val next_name_not_occuring :
- avoid_flags -> name -> identifier list -> name list -> constr -> identifier
-val concrete_name :
- avoid_flags -> identifier list -> name list -> name -> constr ->
- name * identifier list
-val concrete_let_name :
- avoid_flags -> identifier list -> name list -> name -> constr ->
- name * identifier list
-val rename_bound_var : env -> identifier list -> types -> types
-
(* 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 fold_map_rel_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
-val map_rel_context_with_binders :
+val map_rel_context_with_binders :
(int -> constr -> constr) -> rel_context -> rel_context
val fold_named_context_both_sides :
('a -> named_declaration -> named_declaration list -> 'a) ->
named_context -> init:'a -> 'a
val mem_named_context : identifier -> named_context -> bool
-val make_all_name_different : env -> env
+
+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
diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml
index 216a0611..b85c6721 100644
--- a/pretyping/typeclasses.ml
+++ b/pretyping/typeclasses.ml
@@ -7,7 +7,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: typeclasses.ml 12189 2009-06-15 05:08:44Z msozeau $ i*)
+(*i $Id$ i*)
(*i*)
open Names
@@ -21,8 +21,20 @@ open Nametab
open Mod_subst
open Util
open Typeclasses_errors
+open Libobject
(*i*)
+
+let add_instance_hint_ref = ref (fun id 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 register_set_typeclass_transparency =
+ (:=) set_typeclass_transparency_ref
+let set_typeclass_transparency gr c = !set_typeclass_transparency_ref gr c
+
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
@@ -32,284 +44,265 @@ type rels = constr list
(* This module defines type-classes *)
type typeclass = {
(* The class implementation *)
- cl_impl : global_reference;
+ cl_impl : global_reference;
(* Context in which the definitions are typed. Includes both typeclass parameters and superclasses. *)
- cl_context : (global_reference * bool) option list * rel_context;
+ cl_context : (global_reference * bool) option list * rel_context;
(* Context of definitions and properties on defs, will not be shared *)
cl_props : rel_context;
-
+
(* The method implementaions as projections. *)
cl_projs : (identifier * constant option) list;
}
+module Gmap = Fmap.Make(RefOrdered)
-type typeclasses = (global_reference, typeclass) Gmap.t
+type typeclasses = typeclass Gmap.t
type instance = {
is_class: global_reference;
is_pri: int option;
- (* Sections where the instance should be redeclared,
- -1 for discard, 0 for none, mutable to avoid redeclarations
+ (* Sections where the instance should be redeclared,
+ -1 for discard, 0 for none, mutable to avoid redeclarations
when multiple rebuild_object happen. *)
- is_global: int ref;
- is_impl: constant;
+ is_global: int;
+ is_impl: global_reference;
}
-type instances = (global_reference, instance Cmap.t) Gmap.t
+type instances = (instance Gmap.t) Gmap.t
let instance_impl is = is.is_impl
-let new_instance cl pri glob impl =
+let new_instance cl pri glob impl =
let global =
- if Lib.sections_are_opened () then
- if glob then Lib.sections_depth ()
- else -1
- else 0
+ if glob then Lib.sections_depth ()
+ else -1
in
{ is_class = cl.cl_impl;
is_pri = pri ;
- is_global = ref global ;
+ is_global = global ;
is_impl = impl }
-
+
+(*
+ * states management
+ *)
+
let classes : typeclasses ref = ref Gmap.empty
-let methods : (constant, global_reference) Gmap.t ref = ref Gmap.empty
-
let instances : instances ref = ref Gmap.empty
-
-let freeze () = !classes, !methods, !instances
-let unfreeze (cl,m,is) =
+let freeze () = !classes, !instances
+
+let unfreeze (cl,is) =
classes:=cl;
- methods:=m;
instances:=is
-
+
let init () =
- classes:= Gmap.empty;
- methods:= Gmap.empty;
+ classes:= Gmap.empty;
instances:= Gmap.empty
-
-let _ =
+
+let _ =
Summary.declare_summary "classes_and_instances"
{ Summary.freeze_function = freeze;
Summary.unfreeze_function = unfreeze;
- Summary.init_function = init;
- Summary.survive_module = false;
- Summary.survive_section = true }
-
-let gmap_merge old ne =
- Gmap.fold (fun k v acc -> Gmap.add k v acc) old ne
+ Summary.init_function = init }
-let cmap_union = Cmap.fold Cmap.add
+(*
+ * classes persistent object
+ *)
-let gmap_cmap_merge old ne =
- let ne' =
- Gmap.fold (fun cl insts acc ->
- let oldinsts = try Gmap.find cl old with Not_found -> Cmap.empty in
- Gmap.add cl (cmap_union oldinsts insts) acc)
- ne Gmap.empty
- in
- Gmap.fold (fun cl insts acc ->
- if Gmap.mem cl acc then acc
- else Gmap.add cl insts acc)
- old ne'
+let load_class (_, cl) =
+ classes := Gmap.add cl.cl_impl cl !classes
-let add_instance_hint_ref = ref (fun id pri -> assert false)
-let register_add_instance_hint =
- (:=) add_instance_hint_ref
-let add_instance_hint id = !add_instance_hint_ref id
+let cache_class = load_class
-let cache (_, (cl, m, inst)) =
- classes := cl;
- methods := m;
- instances := inst
-
-let load (_, (cl, m, inst)) =
- classes := gmap_merge !classes cl;
- methods := gmap_merge !methods m;
- instances := gmap_cmap_merge !instances inst
-
-open Libobject
-
-let subst (_,subst,(cl,m,inst)) =
+let subst_class (subst,cl) =
let do_subst_con c = fst (Mod_subst.subst_con subst c)
and do_subst c = Mod_subst.subst_mps subst c
- and do_subst_gr gr = fst (subst_global subst gr)
- in
- let do_subst_ctx ctx =
- list_smartmap (fun (na, b, t) ->
- (na, Option.smartmap do_subst b, do_subst t))
- ctx
- in
- let do_subst_context (grs,ctx) =
+ and do_subst_gr gr = fst (subst_global subst gr) in
+ let do_subst_ctx ctx = list_smartmap
+ (fun (na, b, t) -> (na, Option.smartmap do_subst b, do_subst t))
+ ctx in
+ let do_subst_context (grs,ctx) =
list_smartmap (Option.smartmap (fun (gr,b) -> do_subst_gr gr, b)) grs,
- do_subst_ctx ctx
- in
+ 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 subst_class k cl classes =
- let k = do_subst_gr k in
- let cl' = { cl_impl = k;
- cl_context = do_subst_context cl.cl_context;
- cl_props = do_subst_ctx cl.cl_props;
- cl_projs = do_subst_projs cl.cl_projs; }
- in
- let cl' = if cl' = cl then cl else cl' in
- Gmap.add k cl' classes
- in
- let classes = Gmap.fold subst_class cl Gmap.empty in
- let subst_inst k insts instances =
- let k = do_subst_gr k in
- let insts' =
- Cmap.fold (fun cst is acc ->
- let cst = do_subst_con cst in
- let is' = { is with is_class = k; is_impl = cst } in
- Cmap.add cst (if is' = is then is else is') acc) insts Cmap.empty
- in Gmap.add k insts' instances
- in
- let instances = Gmap.fold subst_inst inst Gmap.empty in
- (classes, m, instances)
-
-let rel_of_variable_context ctx =
- List.fold_right (fun (n,_,b,t) (ctx', subst)->
- let decl = (Name n, Option.map (substn_vars 1 subst) b, substn_vars 1 subst t) in
- (decl :: ctx', n :: subst)) ctx ([], [])
-
-let discharge (_,(cl,m,inst)) =
+ { cl_impl = do_subst_gr cl.cl_impl;
+ cl_context = do_subst_context cl.cl_context;
+ cl_props = do_subst_ctx cl.cl_props;
+ cl_projs = do_subst_projs cl.cl_projs; }
+
+let discharge_class (_,cl) =
+ let repl = Lib.replacement_context () in
+ let rel_of_variable_context ctx = List.fold_right
+ ( fun (n,_,b,t) (ctx', subst) ->
+ let decl = (Name n, Option.map (substn_vars 1 subst) b, substn_vars 1 subst t) in
+ (decl :: ctx', n :: subst)
+ ) ctx ([], []) in
let discharge_rel_context subst n rel =
+ let rel = map_rel_context (Cooking.expmod_constr repl) rel in
let ctx, _ =
List.fold_right
- (fun (id, b, t) (ctx, k) ->
- (id, Option.smartmap (substn_vars k subst) b, substn_vars k subst t) :: ctx, succ k)
+ (fun (id, b, t) (ctx, k) ->
+ (id, Option.smartmap (substn_vars k subst) b, substn_vars k subst t) :: ctx, succ k)
rel ([], n)
in ctx
in
let abs_context cl =
match cl.cl_impl with
- | VarRef _ | ConstructRef _ -> assert false
- | ConstRef cst -> Lib.section_segment_of_constant cst
- | IndRef (ind,_) -> Lib.section_segment_of_mutual_inductive ind
- in
+ | VarRef _ | ConstructRef _ -> assert false
+ | ConstRef cst -> Lib.section_segment_of_constant cst
+ | IndRef (ind,_) -> Lib.section_segment_of_mutual_inductive ind in
let discharge_context ctx' subst (grs, ctx) =
- let grs' = List.map (fun _ -> None) subst @
+ let grs' = List.map (fun _ -> None) subst @
list_smartmap (Option.smartmap (fun (gr, b) -> Lib.discharge_global gr, b)) grs
- in grs', discharge_rel_context subst 1 ctx @ ctx'
- in
- let subst_class k cl acc =
- let cl_impl' = Lib.discharge_global cl.cl_impl in
- let cl' =
- if cl_impl' == cl.cl_impl then cl
- else
- let ctx = abs_context cl in
- let ctx', subst = rel_of_variable_context ctx in
- { cl_impl = cl_impl';
- cl_context = discharge_context ctx' subst cl.cl_context;
- cl_props = discharge_rel_context subst (succ (List.length (fst cl.cl_context))) cl.cl_props;
- cl_projs = list_smartmap (fun (x, y) -> x, Option.smartmap Lib.discharge_con y) cl.cl_projs }
- in Gmap.add cl_impl' cl' acc
- in
- let classes = Gmap.fold subst_class cl Gmap.empty in
- let subst_inst k insts acc =
- let k' = Lib.discharge_global k in
- let insts' =
- Cmap.fold (fun k is acc ->
- let impl = Lib.discharge_con is.is_impl in
- let is = { is with is_class = k'; is_impl = impl } in
- Cmap.add impl is acc)
- insts Cmap.empty
- in Gmap.add k' insts' acc
- in
- let instances = Gmap.fold subst_inst inst Gmap.empty in
- Some (classes, m, instances)
-
-let rebuild (cl,m,inst) =
- let inst =
- Gmap.map (fun insts ->
- Cmap.fold (fun k is insts ->
- match !(is.is_global) with
- | -1 -> insts
- | 0 -> Cmap.add k is insts
- | n ->
- add_instance_hint is.is_impl is.is_pri;
- is.is_global := pred n;
- Cmap.add k is insts) insts Cmap.empty)
- inst
- in (cl, m, inst)
-
-let (input,output) =
+ in grs', discharge_rel_context subst 1 ctx @ ctx' in
+ let cl_impl' = Lib.discharge_global cl.cl_impl in
+ if cl_impl' == cl.cl_impl then cl else
+ let ctx = abs_context cl in
+ let ctx, subst = rel_of_variable_context ctx in
+ let context = discharge_context ctx subst cl.cl_context in
+ let props = discharge_rel_context subst (succ (List.length (fst cl.cl_context))) cl.cl_props in
+ { cl_impl = cl_impl';
+ cl_context = context;
+ cl_props = props;
+ cl_projs = list_smartmap (fun (x, y) -> x, Option.smartmap Lib.discharge_con y) cl.cl_projs }
+
+let rebuild_class cl = cl
+
+let (class_input,class_output) =
declare_object
{ (default_object "type classes state") with
- cache_function = cache;
- load_function = (fun _ -> load);
- open_function = (fun _ -> load);
- classify_function = (fun (_,x) -> Substitute x);
- discharge_function = discharge;
- rebuild_function = rebuild;
- subst_function = subst;
- export_function = (fun x -> Some x) }
-
-let update () =
- Lib.add_anonymous_leaf (input (!classes, !methods, !instances))
-
-let add_class c =
- classes := Gmap.add c.cl_impl c !classes;
- methods := List.fold_left (fun acc x ->
- match snd x with
- | Some m -> Gmap.add m c.cl_impl acc
- | None -> acc) !methods c.cl_projs;
- update ()
+ cache_function = cache_class;
+ load_function = (fun _ -> load_class);
+ open_function = (fun _ -> load_class);
+ classify_function = (fun x -> Substitute x);
+ discharge_function = (fun a -> Some (discharge_class a));
+ rebuild_function = rebuild_class;
+ subst_function = subst_class }
+
+let add_class cl =
+ Lib.add_anonymous_leaf (class_input cl)
+
+
+(*
+ * instances persistent object
+ *)
+
+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 subst_instance (subst,inst) =
+ { inst with
+ is_class = fst (subst_global subst inst.is_class);
+ is_impl = fst (subst_global subst inst.is_impl) }
+
+let discharge_instance (_,inst) =
+ if inst.is_global <= 0 then None
+ else Some
+ { inst with
+ is_global = pred inst.is_global;
+ is_class = Lib.discharge_global inst.is_class;
+ is_impl = Lib.discharge_global inst.is_impl }
-let class_info c =
+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 (instance_input,instance_output) =
+ declare_object
+ { (default_object "type classes instances state") with
+ cache_function = cache_instance;
+ load_function = (fun _ -> load_instance);
+ open_function = (fun _ -> load_instance);
+ 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
+
+open Declarations
+
+let add_constant_class cst =
+ let ty = Typeops.type_of_constant (Global.env ()) cst in
+ let ctx, arity = decompose_prod_assum ty in
+ let tc =
+ { cl_impl = ConstRef cst;
+ cl_context = (List.map (const None) ctx, ctx);
+ cl_props = [(Anonymous, None, arity)];
+ cl_projs = []
+ }
+ in add_class tc;
+ set_typeclass_transparency (EvalConstRef cst) false
+
+let add_inductive_class ind =
+ let mind, oneind = Global.lookup_inductive ind in
+ let k =
+ let ctx = oneind.mind_arity_ctxt in
+ let ty = Inductive.type_of_inductive_knowing_parameters
+ (push_rel_context ctx (Global.env ()))
+ oneind (Termops.extended_rel_vect 0 ctx)
+ in
+ { cl_impl = IndRef ind;
+ cl_context = List.map (const None) ctx, ctx;
+ cl_props = [Anonymous, None, ty];
+ cl_projs = [] }
+ in add_class k
+
+(*
+ * interface functions
+ *)
+
+let class_info c =
try Gmap.find c !classes
with _ -> not_a_class (Global.env()) (constr_of_global c)
-let instance_constructor cl args =
- let pars = fst (list_chop (List.length (fst cl.cl_context)) args) in
+let instance_constructor cl args =
+ let lenpars = List.length (List.filter (fun (na, b, t) -> b = None) (snd cl.cl_context)) in
+ let pars = fst (list_chop lenpars args) in
match cl.cl_impl with
| IndRef ind -> applistc (mkConstruct (ind, 1)) args,
applistc (mkInd ind) pars
| ConstRef cst -> list_last args, applistc (mkConst cst) pars
| _ -> assert false
-
+
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 (Cmap.add y.is_impl y l) m
+ Gmap.add x (Gmap.add y.is_impl y l) m
with Not_found ->
- Gmap.add x (Cmap.add y.is_impl y Cmap.empty) m
+ Gmap.add x (Gmap.add y.is_impl y Gmap.empty) m
-let cmap_elements c = Cmap.fold (fun k v acc -> v :: acc) c []
+let cmap_elements c = Gmap.fold (fun k v acc -> v :: acc) c []
-let instances_of c =
+let instances_of c =
try cmap_elements (Gmap.find c.cl_impl !instances) with Not_found -> []
-let add_instance i =
- let cl = class_info i.is_class in
- instances := cmapl_add cl.cl_impl i !instances;
- add_instance_hint i.is_impl i.is_pri;
- update ()
-
let all_instances () =
Gmap.fold (fun k v acc ->
- Cmap.fold (fun k v acc -> v :: acc) v acc)
+ Gmap.fold (fun k v acc -> v :: acc) v acc)
!instances []
let instances r =
- let cl = class_info r in instances_of cl
-
-let method_typeclass ref =
- match ref with
- | ConstRef c ->
- class_info (Gmap.find c !methods)
- | _ -> raise Not_found
+ let cl = class_info r in instances_of cl
let is_class gr =
Gmap.fold (fun k v acc -> acc || v.cl_impl = gr) !classes false
-
-let is_method c =
- Gmap.mem c !methods
let is_instance = function
| ConstRef c ->
@@ -320,18 +313,20 @@ let is_instance = function
(match Decls.variable_kind v with
| IsDefinition Instance -> true
| _ -> false)
+ | ConstructRef (ind,_) ->
+ is_class (IndRef ind)
| _ -> false
-let is_implicit_arg k =
+let is_implicit_arg k =
match k with
- ImplicitArg (ref, (n, id)) -> true
+ ImplicitArg (ref, (n, id), b) -> true
| InternalHole -> true
| _ -> false
-let global_class_of_constr env c =
+let global_class_of_constr env c =
try class_info (global_of_constr c)
with Not_found -> not_a_class env c
-
+
let dest_class_app env c =
let cl, args = decompose_app c in
global_class_of_constr env cl, args
@@ -339,49 +334,49 @@ let dest_class_app env c =
let class_of_constr c = try Some (fst (dest_class_app (Global.env ()) c)) with _ -> None
(* 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]
+ 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. *)
let ((bool_in : bool -> Dyn.t),
(bool_out : Dyn.t -> bool)) = Dyn.create "bool"
-
+
let bool_false = bool_in false
let is_resolvable evi =
match evi.evar_extra with
Some t -> if Dyn.tag t = "bool" then bool_out t else true
| None -> true
-
-let mark_unresolvable evi =
+
+let mark_unresolvable evi =
{ evi with evar_extra = Some bool_false }
-
+
let mark_unresolvables sigma =
Evd.fold (fun ev evi evs ->
Evd.add evs ev (mark_unresolvable evi))
sigma Evd.empty
-
-let rec is_class_type c =
+
+let rec is_class_type evd c =
match kind_of_term c with
- | Prod (_, _, t) -> is_class_type t
+ | Prod (_, _, t) -> is_class_type evd t
+ | Evar (e, _) when is_defined evd e -> is_class_type evd (Evarutil.nf_evar evd c)
| _ -> class_of_constr c <> None
-let is_class_evar evi =
- is_class_type evi.Evd.evar_concl
-
+let is_class_evar evd evi =
+ is_class_type evd evi.Evd.evar_concl
+
let has_typeclasses evd =
- Evd.fold (fun ev evi has -> has ||
- (evi.evar_body = Evar_empty && is_class_evar evi && is_resolvable evi))
+ Evd.fold (fun ev evi has -> has ||
+ (evi.evar_body = Evar_empty && 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.evars_of evd)) then evd
- else
- !solve_instanciations_problem env evd onlyargs split fail
+ 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 d8e15895..997b28c1 100644
--- a/pretyping/typeclasses.mli
+++ b/pretyping/typeclasses.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: typeclasses.mli 11576 2008-11-10 19:13:15Z msozeau $ i*)
+(*i $Id$ i*)
(*i*)
open Names
@@ -24,19 +24,19 @@ open Util
(* 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;
+ 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;
+ cl_context : (global_reference * bool) option list * rel_context;
(* 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
+ (* The methods implementations of the typeclass as projections. Some may be undefinable due to
sorting restrictions. *)
cl_projs : (identifier * constant option) list;
}
@@ -49,7 +49,11 @@ val all_instances : unit -> instance list
val add_class : typeclass -> unit
-val new_instance : typeclass -> int option -> bool -> constant -> instance
+val add_constant_class : constant -> unit
+
+val add_inductive_class : inductive -> unit
+
+val new_instance : typeclass -> int option -> bool -> global_reference -> instance
val add_instance : instance -> unit
val class_info : global_reference -> typeclass (* raises a UserError if not a class *)
@@ -61,11 +65,10 @@ val dest_class_app : env -> constr -> typeclass * constr list
(* Just return None if not a class *)
val class_of_constr : constr -> typeclass option
-val instance_impl : instance -> constant
+val instance_impl : instance -> global_reference
val is_class : global_reference -> bool
val is_instance : global_reference -> bool
-val is_method : constant -> bool
val is_implicit_arg : hole_kind -> bool
@@ -81,15 +84,17 @@ val bool_out : Dyn.t -> bool
val is_resolvable : evar_info -> bool
val mark_unresolvable : evar_info -> evar_info
val mark_unresolvables : evar_map -> evar_map
-val is_class_evar : evar_info -> bool
+val is_class_evar : evar_map -> evar_info -> bool
-val resolve_typeclasses : ?onlyargs:bool -> ?split:bool -> ?fail:bool ->
- env -> evar_defs -> evar_defs
-val resolve_one_typeclass : env -> evar_map -> types -> constr
+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_add_instance_hint : (constant -> int option -> unit) -> unit
-val add_instance_hint : constant -> int option -> unit
+val register_set_typeclass_transparency : (evaluable_global_reference -> bool -> unit) -> unit
+val set_typeclass_transparency : evaluable_global_reference -> bool -> unit
-val solve_instanciations_problem : (env -> evar_defs -> bool -> bool -> bool -> evar_defs) ref
-val solve_instanciation_problem : (env -> evar_map -> types -> constr) ref
+val register_add_instance_hint : (global_reference -> int option -> unit) -> unit
+val add_instance_hint : global_reference -> int option -> 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
diff --git a/pretyping/typeclasses_errors.ml b/pretyping/typeclasses_errors.ml
index 8844baab..1de8b7a5 100644
--- a/pretyping/typeclasses_errors.ml
+++ b/pretyping/typeclasses_errors.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: typeclasses_errors.ml 11282 2008-07-28 11:51:53Z msozeau $ i*)
+(*i $Id$ i*)
(*i*)
open Names
@@ -24,11 +24,11 @@ open Libnames
type contexts = Parameters | Properties
-type typeclass_error =
+type typeclass_error =
| NotAClass of constr
| UnboundMethod of global_reference * identifier located (* Class name, method *)
| NoInstance of identifier located * constr list
- | UnsatisfiableConstraints of evar_defs * (evar_info * hole_kind) option
+ | UnsatisfiableConstraints of evar_map * (existential_key * hole_kind) option
| MismatchedContextInstance of contexts * constr_expr list * rel_context (* found, expected *)
exception TypeClassError of env * typeclass_error
@@ -41,15 +41,19 @@ let unbound_method env cid id = typeclass_error env (UnboundMethod (cid, id))
let no_instance env id args = typeclass_error env (NoInstance (id, args))
-let unsatisfiable_constraints env evd ev =
- let evd = Evd.undefined_evars evd in
- match ev with
- | None ->
- raise (TypeClassError (env, UnsatisfiableConstraints (evd, None)))
- | Some ev ->
- let evi = Evd.find (Evd.evars_of evd) ev in
- let loc, kind = Evd.evar_source ev evd in
- raise (Stdpp.Exc_located (loc, TypeClassError
- (env, UnsatisfiableConstraints (evd, Some (evi, kind)))))
-
+let unsatisfiable_constraints env evd ev =
+ match ev with
+ | None ->
+ raise (TypeClassError (env, UnsatisfiableConstraints (evd, None)))
+ | Some ev ->
+ let loc, kind = Evd.evar_source ev evd in
+ raise (Stdpp.Exc_located (loc, TypeClassError
+ (env, UnsatisfiableConstraints (evd, Some (ev, kind)))))
+
let mismatched_ctx_inst env c n m = typeclass_error env (MismatchedContextInstance (c, n, m))
+
+let rec unsatisfiable_exception exn =
+ match exn with
+ | TypeClassError (_, UnsatisfiableConstraints _) -> true
+ | Stdpp.Exc_located(_, e) -> unsatisfiable_exception e
+ | _ -> false
diff --git a/pretyping/typeclasses_errors.mli b/pretyping/typeclasses_errors.mli
index a79307d0..7fd04e22 100644
--- a/pretyping/typeclasses_errors.mli
+++ b/pretyping/typeclasses_errors.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: typeclasses_errors.mli 11282 2008-07-28 11:51:53Z msozeau $ i*)
+(*i $Id$ i*)
(*i*)
open Names
@@ -24,12 +24,12 @@ open Libnames
type contexts = Parameters | Properties
-type typeclass_error =
- | NotAClass of constr
- | UnboundMethod of global_reference * identifier located (* Class name, method *)
- | NoInstance of identifier located * constr list
- | UnsatisfiableConstraints of evar_defs * (evar_info * hole_kind) option
- | MismatchedContextInstance of contexts * constr_expr list * rel_context (* found, expected *)
+type typeclass_error =
+ | NotAClass of constr
+ | UnboundMethod of global_reference * identifier located (* Class name, method *)
+ | NoInstance of identifier located * constr list
+ | UnsatisfiableConstraints of evar_map * (existential_key * hole_kind) option
+ | MismatchedContextInstance of contexts * constr_expr list * rel_context (* found, expected *)
exception TypeClassError of env * typeclass_error
@@ -39,6 +39,8 @@ val unbound_method : env -> global_reference -> identifier located -> 'a
val no_instance : env -> identifier located -> constr list -> 'a
-val unsatisfiable_constraints : env -> evar_defs -> evar option -> 'a
+val unsatisfiable_constraints : env -> evar_map -> evar option -> 'a
val mismatched_ctx_inst : env -> contexts -> constr_expr list -> rel_context -> 'a
+
+val unsatisfiable_exception : exn -> bool
diff --git a/pretyping/typing.ml b/pretyping/typing.ml
index 43e19ca7..831787a0 100644
--- a/pretyping/typing.ml
+++ b/pretyping/typing.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: typing.ml 10785 2008-04-13 21:41:54Z herbelin $ *)
+(* $Id$ *)
open Util
open Names
@@ -35,155 +35,191 @@ 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_judge_of_apply env evdref funj argjv =
+ let rec apply_rec n typ = function
+ | [] ->
+ { uj_val = mkApp (j_val funj, Array.map j_val argjv);
+ uj_type = typ }
+ | hj::restjl ->
+ match kind_of_term (whd_betadeltaiota env !evdref typ) with
+ | Prod (_,c1,c2) ->
+ if Evarconv.e_cumul env evdref hj.uj_type c1 then
+ apply_rec (n+1) (subst1 hj.uj_val c2) restjl
+ else
+ error_cant_apply_bad_type env (n,c1, hj.uj_type) funj argjv
+ | _ ->
+ 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) =
+ 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)
+ done
+
+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);
+ { uj_val = mkCase (ci, pj.uj_val, cj.uj_val, Array.map j_val lfj);
+ uj_type = rslty }
+
+let e_judge_of_cast env evdref cj k tj =
+ let expected_type = tj.utj_val in
+ if not (Evarconv.e_cumul env evdref cj.uj_type expected_type) then
+ error_actual_type env cj expected_type;
+ { uj_val = mkCast (cj.uj_val, k, expected_type);
+ uj_type = expected_type }
+
(* The typing machine without information, without universes but with
existential variables. *)
(* cstr must be in n.f. w.r.t. evars and execute returns a judgement
where both the term and type are in n.f. *)
-let rec execute env evd cstr =
+let rec execute env evdref cstr =
match kind_of_term cstr with
| Meta n ->
- { uj_val = cstr; uj_type = nf_evar (evars_of evd) (meta_type evd n) }
+ { uj_val = cstr; uj_type = meta_type !evdref n }
| Evar ev ->
- let sigma = Evd.evars_of evd in
- let ty = Evd.existential_type sigma ev in
- let jty = execute env evd (nf_evar (evars_of evd) ty) in
+ let ty = Evd.existential_type !evdref ev in
+ let jty = execute env evdref (whd_evar !evdref ty) in
let jty = assumption_of_judgment env jty in
{ uj_val = cstr; uj_type = jty }
-
- | Rel n ->
- j_nf_evar (evars_of evd) (judge_of_relative env n)
- | Var id ->
- j_nf_evar (evars_of evd) (judge_of_variable env id)
-
+ | Rel n ->
+ judge_of_relative env n
+
+ | Var id ->
+ judge_of_variable env id
+
| Const c ->
- make_judge cstr (nf_evar (evars_of evd) (type_of_constant env c))
-
+ make_judge cstr (type_of_constant env c)
+
| Ind ind ->
- make_judge cstr (nf_evar (evars_of evd) (type_of_inductive env ind))
-
- | Construct cstruct ->
- make_judge cstr
- (nf_evar (evars_of evd) (type_of_constructor env cstruct))
+ make_judge cstr (type_of_inductive env ind)
+
+ | Construct cstruct ->
+ make_judge cstr (type_of_constructor env cstruct)
| Case (ci,p,c,lf) ->
- let cj = execute env evd c in
- let pj = execute env evd p in
- let lfj = execute_array env evd lf in
- let (j,_) = judge_of_case env ci pj cj lfj in
- j
-
+ let cj = execute env evdref c in
+ let pj = execute env evdref p in
+ let lfj = execute_array env evdref lf in
+ e_judge_of_case env evdref ci pj cj lfj
+
| Fix ((vn,i as vni),recdef) ->
- let (_,tys,_ as recdef') = execute_recdef env evd recdef in
+ let (_,tys,_ as recdef') = execute_recdef env evdref recdef in
let fix = (vni,recdef') in
check_fix env fix;
make_judge (mkFix fix) tys.(i)
-
+
| CoFix (i,recdef) ->
- let (_,tys,_ as recdef') = execute_recdef env evd recdef in
+ let (_,tys,_ as recdef') = execute_recdef env evdref recdef in
let cofix = (i,recdef') in
check_cofix env cofix;
make_judge (mkCoFix cofix) tys.(i)
-
- | Sort (Prop c) ->
+
+ | Sort (Prop c) ->
judge_of_prop_contents c
| Sort (Type u) ->
judge_of_type u
-
+
| App (f,args) ->
- let jl = execute_array env evd args in
+ let jl = execute_array env evdref args in
let j =
match kind_of_term f with
| Ind ind ->
(* Sort-polymorphism of inductive types *)
make_judge f
(inductive_type_knowing_parameters env ind
- (jv_nf_evar (evars_of evd) jl))
- | Const cst ->
+ (jv_nf_evar !evdref jl))
+ | Const cst ->
(* Sort-polymorphism of inductive types *)
make_judge f
(constant_type_knowing_parameters env cst
- (jv_nf_evar (evars_of evd) jl))
- | _ ->
- execute env evd f
+ (jv_nf_evar !evdref jl))
+ | _ ->
+ execute env evdref f
in
- fst (judge_of_apply env j jl)
-
- | Lambda (name,c1,c2) ->
- let j = execute env evd c1 in
+ e_judge_of_apply env evdref j jl
+
+ | Lambda (name,c1,c2) ->
+ let j = execute env evdref c1 in
let var = type_judgment env j in
let env1 = push_rel (name,None,var.utj_val) env in
- let j' = execute env1 evd c2 in
+ let j' = execute env1 evdref c2 in
judge_of_abstraction env1 name var j'
-
+
| Prod (name,c1,c2) ->
- let j = execute env evd c1 in
+ let j = execute env evdref c1 in
let varj = type_judgment env j in
let env1 = push_rel (name,None,varj.utj_val) env in
- let j' = execute env1 evd c2 in
+ let j' = execute env1 evdref c2 in
let varj' = type_judgment env1 j' in
judge_of_product env name varj varj'
| LetIn (name,c1,c2,c3) ->
- let j1 = execute env evd c1 in
- let j2 = execute env evd c2 in
+ let j1 = execute env evdref c1 in
+ let j2 = execute env evdref c2 in
let j2 = type_judgment env 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 evd c3 in
+ let j3 = execute env1 evdref c3 in
judge_of_letin env name j1 j2 j3
-
+
| Cast (c,k,t) ->
- let cj = execute env evd c in
- let tj = execute env evd t in
+ let cj = execute env evdref c in
+ let tj = execute env evdref t in
let tj = type_judgment env tj in
- let j, _ = judge_of_cast env cj k tj in
- j
+ e_judge_of_cast env evdref cj k tj
-and execute_recdef env evd (names,lar,vdef) =
- let larj = execute_array env evd lar in
+and execute_recdef env evdref (names,lar,vdef) =
+ let larj = execute_array env evdref lar in
let lara = Array.map (assumption_of_judgment env) larj in
let env1 = push_rec_types (names,lara,vdef) env in
- let vdefj = execute_array env1 evd vdef in
+ let vdefj = execute_array env1 evdref vdef in
let vdefv = Array.map j_val vdefj in
let _ = type_fixpoint env1 names lara vdefj in
(names,lara,vdefv)
-and execute_array env evd = Array.map (execute env evd)
+and execute_array env evdref = Array.map (execute env evdref)
-and execute_list env evd = List.map (execute env evd)
+and execute_list env evdref = List.map (execute env evdref)
-let mcheck env evd c t =
- let sigma = Evd.evars_of evd in
- let j = execute env evd (nf_evar sigma c) in
- if not (is_conv_leq env sigma j.uj_type t) then
- error_actual_type env j (nf_evar sigma t)
+let check env evd c t =
+ let evdref = ref evd in
+ let j = execute env evdref c in
+ if not (Evarconv.e_cumul env evdref j.uj_type t) then
+ error_actual_type env j (nf_evar evd t)
(* Type of a constr *)
-
-let mtype_of env evd c =
- let j = execute env evd (nf_evar (evars_of evd) c) in
+
+let type_of env evd c =
+ let j = execute env (ref evd) c in
(* We are outside the kernel: we take fresh universes *)
(* to avoid tactics and co to refresh universes themselves *)
Termops.refresh_universes j.uj_type
-let msort_of env evd c =
- let j = execute env evd (nf_evar (evars_of evd) c) in
+(* 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
a.utj_type
-let type_of env sigma c =
- mtype_of env (Evd.create_evar_defs sigma) c
-let sort_of env sigma c =
- msort_of env (Evd.create_evar_defs sigma) c
-let check env sigma c =
- mcheck env (Evd.create_evar_defs sigma) c
+(* Try to solve the existential variables by typing *)
-(* The typed type of a judgment. *)
-
-let mtype_of_type env evd constr =
- let j = execute env evd (nf_evar (evars_of evd) constr) in
- assumption_of_judgment env j
+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
+
diff --git a/pretyping/typing.mli b/pretyping/typing.mli
index c9d7d572..e3b3e948 100644
--- a/pretyping/typing.mli
+++ b/pretyping/typing.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: typing.mli 6113 2004-09-17 20:28:19Z barras $ i*)
+(*i $Id$ i*)
(*i*)
open Term
@@ -21,14 +21,11 @@ open Evd
val type_of : env -> evar_map -> constr -> 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
-
-(* The same but with metas... *)
-val mtype_of : env -> evar_defs -> constr -> types
-val msort_of : env -> evar_defs -> types -> sorts
-val mcheck : env -> evar_defs -> constr -> types -> unit
-val meta_type : evar_defs -> metavariable -> types
-
-(* unused typing function... *)
-val mtype_of_type : env -> evar_defs -> types -> types
+
+(* 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
diff --git a/pretyping/unification.ml b/pretyping/unification.ml
index c92f1fc6..a2841ec6 100644
--- a/pretyping/unification.ml
+++ b/pretyping/unification.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: unification.ml 12268 2009-08-11 09:02:16Z herbelin $ *)
+(* $Id$ *)
open Pp
open Util
@@ -14,6 +14,7 @@ open Names
open Nameops
open Term
open Termops
+open Namegen
open Sign
open Environ
open Evd
@@ -25,32 +26,45 @@ open Evarutil
open Pretype_errors
open Retyping
open Coercion.Default
+open Recordops
+
+let occur_meta_or_undefined_evar evd c =
+ let rec occrec c = match kind_of_term c with
+ | Meta _ -> raise Occur
+ | Evar (ev,args) ->
+ (match evar_body (Evd.find evd ev) with
+ | Evar_defined c ->
+ occrec c; Array.iter occrec args
+ | Evar_empty -> raise Occur)
+ | Sort s when is_sort_variable evd s -> raise Occur
+ | _ -> iter_constr occrec c
+ in try occrec c; false with Occur | Not_found -> true
(* if lname_typ is [xn,An;..;x1,A1] and l is a list of terms,
gives [x1:A1]..[xn:An]c' such that c converts to ([x1:A1]..[xn:An]c' l) *)
let abstract_scheme env c l lname_typ =
- List.fold_left2
+ List.fold_left2
(fun t (locc,a) (na,_,ta) ->
let na = match kind_of_term a with Var id -> Name id | _ -> na in
(* [occur_meta ta] test removed for support of eelim/ecase but consequences
are unclear...
if occur_meta ta then error "cannot find a type for the generalisation"
- else *) if occur_meta a then lambda_name env (na,ta,t)
- else lambda_name env (na,ta,subst_term_occ locc a t))
+ else *) if occur_meta a then mkLambda_name env (na,ta,t)
+ else mkLambda_name env (na,ta,subst_term_occ locc a t))
c
(List.rev l)
lname_typ
let abstract_list_all env evd typ c l =
- let ctxt,_ = decomp_n_prod env (evars_of evd) (List.length l) typ in
+ let ctxt,_ = splay_prod_n env evd (List.length l) typ in
let l_with_all_occs = List.map (function a -> (all_occurrences,a)) l in
- let p = abstract_scheme env c l_with_all_occs ctxt in
- try
- if is_conv_leq env (evars_of evd) (Typing.mtype_of env evd p) typ then p
+ let p = abstract_scheme env c l_with_all_occs ctxt in
+ try
+ if is_conv_leq env evd (Typing.type_of env evd p) typ then p
else error "abstract_list_all"
with UserError _ | Type_errors.TypeError _ ->
- error_cannot_find_well_typed_abstraction env (evars_of evd) p l
+ error_cannot_find_well_typed_abstraction env evd p l
(**)
@@ -86,32 +100,50 @@ let rec subst_meta_instances bl c =
| Meta i -> (try assoc_pair i bl with Not_found -> c)
| _ -> map_constr (subst_meta_instances bl) c
-let solve_pattern_eqn_array (env,nb) sigma f l c (metasubst,evarsubst) =
+let pose_all_metas_as_evars env evd t =
+ let evdref = ref evd in
+ let rec aux t = match kind_of_term t with
+ | Meta mv ->
+ (match Evd.meta_opt_fvalue !evdref mv with
+ | Some ({rebus=c},_) -> c
+ | None ->
+ let {rebus=ty;freemetas=mvs} = Evd.meta_ftype evd mv in
+ let ty = if mvs = Evd.Metaset.empty then ty else aux ty in
+ let ev = Evarutil.e_new_evar evdref env ~src:(dummy_loc,GoalEvar) ty in
+ evdref := meta_assign mv (ev,(ConvUpToEta 0,TypeNotProcessed)) !evdref;
+ ev)
+ | _ ->
+ map_constr aux t in
+ let c = aux t in
+ (* side-effect *)
+ (!evdref, c)
+
+let solve_pattern_eqn_array (env,nb) f l c (sigma,metasubst,evarsubst) =
match kind_of_term f with
- | Meta k ->
+ | 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
if noccur_between 1 nb c then
- (k,lift (-nb) c,pb)::metasubst,evarsubst
+ sigma,(k,lift (-nb) c,pb)::metasubst,evarsubst
else error_cannot_unify_local env sigma (mkApp (f, l),c,c)
| Evar ev ->
- (* Currently unused: incompatible with eauto/eassumption backtracking *)
- metasubst,(ev,solve_pattern_eqn env (Array.to_list l) c)::evarsubst
+ 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
| _ -> assert false
let push d (env,n) = (push_rel_assum d env,n+1)
(*******************************)
-(* Unification à l'ordre 0 de m et n: [unify_0 env sigma cv_pb m n]
+(* Unification à l'ordre 0 de m et n: [unify_0 env sigma cv_pb m n]
renvoie deux listes:
- metasubst:(int*constr)list récolte les instances des (Meta k)
- evarsubst:(constr*constr)list récolte les instances des (Const "?k")
+ metasubst:(int*constr)list récolte les instances des (Meta k)
+ evarsubst:(constr*constr)list récolte les instances des (Const "?k")
- Attention : pas d'unification entre les différences instances d'une
- même meta ou evar, il peut rester des doublons *)
+ Attention : pas d'unification entre les différences instances d'une
+ même meta ou evar, il peut rester des doublons *)
(* Unification order: *)
(* Left to right: unifies first argument and then the other arguments *)
@@ -122,93 +154,104 @@ let unify_r2l x = x
let sort_eqns = unify_r2l
*)
-type unify_flags = {
+(* Option introduced and activated in Coq 8.3 *)
+let global_evars_pattern_unification_flag = ref true
+
+open Goptions
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optname = "pattern-unification for existential variables in tactics";
+ optkey = ["Tactic";"Evars";"Pattern";"Unification"];
+ optread = (fun () -> !global_evars_pattern_unification_flag);
+ optwrite = (:=) global_evars_pattern_unification_flag }
+
+type unify_flags = {
modulo_conv_on_closed_terms : Names.transparent_state option;
use_metas_eagerly : bool;
modulo_delta : Names.transparent_state;
+ resolve_evars : bool;
+ use_evars_pattern_unification : bool
}
let default_unify_flags = {
modulo_conv_on_closed_terms = Some full_transparent_state;
use_metas_eagerly = true;
modulo_delta = full_transparent_state;
+ resolve_evars = false;
+ use_evars_pattern_unification = true;
}
let default_no_delta_unify_flags = {
modulo_conv_on_closed_terms = Some full_transparent_state;
use_metas_eagerly = true;
modulo_delta = empty_transparent_state;
+ resolve_evars = false;
+ use_evars_pattern_unification = false;
}
+let use_evars_pattern_unification flags =
+ !global_evars_pattern_unification_flag && flags.use_evars_pattern_unification
+ && Flags.version_strictly_greater Flags.V8_2
+
let expand_key env = function
| Some (ConstKey cst) -> constant_opt_value env cst
- | Some (VarKey id) -> named_body id env
+ | Some (VarKey id) -> (try named_body id env with Not_found -> None)
| Some (RelKey _) -> None
| None -> None
-
+
let key_of flags f =
match kind_of_term f with
| Const cst when is_transparent (ConstKey cst) &&
Cpred.mem cst (snd flags.modulo_delta) ->
- Some (ConstKey cst)
+ Some (ConstKey cst)
| Var id when is_transparent (VarKey id) &&
Idpred.mem id (fst flags.modulo_delta) ->
Some (VarKey id)
| _ -> None
-
+
let oracle_order env cf1 cf2 =
match cf1 with
| None ->
- (match cf2 with
+ (match cf2 with
| None -> None
| Some k2 -> Some false)
- | Some k1 ->
+ | Some k1 ->
match cf2 with
| None -> Some true
| Some k2 -> Some (Conv_oracle.oracle_order k1 k2)
-
-let unify_0_with_initial_metas subst conv_at_top env sigma cv_pb flags m n =
- let trivial_unify curenv pb (metasubst,_) m n =
- let subst = if flags.use_metas_eagerly then metasubst else fst subst in
- match subst_defined_metas subst m with
- | Some m1 ->
- if (match flags.modulo_conv_on_closed_terms with
- Some flags ->
- is_trans_fconv (conv_pb_of pb) flags env sigma m1 n
- | None -> false) then true else
- if (not (is_ground_term (create_evar_defs sigma) m1))
- || occur_meta_or_existential n then false else
- error_cannot_unify curenv sigma (m, n)
- | _ -> false in
- let rec unirec_rec (curenv,nb as curenvnb) pb b ((metasubst,evarsubst) as substn) curm curn =
- let cM = Evarutil.whd_castappevar sigma curm
- and cN = Evarutil.whd_castappevar sigma curn in
+
+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 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 ->
let stM,stN = extract_instance_status pb in
- if k1 < k2
- then (k1,cN,stN)::metasubst,evarsubst
+ if k2 < k1
+ then sigma,(k1,cN,stN)::metasubst,evarsubst
else if k1 = k2 then substn
- else (k2,cM,stM)::metasubst,evarsubst
- | Meta k, _ when not (dependent cM cN) ->
+ else sigma,(k2,cM,stM)::metasubst,evarsubst
+ | Meta k, _ when not (dependent cM cN) ->
(* Here we check that [cN] does not contain any local variables *)
if nb = 0 then
- (k,cN,snd (extract_instance_status pb))::metasubst,evarsubst
+ sigma,(k,cN,snd (extract_instance_status pb))::metasubst,evarsubst
else if noccur_between 1 nb cN then
- (k,lift (-nb) cN,snd (extract_instance_status pb))::metasubst,
- evarsubst
+ (sigma,
+ (k,lift (-nb) cN,snd (extract_instance_status pb))::metasubst,
+ evarsubst)
else error_cannot_unify_local curenv sigma (m,n,cN)
- | _, Meta k when not (dependent cN cM) ->
+ | _, Meta k when not (dependent cN cM) ->
(* Here we check that [cM] does not contain any local variables *)
if nb = 0 then
- (k,cM,snd (extract_instance_status pb))::metasubst,evarsubst
+ (sigma,(k,cM,fst (extract_instance_status pb))::metasubst,evarsubst)
else if noccur_between 1 nb cM
then
- (k,lift (-nb) cM,fst (extract_instance_status pb))::metasubst,
- evarsubst
+ (sigma,(k,lift (-nb) cM,fst (extract_instance_status pb))::metasubst,
+ evarsubst)
else error_cannot_unify_local curenv sigma (m,n,cM)
- | Evar ev, _ -> metasubst,((ev,cN)::evarsubst)
- | _, Evar ev -> metasubst,((ev,cM)::evarsubst)
+ | Evar ev, _ -> sigma,metasubst,((ev,cN)::evarsubst)
+ | _, Evar ev -> sigma,metasubst,((ev,cM)::evarsubst)
| 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
@@ -217,55 +260,93 @@ let unify_0_with_initial_metas subst conv_at_top env sigma cv_pb flags m n =
(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)
-
+
| 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 & is_unification_pattern curenvnb f1 l1 cN &
+ (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 sigma f1 l1 cN substn
+ solve_pattern_eqn_array curenvnb f1 l1 cN substn
| _, App (f2,l2) when
- isMeta f2 & is_unification_pattern curenvnb f2 l2 cM &
+ (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 sigma f2 l2 cM substn
+ solve_pattern_eqn_array curenvnb f2 l2 cM substn
| App (f1,l1), App (f2,l2) ->
- let len1 = Array.length l1
- and len2 = Array.length l2 in
+ 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
+ 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
+ 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 ->
- expand curenvnb pb b substn cM f1 l1 cN f2 l2)
-
+ 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)
+
| _ ->
- 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
-
- and expand (curenv,_ as curenvnb) pb b substn cM f1 l1 cN f2 l2 =
- if trivial_unify curenv pb substn cM cN then substn
- else if b then
+ 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
+
+ and expand (curenv,_ as curenvnb) pb b (sigma,metasubst,_ as substn) cM f1 l1 cN f2 l2 =
+
+ if
+ (* Try full conversion on meta-free terms. *)
+ (* Back to 1995 (later on called trivial_unify in 2002), the
+ heuristic was to apply conversion on meta-free (but not
+ evar-free!) terms in all cases (i.e. for apply but also for
+ auto and rewrite, even though auto and rewrite did not use
+ modulo conversion in the rest of the unification
+ algorithm). By compatibility we need to support this
+ separately from the main unification algorithm *)
+ (* The exploitation of known metas has been added in May 2007
+ (it is used by apply and rewrite); it might now be redundant
+ with the support for delta-expansion (which is used
+ essentially for apply)... *)
+ match flags.modulo_conv_on_closed_terms with
+ | None -> false
+ | Some convflags ->
+ let subst = if flags.use_metas_eagerly 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
+ then true else
+ if is_ground_term sigma m1 && is_ground_term sigma n1 then
+ error_cannot_unify curenv sigma (cM,cN)
+ else false
+ then
+ substn
+ else
+ 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
match oracle_order curenv cf1 cf2 with
| None -> error_cannot_unify curenv sigma (cM,cN)
- | Some true ->
+ | Some true ->
(match expand_key curenv cf1 with
| Some c ->
unirec_rec curenvnb pb b substn
@@ -292,26 +373,70 @@ let unify_0_with_initial_metas subst conv_at_top env sigma cv_pb flags m n =
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
+ 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
+ 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
+ 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)
+ else error_cannot_unify (fst curenvnb) sigma (cM,cN)
+
+ and solve_canonical_projection curenvnb pb b cM f1l1 cN f2l2 (sigma,ms,es) =
+ let (c,bs,(params,params1),(us,us2),(ts,ts1),c1,(n,t2)) =
+ try Evarconv.check_conv_record f1l1 f2l2
+ with Not_found -> error_cannot_unify (fst curenvnb) sigma (cM,cN)
+ in
+ let (evd,ks,_) =
+ List.fold_left
+ (fun (evd,ks,m) b ->
+ if m=n then (evd,t2::ks, m-1) else
+ let mv = new_meta () in
+ let evd' = meta_declare mv (substl ks b) evd in
+ (evd', mkMeta mv :: ks, m - 1))
+ (sigma,[],List.length bs - 1) bs
+ in
+ let unilist2 f substn l l' =
+ try List.fold_left2 f substn l l'
+ with Invalid_argument "List.fold_left2" -> error_cannot_unify (fst curenvnb) sigma (cM,cN)
+ in
+ let substn = unilist2 (fun s u1 u -> unirec_rec curenvnb pb b 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))
+ 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)))
+
in
- if (if occur_meta m 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 else
- if (not (is_ground_term (create_evar_defs sigma) m))
- || occur_meta_or_existential n then false else
- if (match flags.modulo_conv_on_closed_terms, flags.modulo_delta with
+ 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
+ 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
+ else if (match flags.modulo_conv_on_closed_terms, flags.modulo_delta with
| Some (cv_id, cv_k), (dl_id, dl_k) ->
Idpred.subset dl_id cv_id && Cpred.subset dl_k cv_k
| None,(dl_id, dl_k) ->
Idpred.is_empty dl_id && Cpred.is_empty dl_k)
- then error_cannot_unify env sigma (m, n) else false)
- then
- subst
- else
- unirec_rec (env,0) cv_pb conv_at_top subst m n
+ 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
-let unify_0 = unify_0_with_initial_metas ([],[]) true
+let unify_0 env sigma = unify_0_with_initial_metas (sigma,[],[]) true env
let left = true
let right = false
@@ -324,19 +449,19 @@ let rec unify_with_eta keptside flags env sigma k1 k2 c1 c2 =
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 metas,evars = unify_0 env sigma topconv flags t1 t2 in
- let side,status,(metas',evars') =
+ 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,(metas@metas',evars@evars'))
+ 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)
+ 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)
+ unify_with_eta side flags env' sigma (k1-1) (pop k2)
(mkApp (lift 1 c1,[|mkRel 1|])) c2'
| _ ->
(keptside,ConvUpToEta(min k1 k2),
@@ -426,49 +551,31 @@ let merge_instances env sigma flags st1 st2 c1 c2 =
* close it off. But this might not always work,
* since other metavars might also need to be resolved. *)
-let applyHead env evd n c =
+let applyHead env evd n c =
let rec apprec n c cty evd =
- if n = 0 then
+ if n = 0 then
(evd, c)
- else
- match kind_of_term (whd_betadeltaiota env (evars_of evd) cty) with
+ else
+ match kind_of_term (whd_betadeltaiota env evd cty) with
| Prod (_,c1,c2) ->
- let (evd',evar) =
+ 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 (evars_of evd) c) evd
+ in
+ apprec n c (Typing.type_of env evd c) evd
let is_mimick_head f =
match kind_of_term f with
(Const _|Var _|Rel _|Construct _|Ind _) -> true
| _ -> false
-let pose_all_metas_as_evars env evd t =
- let evdref = ref evd in
- let rec aux t = match kind_of_term t with
- | Meta mv ->
- (match Evd.meta_opt_fvalue !evdref mv with
- | Some ({rebus=c},_) -> c
- | None ->
- let {rebus=ty;freemetas=mvs} = Evd.meta_ftype evd mv in
- let ty = if mvs = Evd.Metaset.empty then ty else aux ty in
- let ev = Evarutil.e_new_evar evdref env ~src:(dummy_loc,GoalEvar) ty in
- evdref := meta_assign mv (ev,(ConvUpToEta 0,TypeNotProcessed)) !evdref;
- ev)
- | _ ->
- map_constr aux t in
- let c = aux t in
- (* side-effect *)
- (!evdref, c)
-
let try_to_coerce env evd c cty tycon =
let j = make_judge c cty in
let (evd',j') = inh_conv_coerce_rigid_to dummy_loc env evd j tycon in
let (evd',b) = Evarconv.consider_remaining_unif_problems env evd' in
if b then
- let evd' = Evd.map_metas_fvalue (nf_evar (evars_of evd')) evd' in
+ let evd' = Evd.map_metas_fvalue (nf_evar evd') evd' in
(evd',j'.uj_val)
else
error "Cannot solve unification constraints"
@@ -478,31 +585,39 @@ let w_coerce_to_type env evd c cty mvty =
let tycon = mk_tycon_type mvty in
try try_to_coerce env evd c cty tycon
with e when precatchable_exception e ->
- (* inh_conv_coerce_rigid_to should have reasoned modulo reduction
+ (* 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 (evars_of evd) cty in
+ let cty = Tacred.hnf_constr env evd cty in
try_to_coerce env evd c cty tycon
let w_coerce env evd mv c =
- let cty = get_type_of_with_meta env (evars_of evd) (metas_of evd) c in
+ let cty = get_type_of env evd c in
let mvty = Typing.meta_type evd mv in
w_coerce_to_type env evd c cty mvty
-let unify_to_type env evd flags c u =
- let sigma = evars_of evd in
+let unify_to_type env sigma flags c status u =
let c = refresh_universes c in
- let t = get_type_of_with_meta env sigma (List.map (on_snd (nf_meta evd)) (metas_of evd)) (nf_meta evd c) in
- let t = Tacred.hnf_constr env sigma (nf_betaiota sigma (nf_meta evd t)) 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 unify_0 env sigma Cumul flags t u
- with e when precatchable_exception e -> ([],[])
-
-let unify_type env evd flags mv c =
- let mvty = Typing.meta_type evd mv in
- if occur_meta_or_existential mvty or is_arity env (evars_of evd) mvty then
- unify_to_type env evd flags c mvty
- else ([],[])
+ 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 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,[],[])
(* Move metas that may need coercion at the end of the list of instances *)
@@ -518,7 +633,7 @@ let order_metas metas =
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
- if not b then error_cannot_unify env (evars_of evd) (mkEvar ev,rhs);
+ if not b then error_cannot_unify env evd (mkEvar ev,rhs);
let (evd,b) = Evarconv.consider_remaining_unif_problems env evd in
if not b then error "Cannot solve unification constraints";
evd
@@ -527,30 +642,30 @@ let solve_simple_evar_eqn env evd ev rhs =
or in evars, possibly generating new unification problems; if [b]
is true, unification of types of metas is required *)
-let w_merge env with_types flags (metas,evars) evd =
+let w_merge env with_types flags (evd,metas,evars) =
let rec w_merge_rec evd metas evars eqns =
(* Process evars *)
match evars with
| ((evn,_ as ev),rhs)::evars' ->
if is_defined_evar evd ev then
- let v = Evd.existential_value (evars_of evd) ev in
- let (metas',evars'') =
- unify_0 env (evars_of evd) topconv flags rhs v in
+ let v = Evd.existential_value evd ev in
+ let (evd,metas',evars'') =
+ unify_0 env evd topconv flags rhs v in
w_merge_rec evd (metas'@metas) (evars''@evars') eqns
else begin
let rhs' = subst_meta_instances metas rhs in
match kind_of_term rhs with
| App (f,cl) when is_mimick_head f & occur_meta rhs' ->
if occur_evar evn rhs' then
- error_occur_check env (evars_of evd) evn rhs';
+ error_occur_check env evd evn rhs';
let evd' = mimick_evar evd flags f (Array.length cl) evn in
w_merge_rec evd' metas evars eqns
| _ ->
w_merge_rec (solve_simple_evar_eqn env evd ev rhs')
metas evars' eqns
end
- | [] ->
+ | [] ->
(* Process metas *)
match metas with
@@ -559,52 +674,52 @@ let w_merge env with_types flags (metas,evars) evd =
if with_types & to_type <> TypeProcessed then
if to_type = CoerceToType then
(* Some coercion may have to be inserted *)
- (w_coerce env evd mv c,([],[])),[]
+ (w_coerce env evd mv c,([],[])),eqns
else
(* No coercion needed: delay the unification of types *)
- ((evd,c),([],[])),(mv,c)::eqns
- else
+ ((evd,c),([],[])),(mv,status,c)::eqns
+ else
((evd,c),([],[])),eqns in
- if meta_defined evd mv then
- let {rebus=c'},(status',_) = meta_fvalue evd mv in
- let (take_left,st,(metas',evars')) =
- merge_instances env (evars_of evd) flags status' status c' c
- in
- let evd' =
- if take_left then evd
- else meta_reassign mv (c,(st,TypeProcessed)) evd
- in
- 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
- | [] ->
+ if meta_defined evd mv then
+ let {rebus=c'},(status',_) = meta_fvalue evd mv in
+ let (take_left,st,(evd,metas',evars')) =
+ merge_instances env evd flags status' status c' c
+ in
+ let evd' =
+ if take_left then evd
+ else meta_reassign mv (c,(st,TypeProcessed)) evd
+ in
+ 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
+ | [] ->
(* Process type eqns *)
match eqns with
- | (mv,c)::eqns ->
- let (metas,evars) = unify_type env evd flags mv c in
+ | (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 (evars_of evd) sp in
+ let ev = Evd.find 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 (mc,ec) =
- unify_0 sp_env (evars_of evd') Cumul flags
- (Retyping.get_type_of_with_meta sp_env (evars_of evd') (metas_of evd') c) ev.evar_concl in
- let evd'' = w_merge_rec evd' mc ec [] in
- if (evars_of evd') == (evars_of evd'')
- then Evd.evar_define sp c evd''
- else Evd.evar_define sp (Evarutil.nf_evar (evars_of evd'') c) evd'' in
+ let (evd'',mc,ec) =
+ unify_0 sp_env evd' Cumul flags
+ (Retyping.get_type_of sp_env evd' c) ev.evar_concl in
+ let evd''' = w_merge_rec evd'' mc ec [] in
+ if evd' == evd'''
+ then Evd.define sp c evd'''
+ else Evd.define sp (Evarutil.nf_evar evd''' c) evd''' in
(* merge constraints *)
w_merge_rec evd (order_metas metas) evars []
let w_unify_meta_types env ?(flags=default_unify_flags) evd =
let metas,evd = retract_coercible_metas evd in
- w_merge env true flags (metas,[]) evd
+ w_merge env true flags (evd,metas,[])
(* [w_unify env evd M N]
performs a unification of M and N, generating a bunch of
@@ -616,38 +731,46 @@ let w_unify_meta_types env ?(flags=default_unify_flags) evd =
[clenv_typed_unify M N clenv] expects in addition that expected
types of metavars are unifiable with the types of their instances *)
-let check_types env evd flags subst m n =
- let sigma = evars_of evd in
- if isEvar (fst (whd_stack sigma m)) or isEvar (fst (whd_stack sigma n)) then
- unify_0_with_initial_metas subst true env (evars_of evd) topconv
+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
flags
- (Retyping.get_type_of_with_meta env sigma (metas_of evd) m)
- (Retyping.get_type_of_with_meta env sigma (metas_of evd) n)
- else
- subst
+ (Retyping.get_type_of env sigma n)
+ (Retyping.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
+ flags
+ (Retyping.get_type_of env sigma m)
+ (Retyping.get_type_of env sigma n)
+ else subst
let w_unify_core_0 env with_types cv_pb flags m n evd =
let (mc1,evd') = retract_coercible_metas evd in
- let subst1 = check_types env evd flags (mc1,[]) m n in
+ let (sigma,ms,es) = check_types env flags (evd,mc1,[]) m n in
let subst2 =
- unify_0_with_initial_metas subst1 true env (evars_of evd') cv_pb flags m n
- in
- w_merge env with_types flags subst2 evd'
+ unify_0_with_initial_metas (evd',ms,es) true 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
let w_unify_0 env = w_unify_core_0 env false
let w_typed_unify env = w_unify_core_0 env true
-
(* 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
FAIL because we cannot find a binding *)
let iter_fail f a =
- let n = Array.length a in
+ let n = Array.length a in
let rec ffail i =
- if i = n then error "iter_fail"
+ if i = n then error "iter_fail"
else
- try f a.(i)
+ try f a.(i)
with ex when precatchable_exception ex -> ffail (i+1)
in ffail 0
@@ -657,105 +780,173 @@ let iter_fail f a =
let w_unify_to_subterm env ?(flags=default_unify_flags) (op,cl) evd =
let rec matchrec cl =
let cl = strip_outer_cast cl in
- (try
- if closed0 cl
- then w_unify_0 env topconv flags op cl evd,cl
+ (try
+ if closed0 cl && not (isEvar cl)
+ then w_typed_unify env topconv flags op cl evd,cl
else error "Bound 1"
with ex when precatchable_exception ex ->
- (match kind_of_term cl with
+ (match kind_of_term cl with
| App (f,args) ->
let n = Array.length args in
assert (n>0);
let c1 = mkApp (f,Array.sub args 0 (n-1)) in
let c2 = args.(n-1) in
- (try
+ (try
matchrec c1
- with ex when precatchable_exception ex ->
+ with ex when precatchable_exception ex ->
matchrec c2)
| Case(_,_,c,lf) -> (* does not search in the predicate *)
- (try
+ (try
matchrec c
- with ex when precatchable_exception ex ->
+ with ex when precatchable_exception ex ->
iter_fail matchrec lf)
- | LetIn(_,c1,_,c2) ->
- (try
+ | LetIn(_,c1,_,c2) ->
+ (try
matchrec c1
- with ex when precatchable_exception ex ->
+ with ex when precatchable_exception ex ->
matchrec c2)
- | Fix(_,(_,types,terms)) ->
- (try
+ | Fix(_,(_,types,terms)) ->
+ (try
iter_fail matchrec types
- with ex when precatchable_exception ex ->
+ with ex when precatchable_exception ex ->
iter_fail matchrec terms)
-
- | CoFix(_,(_,types,terms)) ->
- (try
+
+ | CoFix(_,(_,types,terms)) ->
+ (try
iter_fail matchrec types
- with ex when precatchable_exception ex ->
+ with ex when precatchable_exception ex ->
iter_fail matchrec terms)
| Prod (_,t,c) ->
- (try
- matchrec t
- with ex when precatchable_exception ex ->
+ (try
+ matchrec t
+ with ex when precatchable_exception ex ->
matchrec c)
| Lambda (_,t,c) ->
- (try
- matchrec t
- with ex when precatchable_exception ex ->
+ (try
+ matchrec t
+ with ex when precatchable_exception ex ->
matchrec c)
- | _ -> error "Match_subterm"))
- in
+ | _ -> error "Match_subterm"))
+ in
try matchrec cl
with ex when precatchable_exception ex ->
raise (PretypeError (env,NoOccurrenceFound (op, None)))
-let w_unify_to_subterm_list env flags allow_K oplist t evd =
- List.fold_right
+(* 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 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
+ in
+ let fail str _ = error str in
+ let bind f g a =
+ let a1 = try f a
+ with ex
+ when precatchable_exception ex -> a
+ in try g a1
+ with ex
+ when precatchable_exception ex -> a1
+ in
+ let bind_iter f a =
+ let n = Array.length a in
+ let rec ffail i =
+ if i = n then fun a -> a
+ else bind (f a.(i)) (ffail (i+1))
+ in ffail 0
+ in
+ let rec matchrec cl =
+ let cl = strip_outer_cast cl in
+ (bind
+ (if closed0 cl
+ then return (fun () -> w_typed_unify env topconv flags op cl evd,cl)
+ else fail "Bound 1")
+ (match kind_of_term cl with
+ | App (f,args) ->
+ let n = Array.length args in
+ assert (n>0);
+ let c1 = mkApp (f,Array.sub args 0 (n-1)) in
+ let c2 = args.(n-1) in
+ bind (matchrec c1) (matchrec c2)
+
+ | Case(_,_,c,lf) -> (* does not search in the predicate *)
+ bind (matchrec c) (bind_iter matchrec lf)
+
+ | LetIn(_,c1,_,c2) ->
+ bind (matchrec c1) (matchrec c2)
+
+ | Fix(_,(_,types,terms)) ->
+ bind (bind_iter matchrec types) (bind_iter matchrec terms)
+
+ | CoFix(_,(_,types,terms)) ->
+ bind (bind_iter matchrec types) (bind_iter matchrec terms)
+
+ | Prod (_,t,c) ->
+ bind (matchrec t) (matchrec c)
+
+ | Lambda (_,t,c) ->
+ bind (matchrec t) (matchrec c)
+
+ | _ -> fail "Match_subterm"))
+ in
+ let res = matchrec cl [] in
+ if res = [] then
+ raise (PretypeError (env,NoOccurrenceFound (op, None)))
+ else
+ res
+
+let w_unify_to_subterm_list env flags allow_K hdmeta oplist t evd =
+ 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)
- else error "Match_subterm"
+ else error_abstraction_over_meta env evd hdmeta (destMeta op)
else if occur_meta_or_existential op then
let (evd',cl) =
- try
+ 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)
- in
- (evd',cl::l)
+ with PretypeError (env,NoOccurrenceFound _) when allow_K -> (evd,op)
+ in
+ if not allow_K && (* 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
(evd,op::l)
else
(* This is not up to delta ... *)
raise (PretypeError (env,NoOccurrenceFound (op, None))))
- oplist
+ oplist
(evd,[])
let secondOrderAbstraction env flags allow_K typ (p, oplist) evd =
(* 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 oplist typ evd in
+ w_unify_to_subterm_list env flags allow_K p oplist typ evd 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 ([p,pred,(ConvUpToEta 0,TypeProcessed)],[]) evd'
+ w_merge env false flags (evd',[p,pred,(ConvUpToEta 0,TypeProcessed)],[])
let w_unify2 env flags allow_K cv_pb ty1 ty2 evd =
- let c1, oplist1 = whd_stack (evars_of evd) ty1 in
- let c2, oplist2 = whd_stack (evars_of evd) ty2 in
+ 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
+ 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'
| _, Meta p2 ->
(* Find the predicate *)
let evd' =
- secondOrderAbstraction env flags allow_K ty1 (p2, oplist2) evd in
+ 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'
| _ -> error "w_unify2"
@@ -782,29 +973,29 @@ let w_unify2 env flags allow_K cv_pb ty1 ty2 evd =
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 hd1,l1 = whd_stack (evars_of evd) ty1 in
- let hd2,l2 = whd_stack (evars_of evd) ty2 in
+ 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
(* Pattern case *)
| (Meta _, true, Lambda _, _ | Lambda _, _, Meta _, true)
when List.length l1 = List.length l2 ->
- (try
+ (try
w_typed_unify env cv_pb flags ty1 ty2 evd
- with ex when precatchable_exception ex ->
- try
+ 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)
-
+
(* Second order case *)
- | (Meta _, true, _, _ | _, _, Meta _, true) ->
- (try
+ | (Meta _, true, _, _ | _, _, Meta _, true) ->
+ (try
w_unify2 env flags allow_K cv_pb ty1 ty2 evd
with PretypeError (env,NoOccurrenceFound _) as e -> raise e
- | ex when precatchable_exception ex ->
- try
+ | ex when precatchable_exception ex ->
+ try
w_typed_unify env cv_pb flags ty1 ty2 evd
with ex' when precatchable_exception ex' ->
raise ex)
-
+
(* General case: try first order *)
| _ -> w_typed_unify env cv_pb flags ty1 ty2 evd
diff --git a/pretyping/unification.mli b/pretyping/unification.mli
index 89280631..cc62a9bb 100644
--- a/pretyping/unification.mli
+++ b/pretyping/unification.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: unification.mli 10856 2008-04-27 16:15:34Z herbelin $ i*)
+(*i $Id$ i*)
(*i*)
open Term
@@ -14,10 +14,12 @@ open Environ
open Evd
(*i*)
-type unify_flags = {
- modulo_conv_on_closed_terms : Names.transparent_state option;
+type unify_flags = {
+ modulo_conv_on_closed_terms : Names.transparent_state option;
use_metas_eagerly : bool;
modulo_delta : Names.transparent_state;
+ resolve_evars : bool;
+ use_evars_pattern_unification : bool
}
val default_unify_flags : unify_flags
@@ -25,20 +27,23 @@ val default_no_delta_unify_flags : unify_flags
(* The "unique" unification fonction *)
val w_unify :
- bool -> env -> conv_pb -> ?flags:unify_flags -> constr -> constr -> evar_defs -> evar_defs
+ bool -> env -> conv_pb -> ?flags:unify_flags -> constr -> constr -> evar_map -> evar_map
(* [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_defs -> evar_defs * constr
+ env -> ?flags:unify_flags -> constr * constr -> evar_map -> evar_map * constr
-val w_unify_meta_types : env -> ?flags:unify_flags -> evar_defs -> evar_defs
+val w_unify_to_subterm_all :
+ env -> ?flags:unify_flags -> constr * constr -> evar_map -> (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
[ctyp] so that its gets type [typ]; [typ] may contain metavariables *)
-val w_coerce_to_type : env -> evar_defs -> constr -> types -> types ->
- evar_defs * constr
+val w_coerce_to_type : env -> evar_map -> constr -> types -> types ->
+ evar_map * constr
(*i This should be in another module i*)
@@ -46,4 +51,4 @@ val w_coerce_to_type : env -> evar_defs -> constr -> types -> types ->
(* abstracts the terms in l over c to get a term of type t *)
(* (exported for inv.ml) *)
val abstract_list_all :
- env -> evar_defs -> constr -> constr -> constr list -> constr
+ env -> evar_map -> constr -> constr -> constr list -> constr
diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml
index 5d09570e..c894d2b5 100644
--- a/pretyping/vnorm.ml
+++ b/pretyping/vnorm.ml
@@ -6,21 +6,21 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: vnorm.ml 11424 2008-09-30 12:10:28Z jforest $ i*)
+(*i $Id$ i*)
open Names
open Declarations
open Term
open Environ
open Inductive
-open Reduction
+open Reduction
open Vm
(*******************************************)
(* Calcul de la forme normal d'un terme *)
(*******************************************)
-let crazy_type = mkSet
+let crazy_type = mkSet
let decompose_prod env t =
let (name,dom,codom as res) = destProd (whd_betadeltaiota env t) in
@@ -33,18 +33,18 @@ exception Find_at of int
[cst] = true si c'est un constructeur constant *)
let invert_tag cst tag reloc_tbl =
- try
+ try
for j = 0 to Array.length reloc_tbl - 1 do
let tagj,arity = reloc_tbl.(j) in
if tag = tagj && (cst && arity = 0 || not(cst || arity = 0)) then
raise (Find_at j)
else ()
- done;raise Not_found
- with Find_at j -> (j+1)
+ done;raise Not_found
+ with Find_at j -> (j+1)
(* Argggg, ces constructeurs de ... qui commencent a 1*)
let find_rectype_a env c =
- let (t, l) =
+ let (t, l) =
let t = whd_betadeltaiota env c in
try destApp t with _ -> (t,[||]) in
match kind_of_term t with
@@ -53,13 +53,13 @@ let find_rectype_a env c =
(* Instantiate inductives and parameters in constructor type *)
-let type_constructor mind mib typ params =
+let type_constructor mind mib typ params =
let s = ind_subst mind mib in
let ctyp = substl s typ in
let nparams = Array.length params in
if nparams = 0 then ctyp
else
- let _,ctyp = decompose_prod_n nparams ctyp in
+ let _,ctyp = decompose_prod_n nparams ctyp in
substl (List.rev (Array.to_list params)) ctyp
@@ -85,7 +85,7 @@ let construct_of_constr const env tag typ =
let ctyp = type_constructor mind mib (mip.mind_nf_lc.(i-1)) params in
(mkApp(mkConstruct(ind,i), params), ctyp)
-let construct_of_constr_const env tag typ =
+let construct_of_constr_const env tag typ =
fst (construct_of_constr true env tag typ)
let construct_of_constr_block = construct_of_constr false
@@ -94,15 +94,15 @@ let constr_type_of_idkey env idkey =
match idkey with
| ConstKey cst ->
mkConst cst, Typeops.type_of_constant env cst
- | VarKey id ->
- let (_,_,ty) = lookup_named id env in
+ | VarKey id ->
+ let (_,_,ty) = lookup_named id env in
mkVar id, ty
- | RelKey i ->
+ | RelKey i ->
let n = (nb_rel env - i) in
let (_,_,ty) = lookup_rel n env in
mkRel n, lift n ty
-let type_of_ind env ind =
+let type_of_ind env ind =
type_of_inductive env (Inductive.lookup_mind_specif env ind)
let build_branches_type env (mind,_ as _ind) mib mip params dep p =
@@ -116,7 +116,7 @@ let build_branches_type env (mind,_ as _ind) mib mip params dep p =
let nparams = Array.length params in
let carity = snd (rtbl.(i)) in
let crealargs = Array.sub cargs nparams (Array.length cargs - nparams) in
- let codom =
+ let codom =
let papp = mkApp(p,crealargs) in
if dep then
let cstr = ith_constructor_of_inductive ind (i+1) in
@@ -124,17 +124,17 @@ let build_branches_type env (mind,_ as _ind) mib mip params dep p =
let dep_cstr = mkApp(mkApp(mkConstruct cstr,params),relargs) in
mkApp(papp,[|dep_cstr|])
else papp
- in
+ in
decl, codom
in Array.mapi build_one_branch mip.mind_nf_lc
-let build_case_type dep p realargs c =
+let build_case_type dep p realargs c =
if dep then mkApp(mkApp(p, realargs), [|c|])
else mkApp(p, realargs)
(* La fonction de normalisation *)
-let rec nf_val env v t = nf_whd env (whd_val v) t
+let rec nf_val env v t = nf_whd env (whd_val v) t
and nf_vtype env v = nf_val env v crazy_type
@@ -145,18 +145,18 @@ and nf_whd env whd typ =
let dom = nf_vtype env (dom p) in
let name = Name (id_of_string "x") in
let vc = body_of_vfun (nb_rel env) (codom p) in
- let codom = nf_vtype (push_rel (name,None,dom) env) vc in
- mkProd(name,dom,codom)
+ let codom = nf_vtype (push_rel (name,None,dom) env) vc in
+ mkProd(name,dom,codom)
| Vfun f -> nf_fun env f typ
| Vfix(f,None) -> nf_fix env f
| Vfix(f,Some vargs) -> fst (nf_fix_app env f vargs)
- | Vcofix(cf,_,None) -> nf_cofix env cf
- | Vcofix(cf,_,Some vargs) ->
+ | Vcofix(cf,_,None) -> nf_cofix env cf
+ | Vcofix(cf,_,Some vargs) ->
let cfd = nf_cofix env cf in
let i,(_,ta,_) = destCoFix cfd in
let t = ta.(i) in
let _, args = nf_args env vargs t in
- mkApp(cfd,args)
+ mkApp(cfd,args)
| Vconstr_const n -> construct_of_constr_const env n typ
| Vconstr_block b ->
let capp,ctyp = construct_of_constr_block env (btag b) typ in
@@ -168,24 +168,24 @@ and nf_whd env whd typ =
| Vatom_stk(Aiddef(idkey,v), stk) ->
nf_whd env (whd_stack v stk) typ
| Vatom_stk(Aind ind, stk) ->
- nf_stk env (mkInd ind) (type_of_ind env ind) stk
-
+ nf_stk env (mkInd ind) (type_of_ind env ind) stk
+
and nf_stk env c t stk =
match stk with
| [] -> c
| Zapp vargs :: stk ->
let t, args = nf_args env vargs t in
- nf_stk env (mkApp(c,args)) t stk
- | Zfix (f,vargs) :: stk ->
+ nf_stk env (mkApp(c,args)) t stk
+ | Zfix (f,vargs) :: stk ->
let fa, typ = nf_fix_app env f vargs in
let _,_,codom = try decompose_prod env typ with _ -> exit 120 in
nf_stk env (mkApp(fa,[|c|])) (subst1 c codom) stk
- | Zswitch sw :: stk ->
+ | Zswitch sw :: stk ->
let (mind,_ as ind),allargs = find_rectype_a env t in
let (mib,mip) = Inductive.lookup_mind_specif env ind in
let nparams = mib.mind_nparams in
let params,realargs = Util.array_chop nparams allargs in
- let pT =
+ let pT =
hnf_prod_applist env (type_of_ind env ind) (Array.to_list params) in
let pT = whd_betadeltaiota env pT in
let dep, p = nf_predicate env ind mip params (type_of_switch sw) pT in
@@ -195,12 +195,12 @@ and nf_stk env c t stk =
let bsw = branch_of_switch (nb_rel env) sw in
let mkbranch i (n,v) =
let decl,codom = btypes.(i) in
- let env =
- List.fold_right
+ let env =
+ List.fold_right
(fun (name,t) env -> push_rel (name,None,t) env) decl env in
let b = nf_val env v codom in
- compose_lam decl b
- in
+ compose_lam decl b
+ in
let branchs = Array.mapi mkbranch bsw in
let tcase = build_case_type dep p realargs c in
let ci = case_info sw in
@@ -212,10 +212,10 @@ and nf_predicate env ind mip params v pT =
let k = nb_rel env in
let vb = body_of_vfun k f in
let name,dom,codom = try decompose_prod env pT with _ -> exit 121 in
- let dep,body =
+ let dep,body =
nf_predicate (push_rel (name,None,dom) env) ind mip params vb codom in
dep, mkLambda(name,dom,body)
- | Vfun f, _ ->
+ | Vfun f, _ ->
let k = nb_rel env in
let vb = body_of_vfun k f in
let name = Name (id_of_string "c") in
@@ -226,12 +226,12 @@ and nf_predicate env ind mip params v pT =
let body = nf_vtype (push_rel (name,None,dom) env) vb in
true, mkLambda(name,dom,body)
| _, _ -> false, nf_val env v crazy_type
-
+
and nf_args env vargs t =
let t = ref t in
let len = nargs vargs in
- let args =
- Array.init len
+ let args =
+ Array.init len
(fun i ->
let _,dom,codom = try decompose_prod env !t with _ -> exit 123 in
let c = nf_val env (arg vargs i) dom in
@@ -242,8 +242,8 @@ and nf_bargs env b t =
let t = ref t in
let len = bsize b in
let args =
- Array.init len
- (fun i ->
+ Array.init len
+ (fun i ->
let _,dom,codom = try decompose_prod env !t with _ -> exit 124 in
let c = nf_val env (bfield b i) dom in
t := subst1 c codom; c) in
@@ -252,7 +252,7 @@ and nf_bargs env b t =
and nf_fun env f typ =
let k = nb_rel env in
let vb = body_of_vfun k f in
- let name,dom,codom =
+ let name,dom,codom =
try decompose_prod env typ
with _ ->
raise (Type_errors.TypeError(env,Type_errors.ReferenceVariables typ))
@@ -268,17 +268,17 @@ and nf_fix env f =
let ndef = Array.length vt in
let ft = Array.map (fun v -> nf_val env v crazy_type) vt in
let name = Array.init ndef (fun _ -> (Name (id_of_string "Ffix"))) in
- let env = push_rec_types (name,ft,ft) env in
+ let env = push_rec_types (name,ft,ft) env in
let fb = Util.array_map2 (fun v t -> nf_fun env v t) vb ft in
mkFix ((rec_args,init),(name,ft,fb))
-
+
and nf_fix_app env f vargs =
let fd = nf_fix env f in
let (_,i),(_,ta,_) = destFix fd in
let t = ta.(i) in
let t, args = nf_args env vargs t in
mkApp(fd,args),t
-
+
and nf_cofix env cf =
let init = current_cofix cf in
let k = nb_rel env in
@@ -286,15 +286,15 @@ and nf_cofix env cf =
let ndef = Array.length vt in
let cft = Array.map (fun v -> nf_val env v crazy_type) vt in
let name = Array.init ndef (fun _ -> (Name (id_of_string "Fcofix"))) in
- let env = push_rec_types (name,cft,cft) env in
+ let env = push_rec_types (name,cft,cft) env in
let cfb = Util.array_map2 (fun v t -> nf_val env v t) vb cft in
mkCoFix (init,(name,cft,cfb))
-
+
let cbv_vm env c t =
let transp = transp_values () in
- if not transp then set_transp_values true;
+ if not transp then set_transp_values true;
let v = Vconv.val_of_constr env c in
let c = nf_val env v t in
- if not transp then set_transp_values false;
+ if not transp then set_transp_values false;
c
-
+