From a0cfa4f118023d35b767a999d5a2ac4b082857b4 Mon Sep 17 00:00:00 2001 From: Samuel Mimram Date: Fri, 25 Jul 2008 15:12:53 +0200 Subject: Imported Upstream version 8.2~beta3+dfsg --- pretyping/cases.ml | 1298 ++++++++++++++++++++++---------------- pretyping/cases.mli | 21 +- pretyping/classops.ml | 18 +- pretyping/classops.mli | 5 +- pretyping/clenv.ml | 450 ++++++------- pretyping/clenv.mli | 52 +- pretyping/coercion.ml | 181 +++--- pretyping/coercion.mli | 5 +- pretyping/detyping.ml | 139 ++-- pretyping/detyping.mli | 8 +- pretyping/evarconv.ml | 365 +++++------ pretyping/evarutil.ml | 1211 ++++++++++++++++++++++------------- pretyping/evarutil.mli | 45 +- pretyping/evd.ml | 269 ++++++-- pretyping/evd.mli | 102 ++- pretyping/indrec.ml | 152 ++--- pretyping/indrec.mli | 4 +- pretyping/inductiveops.ml | 48 +- pretyping/inductiveops.mli | 8 +- pretyping/matching.ml | 7 +- pretyping/pattern.ml | 20 +- pretyping/pretype_errors.ml | 28 +- pretyping/pretype_errors.mli | 18 +- pretyping/pretyping.ml | 421 +++++++------ pretyping/pretyping.mli | 9 +- pretyping/rawterm.ml | 66 +- pretyping/rawterm.mli | 32 +- pretyping/recordops.ml | 53 +- pretyping/recordops.mli | 17 +- pretyping/reductionops.ml | 144 ++++- pretyping/reductionops.mli | 16 +- pretyping/retyping.ml | 30 +- pretyping/tacred.ml | 524 ++++++++------- pretyping/tacred.mli | 26 +- pretyping/termops.ml | 232 ++++--- pretyping/termops.mli | 71 ++- pretyping/typeclasses.ml | 457 ++++++++++++++ pretyping/typeclasses.mli | 105 +++ pretyping/typeclasses_errors.ml | 55 ++ pretyping/typeclasses_errors.mli | 44 ++ pretyping/typing.ml | 18 +- pretyping/unification.ml | 583 ++++++++++++----- pretyping/unification.mli | 26 +- pretyping/vnorm.ml | 39 +- 44 files changed, 4762 insertions(+), 2660 deletions(-) create mode 100644 pretyping/typeclasses.ml create mode 100644 pretyping/typeclasses.mli create mode 100644 pretyping/typeclasses_errors.ml create mode 100644 pretyping/typeclasses_errors.mli (limited to 'pretyping') diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 1f7bdad3..1d7da3f3 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: cases.ml 10071 2007-08-10 15:34:41Z herbelin $ *) +(* $Id: cases.ml 11085 2008-06-10 08:54:43Z herbelin $ *) open Util open Names @@ -20,12 +20,12 @@ open Sign open Reductionops open Typeops open Type_errors - open Rawterm open Retyping open Pretype_errors open Evarutil open Evarconv +open Evd (* Pattern-matching errors *) @@ -63,16 +63,48 @@ let error_wrong_predicate_arity_loc loc env c n1 n2 = let error_needs_inversion env x t = raise (PatternMatchingError (env, NeedsInversion (x,t))) +(**********************************************************************) +(* Functions to deal with impossible cases *) + +let impossible_default_case = ref None + +let set_impossible_default_clause c = impossible_default_case := Some c + +let coq_unit_judge = + let na1 = Name (id_of_string "A") in + let na2 = Name (id_of_string "H") in + fun () -> + match !impossible_default_case with + | Some (id,type_of_id) -> + make_judge id type_of_id + | None -> + (* In case the constants id/ID are not defined *) + make_judge (mkLambda (na1,mkProp,mkLambda(na2,mkRel 1,mkRel 1))) + (mkProd (na1,mkProp,mkArrow (mkRel 1) (mkRel 2))) + +(**********************************************************************) + module type S = sig val compile_cases : - loc -> - (type_constraint -> env -> rawconstr -> unsafe_judgment) * - Evd.evar_defs ref -> + loc -> case_style -> + (type_constraint -> env -> evar_defs ref -> rawconstr -> unsafe_judgment) * evar_defs ref -> type_constraint -> - env -> rawconstr option * tomatch_tuple * cases_clauses -> + env -> rawconstr option * tomatch_tuples * cases_clauses -> unsafe_judgment end +let rec list_try_compile f = function + | [a] -> f a + | [] -> anomaly "try_find_f" + | h::t -> + try f h + with UserError _ | TypeError _ | PretypeError _ + | Stdpp.Exc_located (_,(UserError _ | TypeError _ | PretypeError _)) -> + list_try_compile f t + +let force_name = + let nx = Name (id_of_string "x") in function Anonymous -> nx | na -> na + (************************************************************************) (* Pattern-matching compilation (Cases) *) (************************************************************************) @@ -86,8 +118,8 @@ let mssg_may_need_inversion () = str "Found a matching with no clauses on a term unknown to have an empty inductive type" (* Utils *) -let make_anonymous_patvars = - list_tabulate (fun _ -> PatVar (dummy_loc,Anonymous)) +let make_anonymous_patvars n = + list_make n (PatVar (dummy_loc,Anonymous)) (* Environment management *) let push_rels vars env = List.fold_right push_rel vars env @@ -112,77 +144,53 @@ type alias_constr = let mkSpecialLetInJudge j (na,(deppat,nondeppat,d,t)) = { uj_val = - (match d with - | DepAlias -> mkLetIn (na,deppat,t,j.uj_val) - | NonDepAlias -> - if (not (dependent (mkRel 1) j.uj_type)) - or (* A leaf: *) isRel deppat - then - (* The body of pat is not needed to type j - see *) - (* insert_aliases - and both deppat and nondeppat have the *) - (* same type, then one can freely substitute one by the other *) - subst1 nondeppat j.uj_val - else - (* The body of pat is not needed to type j but its value *) - (* is dependent in the type of j; our choice is to *) - (* enforce this dependency *) - mkLetIn (na,deppat,t,j.uj_val)); + if + isRel deppat or not (dependent (mkRel 1) j.uj_val) or + d = NonDepAlias & not (dependent (mkRel 1) j.uj_type) + then + (* The body of pat is not needed to type j - see *) + (* insert_aliases - and both deppat and nondeppat have the *) + (* same type, then one can freely substitute one by the other *) + subst1 nondeppat j.uj_val + else + (* The body of pat is not needed to type j but its value *) + (* is dependent in the type of j; our choice is to *) + (* enforce this dependency *) + mkLetIn (na,deppat,t,j.uj_val); uj_type = subst1 deppat j.uj_type } (**********************************************************************) (* Structures used in compiling pattern-matching *) -type 'a lifted = int * 'a - -let insert_lifted a = (0,a);; -type rhs = +type 'a rhs = { rhs_env : env; + rhs_vars : identifier list; avoid_ids : identifier list; - rhs_lift : int; - it : rawconstr } + it : 'a option} -type equation = - { dependencies : constr lifted list; - patterns : cases_pattern list; - rhs : rhs; +type 'a equation = + { patterns : cases_pattern list; + rhs : 'a rhs; alias_stack : name list; eqn_loc : loc; - used : bool ref; - tag : pattern_source } + used : bool ref } + +type 'a matrix = 'a equation list -type matrix = equation list +type dep_status = KnownDep | KnownNotDep | DepUnknown (* 1st argument of IsInd is the original ind before extracting the summary *) type tomatch_type = - | IsInd of types * inductive_type + | IsInd of types * inductive_type * name list | NotInd of constr option * types type tomatch_status = - | Pushed of ((constr * tomatch_type) * int list) + | Pushed of ((constr * tomatch_type) * int list * (name * dep_status)) | Alias of (constr * constr * alias_constr * constr) | Abstract of rel_declaration type tomatch_stack = tomatch_status list -(* The type [predicate_signature] types the terms to match and the rhs: - - - [PrLetIn (names,dep,pred)] types a pushed term ([Pushed]), - if dep<>Anonymous, the term is dependent, let n=|names|, if - n<>0 then the type of the pushed term is necessarily an - inductive with n real arguments. Otherwise, it may be - non inductive, or inductive without real arguments, or inductive - originating from a subterm in which case real args are not dependent; - it accounts for n+1 binders if dep or n binders if not dep - - [PrProd] types abstracted term ([Abstract]); it accounts for one binder - - [PrCcl] types the right-hand-side - - Aliases [Alias] have no trace in [predicate_signature] -*) - -type predicate_signature = - | PrLetIn of (name list * name) * predicate_signature - | PrProd of predicate_signature - | PrCcl of constr - (* We keep a constr for aliases and a cases_pattern for error message *) type alias_builder = @@ -251,7 +259,7 @@ let push_history_pattern n current cont = (* A pattern-matching problem has the following form: - env, isevars |- Cases tomatch of mat end + env, evd |- Cases tomatch of mat end where tomatch is some sequence of "instructions" (t1 ... tn) @@ -281,15 +289,17 @@ let push_history_pattern n current cont = of variables). *) -type pattern_matching_problem = - { env : env; - isevars : Evd.evar_defs ref; - pred : predicate_signature option; - tomatch : tomatch_stack; - history : pattern_continuation; - mat : matrix; - caseloc : loc; - typing_function: type_constraint -> env -> rawconstr -> unsafe_judgment } + +type 'a pattern_matching_problem = + { env : env; + evdref : evar_defs 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 } (*--------------------------------------------------------------------------* * A few functions to infer the inductive type from the patterns instead of * @@ -312,123 +322,129 @@ let rec find_row_ind = function | PatVar _ :: l -> find_row_ind l | PatCstr(loc,c,_,_) :: _ -> Some (loc,c) -let inductive_template isevars env tmloc ind = +let inductive_template evdref env tmloc ind = let arsign = get_full_arity_sign env ind in let hole_source = match tmloc with - | Some loc -> fun i -> (loc, Evd.TomatchTypeParameter (ind,i)) - | None -> fun _ -> (dummy_loc, Evd.InternalHole) in + | Some loc -> fun i -> (loc, TomatchTypeParameter (ind,i)) + | None -> fun _ -> (dummy_loc, InternalHole) in let (_,evarl,_) = List.fold_right (fun (na,b,ty) (subst,evarl,n) -> match b with | None -> let ty' = substl subst ty in - let e = e_new_evar isevars env ~src:(hole_source n) ty' in + let e = e_new_evar evdref env ~src:(hole_source n) ty' in (e::subst,e::evarl,n+1) | Some b -> (b::subst,evarl,n+1)) arsign ([],[],1) in applist (mkInd ind,List.rev evarl) -let inh_coerce_to_ind isevars env ty tyi = - let expected_typ = inductive_template isevars env None tyi in - (* devrait être indifférent d'exiger leq ou pas puisque pour - un inductif cela doit être égal *) - let _ = e_cumul env isevars expected_typ ty in () +let try_find_ind env sigma typ realnames = + let (IndType(_,realargs) as ind) = find_rectype env sigma typ in + let names = + match realnames with + | Some names -> names + | None -> list_make (List.length realargs) Anonymous in + IsInd (typ,ind,names) -let unify_tomatch_with_patterns isevars env loc typ pats = + +let inh_coerce_to_ind evdref env ty tyi = + let expected_typ = inductive_template evdref env None tyi in + (* devrait être indifférent d'exiger leq ou pas puisque pour + un inductif cela doit être égal *) + let _ = e_cumul env evdref expected_typ ty in () + +let unify_tomatch_with_patterns evdref env loc typ pats realnames = match find_row_ind pats with | None -> NotInd (None,typ) | Some (_,(ind,_)) -> - inh_coerce_to_ind isevars env typ ind; - try IsInd (typ,find_rectype env (Evd.evars_of !isevars) typ) + inh_coerce_to_ind evdref env typ ind; + try try_find_ind env (evars_of !evdref) typ realnames with Not_found -> NotInd (None,typ) -let find_tomatch_tycon isevars env loc = function +let find_tomatch_tycon evdref env loc = function (* Try if some 'in I ...' is present and can be used as a constraint *) - | Some (_,ind,_,_) -> mk_tycon (inductive_template isevars env loc ind) - | None -> empty_tycon + | Some (_,ind,_,realnal) -> + mk_tycon (inductive_template evdref env loc ind),Some (List.rev realnal) + | None -> + empty_tycon,None -let coerce_row typing_fun isevars env pats (tomatch,(_,indopt)) = +let coerce_row typing_fun evdref env pats (tomatch,(_,indopt)) = let loc = Some (loc_of_rawconstr tomatch) in - let tycon = find_tomatch_tycon isevars env loc indopt in - let j = typing_fun tycon env tomatch in - let typ = nf_evar (Evd.evars_of !isevars) j.uj_type 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 t = - try IsInd (typ,find_rectype env (Evd.evars_of !isevars) typ) + try try_find_ind env (evars_of !evdref) typ realnames with Not_found -> - unify_tomatch_with_patterns isevars env loc typ pats in + unify_tomatch_with_patterns evdref env loc typ pats realnames in (j.uj_val,t) -let coerce_to_indtype typing_fun isevars env matx tomatchl = +let coerce_to_indtype typing_fun evdref env matx tomatchl = let pats = List.map (fun r -> r.patterns) matx in let matx' = match matrix_transpose pats with | [] -> List.map (fun _ -> []) tomatchl (* no patterns at all *) | m -> m in - List.map2 (coerce_row typing_fun isevars env) matx' tomatchl + List.map2 (coerce_row typing_fun evdref env) matx' tomatchl (************************************************************************) (* Utils *) -let mkExistential env ?(src=(dummy_loc,Evd.InternalHole)) isevars = - e_new_evar isevars env ~src:src (new_Type ()) +let mkExistential env ?(src=(dummy_loc,InternalHole)) evdref = + e_new_evar evdref env ~src:src (new_Type ()) -let evd_comb2 f isevars x y = - let (evd',y) = f !isevars x y in - isevars := evd'; +let evd_comb2 f evdref x y = + let (evd',y) = f !evdref x y in + evdref := evd'; y module Cases_F(Coercion : Coercion.S) : S = struct -let adjust_tomatch_to_pattern pb ((current,typ),deps) = +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 the first pattern type and forget about the others *) - let typ = match typ with IsInd (t,_) -> t | NotInd (_,t) -> t in + let typ,names = + match typ with IsInd(t,_,names) -> t,Some names | NotInd(_,t) -> t,None in let typ = - try IsInd (typ,find_rectype pb.env (Evd.evars_of !(pb.isevars)) typ) + try try_find_ind pb.env (evars_of !(pb.evdref)) typ names with Not_found -> NotInd (None,typ) in - let tomatch = ((current,typ),deps) in + let tomatch = ((current,typ),deps,dep) in match typ with | NotInd (None,typ) -> let tm1 = List.map (fun eqn -> List.hd eqn.patterns) pb.mat in (match find_row_ind tm1 with | None -> tomatch | Some (_,(ind,_)) -> - let indt = inductive_template pb.isevars pb.env None ind in + let indt = inductive_template pb.evdref pb.env None ind in let current = if deps = [] & isEvar typ then (* Don't insert coercions if dependent; only solve evars *) - let _ = e_cumul pb.env pb.isevars indt typ in + let _ = e_cumul pb.env pb.evdref indt typ in current else (evd_comb2 (Coercion.inh_conv_coerce_to dummy_loc pb.env) - pb.isevars (make_judge current typ) (mk_tycon_type indt)).uj_val in - let sigma = Evd.evars_of !(pb.isevars) in - let typ = IsInd (indt,find_rectype pb.env sigma indt) in - ((current,typ),deps)) + pb.evdref (make_judge current typ) (mk_tycon_type indt)).uj_val in + let sigma = evars_of !(pb.evdref) in + let typ = try_find_ind pb.env sigma indt names in + ((current,typ),deps,dep)) | _ -> tomatch - (* extract some ind from [t], possibly coercing from constructors in [tm] *) -let to_mutind env isevars tm c t = -(* match c with - | Some body -> *) NotInd (c,t) -(* | None -> unify_tomatch_with_patterns isevars env t tm*) - let type_of_tomatch = function - | IsInd (t,_) -> t + | IsInd (t,_,_) -> t | NotInd (_,t) -> t let mkDeclTomatch na = function - | IsInd (t,_) -> (na,None,t) + | IsInd (t,_,_) -> (na,None,t) | NotInd (c,t) -> (na,c,t) let map_tomatch_type f = function - | IsInd (t,ind) -> IsInd (f t,map_inductive_type f ind) - | NotInd (c,t) -> NotInd (option_map f c, f t) + | IsInd (t,ind,names) -> IsInd (f t,map_inductive_type f ind,names) + | NotInd (c,t) -> NotInd (Option.map f c, f t) let liftn_tomatch_type n depth = map_tomatch_type (liftn n depth) let lift_tomatch_type n = liftn_tomatch_type n 1 @@ -464,25 +480,6 @@ let remove_current_pattern eqn = let prepend_pattern tms eqn = {eqn with patterns = tms@eqn.patterns } -(**********************************************************************) -(* Dealing with regular and default patterns *) -let is_regular eqn = eqn.tag = RegularPat - -let lower_pattern_status = function - | RegularPat -> DefaultPat 0 - | DefaultPat n -> DefaultPat (n+1) - -let pattern_status pats = - if array_exists ((=) RegularPat) pats then RegularPat - else - let min = - Array.fold_right - (fun pat n -> match pat with - | DefaultPat i when i i - | _ -> n) - pats 0 in - DefaultPat min - (**********************************************************************) (* Well-formedness tests *) (* Partial check on patterns *) @@ -541,7 +538,7 @@ let extract_rhs pb = | [] -> errorlabstrm "build_leaf" (mssg_may_need_inversion()) | eqn::_ -> set_used_pattern eqn; - eqn.tag, eqn.rhs + eqn.rhs (**********************************************************************) (* Functions to deal with matrix factorization *) @@ -549,23 +546,35 @@ let extract_rhs pb = let occur_in_rhs na rhs = match na with | Anonymous -> false - | Name id -> occur_rawconstr id rhs.it + | Name id -> List.mem id rhs.rhs_vars -let is_dep_patt eqn = function +let is_dep_patt_in eqn = function | PatVar (_,name) -> occur_in_rhs name eqn.rhs | PatCstr _ -> true -let dependencies_in_rhs nargs eqns = - if eqns = [] then list_tabulate (fun _ -> false) nargs (* Only "_" patts *) - else - let deps = List.map (fun (tms,eqn) -> List.map (is_dep_patt eqn) tms) eqns in - let columns = matrix_transpose deps in - List.map (List.exists ((=) true)) columns +let mk_dep_patt_row (pats,eqn) = + List.map (is_dep_patt_in eqn) pats + +let dependencies_in_pure_rhs nargs eqns = + if eqns = [] then list_make nargs false (* Only "_" patts *) else + let deps_rows = List.map mk_dep_patt_row eqns in + let deps_columns = matrix_transpose deps_rows in + List.map (List.exists ((=) true)) deps_columns let dependent_decl a = function | (na,None,t) -> dependent a t | (na,Some c,t) -> dependent a t || dependent a c +let rec dep_in_tomatch n = function + | (Pushed _ | Alias _) :: l -> dep_in_tomatch n l + | Abstract d :: l -> dependent_decl (mkRel n) d or dep_in_tomatch (n+1) l + | [] -> false + +let dependencies_in_rhs nargs current tms eqns = + match kind_of_term current with + | Rel n when dep_in_tomatch n tms -> list_make nargs true + | _ -> dependencies_in_pure_rhs nargs eqns + (* Computing the matrix of dependencies *) (* We are in context d1...dn |- and [find_dependencies k 1 nextlist] @@ -609,36 +618,41 @@ let find_dependencies_signature deps_in_rhs typs = let regeneralize_index_tomatch n = let rec genrec depth = function - | [] -> [] - | Pushed ((c,tm),l)::rest -> + | [] -> + [] + | Pushed ((c,tm),l,dep) :: rest -> let c = regeneralize_index n depth c in let tm = map_tomatch_type (regeneralize_index n depth) tm in let l = List.map (regeneralize_rel n depth) l in - Pushed ((c,tm),l)::(genrec depth rest) - | Alias (c1,c2,d,t)::rest -> - Alias (regeneralize_index n depth c1,c2,d,t)::(genrec depth rest) - | Abstract d::rest -> + Pushed ((c,tm),l,dep) :: genrec depth rest + | Alias (c1,c2,d,t) :: rest -> + Alias (regeneralize_index n depth c1,c2,d,t) :: genrec depth rest + | Abstract d :: rest -> Abstract (map_rel_declaration (regeneralize_index n depth) d) - ::(genrec (depth+1) rest) in + :: genrec (depth+1) rest in genrec 0 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 +let length_of_tomatch_type_sign (dep,_) = function + | NotInd _ -> if dep<>Anonymous then 1 else 0 + | IsInd (_,_,names) -> List.length names + if dep<>Anonymous then 1 else 0 + let replace_tomatch n c = let rec replrec depth = function | [] -> [] - | Pushed ((b,tm),l)::rest -> + | Pushed ((b,tm),l,dep) :: rest -> let b = replace_term n c depth b in let tm = map_tomatch_type (replace_term n c depth) tm in List.iter (fun i -> if i=n+depth then anomaly "replace_tomatch") l; - Pushed ((b,tm),l)::(replrec depth rest) - | Alias (c1,c2,d,t)::rest -> - Alias (replace_term n c depth c1,c2,d,t)::(replrec depth rest) - | Abstract d::rest -> + Pushed ((b,tm),l,dep) :: replrec depth rest + | Alias (c1,c2,d,t) :: rest -> + Alias (replace_term n c depth c1,c2,d,t) :: replrec depth rest + | Abstract d :: rest -> Abstract (map_rel_declaration (replace_term n c depth) d) - ::(replrec (depth+1) rest) in + :: replrec (depth+1) rest in replrec 0 let liftn_rel_declaration n k = map_rel_declaration (liftn n k) @@ -646,11 +660,11 @@ let substnl_rel_declaration sigma k = map_rel_declaration (substnl sigma k) let rec liftn_tomatch_stack n depth = function | [] -> [] - | Pushed ((c,tm),l)::rest -> + | Pushed ((c,tm),l,dep)::rest -> let c = liftn n depth c in let tm = liftn_tomatch_type n depth tm in let l = List.map (fun i -> if i Alias (liftn n depth c1,liftn n depth c2,d,liftn n depth t) ::(liftn_tomatch_stack n depth rest) @@ -658,7 +672,6 @@ let rec liftn_tomatch_stack n depth = function Abstract (map_rel_declaration (liftn n depth) d) ::(liftn_tomatch_stack n (depth+1) rest) - let lift_tomatch_stack n = liftn_tomatch_stack n 1 (* if [current] has type [I(p1...pn u1...um)] and we consider the case @@ -686,7 +699,7 @@ let merge_name get_name obj = function let merge_names get_name = List.map2 (merge_name get_name) let get_names env sign eqns = - let names1 = list_tabulate (fun _ -> Anonymous) (List.length sign) in + let names1 = list_make (List.length sign) Anonymous in (* If any, we prefer names used in pats, from top to bottom *) let names2 = List.fold_right @@ -753,17 +766,17 @@ let insert_aliases_eqn sign eqnnames alias_rest eqn = rhs = {eqn.rhs with rhs_env = push_rels thissign eqn.rhs.rhs_env } } let insert_aliases env sigma alias eqns = - (* Là, y a une faiblesse, si un alias est utilisé dans un cas par *) - (* défaut présent mais inutile, ce qui est le cas général, l'alias *) - (* est introduit même s'il n'est pas utilisé dans les cas réguliers *) + (* Là, y a une faiblesse, si un alias est utilisé dans un cas par *) + (* défaut présent mais inutile, ce qui est le cas général, l'alias *) + (* est introduit même s'il n'est pas utilisé dans les cas réguliers *) let eqnsnames = List.map (fun eqn -> List.hd eqn.alias_stack) eqns in let alias_rests = List.map (fun eqn -> List.tl eqn.alias_stack) eqns in - (* names2 takes the meet of all needed aliases *) - let names2 = + (* name2 takes the meet of all needed aliases *) + let name2 = List.fold_right (merge_name (fun x -> x)) eqnsnames Anonymous in (* Only needed aliases are kept by build_aliases_context *) let eqnsnames, sign1, sign2, env = - build_aliases_context env sigma [names2] eqnsnames [alias] in + build_aliases_context env sigma [name2] eqnsnames [alias] in let eqns = list_map3 (insert_aliases_eqn sign1) eqnsnames alias_rests eqns in sign2, env, eqns @@ -771,28 +784,13 @@ let insert_aliases env sigma alias eqns = (* Functions to deal with elimination predicate *) exception Occur -let noccur_between_without_evar n m term = +let noccur_between_without_evar n m term = let rec occur_rec n c = match kind_of_term c with | Rel p -> if n<=p && p () | _ -> iter_constr_with_binders succ occur_rec n c in - try occur_rec n term; true with Occur -> false - -(* Infering the predicate *) -let prepare_unif_pb typ cs = - let n = List.length (assums_of_rel_context cs.cs_args) in - - (* We may need to invert ci if its parameters occur in typ *) - let typ' = - if noccur_between_without_evar 1 n typ then lift (-n) typ - else (* TODO4-1 *) - error "Unable to infer return clause of this pattern-matching problem" in - let args = extended_rel_list (-n) cs.cs_args in - let ci = applist (mkConstruct cs.cs_cstr, cs.cs_params@args) in - - (* This is the problem: finding P s.t. cs_args |- (P realargs ci) = typ' *) - (Array.map (lift (-n)) cs.cs_concl_realargs, ci, typ') + (m = 0) or (try occur_rec n term; true with Occur -> false) (* Infering the predicate *) @@ -823,8 +821,9 @@ the following n+1 equations: Some hints: -- Clearly, if xij occurs in Ti, then, a "match z with (Ci xi1..xipi) => ..." - should be inserted somewhere in Ti. +- Clearly, if xij occurs in Ti, then, a "match z with (Ci xi1..xipi) + => ..." or a "psi(yk)", with psi extracting xij from uik, should be + inserted somewhere in Ti. - If T is undefined, an easy solution is to insert a "match z with (Ci xi1..xipi) => ..." in front of each Ti @@ -834,111 +833,21 @@ Some hints: - The main problem is what to do when an existential variables is encountered -let prepare_unif_pb typ cs = - let n = cs.cs_nargs in - let _,p = decompose_prod_n n typ in - let ci = build_dependent_constructor cs in - (* This is the problem: finding P s.t. cs_args |- (P realargs ci) = p *) - (n, cs.cs_concl_realargs, ci, p) - -let eq_operator_lift k (n,n') = function - | OpRel p, OpRel p' when p > k & p' > k -> - if p < k+n or p' < k+n' then false else p - n = p' - n' - | op, op' -> op = op' - -let rec transpose_args n = - if n=0 then [] - else - (Array.map (fun l -> List.hd l) lv):: - (transpose_args (m-1) (Array.init (fun l -> List.tl l))) - -let shift_operator k = function OpLambda _ | OpProd _ -> k+1 | _ -> k - -let reloc_operator (k,n) = function OpRel p when p > k -> -let rec unify_clauses k pv = - let pv'= Array.map (fun (n,sign,_,p) -> n,splay_constr (whd_betaiotaevar (push_rels (List.rev sign) env) (Evd.evars_of isevars)) p) pv in - let n1,op1 = let (n1,(op1,args1)) = pv'.(0) in n1,op1 in - if Array.for_all (fun (ni,(opi,_)) -> eq_operator_lift k (n1,ni) (op1,opi)) pv' - then - let argvl = transpose_args (List.length args1) pv' in - let k' = shift_operator k op1 in - let argl = List.map (unify_clauses k') argvl in - gather_constr (reloc_operator (k,n1) op1) argl -*) - -let abstract_conclusion typ cs = - let n = List.length (assums_of_rel_context cs.cs_args) in - let (sign,p) = decompose_prod_n n typ in - lam_it p sign - -let infer_predicate loc env isevars typs cstrs indf = - (* Il faudra substituer les isevars a un certain moment *) - if Array.length cstrs = 0 then (* "TODO4-3" *) - error "Inference of annotation for empty inductive types not implemented" - else - (* Empiric normalization: p may depend in a irrelevant way on args of the*) - (* cstr as in [c:{_:Alpha & Beta}] match c with (existS a b)=>(a,b) end *) - let typs = - Array.map (local_strong (whd_betaevar empty_env (Evd.evars_of !isevars))) typs - in - let eqns = array_map2 prepare_unif_pb typs cstrs in - (* First strategy: no dependencies at all *) -(* - let (mis,_) = dest_ind_family indf in - let (cclargs,_,typn) = eqns.(mis_nconstr mis -1) in -*) - let (sign,_) = get_arity env indf in - let mtyp = - if array_exists is_Type typs then - (* Heuristic to avoid comparison between non-variables algebric univs*) - new_Type () - else - mkExistential env ~src:(loc, Evd.CasesType) isevars - in - if array_for_all (fun (_,_,typ) -> e_cumul env isevars typ mtyp) eqns - then - (* Non dependent case -> turn it into a (dummy) dependent one *) - let sign = (Anonymous,None,build_dependent_inductive env indf)::sign in - let pred = it_mkLambda_or_LetIn (lift (List.length sign) mtyp) sign in - (true,pred) (* true = dependent -- par défaut *) - else -(* - let s = get_sort_of env (evars_of isevars) typs.(0) in - let predpred = it_mkLambda_or_LetIn (mkSort s) sign in - let caseinfo = make_default_case_info mis in - let brs = array_map2 abstract_conclusion typs cstrs in - let predbody = mkCase (caseinfo, (nf_betaiota predpred), mkRel 1, brs) in - let pred = it_mkLambda_or_LetIn (lift (List.length sign) mtyp) sign in -*) - (* "TODO4-2" *) - (* We skip parameters *) - let cis = - Array.map - (fun cs -> - applist (mkConstruct cs.cs_cstr, extended_rel_list 0 cs.cs_args)) - cstrs in - let ct = array_map2 (fun ci (_,_,t) -> (ci,t)) cis eqns in - raise_pattern_matching_error (loc,env, CannotInferPredicate ct) -(* - (true,pred) *) (* Propagation of user-provided predicate through compilation steps *) -let rec map_predicate f k = function - | PrCcl ccl -> PrCcl (f k ccl) - | PrProd pred -> - PrProd (map_predicate f (k+1) pred) - | PrLetIn ((names,dep as tm),pred) -> - let k' = List.length names + (if dep<>Anonymous then 1 else 0) in - PrLetIn (tm, map_predicate f (k+k') pred) - -let rec noccurn_predicate k = function - | PrCcl ccl -> noccurn k ccl - | PrProd pred -> noccurn_predicate (k+1) pred - | PrLetIn ((names,dep),pred) -> - let k' = List.length names + (if dep<>Anonymous then 1 else 0) in - noccurn_predicate (k+k') pred +let rec map_predicate f k ccl = function + | [] -> f k ccl + | Pushed ((_,tm),_,dep) :: rest -> + let k' = length_of_tomatch_type_sign dep tm in + map_predicate f (k+k') ccl rest + | Alias _ :: rest -> + map_predicate f k ccl rest + | Abstract _ :: rest -> + map_predicate f (k+1) ccl rest + +let noccurn_predicate = map_predicate noccurn let liftn_predicate n = map_predicate (liftn n) @@ -949,26 +858,19 @@ let regeneralize_index_predicate n = map_predicate (regeneralize_index n) 0 let substnl_predicate sigma = map_predicate (substnl sigma) (* This is parallel bindings *) -let subst_predicate (args,copt) pred = +let subst_predicate (args,copt) ccl tms = let sigma = match copt with | None -> List.rev args | Some c -> c::(List.rev args) in - substnl_predicate sigma 0 pred - -let specialize_predicate_var (cur,typ) = function - | PrProd _ | PrCcl _ -> - anomaly "specialize_predicate_var: a pattern-variable must be pushed" - | PrLetIn (([],dep),pred) -> - subst_predicate ([],if dep<>Anonymous then Some cur else None) pred - | PrLetIn ((_,dep),pred) -> - (match typ with - | IsInd (_,IndType (_,realargs)) -> - subst_predicate (realargs,if dep<>Anonymous then Some cur else None) pred - | _ -> anomaly "specialize_predicate_var") - -let ungeneralize_predicate = function - | PrLetIn _ | PrCcl _ -> anomaly "ungeneralize_predicate: expects a product" - | PrProd pred -> pred + substnl_predicate sigma 0 ccl tms + +let specialize_predicate_var (cur,typ,dep) tms ccl = + let c = if dep<>Anonymous then Some cur else None in + let l = + match typ with + | IsInd (_,IndType(_,realargs),names) -> if names<>[] then realargs else [] + | NotInd _ -> [] in + subst_predicate (l,c) ccl tms (*****************************************************************************) (* We have pred = [X:=realargs;x:=c]P typed in Gamma1, x:I(realargs), Gamma2 *) @@ -979,85 +881,77 @@ let ungeneralize_predicate = function (* We first need to lift t(x) s.t. it is typed in Gamma, X:=rargs, x' *) (* then we have to replace x by x' in t(x) and y by y' in P *) (*****************************************************************************) -let generalize_predicate c ny d = function - | PrLetIn ((names,dep as tm),pred) -> - if dep=Anonymous then anomaly "Undetected dependency"; - let p = List.length names + 1 in - let pred = lift_predicate 1 pred in - let pred = regeneralize_index_predicate (ny+p+1) pred in - PrLetIn (tm, PrProd pred) - | PrProd _ | PrCcl _ -> - anomaly "generalize_predicate: expects a non trivial pattern" - -let rec extract_predicate l = function - | pred, Alias (deppat,nondeppat,_,_)::tms -> +let generalize_predicate (names,(nadep,_)) ny d tms ccl = + if nadep=Anonymous then anomaly "Undetected dependency"; + let p = List.length names + 1 in + let ccl = lift_predicate 1 ccl tms in + regeneralize_index_predicate (ny+p+1) ccl tms + +let rec extract_predicate l ccl = function + | Alias (deppat,nondeppat,_,_)::tms -> let tms' = match kind_of_term nondeppat with | Rel i -> replace_tomatch i deppat tms | _ -> (* initial terms are not dependent *) tms in - extract_predicate l (pred,tms') - | PrProd pred, Abstract d'::tms -> + extract_predicate l ccl tms' + | Abstract d'::tms -> let d' = map_rel_declaration (lift (List.length l)) d' in - substl l (mkProd_or_LetIn d' (extract_predicate [] (pred,tms))) - | PrLetIn (([],dep),pred), Pushed ((cur,_),_)::tms -> - extract_predicate (if dep<>Anonymous then cur::l else l) (pred,tms) - | PrLetIn ((_,dep),pred), Pushed ((cur,IsInd (_,(IndType(_,realargs)))),_)::tms -> + substl l (mkProd_or_LetIn d' (extract_predicate [] ccl tms)) + | Pushed ((cur,NotInd _),_,(dep,_))::tms -> + extract_predicate (if dep<>Anonymous then cur::l else l) ccl tms + | Pushed ((cur,IsInd (_,IndType(_,realargs),_)),_,(dep,_))::tms -> let l = List.rev realargs@l in - extract_predicate (if dep<>Anonymous then cur::l else l) (pred,tms) - | PrCcl ccl, [] -> + extract_predicate (if dep<>Anonymous then cur::l else l) ccl tms + | [] -> substl l ccl - | _ -> anomaly"extract_predicate: predicate inconsistent with terms to match" - -let abstract_predicate env sigma indf cur tms = function - | (PrProd _ | PrCcl _) -> anomaly "abstract_predicate: must be some LetIn" - | PrLetIn ((names,dep),pred) -> - let sign = make_arity_signature env true indf in - (* n is the number of real args + 1 *) - let n = List.length sign in - let tms = lift_tomatch_stack n tms in - let tms = - match kind_of_term cur with - | Rel i -> regeneralize_index_tomatch (i+n) tms - | _ -> (* Initial case *) tms in - (* Depending on whether the predicate is dependent or not, and has real - args or not, we lift it to make room for [sign] *) - (* Even if not intrinsically dep, we move the predicate into a dep one *) - let sign,q = - if names = [] & n <> 1 then - (* Real args were not considered *) - (if dep<>Anonymous then - (let (_,c,t) = List.hd sign in (dep,c,t)::List.tl sign) - else - sign),n - else - (* Real args are OK *) - (List.map2 (fun na (_,c,t) -> (na,c,t)) (dep::names) sign,1) in - let q,k = if dep <> Anonymous then (q-1,2) else (q,1) in - let pred = liftn_predicate q k pred in - let pred = extract_predicate [] (pred,tms) in - (true, it_mkLambda_or_LetIn_name env pred sign) - -let rec known_dependent = function - | None -> false - | Some (PrLetIn ((_,dep),_)) -> dep<>Anonymous - | Some (PrCcl _) -> false - | Some (PrProd _) -> - anomaly "known_dependent: can only be used when patterns remain" + +let abstract_predicate env sigma indf cur (names,(nadep,_)) tms ccl = + let sign = make_arity_signature env true indf in + (* n is the number of real args + 1 *) + let n = List.length sign in + let tms = lift_tomatch_stack n tms in + let tms = + match kind_of_term cur with + | Rel i -> regeneralize_index_tomatch (i+n) tms + | _ -> (* Initial case *) tms in + let sign = List.map2 (fun na (_,c,t) -> (na,c,t)) (nadep::names) sign in + let ccl = if nadep <> Anonymous then ccl else lift_predicate 1 ccl tms in + let pred = extract_predicate [] ccl tms in + it_mkLambda_or_LetIn_name env pred sign + +let known_dependent (_,dep) = (dep = KnownDep) (* [expand_arg] is used by [specialize_predicate] it replaces gamma, x1...xn, x1...xk |- pred by gamma, x1...xn, x1...xk-1 |- [X=realargs,xk=xk]pred (if dep) or by gamma, x1...xn, x1...xk-1 |- [X=realargs]pred (if not dep) *) -let expand_arg n alreadydep (na,t) deps (k,pred) = - (* current can occur in pred even if the original problem is not dependent *) - let dep = - if alreadydep<>Anonymous then alreadydep - else if deps = [] && noccurn_predicate 1 pred then Anonymous - else Name (id_of_string "x") in - let pred = if dep<>Anonymous then pred else lift_predicate (-1) pred in - (* There is no dependency in realargs for subpattern *) - (k-1, PrLetIn (([],dep), pred)) +let expand_arg tms ccl ((_,t),_,na) = + 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 + | 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); + (* 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 = []; + it = None }; + alias_stack = Anonymous::aliasnames; + eqn_loc = dummy_loc; + used = ref false } ] + | _ -> + submat + else + submat (*****************************************************************************) (* pred = [X:=realargs;x:=c]P types the following problem: *) @@ -1073,37 +967,35 @@ let expand_arg n alreadydep (na,t) deps (k,pred) = (* s.t Gamma,x1'..xn' |- match Pushed(x1')..Pushed(xn') rest with..end :pred'*) (* *) (*****************************************************************************) -let specialize_predicate tomatchs deps cs = function - | (PrProd _ | PrCcl _) -> - anomaly "specialize_predicate: a matched pattern must be pushed" - | PrLetIn ((names,isdep),pred) -> - (* Assume some gamma st: gamma, (X,x:=realargs,copt) |- pred *) - let nrealargs = List.length names in - let k = nrealargs + (if isdep<>Anonymous then 1 else 0) in - (* We adjust pred st: gamma, x1..xn, (X,x:=realargs,copt) |- pred' *) - let n = cs.cs_nargs in - let pred' = liftn_predicate n (k+1) pred in - let argsi = if nrealargs <> 0 then Array.to_list cs.cs_concl_realargs else [] in - let copti = if isdep<>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 |- pred'' *) - let pred'' = subst_predicate (argsi, copti) pred' in - (* We adjust pred st: gamma, x1..xn, x1..xn |- pred'' *) - let pred''' = liftn_predicate n (n+1) pred'' in - (* We finally get gamma,x1..xn |- [X1,x1:=R1,x1]..[Xn,xn:=Rn,xn]pred'''*) - snd (List.fold_right2 (expand_arg n isdep) tomatchs deps (n,pred''')) - -let find_predicate loc env isevars p typs cstrs current - (IndType (indf,realargs)) tms = - let (dep,pred) = - match p with - | Some p -> abstract_predicate env (Evd.evars_of !isevars) indf current tms p - | None -> infer_predicate loc env isevars typs cstrs indf in - let typ = whd_beta (applist (pred, realargs)) in - if dep then - (pred, whd_beta (applist (typ, [current])), new_Type ()) - else - (pred, typ, new_Type ()) +let specialize_predicate newtomatchs (names,(depna,_)) 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 copti = if depna<>Anonymous then Some (build_dependent_constructor cs) else None in + (* The substituends argsi, copti are all defined in gamma, x1...xn *) + (* We need _parallel_ bindings to get gamma, x1...xn, tms |- ccl'' *) + let ccl'' = whd_betaiota (subst_predicate (argsi, copti) ccl' tms) in + (* We adjust ccl st: gamma, x1..xn, x1..xn, tms |- ccl'' *) + let ccl''' = liftn_predicate n (n+1) ccl'' tms in + (* We finally get gamma,x1..xn |- [X1,x1:=R1,x1]..[Xn,xn:=Rn,xn]pred'''*) + List.fold_left (expand_arg tms) ccl''' newtomatchs + +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 (applist (pred, realargs@[current])), new_Type ()) + +let adjust_predicate_from_tomatch ((_,oldtyp),_,(nadep,_)) typ pb = + match typ, oldtyp with + | IsInd (_,_,names), NotInd _ -> + let k = if nadep <> Anonymous then 2 else 1 in + let n = List.length names in + { pb with pred = liftn_predicate n k pb.pred pb.tomatch } + | _ -> + pb (************************************************************************) (* Sorting equations by constructor *) @@ -1144,7 +1036,6 @@ let group_equations pb ind current cstrs mat = (* This is a default clause that we expand *) for i=1 to Array.length cstrs do let args = make_anonymous_patvars cstrs.(i-1).cs_nargs in - let rest = {rest with tag = lower_pattern_status rest.tag} in brs.(i-1) <- (args, rest) :: brs.(i-1) done | PatCstr (loc,((_,i)),args,_) -> @@ -1157,40 +1048,37 @@ let group_equations pb ind current cstrs mat = (* Here starts the pattern-matching compilation algorithm *) (* Abstracting over dependent subterms to match *) -let rec generalize_problem pb current = function +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 pb' = generalize_problem pb current l in + 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 - { pb with + { pb' with tomatch = Abstract d :: tomatch; - pred = option_map (generalize_predicate current i d) pb'.pred } + pred = generalize_predicate names i d pb.tomatch pb'.pred } (* No more patterns: typing the right-hand-side of equations *) let build_leaf pb = - let tag, rhs = extract_rhs pb in - let tycon = match pb.pred with - | None -> empty_tycon - | Some (PrCcl typ) -> mk_tycon typ - | Some _ -> anomaly "not all parameters of pred have been consumed" in - tag, pb.typing_function tycon rhs.rhs_env rhs.it + 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 (* Building the sub-problem when all patterns are variables *) -let shift_problem (current,t) pb = +let shift_problem ((current,t),_,(nadep,_)) pb = {pb with tomatch = Alias (current,current,NonDepAlias,type_of_tomatch t)::pb.tomatch; - pred = option_map (specialize_predicate_var (current,t)) pb.pred; + pred = specialize_predicate_var (current,t,nadep) pb.tomatch pb.pred; history = push_history_pattern 0 AliasLeaf pb.history; mat = List.map remove_current_pattern pb.mat } (* Building the sub-pattern-matching problem for a given branch *) -let build_branch current deps pb eqns const_info = +let build_branch current deps (realnames,dep) pb 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 pb.pred) & deps = [] + & not (known_dependent dep) & deps = [] then NonDepAlias else @@ -1202,24 +1090,21 @@ let build_branch current deps pb eqns const_info = pb.history in (* We find matching clauses *) - let cs_args = (*assums_of_rel_context*) const_info.cs_args in + let cs_args = const_info.cs_args in let names = get_names pb.env cs_args eqns in let submat = List.map (fun (tms,eqn) -> prepend_pattern tms eqn) eqns in - if submat = [] then - raise_pattern_matching_error - (dummy_loc, pb.env, NonExhaustive (complete_history history)); let typs = List.map2 (fun (_,c,t) na -> (na,c,t)) cs_args names in let _,typs',_ = List.fold_right (fun (na,c,t as d) (env,typs,tms) -> - let tm1 = List.map List.hd tms in let tms = List.map List.tl tms in - (push_rel d env, (na,to_mutind env pb.isevars tm1 c t)::typs,tms)) + (push_rel d env, (na,NotInd(c,t))::typs,tms)) typs (pb.env,[],List.map fst eqns) in let dep_sign = find_dependencies_signature - (dependencies_in_rhs const_info.cs_nargs eqns) (List.rev typs) in + (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 terms is relative to the current context enriched by topushs *) @@ -1231,11 +1116,25 @@ let build_branch current deps pb eqns const_info = (* into "Gamma; typs; curalias |- tms" *) let tomatch = lift_tomatch_stack const_info.cs_nargs pb.tomatch in - let currents = + let typs'' = list_map2_i - (fun i (na,t) deps -> Pushed ((mkRel i, lift_tomatch_type i t), deps)) + (fun i (na,t) deps -> + let dep = match dep with + | Name _ as na',k -> (if na <> Anonymous then na else na'),k + | Anonymous,KnownNotDep -> + if deps = [] && noccurn_predicate 1 pb.pred tomatch then + (Anonymous,KnownNotDep) + else + (force_name na,KnownDep) + | _,_ -> anomaly "Inconsistent dependency" in + ((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 currents = List.map (fun x -> Pushed x) typs'' in + let sign = List.map (fun (na,t) -> mkDeclTomatch na t) typs' in let ind = @@ -1246,12 +1145,18 @@ let build_branch current deps pb eqns const_info = let cur_alias = lift (List.length sign) current in let currents = Alias (ci,cur_alias,alias_type,ind) :: currents in + let tomatch = List.rev_append currents tomatch in + + let submat = adjust_impossible_cases pb pred tomatch submat in + if submat = [] then + raise_pattern_matching_error + (dummy_loc, pb.env, NonExhaustive (complete_history history)); sign, { pb with env = push_rels sign pb.env; - tomatch = List.rev_append currents tomatch; - pred = option_map (specialize_predicate (List.rev typs') dep_sign const_info) pb.pred; + tomatch = tomatch; + pred = pred; history = history; mat = List.map (push_rels_eqn_with_names sign) submat } @@ -1264,9 +1169,6 @@ let build_branch current deps pb eqns const_info = "Pushed" terms and types are relative to env "Abstract" types are relative to env enriched by the previous terms to match - Concretely, each term "c" or type "T" comes with a delayed lift - index, but it works as if the lifting were effective. - *) (**********************************************************************) @@ -1279,12 +1181,13 @@ let rec compile pb = | [] -> build_leaf pb and match_current pb tomatch = - let ((current,typ as ct),deps) = adjust_tomatch_to_pattern pb tomatch in + let ((current,typ),deps,dep as ct) = adjust_tomatch_to_pattern pb tomatch in + let pb = adjust_predicate_from_tomatch tomatch typ pb in match typ with | NotInd (_,typ) -> check_all_variables typ pb.mat; compile (shift_problem ct pb) - | IsInd (_,(IndType(indf,realargs) as indt)) -> + | IsInd (_,(IndType(indf,realargs) as indt),names) -> let mind,_ = dest_ind_family indf in let cstrs = get_constructors pb.env indf in let eqns,onlydflt = group_equations pb mind current cstrs pb.mat in @@ -1294,46 +1197,41 @@ and match_current pb tomatch = let _constraints = Array.map (solve_constraints indt) cstrs in (* We generalize over terms depending on current term to match *) - let pb = generalize_problem pb current deps in + let pb = generalize_problem (names,dep) pb deps in (* We compile branches *) - let brs = array_map2 (compile_branch current deps pb) eqns cstrs in + let brs = array_map2 (compile_branch current (names,dep) deps pb) eqns cstrs in (* We build the (elementary) case analysis *) - let tags = Array.map (fun (t,_,_) -> t) brs in - let brvals = Array.map (fun (_,v,_) -> v) brs in - let brtyps = Array.map (fun (_,_,t) -> t) brs in + let brvals = Array.map (fun (v,_) -> v) brs in let (pred,typ,s) = - find_predicate pb.caseloc pb.env pb.isevars - pb.pred brtyps cstrs current indt pb.tomatch in - let ci = make_case_info pb.env mind RegularStyle tags in + 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 pred,current,brvals) in let inst = List.map mkRel deps in - pattern_status tags, { uj_val = applist (case, inst); uj_type = substl inst typ } -and compile_branch current deps pb eqn cstr = - let sign, pb = build_branch current deps pb eqn cstr in - let tag, j = compile pb in - (tag, it_mkLambda_or_LetIn j.uj_val sign, j.uj_type) +and compile_branch current names deps pb eqn cstr = + let sign, pb = build_branch current deps names pb eqn cstr in + let j = compile pb in + (it_mkLambda_or_LetIn j.uj_val sign, j.uj_type) and compile_generalization pb d rest = let pb = { pb with env = push_rel d pb.env; tomatch = rest; - pred = option_map ungeneralize_predicate pb.pred; mat = List.map (push_rels_eqn [d]) pb.mat } in - let patstat,j = compile pb in - patstat, + let j = compile pb in { uj_val = mkLambda_or_LetIn d j.uj_val; uj_type = mkProd_or_LetIn d j.uj_type } and compile_alias pb (deppat,nondeppat,d,t) rest = let history = simplify_history pb.history in let sign, newenv, mat = - insert_aliases pb.env (Evd.evars_of !(pb.isevars)) (deppat,nondeppat,d,t) pb.mat in + insert_aliases pb.env (evars_of !(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 *) @@ -1349,15 +1247,14 @@ and compile_alias pb (deppat,nondeppat,d,t) rest = {pb with env = newenv; tomatch = tomatch; - pred = option_map (lift_predicate n) pb.pred; + pred = lift_predicate n pb.pred tomatch; history = history; mat = mat } in - let patstat,j = compile pb in - patstat, + let j = compile pb in List.fold_left mkSpecialLetInJudge j sign (* pour les alias des initiaux, enrichir les env de ce qu'il faut et -substituer après par les initiaux *) +substituer après par les initiaux *) (**************************************************************************) (* Preparation of the pattern-matching problem *) @@ -1371,12 +1268,10 @@ let matx_of_eqns env tomatchl eqns = let initial_rhs = rhs in let rhs = { rhs_env = env; + rhs_vars = free_rawvars initial_rhs; avoid_ids = ids@(ids_of_named_context (named_context env)); - rhs_lift = 0; - it = initial_rhs } in - { dependencies = []; - patterns = initial_lpat; - tag = RegularPat; + it = Some initial_rhs } in + { patterns = initial_lpat; alias_stack = []; eqn_loc = loc; used = ref false; @@ -1386,9 +1281,9 @@ let matx_of_eqns env tomatchl eqns = (************************************************************************) (* preparing the elimination predicate if any *) -let build_expected_arity env isevars isdep tomatchl = +let build_expected_arity env evdref isdep tomatchl = let cook n = function - | _,IsInd (_,IndType(indf,_)) -> + | _,IsInd (_,IndType(indf,_),_) -> let indf' = lift_inductive_family n indf in Some (build_dependent_inductive env indf', fst (get_arity env indf')) | _,NotInd _ -> None @@ -1415,7 +1310,7 @@ let build_expected_arity env isevars isdep tomatchl = let extract_predicate_conclusion isdep tomatchl pred = let cook = function - | _,IsInd (_,IndType(_,args)) -> Some (List.length args) + | _,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 @@ -1445,9 +1340,9 @@ let set_arity_signature dep n arsign tomatchl pred x = 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) -> + | 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 + | 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) @@ -1458,7 +1353,7 @@ let set_arity_signature dep n arsign tomatchl pred x = | _ -> (RApp (dummy_loc,p,[a]))) in let rec decomp_block avoid p = function | ([], _) -> x := Some p - | ((_,IsInd (_,IndType(indf,realargs)))::l),(y::l') -> + | ((_,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 @@ -1479,9 +1374,240 @@ let set_arity_signature dep n arsign tomatchl pred x = in decomp_block [] pred (tomatchl,arsign) -let prepare_predicate_from_tycon loc dep env isevars tomatchs sign c = +(***************** Building an inversion predicate ************************) + +(* Let "match t1 in I1 u11..u1n_1 ... tm in Im um1..umn_m with ... end : T" + be a pattern-matching problem. We assume that the each uij can be + decomposed under the form pij(vij1..vijq_ij) where pij(aij1..aijq_ij) + is a pattern depending on some variables aijk and the vijk are + instances of these variables. We also assume that each ti has the + form of a pattern qi(wi1..wiq_i) where qi(bi1..biq_i) is a pattern + depending on some variables bik and the wik are instances of these + 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 + T = U(..v1jk..t1 .. ..vmjk..tm). This a higher-order matching + problem with a priori different solution (one of them if T itself!). + + We finally invert the uij and the ti and build the return clause + + phi(x11..x1n_1y1..xm1..xmn_mym) = + match x11..x1n_1 y1 .. xm1..xmn_m ym with + | p11..p1n_1 q1 .. pm1..pmn_m qm => U(..a1jk..b1 .. ..amjk..bm) + | _ .. _ _ .. _ .. _ _ => True + end + + so that "phi(u11..u1n_1t1..um1..umn_mtm) = T" (note that the clause + returning True never happens and any inhabited type can be put instead). +*) + +let adjust_to_extended_env_and_remove_deps env extenv subst t = + let 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 + 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 + [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 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)) + | None -> + (n,ty) in + let (p,ty) = aux p (lookup_rel p (rel_context extenv)) in + if noccur_between_without_evar 1 (n'-p-n+1) ty + then + let u = lift (n'-n) u in + (p,u,(expand_vars_in_term extenv u,lift p ty)) + else + failwith "") subst in + let t0 = lift (n'-n) t in + (subst0,t0) + +(* 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 + * "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 + * a subset u1..ul of the vijk and ti by a special evar + * ?id(x=t;c1:=c1,..,cl=cl) defined in context Gamma0,x,c1,...,cl |- ?id + * (where the c1..cl are the aijk and bi matching the u1..ul), and + * similarly for each ti. +*) + +let abstract_tycon loc env evdref subst _tycon extenv t = + let t = nf_betaiota t in (* it helps in some cases to remove K-redex... *) + let sigma = evars_of !evdref in + let subst0,t0 = adjust_to_extended_env_and_remove_deps env extenv subst t in + (* We traverse the type T of the original problem Xi looking for subterms + that match the non-constructor part of the constraints (this part + is in subst); these subterms are the "good" subterms and we replace them + by an evar that may depend (and only depend) on the corresponding + convertible subterms of the substitution *) + let rec aux (k,env,subst as x) t = + let 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 = + list_map_i + (fun i _ -> if List.mem i vl then u else mkRel i) 1 + (rel_context extenv) in + let rel_filter = + List.map (fun a -> not (isRel a) or dependent a u) inst in + let named_filter = + List.map (fun (id,_,_) -> dependent (mkVar id) u) + (named_context extenv) in + let filter = rel_filter@named_filter in + let ev = + e_new_evar evdref extenv ~src:(loc, CasesType) ~filter:filter ty in + evdref := add_conv_pb (Reduction.CONV,extenv,substl inst ev,u) !evdref; + 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 + 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 + 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 = + 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 + +(* 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 + * 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 + * impossible cases. See above "Building an inversion predicate" for + * further explanations + *) + +let build_inversion_problem loc env evdref tms t = + let sigma = evars_of !evdref in + 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 + let rec reveal_pattern t (subst,avoid as acc) = + match kind_of_term (whd_betadeltaiota env sigma t) with + | Construct cstr -> PatCstr (dummy_loc,cstr,[],Anonymous), acc + | App (f,v) when isConstruct f -> + let cstr = destConstruct f in + let n = constructor_nrealargs env cstr in + let l = list_lastn n (Array.to_list v) in + let l,acc = list_fold_map' reveal_pattern l acc in + PatCstr (dummy_loc,cstr,l,Anonymous), acc + | _ -> make_patvar t acc in + let rec aux n env acc_sign tms acc = + match tms with + | [] -> [], acc_sign, acc + | (t, IsInd (_,IndType(indf,realargs),_)) :: tms -> + let patl,acc = list_fold_map' reveal_pattern realargs acc in + let pat,acc = make_patvar t acc in + let indf' = lift_inductive_family n indf in + let sign = make_arity_signature env true indf' in + let p = List.length realargs in + let env' = push_rels sign env in + let patl',acc_sign,acc = aux (n+p+1) env' (sign@acc_sign) tms acc in + patl@pat::patl',acc_sign,acc + | (t, NotInd (bo,typ)) :: tms -> + aux n env acc_sign tms acc in + let avoid0 = ids_of_context env in + (* [patl] is a list of patterns revealing the substructure of + constructors present in the constraints on the type of the + multiple terms t1..tn that are matched in the original problem; + [subst] is the substitution of the free pattern variables in + [patl] that returns the non-constructor parts of the constraints. + Especially, if the ti has type I ui1..uin_i, and the patterns associated + to ti are pi1..pin_i, then subst(pij) is uij; the substitution is + useful to recognize which subterms of the whole type T of the original + problem have to be abstracted *) + let patl,sign,(subst,avoid) = aux 0 env [] tms ([],avoid0) in + let n = List.length sign in + let (pb_env,_),sub_tms = + list_fold_map (fun (env,i) (na,b,t as d) -> + let typ = + if b<>None then NotInd(None,t) else + try try_find_ind env sigma t None + with Not_found -> NotInd (None,t) in + let ty = lift_tomatch_type (n-i) typ in + let tm = Pushed ((mkRel (n-i),ty),[],(Anonymous,KnownNotDep)) in + ((push_rel d env,i+1),tm)) + (env,0) (List.rev sign) in + let subst = List.map (fun (na,t) -> (na,lift n t)) subst in + (* [eqn1] is the first clause of the auxiliary pattern-matching that + serves as skeleton for the return type: [patl] is the + substructure of constructors extracted from the list of + constraints on the inductive types of the multiple terms matched + in the original pattern-matching problem Xi *) + let eqn1 = + { patterns = patl; + alias_stack = []; + eqn_loc = dummy_loc; + used = ref false; + rhs = { rhs_env = pb_env; + (* we assume all vars are used; in practice we discard dependent + vars so that the field rhs_vars is normally not used *) + rhs_vars = List.map fst subst; + avoid_ids = avoid; + it = Some (lift n t) } } in + (* [eqn2] is the default clause of the auxiliary pattern-matching: it will + catch the clauses of the original pattern-matching problem Xi whose + type constraints are incompatible with the constraints on the + inductive types of the multiple terms matched in Xi *) + let eqn2 = + { patterns = List.map (fun _ -> PatVar (dummy_loc,Anonymous)) patl; + alias_stack = []; + eqn_loc = dummy_loc; + used = ref false; + rhs = { rhs_env = pb_env; + rhs_vars = []; + avoid_ids = avoid0; + it = None } } in + (* [pb] is the auxiliary pattern-matching serving as skeleton for the + return type of the original problem Xi *) + let pb = + { env = pb_env; + evdref = evdref; + pred = new_Type(); + tomatch = sub_tms; + history = start_history n; + mat = [eqn1;eqn2]; + caseloc = loc; + casestyle = RegularStyle; + typing_function = build_tycon loc env pb_env subst} in + (compile pb).uj_val + +let prepare_predicate_from_tycon loc dep env evdref tomatchs sign c = let cook (n, l, env, signs) = function - | c,IsInd (_,IndType(indf,realargs)) -> + | 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 @@ -1489,59 +1615,47 @@ let prepare_predicate_from_tycon loc dep env isevars tomatchs sign c = (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 _ -> - (n, l, env, []::signs) in - let n, allargs, env, signs = List.fold_left cook (0, [], env, []) tomatchs in + | 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 - let allargs = - List.map (fun c -> lift n (nf_betadeltaiota env (Evd.evars_of !isevars) c)) allargs in - let rec build_skeleton env c = - (* Don't put into normal form, it has effects on the synthesis of evars *) - (* let c = whd_betadeltaiota env (evars_of isevars) c in *) - (* We turn all subterms possibly dependent into an evar with maximum ctxt*) - if isEvar c or List.exists (eq_constr c) allargs then - e_new_evar isevars env ~src:(loc, Evd.CasesType) - (Retyping.get_type_of env (Evd.evars_of !isevars) c) - else - map_constr_with_full_binders push_rel build_skeleton env c - in - names, build_skeleton env (lift n c) + 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 isdep allnames pred = +let build_initial_predicate knowndep allnames pred = let nar = List.fold_left (fun n names -> List.length names + n) 0 allnames in - let rec buildrec n pred = function - | [] -> PrCcl pred + let rec buildrec n pred nal = function + | [] -> List.rev nal,pred | names::lnames -> - let names' = if isdep then List.tl names else names in + let names' = List.tl names in let n' = n + List.length names' in - let pred, p, user_p = - if isdep then - if dependent (mkRel (nar-n')) pred then pred, 1, 1 - else liftn (-1) (nar-n') pred, 0, 1 - else pred, 0, 0 in + let pred, p = + if dependent (mkRel (nar-n')) pred then pred, 1 + else liftn (-1) (nar-n') pred, 0 in let na = if p=1 then let na = List.hd names in - if na = Anonymous then - (* peut arriver en raison des evars *) + ((if na = Anonymous then + (* can happen if evars occur in the return clause *) Name (id_of_string "x") (*Hum*) - else na - else Anonymous in - PrLetIn ((names',na), buildrec (n'+user_p) pred lnames) - in buildrec 0 pred allnames + else na),knowndep) + else (Anonymous,KnownNotDep) in + buildrec (n'+1) pred (na::nal) lnames + in buildrec 0 pred [] allnames let extract_arity_signature env0 tomatchl tmsign = let get_one_sign n tm (na,t) = match tm with | NotInd (bo,typ) -> (match t with - | None -> [na,option_map (lift n) bo,lift n typ] + | None -> [na,Option.map (lift n) bo,lift n typ] | Some (loc,_,_,_) -> user_err_loc (loc,"", str "Unexpected type annotation for a term of non inductive type")) - | IsInd (_,IndType(indf,realargs)) -> + | 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 @@ -1554,8 +1668,16 @@ let extract_arity_signature env0 tomatchl tmsign = or nrealargs <> List.length realnal then anomaly "Ill-formed 'in' clause in cases"; List.rev realnal - | None -> list_tabulate (fun _ -> Anonymous) nrealargs in + | None -> list_make nrealargs Anonymous in let arsign = fst (get_arity env0 indf') in +(* let na = *) +(* match na with *) +(* | Name _ -> na *) +(* | Anonymous -> *) +(* match kind_of_term term with *) +(* | Rel n -> pi1 (lookup_rel n (Environ.rel_context env0)) *) +(* | _ -> Anonymous *) +(* in *) (na,None,build_dependent_inductive env0 indf') ::(List.map2 (fun x (_,c,t) ->(x,c,t)) realnal arsign) in let rec buildrec n = function @@ -1566,14 +1688,55 @@ let extract_arity_signature env0 tomatchl tmsign = | _ -> assert false in List.rev (buildrec 0 (tomatchl,tmsign)) -let inh_conv_coerce_to_tycon loc env isevars j tycon = +let inh_conv_coerce_to_tycon loc env evdref j tycon = match tycon with | Some p -> - let (evd',j) = Coercion.inh_conv_coerce_to loc env !isevars j p in - isevars := evd'; + let (evd',j) = Coercion.inh_conv_coerce_to loc env !evdref j p in + evdref := evd'; j | None -> j +(* We put the tycon inside the arity signature, possibly discovering dependencies. *) + +let prepare_predicate_from_arsign_tycon loc env tomatchs sign arsign c = + let nar = List.fold_left (fun n sign -> List.length sign + n) 0 arsign in + let subst, len = + List.fold_left2 (fun (subst, len) (tm, tmtype) sign -> + let signlen = List.length sign in + match kind_of_term tm with + | Rel n when dependent tm c + && signlen = 1 (* The term to match is not of a dependent type itself *) -> + ((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. *) -> + (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 + | Rel n when dependent arg c -> + ((n, len) :: subst, pred len) + | _ -> (subst, pred len)) + (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 + (* Make the predicate dependent on the matched variable *) + let idx = List.assoc (n - lift) subst in + mkRel (idx + lift) + with Not_found -> + (* A variable that is not matched, lift over the arsign. *) + mkRel (n + nar)) + | _ -> + map_constr_with_binders succ predicate lift c + in predicate 0 c + (* Builds the predicate. If the predicate is dependent, its context is * made of 1+nrealargs assumptions for each matched term in an inductive @@ -1582,18 +1745,58 @@ let inh_conv_coerce_to_tycon loc env isevars j tycon = * Each matched terms are independently considered dependent or not. - * A type constraint but no annotation case: it is assumed non dependent. + * A type constraint but no annotation case: we try to specialize the + * tycon to make the predicate if it is not closed. *) -let prepare_predicate loc typing_fun isevars env tomatchs sign tycon = function +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 false env isevars tomatchs sign t - in Some (build_initial_predicate false names pred) - | _ -> None) + 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 nal2,pred2 = build_initial_predicate KnownNotDep names2 t2 in + [evdref, nal1, pred1; evdref2, nal2, pred2]) (* Some type annotation *) | Some rtntyp -> @@ -1601,50 +1804,69 @@ let prepare_predicate loc typing_fun isevars env tomatchs sign tycon = function 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 rtntyp in + let predcclj = typing_fun (mk_tycon (new_Type ())) env evdref rtntyp in let _ = - option_map (fun tycon -> - isevars := Coercion.inh_conv_coerces_to loc env !isevars predcclj.uj_val + 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 !isevars predcclj).uj_val in - Some (build_initial_predicate true allnames predccl) + let predccl = (j_nf_isevar !evdref predcclj).uj_val in + let nal,pred = build_initial_predicate KnownDep allnames predccl in + [evdref, nal, pred] (**************************************************************************) (* Main entry of the matching compilation *) -let compile_cases loc (typing_fun, isevars) (tycon : Evarutil.type_constraint) env (predopt, tomatchl, eqns)= - +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 isevars env matx tomatchl in - - (* We build the elimination predicate if any and check its consistency *) - (* with the type of arguments to match *) - let tmsign = List.map snd tomatchl in - let pred = prepare_predicate loc typing_fun isevars env tomatchs tmsign tycon predopt in - - (* We push the initial terms to match and push their alias to rhs' envs *) - (* names of aliases will be recovered from patterns (hence Anonymous here) *) - let initial_pushed = List.map (fun tm -> Pushed (tm,[])) tomatchs in + let tomatchs = coerce_to_indtype typing_fun evdref env matx tomatchl in - let pb = - { env = env; - isevars = isevars; - pred = pred; - tomatch = initial_pushed; - history = start_history (List.length initial_pushed); - mat = matx; - caseloc = loc; - typing_function = typing_fun } 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 _, j = compile pb in - (* We check for unused patterns *) - List.iter (check_unused_pattern env) matx; - inh_conv_coerce_to_tycon loc env isevars j tycon + let compile_for_one_predicate (myevdref,nal,pred) = + (* We push the initial terms to match and push their alias to rhs' envs *) + (* names of aliases will be recovered from patterns (hence Anonymous *) + (* here) *) + let initial_pushed = List.map2 (fun tm na -> Pushed(tm,[],na)) tomatchs nal in + + (* A typing function that provides with a canonical term for absurd cases*) + let typing_fun tycon env evdref = function + | Some t -> typing_fun tycon env evdref t + | None -> coq_unit_judge () in + + let pb = + { env = env; + evdref = myevdref; + pred = pred; + tomatch = initial_pushed; + history = start_history (List.length initial_pushed); + mat = matx; + caseloc = loc; + casestyle = style; + typing_function = typing_fun } in + + let j = compile pb in + evdref := !myevdref; + j in + + (* Return the term compiled with the first possible elimination *) + (* predicate for which the compilation succeeds *) + let j = list_try_compile compile_for_one_predicate preds in + + (* We check for unused patterns *) + List.iter (check_unused_pattern env) matx; + + (* We coerce to the tycon (if an elim predicate was provided) *) + inh_conv_coerce_to_tycon loc env evdref j tycon + end - diff --git a/pretyping/cases.mli b/pretyping/cases.mli index 30f68083..98923b2a 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 9976 2007-07-12 11:58:30Z msozeau $ i*) +(*i $Id: cases.mli 11014 2008-05-28 19:09:32Z herbelin $ i*) (*i*) open Util @@ -46,14 +46,27 @@ val error_wrong_predicate_arity_loc : loc -> env -> constr -> constr -> constr - val error_needs_inversion : env -> constr -> types -> 'a +val set_impossible_default_clause : constr * types -> unit (*s Compilation of pattern-matching. *) +type alias_constr = + | DepAlias + | NonDepAlias +type dep_status = KnownDep | KnownNotDep | DepUnknown +type tomatch_type = + | IsInd of types * inductive_type * name list + | NotInd of constr option * types +type tomatch_status = + | Pushed of ((constr * tomatch_type) * int list * (name * dep_status)) + | Alias of (constr * constr * alias_constr * constr) + | Abstract of rel_declaration + module type S = sig val compile_cases : - loc -> - (type_constraint -> env -> rawconstr -> unsafe_judgment) * evar_defs ref -> + loc -> case_style -> + (type_constraint -> env -> evar_defs ref -> rawconstr -> unsafe_judgment) * evar_defs ref -> type_constraint -> - env -> rawconstr option * tomatch_tuple * cases_clauses -> + env -> rawconstr option * tomatch_tuples * cases_clauses -> unsafe_judgment end diff --git a/pretyping/classops.ml b/pretyping/classops.ml index bbad005c..83ba05bb 100644 --- a/pretyping/classops.ml +++ b/pretyping/classops.ml @@ -6,11 +6,11 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: classops.ml 9257 2006-10-21 17:28:28Z herbelin $ *) +(* $Id: classops.ml 10840 2008-04-23 21:29:34Z herbelin $ *) open Util open Pp -open Options +open Flags open Names open Libnames open Nametab @@ -44,7 +44,7 @@ type coe_typ = global_reference type coe_info_typ = { coe_value : constr; coe_type : types; - coe_strength : strength; + coe_strength : locality; coe_is_identity : bool; coe_param : int } @@ -290,7 +290,7 @@ let add_coercion_in_graph (ic,source,target) = if (!ambig_paths <> []) && is_verbose () then ppnl (message_ambig !ambig_paths) -type coercion = coe_typ * strength * bool * cl_typ * cl_typ * int +type coercion = coe_typ * locality * bool * cl_typ * cl_typ * int (* Calcul de l'arité d'une classe *) @@ -304,7 +304,7 @@ let class_params = function | CL_SECVAR sp -> reference_arity_length (VarRef sp) | CL_IND sp -> reference_arity_length (IndRef sp) -(* add_class : cl_typ -> strength option -> bool -> unit *) +(* add_class : cl_typ -> locality_flag option -> bool -> unit *) let add_class cl = add_new_class cl { cl_param = class_params cl } @@ -366,12 +366,14 @@ let coercion_identity v = v.coe_is_identity (* For printing purpose *) let get_coercion_value v = v.coe_value +let pr_cl_index n = int n + let classes () = Bijint.dom !class_tab let coercions () = Gmap.rng !coercion_tab let inheritance_graph () = Gmap.to_list !inheritance_graph -let coercion_of_qualid qid = - let ref = Nametab.global qid in +let coercion_of_reference r = + let ref = Nametab.global r in if not (coercion_exists ref) then errorlabstrm "try_add_coercion" (Nametab.pr_global_env Idset.empty ref ++ str" is not a coercion"); @@ -380,7 +382,7 @@ let coercion_of_qualid qid = module CoercionPrinting = struct type t = coe_typ - let encode = coercion_of_qualid + 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") diff --git a/pretyping/classops.mli b/pretyping/classops.mli index 276b14d1..1436a11b 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 6748 2005-02-18 22:17:50Z herbelin $ i*) +(*i $Id: classops.mli 10840 2008-04-23 21:29:34Z herbelin $ i*) (*i*) open Names @@ -66,7 +66,7 @@ val class_args_of : constr -> constr list (*s [declare_coercion] adds a coercion in the graph of coercion paths *) val declare_coercion : - coe_typ -> strength -> isid:bool -> + coe_typ -> locality -> isid:bool -> src:cl_typ -> target:cl_typ -> params:int -> unit (*s Access to coercions infos *) @@ -90,6 +90,7 @@ val install_path_printer : (*s This is for printing purpose *) val string_of_class : cl_typ -> string val pr_class : cl_typ -> std_ppcmds +val pr_cl_index : cl_index -> std_ppcmds val get_coercion_value : coe_index -> constr val inheritance_graph : unit -> ((cl_index * cl_index) * inheritance_path) list val classes : unit -> cl_typ list diff --git a/pretyping/clenv.ml b/pretyping/clenv.ml index 29dbe83d..03f84809 100644 --- a/pretyping/clenv.ml +++ b/pretyping/clenv.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: clenv.ml 9665 2007-02-21 17:08:10Z herbelin $ *) +(* $Id: clenv.ml 11166 2008-06-22 13:23:35Z herbelin $ *) open Pp open Util @@ -21,18 +21,14 @@ open Reduction open Reductionops open Rawterm open Pattern -open Tacexpr open Tacred open Pretype_errors open Evarutil open Unification open Mod_subst +open Coercion.Default -(* *) -let w_coerce env c ctyp target evd = - let j = make_judge c ctyp in - let (evd',j') = Coercion.Default.inh_conv_coerce_to dummy_loc env evd j (mk_tycon_type target) in - (evd',j'.uj_val) +(* Abbreviations *) let pf_env gls = Global.env_of_context gls.it.evar_hyps let pf_type_of gls c = Typing.type_of (pf_env gls) gls.sigma c @@ -43,36 +39,31 @@ let pf_concl gl = gl.it.evar_concl (* Clausal environments *) type clausenv = { - templenv : env; - env : evar_defs; + env : env; + evd : evar_defs; templval : constr freelisted; templtyp : constr freelisted } -let cl_env ce = ce.templenv -let cl_sigma ce = evars_of ce.env +let cl_env ce = ce.env +let cl_sigma ce = evars_of ce.evd let subst_clenv sub clenv = { templval = map_fl (subst_mps sub) clenv.templval; templtyp = map_fl (subst_mps sub) clenv.templtyp; - env = subst_evar_defs_light sub clenv.env; - templenv = clenv.templenv } + evd = subst_evar_defs_light sub clenv.evd; + env = clenv.env } -let clenv_nf_meta clenv c = nf_meta clenv.env c -let clenv_meta_type clenv mv = Typing.meta_type clenv.env mv -let clenv_value clenv = meta_instance clenv.env clenv.templval -let clenv_type clenv = meta_instance clenv.env clenv.templtyp +let clenv_nf_meta clenv c = nf_meta clenv.evd c +let clenv_term clenv c = meta_instance clenv.evd c +let clenv_meta_type clenv mv = Typing.meta_type clenv.evd mv +let clenv_value clenv = meta_instance clenv.evd clenv.templval +let clenv_type clenv = meta_instance clenv.evd clenv.templtyp let clenv_hnf_constr ce t = hnf_constr (cl_env ce) (cl_sigma ce) t let clenv_get_type_of ce c = - let metamap = - List.map - (function - | (n,Clval(_,_,typ)) -> (n,typ.rebus) - | (n,Cltyp (_,typ)) -> (n,typ.rebus)) - (meta_list ce.env) in - Retyping.get_type_of_with_meta (cl_env ce) (cl_sigma ce) metamap c + Retyping.get_type_of_with_meta (cl_env ce) (cl_sigma ce) (metas_of ce.evd) c exception NotExtensibleClause @@ -84,62 +75,76 @@ let clenv_push_prod cl = let mv = new_meta () in let dep = dependent (mkRel 1) u in let na' = if dep then na else Anonymous in - let e' = meta_declare mv t ~name:na' cl.env in + let e' = meta_declare mv t ~name:na' cl.evd in let concl = if dep then subst1 (mkMeta mv) u else u in let def = applist (cl.templval.rebus,[mkMeta mv]) in { templval = mk_freelisted def; templtyp = mk_freelisted concl; - env = e'; - templenv = cl.templenv } + evd = e'; + env = cl.env } | _ -> raise NotExtensibleClause in clrec typ -let clenv_environments evd bound c = - let rec clrec (e,metas) n c = - match n, kind_of_term c with - | (Some 0, _) -> (e, List.rev metas, c) - | (n, Cast (c,_,_)) -> clrec (e,metas) n c - | (n, Prod (na,c1,c2)) -> +(* Instantiate the first [bound] products of [t] with metas (all products if + [bound] is [None]; unfold local defs *) + +let clenv_environments evd bound t = + let rec clrec (e,metas) n t = + match n, kind_of_term t with + | (Some 0, _) -> (e, List.rev metas, t) + | (n, Cast (t,_,_)) -> clrec (e,metas) n t + | (n, Prod (na,t1,t2)) -> let mv = new_meta () in - let dep = dependent (mkRel 1) c2 in + let dep = dependent (mkRel 1) t2 in let na' = if dep then na else Anonymous in - let e' = meta_declare mv c1 ~name:na' e in - clrec (e', (mkMeta mv)::metas) (option_map ((+) (-1)) n) - (if dep then (subst1 (mkMeta mv) c2) else c2) - | (n, LetIn (na,b,_,c)) -> - clrec (e,metas) (option_map ((+) (-1)) n) (subst1 b c) - | (n, _) -> (e, List.rev metas, c) + let e' = meta_declare mv t1 ~name:na' e in + clrec (e', (mkMeta mv)::metas) (Option.map ((+) (-1)) n) + (if dep then (subst1 (mkMeta mv) t2) else t2) + | (n, LetIn (na,b,_,t)) -> clrec (e,metas) n (subst1 b t) + | (n, _) -> (e, List.rev metas, t) in - clrec (evd,[]) bound c - -let clenv_environments_evars env evd bound c = - let rec clrec (e,ts) n c = - match n, kind_of_term c with - | (Some 0, _) -> (e, List.rev ts, c) - | (n, Cast (c,_,_)) -> clrec (e,ts) n c - | (n, Prod (na,c1,c2)) -> - let e',constr = Evarutil.new_evar e env c1 in - let dep = dependent (mkRel 1) c2 in - clrec (e', constr::ts) (option_map ((+) (-1)) n) - (if dep then (subst1 constr c2) else c2) - | (n, LetIn (na,b,_,c)) -> - clrec (e,ts) (option_map ((+) (-1)) n) (subst1 b c) - | (n, _) -> (e, List.rev ts, c) + clrec (evd,[]) bound t + +(* Instantiate the first [bound] products of [t] with evars (all products if + [bound] is [None]; unfold local defs *) + +let clenv_environments_evars env evd bound t = + let rec clrec (e,ts) n t = + match n, kind_of_term t with + | (Some 0, _) -> (e, List.rev ts, t) + | (n, Cast (t,_,_)) -> clrec (e,ts) n t + | (n, Prod (na,t1,t2)) -> + let e',constr = Evarutil.new_evar e env t1 in + let dep = dependent (mkRel 1) t2 in + clrec (e', constr::ts) (Option.map ((+) (-1)) n) + (if dep then (subst1 constr t2) else t2) + | (n, LetIn (na,b,_,t)) -> clrec (e,ts) n (subst1 b t) + | (n, _) -> (e, List.rev ts, t) in - clrec (evd,[]) bound c - -let mk_clenv_from_n gls n (c,cty) = - let evd = create_evar_defs gls.sigma in + clrec (evd,[]) bound t + +let clenv_conv_leq env sigma t c bound = + let ty = Retyping.get_type_of env sigma c in + let evd = Evd.create_goal_evar_defs sigma in + let evars,args,_ = clenv_environments_evars env evd (Some bound) ty in + let evars = Evarconv.the_conv_x_leq env t (applist (c,args)) evars in + let evars,_ = Evarconv.consider_remaining_unif_problems env evars in + let args = List.map (whd_evar (Evd.evars_of 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 { templval = mk_freelisted (match args with [] -> c | _ -> applist (c,args)); templtyp = mk_freelisted concl; - env = env; - templenv = Global.env_of_context gls.it.evar_hyps } + evd = env; + env = environ } -let mk_clenv_from gls = mk_clenv_from_n gls None +let mk_clenv_from_n gls n (c,cty) = + mk_clenv_from_env (Global.env_of_context gls.it.evar_hyps) gls.sigma n (c, cty) -let mk_clenv_rename_from gls (c,t) = - mk_clenv_from gls (c,rename_bound_var (pf_env gls) [] t) +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) @@ -156,15 +161,17 @@ let mentions clenv mv0 = let rec menrec mv1 = mv0 = mv1 || let mlist = - try (meta_fvalue clenv.env mv1).freemetas - with Anomaly _ | Not_found -> Metaset.empty in + try match meta_opt_fvalue clenv.evd mv1 with + | Some (b,_) -> b.freemetas + | None -> Metaset.empty + with Not_found -> Metaset.empty in meta_exists menrec mlist in menrec -let clenv_defined clenv mv = meta_defined clenv.env mv +let clenv_defined clenv mv = meta_defined clenv.evd mv let error_incompatible_inst clenv mv = - let na = meta_name clenv.env mv in + let na = meta_name clenv.evd mv in match na with Name id -> errorlabstrm "clenv_assign" @@ -179,17 +186,19 @@ let clenv_assign mv rhs clenv = if meta_exists (mentions clenv mv) rhs_fls.freemetas then error "clenv_assign: circularity in unification"; try - if meta_defined clenv.env mv then - if not (eq_constr (meta_fvalue clenv.env mv).rebus rhs) then + if meta_defined clenv.evd mv then + if not (eq_constr (fst (meta_fvalue clenv.evd mv)).rebus rhs) then error_incompatible_inst clenv mv else clenv - else {clenv with env = meta_assign mv rhs_fls.rebus clenv.env} + else + let st = (ConvUpToEta 0,TypeNotProcessed) in + {clenv with evd = meta_assign mv (rhs_fls.rebus,st) clenv.evd} with Not_found -> error "clenv_assign: undefined meta" -let clenv_wtactic f clenv = {clenv with env = f clenv.env } +let clenv_wtactic f clenv = {clenv with evd = f clenv.evd } (* [clenv_dependent hyps_only clenv] @@ -200,74 +209,108 @@ let clenv_wtactic f clenv = {clenv with env = f clenv.env } * type of clenv. * If [hyps_only] then metavariables occurring in the type are _excluded_ *) -(* collects all metavar occurences, in left-to-right order, preserving - * repetitions and all. *) - -let collect_metas c = - let rec collrec acc c = - match kind_of_term c with - | Meta mv -> mv::acc - | _ -> fold_constr collrec acc c - in - List.rev (collrec [] c) - (* [clenv_metavars clenv mv] * returns a list of the metavars which appear in the type of * the metavar mv. The list is unordered. *) -let clenv_metavars clenv mv = (meta_ftype clenv mv).freemetas +let clenv_metavars evd mv = + (mk_freelisted (meta_instance evd (meta_ftype evd mv))).freemetas let dependent_metas clenv mvs conclmetas = List.fold_right (fun mv deps -> - Metaset.union deps (clenv_metavars clenv.env mv)) + Metaset.union deps (clenv_metavars clenv.evd mv)) mvs conclmetas +let duplicated_metas c = + let rec collrec (one,more as acc) c = + match kind_of_term c with + | Meta mv -> if List.mem mv one then (one,mv::more) else (mv::one,more) + | _ -> fold_constr collrec acc c + in + snd (collrec ([],[]) c) + let clenv_dependent hyps_only clenv = - let mvs = collect_metas (clenv_value clenv) in + let mvs = undefined_metas clenv.evd in let ctyp_mvs = (mk_freelisted (clenv_type clenv)).freemetas in let deps = dependent_metas clenv mvs ctyp_mvs in + let nonlinear = duplicated_metas (clenv_value clenv) in + (* Make the assumption that duplicated metas have internal dependencies *) List.filter - (fun mv -> Metaset.mem mv deps && - not (hyps_only && Metaset.mem mv ctyp_mvs)) + (fun mv -> (Metaset.mem mv deps && + not (hyps_only && Metaset.mem mv ctyp_mvs)) + or List.mem mv nonlinear) mvs let clenv_missing ce = clenv_dependent true ce (******************************************************************) -let clenv_unify allow_K cv_pb t1 t2 clenv = - { clenv with env = w_unify allow_K clenv.templenv cv_pb t1 t2 clenv.env } +let clenv_unify allow_K ?(flags=default_unify_flags) cv_pb t1 t2 clenv = + { clenv with + evd = w_unify allow_K ~flags:flags clenv.env cv_pb t1 t2 clenv.evd } -let clenv_unique_resolver allow_K clause gl = - clenv_unify allow_K CUMUL (clenv_type clause) (pf_concl gl) clause +let clenv_unify_meta_types ?(flags=default_unify_flags) clenv = + { clenv with evd = w_unify_meta_types ~flags:flags clenv.env clenv.evd } +let clenv_unique_resolver allow_K ?(flags=default_unify_flags) clenv gl = + if isMeta (fst (whd_stack clenv.templtyp.rebus)) then + clenv_unify allow_K CUMUL ~flags:flags (clenv_type clenv) (pf_concl gl) + (clenv_unify_meta_types ~flags:flags clenv) + else + clenv_unify allow_K CUMUL ~flags:flags + (meta_reducible_instance clenv.evd clenv.templtyp) (pf_concl gl) clenv -(* [clenv_pose_dependent_evars clenv] +(* [clenv_pose_metas_as_evars clenv dep_mvs] * For each dependent evar in the clause-env which does not have a value, * pose a value for it by constructing a fresh evar. We do this in * left-to-right order, so that every evar's type is always closed w.r.t. - * metas. *) -let clenv_pose_dependent_evars clenv = - let dep_mvs = clenv_dependent false clenv in - List.fold_left - (fun clenv mv -> - let ty = clenv_meta_type clenv mv in - let (evd,evar) = new_evar clenv.env (cl_env clenv) ty in - clenv_assign mv evar {clenv with env=evd}) - clenv - dep_mvs + * metas. + + * Node added 14/4/08 [HH]: before this date, evars were collected in + clenv_dependent by collect_metas in the fold_constr order which is + (almost) the left-to-right order of dependencies in term. However, + due to K-redexes, collect_metas was sometimes missing some metas. + The call to collect_metas has been replaced by a call to + undefined_metas, but then the order was the one of definition of + the metas (numbers in increasing order) which is _not_ the + dependency order when a clenv_fchain occurs (because clenv_fchain + plugs a term with a list of consecutive metas in place of a - a priori - + arbitrary metavariable belonging to another sequence of consecutive metas: + e.g., clenv_fchain may plug (H ?1 ?2) at the position ?6 of + (nat_ind ?3 ?4 ?5 ?6), leading to a dependency order 3<4<5<1<2). + To ensure the dependency order, we check that the type of each meta + to pose is already meta-free, otherwise we postpone the transformation, + hoping that no cycle may happen. + + Another approach could have been to use decimal numbers for metas so that + in the example above, (H ?1 ?2) would have been renumbered (H ?6.1 ?6.2) + then making the numeric order match the dependency order. +*) -let evar_clenv_unique_resolver clenv gls = - clenv_pose_dependent_evars (clenv_unique_resolver false clenv gls) +let clenv_pose_metas_as_evars clenv dep_mvs = + let rec fold clenv = function + | [] -> clenv + | mv::mvs -> + let ty = clenv_meta_type clenv mv in + (* Postpone the evar-ization if dependent on another meta *) + (* This assumes no cycle in the dependencies - is it correct ? *) + if occur_meta ty then fold clenv (mvs@[mv]) + else + let (evd,evar) = + new_evar clenv.evd (cl_env clenv) ~src:(dummy_loc,GoalEvar) ty in + let clenv = clenv_assign mv evar {clenv with evd=evd} in + fold clenv mvs in + fold clenv dep_mvs +let evar_clenv_unique_resolver = clenv_unique_resolver (******************************************************************) let connect_clenv gls clenv = { clenv with - env = evars_reset_evd gls.sigma clenv.env; - templenv = Global.env_of_context gls.it.evar_hyps } + evd = evars_reset_evd gls.sigma clenv.evd; + env = Global.env_of_context gls.it.evar_hyps } (* [clenv_fchain mv clenv clenv'] * @@ -292,29 +335,30 @@ let connect_clenv gls clenv = In particular, it assumes that [env'] and [sigma'] extend [env] and [sigma]. *) -let clenv_fchain mv clenv nextclenv = +let clenv_fchain ?(allow_K=true) ?(flags=default_unify_flags) mv clenv nextclenv = (* Add the metavars of [nextclenv] to [clenv], with their name-environment *) let clenv' = { templval = clenv.templval; templtyp = clenv.templtyp; - env = meta_merge clenv.env nextclenv.env; - templenv = nextclenv.templenv } in + evd = + evar_merge (meta_merge clenv.evd nextclenv.evd) (evars_of clenv.evd); + env = nextclenv.env } in (* unify the type of the template of [nextclenv] with the type of [mv] *) let clenv'' = - clenv_unify true CUMUL - (clenv_nf_meta clenv' nextclenv.templtyp.rebus) + clenv_unify allow_K ~flags:flags CUMUL + (clenv_term clenv' nextclenv.templtyp) (clenv_meta_type clenv' mv) clenv' in (* assign the metavar *) let clenv''' = - clenv_assign mv (clenv_nf_meta clenv' nextclenv.templval.rebus) clenv'' + clenv_assign mv (clenv_term clenv' nextclenv.templval) clenv'' in clenv''' (***************************************************************) (* Bindings *) -type arg_bindings = (int * constr) list +type arg_bindings = open_constr explicit_bindings (* [clenv_independent clenv] * returns a list of metavariables which appear in the term cval, @@ -328,119 +372,87 @@ let clenv_independent clenv = let deps = dependent_metas clenv mvs ctyp_mvs in List.filter (fun mv -> not (Metaset.mem mv deps)) mvs -let meta_of_binder clause loc b t mvs = - match b with - | NamedHyp s -> - if List.exists (fun (_,b',_) -> b=b') t then - errorlabstrm "clenv_match_args" - (str "The variable " ++ pr_id s ++ - str " occurs more than once in binding"); - meta_with_name clause.env s - | AnonHyp n -> - if List.exists (fun (_,b',_) -> b=b') t then - errorlabstrm "clenv_match_args" - (str "The position " ++ int n ++ - str " occurs more than once in binding"); - try List.nth mvs (n-1) - with (Failure _|Invalid_argument _) -> - errorlabstrm "clenv_match_args" (str "No such binder") +let check_bindings bl = + match list_duplicates (List.map pi2 bl) with + | NamedHyp s :: _ -> + errorlabstrm "" + (str "The variable " ++ pr_id s ++ + str " occurs more than once in binding list"); + | AnonHyp n :: _ -> + errorlabstrm "" + (str "The position " ++ int n ++ + str " occurs more than once in binding list") + | [] -> () + +let meta_of_binder clause loc mvs = function + | NamedHyp s -> meta_with_name clause.evd s + | AnonHyp n -> + try List.nth mvs (n-1) + with (Failure _|Invalid_argument _) -> + errorlabstrm "" (str "No such binder") let error_already_defined b = match b with - NamedHyp id -> - errorlabstrm "clenv_match_args" + | NamedHyp id -> + errorlabstrm "" (str "Binder name \"" ++ pr_id id ++ str"\" already defined with incompatible value") | AnonHyp n -> - anomalylabstrm "clenv_match_args" + anomalylabstrm "" (str "Position " ++ int n ++ str" already defined") -let clenv_match_args s clause = - let mvs = clenv_independent clause in - let rec matchrec clause = function - | [] -> clause - | (loc,b,c)::t -> - let k = meta_of_binder clause loc b t mvs in - if meta_defined clause.env k then - if eq_constr (meta_fvalue clause.env k).rebus c then - matchrec clause t +let clenv_unify_binding_type clenv c t u = + if isMeta (fst (whd_stack u)) then + (* Not enough information to know if some subtyping is needed *) + CoerceToType, clenv, c + else + (* Enough information so as to try a coercion now *) + try + let evd,c = w_coerce_to_type (cl_env clenv) clenv.evd c t u in + TypeProcessed, { clenv with evd = evd }, c + with e when precatchable_exception e -> + TypeNotProcessed, clenv, c + +let clenv_assign_binding clenv k (sigma,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 (clenv_get_type_of clenv' c) in + let status,clenv'',c = clenv_unify_binding_type clenv' c c_typ k_typ in + { clenv'' with evd = meta_assign k (c,(UserGiven,status)) clenv''.evd } + +let clenv_match_args bl clenv = + if bl = [] then + clenv + else + let mvs = clenv_independent clenv in + check_bindings bl; + List.fold_left + (fun clenv (loc,b,(sigma,c as sc)) -> + 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 - let k_typ = clenv_hnf_constr clause (clenv_meta_type clause k) - (* nf_betaiota was before in type_of - useful to reduce - types like (x:A)([x]P u) *) - and c_typ = - clenv_hnf_constr clause - (nf_betaiota (clenv_get_type_of clause c)) in - let cl = - (* Try to infer some Meta/Evar from the type of [c] *) - try clenv_assign k c (clenv_unify true CUMUL c_typ k_typ clause) - with e when precatchable_exception e -> - (* Try to coerce to the type of [k]; cannot merge with the - previous case because Coercion does not handle Meta *) - let (_,c') = w_coerce (cl_env clause) c c_typ k_typ clause.env in - try clenv_unify true CONV (mkMeta k) c' clause - with PretypeError (env,CannotUnify (m,n)) -> - Stdpp.raise_with_loc loc - (PretypeError (env,CannotUnifyBindingType (m,n))) - in matchrec cl t - in - matchrec clause s - - -let clenv_constrain_with_bindings bl clause = - if bl = [] then - clause - else - let all_mvs = collect_metas clause.templval.rebus in - let rec matchrec clause = function - | [] -> clause - | (n,c)::t -> - let k = - (try - if n > 0 then - List.nth all_mvs (n-1) - else if n < 0 then - List.nth (List.rev all_mvs) (-n-1) - else error "clenv_constrain_with_bindings" - with Failure _ -> - errorlabstrm "clenv_constrain_with_bindings" - (str"Clause did not have " ++ int n ++ str"-th" ++ - str" absolute argument")) in - let k_typ = nf_betaiota (clenv_meta_type clause k) in - let c_typ = nf_betaiota (clenv_get_type_of clause c) in - matchrec - (clenv_assign k c (clenv_unify true CUMUL c_typ k_typ clause)) t - in - matchrec clause bl - - -(* not exported: maybe useful ? *) -let clenv_constrain_dep_args hyps_only clause = function - | [] -> clause - | mlist -> - let occlist = clenv_dependent hyps_only clause in - if List.length occlist = List.length mlist then - List.fold_left2 - (fun clenv k c -> - try - let k_typ = - clenv_hnf_constr clause (clenv_meta_type clause k) in - let c_typ = - clenv_hnf_constr clause (clenv_get_type_of clause c) in - (* faire quelque chose avec le sigma retourne ? *) - let (_,c') = w_coerce (cl_env clenv) c c_typ k_typ clenv.env in - clenv_unify true CONV (mkMeta k) c' clenv - with _ -> - clenv_unify true CONV (mkMeta k) c clenv) - clause occlist mlist - else - error ("Not the right number of missing arguments (expected " - ^(string_of_int (List.length occlist))^")") - -let clenv_constrain_missing_args mlist clause = - clenv_constrain_dep_args true clause mlist - + clenv_assign_binding clenv k sc) + clenv bl + +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 _ -> error "clenv_constrain_with_bindings" in + clenv_assign_binding clenv k (Evd.empty,c) + +let clenv_constrain_dep_args hyps_only bl clenv = + if bl = [] then + clenv + else + let occlist = clenv_dependent hyps_only clenv in + if List.length occlist = List.length bl then + List.fold_left2 clenv_assign_binding clenv occlist bl + else + error ("Not the right number of missing arguments (expected " + ^(string_of_int (List.length occlist))^")") (****************************************************************) (* Clausal environment for an application *) @@ -448,7 +460,7 @@ let clenv_constrain_missing_args mlist clause = let make_clenv_binding_gen hyps_only n gls (c,t) = function | ImplicitBindings largs -> let clause = mk_clenv_from_n gls n (c,t) in - clenv_constrain_dep_args hyps_only clause largs + clenv_constrain_dep_args hyps_only largs clause | ExplicitBindings lbind -> let clause = mk_clenv_rename_from_n gls n (c,t) in clenv_match_args lbind clause @@ -465,4 +477,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.env) + pr_evar_defs clenv.evd) diff --git a/pretyping/clenv.mli b/pretyping/clenv.mli index b5433cac..9b2d6e29 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 9277 2006-10-25 13:02:22Z herbelin $ i*) +(*i $Id: clenv.mli 10856 2008-04-27 16:15:34Z herbelin $ i*) (*i*) open Util @@ -18,22 +18,21 @@ open Evd open Evarutil open Mod_subst open Rawterm +open Unification (*i*) (***************************************************************) (* The Type of Constructions clausale environments. *) -(* [templenv] is the typing context - * [env] is the mapping from metavar and evar numbers to their types +(* [env] is the typing context + * [evd] is the mapping from metavar and evar numbers to their types * and values. * [templval] is the template which we are trying to fill out. * [templtyp] is its type. - * [namenv] is a mapping from metavar numbers to names, for - * use in instantiating metavars by name. *) type clausenv = { - templenv : env; - env : evar_defs; + env : env; + evd : evar_defs; templval : constr freelisted; templtyp : constr freelisted } @@ -54,48 +53,53 @@ val clenv_meta_type : clausenv -> metavariable -> types val mk_clenv_from : evar_info sigma -> constr * types -> clausenv val mk_clenv_from_n : evar_info sigma -> int option -> constr * types -> clausenv -val mk_clenv_rename_from : evar_info sigma -> constr * types -> clausenv -val mk_clenv_rename_from_n : - evar_info sigma -> int option -> constr * types -> clausenv val mk_clenv_type_of : evar_info sigma -> constr -> clausenv +val mk_clenv_from_env : env -> evar_map -> int option -> constr * types -> clausenv (***************************************************************) (* linking of clenvs *) val connect_clenv : evar_info sigma -> clausenv -> clausenv -val clenv_fchain : metavariable -> clausenv -> clausenv -> clausenv +val clenv_fchain : + ?allow_K:bool -> ?flags:unify_flags -> metavariable -> clausenv -> clausenv -> clausenv (***************************************************************) (* Unification with clenvs *) (* Unifies two terms in a clenv. The boolean is [allow_K] (see [Unification]) *) val clenv_unify : - bool -> conv_pb -> constr -> constr -> clausenv -> clausenv + bool -> ?flags:unify_flags -> conv_pb -> constr -> constr -> clausenv -> clausenv (* unifies the concl of the goal with the type of the clenv *) val clenv_unique_resolver : - bool -> clausenv -> evar_info sigma -> clausenv + bool -> ?flags:unify_flags -> clausenv -> evar_info sigma -> clausenv (* same as above ([allow_K=false]) but replaces remaining metas - with fresh evars *) + with fresh evars if [evars_flag] is [true] *) val evar_clenv_unique_resolver : - clausenv -> evar_info sigma -> clausenv + bool -> ?flags:unify_flags -> clausenv -> evar_info sigma -> clausenv + +val clenv_dependent : bool -> clausenv -> metavariable list + +val clenv_pose_metas_as_evars : clausenv -> metavariable list -> clausenv (***************************************************************) (* Bindings *) +type arg_bindings = open_constr explicit_bindings + (* bindings where the key is the position in the template of the clenv (dependent or not). Positions can be negative meaning to start from the rightmost argument of the template. *) -type arg_bindings = (int * constr) list - val clenv_independent : clausenv -> metavariable list val clenv_missing : clausenv -> metavariable list +val clenv_constrain_last_binding : constr -> clausenv -> clausenv + (* defines metas corresponding to the name of the bindings *) -val clenv_match_args : - constr explicit_bindings -> clausenv -> clausenv -val clenv_constrain_with_bindings : arg_bindings -> clausenv -> clausenv +val clenv_match_args : arg_bindings -> clausenv -> clausenv + +val clenv_unify_meta_types : ?flags:unify_flags -> clausenv -> clausenv (* start with a clenv to refine with a given term with bindings *) @@ -103,10 +107,10 @@ val clenv_constrain_with_bindings : arg_bindings -> 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 -> constr bindings -> + evar_info sigma -> int option -> constr * constr -> open_constr bindings -> clausenv val make_clenv_binding : - evar_info sigma -> constr * constr -> constr bindings -> clausenv + evar_info sigma -> constr * constr -> open_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 @@ -125,6 +129,10 @@ val clenv_environments : val clenv_environments_evars : env -> evar_defs -> int option -> types -> evar_defs * constr list * types +(* [clenv_conv_leq env sigma t c n] looks for c1...cn s.t. [t <= c c1...cn] *) +val clenv_conv_leq : + env -> evar_map -> types -> constr -> int -> constr list + (* if the clause is a product, add an extra meta for this product *) exception NotExtensibleClause val clenv_push_prod : clausenv -> clausenv diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index cc74b0ad..3c37a49f 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 9257 2006-10-21 17:28:28Z herbelin $ *) +(* $Id: coercion.ml 10883 2008-05-05 13:55:24Z herbelin $ *) open Util open Names @@ -25,38 +25,40 @@ open Termops module type S = sig (*s Coercions. *) - (* [inh_app_fun env isevars j] coerces [j] to a function; i.e. it + (* [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 - (* [inh_coerce_to_sort env isevars j] coerces [j] to a type; i.e. it + (* [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 - (* [inh_coerce_to_base env isevars j] coerces [j] to its base type; i.e. it + (* [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_conv_coerce_to loc env isevars 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 - - (* [inh_conv_coerces_to loc env isevars t t'] checks if an object of type [t] + 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 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 - (* [inh_pattern_coerce_to loc env isevars pat ind1 ind2] coerces the Cases + (* [inh_pattern_coerce_to loc env evd pat ind1 ind2] coerces the Cases pattern [pat] typed in [ind1] into a pattern typed in [ind2]; raises [Not_found] if no coercion found *) val inh_pattern_coerce_to : @@ -72,8 +74,8 @@ module Default = struct | App (f,l) -> mkApp (whd_evar sigma f,l) | _ -> whd_evar sigma t - let class_of1 env isevars t = - let sigma = evars_of isevars in + let class_of1 env evd t = + let sigma = evars_of evd in class_of env sigma (whd_app_evar sigma t) (* Here, funj is a coercion therefore already typed in global context *) @@ -122,47 +124,51 @@ module Default = struct (hj,typ_cl) p) with _ -> anomaly "apply_coercion" - let inh_app_fun env isevars j = - let t = whd_betadeltaiota env (evars_of isevars) j.uj_type in + let inh_app_fun env evd j = + let t = whd_betadeltaiota env (evars_of evd) j.uj_type in match kind_of_term t with - | Prod (_,_,_) -> (isevars,j) + | Prod (_,_,_) -> (evd,j) | Evar ev -> - let (isevars',t) = define_evar_as_arrow isevars ev in - (isevars',{ uj_val = j.uj_val; uj_type = t }) + let (evd',t) = define_evar_as_product evd ev in + (evd',{ uj_val = j.uj_val; uj_type = t }) | _ -> (try - let t,i1 = class_of1 env isevars j.uj_type in + let t,i1 = class_of1 env evd j.uj_type in let p = lookup_path_to_fun_from i1 in - (isevars,apply_coercion env p j t) - with Not_found -> (isevars,j)) + (evd,apply_coercion env p j t) + with Not_found -> (evd,j)) - let inh_tosort_force loc env isevars j = + let inh_tosort_force loc env evd j = try - let t,i1 = class_of1 env isevars j.uj_type in + let t,i1 = class_of1 env evd j.uj_type in let p = lookup_path_to_sort_from i1 in let j1 = apply_coercion env p j t in - let j2 = on_judgment_type (whd_evar (evars_of isevars)) j1 in - (isevars,type_judgment env j2) + let j2 = on_judgment_type (whd_evar (evars_of evd)) j1 in + (evd,type_judgment env j2) with Not_found -> - error_not_a_type_loc loc env (evars_of isevars) j + error_not_a_type_loc loc env (evars_of evd) j - let inh_coerce_to_sort loc env isevars j = - let typ = whd_betadeltaiota env (evars_of isevars) j.uj_type in + let inh_coerce_to_sort loc env evd j = + let typ = whd_betadeltaiota env (evars_of evd) j.uj_type in match kind_of_term typ with - | Sort s -> (isevars,{ utj_val = j.uj_val; utj_type = s }) - | Evar ev when not (is_defined_evar isevars ev) -> - let (isevars',s) = define_evar_as_sort isevars ev in - (isevars',{ utj_val = j.uj_val; utj_type = s }) + | Sort s -> (evd,{ utj_val = j.uj_val; utj_type = s }) + | Evar ev when not (is_defined_evar evd ev) -> + let (evd',s) = define_evar_as_sort evd ev in + (evd',{ utj_val = j.uj_val; utj_type = s }) | _ -> - inh_tosort_force loc env isevars j + inh_tosort_force loc env evd j - let inh_coerce_to_base loc env isevars j = (isevars, j) + let inh_coerce_to_base loc env evd j = (evd, j) - let inh_coerce_to_fail env isevars c1 v t = + let inh_coerce_to_fail env evd rigidonly v t c1 = + if rigidonly & not (Heads.is_rigid env c1 && Heads.is_rigid env t) + then + raise NoCoercion + else let v', t' = try - let t1,i1 = class_of1 env isevars c1 in - let t2,i2 = class_of1 env isevars t in + let t1,i1 = class_of1 env evd c1 in + let t2,i2 = class_of1 env evd t in let p = lookup_path_between (i2,i1) in match v with Some v -> @@ -171,86 +177,57 @@ module Default = struct | None -> None, t with Not_found -> raise NoCoercion in - try (the_conv_x_leq env t' c1 isevars, v', t') + try (the_conv_x_leq env t' c1 evd, v') with Reduction.NotConvertible -> raise NoCoercion - let rec inh_conv_coerce_to_fail loc env isevars v t c1 = - try (the_conv_x_leq env t c1 isevars, v, t) + let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = + try (the_conv_x_leq env t c1 evd, v) with Reduction.NotConvertible -> - try inh_coerce_to_fail env isevars c1 v t + try inh_coerce_to_fail env evd rigidonly v t c1 with NoCoercion -> match - kind_of_term (whd_betadeltaiota env (evars_of isevars) t), - kind_of_term (whd_betadeltaiota env (evars_of isevars) c1) + kind_of_term (whd_betadeltaiota env (evars_of evd) t), + kind_of_term (whd_betadeltaiota env (evars_of evd) c1) with - | Prod (_,t1,t2), Prod (name,u1,u2) -> - let v' = option_map (whd_betadeltaiota env (evars_of isevars)) v in - let (evd',b) = - match v' with - | Some v' -> - (match kind_of_term v' with - | Lambda (x,v1,v2) -> - (* sous-typage non contravariant: pas de "leq v1 u1" *) - (try the_conv_x env v1 u1 isevars, Some (x, v1, v2) - with Reduction.NotConvertible -> (isevars, None)) - | _ -> (isevars, None)) - | None -> (isevars, None) - in - (match b with - | Some (x, v1, v2) -> - let env1 = push_rel (x,None,v1) env in - let (evd'', v2', t2') = inh_conv_coerce_to_fail loc env1 evd' - (Some v2) t2 u2 in - (evd'', option_map (fun v2' -> mkLambda (x, v1, v2')) v2', - mkProd (x, v1, t2')) - | None -> - (* Mismatch on t1 and u1 or not a lambda: we eta-expand *) - (* we look for a coercion c:u1->t1 s.t. [name:u1](v' (c x)) *) - (* has type (name:u1)u2 (with v' recursively obtained) *) - let name = match name with - | Anonymous -> Name (id_of_string "x") - | _ -> name - in - let env1 = push_rel (name,None,u1) env in - let (evd', v1', t1') = - inh_conv_coerce_to_fail loc env1 isevars - (Some (mkRel 1)) (lift 1 u1) (lift 1 t1) - in - let (evd'', v2', t2') = - let v2 = - match v with - | Some v -> option_map (fun x -> mkApp(lift 1 v,[|x|])) v1' - | None -> None - and evd', t2 = - match v1' with - | Some v1' -> evd', subst_term v1' t2 - | None -> - let evd', ev = - new_evar evd' env ~src:(loc, InternalHole) t1' in - evd', subst_term ev t2 - in - inh_conv_coerce_to_fail loc env1 evd' v2 t2 u2 - in - (evd'', option_map (fun v2' -> mkLambda (name, u1, v2')) v2', - mkProd (name, u1, t2'))) + | 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 + | Anonymous -> Name (id_of_string "x") + | _ -> name in + let env1 = push_rel (name,None,u1) env in + let (evd', v1) = + inh_conv_coerce_to_fail loc env1 evd rigidonly + (Some (mkRel 1)) (lift 1 u1) (lift 1 t1) in + let v1 = Option.get v1 in + let v2 = Option.map (fun v -> beta_applist (lift 1 v,[v1])) v in + let t2 = subst_term v1 t2 in + let (evd'',v2') = inh_conv_coerce_to_fail loc env1 evd' rigidonly v2 t2 u2 in + (evd'', Option.map (fun v2' -> mkLambda (name, u1, v2')) v2') | _ -> raise NoCoercion (* Look for cj' obtained from cj by inserting coercions, s.t. cj'.typ = t *) - let inh_conv_coerce_to loc env isevars cj (n, t) = + let inh_conv_coerce_to_gen rigidonly loc env evd cj (n, t) = match n with None -> - let (evd', val', type') = + let (evd', val') = try - inh_conv_coerce_to_fail loc env isevars (Some cj.uj_val) cj.uj_type t + inh_conv_coerce_to_fail loc env evd rigidonly (Some cj.uj_val) cj.uj_type t with NoCoercion -> - let sigma = evars_of isevars in + let sigma = evars_of evd in error_actual_type_loc loc env sigma cj t in let val' = match val' with Some v -> v | None -> assert(false) in (evd',{ uj_val = val'; uj_type = t }) - | Some (init, cur) -> (isevars, cj) + | Some (init, cur) -> (evd, cj) + + 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 (isevars : evar_defs) t (abs, t') = isevars + let inh_conv_coerces_to loc env (evd : evar_defs) t (abs, t') = evd (* Still problematic, as it changes unification let nabsinit, nabs = match abs with @@ -262,7 +239,7 @@ module Default = struct (* a little more effort to get products is needed *) try decompose_prod_n nabs t with _ -> - if !Options.debug then + if !Flags.debug then msg_warning (str "decompose_prod_n failed"); raise (Invalid_argument "Coercion.inh_conv_coerces_to") in @@ -274,11 +251,11 @@ module Default = struct env', rng, lift nabs t' in try - pi1 (inh_conv_coerce_to_fail loc env' isevars None t t') + pi1 (inh_conv_coerce_to_fail loc env' evd None t t') with NoCoercion -> - isevars) (* Maybe not enough information to unify *) - (*let sigma = evars_of isevars in + evd) (* Maybe not enough information to unify *) + (*let sigma = evars_of evd in error_cannot_coerce env' sigma (t, t'))*) - else isevars - with Invalid_argument _ -> isevars *) + else evd + with Invalid_argument _ -> evd *) end diff --git a/pretyping/coercion.mli b/pretyping/coercion.mli index 42ce27fd..b1c8e893 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 8875 2006-05-29 19:59:11Z msozeau $ i*) +(*i $Id: coercion.mli 10840 2008-04-23 21:29:34Z herbelin $ i*) (*i*) open Util @@ -45,6 +45,9 @@ module type S = sig [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_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; diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index 29ec7904..81bb41ef 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: detyping.ml 10135 2007-09-21 14:28:12Z herbelin $ *) +(* $Id: detyping.ml 11073 2008-06-08 20:24:51Z herbelin $ *) open Pp open Util @@ -31,8 +31,8 @@ let dl = dummy_loc (****************************************************************************) (* Tools for printing of Cases *) -let encode_inductive qid = - let indsp = global_inductive qid in +let encode_inductive r = + let indsp = inductive_of_reference r in let constr_lengths = mis_constr_nargs indsp in (indsp,constr_lengths) @@ -108,14 +108,7 @@ module PrintingCasesLet = module PrintingIf = Goptions.MakeRefTable(PrintingCasesIf) module PrintingLet = Goptions.MakeRefTable(PrintingCasesLet) -let force_let ci = - let indsp = ci.ci_ind in - let lc = mis_constr_nargs indsp in PrintingLet.active (indsp,lc) -let force_if ci = - let indsp = ci.ci_ind in - let lc = mis_constr_nargs indsp in PrintingIf.active (indsp,lc) - -(* Options for printing or not wildcard and synthetisable types *) +(* Flags.for printing or not wildcard and synthetisable types *) open Goptions @@ -174,17 +167,18 @@ let computable p k = 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 | Prod (name,_,c') -> - (match concrete_name true avoid env_names name c' with + (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')) | LetIn (name,_,_,c') -> - (match concrete_name true avoid env_names name c' with + (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' @@ -196,15 +190,28 @@ let lookup_name_as_renamed env t s = 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 true [] empty_names_context name c' with + (match concrete_name (Some true) [] empty_names_context name c' with (Name _,_) -> lookup n (d+1) c' - | (Anonymous,_) -> if n=1 then Some d else lookup (n-1) (d+1) c') + | (Anonymous,_) -> + if n=0 then + Some (d-1) + else if n=1 then + Some d + else + lookup (n-1) (d+1) c') | LetIn (name,_,_,c') -> - (match concrete_name true [] empty_names_context name c' with + (match concrete_name (Some true) [] empty_names_context name c' with | (Name _,_) -> lookup n (d+1) c' - | (Anonymous,_) -> if n=1 then Some d else lookup (n-1) (d+1) c') + | (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 - | _ -> None + | _ -> if n=0 then Some (d-1) else None in lookup n 1 t (**********************************************************************) @@ -227,7 +234,7 @@ let rec decomp_branch n nal b (avoid,env as e) c = | _ -> Name (id_of_string "x"),(applist (lift 1 c, [mkRel 1])), concrete_name in - let na',avoid' = f b avoid env na c in + let na',avoid' = f (Some b) avoid env 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 = @@ -266,15 +273,15 @@ and contract_branch isgoal e (cn,mkpat,b) = let is_nondep_branch c n = try - let _,ccl = decompose_lam_n_assum n c in - noccur_between 1 n ccl + let sign,ccl = decompose_lam_n_assum n c in + noccur_between 1 (rel_context_length sign) ccl with _ -> (* Not eta-expanded or not reduced *) false let extract_nondep_branches test c b n = let rec strip n r = if n=0 then r else match r with - | RLambda (_,_,_,t) -> strip (n-1) t + | RLambda (_,_,_,_,t) -> strip (n-1) t | RLetIn (_,_,_,t) -> strip (n-1) t | _ -> assert false in if test c n then Some (strip n b) else None @@ -282,12 +289,14 @@ let extract_nondep_branches test c b n = let it_destRLambda_or_LetIn_names n c = let rec aux n nal c = if n=0 then (List.rev nal,c) else match c with - | RLambda (_,na,_,c) -> aux (n-1) (na::nal) c + | RLambda (_,na,_,_,c) -> aux (n-1) (na::nal) c | RLetIn (_,na,_,c) -> aux (n-1) (na::nal) c | _ -> (* eta-expansion *) let rec next l = let x = Nameops.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 @@ -303,16 +312,16 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl = let synth_type = synthetize_type () in let tomatch = detype c in let alias, aliastyp, pred= - if (not !Options.raw_print) & synth_type & computable & Array.length bl<>0 + if (not !Flags.raw_print) & synth_type & computable & Array.length bl<>0 then Anonymous, None, None else - match option_map detype p with + 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 - | RLambda (_,x,t,c) -> x, c + | RLambda (_,x,_,t,c) -> x, c | _ -> Anonymous, typ in let aliastyp = if List.for_all ((=) Anonymous) nl then None @@ -323,8 +332,10 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl = let eqnl = detype_eqns constructs consnargsl bl in let tag = try - if !Options.raw_print then + if !Flags.raw_print then RegularStyle + else if st = LetPatternStyle then + st else if PrintingLet.active (indsp,consnargsl) then LetStyle else if PrintingIf.active (indsp,consnargsl) then @@ -344,11 +355,11 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl = array_map3 (extract_nondep_branches testdep) bl bl' consnargsl in if array_for_all ((<>) None) nondepbrs then RIf (dl,tomatch,(alias,pred), - out_some nondepbrs.(0),out_some nondepbrs.(1)) + Option.get nondepbrs.(0),Option.get nondepbrs.(1)) else - RCases (dl,pred,[tomatch,(alias,aliastyp)],eqnl) + RCases (dl,tag,pred,[tomatch,(alias,aliastyp)],eqnl) | _ -> - RCases (dl,pred,[tomatch,(alias,aliastyp)],eqnl) + RCases (dl,tag,pred,[tomatch,(alias,aliastyp)],eqnl) let detype_sort = function | Prop c -> RProp c @@ -448,14 +459,14 @@ and share_names isgoal n l avoid env c t = let t = detype isgoal avoid env t in let id = next_name_away na avoid in let avoid = id::avoid and env = add_name (Name id) env in - share_names isgoal (n-1) ((Name id,None,t)::l) avoid env c c' + 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 avoid = id::avoid and env = add_name (Name id) env in - share_names isgoal n ((Name id,Some b,t')::l) avoid env c t + 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 *) | _, LetIn (_,b,_,t) when n > 0 -> share_names isgoal n l avoid env c (subst1 b t) @@ -465,7 +476,7 @@ and share_names isgoal n l avoid env c t = 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,None,t')::l) avoid env appc c' + share_names isgoal (n-1) ((Name id,Explicit,None,t')::l) avoid env appc c' (* If built with the f/n notation: we renounce to share names *) | _ -> if n>0 then warning "Detyping.detype: cannot factorize fix enough"; @@ -475,7 +486,7 @@ and share_names isgoal n l avoid env c t = and detype_eqns isgoal avoid env ci computable constructs consnargsl bl = try - if !Options.raw_print or not (reverse_matching ()) then raise Exit; + if !Flags.raw_print or not (reverse_matching ()) then raise Exit; let mat = build_tree Anonymous isgoal (avoid,env) ci bl in List.map (fun (pat,((avoid,env),c)) -> (dl,[],[pat],detype isgoal avoid env c)) mat @@ -488,7 +499,7 @@ and detype_eqn isgoal avoid env constr construct_nargs branch = if force_wildcard () & noccurn 1 b then PatVar (dl,Anonymous),avoid,(add_name Anonymous env),ids else - let id = next_name_away_with_default "x" x avoid in + 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 = @@ -523,15 +534,31 @@ and detype_eqn isgoal avoid env constr construct_nargs branch = and detype_binder isgoal bk avoid env na ty c = let na',avoid' = if bk = BLetIn then - concrete_let_name isgoal avoid env na c + concrete_let_name (avoid_flag isgoal) avoid env na c else - concrete_name isgoal avoid env na c in + concrete_name (avoid_flag isgoal) avoid env na c in let r = detype isgoal avoid' (add_name na' env) c in match bk with - | BProd -> RProd (dl, na',detype isgoal avoid env ty, r) - | BLambda -> RLambda (dl, na',detype isgoal avoid env ty, r) + | 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) +let rec detype_rel_context where avoid env sign = + let where = Option.map (fun c -> it_mkLambda_or_LetIn c sign) where in + let rec aux avoid env = function + | [] -> [] + | (na,b,t)::rest -> + let na',avoid' = + 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 + 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 + in aux avoid env (List.rev sign) + (**********************************************************************) (* Module substitution: relies on detyping *) @@ -561,27 +588,27 @@ let rec subst_rawconstr subst raw = if r' == r && rl' == rl then raw else RApp(loc,r',rl') - | RLambda (loc,n,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,r1',r2') + RLambda (loc,n,bk,r1',r2') - | RProd (loc,n,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,r1',r2') + RProd (loc,n,bk,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,rtno,rl,branches) -> - let rtno' = option_smartmap (subst_rawconstr subst) rtno + | 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 topt' = option_smartmap + let topt' = Option.smartmap (fun (loc,(sp,i),x,y as t) -> let sp' = subst_kn subst sp in if sp == sp' then t else (loc,(sp',i),x,y)) topt in @@ -596,17 +623,17 @@ let rec subst_rawconstr subst raw = branches in if rtno' == rtno && rl' == rl && branches' == branches then raw else - RCases (loc,rtno',rl',branches') + RCases (loc,sty,rtno',rl',branches') | RLetTuple (loc,nal,(na,po),b,c) -> - let po' = option_smartmap (subst_rawconstr subst) po + let po' = Option.smartmap (subst_rawconstr subst) po 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 + let po' = Option.smartmap (subst_rawconstr subst) po and b1' = subst_rawconstr subst b1 and b2' = subst_rawconstr subst b2 and c' = subst_rawconstr subst c in @@ -617,10 +644,10 @@ let rec subst_rawconstr subst raw = let ra1' = array_smartmap (subst_rawconstr subst) ra1 and ra2' = array_smartmap (subst_rawconstr subst) ra2 in let bl' = array_smartmap - (list_smartmap (fun (na,obd,ty as dcl) -> + (list_smartmap (fun (na,k,obd,ty as dcl) -> let ty' = subst_rawconstr subst ty in - let obd' = option_smartmap (subst_rawconstr subst) obd in - if ty'==ty & obd'==obd then dcl else (na,obd',ty'))) + let obd' = Option.smartmap (subst_rawconstr subst) obd in + if ty'==ty & obd'==obd then dcl else (na,k,obd',ty'))) bl in if ra1' == ra1 && ra2' == ra2 && bl'==bl then raw else RRec (loc,fix,ida,bl',ra1',ra2') @@ -631,8 +658,8 @@ let rec subst_rawconstr subst raw = let ref',_ = subst_global subst ref in if ref' == ref then raw else RHole (loc,InternalHole) - | RHole (loc, (BinderType _ | QuestionMark _ | CasesType | - InternalHole | TomatchTypeParameter _)) -> raw + | RHole (loc, (BinderType _ | QuestionMark _ | CasesType | InternalHole | + TomatchTypeParameter _ | GoalEvar | ImpossibleCase)) -> raw | RCast (loc,r1,k) -> (match k with diff --git a/pretyping/detyping.mli b/pretyping/detyping.mli index 7ac7162f..23090858 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 9976 2007-07-12 11:58:30Z msozeau $ i*) +(*i $Id: detyping.mli 10410 2007-12-31 13:11:55Z msozeau $ i*) (*i*) open Util @@ -40,6 +40,9 @@ val detype_case : val detype_sort : sorts -> rawsort +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_index_as_renamed : env -> constr -> int -> int option @@ -47,8 +50,6 @@ val lookup_index_as_renamed : env -> constr -> int -> int option val set_detype_anonymous : (loc -> int -> rawconstr) -> unit val force_wildcard : unit -> bool val synthetize_type : unit -> bool -val force_if : case_info -> bool -val force_let : case_info -> bool (* Utilities to transform kernel cases to simple pattern-matching problem *) @@ -57,4 +58,3 @@ 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 2764633b..58951302 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: evarconv.ml 9665 2007-02-21 17:08:10Z herbelin $ *) +(* $Id: evarconv.ml 11157 2008-06-21 10:45:51Z herbelin $ *) open Pp open Util @@ -43,55 +43,28 @@ let eval_flexible_term env c = match kind_of_term c with | Const c -> constant_opt_value env c | Rel n -> - (try let (_,v,_) = lookup_rel n env in option_map (lift n) v + (try let (_,v,_) = lookup_rel n env in Option.map (lift n) v with Not_found -> None) | Var id -> (try let (_,v,_) = lookup_named id env in v with Not_found -> None) | LetIn (_,b,_,c) -> Some (subst1 b c) | Lambda _ -> Some c | _ -> assert false -(* -let rec apprec_nobeta env sigma s = - let (t,stack as s) = whd_state s in - match kind_of_term (fst s) 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 - apprec_nobeta env sigma (rslt, stack) - else - s - | Fix fix -> - (match reduce_fix (whd_betadeltaiota_state env sigma) fix stack with - | Reduced s -> apprec_nobeta env sigma s - | NotReducible -> s) - | _ -> s - -let evar_apprec_nobeta env isevars stack c = - let rec aux s = - let (t,stack as s') = apprec_nobeta env (evars_of isevars) s in - match kind_of_term t with - | Evar (n,_ as ev) when Evd.is_defined (evars_of isevars) n -> - aux (Evd.existential_value (evars_of isevars) ev, stack) - | _ -> (t, list_of_stack stack) - in aux (c, append_stack (Array.of_list stack) empty_stack) -*) -let evar_apprec env isevars stack c = - let sigma = evars_of isevars in +let evar_apprec env evd stack c = + let sigma = evars_of evd in let rec aux s = - let (t,stack) = Reductionops.apprec env sigma s in + let (t,stack) = whd_betaiota_deltazeta_for_iota_state env sigma s in match kind_of_term t with - | Evar (n,_ as ev) when Evd.is_defined sigma n -> + | Evar (evk,_ as ev) when Evd.is_defined sigma evk -> aux (Evd.existential_value sigma ev, stack) | _ -> (t, list_of_stack stack) in aux (c, append_stack_list stack empty_stack) -let apprec_nohdbeta env isevars c = - let (t,stack as s) = Reductionops.whd_stack c in - match kind_of_term t with - | (Case _ | Fix _) -> evar_apprec env isevars [] c - | _ -> s +let apprec_nohdbeta env evd c = + match kind_of_term (fst (Reductionops.whd_stack c)) with + | (Case _ | Fix _) -> applist (evar_apprec env evd [] c) + | _ -> c (* [check_conv_record (t1,l1) (t2,l2)] tries to decompose the problem (t1 l1) = (t2 l2) into a problem @@ -116,84 +89,96 @@ let apprec_nohdbeta env isevars c = let check_conv_record (t1,l1) (t2,l2) = try let proji = global_of_constr t1 in - let cstr = global_of_constr t2 in - let { o_DEF = c; o_TABS = bs; o_TPARAMS = params; o_TCOMPS = us } = - lookup_canonical_conversion (proji, cstr) in + let canon_s,l2_effective = + try + match kind_of_term t2 with + Prod (_,a,b) -> (* assert (l2=[]); *) + if dependent (mkRel 1) b then raise Not_found + else lookup_canonical_conversion (proji, Prod_cs),[a;pop b] + | Sort s -> + lookup_canonical_conversion + (proji, Sort_cs (family_of_sort s)),[] + | _ -> + let c2 = try global_of_constr t2 with _ -> raise Not_found in + lookup_canonical_conversion (proji, Const_cs c2),l2 + with Not_found -> + lookup_canonical_conversion (proji,Default_cs),[] + in + let { o_DEF = c; o_INJ=n; o_TABS = bs; o_TPARAMS = params; o_TCOMPS = us } + = canon_s in let params1, c1, extra_args1 = match list_chop (List.length params) l1 with | params1, c1::extra_args1 -> params1, c1, extra_args1 | _ -> assert false in - let us2,extra_args2 = list_chop (List.length us) l2 in - c,bs,(params,params1),(us,us2),(extra_args1,extra_args2),c1 + 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 _ -> raise Not_found (* Precondition: one of the terms of the pb is an uninstantiated evar, * possibly applied to arguments. *) -let rec ise_try isevars = function +let rec ise_try evd = function [] -> assert false - | [f] -> f isevars + | [f] -> f evd | f1::l -> - let (isevars',b) = f1 isevars in - if b then (isevars',b) else ise_try isevars l + let (evd',b) = f1 evd in + if b then (evd',b) else ise_try evd l -let ise_and isevars l = +let ise_and evd l = let rec ise_and i = function [] -> assert false | [f] -> f i | f1::l -> let (i',b) = f1 i in - if b then ise_and i' l else (isevars,false) in - ise_and isevars l + if b then ise_and i' l else (evd,false) in + ise_and evd l -let ise_list2 isevars f l1 l2 = +let ise_list2 evd f l1 l2 = let rec ise_list2 i l1 l2 = match l1,l2 with [], [] -> (i, true) | [x], [y] -> f i x y | x::l1, y::l2 -> let (i',b) = f i x y in - if b then ise_list2 i' l1 l2 else (isevars,false) - | _ -> (isevars, false) in - ise_list2 isevars l1 l2 + if b then ise_list2 i' l1 l2 else (evd,false) + | _ -> (evd, false) in + ise_list2 evd l1 l2 -let ise_array2 isevars f v1 v2 = +let ise_array2 evd f v1 v2 = let rec allrec i = function | -1 -> (i,true) | n -> let (i',b) = f i v1.(n) v2.(n) in - if b then allrec i' (n-1) else (isevars,false) + if b then allrec i' (n-1) else (evd,false) in let lv1 = Array.length v1 in - if lv1 = Array.length v2 then allrec isevars (pred lv1) - else (isevars,false) + if lv1 = Array.length v2 then allrec evd (pred lv1) + else (evd,false) -let rec evar_conv_x env isevars pbty term1 term2 = - let sigma = evars_of isevars in +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 -(* - if eq_constr term1 term2 then - true - else -*) (* 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... *) - if is_ground_term isevars term1 && is_ground_term isevars term2 & - is_fconv pbty env (evars_of isevars) term1 term2 then - (isevars,true) - else if is_undefined_evar isevars term1 then - solve_simple_eqn evar_conv_x env isevars (pbty,destEvar term1,term2) - else if is_undefined_evar isevars term2 then - solve_simple_eqn evar_conv_x env isevars (pbty,destEvar term2,term1) - else - let (t1,l1) = apprec_nohdbeta env isevars term1 in - let (t2,l2) = apprec_nohdbeta env isevars term2 in - evar_eqappr_x env isevars pbty (t1,l1) (t2,l2) - -and evar_eqappr_x env isevars pbty (term1,l1 as appr1) (term2,l2 as appr2) = + if is_ground_term evd term1 && is_ground_term evd term2 & + is_fconv pbty env (evars_of evd) term1 term2 + then + (evd,true) + else + let term1 = apprec_nohdbeta env evd term1 in + let term2 = apprec_nohdbeta env evd term2 in + if is_undefined_evar evd term1 then + solve_simple_eqn evar_conv_x env evd (pbty,destEvar term1,term2) + else if is_undefined_evar evd term2 then + solve_simple_eqn evar_conv_x env evd (pbty,destEvar term2,term1) + else + evar_eqappr_x env evd pbty (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 *) match (flex_kind_of_term term1 l1, flex_kind_of_term term2 l2) with | Flexible (sp1,al1 as ev1), Flexible (sp2,al2 as ev2) -> @@ -215,25 +200,25 @@ and evar_eqappr_x env isevars pbty (term1,l1 as appr1) (term2,l2 as appr2) = and f2 i = if sp1 = sp2 then ise_and i - [(fun i -> ise_array2 i - (fun i -> evar_conv_x env i CONV) al1 al2); - (fun i -> ise_list2 i - (fun i -> evar_conv_x env i CONV) l1 l2)] + [(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 - ise_try isevars [f1; f2] + ise_try evd [f1; f2] | Flexible ev1, MaybeFlexible flex2 -> let f1 i = if - is_unification_pattern_evar ev1 l1 & - not (occur_evar (fst ev1) (applist (term2,l2))) + is_unification_pattern_evar env ev1 l1 & + 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 isevars) (applist(term2,l2)) in + let t2 = nf_evar (evars_of evd) (applist appr2) in let t2 = solve_pattern_eqn env l1 t2 in - solve_simple_eqn evar_conv_x env isevars (pbty,ev1,t2) + solve_simple_eqn evar_conv_x env evd (pbty,ev1,t2) else if List.length l1 <= List.length l2 then @@ -253,19 +238,19 @@ and evar_eqappr_x env isevars pbty (term1,l1 as appr1) (term2,l2 as appr2) = evar_eqappr_x env i pbty appr1 (evar_apprec env i l2 v2) | None -> (i,false) in - ise_try isevars [f1; f4] + ise_try evd [f1; f4] | MaybeFlexible flex1, Flexible ev2 -> let f1 i = if - is_unification_pattern_evar ev2 l2 & - not (occur_evar (fst ev2) (applist (term1,l1))) + is_unification_pattern_evar env ev2 l2 & + 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 isevars) (applist(term1,l1)) in + let t1 = nf_evar (evars_of evd) (applist appr1) in let t1 = solve_pattern_eqn env l2 t1 in - solve_simple_eqn evar_conv_x env isevars (pbty,ev2,t1) + solve_simple_eqn evar_conv_x env evd (pbty,ev2,t1) else if List.length l2 <= List.length l1 then @@ -284,7 +269,7 @@ and evar_eqappr_x env isevars pbty (term1,l1 as appr1) (term2,l2 as appr2) = evar_eqappr_x env i pbty (evar_apprec env i l1 v1) appr2 | None -> (i,false) in - ise_try isevars [f1; f4] + ise_try evd [f1; f4] | MaybeFlexible flex1, MaybeFlexible flex2 -> let f2 i = @@ -313,36 +298,36 @@ and evar_eqappr_x env isevars pbty (term1,l1 as appr1) (term2,l2 as appr2) = evar_eqappr_x env i pbty (evar_apprec env i l1 v1) appr2 | None -> (i,false) in - ise_try isevars [f2; f3; f4] + ise_try evd [f2; f3; f4] | Flexible ev1, Rigid _ -> if - is_unification_pattern_evar ev1 l1 & - not (occur_evar (fst ev1) (applist (term2,l2))) + is_unification_pattern_evar env ev1 l1 & + 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 isevars) (applist(term2,l2)) in + let t2 = nf_evar (evars_of evd) (applist appr2) in let t2 = solve_pattern_eqn env l1 t2 in - solve_simple_eqn evar_conv_x env isevars (pbty,ev1,t2) + solve_simple_eqn evar_conv_x env evd (pbty,ev1,t2) else (* Postpone the use of an heuristic *) - add_conv_pb (pbty,env,applist(term1,l1),applist(term2,l2)) isevars, + add_conv_pb (pbty,env,applist appr1,applist appr2) evd, true | Rigid _, Flexible ev2 -> if - is_unification_pattern_evar ev2 l2 & - not (occur_evar (fst ev2) (applist (term1,l1))) + is_unification_pattern_evar env ev2 l2 & + 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 isevars) (applist(term1,l1)) in + let t1 = nf_evar (evars_of evd) (applist appr1) in let t1 = solve_pattern_eqn env l2 t1 in - solve_simple_eqn evar_conv_x env isevars (pbty,ev2,t1) + solve_simple_eqn evar_conv_x env evd (pbty,ev2,t1) else (* Postpone the use of an heuristic *) - add_conv_pb (pbty,env,applist(term1,l1),applist(term2,l2)) isevars, + add_conv_pb (pbty,env,applist appr1,applist appr2) evd, true | MaybeFlexible flex1, Rigid _ -> @@ -355,7 +340,7 @@ and evar_eqappr_x env isevars pbty (term1,l1 as appr1) (term2,l2 as appr2) = evar_eqappr_x env i pbty (evar_apprec env i l1 v1) appr2 | None -> (i,false) in - ise_try isevars [f3; f4] + ise_try evd [f3; f4] | Rigid _ , MaybeFlexible flex2 -> let f3 i = @@ -367,19 +352,19 @@ and evar_eqappr_x env isevars pbty (term1,l1 as appr1) (term2,l2 as appr2) = evar_eqappr_x env i pbty appr1 (evar_apprec env i l2 v2) | None -> (i,false) in - ise_try isevars [f3; f4] + 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 isevars pbty (c1,l1) appr2 + | Cast (c1,_,_), _ -> evar_eqappr_x env evd pbty (c1,l1) appr2 - | _, Cast (c2,_,_) -> evar_eqappr_x env isevars pbty appr1 (c2,l2) + | _, Cast (c2,_,_) -> evar_eqappr_x env evd pbty appr1 (c2,l2) | Sort s1, Sort s2 when l1=[] & l2=[] -> - (isevars,base_sort_cmp pbty s1 s2) + (evd,base_sort_cmp pbty s1 s2) | Lambda (na,c1,c'1), Lambda (_,c2,c'2) when l1=[] & l2=[] -> - ise_and isevars + ise_and evd [(fun i -> evar_conv_x env i CONV c1 c2); (fun i -> let c = nf_evar (evars_of i) c1 in @@ -400,18 +385,18 @@ and evar_eqappr_x env isevars pbty (term1,l1 as appr1) (term2,l2 as appr2) = and appr2 = evar_apprec env i l2 (subst1 b2 c'2) in evar_eqappr_x env i pbty appr1 appr2 in - ise_try isevars [f1; f2] + ise_try evd [f1; f2] | LetIn (_,b1,_,c'1), _ ->(* On fait commuter les args avec le Let *) - let appr1 = evar_apprec env isevars l1 (subst1 b1 c'1) - in evar_eqappr_x env isevars pbty appr1 appr2 + let appr1 = evar_apprec env evd l1 (subst1 b1 c'1) + in evar_eqappr_x env evd pbty appr1 appr2 | _, LetIn (_,b2,_,c'2) -> - let appr2 = evar_apprec env isevars l2 (subst1 b2 c'2) - in evar_eqappr_x env isevars pbty appr1 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=[] -> - ise_and isevars + ise_and evd [(fun i -> evar_conv_x env i CONV c1 c2); (fun i -> let c = nf_evar (evars_of i) c1 in @@ -419,16 +404,16 @@ and evar_eqappr_x env isevars pbty (term1,l1 as appr1) (term2,l2 as appr2) = | Ind sp1, Ind sp2 -> if sp1=sp2 then - ise_list2 isevars (fun i -> evar_conv_x env i CONV) l1 l2 - else (isevars, false) + ise_list2 evd (fun i -> evar_conv_x env i CONV) l1 l2 + else (evd, false) | Construct sp1, Construct sp2 -> if sp1=sp2 then - ise_list2 isevars (fun i -> evar_conv_x env i CONV) l1 l2 - else (isevars, false) + ise_list2 evd (fun i -> evar_conv_x env i CONV) l1 l2 + else (evd, false) | Case (_,p1,c1,cl1), Case (_,p2,c2,cl2) -> - ise_and isevars + ise_and evd [(fun i -> evar_conv_x env i CONV p1 p2); (fun i -> evar_conv_x env i CONV c1 c2); (fun i -> ise_array2 i @@ -437,7 +422,7 @@ and evar_eqappr_x env isevars pbty (term1,l1 as appr1) (term2,l2 as appr2) = | Fix (li1,(_,tys1,bds1 as recdef1)), Fix (li2,(_,tys2,bds2)) -> if li1=li2 then - ise_and isevars + ise_and evd [(fun i -> ise_array2 i (fun i -> evar_conv_x env i CONV) tys1 tys2); (fun i -> ise_array2 i @@ -445,10 +430,10 @@ and evar_eqappr_x env isevars pbty (term1,l1 as appr1) (term2,l2 as appr2) = bds1 bds2); (fun i -> ise_list2 i (fun i -> evar_conv_x env i CONV) l1 l2)] - else (isevars,false) + else (evd,false) | CoFix (i1,(_,tys1,bds1 as recdef1)), CoFix (i2,(_,tys2,bds2)) -> if i1=i2 then - ise_and isevars + ise_and evd [(fun i -> ise_array2 i (fun i -> evar_conv_x env i CONV) tys1 tys2); (fun i -> ise_array2 i @@ -456,30 +441,31 @@ and evar_eqappr_x env isevars pbty (term1,l1 as appr1) (term2,l2 as appr2) = bds1 bds2); (fun i -> ise_list2 i (fun i -> evar_conv_x env i CONV) l1 l2)] - else (isevars,false) + else (evd,false) - | (Meta _ | Lambda _), _ -> (isevars,false) - | _, (Meta _ | Lambda _) -> (isevars,false) + | (Meta _ | Lambda _), _ -> (evd,false) + | _, (Meta _ | Lambda _) -> (evd,false) - | (Ind _ | Construct _ | Sort _ | Prod _), _ -> (isevars,false) - | _, (Ind _ | Construct _ | Sort _ | Prod _) -> (isevars,false) + | (Ind _ | Construct _ | Sort _ | Prod _), _ -> (evd,false) + | _, (Ind _ | Construct _ | Sort _ | Prod _) -> (evd,false) | (App _ | Case _ | Fix _ | CoFix _), - (App _ | Case _ | Fix _ | CoFix _) -> (isevars,false) + (App _ | Case _ | Fix _ | CoFix _) -> (evd,false) | (Rel _ | Var _ | Const _ | Evar _), _ -> assert false | _, (Rel _ | Var _ | Const _ | Evar _) -> assert false -and conv_record env isevars (c,bs,(params,params1),(us,us2),(ts,ts1),c1) = - let (isevars',ks) = +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) b -> + (fun (i,ks,m) b -> + if m=0 then (i,t2::ks, n-1) else let dloc = (dummy_loc,InternalHole) in let (i',ev) = new_evar i env ~src:dloc (substl ks b) in - (i', ev :: ks)) - (isevars,[]) bs + (i', ev :: ks, n - 1)) + (evd,[],n) bs in - ise_and isevars' + ise_and evd' [(fun i -> ise_list2 i (fun i u1 u -> evar_conv_x env i CONV u1 (substl ks u)) @@ -491,65 +477,80 @@ and conv_record env isevars (c,bs,(params,params1),(us,us2),(ts,ts1),c1) = (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))))] -let first_order_unification env isevars pbty (term1,l1) (term2,l2) = +(* We assume here |l1| <= |l2| *) + +let first_order_unification env evd (ev1,l1) (term2,l2) = + let (deb2,rest2) = list_chop (List.length l2-List.length l1) l2 in + ise_and evd + (* First compare extra args for better failure message *) + [(fun i -> ise_list2 i (fun i -> evar_conv_x env i CONV) rest2 l1); + (fun i -> + (* Then instantiate evar unless already done by unifying args *) + let t2 = applist(term2,deb2) in + if is_defined_evar i ev1 then + evar_conv_x env i CONV t2 (mkEvar ev1) + else + solve_simple_eqn evar_conv_x env i (CONV,ev1,t2))] + +let choose_less_dependent_instance evk evd term args = + let evi = Evd.find (evars_of 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 + +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 (term1,l1 as appr1) = decompose_app t1 in + let (term2,l2 as appr2) = decompose_app t2 in match kind_of_term term1, kind_of_term term2 with - | Evar ev1,_ when List.length l1 <= List.length l2 -> - let (deb2,rest2) = list_chop (List.length l2-List.length l1) l2 in - ise_and isevars - (* First compare extra args for better failure message *) - [(fun i -> ise_list2 i (fun i -> evar_conv_x env i CONV) rest2 l1); - (fun i -> - (* Then instantiate evar unless already done by unifying args *) - let t2 = applist(term2,deb2) in - if is_defined_evar i ev1 then - evar_conv_x env i pbty t2 (mkEvar ev1) - else - solve_simple_eqn evar_conv_x env i (pbty,ev1,t2))] - | _,Evar ev2 when List.length l2 <= List.length l1 -> - let (deb1,rest1) = list_chop (List.length l1-List.length l2) l1 in - ise_and isevars - (* First compare extra args for better failure message *) - [(fun i -> ise_list2 i (fun i -> evar_conv_x env i CONV) rest1 l2); - (fun i -> - (* Then instantiate evar unless already done by unifying args *) - let t1 = applist(term1,deb1) in - if is_defined_evar i ev2 then - evar_conv_x env i pbty t1 (mkEvar ev2) - else - solve_simple_eqn evar_conv_x env i (pbty,ev2,t1))] - | _ -> - (* Some head evar have been instantiated *) - evar_conv_x env isevars pbty (applist(term1,l1)) (applist(term2,l2)) - -let consider_remaining_unif_problems env isevars = - let (isevars,pbs) = get_conv_pbs isevars (fun _ -> true) in + | Evar (evk1,args1), (Rel _|Var _) when l1 = [] & l2 = [] -> + (* The typical kind of constraint coming from patter-matching return + type inference *) + assert (array_for_all (fun a -> a = term2 or isEvar a) args1); + choose_less_dependent_instance evk1 evd term2 args1, true + | (Rel _|Var _), Evar (evk2,args2) when l1 = [] & l2 = [] -> + (* The typical kind of constraint coming from patter-matching return + type inference *) + assert (array_for_all ((=) term1) args2); + choose_less_dependent_instance evk2 evd term1 args2, true + | Evar ev1,_ when List.length l1 <= List.length l2 -> + (* On "?n t1 .. tn = u u1 .. u(n+p)", try first-order unification *) + first_order_unification env evd (ev1,l1) appr2 + | _,Evar ev2 when List.length l2 <= List.length l1 -> + (* On "u u1 .. u(n+p) = ?n t1 .. tn", try first-order unification *) + first_order_unification env evd (ev2,l2) appr1 + | _ -> + (* Some head evar have been instantiated, or unknown kind of problem *) + evar_conv_x env evd pbty t1 t2 + +let consider_remaining_unif_problems env evd = + let (evd,pbs) = extract_all_conv_pbs evd in List.fold_left - (fun (isevars,b as p) (pbty,env,t1,t2) -> - if b then first_order_unification env isevars pbty - (apprec_nohdbeta env isevars (whd_castappevar (evars_of isevars) t1)) - (apprec_nohdbeta env isevars (whd_castappevar (evars_of isevars) t2)) - else p) - (isevars,true) + (fun (evd,b as p) (pbty,env,t1,t2) -> + if b then apply_conversion_problem_heuristic env evd pbty t1 t2 else p) + (evd,true) pbs (* Main entry points *) -let the_conv_x env t1 t2 isevars = - match evar_conv_x env isevars CONV t1 t2 with +let the_conv_x env t1 t2 evd = + match evar_conv_x env evd CONV t1 t2 with (evd',true) -> evd' | _ -> raise Reduction.NotConvertible -let the_conv_x_leq env t1 t2 isevars = - match evar_conv_x env isevars CUMUL t1 t2 with +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 isevars t1 t2 = - match evar_conv_x env !isevars CONV t1 t2 with - (evd',true) -> isevars := evd'; true +let e_conv env evd t1 t2 = + match evar_conv_x env !evd CONV t1 t2 with + (evd',true) -> evd := evd'; true | _ -> false -let e_cumul env isevars t1 t2 = - match evar_conv_x env !isevars CUMUL t1 t2 with - (evd',true) -> isevars := evd'; true +let e_cumul env evd t1 t2 = + match evar_conv_x env !evd CUMUL t1 t2 with + (evd',true) -> evd := evd'; true | _ -> false diff --git a/pretyping/evarutil.ml b/pretyping/evarutil.ml index 16059d94..fafce268 100644 --- a/pretyping/evarutil.ml +++ b/pretyping/evarutil.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: evarutil.ml 10374 2007-12-13 13:16:52Z notin $ *) +(* $Id: evarutil.ml 11127 2008-06-14 15:39:46Z herbelin $ *) open Util open Pp @@ -16,10 +16,12 @@ open Univ open Term open Termops open Sign +open Pre_env open Environ open Evd open Reductionops open Pretype_errors +open Retyping (* Expanding existential variables (pretyping.ml) *) (* 1- whd_ise fails if an existential is undefined *) @@ -28,10 +30,10 @@ exception Uninstantiated_evar of existential_key let rec whd_ise sigma c = match kind_of_term c with - | Evar (ev,args) when Evd.mem sigma ev -> - if Evd.is_defined sigma ev then - whd_ise sigma (existential_value sigma (ev,args)) - else raise (Uninstantiated_evar ev) + | 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 @@ -39,8 +41,8 @@ let rec whd_ise sigma c = let whd_castappevar_stack sigma c = let rec whrec (c, l as s) = match kind_of_term c with - | Evar (ev,args) when Evd.mem sigma ev & Evd.is_defined sigma ev -> - whrec (existential_value sigma (ev,args), l) + | 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 @@ -55,6 +57,17 @@ 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 = + Sign.map_named_context (Reductionops.nf_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 nc' = nf_named_context_evar sigma (Environ.named_context env) in + let rel' = nf_rel_context_evar sigma (Environ.rel_context env) in + push_rel_context rel' (reset_with_named_context (val_of_named_context nc') env) + let nf_evar_info evc info = { info with evar_concl = Reductionops.nf_evar evc info.evar_concl; @@ -63,13 +76,13 @@ let nf_evar_info evc info = 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 isevars = Evd.evars_reset_evd (nf_evars (Evd.evars_of isevars)) isevars +let nf_evar_defs evd = Evd.evars_reset_evd (nf_evars (Evd.evars_of evd)) evd -let nf_isevar isevars = nf_evar (Evd.evars_of isevars) -let j_nf_isevar isevars = j_nf_evar (Evd.evars_of isevars) -let jl_nf_isevar isevars = jl_nf_evar (Evd.evars_of isevars) -let jv_nf_isevar isevars = jv_nf_evar (Evd.evars_of isevars) -let tj_nf_isevar isevars = tj_nf_evar (Evd.evars_of isevars) +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) (**********************) (* Creating new metas *) @@ -85,8 +98,8 @@ let mk_new_meta () = mkMeta(new_meta()) let collect_evars emap c = let rec collrec acc c = match kind_of_term c with - | Evar (k,_) -> - if Evd.mem emap k & not (Evd.is_defined emap k) then k::acc + | Evar (evk,_) -> + if Evd.mem emap evk & not (Evd.is_defined emap evk) then evk::acc else (* No recursion on the evar instantiation *) acc | _ -> fold_constr collrec acc c in @@ -111,7 +124,7 @@ let evars_to_metas sigma (emap, c) = mkCast (mkMeta n, DEFAULTcast, ty) in let rec replace c = match kind_of_term c with - Evar (k,_ as ev) when Evd.mem emap' k -> change_exist ev + | Evar (evk,_ as ev) when Evd.mem emap' evk -> change_exist ev | _ -> map_constr replace c in (sigma', replace c) @@ -120,35 +133,11 @@ let evars_to_metas sigma (emap, c) = let non_instantiated sigma = let listev = to_list sigma in List.fold_left - (fun l (ev,evd) -> - if evd.evar_body = Evar_empty then - ((ev,nf_evar_info sigma evd)::l) else l) + (fun l (ev,evi) -> + if evi.evar_body = Evar_empty then + ((ev,nf_evar_info sigma evi)::l) else l) [] listev -(*************************************) -(* Metas *) - -let meta_value evd mv = - let rec valrec mv = - try - let b = meta_fvalue evd mv in - instance - (List.map (fun mv' -> (mv',valrec mv')) (Metaset.elements b.freemetas)) - b.rebus - with Anomaly _ | Not_found -> - mkMeta mv - in - valrec mv - -let meta_instance env b = - let c_sigma = - List.map - (fun mv -> (mv,meta_value env mv)) (Metaset.elements b.freemetas) - in - instance c_sigma b.rebus - -let nf_meta env c = meta_instance env (mk_freelisted c) - (**********************) (* Creating new evars *) (**********************) @@ -162,63 +151,108 @@ let new_untyped_evar = * functional operations on evar sets * *------------------------------------*) -(* All ids of sign must be distincts! *) -let new_evar_instance sign evd typ ?(src=(dummy_loc,InternalHole)) instance = - let ctxt = named_context_of_val sign in - assert (List.length instance = named_context_length ctxt); - if not (list_distinct (ids_of_named_context ctxt)) then - anomaly "new_evar_instance: two vars have the same name"; - let newev = new_untyped_evar() in - (evar_declare sign newev typ ~src:src evd, - mkEvar (newev,Array.of_list instance)) - -let make_evar_instance_with_rel env = - let n = rel_context_length (rel_context env) in - let vars = - fold_named_context - (fun env (id,b,_) l -> (* if b=None then *) mkVar id :: l (*else l*)) - env ~init:[] in - snd (fold_rel_context - (fun env (_,b,_) (i,l) -> - (i-1, (*if b=None then *) mkRel i :: l (*else l*))) - env ~init:(n,vars)) - -let make_subst env args = - snd (fold_named_context - (fun env (id,b,c) (args,l) -> - match b, args with - | (* None *) _ , a::rest -> (rest, (id,a)::l) -(* | Some _, _ -> g*) - | _ -> anomaly "Instance does not match its signature") - env ~init:(List.rev args,[])) - -(* [new_isevar] declares a new existential in an env env with type typ *) +let new_evar_instance sign evd typ ?(src=(dummy_loc,InternalHole)) ?filter instance = + let instance = + match filter with + | None -> instance + | Some filter -> snd (list_filter2 (fun b c -> b) (filter,instance)) in + assert + (let ctxt = named_context_of_val sign in + list_distinct (ids_of_named_context ctxt)); + let newevk = new_untyped_evar() in + let evd = evar_declare sign newevk typ ~src:src ?filter evd in + (evd,mkEvar (newevk,Array.of_list instance)) + +(* Knowing that [Gamma |- ev : T] and that [ev] is applied to [args], + * [make_projectable_subst ev args] builds the substitution [Gamma:=args]. + * If a variable and an alias of it are bound to the same instance, we skip + * the alias (we just use eq_constr -- instead of conv --, since anyway, + * only instances that are variables -- or evars -- are later considered; + * morever, we can bet that similar instances came at some time from + * the very same substitution. The removal of aliased duplicates is + * useful to ensure the uniqueness of a projection. +*) + +let make_projectable_subst 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 + 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) + | _ -> anomaly "Instance does not match its signature") + sign (List.rev (Array.to_list args),[])) + +let make_pure_subst evi args = + snd (List.fold_right + (fun (id,b,c) (args,l) -> + match args with + | a::rest -> (rest, (id,a)::l) + | _ -> anomaly "Instance does not match its signature") + (evar_filtered_context evi) (List.rev (Array.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 + * + * a1 ... an xp ... x1 + * \----------- named part ------------/ + * + * 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). + * 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 + * we have the property that u and phi(t) are convertible in env. + *) + +let push_rel_context_to_named_context env typ = + (* compute the instances relative to the named context and rel_context *) + let ids = List.map pi1 (named_context env) in + let inst_vars = List.map mkVar ids in + let inst_rels = List.rev (rel_list 0 (nb_rel env)) in + (* move the rel context to a named context and extend the named instance *) + (* with vars of the rel context *) + (* We do keep the instances corresponding to local definition (see above) *) + let (subst, _, env) = + Sign.fold_rel_context + (fun (na,c,t) (subst, avoid, env) -> + let id = next_name_away na avoid in + let d = (id,Option.map (substl subst) c,substl subst t) in + (mkVar id :: subst, id::avoid, push_named d env)) + (rel_context env) ~init:([], ids, env) in + (named_context_val env, substl subst typ, inst_rels@inst_vars) + +(* [new_evar] declares a new existential in an env env with type typ *) (* Converting the env into the sign of the evar to define *) -let push_rel_context_to_named_context env = - let (subst,_,env) = - Sign.fold_rel_context - (fun (na,c,t) (subst,avoid,env) -> - let id = next_name_away na avoid in - ((mkVar id)::subst, - id::avoid, - push_named (id,option_map (substl subst) c, - type_app (substl subst) t) - env)) - (rel_context env) ~init:([],ids_of_named_context (named_context env),env) - in (subst, (named_context_val env)) - -let new_evar evd env ?(src=(dummy_loc,InternalHole)) typ = - let subst,sign = push_rel_context_to_named_context env in - let typ' = substl subst typ in - let instance = make_evar_instance_with_rel env in - new_evar_instance sign evd typ' ~src:src instance +let new_evar evd env ?(src=(dummy_loc,InternalHole)) ?filter typ = + let sign,typ',instance = push_rel_context_to_named_context env typ in + new_evar_instance sign evd typ' ~src:src ?filter instance (* The same using side-effect *) -let e_new_evar evd env ?(src=(dummy_loc,InternalHole)) ty = - let (evd',ev) = new_evar !evd env ~src:src ty in - evd := evd'; - ev +let e_new_evar evdref env ?(src=(dummy_loc,InternalHole)) ?filter ty = + let (evd',ev) = new_evar !evdref env ~src:src ?filter ty in + evdref := evd'; + ev (*------------------------------------* * operations on the evar constraints * @@ -234,7 +268,7 @@ let is_pattern inst = let evar_well_typed_body evd ev evi body = try - let env = evar_env evi in + let env = evar_unfiltered_env evi in let ty = evi.evar_concl in Typing.check env (evars_of evd) body ty; true @@ -246,92 +280,111 @@ let evar_well_typed_body evd ev evi body = print_constr body); false -let strict_inverse = false - -let inverse_instance env isevars ev evi inst rhs = - let subst = make_subst (evar_env evi) (Array.to_list inst) in - let subst = List.map (fun (x,y) -> (y,mkVar x)) subst in - let evd = ref isevars in - let error () = - error_not_clean env (evars_of !evd) ev rhs (evar_source ev !evd) in - let rec subs rigid k t = - match kind_of_term t with - | Rel i -> - if i<=k then t - else - (try List.assoc (mkRel (i-k)) subst - with Not_found -> - if rigid then error() - else if strict_inverse then - failwith "cannot solve pb yet" - else t) - | Var id -> - (try List.assoc t subst - with Not_found -> - if rigid then error() - else if - not strict_inverse && - List.exists (fun (id',_,_) -> id=id') (evar_context evi) - then - failwith "cannot solve pb yet" - else t) - | Evar (ev,args) -> - if Evd.is_defined_evar !evd (ev,args) then - subs rigid k (existential_value (evars_of !evd) (ev,args)) - else - let args' = Array.map (subs false k) args in - mkEvar (ev,args') - | _ -> map_constr_with_binders succ (subs rigid) k t in - let body = subs true 0 (nf_evar (evars_of isevars) rhs) in - (!evd,body) - - -let is_defined_equation env evd (ev,inst) rhs = - is_pattern inst && - not (occur_evar ev rhs) && - try - let evi = Evd.find (evars_of evd) ev in - let (evd',body) = inverse_instance env evd ev evi inst rhs in - evar_well_typed_body evd' ev evi body - with Failure _ -> 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 + * virtually have x1..xq, y1..yk | ?e1' and the equation + * Γ, y1..yk |- ?e1'[u1..uq y1..yk] = c. + * What we do is to formally introduce ?e1' in context x1..xq, Γ, y1..yk, + * but forbidding it to use the variables of Γ (otherwise said, + * Γ is here only for ensuring the correct typing of ?e1'). + * + * In fact, we optimize a little and try to compute a maximum + * common subpart of x1..xq and Γ. This is done by detecting the + * longest subcontext x1..xp such that Γ = x1'..xp' z1..zm and + * u1..up = x1'..xp'. + * + * At the end, we return ?e1'[x1..xn z1..zm y1..yk] so that ?e1 can be + * instantiated by (...\y1 ... \yk ... ?e1[x1..xn z1..zm y1..yk]) and the + * new problem is Σ; Γ, y1..yk |- ?e1'[u1..un z1..zm y1..yk] = c, + * making the z1..zm unavailable. + * + * This is what [extend_evar Γ evd k (?e1[u1..uq]) c] does. + *) +let shrink_context env subst ty = + let rev_named_sign = List.rev (named_context env) in + let rel_sign = rel_context env in + (* We merge the contexts (optimization) *) + let rec shrink_rel i subst rel_subst rev_rel_sign = + match subst,rev_rel_sign with + | (id,c)::subst,_::rev_rel_sign when c = mkRel i -> + shrink_rel (i-1) subst (mkVar id::rel_subst) rev_rel_sign + | _ -> + substl_rel_context rel_subst (List.rev rev_rel_sign), + substl rel_subst ty + in + let rec shrink_named subst named_subst rev_named_sign = + match subst,rev_named_sign with + | (id,c)::subst,(id',b',t')::rev_named_sign when c = mkVar id' -> + shrink_named subst ((id',mkVar id)::named_subst) rev_named_sign + | _::_, [] -> + let nrel = List.length rel_sign in + let rel_sign, ty = shrink_rel nrel subst [] (List.rev rel_sign) in + [], map_rel_context (replace_vars named_subst) rel_sign, + replace_vars named_subst ty + | _ -> + map_named_context (replace_vars named_subst) (List.rev rev_named_sign), + rel_sign, ty + in + 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 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 named_sign',rel_sign',ty = + if k = 0 then [], [], ty + else shrink_context env (List.rev (make_pure_subst evi1 args1)) ty in + let extenv = + List.fold_right push_rel rel_sign' + (List.fold_right push_named named_sign' (evar_unfiltered_env evi1)) in + let nb_to_hide = rel_context_length rel_sign' - k in + let rel_filter = list_map_i (fun i _ -> i > nb_to_hide) 1 rel_sign' in + let named_filter1 = List.map (fun _ -> true) (evar_context evi1) in + let named_filter2 = List.map (fun _ -> false) named_sign' in + let filter = rel_filter@named_filter2@named_filter1 in + let evar1' = e_new_evar evdref extenv ~filter:filter ty in + let evk1',args1'_in_env = destEvar evar1' in + let args1'_in_extenv = Array.map (lift k) (overwrite_first args1'_in_env args1) in + (evar1',(evk1',args1'_in_extenv)) + +let subfilter p filter l = + let (filter,_,l) = + List.fold_left (fun (filter,l,newl) b -> + if b then + let a,l' = match l with a::args -> a,args | _ -> assert false in + if p a then (true::filter,l',a::newl) else (false::filter,l',newl) + else (false::filter,l,newl)) + ([],l,[]) filter in + (List.rev filter,List.rev l) + +let restrict_upon_filter evd evi evk p args = + let filter = evar_filter evi in + let newfilter,newargs = subfilter p filter args in + if newfilter <> filter then + let (evd,newev) = new_evar evd (evar_unfiltered_env evi) ~src:(evar_source evk evd) + ~filter:newfilter evi.evar_concl in + let evd = Evd.evar_define evk newev evd in + evd,fst (destEvar newev),newargs + else + evd,evk,args -(* Redefines an evar with a smaller context (i.e. it may depend on less - * variables) such that c becomes closed. - * Example: in [x:?1; y:(list ?2)] x=y /\ x=(nil bool) - * ?3 <-- ?1 no pb: env of ?3 is larger than ?1's - * ?1 <-- (list ?2) pb: ?2 may depend on x, but not ?1. - * What we do is that ?2 is defined by a new evar ?4 whose context will be - * a prefix of ?2's env, included in ?1's env. +exception Dependency_error of identifier - Concretely, the assumptions are "env |- ev : T" and "Gamma |- - ev[hyps:=args]" for some Gamma whose de Bruijn context has length k. - We create "env' |- ev' : T" for some env' <= env and define ev:=ev' -*) +module EvkOrd = +struct + type t = Term.existential_key + let compare = Pervasives.compare +end -let do_restrict_hyps env k evd ev args = - let args = Array.to_list args in - let evi = Evd.find (evars_of !evd) ev in - let hyps = evar_context evi in - let (hyps',ncargs) = list_filter2 (fun _ a -> closedn k a) (hyps,args) in - (* No care is taken in case the evar type uses vars filtered out! - Assuming that the restriction comes from a well-typed Flex/Flex - unification problem (see real_clean), the type of the evar cannot - depend on variables that are not in the scope of the other evar, - since this other evar has the same type (up to unification). - Since moreover, the evar contexts uses names only, the - restriction raise no de Bruijn reallocation problem *) - let env' = - Sign.fold_named_context push_named hyps' ~init:(reset_context env) in - let nc = e_new_evar evd env' ~src:(evar_source ev !evd) evi.evar_concl in - evd := Evd.evar_define ev nc !evd; - let (evn,_) = destEvar nc in - mkEvar(evn,Array.of_list ncargs) - +module EvkSet = Set.Make(EvkOrd) -exception Dependency_error of identifier - -let rec check_and_clear_in_constr evd c ids hist = +let rec check_and_clear_in_constr evdref c ids hist = (* returns a new constr where all the evars have been 'cleaned' (ie the hypotheses ids have been removed from the contexts of evars *) @@ -341,234 +394,542 @@ let rec check_and_clear_in_constr evd c ids hist = in match kind_of_term c with | ( Rel _ | Meta _ | Sort _ ) -> c + | ( Const _ | Ind _ | Construct _ ) -> let vars = Environ.vars_of_global (Global.env()) c in List.iter check vars; c + | Var id' -> check id'; mkVar id' - | Evar (e,l) -> - if (List.mem e hist) then + + | 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 = nf_evar (evars_of !evdref) c in + (check_and_clear_in_constr evdref nc ids hist) + else if EvkSet.mem evk hist then + (* Loop detection => do nothing *) c - else - begin - if Evd.is_defined_evar !evd (e,l) then - (* If e is already defined we replace it by its definition *) - let nc = nf_evar (evars_of !evd) c in - (check_and_clear_in_constr evd nc ids hist) - 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 !evd) e in - let nconcl = check_and_clear_in_constr evd (evar_concl evi) ids (e::hist) in - let (nhyps,nargs) = - List.fold_right2 - (fun (id,ob,c) i (hy,ar) -> - if List.mem id ids then - (hy,ar) - else - let d' = (id, - (match ob with - None -> None - | Some b -> Some (check_and_clear_in_constr evd b ids (e::hist))), - check_and_clear_in_constr evd c ids (e::hist)) in - let i' = check_and_clear_in_constr evd i ids (e::hist) in - (d'::hy, i'::ar) - ) - (evar_context evi) (Array.to_list l) ([],[]) in - let env = Sign.fold_named_context push_named nhyps ~init:(empty_env) in - let ev'= e_new_evar evd env ~src:(evar_source e !evd) nconcl in - evd := Evd.evar_define e ev' !evd; - let (e',_) = destEvar ev' in - mkEvar(e', Array.of_list nargs) - end - | _ -> map_constr (fun c -> check_and_clear_in_constr evd c ids hist) c - -and clear_hyps_in_evi evd evi ids = - (* clear_evar_hyps erases hypotheses ids in evi, checking if some + 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 ctxt = Evd.evar_filtered_context evi in + let (nhyps,nargs,rids) = + List.fold_right2 + (fun (rid,ob,c as h) a (hy,ar,ri) -> + match kind_of_term a with + | Var id -> if List.mem id ids then (hy,ar,id::ri) else (h::hy,a::ar,ri) + | _ -> (h::hy,a::ar,ri) + ) + ctxt (Array.to_list l) ([],[],[]) in + (* nconcl must be computed ids (non instanciated hyps) *) + let nconcl = check_and_clear_in_constr evdref (evar_concl evi) rids (EvkSet.add evk hist) 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; + let (evk',_) = destEvar ev' in + mkEvar(evk', Array.of_list nargs) + + | _ -> map_constr (fun c -> check_and_clear_in_constr evdref c ids hist) c + +let clear_hyps_in_evi evdref hyps concl ids = + (* clear_evar_hyps erases hypotheses ids in hyps, checking if some hypothesis does not depend on a element of ids, and erases ids in the contexts of the evars occuring in evi *) - let nconcl = try check_and_clear_in_constr evd (evar_concl evi) ids [] + let nconcl = try check_and_clear_in_constr evdref concl ids EvkSet.empty with Dependency_error id' -> error (string_of_id id' ^ " is used in conclusion") in let (nhyps,_) = - let aux (id,ob,c) = + let check_context (id,ob,c) = try (id, (match ob with None -> None - | Some b -> Some (check_and_clear_in_constr evd b ids [])), - check_and_clear_in_constr evd c ids []) + | Some b -> Some (check_and_clear_in_constr evdref b ids EvkSet.empty)), + check_and_clear_in_constr evdref c ids EvkSet.empty) with Dependency_error id' -> error (string_of_id id' ^ " is used in hypothesis " - ^ string_of_id id) + ^ string_of_id id) in - remove_hyps ids aux (evar_hyps evi) - in - { evi with - evar_concl = nconcl; - evar_hyps = nhyps} - + let check_value vk = + match !vk with + | VKnone -> vk + | VKvalue (v,d) -> + if (List.for_all (fun e -> not (Idset.mem e d)) ids) then + (* v does depend on any of ids, it's ok *) + vk + else + (* v depends on one of the cleared hyps: we forget the computed value *) + ref VKnone + in + remove_hyps ids check_context check_value hyps + in + (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 rec expand_var env x = match kind_of_term x with + | Rel n -> + begin try match pi2 (lookup_rel n env) with + | Some t when isRel t -> expand_var env (lift n t) + | _ -> x + with Not_found -> x + end + | Var id -> + begin match pi2 (lookup_named id env) with + | Some t when isVar t -> expand_var env t + | _ -> x + end + | _ -> x + +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 + +(* [find_projectable_vars env sigma y subst] finds all vars of [subst] + * that project on [y]. It is able to find solutions to the following + * two kinds of problems: + * + * - ?n[...;x:=y;...] = y + * - ?n[...;x:=?m[args];...] = y with ?m[args] = y recursively solvable + * + * (see test-suite/success/Fixpoint.v for an example of application of + * the second kind of problem). + * + * The seek for [y] is up to variable aliasing. In case of solutions that + * differ only up to aliasing, the binding that requires the less + * steps of alias reduction is kept. At the end, only one solution up + * to aliasing is kept. + * + * [find_projectable_vars] also unifies against evars that themselves mention + * [y] and recursively. + * + * In short, the following situations give the following solutions: + * + * problem evar ctxt soluce remark + * z1; z2:=z1 |- ?ev[z1;z2] = z1 y1:A; y2:=y1 y1 \ thanks to defs kept in + * z1; z2:=z1 |- ?ev[z1;z2] = z2 y1:A; y2:=y1 y2 / subst and preferring = + * z1; z2:=z1 |- ?ev[z1] = z2 y1:A y1 thanks to expand_var + * z1; z2:=z1 |- ?ev[z2] = z1 y1:A y1 thanks to expand_var + * z3 |- ?ev[z3;z3] = z3 y1:A; y2:=y1 y2 see make_projectable_subst + * + * Remark: [find_projectable_vars] assumes that identical instances of + * variables in the same set of aliased variables are already removed (see + * [make_projectable_subst]) + *) -let need_restriction k args = not (array_for_all (closedn k) args) +exception NotUnique + +type evar_projection = +| ProjectVar +| ProjectEvar of existential * evar_info * identifier * evar_projection + +let rec find_projectable_vars 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 isEvar y' then + 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 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 + +(* [filter_solution] checks if one and only one possible projection exists + * among a set of solutions to a projection problem *) + +let filter_solution = function + | [] -> raise Not_found + | (id,p)::_::_ -> raise NotUnique + | [id,p] -> (mkVar id, p) + +let project_with_effects env sigma effects t subst = + let c, p = filter_solution (find_projectable_vars env sigma t subst) in + effects := p :: !effects; + c + +(* In case the solution to a projection problem requires the instantiation of + * subsidiary evars, [do_projection_effects] performs them; it + * also try to instantiate the type of those subsidiary evars if their + * type is an evar too. + * + * Note: typing creates new evar problems, which induces a recursive dependency + * with [evar_define]. To avoid a too large set of recursive functions, we + * pass [evar_define] to [do_projection_effects] as a parameter. + *) -(* 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 only Miller-Pfenning patterns unification) +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 + (* 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 + if not (isSort ty) then + (* Don't try to instantiate if a sort because if evar_concl is an + evar it may commit to a univ level which is not the right + one (however, regarding coercions, because t is obtained by + unif, we know that no coercion can be inserted) *) + let subst = make_pure_subst evi argsv in + let ty' = replace_vars subst evi.evar_concl in + let ty' = whd_evar (evars_of 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. + * The strategy is to imitate the structure of c and then to invert + * the variables of c (i.e. rels or vars of Γ) using the algorithm + * implemented by project_with_effects/find_projectable_vars. + * It returns either a unique solution or says whether 0 or more than + * 1 solutions is found. + * + * Precondition: Σ; Γ, y1..yk |- c /\ Σ; Γ |- u1..un + * Postcondition: if φ(x1..xn) is returned then + * Σ; Γ, y1..yk |- φ(u1..un) = c /\ x1..xn |- φ(x1..xn) + * + * The effects correspond to evars instantiated while trying to project. + * + * [invert_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]. + *) - * 1) Let a unification problem "env |- ev[hyps:=args] = rhs" - * 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 failing if some Rel/Var not in scope +type projectibility_kind = + | NoUniqueProjection + | UniqueProjection of constr * evar_projection list - * Note: we don't assume rhs in normal form, it may fail while it would - * have succeeded after some reductions +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 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 + | Var id -> + project_with_effects env sigma effects t subst_in_env + | _ -> + map_constr_with_binders succ aux k t in + 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 + invert_arg_from_subst env k sigma subst_in_env c_in_env_extended_with_k_binders + +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 + match projs with + | NoUniqueProjection -> raise NotUnique + | UniqueProjection (c,effects) -> + (List.fold_left (do_projection_effects f env ty) evd effects, c) + +let filter_of_projection = function CannotInvert -> false | _ -> true + +let filter_along f projs v = + let l = Array.to_list v in + let _,l = list_filter2 (fun b c -> f b) (projs,l) in + Array.of_list l + +(* Redefines an evar with a smaller context (i.e. it may depend on less + * variables) such that c becomes closed. + * Example: in "fun (x:?1) (y:list ?2[x]) => x = y :> ?3[x,y] /\ x = nil bool" + * ?3 <-- ?1 no pb: env of ?3 is larger than ?1's + * ?1 <-- list ?2 pb: ?2 may depend on x, but not ?1. + * What we do is that ?2 is defined by a new evar ?4 whose context will be + * a prefix of ?2's env, included in ?1's env. + * + * If "hyps |- ?e : T" and "filter" selects a subset hyps' of hyps then + * [do_restrict_hyps evd ?e filter] sets ?e:=?e'[hyps'] and returns ?e' + * such that "hyps' |- ?e : T" *) -(* Note: error_not_clean should not be an error: it simply means that the - * conversion test that lead to the faulty call to [real_clean] should return - * false. The problem is that we won't get the right error message. + +let do_restrict_hyps_virtual evd evk filter = + (* What to do with dependencies? + Assume we have x:A, y:B(x), z:C(x,y) |- ?e:T(x,y,z) and restrict on y. + - If y is in a non-erasable position in C(x,y) (i.e. it is not below an + occurrence of x in the hnf of C), then z should be removed too. + - If y is in a non-erasable position in T(x,y,z) then the problem is + unsolvable. + Computing whether y is erasable or not may be costly and the + interest for this early detection in practice is not obvious. We let + it for future work. Anyway, thanks to the use of filters, the whole + context remains consistent. *) + let evi = Evd.find (evars_of 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) -> + if oldb then List.hd filter::l,List.tl filter else (false::l,filter)) + oldfilter ([],List.rev filter) in + new_evar evd env ~src:(evar_source evk evd) + ~filter:filter evi.evar_concl + +let do_restrict_hyps evd evk projs = + let filter = List.map filter_of_projection projs in + if List.for_all (fun x -> x) filter then + evd,evk + else + let evd,nc = do_restrict_hyps_virtual evd evk filter in + let evd = Evd.evar_define evk nc evd in + let evk',_ = destEvar nc in + evd,evk' + +(* [postpone_evar_term] postpones an equation of the form ?e[σ] = c *) + +let postpone_evar_term env evd (evk,argsv) rhs = + let rhs = expand_vars_in_term env rhs in + let evi = Evd.find (evars_of evd) evk in + let evd,evk,args = + restrict_upon_filter evd evi evk + (* Keep only variables that depends in rhs *) + (* This is not safe: is the variable is a local def, its body *) + (* may contain references to variables that are removed, leading to *) + (* a ill-formed context. We would actually need a notion of filter *) + (* that says that the body is hidden. Note that expand_vars_in_term *) + (* expands only rels and vars aliases, not rels or vars bound to an *) + (* arbitrary complex term *) + (fun a -> not (isRel a || isVar a) || dependent a rhs) + (Array.to_list argsv) in + let args = Array.of_list args in + let pb = (Reduction.CONV,env,mkEvar(evk,args),rhs) in + Evd.add_conv_pb pb evd + +(* [postpone_evar_evar] postpones an equation of the form ?e1[σ1] = ?e2[σ2] *) + +let postpone_evar_evar env evd projs1 (evk1,args1) projs2 (evk2,args2) = + (* Leave an equation between (restrictions of) ev1 andv ev2 *) + let args1' = filter_along filter_of_projection projs1 args1 in + let evd,evk1' = do_restrict_hyps evd evk1 projs1 in + let args2' = filter_along filter_of_projection projs2 args2 in + let evd,evk2' = do_restrict_hyps evd evk2 projs2 in + 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 + * 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 + * 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), + * - otherwise, each position i s.t. ui does not occur in v1..vp has to + * be restricted and similarly for the vi, and we leave the equation + * as an open equation (performed by [postpone_evar]) + * + * Warning: the notion of unique φj is relative to some given class + * of unification problems + * + * Note: argument f is the function used to instantiate evars. *) -exception NotClean of constr +exception CannotProject of projectibility_status list -let real_clean env isevars ev evi args rhs = - let evd = ref isevars in - let subst = List.map (fun (x,y) -> (y,mkVar x)) (list_uniquize args) in - let rec subs rigid k t = - match kind_of_term t with - | Rel i -> - if i<=k then t - else - (* Flex/Rel problem: unifiable as a pattern iff Rel in ev scope *) - (try List.assoc (mkRel (i-k)) subst - with Not_found -> if rigid then raise (NotClean t) else t) - | Evar (ev,args) -> - if Evd.is_defined_evar !evd (ev,args) then - subs rigid k (existential_value (evars_of !evd) (ev,args)) - else - (* Flex/Flex problem: restriction to a common scope *) - let args' = Array.map (subs false k) args in - if need_restriction k args' then - do_restrict_hyps (reset_context env) k evd ev args' - else - mkEvar (ev,args') - | Var id -> - (* Flex/Var problem: unifiable as a pattern iff Var in scope of ev *) - (try List.assoc t subst - with Not_found -> - if - not rigid - (* I don't understand this line: vars from evar_context evi - are private (especially some of them are freshly - generated in push_rel_context_to_named_context). They - have a priori nothing to do with the vars in env. I - remove the test [HH 25/8/06] - - or List.exists (fun (id',_,_) -> id=id') (evar_context evi) - *) - then t - else raise (NotClean t)) - - | _ -> - (* Flex/Rigid problem (or assimilated if not normal): we "imitate" *) - map_constr_with_binders succ (subs rigid) k t - in - let rhs = nf_evar (evars_of isevars) rhs in - let rhs = whd_beta rhs (* heuristic *) in - let body = - try subs true 0 rhs - with NotClean t -> - error_not_clean env (evars_of !evd) ev t (evar_source ev !evd) in - (!evd,body) - -(* [evar_define] solves the problem lhs = rhs when lhs is an uninstantiated - * evar, i.e. tries to find the body ?sp for lhs=mkEvar (sp,args) - * ?sp [ sp.hyps \ args ] unifies to rhs - * ?sp must be a closed term, not referring to itself. - * Not so trivial because some terms of args may be terms that are not - * variables. In this case, the non-var-or-Rels arguments are replaced - * by . [clean_rhs] will ignore this part of the subtitution. - * This leads to incompleteness (we don't deal with pbs that require - * inference of dependent types), but it seems sensible. +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 + 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 + with NotUnique -> + raise (CannotProject proj1) + +let solve_evar_evar f env evd ev1 ev2 = + try solve_evar_evar_l2r f env evd ev1 ev2 + with CannotProject projs1 -> + try solve_evar_evar_l2r f env evd ev2 ev1 + with CannotProject projs2 -> + postpone_evar_evar env evd projs1 ev1 projs2 ev2 + +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 + +(* 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) * - * If after cleaning, some free vars still occur, the function [restrict_hyps] - * tries to narrow the env of the evars that depend on Rels. + * 1) Let "env |- ?ev[hyps:=args] = rhs" be the unification problem + * 2) We limit it to a patterns unification problem "env |- ev[subst] = rhs" + * where only Rel's and Var's are relevant in subst + * 3) We recur on rhs, "imitating" the term, and failing if some Rel/Var is + * not in the scope of ?ev. For instance, the problem + * "y:nat |- ?x[] = y" where "|- ?1:nat" is not satisfiable because + * ?1 would be instantiated by y which is not in the scope of ?1. + * 4) We try to "project" the term if the process of imitation fails + * and that only one projection is possible * - * If after that free Rels still occur it means that the instantiation - * cannot be done, as in [x:?1; y:nat; z:(le y y)] x=z - * ?1 would be instantiated by (le y y) but y is not in the scope of ?1 + * 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] + * Precondition: Σ; Γ, y1..yk |- c /\ Σ; Γ |- u1..un + * Postcondition: if φ(x1..xn) is returned then + * Σ; Γ, y1..yk |- φ(u1..un) = c /\ x1..xn |- φ(x1..xn) *) -(* env needed for error messages... *) -let evar_define env (ev,argsv) rhs isevars = - if occur_evar ev rhs - then error_occur_check env (evars_of isevars) ev rhs; - let args = Array.to_list argsv in - let evi = Evd.find (evars_of isevars) ev in - (* the bindings to invert *) - let worklist = make_subst (evar_env evi) args in - let (isevars',body) = real_clean env isevars ev evi worklist rhs in - if occur_meta body then error "Meta cannot occur in evar body" - else +exception NotInvertibleUsingOurAlgorithm of constr +exception NotEnoughInformationToProgress + +let rec invert_definition env evd (evk,argsv as ev) rhs = + 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 + + (* Projection *) + let project_variable t = + (* Evar/Var problem: unifiable iff variable projectable from ev subst *) + try + let sols = find_projectable_vars env (evars_of !evdref) t subst in + let c, p = filter_solution sols in + let ty = lazy (Retyping.get_type_of env (evars_of !evdref) t) in + let evd = do_projection_effects evar_define env ty !evdref p in + evdref := evd; + c + with + | Not_found -> raise (NotInvertibleUsingOurAlgorithm t) + | NotUnique -> + if not !progress then raise NotEnoughInformationToProgress; + (* No unique projection but still restrict to where it is possible *) + let filter = array_map_to_list (fun c -> isEvar c or c = t) argsv in + let args' = filter_along (fun x -> x) filter argsv in + let evd,evar = do_restrict_hyps_virtual !evdref evk filter in + let evk',_ = destEvar evar in + let pb = (Reduction.CONV,env,mkEvar(evk',args'),t) in + evdref := Evd.add_conv_pb pb evd; + evar in + + let rec imitate (env',k as envk) t = + let t = whd_evar (evars_of !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; + (* 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' + in + (try + (* Try to project (a restriction of) the right evar *) + let eprojs' = effective_projections projs' in + let evd,args' = + list_fold_map (instance_of_projection evar_define env' t) + !evdref eprojs' in + let evd,evk' = do_restrict_hyps evd evk' projs' in + evdref := evd; + mkEvar (evk',Array.of_list args') + with NotUnique -> + assert !progress; + (* Make the virtual left evar real *) + let (evar'',ev'') = extend_evar env' evdref k ev t in + let evd = + (* Try to project (a restriction of) the left evar ... *) + try solve_evar_evar_l2r evar_define env' !evdref ev'' ev' + with CannotProject projs'' -> + (* ... or postpone the problem *) + postpone_evar_evar env' !evdref projs'' ev'' projs' ev' in + evdref := evd; + evar'') + | _ -> + progress := true; + (* Evar/Rigid problem (or assimilated if not normal): we "imitate" *) + map_constr_with_full_binders (fun d (env,k) -> push_rel d env, k+1) + imitate envk t in + + let rhs = whd_beta 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 + * an (uninstantiated) evar such that "hyps |- ?ev : typ". Otherwise said, + * [evar_define] tries to find an instance lhs such that + * "lhs [hyps:=args]" unifies to rhs. The term "lhs" must be closed in + * context "hyps" and not referring to itself. + *) + +and evar_define env (evk,_ as ev) rhs evd = + try + let (evd',body) = invert_definition 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; (* needed only if an inferred type *) let body = refresh_universes body in (* Cannot strictly type instantiations since the unification algorithm - * does not unifies applications from left to right. + * does not unify applications from left to right. * e.g problem f x == g y yields x==y and f==g (in that order) * Another problem is that type variables are evars of type Type let _ = try let env = evar_env evi in let ty = evi.evar_concl in - Typing.check env (evars_of isevars') body ty + Typing.check env (evars_of evd') body ty with e -> pperrnl (str "Ill-typed evar instantiation: " ++ fnl() ++ - pr_evar_defs isevars' ++ fnl() ++ + pr_evar_defs evd' ++ fnl() ++ str "----> " ++ int ev ++ str " := " ++ print_constr body); raise e in*) - let isevars'' = Evd.evar_define ev body isevars' in - isevars'',[ev] - - + Evd.evar_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) (*-------------------*) (* Auxiliary functions for the conversion algorithms modulo evars *) -let has_undefined_evars isevars t = - try let _ = local_strong (whd_ise (evars_of isevars)) t in false +let has_undefined_evars evd t = + try let _ = local_strong (whd_ise (evars_of evd)) t in false with Uninstantiated_evar _ -> true -let is_ground_term isevars t = - not (has_undefined_evars isevars t) - -let head_is_evar isevars = - let rec hrec k = match kind_of_term k with - | Evar n -> not (Evd.is_defined_evar isevars n) - | App (f,_) -> hrec f - | Cast (c,_,_) -> hrec c - | _ -> false - in - hrec - -let rec is_eliminator c = match kind_of_term c with - | App _ -> true - | Case _ -> true - | Cast (c,_,_) -> is_eliminator c - | _ -> false - -let head_is_embedded_evar isevars c = - (head_is_evar isevars c) & (is_eliminator c) +let is_ground_term evd t = + not (has_undefined_evars evd t) let head_evar = let rec hrec c = match kind_of_term c with - | Evar (ev,_) -> ev + | Evar (evk,_) -> evk | Case (_,_,c,_) -> hrec c - | App (c,_) -> hrec c - | Cast (c,_,_) -> hrec c - | _ -> failwith "headconstant" + | App (c,_) -> hrec c + | Cast (c,_,_) -> hrec c + | _ -> failwith "headconstant" in hrec @@ -576,20 +937,22 @@ let head_evar = that we don't care whether args itself contains Rel's or even Rel's distinct from the ones in l *) -let is_unification_pattern_evar (_,args) l = +let is_unification_pattern_evar env (_,args) l = let l' = Array.to_list args @ l in + let l' = List.map (expand_var env) l' in List.for_all (fun a -> isRel a or isVar a) l' & list_distinct l' -let is_unification_pattern f l = +let is_unification_pattern env f l = match kind_of_term f with | Meta _ -> array_for_all isRel l & array_distinct l - | Evar ev -> is_unification_pattern_evar ev (Array.to_list l) + | Evar ev -> is_unification_pattern_evar env ev (Array.to_list l) | _ -> false (* From a unification problem "?X l1 = term1 l2" such that l1 is made of distinct rel's, build "\x1...xn.(term1 l2)" (patterns unification) *) let solve_pattern_eqn env l1 c = + let l1 = List.map (expand_var 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 @@ -598,6 +961,8 @@ let solve_pattern_eqn env l1 c = | Var id -> let (id,_,t) = lookup_named id env in mkNamedLambda id t c' | _ -> assert false) l1 c in + (* Warning: we may miss some opportunity to eta-reduce more since c' + is not in normal form *) whd_eta c' (* This code (i.e. solve_pb, etc.) takes a unification @@ -637,28 +1002,19 @@ let status_changed lev (pbty,_,t1,t2) = * 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 isevars ev argsv1 argsv2 = - if argsv1 = argsv2 then (isevars,[]) else - let evd = Evd.find (evars_of isevars) ev in - let hyps = evar_context evd in - let (isevars',_,rsign) = - array_fold_left2 - (fun (isevars,sgn,rsgn) a1 a2 -> - let (isevars',b) = conv_algo env isevars Reduction.CONV a1 a2 in - if b then - (isevars',List.tl sgn, add_named_decl (List.hd sgn) rsgn) - else - (isevars,List.tl sgn, rsgn)) - (isevars,hyps,[]) argsv1 argsv2 - in - let nsign = List.rev rsign in - let (evd',newev) = - let env = - Sign.fold_named_context push_named nsign ~init:(reset_context env) in - new_evar isevars env ~src:(evar_source ev isevars) evd.evar_concl in - let evd'' = Evd.evar_define ev newev evd' in - evd'', [ev] - +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 (* Tries to solve problem t1 = t2. * Precondition: t1 is an uninstantiated evar @@ -666,45 +1022,57 @@ let solve_refl conv_algo env isevars ev argsv1 argsv2 = * if the problem couldn't be solved. *) (* Rq: uncomplete algorithm if pbty = CONV_X_LEQ ! *) -let solve_simple_eqn conv_algo env isevars (pbty,(n1,args1 as ev1),t2) = +let solve_simple_eqn conv_algo env evd (pbty,(evk1,args1 as ev1),t2) = try - let t2 = nf_evar (evars_of isevars) t2 in - let (isevars,lsp) = match kind_of_term t2 with - | Evar (n2,args2 as ev2) -> - if n1 = n2 then - solve_refl conv_algo env isevars n1 args1 args2 + let t2 = whd_evar (evars_of evd) t2 in + let evd = match kind_of_term t2 with + | Evar (evk2,args2 as ev2) -> + if evk1 = evk2 then + solve_refl conv_algo env evd evk1 args1 args2 else - (try evar_define env ev1 t2 isevars - with e when precatchable_exception e -> - evar_define env ev2 (mkEvar ev1) isevars) -(* if Array.length args1 < Array.length args2 then - evar_define env ev2 (mkEvar ev1) isevars - else - evar_define env ev1 t2 isevars*) + if pbty = Reduction.CONV + then solve_evar_evar evar_define env evd ev1 ev2 + else add_conv_pb (pbty,env,mkEvar ev1,t2) evd | _ -> - evar_define env ev1 t2 isevars in - let (isevars,pbs) = get_conv_pbs isevars (status_changed lsp) in + evar_define env ev1 t2 evd in + let (evd,pbs) = extract_changed_conv_pbs evd status_changed in List.fold_left - (fun (isevars,b as p) (pbty,env,t1,t2) -> - if b then conv_algo env isevars pbty t1 t2 else p) (isevars,true) + (fun (evd,b as p) (pbty,env,t1,t2) -> + if b then conv_algo env evd pbty t1 t2 else p) (evd,true) pbs with e when precatchable_exception e -> - (isevars,false) + (evd,false) (* [check_evars] fails if some unresolved evar remains *) (* it assumes that the defined existentials have already been substituted *) -let check_evars env initial_sigma isevars c = - let sigma = evars_of isevars in +let check_evars env initial_sigma evd c = + let sigma = evars_of evd in let c = nf_evar sigma c in let rec proc_rec c = match kind_of_term c with - | Evar (ev,args) -> - assert (Evd.mem sigma ev); - if not (Evd.mem initial_sigma ev) then - let (loc,k) = evar_source ev isevars in - error_unsolvable_implicit loc env sigma k + | 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 + let explain = + let f (_,_,t1,t2) = + (try head_evar t1 = evk with Failure _ -> false) + or (try head_evar t2 = evk with Failure _ -> false) in + let check_several c inst = + let _,argsv = destEvar c in + let l = List.filter (eq_constr inst) (Array.to_list argsv) in + let n = List.length l in + (* Maybe should we select one instead of failing ... *) + if n >= 2 then Some (SeveralInstancesFound n) else None in + match List.filter f (snd (extract_all_conv_pbs evd)) with + | (_,_,t1,t2)::_ -> + if isEvar t2 then check_several t2 t1 + else check_several t1 t2 + | [] -> None + in error_unsolvable_implicit loc env sigma evi k explain | _ -> iter_constr proc_rec c in proc_rec c @@ -754,14 +1122,15 @@ let mk_valcon c = Some c 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 evenv = evar_env evi in - let (evd1,dom) = new_evar evd evenv (new_Type()) in + let evenv = evar_unfiltered_env evi in + let (evd1,dom) = new_evar evd evenv (new_Type()) ~filter:(evar_filter evi) in let nvar = next_ident_away (id_of_string "x") (ids_of_named_context (evar_context evi)) in let newenv = push_named (nvar, None, dom) evenv in let (evd2,rng) = - new_evar evd1 newenv ~src:(evar_source ev evd1) (new_Type()) in + 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 evdom = fst (destEvar dom), args in @@ -770,15 +1139,15 @@ let define_evar_as_abstraction abs evd (ev,args) = let prod' = abs (Name nvar, mkEvar evdom, mkEvar evrng) in (evd3,prod') -let define_evar_as_arrow evd (ev,args) = +let define_evar_as_product evd (ev,args) = define_evar_as_abstraction (fun t -> mkProd t) evd (ev,args) let define_evar_as_lambda evd (ev,args) = define_evar_as_abstraction (fun t -> mkLambda t) evd (ev,args) -let define_evar_as_sort isevars (ev,args) = +let define_evar_as_sort evd (ev,args) = let s = new_Type () in - Evd.evar_define ev s isevars, destSort s + Evd.evar_define ev s evd, destSort s (* We don't try to guess in which sort the type should be defined, since @@ -791,33 +1160,33 @@ 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 isevars tycon = +let split_tycon loc env evd tycon = let rec real_split c = - let sigma = evars_of isevars in + let sigma = evars_of evd in let t = whd_betadeltaiota env sigma c in match kind_of_term t with - | Prod (na,dom,rng) -> isevars, (na, dom, rng) - | Evar ev when not (Evd.is_defined_evar isevars ev) -> - let (isevars',prod) = define_evar_as_arrow isevars ev in + | 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 - isevars',(Anonymous, dom, rng) + evd',(Anonymous, dom, rng) | _ -> error_not_product_loc loc env sigma c in match tycon with - | None -> isevars,(Anonymous,None,None) + | None -> evd,(Anonymous,None,None) | Some (abs, c) -> (match abs with None -> - let isevars', (n, dom, rng) = real_split c in - isevars', (n, mk_tycon dom, mk_tycon rng) + let evd', (n, dom, rng) = real_split c in + evd', (n, mk_tycon dom, mk_tycon rng) | Some (init, cur) -> if cur = 0 then - let isevars', (x, dom, rng) = real_split c in - isevars, (Anonymous, + let evd', (x, dom, rng) = real_split c in + evd, (Anonymous, Some (Some (init, 0), dom), Some (Some (init, 0), rng)) else - isevars, (Anonymous, None, Some (Some (init, pred cur), c))) + evd, (Anonymous, None, Some (Some (init, pred cur), c))) let valcon_of_tycon x = match x with @@ -833,7 +1202,7 @@ let lift_abstr_tycon_type n (abs, t) = else (Some (init, abs'), t) let lift_tycon_type n (abs, t) = (abs, lift n t) -let lift_tycon n = option_map (lift_tycon_type n) +let lift_tycon n = Option.map (lift_tycon_type n) let pr_tycon_type env (abs, t) = match abs with diff --git a/pretyping/evarutil.mli b/pretyping/evarutil.mli index 896bf26c..c915d4b0 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 9573 2007-01-31 20:18:18Z notin $ i*) +(*i $Id: evarutil.mli 11136 2008-06-18 10:41:34Z herbelin $ i*) (*i*) open Util @@ -34,10 +34,10 @@ 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 -> types -> evar_defs * constr + evar_defs -> env -> ?src:loc * hole_kind -> ?filter:bool list -> types -> evar_defs * constr (* the same with side-effects *) val e_new_evar : - evar_defs ref -> env -> ?src:loc * hole_kind -> types -> constr + evar_defs 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 @@ -46,16 +46,18 @@ val e_new_evar : of [inst] are typed in the occurrence context and their type (seen as a telescope) is [sign] *) val new_evar_instance : - named_context_val -> evar_defs -> types -> ?src:loc * hole_kind -> - constr list -> evar_defs * constr + named_context_val -> evar_defs -> types -> ?src:loc * hole_kind -> ?filter:bool list -> constr list -> evar_defs * constr -(***********************************************************) -(* Instanciate evars *) +val make_pure_subst : evar_info -> constr array -> (identifier * constr) list -(* suspicious env ? *) -val evar_define : - env -> existential -> constr -> evar_defs -> evar_defs * evar list +(***********************************************************) +(* Instantiate evars *) +(* [evar_define env ev c] try to instantiate [ev] with [c] (typed in [env]), + possibly solving related unification problems, possibly leaving open + some problems that cannot be solved in a unique way; + failed if the instance is not valid for the given [ev] *) +val evar_define : env -> existential -> constr -> evar_defs -> evar_defs (***********************************************************) (* Evars/Metas switching... *) @@ -71,8 +73,10 @@ val non_instantiated : evar_map -> (evar * evar_info) list (* Unification utils *) val is_ground_term : evar_defs -> constr -> bool -val is_eliminator : constr -> bool -val head_is_embedded_evar : evar_defs -> constr -> 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 val solve_simple_eqn : (env -> evar_defs -> conv_pb -> constr -> constr -> evar_defs * bool) -> env -> evar_defs -> conv_pb * existential * constr -> @@ -82,12 +86,12 @@ val solve_simple_eqn : new unresolved evar remains in [c] *) val check_evars : env -> evar_map -> evar_defs -> constr -> unit -val define_evar_as_arrow : evar_defs -> existential -> evar_defs * types +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 is_unification_pattern_evar : existential -> constr list -> bool -val is_unification_pattern : constr -> constr array -> bool +val is_unification_pattern_evar : env -> existential -> constr list -> bool +val is_unification_pattern : env -> constr -> constr array -> bool val solve_pattern_eqn : env -> constr list -> constr -> constr (***********************************************************) @@ -136,6 +140,10 @@ val tj_nf_evar : val nf_evar_info : evar_map -> evar_info -> evar_info val nf_evars : evar_map -> evar_map +val nf_named_context_evar : evar_map -> named_context -> named_context +val nf_rel_context_evar : evar_map -> rel_context -> rel_context +val nf_env_evar : evar_map -> env -> env + (* Same for evar defs *) val nf_isevar : evar_defs -> constr -> constr val j_nf_isevar : evar_defs -> unsafe_judgment -> unsafe_judgment @@ -153,6 +161,10 @@ exception Uninstantiated_evar of existential_key val whd_ise : evar_map -> constr -> constr val whd_castappevar : 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 + (*********************************************************************) (* debug pretty-printer: *) @@ -162,4 +174,5 @@ val pr_tycon : env -> type_constraint -> Pp.std_ppcmds (**********************************) (* Removing hyps in evars'context *) -val clear_hyps_in_evi : evar_defs ref -> evar_info -> identifier list -> evar_info +val clear_hyps_in_evi : evar_defs ref -> named_context_val -> types -> + identifier list -> named_context_val * types diff --git a/pretyping/evd.ml b/pretyping/evd.ml index 69d4352f..76a5ff26 100644 --- a/pretyping/evd.ml +++ b/pretyping/evd.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: evd.ml 9976 2007-07-12 11:58:30Z msozeau $ *) +(* $Id: evd.ml 11166 2008-06-22 13:23:35Z herbelin $ *) open Pp open Util @@ -31,9 +31,28 @@ type evar_info = { evar_concl : constr; evar_hyps : named_context_val; evar_body : evar_body; + evar_filter : bool list; evar_extra : Dyn.t option} +let make_evar hyps ccl = { + evar_concl = ccl; + evar_hyps = hyps; + evar_body = Evar_empty; + evar_filter = List.map (fun _ -> true) (named_context_of_val hyps); + evar_extra = None +} + +let evar_concl evi = evi.evar_concl +let evar_hyps evi = evi.evar_hyps let evar_context evi = named_context_of_val evi.evar_hyps +let evar_body evi = evi.evar_body +let evar_filter evi = evi.evar_filter +let evar_unfiltered_env evi = Global.env_of_context evi.evar_hyps +let evar_filtered_context evi = + snd (list_filter2 (fun b c -> b) (evar_filter evi,evar_context 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 || @@ -49,42 +68,37 @@ let empty = Evarmap.empty let to_list evc = (* Workaround for change in Map.fold behavior *) let l = ref [] in - Evarmap.iter (fun ev x -> l:=(ev,x)::!l) evc; - !l + Evarmap.iter (fun evk x -> l := (evk,x)::!l) evc; + !l -let dom evc = Evarmap.fold (fun ev _ acc -> ev::acc) evc [] +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 ev newinfo = Evarmap.add ev newinfo evd +let add evd evk newinfo = Evarmap.add evk newinfo evd -let define evd ev body = +let define evd evk body = let oldinfo = - try find evd ev + 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 ev newinfo evd - | _ -> anomaly "Evd.define: cannot define an isevar twice" - -let is_evar sigma ev = mem sigma ev + | Evar_empty -> Evarmap.add evk newinfo evd + | _ -> anomaly "Evd.define: cannot define an evar twice" -let is_defined sigma ev = - let info = find sigma ev in - not (info.evar_body = Evar_empty) +let is_evar sigma evk = mem sigma evk -let evar_concl ev = ev.evar_concl -let evar_hyps ev = ev.evar_hyps -let evar_body ev = ev.evar_body -let evar_env evd = Global.env_of_context evd.evar_hyps +let is_defined sigma evk = + let info = find sigma evk in + not (info.evar_body = Evar_empty) -let string_of_existential ev = "?" ^ string_of_int ev +let string_of_existential evk = "?" ^ string_of_int evk -let existential_of_int ev = ev +let existential_of_int evk = evk (*******************************************************************) (* Formerly Instantiate module *) @@ -120,14 +134,14 @@ let existential_type sigma (n,args) = try find sigma n with Not_found -> anomaly ("Evar "^(string_of_existential n)^" was not declared") in - let hyps = evar_context info in + let hyps = evar_filtered_context info in instantiate_evar hyps info.evar_concl (Array.to_list args) exception NotInstantiatedEvar let existential_value sigma (n,args) = let info = find sigma n in - let hyps = evar_context info in + let hyps = evar_filtered_context info in match evar_body info with | Evar_defined c -> instantiate_evar hyps c (Array.to_list args) @@ -287,6 +301,8 @@ let existential_value (sigma,_) = existential_value sigma let existential_type (sigma,_) = existential_type sigma let existential_opt_value (sigma,_) = existential_opt_value sigma +let merge e e' = fold (fun n v sigma -> add sigma n v) e' e + (*******************************************************************) type open_constr = evar_map * constr @@ -325,16 +341,45 @@ let mk_freelisted c = let map_fl f cfl = { cfl with rebus=f cfl.rebus } +(* Status of an instance found by unification wrt to the meta it solves: + - a supertype of the meta (e.g. the solution to ?X <= T is a supertype of ?X) + - a subtype of the meta (e.g. the solution to T <= ?X is a supertype of ?X) + - a term that can be eta-expanded n times while still being a solution + (e.g. the solution [P] to [?X u v = P u v] can be eta-expanded twice) +*) + +type instance_constraint = + IsSuperType | IsSubType | ConvUpToEta of int | UserGiven + +(* Status of the unification of the type of an instance against the type of + the meta it instantiates: + - CoerceToType means that the unification of types has not been done + and that a coercion can still be inserted: the meta should not be + substituted freely (this happens for instance given via the + "with" binding clause). + - TypeProcessed means that the information obtainable from the + unification of types has been extracted. + - TypeNotProcessed means that the unification of types has not been + done but it is known that no coercion may be inserted: the meta + can be substituted freely. +*) + +type instance_typing_status = + CoerceToType | TypeNotProcessed | TypeProcessed + +(* Status of an instance together with the status of its type unification *) + +type instance_status = instance_constraint * instance_typing_status (* Clausal environments *) type clbinding = | Cltyp of name * constr freelisted - | Clval of name * constr freelisted * constr freelisted + | Clval of name * (constr freelisted * instance_status) * constr freelisted let map_clb f = function | Cltyp (na,cfl) -> Cltyp (na,map_fl f cfl) - | Clval (na,cfl1,cfl2) -> Clval (na,map_fl f cfl1,map_fl f cfl2) + | Clval (na,(cfl1,pb),cfl2) -> Clval (na,(map_fl f cfl1,pb),map_fl f cfl2) (* name of defined is erased (but it is pretty-printed) *) let clb_name = function @@ -362,12 +407,15 @@ type hole_kind = | 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; conv_pbs : evar_constraint list; + last_mods : existential_key list; history : (existential_key * (loc * hole_kind)) list; metas : clbinding Metamap.t } @@ -378,46 +426,61 @@ let subst_evar_defs_light sub evd = metas = Metamap.map (map_clb (subst_mps sub)) evd.metas } let create_evar_defs sigma = - { evars=sigma; conv_pbs=[]; history=[]; metas=Metamap.empty } + { 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} let add_conv_pb pb d = {d with conv_pbs = pb::d.conv_pbs} -let evar_source ev d = - try List.assoc ev d.history +let evar_source evk d = + try List.assoc evk d.history with Not_found -> (dummy_loc, InternalHole) (* define the existential of section path sp as the constr body *) -let evar_define sp body isevars = - {isevars with evars = define isevars.evars sp body} - -let evar_declare hyps evn ty ?(src=(dummy_loc,InternalHole)) evd = +let evar_define evk body evd = { evd with - evars = add evd.evars evn - {evar_hyps=hyps; - evar_concl=ty; - evar_body=Evar_empty; - evar_extra=None}; - history = (evn,src)::evd.history } + evars = define evd.evars evk body; + last_mods = evk :: evd.last_mods } + +let evar_declare hyps evk ty ?(src=(dummy_loc,InternalHole)) ?filter evd = + let filter = + if filter = None then + List.map (fun _ -> true) (named_context_of_val hyps) + else + (let filter = Option.get filter in + assert (List.length filter = List.length (named_context_of_val hyps)); + filter) + in + { evd with + evars = 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 } -let is_defined_evar isevars (n,_) = is_defined isevars.evars n +let is_defined_evar evd (evk,_) = is_defined evd.evars evk (* Does k corresponds to an (un)defined existential ? *) -let is_undefined_evar isevars c = match kind_of_term c with - | Evar ev -> not (is_defined_evar isevars ev) +let is_undefined_evar evd c = match kind_of_term c with + | Evar ev -> not (is_defined_evar evd ev) | _ -> false -let undefined_evars isevars = - let evd = - fold (fun ev evi sigma -> if evi.evar_body = Evar_empty then - add sigma ev evi else sigma) - isevars.evars empty +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 - { isevars with evars = evd } + { evd with evars = evars } (* extracts conversion problems that satisfy predicate p *) (* Note: conv_pbs not satisying p are stored back in reverse order *) -let get_conv_pbs isevars p = +let extract_conv_pbs evd p = let (pbs,pbs1) = List.fold_left (fun (pbs,pbs1) pb -> @@ -426,11 +489,19 @@ let get_conv_pbs isevars p = else (pbs,pb::pbs1)) ([],[]) - isevars.conv_pbs + evd.conv_pbs in - {isevars with conv_pbs = pbs1}, + {evd with conv_pbs = pbs1; last_mods = []}, pbs +let extract_changed_conv_pbs evd p = + extract_conv_pbs evd (p evd.last_mods) + +let extract_all_conv_pbs evd = + extract_conv_pbs evd (fun _ -> true) + +let evar_merge evd evars = + { evd with evars = merge evd.evars evars } (**********************************************************) (* Sort variables *) @@ -452,11 +523,35 @@ let pr_sort_constraints (_,sm) = pr_sort_cstrs sm let meta_list evd = metamap_to_list evd.metas +let undefined_metas evd = + List.sort Pervasives.compare (map_succeed (function + | (n,Clval(_,_,typ)) -> failwith "" + | (n,Cltyp (_,typ)) -> n) + (meta_list evd)) + +let metas_of evd = + List.map (function + | (n,Clval(_,_,typ)) -> (n,typ.rebus) + | (n,Cltyp (_,typ)) -> (n,typ.rebus)) + (meta_list evd) + +let map_metas_fvalue f evd = + { evd with metas = + Metamap.map + (function + | Clval(id,(c,s),typ) -> Clval(id,(mk_freelisted (f c.rebus),s),typ) + | x -> x) evd.metas } + +let meta_opt_fvalue evd mv = + match Metamap.find mv evd.metas with + | Clval(_,b,_) -> Some b + | Cltyp _ -> None + let meta_defined evd mv = match Metamap.find mv evd.metas with | Clval _ -> true | Cltyp _ -> false - + let meta_fvalue evd mv = match Metamap.find mv evd.metas with | Clval(_,b,_) -> b @@ -470,12 +565,19 @@ let meta_ftype evd mv = 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 evd = +let meta_assign mv (v,pb) evd = + match Metamap.find mv evd.metas with + | Cltyp(na,ty) -> + { evd with + metas = Metamap.add mv (Clval(na,(mk_freelisted v,pb),ty)) evd.metas } + | _ -> anomaly "meta_assign: already defined" + +let meta_reassign mv (v,pb) evd = match Metamap.find mv evd.metas with - Cltyp(na,ty) -> - { evd with - metas = Metamap.add mv (Clval(na,mk_freelisted v, ty)) evd.metas } - | _ -> anomaly "meta_assign: already defined" + | Clval(na,_,ty) -> + { evd with + metas = Metamap.add mv (Clval(na,(mk_freelisted v,pb),ty)) evd.metas } + | _ -> anomaly "meta_reassign: not yet defined" (* If the meta is defined then forget its name *) let meta_name evd mv = @@ -510,10 +612,51 @@ let meta_merge evd1 evd2 = 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) -> + match v with + | Clval (na,(b,(UserGiven,CoerceToType as s)),typ) -> + (n,b.rebus,s)::mc, Metamap.add n (Cltyp (na,typ)) ml + | v -> + mc, Metamap.add n v ml) + evd.metas ([],Metamap.empty) in + mc, { evd with metas = ml } + +let rec list_assoc_in_triple x = function + [] -> raise Not_found + | (a,b,_)::l -> if compare a x = 0 then b else list_assoc_in_triple x l + +let subst_defined_metas bl c = + let 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 + +(**********************************************************) +(* Failure explanation *) + +type unsolvability_explanation = SeveralInstancesFound of int (**********************************************************) (* Pretty-printing *) +let pr_instance_status (sc,typ) = + begin match sc with + | IsSubType -> str " [or a subtype of it]" + | IsSuperType -> str " [or a supertype of it]" + | ConvUpToEta 0 -> mt () + | UserGiven -> mt () + | ConvUpToEta n -> str" [or an eta-expansion up to " ++ int n ++ str" of it]" + end ++ + begin match typ with + | CoerceToType -> str " [up to coercion]" + | TypeNotProcessed -> mt () + | TypeProcessed -> str " [type is checked]" + end + let pr_meta_map mmap = let pr_name = function Name id -> str"[" ++ pr_id id ++ str"]" @@ -523,17 +666,22 @@ let pr_meta_map mmap = hov 0 (pr_meta mv ++ pr_name na ++ str " : " ++ print_constr b.rebus ++ fnl ()) - | (mv,Clval(na,b,_)) -> + | (mv,Clval(na,(b,s),_)) -> hov 0 (pr_meta mv ++ pr_name na ++ str " := " ++ - print_constr b.rebus ++ fnl ()) + print_constr b.rebus ++ spc () ++ pr_instance_status s ++ fnl ()) in prlist pr_meta_binding (metamap_to_list mmap) -let pr_idl idl = prlist_with_sep pr_spc pr_id idl +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 ":=" ++ + print_constr c ++ str (if ok then ")" else "}") let pr_evar_info evi = - let phyps = pr_idl (List.rev (ids_of_named_context (evar_context evi))) in + let decls = List.combine (evar_context evi) (evar_filter evi) in + let phyps = prlist_with_sep pr_spc pr_decl (List.rev decls) in let pty = print_constr evi.evar_concl in let pb = match evi.evar_body with @@ -568,3 +716,6 @@ let pr_evar_defs evd = 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 = + str "[" ++ prlist_with_sep spc pr_meta (Metaset.elements metas) ++ str "]" diff --git a/pretyping/evd.mli b/pretyping/evd.mli index ef6a3d7b..1acc811b 100644 --- a/pretyping/evd.mli +++ b/pretyping/evd.mli @@ -1,3 +1,4 @@ + (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) (* 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 empty : evar_map @@ -50,17 +64,14 @@ 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 evar_concl : evar_info -> constr -val evar_hyps : evar_info -> Environ.named_context_val -val evar_body : evar_info -> evar_body -val evar_env : evar_info -> Environ.env - val string_of_existential : evar -> string val existential_of_int : int -> evar @@ -101,9 +112,41 @@ val metavars_of : constr -> Metaset.t val mk_freelisted : constr -> constr freelisted val map_fl : ('a -> 'b) -> 'a freelisted -> 'b freelisted +(* Status of an instance found by unification wrt to the meta it solves: + - a supertype of the meta (e.g. the solution to ?X <= T is a supertype of ?X) + - a subtype of the meta (e.g. the solution to T <= ?X is a supertype of ?X) + - a term that can be eta-expanded n times while still being a solution + (e.g. the solution [P] to [?X u v = P u v] can be eta-expanded twice) +*) + +type instance_constraint = + IsSuperType | IsSubType | ConvUpToEta of int | UserGiven + +(* Status of the unification of the type of an instance against the type of + the meta it instantiates: + - CoerceToType means that the unification of types has not been done + and that a coercion can still be inserted: the meta should not be + substituted freely (this happens for instance given via the + "with" binding clause). + - TypeProcessed means that the information obtainable from the + unification of types has been extracted. + - TypeNotProcessed means that the unification of types has not been + done but it is known that no coercion may be inserted: the meta + can be substituted freely. +*) + +type instance_typing_status = + CoerceToType | TypeNotProcessed | TypeProcessed + +(* Status of an instance together with the status of its type unification *) + +type instance_status = instance_constraint * instance_typing_status + +(* Clausal environments *) + type clbinding = | Cltyp of name * constr freelisted - | Clval of name * constr freelisted * constr freelisted + | Clval of name * (constr freelisted * instance_status) * constr freelisted val map_clb : (constr -> constr) -> clbinding -> clbinding @@ -115,7 +158,9 @@ type evar_defs 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_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 @@ -127,38 +172,57 @@ type hole_kind = | CasesType | InternalHole | 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 val evar_declare : - Environ.named_context_val -> evar -> types -> ?src:loc * hole_kind -> - evar_defs -> evar_defs + 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 +(* [evar_merge evd evars] extends the evars of [evd] with [evars] *) +val evar_merge : evar_defs -> evar_map -> evar_defs + (* Unification constraints *) type conv_pb = Reduction.conv_pb -type evar_constraint = conv_pb * Environ.env * constr * constr +type evar_constraint = conv_pb * env * constr * constr val add_conv_pb : evar_constraint -> evar_defs -> evar_defs -val get_conv_pbs : evar_defs -> (evar_constraint -> bool) -> - evar_defs * evar_constraint list +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 + (* Metas *) val meta_list : evar_defs -> (metavariable * clbinding) list val meta_defined : evar_defs -> 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 +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 val meta_declare : metavariable -> types -> ?name:name -> evar_defs -> evar_defs -val meta_assign : metavariable -> constr -> 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 (* [meta_merge evd1 evd2] returns [evd2] extended with the metas of [evd1] *) val meta_merge : evar_defs -> evar_defs -> evar_defs +val undefined_metas : evar_defs -> metavariable list +val metas_of : evar_defs -> metamap +val map_metas_fvalue : (constr -> constr) -> evar_defs -> evar_defs + +type metabinding = metavariable * constr * instance_status + +val retract_coercible_metas : evar_defs -> metabinding list * evar_defs +val subst_defined_metas : metabinding list -> constr -> constr option + (**********************************************************) (* Sort variables *) @@ -168,6 +232,11 @@ 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 +(**********************************************************) +(* Failure explanation *) + +type unsolvability_explanation = SeveralInstancesFound of int + (*********************************************************************) (* debug pretty-printer: *) @@ -175,3 +244,4 @@ 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_sort_constraints : evar_map -> Pp.std_ppcmds +val pr_metaset : Metaset.t -> Pp.std_ppcmds diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml index eeddcb64..2325baec 100644 --- a/pretyping/indrec.ml +++ b/pretyping/indrec.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: indrec.ml 9519 2007-01-22 18:13:29Z notin $ *) +(* $Id: indrec.ml 10410 2007-12-31 13:11:55Z msozeau $ *) open Pp open Util @@ -31,7 +31,7 @@ open Sign type recursion_scheme_error = | NotAllowedCaseAnalysis of bool * sorts * inductive | BadInduction of bool * identifier * sorts - | NotMutualInScheme + | NotMutualInScheme of inductive * inductive exception RecursionSchemeError of recursion_scheme_error @@ -78,7 +78,7 @@ let mis_make_case_com depopt env sigma ind (mib,mip as specif) kind = let depind = build_dependent_inductive env indf' in let deparsign = (Anonymous,None,depind)::arsign in - let ci = make_default_case_info env RegularStyle ind in + let ci = make_case_info env ind RegularStyle in let pbody = appvect (mkRel (ndepar + nbprod), @@ -157,7 +157,7 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs = (match dest_recarg ra with | Mrec j when is_rec -> (depPvect.(j),rest) | Imbr _ -> - Options.if_verbose warning "Ignoring recursive call"; + Flags.if_verbose warning "Ignoring recursive call"; (None,rest) | _ -> (None, rest)) in @@ -275,19 +275,18 @@ let mis_make_indrec env sigma listdepkind mib = Array.create mib.mind_ntypes (None : (bool * constr) option) in let _ = let rec - assign k = function - | [] -> () - | (indi,mibi,mipi,dep,_)::rest -> - (Array.set depPvec (snd indi) (Some(dep,mkRel k)); - assign (k-1) rest) + assign k = function + | [] -> () + | (indi,mibi,mipi,dep,_)::rest -> + (Array.set depPvec (snd indi) (Some(dep,mkRel k)); + assign (k-1) rest) in - assign nrec listdepkind in + assign nrec listdepkind in let recargsvec = Array.map (fun mip -> mip.mind_recargs) mib.mind_packets in - (* recarg information for non recursive parameters *) + (* recarg information for non recursive parameters *) let rec recargparn l n = - if n = 0 then l else recargparn (mk_norec::l) (n-1) - in + if n = 0 then l else recargparn (mk_norec::l) (n-1) in let recargpar = recargparn [] (nparams-nparrec) in let make_one_rec p = let makefix nbconstruct = @@ -296,97 +295,97 @@ 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 *) + 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 *) + + (* 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 constrs = get_constructors env indf' in let fi = rel_vect (dect-i-nctyi) nctyi in let vecfi = Array.map - (fun f -> appvect (f,extended_rel_vect ndepar lnonparrec)) - fi + (fun f -> appvect (f,extended_rel_vect ndepar lnonparrec)) + fi in - array_map3 - (make_rec_branch_arg env sigma - (nparrec,depPvec,larsign)) - vecfi constrs (dest_subterms recargsvec.(tyi)) + array_map3 + (make_rec_branch_arg env sigma + (nparrec,depPvec,larsign)) + vecfi constrs (dest_subterms recargsvec.(tyi)) in - + let j = (match depPvec.(tyi) with - | Some (_,c) when isRel c -> destRel c - | _ -> assert false) + | Some (_,c) when isRel c -> destRel c + | _ -> assert false) in - - (* Predicate in the context of the case *) - + + (* 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 and nrar = if dep then extended_rel_list 0 deparsign' - else extended_rel_list 1 arsign' + else extended_rel_list 1 arsign' in nrpar@nrar - + in - (* body of i-th component of the mutual fixpoint *) + (* body of i-th component of the mutual fixpoint *) let deftyi = - let ci = make_default_case_info env RegularStyle indi in + let ci = make_case_info env indi RegularStyle in let concl = applist (mkRel (dect+j+ndepar),pargs) in let pred = it_mkLambda_or_LetIn_name env ((if dep then mkLambda_name env else mkLambda) - (Anonymous,depind',concl)) + (Anonymous,depind',concl)) arsign' in - it_mkLambda_or_LetIn_name env - (mkCase (ci, pred, - mkRel 1, - branches)) - (lift_rel_context nrec deparsign) + it_mkLambda_or_LetIn_name env + (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 pargs = if dep then extended_rel_vect 0 deparsign - else extended_rel_vect 1 arsign - in appvect (mkRel (nbconstruct+ndepar+nonrecpar+j),pargs) - in it_mkProd_or_LetIn_name env + 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) + in it_mkProd_or_LetIn_name env concl deparsign in - mrec (i+nctyi) (rel_context_nhyps arsign ::ln) (typtyi::ltyp) - (deftyi::ldef) rest + 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 names = Array.create nrec (Name(id_of_string "F")) in - mkFix ((fixn,p),(names,fixtyi,fixdef)) + mkFix ((fixn,p),(names,fixtyi,fixdef)) in - mrec 0 [] [] [] + mrec 0 [] [] [] in let rec make_branch env i = function | (indi,mibi,mipi,dep,_)::rest -> @@ -404,27 +403,28 @@ let mis_make_indrec env sigma listdepkind mib = type_rec_branch true dep env sigma (vargs,depPvec,i+j) tyi cs recarg in - mkLambda_string "f" p_0 - (onerec (push_rel (Anonymous,None,p_0) env) (j+1)) + mkLambda_string "f" p_0 + (onerec (push_rel (Anonymous,None,p_0) env) (j+1)) in onerec env 0 | [] -> makefix i listdepkind - in + in let rec put_arity env i = function | (indi,_,_,dep,kinds)::rest -> let indf = make_ind_family (indi,extended_rel_list i lnamesparrec) in let typP = make_arity env dep indf (new_sort_in_family kinds) in - mkLambda_string "P" typP - (put_arity (push_rel (Anonymous,None,typP) env) (i+1) rest) + mkLambda_string "P" typP + (put_arity (push_rel (Anonymous,None,typP) env) (i+1) rest) | [] -> make_branch env 0 listdepkind - in - - (* Body on make_one_rec *) + 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) + (List.map (fun (indi,_,_,_,_) -> snd indi) listdepkind) + mipi.mind_recargs) then let env' = push_rel_context lnamesparrec env in it_mkLambda_or_LetIn_name env (put_arity env' 0 listdepkind) @@ -432,15 +432,15 @@ let mis_make_indrec env sigma listdepkind mib = else mis_make_case_com (Some dep) env sigma indi (mibi,mipi) kind in - (* Body of mis_make_indrec *) - list_tabulate make_one_rec nrec + (* 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 + mis_make_case_com depopt 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 @@ -459,7 +459,7 @@ let change_sort_arity sort = | Sort _ -> mkSort sort | _ -> assert false in - drec + drec (* [npar] is the number of expected arguments (then excluding letin's) *) let instantiate_indrec_scheme sort = @@ -501,13 +501,13 @@ let instantiate_type_indrec_scheme sort npars term = let check_arities listdepkind = let _ = List.fold_left - (fun ln ((_,ni),mibi,mipi,dep,kind) -> + (fun ln ((_,ni as mind),mibi,mipi,dep,kind) -> let id = mipi.mind_typename in let kelim = elim_sorts (mibi,mipi) in if not (List.exists ((=) kind) kelim) then raise (RecursionSchemeError (BadInduction (dep,id,new_sort_in_family kind))) else if List.mem ni ln then raise - (RecursionSchemeError NotMutualInScheme) + (RecursionSchemeError (NotMutualInScheme (mind,mind))) else ni::ln) [] listdepkind in true @@ -520,11 +520,11 @@ let build_mutual_indrec env sigma = function (List.map (function (mind',mibi',mipi',dep',s') -> let (sp',_) = mind' in - if sp=sp' then + if sp=sp' then let (mibi',mipi') = lookup_mind_specif env mind' in - (mind',mibi',mipi',dep',s') - else - raise (RecursionSchemeError NotMutualInScheme)) + (mind',mibi',mipi',dep',s') + else + raise (RecursionSchemeError (NotMutualInScheme (mind,mind')))) lrecspec) in let _ = check_arities listdepkind in @@ -535,7 +535,7 @@ let build_indrec env sigma ind = 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) + List.hd (mis_make_indrec env sigma [(ind,mib,mip,dep,kind)] mib) (**********************************************************************) (* To handle old Case/Match syntax in Pretyping *) diff --git a/pretyping/indrec.mli b/pretyping/indrec.mli index e5eb07f5..6f177474 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 7660 2005-12-17 21:13:48Z herbelin $ i*) +(*i $Id: indrec.mli 9831 2007-05-17 18:55:42Z herbelin $ i*) (*i*) open Names @@ -22,7 +22,7 @@ open Evd type recursion_scheme_error = | NotAllowedCaseAnalysis of bool * sorts * inductive | BadInduction of bool * identifier * sorts - | NotMutualInScheme + | NotMutualInScheme of inductive * inductive exception RecursionSchemeError of recursion_scheme_error diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index 14136f61..0daff713 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: inductiveops.ml 9310 2006-10-28 19:35:09Z herbelin $ *) +(* $Id: inductiveops.ml 10114 2007-09-06 07:36:14Z herbelin $ *) open Util open Names @@ -131,21 +131,16 @@ let allowed_sorts env (kn,i as ind) = mip.mind_kelim (* Annotation for cases *) -let make_case_info env ind style pats_source = +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; - source = pats_source } in + let print_info = { ind_nargs = mip.mind_nrealargs; style = style } in { ci_ind = ind; ci_npar = mib.mind_nparams; ci_cstr_nargs = mip.mind_consnrealdecls; ci_pp_info = print_info } let make_default_case_info env style ind = - let (mib,mip) = Inductive.lookup_mind_specif env ind in make_case_info env ind style - (Array.map (fun _ -> RegularPat) mip.mind_consnames) (*s Useful functions *) @@ -398,30 +393,19 @@ let arity_of_case_predicate env (ind,params) dep k = (* A function which checks that a term well typed verifies both syntactic conditions *) -let control_only_guard env = - let rec control_rec c = match kind_of_term c with - | Rel _ | Var _ -> () - | Sort _ | Meta _ -> () - | Ind _ -> () - | Construct _ -> () - | Const _ -> () - | CoFix (_,(_,tys,bds) as cofix) -> - Inductive.check_cofix env cofix; - Array.iter control_rec tys; - Array.iter control_rec bds; - | Fix (_,(_,tys,bds) as fix) -> - Inductive.check_fix env fix; - Array.iter control_rec tys; - Array.iter control_rec bds; - | Case(_,p,c,b) -> control_rec p;control_rec c;Array.iter control_rec b - | Evar (_,cl) -> Array.iter control_rec cl - | App (c,cl) -> control_rec c; Array.iter control_rec cl - | Cast (c1,_, c2) -> control_rec c1; control_rec c2 - | Prod (_,c1,c2) -> control_rec c1; control_rec c2 - | Lambda (_,c1,c2) -> control_rec c1; control_rec c2 - | LetIn (_,c1,c2,c3) -> control_rec c1; control_rec c2; control_rec c3 - in - control_rec +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 + | Fix (_,(_,_,_) as fix) -> + Inductive.check_fix e fix + | _ -> () + in + 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 diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli index d49b64d9..9e370fec 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 9194 2006-10-01 09:25:19Z herbelin $ i*) +(*i $Id: inductiveops.mli 9707 2007-03-15 16:36:15Z herbelin $ i*) open Names open Term @@ -105,9 +105,11 @@ val arity_of_case_predicate : val type_case_branches_with_names : env -> inductive * constr list -> unsafe_judgment -> constr -> types array * types -val make_case_info : - env -> inductive -> case_style -> pattern_source array -> case_info +val make_case_info : env -> inductive -> case_style -> case_info + +(*i Compatibility val make_default_case_info : env -> case_style -> inductive -> case_info +i*) (********************) val control_only_guard : env -> types -> unit diff --git a/pretyping/matching.ml b/pretyping/matching.ml index 65ce2ef4..aaa4c3f0 100644 --- a/pretyping/matching.ml +++ b/pretyping/matching.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: matching.ml 9280 2006-10-25 21:37:37Z herbelin $ *) +(* $Id: matching.ml 10451 2008-01-18 17:20:28Z barras $ *) (*i*) open Util @@ -153,7 +153,7 @@ let matches_core convert allow_partial_app pat c = sorec ((na2,t2)::stk) (sorec stk sigma c1 c2) d1 d2 | PRef (ConstRef _ as ref), _ when convert <> None -> - let (env,evars) = out_some convert in + let (env,evars) = Option.get convert in let c = constr_of_global ref in if is_conv env evars c cT then sigma else raise PatternMatchingFailure @@ -161,7 +161,8 @@ let matches_core convert allow_partial_app pat c = | PIf (a1,b1,b1'), Case (ci,_,a2,[|b2;b2'|]) -> let ctx,b2 = decompose_lam_n_assum ci.ci_cstr_nargs.(0) b2 in let ctx',b2' = decompose_lam_n_assum ci.ci_cstr_nargs.(1) b2' in - let n = List.length ctx and n' = List.length ctx' in + let n = rel_context_length ctx in + let n' = rel_context_length ctx' in if noccur_between 1 n b2 & noccur_between 1 n' b2' then let s = List.fold_left (fun l (na,_,t) -> (na,t)::l) stk ctx in let s' = List.fold_left (fun l (na,_,t) -> (na,t)::l) stk ctx' in diff --git a/pretyping/pattern.ml b/pretyping/pattern.ml index 3060ee03..5df5c9fb 100644 --- a/pretyping/pattern.ml +++ b/pretyping/pattern.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: pattern.ml 9976 2007-07-12 11:58:30Z msozeau $ *) +(* $Id: pattern.ml 10785 2008-04-13 21:41:54Z herbelin $ *) open Util open Names @@ -92,7 +92,7 @@ let head_of_constr_reference c = match kind_of_term c with let rec pattern_of_constr t = match kind_of_term t with | Rel n -> PRel n - | Meta n -> PMeta (Some (id_of_string (string_of_int n))) + | Meta n -> PMeta (Some (id_of_string ("META" ^ string_of_int n))) | Var id -> PVar id | Sort (Prop c) -> PSort (RProp c) | Sort (Type _) -> PSort (RType None) @@ -185,7 +185,7 @@ let rec subst_pattern subst pat = match pat with if c' == c && c1' == c1 && c2' == c2 then pat else 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 ind' = Option.smartmap (Inductiveops.subst_inductive subst) ind in let typ' = subst_pattern subst typ in let c' = subst_pattern subst c in let branches' = array_smartmap (subst_pattern subst) branches in @@ -216,14 +216,14 @@ let rec pat_of_raw metas vars = function PRef r (* Hack pour ne pas réécrire une interprétation complète des patterns*) | RApp (_, RPatVar (_,(true,n)), cl) -> - PSoApp (n, List.map (pat_of_raw metas vars) cl) + metas := n::!metas; PSoApp (n, List.map (pat_of_raw metas vars) cl) | RApp (_,c,cl) -> PApp (pat_of_raw metas vars c, Array.of_list (List.map (pat_of_raw metas vars) cl)) - | RLambda (_,na,c1,c2) -> + | RLambda (_,na,bk,c1,c2) -> PLambda (na, pat_of_raw metas vars c1, pat_of_raw metas (na::vars) c2) - | RProd (_,na,c1,c2) -> + | RProd (_,na,bk,c1,c2) -> PProd (na, pat_of_raw metas vars c1, pat_of_raw metas (na::vars) c2) | RLetIn (_,na,c1,c2) -> @@ -234,18 +234,18 @@ let rec pat_of_raw metas vars = function | RHole _ -> PMeta None | RCast (_,c,_) -> - Options.if_verbose + Flags.if_verbose Pp.warning "Cast not taken into account in constr pattern"; pat_of_raw metas vars c | RIf (_,c,(_,None),b1,b2) -> PIf (pat_of_raw metas vars c, pat_of_raw metas vars b1,pat_of_raw metas vars b2) | RLetTuple (loc,nal,(_,None),b,c) -> - let mkRLambda c na = RLambda (loc,na,RHole (loc,Evd.InternalHole),c) in + let mkRLambda c na = RLambda (loc,na,Explicit,RHole (loc,Evd.InternalHole),c) in let c = List.fold_left mkRLambda c nal in PCase ((LetStyle,[|1|],None,None),PMeta None,pat_of_raw metas vars b, [|pat_of_raw metas vars c|]) - | RCases (loc,p,[c,(na,indnames)],brs) -> + | RCases (loc,sty,p,[c,(na,indnames)],brs) -> let pred,ind_nargs, ind = match p,indnames with | Some p, Some (_,ind,n,nal) -> rev_it_mkPLambda nal (mkPLambda na (pat_of_raw metas vars p)), @@ -259,7 +259,7 @@ let rec pat_of_raw metas vars = function Array.init (List.length brs) (pat_of_raw_branch loc metas vars ind brs) in let cstr_nargs,brs = (Array.map fst cbrs, Array.map snd cbrs) in - PCase ((RegularStyle,cstr_nargs,ind,ind_nargs), pred, + PCase ((sty,cstr_nargs,ind,ind_nargs), pred, pat_of_raw metas vars c, brs) | r -> diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml index 59cdad04..a513d558 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 9217 2006-10-05 17:31:23Z notin $ *) +(* $Id: pretype_errors.ml 10860 2008-04-27 21:39:08Z herbelin $ *) open Util open Stdpp @@ -25,12 +25,14 @@ type pretype_error = (* Unification *) | OccurCheck of existential_key * constr | NotClean of existential_key * constr * Evd.hole_kind - | UnsolvableImplicit of Evd.hole_kind + | UnsolvableImplicit of Evd.evar_info * Evd.hole_kind * + Evd.unsolvability_explanation option | CannotUnify of constr * constr - | CannotUnifyLocal of Environ.env * constr * constr * constr + | CannotUnifyLocal of constr * constr * constr | CannotUnifyBindingType of constr * constr | CannotGeneralize of constr - | NoOccurrenceFound of constr + | NoOccurrenceFound of constr * identifier option + | CannotFindWellTypedAbstraction of constr * constr list (* Pretyping *) | VarNotFound of identifier | UnexpectedType of constr * constr @@ -51,7 +53,7 @@ let j_nf_evar sigma j = let jl_nf_evar sigma jl = List.map (j_nf_evar 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=type_app (nf_evar sigma) v;utj_type=t} + {utj_val=nf_evar sigma v;utj_type=t} let env_ise sigma env = let sign = named_context_val env in @@ -60,7 +62,7 @@ let env_ise sigma env = Sign.fold_rel_context (fun (na,b,ty) e -> push_rel - (na, option_map (nf_evar sigma) b, nf_evar sigma ty) + (na, Option.map (nf_evar sigma) b, nf_evar sigma ty) e) ctxt ~init:env0 @@ -76,7 +78,7 @@ let contract env lc = env | _ -> let t' = substl !l t in - let c' = option_map (substl !l) c in + let c' = Option.map (substl !l) c in let na' = named_hd env t' na in l := (mkRel 1) :: List.map (lift 1) !l; push_rel (na',c',t') env in @@ -152,18 +154,22 @@ let error_not_clean env sigma ev c (loc,k) = raise_with_loc loc (PretypeError (env_ise sigma env, NotClean (ev,c,k))) -let error_unsolvable_implicit loc env sigma e = - raise_with_loc loc (PretypeError (env_ise sigma env, UnsolvableImplicit e)) +let error_unsolvable_implicit loc env sigma evi e explain = + raise_with_loc loc + (PretypeError (env_ise sigma env, UnsolvableImplicit (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 (e,m,n,sn) = - raise (PretypeError (env_ise sigma env,CannotUnifyLocal (e,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) = raise (PretypeError (env_ise sigma env,CannotUnify (m,n))) +let error_cannot_find_well_typed_abstraction env sigma p l = + raise (PretypeError (env_ise sigma env,CannotFindWellTypedAbstraction (p,l))) + (*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 137ef639..6ad2793f 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 9217 2006-10-05 17:31:23Z notin $ i*) +(*i $Id: pretype_errors.mli 10860 2008-04-27 21:39:08Z herbelin $ i*) (*i*) open Pp @@ -27,12 +27,14 @@ type pretype_error = (* Unification *) | OccurCheck of existential_key * constr | NotClean of existential_key * constr * Evd.hole_kind - | UnsolvableImplicit of Evd.hole_kind + | UnsolvableImplicit of Evd.evar_info * Evd.hole_kind * + Evd.unsolvability_explanation option | CannotUnify of constr * constr - | CannotUnifyLocal of Environ.env * constr * constr * constr + | CannotUnifyLocal of constr * constr * constr | CannotUnifyBindingType of constr * constr | CannotGeneralize of constr - | NoOccurrenceFound of constr + | NoOccurrenceFound of constr * identifier option + | CannotFindWellTypedAbstraction of constr * constr list (* Pretyping *) | VarNotFound of identifier | UnexpectedType of constr * constr @@ -93,11 +95,15 @@ 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.hole_kind -> 'b + 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 -val error_cannot_unify_local : env -> Evd.evar_map -> Environ.env * constr * constr * constr -> 'b +val error_cannot_unify_local : env -> Evd.evar_map -> constr * constr * constr -> 'b + +val error_cannot_find_well_typed_abstraction : env -> Evd.evar_map -> + constr -> constr list -> 'b (*s Ml Case errors *) diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 0db64a52..5f0999cb 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: pretyping.ml 9976 2007-07-12 11:58:30Z msozeau $ *) +(* $Id: pretyping.ml 11047 2008-06-03 23:08:00Z msozeau $ *) open Pp open Util @@ -43,6 +43,34 @@ open Inductiveops (************************************************************************) +(* An auxiliary function for searching for fixpoint guard indexes *) + +exception Found of int array + +let search_guard loc env possible_indexes fixdefs = + (* Standard situation with only one possibility for each fix. *) + (* We treat it separately in order to get proper error msg. *) + if List.for_all (fun l->1=List.length l) possible_indexes then + let indexes = Array.of_list (List.map List.hd possible_indexes) in + let fix = ((indexes, 0),fixdefs) in + (try check_fix env fix with + | e -> if loc = dummy_loc then raise e else Stdpp.raise_with_loc loc e); + indexes + else + (* we now search recursively amoungst all combinations *) + (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) + with TypeError _ -> ()) + (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" @@ -70,7 +98,7 @@ sig unresolved holes as evars and returning the typing contexts of these evars. Work as [understand_gen] for the rest. *) - val understand_tcc : + val understand_tcc : ?resolve_classes:bool -> evar_map -> env -> ?expected_type:types -> rawconstr -> open_constr val understand_tcc_evars : @@ -129,7 +157,7 @@ sig rawconstr -> unsafe_type_judgment val pretype_gen : - evar_defs ref -> env -> + evar_defs ref -> env -> var_map * (identifier * identifier option) list -> typing_constraint -> rawconstr -> constr @@ -143,90 +171,83 @@ module Pretyping_F (Coercion : Coercion.S) = struct (* Allow references to syntaxically inexistent variables (i.e., if applied on an inductive) *) let allow_anonymous_refs = ref false - let evd_comb0 f isevars = - let (evd',x) = f !isevars in - isevars := evd'; + let evd_comb0 f evdref = + let (evd',x) = f !evdref in + evdref := evd'; x - let evd_comb1 f isevars x = - let (evd',y) = f !isevars x in - isevars := evd'; + let evd_comb1 f evdref x = + let (evd',y) = f !evdref x in + evdref := evd'; y - let evd_comb2 f isevars x y = - let (evd',z) = f !isevars x y in - isevars := evd'; + let evd_comb2 f evdref x y = + let (evd',z) = f !evdref x y in + evdref := evd'; z - let evd_comb3 f isevars x y z = - let (evd',t) = f !isevars x y z in - isevars := evd'; + let evd_comb3 f evdref x y z = + let (evd',t) = f !evdref x y z in + evdref := evd'; t let mt_evd = Evd.empty - let vect_lift_type = Array.mapi (fun i t -> type_app (lift i) t) - (* 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 isevars lna lar vdefj = + let evar_type_fixpoint loc env evdref lna lar vdefj = let lt = Array.length vdefj in if Array.length lar = lt then for i = 0 to lt-1 do - if not (e_cumul env isevars (vdefj.(i)).uj_type + 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 !isevars) + error_ill_typed_rec_body_loc loc env (evars_of !evdref) i lna vdefj lar done - let check_branches_message loc env isevars 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 isevars lft.(i) explft.(i)) then - let sigma = evars_of !isevars in + if not (e_cumul env evdref lft.(i) explft.(i)) then + let sigma = evars_of !evdref in error_ill_formed_branch_loc loc env sigma c i lft.(i) explft.(i) done (* coerce to tycon if any *) - let inh_conv_coerce_to_tycon loc env isevars j = function + let inh_conv_coerce_to_tycon loc env evdref j = function | None -> j - | Some t -> evd_comb2 (Coercion.inh_conv_coerce_to loc env) isevars j t + | Some t -> evd_comb2 (Coercion.inh_conv_coerce_to loc env) evdref j t let push_rels vars env = List.fold_right push_rel vars env - (* - let evar_type_case isevars env ct pt lft p c = - let (mind,bty,rslty) = type_case_branches env (evars_of isevars) ct pt p c - in check_branches_message isevars env (c,ct) (bty,lft); (mind,rslty) - *) + (* used to enforce a name in Lambda when the type constraints itself + is named, hence possibly dependent *) - let strip_meta id = (* For Grammar v7 compatibility *) - let s = string_of_id id in - if s.[0]='$' then id_of_string (String.sub s 1 (String.length s - 1)) - else id + let orelse_name name name' = match name with + | Anonymous -> name' + | _ -> name let pretype_id loc env (lvar,unbndltacvars) id = - let id = strip_meta id in (* May happen in tactics defined by Grammar *) - try - let (n,typ) = lookup_rel_id id (rel_context env) in - { uj_val = mkRel n; uj_type = type_app (lift n) typ } - with Not_found -> - try - List.assoc id lvar - with Not_found -> - try - let (_,_,typ) = lookup_named id env in - { uj_val = mkVar id; uj_type = typ } - with Not_found -> - try (* To build a nicer ltac error message *) - match List.assoc id unbndltacvars with - | None -> user_err_loc (loc,"", - str "variable " ++ pr_id id ++ str " should be bound to a term") - | Some id0 -> Pretype_errors.error_var_not_found_loc loc id0 - with Not_found -> - error_var_not_found_loc loc id + try + let (n,typ) = lookup_rel_id id (rel_context env) in + { uj_val = mkRel n; uj_type = lift n typ } + with Not_found -> + try + List.assoc id lvar + with Not_found -> + try + let (_,_,typ) = lookup_named id env in + { uj_val = mkVar id; uj_type = typ } + with Not_found -> + try (* To build a nicer ltac error message *) + match List.assoc id unbndltacvars with + | None -> user_err_loc (loc,"", + str "variable " ++ pr_id id ++ str " should be bound to a term") + | Some id0 -> Pretype_errors.error_var_not_found_loc loc id0 + with Not_found -> + error_var_not_found_loc loc id (* make a dependent predicate from an undependent one *) @@ -251,7 +272,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct (*************************************************************************) (* Main pretyping function *) - let pretype_ref isevars 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) @@ -259,30 +280,32 @@ module Pretyping_F (Coercion : Coercion.S) = struct | RProp c -> judge_of_prop_contents c | RType _ -> judge_of_new_Type () - (* [pretype tycon env isevars lvar lmeta cstr] attempts to type [cstr] *) - (* in environment [env], with existential variables [(evars_of isevars)] and *) + exception Found of fixpoint + + (* [pretype tycon env evdref lvar lmeta cstr] attempts to type [cstr] *) + (* in environment [env], with existential variables [evdref] and *) (* the type constraint tycon *) - let rec pretype (tycon : type_constraint) env isevars lvar = function + let rec pretype (tycon : type_constraint) env evdref lvar = function | RRef (loc,ref) -> - inh_conv_coerce_to_tycon loc env isevars - (pretype_ref isevars env ref) + inh_conv_coerce_to_tycon loc env evdref + (pretype_ref evdref env ref) tycon | RVar (loc, id) -> - inh_conv_coerce_to_tycon loc env isevars + inh_conv_coerce_to_tycon loc env evdref (pretype_id loc env lvar id) tycon - | REvar (loc, ev, instopt) -> + | 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_context (Evd.find (evars_of !isevars) ev) in + let hyps = evar_filtered_context (Evd.find (evars_of !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 (ev, args) in - let j = (Retyping.get_judgment_of env (evars_of !isevars) c) in - inh_conv_coerce_to_tycon loc env isevars j tycon + let c = mkEvar (evk, args) in + let j = (Retyping.get_judgment_of env (evars_of !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" @@ -292,26 +315,26 @@ module Pretyping_F (Coercion : Coercion.S) = struct match tycon with | Some (None, ty) -> ty | None | Some _ -> - e_new_evar isevars env ~src:(loc,InternalHole) (new_Type ()) in - { uj_val = e_new_evar isevars env ~src:(loc,k) ty; uj_type = ty } + e_new_evar evdref env ~src:(loc,InternalHole) (new_Type ()) in + { uj_val = e_new_evar evdref env ~src:(loc,k) ty; uj_type = ty } | RRec (loc,fixkind,names,bl,lar,vdef) -> let rec type_bl env ctxt = function [] -> ctxt - | (na,None,ty)::bl -> - let ty' = pretype_type empty_valcon env isevars lvar ty in + | (na,bk,None,ty)::bl -> + let ty' = pretype_type empty_valcon env evdref lvar ty in let dcl = (na,None,ty'.utj_val) in type_bl (push_rel dcl env) (add_rel_decl dcl ctxt) bl - | (na,Some bd,ty)::bl -> - let ty' = pretype_type empty_valcon env isevars lvar ty in - let bd' = pretype (mk_tycon ty'.utj_val) env isevars lvar ty in + | (na,bk,Some bd,ty)::bl -> + let ty' = pretype_type empty_valcon env evdref lvar ty in + let bd' = pretype (mk_tycon ty'.utj_val) env evdref lvar ty in let dcl = (na,Some bd'.uj_val,ty'.utj_val) in type_bl (push_rel dcl env) (add_rel_decl dcl ctxt) bl in let ctxtv = Array.map (type_bl env empty_rel_context) bl in let larj = array_map2 (fun e ar -> - pretype_type empty_valcon (push_rel_context e env) isevars lvar ar) + pretype_type empty_valcon (push_rel_context e env) evdref lvar ar) ctxtv lar in let lara = Array.map (fun a -> a.utj_val) larj in let ftys = array_map2 (fun e a -> it_mkProd_or_LetIn a e) ctxtv lara in @@ -328,115 +351,111 @@ module Pretyping_F (Coercion : Coercion.S) = struct decompose_prod_n_assum (rel_context_length ctxt) (lift nbfix ftys.(i)) in let nenv = push_rel_context ctxt newenv in - let j = pretype (mk_tycon ty) nenv isevars lvar def in + let j = pretype (mk_tycon ty) nenv evdref lvar def in { uj_val = it_mkLambda_or_LetIn j.uj_val ctxt; uj_type = it_mkProd_or_LetIn j.uj_type ctxt }) ctxtv vdef in - evar_type_fixpoint loc env isevars names ftys vdefj; + evar_type_fixpoint loc env evdref names ftys vdefj; let fixj = match fixkind with | RFix (vn,i) -> - let guard_indexes = Array.mapi + (* First, let's find the guard indexes. *) + (* If recursive argument was not given by user, we try all args. + An earlier approach was to look only for inductive arguments, + but doing it properly involves delta-reduction, and it finally + doesn't seem worth the effort (except for huge mutual + fixpoints ?) *) + let possible_indexes = Array.to_list (Array.mapi (fun i (n,_) -> match n with - | Some n -> n - | None -> - (* Recursive argument was not given by the user : We - check that there is only one inductive argument *) - let ctx = ctxtv.(i) in - let isIndApp t = - isInd (fst (decompose_app (strip_head_cast t))) in - (* This could be more precise (e.g. do some delta) *) - let lb = List.rev_map (fun (_,_,t) -> isIndApp t) ctx in - try (list_unique_index true lb) - 1 - with Not_found -> - Util.user_err_loc - (loc,"pretype", - Pp.str "cannot guess decreasing argument of fix")) - vn - in - let fix = ((guard_indexes, i),(names,ftys,Array.map j_val vdefj)) in - (try check_fix env fix with e -> Stdpp.raise_with_loc loc e); - make_judge (mkFix fix) ftys.(i) - | RCoFix i -> + | Some n -> [n] + | None -> list_map_i (fun i _ -> i) 0 ctxtv.(i)) + vn) + in + let fixdecls = (names,ftys,Array.map j_val vdefj) 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,Array.map j_val vdefj)) in (try check_cofix env cofix with e -> Stdpp.raise_with_loc loc e); make_judge (mkCoFix cofix) ftys.(i) in - inh_conv_coerce_to_tycon loc env isevars fixj tycon + inh_conv_coerce_to_tycon loc env evdref fixj tycon | RSort (loc,s) -> - inh_conv_coerce_to_tycon loc env isevars (pretype_sort s) tycon + inh_conv_coerce_to_tycon loc env evdref (pretype_sort s) tycon | RApp (loc,f,args) -> - let fj = pretype empty_tycon env isevars lvar f in + 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 | [] -> resj | c::rest -> let argloc = loc_of_rawconstr c in - let resj = evd_comb1 (Coercion.inh_app_fun env) isevars resj in - let resty = whd_betadeltaiota env (evars_of !isevars) resj.uj_type 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 match kind_of_term resty with | Prod (na,c1,c2) -> - let hj = pretype (mk_tycon c1) env isevars lvar c in + let hj = pretype (mk_tycon c1) env evdref lvar c in let value, typ = applist (j_val resj, [j_val hj]), subst1 hj.uj_val c2 in apply_rec env (n+1) { uj_val = value; uj_type = typ } rest | _ -> - let hj = pretype empty_tycon env isevars lvar c in + let hj = pretype empty_tycon env evdref lvar c in error_cant_apply_not_functional_loc - (join_loc floc argloc) env (evars_of !isevars) + (join_loc floc argloc) env (evars_of !evdref) resj [hj] in let resj = apply_rec env 1 fj args in let resj = - match evar_kind_of_term !isevars resj.uj_val with + match evar_kind_of_term !evdref resj.uj_val with | App (f,args) -> - let f = whd_evar (Evd.evars_of !isevars) f in + let f = whd_evar (Evd.evars_of !evdref) f in begin match kind_of_term f with - | Ind _ (* | Const _ *) -> - let sigma = evars_of !isevars in + | Ind _ | Const _ + when isInd f or has_polymorphic_type (destConst f) + -> + let sigma = evars_of !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 | _ -> resj end | _ -> resj in - inh_conv_coerce_to_tycon loc env isevars resj tycon + inh_conv_coerce_to_tycon loc env evdref resj tycon - | RLambda(loc,name,c1,c2) -> - let (name',dom,rng) = evd_comb1 (split_tycon loc env) isevars tycon in + | RLambda(loc,name,bk,c1,c2) -> + let (name',dom,rng) = evd_comb1 (split_tycon loc env) evdref tycon in let dom_valcon = valcon_of_tycon dom in - let j = pretype_type dom_valcon env isevars lvar c1 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) isevars lvar c2 in - judge_of_abstraction env name j j' + 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,c1,c2) -> - let j = pretype_type empty_valcon env isevars lvar c1 in + | 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' isevars lvar c2 in + let j' = 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 isevars resj tycon + inh_conv_coerce_to_tycon loc env evdref resj tycon | RLetIn(loc,name,c1,c2) -> - let j = pretype empty_tycon env isevars lvar c1 in + let j = pretype empty_tycon env evdref lvar c1 in let t = refresh_universes j.uj_type in let var = (name,Some j.uj_val,t) in let tycon = lift_tycon 1 tycon in - let j' = pretype tycon (push_rel var env) isevars lvar c2 in + let j' = pretype tycon (push_rel var env) evdref lvar c2 in { uj_val = mkLetIn (name, j.uj_val, t, j'.uj_val) ; uj_type = subst1 j.uj_val j'.uj_type } | RLetTuple (loc,nal,(na,po),c,d) -> - let cj = pretype empty_tycon env isevars lvar c in + let cj = pretype empty_tycon env evdref lvar c in let (IndType (indf,realargs)) = - try find_rectype env (evars_of !isevars) cj.uj_type + try find_rectype env (evars_of !evdref) cj.uj_type with Not_found -> let cloc = loc_of_rawconstr c in - error_case_not_inductive_loc cloc env (evars_of !isevars) cj + error_case_not_inductive_loc cloc env (evars_of !evdref) cj in let cstrs = get_constructors env indf in if Array.length cstrs <> 1 then @@ -459,50 +478,50 @@ module Pretyping_F (Coercion : Coercion.S) = struct (match po with | Some p -> let env_p = push_rels psign env in - let pj = pretype_type empty_valcon env_p isevars lvar p in - let ccl = nf_isevar !isevars pj.utj_val in + let pj = pretype_type empty_valcon env_p evdref lvar p in + let ccl = nf_isevar !evdref pj.utj_val in let psign = make_arity_signature env true indf in (* with names *) let p = it_mkLambda_or_LetIn ccl psign in let inst = (Array.to_list cs.cs_concl_realargs) @[build_dependent_constructor cs] in let lp = lift cs.cs_nargs p in - let fty = hnf_lam_applist env (evars_of !isevars) lp inst in - let fj = pretype (mk_tycon fty) env_f isevars lvar d in + let fty = hnf_lam_applist env (evars_of !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_default_case_info env LetStyle mis in + let ci = make_case_info env mis LetStyle in mkCase (ci, p, cj.uj_val,[|f|]) in { uj_val = v; uj_type = substl (realargs@[cj.uj_val]) ccl } | None -> let tycon = lift_tycon cs.cs_nargs tycon in - let fj = pretype tycon env_f isevars lvar d 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 !isevars fj.uj_type in + let ccl = nf_isevar !evdref fj.uj_type in let ccl = if noccur_between 1 cs.cs_nargs ccl then lift (- cs.cs_nargs) ccl else - error_cant_find_case_type_loc loc env (evars_of !isevars) + error_cant_find_case_type_loc loc env (evars_of !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_default_case_info env LetStyle mis in + let ci = make_case_info env mis LetStyle in 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 isevars lvar c in + let cj = pretype empty_tycon env evdref lvar c in let (IndType (indf,realargs)) = - try find_rectype env (evars_of !isevars) cj.uj_type + try find_rectype env (evars_of !evdref) cj.uj_type with Not_found -> let cloc = loc_of_rawconstr c in - error_case_not_inductive_loc cloc env (evars_of !isevars) cj in + error_case_not_inductive_loc cloc env (evars_of !evdref) cj in let cstrs = get_constructors env indf in if Array.length cstrs <> 2 then user_err_loc (loc,"", @@ -520,11 +539,11 @@ module Pretyping_F (Coercion : Coercion.S) = struct let pred,p = match po with | Some p -> let env_p = push_rels psign env in - let pj = pretype_type empty_valcon env_p isevars lvar p in - let ccl = nf_evar (evars_of !isevars) pj.utj_val 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 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 isevars {uj_val = pred; + let jtyp = inh_conv_coerce_to_tycon loc env evdref {uj_val = pred; uj_type = typ} tycon in jtyp.uj_val, jtyp.uj_type @@ -532,11 +551,11 @@ module Pretyping_F (Coercion : Coercion.S) = struct let p = match tycon with | Some (None, ty) -> ty | None | Some _ -> - e_new_evar isevars env ~src:(loc,InternalHole) (new_Type ()) + 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 !isevars) pred in - let p = nf_evar (evars_of !isevars) 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 f cs b = let n = rel_context_length cs.cs_args in @@ -555,118 +574,129 @@ module Pretyping_F (Coercion : Coercion.S) = struct in let env_c = push_rels csgn env in (* msgnl (str "Pi is: " ++ Termops.print_constr_env env_c pi); *) - let bj = pretype (mk_tycon pi) env_c isevars lvar b in + let bj = pretype (mk_tycon pi) env_c evdref lvar b in it_mkLambda_or_LetIn bj.uj_val cs.cs_args in let b1 = f cstrs.(0) b1 in let b2 = f cstrs.(1) b2 in let v = let mis,_ = dest_ind_family indf in - let ci = make_default_case_info env IfStyle mis in + let ci = make_case_info env mis IfStyle in mkCase (ci, pred, cj.uj_val, [|b1;b2|]) in { uj_val = v; uj_type = p } - - | RCases (loc,po,tml,eqns) -> - Cases.compile_cases loc - ((fun vtyc env -> pretype vtyc env isevars lvar),isevars) + + | RCases (loc,sty,po,tml,eqns) -> + Cases.compile_cases loc sty + ((fun vtyc env evdref -> pretype vtyc env evdref lvar),evdref) tycon env (* loc *) (po,tml,eqns) | RCast (loc,c,k) -> let cj = match k with CastCoerce -> - let cj = pretype empty_tycon env isevars lvar c in - evd_comb1 (Coercion.inh_coerce_to_base loc env) isevars cj + let cj = pretype empty_tycon env evdref lvar c in + evd_comb1 (Coercion.inh_coerce_to_base loc env) evdref cj | CastConv (k,t) -> - let tj = pretype_type empty_valcon env isevars lvar t in - let cj = pretype (mk_tycon tj.utj_val) env isevars lvar c in + 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 } in - inh_conv_coerce_to_tycon loc env isevars cj tycon + inh_conv_coerce_to_tycon loc env evdref cj tycon | RDynamic (loc,d) -> if (tag d) = "constr" then let c = constr_out d in - let j = (Retyping.get_judgment_of env (evars_of !isevars) c) in + let j = (Retyping.get_judgment_of env (evars_of !evdref) c) in j - (*inh_conv_coerce_to_tycon loc env isevars j tycon*) + (*inh_conv_coerce_to_tycon loc env evdref j tycon*) else user_err_loc (loc,"pretype",(str "Not a constr tagged Dynamic")) - (* [pretype_type valcon env isevars lvar c] coerces [c] into a type *) - and pretype_type valcon env isevars lvar = function + (* [pretype_type valcon env evdref lvar c] coerces [c] into a type *) + and pretype_type valcon env evdref lvar = function | RHole loc -> (match valcon with | Some v -> let s = - let sigma = evars_of !isevars in + let sigma = evars_of !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 v when is_Type (existential_type sigma v) -> - evd_comb1 (define_evar_as_sort) isevars v + | Evar ev when is_Type (existential_type sigma ev) -> + evd_comb1 (define_evar_as_sort) evdref ev | _ -> anomaly "Found a type constraint which is not a type" in { utj_val = v; utj_type = s } | None -> let s = new_Type_sort () in - { utj_val = e_new_evar isevars env ~src:loc (mkSort s); + { utj_val = e_new_evar evdref env ~src:loc (mkSort s); utj_type = s}) | c -> - let j = pretype empty_tycon env isevars lvar c in + let j = pretype empty_tycon env evdref lvar c in let loc = loc_of_rawconstr c in - let tj = evd_comb1 (Coercion.inh_coerce_to_sort loc env) isevars j in + let tj = evd_comb1 (Coercion.inh_coerce_to_sort loc env) evdref j in match valcon with | None -> tj | Some v -> - if e_cumul env isevars v tj.utj_val then tj + if e_cumul env evdref v tj.utj_val then tj else error_unexpected_type_loc - (loc_of_rawconstr c) env (evars_of !isevars) tj.utj_val v + (loc_of_rawconstr c) env (evars_of !evdref) tj.utj_val v - let pretype_gen isevars env lvar kind c = + let pretype_gen_aux 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 isevars lvar c).uj_val + (pretype tycon env evdref lvar c).uj_val | IsType -> - (pretype_type empty_valcon env isevars lvar c).utj_val in - nf_evar (evars_of !isevars) c' - + (pretype_type empty_valcon env evdref lvar c).utj_val in + let evd,_ = consider_remaining_unif_problems env !evdref in + evdref := evd; 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_evar (evars_of !evdref) c + (* TODO: comment faire remonter l'information si le typage a resolu des variables du sigma original. il faudrait que la fonction de typage retourne aussi le nouveau sigma... *) let understand_judgment sigma env c = - let isevars = ref (create_evar_defs sigma) in - let j = pretype empty_tycon env isevars ([],[]) c in - let isevars,_ = consider_remaining_unif_problems env !isevars in - let j = j_nf_evar (evars_of isevars) j in - check_evars env sigma isevars (mkCast(j.uj_val,DEFAULTcast, j.uj_type)); - j - - let understand_judgment_tcc isevars env c = - let j = pretype empty_tycon env isevars ([],[]) c in - let sigma = evars_of !isevars in + 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 + 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 (* 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 isevars = ref (Evd.create_evar_defs sigma) in - let c = pretype_gen isevars env lvar kind c in - let isevars,_ = consider_remaining_unif_problems env !isevars in - let c = nf_evar (evars_of isevars) c in - if fail_evar then check_evars env sigma isevars c; - isevars, c + let evdref = ref (Evd.create_evar_defs sigma) in + let c = pretype_gen evdref env lvar kind c in + let evd,_ = consider_remaining_unif_problems env !evdref in + if fail_evar then + let evd = Typeclasses.resolve_typeclasses ~onlyargs:false ~fail:true env evd in + let c = Evarutil.nf_isevar evd c in + check_evars env Evd.empty evd c; + evd, c + else evd, c (** Entry points of the high-level type synthesis algorithm *) @@ -682,12 +712,19 @@ module Pretyping_F (Coercion : Coercion.S) = struct let understand_ltac sigma env lvar kind c = ise_pretype_gen false sigma env lvar kind c - let understand_tcc_evars isevars env kind c = - pretype_gen isevars env ([],[]) kind c - - let understand_tcc sigma env ?expected_type:exptyp c = - let ev, t = ise_pretype_gen false sigma env ([],[]) (OfType exptyp) c in - Evd.evars_of ev, t + let understand_tcc_evars evdref env kind c = + pretype_gen evdref env ([],[]) kind c + + let understand_tcc ?(resolve_classes=true) sigma env ?expected_type:exptyp c = + let evd, t = + if resolve_classes then + ise_pretype_gen false sigma env ([],[]) (OfType exptyp) c + else + let evdref = ref (Evd.create_evar_defs sigma) in + let c = pretype_gen_aux evdref env ([],[]) (OfType exptyp) c in + !evdref, nf_isevar !evdref c + in + Evd.evars_of evd, t end module Default : S = Pretyping_F(Coercion.Default) diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli index b1ed20c2..4f116053 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 9141 2006-09-15 10:07:01Z herbelin $ i*) +(*i $Id: pretyping.mli 11047 2008-06-03 23:08:00Z msozeau $ i*) (*i*) open Names @@ -18,6 +18,11 @@ open Rawterm open Evarutil (*i*) +(* An auxiliary function for searching for fixpoint guard indexes *) + +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 @@ -35,7 +40,7 @@ sig unresolved holes as evars and returning the typing contexts of these evars. Work as [understand_gen] for the rest. *) - val understand_tcc : + val understand_tcc : ?resolve_classes:bool -> evar_map -> env -> ?expected_type:types -> rawconstr -> open_constr val understand_tcc_evars : diff --git a/pretyping/rawterm.ml b/pretyping/rawterm.ml index aaf9e63d..62798a45 100644 --- a/pretyping/rawterm.ml +++ b/pretyping/rawterm.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: rawterm.ml 9976 2007-07-12 11:58:30Z msozeau $ *) +(* $Id: rawterm.ml 11094 2008-06-10 19:35:23Z herbelin $ *) (*i*) open Util @@ -36,6 +36,8 @@ type rawsort = RProp of Term.contents | RType of Univ.universe option type binder_kind = BProd | BLambda | BLetIn +type binding_kind = Explicit | Implicit + type quantified_hypothesis = AnonHyp of int | NamedHyp of identifier type 'a explicit_bindings = (loc * quantified_hypothesis * 'a) list @@ -57,10 +59,10 @@ type rawconstr = | REvar of loc * existential_key * rawconstr list option | RPatVar of loc * (bool * patvar) (* Used for patterns only *) | RApp of loc * rawconstr * rawconstr list - | RLambda of loc * name * rawconstr * rawconstr - | RProd of loc * name * rawconstr * rawconstr + | RLambda of loc * name * binding_kind * rawconstr * rawconstr + | RProd of loc * name * binding_kind * rawconstr * rawconstr | RLetIn of loc * name * rawconstr * rawconstr - | RCases of loc * rawconstr option * tomatch_tuple * cases_clauses + | RCases of loc * case_style * rawconstr option * tomatch_tuples * cases_clauses | RLetTuple of loc * name list * (name * rawconstr option) * rawconstr * rawconstr | RIf of loc * rawconstr * (name * rawconstr option) * rawconstr * rawconstr @@ -71,7 +73,7 @@ type rawconstr = | RCast of loc * rawconstr * rawconstr cast_type | RDynamic of loc * Dyn.t -and rawdecl = name * rawconstr option * rawconstr +and rawdecl = name * binding_kind * rawconstr option * rawconstr and fix_recursion_order = RStructRec | RWfRec of rawconstr | RMeasureRec of rawconstr @@ -82,10 +84,13 @@ and fix_kind = and predicate_pattern = name * (loc * inductive * int * name list) option -and tomatch_tuple = (rawconstr * predicate_pattern) list +and tomatch_tuple = (rawconstr * predicate_pattern) + +and tomatch_tuples = tomatch_tuple list + +and cases_clause = (loc * identifier list * cases_pattern list * rawconstr) -and cases_clauses = - (loc * identifier list * cases_pattern list * rawconstr) list +and cases_clauses = cases_clause list let cases_predicate_names tml = List.flatten (List.map (function @@ -101,22 +106,22 @@ let cases_predicate_names tml = - boolean in POldCase means it is recursive i*) -let map_rawdecl f (na,obd,ty) = (na,option_map f obd,f ty) +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,ty,c) -> RLambda (loc,na,f ty,f c) - | RProd (loc,na,ty,c) -> RProd (loc,na,f ty,f c) + | 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) - | RCases (loc,rtntypopt,tml,pl) -> - RCases (loc,option_map f rtntypopt, + | 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) | RLetTuple (loc,nal,(na,po),b,c) -> - RLetTuple (loc,nal,(na,option_map f po),f b,f c) + RLetTuple (loc,nal,(na,Option.map f po),f b,f c) | RIf (loc,c,(na,po),b1,b2) -> - RIf (loc,f c,(na,option_map f po),f b1,f b2) + RIf (loc,f c,(na,Option.map f po),f b1,f b2) | 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) @@ -149,7 +154,7 @@ let map_rawconstr_with_binders_loc loc g f e = function let g' id e = snd (g id e) in let h (_,idl,p,c) = (loc,idl,p,f (List.fold_right g' idl e) c) in RCases - (loc,option_map (f e) tyopt,List.map (f e) tml, List.map h pl) + (loc,Option.map (f e) tyopt,List.map (f e) tml, List.map h pl) | RRec (_,fk,idl,tyl,bv) -> let idl',e' = fold_ident g idl e in RRec (loc,fk,idl',Array.map (f e) tyl,Array.map (f e') bv) @@ -166,10 +171,10 @@ let occur_rawconstr id = let rec occur = function | RVar (loc,id') -> id = id' | RApp (loc,f,args) -> (occur f) or (List.exists occur args) - | RLambda (loc,na,ty,c) -> (occur ty) or ((na <> Name id) & (occur c)) - | RProd (loc,na,ty,c) -> (occur ty) or ((na <> Name id) & (occur c)) + | RLambda (loc,na,bk,ty,c) -> (occur ty) or ((na <> Name id) & (occur c)) + | RProd (loc,na,bk,ty,c) -> (occur ty) or ((na <> Name id) & (occur c)) | RLetIn (loc,na,b,c) -> (occur b) or ((na <> Name id) & (occur c)) - | RCases (loc,rtntypopt,tml,pl) -> + | RCases (loc,sty,rtntypopt,tml,pl) -> (occur_option rtntypopt) or (List.exists (fun (tm,_) -> occur tm) tml) or (List.exists occur_pattern pl) @@ -182,7 +187,7 @@ let occur_rawconstr id = not (array_for_all4 (fun fid bl ty bd -> let rec occur_fix = function [] -> not (occur ty) && (fid=id or not(occur bd)) - | (na,bbd,bty)::bl -> + | (na,k,bbd,bty)::bl -> not (occur bty) && (match bbd with Some bd -> not (occur bd) @@ -211,11 +216,11 @@ 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) -> + | 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,rtntypopt,tml,pl) -> + | 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 List.fold_left (vars_pattern bounded) vs2 pl @@ -234,7 +239,7 @@ let free_rawvars = let vars_fix i vs fid = let vs1,bounded1 = List.fold_left - (fun (vs,bounded) (na,bbd,bty) -> + (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 @@ -272,10 +277,10 @@ let loc_of_rawconstr = function | REvar (loc,_,_) -> loc | RPatVar (loc,_) -> loc | RApp (loc,_,_) -> loc - | RLambda (loc,_,_,_) -> loc - | RProd (loc,_,_,_) -> loc + | RLambda (loc,_,_,_,_) -> loc + | RProd (loc,_,_,_,_) -> loc | RLetIn (loc,_,_,_) -> loc - | RCases (loc,_,_,_) -> loc + | RCases (loc,_,_,_,_) -> loc | RLetTuple (loc,_,_,_,_) -> loc | RIf (loc,_,_,_,_) -> loc | RRec (loc,_,_,_,_,_) -> loc @@ -330,7 +335,14 @@ let all_flags = type 'a or_var = ArgArg of 'a | ArgVar of identifier located -type 'a with_occurrences = int or_var list * 'a +type occurrences_expr = bool * int or_var list + +let all_occurrences_expr_but l = (false,l) +let no_occurrences_expr_but l = (true,l) +let all_occurrences_expr = (false,[]) +let no_occurrences_expr = (true,[]) + +type 'a with_occurrences = occurrences_expr * 'a type ('a,'b) red_expr_gen = | Red of bool diff --git a/pretyping/rawterm.mli b/pretyping/rawterm.mli index 546b36b0..2ba8022f 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 9976 2007-07-12 11:58:30Z msozeau $ i*) +(*i $Id: rawterm.mli 11094 2008-06-10 19:35:23Z herbelin $ i*) (*i*) open Util @@ -40,6 +40,8 @@ type rawsort = RProp of Term.contents | RType of Univ.universe option type binder_kind = BProd | BLambda | BLetIn +type binding_kind = Explicit | Implicit + type quantified_hypothesis = AnonHyp of int | NamedHyp of identifier type 'a explicit_bindings = (loc * quantified_hypothesis * 'a) list @@ -61,10 +63,10 @@ type rawconstr = | REvar of loc * existential_key * rawconstr list option | RPatVar of loc * (bool * patvar) (* Used for patterns only *) | RApp of loc * rawconstr * rawconstr list - | RLambda of loc * name * rawconstr * rawconstr - | RProd of loc * name * rawconstr * rawconstr + | RLambda of loc * name * binding_kind * rawconstr * rawconstr + | RProd of loc * name * binding_kind * rawconstr * rawconstr | RLetIn of loc * name * rawconstr * rawconstr - | RCases of loc * rawconstr option * tomatch_tuple * cases_clauses + | RCases of loc * case_style * rawconstr option * tomatch_tuples * cases_clauses | RLetTuple of loc * name list * (name * rawconstr option) * rawconstr * rawconstr | RIf of loc * rawconstr * (name * rawconstr option) * rawconstr * rawconstr @@ -75,7 +77,7 @@ type rawconstr = | RCast of loc * rawconstr * rawconstr cast_type | RDynamic of loc * Dyn.t -and rawdecl = name * rawconstr option * rawconstr +and rawdecl = name * binding_kind * rawconstr option * rawconstr and fix_recursion_order = RStructRec | RWfRec of rawconstr | RMeasureRec of rawconstr @@ -86,12 +88,15 @@ and fix_kind = and predicate_pattern = name * (loc * inductive * int * name list) option -and tomatch_tuple = (rawconstr * predicate_pattern) list +and tomatch_tuple = (rawconstr * predicate_pattern) + +and tomatch_tuples = tomatch_tuple list + +and cases_clause = (loc * identifier list * cases_pattern list * rawconstr) -and cases_clauses = - (loc * identifier list * cases_pattern list * rawconstr) list +and cases_clauses = cases_clause list -val cases_predicate_names : tomatch_tuple -> name 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 @@ -141,7 +146,14 @@ val all_flags : 'a raw_red_flag type 'a or_var = ArgArg of 'a | ArgVar of identifier located -type 'a with_occurrences = int or_var list * 'a +type occurrences_expr = bool * int or_var list + +val all_occurrences_expr_but : int or_var list -> occurrences_expr +val no_occurrences_expr_but : int or_var list -> occurrences_expr +val all_occurrences_expr : occurrences_expr +val no_occurrences_expr : occurrences_expr + +type 'a with_occurrences = occurrences_expr * 'a type ('a,'b) red_expr_gen = | Red of bool diff --git a/pretyping/recordops.ml b/pretyping/recordops.ml index 5bbaa207..bd73740f 100644 --- a/pretyping/recordops.ml +++ b/pretyping/recordops.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: recordops.ml 9166 2006-09-23 11:20:06Z herbelin $ *) +(* $Id: recordops.ml 10577 2008-02-19 10:18:19Z corbinea $ *) open Util open Pp @@ -40,15 +40,13 @@ type struc_typ = { let structure_table = ref (Indmap.empty : struc_typ Indmap.t) let projection_table = ref Cmap.empty -let option_fold_right f p e = match p with Some a -> f a e | None -> e - let load_structure i (_,(ind,id,kl,projs)) = let n = (fst (Global.lookup_inductive ind)).Declarations.mind_nparams in 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 := - List.fold_right (option_fold_right (fun proj -> Cmap.add proj struc)) + List.fold_right (Option.fold_right (fun proj -> Cmap.add proj struc)) projs !projection_table let cache_structure o = @@ -60,7 +58,7 @@ let subst_structure (_,subst,((kn,i),id,kl,projs as obj)) = (* invariant: struc.s_PROJ is an evaluable reference. Thus we can take *) (* the first component of subst_con. *) list_smartmap - (option_smartmap (fun kn -> fst (subst_con subst kn))) + (Option.smartmap (fun kn -> fst (subst_con subst kn))) projs in if projs' == projs && kn' == kn then obj else @@ -68,7 +66,7 @@ let subst_structure (_,subst,((kn,i),id,kl,projs as obj)) = let discharge_structure (_,(ind,id,kl,projs)) = Some (Lib.discharge_inductive ind, id, kl, - List.map (option_map Lib.discharge_con) projs) + List.map (Option.map Lib.discharge_con) projs) let (inStruc,outStruc) = declare_object {(default_object "STRUCTURE") with @@ -117,12 +115,19 @@ that maps the pair (Li,ci) to the following data type obj_typ = { o_DEF : constr; + o_INJ : int; (* position of trivial argument (negative= none) *) o_TABS : constr list; (* ordered *) o_TPARAMS : constr list; (* ordered *) o_TCOMPS : constr list } (* ordered *) +type cs_pattern = + Const_cs of global_reference + | Prod_cs + | Sort_cs of sorts_family + | Default_cs + let object_table = - (ref [] : ((global_reference * global_reference) * obj_typ) list ref) + (ref [] : ((global_reference * cs_pattern) * obj_typ) list ref) let canonical_projections () = !object_table @@ -130,14 +135,31 @@ let keep_true_projections projs kinds = map_succeed (function (p,true) -> p | _ -> failwith "") (List.combine projs kinds) -(* Intended to always success *) +let cs_pattern_of_constr t = + match kind_of_term t with + App (f,vargs) -> + begin + try Const_cs (global_of_constr f) , -1, Array.to_list vargs with + _ -> 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 + try Const_cs (global_of_constr t) , -1, [] with + _ -> 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 = List.rev (List.map snd lt) in let args = snd (decompose_app t) in - let { s_EXPECTEDPARAM = p; s_PROJ = lpj; s_PROJKIND = kl } = lookup_structure ind in + let { s_EXPECTEDPARAM = p; s_PROJ = lpj; s_PROJKIND = kl } = + lookup_structure ind in let params, projs = list_chop p args in let lpj = keep_true_projections lpj kl in let lps = List.combine lpj projs in @@ -146,13 +168,16 @@ let compute_canonical_projections (con,ind) = (fun l (spopt,t) -> (* comp=components *) match spopt with | Some proji_sp -> - let c, args = decompose_app t in - (try (ConstRef proji_sp, global_of_constr c, args) :: l - with Not_found -> l) + begin + try + let patt, n , args = cs_pattern_of_constr t in + ((ConstRef proji_sp, patt, n, args) :: l) + with Not_found -> l + end | _ -> l) [] lps in - List.map (fun (refi,c,argj) -> - (refi,c),{o_DEF=v; o_TABS=lt; o_TPARAMS=params; o_TCOMPS=argj}) + List.map (fun (refi,c,inj,argj) -> + (refi,c),{o_DEF=v; o_INJ=inj; o_TABS=lt; o_TPARAMS=params; o_TCOMPS=argj}) comp let open_canonical_structure i (_,o) = diff --git a/pretyping/recordops.mli b/pretyping/recordops.mli index 0a05aef6..7a7d941d 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 9082 2006-08-24 17:03:28Z herbelin $ i*) +(*i $Id: recordops.mli 10577 2008-02-19 10:18:19Z corbinea $ i*) (*i*) open Names @@ -36,13 +36,22 @@ val find_projection_nparams : global_reference -> int (* the effective components of a structure and the projections of the *) (* structure *) +type cs_pattern = + Const_cs of global_reference + | Prod_cs + | Sort_cs of sorts_family + | Default_cs + type obj_typ = { o_DEF : constr; + o_INJ : int; (* position of trivial argument *) o_TABS : constr list; (* ordered *) o_TPARAMS : constr list; (* ordered *) o_TCOMPS : constr list } (* ordered *) - -val lookup_canonical_conversion : (global_reference * global_reference) -> obj_typ + +val cs_pattern_of_constr : constr -> cs_pattern * int * constr list + +val lookup_canonical_conversion : (global_reference * cs_pattern) -> obj_typ val declare_canonical_structure : global_reference -> unit val canonical_projections : unit -> - ((global_reference * global_reference) * obj_typ) list + ((global_reference * cs_pattern) * obj_typ) list diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index 19f515d0..6fc30aae 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: reductionops.ml 9106 2006-09-01 11:18:17Z herbelin $ *) +(* $Id: reductionops.ml 11010 2008-05-28 15:25:19Z barras $ *) open Pp open Util @@ -601,6 +601,9 @@ let fakey = Profile.declare_profile "fhnf_apply";; let fhnf_apply info k h a = Profile.profile4 fakey fhnf_apply info k h a;; *) +let is_transparent k = + Conv_oracle.get_strategy k <> Conv_oracle.Opaque + (* Conversion utility functions *) type conversion_test = constraints -> constraints @@ -611,26 +614,7 @@ let pb_equal = function | CUMUL -> CONV | CONV -> CONV -let sort_cmp pb s0 s1 cuniv = - match (s0,s1) with - | (Prop c1, Prop c2) -> if c1 = c2 then cuniv else raise NotConvertible - | (Prop c1, Type u) -> - (match pb with - CUMUL -> cuniv - | _ -> raise NotConvertible) - | (Type u1, Type u2) -> - (match pb with - | CONV -> enforce_eq u1 u2 cuniv - | CUMUL -> enforce_geq u2 u1 cuniv) - | (_, _) -> raise NotConvertible - -let base_sort_cmp pb s0 s1 = - match (s0,s1) with - | (Prop c1, Prop c2) -> c1 = c2 - | (Prop c1, Type u) -> pb = CUMUL - | (Type u1, Type u2) -> true - | (_, _) -> false - +let sort_cmp = sort_cmp let test_conversion f env sigma x y = try let _ = f env (nf_evar sigma x) (nf_evar sigma y) in true @@ -640,6 +624,14 @@ let is_conv env sigma = test_conversion Reduction.conv env sigma let is_conv_leq env sigma = test_conversion Reduction.conv_leq env sigma let is_fconv = function | CONV -> is_conv | CUMUL -> is_conv_leq +let test_trans_conversion f reds env sigma x y = + try let _ = f reds env (nf_evar sigma x) (nf_evar sigma y) in true + 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_fconv = function | CONV -> is_trans_conv | CUMUL -> is_trans_conv_leq + (********************************************************************) (* Special-Purpose Reduction *) (********************************************************************) @@ -675,10 +667,42 @@ let plain_instance s c = in if s = [] then c else irec c -(* Pourquoi ne fait-on pas nf_betaiota si s=[] ? *) -let instance s c = - if s = [] then c else local_strong whd_betaiota (plain_instance s c) - +(* [instance] is used for [res_pf]; the call to [local_strong whd_betaiota] + has (unfortunately) different subtle side effects: + + - ** Order of subgoals ** + If the lemma is a case analysis with parameters, it will move the + parameters as first subgoals (e.g. "case H" applied on + "H:D->A/\B|-C" will present the subgoal |-D first while w/o + betaiota the subgoal |-D would have come last). + + - ** Betaiota-contraction in statement ** + If the lemma has a parameter which is a function and this + function is applied in the lemma, then the _strong_ betaiota will + contract the application of the function to its argument (e.g. + "apply (H (fun x => x))" in "H:forall f, f 0 = 0 |- 0=0" will + result in applying the lemma 0=0 in which "(fun x => x) 0" has + been contracted). A goal to rewrite may then fail or succeed + differently. + + - ** 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 + generate a subgoal "forall a:A, P a" (and intro will use name + "a"), while with reduction, it will generate a subgoal "forall + b:A, P b" (and intro will use name "b"). + + - ** First-order pattern-matching ** + If a lemma has the type "(fun x => p) t" then rewriting t may fail + if the type of the lemma is first beta-reduced (this typically happens + when rewriting a single variable and the type of the lemma is obtained + by meta_instance (with empty map) which itself call instance with this + empty map. + *) + +let instance s c = + (* if s = [] then c else *) local_strong whd_betaiota (plain_instance s c) (* pseudo-reduction rule: * [hnf_prod_app env s (Prod(_,B)) N --> B[N] @@ -777,21 +801,23 @@ let is_sort env sigma arity = (* reduction to head-normal-form allowing delta/zeta only in argument of case/fix (heuristic used by evar_conv) *) -let rec apprec env sigma s = - let (t, stack as s) = whd_betaiota_state s in - match kind_of_term t with +let whd_betaiota_deltazeta_for_iota_state env sigma s = + let rec whrec s = + let (t, stack as s) = whd_betaiota_state 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 - apprec env sigma (rslt, stack) + whrec (rslt, stack) else s | Fix fix -> (match reduce_fix (whd_betadeltaiota_state env sigma) fix stack with - | Reduced s -> apprec env sigma s + | Reduced s -> whrec s | NotReducible -> s) | _ -> s + in whrec s (* A reduction function like whd_betaiota but which keeps casts * and does not reduce redexes containing existential variables. @@ -861,13 +887,12 @@ let is_arity env sigma c = let meta_value evd mv = let rec valrec mv = - try - let b = meta_fvalue evd mv in + match meta_opt_fvalue evd mv with + | Some (b,_) -> instance (List.map (fun mv' -> (mv',valrec mv')) (Metaset.elements b.freemetas)) b.rebus - with Anomaly _ | Not_found -> - mkMeta mv + | None -> mkMeta mv in valrec mv @@ -876,6 +901,55 @@ let meta_instance env b = List.map (fun mv -> (mv,meta_value env mv)) (Metaset.elements b.freemetas) in - instance c_sigma b.rebus + if c_sigma = [] then b.rebus else instance c_sigma b.rebus let nf_meta env c = meta_instance env (mk_freelisted c) + +(* Instantiate metas that create beta/iota redexes *) + +let meta_value evd mv = + let rec valrec mv = + match meta_opt_fvalue evd mv with + | Some (b,_) -> + instance + (List.map (fun mv' -> (mv',valrec mv')) (Metaset.elements b.freemetas)) + b.rebus + | None -> mkMeta mv + in + valrec mv + +let meta_reducible_instance evd b = + let fm = Metaset.elements b.freemetas in + let metas = List.fold_left (fun l mv -> + match (try meta_opt_fvalue evd mv with Not_found -> None) with + | Some (g,(_,s)) -> (mv,(g.rebus,s))::l + | None -> l) [] fm in + let rec irec u = + let u = whd_betaiota u in + match kind_of_term u with + | Case (ci,p,c,bl) when isMeta c or isCast c & isMeta (pi1 (destCast c)) -> + let m = try destMeta c with _ -> destMeta (pi1 (destCast c)) in + (match + try + let g,s = List.assoc m metas in + if isConstruct g or s <> CoerceToType then Some g else None + with Not_found -> None + with + | Some g -> irec (mkCase (ci,p,g,bl)) + | None -> mkCase (ci,irec p,c,Array.map irec bl)) + | App (f,l) when isMeta f or isCast f & isMeta (pi1 (destCast f)) -> + let m = try destMeta f with _ -> destMeta (pi1 (destCast f)) in + (match + try + let g,s = List.assoc m metas in + if isLambda g or s <> CoerceToType then Some g else None + with Not_found -> None + with + | Some g -> irec (mkApp (g,l)) + | None -> mkApp (f,Array.map irec l)) + | Meta m -> + (try let g,s = List.assoc m metas in if s<>CoerceToType then irec g else u + with Not_found -> u) + | _ -> map_constr irec u + in + if fm = [] then (* nf_betaiota? *) b.rebus else irec b.rebus diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index 1e9b3b5b..d48055cf 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 9106 2006-09-01 11:18:17Z herbelin $ i*) +(*i $Id: reductionops.mli 11010 2008-05-28 15:25:19Z barras $ i*) (*i*) open Names @@ -181,7 +181,8 @@ val is_arity : env -> evar_map -> constr -> bool val whd_programs : reduction_function -(* [reduce_fix] contracts a fix redex if it is actually reducible *) +(* [reduce_fix redfun fix stk] contracts [fix stk] if it is actually + reducible; the structural argument is reduced by [redfun] *) type fix_reduction_result = NotReducible | Reduced of state @@ -189,6 +190,9 @@ val fix_recarg : fixpoint -> constr stack -> (int * constr) option val reduce_fix : local_state_reduction_function -> fixpoint -> constr stack -> fix_reduction_result +(*s Querying the kernel conversion oracle: opaque/transparent constants *) +val is_transparent : 'a tableKey -> bool + (*s Conversion Functions (uses closures, lazy strategy) *) type conversion_test = constraints -> constraints @@ -197,12 +201,15 @@ val pb_is_equal : conv_pb -> bool val pb_equal : conv_pb -> conv_pb val sort_cmp : conv_pb -> sorts -> sorts -> conversion_test -val base_sort_cmp : conv_pb -> sorts -> sorts -> bool val is_conv : env -> evar_map -> constr -> constr -> bool val is_conv_leq : env -> evar_map -> constr -> constr -> bool val is_fconv : conv_pb -> env -> evar_map -> constr -> constr -> bool +val is_trans_conv : transparent_state -> env -> evar_map -> constr -> constr -> bool +val is_trans_conv_leq : transparent_state -> env -> evar_map -> constr -> constr -> bool +val is_trans_fconv : conv_pb -> transparent_state -> env -> evar_map -> constr -> constr -> bool + (*s Special-Purpose Reduction Functions *) val whd_meta : (metavariable * constr) list -> constr -> constr @@ -211,8 +218,9 @@ val instance : (metavariable * constr) list -> constr -> constr (*s Heuristic for Conversion with Evar *) -val apprec : state_reduction_function +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 diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index ecead438..82c2668c 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: retyping.ml 9314 2006-10-29 20:11:08Z herbelin $ *) +(* $Id: retyping.ml 11077 2008-06-09 11:26:32Z herbelin $ *) open Util open Term @@ -44,7 +44,7 @@ let type_of_var env id = with Not_found -> anomaly ("type_of: variable "^(string_of_id id)^" unbound") -let typeur sigma metamap = +let retype sigma metamap = let rec type_of env cstr= match kind_of_term cstr with | Meta n -> @@ -84,7 +84,7 @@ let typeur sigma metamap = and sort_of env t = match kind_of_term t with | Cast (c,_, s) when isSort s -> destSort s - | Sort (Prop c) -> type_0 + | Sort (Prop c) -> type1_sort | Sort (Type u) -> Type (Univ.super u) | Prod (name,t,c2) -> (match (sort_of env t, sort_of (push_rel (name,None,t) env) c2) with @@ -92,10 +92,12 @@ let typeur sigma metamap = | Prop _, (Prop Pos as s) -> s | Type _, (Prop Pos as s) when Environ.engagement env = Some ImpredicativeSet -> s - | Type u1, Prop Pos -> Type (Univ.sup u1 Univ.base_univ) - | Prop Pos, (Type u2) -> Type (Univ.sup Univ.base_univ u2) + | (Type _, _) | (_, Type _) -> new_Type_sort () +(* + | Type u1, Prop Pos -> Type (Univ.sup u1 Univ.type0_univ) + | Prop Pos, (Type u2) -> Type (Univ.sup Univ.type0_univ u2) | Prop Null, (Type _ as s) -> s - | Type u1, Type u2 -> Type (Univ.sup u1 u2)) + | Type u1, Type u2 -> Type (Univ.sup u1 u2)*)) | App(f,args) when isGlobalRef f -> let t = type_of_global_reference_knowing_parameters env f args in sort_of_atomic_type env sigma t args @@ -132,14 +134,18 @@ let typeur sigma metamap = in type_of, sort_of, sort_family_of, type_of_global_reference_knowing_parameters -let get_type_of env sigma c = let f,_,_,_ = typeur sigma [] in f env c -let get_sort_of env sigma t = let _,f,_,_ = typeur sigma [] in f env t -let get_sort_family_of env sigma c = let _,_,f,_ = typeur 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 = typeur sigma [] in f env c args + let _,_,_,f = retype sigma [] in f env c args -let get_type_of_with_meta env sigma metamap = - let f,_,_,_ = typeur sigma metamap in f env +(* 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) (* Makes an assumption from a constr *) let get_assumption_of env evc c = c diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 64e9ebec..50401502 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: tacred.ml 10135 2007-09-21 14:28:12Z herbelin $ *) +(* $Id: tacred.ml 11094 2008-06-10 19:35:23Z herbelin $ *) open Pp open Util @@ -35,18 +35,23 @@ exception ReductionTacticError of reduction_tactic_error exception Elimconst exception Redelimination -let is_evaluable env ref = - match ref with - EvalConstRef kn -> - let (ids,kns) = Conv_oracle.freeze() in - Cpred.mem kn kns & - let cb = Environ.lookup_constant kn env in - cb.const_body <> None & not cb.const_opaque - | EvalVarRef id -> - let (ids,sps) = Conv_oracle.freeze() in - Idpred.mem id ids & - let (_,value,_) = Environ.lookup_named id env in - value <> None +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 + +let value_of_evaluable_ref env = function + | EvalConstRef con -> constant_value env con + | EvalVarRef id -> Option.get (pi2 (lookup_named id env)) + +let constr_of_evaluable_ref = function + | EvalConstRef con -> mkConst con + | EvalVarRef id -> mkVar id type evaluable_reference = | EvalConst of constant @@ -80,7 +85,7 @@ let reference_opt_value sigma env = function v | EvalRel n -> let (_,v,_) = lookup_rel n env in - option_map (lift n) v + Option.map (lift n) v | EvalEvar ev -> Evd.existential_opt_value sigma ev exception NotEvaluable @@ -90,8 +95,8 @@ let reference_value sigma env c = | Some d -> d (************************************************************************) -(* Reduction of constant hiding fixpoints (e.g. for Simpl). The trick *) -(* is to reuse the name of the function after reduction of the fixpoint *) +(* 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 = | EliminationFix of int * (int * (int * constr) list * int) @@ -103,7 +108,6 @@ type constant_evaluation = (* We use a cache registered as a global table *) - module CstOrdered = struct type t = constant @@ -325,20 +329,18 @@ let make_elim_fun (names,(nbfix,lv,n)) largs = list_fold_left_i (fun q (* j from comment is 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) + mkLambda (x,tij',c)) 1 body (List.rev lv) in Some g -(* [f] is convertible to [Fix(recindices,bodynum),bodyvect)] make - the reduction using this extra information *) +(* [f] is convertible to [Fix(recindices,bodynum),bodyvect)]: + do so that the reduction uses this extra information *) let contract_fix_use_function f - ((recindices,bodynum),(types,names,bodies as typedbodies)) = + ((recindices,bodynum),(_names,_types,bodies as typedbodies)) = let nbodies = Array.length recindices in let make_Fi j = match f j with | None -> mkFix((recindices,j),typedbodies) | Some c -> c in -(* match List.nth names j with Name id -> f id | _ -> assert false in*) let lbodies = list_tabulate make_Fi nbodies in substl (List.rev lbodies) bodies.(bodynum) @@ -378,7 +380,7 @@ let reduce_mind_case_use_function func env mia = fun i -> if i = bodynum then Some func else match names.(i) with - | Anonymous -> None + | Anonymous -> None | Name id -> (* In case of a call to another component of a block of mutual inductive, try to reuse the global name if @@ -419,20 +421,23 @@ let special_red_case sigma env whfun (ci, p, c, lf) = in redrec (c, empty_stack) +(* [red_elim_const] contracts iota/fix/cofix redexes hidden behind + constants by keeping the name of the constants in the recursive calls; + it fails if no redex is around *) let rec red_elim_const env sigma ref largs = match reference_eval sigma env ref with | EliminationCases n when stack_args_size largs >= n -> let c = reference_value sigma env ref in let c', lrest = whd_betadelta_state env sigma (c,largs) in - (special_red_case sigma env (construct_const env sigma) (destCase c'), - lrest) + let whfun = whd_simpl_state env sigma in + (special_red_case sigma env whfun (destCase c'), lrest) | EliminationFix (min,infos) when stack_args_size largs >=min -> let c = reference_value sigma env ref in let d, lrest = whd_betadelta_state env sigma (c,largs) in let f = make_elim_fun ([|Some ref|],infos) largs in - let co = construct_const env sigma in - (match reduce_fix_use_function f co (destFix d) lrest with + let whfun = whd_construct_state env sigma in + (match reduce_fix_use_function f whfun (destFix d) lrest with | NotReducible -> raise Redelimination | Reduced (c,rest) -> (nf_beta c, rest)) | EliminationMutualFix (min,refgoal,refinfos) @@ -447,53 +452,65 @@ let rec red_elim_const env sigma ref largs = let (_, midargs as s) = descend ref largs in let d, lrest = whd_betadelta_state env sigma s in let f = make_elim_fun refinfos midargs in - let co = construct_const env sigma in - (match reduce_fix_use_function f co (destFix d) lrest with + let whfun = whd_construct_state env sigma in + (match reduce_fix_use_function f whfun (destFix d) lrest with | NotReducible -> raise Redelimination | Reduced (c,rest) -> (nf_beta c, rest)) | _ -> raise Redelimination -and construct_const env sigma = - let rec hnfstack (x, stack as s) = +(* reduce to whd normal form or to an applied constant that does not hide + a reducible iota/fix/cofix redex (the "simpl" tactic) *) + +and whd_simpl_state env sigma s = + let rec redrec (x, stack as s) = match kind_of_term x with - | Cast (c,_,_) -> hnfstack (c, stack) - | App (f,cl) -> hnfstack (f, append_stack cl stack) - | Lambda (id,t,c) -> + | Lambda (na,t,c) -> (match decomp_stack stack with - | None -> assert false - | Some (c',rest) -> - stacklam hnfstack [c'] c rest) - | LetIn (n,b,t,c) -> stacklam hnfstack [b] c stack + | 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) -> - hnfstack - (special_red_case sigma env - (construct_const env sigma) (ci,p,c,lf), stack) - | Construct _ -> s - | CoFix _ -> s - | Fix fix -> - (match reduce_fix hnfstack fix stack with - | Reduced s' -> hnfstack s' - | NotReducible -> raise Redelimination) + (try + redrec (special_red_case sigma env redrec (ci,p,c,lf), stack) + with + Redelimination -> s) + | Fix fix -> + (try match reduce_fix (whd_construct_state env sigma) fix stack with + | Reduced s' -> redrec s' + | NotReducible -> s + with Redelimination -> s) | _ when isEvalRef env x -> let ref = destEvalRef x in (try - hnfstack (red_elim_const env sigma ref stack) + redrec (red_elim_const env sigma ref stack) with Redelimination -> - (match reference_opt_value sigma env ref with - | Some cval -> - (match kind_of_term cval with - | CoFix _ -> s - | _ -> hnfstack (cval, stack)) - | None -> - raise Redelimination)) - | _ -> raise Redelimination + s) + | _ -> s in - hnfstack + 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 + if reducible_mind_case constr then s' + else if isEvalRef env constr then + let ref = destEvalRef constr in + match reference_opt_value sigma env ref with + | None -> raise Redelimination + | Some gvalue -> whd_construct_state env sigma (gvalue, cargs) + else + raise Redelimination (************************************************************************) (* Special Purpose Reduction Strategies *) -(* Red reduction tactic: reduction to a product *) +(* Red reduction tactic: one step of delta reduction + full + beta-iota-fix-cofix-zeta-cast at the head of the conclusion of a + sequence of products; fails if no delta redex is around +*) let try_red_product env sigma c = let simpfun = clos_norm_flags betaiotazeta env sigma in @@ -528,97 +545,105 @@ let red_product env sigma c = try try_red_product env sigma c with Redelimination -> error "Not reducible" -let hnf_constr env sigma c = - let rec redrec (x, largs as s) = +(* +(* 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. + + returned + + (fix plus (n m : nat) {struct n} : nat := + match n with + | O => m + | S p => S (plus p m) + end) n 0 + + while the new version returns (plus n O) + *) + +let whd_simpl_orelse_delta_but_fix_old env sigma c = + let whd_all = whd_betadeltaiota_state env sigma in + let rec redrec (x, stack as s) = match kind_of_term x with - | Lambda (n,t,c) -> - (match decomp_stack largs with - | None -> app_stack s - | Some (a,rest) -> - stacklam redrec [a] c rest) - | LetIn (n,b,t,c) -> stacklam redrec [b] c largs - | App (f,cl) -> redrec (f, append_stack cl largs) - | Cast (c,_,_) -> redrec (c, largs) - | Case (ci,p,c,lf) -> + | Lambda (na,t,c) -> + (match decomp_stack stack with + | None -> s + | Some (a,rest) -> stacklam redrec [a] c rest) + | LetIn (n,b,t,c) -> stacklam redrec [b] c stack + | App (f,cl) -> redrec (f, append_stack cl stack) + | Cast (c,_,_) -> redrec (c, stack) + | Case (ci,p,d,lf) -> (try - redrec - (special_red_case sigma env (whd_betadeltaiota_state env sigma) - (ci, p, c, lf), largs) + redrec (special_red_case sigma env whd_all (ci,p,d,lf), stack) with Redelimination -> - app_stack s) + s) | Fix fix -> - (match reduce_fix (whd_betadeltaiota_state env sigma) fix largs with + (match reduce_fix whd_all fix stack with | Reduced s' -> redrec s' - | NotReducible -> app_stack s) + | NotReducible -> s) | _ when isEvalRef env x -> let ref = destEvalRef x in (try - let (c',lrest) = red_elim_const env sigma ref largs in - redrec (c', lrest) + redrec (red_elim_const env sigma ref stack) with Redelimination -> match reference_opt_value sigma env ref with | Some c -> (match kind_of_term (snd (decompose_lam c)) with - | CoFix _ | Fix _ -> app_stack (x,largs) - | _ -> redrec (c, largs)) - | None -> app_stack s) - | _ -> app_stack s - in - redrec (c, empty_stack) + | CoFix _ | Fix _ -> s + | _ -> redrec (c, stack)) + | None -> s) + | _ -> s + in app_stack (redrec (c, empty_stack)) +*) -(* Simpl reduction tactic: same as simplify, but also reduces - elimination constants *) +(* Same as [whd_simpl] but also reduces constants that do not hide a + reducible fix, but does this reduction of constants only until it + it immediately hides a non reducible fix or a cofix *) -let whd_nf env sigma c = - let rec nf_app (c, stack as s) = - match kind_of_term c with - | Lambda (name,c1,c2) -> - (match decomp_stack stack with - | None -> (c,empty_stack) - | Some (a1,rest) -> - stacklam nf_app [a1] c2 rest) - | LetIn (n,b,t,c) -> stacklam nf_app [b] c stack - | App (f,cl) -> nf_app (f, append_stack cl stack) - | Cast (c,_,_) -> nf_app (c, stack) - | Case (ci,p,d,lf) -> - (try - nf_app (special_red_case sigma env nf_app (ci,p,d,lf), stack) - with Redelimination -> - s) - | Fix fix -> - (match reduce_fix nf_app fix stack with - | Reduced s' -> nf_app s' - | NotReducible -> s) - | _ when isEvalRef env c -> - (try - nf_app (red_elim_const env sigma (destEvalRef c) stack) - with Redelimination -> - s) - | _ -> s - in - app_stack (nf_app (c, empty_stack)) +let whd_simpl_orelse_delta_but_fix env sigma c = + let rec redrec s = + let (constr, stack as s') = whd_simpl_state env sigma s in + if isEvalRef env constr then + match reference_opt_value sigma env (destEvalRef constr) with + | Some c -> + (match kind_of_term (snd (decompose_lam c)) with + | CoFix _ | Fix _ -> s' + | _ -> redrec (c, stack)) + | None -> s' + else s' + in app_stack (redrec (c, empty_stack)) + +let hnf_constr = whd_simpl_orelse_delta_but_fix + +(* The "simpl" reduction tactic *) + +let whd_simpl env sigma c = + app_stack (whd_simpl_state env sigma (c, empty_stack)) -let nf env sigma c = strong whd_nf 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 = match kind_of_term t with | App (f,_) -> f = c | _ -> false -let contextually byhead (locs,c) f env sigma t = +let contextually byhead ((nowhere_except_in,locs),c) f env sigma t = let maxocc = List.fold_right max locs 0 in let pos = ref 1 in - let except = List.exists (fun n -> n<0) locs in - if except & (List.exists (fun n -> n>=0) locs) - then error "mixing of positive and negative occurences" - else - let rec traverse (env,c as envc) t = - if locs <> [] & (not except) & (!pos > maxocc) then 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 = - if except then not (List.mem (- !pos) locs) - else (locs = [] or List.mem !pos locs) in + 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 @@ -634,110 +659,34 @@ let contextually byhead (locs,c) f env sigma t = traverse envc t in let t' = traverse (env,c) t in - if locs <> [] & List.exists (fun o -> o >= !pos or o <= - !pos) locs then - error_invalid_occurrence locs; + if List.exists (fun o -> o >= !pos) locs then error_invalid_occurrence locs; t' (* linear bindings (following pretty-printer) of the value of name in c. * n is the number of the next occurence of name. * ol is the occurence list to find. *) -let rec substlin env name n ol c = - match kind_of_term c with - | Const kn when EvalConstRef kn = name -> - if List.hd ol = n then - try - (n+1, List.tl ol, constant_value env kn) - with - NotEvaluableConst _ -> - errorlabstrm "substlin" - (pr_con kn ++ str " is not a defined constant") - else - ((n+1), ol, c) - - | Var id when EvalVarRef id = name -> - if List.hd ol = n then - match lookup_named id env with - | (_,Some c,_) -> (n+1, List.tl ol, c) - | _ -> - errorlabstrm "substlin" - (pr_id id ++ str " is not a defined constant") - else - ((n+1), ol, c) - - (* INEFFICIENT: OPTIMIZE *) - | App (c1,cl) -> - Array.fold_left - (fun (n1,ol1,c1') c2 -> - (match ol1 with - | [] -> (n1,[],applist(c1',[c2])) - | _ -> - let (n2,ol2,c2') = substlin env name n1 ol1 c2 in - (n2,ol2,applist(c1',[c2'])))) - (substlin env name n ol c1) cl - - | Lambda (na,c1,c2) -> - let (n1,ol1,c1') = substlin env name n ol c1 in - (match ol1 with - | [] -> (n1,[],mkLambda (na,c1',c2)) - | _ -> - let (n2,ol2,c2') = substlin env name n1 ol1 c2 in - (n2,ol2,mkLambda (na,c1',c2'))) - - | LetIn (na,c1,t,c2) -> - let (n1,ol1,c1') = substlin env name n ol c1 in - (match ol1 with - | [] -> (n1,[],mkLetIn (na,c1',t,c2)) - | _ -> - let (n2,ol2,c2') = substlin env name n1 ol1 c2 in - (n2,ol2,mkLetIn (na,c1',t,c2'))) - - | Prod (na,c1,c2) -> - let (n1,ol1,c1') = substlin env name n ol c1 in - (match ol1 with - | [] -> (n1,[],mkProd (na,c1',c2)) - | _ -> - let (n2,ol2,c2') = substlin env name n1 ol1 c2 in - (n2,ol2,mkProd (na,c1',c2'))) - - | Case (ci,p,d,llf) -> - let rec substlist nn oll = function - | [] -> (nn,oll,[]) - | f::lfe -> - let (nn1,oll1,f') = substlin env name nn oll f in - (match oll1 with - | [] -> (nn1,[],f'::lfe) - | _ -> - let (nn2,oll2,lfe') = substlist nn1 oll1 lfe in - (nn2,oll2,f'::lfe')) - in - (* p is printed after d in v8 syntax *) - let (n1,ol1,d') = substlin env name n ol d in - (match ol1 with - | [] -> (n1,[],mkCase (ci, p, d', llf)) - | _ -> - let (n2,ol2,p') = substlin env name n1 ol1 p in - (match ol2 with - | [] -> (n2,[],mkCase (ci, p', d', llf)) - | _ -> - let (n3,ol3,lf') = substlist n2 ol2 (Array.to_list llf) - in (n3,ol3,mkCase (ci, p', d', Array.of_list lf')))) - - | Cast (c1,k,c2) -> - let (n1,ol1,c1') = substlin env name n ol c1 in - (match ol1 with - | [] -> (n1,[],mkCast (c1',k,c2)) - | _ -> - let (n2,ol2,c2') = substlin env name n1 ol1 c2 in - (n2,ol2,mkCast (c1',k,c2'))) - - | Fix _ -> - (warning "do not consider occurrences inside fixpoints"; (n,ol,c)) - - | CoFix _ -> - (warning "do not consider occurrences inside cofixpoints"; (n,ol,c)) - | (Rel _|Meta _|Var _|Sort _ - |Evar _|Const _|Ind _|Construct _) -> (n,ol,c) +let substlin env evalref n (nowhere_except_in,locs) c = + let maxocc = List.fold_right max locs 0 in + let pos = ref n in + assert (List.for_all (fun x -> x >= 0) locs); + let value = value_of_evaluable_ref env evalref in + let term = constr_of_evaluable_ref evalref in + let rec substrec () c = + if nowhere_except_in & !pos > maxocc then c + else if c = term then + let ok = + if nowhere_except_in then List.mem !pos locs + else not (List.mem !pos locs) in + incr pos; + if ok then value else c + else + map_constr_with_binders_left_to_right + (fun _ () -> ()) + substrec () c + in + let t' = substrec () c in + (!pos, t') let string_of_evaluable_ref env = function | EvalVarRef id -> string_of_id id @@ -755,16 +704,15 @@ let unfold env sigma name = * Unfolds the constant name in a term c following a list of occurrences occl. * at the occurrences of occ_list. If occ_list is empty, unfold all occurences. * Performs a betaiota reduction after unfolding. *) -let unfoldoccs env sigma (occl,name) c = - match occl with - | [] -> unfold env sigma name c - | l -> - match substlin env name 1 (Sort.list (<) l) c with - | (_,[],uc) -> nf_betaiota uc - | (1,_,_) -> - error ((string_of_evaluable_ref env name)^" does not occur") - | (_,l,_) -> - error_invalid_occurrence l +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 + 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 uc (* Unfold reduction tactic: *) let unfoldn loccname env sigma c = @@ -775,7 +723,17 @@ let fold_one_com com env sigma c = let rcom = try red_product env sigma com with Redelimination -> error "Not reducible" in - subst1 com (subst_term rcom c) + (* Reason first on the beta-iota-zeta normal form of the constant as + unfold produces it, so that the "unfold f; fold f" configuration works + to refold fix expressions *) + let a = subst_term (clos_norm_flags unfold_side_red env sigma rcom) c in + if not (eq_constr a c) then + subst1 com a + else + (* Then reason on the non beta-iota-zeta form for compatibility - + even if it is probably a useless configuration *) + let a = subst_term rcom c in + subst1 com a let fold_commands cl env sigma c = List.fold_right (fun com -> fold_one_com com env sigma) (List.rev cl) c @@ -815,70 +773,81 @@ let pattern_occs loccs_trm env sigma c = (* Used in several tactics. *) +(* 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 rec elimrec env t l = + let t = hnf_constr env sigma t in + match kind_of_term (fst (decompose_app t)) with + | Ind ind-> (ind, it_mkProd_or_LetIn t l) + | Prod (n,ty,t') -> + if allow_product then + elimrec (push_rel (n,None,ty) env) t' ((n,None,ty)::l) + else + errorlabstrm "" (str"Not an inductive definition") + | _ -> + (* Last chance: we allow to bypass the Opaque flag (as it + was partially the case between V5.10 and V8.1 *) + let t' = whd_betadeltaiota env sigma t in + match kind_of_term (fst (decompose_app t')) with + | Ind ind-> (ind, it_mkProd_or_LetIn t' l) + | _ -> errorlabstrm "" (str"Not an inductive product") + in + elimrec env t [] + +let reduce_to_quantified_ind x = reduce_to_ind_gen true x +let reduce_to_atomic_ind x = reduce_to_ind_gen false x + +(* Reduce the weak-head redex [beta,iota/fix/cofix[all],cast,zeta,simpl/delta] + or raise [NotStepReducible] if not a weak-head redex *) + exception NotStepReducible let one_step_reduce env sigma c = - let rec redrec (x, largs) = + let rec redrec (x, stack) = match kind_of_term x with | Lambda (n,t,c) -> - (match decomp_stack largs with + (match decomp_stack stack with | None -> raise NotStepReducible | Some (a,rest) -> (subst1 a c, rest)) - | App (f,cl) -> redrec (f, append_stack cl largs) - | LetIn (_,f,_,cl) -> (subst1 f cl,largs) + | App (f,cl) -> redrec (f, append_stack cl stack) + | LetIn (_,f,_,cl) -> (subst1 f cl,stack) + | Cast (c,_,_) -> redrec (c,stack) | Case (ci,p,c,lf) -> (try - (special_red_case sigma env (whd_betadeltaiota_state env sigma) - (ci,p,c,lf), largs) + (special_red_case sigma env (whd_simpl_state env sigma) + (ci,p,c,lf), stack) with Redelimination -> raise NotStepReducible) | Fix fix -> - (match reduce_fix (whd_betadeltaiota_state env sigma) fix largs with + (match reduce_fix (whd_construct_state env sigma) fix stack with | Reduced s' -> s' | NotReducible -> raise NotStepReducible) - | Cast (c,_,_) -> redrec (c,largs) | _ when isEvalRef env x -> - let ref = - try destEvalRef x - with Redelimination -> raise NotStepReducible in + let ref = destEvalRef x in (try - red_elim_const env sigma ref largs + red_elim_const env sigma ref stack with Redelimination -> match reference_opt_value sigma env ref with - | Some d -> d, largs + | Some d -> d, stack | None -> raise NotStepReducible) | _ -> raise NotStepReducible in app_stack (redrec (c, empty_stack)) -(* put t as t'=(x1:A1)..(xn:An)B with B an inductive definition of name name - return name, B and t' *) +let isIndRef = function IndRef _ -> true | _ -> false -let reduce_to_ind_gen allow_product env sigma t = - let rec elimrec env t l = - let c, _ = Reductionops.whd_stack t in - match kind_of_term c with - | Ind (mind,args) -> ((mind,args),it_mkProd_or_LetIn t l) - | Prod (n,ty,t') -> - if allow_product then - elimrec (push_rel (n,None,t) env) t' ((n,None,ty)::l) - else - errorlabstrm "tactics__reduce_to_mind" - (str"Not an inductive definition") - | _ -> - (try - let t' = nf_betaiota (one_step_reduce env sigma t) in - elimrec env t' l - with NotStepReducible -> - errorlabstrm "tactics__reduce_to_mind" - (str"Not an inductive product")) - in - elimrec env t [] - -let reduce_to_quantified_ind x = reduce_to_ind_gen true x -let reduce_to_atomic_ind x = reduce_to_ind_gen false x - -let reduce_to_ref_gen allow_product env sigma ref t = +let reduce_to_ref_gen allow_product env sigma ref t = + if isIndRef ref then + let (mind,t) = reduce_to_ind_gen allow_product env sigma t in + if IndRef mind <> ref then + errorlabstrm "" (str "Cannot recognize a statement based on " ++ + Nametab.pr_global_env Idset.empty ref) + else + t + else + (* lazily reduces to match the head of [t] with the expected [ref] *) let rec elimrec env t l = let c, _ = Reductionops.whd_stack t in match kind_of_term c with @@ -886,8 +855,9 @@ let reduce_to_ref_gen allow_product env sigma ref t = if allow_product then elimrec (push_rel (n,None,t) env) t' ((n,None,ty)::l) else - errorlabstrm "Tactics.reduce_to_ref_gen" - (str"Not an induction object of atomic type") + errorlabstrm "" + (str "Cannot recognize an atomic statement based on " ++ + Nametab.pr_global_env Idset.empty ref) | _ -> try if global_of_constr c = ref @@ -899,8 +869,8 @@ let reduce_to_ref_gen allow_product env sigma ref t = elimrec env t' l with NotStepReducible -> errorlabstrm "" - (str "Not a statement of conclusion " ++ - Nametab.pr_global_env Idset.empty ref) + (str "Cannot recognize a statement based on " ++ + Nametab.pr_global_env Idset.empty ref) in elimrec env t [] diff --git a/pretyping/tacred.mli b/pretyping/tacred.mli index 691fdf01..e12d6ad2 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 8878 2006-05-30 16:44:25Z herbelin $ i*) +(*i $Id: tacred.mli 11094 2008-06-10 19:35:23Z herbelin $ i*) (*i*) open Names @@ -16,6 +16,7 @@ open Evd open Reductionops open Closure open Rawterm +open Termops (*i*) type reduction_tactic_error = @@ -35,21 +36,25 @@ val red_product : reduction_function (* Red (raise Redelimination if nothing reducible) *) val try_red_product : reduction_function -(* Hnf *) -val hnf_constr : reduction_function - (* Simpl *) -val nf : reduction_function +val simpl : reduction_function + +(* Simpl only at the head *) +val whd_simpl : reduction_function + +(* Hnf: like whd_simpl but force delta-reduction of constants that do + not immediately hide a non reducible fix or cofix *) +val hnf_constr : reduction_function (* Unfold *) val unfoldn : - (int list * evaluable_global_reference) list -> reduction_function + (occurrences * evaluable_global_reference) list -> reduction_function (* Fold *) val fold_commands : constr list -> reduction_function (* Pattern *) -val pattern_occs : (int list * constr) list -> reduction_function +val pattern_occs : (occurrences * constr) list -> reduction_function (* Rem: Lazy strategies are defined in Reduction *) (* Call by value strategy (uses Closures) *) @@ -77,5 +82,10 @@ val reduce_to_quantified_ref : val reduce_to_atomic_ref : env -> evar_map -> Libnames.global_reference -> types -> types -val contextually : bool -> int list * constr -> reduction_function +val contextually : bool -> occurrences * constr -> reduction_function -> reduction_function + +(* Compatibility *) +(* use [simpl] instead of [nf] *) +val nf : reduction_function + diff --git a/pretyping/termops.ml b/pretyping/termops.ml index 6b7e1fb7..e31024bb 100644 --- a/pretyping/termops.ml +++ b/pretyping/termops.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: termops.ml 9429 2006-12-12 08:01:03Z herbelin $ *) +(* $Id: termops.ml 11166 2008-06-22 13:23:35Z herbelin $ *) open Pp open Util @@ -160,18 +160,22 @@ let new_Type_sort () = Type (new_univ ()) (* This refreshes universes in types; works only for inferred types (i.e. for types of the form (x1:A1)...(xn:An)B with B a sort or an atom in head normal form) *) -let refresh_universes t = +let refresh_universes_gen strict t = let modified = ref false in let rec refresh t = match kind_of_term t with - | Sort (Type _) -> modified := true; new_Type () + | Sort (Type u) when strict or u <> Univ.type0m_univ -> + modified := true; new_Type () | Prod (na,u,v) -> mkProd (na,u,refresh v) | _ -> t in let t' = refresh t in if !modified then t' else t +let refresh_universes = refresh_universes_gen false +let refresh_universes_strict = refresh_universes_gen true + let new_sort_in_family = function - | InProp -> mk_Prop - | InSet -> mk_Set + | InProp -> prop_sort + | InSet -> set_sort | InType -> Type (new_univ ()) @@ -214,7 +218,7 @@ let push_named_rec_types (lna,typarray,_) env = array_map2_i (fun i na t -> match na with - | Name id -> (id, None, type_app (lift i) t) + | Name id -> (id, None, lift i t) | Anonymous -> anomaly "Fix declarations must be named") lna typarray in Array.fold_left @@ -271,6 +275,11 @@ let rec strip_head_cast c = match kind_of_term c with | Cast (c,_,_) -> strip_head_cast c | _ -> c +(* Get the last arg of an application *) +let last_arg c = match kind_of_term c with + | App (f,cl) -> array_last cl + | _ -> anomaly "last_arg" + (* [map_constr_with_named_binders g f l c] maps [f l] on the immediate subterms of [c]; it carries an extra data [l] (typically a name list) which is processed by [g na] (which typically cons [na] to @@ -305,11 +314,8 @@ let map_constr_with_named_binders g f l c = match kind_of_term c with (co-)fixpoint) *) let fold_rec_types g (lna,typarray,_) e = - let ctxt = - array_map2_i - (fun i na t -> (na, None, type_app (lift i) t)) lna typarray in - Array.fold_left - (fun e assum -> g assum e) e ctxt + let ctxt = array_map2_i (fun i na t -> (na, None, lift i t)) lna typarray in + Array.fold_left (fun e assum -> g assum e) e ctxt let map_constr_with_binders_left_to_right g f l c = match kind_of_term c with @@ -514,6 +520,16 @@ let free_rels m = in frec 1 Intset.empty m +(* collects all metavar occurences, in left-to-right order, preserving + * repetitions and all. *) + +let collect_metas c = + let rec collrec acc c = + match kind_of_term c with + | Meta mv -> list_add_set mv acc + | _ -> fold_constr collrec acc 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 *) @@ -610,32 +626,34 @@ let subst_term = subst_term_gen eq_constr let replace_term = replace_term_gen eq_constr -(* Substitute only a list of locations locs, the empty list is - interpreted as substitute all, if 0 is in the list then no - bindings is done. The list may contain only negative occurrences - that will not be substituted. *) +(* Substitute only at a list of locations or excluding a list of + locations; in the occurrences list (b,l), b=true means no + occurrence except the ones in l and b=false, means all occurrences + except the ones in l *) + +type occurrences = bool * int list +let all_occurrences = (false,[]) +let no_occurrences_in_set = (true,[]) let error_invalid_occurrence l = + let l = list_uniquize (List.sort Pervasives.compare l) in errorlabstrm "" (str ("Invalid occurrence " ^ plural (List.length l) "number" ^": ") ++ prlist_with_sep spc int l) -let subst_term_occ_gen locs occ c t = +let subst_term_occ_gen (nowhere_except_in,locs) occ c t = let maxocc = List.fold_right max locs 0 in let pos = ref occ in - let except = List.exists (fun n -> n<0) locs in - if except & (List.exists (fun n -> n>=0) locs) - then error "mixing of positive and negative occurences" - else - let rec substrec (k,c as kc) t = - if (not except) & (!pos > maxocc) then t + assert (List.for_all (fun x -> x >= 0) locs); + let rec substrec (k,c as kc) t = + if nowhere_except_in & !pos > maxocc then t else if eq_constr c t then let r = - if except then - if List.mem (- !pos) locs then t else (mkRel k) - else + if nowhere_except_in then if List.mem !pos locs then (mkRel k) else t + else + if List.mem !pos locs then t else (mkRel k) in incr pos; r else map_constr_with_binders_left_to_right @@ -645,40 +663,33 @@ let subst_term_occ_gen locs occ c t = let t' = substrec (1,c) t in (!pos, t') -let subst_term_occ locs c t = - if locs = [] then subst_term c t - else if List.mem 0 locs then - t +let subst_term_occ (nowhere_except_in,locs as plocs) c t = + if locs = [] then if nowhere_except_in then t else subst_term c t else - let (nbocc,t') = subst_term_occ_gen locs 1 c t in - let rest = List.filter (fun o -> o >= nbocc or o <= -nbocc) locs in + let (nbocc,t') = subst_term_occ_gen plocs 1 c t in + let rest = List.filter (fun o -> o >= nbocc) locs in if rest <> [] then error_invalid_occurrence rest; t' -let subst_term_occ_decl locs c (id,bodyopt,typ as d) = +let subst_term_occ_decl (nowhere_except_in,locs as plocs) c (id,bodyopt,typ as d) = match bodyopt with - | None -> (id,None,subst_term_occ locs c typ) + | None -> (id,None,subst_term_occ plocs c typ) | Some body -> if locs = [] then - (id,Some (subst_term c body),type_app (subst_term c) typ) - else if List.mem 0 locs then - d + if nowhere_except_in then d + else (id,Some (subst_term c body),subst_term c typ) else - let (nbocc,body') = subst_term_occ_gen locs 1 c body in - let (nbocc',t') = subst_term_occ_gen locs nbocc c typ in - let rest = List.filter (fun o -> o >= nbocc' or o <= -nbocc') locs in + 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 first_char id = - let id = string_of_id id in - assert (id <> ""); - String.make 1 id.[0] - -let lowercase_first_char id = String.lowercase (first_char id) +let lowercase_first_char id = + lowercase_first_char_utf8 (string_of_id id) let vars_of_env env = let s = @@ -707,8 +718,7 @@ let hdchar env c = | Cast (c,_,_) -> hdrec k c | App (f,l) -> hdrec k f | Const kn -> - let c = lowercase_first_char (id_of_label (con_label kn)) in - if c = "?" then "y" else c + 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)) @@ -743,20 +753,18 @@ let named_hd env a = function | Anonymous -> Name (id_of_string (hdchar env a)) | x -> x -let named_hd_type env a = named_hd env (body_of_type a) - -let mkProd_name env (n,a,b) = mkProd (named_hd_type env a n, a, b) -let mkLambda_name env (n,a,b) = mkLambda (named_hd_type env a n, a, b) +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_type env a Anonymous, a, b) -let lambda_create env (a,b) = mkLambda (named_hd_type env a Anonymous, a, b) +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_type env t na, None, t) + | None -> (named_hd env t na, None, t) | Some body -> (named_hd env body na, c, t) let name_context env hyps = @@ -834,6 +842,14 @@ let is_global id = 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 @@ -860,6 +876,11 @@ let isGlobalRef c = | Const _ | Ind _ | Construct _ | Var _ -> true | _ -> false +let has_polymorphic_type c = + match (Global.lookup_constant c).Declarations.const_type with + | Declarations.PolymorphicArity _ -> true + | _ -> false + (* nouvelle version de renommage des variables (DEC 98) *) (* This is the algorithm to display distinct bound variables @@ -898,11 +919,19 @@ let occur_id nenv id0 c = with Occur -> true | Not_found -> false (* Case when a global is not in the env *) -let next_name_not_occuring is_goal_ccl name l env_names t = +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 - (* To be consistent with intro mechanism *) - (is_goal_ccl & is_global id & not (is_section_variable id)) + (* 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 @@ -914,6 +943,22 @@ let next_name_not_occuring is_goal_ccl name l env_names t = (* 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 *) + | (Prop c1, Type u) -> pb = Reduction.CUMUL + | (Type u1, Type u2) -> true + | _ -> 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 + Sort s1, Sort s2 -> base_sort_cmp cv_pb s1 s2 + | _ -> false) + || compare_constr (constr_cmp cv_pb) t1 t2 + +let eq_constr = constr_cmp Reduction.CONV + (* 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 *) @@ -939,6 +984,7 @@ let rec eta_reduce_head c = | _ -> c) | _ -> c + (* alpha-eta conversion : ignore print names and casts *) let eta_eq_constr = let rec aux t1 t2 = @@ -963,14 +1009,18 @@ let assums_of_rel_context sign = | None -> (na, t)::l) sign ~init:[] -let lift_rel_context n sign = - let rec liftrec k = function - | (na,c,t)::sign -> - (na,option_map (liftn n k) c,type_app (liftn n k) t) - ::(liftrec (k-1) sign) +let map_rel_context_with_binders f sign = + let rec aux k = function + | d::sign -> map_rel_declaration (f k) d :: aux (k-1) sign | [] -> [] in - liftrec (rel_context_length sign) sign + aux (rel_context_length sign) sign + +let substl_rel_context l = + map_rel_context_with_binders (fun k -> substnl l (k-1)) + +let lift_rel_context n = + map_rel_context_with_binders (liftn n) let fold_named_context_both_sides f l ~init = list_fold_right_and_left f l init @@ -996,34 +1046,42 @@ let global_vars_set_of_decl env = function Idset.union (global_vars_set env t) (global_vars_set env c) +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 is_goal_ccl l env_names n c = +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 is_goal_ccl n l env_names c in + 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 is_goal_ccl l env_names n c = - let fresh_id = next_name_not_occuring is_goal_ccl n l env_names c in +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 l c = - match kind_of_term c with - | Prod (Name s,c1,c2) -> - if noccurn 1 c2 then - let env' = push_rel (Name s,None,c1) env in - mkProd (Name s, c1, rename_bound_var env' l c2) - else - let s' = next_ident_away s (global_vars env c2@l) in - let env' = push_rel (Name s',None,c1) env in - mkProd (Name s', c1, rename_bound_var env' (s'::l) c2) - | Prod (Anonymous,c1,c2) -> - let env' = push_rel (Anonymous,None,c1) env in - mkProd (Anonymous, c1, rename_bound_var env' l c2) - | Cast (c,k,t) -> mkCast (rename_bound_var env l c, k,t) - | x -> c +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 *) @@ -1031,3 +1089,13 @@ 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 + variables *) +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) + | (n, (h::t)) -> chop_aux (h::acc) (pred n, t) + | (_, []) -> anomaly "context_chop" + in chop_aux [] (k,ctx) + diff --git a/pretyping/termops.mli b/pretyping/termops.mli index e406b148..94aff66f 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 9314 2006-10-29 20:11:08Z herbelin $ i*) +(*i $Id: termops.mli 11166 2008-06-22 13:23:35Z herbelin $ i*) open Util open Pp @@ -21,6 +21,7 @@ val new_sort_in_family : sorts_family -> sorts val new_Type : unit -> types val new_Type_sort : unit -> sorts val refresh_universes : types -> types +val refresh_universes_strict : types -> types (* printers *) val print_sort : sorts -> std_ppcmds @@ -30,6 +31,7 @@ val set_print_constr : (env -> constr -> std_ppcmds) -> unit val print_constr : constr -> std_ppcmds val print_constr_env : env -> constr -> std_ppcmds val print_named_context : env -> std_ppcmds +val pr_rel_decl : env -> rel_declaration -> std_ppcmds val print_rel_context : env -> std_ppcmds val print_env : env -> std_ppcmds @@ -100,7 +102,8 @@ val occur_var_in_decl : identifier -> 'a * types option * types -> bool val occur_term : constr -> constr -> bool val free_rels : constr -> Intset.t - +val dependent : constr -> constr -> bool +val collect_metas : constr -> int list (* Substitution of metavariables *) type metamap = (metavariable * constr) list val subst_meta : metamap -> constr -> constr @@ -108,37 +111,62 @@ val subst_meta : metamap -> constr -> constr (* [pop c] lifts by -1 the positive indexes in [c] *) val pop : constr -> constr -(* bindings of an arbitrary large term. Uses equality modulo +(*s Substitution of an arbitrary large term. Uses equality modulo reduction of let *) -val dependent : constr -> constr -> bool + +(* [subst_term_gen eq d c] replaces [Rel 1] by [d] in [c] using [eq] + as equality *) val subst_term_gen : (constr -> constr -> bool) -> constr -> constr -> constr + +(* [replace_term_gen eq d e c] replaces [d] by [e] in [c] using [eq] + as equality *) val replace_term_gen : (constr -> constr -> bool) -> constr -> constr -> constr -> constr + +(* [subst_term d c] replaces [Rel 1] by [d] in [c] *) val subst_term : constr -> constr -> constr + +(* [replace_term d e c] replaces [d] by [e] in [c] *) val replace_term : constr -> constr -> constr -> constr -val subst_term_occ_gen : - int list -> int -> constr -> types -> int * types -val subst_term_occ : int list -> constr -> types -> types + +(* In occurrences sets, false = everywhere except and true = nowhere except *) +type occurrences = bool * int list +val all_occurrences : occurrences +val no_occurrences_in_set : occurrences + +(* [subst_term_occ_gen occl n c d] replaces occurrences of [c] at + positions [occl], counting from [n], by [Rel 1] in [d] *) +val subst_term_occ_gen : + occurrences -> int -> constr -> types -> int * types + +(* [subst_term_occ occl c d] replaces occurrences of [c] at + positions [occl] by [Rel 1] in [d] (see also Note OCC) *) +val subst_term_occ : occurrences -> constr -> constr -> constr + +(* [subst_term_occ_decl occl c decl] replaces occurrences of [c] at + positions [occl] by [Rel 1] in [decl] *) val subst_term_occ_decl : - int list -> constr -> named_declaration -> named_declaration + occurrences -> 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 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 first_char : identifier -> string 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 named_hd_type : env -> types -> name -> name val mkProd_name : env -> name * types * types -> types val mkLambda_name : env -> name * types * constr -> constr @@ -157,6 +185,9 @@ 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 +(* Get the last arg of a constr intended to be nn application *) +val last_arg : constr -> constr + (* name contexts *) type names_context = name list val add_name : name -> names_context -> names_context @@ -168,6 +199,8 @@ val ids_of_named_context : named_context -> identifier list val ids_of_context : env -> identifier list val names_of_rel_context : env -> names_context +val context_chop : int -> rel_context -> (rel_context*rel_context) + (* Set of local names *) val vars_of_env: env -> Idset.t val add_vname : Idset.t -> name -> Idset.t @@ -177,21 +210,31 @@ 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 : - bool -> name -> identifier list -> name list -> constr -> identifier + avoid_flags -> name -> identifier list -> name list -> constr -> identifier val concrete_name : - bool -> identifier list -> name list -> name -> constr -> + avoid_flags -> identifier list -> name list -> name -> constr -> name * identifier list val concrete_let_name : - bool -> identifier list -> name list -> name -> constr -> name * identifier list + 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 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 @@ -206,6 +249,8 @@ val is_section_variable : identifier -> bool val isGlobalRef : constr -> bool +val has_polymorphic_type : constant -> bool + (* Combinators on judgments *) val on_judgment : (types -> types) -> unsafe_judgment -> unsafe_judgment diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml new file mode 100644 index 00000000..a2392033 --- /dev/null +++ b/pretyping/typeclasses.ml @@ -0,0 +1,457 @@ +(* -*- compile-command: "make -C .. bin/coqtop.byte" -*- *) +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Gmap.add k v acc) old ne + +let cmap_union = Cmap.fold Cmap.add + +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) + Gmap.empty ne + in + Gmap.fold (fun cl insts acc -> + if Gmap.mem cl ne' then acc + else Gmap.add cl insts acc) + ne' old + +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 (_, (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 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_named ctx = + list_smartmap (fun (na, b, t) -> + (na, Option.smartmap do_subst b, do_subst t)) + ctx + in + let do_subst_ctx ctx = + list_smartmap (fun (cl, (na, b, t)) -> + (Option.smartmap (fun (gr,b) -> do_subst_gr gr, b) cl, + (na, Option.smartmap do_subst b, do_subst t))) + ctx + in + let do_subst_projs projs = list_smartmap (fun (x, y) -> (x, do_subst_con y)) projs in + let subst_class k cl classes = + let k = do_subst_gr k in + let cl' = { cl with cl_impl = k; + cl_context = do_subst_ctx cl.cl_context; + cl_props = do_subst_named 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 discharge (_,(cl,m,inst)) = + let discharge_named (cl, r) = + Option.smartmap (fun (gr, b) -> Lib.discharge_global gr, b) cl, r + 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 + 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 + { cl with cl_impl = cl_impl'; + cl_context = + List.map (fun x -> None, x) (abs_context cl) @ + (list_smartmap discharge_named cl.cl_context); + cl_projs = list_smartmap (fun (x, y) -> x, 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 + | None -> insts + | Some 0 -> Cmap.add k is insts + | Some n -> + add_instance_hint is.is_impl is.is_pri; + let is' = { is with is_global = Some (pred n) } + in Cmap.add k is' insts) insts Cmap.empty) + inst + in (cl, m, inst) + +let (input,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 -> Gmap.add (snd x) c.cl_impl acc) !methods c.cl_projs; + update () + +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 cl.cl_context) 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 + with Not_found -> + Gmap.add x (Cmap.add y.is_impl y Cmap.empty) m + +let cmap_elements c = Cmap.fold (fun k v acc -> v :: acc) 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 instances r = + let cl = class_info r in instances_of cl + +let solve_instanciation_problem = ref (fun _ _ _ _ -> assert false) +let solve_instanciations_problem = ref (fun _ _ _ _ _ -> assert false) + +let resolve_typeclass env ev evi (evd, defined as acc) = + try + assert(evi.evar_body = Evar_empty); + !solve_instanciation_problem env evd ev evi + with Exit -> acc + +let resolve_one_typeclass env types = + try + let evi = Evd.make_evar (Environ.named_context_val env) types in + let ev = 1 in + let evm = Evd.add Evd.empty ev evi in + let evd = Evd.create_evar_defs evm in + let evd', b = !solve_instanciation_problem env evd ev evi in + if b then + let evm' = Evd.evars_of evd' in + match Evd.evar_body (Evd.find evm' ev) with + Evar_empty -> raise Not_found + | Evar_defined c -> c + else raise Not_found + with Exit -> raise Not_found + +let resolve_one_typeclass_evd env evd types = + try + let ev = Evarutil.e_new_evar evd env types in + let (ev,_) = destEvar ev in + let evd', b = + !solve_instanciation_problem env !evd ev (Evd.find (Evd.evars_of !evd) ev) + in + evd := evd'; + if b then + let evm' = Evd.evars_of evd' in + match Evd.evar_body (Evd.find evm' ev) with + Evar_empty -> raise Not_found + | Evar_defined c -> Evarutil.nf_evar evm' c + else raise Not_found + with Exit -> raise Not_found + +let method_typeclass ref = + match ref with + | ConstRef c -> + class_info (Gmap.find c !methods) + | _ -> raise Not_found + +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 -> + (match Decls.constant_kind c with + | IsDefinition Instance -> true + | _ -> false) + | VarRef v -> + (match Decls.variable_kind v with + | IsDefinition Instance -> true + | _ -> false) + | _ -> false + +let is_implicit_arg k = + match k with + ImplicitArg (ref, (n, id)) -> true + | InternalHole -> true + | _ -> false + +let destClassApp c = + match kind_of_term c with + | App (ki, args) when isInd ki -> + Some (destInd ki, args) + | _ when isInd c -> Some (destInd c, [||]) + | _ -> None + +let is_independent evm ev = + Evd.fold (fun ev' evi indep -> indep && + (ev = ev' || + evi.evar_body <> Evar_empty || + not (Termops.occur_evar ev evi.evar_concl))) + evm true + + +(* try !solve_instanciations_problem env (Evarutil.nf_evar_defs evd) *) +(* with _ -> *) +(* let evm = Evd.evars_of evd in *) +(* let tc_evars = *) +(* let f ev evi acc = *) +(* let (l, k) = Evd.evar_source ev evd in *) +(* if not check || is_implicit_arg k then *) +(* match destClassApp evi.evar_concl with *) +(* | Some (i, args) when is_class i -> *) +(* Evd.add acc ev evi *) +(* | _ -> acc *) +(* else acc *) +(* in Evd.fold f evm Evd.empty *) +(* in *) +(* let rec sat evars = *) +(* let evm = Evd.evars_of evars in *) +(* let (evars', progress) = *) +(* Evd.fold *) +(* (fun ev evi acc -> *) +(* if (Evd.mem tc_evars ev || not (Evd.mem evm ev)) *) +(* && evi.evar_body = Evar_empty then *) +(* resolve_typeclass env ev evi acc *) +(* else acc) *) +(* evm (evars, false) *) +(* in *) +(* if not progress then evars' *) +(* else *) +(* sat (Evarutil.nf_evar_defs evars') *) +(* in sat (Evarutil.nf_evar_defs evd) *) + +let class_of_constr c = + let extract_cl c = + try Some (class_info (global_of_constr c)) with _ -> None + in + match kind_of_term c with + App (c, _) -> extract_cl c + | _ -> extract_cl c + +let dest_class_app c = + let cl c = class_info (global_of_constr c) in + match kind_of_term c with + App (c, args) -> cl c, args + | _ -> cl c, [||] + +(* To embed a boolean for resolvability status. + This is essentially a hack to mark which evars correspond to + goals and do not need to be resolved when we have nested [resolve_all_evars] + calls (e.g. when doing apply in an External hint in typeclass_instances). + Would be solved by having real evars-as-goals. *) + +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 = + { 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 has_typeclasses evd = + Evd.fold (fun ev evi has -> has || + (evi.evar_body = Evar_empty && class_of_constr evi.evar_concl <> None + && is_resolvable evi)) + evd 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 (Evarutil.nf_evar_defs evd) onlyargs split fail + +type substitution = (identifier * constr) list + +let substitution_of_named_context isevars env id n subst l = + List.fold_right + (fun (na, _, t) subst -> + let t' = replace_vars subst t in + let b = Evarutil.e_new_evar isevars env ~src:(dummy_loc, ImplicitArg (VarRef id, (n, Some na))) t' in + (na, b) :: subst) + l subst + +let nf_substitution sigma subst = + List.map (function (na, c) -> na, Reductionops.nf_evar sigma c) subst diff --git a/pretyping/typeclasses.mli b/pretyping/typeclasses.mli new file mode 100644 index 00000000..a3118053 --- /dev/null +++ b/pretyping/typeclasses.mli @@ -0,0 +1,105 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* constant + +val new_instance : typeclass -> int option -> bool -> constant -> instance + +val instances : global_reference -> instance list +val typeclasses : unit -> typeclass list +val add_class : typeclass -> unit +val add_instance : instance -> unit + +val register_add_instance_hint : (constant -> int option -> unit) -> unit +val add_instance_hint : constant -> int option -> unit + +val class_info : global_reference -> typeclass (* raises a UserError if not a class *) +val is_class : global_reference -> bool +val class_of_constr : constr -> typeclass option +val dest_class_app : constr -> typeclass * constr array (* raises a UserError if not a class *) + +val is_instance : global_reference -> bool +val is_method : constant -> bool + +(* Returns the term and type for the given instance of the parameters and fields + of the type class. *) +val instance_constructor : typeclass -> constr list -> constr * types + +val resolve_one_typeclass : env -> types -> types (* Raises Not_found *) +val resolve_one_typeclass_evd : env -> evar_defs ref -> types -> types (* Raises Not_found *) +val resolve_typeclass : env -> evar -> evar_info -> evar_defs * bool -> evar_defs * bool + +(* Use evar_extra for marking resolvable evars. *) +val bool_in : bool -> Dyn.t +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 resolve_typeclasses : ?onlyargs:bool -> ?split:bool -> ?fail:bool -> env -> evar_defs -> evar_defs + +val solve_instanciation_problem : (env -> evar_defs -> existential_key -> evar_info -> evar_defs * bool) ref +val solve_instanciations_problem : (env -> evar_defs -> bool -> bool -> bool -> evar_defs) ref + +type substitution = (identifier * constr) list + +val substitution_of_named_context : + evar_defs ref -> env -> identifier -> int -> + substitution -> named_context -> substitution + +val nf_substitution : evar_map -> substitution -> substitution + +val is_implicit_arg : hole_kind -> bool + +(* debug *) + +(* val subst : *) +(* 'a * Mod_subst.substitution * *) +(* ((Libnames.global_reference, typeclass) Gmap.t * 'b * *) +(* ('c, instance list) Gmap.t) -> *) +(* (Libnames.global_reference, typeclass) Gmap.t * 'b * *) +(* ('c, instance list) Gmap.t *) + diff --git a/pretyping/typeclasses_errors.ml b/pretyping/typeclasses_errors.ml new file mode 100644 index 00000000..9faac94d --- /dev/null +++ b/pretyping/typeclasses_errors.ml @@ -0,0 +1,55 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* + 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 mismatched_ctx_inst env c n m = typeclass_error env (MismatchedContextInstance (c, n, m)) diff --git a/pretyping/typeclasses_errors.mli b/pretyping/typeclasses_errors.mli new file mode 100644 index 00000000..6491372f --- /dev/null +++ b/pretyping/typeclasses_errors.mli @@ -0,0 +1,44 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* constr -> 'a + +val unbound_method : env -> global_reference -> identifier located -> 'a + +val no_instance : env -> identifier located -> constr list -> 'a + +val unsatisfiable_constraints : env -> evar_defs -> evar option -> 'a + +val mismatched_ctx_inst : env -> contexts -> constr_expr list -> named_context -> 'a diff --git a/pretyping/typing.ml b/pretyping/typing.ml index 6fa05fa5..43e19ca7 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: typing.ml 9511 2007-01-22 08:27:31Z herbelin $ *) +(* $Id: typing.ml 10785 2008-04-13 21:41:54Z herbelin $ *) open Util open Names @@ -20,14 +20,11 @@ open Inductiveops open Typeops open Evd -let meta_type env mv = +let meta_type evd mv = let ty = - try Evd.meta_ftype env mv - with Not_found -> error ("unknown meta ?"^string_of_int mv) in - meta_instance env ty - -let vect_lift = Array.mapi lift -let vect_lift_type = Array.mapi (fun i t -> type_app (lift i) t) + try Evd.meta_ftype evd mv + with Not_found -> anomaly ("unknown meta ?"^Nameops.string_of_meta mv) in + meta_instance evd ty let constant_type_knowing_parameters env cst jl = let paramstyp = Array.map (fun j -> j.uj_type) jl in @@ -169,8 +166,9 @@ let mcheck env evd c t = let mtype_of env evd c = let j = execute env evd (nf_evar (evars_of evd) c) in - (* No normalization: it breaks Pattern! *) - (*nf_betaiota*) (body_of_type j.uj_type) + (* 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 diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 5fb8054f..fc812594 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: unification.ml 9517 2007-01-22 17:37:58Z herbelin $ *) +(* $Id: unification.ml 11157 2008-06-21 10:45:51Z herbelin $ *) open Pp open Util @@ -23,6 +23,8 @@ open Rawterm open Pattern open Evarutil open Pretype_errors +open Retyping +open Coercion.Default (* 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) *) @@ -31,31 +33,69 @@ let abstract_scheme env c l lname_typ = 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 *) if occur_meta a then lambda_name env (na,ta,t) else lambda_name env (na,ta,subst_term_occ locc a t)) c (List.rev l) lname_typ -let abstract_list_all env sigma typ c l = - let ctxt,_ = decomp_n_prod env sigma (List.length l) typ in - let p = abstract_scheme env c (List.map (function a -> [],a) l) ctxt in +let abstract_list_all env evd typ c l = + let ctxt,_ = decomp_n_prod env (evars_of 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 sigma (Typing.type_of env sigma p) typ then p + if is_conv_leq env (evars_of evd) (Typing.mtype_of env evd p) typ then p else error "abstract_list_all" - with UserError _ -> - raise (PretypeError (env,CannotGeneralize typ)) + with UserError _ | Type_errors.TypeError _ -> + error_cannot_find_well_typed_abstraction env (evars_of evd) p l (**) +(* A refinement of [conv_pb]: the integers tells how many arguments + were applied in the context of the conversion problem; if the number + is non zero, steps of eta-expansion will be allowed +*) + +type conv_pb_up_to_eta = Cumul | ConvUnderApp of int * int + +let topconv = ConvUnderApp (0,0) +let of_conv_pb = function CONV -> topconv | CUMUL -> Cumul +let conv_pb_of = function ConvUnderApp _ -> CONV | Cumul -> CUMUL +let prod_pb = function ConvUnderApp _ -> topconv | pb -> pb + +let opp_status = function + | IsSuperType -> IsSubType + | IsSubType -> IsSuperType + | ConvUpToEta _ | UserGiven as x -> x + +let add_type_status (x,y) = ((x,TypeNotProcessed),(y,TypeNotProcessed)) + +let extract_instance_status = function + | Cumul -> add_type_status (IsSubType, IsSuperType) + | ConvUnderApp (n1,n2) -> add_type_status (ConvUpToEta n1, ConvUpToEta n2) + +let rec assoc_pair x = function + [] -> raise Not_found + | (a,b,_)::l -> if compare a x = 0 then b else assoc_pair x l + +let rec subst_meta_instances bl c = + match kind_of_term c with + | Meta i -> (try assoc_pair i bl with Not_found -> c) + | _ -> map_constr (subst_meta_instances bl) c + let solve_pattern_eqn_array env f l c (metasubst,evarsubst) = match kind_of_term f with | Meta k -> - (k,solve_pattern_eqn env (Array.to_list l) c)::metasubst,evarsubst + 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 + (k,c,pb)::metasubst,evarsubst | Evar ev -> (* Currently unused: incompatible with eauto/eassumption backtracking *) - metasubst,(f,solve_pattern_eqn env (Array.to_list l) c)::evarsubst + metasubst,(ev,solve_pattern_eqn env (Array.to_list l) c)::evarsubst | _ -> assert false (*******************************) @@ -78,79 +118,207 @@ let unify_r2l x = x let sort_eqns = unify_r2l *) -let unify_0 env sigma cv_pb mod_delta m n = - let trivial_unify pb substn m n = - if (not(occur_meta m)) && - (if mod_delta then is_fconv pb env sigma m n else eq_constr m n) - then substn - else error_cannot_unify env sigma (m,n) in - let rec unirec_rec curenv pb ((metasubst,evarsubst) as substn) curm curn = +type unify_flags = { + modulo_conv_on_closed_terms : Names.transparent_state option; + use_metas_eagerly : bool; + modulo_delta : Names.transparent_state; +} + +let default_unify_flags = { + modulo_conv_on_closed_terms = Some full_transparent_state; + use_metas_eagerly = true; + modulo_delta = full_transparent_state; +} + +let default_no_delta_unify_flags = { + modulo_conv_on_closed_terms = Some full_transparent_state; + use_metas_eagerly = true; + modulo_delta = empty_transparent_state; +} + +let expand_constant env flags c = + match kind_of_term c with + | Const cst when is_transparent (ConstKey cst) && + Cpred.mem cst (snd flags.modulo_delta) -> + constant_opt_value env cst + | Var id when is_transparent (VarKey id) && + Idpred.mem id (fst flags.modulo_delta) -> + named_body id env + | _ -> None + +let unify_0_with_initial_metas subst conv_at_top env sigma cv_pb flags m n = + let nb = nb_rel env in + let trivial_unify 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 m -> + (match flags.modulo_conv_on_closed_terms with + Some flags -> + is_trans_fconv (conv_pb_of pb) flags env sigma m n + | None -> constr_cmp (conv_pb_of cv_pb) m n) + | _ -> constr_cmp (conv_pb_of cv_pb) m n in + let rec unirec_rec curenv pb b ((metasubst,evarsubst) as substn) curm curn = let cM = Evarutil.whd_castappevar sigma curm and cN = Evarutil.whd_castappevar 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)::metasubst,evarsubst - else if k1 = k2 then substn else (k2,cM)::metasubst,evarsubst + then (k1,cN,stN)::metasubst,evarsubst + else if k1 = k2 then substn + else (k2,cM,stM)::metasubst,evarsubst | Meta k, _ -> (* Here we check that [cN] does not contain any local variables *) - if (closedn (nb_rel env) cN) - then (k,cN)::metasubst,evarsubst - else error_cannot_unify_local curenv sigma (curenv,m,n,cN) + if (closedn nb cN) + then (k,cN,snd (extract_instance_status pb))::metasubst,evarsubst + else error_cannot_unify_local curenv sigma (m,n,cN) | _, Meta k -> (* Here we check that [cM] does not contain any local variables *) - if (closedn (nb_rel env) cM) - then (k,cM)::metasubst,evarsubst - else error_cannot_unify_local curenv sigma (curenv,m,n,cM) - | Evar _, _ -> metasubst,((cM,cN)::evarsubst) - | _, Evar _ -> metasubst,((cN,cM)::evarsubst) + if (closedn nb cM) + then (k,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) | Lambda (na,t1,c1), Lambda (_,t2,c2) -> - unirec_rec (push_rel_assum (na,t1) curenv) CONV - (unirec_rec curenv CONV substn t1 t2) c1 c2 + unirec_rec (push_rel_assum (na,t1) curenv) topconv true + (unirec_rec curenv topconv true substn t1 t2) c1 c2 | Prod (na,t1,c1), Prod (_,t2,c2) -> - unirec_rec (push_rel_assum (na,t1) curenv) pb - (unirec_rec curenv CONV substn t1 t2) c1 c2 - | LetIn (_,b,_,c), _ -> unirec_rec curenv pb substn (subst1 b c) cN - | _, LetIn (_,b,_,c) -> unirec_rec curenv pb substn cM (subst1 b c) + unirec_rec (push_rel_assum (na,t1) curenv) (prod_pb pb) true + (unirec_rec curenv topconv true substn t1 t2) c1 c2 + | LetIn (_,a,_,c), _ -> unirec_rec curenv pb b substn (subst1 a c) cN + | _, LetIn (_,a,_,c) -> unirec_rec curenv pb b substn cM (subst1 a c) + | Case (_,p1,c1,cl1), Case (_,p2,c2,cl2) -> + array_fold_left2 (unirec_rec curenv topconv true) + (unirec_rec curenv topconv true + (unirec_rec curenv topconv true substn p1 p2) c1 c2) cl1 cl2 + | App (f1,l1), _ when - isMeta f1 & is_unification_pattern f1 l1 & not (dependent f1 cN) -> + isMeta f1 & is_unification_pattern curenv f1 l1 & + not (dependent f1 cN) -> solve_pattern_eqn_array curenv f1 l1 cN substn | _, App (f2,l2) when - isMeta f2 & is_unification_pattern f2 l2 & not (dependent f2 cM) -> + isMeta f2 & is_unification_pattern curenv f2 l2 & + not (dependent f2 cM) -> solve_pattern_eqn_array curenv f2 l2 cM substn | App (f1,l1), App (f2,l2) -> let len1 = Array.length l1 and len2 = Array.length l2 in - let (f1,l1,f2,l2) = - if len1 = len2 then (f1,l1,f2,l2) - else if len1 < len2 then - let extras,restl2 = array_chop (len2-len1) l2 in - (f1, l1, appvect (f2,extras), restl2) - else - let extras,restl1 = array_chop (len1-len2) l1 in - (appvect (f1,extras), restl1, f2, l2) in - (try - array_fold_left2 (unirec_rec curenv CONV) - (unirec_rec curenv CONV substn f1 f2) l1 l2 - with ex when precatchable_exception ex -> - trivial_unify pb substn cM cN) - | Case (_,p1,c1,cl1), Case (_,p2,c2,cl2) -> - array_fold_left2 (unirec_rec curenv CONV) - (unirec_rec curenv CONV - (unirec_rec curenv CONV substn p1 p2) c1 c2) cl1 cl2 - | _ -> trivial_unify pb substn cM cN + (try + let (f1,l1,f2,l2) = + if len1 = len2 then (f1,l1,f2,l2) + else if len1 < len2 then + let extras,restl2 = array_chop (len2-len1) l2 in + (f1, l1, appvect (f2,extras), restl2) + else + let extras,restl1 = array_chop (len1-len2) l1 in + (appvect (f1,extras), restl1, f2, l2) in + let pb = ConvUnderApp (len1,len2) in + array_fold_left2 (unirec_rec curenv topconv true) + (unirec_rec curenv pb true substn f1 f2) l1 l2 + with ex when precatchable_exception ex -> + expand curenv pb b substn cM f1 l1 cN f2 l2) + + | _ -> + 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 curenv pb b substn cM f1 l1 cN f2 l2 + + and expand curenv pb b substn cM f1 l1 cN f2 l2 = + if trivial_unify pb substn cM cN then substn + else if b then + match expand_constant curenv flags f1 with + | Some c -> + unirec_rec curenv pb b substn (whd_betaiotazeta (mkApp(c,l1))) cN + | None -> + match expand_constant curenv flags f2 with + | Some c -> + unirec_rec curenv pb b substn cM (whd_betaiotazeta (mkApp(c,l2))) + | None -> + error_cannot_unify env sigma (cM,cN) + else + error_cannot_unify env sigma (cM,cN) + in if (not(occur_meta m)) && - (if mod_delta then is_fconv cv_pb env sigma m n else eq_constr m n) + (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 - ([],[]) + subst else - let (mc,ec) = unirec_rec env cv_pb ([],[]) m n in - ((*sort_eqns*) mc, (*sort_eqns*) ec) + unirec_rec env cv_pb conv_at_top subst m n + +let unify_0 = unify_0_with_initial_metas ([],[]) true +let left = true +let right = false + +let pop k = if k=0 then 0 else k-1 + +let rec unify_with_eta keptside flags env sigma k1 k2 c1 c2 = + (* Reason up to limited eta-expansion: ci is allowed to start with ki lam *) + (* Question: try whd_betadeltaiota on ci if ki>0 ? *) + 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') = + unify_with_eta keptside flags env' sigma (pop k1) (pop k2) c1' c2' + in (side,status,(metas@metas',evars@evars')) + | (Lambda (na,t,c1'),_) when k2 > 0 -> + let env' = push_rel_assum (na,t) env in + let side = left in (* expansion on the right: we keep the left side *) + unify_with_eta side flags env' sigma (pop k1) (k2-1) + c1' (mkApp (lift 1 c2,[|mkRel 1|])) + | (_,Lambda (na,t,c2')) when k1 > 0 -> + let env' = push_rel_assum (na,t) env in + let side = right in (* expansion on the left: we keep the right side *) + unify_with_eta side flags env' sigma (k1-1) (pop k2) + (mkApp (lift 1 c1,[|mkRel 1|])) c2' + | _ -> + (keptside,ConvUpToEta(min k1 k2), + unify_0 env sigma topconv flags c1 c2) + +(* We solved problems [?n =_pb u] (i.e. [u =_(opp pb) ?n]) and [?n =_pb' u'], + we now compute the problem on [u =? u'] and decide which of u or u' is kept + + Rem: the upper constraint is lost in case u <= ?n <= u' (and symmetrically + in the case u' <= ?n <= u) + *) + +let merge_instances env sigma flags st1 st2 c1 c2 = + match (opp_status st1, st2) with + | (UserGiven, ConvUpToEta n2) -> + unify_with_eta left flags env sigma 0 n2 c1 c2 + | (ConvUpToEta n1, UserGiven) -> + unify_with_eta right flags env sigma n1 0 c1 c2 + | (ConvUpToEta n1, ConvUpToEta n2) -> + let side = left (* arbitrary choice, but agrees with compatibility *) in + unify_with_eta side flags env sigma n1 n2 c1 c2 + | ((IsSubType | ConvUpToEta _ | UserGiven as oppst1), + (IsSubType | ConvUpToEta _ | UserGiven)) -> + let res = unify_0 env sigma Cumul flags c2 c1 in + if oppst1=st2 then (* arbitrary choice *) (left, st1, res) + else if st2=IsSubType or st1=UserGiven then (left, st1, res) + else (right, st2, res) + | ((IsSuperType | ConvUpToEta _ | UserGiven as oppst1), + (IsSuperType | ConvUpToEta _ | UserGiven)) -> + let res = unify_0 env sigma Cumul flags c1 c2 in + if oppst1=st2 then (* arbitrary choice *) (left, st1, res) + else if st2=IsSuperType or st1=UserGiven then (left, st1, res) + else (right, st2, res) + | (IsSuperType,IsSubType) -> + (try (left, IsSubType, unify_0 env sigma Cumul flags c2 c1) + with _ -> (right, IsSubType, unify_0 env sigma Cumul flags c1 c2)) + | (IsSubType,IsSuperType) -> + (try (left, IsSuperType, unify_0 env sigma Cumul flags c1 c2) + with _ -> (right, IsSuperType, unify_0 env sigma Cumul flags c2 c1)) (* Unification * @@ -208,7 +376,8 @@ let applyHead env evd n c = else match kind_of_term (whd_betadeltaiota env (evars_of evd) cty) with | Prod (_,c1,c2) -> - let (evd',evar) = Evarutil.new_evar evd env c1 in + 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 @@ -218,94 +387,167 @@ 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 + (evd',j'.uj_val) + else + error "Cannot solve unification constraints" + +let w_coerce_to_type env evd c cty mvty = + let evd,mvty = pose_all_metas_as_evars env evd mvty in + let tycon = mk_tycon_type mvty in + try try_to_coerce env evd c cty tycon + with e when precatchable_exception e -> + (* 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 + try_to_coerce env evd c cty tycon + +let w_coerce env evd mv c = + let cty = get_type_of env (evars_of 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 c = refresh_universes c in + let t = get_type_of_with_meta env sigma (metas_of evd) c in + let t = Tacred.hnf_constr env sigma (nf_betaiota (nf_meta evd 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 mvty then + unify_to_type env evd flags c mvty + else ([],[]) + +(* Move metas that may need coercion at the end of the list of instances *) + +let order_metas metas = + let rec order latemetas = function + | [] -> List.rev latemetas + | (_,_,(status,to_type) as meta)::metas -> + if to_type = CoerceToType then order (meta::latemetas) metas + else meta :: order latemetas metas + in order [] metas + +(* Solve an equation ?n[x1=u1..xn=un] = t where ?n is an evar *) + +let solve_simple_evar_eqn env evd ev rhs = + let evd,b = solve_simple_eqn Evarconv.evar_conv_x env evd (CONV,ev,rhs) in + if not b then error_cannot_unify env (evars_of 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 + (* [w_merge env sigma b metas evars] merges common instances in metas or in evars, possibly generating new unification problems; if [b] is true, unification of types of metas is required *) -let w_merge env with_types mod_delta metas evars evd = - let ty_metas = ref [] in - let ty_evars = ref [] in - let rec w_merge_rec evd metas evars = - match (evars,metas) with - | ([], []) -> evd - - | ((lhs,rhs)::t, metas) -> - (match kind_of_term rhs with - - | Meta k -> w_merge_rec evd ((k,lhs)::metas) t - - | krhs -> - (match kind_of_term lhs with - - | Evar (evn,_ as ev) -> - if is_defined_evar evd ev then - let (metas',evars') = - unify_0 env (evars_of evd) CONV mod_delta rhs lhs in - w_merge_rec evd (metas'@metas) (evars'@t) - else begin - let rhs' = - if occur_meta rhs then subst_meta metas rhs else rhs - in - if occur_evar evn rhs' then - error "w_merge: recursive equation"; - match krhs with - | App (f,cl) when is_mimick_head f -> - (try - w_merge_rec (fst (evar_define env ev rhs' evd)) metas t - with ex when precatchable_exception ex -> - let evd' = - mimick_evar evd mod_delta f (Array.length cl) evn in - w_merge_rec evd' metas evars) - | _ -> - (* ensure tail recursion in non-mimickable case! *) - w_merge_rec (fst (evar_define env ev rhs' evd)) metas t - end - - | _ -> anomaly "w_merge_rec")) - - | ([], (mv,n)::t) -> - if meta_defined evd mv then - let (metas',evars') = - unify_0 env (evars_of evd) CONV mod_delta - (meta_fvalue evd mv).rebus n in - w_merge_rec evd (metas'@t) evars' - else - begin - if with_types (* or occur_meta mvty *) then - (let mvty = Typing.meta_type evd mv in - try - let sigma = evars_of evd in - (* why not typing with the metamap ? *) - let nty = Typing.type_of env sigma (nf_meta evd n) in - let (mc,ec) = unify_0 env sigma CUMUL mod_delta nty mvty in - ty_metas := mc @ !ty_metas; - ty_evars := ec @ !ty_evars - with e when precatchable_exception e -> ()); - let evd' = meta_assign mv n evd in - w_merge_rec evd' t [] - end - - and mimick_evar evd mod_delta hdc nargs sp = +let w_merge env with_types flags (metas,evars) evd = + 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 + 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'; + 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 + | (mv,c,(status,to_type))::metas -> + let ((evd,c),(metas'',evars'')),eqns = + 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,([],[])),[] + else + (* No coercion needed: delay the unification of types *) + ((evd,c),([],[])),(mv,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 + | [] -> + + (* Process type eqns *) + match eqns with + | (mv,c)::eqns -> + let (metas,evars) = unify_type env evd flags mv 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 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 mod_delta + unify_0 sp_env (evars_of evd') Cumul flags (Retyping.get_type_of sp_env (evars_of evd') c) ev.evar_concl in - let evd'' = w_merge_rec evd' mc ec 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 (* merge constraints *) - let evd' = w_merge_rec evd metas evars in - if with_types then - (* merge constraints about types: if they fail, don't worry *) - try w_merge_rec evd' !ty_metas !ty_evars - with e when precatchable_exception e -> evd' - else - evd' + 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_unify env evd M N] performs a unification of M and N, generating a bunch of @@ -317,9 +559,22 @@ let w_merge env with_types mod_delta metas evars evd = [clenv_typed_unify M N clenv] expects in addition that expected types of metavars are unifiable with the types of their instances *) -let w_unify_core_0 env with_types cv_pb mod_delta m n evd = - let (mc,ec) = unify_0 env (evars_of evd) cv_pb mod_delta m n in - w_merge env with_types mod_delta mc ec evd +let check_types env evd subst m n = + if isEvar (fst (whd_stack m)) or isEvar (fst (whd_stack n)) then + unify_0_with_initial_metas subst true env (evars_of evd) topconv + default_unify_flags + (Retyping.get_type_of_with_meta env (evars_of evd) (metas_of evd) m) + (Retyping.get_type_of_with_meta env (evars_of evd) (metas_of evd) 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 (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' let w_unify_0 env = w_unify_core_0 env false let w_typed_unify env = w_unify_core_0 env true @@ -341,12 +596,12 @@ let iter_fail f a = (* Tries to find an instance of term [cl] in term [op]. Unifies [cl] to every subterm of [op] until it finds a match. Fails if no match is found *) -let w_unify_to_subterm env ?(mod_delta=true) (op,cl) evd = +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 CONV mod_delta op cl evd,cl + then w_unify_0 env topconv flags op cl evd,cl else error "Bound 1" with ex when precatchable_exception ex -> (match kind_of_term cl with @@ -396,9 +651,9 @@ let w_unify_to_subterm env ?(mod_delta=true) (op,cl) evd = in try matchrec cl with ex when precatchable_exception ex -> - raise (PretypeError (env,NoOccurrenceFound op)) + raise (PretypeError (env,NoOccurrenceFound (op, None))) -let w_unify_to_subterm_list env mod_delta allow_K oplist t evd = +let w_unify_to_subterm_list env flags allow_K oplist t evd = List.fold_right (fun op (evd,l) -> if isMeta op then @@ -408,7 +663,7 @@ let w_unify_to_subterm_list env mod_delta allow_K oplist t evd = let (evd',cl) = try (* This is up to delta for subterms w/o metas ... *) - w_unify_to_subterm env ~mod_delta (strip_outer_cast op,t) evd + 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) @@ -416,37 +671,37 @@ let w_unify_to_subterm_list env mod_delta allow_K oplist t evd = (evd,op::l) else (* This is not up to delta ... *) - raise (PretypeError (env,NoOccurrenceFound op))) + raise (PretypeError (env,NoOccurrenceFound (op, None)))) oplist (evd,[]) -let secondOrderAbstraction env mod_delta allow_K typ (p, oplist) evd = - let sigma = evars_of evd in +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 mod_delta allow_K oplist typ evd in + w_unify_to_subterm_list env flags allow_K oplist typ evd in let typp = Typing.meta_type evd' p in - let pred = abstract_list_all env sigma typp typ cllist in - w_unify_0 env CONV mod_delta (mkMeta p) pred evd' + let pred = abstract_list_all env evd' typp typ cllist in + w_merge env false flags ([p,pred,(ConvUpToEta 0,TypeProcessed)],[]) evd' -let w_unify2 env mod_delta allow_K cv_pb ty1 ty2 evd = +let w_unify2 env flags allow_K cv_pb ty1 ty2 evd = let c1, oplist1 = whd_stack ty1 in let c2, oplist2 = whd_stack ty2 in match kind_of_term c1, kind_of_term c2 with | Meta p1, _ -> (* Find the predicate *) let evd' = - secondOrderAbstraction env mod_delta 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 mod_delta (nf_meta evd' ty1) ty2 evd' + w_unify_0 env cv_pb flags (nf_meta evd' ty1) ty2 evd' | _, Meta p2 -> (* Find the predicate *) let evd' = - secondOrderAbstraction env mod_delta 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 mod_delta ty1 (nf_meta evd' ty2) evd' + w_unify_0 env cv_pb flags ty1 (nf_meta evd' ty2) evd' | _ -> error "w_unify2" - (* The unique unification algorithm works like this: If the pattern is flexible, and the goal has a lambda-abstraction at the head, then we do a first-order unification. @@ -467,7 +722,8 @@ let w_unify2 env mod_delta allow_K cv_pb ty1 ty2 evd = Before, second-order was used if the type of Meta(1) and [x:A]t was convertible and first-order otherwise. But if failed if e.g. the type of Meta(1) had meta-variables in it. *) -let w_unify allow_K env cv_pb ?(mod_delta=true) ty1 ty2 evd = +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 ty1 in let hd2,l2 = whd_stack ty2 in match kind_of_term hd1, l1<>[], kind_of_term hd2, l2<>[] with @@ -475,25 +731,22 @@ let w_unify allow_K env cv_pb ?(mod_delta=true) ty1 ty2 evd = | (Meta _, true, Lambda _, _ | Lambda _, _, Meta _, true) when List.length l1 = List.length l2 -> (try - w_typed_unify env cv_pb mod_delta ty1 ty2 evd + w_typed_unify env cv_pb flags ty1 ty2 evd with ex when precatchable_exception ex -> try - w_unify2 env mod_delta allow_K cv_pb ty1 ty2 evd - with PretypeError (env,NoOccurrenceFound c) as e -> raise e - | ex when precatchable_exception ex -> - error "Cannot solve a second-order unification problem") + 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 - w_unify2 env mod_delta allow_K cv_pb ty1 ty2 evd - with PretypeError (env,NoOccurrenceFound c) as e -> raise e + 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 - w_typed_unify env cv_pb mod_delta ty1 ty2 evd - with ex when precatchable_exception ex -> - error "Cannot solve a second-order unification problem") + w_typed_unify env cv_pb flags ty1 ty2 evd + with ex' when precatchable_exception ex' -> + raise ex) (* General case: try first order *) - | _ -> w_unify_0 env cv_pb mod_delta ty1 ty2 evd - + | _ -> w_typed_unify env cv_pb flags ty1 ty2 evd diff --git a/pretyping/unification.mli b/pretyping/unification.mli index 6be530be..89280631 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 6142 2004-09-27 19:33:01Z sacerdot $ i*) +(*i $Id: unification.mli 10856 2008-04-27 16:15:34Z herbelin $ i*) (*i*) open Term @@ -14,20 +14,36 @@ open Environ open Evd (*i*) +type unify_flags = { + modulo_conv_on_closed_terms : Names.transparent_state option; + use_metas_eagerly : bool; + modulo_delta : Names.transparent_state; +} + +val default_unify_flags : unify_flags +val default_no_delta_unify_flags : unify_flags + (* The "unique" unification fonction *) val w_unify : - bool -> env -> conv_pb -> ?mod_delta:bool -> constr -> constr -> evar_defs -> evar_defs + bool -> env -> conv_pb -> ?flags:unify_flags -> constr -> constr -> evar_defs -> evar_defs (* [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 -> ?mod_delta:bool -> constr * constr -> evar_defs -> evar_defs * constr + env -> ?flags:unify_flags -> constr * constr -> evar_defs -> evar_defs * constr + +val w_unify_meta_types : env -> ?flags:unify_flags -> evar_defs -> evar_defs + +(* [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 (*i This should be in another module i*) -(* [abstract_list_all env sigma t c l] *) +(* [abstract_list_all env evd t c l] *) (* abstracts the terms in l over c to get a term of type t *) (* (exported for inv.ml) *) val abstract_list_all : - env -> evar_map -> constr -> constr -> constr list -> constr + env -> evar_defs -> constr -> constr -> constr list -> constr diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml index 46d67406..465c062b 100644 --- a/pretyping/vnorm.ml +++ b/pretyping/vnorm.ml @@ -1,3 +1,13 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* + let mib,mip = lookup_mind_specif env ind in + let nparams = mib.mind_nparams in + let i = invert_tag const tag mip.mind_reloc_tbl in + let params = Array.sub allargs 0 nparams in + let ctyp = type_constructor mind mib (mip.mind_nf_lc.(i-1)) params in + (mkApp(mkConstruct(ind,i), params), ctyp) let construct_of_constr_const env tag typ = fst (construct_of_constr true env tag typ) @@ -196,7 +220,8 @@ and nf_predicate env ind mip params v pT = let name = Name (id_of_string "c") in let n = mip.mind_nrealargs in let rargs = Array.init n (fun i -> mkRel (n-i)) in - let dom = mkApp(mkApp(mkInd ind,params),rargs) in + let params = if n=0 then params else Array.map (lift n) params in + let dom = mkApp(mkInd ind,Array.append params rargs) in let body = nf_vtype (push_rel (name,None,dom) env) vb in true, mkLambda(name,dom,body) | _, _ -> false, nf_val env v crazy_type -- cgit v1.2.3