From 34ef02fac1110673ae74c41c185c228ff7876de2 Mon Sep 17 00:00:00 2001 From: Matej Kosik Date: Fri, 29 Jan 2016 10:13:12 +0100 Subject: CLEANUP: Context.{Rel,Named}.Declaration.t Originally, rel-context was represented as: Context.rel_context = Names.Name.t * Constr.t option * Constr.t Now it is represented as: Context.Rel.t = LocalAssum of Names.Name.t * Constr.t | LocalDef of Names.Name.t * Constr.t * Constr.t Originally, named-context was represented as: Context.named_context = Names.Id.t * Constr.t option * Constr.t Now it is represented as: Context.Named.t = LocalAssum of Names.Id.t * Constr.t | LocalDef of Names.Id.t * Constr.t * Constr.t Motivation: (1) In "tactics/hipattern.ml4" file we define "test_strict_disjunction" function which looked like this: let test_strict_disjunction n lc = Array.for_all_i (fun i c -> match (prod_assum (snd (decompose_prod_n_assum n c))) with | [_,None,c] -> isRel c && Int.equal (destRel c) (n - i) | _ -> false) 0 lc Suppose that you do not know about rel-context and named-context. (that is the case of people who just started to read the source code) Merlin would tell you that the type of the value you are destructing by "match" is: 'a * 'b option * Constr.t (* worst-case scenario *) or Named.Name.t * Constr.t option * Constr.t (* best-case scenario (?) *) To me, this is akin to wearing an opaque veil. It is hard to figure out the meaning of the values you are looking at. In particular, it is hard to discover the connection between the value we are destructing above and the datatypes and functions defined in the "kernel/context.ml" file. In this case, the connection is there, but it is not visible (between the function above and the "Context" module). ------------------------------------------------------------------------ Now consider, what happens when the reader see the same function presented in the following form: let test_strict_disjunction n lc = Array.for_all_i (fun i c -> match (prod_assum (snd (decompose_prod_n_assum n c))) with | [LocalAssum (_,c)] -> isRel c && Int.equal (destRel c) (n - i) | _ -> false) 0 lc If the reader haven't seen "LocalAssum" before, (s)he can use Merlin to jump to the corresponding definition and learn more. In this case, the connection is there, and it is directly visible (between the function above and the "Context" module). (2) Also, if we already have the concepts such as: - local declaration - local assumption - local definition and we describe these notions meticulously in the Reference Manual, then it is a real pity not to reinforce the connection of the actual code with the abstract description we published. --- pretyping/cases.ml | 196 +++++++++++++++++++++++-------------------- pretyping/coercion.ml | 11 ++- pretyping/constr_matching.ml | 15 ++-- pretyping/detyping.ml | 36 ++++++-- pretyping/evarconv.ml | 34 ++++---- pretyping/evarsolve.ml | 131 +++++++++++++++++------------ pretyping/evarutil.ml | 94 +++++++++++++-------- pretyping/find_subterm.ml | 23 ++--- pretyping/indrec.ml | 48 ++++++----- pretyping/inductiveops.ml | 11 +-- pretyping/nativenorm.ml | 21 ++--- pretyping/patternops.ml | 7 +- pretyping/pretyping.ml | 62 +++++++------- pretyping/reductionops.ml | 36 ++++---- pretyping/retyping.ml | 22 ++--- pretyping/tacred.ml | 54 +++++++----- pretyping/typeclasses.ml | 37 ++++---- pretyping/typing.ml | 13 +-- pretyping/unification.ml | 27 +++--- pretyping/vnorm.ml | 14 ++-- 20 files changed, 504 insertions(+), 388 deletions(-) (limited to 'pretyping') diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 2cbf3a265..dd5859092 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -28,6 +28,7 @@ open Evarutil open Evarsolve open Evarconv open Evd +open Context.Rel.Declaration (* Pattern-matching errors *) @@ -272,13 +273,13 @@ let inductive_template evdref env tmloc ind = | None -> fun _ -> (Loc.ghost, Evar_kinds.InternalHole) in let (_,evarl,_) = List.fold_right - (fun (na,b,ty) (subst,evarl,n) -> - match b with - | None -> + (fun decl (subst,evarl,n) -> + match decl with + | LocalAssum (na,ty) -> let ty' = substl subst ty in let e = e_new_evar env evdref ~src:(hole_source n) ty' in (e::subst,e::evarl,n+1) - | Some b -> + | LocalDef (na,b,ty) -> (substl subst b::subst,evarl,n+1)) arsign ([],[],1) in applist (mkIndU indu,List.rev evarl) @@ -306,15 +307,15 @@ let binding_vars_of_inductive = function | NotInd _ -> [] | IsInd (_,IndType(_,realargs),_) -> List.filter isRel realargs -let extract_inductive_data env sigma (_,b,t) = - match b with - | None -> +let extract_inductive_data env sigma decl = + match decl with + | LocalAssum (_,t) -> let tmtyp = try try_find_ind env sigma t None with Not_found -> NotInd (None,t) in let tmtypvars = binding_vars_of_inductive tmtyp in (tmtyp,tmtypvars) - | Some _ -> + | LocalDef (_,_,t) -> (NotInd (None, t), []) let unify_tomatch_with_patterns evdref env loc typ pats realnames = @@ -427,7 +428,7 @@ let remove_current_pattern eqn = let push_current_pattern (cur,ty) eqn = match eqn.patterns with | pat::pats -> - let rhs_env = push_rel (alias_of_pat pat,Some cur,ty) eqn.rhs.rhs_env in + let rhs_env = push_rel (LocalDef (alias_of_pat pat,cur,ty)) eqn.rhs.rhs_env in { eqn with rhs = { eqn.rhs with rhs_env = rhs_env }; patterns = pats } @@ -454,9 +455,9 @@ let prepend_pattern tms eqn = {eqn with patterns = tms@eqn.patterns } exception NotAdjustable let rec adjust_local_defs loc = function - | (pat :: pats, (_,None,_) :: decls) -> + | (pat :: pats, LocalAssum _ :: decls) -> pat :: adjust_local_defs loc (pats,decls) - | (pats, (_,Some _,_) :: decls) -> + | (pats, LocalDef _ :: decls) -> PatVar (loc, Anonymous) :: adjust_local_defs loc (pats,decls) | [], [] -> [] | _ -> raise NotAdjustable @@ -528,9 +529,10 @@ let dependencies_in_pure_rhs nargs eqns = let deps_columns = matrix_transpose deps_rows in List.map (List.exists (fun x -> x)) deps_columns -let dependent_decl a = function - | (na,None,t) -> dependent a t - | (na,Some c,t) -> dependent a t || dependent a c +let dependent_decl a = + function + | LocalAssum (na,t) -> dependent a t + | LocalDef (na,c,t) -> dependent a t || dependent a c let rec dep_in_tomatch n = function | (Pushed _ | Alias _ | NonDepAlias) :: l -> dep_in_tomatch n l @@ -601,7 +603,7 @@ let relocate_index_tomatch n1 n2 = NonDepAlias :: genrec depth rest | Abstract (i,d) :: rest -> let i = relocate_rel n1 n2 depth i in - Abstract (i, Context.Rel.Declaration.map (relocate_index n1 n2 depth) d) + Abstract (i, map_constr (relocate_index n1 n2 depth) d) :: genrec (depth+1) rest in genrec 0 @@ -634,7 +636,7 @@ let replace_tomatch n c = | NonDepAlias :: rest -> NonDepAlias :: replrec depth rest | Abstract (i,d) :: rest -> - Abstract (i, Context.Rel.Declaration.map (replace_term n c depth) d) + Abstract (i, map_constr (replace_term n c depth) d) :: replrec (depth+1) rest in replrec 0 @@ -659,7 +661,7 @@ let rec liftn_tomatch_stack n depth = function NonDepAlias :: liftn_tomatch_stack n depth rest | Abstract (i,d)::rest -> let i = if i let na = merge_name - (fun (na,_,t) -> Name (next_name_away (named_hd env t na) avoid)) + (fun (LocalAssum (na,t) | LocalDef (na,_,t)) -> Name (next_name_away (named_hd env t na) avoid)) d na in (na::l,(out_name na)::avoid)) @@ -727,18 +729,16 @@ let get_names env sign eqns = (* We now replace the names y1 .. yn y by the actual names *) (* xi1 .. xin xi to be found in the i-th clause of the matrix *) -let set_declaration_name x (_,c,t) = (x,c,t) - -let recover_initial_subpattern_names = List.map2 set_declaration_name +let recover_initial_subpattern_names = List.map2 set_name let recover_and_adjust_alias_names names sign = let rec aux = function | [],[] -> [] - | x::names, (_,None,t)::sign -> - (x,(alias_of_pat x,None,t)) :: aux (names,sign) - | names, (na,(Some _ as c),t)::sign -> - (PatVar (Loc.ghost,na),(na,c,t)) :: aux (names,sign) + | x::names, LocalAssum (_,t)::sign -> + (x, LocalAssum (alias_of_pat x,t)) :: aux (names,sign) + | names, (LocalDef (na,_,_) as decl)::sign -> + (PatVar (Loc.ghost,na), decl) :: aux (names,sign) | _ -> assert false in List.split (aux (names,sign)) @@ -753,11 +753,12 @@ let push_rels_eqn_with_names sign eqn = let sign = recover_initial_subpattern_names subpatnames sign in push_rels_eqn sign eqn -let push_generalized_decl_eqn env n (na,c,t) eqn = - let na = match na with - | Anonymous -> Anonymous - | Name id -> pi1 (Environ.lookup_rel n eqn.rhs.rhs_env) in - push_rels_eqn [(na,c,t)] eqn +let push_generalized_decl_eqn env n decl eqn = + match get_name decl with + | Anonymous -> + push_rels_eqn [decl] eqn + | Name _ -> + push_rels_eqn [set_name (get_name (Environ.lookup_rel n eqn.rhs.rhs_env)) decl] eqn let drop_alias_eqn eqn = { eqn with alias_stack = List.tl eqn.alias_stack } @@ -765,7 +766,7 @@ let drop_alias_eqn eqn = let push_alias_eqn alias eqn = let aliasname = List.hd eqn.alias_stack in let eqn = drop_alias_eqn eqn in - let alias = set_declaration_name aliasname alias in + let alias = set_name aliasname alias in push_rels_eqn [alias] eqn (**********************************************************************) @@ -931,7 +932,7 @@ let abstract_predicate env sigma indf cur realargs (names,na) tms ccl = in let pred = extract_predicate ccl tms in (* Build the predicate properly speaking *) - let sign = List.map2 set_declaration_name (na::names) sign in + let sign = List.map2 set_name (na::names) sign in it_mkLambda_or_LetIn_name env pred sign (* [expand_arg] is used by [specialize_predicate] @@ -1117,14 +1118,14 @@ let postprocess_dependencies evd tocheck brs tomatch pred deps cs = let rec aux k brs tomatch pred tocheck deps = match deps, tomatch with | [], _ -> brs,tomatch,pred,[] | n::deps, Abstract (i,d) :: tomatch -> - let d = Context.Rel.Declaration.map (nf_evar evd) d in - let is_d = match d with (_, None, _) -> false | _ -> true in + let d = map_constr (nf_evar evd) d in + let is_d = match d with LocalAssum _ -> false | LocalDef _ -> true in if is_d || List.exists (fun c -> dependent_decl (lift k c) d) tocheck && Array.exists (is_dependent_branch k) brs then (* Dependency in the current term to match and its dependencies is real *) let brs,tomatch,pred,inst = aux (k+1) brs tomatch pred (mkRel n::tocheck) deps in let inst = match d with - | (_, None, _) -> mkRel n :: inst + | LocalAssum _ -> mkRel n :: inst | _ -> inst in brs, Abstract (i,d) :: tomatch, pred, inst @@ -1186,12 +1187,13 @@ let group_equations pb ind current cstrs mat = let rec generalize_problem names pb = function | [] -> pb, [] | i::l -> - let (na,b,t as d) = Context.Rel.Declaration.map (lift i) (Environ.lookup_rel i pb.env) in let pb',deps = generalize_problem names pb l in - begin match (na, b) with - | Anonymous, Some _ -> pb', deps + let d = map_constr (lift i) (Environ.lookup_rel i pb.env) in + begin match d with + | LocalDef (Anonymous,_,_) -> pb', deps | _ -> - let d = on_pi3 (whd_betaiota !(pb.evdref)) d in (* for better rendering *) + (* for better rendering *) + let d = map_type (whd_betaiota !(pb.evdref)) d in let tomatch = lift_tomatch_stack 1 pb'.tomatch in let tomatch = relocate_index_tomatch (i+1) 1 tomatch in { pb' with @@ -1219,7 +1221,8 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn (* that had matched constructor C *) let cs_args = const_info.cs_args in let names,aliasname = get_names pb.env cs_args eqns in - let typs = List.map2 (fun (_,c,t) na -> (na,c,t)) cs_args names in + let typs = List.map2 set_name names cs_args + in (* We build the matrix obtained by expanding the matching on *) (* "C x1..xn as x" followed by a residual matching on eqn into *) @@ -1229,7 +1232,7 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn (* We adjust the terms to match in the context they will be once the *) (* context [x1:T1,..,xn:Tn] will have been pushed on the current env *) let typs' = - List.map_i (fun i d -> (mkRel i, Context.Rel.Declaration.map (lift i) d)) 1 typs in + List.map_i (fun i d -> (mkRel i, map_constr (lift i) d)) 1 typs in let extenv = push_rel_context typs pb.env in @@ -1267,7 +1270,8 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn let typs' = List.map2 - (fun (tm,(tmtyp,_),(na,_,_)) deps -> + (fun (tm, (tmtyp,_), decl) deps -> + let na = get_name decl in let na = match curname, na with | Name _, Anonymous -> curname | Name _, Name _ -> na @@ -1391,7 +1395,7 @@ and shift_problem ((current,t),_,na) pb = let pred = specialize_predicate_var (current,t,na) pb.tomatch pb.pred in let pb = { pb with - env = push_rel (na,Some current,ty) pb.env; + env = push_rel (LocalDef (na,current,ty)) pb.env; tomatch = tomatch; pred = lift_predicate 1 pred tomatch; history = pop_history pb.history; @@ -1439,7 +1443,7 @@ and compile_generalization pb i d rest = ([false]). *) and compile_alias initial pb (na,orig,(expanded,expanded_typ)) rest = let f c t = - let alias = (na,Some c,t) in + let alias = LocalDef (na,c,t) in let pb = { pb with env = push_rel alias pb.env; @@ -1575,9 +1579,9 @@ let adjust_to_extended_env_and_remove_deps env extenv subst t = (* \--------------extenv------------/ *) let (p, _, _) = lookup_rel_id x (rel_context extenv) in let rec traverse_local_defs p = - match pi2 (lookup_rel p extenv) with - | Some c -> assert (isRel c); traverse_local_defs (p + destRel c) - | None -> p in + match lookup_rel p extenv with + | LocalDef (_,c,_) -> assert (isRel c); traverse_local_defs (p + destRel c) + | LocalAssum _ -> p in let p = traverse_local_defs p in let u = lift (n' - n) u in try Some (p, u, expand_vars_in_term extenv u) @@ -1622,7 +1626,7 @@ let abstract_tycon loc env evdref subst tycon extenv t = convertible subterms of the substitution *) let rec aux (k,env,subst as x) t = let t = whd_evar !evdref t in match kind_of_term t with - | Rel n when pi2 (lookup_rel n env) != None -> t + | Rel n when is_local_def (lookup_rel n env) -> t | Evar ev -> let ty = get_type_of env !evdref t in let ty = Evarutil.evd_comb1 (refresh_universes (Some false) env) evdref ty in @@ -1658,7 +1662,8 @@ let abstract_tycon loc env evdref subst tycon extenv t = List.map (fun a -> not (isRel a) || dependent a u || Int.Set.mem (destRel a) depvl) inst in let named_filter = - List.map (fun (id,_,_) -> dependent (mkVar id) u) + let open Context.Named.Declaration in + List.map (fun d -> dependent (mkVar (get_id d)) u) (named_context extenv) in let filter = Filter.make (rel_filter @ named_filter) in let candidates = u :: List.map mkRel vl in @@ -1726,7 +1731,7 @@ let build_inversion_problem loc env sigma tms t = List.rev_append patl patl',acc_sign,acc | (t, NotInd (bo,typ)) :: tms -> let pat,acc = make_patvar t acc in - let d = (alias_of_pat pat,None,typ) in + let d = LocalAssum (alias_of_pat pat,typ) in let patl,acc_sign,acc = aux (n+1) (push_rel d env) (d::acc_sign) tms acc in pat::patl,acc_sign,acc in let avoid0 = ids_of_context env in @@ -1743,7 +1748,7 @@ let build_inversion_problem loc env sigma tms t = let n = List.length sign in let decls = - List.map_i (fun i d -> (mkRel i, Context.Rel.Declaration.map (lift i) d)) 1 sign in + List.map_i (fun i d -> (mkRel i, map_constr (lift i) d)) 1 sign in let pb_env = push_rel_context sign env in let decls = @@ -1753,8 +1758,8 @@ let build_inversion_problem loc env sigma tms t = let dep_sign = find_dependencies_signature (List.make n true) decls in let sub_tms = - List.map2 (fun deps (tm,(tmtyp,_),(na,b,t)) -> - let na = if List.is_empty deps then Anonymous else force_name na in + List.map2 (fun deps (tm, (tmtyp,_), decl) -> + let na = if List.is_empty deps then Anonymous else force_name (get_name decl) in Pushed (true,((tm,tmtyp),deps,na))) dep_sign decls in let subst = List.map (fun (na,t) -> (na,lift n t)) subst in @@ -1815,7 +1820,8 @@ let build_inversion_problem loc env sigma tms t = let build_initial_predicate arsign pred = let rec buildrec n pred tmnames = function | [] -> List.rev tmnames,pred - | ((na,c,t)::realdecls)::lnames -> + | (decl::realdecls)::lnames -> + let na = get_name decl in let n' = n + List.length realdecls in buildrec (n'+1) pred (force_name na::tmnames) lnames | _ -> assert false @@ -1827,7 +1833,9 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign = match tm with | NotInd (bo,typ) -> (match t with - | None -> [na,Option.map (lift n) bo,lift n typ] + | None -> (match bo with + | None -> [LocalAssum (na, lift n typ)] + | Some b -> [LocalDef (na, lift n b, lift n typ)]) | Some (loc,_,_) -> user_err_loc (loc,"", str"Unexpected type annotation for a term of non inductive type.")) @@ -1845,8 +1853,8 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign = anomaly (Pp.str "Ill-formed 'in' clause in cases"); List.rev realnal | None -> List.make nrealargs_ctxt Anonymous in - (na,None,build_dependent_inductive env0 indf') - ::(List.map2 (fun x (_,c,t) ->(x,c,t)) realnal arsign) in + LocalAssum (na, build_dependent_inductive env0 indf') + ::(List.map2 set_name realnal arsign) in let rec buildrec n = function | [],[] -> [] | (_,tm)::ltm, (_,x)::tmsign -> @@ -2027,7 +2035,7 @@ let constr_of_pat env evdref arsign pat avoid = let previd, id = prime avoid (Name (Id.of_string "wildcard")) in Name id, id :: avoid in - (PatVar (l, name), [name, None, ty] @ realargs, mkRel 1, ty, + (PatVar (l, name), [LocalAssum (name, ty)] @ realargs, mkRel 1, ty, (List.map (fun x -> mkRel 1) realargs), 1, avoid) | PatCstr (l,((_, i) as cstr),args,alias) -> let cind = inductive_of_constructor cstr in @@ -2044,7 +2052,8 @@ let constr_of_pat env evdref arsign pat avoid = assert (Int.equal nb_args_constr (List.length args)); let patargs, args, sign, env, n, m, avoid = List.fold_right2 - (fun (na, c, t) ua (patargs, args, sign, env, n, m, avoid) -> + (fun decl ua (patargs, args, sign, env, n, m, avoid) -> + let t = get_type decl in let pat', sign', arg', typ', argtypargs, n', avoid = let liftt = liftn (List.length sign) (succ (List.length args)) t in typ env (substl args liftt, []) ua avoid @@ -2066,7 +2075,7 @@ let constr_of_pat env evdref arsign pat avoid = Anonymous -> pat', sign, app, apptype, realargs, n, avoid | Name id -> - let sign = (alias, None, lift m ty) :: sign in + let sign = LocalAssum (alias, lift m ty) :: sign in let avoid = id :: avoid in let sign, i, avoid = try @@ -2078,14 +2087,14 @@ let constr_of_pat env evdref arsign pat avoid = (lift 1 app) (* aliased term *) in let neq = eq_id avoid id in - (Name neq, Some (mkRel 0), eq_t) :: sign, 2, neq :: avoid + LocalDef (Name neq, mkRel 0, eq_t) :: sign, 2, neq :: avoid with Reduction.NotConvertible -> sign, 1, avoid in (* Mark the equality as a hole *) pat', sign, lift i app, lift i apptype, realargs, n + i, avoid in - let pat', sign, patc, patty, args, z, avoid = typ env (pi3 (List.hd arsign), List.tl arsign) pat avoid in - pat', (sign, patc, (pi3 (List.hd arsign), args), pat'), avoid + let pat', sign, patc, patty, args, z, avoid = typ env (get_type (List.hd arsign), List.tl arsign) pat avoid in + pat', (sign, patc, (get_type (List.hd arsign), args), pat'), avoid (* shadows functional version *) @@ -2100,23 +2109,23 @@ match kind_of_term t with | Rel 0 -> true | _ -> false -let rels_of_patsign l = - List.map (fun ((na, b, t) as x) -> - match b with - | Some t' when is_topvar t' -> (na, None, t) - | _ -> x) l +let rels_of_patsign = + List.map (fun decl -> + match decl with + | LocalDef (na,t',t) when is_topvar t' -> LocalAssum (na,t) + | _ -> decl) let vars_of_ctx ctx = let _, y = - List.fold_right (fun (na, b, t) (prev, vars) -> - match b with - | Some t' when is_topvar t' -> + List.fold_right (fun decl (prev, vars) -> + match decl with + | LocalDef (na,t',t) when is_topvar t' -> prev, (GApp (Loc.ghost, (GRef (Loc.ghost, delayed_force coq_eq_refl_ref, None)), [hole; GVar (Loc.ghost, prev)])) :: vars | _ -> - match na with + match get_name decl with Anonymous -> invalid_arg "vars_of_ctx" | Name n -> n, GVar (Loc.ghost, n) :: vars) ctx (Id.of_string "vars_of_ctx_error", []) @@ -2225,7 +2234,7 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity = match ineqs with | None -> [], arity | Some ineqs -> - [Anonymous, None, ineqs], lift 1 arity + [LocalAssum (Anonymous, ineqs)], lift 1 arity in let eqs_rels, arity = decompose_prod_n_assum neqs arity in eqs_rels @ neqs_rels @ rhs_rels', arity @@ -2236,7 +2245,7 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity = and btype = it_mkProd_or_LetIn j.uj_type rhs_rels' in let _btype = evd_comb1 (Typing.type_of env) evdref bbody in let branch_name = Id.of_string ("program_branch_" ^ (string_of_int !i)) in - let branch_decl = (Name branch_name, Some (lift !i bbody), (lift !i btype)) in + let branch_decl = LocalDef (Name branch_name, lift !i bbody, lift !i btype) in let branch = let bref = GVar (Loc.ghost, branch_name) in match vars_of_ctx rhs_rels with @@ -2285,7 +2294,7 @@ let abstract_tomatch env tomatchs tycon = (fun t -> subst_term (lift 1 c) (lift 1 t)) tycon in let name = next_ident_away (Id.of_string "filtered_var") names in (mkRel 1, lift_tomatch_type (succ lenctx) t) :: lift_ctx 1 prev, - (Name name, Some (lift lenctx c), lift lenctx $ type_of_tomatch t) :: ctx, + LocalDef (Name name, lift lenctx c, lift lenctx $ type_of_tomatch t) :: ctx, name :: names, tycon) ([], [], [], tycon) tomatchs in List.rev prev, ctx, tycon @@ -2293,7 +2302,7 @@ let abstract_tomatch env tomatchs tycon = let build_dependent_signature env evdref avoid tomatchs arsign = let avoid = ref avoid in let arsign = List.rev arsign in - let allnames = List.rev_map (List.map pi1) arsign in + let allnames = List.rev_map (List.map get_name) arsign in let nar = List.fold_left (fun n names -> List.length names + n) 0 allnames in let eqs, neqs, refls, slift, arsign' = List.fold_left2 @@ -2309,11 +2318,15 @@ let build_dependent_signature env evdref avoid tomatchs arsign = (* Build the arity signature following the names in matched terms as much as possible *) let argsign = List.tl arsign in (* arguments in inverse application order *) - let (appn, appb, appt) as _appsign = List.hd arsign in (* The matched argument *) + let app_decl = List.hd arsign in (* The matched argument *) + let appn = get_name app_decl in + let appt = get_type app_decl in let argsign = List.rev argsign in (* arguments in application order *) let env', nargeqs, argeqs, refl_args, slift, argsign' = List.fold_left2 - (fun (env, nargeqs, argeqs, refl_args, slift, argsign') arg (name, b, t) -> + (fun (env, nargeqs, argeqs, refl_args, slift, argsign') arg decl -> + let name = get_name decl in + let t = get_type decl in let argt = Retyping.get_type_of env !evdref arg in let eq, refl_arg = if Reductionops.is_conv env !evdref argt t then @@ -2331,16 +2344,16 @@ let build_dependent_signature env evdref avoid tomatchs arsign = let previd, id = let name = match kind_of_term arg with - Rel n -> pi1 (lookup_rel n env) + Rel n -> get_name (lookup_rel n env) | _ -> name in make_prime avoid name in (env, succ nargeqs, - (Name (eq_id avoid previd), None, eq) :: argeqs, + (LocalAssum (Name (eq_id avoid previd), eq)) :: argeqs, refl_arg :: refl_args, pred slift, - (Name id, b, t) :: argsign')) + set_name (Name id) decl :: argsign')) (env, neqs, [], [], slift, []) args argsign in let eq = mk_JMeq evdref @@ -2351,22 +2364,23 @@ let build_dependent_signature env evdref avoid tomatchs arsign = in let refl_eq = mk_JMeq_refl evdref ty tm in let previd, id = make_prime avoid appn in - (((Name (eq_id avoid previd), None, eq) :: argeqs) :: eqs, + ((LocalAssum (Name (eq_id avoid previd), eq) :: argeqs) :: eqs, succ nargeqs, refl_eq :: refl_args, pred slift, - (((Name id, appb, appt) :: argsign') :: arsigns)) + ((set_name (Name id) app_decl :: argsign') :: arsigns)) | _ -> (* Non dependent inductive or not inductive, just use a regular equality *) - let (name, b, typ) = match arsign with [x] -> x | _ -> assert(false) in + let decl = match arsign with [x] -> x | _ -> assert(false) in + let name = get_name decl in let previd, id = make_prime avoid name in - let arsign' = (Name id, b, typ) in + let arsign' = set_name (Name id) decl in let tomatch_ty = type_of_tomatch ty in let eq = mk_eq evdref (lift nar tomatch_ty) (mkRel slift) (lift nar tm) in - ([(Name (eq_id avoid previd), None, eq)] :: eqs, succ neqs, + ([LocalAssum (Name (eq_id avoid previd), eq)] :: eqs, succ neqs, (mk_eq_refl evdref tomatch_ty tm) :: refl_args, pred slift, (arsign' :: []) :: arsigns)) ([], 0, [], nar, []) tomatchs arsign @@ -2440,7 +2454,9 @@ let compile_program_cases loc style (typing_function, evdref) tycon env (* We push the initial terms to match and push their alias to rhs' envs *) (* names of aliases will be recovered from patterns (hence Anonymous here) *) - let out_tmt na = function NotInd (c,t) -> (na,c,t) | IsInd (typ,_,_) -> (na,None,typ) in + let out_tmt na = function NotInd (None,t) -> LocalAssum (na,t) + | NotInd (Some b, t) -> LocalDef (na,b,t) + | IsInd (typ,_,_) -> LocalAssum (na,typ) in let typs = List.map2 (fun na (tm,tmt) -> (tm,out_tmt na tmt)) nal tomatchs in let typs = @@ -2513,7 +2529,9 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e (* names of aliases will be recovered from patterns (hence Anonymous *) (* here) *) - let out_tmt na = function NotInd (c,t) -> (na,c,t) | IsInd (typ,_,_) -> (na,None,typ) in + let out_tmt na = function NotInd (None,t) -> LocalAssum (na,t) + | NotInd (Some b,t) -> LocalDef (na,b,t) + | IsInd (typ,_,_) -> LocalAssum (na,typ) in let typs = List.map2 (fun na (tm,tmt) -> (tm,out_tmt na tmt)) nal tomatchs in let typs = diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index 489a311bc..9d5a6006d 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -142,6 +142,7 @@ let mu env evdref t = and coerce loc env evdref (x : Term.constr) (y : Term.constr) : (Term.constr -> Term.constr) option = + let open Context.Rel.Declaration in let rec coerce_unify env x y = let x = hnf env !evdref x and y = hnf env !evdref y in try @@ -151,8 +152,9 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr) and coerce' env x y : (Term.constr -> Term.constr) option = let subco () = subset_coerce env evdref x y in let dest_prod c = + let open Context.Rel.Declaration in match Reductionops.splay_prod_n env ( !evdref) 1 c with - | [(na,b,t)], c -> (na,t), c + | [LocalAssum (na,t) | LocalDef (na,_,t)], c -> (na,t), c | _ -> raise NoSubtacCoercion in let coerce_application typ typ' c c' l l' = @@ -205,7 +207,7 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr) let name' = Name (Namegen.next_ident_away Namegen.default_dependent_ident (Termops.ids_of_context env)) in - let env' = push_rel (name', None, a') env in + let env' = push_rel (LocalAssum (name', a')) env in let c1 = coerce_unify env' (lift 1 a') (lift 1 a) in (* env, x : a' |- c1 : lift 1 a' > lift 1 a *) let coec1 = app_opt env' evdref c1 (mkRel 1) in @@ -255,7 +257,7 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr) | _ -> raise NoSubtacCoercion in let (pb, b), (pb', b') = remove_head a pb, remove_head a' pb' in - let env' = push_rel (Name Namegen.default_dependent_ident, None, a) env in + let env' = push_rel (LocalAssum (Name Namegen.default_dependent_ident, a)) env in let c2 = coerce_unify env' b b' in match c1, c2 with | None, None -> None @@ -475,7 +477,8 @@ let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = let name = match name with | Anonymous -> Name Namegen.default_dependent_ident | _ -> name in - let env1 = push_rel (name,None,u1) env in + let open Context.Rel.Declaration in + let env1 = push_rel (LocalAssum (name,u1)) env in let (evd', v1) = inh_conv_coerce_to_fail loc env1 evd rigidonly (Some (mkRel 1)) (lift 1 u1) (lift 1 t1) in diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index 69c1dfb47..4fb411202 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -20,6 +20,7 @@ open Vars open Pattern open Patternops open Misctypes +open Context.Rel.Declaration (*i*) (* Given a term with second-order variables in it, @@ -254,15 +255,15 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels sorec ctx env subst c1 c2 | PProd (na1,c1,d1), Prod(na2,c2,d2) -> - sorec ((na1,na2,c2)::ctx) (Environ.push_rel (na2,None,c2) env) + sorec ((na1,na2,c2)::ctx) (Environ.push_rel (LocalAssum (na2,c2)) env) (add_binders na1 na2 binding_vars (sorec ctx env subst c1 c2)) d1 d2 | PLambda (na1,c1,d1), Lambda(na2,c2,d2) -> - sorec ((na1,na2,c2)::ctx) (Environ.push_rel (na2,None,c2) env) + sorec ((na1,na2,c2)::ctx) (Environ.push_rel (LocalAssum (na2,c2)) env) (add_binders na1 na2 binding_vars (sorec ctx env subst c1 c2)) d1 d2 | PLetIn (na1,c1,d1), LetIn(na2,c2,t2,d2) -> - sorec ((na1,na2,t2)::ctx) (Environ.push_rel (na2,Some c2,t2) env) + sorec ((na1,na2,t2)::ctx) (Environ.push_rel (LocalDef (na2,c2,t2)) env) (add_binders na1 na2 binding_vars (sorec ctx env subst c1 c2)) d1 d2 | PIf (a1,b1,b1'), Case (ci,_,a2,[|b2;b2'|]) -> @@ -271,7 +272,7 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels let n = Context.Rel.length ctx_b2 in let n' = Context.Rel.length ctx_b2' in if noccur_between 1 n b2 && noccur_between 1 n' b2' then - let f l (na,_,t) = (Anonymous,na,t)::l in + let f l (LocalAssum (na,t) | LocalDef (na,_,t)) = (Anonymous,na,t)::l in let ctx_br = List.fold_left f ctx ctx_b2 in let ctx_br' = List.fold_left f ctx ctx_b2' in let b1 = lift_pattern n b1 and b1' = lift_pattern n' b1' in @@ -367,21 +368,21 @@ let sub_match ?(partial_app=false) ?(closed=true) env sigma pat c = | [c1; c2] -> mk_ctx (mkLambda (x, c1, c2)) | _ -> assert false in - let env' = Environ.push_rel (x,None,c1) env in + let env' = Environ.push_rel (LocalAssum (x,c1)) env in try_aux [(env, c1); (env', c2)] next_mk_ctx next | Prod (x,c1,c2) -> let next_mk_ctx = function | [c1; c2] -> mk_ctx (mkProd (x, c1, c2)) | _ -> assert false in - let env' = Environ.push_rel (x,None,c1) env in + let env' = Environ.push_rel (LocalAssum (x,c1)) env in try_aux [(env, c1); (env', c2)] next_mk_ctx next | LetIn (x,c1,t,c2) -> let next_mk_ctx = function | [c1; c2] -> mk_ctx (mkLetIn (x, c1, t, c2)) | _ -> assert false in - let env' = Environ.push_rel (x,Some c1,t) env in + let env' = Environ.push_rel (LocalDef (x,c1,t)) env in try_aux [(env, c1); (env', c2)] next_mk_ctx next | App (c1,lc) -> let topdown = true in diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index c8ecf91d3..67a8f01aa 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -24,6 +24,7 @@ open Nametab open Mod_subst open Misctypes open Decl_kinds +open Context.Named.Declaration let dl = Loc.ghost @@ -33,8 +34,15 @@ let print_universes = Flags.univ_print (** If true, prints local context of evars, whatever print_arguments *) let print_evar_arguments = ref false -let add_name na b t (nenv, env) = add_name na nenv, push_rel (na, b, t) env -let add_name_opt na b t (nenv, env) = +let add_name na b t (nenv, env) = + let open Context.Rel.Declaration in + add_name na nenv, push_rel (match b with + | None -> LocalAssum (na,t) + | Some b -> LocalDef (na,b,t) + ) + env + +let add_name_opt na b t (nenv, env) = match t with | None -> Termops.add_name na nenv, env | Some t -> add_name na b t (nenv, env) @@ -510,11 +518,14 @@ let rec detype flags avoid env sigma t = else noparams () | Evar (evk,cl) -> - let bound_to_itself_or_letin (id,b,_) c = - b != None || - try let n = List.index Name.equal (Name id) (fst env) in - isRelN n c - with Not_found -> isVarId id c in + let bound_to_itself_or_letin decl c = + match decl with + | LocalDef _ -> true + | LocalAssum (id,_) -> + try let n = List.index Name.equal (Name id) (fst env) in + isRelN n c + with Not_found -> isVarId id c + in let id,l = try let id = Evd.evar_ident evk sigma in @@ -684,17 +695,24 @@ let detype_rel_context ?(lax=false) where avoid env sigma sign = let where = Option.map (fun c -> it_mkLambda_or_LetIn c sign) where in let rec aux avoid env = function | [] -> [] - | (na,b,t)::rest -> + | decl::rest -> + let open Context.Rel.Declaration in + let na = get_name decl in + let t = get_type decl in let na',avoid' = match where with | None -> na,avoid | Some c -> - if b != None then + if is_local_def decl then compute_displayed_let_name_in (RenamingElsewhereFor (fst env,c)) avoid na c else compute_displayed_name_in (RenamingElsewhereFor (fst env,c)) avoid na c in + let b = match decl with + | LocalAssum _ -> None + | LocalDef (_,b,_) -> Some b + in let b' = Option.map (detype (lax,false) avoid env sigma) b in let t' = detype (lax,false) avoid env sigma t in (na',Explicit,b',t') :: aux avoid' (add_name na' b t env) rest diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 99e51330e..020f998aa 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -23,6 +23,7 @@ open Globnames open Evd open Pretype_errors open Sigma.Notations +open Context.Rel.Declaration type unify_fun = transparent_state -> env -> evar_map -> conv_pb -> constr -> constr -> Evarsolve.unification_result @@ -55,12 +56,15 @@ let eval_flexible_term ts env evd c = then constant_opt_value_in env cu else None | Rel n -> - (try let (_,v,_) = lookup_rel n env in Option.map (lift n) v - with Not_found -> None) + (try match lookup_rel n env with + | LocalAssum _ -> None + | LocalDef (_,v,_) -> Some (lift n v) + with Not_found -> None) | Var id -> (try if is_transparent_variable ts id then - let (_,v,_) = lookup_named id env in v + let open Context.Named.Declaration in + lookup_named id env |> get_value else None with Not_found -> None) | LetIn (_,b,_,c) -> Some (subst1 b c) @@ -394,7 +398,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty assert (match sk with [] -> true | _ -> false); let (na,c1,c'1) = destLambda term in let c = nf_evar evd c1 in - let env' = push_rel (na,None,c) env in + let env' = push_rel (LocalAssum (na,c)) env in let out1 = whd_betaiota_deltazeta_for_iota_state (fst ts) env' evd Cst_stack.empty (c'1, Stack.empty) in let out2 = whd_nored_state evd @@ -561,7 +565,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty let b = nf_evar i b1 in let t = nf_evar i t1 in let na = Nameops.name_max na1 na2 in - evar_conv_x ts (push_rel (na,Some b,t) env) i pbty c'1 c'2); + evar_conv_x ts (push_rel (LocalDef (na,b,t)) env) i pbty c'1 c'2); (fun i -> exact_ise_stack2 env i (evar_conv_x ts) sk1 sk2)] and f2 i = let out1 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts1 (v1,sk1) @@ -676,7 +680,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty (fun i -> let c = nf_evar i c1 in let na = Nameops.name_max na1 na2 in - evar_conv_x ts (push_rel (na,None,c) env) i CONV c'1 c'2)] + evar_conv_x ts (push_rel (LocalAssum (na,c)) env) i CONV c'1 c'2)] | Flexible ev1, Rigid -> flex_rigid true ev1 appr1 appr2 | Rigid, Flexible ev2 -> flex_rigid false ev2 appr2 appr1 @@ -735,7 +739,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty (fun i -> let c = nf_evar i c1 in let na = Nameops.name_max n1 n2 in - evar_conv_x ts (push_rel (na,None,c) env) i pbty c'1 c'2)] + evar_conv_x ts (push_rel (LocalAssum (na,c)) env) i pbty c'1 c'2)] | Rel x1, Rel x2 -> if Int.equal x1 x2 then @@ -912,6 +916,7 @@ let choose_less_dependent_instance evk evd term args = | [] -> None | (id, _) :: _ -> Some (Evd.define evk (mkVar id) evd) +open Context.Named.Declaration let apply_on_subterm env evdref f c t = let rec applyrec (env,(k,c) as acc) t = (* By using eq_constr, we make an approximation, for instance, we *) @@ -922,7 +927,7 @@ let apply_on_subterm env evdref f c t = match kind_of_term t with | Evar (evk,args) when Evd.is_undefined !evdref evk -> let ctx = evar_filtered_context (Evd.find_undefined !evdref evk) in - let g (_,b,_) a = if Option.is_empty b then applyrec acc a else a in + let g decl a = if is_local_assum decl then applyrec acc a else a in mkEvar (evk, Array.of_list (List.map2 g ctx (Array.to_list args))) | _ -> map_constr_with_binders_left_to_right @@ -939,17 +944,17 @@ let filter_possible_projections c ty ctxt args = let fv2 = collect_vars (mkApp (c,args)) in let len = Array.length args in let tyvars = collect_vars ty in - List.map_i (fun i (id,b,_) -> + List.map_i (fun i decl -> let () = assert (i < len) in let a = Array.unsafe_get args i in - (match b with None -> false | Some c -> not (isRel c || isVar c)) || + (match decl with LocalAssum _ -> false | LocalDef (_,c,_) -> not (isRel c || isVar c)) || a == c || (* Here we make an approximation, for instance, we could also be *) (* interested in finding a term u convertible to c such that a occurs *) (* in u *) isRel a && Int.Set.mem (destRel a) fv1 || isVar a && Id.Set.mem (destVar a) fv2 || - Id.Set.mem id tyvars) + Id.Set.mem (get_id decl) tyvars) 0 ctxt let solve_evars = ref (fun _ -> failwith "solve_evars not installed") @@ -980,17 +985,18 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs = let env_evar = evar_filtered_env evi in let sign = named_context_val env_evar in let ctxt = evar_filtered_context evi in - let instance = List.map mkVar (List.map pi1 ctxt) in + let instance = List.map mkVar (List.map get_id ctxt) in let rec make_subst = function - | (id,_,t)::ctxt', c::l, occs::occsl when isVarId id c -> + | decl'::ctxt', c::l, occs::occsl when isVarId (get_id decl') c -> begin match occs with | Some _ -> error "Cannot force abstraction on identity instance." | None -> make_subst (ctxt',l,occsl) end - | (id,_,t)::ctxt', c::l, occs::occsl -> + | decl'::ctxt', c::l, occs::occsl -> + let (id,_,t) = to_tuple decl' in let evs = ref [] in let ty = Retyping.get_type_of env_rhs evd c in let filter' = filter_possible_projections c ty ctxt args in diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index 0dd0ad2e0..a65394e17 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -19,6 +19,7 @@ open Retyping open Reductionops open Evarutil open Pretype_errors +open Context.Rel.Declaration let normalize_evar evd ev = match kind_of_term (whd_evar evd (mkEvar ev)) with @@ -79,7 +80,7 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) pbty env evd t = if !modified then evdref := Evd.add !evdref ev {evi with evar_concl = ty'} else () - | _ -> iter_constr (refresh_term_evars onevars false) t + | _ -> Constr.iter (refresh_term_evars onevars false) t and refresh_polymorphic_positions args pos = let rec aux i = function | Some l :: ls -> @@ -162,7 +163,8 @@ type 'a update = | UpdateWith of 'a | NoUpdate -let inst_of_vars sign = Array.map_of_list (fun (id,_,_) -> mkVar id) sign +open Context.Named.Declaration +let inst_of_vars sign = Array.map_of_list (mkVar % get_id) sign let restrict_evar_key evd evk filter candidates = match filter, candidates with @@ -205,6 +207,7 @@ let restrict_instance evd evk filter argsv = let evi = Evd.find evd evk in Filter.filter_array (Filter.compose (evar_filter evi) filter) argsv +open Context.Rel.Declaration let noccur_evar env evd evk c = let cache = ref Int.Set.empty (* cache for let-ins *) in let rec occur_rec (k, env as acc) c = @@ -217,9 +220,9 @@ let noccur_evar env evd evk c = else Array.iter (occur_rec acc) args') | Rel i when i > k -> if not (Int.Set.mem (i-k) !cache) then - (match pi2 (Environ.lookup_rel i env) with - | None -> () - | Some b -> cache := Int.Set.add (i-k) !cache; occur_rec acc (lift i b)) + (match Environ.lookup_rel i env with + | LocalAssum _ -> () + | LocalDef (_,b,_) -> cache := Int.Set.add (i-k) !cache; occur_rec acc (lift i b)) | Proj (p,c) -> let c = try Retyping.expand_projection env evd p c [] @@ -241,9 +244,11 @@ let noccur_evar env evd evk c = variable in its family of aliased variables *) let compute_var_aliases sign = - List.fold_right (fun (id,b,c) aliases -> - match b with - | Some t -> + let open Context.Named.Declaration in + List.fold_right (fun decl aliases -> + let id = get_id decl in + match decl with + | LocalDef (_,t,_) -> (match kind_of_term t with | Var id' -> let aliases_of_id = @@ -251,27 +256,30 @@ let compute_var_aliases sign = Id.Map.add id (aliases_of_id@[t]) aliases | _ -> Id.Map.add id [t] aliases) - | None -> aliases) + | LocalAssum _ -> aliases) sign Id.Map.empty let compute_rel_aliases var_aliases rels = - snd (List.fold_right (fun (_,b,u) (n,aliases) -> - (n-1, - match b with - | Some t -> - (match kind_of_term t with - | Var id' -> - let aliases_of_n = - try Id.Map.find id' var_aliases with Not_found -> [] in - Int.Map.add n (aliases_of_n@[t]) aliases - | Rel p -> - let aliases_of_n = - try Int.Map.find (p+n) aliases with Not_found -> [] in - Int.Map.add n (aliases_of_n@[mkRel (p+n)]) aliases - | _ -> - Int.Map.add n [lift n (mkCast(t,DEFAULTcast,u))] aliases) - | None -> aliases)) - rels (List.length rels,Int.Map.empty)) + snd (List.fold_right + (fun decl (n,aliases) -> + (n-1, + match decl with + | LocalDef (_,t,u) -> + (match kind_of_term t with + | Var id' -> + let aliases_of_n = + try Id.Map.find id' var_aliases with Not_found -> [] in + Int.Map.add n (aliases_of_n@[t]) aliases + | Rel p -> + let aliases_of_n = + try Int.Map.find (p+n) aliases with Not_found -> [] in + Int.Map.add n (aliases_of_n@[mkRel (p+n)]) aliases + | _ -> + Int.Map.add n [lift n (mkCast(t,DEFAULTcast,u))] aliases) + | LocalAssum _ -> aliases) + ) + rels + (List.length rels,Int.Map.empty)) let make_alias_map env = (* We compute the chain of aliases for each var and rel *) @@ -305,13 +313,13 @@ let normalize_alias aliases x = let normalize_alias_var var_aliases id = destVar (normalize_alias (var_aliases,Int.Map.empty) (mkVar id)) -let extend_alias (_,b,_) (var_aliases,rel_aliases) = +let extend_alias decl (var_aliases,rel_aliases) = let rel_aliases = Int.Map.fold (fun n l -> Int.Map.add (n+1) (List.map (lift 1) l)) rel_aliases Int.Map.empty in let rel_aliases = - match b with - | Some t -> + match decl with + | LocalDef(_,t,_) -> (match kind_of_term t with | Var id' -> let aliases_of_binder = @@ -323,7 +331,7 @@ let extend_alias (_,b,_) (var_aliases,rel_aliases) = Int.Map.add 1 (aliases_of_binder@[mkRel (p+1)]) rel_aliases | _ -> Int.Map.add 1 [lift 1 t] rel_aliases) - | None -> rel_aliases in + | LocalAssum _ -> rel_aliases in (var_aliases, rel_aliases) let expand_alias_once aliases x = @@ -429,16 +437,17 @@ let get_actual_deps aliases l t = | Rel n -> Int.Set.mem n fv_rels | _ -> assert false) l +open Context.Named.Declaration let remove_instance_local_defs evd evk args = let evi = Evd.find evd evk in let len = Array.length args in let rec aux sign i = match sign with | [] -> let () = assert (i = len) in [] - | (_, None, _) :: sign -> + | LocalAssum _ :: sign -> let () = assert (i < len) in (Array.unsafe_get args i) :: aux sign (succ i) - | (_, Some _, _) :: sign -> + | LocalDef _ :: sign -> aux sign (succ i) in aux (evar_filtered_context evi) 0 @@ -500,7 +509,8 @@ let solve_pattern_eqn env l c = match kind_of_term a with (* Rem: if [a] links to a let-in, do as if it were an assumption *) | Rel n -> - let d = Context.Rel.Declaration.map (lift n) (lookup_rel n env) in + let open Context.Rel.Declaration in + let d = map_constr (lift n) (lookup_rel n env) in mkLambda_or_LetIn d c' | Var id -> let d = lookup_named id env in mkNamedLambda_or_LetIn d c' @@ -529,9 +539,9 @@ let make_projectable_subst aliases sigma evi args = let evar_aliases = compute_var_aliases sign in let (_,full_subst,cstr_subst) = List.fold_right - (fun (id,b,c) (args,all,cstrs) -> - match b,args with - | None, a::rest -> + (fun decl (args,all,cstrs) -> + match decl,args with + | LocalAssum (id,c), a::rest -> let a = whd_evar sigma a in let cstrs = let a',args = decompose_app_vect a in @@ -541,7 +551,7 @@ let make_projectable_subst aliases sigma evi args = Constrmap.add (fst cstr) ((args,id)::l) cstrs | _ -> cstrs in (rest,Id.Map.add id [a,normalize_alias_opt aliases a,id] all,cstrs) - | Some c, a::rest -> + | LocalDef (id,c,_), a::rest -> let a = whd_evar sigma a in (match kind_of_term c with | Var id' -> @@ -601,10 +611,12 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env = let sign1 = evar_hyps evi1 in let filter1 = evar_filter evi1 in let src = subterm_source evk1 evi1.evar_source in - let ids1 = List.map pi1 (named_context_of_val sign1) in + let ids1 = List.map get_id (named_context_of_val sign1) in let inst_in_sign = List.map mkVar (Filter.filter_list filter1 ids1) in + let open Context.Rel.Declaration in let (sign2,filter2,inst2_in_env,inst2_in_sign,_,evd,_) = - List.fold_right (fun (na,b,t_in_env as d) (sign,filter,inst_in_env,inst_in_sign,env,evd,avoid) -> + List.fold_right (fun d (sign,filter,inst_in_env,inst_in_sign,env,evd,avoid) -> + let LocalAssum (na,t_in_env) | LocalDef (na,_,t_in_env) = d in let id = next_name_away na avoid in let evd,t_in_sign = let s = Retyping.get_sort_of env evd t_in_env in @@ -612,13 +624,13 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env = ~status:univ_flexible (Some false) env evd (mkSort s) in define_evar_from_virtual_equation define_fun env evd src t_in_env ty_t_in_sign sign filter inst_in_env in - let evd,b_in_sign = match b with - | None -> evd,None - | Some b -> + let evd,b_in_sign = match d with + | LocalAssum _ -> evd,None + | LocalDef (_,b,_) -> let evd,b = define_evar_from_virtual_equation define_fun env evd src b t_in_sign sign filter inst_in_env in evd,Some b in - (push_named_context_val (id,b_in_sign,t_in_sign) sign, Filter.extend 1 filter, + (push_named_context_val (Context.Named.Declaration.of_tuple (id,b_in_sign,t_in_sign)) sign, Filter.extend 1 filter, (mkRel 1)::(List.map (lift 1) inst_in_env), (mkRel 1)::(List.map (lift 1) inst_in_sign), push_rel d env,evd,id::avoid)) @@ -756,9 +768,10 @@ let project_with_effects aliases sigma effects t subst = effects := p :: !effects; c +open Context.Named.Declaration let rec find_solution_type evarenv = function - | (id,ProjectVar)::l -> pi3 (lookup_named id evarenv) - | [id,ProjectEvar _] -> (* bugged *) pi3 (lookup_named id evarenv) + | (id,ProjectVar)::l -> get_type (lookup_named id evarenv) + | [id,ProjectEvar _] -> (* bugged *) get_type (lookup_named id evarenv) | (id,ProjectEvar _)::l -> find_solution_type evarenv l | [] -> assert false @@ -892,7 +905,7 @@ let invert_invertible_arg fullenv evd aliases k (evk,argsv) args' = *) let set_of_evctx l = - List.fold_left (fun s (id,_,_) -> Id.Set.add id s) Id.Set.empty l + List.fold_left (fun s decl -> Id.Set.add (get_id decl) s) Id.Set.empty l let filter_effective_candidates evi filter candidates = match filter with @@ -924,7 +937,13 @@ let closure_of_filter evd evk = function | Some filter -> let evi = Evd.find_undefined evd evk in let vars = collect_vars (Evarutil.nf_evar evd (evar_concl evi)) in - let test b (id,c,_) = b || Idset.mem id vars || match c with None -> false | Some c -> not (isRel c || isVar c) in + let test b decl = b || Idset.mem (get_id decl) vars || + match decl with + | LocalAssum _ -> + false + | LocalDef (_,c,_) -> + not (isRel c || isVar c) + in let newfilter = Filter.map_along test filter (evar_context evi) in (* Now ensure that restriction is at least what is was originally *) let newfilter = Option.cata (Filter.map_along (&&) newfilter) newfilter (Filter.repr (evar_filter evi)) in @@ -1280,7 +1299,7 @@ let occur_evar_upto_types sigma n c = seen := Evar.Set.add sp !seen; Option.iter occur_rec (existential_opt_value sigma e); occur_rec (existential_type sigma e)) - | _ -> iter_constr occur_rec c + | _ -> Constr.iter occur_rec c in try occur_rec c; false with Occur -> true @@ -1365,15 +1384,16 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = let t = whd_evar !evdref t in match kind_of_term t with | Rel i when i>k -> - (match pi2 (Environ.lookup_rel (i-k) env') with - | None -> project_variable (mkRel (i-k)) - | Some b -> + let open Context.Rel.Declaration in + (match Environ.lookup_rel (i-k) env' with + | LocalAssum _ -> project_variable (mkRel (i-k)) + | LocalDef (_,b,_) -> try project_variable (mkRel (i-k)) with NotInvertibleUsingOurAlgorithm _ -> imitate envk (lift i b)) | Var id -> - (match pi2 (Environ.lookup_named id env') with - | None -> project_variable t - | Some b -> + (match Environ.lookup_named id env' with + | LocalAssum _ -> project_variable t + | LocalDef (_,b,_) -> try project_variable t with NotInvertibleUsingOurAlgorithm _ -> imitate envk b) | LetIn (na,b,u,c) -> @@ -1453,7 +1473,8 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = let names = ref Idset.empty in let rec is_id_subst ctxt s = match ctxt, s with - | ((id, _, _) :: ctxt'), (c :: s') -> + | (decl :: ctxt'), (c :: s') -> + let id = get_id decl in names := Idset.add id !names; isVarId id c && is_id_subst ctxt' s' | [], [] -> true diff --git a/pretyping/evarutil.ml b/pretyping/evarutil.ml index 8c210e283..343d3ef90 100644 --- a/pretyping/evarutil.ml +++ b/pretyping/evarutil.ml @@ -76,13 +76,15 @@ let tj_nf_evar sigma {utj_val=v;utj_type=t} = {utj_val=nf_evar sigma v;utj_type=t} let env_nf_evar sigma env = + let open Context.Rel.Declaration in process_rel_context - (fun d e -> push_rel (Context.Rel.Declaration.map (nf_evar sigma) d) e) env + (fun d e -> push_rel (map_constr (nf_evar sigma) d) e) env let env_nf_betaiotaevar sigma env = + let open Context.Rel.Declaration in process_rel_context (fun d e -> - push_rel (Context.Rel.Declaration.map (Reductionops.nf_betaiota sigma) d) e) env + push_rel (map_constr (Reductionops.nf_betaiota sigma) d) e) env let nf_evars_universes evm = Universes.nf_evars_and_universes_opt_subst (Reductionops.safe_evar_value evm) @@ -149,11 +151,16 @@ let is_ground_term evd t = not (has_undefined_evars evd t) let is_ground_env evd env = - let is_ground_decl = function - (_,Some b,_) -> is_ground_term evd b + let open Context.Rel.Declaration in + let is_ground_rel_decl = function + | LocalDef (_,b,_) -> is_ground_term evd b | _ -> true in - List.for_all is_ground_decl (rel_context env) && - List.for_all is_ground_decl (named_context env) + let open Context.Named.Declaration in + let is_ground_named_decl = function + | LocalDef (_,b,_) -> is_ground_term evd b + | _ -> true in + List.for_all is_ground_rel_decl (rel_context env) && + List.for_all is_ground_named_decl (named_context env) (* Memoization is safe since evar_map and environ are applicative structures *) @@ -231,10 +238,11 @@ let non_instantiated sigma = (************************) let make_pure_subst evi args = + let open Context.Named.Declaration in snd (List.fold_right - (fun (id,b,c) (args,l) -> + (fun decl (args,l) -> match args with - | a::rest -> (rest, (id,a)::l) + | a::rest -> (rest, (get_id decl, a)::l) | _ -> anomaly (Pp.str "Instance does not match its signature")) (evar_filtered_context evi) (Array.rev_to_list args,[])) @@ -276,17 +284,15 @@ let subst2 subst vsubst c = let push_rel_context_to_named_context env typ = (* compute the instances relative to the named context and rel_context *) - let ids = List.map pi1 (named_context env) in + let open Context.Named.Declaration in + let ids = List.map get_id (named_context env) in let inst_vars = List.map mkVar ids in let inst_rels = List.rev (rel_list 0 (nb_rel env)) in - let replace_var_named_declaration id0 id (id',b,t) = + let replace_var_named_declaration id0 id decl = + let id' = get_id decl in let id' = if Id.equal id0 id' then id else id' in let vsubst = [id0 , mkVar id] in - let b = match b with - | None -> None - | Some c -> Some (replace_vars vsubst c) - in - id', b, replace_vars vsubst t + decl |> set_id id' |> map_constr (replace_vars vsubst) in let replace_var_named_context id0 id env = let nc = Environ.named_context env in @@ -303,7 +309,12 @@ let push_rel_context_to_named_context env typ = (* We do keep the instances corresponding to local definition (see above) *) let (subst, vsubst, _, env) = Context.Rel.fold_outside - (fun (na,c,t) (subst, vsubst, avoid, env) -> + (fun decl (subst, vsubst, avoid, env) -> + let open Context.Rel.Declaration in + let na = get_name decl in + let c = get_value decl in + let t = get_type decl in + let open Context.Named.Declaration in let id = (* ppedrot: we want to infer nicer names for the refine tactic, but keeping at the same time backward compatibility in other code @@ -321,7 +332,10 @@ let push_rel_context_to_named_context env typ = context. Unless [id] is a section variable. *) let subst = List.map (replace_vars [id0,mkVar id]) subst in let vsubst = (id0,mkVar id)::vsubst in - let d = (id0, Option.map (subst2 subst vsubst) c, subst2 subst vsubst t) in + let d = match c with + | None -> LocalAssum (id0, subst2 subst vsubst t) + | Some c -> LocalDef (id0, subst2 subst vsubst c, subst2 subst vsubst t) + in let env = replace_var_named_context id0 id env in (mkVar id0 :: subst, vsubst, id::avoid, push_named d env) | _ -> @@ -329,7 +343,10 @@ let push_rel_context_to_named_context env typ = incorrect. We revert to a less robust behaviour where the new binder has name [id]. Which amounts to the same behaviour than when [id=id0]. *) - let d = (id,Option.map (subst2 subst vsubst) c,subst2 subst vsubst t) in + let d = match c with + | None -> LocalAssum (id, subst2 subst vsubst t) + | Some c -> LocalDef (id, subst2 subst vsubst c, subst2 subst vsubst t) + in (mkVar id :: subst, vsubst, id::avoid, push_named d env) ) (rel_context env) ~init:([], [], ids, env) in @@ -477,7 +494,7 @@ let rec check_and_clear_in_constr env evdref err ids c = let ctxt = Evd.evar_filtered_context evi in let (rids,filter) = List.fold_right2 - (fun (rid, ob,c as h) a (ri,filter) -> + (fun h a (ri,filter) -> try (* Check if some id to clear occurs in the instance a of rid in ev and remember the dependency *) @@ -493,7 +510,8 @@ let rec check_and_clear_in_constr env evdref err ids c = let () = Id.Map.iter check ri in (* No dependency at all, we can keep this ev's context hyp *) (ri, true::filter) - with Depends id -> (Id.Map.add rid id ri, false::filter)) + with Depends id -> let open Context.Named.Declaration in + (Id.Map.add (get_id h) id ri, false::filter)) ctxt (Array.to_list l) (Id.Map.empty,[]) in (* Check if some rid to clear in the context of ev has dependencies in the type of ev and adjust the source of the dependency *) @@ -528,11 +546,10 @@ let clear_hyps_in_evi_main env evdref hyps terms ids = let terms = List.map (check_and_clear_in_constr env evdref (OccurHypInSimpleClause None) ids) terms in let nhyps = - let check_context ((id,ob,c) as decl) = - let err = OccurHypInSimpleClause (Some id) in - let ob' = Option.smartmap (fun c -> check_and_clear_in_constr env evdref err ids c) ob in - let c' = check_and_clear_in_constr env evdref err ids c in - if ob == ob' && c == c' then decl else (id, ob', c') + let open Context.Named.Declaration in + let check_context decl = + let err = OccurHypInSimpleClause (Some (get_id decl)) in + map_constr (check_and_clear_in_constr env evdref err ids) decl in let check_value vk = match force_lazy_val vk with | None -> vk @@ -570,11 +587,12 @@ let process_dependent_evar q acc evm is_dependent e = (* Queues evars appearing in the types of the goal (conclusion, then hypotheses), they are all dependent. *) queue_term q true evi.evar_concl; - List.iter begin fun (_,b,t) -> - queue_term q true t; - match b with - | None -> () - | Some b -> queue_term q true b + List.iter begin fun decl -> + let open Context.Named.Declaration in + queue_term q true (get_type decl); + match decl with + | LocalAssum _ -> () + | LocalDef (_,b,_) -> queue_term q true b end (Environ.named_context_of_val evi.evar_hyps); match evi.evar_body with | Evar_empty -> @@ -625,11 +643,11 @@ let undefined_evars_of_term evd t = evrec Evar.Set.empty t let undefined_evars_of_named_context evd nc = - List.fold_right (fun (_, b, t) s -> - Option.fold_left (fun s t -> - Evar.Set.union s (undefined_evars_of_term evd t)) - (Evar.Set.union s (undefined_evars_of_term evd t)) b) - nc Evar.Set.empty + let open Context.Named.Declaration in + Context.Named.fold_outside + (fold (fun c s -> Evar.Set.union s (undefined_evars_of_term evd c))) + nc + ~init:Evar.Set.empty let undefined_evars_of_evar_info evd evi = Evar.Set.union (undefined_evars_of_term evd evi.evar_concl) @@ -709,6 +727,7 @@ let idx = Namegen.default_dependent_ident (* Refining an evar to a product *) let define_pure_evar_as_product evd evk = + let open Context.Named.Declaration in let evi = Evd.find_undefined evd evk in let evenv = evar_env evi in let id = next_ident_away idx (ids_of_named_context (evar_context evi)) in @@ -717,7 +736,7 @@ let define_pure_evar_as_product evd evk = let evd1,(dom,u1) = new_type_evar evenv evd univ_flexible_alg ~filter:(evar_filter evi) in let evd2,rng = - let newenv = push_named (id, None, dom) evenv in + let newenv = push_named (LocalAssum (id, dom)) evenv in let src = evar_source evk evd1 in let filter = Filter.extend 1 (evar_filter evi) in if is_prop_sort s then @@ -756,6 +775,7 @@ let define_evar_as_product evd (evk,args) = *) let define_pure_evar_as_lambda env evd evk = + let open Context.Named.Declaration in let evi = Evd.find_undefined evd evk in let evenv = evar_env evi in let typ = whd_betadeltaiota evenv evd (evar_concl evi) in @@ -766,7 +786,7 @@ let define_pure_evar_as_lambda env evd evk = let avoid = ids_of_named_context (evar_context evi) in let id = next_name_away_with_default_using_types "x" na avoid (whd_evar evd dom) in - let newenv = push_named (id, None, dom) evenv in + let newenv = push_named (LocalAssum (id, dom)) evenv in let filter = Filter.extend 1 (evar_filter evi) in let src = evar_source evk evd1 in let evd2,body = new_evar_unsafe newenv evd1 ~src (subst1 (mkVar id) rng) ~filter in diff --git a/pretyping/find_subterm.ml b/pretyping/find_subterm.ml index 6733b7fca..ae8b91c34 100644 --- a/pretyping/find_subterm.ml +++ b/pretyping/find_subterm.ml @@ -59,19 +59,22 @@ let proceed_with_occurrences f occs x = (** Applying a function over a named_declaration with an hypothesis location request *) -let map_named_declaration_with_hyploc f hyploc acc (id,bodyopt,typ) = - let f = f (Some (id,hyploc)) in - match bodyopt,hyploc with - | None, InHypValueOnly -> +let map_named_declaration_with_hyploc f hyploc acc decl = + let open Context.Named.Declaration in + let f = f (Some (get_id decl, hyploc)) in + match decl,hyploc with + | LocalAssum (id,_), InHypValueOnly -> error_occurrences_error (IncorrectInValueOccurrence id) - | None, _ | Some _, InHypTypeOnly -> - let acc,typ = f acc typ in acc,(id,bodyopt,typ) - | Some body, InHypValueOnly -> - let acc,body = f acc body in acc,(id,Some body,typ) - | Some body, InHyp -> + | LocalAssum (id,typ), _ -> + let acc,typ = f acc typ in acc, LocalAssum (id,typ) + | LocalDef (id,body,typ), InHypTypeOnly -> + let acc,typ = f acc typ in acc, LocalDef (id,body,typ) + | LocalDef (id,body,typ), InHypValueOnly -> + let acc,body = f acc body in acc, LocalDef (id,body,typ) + | LocalDef (id,body,typ), InHyp -> let acc,body = f acc body in let acc,typ = f acc typ in - acc,(id,Some body,typ) + acc, LocalDef (id,body,typ) (** Finding a subterm up to some testing function *) diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml index fb45be663..713c99597 100644 --- a/pretyping/indrec.ml +++ b/pretyping/indrec.ml @@ -28,6 +28,7 @@ open Environ open Reductionops open Nametab open Sigma.Notations +open Context.Rel.Declaration type dep_flag = bool @@ -77,7 +78,6 @@ let mis_make_case_com dep env sigma (ind, u as pind) (mib,mip as specif) kind = (* mais pas très joli ... (mais manque get_sort_of à ce niveau) *) let env' = push_rel_context lnamespar env in - let rec add_branch env k = if Int.equal k (Array.length mip.mind_consnames) then let nbprod = k+1 in @@ -85,7 +85,7 @@ let mis_make_case_com dep env sigma (ind, u as pind) (mib,mip as specif) kind = let indf' = lift_inductive_family nbprod indf in let arsign,_ = get_arity env indf' in let depind = build_dependent_inductive env indf' in - let deparsign = (Anonymous,None,depind)::arsign in + let deparsign = LocalAssum (Anonymous,depind)::arsign in let ci = make_case_info env (fst pind) RegularStyle in let pbody = @@ -118,14 +118,14 @@ let mis_make_case_com dep env sigma (ind, u as pind) (mib,mip as specif) kind = let cs = lift_constructor (k+1) constrs.(k) in let t = build_branch_type env dep (mkRel (k+1)) cs in mkLambda_string "f" t - (add_branch (push_rel (Anonymous, None, t) env) (k+1)) + (add_branch (push_rel (LocalAssum (Anonymous, t)) env) (k+1)) in let Sigma (s, sigma, p) = Sigma.fresh_sort_in_family ~rigid:Evd.univ_flexible_alg env sigma kind in let typP = make_arity env' dep indf s in let c = it_mkLambda_or_LetIn_name env (mkLambda_string "P" typP - (add_branch (push_rel (Anonymous,None,typP) env') 0)) lnamespar + (add_branch (push_rel (LocalAssum (Anonymous,typP)) env') 0)) lnamespar in Sigma (c, sigma, p) @@ -154,10 +154,10 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs = let p',largs = whd_betadeltaiota_nolet_stack env sigma p in match kind_of_term p' with | Prod (n,t,c) -> - let d = (n,None,t) in + let d = LocalAssum (n,t) in make_prod env (n,t,prec (push_rel d env) (i+1) (d::sign) c) | LetIn (n,b,t,c) when List.is_empty largs -> - let d = (n,Some b,t) in + let d = LocalDef (n,b,t) in mkLetIn (n,b,t,prec (push_rel d env) (i+1) (d::sign) c) | Ind (_,_) -> let realargs = List.skipn nparams largs in @@ -192,22 +192,22 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs = | None -> make_prod env (n,t, - process_constr (push_rel (n,None,t) env) (i+1) c_0 rest + process_constr (push_rel (LocalAssum (n,t)) env) (i+1) c_0 rest (nhyps-1) (i::li)) | Some(dep',p) -> let nP = lift (i+1+decP) p in - let env' = push_rel (n,None,t) env in + let env' = push_rel (LocalAssum (n,t)) env in let t_0 = process_pos env' dep' nP (lift 1 t) in make_prod_dep (dep || dep') env (n,t, mkArrow t_0 (process_constr - (push_rel (Anonymous,None,t_0) env') + (push_rel (LocalAssum (Anonymous,t_0)) env') (i+2) (lift 1 c_0) rest (nhyps-1) (i::li)))) | LetIn (n,b,t,c_0) -> mkLetIn (n,b,t, process_constr - (push_rel (n,Some b,t) env) + (push_rel (LocalDef (n,b,t)) env) (i+1) c_0 recargs (nhyps-1) li) | _ -> assert false else @@ -232,10 +232,10 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs = let p',largs = whd_betadeltaiota_nolet_stack env sigma p in match kind_of_term p' with | Prod (n,t,c) -> - let d = (n,None,t) in + let d = LocalAssum (n,t) in mkLambda_name env (n,t,prec (push_rel d env) (i+1) (d::hyps) c) | LetIn (n,b,t,c) when List.is_empty largs -> - let d = (n,Some b,t) in + let d = LocalDef (n,b,t) in mkLetIn (n,b,t,prec (push_rel d env) (i+1) (d::hyps) c) | Ind _ -> let realargs = List.skipn nparrec largs @@ -250,7 +250,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs = in (* ici, cstrprods est la liste des produits du constructeur instantié *) let rec process_constr env i f = function - | (n,None,t as d)::cprest, recarg::rest -> + | (LocalAssum (n,t) as d)::cprest, recarg::rest -> let optionpos = match dest_recarg recarg with | Norec -> None @@ -271,7 +271,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs = (n,t,process_constr env' (i+1) (whd_beta Evd.empty (applist (lift 1 f, [(mkRel 1); arg]))) (cprest,rest))) - | (n,Some c,t as d)::cprest, rest -> + | (LocalDef (n,c,t) as d)::cprest, rest -> mkLetIn (n,c,t, process_constr (push_rel d env) (i+1) (lift 1 f) @@ -322,7 +322,7 @@ let mis_make_indrec env sigma listdepkind mib u = let arsign,_ = get_arity env indf in let depind = build_dependent_inductive env indf in - let deparsign = (Anonymous,None,depind)::arsign in + let deparsign = LocalAssum (Anonymous,depind)::arsign in let nonrecpar = Context.Rel.length lnonparrec in let larsign = Context.Rel.length deparsign in @@ -357,7 +357,7 @@ let mis_make_indrec env sigma listdepkind mib u = let depind' = build_dependent_inductive env indf' in let arsign',_ = get_arity env indf' in - let deparsign' = (Anonymous,None,depind')::arsign' in + let deparsign' = LocalAssum (Anonymous,depind')::arsign' in let pargs = let nrpar = Context.Rel.to_extended_list (2*ndepar) lnonparrec @@ -387,11 +387,13 @@ let mis_make_indrec env sigma listdepkind mib u = let branch = branches.(0) in let ctx, br = decompose_lam_assum branch in let n, subst = - List.fold_right (fun (na,b,t) (i, subst) -> - if b == None then - let t = mkProj (Projection.make ps.(i) true, mkRel 1) in - (i + 1, t :: subst) - else (i, mkRel 0 :: subst)) + List.fold_right (fun decl (i, subst) -> + match decl with + | LocalAssum (na,t) -> + let t = mkProj (Projection.make ps.(i) true, mkRel 1) in + i + 1, t :: subst + | LocalDef (na,b,t) -> + i, mkRel 0 :: subst) ctx (0, []) in let term = substl subst br in @@ -440,7 +442,7 @@ let mis_make_indrec env sigma listdepkind mib u = true dep env !evdref (vargs,depPvec,i+j) tyi cs recarg in mkLambda_string "f" p_0 - (onerec (push_rel (Anonymous,None,p_0) env) (j+1)) + (onerec (push_rel (LocalAssum (Anonymous,p_0)) env) (j+1)) in onerec env 0 | [] -> makefix i listdepkind @@ -454,7 +456,7 @@ let mis_make_indrec env sigma listdepkind mib u = in let typP = make_arity env dep indf s in mkLambda_string "P" typP - (put_arity (push_rel (Anonymous,None,typP) env) (i+1) rest) + (put_arity (push_rel (LocalAssum (Anonymous,typP)) env) (i+1) rest) | [] -> make_branch env 0 listdepkind in diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index bb38c72f2..80f1988a9 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -17,6 +17,7 @@ open Declarations open Declareops open Environ open Reductionops +open Context.Rel.Declaration (* The following three functions are similar to the ones defined in Inductive, but they expect an env *) @@ -389,7 +390,7 @@ let make_arity_signature env dep indf = if dep then (* We need names everywhere *) Namegen.name_context env - ((Anonymous,None,build_dependent_inductive env indf)::arsign) + ((LocalAssum (Anonymous,build_dependent_inductive env indf))::arsign) (* Costly: would be better to name once for all at definition time *) else (* No need to enforce names *) @@ -459,7 +460,7 @@ let is_predicate_explicitly_dep env pred arsign = let rec srec env pval arsign = let pv' = whd_betadeltaiota env Evd.empty pval in match kind_of_term pv', arsign with - | Lambda (na,t,b), (_,None,_)::arsign -> + | Lambda (na,t,b), (LocalAssum _)::arsign -> srec (push_rel_assum (na,t) env) b arsign | Lambda (na,_,t), _ -> @@ -539,11 +540,11 @@ let arity_of_case_predicate env (ind,params) dep k = that appear in the type of the inductive by the sort of the conclusion, and the other ones by fresh universes. *) let rec instantiate_universes env evdref scl is = function - | (_,Some _,_ as d)::sign, exp -> + | (LocalDef _ as d)::sign, exp -> d :: instantiate_universes env evdref scl is (sign, exp) | d::sign, None::exp -> d :: instantiate_universes env evdref scl is (sign, exp) - | (na,None,ty)::sign, Some l::exp -> + | (LocalAssum (na,ty))::sign, Some l::exp -> let ctx,_ = Reduction.dest_arity env ty in let u = Univ.Universe.make l in let s = @@ -557,7 +558,7 @@ let rec instantiate_universes env evdref scl is = function let evm = Evd.set_leq_sort env evm s (Sorts.sort_of_univ u) in evdref := evm; s in - (na,None,mkArity(ctx,s)):: instantiate_universes env evdref scl is (sign, exp) + (LocalAssum (na,mkArity(ctx,s))) :: instantiate_universes env evdref scl is (sign, exp) | sign, [] -> sign (* Uniform parameters are exhausted *) | [], _ -> assert false diff --git a/pretyping/nativenorm.ml b/pretyping/nativenorm.ml index 6d09d5698..8ddfeaf2f 100644 --- a/pretyping/nativenorm.ml +++ b/pretyping/nativenorm.ml @@ -19,6 +19,7 @@ open Util open Nativecode open Nativevalues open Nativelambda +open Context.Rel.Declaration (** This module implements normalization by evaluation to OCaml code *) @@ -121,9 +122,8 @@ let build_case_type dep p realargs c = else mkApp(p, realargs) (* TODO move this function *) -let type_of_rel env n = - let (_,_,ty) = lookup_rel n env in - lift n ty +let type_of_rel env n = + lookup_rel n env |> get_type |> lift n let type_of_prop = mkSort type1_sort @@ -132,8 +132,9 @@ let type_of_sort s = | Prop _ -> type_of_prop | Type u -> mkType (Univ.super u) -let type_of_var env id = - try let (_,_,ty) = lookup_named id env in ty +let type_of_var env id = + let open Context.Named.Declaration in + try lookup_named id env |> get_type with Not_found -> anomaly ~label:"type_of_var" (str "variable " ++ Id.print id ++ str " unbound") @@ -181,7 +182,7 @@ let rec nf_val env v typ = Errors.anomaly (Pp.strbrk "Returned a functional value in a type not recognized as a product type.") in - let env = push_rel (name,None,dom) env in + let env = push_rel (LocalAssum (name,dom)) env in let body = nf_val env (f (mk_rel_accu lvl)) codom in mkLambda(name,dom,body) | Vconst n -> construct_of_constr_const env n typ @@ -257,7 +258,7 @@ and nf_atom env atom = | Aprod(n,dom,codom) -> let dom = nf_type env dom in let vn = mk_rel_accu (nb_rel env) in - let env = push_rel (n,None,dom) env in + let env = push_rel (LocalAssum (n,dom)) env in let codom = nf_type env (codom vn) in mkProd(n,dom,codom) | Ameta (mv,_) -> mkMeta mv @@ -328,7 +329,7 @@ and nf_atom_type env atom = | Aprod(n,dom,codom) -> let dom,s1 = nf_type_sort env dom in let vn = mk_rel_accu (nb_rel env) in - let env = push_rel (n,None,dom) env in + let env = push_rel (LocalAssum (n,dom)) env in let codom,s2 = nf_type_sort env (codom vn) in mkProd(n,dom,codom), mkSort (sort_of_product env s1 s2) | Aevar(ev,ty) -> @@ -356,7 +357,7 @@ and nf_predicate env ind mip params v pT = (Pp.strbrk "Returned a functional value in a type not recognized as a product type.") in let dep,body = - nf_predicate (push_rel (name,None,dom) env) ind mip params vb codom in + nf_predicate (push_rel (LocalAssum (name,dom)) env) ind mip params vb codom in dep, mkLambda(name,dom,body) | Vfun f, _ -> let k = nb_rel env in @@ -366,7 +367,7 @@ and nf_predicate env ind mip params v pT = let rargs = Array.init n (fun i -> mkRel (n-i)) in let params = if Int.equal n 0 then params else Array.map (lift n) params in let dom = mkApp(mkIndU ind,Array.append params rargs) in - let body = nf_type (push_rel (name,None,dom) env) vb in + let body = nf_type (push_rel (LocalAssum (name,dom)) env) vb in true, mkLambda(name,dom,body) | _, _ -> false, nf_type env v diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index af46c390a..827071054 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -123,6 +123,7 @@ let head_of_constr_reference c = match kind_of_term c with let pattern_of_constr env sigma t = let rec pattern_of_constr env t = + let open Context.Rel.Declaration in match kind_of_term t with | Rel n -> PRel n | Meta n -> PMeta (Some (Id.of_string ("META" ^ string_of_int n))) @@ -132,11 +133,11 @@ let pattern_of_constr env sigma t = | Sort (Type _) -> PSort (GType []) | Cast (c,_,_) -> pattern_of_constr env c | LetIn (na,c,t,b) -> PLetIn (na,pattern_of_constr env c, - pattern_of_constr (push_rel (na,Some c,t) env) b) + pattern_of_constr (push_rel (LocalDef (na,c,t)) env) b) | Prod (na,c,b) -> PProd (na,pattern_of_constr env c, - pattern_of_constr (push_rel (na, None, c) env) b) + pattern_of_constr (push_rel (LocalAssum (na, c)) env) b) | Lambda (na,c,b) -> PLambda (na,pattern_of_constr env c, - pattern_of_constr (push_rel (na, None, c) env) b) + pattern_of_constr (push_rel (LocalAssum (na, c)) env) b) | App (f,a) -> (match match kind_of_term f with diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 11fba7b94..7c91b1a93 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -42,6 +42,7 @@ open Glob_ops open Evarconv open Pattern open Misctypes +open Context.Named.Declaration type typing_constraint = OfType of types | IsType | WithoutTypeConstraint type var_map = constr_under_binders Id.Map.t @@ -319,7 +320,7 @@ let ltac_interp_name_env k0 lvar env = let n = Context.Rel.length (rel_context env) - k0 in let ctxt,_ = List.chop n (rel_context env) in let env = pop_rel_context n env in - let ctxt = List.map (fun (na,c,t) -> ltac_interp_name lvar na,c,t) ctxt in + let ctxt = List.map (Context.Rel.Declaration.map_name (ltac_interp_name lvar)) ctxt in push_rel_context ctxt env let invert_ltac_bound_name lvar env id0 id = @@ -372,8 +373,7 @@ let pretype_id pretype k0 loc env evdref lvar id = str "Variable " ++ pr_id id ++ str " should be bound to a term."); (* Check if [id] is a section or goal variable *) try - let (_,_,typ) = lookup_named id env in - { uj_val = mkVar id; uj_type = typ } + { uj_val = mkVar id; uj_type = (get_type (lookup_named id env)) } with Not_found -> (* [id] not found, standard error message *) error_var_not_found_loc loc id @@ -418,8 +418,7 @@ let pretype_ref loc evdref env ref us = match ref with | VarRef id -> (* Section variable *) - (try let (_,_,ty) = lookup_named id env in - make_judge (mkVar id) ty + (try make_judge (mkVar id) (get_type (lookup_named id env)) with Not_found -> (* This may happen if env is a goal env and section variables have been cleared - section variables should be different from goal @@ -459,6 +458,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) env evdref (lvar : ltac_ let inh_conv_coerce_to_tycon = inh_conv_coerce_to_tycon resolve_tc in let pretype_type = pretype_type k0 resolve_tc in let pretype = pretype k0 resolve_tc in + let open Context.Rel.Declaration in match t with | GRef (loc,ref,u) -> inh_conv_coerce_to_tycon loc env evdref @@ -518,14 +518,14 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) env evdref (lvar : ltac_ [] -> ctxt | (na,bk,None,ty)::bl -> let ty' = pretype_type empty_valcon env evdref lvar ty in - let dcl = (na,None,ty'.utj_val) in - let dcl' = (ltac_interp_name lvar na,None,ty'.utj_val) in + let dcl = LocalAssum (na, ty'.utj_val) in + let dcl' = LocalAssum (ltac_interp_name lvar na,ty'.utj_val) in type_bl (push_rel dcl env) (Context.Rel.add dcl' ctxt) bl | (na,bk,Some bd,ty)::bl -> let ty' = pretype_type empty_valcon env evdref lvar ty in let bd' = pretype (mk_tycon ty'.utj_val) env evdref lvar bd in - let dcl = (na,Some bd'.uj_val,ty'.utj_val) in - let dcl' = (ltac_interp_name lvar na,Some bd'.uj_val,ty'.utj_val) in + let dcl = LocalDef (na, bd'.uj_val, ty'.utj_val) in + let dcl' = LocalDef (ltac_interp_name lvar na, bd'.uj_val, ty'.utj_val) in type_bl (push_rel dcl env) (Context.Rel.add dcl' ctxt) bl in let ctxtv = Array.map (type_bl env Context.Rel.empty) bl in let larj = @@ -694,7 +694,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) env evdref (lvar : ltac_ (* The name specified by ltac is used also to create bindings. So the substitution must also be applied on variables before they are looked up in the rel context. *) - let var = (name,None,j.utj_val) in + let var = LocalAssum (name, j.utj_val) in let j' = pretype rng (push_rel var env) evdref lvar c2 in let name = ltac_interp_name lvar name in let resj = judge_of_abstraction env (orelse_name name name') j j' in @@ -738,7 +738,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) env evdref (lvar : ltac_ (* The name specified by ltac is used also to create bindings. So the substitution must also be applied on variables before they are looked up in the rel context. *) - let var = (name,Some j.uj_val,t) in + let var = LocalDef (name, j.uj_val, t) in let tycon = lift_tycon 1 tycon in let j' = pretype tycon (push_rel var env) evdref lvar c2 in let name = ltac_interp_name lvar name in @@ -763,17 +763,17 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) env evdref (lvar : ltac_ int cs.cs_nargs ++ str " variables."); let fsign, record = match get_projections env indf with - | None -> List.map2 (fun na (_,c,t) -> (na,c,t)) - (List.rev nal) cs.cs_args, false + | None -> + List.map2 set_name (List.rev nal) cs.cs_args, false | Some ps -> let rec aux n k names l = match names, l with - | na :: names, ((_, None, t) :: l) -> + | na :: names, (LocalAssum (_,t) :: l) -> let proj = Projection.make ps.(cs.cs_nargs - k) true in - (na, Some (lift (cs.cs_nargs - n) (mkProj (proj, cj.uj_val))), t) + LocalDef (na, lift (cs.cs_nargs - n) (mkProj (proj, cj.uj_val)), t) :: aux (n+1) (k + 1) names l - | na :: names, ((_, c, t) :: l) -> - (na, c, t) :: aux (n+1) k names l + | na :: names, (decl :: l) -> + set_name na decl :: aux (n+1) k names l | [], [] -> [] | _ -> assert false in aux 1 1 (List.rev nal) cs.cs_args, true in @@ -781,7 +781,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) env evdref (lvar : ltac_ if not record then let nal = List.map (fun na -> ltac_interp_name lvar na) nal in let nal = List.rev nal in - let fsign = List.map2 (fun na (_,b,t) -> (na,b,t)) nal fsign in + let fsign = List.map2 set_name nal fsign in let f = it_mkLambda_or_LetIn f fsign in let ci = make_case_info env (fst ind) LetStyle in mkCase (ci, p, cj.uj_val,[|f|]) @@ -792,10 +792,10 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) env evdref (lvar : ltac_ let arsgn = let arsgn,_ = get_arity env indf in if not !allow_anonymous_refs then - List.map (fun (_,b,t) -> (Anonymous,b,t)) arsgn + List.map (set_name Anonymous) arsgn else arsgn in - let psign = (na,None,build_dependent_inductive env indf)::arsgn in + let psign = LocalAssum (na, build_dependent_inductive env indf) :: arsgn in let nar = List.length arsgn in (match po with | Some p -> @@ -851,11 +851,11 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) env evdref (lvar : ltac_ let arsgn,_ = get_arity env indf in if not !allow_anonymous_refs then (* Make dependencies from arity signature impossible *) - List.map (fun (_,b,t) -> (Anonymous,b,t)) arsgn + List.map (set_name Anonymous) arsgn else arsgn in let nar = List.length arsgn in - let psign = (na,None,build_dependent_inductive env indf)::arsgn in + let psign = LocalAssum (na, build_dependent_inductive env indf) :: arsgn in let pred,p = match po with | Some p -> let env_p = push_rel_context psign env in @@ -880,14 +880,11 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) env evdref (lvar : ltac_ let pi = beta_applist (pi, [build_dependent_constructor cs]) in let csgn = if not !allow_anonymous_refs then - List.map (fun (_,b,t) -> (Anonymous,b,t)) cs.cs_args + List.map (set_name Anonymous) cs.cs_args else - List.map - (fun (n, b, t) -> - match n with - Name _ -> (n, b, t) - | Anonymous -> (Name Namegen.default_non_dependent_ident, b, t)) - cs.cs_args + List.map (map_name (function Name _ as n -> n + | Anonymous -> Name Namegen.default_non_dependent_ident)) + cs.cs_args in let env_c = push_rel_context csgn env in let bj = pretype (mk_tycon pi) env_c evdref lvar b in @@ -949,8 +946,9 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) env evdref (lvar : ltac_ in inh_conv_coerce_to_tycon loc env evdref cj tycon and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update = - let f (id,_,t) (subst,update) = - let t = replace_vars subst t in + let f decl (subst,update) = + let id = get_id decl in + let t = replace_vars subst (get_type decl) in let c, update = try let c = List.assoc id update in @@ -962,7 +960,7 @@ and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update = if is_conv env !evdref t t' then mkRel n, update else raise Not_found with Not_found -> try - let (_,_,t') = lookup_named id env in + let t' = lookup_named id env |> get_type in if is_conv env !evdref t t' then mkVar id, update else raise Not_found with Not_found -> user_err_loc (loc,"",str "Cannot interpret " ++ diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index 5e21154a6..d7637d1c2 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -15,6 +15,7 @@ open Termops open Univ open Evd open Environ +open Context.Rel.Declaration exception Elimconst @@ -607,7 +608,7 @@ let strong whdfun env sigma t = strongrec env t let local_strong whdfun sigma = - let rec strongrec t = map_constr strongrec (whdfun sigma t) in + let rec strongrec t = Constr.map strongrec (whdfun sigma t) in strongrec let rec strong_prodspine redfun sigma c = @@ -799,6 +800,7 @@ let equal_stacks (x, l) (y, l') = | Some (lft1,lft2) -> f_equal (x, lft1) (y, lft2) let rec whd_state_gen ?csts tactic_mode flags env sigma = + let open Context.Named.Declaration in let rec whrec cst_l (x, stack as s) = let () = if !debug_RAKAM then let open Pp in @@ -815,11 +817,11 @@ let rec whd_state_gen ?csts tactic_mode flags env sigma = match kind_of_term x with | Rel n when Closure.RedFlags.red_set flags Closure.RedFlags.fDELTA -> (match lookup_rel n env with - | (_,Some body,_) -> whrec Cst_stack.empty (lift n body, stack) + | LocalDef (_,body,_) -> whrec Cst_stack.empty (lift n body, stack) | _ -> fold ()) | Var id when Closure.RedFlags.red_set flags (Closure.RedFlags.fVAR id) -> (match lookup_named id env with - | (_,Some body,_) -> whrec (Cst_stack.add_cst (mkVar id) cst_l) (body, stack) + | LocalDef (_,body,_) -> whrec (Cst_stack.add_cst (mkVar id) cst_l) (body, stack) | _ -> fold ()) | Evar ev -> (match safe_evar_value sigma ev with @@ -922,7 +924,7 @@ let rec whd_state_gen ?csts tactic_mode flags env sigma = | Some _ when Closure.RedFlags.red_set flags Closure.RedFlags.fBETA -> apply_subst whrec [] cst_l x stack | None when Closure.RedFlags.red_set flags Closure.RedFlags.fETA -> - let env' = push_rel (na,None,t) env in + let env' = push_rel (LocalAssum (na,t)) env in let whrec' = whd_state_gen tactic_mode flags env' sigma in (match kind_of_term (Stack.zip ~refold:true (fst (whrec' (c, Stack.empty)))) with | App (f,cl) -> @@ -1442,7 +1444,7 @@ let splay_prod env sigma = let t = whd_betadeltaiota env sigma c in match kind_of_term t with | Prod (n,a,c0) -> - decrec (push_rel (n,None,a) env) + decrec (push_rel (LocalAssum (n,a)) env) ((n,a)::m) c0 | _ -> m,t in @@ -1453,7 +1455,7 @@ let splay_lam env sigma = let t = whd_betadeltaiota env sigma c in match kind_of_term t with | Lambda (n,a,c0) -> - decrec (push_rel (n,None,a) env) + decrec (push_rel (LocalAssum (n,a)) env) ((n,a)::m) c0 | _ -> m,t in @@ -1464,11 +1466,11 @@ let splay_prod_assum env sigma = let t = whd_betadeltaiota_nolet env sigma c in match kind_of_term t with | Prod (x,t,c) -> - prodec_rec (push_rel (x,None,t) env) - (Context.Rel.add (x, None, t) l) c + prodec_rec (push_rel (LocalAssum (x,t)) env) + (Context.Rel.add (LocalAssum (x,t)) l) c | LetIn (x,b,t,c) -> - prodec_rec (push_rel (x, Some b, t) env) - (Context.Rel.add (x, Some b, t) l) c + prodec_rec (push_rel (LocalDef (x,b,t)) env) + (Context.Rel.add (LocalDef (x,b,t)) l) c | Cast (c,_,_) -> prodec_rec env l c | _ -> let t' = whd_betadeltaiota env sigma t in @@ -1489,8 +1491,8 @@ let splay_prod_n env sigma n = let rec decrec env m ln c = if Int.equal m 0 then (ln,c) else match kind_of_term (whd_betadeltaiota env sigma c) with | Prod (n,a,c0) -> - decrec (push_rel (n,None,a) env) - (m-1) (Context.Rel.add (n,None,a) ln) c0 + decrec (push_rel (LocalAssum (n,a)) env) + (m-1) (Context.Rel.add (LocalAssum (n,a)) ln) c0 | _ -> invalid_arg "splay_prod_n" in decrec env n Context.Rel.empty @@ -1499,8 +1501,8 @@ let splay_lam_n env sigma n = let rec decrec env m ln c = if Int.equal m 0 then (ln,c) else match kind_of_term (whd_betadeltaiota env sigma c) with | Lambda (n,a,c0) -> - decrec (push_rel (n,None,a) env) - (m-1) (Context.Rel.add (n,None,a) ln) c0 + decrec (push_rel (LocalAssum (n,a)) env) + (m-1) (Context.Rel.add (LocalAssum (n,a)) ln) c0 | _ -> invalid_arg "splay_lam_n" in decrec env n Context.Rel.empty @@ -1538,8 +1540,8 @@ let find_conclusion env sigma = let rec decrec env c = let t = whd_betadeltaiota env sigma c in match kind_of_term t with - | Prod (x,t,c0) -> decrec (push_rel (x,None,t) env) c0 - | Lambda (x,t,c0) -> decrec (push_rel (x,None,t) env) c0 + | Prod (x,t,c0) -> decrec (push_rel (LocalAssum (x,t)) env) c0 + | Lambda (x,t,c0) -> decrec (push_rel (LocalAssum (x,t)) env) c0 | t -> t in decrec env @@ -1623,7 +1625,7 @@ let meta_reducible_instance evd b = with | Some g -> irec (mkProj (p,g)) | None -> mkProj (p,c)) - | _ -> map_constr irec u + | _ -> Constr.map irec u in if Metaset.is_empty fm then (* nf_betaiota? *) b.rebus else irec b.rebus diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index cb4e588ee..1a6f7832a 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -18,6 +18,7 @@ open Reductionops open Environ open Termops open Arguments_renaming +open Context.Rel.Declaration type retype_error = | NotASort @@ -71,13 +72,14 @@ let rec subst_type env sigma typ = function let sort_of_atomic_type env sigma ft args = let rec concl_of_arity env n ar args = match kind_of_term (whd_betadeltaiota env sigma ar), args with - | Prod (na, t, b), h::l -> concl_of_arity (push_rel (na,Some (lift n h),t) env) (n + 1) b l + | Prod (na, t, b), h::l -> concl_of_arity (push_rel (LocalDef (na, lift n h, t)) env) (n + 1) b l | Sort s, [] -> s | _ -> retype_error NotASort in concl_of_arity env 0 ft (Array.to_list args) let type_of_var env id = - try let (_,_,ty) = lookup_named id env in ty + let open Context.Named.Declaration in + try get_type (lookup_named id env) with Not_found -> retype_error (BadVariable id) let decomp_sort env sigma t = @@ -86,13 +88,13 @@ let decomp_sort env sigma t = | _ -> retype_error NotASort let retype ?(polyprop=true) sigma = - let rec type_of env cstr= + let rec type_of env cstr = match kind_of_term cstr with | Meta n -> (try strip_outer_cast (Evd.meta_ftype sigma n).Evd.rebus with Not_found -> retype_error (BadMeta n)) | Rel n -> - let (_,_,ty) = lookup_rel n env in + let ty = get_type (lookup_rel n env) in lift n ty | Var id -> type_of_var env id | Const cst -> rename_type_of_constant env cst @@ -115,9 +117,9 @@ let retype ?(polyprop=true) sigma = | Prod _ -> whd_beta sigma (applist (t, [c])) | _ -> t) | Lambda (name,c1,c2) -> - mkProd (name, c1, type_of (push_rel (name,None,c1) env) c2) + mkProd (name, c1, type_of (push_rel (LocalAssum (name,c1)) env) c2) | LetIn (name,b,c1,c2) -> - subst1 b (type_of (push_rel (name,Some b,c1) env) c2) + subst1 b (type_of (push_rel (LocalDef (name,b,c1)) env) c2) | Fix ((_,i),(_,tys,_)) -> tys.(i) | CoFix (i,(_,tys,_)) -> tys.(i) | App(f,args) when is_template_polymorphic env f -> @@ -140,7 +142,7 @@ let retype ?(polyprop=true) sigma = | 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 + (match (sort_of env t, sort_of (push_rel (LocalAssum (name,t)) env) c2) with | _, (Prop Null as s) -> s | Prop _, (Prop Pos as s) -> s | Type _, (Prop Pos as s) when is_impredicative_set env -> s @@ -161,7 +163,7 @@ let retype ?(polyprop=true) sigma = | Sort (Prop c) -> InType | Sort (Type u) -> InType | Prod (name,t,c2) -> - let s2 = sort_family_of (push_rel (name,None,t) env) c2 in + let s2 = sort_family_of (push_rel (LocalAssum (name,t)) env) c2 in if not (is_impredicative_set env) && s2 == InSet && sort_family_of env t == InType then InType else s2 | App(f,args) when is_template_polymorphic env f -> @@ -235,9 +237,9 @@ let get_judgment_of env evc c = { uj_val = c; uj_type = get_type_of env evc c } let sorts_of_context env evc ctxt = let rec aux = function | [] -> env,[] - | (_,_,t as d)::ctxt -> + | d :: ctxt -> let env,sorts = aux ctxt in - let s = get_sort_of env evc t in + let s = get_sort_of env evc (get_type d) in (push_rel d env,s::sorts) in snd (aux ctxt) diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index bd46911c9..ae224cf0d 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -53,12 +53,13 @@ let is_evaluable env = function | EvalVarRef id -> is_evaluable_var env id let value_of_evaluable_ref env evref u = + let open Context.Named.Declaration in match evref with | EvalConstRef con -> (try constant_value_in env (con,u) with NotEvaluableConst IsProj -> raise (Invalid_argument "value_of_evaluable_ref")) - | EvalVarRef id -> Option.get (pi2 (lookup_named id env)) + | EvalVarRef id -> lookup_named id env |> get_value |> Option.get let evaluable_of_global_reference env = function | ConstRef cst when is_evaluable_const env cst -> EvalConstRef cst @@ -103,29 +104,29 @@ let destEvalRefU c = match kind_of_term c with | Evar ev -> (EvalEvar ev, Univ.Instance.empty) | _ -> anomaly (Pp.str "Not an unfoldable reference") -let unsafe_reference_opt_value env sigma eval = +let unsafe_reference_opt_value env sigma eval = match eval with | EvalConst cst -> (match (lookup_constant cst env).Declarations.const_body with | Declarations.Def c -> Some (Mod_subst.force_constr c) | _ -> None) | EvalVar id -> - let (_,v,_) = lookup_named id env in - v + let open Context.Named.Declaration in + lookup_named id env |> get_value | EvalRel n -> - let (_,v,_) = lookup_rel n env in - Option.map (lift n) v + let open Context.Rel.Declaration in + lookup_rel n env |> map_value (lift n) |> get_value | EvalEvar ev -> Evd.existential_opt_value sigma ev let reference_opt_value env sigma eval u = match eval with | EvalConst cst -> constant_opt_value_in env (cst,u) | EvalVar id -> - let (_,v,_) = lookup_named id env in - v + let open Context.Named.Declaration in + lookup_named id env |> get_value | EvalRel n -> - let (_,v,_) = lookup_rel n env in - Option.map (lift n) v + let open Context.Rel.Declaration in + lookup_rel n env |> map_value (lift n) |> get_value | EvalEvar ev -> Evd.existential_opt_value sigma ev exception NotEvaluable @@ -258,7 +259,8 @@ let compute_consteval_direct env sigma ref = let c',l = whd_betadelta_stack env sigma c in match kind_of_term c' with | Lambda (id,t,g) when List.is_empty l && not onlyproj -> - srec (push_rel (id,None,t) env) (n+1) (t::labs) onlyproj g + let open Context.Rel.Declaration in + srec (push_rel (LocalAssum (id,t)) env) (n+1) (t::labs) onlyproj g | Fix fix when not onlyproj -> (try check_fix_reversibility labs l fix with Elimconst -> NotAnElimination) @@ -277,7 +279,8 @@ let compute_consteval_mutual_fix env sigma ref = let nargs = List.length l in match kind_of_term c' with | Lambda (na,t,g) when List.is_empty l -> - srec (push_rel (na,None,t) env) (minarg+1) (t::labs) ref g + let open Context.Rel.Declaration in + srec (push_rel (LocalAssum (na,t)) env) (minarg+1) (t::labs) ref g | Fix ((lv,i),(names,_,_)) -> (* Last known constant wrapping Fix is ref = [labs](Fix l) *) (match compute_consteval_direct env sigma ref with @@ -371,7 +374,8 @@ let make_elim_fun (names,(nbfix,lv,n)) u largs = let dummy = mkProp let vfx = Id.of_string "_expanded_fix_" let vfun = Id.of_string "_eliminator_function_" -let venv = val_of_named_context [(vfx, None, dummy); (vfun, None, dummy)] +let venv = let open Context.Named.Declaration in + val_of_named_context [LocalAssum (vfx, dummy); LocalAssum (vfun, dummy)] (* Mark every occurrence of substituted vars (associated to a function) as a problem variable: an evar that can be instantiated either by @@ -534,9 +538,11 @@ let match_eval_ref_value env sigma constr = | Const (sp, u) when is_evaluable env (EvalConstRef sp) -> Some (constant_value_in env (sp, u)) | Var id when is_evaluable env (EvalVarRef id) -> - let (_,v,_) = lookup_named id env in v - | Rel n -> let (_,v,_) = lookup_rel n env in - Option.map (lift n) v + let open Context.Named.Declaration in + lookup_named id env |> get_value + | Rel n -> + let open Context.Rel.Declaration in + lookup_rel n env |> map_value (lift n) |> get_value | Evar ev -> Evd.existential_opt_value sigma ev | _ -> None @@ -601,12 +607,14 @@ let whd_nothing_for_iota env sigma s = let rec whrec (x, stack as s) = match kind_of_term x with | Rel n -> + let open Context.Rel.Declaration in (match lookup_rel n env with - | (_,Some body,_) -> whrec (lift n body, stack) + | LocalDef (_,body,_) -> whrec (lift n body, stack) | _ -> s) | Var id -> + let open Context.Named.Declaration in (match lookup_named id env with - | (_,Some body,_) -> whrec (body, stack) + | LocalDef (_,body,_) -> whrec (body, stack) | _ -> s) | Evar ev -> (try whrec (Evd.existential_value sigma ev, stack) @@ -809,7 +817,9 @@ let try_red_product env sigma c = simpfun (Stack.zip (f,stack'))) | _ -> simpfun (appvect (redrec env f, l))) | Cast (c,_,_) -> redrec env c - | Prod (x,a,b) -> mkProd (x, a, redrec (push_rel (x,None,a) env) b) + | Prod (x,a,b) -> + let open Context.Rel.Declaration in + mkProd (x, a, redrec (push_rel (LocalAssum (x,a)) env) b) | LetIn (x,a,b,t) -> redrec env (subst1 a t) | Case (ci,p,d,lf) -> simpfun (mkCase (ci,p,redrec env d,lf)) | Proj (p, c) -> @@ -1157,8 +1167,9 @@ let reduce_to_ind_gen allow_product env sigma t = match kind_of_term (fst (decompose_app t)) with | Ind ind-> (check_privacy env ind, it_mkProd_or_LetIn t l) | Prod (n,ty,t') -> + let open Context.Rel.Declaration in if allow_product then - elimrec (push_rel (n,None,ty) env) t' ((n,None,ty)::l) + elimrec (push_rel (LocalAssum (n,ty)) env) t' ((LocalAssum (n,ty))::l) else errorlabstrm "" (str"Not an inductive definition.") | _ -> @@ -1235,7 +1246,8 @@ let reduce_to_ref_gen allow_product env sigma ref t = match kind_of_term c with | Prod (n,ty,t') -> if allow_product then - elimrec (push_rel (n,None,t) env) t' ((n,None,ty)::l) + let open Context.Rel.Declaration in + elimrec (push_rel (LocalAssum (n,t)) env) t' ((LocalAssum (n,ty))::l) else error_cannot_recognize ref | _ -> diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml index 3d6196c35..6c62bd08f 100644 --- a/pretyping/typeclasses.ml +++ b/pretyping/typeclasses.ml @@ -16,6 +16,7 @@ open Evd open Util open Typeclasses_errors open Libobject +open Context.Rel.Declaration (*i*) let typeclasses_unique_solutions = ref false @@ -180,9 +181,7 @@ let subst_class (subst,cl) = let do_subst_con c = Mod_subst.subst_constant subst c and do_subst c = Mod_subst.subst_mps subst c and do_subst_gr gr = fst (subst_global subst gr) in - let do_subst_ctx ctx = List.smartmap - (fun (na, b, t) -> (na, Option.smartmap do_subst b, do_subst t)) - ctx in + let do_subst_ctx = List.smartmap (map_constr do_subst) in let do_subst_context (grs,ctx) = List.smartmap (Option.smartmap (fun (gr,b) -> do_subst_gr gr, b)) grs, do_subst_ctx ctx in @@ -199,15 +198,19 @@ let discharge_class (_,cl) = let repl = Lib.replacement_context () in let rel_of_variable_context ctx = List.fold_right ( fun (n,_,b,t) (ctx', subst) -> - let decl = (Name n, Option.map (substn_vars 1 subst) b, substn_vars 1 subst t) in + let decl = match b with + | None -> LocalAssum (Name n, substn_vars 1 subst t) + | Some b -> LocalDef (Name n, substn_vars 1 subst b, substn_vars 1 subst t) + in (decl :: ctx', n :: subst) ) ctx ([], []) in let discharge_rel_context subst n rel = let rel = Context.Rel.map (Cooking.expmod_constr repl) rel in let ctx, _ = List.fold_right - (fun (id, b, t) (ctx, k) -> - (id, Option.smartmap (substn_vars k subst) b, substn_vars k subst t) :: ctx, succ k) + (fun decl (ctx, k) -> + map_constr (substn_vars k subst) decl :: ctx, succ k + ) rel ([], n) in ctx in @@ -217,15 +220,15 @@ let discharge_class (_,cl) = | ConstRef cst -> Lib.section_segment_of_constant cst | IndRef (ind,_) -> Lib.section_segment_of_mutual_inductive ind in let discharge_context ctx' subst (grs, ctx) = - let grs' = - let newgrs = List.map (fun (_, _, t) -> - match class_of_constr t with - | None -> None - | Some (_, ((tc,_), _)) -> Some (tc.cl_impl, true)) - ctx' + let grs' = + let newgrs = List.map (fun decl -> + match decl |> get_type |> class_of_constr with + | None -> None + | Some (_, ((tc,_), _)) -> Some (tc.cl_impl, true)) + ctx' in - List.smartmap (Option.smartmap (fun (gr, b) -> Lib.discharge_global gr, b)) grs - @ newgrs + List.smartmap (Option.smartmap (fun (gr, b) -> Lib.discharge_global gr, b)) grs + @ newgrs in grs', discharge_rel_context subst 1 ctx @ ctx' in let cl_impl' = Lib.discharge_global cl.cl_impl in if cl_impl' == cl.cl_impl then cl else @@ -431,11 +434,7 @@ let add_class cl = *) let instance_constructor (cl,u) args = - let filter (_, b, _) = match b with - | None -> true - | Some _ -> false - in - let lenpars = List.count filter (snd cl.cl_context) in + let lenpars = List.count is_local_assum (snd cl.cl_context) in let pars = fst (List.chop lenpars args) in match cl.cl_impl with | IndRef ind -> diff --git a/pretyping/typing.ml b/pretyping/typing.ml index fb0c49320..8be28a620 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -18,6 +18,7 @@ open Inductive open Inductiveops open Typeops open Arguments_renaming +open Context.Rel.Declaration let meta_type evd mv = let ty = @@ -88,16 +89,16 @@ let e_is_correct_arity env evdref c pj ind specif params = let rec srec env pt ar = let pt' = whd_betadeltaiota env !evdref pt in match kind_of_term pt', ar with - | Prod (na1,a1,t), (_,None,a1')::ar' -> + | Prod (na1,a1,t), (LocalAssum (_,a1'))::ar' -> if not (Evarconv.e_cumul env evdref a1 a1') then error (); - srec (push_rel (na1,None,a1) env) t ar' + srec (push_rel (LocalAssum (na1,a1)) env) t ar' | Sort s, [] -> if not (Sorts.List.mem (Sorts.family s) allowed_sorts) then error () | Evar (ev,_), [] -> let evd, s = Evd.fresh_sort_in_family env !evdref (max_sort allowed_sorts) in evdref := Evd.define ev (mkSort s) evd - | _, (_,Some _,_ as d)::ar' -> + | _, (LocalDef _ as d)::ar' -> srec (push_rel d env) (lift 1 pt') ar' | _ -> error () @@ -229,14 +230,14 @@ let rec execute env evdref cstr = | Lambda (name,c1,c2) -> let j = execute env evdref c1 in let var = e_type_judgment env evdref j in - let env1 = push_rel (name,None,var.utj_val) env in + let env1 = push_rel (LocalAssum (name, var.utj_val)) env in let j' = execute env1 evdref c2 in judge_of_abstraction env1 name var j' | Prod (name,c1,c2) -> let j = execute env evdref c1 in let varj = e_type_judgment env evdref j in - let env1 = push_rel (name,None,varj.utj_val) env in + let env1 = push_rel (LocalAssum (name, varj.utj_val)) env in let j' = execute env1 evdref c2 in let varj' = e_type_judgment env1 evdref j' in judge_of_product env name varj varj' @@ -246,7 +247,7 @@ let rec execute env evdref cstr = let j2 = execute env evdref c2 in let j2 = e_type_judgment env evdref j2 in let _ = e_judge_of_cast env evdref j1 DEFAULTcast j2 in - let env1 = push_rel (name,Some j1.uj_val,j2.utj_val) env in + let env1 = push_rel (LocalDef (name, j1.uj_val, j2.utj_val)) env in let j3 = execute env1 evdref c3 in judge_of_letin env name j1 j2 j3 diff --git a/pretyping/unification.ml b/pretyping/unification.ml index b5e882bc4..6614749d0 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -28,6 +28,7 @@ open Locus open Locusops open Find_subterm open Sigma.Notations +open Context.Named.Declaration let keyed_unification = ref (false) let _ = Goptions.declare_bool_option { @@ -58,7 +59,7 @@ let occur_meta_or_undefined_evar evd c = | Evar_defined c -> occrec c; Array.iter occrec args | Evar_empty -> raise Occur) - | _ -> iter_constr occrec c + | _ -> Constr.iter occrec c in try occrec c; false with Occur | Not_found -> true let occur_meta_evd sigma mv c = @@ -67,7 +68,7 @@ let occur_meta_evd sigma mv c = let c = whd_evar sigma (whd_meta sigma c) in match kind_of_term c with | Meta mv' when Int.equal mv mv' -> raise Occur - | _ -> iter_constr occrec c + | _ -> Constr.iter occrec c in try occrec c; false with Occur -> true (* if lname_typ is [xn,An;..;x1,A1] and l is a list of terms, @@ -75,7 +76,10 @@ let occur_meta_evd sigma mv c = let abstract_scheme env evd c l lname_typ = List.fold_left2 - (fun (t,evd) (locc,a) (na,_,ta) -> + (fun (t,evd) (locc,a) decl -> + let open Context.Rel.Declaration in + let na = get_name decl in + let ta = get_type decl in let na = match kind_of_term a with Var id -> Name id | _ -> na in (* [occur_meta ta] test removed for support of eelim/ecase but consequences are unclear... @@ -146,7 +150,7 @@ let rec subst_meta_instances bl c = | Meta i -> let select (j,_,_) = Int.equal i j in (try pi2 (List.find select bl) with Not_found -> c) - | _ -> map_constr (subst_meta_instances bl) c + | _ -> Constr.map (subst_meta_instances bl) c (** [env] should be the context in which the metas live *) @@ -164,7 +168,7 @@ let pose_all_metas_as_evars env evd t = evdref := meta_assign mv (ev,(Conv,TypeNotProcessed)) !evdref; ev) | _ -> - map_constr aux t in + Constr.map aux t in let c = aux t in (* side-effect *) (!evdref, c) @@ -568,8 +572,8 @@ let subst_defined_metas_evars (bl,el) c = | Evar (evk,args) -> let select (_,(evk',args'),_) = Evar.equal evk evk' && Array.equal Constr.equal args args' in (try substrec (pi3 (List.find select el)) - with Not_found -> map_constr substrec c) - | _ -> map_constr substrec c + with Not_found -> Constr.map substrec c) + | _ -> Constr.map substrec c in try Some (substrec c) with Not_found -> None let check_compatibility env pbty flags (sigma,metasubst,evarsubst) tyM tyN = @@ -1448,10 +1452,10 @@ let indirectly_dependent c d decls = it is needed otherwise, as e.g. when abstracting over "2" in "forall H:0=2, H=H:>(0=1+1) -> 0=2." where there is now obvious way to see that the second hypothesis depends indirectly over 2 *) - List.exists (fun (id,_,_) -> dependent_in_decl (mkVar id) d) decls + List.exists (fun d' -> dependent_in_decl (mkVar (get_id d')) d) decls let indirect_dependency d decls = - pi1 (List.hd (List.filter (fun (id,_,_) -> dependent_in_decl (mkVar id) d) decls)) + decls |> List.filter (fun d' -> dependent_in_decl (mkVar (get_id d')) d) |> List.hd |> get_id let finish_evar_resolution ?(flags=Pretyping.all_and_fail_flags) env current_sigma (pending,c) = let current_sigma = Sigma.to_evar_map current_sigma in @@ -1570,7 +1574,8 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = in let likefirst = clause_with_generic_occurrences occs in let mkvarid () = mkVar id in - let compute_dependency _ (hyp,_,_ as d) (sign,depdecls) = + let compute_dependency _ d (sign,depdecls) = + let hyp = get_id d in match occurrences_of_hyp hyp occs with | NoOccurrences, InHyp -> if indirectly_dependent c d depdecls then @@ -1607,7 +1612,7 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = replace_term_occ_modulo occ test mkvarid concl in let lastlhyp = - if List.is_empty depdecls then None else Some (pi1(List.last depdecls)) in + if List.is_empty depdecls then None else Some (get_id (List.last depdecls)) in let res = match out test with | None -> None | Some (sigma, c) -> Some (Sigma.Unsafe.of_pair (c, sigma)) diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml index 8b9c2d6c9..7ea9b9063 100644 --- a/pretyping/vnorm.ml +++ b/pretyping/vnorm.ml @@ -15,6 +15,7 @@ open Environ open Inductive open Reduction open Vm +open Context.Rel.Declaration (*******************************************) (* Calcul de la forme normal d'un terme *) @@ -134,7 +135,7 @@ and nf_whd env whd typ = let dom = nf_vtype env (dom p) in let name = Name (Id.of_string "x") in let vc = body_of_vfun (nb_rel env) (codom p) in - let codom = nf_vtype (push_rel (name,None,dom) env) vc in + let codom = nf_vtype (push_rel (LocalAssum (name,dom)) env) vc in mkProd(name,dom,codom) | Vfun f -> nf_fun env f typ | Vfix(f,None) -> nf_fix env f @@ -202,11 +203,12 @@ and constr_type_of_idkey env (idkey : Vars.id_key) stk = in nf_univ_args ~nb_univs mk env stk | VarKey id -> - let (_,_,ty) = lookup_named id env in + let open Context.Named.Declaration in + let ty = get_type (lookup_named id env) in nf_stk env (mkVar id) ty stk | RelKey i -> let n = (nb_rel env - i) in - let (_,_,ty) = lookup_rel n env in + let ty = get_type (lookup_rel n env) in nf_stk env (mkRel n) (lift n ty) stk and nf_stk ?from:(from=0) env c t stk = @@ -260,7 +262,7 @@ and nf_predicate env ind mip params v pT = let vb = body_of_vfun k f in let name,dom,codom = decompose_prod env pT in let dep,body = - nf_predicate (push_rel (name,None,dom) env) ind mip params vb codom in + nf_predicate (push_rel (LocalAssum (name,dom)) env) ind mip params vb codom in dep, mkLambda(name,dom,body) | Vfun f, _ -> let k = nb_rel env in @@ -270,7 +272,7 @@ and nf_predicate env ind mip params v pT = let rargs = Array.init n (fun i -> mkRel (n-i)) in let params = if Int.equal n 0 then params else Array.map (lift n) params in let dom = mkApp(mkIndU ind,Array.append params rargs) in - let body = nf_vtype (push_rel (name,None,dom) env) vb in + let body = nf_vtype (push_rel (LocalAssum (name,dom)) env) vb in true, mkLambda(name,dom,body) | _, _ -> false, nf_val env v crazy_type @@ -306,7 +308,7 @@ and nf_fun env f typ = Errors.anomaly (Pp.strbrk "Returned a functional value in a type not recognized as a product type.") in - let body = nf_val (push_rel (name,None,dom) env) vb codom in + let body = nf_val (push_rel (LocalAssum (name,dom)) env) vb codom in mkLambda(name,dom,body) and nf_fix env f = -- cgit v1.2.3