aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Matej Kosik <m4tej.kosik@gmail.com>2016-01-29 10:13:12 +0100
committerGravatar Matej Kosik <m4tej.kosik@gmail.com>2016-02-09 15:58:17 +0100
commit34ef02fac1110673ae74c41c185c228ff7876de2 (patch)
treea688eb9e2c23fc5353391f0c8b4ba1d7ba327844
parente9675e068f9e0e92bab05c030fb4722b146123b8 (diff)
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.
-rw-r--r--checker/cic.mli3
-rw-r--r--checker/closure.ml8
-rw-r--r--checker/declarations.ml13
-rw-r--r--checker/environ.ml2
-rw-r--r--checker/indtypes.ml20
-rw-r--r--checker/inductive.ml48
-rw-r--r--checker/reduction.ml10
-rw-r--r--checker/term.ml47
-rw-r--r--checker/term.mli2
-rw-r--r--checker/typeops.ml16
-rw-r--r--checker/values.ml6
-rw-r--r--dev/doc/changes.txt31
-rw-r--r--engine/evd.ml27
-rw-r--r--engine/namegen.ml15
-rw-r--r--engine/termops.ml199
-rw-r--r--engine/termops.mli2
-rw-r--r--ide/ide_slave.ml11
-rw-r--r--interp/constrextern.ml4
-rw-r--r--interp/constrintern.ml9
-rw-r--r--interp/implicit_quantifiers.ml19
-rw-r--r--interp/implicit_quantifiers.mli4
-rw-r--r--kernel/closure.ml10
-rw-r--r--kernel/context.ml364
-rw-r--r--kernel/context.mli151
-rw-r--r--kernel/cooking.ml7
-rw-r--r--kernel/csymtable.ml10
-rw-r--r--kernel/declareops.ml35
-rw-r--r--kernel/environ.ml64
-rw-r--r--kernel/fast_typeops.ml17
-rw-r--r--kernel/indtypes.ml35
-rw-r--r--kernel/inductive.ml50
-rw-r--r--kernel/nativecode.ml17
-rw-r--r--kernel/nativelambda.ml3
-rw-r--r--kernel/pre_env.ml7
-rw-r--r--kernel/reduction.ml11
-rw-r--r--kernel/safe_typing.ml8
-rw-r--r--kernel/term.ml111
-rw-r--r--kernel/term_typing.ml24
-rw-r--r--kernel/typeops.ml39
-rw-r--r--kernel/vars.ml11
-rw-r--r--lib/util.ml1
-rw-r--r--lib/util.mli1
-rw-r--r--library/decls.ml10
-rw-r--r--library/heads.ml11
-rw-r--r--library/impargs.ml15
-rw-r--r--library/lib.ml9
-rw-r--r--plugins/cc/ccalgo.ml2
-rw-r--r--plugins/cc/cctac.ml8
-rw-r--r--plugins/decl_mode/decl_interp.ml2
-rw-r--r--plugins/decl_mode/decl_mode.ml2
-rw-r--r--plugins/decl_mode/decl_proof_instr.ml25
-rw-r--r--plugins/derive/derive.ml4
-rw-r--r--plugins/extraction/extraction.ml9
-rw-r--r--plugins/firstorder/formula.ml11
-rw-r--r--plugins/firstorder/instances.ml3
-rw-r--r--plugins/firstorder/rules.ml6
-rw-r--r--plugins/fourier/fourierR.ml6
-rw-r--r--plugins/funind/functional_principles_proofs.ml60
-rw-r--r--plugins/funind/functional_principles_types.ml39
-rw-r--r--plugins/funind/glob_term_to_relation.ml52
-rw-r--r--plugins/funind/indfun.ml8
-rw-r--r--plugins/funind/invfun.ml40
-rw-r--r--plugins/funind/merge.ml24
-rw-r--r--plugins/funind/recdef.ml18
-rw-r--r--plugins/omega/coq_omega.ml37
-rw-r--r--plugins/rtauto/refl_tauto.ml5
-rw-r--r--plugins/rtauto/refl_tauto.mli2
-rw-r--r--pretyping/cases.ml196
-rw-r--r--pretyping/coercion.ml11
-rw-r--r--pretyping/constr_matching.ml15
-rw-r--r--pretyping/detyping.ml36
-rw-r--r--pretyping/evarconv.ml34
-rw-r--r--pretyping/evarsolve.ml131
-rw-r--r--pretyping/evarutil.ml94
-rw-r--r--pretyping/find_subterm.ml23
-rw-r--r--pretyping/indrec.ml48
-rw-r--r--pretyping/inductiveops.ml11
-rw-r--r--pretyping/nativenorm.ml21
-rw-r--r--pretyping/patternops.ml7
-rw-r--r--pretyping/pretyping.ml62
-rw-r--r--pretyping/reductionops.ml36
-rw-r--r--pretyping/retyping.ml22
-rw-r--r--pretyping/tacred.ml54
-rw-r--r--pretyping/typeclasses.ml37
-rw-r--r--pretyping/typing.ml13
-rw-r--r--pretyping/unification.ml27
-rw-r--r--pretyping/vnorm.ml14
-rw-r--r--printing/prettyp.ml31
-rw-r--r--printing/prettyp.mli2
-rw-r--r--printing/printer.ml18
-rw-r--r--proofs/goal.ml5
-rw-r--r--proofs/logic.ml51
-rw-r--r--proofs/proof_global.ml7
-rw-r--r--proofs/proof_using.ml8
-rw-r--r--proofs/proofview.ml22
-rw-r--r--proofs/refiner.ml6
-rw-r--r--proofs/tacmach.ml17
-rw-r--r--stm/lemmas.ml11
-rw-r--r--stm/stm.ml6
-rw-r--r--tactics/auto.ml2
-rw-r--r--tactics/autorewrite.ml5
-rw-r--r--tactics/class_tactics.ml11
-rw-r--r--tactics/contradiction.ml8
-rw-r--r--tactics/elim.ml4
-rw-r--r--tactics/eqschemes.ml29
-rw-r--r--tactics/equality.ml36
-rw-r--r--tactics/evar_tactics.ml7
-rw-r--r--tactics/hints.ml8
-rw-r--r--tactics/hipattern.ml410
-rw-r--r--tactics/inv.ml21
-rw-r--r--tactics/leminv.ml18
-rw-r--r--tactics/rewrite.ml40
-rw-r--r--tactics/tacinterp.ml6
-rw-r--r--tactics/tactic_matching.ml14
-rw-r--r--tactics/tacticals.ml12
-rw-r--r--tactics/tactics.ml327
-rw-r--r--toplevel/assumptions.ml7
-rw-r--r--toplevel/auto_ind_decl.ml43
-rw-r--r--toplevel/classes.ml44
-rw-r--r--toplevel/command.ml78
-rw-r--r--toplevel/discharge.ml12
-rw-r--r--toplevel/himsg.ml24
-rw-r--r--toplevel/indschemes.ml3
-rw-r--r--toplevel/obligations.ml33
-rw-r--r--toplevel/record.ml74
-rw-r--r--toplevel/search.ml9
-rw-r--r--toplevel/vernacentries.ml14
127 files changed, 2283 insertions, 1572 deletions
diff --git a/checker/cic.mli b/checker/cic.mli
index 041394d46..00ac2f56c 100644
--- a/checker/cic.mli
+++ b/checker/cic.mli
@@ -111,7 +111,8 @@ type cofixpoint = constr pcofixpoint
(** {6 Type of assumptions and contexts} *)
-type rel_declaration = Name.t * constr option * constr
+type rel_declaration = LocalAssum of Name.t * constr (* name, type *)
+ | LocalDef of Name.t * constr * constr (* name, value, type *)
type rel_context = rel_declaration list
(** The declarations below in .vo should be outside sections,
diff --git a/checker/closure.ml b/checker/closure.ml
index 400a535cf..c2708e97d 100644
--- a/checker/closure.ml
+++ b/checker/closure.ml
@@ -217,10 +217,10 @@ let ref_value_cache info ref =
let defined_rels flags env =
(* if red_local_const (snd flags) then*)
fold_rel_context
- (fun (id,b,t) (i,subs) ->
- match b with
- | None -> (i+1, subs)
- | Some body -> (i+1, (i,body) :: subs))
+ (fun decl (i,subs) ->
+ match decl with
+ | LocalAssum _ -> (i+1, subs)
+ | LocalDef (_,body,_) -> (i+1, (i,body) :: subs))
(rel_context env) ~init:(0,[])
(* else (0,[])*)
diff --git a/checker/declarations.ml b/checker/declarations.ml
index 32d1713a8..2f6eeba1d 100644
--- a/checker/declarations.ml
+++ b/checker/declarations.ml
@@ -517,11 +517,14 @@ let map_decl_arity f g = function
| RegularArity a -> RegularArity (f a)
| TemplateArity a -> TemplateArity (g a)
-
-let subst_rel_declaration sub (id,copt,t as x) =
- let copt' = Option.smartmap (subst_mps sub) copt in
- let t' = subst_mps sub t in
- if copt == copt' && t == t' then x else (id,copt',t')
+let subst_rel_declaration sub = function
+ | LocalAssum (id,t) as x ->
+ let t' = subst_mps sub t in
+ if t == t' then x else LocalAssum (id,t')
+ | LocalDef (id,c,t) as x ->
+ let c' = subst_mps sub c in
+ let t' = subst_mps sub t in
+ if c == c' && t == t' then x else LocalDef (id,c',t')
let subst_rel_context sub = List.smartmap (subst_rel_declaration sub)
diff --git a/checker/environ.ml b/checker/environ.ml
index f8f5c29b7..7040fdda4 100644
--- a/checker/environ.ml
+++ b/checker/environ.ml
@@ -80,7 +80,7 @@ let push_rel d env =
let push_rel_context ctxt x = fold_rel_context push_rel ctxt ~init:x
let push_rec_types (lna,typarray,_) env =
- let ctxt = Array.map2_i (fun i na t -> (na, None, lift i t)) lna typarray in
+ let ctxt = Array.map2_i (fun i na t -> LocalAssum (na, lift i t)) lna typarray in
Array.fold_left (fun e assum -> push_rel assum e) env ctxt
(* Universe constraints *)
diff --git a/checker/indtypes.ml b/checker/indtypes.ml
index 2865f5bd4..f11fa5a7a 100644
--- a/checker/indtypes.ml
+++ b/checker/indtypes.ml
@@ -56,10 +56,10 @@ let is_constructor_head t =
let conv_ctxt_prefix env (ctx1:rel_context) ctx2 =
let rec chk env rctx1 rctx2 =
match rctx1, rctx2 with
- (_,None,ty1 as d1)::rctx1', (_,None,ty2)::rctx2' ->
+ (LocalAssum (_,ty1) as d1)::rctx1', LocalAssum (_,ty2)::rctx2' ->
conv env ty1 ty2;
chk (push_rel d1 env) rctx1' rctx2'
- | (_,Some bd1,ty1 as d1)::rctx1', (_,Some bd2,ty2)::rctx2' ->
+ | (LocalDef (_,bd1,ty1) as d1)::rctx1', LocalDef (_,bd2,ty2)::rctx2' ->
conv env ty1 ty2;
conv env bd1 bd2;
chk (push_rel d1 env) rctx1' rctx2'
@@ -94,10 +94,10 @@ let rec sorts_of_constr_args env t =
match t with
| Prod (name,c1,c2) ->
let varj = infer_type env c1 in
- let env1 = push_rel (name,None,c1) env in
+ let env1 = push_rel (LocalAssum (name,c1)) env in
varj :: sorts_of_constr_args env1 c2
| LetIn (name,def,ty,c) ->
- let env1 = push_rel (name,Some def,ty) env in
+ let env1 = push_rel (LocalDef (name,def,ty)) env in
sorts_of_constr_args env1 c
| _ when is_constructor_head t -> []
| _ -> anomaly ~label:"infos_and_sort" (Pp.str "not a positive constructor")
@@ -167,7 +167,7 @@ let typecheck_arity env params inds =
full_arity is used as argument or subject to cast, an
upper universe will be generated *)
let id = ind.mind_typename in
- let env_ar' = push_rel (Name id, None, arity) env_ar in
+ let env_ar' = push_rel (LocalAssum (Name id, arity)) env_ar in
env_ar')
env
inds in
@@ -319,7 +319,7 @@ let check_correct_par (env,n,ntypes,_) hyps l largs =
let nhyps = List.length hyps in
let rec check k index = function
| [] -> ()
- | (_,Some _,_)::hyps -> check k (index+1) hyps
+ | LocalDef (_,_,_) :: hyps -> check k (index+1) hyps
| _::hyps ->
match whd_betadeltaiota env lpar.(k) with
| Rel w when w = index -> check (k-1) (index+1) hyps
@@ -340,7 +340,7 @@ let check_rec_par (env,n,_,_) hyps nrecp largs =
| ([],_) -> ()
| (_,[]) ->
failwith "number of recursive parameters cannot be greater than the number of parameters."
- | (lp,(_,Some _,_)::hyps) -> find (index-1) (lp,hyps)
+ | (lp,LocalDef _ :: hyps) -> find (index-1) (lp,hyps)
| (p::lp,_::hyps) ->
(match whd_betadeltaiota env p with
| Rel w when w = index -> find (index-1) (lp,hyps)
@@ -370,14 +370,14 @@ let abstract_mind_lc env ntyps npars lc =
[lra] is the list of recursive tree of each variable
*)
let ienv_push_var (env, n, ntypes, lra) (x,a,ra) =
- (push_rel (x,None,a) env, n+1, ntypes, (Norec,ra)::lra)
+ (push_rel (LocalAssum (x,a)) env, n+1, ntypes, (Norec,ra)::lra)
let ienv_push_inductive (env, n, ntypes, ra_env) ((mi,u),lpar) =
let auxntyp = 1 in
let specif = lookup_mind_specif env mi in
let env' =
- push_rel (Anonymous,None,
- hnf_prod_applist env (type_of_inductive env (specif,u)) lpar) env in
+ push_rel (LocalAssum (Anonymous,
+ hnf_prod_applist env (type_of_inductive env (specif,u)) lpar)) env in
let ra_env' =
(Imbr mi,(Rtree.mk_rec_calls 1).(0)) ::
List.map (fun (r,t) -> (r,Rtree.lift 1 t)) ra_env in
diff --git a/checker/inductive.ml b/checker/inductive.ml
index 79dba4fac..948012421 100644
--- a/checker/inductive.ml
+++ b/checker/inductive.ml
@@ -88,10 +88,10 @@ let instantiate_params full t u args sign =
anomaly ~label:"instantiate_params" (Pp.str "type, ctxt and args mismatch") in
let (rem_args, subs, ty) =
fold_rel_context
- (fun (_,copt,_) (largs,subs,ty) ->
- match (copt, largs, ty) with
- | (None, a::args, Prod(_,_,t)) -> (args, a::subs, t)
- | (Some b,_,LetIn(_,_,_,t)) ->
+ (fun decl (largs,subs,ty) ->
+ match (decl, largs, ty) with
+ | (LocalAssum _, a::args, Prod(_,_,t)) -> (args, a::subs, t)
+ | (LocalDef (_,b,_),_,LetIn(_,_,_,t)) ->
(largs, (substl subs (subst_instance_constr u b))::subs, t)
| (_,[],_) -> if full then fail() else ([], subs, ty)
| _ -> fail ())
@@ -161,7 +161,7 @@ let remember_subst u subst =
(* Propagate the new levels in the signature *)
let rec make_subst env =
let rec make subst = function
- | (_,Some _,_)::sign, exp, args ->
+ | LocalDef _ :: sign, exp, args ->
make subst (sign, exp, args)
| d::sign, None::exp, args ->
let args = match args with _::args -> args | [] -> [] in
@@ -174,7 +174,7 @@ let rec make_subst env =
(* a useless extra constraint *)
let s = sort_as_univ (snd (dest_arity env a)) in
make (cons_subst u s subst) (sign, exp, args)
- | (na,None,t)::sign, Some u::exp, [] ->
+ | LocalAssum (na,t) :: sign, Some u::exp, [] ->
(* No more argument here: we add the remaining universes to the *)
(* substitution (when [u] is distinct from all other universes in the *)
(* template, it is identity substitution otherwise (ie. when u is *)
@@ -319,8 +319,8 @@ let elim_sorts (_,mip) = mip.mind_kelim
let extended_rel_list n hyps =
let rec reln l p = function
- | (_,None,_) :: hyps -> reln (Rel (n+p) :: l) (p+1) hyps
- | (_,Some _,_) :: hyps -> reln l (p+1) hyps
+ | LocalAssum _ :: hyps -> reln (Rel (n+p) :: l) (p+1) hyps
+ | LocalDef _ :: hyps -> reln l (p+1) hyps
| [] -> l
in
reln [] 1 hyps
@@ -345,12 +345,12 @@ let is_correct_arity env c (p,pj) ind specif params =
let rec srec env pt ar =
let pt' = whd_betadeltaiota env pt in
match pt', ar with
- | Prod (na1,a1,t), (_,None,a1')::ar' ->
+ | Prod (na1,a1,t), LocalAssum (_,a1')::ar' ->
(try conv env a1 a1'
with NotConvertible -> raise (LocalArity None));
- srec (push_rel (na1,None,a1) env) t ar'
+ srec (push_rel (LocalAssum (na1,a1)) env) t ar'
| Prod (na1,a1,a2), [] -> (* whnf of t was not needed here! *)
- let env' = push_rel (na1,None,a1) env in
+ let env' = push_rel (LocalAssum (na1,a1)) env in
let ksort = match (whd_betadeltaiota env' a2) with
| Sort s -> family_of_sort s
| _ -> raise (LocalArity None) in
@@ -362,8 +362,8 @@ let is_correct_arity env c (p,pj) ind specif params =
| Sort s', [] ->
check_allowed_sort (family_of_sort s') specif;
false
- | _, (_,Some _,_ as d)::ar' ->
- srec (push_rel d env) (lift 1 pt') ar'
+ | _, (LocalDef _ as d)::ar' ->
+ srec (push_rel d env) (lift 1 pt') ar'
| _ ->
raise (LocalArity None)
in
@@ -530,7 +530,7 @@ let make_renv env recarg tree =
genv = [Lazy.lazy_from_val(Subterm(Large,tree))] }
let push_var renv (x,ty,spec) =
- { env = push_rel (x,None,ty) renv.env;
+ { env = push_rel (LocalAssum (x,ty)) renv.env;
rel_min = renv.rel_min+1;
genv = spec:: renv.genv }
@@ -628,14 +628,14 @@ let check_inductive_codomain env p =
(* The following functions are almost duplicated from indtypes.ml, except
that they carry here a poorer environment (containing less information). *)
let ienv_push_var (env, lra) (x,a,ra) =
-(push_rel (x,None,a) env, (Norec,ra)::lra)
+(push_rel (LocalAssum (x,a)) env, (Norec,ra)::lra)
let ienv_push_inductive (env, ra_env) ((mind,u),lpar) =
let mib = Environ.lookup_mind mind env in
let ntypes = mib.mind_ntypes in
let push_ind specif env =
- push_rel (Anonymous,None,
- hnf_prod_applist env (type_of_inductive env ((mib,specif),u)) lpar) env
+ push_rel (LocalAssum (Anonymous,
+ hnf_prod_applist env (type_of_inductive env ((mib,specif),u)) lpar)) env
in
let env = Array.fold_right push_ind mib.mind_packets env in
let rc = Array.mapi (fun j t -> (Imbr (mind,j),t)) (Rtree.mk_rec_calls ntypes) in
@@ -902,7 +902,7 @@ let filter_stack_domain env ci p stack =
let t = whd_betadeltaiota env ar in
match stack, t with
| elt :: stack', Prod (n,a,c0) ->
- let d = (n,None,a) in
+ let d = LocalAssum (n,a) in
let ty, args = decompose_app (whd_betadeltaiota env a) in
let elt = match ty with
| Ind ind ->
@@ -956,10 +956,10 @@ let check_one_fix renv recpos trees def =
end
else
begin
- match pi2 (lookup_rel p renv.env) with
- | None ->
+ match lookup_rel p renv.env with
+ | LocalAssum _ ->
List.iter (check_rec_call renv []) l
- | Some c ->
+ | LocalDef (_,c,_) ->
try List.iter (check_rec_call renv []) l
with FixGuardError _ ->
check_rec_call renv stack (applist(lift p c,l))
@@ -1078,7 +1078,7 @@ let inductive_of_mutfix env ((nvect,bodynum),(names,types,bodies as recdef)) =
match (whd_betadeltaiota env def) with
| Lambda (x,a,b) ->
if noccur_with_meta n nbfix a then
- let env' = push_rel (x, None, a) env in
+ let env' = push_rel (LocalAssum (x,a)) env in
if n = k+1 then
(* get the inductive type of the fixpoint *)
let (mind, _) =
@@ -1127,7 +1127,7 @@ let rec codomain_is_coind env c =
let b = whd_betadeltaiota env c in
match b with
| Prod (x,a,b) ->
- codomain_is_coind (push_rel (x, None, a) env) b
+ codomain_is_coind (push_rel (LocalAssum (x,a)) env) b
| _ ->
(try find_coinductive env b
with Not_found ->
@@ -1168,7 +1168,7 @@ let check_one_cofix env nbfix def deftype =
| Lambda (x,a,b) ->
assert (args = []);
if noccur_with_meta n nbfix a then
- let env' = push_rel (x, None, a) env in
+ let env' = push_rel (LocalAssum (x,a)) env in
check_rec_call env' alreadygrd (n+1) tree vlra b
else
raise (CoFixGuardError (env,RecCallInTypeOfAbstraction a))
diff --git a/checker/reduction.ml b/checker/reduction.ml
index 3a666a60a..f1aa5d919 100644
--- a/checker/reduction.ml
+++ b/checker/reduction.ml
@@ -490,7 +490,7 @@ let dest_prod env =
let t = whd_betadeltaiota env c in
match t with
| Prod (n,a,c0) ->
- let d = (n,None,a) in
+ let d = LocalAssum (n,a) in
decrec (push_rel d env) (d::m) c0
| _ -> m,t
in
@@ -502,10 +502,10 @@ let dest_prod_assum env =
let rty = whd_betadeltaiota_nolet env ty in
match rty with
| Prod (x,t,c) ->
- let d = (x,None,t) in
+ let d = LocalAssum (x,t) in
prodec_rec (push_rel d env) (d::l) c
| LetIn (x,b,t,c) ->
- let d = (x,Some b,t) in
+ let d = LocalDef (x,b,t) in
prodec_rec (push_rel d env) (d::l) c
| Cast (c,_,_) -> prodec_rec env l c
| _ ->
@@ -520,10 +520,10 @@ let dest_lam_assum env =
let rty = whd_betadeltaiota_nolet env ty in
match rty with
| Lambda (x,t,c) ->
- let d = (x,None,t) in
+ let d = LocalAssum (x,t) in
lamec_rec (push_rel d env) (d::l) c
| LetIn (x,b,t,c) ->
- let d = (x,Some b,t) in
+ let d = LocalDef (x,b,t) in
lamec_rec (push_rel d env) (d::l) c
| Cast (c,_,_) -> lamec_rec env l c
| _ -> l,rty
diff --git a/checker/term.ml b/checker/term.ml
index 6487d1a15..181d292ad 100644
--- a/checker/term.ml
+++ b/checker/term.ml
@@ -222,24 +222,29 @@ let rel_context_length = List.length
let rel_context_nhyps hyps =
let rec nhyps acc = function
| [] -> acc
- | (_,None,_)::hyps -> nhyps (1+acc) hyps
- | (_,Some _,_)::hyps -> nhyps acc hyps in
+ | LocalAssum _ :: hyps -> nhyps (1+acc) hyps
+ | LocalDef _ :: hyps -> nhyps acc hyps in
nhyps 0 hyps
let fold_rel_context f l ~init = List.fold_right f l init
let map_rel_context f l =
- let map_decl (n, body_o, typ as decl) =
- let body_o' = Option.smartmap f body_o in
- let typ' = f typ in
- if body_o' == body_o && typ' == typ then decl else
- (n, body_o', typ')
+ let map_decl = function
+ | LocalAssum (n, typ) as decl ->
+ let typ' = f typ in
+ if typ' == typ then decl else
+ LocalAssum (n, typ')
+ | LocalDef (n, body, typ) as decl ->
+ let body' = f body in
+ let typ' = f typ in
+ if body' == body && typ' == typ then decl else
+ LocalDef (n, body', typ')
in
List.smartmap map_decl l
let extended_rel_list n hyps =
let rec reln l p = function
- | (_,None,_) :: hyps -> reln (Rel (n+p) :: l) (p+1) hyps
- | (_,Some _,_) :: hyps -> reln l (p+1) hyps
+ | LocalAssum _ :: hyps -> reln (Rel (n+p) :: l) (p+1) hyps
+ | LocalDef _ :: hyps -> reln l (p+1) hyps
| [] -> l
in
reln [] 1 hyps
@@ -272,8 +277,8 @@ let decompose_lam_n_assum n =
let rec lamdec_rec l n c =
if Int.equal n 0 then l,c
else match c with
- | Lambda (x,t,c) -> lamdec_rec ((x,None,t) :: l) (n-1) c
- | LetIn (x,b,t,c) -> lamdec_rec ((x,Some b,t) :: l) n c
+ | Lambda (x,t,c) -> lamdec_rec (LocalAssum (x,t) :: l) (n-1) c
+ | LetIn (x,b,t,c) -> lamdec_rec (LocalDef (x,b,t) :: l) n c
| Cast (c,_,_) -> lamdec_rec l n c
| c -> error "decompose_lam_n_assum: not enough abstractions"
in
@@ -282,18 +287,18 @@ let decompose_lam_n_assum n =
(* Iterate products, with or without lets *)
(* Constructs either [(x:t)c] or [[x=b:t]c] *)
-let mkProd_or_LetIn (na,body,t) c =
- match body with
- | None -> Prod (na, t, c)
- | Some b -> LetIn (na, b, t, c)
+let mkProd_or_LetIn decl c =
+ match decl with
+ | LocalAssum (na,t) -> Prod (na, t, c)
+ | LocalDef (na,b,t) -> LetIn (na, b, t, c)
let it_mkProd_or_LetIn = List.fold_left (fun c d -> mkProd_or_LetIn d c)
let decompose_prod_assum =
let rec prodec_rec l c =
match c with
- | Prod (x,t,c) -> prodec_rec ((x,None,t) :: l) c
- | LetIn (x,b,t,c) -> prodec_rec ((x,Some b,t) :: l) c
+ | Prod (x,t,c) -> prodec_rec (LocalAssum (x,t) :: l) c
+ | LetIn (x,b,t,c) -> prodec_rec (LocalDef (x,b,t) :: l) c
| Cast (c,_,_) -> prodec_rec l c
| _ -> l,c
in
@@ -305,8 +310,8 @@ let decompose_prod_n_assum n =
let rec prodec_rec l n c =
if Int.equal n 0 then l,c
else match c with
- | Prod (x,t,c) -> prodec_rec ((x,None,t) :: l) (n-1) c
- | LetIn (x,b,t,c) -> prodec_rec ((x,Some b,t) :: l) (n-1) c
+ | Prod (x,t,c) -> prodec_rec (LocalAssum (x,t) :: l) (n-1) c
+ | LetIn (x,b,t,c) -> prodec_rec (LocalDef (x,b,t) :: l) (n-1) c
| Cast (c,_,_) -> prodec_rec l n c
| c -> error "decompose_prod_n_assum: not enough assumptions"
in
@@ -324,8 +329,8 @@ let mkArity (sign,s) = it_mkProd_or_LetIn (Sort s) sign
let destArity =
let rec prodec_rec l c =
match c with
- | Prod (x,t,c) -> prodec_rec ((x,None,t)::l) c
- | LetIn (x,b,t,c) -> prodec_rec ((x,Some b,t)::l) c
+ | Prod (x,t,c) -> prodec_rec (LocalAssum (x,t)::l) c
+ | LetIn (x,b,t,c) -> prodec_rec (LocalDef (x,b,t)::l) c
| Cast (c,_,_) -> prodec_rec l c
| Sort s -> l,s
| _ -> anomaly ~label:"destArity" (Pp.str "not an arity")
diff --git a/checker/term.mli b/checker/term.mli
index ab488b2b7..d6455e23f 100644
--- a/checker/term.mli
+++ b/checker/term.mli
@@ -40,7 +40,7 @@ val extended_rel_list : int -> rel_context -> constr list
val compose_lam : (name * constr) list -> constr -> constr
val decompose_lam : constr -> (name * constr) list * constr
val decompose_lam_n_assum : int -> constr -> rel_context * constr
-val mkProd_or_LetIn : name * constr option * constr -> constr -> constr
+val mkProd_or_LetIn : rel_declaration -> constr -> constr
val it_mkProd_or_LetIn : constr -> rel_context -> constr
val decompose_prod_assum : constr -> rel_context * constr
val decompose_prod_n_assum : int -> constr -> rel_context * constr
diff --git a/checker/typeops.ml b/checker/typeops.ml
index d49c40a8b..64afd21b2 100644
--- a/checker/typeops.ml
+++ b/checker/typeops.ml
@@ -62,7 +62,7 @@ let judge_of_type u = Sort (Type (Univ.super u))
let judge_of_relative env n =
try
- let (_,_,typ) = lookup_rel n env in
+ let LocalAssum (_,typ) | LocalDef (_,_,typ) = lookup_rel n env in
lift n typ
with Not_found ->
error_unbound_rel env n
@@ -296,13 +296,13 @@ let rec execute env cstr =
| Lambda (name,c1,c2) ->
let _ = execute_type env c1 in
- let env1 = push_rel (name,None,c1) env in
+ let env1 = push_rel (LocalAssum (name,c1)) env in
let j' = execute env1 c2 in
Prod(name,c1,j')
| Prod (name,c1,c2) ->
let varj = execute_type env c1 in
- let env1 = push_rel (name,None,c1) env in
+ let env1 = push_rel (LocalAssum (name,c1)) env in
let varj' = execute_type env1 c2 in
Sort (sort_of_product env varj varj')
@@ -314,7 +314,7 @@ let rec execute env cstr =
let env',c2' = (* refresh_arity env *) env, c2 in
let _ = execute_type env' c2' in
judge_of_cast env' (c1,j1) DEFAULTcast c2' in
- let env1 = push_rel (name,Some c1,c2) env in
+ let env1 = push_rel (LocalDef (name,c1,c2)) env in
let j' = execute env1 c3 in
subst1 c1 j'
@@ -378,10 +378,10 @@ let infer_type env constr = execute_type env constr
let check_ctxt env rels =
fold_rel_context (fun d env ->
match d with
- (_,None,ty) ->
+ | LocalAssum (_,ty) ->
let _ = infer_type env ty in
push_rel d env
- | (_,Some bd,ty) ->
+ | LocalDef (_,bd,ty) ->
let j1 = infer env bd in
let _ = infer env ty in
conv_leq env j1 ty;
@@ -399,9 +399,9 @@ let check_polymorphic_arity env params par =
let pl = par.template_param_levels in
let rec check_p env pl params =
match pl, params with
- Some u::pl, (na,None,ty)::params ->
+ Some u::pl, LocalAssum (na,ty)::params ->
check_kind env ty u;
- check_p (push_rel (na,None,ty) env) pl params
+ check_p (push_rel (LocalAssum (na,ty)) env) pl params
| None::pl,d::params -> check_p (push_rel d env) pl params
| [], _ -> ()
| _ -> failwith "check_poly: not the right number of params" in
diff --git a/checker/values.ml b/checker/values.ml
index c14e9223d..19cbb5060 100644
--- a/checker/values.ml
+++ b/checker/values.ml
@@ -13,7 +13,7 @@
To ensure this file is up-to-date, 'make' now compares the md5 of cic.mli
with a copy we maintain here:
-MD5 7c050ff1db22f14ee3a4c44aae533082 checker/cic.mli
+MD5 9f7fd499f812b6548a55f7067e6a9d06 checker/cic.mli
*)
@@ -154,8 +154,8 @@ and v_prec = Tuple ("prec_declaration",
and v_fix = Tuple ("pfixpoint", [|Tuple ("fix2",[|Array Int;Int|]);v_prec|])
and v_cofix = Tuple ("pcofixpoint",[|Int;v_prec|])
-
-let v_rdecl = v_tuple "rel_declaration" [|v_name;Opt v_constr;v_constr|]
+let v_rdecl = v_sum "rel_declaration" 0 [| [|v_name; v_constr|]; (* LocalAssum *)
+ [|v_name; v_constr; v_constr|] |] (* LocalDef *)
let v_rctxt = List v_rdecl
let v_section_ctxt = v_enum "emptylist" 1
diff --git a/dev/doc/changes.txt b/dev/doc/changes.txt
index c143afd37..0581a5f85 100644
--- a/dev/doc/changes.txt
+++ b/dev/doc/changes.txt
@@ -1,5 +1,5 @@
=========================================
-= CHANGES BETWEEN COQ V8.5 AND CQQ V8.6 =
+= CHANGES BETWEEN COQ V8.5 AND COQ V8.6 =
=========================================
- The interface of the Context module was changed.
@@ -9,9 +9,9 @@
Context.named_declaration ---> Context.Named.Declaration.t
Context.named_list_declaration ---> Context.NamedList.Declaration.t
Context.rel_declaration ---> Context.Rel.Declaration.t
- Context.map_named_declaration ---> Context.Named.Declaration.map
+ Context.map_named_declaration ---> Context.Named.Declaration.map_constr
Context.map_named_list_declaration ---> Context.NamedList.Declaration.map
- Context.map_rel_declaration ---> Context.Rel.Declaration.map
+ Context.map_rel_declaration ---> Context.Rel.Declaration.map_constr
Context.fold_named_declaration ---> Context.Named.Declaration.fold
Context.fold_rel_declaration ---> Context.Rel.Declaration.fold
Context.exists_named_declaration ---> Context.Named.Declaration.exists
@@ -37,8 +37,8 @@
Context.extended_rel_vect ---> Context.Rel.to_extended_vect
Context.fold_rel_context ---> Context.Rel.fold_outside
Context.fold_rel_context_reverse ---> Context.Rel.fold_inside
- Context.map_rel_context ---> Context.Rel.map
- Context.map_named_context ---> Context.Named.map
+ Context.map_rel_context ---> Context.Rel.map_constr
+ Context.map_named_context ---> Context.Named.map_constr
Context.iter_rel_context ---> Context.Rel.iter
Context.iter_named_context ---> Context.Named.iter
Context.empty_rel_context ---> Context.Rel.empty
@@ -48,8 +48,27 @@
Context.rel_context_nhyps ---> Context.Rel.nhyps
Context.rel_context_tags ---> Context.Rel.to_tags
+- 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
+
+
=========================================
-= CHANGES BETWEEN COQ V8.4 AND CQQ V8.5 =
+= CHANGES BETWEEN COQ V8.4 AND COQ V8.5 =
=========================================
** Refactoring : more mli interfaces and simpler grammar.cma **
diff --git a/engine/evd.ml b/engine/evd.ml
index 8f8b29d10..f1f65bd8a 100644
--- a/engine/evd.ml
+++ b/engine/evd.ml
@@ -16,6 +16,7 @@ open Vars
open Termops
open Environ
open Globnames
+open Context.Named.Declaration
(** Generic filters *)
module Filter :
@@ -230,20 +231,20 @@ let evar_instance_array test_id info args =
else instance_mismatch ()
| false :: filter, _ :: ctxt ->
instrec filter ctxt i
- | true :: filter, (id,_,_ as d) :: ctxt ->
+ | true :: filter, d :: ctxt ->
if i < len then
let c = Array.unsafe_get args i in
if test_id d c then instrec filter ctxt (succ i)
- else (id, c) :: instrec filter ctxt (succ i)
+ else (get_id d, c) :: instrec filter ctxt (succ i)
else instance_mismatch ()
| _ -> instance_mismatch ()
in
match Filter.repr (evar_filter info) with
| None ->
- let map i (id,_,_ as d) =
+ let map i d =
if (i < len) then
let c = Array.unsafe_get args i in
- if test_id d c then None else Some (id,c)
+ if test_id d c then None else Some (get_id d, c)
else instance_mismatch ()
in
List.map_filter_i map (evar_context info)
@@ -251,7 +252,7 @@ let evar_instance_array test_id info args =
instrec filter (evar_context info) 0
let make_evar_instance_array info args =
- evar_instance_array (fun (id,_,_) -> isVarId id) info args
+ evar_instance_array (isVarId % get_id) info args
let instantiate_evar_array info c args =
let inst = make_evar_instance_array info args in
@@ -660,7 +661,7 @@ let restrict evk filter ?candidates evd =
evar_extra = Store.empty } in
let evar_names = reassign_name_defined evk evk' evd.evar_names in
let ctxt = Filter.filter_list filter (evar_context evar_info) in
- let id_inst = Array.map_of_list (fun (id,_,_) -> mkVar id) ctxt in
+ let id_inst = Array.map_of_list (mkVar % get_id) ctxt in
let body = mkEvar(evk',id_inst) in
let (defn_evars, undf_evars) = define_aux evd.defn_evars evd.undf_evars evk body in
{ evd with undf_evars = EvMap.add evk' evar_info' undf_evars;
@@ -722,10 +723,10 @@ let evars_of_term c =
evrec Evar.Set.empty c
let evars_of_named_context nc =
- List.fold_right (fun (_, b, t) s ->
+ List.fold_right (fun decl s ->
Option.fold_left (fun s t ->
Evar.Set.union s (evars_of_term t))
- (Evar.Set.union s (evars_of_term t)) b)
+ (Evar.Set.union s (evars_of_term (get_type decl))) (get_value decl))
nc Evar.Set.empty
let evars_of_filtered_evar_info evi =
@@ -1228,8 +1229,9 @@ let pr_meta_map mmap =
in
prlist pr_meta_binding (metamap_to_list mmap)
-let pr_decl ((id,b,_),ok) =
- match b with
+let pr_decl (decl,ok) =
+ let id = get_id decl in
+ match get_value decl with
| None -> if ok then pr_id id else (str "{" ++ pr_id id ++ str "}")
| Some c -> str (if ok then "(" else "{") ++ pr_id id ++ str ":=" ++
print_constr c ++ str (if ok then ")" else "}")
@@ -1346,8 +1348,9 @@ let print_env_short env =
let pr_body n = function
| None -> pr_name n
| Some b -> str "(" ++ pr_name n ++ str " := " ++ print_constr b ++ str ")" in
- let pr_named_decl (n, b, _) = pr_body (Name n) b in
- let pr_rel_decl (n, b, _) = pr_body n b in
+ let pr_named_decl decl = pr_body (Name (get_id decl)) (get_value decl) in
+ let pr_rel_decl decl = let open Context.Rel.Declaration in
+ pr_body (get_name decl) (get_value decl) in
let nc = List.rev (named_context env) in
let rc = List.rev (rel_context env) in
str "[" ++ pr_sequence pr_named_decl nc ++ str "]" ++ spc () ++
diff --git a/engine/namegen.ml b/engine/namegen.ml
index fc3f0cc75..6b2b58531 100644
--- a/engine/namegen.ml
+++ b/engine/namegen.ml
@@ -22,6 +22,7 @@ open Libnames
open Globnames
open Environ
open Termops
+open Context.Rel.Declaration
(**********************************************************************)
(* Conventional names *)
@@ -113,7 +114,7 @@ let hdchar env c =
| Rel n ->
(if n<=k then "p" (* the initial term is flexible product/function *)
else
- try match Environ.lookup_rel (n-k) env with
+ try match Environ.lookup_rel (n-k) env |> to_tuple with
| (Name id,_,_) -> lowercase_first_char id
| (Anonymous,_,t) -> hdrec 0 (lift (n-k) t)
with Not_found -> "y")
@@ -142,10 +143,9 @@ let prod_name = mkProd_name
let prod_create env (a,b) = mkProd (named_hd env a Anonymous, a, b)
let lambda_create env (a,b) = mkLambda (named_hd env a Anonymous, a, b)
-let name_assumption env (na,c,t) =
- match c with
- | None -> (named_hd env t na, None, t)
- | Some body -> (named_hd env body na, c, t)
+let name_assumption env = function
+ | LocalAssum (na,t) -> LocalAssum (named_hd env t na, t)
+ | LocalDef (na,c,t) -> LocalDef (named_hd env c na, c, t)
let name_context env hyps =
snd
@@ -277,11 +277,12 @@ let next_name_away = next_name_away_with_default default_non_dependent_string
let make_all_name_different env =
let avoid = ref (ids_of_named_context (named_context env)) in
process_rel_context
- (fun (na,c,t) newenv ->
+ (fun decl newenv ->
+ let (na,_,t) = to_tuple decl in
let na = named_hd newenv t na in
let id = next_name_away na !avoid in
avoid := id::!avoid;
- push_rel (Name id,c,t) newenv)
+ push_rel (set_name (Name id) decl) newenv)
env
(* 5- Looks for next fresh name outside a list; avoids also to use names that
diff --git a/engine/termops.ml b/engine/termops.ml
index b7d89ba7b..f698f8151 100644
--- a/engine/termops.ml
+++ b/engine/termops.ml
@@ -15,6 +15,9 @@ open Term
open Vars
open Environ
+module RelDecl = Context.Rel.Declaration
+module NamedDecl = Context.Named.Declaration
+
(* Sorts and sort family *)
let print_sort = function
@@ -98,26 +101,28 @@ let print_constr_env t = !term_printer t
let print_constr t = !term_printer (Global.env()) t
let set_print_constr f = term_printer := f
-let pr_var_decl env (id,c,typ) =
- let pbody = match c with
- | None -> (mt ())
- | Some c ->
+let pr_var_decl env decl =
+ let open NamedDecl in
+ let pbody = match decl with
+ | LocalAssum _ -> mt ()
+ | LocalDef (_,c,_) ->
(* Force evaluation *)
let pb = print_constr_env env c in
(str" := " ++ pb ++ cut () ) in
- let pt = print_constr_env env typ in
+ let pt = print_constr_env env (get_type decl) in
let ptyp = (str" : " ++ pt) in
- (pr_id id ++ hov 0 (pbody ++ ptyp))
+ (pr_id (get_id decl) ++ hov 0 (pbody ++ ptyp))
-let pr_rel_decl env (na,c,typ) =
- let pbody = match c with
- | None -> mt ()
- | Some c ->
+let pr_rel_decl env decl =
+ let open RelDecl in
+ let pbody = match decl with
+ | LocalAssum _ -> mt ()
+ | LocalDef (_,c,_) ->
(* Force evaluation *)
let pb = print_constr_env env c in
(str":=" ++ spc () ++ pb ++ spc ()) in
- let ptyp = print_constr_env env typ in
- match na with
+ let ptyp = print_constr_env env (get_type decl) in
+ match get_name decl with
| Anonymous -> hov 0 (str"<>" ++ spc () ++ pbody ++ str":" ++ spc () ++ ptyp)
| Name id -> hov 0 (pr_id id ++ spc () ++ pbody ++ str":" ++ spc () ++ ptyp)
@@ -157,42 +162,53 @@ let rel_list n m =
in
reln [] 1
-let push_rel_assum (x,t) env = push_rel (x,None,t) env
+let push_rel_assum (x,t) env =
+ let open RelDecl in
+ push_rel (LocalAssum (x,t)) env
let push_rels_assum assums =
- push_rel_context (List.map (fun (x,t) -> (x,None,t)) assums)
+ let open RelDecl in
+ push_rel_context (List.map (fun (x,t) -> LocalAssum (x,t)) assums)
let push_named_rec_types (lna,typarray,_) env =
+ let open NamedDecl in
let ctxt =
Array.map2_i
(fun i na t ->
match na with
- | Name id -> (id, None, lift i t)
+ | Name id -> LocalAssum (id, lift i t)
| Anonymous -> anomaly (Pp.str "Fix declarations must be named"))
lna typarray in
Array.fold_left
(fun e assum -> push_named assum e) env ctxt
let lookup_rel_id id sign =
+ let open RelDecl in
let rec lookrec n = function
- | [] -> raise Not_found
- | (Anonymous, _, _) :: l -> lookrec (n + 1) l
- | (Name id', b, t) :: l ->
- if Names.Id.equal id' id then (n, b, t) else lookrec (n + 1) l
+ | [] ->
+ raise Not_found
+ | (LocalAssum (Anonymous, _) | LocalDef (Anonymous,_,_)) :: l ->
+ lookrec (n + 1) l
+ | LocalAssum (Name id', t) :: l ->
+ if Names.Id.equal id' id then (n,None,t) else lookrec (n + 1) l
+ | LocalDef (Name id', b, t) :: l ->
+ if Names.Id.equal id' id then (n,Some b,t) else lookrec (n + 1) l
in
lookrec 1 sign
(* Constructs either [forall x:t, c] or [let x:=b:t in c] *)
-let mkProd_or_LetIn (na,body,t) c =
- match body with
- | None -> mkProd (na, t, c)
- | Some b -> mkLetIn (na, b, t, c)
+let mkProd_or_LetIn decl c =
+ let open RelDecl in
+ match decl with
+ | LocalAssum (na,t) -> mkProd (na, t, c)
+ | LocalDef (na,b,t) -> mkLetIn (na, b, t, c)
(* Constructs either [forall x:t, c] or [c] in which [x] is replaced by [b] *)
-let mkProd_wo_LetIn (na,body,t) c =
- match body with
- | None -> mkProd (na, t, c)
- | Some b -> subst1 b c
+let mkProd_wo_LetIn decl c =
+ let open RelDecl in
+ match decl with
+ | LocalAssum (na,t) -> mkProd (na, t, c)
+ | LocalDef (_,b,_) -> subst1 b c
let it_mkProd init = List.fold_left (fun c (n,t) -> mkProd (n, t, c)) init
let it_mkLambda init = List.fold_left (fun c (n,t) -> mkLambda (n, t, c)) init
@@ -208,10 +224,11 @@ let it_mkNamedProd_wo_LetIn init = it_named_context_quantifier mkNamedProd_wo_Le
let it_mkNamedLambda_or_LetIn init = it_named_context_quantifier mkNamedLambda_or_LetIn ~init
let it_mkLambda_or_LetIn_from_no_LetIn c decls =
+ let open RelDecl in
let rec aux k decls c = match decls with
| [] -> c
- | (na,Some b,t)::decls -> mkLetIn (na,b,t,aux (k-1) decls (liftn 1 k c))
- | (na,None,t)::decls -> mkLambda (na,t,aux (k-1) decls c)
+ | LocalDef (na,b,t) :: decls -> mkLetIn (na,b,t,aux (k-1) decls (liftn 1 k c))
+ | LocalAssum (na,t) :: decls -> mkLambda (na,t,aux (k-1) decls c)
in aux (List.length decls) (List.rev decls) c
(* *)
@@ -302,7 +319,7 @@ let map_constr_with_named_binders g f l c = match kind_of_term c with
(co-)fixpoint) *)
let fold_rec_types g (lna,typarray,_) e =
- let ctxt = Array.map2_i (fun i na t -> (na, None, lift i t)) lna typarray in
+ let ctxt = Array.map2_i (fun i na t -> RelDecl.LocalAssum (na, lift i t)) lna typarray in
Array.fold_left (fun e assum -> g assum e) e ctxt
let map_left2 f a g b =
@@ -317,7 +334,9 @@ let map_left2 f a g b =
r, s
end
-let map_constr_with_binders_left_to_right g f l c = match kind_of_term c with
+let map_constr_with_binders_left_to_right g f l c =
+ let open RelDecl in
+ match kind_of_term c with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> c
| Cast (b,k,t) ->
@@ -327,18 +346,18 @@ let map_constr_with_binders_left_to_right g f l c = match kind_of_term c with
else mkCast (b',k,t')
| Prod (na,t,b) ->
let t' = f l t in
- let b' = f (g (na,None,t) l) b in
+ let b' = f (g (LocalAssum (na,t)) l) b in
if t' == t && b' == b then c
else mkProd (na, t', b')
| Lambda (na,t,b) ->
let t' = f l t in
- let b' = f (g (na,None,t) l) b in
+ let b' = f (g (LocalAssum (na,t)) l) b in
if t' == t && b' == b then c
else mkLambda (na, t', b')
| LetIn (na,bo,t,b) ->
let bo' = f l bo in
let t' = f l t in
- let b' = f (g (na,Some bo,t) l) b in
+ let b' = f (g (LocalDef (na,bo,t)) l) b in
if bo' == bo && t' == t && b' == b then c
else mkLetIn (na, bo', t', b')
| App (c,[||]) -> assert false
@@ -379,7 +398,9 @@ let map_constr_with_binders_left_to_right g f l c = match kind_of_term c with
else mkCoFix (ln,(lna,tl',bl'))
(* strong *)
-let map_constr_with_full_binders g f l cstr = match kind_of_term cstr with
+let map_constr_with_full_binders g f l cstr =
+ let open RelDecl in
+ match kind_of_term cstr with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> cstr
| Cast (c,k, t) ->
@@ -388,16 +409,16 @@ let map_constr_with_full_binders g f l cstr = match kind_of_term cstr with
if c==c' && t==t' then cstr else mkCast (c', k, t')
| Prod (na,t,c) ->
let t' = f l t in
- let c' = f (g (na,None,t) l) c in
+ let c' = f (g (LocalAssum (na,t)) l) c in
if t==t' && c==c' then cstr else mkProd (na, t', c')
| Lambda (na,t,c) ->
let t' = f l t in
- let c' = f (g (na,None,t) l) c in
+ let c' = f (g (LocalAssum (na,t)) l) c in
if t==t' && c==c' then cstr else mkLambda (na, t', c')
| LetIn (na,b,t,c) ->
let b' = f l b in
let t' = f l t in
- let c' = f (g (na,Some b,t) l) c in
+ let c' = f (g (LocalDef (na,b,t)) l) c in
if b==b' && t==t' && c==c' then cstr else mkLetIn (na, b', t', c')
| App (c,al) ->
let c' = f l c in
@@ -418,7 +439,7 @@ let map_constr_with_full_binders g f l cstr = match kind_of_term cstr with
| Fix (ln,(lna,tl,bl)) ->
let tl' = Array.map (f l) tl in
let l' =
- Array.fold_left2 (fun l na t -> g (na,None,t) l) l lna tl in
+ Array.fold_left2 (fun l na t -> g (LocalAssum (na,t)) l) l lna tl in
let bl' = Array.map (f l') bl in
if Array.for_all2 (==) tl tl' && Array.for_all2 (==) bl bl'
then cstr
@@ -426,7 +447,7 @@ let map_constr_with_full_binders g f l cstr = match kind_of_term cstr with
| CoFix(ln,(lna,tl,bl)) ->
let tl' = Array.map (f l) tl in
let l' =
- Array.fold_left2 (fun l na t -> g (na,None,t) l) l lna tl in
+ Array.fold_left2 (fun l na t -> g (LocalAssum (na,t)) l) l lna tl in
let bl' = Array.map (f l') bl in
if Array.for_all2 (==) tl tl' && Array.for_all2 (==) bl bl'
then cstr
@@ -439,23 +460,25 @@ let map_constr_with_full_binders g f l cstr = match kind_of_term cstr with
index) which is processed by [g] (which typically add 1 to [n]) at
each binder traversal; it is not recursive *)
-let fold_constr_with_full_binders g f n acc c = match kind_of_term c with
+let fold_constr_with_full_binders g f n acc c =
+ let open RelDecl in
+ match kind_of_term c with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> acc
| Cast (c,_, t) -> f n (f n acc c) t
- | Prod (na,t,c) -> f (g (na,None,t) n) (f n acc t) c
- | Lambda (na,t,c) -> f (g (na,None,t) n) (f n acc t) c
- | LetIn (na,b,t,c) -> f (g (na,Some b,t) n) (f n (f n acc b) t) c
+ | Prod (na,t,c) -> f (g (LocalAssum (na,t)) n) (f n acc t) c
+ | Lambda (na,t,c) -> f (g (LocalAssum (na,t)) n) (f n acc t) c
+ | LetIn (na,b,t,c) -> f (g (LocalDef (na,b,t)) n) (f n (f n acc b) t) c
| App (c,l) -> Array.fold_left (f n) (f n acc c) l
| Proj (p,c) -> f n acc c
| Evar (_,l) -> Array.fold_left (f n) acc l
| Case (_,p,c,bl) -> Array.fold_left (f n) (f n (f n acc p) c) bl
| Fix (_,(lna,tl,bl)) ->
- let n' = CArray.fold_left2 (fun c n t -> g (n,None,t) c) n lna tl in
+ let n' = CArray.fold_left2 (fun c n t -> g (LocalAssum (n,t)) c) n lna tl in
let fd = Array.map2 (fun t b -> (t,b)) tl bl in
Array.fold_left (fun acc (t,b) -> f n' (f n acc t) b) acc fd
| CoFix (_,(lna,tl,bl)) ->
- let n' = CArray.fold_left2 (fun c n t -> g (n,None,t) c) n lna tl in
+ let n' = CArray.fold_left2 (fun c n t -> g (LocalAssum (n,t)) c) n lna tl in
let fd = Array.map2 (fun t b -> (t,b)) tl bl in
Array.fold_left (fun acc (t,b) -> f n' (f n acc t) b) acc fd
@@ -467,23 +490,25 @@ let fold_constr_with_binders g f n acc c =
each binder traversal; it is not recursive and the order with which
subterms are processed is not specified *)
-let iter_constr_with_full_binders g f l c = match kind_of_term c with
+let iter_constr_with_full_binders g f l c =
+ let open RelDecl in
+ match kind_of_term c with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> ()
| Cast (c,_, t) -> f l c; f l t
- | Prod (na,t,c) -> f l t; f (g (na,None,t) l) c
- | Lambda (na,t,c) -> f l t; f (g (na,None,t) l) c
- | LetIn (na,b,t,c) -> f l b; f l t; f (g (na,Some b,t) l) c
+ | Prod (na,t,c) -> f l t; f (g (LocalAssum (na,t)) l) c
+ | Lambda (na,t,c) -> f l t; f (g (LocalAssum (na,t)) l) c
+ | LetIn (na,b,t,c) -> f l b; f l t; f (g (LocalDef (na,b,t)) l) c
| App (c,args) -> f l c; Array.iter (f l) args
| Proj (p,c) -> f l c
| Evar (_,args) -> Array.iter (f l) args
| Case (_,p,c,bl) -> f l p; f l c; Array.iter (f l) bl
| Fix (_,(lna,tl,bl)) ->
- let l' = Array.fold_left2 (fun l na t -> g (na,None,t) l) l lna tl in
+ let l' = Array.fold_left2 (fun l na t -> g (LocalAssum (na,t)) l) l lna tl in
Array.iter (f l) tl;
Array.iter (f l') bl
| CoFix (_,(lna,tl,bl)) ->
- let l' = Array.fold_left2 (fun l na t -> g (na,None,t) l) l lna tl in
+ let l' = Array.fold_left2 (fun l na t -> g (LocalAssum (na,t)) l) l lna tl in
Array.iter (f l) tl;
Array.iter (f l') bl
@@ -531,10 +556,11 @@ let occur_var env id c =
in
try occur_rec c; false with Occur -> true
-let occur_var_in_decl env hyp (_,c,typ) =
- match c with
- | None -> occur_var env hyp typ
- | Some body ->
+let occur_var_in_decl env hyp decl =
+ let open NamedDecl in
+ match decl with
+ | LocalAssum (_,typ) -> occur_var env hyp typ
+ | LocalDef (_, body, typ) ->
occur_var env hyp typ ||
occur_var env hyp body
@@ -593,10 +619,11 @@ let dependent_no_evar = dependent_main true false
let dependent_univs = dependent_main false true
let dependent_univs_no_evar = dependent_main true true
-let dependent_in_decl a (_,c,t) =
- match c with
- | None -> dependent a t
- | Some body -> dependent a body || dependent a t
+let dependent_in_decl a decl =
+ let open NamedDecl in
+ match decl with
+ | LocalAssum (_,t) -> dependent a t
+ | LocalDef (_, body, t) -> dependent a body || dependent a t
let count_occurrences m t =
let n = ref 0 in
@@ -699,10 +726,10 @@ let replace_term = replace_term_gen eq_constr
let vars_of_env env =
let s =
- Context.Named.fold_outside (fun (id,_,_) s -> Id.Set.add id s)
+ Context.Named.fold_outside (fun decl s -> Id.Set.add (NamedDecl.get_id decl) s)
(named_context env) ~init:Id.Set.empty in
Context.Rel.fold_outside
- (fun (na,_,_) s -> match na with Name id -> Id.Set.add id s | _ -> s)
+ (fun decl s -> match RelDecl.get_name decl with Name id -> Id.Set.add id s | _ -> s)
(rel_context env) ~init:s
let add_vname vars = function
@@ -728,11 +755,11 @@ let empty_names_context = []
let ids_of_rel_context sign =
Context.Rel.fold_outside
- (fun (na,_,_) l -> match na with Name id -> id::l | Anonymous -> l)
+ (fun decl l -> match RelDecl.get_name decl with Name id -> id::l | Anonymous -> l)
sign ~init:[]
let ids_of_named_context sign =
- Context.Named.fold_outside (fun (id,_,_) idl -> id::idl) sign ~init:[]
+ Context.Named.fold_outside (fun decl idl -> NamedDecl.get_id decl :: idl) sign ~init:[]
let ids_of_context env =
(ids_of_rel_context (rel_context env))
@@ -740,7 +767,7 @@ let ids_of_context env =
let names_of_rel_context env =
- List.map (fun (na,_,_) -> na) (rel_context env)
+ List.map RelDecl.get_name (rel_context env)
let is_section_variable id =
try let _ = Global.lookup_named id in true
@@ -813,7 +840,7 @@ let filtering env cv_pb c1 c2 =
end
| Prod (n,t1,c1), Prod (_,t2,c2) ->
aux env cv_pb t1 t2;
- aux ((n,None,t1)::env) cv_pb c1 c2
+ aux (RelDecl.LocalAssum (n,t1) :: env) cv_pb c1 c2
| _, Evar (ev,_) -> define cv_pb env ev c1
| Evar (ev,_), _ -> define cv_pb env ev c2
| _ ->
@@ -826,8 +853,8 @@ let filtering env cv_pb c1 c2 =
let decompose_prod_letin : constr -> int * Context.Rel.t * constr =
let rec prodec_rec i l c = match kind_of_term c with
- | Prod (n,t,c) -> prodec_rec (succ i) ((n,None,t)::l) c
- | LetIn (n,d,t,c) -> prodec_rec (succ i) ((n,Some d,t)::l) c
+ | Prod (n,t,c) -> prodec_rec (succ i) (RelDecl.LocalAssum (n,t)::l) c
+ | LetIn (n,d,t,c) -> prodec_rec (succ i) (RelDecl.LocalDef (n,d,t)::l) c
| Cast (c,_,_) -> prodec_rec i l c
| _ -> i,l,c in
prodec_rec 0 []
@@ -902,16 +929,16 @@ let process_rel_context f env =
let assums_of_rel_context sign =
Context.Rel.fold_outside
- (fun (na,c,t) l ->
- match c with
- Some _ -> l
- | None -> (na, t)::l)
+ (fun decl l ->
+ match decl with
+ | RelDecl.LocalDef _ -> l
+ | RelDecl.LocalAssum (na,t) -> (na, t)::l)
sign ~init:[]
let map_rel_context_in_env f env sign =
let rec aux env acc = function
| d::sign ->
- aux (push_rel d env) (Context.Rel.Declaration.map (f env) d :: acc) sign
+ aux (push_rel d env) (RelDecl.map_constr (f env) d :: acc) sign
| [] ->
acc
in
@@ -919,7 +946,7 @@ let map_rel_context_in_env f env sign =
let map_rel_context_with_binders f sign =
let rec aux k = function
- | d::sign -> Context.Rel.Declaration.map (f k) d :: aux (k-1) sign
+ | d::sign -> RelDecl.map_constr (f k) d :: aux (k-1) sign
| [] -> []
in
aux (Context.Rel.length sign) sign
@@ -933,21 +960,23 @@ let lift_rel_context n =
let smash_rel_context sign =
let rec aux acc = function
| [] -> acc
- | (_,None,_ as d) :: l -> aux (d::acc) l
- | (_,Some b,_) :: l ->
+ | (RelDecl.LocalAssum _ as d) :: l -> aux (d::acc) l
+ | RelDecl.LocalDef (_,b,_) :: l ->
(* Quadratic in the number of let but there are probably a few of them *)
aux (List.rev (substl_rel_context [b] (List.rev acc))) l
in List.rev (aux [] sign)
let fold_named_context_both_sides f l ~init = List.fold_right_and_left f l init
-let rec mem_named_context id = function
- | (id',_,_) :: _ when Id.equal id id' -> true
+let rec mem_named_context id ctxt =
+ match ctxt with
+ | decl :: _ when Id.equal id (NamedDecl.get_id decl) -> true
| _ :: sign -> mem_named_context id sign
| [] -> false
let compact_named_context_reverse sign =
- let compact l (i1,c1,t1) =
+ let compact l decl =
+ let (i1,c1,t1) = NamedDecl.to_tuple decl in
match l with
| [] -> [[i1],c1,t1]
| (l2,c2,t2)::q ->
@@ -959,16 +988,17 @@ let compact_named_context_reverse sign =
let compact_named_context sign = List.rev (compact_named_context_reverse sign)
let clear_named_body id env =
+ let open NamedDecl in
let aux _ = function
- | (id',Some c,t) when Id.equal id id' -> push_named (id,None,t)
+ | LocalDef (id',c,t) when Id.equal id id' -> push_named (LocalAssum (id,t))
| d -> push_named d in
fold_named_context aux env ~init:(reset_context env)
let global_vars env ids = Id.Set.elements (global_vars_set env ids)
let global_vars_set_of_decl env = function
- | (_,None,t) -> global_vars_set env t
- | (_,Some c,t) ->
+ | NamedDecl.LocalAssum (_,t) -> global_vars_set env t
+ | NamedDecl.LocalDef (_,c,t) ->
Id.Set.union (global_vars_set env t)
(global_vars_set env c)
@@ -976,7 +1006,8 @@ let dependency_closure env sign hyps =
if Id.Set.is_empty hyps then [] else
let (_,lh) =
Context.Named.fold_inside
- (fun (hs,hl) (x,_,_ as d) ->
+ (fun (hs,hl) d ->
+ let x = NamedDecl.get_id d in
if Id.Set.mem x hs then
(Id.Set.union (global_vars_set_of_decl env d) (Id.Set.remove x hs),
x::hl)
@@ -996,7 +1027,7 @@ let on_judgment_type f j = { j with uj_type = f j.uj_type }
let context_chop k ctx =
let rec chop_aux acc = function
| (0, l2) -> (List.rev acc, l2)
- | (n, ((_,Some _,_ as h)::t)) -> chop_aux (h::acc) (n, t)
+ | (n, (RelDecl.LocalDef _ as h)::t) -> chop_aux (h::acc) (n, t)
| (n, (h::t)) -> chop_aux (h::acc) (pred n, t)
| (_, []) -> anomaly (Pp.str "context_chop")
in chop_aux [] (k,ctx)
diff --git a/engine/termops.mli b/engine/termops.mli
index 720ed3bd6..c2a4f3323 100644
--- a/engine/termops.mli
+++ b/engine/termops.mli
@@ -101,7 +101,7 @@ val occur_evar : existential_key -> types -> bool
val occur_var : env -> Id.t -> types -> bool
val occur_var_in_decl :
env ->
- Id.t -> 'a * types option * types -> bool
+ Id.t -> Context.Named.Declaration.t -> bool
val free_rels : constr -> Int.Set.t
(** [dependent m t] tests whether [m] is a subterm of [t] *)
diff --git a/ide/ide_slave.ml b/ide/ide_slave.ml
index 20cf48d7f..12ef0e075 100644
--- a/ide/ide_slave.ml
+++ b/ide/ide_slave.ml
@@ -130,7 +130,8 @@ let annotate phrase =
(** Goal display *)
-let hyp_next_tac sigma env (id,_,ast) =
+let hyp_next_tac sigma env decl =
+ let (id,_,ast) = Context.Named.Declaration.to_tuple decl in
let id_s = Names.Id.to_string id in
let type_s = string_of_ppcmds (pr_ltype_env env sigma ast) in
[
@@ -187,8 +188,12 @@ let process_goal sigma g =
Richpp.richpp_of_pp (pr_goal_concl_style_env env sigma norm_constr)
in
let process_hyp d (env,l) =
- let d = Context.NamedList.Declaration.map (Reductionops.nf_evar sigma) d in
- let d' = List.map (fun x -> (x, pi2 d, pi3 d)) (pi1 d) in
+ let d = Context.NamedList.Declaration.map_constr (Reductionops.nf_evar sigma) d in
+ let d' = List.map (fun name -> let open Context.Named.Declaration in
+ match pi2 d with
+ | None -> LocalAssum (name, pi3 d)
+ | Some value -> LocalDef (name, value, pi3 d))
+ (pi1 d) in
(List.fold_right Environ.push_named d' env,
(Richpp.richpp_of_pp (pr_var_list_decl env sigma d)) :: l) in
let (_env, hyps) =
diff --git a/interp/constrextern.ml b/interp/constrextern.ml
index 70a35c613..3a8c506cb 100644
--- a/interp/constrextern.ml
+++ b/interp/constrextern.ml
@@ -29,6 +29,8 @@ open Notation
open Detyping
open Misctypes
open Decl_kinds
+
+module NamedDecl = Context.Named.Declaration
(*i*)
(* Translation from glob_constr to front constr *)
@@ -980,7 +982,7 @@ let rec glob_of_pat env sigma = function
| PRef ref -> GRef (loc,ref,None)
| PVar id -> GVar (loc,id)
| PEvar (evk,l) ->
- let test (id,_,_) = function PVar id' -> Id.equal id id' | _ -> false in
+ let test decl = function PVar id' -> Id.equal (NamedDecl.get_id decl) id' | _ -> false in
let l = Evd.evar_instance_array test (Evd.find sigma evk) l in
let id = Evd.evar_ident evk sigma in
GEvar (loc,id,List.map (on_snd (glob_of_pat env sigma)) l)
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index 964ed0514..70802d5cb 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -29,6 +29,7 @@ open Nametab
open Notation
open Inductiveops
open Decl_kinds
+open Context.Rel.Declaration
(** constr_expr -> glob_constr translation:
- it adds holes for implicit arguments
@@ -1645,14 +1646,14 @@ let internalize globalenv env allow_patvar lvar c =
|loc,(Name y as x) -> (y,PatVar(loc,x)) :: l in
match case_rel_ctxt,arg_pats with
(* LetIn in the rel_context *)
- |(_,Some _,_)::t, l when not with_letin ->
+ | LocalDef _ :: t, l when not with_letin ->
canonize_args t l forbidden_names match_acc ((Loc.ghost,Anonymous)::var_acc)
|[],[] ->
(add_name match_acc na, var_acc)
|_::t,PatVar (loc,x)::tt ->
canonize_args t tt forbidden_names
(add_name match_acc (loc,x)) ((loc,x)::var_acc)
- |(cano_name,_,ty)::t,c::tt ->
+ | (LocalAssum (cano_name,ty) | LocalDef (cano_name,_,ty)) :: t, c::tt ->
let fresh =
Namegen.next_name_away_with_default_using_types "iV" cano_name forbidden_names ty in
canonize_args t tt (fresh::forbidden_names)
@@ -1894,7 +1895,7 @@ let interp_rawcontext_evars env evdref k bl =
let t' = locate_if_hole (loc_of_glob_constr t) na t in
let t =
understand_tcc_evars env evdref ~expected_type:IsType t' in
- let d = (na,None,t) in
+ let d = LocalAssum (na,t) in
let impls =
if k == Implicit then
let na = match na with Name n -> Some n | Anonymous -> None in
@@ -1904,7 +1905,7 @@ let interp_rawcontext_evars env evdref k bl =
(push_rel d env, d::params, succ n, impls)
| Some b ->
let c = understand_judgment_tcc env evdref b in
- let d = (na, Some c.uj_val, c.uj_type) in
+ let d = LocalDef (na, c.uj_val, c.uj_type) in
(push_rel d env, d::params, n, impls))
(env,[],k+1,[]) (List.rev bl)
in (env, par), impls
diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml
index 391c600ed..751b03a4a 100644
--- a/interp/implicit_quantifiers.ml
+++ b/interp/implicit_quantifiers.ml
@@ -20,6 +20,7 @@ open Pp
open Libobject
open Nameops
open Misctypes
+open Context.Rel.Declaration
(*i*)
let generalizable_table = Summary.ref Id.Pred.empty ~name:"generalizable-ident"
@@ -196,7 +197,7 @@ let combine_params avoid fn applied needed =
List.partition
(function
(t, Some (loc, ExplByName id)) ->
- let is_id (_, (na, _, _)) = match na with
+ let is_id (_, decl) = match get_name decl with
| Name id' -> Id.equal id id'
| Anonymous -> false
in
@@ -209,22 +210,22 @@ let combine_params avoid fn applied needed =
(fun x -> match x with (t, Some (loc, ExplByName id)) -> id, t | _ -> assert false)
named
in
- let is_unset (_, (_, b, _)) = match b with
- | None -> true
- | Some _ -> false
+ let is_unset (_, decl) = match decl with
+ | LocalAssum _ -> true
+ | LocalDef _ -> false
in
let needed = List.filter is_unset needed in
let rec aux ids avoid app need =
match app, need with
[], [] -> List.rev ids, avoid
- | app, (_, (Name id, _, _)) :: need when Id.List.mem_assoc id named ->
+ | app, (_, (LocalAssum (Name id, _) | LocalDef (Name id, _, _))) :: need when Id.List.mem_assoc id named ->
aux (Id.List.assoc id named :: ids) avoid app need
- | (x, None) :: app, (None, (Name id, _, _)) :: need ->
+ | (x, None) :: app, (None, (LocalAssum (Name id, _) | LocalDef (Name id, _, _))) :: need ->
aux (x :: ids) avoid app need
- | _, (Some cl, (_, _, _) as d) :: need ->
+ | _, (Some cl, _ as d) :: need ->
let t', avoid' = fn avoid d in
aux (t' :: ids) avoid' app need
@@ -239,8 +240,8 @@ let combine_params avoid fn applied needed =
in aux [] avoid applied needed
let combine_params_freevar =
- fun avoid (_, (na, _, _)) ->
- let id' = next_name_away_from na avoid in
+ fun avoid (_, decl) ->
+ let id' = next_name_away_from (get_name decl) avoid in
(CRef (Ident (Loc.ghost, id'),None), Id.Set.add id' avoid)
let destClassApp cl =
diff --git a/interp/implicit_quantifiers.mli b/interp/implicit_quantifiers.mli
index b226bfa0a..d0327e506 100644
--- a/interp/implicit_quantifiers.mli
+++ b/interp/implicit_quantifiers.mli
@@ -38,10 +38,10 @@ val make_fresh : Id.Set.t -> Environ.env -> Id.t -> Id.t
val implicits_of_glob_constr : ?with_products:bool -> Glob_term.glob_constr -> Impargs.manual_implicits
val combine_params_freevar :
- Id.Set.t -> (global_reference * bool) option * (Name.t * Term.constr option * Term.types) ->
+ Id.Set.t -> (global_reference * bool) option * Context.Rel.Declaration.t ->
Constrexpr.constr_expr * Id.Set.t
val implicit_application : Id.Set.t -> ?allow_partial:bool ->
- (Id.Set.t -> (global_reference * bool) option * (Name.t * Term.constr option * Term.types) ->
+ (Id.Set.t -> (global_reference * bool) option * Context.Rel.Declaration.t ->
Constrexpr.constr_expr * Id.Set.t) ->
constr_expr -> constr_expr * Id.Set.t
diff --git a/kernel/closure.ml b/kernel/closure.ml
index 9bc67b5ad..dc98cc65d 100644
--- a/kernel/closure.ml
+++ b/kernel/closure.ml
@@ -247,8 +247,8 @@ let info_env info = info.i_cache.i_env
let rec assoc_defined id = function
| [] -> raise Not_found
-| (_, None, _) :: ctxt -> assoc_defined id ctxt
-| (id', Some c, _) :: ctxt ->
+| Context.Named.Declaration.LocalAssum _ :: ctxt -> assoc_defined id ctxt
+| Context.Named.Declaration.LocalDef (id', c, _) :: ctxt ->
if Id.equal id id' then c else assoc_defined id ctxt
let ref_value_cache ({i_cache = cache} as infos) ref =
@@ -285,9 +285,9 @@ let defined_rels flags env =
let ctx = rel_context env in
let len = List.length ctx in
let ans = Array.make len None in
- let iter i (_, b, _) = match b with
- | None -> ()
- | Some _ -> Array.unsafe_set ans i b
+ let iter i = function
+ | Context.Rel.Declaration.LocalAssum _ -> ()
+ | Context.Rel.Declaration.LocalDef (_,b,_) -> Array.unsafe_set ans i (Some b)
in
let () = List.iteri iter ctx in
ans
diff --git a/kernel/context.ml b/kernel/context.ml
index 3be1e8323..cc1e6f176 100644
--- a/kernel/context.ml
+++ b/kernel/context.ml
@@ -33,33 +33,122 @@ open Names
Individual declarations are then designated by de Bruijn indexes. *)
module Rel =
struct
- (** Representation of {e local declarations}.
-
- [(name, None, typ)] represents a {e local assumption}.
- In the Reference Manual we denote them as [(name:typ)].
-
- [(name, Some value, typ)] represents a {e local definition}.
- In the Reference Manual we denote them as [(name := value : typ)].
- *)
+ (** Representation of {e local declarations}. *)
module Declaration =
struct
- type t = Name.t * Constr.t option * Constr.t
-
- (** Map all terms in a given declaration. *)
- let map f (n, v, ty) = (n, Option.map f v, f ty)
-
- (** Reduce all terms in a given declaration to a single value. *)
- let fold f (_, v, ty) a = f ty (Option.fold_right f v a)
+ (* local declaration *)
+ type t = LocalAssum of Name.t * Constr.t (* local assumption *)
+ | LocalDef of Name.t * Constr.t * Constr.t (* local definition *)
+
+ (** Return the name bound by a given declaration. *)
+ let get_name = function
+ | LocalAssum (na,_)
+ | LocalDef (na,_,_) -> na
+
+ (** Return [Some value] for local-declarations and [None] for local-assumptions. *)
+ let get_value = function
+ | LocalAssum _ -> None
+ | LocalDef (_,v,_) -> Some v
+
+ (** Return the type of the name bound by a given declaration. *)
+ let get_type = function
+ | LocalAssum (_,ty)
+ | LocalDef (_,_,ty) -> ty
+
+ (** Set the name that is bound by a given declaration. *)
+ let set_name na = function
+ | LocalAssum (_,ty) -> LocalAssum (na, ty)
+ | LocalDef (_,v,ty) -> LocalDef (na, v, ty)
+
+ (** Set the type of the bound variable in a given declaration. *)
+ let set_type ty = function
+ | LocalAssum (na,_) -> LocalAssum (na, ty)
+ | LocalDef (na,v,_) -> LocalDef (na, v, ty)
+
+ (** Return [true] iff a given declaration is a local assumption. *)
+ let is_local_assum = function
+ | LocalAssum _ -> true
+ | LocalDef _ -> false
+
+ (** Return [true] iff a given declaration is a local definition. *)
+ let is_local_def = function
+ | LocalAssum _ -> false
+ | LocalDef _ -> true
(** Check whether any term in a given declaration satisfies a given predicate. *)
- let exists f (_, v, ty) = Option.cata f false v || f ty
+ let exists f = function
+ | LocalAssum (_, ty) -> f ty
+ | LocalDef (_, v, ty) -> f v || f ty
(** Check whether all terms in a given declaration satisfy a given predicate. *)
- let for_all f (_, v, ty) = Option.cata f true v && f ty
+ let for_all f = function
+ | LocalAssum (_, ty) -> f ty
+ | LocalDef (_, v, ty) -> f v && f ty
(** Check whether the two given declarations are equal. *)
- let equal (n1, v1, ty1) (n2, v2, ty2) =
- Name.equal n1 n2 && Option.equal Constr.equal v1 v2 && Constr.equal ty1 ty2
+ let equal decl1 decl2 =
+ match decl1, decl2 with
+ | LocalAssum (n1,ty1), LocalAssum (n2, ty2) ->
+ Name.equal n1 n2 && Constr.equal ty1 ty2
+ | LocalDef (n1,v1,ty1), LocalDef (n2,v2,ty2) ->
+ Name.equal n1 n2 && Constr.equal v1 v2 && Constr.equal ty1 ty2
+ | _ ->
+ false
+
+ (** Map the name bound by a given declaration. *)
+ let map_name f = function
+ | LocalAssum (na, ty) as decl ->
+ let na' = f na in
+ if na == na' then decl else LocalAssum (na', ty)
+ | LocalDef (na, v, ty) as decl ->
+ let na' = f na in
+ if na == na' then decl else LocalDef (na', v, ty)
+
+ (** For local assumptions, this function returns the original local assumptions.
+ For local definitions, this function maps the value in the local definition. *)
+ let map_value f = function
+ | LocalAssum _ as decl -> decl
+ | LocalDef (na, v, t) as decl ->
+ let v' = f v in
+ if v == v' then decl else LocalDef (na, v', t)
+
+ (** Map the type of the name bound by a given declaration. *)
+ let map_type f = function
+ | LocalAssum (na, ty) as decl ->
+ let ty' = f ty in
+ if ty == ty' then decl else LocalAssum (na, ty')
+ | LocalDef (na, v, ty) as decl ->
+ let ty' = f ty in
+ if ty == ty' then decl else LocalDef (na, v, ty')
+
+ (** Map all terms in a given declaration. *)
+ let map_constr f = function
+ | LocalAssum (na, ty) as decl ->
+ let ty' = f ty in
+ if ty == ty' then decl else LocalAssum (na, ty')
+ | LocalDef (na, v, ty) as decl ->
+ let v' = f v in
+ let ty' = f ty in
+ if v == v' && ty == ty' then decl else LocalDef (na, v', ty')
+
+ (** Perform a given action on all terms in a given declaration. *)
+ let iter_constr f = function
+ | LocalAssum (_,ty) -> f ty
+ | LocalDef (_,v,ty) -> f v; f ty
+
+ (** Reduce all terms in a given declaration to a single value. *)
+ let fold f decl acc =
+ match decl with
+ | LocalAssum (n,ty) -> f ty acc
+ | LocalDef (n,v,ty) -> f ty (f v acc)
+
+ let to_tuple = function
+ | LocalAssum (na, ty) -> na, None, ty
+ | LocalDef (na, v, ty) -> na, Some v, ty
+
+ let of_tuple = function
+ | n, None, ty -> LocalAssum (n,ty)
+ | n, Some v, ty -> LocalDef (n,v,ty)
end
(** Rel-context is represented as a list of declarations.
@@ -73,6 +162,21 @@ module Rel =
(** Return a new rel-context enriched by with a given inner-most declaration. *)
let add d ctx = d :: ctx
+ (** Return the number of {e local declarations} in a given context. *)
+ let length = List.length
+
+ (** [extended_rel_list n Γ] builds an instance [args] such that [Γ,Δ ⊢ args:Γ]
+ with n = |Δ| and with the local definitions of [Γ] skipped in
+ [args]. Example: for [x:T,y:=c,z:U] and [n]=2, it gives [Rel 5, Rel 3]. *)
+ let nhyps =
+ let open Declaration in
+ let rec nhyps acc = function
+ | [] -> acc
+ | LocalAssum _ :: hyps -> nhyps (succ acc) hyps
+ | LocalDef _ :: hyps -> nhyps acc hyps
+ in
+ nhyps 0
+
(** Return a declaration designated by a given de Bruijn index.
@raise Not_found if the designated de Bruijn index is not present in the designated rel-context. *)
let rec lookup n ctx =
@@ -81,15 +185,14 @@ module Rel =
| n, _ :: sign -> lookup (n-1) sign
| _, [] -> raise Not_found
+ (** Check whether given two rel-contexts are equal. *)
+ let equal = List.equal Declaration.equal
+
(** Map all terms in a given rel-context. *)
- let map f =
- let map_decl (n, body_o, typ as decl) =
- let body_o' = Option.smartmap f body_o in
- let typ' = f typ in
- if body_o' == body_o && typ' == typ then decl else
- (n, body_o', typ')
- in
- List.smartmap map_decl
+ let map f = List.smartmap (Declaration.map_constr f)
+
+ (** Perform a given action on every declaration in a given rel-context. *)
+ let iter f = List.iter (Declaration.iter_constr f)
(** Reduce all terms in a given rel-context to a single value.
Innermost declarations are processed first. *)
@@ -99,29 +202,13 @@ module Rel =
Outermost declarations are processed first. *)
let fold_outside f l ~init = List.fold_right f l init
- (** Perform a given action on every declaration in a given rel-context. *)
- let iter f = List.iter (fun (_,b,t) -> f t; Option.iter f b)
-
- (** Return the number of {e local declarations} in a given context. *)
- let length = List.length
-
- (** [extended_rel_list n Γ] builds an instance [args] such that [Γ,Δ ⊢ args:Γ]
- with n = |Δ| and with the local definitions of [Γ] skipped in
- [args]. Example: for [x:T,y:=c,z:U] and [n]=2, it gives [Rel 5, Rel 3]. *)
- let nhyps =
- let rec nhyps acc = function
- | [] -> acc
- | (_,None,_)::hyps -> nhyps (1+acc) hyps
- | (_,Some _,_)::hyps -> nhyps acc hyps in
- nhyps 0
-
(** Map a given rel-context to a list where each {e local assumption} is mapped to [true]
and each {e local definition} is mapped to [false]. *)
let to_tags =
let rec aux l = function
| [] -> l
- | (_,Some _,_)::ctx -> aux (true::l) ctx
- | (_,None,_)::ctx -> aux (false::l) ctx
+ | Declaration.LocalDef _ :: ctx -> aux (true::l) ctx
+ | Declaration.LocalAssum _ :: ctx -> aux (false::l) ctx
in aux []
(** [extended_list n Γ] builds an instance [args] such that [Γ,Δ ⊢ args:Γ]
@@ -129,8 +216,8 @@ module Rel =
[args]. Example: for [x:T, y:=c, z:U] and [n]=2, it gives [Rel 5, Rel 3]. *)
let to_extended_list n =
let rec reln l p = function
- | (_, None, _) :: hyps -> reln (Constr.mkRel (n+p) :: l) (p+1) hyps
- | (_, Some _, _) :: hyps -> reln l (p+1) hyps
+ | Declaration.LocalAssum _ :: hyps -> reln (Constr.mkRel (n+p) :: l) (p+1) hyps
+ | Declaration.LocalDef _ :: hyps -> reln l (p+1) hyps
| [] -> l
in
reln [] 1
@@ -143,38 +230,127 @@ module Rel =
Individual declarations are then designated by the identifiers they bind. *)
module Named =
struct
- (** Representation of {e local declarations}.
-
- [(id, None, typ)] represents a {e local assumption}.
- In the Reference Manual we denote them as [(name:typ)].
-
- [(id, Some value, typ)] represents a {e local definition}.
- In the Reference Manual we denote them as [(name := value : typ)].
- *)
+ (** Representation of {e local declarations}. *)
module Declaration =
struct
- (** Named-context is represented as a list of declarations.
- Inner-most declarations are at the beginning of the list.
- Outer-most declarations are at the end of the list. *)
- type t = Id.t * Constr.t option * Constr.t
-
- (** Map all terms in a given declaration. *)
- let map = Rel.Declaration.map
-
- (** Reduce all terms in a given declaration to a single value. *)
- let fold f (_, v, ty) a = f ty (Option.fold_right f v a)
+ (** local declaration *)
+ type t = LocalAssum of Id.t * Constr.t
+ | LocalDef of Id.t * Constr.t * Constr.t
+
+ (** Return the identifier bound by a given declaration. *)
+ let get_id = function
+ | LocalAssum (id,_) -> id
+ | LocalDef (id,_,_) -> id
+
+ (** Return [Some value] for local-declarations and [None] for local-assumptions. *)
+ let get_value = function
+ | LocalAssum _ -> None
+ | LocalDef (_,v,_) -> Some v
+
+ (** Return the type of the name bound by a given declaration. *)
+ let get_type = function
+ | LocalAssum (_,ty)
+ | LocalDef (_,_,ty) -> ty
+
+ (** Set the identifier that is bound by a given declaration. *)
+ let set_id id = function
+ | LocalAssum (_,ty) -> LocalAssum (id, ty)
+ | LocalDef (_, v, ty) -> LocalDef (id, v, ty)
+
+ (** Set the type of the bound variable in a given declaration. *)
+ let set_type ty = function
+ | LocalAssum (id,_) -> LocalAssum (id, ty)
+ | LocalDef (id,v,_) -> LocalDef (id, v, ty)
+
+ (** Return [true] iff a given declaration is a local assumption. *)
+ let is_local_assum = function
+ | LocalAssum _ -> true
+ | LocalDef _ -> false
+
+ (** Return [true] iff a given declaration is a local definition. *)
+ let is_local_def = function
+ | LocalDef _ -> true
+ | LocalAssum _ -> false
(** Check whether any term in a given declaration satisfies a given predicate. *)
- let exists f (_, v, ty) = Option.cata f false v || f ty
+ let exists f = function
+ | LocalAssum (_, ty) -> f ty
+ | LocalDef (_, v, ty) -> f v || f ty
(** Check whether all terms in a given declaration satisfy a given predicate. *)
- let for_all f (_, v, ty) = Option.cata f true v && f ty
+ let for_all f = function
+ | LocalAssum (_, ty) -> f ty
+ | LocalDef (_, v, ty) -> f v && f ty
(** Check whether the two given declarations are equal. *)
- let equal (i1, v1, ty1) (i2, v2, ty2) =
- Id.equal i1 i2 && Option.equal Constr.equal v1 v2 && Constr.equal ty1 ty2
+ let equal decl1 decl2 =
+ match decl1, decl2 with
+ | LocalAssum (id1, ty1), LocalAssum (id2, ty2) ->
+ Id.equal id1 id2 && Constr.equal ty1 ty2
+ | LocalDef (id1, v1, ty1), LocalDef (id2, v2, ty2) ->
+ Id.equal id1 id2 && Constr.equal v1 v2 && Constr.equal ty1 ty2
+ | _ ->
+ false
+
+ (** Map the identifier bound by a given declaration. *)
+ let map_id f = function
+ | LocalAssum (id, ty) as decl ->
+ let id' = f id in
+ if id == id' then decl else LocalAssum (id', ty)
+ | LocalDef (id, v, ty) as decl ->
+ let id' = f id in
+ if id == id' then decl else LocalDef (id', v, ty)
+
+ (** For local assumptions, this function returns the original local assumptions.
+ For local definitions, this function maps the value in the local definition. *)
+ let map_value f = function
+ | LocalAssum _ as decl -> decl
+ | LocalDef (na, v, t) as decl ->
+ let v' = f v in
+ if v == v' then decl else LocalDef (na, v', t)
+
+ (** Map the type of the name bound by a given declaration. *)
+ let map_type f = function
+ | LocalAssum (id, ty) as decl ->
+ let ty' = f ty in
+ if ty == ty' then decl else LocalAssum (id, ty')
+ | LocalDef (id, v, ty) as decl ->
+ let ty' = f ty in
+ if ty == ty' then decl else LocalDef (id, v, ty')
+
+ (** Map all terms in a given declaration. *)
+ let map_constr f = function
+ | LocalAssum (id, ty) as decl ->
+ let ty' = f ty in
+ if ty == ty' then decl else LocalAssum (id, ty')
+ | LocalDef (id, v, ty) as decl ->
+ let v' = f v in
+ let ty' = f ty in
+ if v == v' && ty == ty' then decl else LocalDef (id, v', ty')
+
+ (** Perform a given action on all terms in a given declaration. *)
+ let iter_constr f = function
+ | LocalAssum (_, ty) -> f ty
+ | LocalDef (_, v, ty) -> f v; f ty
+
+ (** Reduce all terms in a given declaration to a single value. *)
+ let fold f decl a =
+ match decl with
+ | LocalAssum (_, ty) -> f ty a
+ | LocalDef (_, v, ty) -> a |> f v |> f ty
+
+ let to_tuple = function
+ | LocalAssum (id, ty) -> id, None, ty
+ | LocalDef (id, v, ty) -> id, Some v, ty
+
+ let of_tuple = function
+ | id, None, ty -> LocalAssum (id, ty)
+ | id, Some v, ty -> LocalDef (id, v, ty)
end
+ (** Named-context is represented as a list of declarations.
+ Inner-most declarations are at the beginning of the list.
+ Outer-most declarations are at the end of the list. *)
type t = Declaration.t list
(** empty named-context *)
@@ -183,22 +359,23 @@ module Named =
(** empty named-context *)
let add d ctx = d :: ctx
+ (** Return the number of {e local declarations} in a given named-context. *)
+ let length = List.length
+
(** Return a declaration designated by a given de Bruijn index.
- @raise Not_found if the designated identifier is not present in the designated named-context. *)
- let rec lookup id = function
- | (id',_,_ as decl) :: _ when Id.equal id id' -> decl
- | _ :: sign -> lookup id sign
- | [] -> raise Not_found
+ @raise Not_found if the designated identifier is not present in the designated named-context. *) let rec lookup id = function
+ | decl :: _ when Id.equal id (Declaration.get_id decl) -> decl
+ | _ :: sign -> lookup id sign
+ | [] -> raise Not_found
+
+ (** Check whether given two named-contexts are equal. *)
+ let equal = List.equal Declaration.equal
(** Map all terms in a given named-context. *)
- let map f =
- let map_decl (n, body_o, typ as decl) =
- let body_o' = Option.smartmap f body_o in
- let typ' = f typ in
- if body_o' == body_o && typ' == typ then decl else
- (n, body_o', typ')
- in
- List.smartmap map_decl
+ let map f = List.smartmap (Declaration.map_constr f)
+
+ (** Perform a given action on every declaration in a given named-context. *)
+ let iter f = List.iter (Declaration.iter_constr f)
(** Reduce all terms in a given named-context to a single value.
Innermost declarations are processed first. *)
@@ -208,18 +385,9 @@ module Named =
Outermost declarations are processed first. *)
let fold_outside f l ~init = List.fold_right f l init
- (** Perform a given action on every declaration in a given named-context. *)
- let iter f = List.iter (fun (_,b,t) -> f t; Option.iter f b)
-
- (** Return the number of {e local declarations} in a given named-context. *)
- let length = List.length
-
- (** Check whether given two named-contexts are equal. *)
- let equal = List.equal Declaration.equal
-
(** Return the set of all identifiers bound in a given named-context. *)
let to_vars =
- List.fold_left (fun accu (id, _, _) -> Id.Set.add id accu) Id.Set.empty
+ List.fold_left (fun accu decl -> Id.Set.add (Declaration.get_id decl) accu) Id.Set.empty
(** [instance_from_named_context Ω] builds an instance [args] such
that [Ω ⊢ args:Ω] where [Ω] is a named context and with the local
@@ -227,8 +395,8 @@ module Named =
gives [Var id1, Var id3]. All [idj] are supposed distinct. *)
let to_instance =
let filter = function
- | (id, None, _) -> Some (Constr.mkVar id)
- | (_, Some _, _) -> None
+ | Declaration.LocalAssum (id, _) -> Some (Constr.mkVar id)
+ | _ -> None
in
List.map_filter filter
end
@@ -238,9 +406,15 @@ module NamedList =
module Declaration =
struct
type t = Id.t list * Constr.t option * Constr.t
- let map = Named.Declaration.map
+
+ let map_constr f (ids, copt, ty as decl) =
+ let copt' = Option.map f copt in
+ let ty' = f ty in
+ if copt == copt' && ty == ty' then decl else (ids, copt', ty')
end
+
type t = Declaration.t list
+
let fold f l ~init = List.fold_right f l init
end
diff --git a/kernel/context.mli b/kernel/context.mli
index 1976e46d3..a69754cc2 100644
--- a/kernel/context.mli
+++ b/kernel/context.mli
@@ -26,21 +26,32 @@ open Names
Individual declarations are then designated by de Bruijn indexes. *)
module Rel :
sig
- (** Representation of {e local declarations}.
-
- [(name, None, typ)] represents a {e local assumption}.
-
- [(name, Some value, typ)] represents a {e local definition}.
- *)
module Declaration :
sig
- type t = Name.t * Constr.t option * Constr.t
+ (* local declaration *)
+ type t = LocalAssum of Name.t * Constr.t (* local assumption *)
+ | LocalDef of Name.t * Constr.t * Constr.t (* local definition *)
- (** Map all terms in a given declaration. *)
- val map : (Constr.t -> Constr.t) -> t -> t
+ (** Return the name bound by a given declaration. *)
+ val get_name : t -> Name.t
- (** Reduce all terms in a given declaration to a single value. *)
- val fold : (Constr.t -> 'a -> 'a) -> t -> 'a -> 'a
+ (** Return [Some value] for local-declarations and [None] for local-assumptions. *)
+ val get_value : t -> Constr.t option
+
+ (** Return the type of the name bound by a given declaration. *)
+ val get_type : t -> Constr.t
+
+ (** Set the name that is bound by a given declaration. *)
+ val set_name : Name.t -> t -> t
+
+ (** Set the type of the bound variable in a given declaration. *)
+ val set_type : Constr.t -> t -> t
+
+ (** Return [true] iff a given declaration is a local assumption. *)
+ val is_local_assum : t -> bool
+
+ (** Return [true] iff a given declaration is a local definition. *)
+ val is_local_def : t -> bool
(** Check whether any term in a given declaration satisfies a given predicate. *)
val exists : (Constr.t -> bool) -> t -> bool
@@ -50,6 +61,28 @@ sig
(** Check whether the two given declarations are equal. *)
val equal : t -> t -> bool
+
+ (** Map the name bound by a given declaration. *)
+ val map_name : (Name.t -> Name.t) -> t -> t
+
+ (** For local assumptions, this function returns the original local assumptions.
+ For local definitions, this function maps the value in the local definition. *)
+ val map_value : (Constr.t -> Constr.t) -> t -> t
+
+ (** Map the type of the name bound by a given declaration. *)
+ val map_type : (Constr.t -> Constr.t) -> t -> t
+
+ (** Map all terms in a given declaration. *)
+ val map_constr : (Constr.t -> Constr.t) -> t -> t
+
+ (** Perform a given action on all terms in a given declaration. *)
+ val iter_constr : (Constr.t -> unit) -> t -> unit
+
+ (** Reduce all terms in a given declaration to a single value. *)
+ val fold : (Constr.t -> 'a -> 'a) -> t -> 'a -> 'a
+
+ val to_tuple : t -> Name.t * Constr.t option * Constr.t
+ val of_tuple : Name.t * Constr.t option * Constr.t -> t
end
(** Rel-context is represented as a list of declarations.
@@ -63,6 +96,15 @@ sig
(** Return a new rel-context enriched by with a given inner-most declaration. *)
val add : Declaration.t -> t -> t
+ (** Return the number of {e local declarations} in a given context. *)
+ val length : t -> int
+
+ (** Check whether given two rel-contexts are equal. *)
+ val equal : t -> t -> bool
+
+ (** Return the number of {e local assumptions} in a given rel-context. *)
+ val nhyps : t -> int
+
(** Return a declaration designated by a given de Bruijn index.
@raise Not_found if the designated de Bruijn index outside the range. *)
val lookup : int -> t -> Declaration.t
@@ -70,6 +112,9 @@ sig
(** Map all terms in a given rel-context. *)
val map : (Constr.t -> Constr.t) -> t -> t
+ (** Perform a given action on every declaration in a given rel-context. *)
+ val iter : (Constr.t -> unit) -> t -> unit
+
(** Reduce all terms in a given rel-context to a single value.
Innermost declarations are processed first. *)
val fold_inside : ('a -> Declaration.t -> 'a) -> init:'a -> t -> 'a
@@ -78,15 +123,6 @@ sig
Outermost declarations are processed first. *)
val fold_outside : (Declaration.t -> 'a -> 'a) -> t -> init:'a -> 'a
- (** Perform a given action on every declaration in a given rel-context. *)
- val iter : (Constr.t -> unit) -> t -> unit
-
- (** Return the number of {e local declarations} in a given context. *)
- val length : t -> int
-
- (** Return the number of {e local assumptions} in a given rel-context. *)
- val nhyps : t -> int
-
(** Map a given rel-context to a list where each {e local assumption} is mapped to [true]
and each {e local definition} is mapped to [false]. *)
val to_tags : t -> bool list
@@ -104,21 +140,32 @@ end
Individual declarations are then designated by the identifiers they bind. *)
module Named :
sig
- (** Representation of {e local declarations}.
-
- [(id, None, typ)] represents a {e local assumption}.
-
- [(id, Some value, typ)] represents a {e local definition}.
- *)
+ (** Representation of {e local declarations}. *)
module Declaration :
sig
- type t = Id.t * Constr.t option * Constr.t
+ type t = LocalAssum of Id.t * Constr.t
+ | LocalDef of Id.t * Constr.t * Constr.t
- (** Map all terms in a given declaration. *)
- val map : (Constr.t -> Constr.t) -> t -> t
+ (** Return the identifier bound by a given declaration. *)
+ val get_id : t -> Id.t
- (** Reduce all terms in a given declaration to a single value. *)
- val fold : (Constr.t -> 'a -> 'a) -> t -> 'a -> 'a
+ (** Return [Some value] for local-declarations and [None] for local-assumptions. *)
+ val get_value : t -> Constr.t option
+
+ (** Return the type of the name bound by a given declaration. *)
+ val get_type : t -> Constr.t
+
+ (** Set the identifier that is bound by a given declaration. *)
+ val set_id : Id.t -> t -> t
+
+ (** Set the type of the bound variable in a given declaration. *)
+ val set_type : Constr.t -> t -> t
+
+ (** Return [true] iff a given declaration is a local assumption. *)
+ val is_local_assum : t -> bool
+
+ (** Return [true] iff a given declaration is a local definition. *)
+ val is_local_def : t -> bool
(** Check whether any term in a given declaration satisfies a given predicate. *)
val exists : (Constr.t -> bool) -> t -> bool
@@ -128,6 +175,28 @@ sig
(** Check whether the two given declarations are equal. *)
val equal : t -> t -> bool
+
+ (** Map the identifier bound by a given declaration. *)
+ val map_id : (Id.t -> Id.t) -> t -> t
+
+ (** For local assumptions, this function returns the original local assumptions.
+ For local definitions, this function maps the value in the local definition. *)
+ val map_value : (Constr.t -> Constr.t) -> t -> t
+
+ (** Map the type of the name bound by a given declaration. *)
+ val map_type : (Constr.t -> Constr.t) -> t -> t
+
+ (** Map all terms in a given declaration. *)
+ val map_constr : (Constr.t -> Constr.t) -> t -> t
+
+ (** Perform a given action on all terms in a given declaration. *)
+ val iter_constr : (Constr.t -> unit) -> t -> unit
+
+ (** Reduce all terms in a given declaration to a single value. *)
+ val fold : (Constr.t -> 'a -> 'a) -> t -> 'a -> 'a
+
+ val to_tuple : t -> Id.t * Constr.t option * Constr.t
+ val of_tuple : Id.t * Constr.t option * Constr.t -> t
end
(** Rel-context is represented as a list of declarations.
@@ -141,13 +210,22 @@ sig
(** Return a new rel-context enriched by with a given inner-most declaration. *)
val add : Declaration.t -> t -> t
+ (** Return the number of {e local declarations} in a given named-context. *)
+ val length : t -> int
+
(** Return a declaration designated by an identifier of the variable bound in that declaration.
@raise Not_found if the designated identifier is not bound in a given named-context. *)
val lookup : Id.t -> t -> Declaration.t
+ (** Check whether given two rel-contexts are equal. *)
+ val equal : t -> t -> bool
+
(** Map all terms in a given named-context. *)
val map : (Constr.t -> Constr.t) -> t -> t
+ (** Perform a given action on every declaration in a given named-context. *)
+ val iter : (Constr.t -> unit) -> t -> unit
+
(** Reduce all terms in a given named-context to a single value.
Innermost declarations are processed first. *)
val fold_inside : ('a -> Declaration.t -> 'a) -> init:'a -> t -> 'a
@@ -156,15 +234,6 @@ sig
Outermost declarations are processed first. *)
val fold_outside : (Declaration.t -> 'a -> 'a) -> t -> init:'a -> 'a
- (** Perform a given action on every declaration in a given named-context. *)
- val iter : (Constr.t -> unit) -> t -> unit
-
- (** Return the number of {e local declarations} in a given named-context. *)
- val length : t -> int
-
- (** Check whether given two named-contexts are equal. *)
- val equal : t -> t -> bool
-
(** Return the set of all identifiers bound in a given named-context. *)
val to_vars : t -> Id.Set.t
@@ -180,7 +249,7 @@ sig
module Declaration :
sig
type t = Id.t list * Constr.t option * Constr.t
- val map : (Constr.t -> Constr.t) -> t -> t
+ val map_constr : (Constr.t -> Constr.t) -> t -> t
end
type t = Declaration.t list
diff --git a/kernel/cooking.ml b/kernel/cooking.ml
index 3ab6983d8..d2106f860 100644
--- a/kernel/cooking.ml
+++ b/kernel/cooking.ml
@@ -201,8 +201,11 @@ let cook_constant env { from = cb; info } =
cb.const_body
in
let const_hyps =
- Context.Named.fold_outside (fun (h,_,_) hyps ->
- List.filter (fun (id,_,_) -> not (Id.equal id h)) hyps)
+ Context.Named.fold_outside (fun decl hyps ->
+ let open Context.Named.Declaration in
+ let h = get_id decl in
+ List.filter (fun decl -> let id = get_id decl in
+ not (Id.equal id h)) hyps)
hyps ~init:cb.const_hyps in
let typ = match cb.const_type with
| RegularArity t ->
diff --git a/kernel/csymtable.ml b/kernel/csymtable.ml
index 8f60216af..cfbb89f06 100644
--- a/kernel/csymtable.ml
+++ b/kernel/csymtable.ml
@@ -189,16 +189,18 @@ and slot_for_fv env fv =
let nv = Pre_env.lookup_named_val id env in
begin match force_lazy_val nv with
| None ->
- let _, b, _ = Context.Named.lookup id env.env_named_context in
- fill_fv_cache nv id val_of_named idfun b
+ let open Context.Named in
+ let open Context.Named.Declaration in
+ fill_fv_cache nv id val_of_named idfun (lookup id env.env_named_context |> get_value)
| Some (v, _) -> v
end
| FVrel i ->
let rv = Pre_env.lookup_rel_val i env in
begin match force_lazy_val rv with
| None ->
- let _, b, _ = Context.Rel.lookup i env.env_rel_context in
- fill_fv_cache rv i val_of_rel env_of_rel b
+ let open Context.Rel in
+ let open Context.Rel.Declaration in
+ fill_fv_cache rv i val_of_rel env_of_rel (lookup i env.env_rel_context |> get_value)
| Some (v, _) -> v
end
| FVuniv_var idu ->
diff --git a/kernel/declareops.ml b/kernel/declareops.ml
index f73eea030..cb67135ad 100644
--- a/kernel/declareops.ml
+++ b/kernel/declareops.ml
@@ -87,10 +87,18 @@ let is_opaque cb = match cb.const_body with
(** {7 Constant substitutions } *)
-let subst_rel_declaration sub (id,copt,t as x) =
- let copt' = Option.smartmap (subst_mps sub) copt in
- let t' = subst_mps sub t in
- if copt == copt' && t == t' then x else (id,copt',t')
+let subst_rel_declaration sub x =
+ let open Context.Rel.Declaration in
+ match x with
+ | LocalAssum (id,t) ->
+ let t' = subst_mps sub t in
+ if t == t' then x
+ else LocalAssum (id,t')
+ | LocalDef (id,v,t) ->
+ let v' = subst_mps sub v in
+ let t' = subst_mps sub t in
+ if v == v' && t == t' then x
+ else LocalDef (id,v',t')
let subst_rel_context sub = List.smartmap (subst_rel_declaration sub)
@@ -140,11 +148,20 @@ let subst_const_body sub cb =
share internal fields (e.g. constr), and not the records
themselves. But would it really bring substantial gains ? *)
-let hcons_rel_decl ((n,oc,t) as d) =
- let n' = Names.Name.hcons n
- and oc' = Option.smartmap Term.hcons_constr oc
- and t' = Term.hcons_types t
- in if n' == n && oc' == oc && t' == t then d else (n',oc',t')
+let hcons_rel_decl d =
+ let open Context.Rel.Declaration in
+ match d with
+ | LocalAssum (n,t) ->
+ let n' = Names.Name.hcons n
+ and t' = Term.hcons_types t in
+ if n' == n && t' == t then d
+ else LocalAssum (n',t')
+ | LocalDef (n,v,t) ->
+ let n' = Names.Name.hcons n
+ and v' = Term.hcons_constr v
+ and t' = Term.hcons_types t in
+ if n' == n && v' == v && t' == t then d
+ else LocalDef (n',v',t')
let hcons_rel_context l = List.smartmap hcons_rel_decl l
diff --git a/kernel/environ.ml b/kernel/environ.ml
index 847e1d08f..1089dff92 100644
--- a/kernel/environ.ml
+++ b/kernel/environ.ml
@@ -72,9 +72,8 @@ let lookup_rel n env =
Context.Rel.lookup n env.env_rel_context
let evaluable_rel n env =
- match lookup_rel n env with
- | (_,Some _,_) -> true
- | _ -> false
+ let open Context.Rel.Declaration in
+ lookup_rel n env |> is_local_def
let nb_rel env = env.env_nb_rel
@@ -83,7 +82,8 @@ let push_rel = push_rel
let push_rel_context ctxt x = Context.Rel.fold_outside push_rel ctxt ~init:x
let push_rec_types (lna,typarray,_) env =
- let ctxt = Array.map2_i (fun i na t -> (na, None, lift i t)) lna typarray in
+ let open Context.Rel.Declaration in
+ let ctxt = Array.map2_i (fun i na t -> LocalAssum (na, lift i t)) lna typarray in
Array.fold_left (fun e assum -> push_rel assum e) env ctxt
let fold_rel_context f env ~init =
@@ -107,17 +107,8 @@ let named_vals_of_val = snd
(* [map_named_val f ctxt] apply [f] to the body and the type of
each declarations.
*** /!\ *** [f t] should be convertible with t *)
-let map_named_val f (ctxt,ctxtv) =
- let rec map ctx = match ctx with
- | [] -> []
- | (id, body, typ) :: rem ->
- let body' = Option.smartmap f body in
- let typ' = f typ in
- let rem' = map rem in
- if body' == body && typ' == typ && rem' == rem then ctx
- else (id, body', typ') :: rem'
- in
- (map ctxt, ctxtv)
+let map_named_val f =
+ on_fst (Context.Named.map f)
let empty_named_context = Context.Named.empty
@@ -138,10 +129,10 @@ let eq_named_context_val c1 c2 =
(* A local const is evaluable if it is defined *)
let named_type id env =
- let (_,_,t) = lookup_named id env in t
+ lookup_named id env |> Context.Named.Declaration.get_type
let named_body id env =
- let (_,b,_) = lookup_named id env in b
+ lookup_named id env |> Context.Named.Declaration.get_value
let evaluable_named id env =
match named_body id env with
@@ -426,15 +417,16 @@ let global_vars_set env constr =
contained in the types of the needed variables. *)
let really_needed env needed =
+ let open Context.Named.Declaration in
Context.Named.fold_inside
- (fun need (id,copt,t) ->
- if Id.Set.mem id need then
+ (fun need decl ->
+ if Id.Set.mem (get_id decl) need then
let globc =
- match copt with
- | None -> Id.Set.empty
- | Some c -> global_vars_set env c in
+ match decl with
+ | LocalAssum _ -> Id.Set.empty
+ | LocalDef (_,c,_) -> global_vars_set env c in
Id.Set.union
- (global_vars_set env t)
+ (global_vars_set env (get_type decl))
(Id.Set.union globc need)
else need)
~init:needed
@@ -443,8 +435,9 @@ let really_needed env needed =
let keep_hyps env needed =
let really_needed = really_needed env needed in
Context.Named.fold_outside
- (fun (id,_,_ as d) nsign ->
- if Id.Set.mem id really_needed then Context.Named.add d nsign
+ (fun d nsign ->
+ let open Context.Named.Declaration in
+ if Id.Set.mem (get_id d) really_needed then Context.Named.add d nsign
else nsign)
(named_context env)
~init:empty_named_context
@@ -494,11 +487,12 @@ let compile_constant_body = Cbytegen.compile_constant_body false
exception Hyp_not_found
let apply_to_hyp (ctxt,vals) id f =
+ let open Context.Named.Declaration in
let rec aux rtail ctxt vals =
match ctxt, vals with
- | (idc,c,ct as d)::ctxt, v::vals ->
- if Id.equal idc id then
- (f ctxt d rtail)::ctxt, v::vals
+ | d::ctxt, v::vals ->
+ if Id.equal (get_id d) id then
+ (f ctxt d rtail)::ctxt, v::vals
else
let ctxt',vals' = aux (d::rtail) ctxt vals in
d::ctxt', v::vals'
@@ -507,10 +501,11 @@ let apply_to_hyp (ctxt,vals) id f =
in aux [] ctxt vals
let apply_to_hyp_and_dependent_on (ctxt,vals) id f g =
+ let open Context.Named.Declaration in
let rec aux ctxt vals =
match ctxt,vals with
- | (idc,c,ct as d)::ctxt, v::vals ->
- if Id.equal idc id then
+ | d::ctxt, v::vals ->
+ if Id.equal (get_id d) id then
let sign = ctxt,vals in
push_named_context_val (f d sign) sign
else
@@ -521,10 +516,11 @@ let apply_to_hyp_and_dependent_on (ctxt,vals) id f g =
in aux ctxt vals
let insert_after_hyp (ctxt,vals) id d check =
+ let open Context.Named.Declaration in
let rec aux ctxt vals =
match ctxt, vals with
- | (idc,c,ct)::ctxt', v::vals' ->
- if Id.equal idc id then begin
+ | decl::ctxt', v::vals' ->
+ if Id.equal (get_id decl) id then begin
check ctxt;
push_named_context_val d (ctxt,vals)
end else
@@ -537,12 +533,12 @@ let insert_after_hyp (ctxt,vals) id d check =
(* To be used in Logic.clear_hyps *)
let remove_hyps ids check_context check_value (ctxt, vals) =
+ let open Context.Named.Declaration in
let rec remove_hyps ctxt vals = match ctxt, vals with
| [], [] -> [], []
| d :: rctxt, (nid, v) :: rvals ->
- let (id, _, _) = d in
let ans = remove_hyps rctxt rvals in
- if Id.Set.mem id ids then ans
+ if Id.Set.mem (get_id d) ids then ans
else
let (rctxt', rvals') = ans in
let d' = check_context d in
diff --git a/kernel/fast_typeops.ml b/kernel/fast_typeops.ml
index ebc1853d9..df95c93dc 100644
--- a/kernel/fast_typeops.ml
+++ b/kernel/fast_typeops.ml
@@ -73,8 +73,9 @@ let judge_of_type u =
let judge_of_relative env n =
try
- let (_,_,typ) = lookup_rel n env in
- lift n typ
+ let open Context.Rel.Declaration in
+ let typ = get_type (lookup_rel n env) in
+ lift n typ
with Not_found ->
error_unbound_rel env n
@@ -91,7 +92,10 @@ let judge_of_variable env id =
(* TODO: check order? *)
let check_hyps_inclusion env f c sign =
Context.Named.fold_outside
- (fun (id,_,ty1) () ->
+ (fun decl () ->
+ let open Context.Named.Declaration in
+ let id = get_id decl in
+ let ty1 = get_type decl in
try
let ty2 = named_type id env in
if not (eq_constr ty2 ty1) then raise Exit
@@ -325,6 +329,7 @@ let type_fixpoint env lna lar vdef vdeft =
Ind et Constructsi un jour cela devient des constructions
arbitraires et non plus des variables *)
let rec execute env cstr =
+ let open Context.Rel.Declaration in
match kind_of_term cstr with
(* Atomic terms *)
| Sort (Prop c) ->
@@ -368,13 +373,13 @@ let rec execute env cstr =
| Lambda (name,c1,c2) ->
let _ = execute_is_type env c1 in
- let env1 = push_rel (name,None,c1) env in
+ let env1 = push_rel (LocalAssum (name,c1)) env in
let c2t = execute env1 c2 in
judge_of_abstraction env name c1 c2t
| Prod (name,c1,c2) ->
let vars = execute_is_type env c1 in
- let env1 = push_rel (name,None,c1) env in
+ let env1 = push_rel (LocalAssum (name,c1)) env in
let vars' = execute_is_type env1 c2 in
judge_of_product env name vars vars'
@@ -382,7 +387,7 @@ let rec execute env cstr =
let c1t = execute env c1 in
let _c2s = execute_is_type env c2 in
let _ = judge_of_cast env c1 c1t DEFAULTcast c2 in
- let env1 = push_rel (name,Some c1,c2) env in
+ let env1 = push_rel (LocalDef (name,c1,c2)) env in
let c3t = execute env1 c3 in
subst1 c1 c3t
diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml
index a8625009c..4834f95d1 100644
--- a/kernel/indtypes.ml
+++ b/kernel/indtypes.ml
@@ -20,6 +20,7 @@ open Reduction
open Typeops
open Entries
open Pp
+open Context.Rel.Declaration
(* Tell if indices (aka real arguments) contribute to size of inductive type *)
(* If yes, this is compatible with the univalent model *)
@@ -122,7 +123,7 @@ let infos_and_sort env t =
match kind_of_term t with
| Prod (name,c1,c2) ->
let varj = infer_type env c1 in
- let env1 = Environ.push_rel (name,None,varj.utj_val) env in
+ let env1 = Environ.push_rel (LocalAssum (name,varj.utj_val)) env in
let max = Universe.sup max (univ_of_sort varj.utj_type) in
aux env1 c2 max
| _ when is_constructor_head t -> max
@@ -168,12 +169,14 @@ let infer_constructor_packet env_ar_par params lc =
(* If indices matter *)
let cumulate_arity_large_levels env sign =
fst (List.fold_right
- (fun (_,b,t as d) (lev,env) ->
- if Option.is_empty b then
+ (fun d (lev,env) ->
+ match d with
+ | LocalAssum (_,t) ->
let tj = infer_type env t in
let u = univ_of_sort tj.utj_type in
(Universe.sup u lev, push_rel d env)
- else lev, push_rel d env)
+ | LocalDef _ ->
+ lev, push_rel d env)
sign (Universe.type0m,env))
let is_impredicative env u =
@@ -184,12 +187,12 @@ let is_impredicative env u =
from the most recent and ignoring let-definitions) is not contributing
or is Some u_k if its level is u_k and is contributing. *)
let param_ccls params =
- let fold acc = function (_, None, p) ->
+ let fold acc = function (LocalAssum (_, p)) ->
(let c = strip_prod_assum p in
match kind_of_term c with
| Sort (Type u) -> Univ.Universe.level u
| _ -> None) :: acc
- | _ -> acc
+ | LocalDef _ -> acc
in
List.fold_left fold [] params
@@ -249,7 +252,7 @@ let typecheck_inductive env mie =
let full_arity = it_mkProd_or_LetIn arity params in
let id = ind.mind_entry_typename in
let env_ar' =
- push_rel (Name id, None, full_arity) env_ar in
+ push_rel (LocalAssum (Name id, full_arity)) env_ar in
(* (add_constraints cst2 env_ar) in *)
(env_ar', (id,full_arity,sign @ params,expltype,deflev,inflev)::l))
(env',[])
@@ -390,7 +393,7 @@ let check_correct_par (env,n,ntypes,_) hyps l largs =
let nhyps = List.length hyps in
let rec check k index = function
| [] -> ()
- | (_,Some _,_)::hyps -> check k (index+1) hyps
+ | LocalDef _ :: hyps -> check k (index+1) hyps
| _::hyps ->
match kind_of_term (whd_betadeltaiota env lpar.(k)) with
| Rel w when Int.equal w index -> check (k-1) (index+1) hyps
@@ -412,7 +415,7 @@ if Int.equal nmr 0 then 0 else
function
([],_) -> nmr
| (_,[]) -> assert false (* |hyps|>=nmr *)
- | (lp,(_,Some _,_)::hyps) -> find k (index-1) (lp,hyps)
+ | (lp, LocalDef _ :: hyps) -> find k (index-1) (lp,hyps)
| (p::lp,_::hyps) ->
( match kind_of_term (whd_betadeltaiota env p) with
| Rel w when Int.equal w index -> find (k+1) (index-1) (lp,hyps)
@@ -426,15 +429,15 @@ if Int.equal nmr 0 then 0 else
[lra] is the list of recursive tree of each variable
*)
let ienv_push_var (env, n, ntypes, lra) (x,a,ra) =
- (push_rel (x,None,a) env, n+1, ntypes, (Norec,ra)::lra)
+ (push_rel (LocalAssum (x,a)) env, n+1, ntypes, (Norec,ra)::lra)
let ienv_push_inductive (env, n, ntypes, ra_env) ((mi,u),lpar) =
let auxntyp = 1 in
let specif = (lookup_mind_specif env mi, u) in
let ty = type_of_inductive env specif in
let env' =
- push_rel (Anonymous,None,
- hnf_prod_applist env ty lpar) env in
+ let decl = LocalAssum (Anonymous, hnf_prod_applist env ty lpar) in
+ push_rel decl env in
let ra_env' =
(Imbr mi,(Rtree.mk_rec_calls 1).(0)) ::
List.map (fun (r,t) -> (r,Rtree.lift 1 t)) ra_env in
@@ -726,9 +729,9 @@ let compute_projections ((kn, _ as ind), u as indu) n x nparamargs params
let body = mkCase (ci, p, mkRel 1, [|lift 1 branch|]) in
it_mkLambda_or_LetIn (mkLambda (x,indty,body)) params
in
- let projections (na, b, t) (i, j, kns, pbs, subst, letsubst) =
- match b with
- | Some c ->
+ let projections decl (i, j, kns, pbs, subst, letsubst) =
+ match decl with
+ | LocalDef (na,c,t) ->
(* From [params, field1,..,fieldj |- c(params,field1,..,fieldj)]
to [params, x:I, field1,..,fieldj |- c(params,field1,..,fieldj)] *)
let c = liftn 1 j c in
@@ -746,7 +749,7 @@ let compute_projections ((kn, _ as ind), u as indu) n x nparamargs params
to [params-wo-let, x:I |- subst:(params, x:I, field1,..,fieldj+1)] *)
let letsubst = c2 :: letsubst in
(i, j+1, kns, pbs, subst, letsubst)
- | None ->
+ | LocalAssum (na,t) ->
match na with
| Name id ->
let kn = Constant.make1 (KerName.make mp dp (Label.of_id id)) in
diff --git a/kernel/inductive.ml b/kernel/inductive.ml
index dd49c4a1b..ca29d83f6 100644
--- a/kernel/inductive.ml
+++ b/kernel/inductive.ml
@@ -17,6 +17,7 @@ open Declareops
open Environ
open Reduction
open Type_errors
+open Context.Rel.Declaration
type mind_specif = mutual_inductive_body * one_inductive_body
@@ -77,10 +78,10 @@ let instantiate_params full t u args sign =
anomaly ~label:"instantiate_params" (Pp.str "type, ctxt and args mismatch") in
let (rem_args, subs, ty) =
Context.Rel.fold_outside
- (fun (_,copt,_) (largs,subs,ty) ->
- match (copt, largs, kind_of_term ty) with
- | (None, a::args, Prod(_,_,t)) -> (args, a::subs, t)
- | (Some b,_,LetIn(_,_,_,t)) ->
+ (fun decl (largs,subs,ty) ->
+ match (decl, largs, kind_of_term ty) with
+ | (LocalAssum _, a::args, Prod(_,_,t)) -> (args, a::subs, t)
+ | (LocalDef (_,b,_), _, LetIn(_,_,_,t)) ->
(largs, (substl subs (subst_instance_constr u b))::subs, t)
| (_,[],_) -> if full then fail() else ([], subs, ty)
| _ -> fail ())
@@ -152,7 +153,7 @@ let remember_subst u subst =
(* Propagate the new levels in the signature *)
let make_subst env =
let rec make subst = function
- | (_,Some _,_)::sign, exp, args ->
+ | LocalDef _ :: sign, exp, args ->
make subst (sign, exp, args)
| d::sign, None::exp, args ->
let args = match args with _::args -> args | [] -> [] in
@@ -165,7 +166,7 @@ let make_subst env =
(* a useless extra constraint *)
let s = sort_as_univ (snd (dest_arity env (Lazy.force a))) in
make (cons_subst u s subst) (sign, exp, args)
- | (na,None,t)::sign, Some u::exp, [] ->
+ | LocalAssum (na,t) :: sign, Some u::exp, [] ->
(* No more argument here: we add the remaining universes to the *)
(* substitution (when [u] is distinct from all other universes in the *)
(* template, it is identity substitution otherwise (ie. when u is *)
@@ -314,14 +315,14 @@ let is_correct_arity env c pj ind specif params =
let rec srec env pt ar =
let pt' = whd_betadeltaiota env pt in
match kind_of_term pt', ar with
- | Prod (na1,a1,t), (_,None,a1')::ar' ->
+ | Prod (na1,a1,t), (LocalAssum (_,a1'))::ar' ->
let () =
try conv env a1 a1'
with NotConvertible -> raise (LocalArity None) in
- srec (push_rel (na1,None,a1) env) t ar'
+ srec (push_rel (LocalAssum (na1,a1)) env) t ar'
(* The last Prod domain is the type of the scrutinee *)
| Prod (na1,a1,a2), [] -> (* whnf of t was not needed here! *)
- let env' = push_rel (na1,None,a1) env in
+ let env' = push_rel (LocalAssum (na1,a1)) env in
let ksort = match kind_of_term (whd_betadeltaiota env' a2) with
| Sort s -> family_of_sort s
| _ -> raise (LocalArity None) in
@@ -330,7 +331,7 @@ let is_correct_arity env c pj ind specif params =
try conv env a1 dep_ind
with NotConvertible -> raise (LocalArity None) in
check_allowed_sort ksort specif
- | _, (_,Some _,_ as d)::ar' ->
+ | _, (LocalDef _ as d)::ar' ->
srec (push_rel d env) (lift 1 pt') ar'
| _ ->
raise (LocalArity None)
@@ -482,7 +483,7 @@ let make_renv env recarg tree =
genv = [Lazy.lazy_from_val(Subterm(Large,tree))] }
let push_var renv (x,ty,spec) =
- { env = push_rel (x,None,ty) renv.env;
+ { env = push_rel (LocalAssum (x,ty)) renv.env;
rel_min = renv.rel_min+1;
genv = spec:: renv.genv }
@@ -568,14 +569,14 @@ let check_inductive_codomain env p =
(* The following functions are almost duplicated from indtypes.ml, except
that they carry here a poorer environment (containing less information). *)
let ienv_push_var (env, lra) (x,a,ra) =
- (push_rel (x,None,a) env, (Norec,ra)::lra)
+ (push_rel (LocalAssum (x,a)) env, (Norec,ra)::lra)
let ienv_push_inductive (env, ra_env) ((mind,u),lpar) =
let mib = Environ.lookup_mind mind env in
let ntypes = mib.mind_ntypes in
let push_ind specif env =
- push_rel (Anonymous,None,
- hnf_prod_applist env (type_of_inductive env ((mib,specif),u)) lpar) env
+ let decl = LocalAssum (Anonymous, hnf_prod_applist env (type_of_inductive env ((mib,specif),u)) lpar) in
+ push_rel decl env
in
let env = Array.fold_right push_ind mib.mind_packets env in
let rc = Array.mapi (fun j t -> (Imbr (mind,j),t)) (Rtree.mk_rec_calls ntypes) in
@@ -848,7 +849,7 @@ let filter_stack_domain env ci p stack =
let t = whd_betadeltaiota env ar in
match stack, kind_of_term t with
| elt :: stack', Prod (n,a,c0) ->
- let d = (n,None,a) in
+ let d = LocalAssum (n,a) in
let ty, args = decompose_app (whd_betadeltaiota env a) in
let elt = match kind_of_term ty with
| Ind ind ->
@@ -905,10 +906,10 @@ let check_one_fix renv recpos trees def =
end
else
begin
- match pi2 (lookup_rel p renv.env) with
- | None ->
+ match lookup_rel p renv.env with
+ | LocalAssum _ ->
List.iter (check_rec_call renv []) l
- | Some c ->
+ | LocalDef (_,c,_) ->
try List.iter (check_rec_call renv []) l
with FixGuardError _ ->
check_rec_call renv stack (applist(lift p c,l))
@@ -983,10 +984,11 @@ let check_one_fix renv recpos trees def =
| Var id ->
begin
- match pi2 (lookup_named id renv.env) with
- | None ->
+ let open Context.Named.Declaration in
+ match lookup_named id renv.env with
+ | LocalAssum _ ->
List.iter (check_rec_call renv []) l
- | Some c ->
+ | LocalDef (_,c,_) ->
try List.iter (check_rec_call renv []) l
with (FixGuardError _) ->
check_rec_call renv stack (applist(c,l))
@@ -1040,7 +1042,7 @@ let inductive_of_mutfix env ((nvect,bodynum),(names,types,bodies as recdef)) =
match kind_of_term (whd_betadeltaiota env def) with
| Lambda (x,a,b) ->
if noccur_with_meta n nbfix a then
- let env' = push_rel (x, None, a) env in
+ let env' = push_rel (LocalAssum (x,a)) env in
if Int.equal n (k + 1) then
(* get the inductive type of the fixpoint *)
let (mind, _) =
@@ -1090,7 +1092,7 @@ let rec codomain_is_coind env c =
let b = whd_betadeltaiota env c in
match kind_of_term b with
| Prod (x,a,b) ->
- codomain_is_coind (push_rel (x, None, a) env) b
+ codomain_is_coind (push_rel (LocalAssum (x,a)) env) b
| _ ->
(try find_coinductive env b
with Not_found ->
@@ -1131,7 +1133,7 @@ let check_one_cofix env nbfix def deftype =
| Lambda (x,a,b) ->
let () = assert (List.is_empty args) in
if noccur_with_meta n nbfix a then
- let env' = push_rel (x, None, a) env in
+ let env' = push_rel (LocalAssum (x,a)) env in
check_rec_call env' alreadygrd (n+1) tree vlra b
else
raise (CoFixGuardError (env,RecCallInTypeOfAbstraction a))
diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml
index 711096b2b..47274a5cd 100644
--- a/kernel/nativecode.ml
+++ b/kernel/nativecode.ml
@@ -1832,24 +1832,25 @@ and apply_fv env sigma univ (fv_named,fv_rel) auxdefs ml =
auxdefs, MLlet(aux_name, ml, mkMLapp (MLlocal aux_name) (Array.of_list (fv_rel@fv_named)))
and compile_rel env sigma univ auxdefs n =
- let (_,body,_) = Context.Rel.lookup n env.env_rel_context in
+ let decl = Context.Rel.lookup n env.env_rel_context in
let n = Context.Rel.length env.env_rel_context - n in
- match body with
- | Some t ->
+ let open Context.Rel.Declaration in
+ match decl with
+ | LocalDef (_,t,_) ->
let code = lambda_of_constr env sigma t in
let auxdefs,code = compile_with_fv env sigma univ auxdefs None code in
Glet(Grel n, code)::auxdefs
- | None ->
+ | LocalAssum _ ->
Glet(Grel n, MLprimitive (Mk_rel n))::auxdefs
and compile_named env sigma univ auxdefs id =
- let (_,body,_) = Context.Named.lookup id env.env_named_context in
- match body with
- | Some t ->
+ let open Context.Named.Declaration in
+ match Context.Named.lookup id env.env_named_context with
+ | LocalDef (_,t,_) ->
let code = lambda_of_constr env sigma t in
let auxdefs,code = compile_with_fv env sigma univ auxdefs None code in
Glet(Gnamed id, code)::auxdefs
- | None ->
+ | LocalAssum _ ->
Glet(Gnamed id, MLprimitive (Mk_var id))::auxdefs
let compile_constant env sigma prefix ~interactive con cb =
diff --git a/kernel/nativelambda.ml b/kernel/nativelambda.ml
index 01f59df15..91b40be7e 100644
--- a/kernel/nativelambda.ml
+++ b/kernel/nativelambda.ml
@@ -727,7 +727,8 @@ let optimize lam =
let lambda_of_constr env sigma c =
set_global_env env;
let env = Renv.make () in
- let ids = List.rev_map (fun (id, _, _) -> id) !global_env.env_rel_context in
+ let open Context.Rel.Declaration in
+ let ids = List.rev_map get_name !global_env.env_rel_context in
Renv.push_rels env (Array.of_list ids);
let lam = lambda_of_constr env sigma c in
(* if Flags.vm_draw_opt () then begin
diff --git a/kernel/pre_env.ml b/kernel/pre_env.ml
index 4c1b2c5a6..99d99e694 100644
--- a/kernel/pre_env.ml
+++ b/kernel/pre_env.ml
@@ -18,6 +18,7 @@ open Names
open Univ
open Term
open Declarations
+open Context.Named.Declaration
(* The type of environments. *)
@@ -124,18 +125,16 @@ let env_of_rel n env =
(* Named context *)
let push_named_context_val d (ctxt,vals) =
- let id,_,_ = d in
let rval = ref VKnone in
- Context.Named.add d ctxt, (id,rval)::vals
+ Context.Named.add d ctxt, (get_id d,rval)::vals
let push_named d env =
(* if not (env.env_rel_context = []) then raise (ASSERT env.env_rel_context);
assert (env.env_rel_context = []); *)
- let id,body,_ = d in
let rval = ref VKnone in
{ env_globals = env.env_globals;
env_named_context = Context.Named.add d env.env_named_context;
- env_named_vals = (id, rval) :: env.env_named_vals;
+ env_named_vals = (get_id d, rval) :: env.env_named_vals;
env_rel_context = env.env_rel_context;
env_rel_val = env.env_rel_val;
env_nb_rel = env.env_nb_rel;
diff --git a/kernel/reduction.ml b/kernel/reduction.ml
index 40b80cc5e..cfc286135 100644
--- a/kernel/reduction.ml
+++ b/kernel/reduction.ml
@@ -24,6 +24,7 @@ open Univ
open Environ
open Closure
open Esubst
+open Context.Rel.Declaration
let rec is_empty_stack = function
[] -> true
@@ -739,7 +740,7 @@ let dest_prod env =
let t = whd_betadeltaiota env c in
match kind_of_term t with
| Prod (n,a,c0) ->
- let d = (n,None,a) in
+ let d = LocalAssum (n,a) in
decrec (push_rel d env) (Context.Rel.add d m) c0
| _ -> m,t
in
@@ -751,10 +752,10 @@ let dest_prod_assum env =
let rty = whd_betadeltaiota_nolet env ty in
match kind_of_term rty with
| Prod (x,t,c) ->
- let d = (x,None,t) in
+ let d = LocalAssum (x,t) in
prodec_rec (push_rel d env) (Context.Rel.add d l) c
| LetIn (x,b,t,c) ->
- let d = (x,Some b,t) in
+ let d = LocalDef (x,b,t) in
prodec_rec (push_rel d env) (Context.Rel.add d l) c
| Cast (c,_,_) -> prodec_rec env l c
| _ ->
@@ -769,10 +770,10 @@ let dest_lam_assum env =
let rty = whd_betadeltaiota_nolet env ty in
match kind_of_term rty with
| Lambda (x,t,c) ->
- let d = (x,None,t) in
+ let d = LocalAssum (x,t) in
lamec_rec (push_rel d env) (Context.Rel.add d l) c
| LetIn (x,b,t,c) ->
- let d = (x,Some b,t) in
+ let d = LocalDef (x,b,t) in
lamec_rec (push_rel d env) (Context.Rel.add d l) c
| Cast (c,_,_) -> lamec_rec env l c
| _ -> l,rty
diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml
index e56a6e099..8a402322f 100644
--- a/kernel/safe_typing.ml
+++ b/kernel/safe_typing.ml
@@ -60,6 +60,7 @@
open Util
open Names
open Declarations
+open Context.Named.Declaration
(** {6 Safe environments }
@@ -362,7 +363,8 @@ let check_required current_libs needed =
hypothesis many many times, and the check performed here would
cost too much. *)
-let safe_push_named (id,_,_ as d) env =
+let safe_push_named d env =
+ let id = get_id d in
let _ =
try
let _ = Environ.lookup_named id env in
@@ -383,13 +385,13 @@ let push_named_def (id,de) senv =
(Opaqueproof.force_constraints (Environ.opaque_tables senv.env) o)
| _ -> assert false in
let senv' = push_context_set poly univs senv in
- let env'' = safe_push_named (id,Some c,typ) senv'.env in
+ let env'' = safe_push_named (LocalDef (id,c,typ)) senv'.env in
univs, {senv' with env=env''}
let push_named_assum ((id,t,poly),ctx) senv =
let senv' = push_context_set poly ctx senv in
let t = Term_typing.translate_local_assum senv'.env t in
- let env'' = safe_push_named (id,None,t) senv'.env in
+ let env'' = safe_push_named (LocalAssum (id,t)) senv'.env in
{senv' with env=env''}
diff --git a/kernel/term.ml b/kernel/term.ml
index 9ba45f540..4416770fe 100644
--- a/kernel/term.ml
+++ b/kernel/term.ml
@@ -383,40 +383,46 @@ let mkNamedLambda id typ c = mkLambda (Name id, typ, subst_var id c)
let mkNamedLetIn id c1 t c2 = mkLetIn (Name id, c1, t, subst_var id c2)
(* Constructs either [(x:t)c] or [[x=b:t]c] *)
-let mkProd_or_LetIn (na,body,t) c =
- match body with
- | None -> mkProd (na, t, c)
- | Some b -> mkLetIn (na, b, t, c)
-
-let mkNamedProd_or_LetIn (id,body,t) c =
- match body with
- | None -> mkNamedProd id t c
- | Some b -> mkNamedLetIn id b t c
+let mkProd_or_LetIn decl c =
+ let open Context.Rel.Declaration in
+ match decl with
+ | LocalAssum (na,t) -> mkProd (na, t, c)
+ | LocalDef (na,b,t) -> mkLetIn (na, b, t, c)
+
+let mkNamedProd_or_LetIn decl c =
+ let open Context.Named.Declaration in
+ match decl with
+ | LocalAssum (id,t) -> mkNamedProd id t c
+ | LocalDef (id,b,t) -> mkNamedLetIn id b t c
(* Constructs either [(x:t)c] or [c] where [x] is replaced by [b] *)
-let mkProd_wo_LetIn (na,body,t) c =
- match body with
- | None -> mkProd (na, t, c)
- | Some b -> subst1 b c
-
-let mkNamedProd_wo_LetIn (id,body,t) c =
- match body with
- | None -> mkNamedProd id t c
- | Some b -> subst1 b (subst_var id c)
+let mkProd_wo_LetIn decl c =
+ let open Context.Rel.Declaration in
+ match decl with
+ | LocalAssum (na,t) -> mkProd (na, t, c)
+ | LocalDef (na,b,t) -> subst1 b c
+
+let mkNamedProd_wo_LetIn decl c =
+ let open Context.Named.Declaration in
+ match decl with
+ | LocalAssum (id,t) -> mkNamedProd id t c
+ | LocalDef (id,b,t) -> subst1 b (subst_var id c)
(* non-dependent product t1 -> t2 *)
let mkArrow t1 t2 = mkProd (Anonymous, t1, t2)
(* Constructs either [[x:t]c] or [[x=b:t]c] *)
-let mkLambda_or_LetIn (na,body,t) c =
- match body with
- | None -> mkLambda (na, t, c)
- | Some b -> mkLetIn (na, b, t, c)
-
-let mkNamedLambda_or_LetIn (id,body,t) c =
- match body with
- | None -> mkNamedLambda id t c
- | Some b -> mkNamedLetIn id b t c
+let mkLambda_or_LetIn decl c =
+ let open Context.Rel.Declaration in
+ match decl with
+ | LocalAssum (na,t) -> mkLambda (na, t, c)
+ | LocalDef (na,b,t) -> mkLetIn (na, b, t, c)
+
+let mkNamedLambda_or_LetIn decl c =
+ let open Context.Named.Declaration in
+ match decl with
+ | LocalAssum (id,t) -> mkNamedLambda id t c
+ | LocalDef (id,b,t) -> mkNamedLetIn id b t c
(* prodn n [xn:Tn;..;x1:T1;Gamma] b = (x1:T1)..(xn:Tn)b *)
let prodn n env b =
@@ -576,10 +582,11 @@ let decompose_lam_n n =
(* Transforms a product term (x1:T1)..(xn:Tn)T into the pair
([(xn,Tn);...;(x1,T1)],T), where T is not a product *)
let decompose_prod_assum =
+ let open Context.Rel.Declaration in
let rec prodec_rec l c =
match kind_of_term c with
- | Prod (x,t,c) -> prodec_rec (Context.Rel.add (x,None,t) l) c
- | LetIn (x,b,t,c) -> prodec_rec (Context.Rel.add (x,Some b,t) l) c
+ | Prod (x,t,c) -> prodec_rec (Context.Rel.add (LocalAssum (x,t)) l) c
+ | LetIn (x,b,t,c) -> prodec_rec (Context.Rel.add (LocalDef (x,b,t)) l) c
| Cast (c,_,_) -> prodec_rec l c
| _ -> l,c
in
@@ -589,9 +596,10 @@ let decompose_prod_assum =
([(xn,Tn);...;(x1,T1)],T), where T is not a lambda *)
let decompose_lam_assum =
let rec lamdec_rec l c =
+ let open Context.Rel.Declaration in
match kind_of_term c with
- | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (x,None,t) l) c
- | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (x,Some b,t) l) c
+ | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (LocalAssum (x,t)) l) c
+ | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (LocalDef (x,b,t)) l) c
| Cast (c,_,_) -> lamdec_rec l c
| _ -> l,c
in
@@ -606,11 +614,13 @@ let decompose_prod_n_assum n =
error "decompose_prod_n_assum: integer parameter must be positive";
let rec prodec_rec l n c =
if Int.equal n 0 then l,c
- else match kind_of_term c with
- | Prod (x,t,c) -> prodec_rec (Context.Rel.add (x,None,t) l) (n-1) c
- | LetIn (x,b,t,c) -> prodec_rec (Context.Rel.add (x,Some b,t) l) (n-1) c
- | Cast (c,_,_) -> prodec_rec l n c
- | c -> error "decompose_prod_n_assum: not enough assumptions"
+ else
+ let open Context.Rel.Declaration in
+ match kind_of_term c with
+ | Prod (x,t,c) -> prodec_rec (Context.Rel.add (LocalAssum (x,t)) l) (n-1) c
+ | LetIn (x,b,t,c) -> prodec_rec (Context.Rel.add (LocalDef (x,b,t)) l) (n-1) c
+ | Cast (c,_,_) -> prodec_rec l n c
+ | c -> error "decompose_prod_n_assum: not enough assumptions"
in
prodec_rec Context.Rel.empty n
@@ -625,11 +635,13 @@ let decompose_lam_n_assum n =
error "decompose_lam_n_assum: integer parameter must be positive";
let rec lamdec_rec l n c =
if Int.equal n 0 then l,c
- else match kind_of_term c with
- | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (x,None,t) l) (n-1) c
- | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (x,Some b,t) l) n c
- | Cast (c,_,_) -> lamdec_rec l n c
- | c -> error "decompose_lam_n_assum: not enough abstractions"
+ else
+ let open Context.Rel.Declaration in
+ match kind_of_term c with
+ | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (LocalAssum (x,t)) l) (n-1) c
+ | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (LocalDef (x,b,t)) l) n c
+ | Cast (c,_,_) -> lamdec_rec l n c
+ | c -> error "decompose_lam_n_assum: not enough abstractions"
in
lamdec_rec Context.Rel.empty n
@@ -639,11 +651,13 @@ let decompose_lam_n_decls n =
error "decompose_lam_n_decls: integer parameter must be positive";
let rec lamdec_rec l n c =
if Int.equal n 0 then l,c
- else match kind_of_term c with
- | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (x,None,t) l) (n-1) c
- | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (x,Some b,t) l) (n-1) c
- | Cast (c,_,_) -> lamdec_rec l n c
- | c -> error "decompose_lam_n_decls: not enough abstractions"
+ else
+ let open Context.Rel.Declaration in
+ match kind_of_term c with
+ | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (LocalAssum (x,t)) l) (n-1) c
+ | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (LocalDef (x,b,t)) l) (n-1) c
+ | Cast (c,_,_) -> lamdec_rec l n c
+ | c -> error "decompose_lam_n_decls: not enough abstractions"
in
lamdec_rec Context.Rel.empty n
@@ -669,10 +683,11 @@ let strip_lam_n n t = snd (decompose_lam_n n t)
type arity = Context.Rel.t * sorts
let destArity =
+ let open Context.Rel.Declaration in
let rec prodec_rec l c =
match kind_of_term c with
- | Prod (x,t,c) -> prodec_rec ((x,None,t)::l) c
- | LetIn (x,b,t,c) -> prodec_rec ((x,Some b,t)::l) c
+ | Prod (x,t,c) -> prodec_rec (LocalAssum (x,t) :: l) c
+ | LetIn (x,b,t,c) -> prodec_rec (LocalDef (x,b,t) :: l) c
| Cast (c,_,_) -> prodec_rec l c
| Sort s -> l,s
| _ -> anomaly ~label:"destArity" (Pp.str "not an arity")
diff --git a/kernel/term_typing.ml b/kernel/term_typing.ml
index 979165e49..2a3ac957f 100644
--- a/kernel/term_typing.ml
+++ b/kernel/term_typing.ml
@@ -138,16 +138,17 @@ let check_signatures curmb sl =
let skip_trusted_seff sl b e =
let rec aux sl b e acc =
+ let open Context.Rel.Declaration in
match sl, kind_of_term b with
| (None|Some 0), _ -> b, e, acc
| Some sl, LetIn (n,c,ty,bo) ->
aux (Some (sl-1)) bo
- (Environ.push_rel (n,Some c,ty) e) (`Let(n,c,ty)::acc)
+ (Environ.push_rel (LocalDef (n,c,ty)) e) (`Let(n,c,ty)::acc)
| Some sl, App(hd,arg) ->
begin match kind_of_term hd with
| Lambda (n,ty,bo) ->
aux (Some (sl-1)) bo
- (Environ.push_rel (n,None,ty) e) (`Cut(n,ty,arg)::acc)
+ (Environ.push_rel (LocalAssum (n,ty)) e) (`Cut(n,ty,arg)::acc)
| _ -> assert false
end
| _ -> assert false
@@ -251,11 +252,13 @@ let global_vars_set_constant_type env = function
ctx ~init:Id.Set.empty
let record_aux env s_ty s_bo suggested_expr =
+ let open Context.Named.Declaration in
let in_ty = keep_hyps env s_ty in
let v =
String.concat " "
- (CList.map_filter (fun (id, _,_) ->
- if List.exists (fun (id',_,_) -> Id.equal id id') in_ty then None
+ (CList.map_filter (fun decl ->
+ let id = get_id decl in
+ if List.exists (Id.equal id % get_id) in_ty then None
else Some (Id.to_string id))
(keep_hyps env s_bo)) in
Aux_file.record_in_aux "context_used" (v ^ ";" ^ suggested_expr)
@@ -264,8 +267,9 @@ let suggest_proof_using = ref (fun _ _ _ _ _ -> "")
let set_suggest_proof_using f = suggest_proof_using := f
let build_constant_declaration kn env (def,typ,proj,poly,univs,inline_code,ctx) =
+ let open Context.Named.Declaration in
let check declared inferred =
- let mk_set l = List.fold_right Id.Set.add (List.map pi1 l) Id.Set.empty in
+ let mk_set l = List.fold_right Id.Set.add (List.map get_id l) Id.Set.empty in
let inferred_set, declared_set = mk_set inferred, mk_set declared in
if not (Id.Set.subset inferred_set declared_set) then
let l = Id.Set.elements (Idset.diff inferred_set declared_set) in
@@ -276,12 +280,13 @@ let build_constant_declaration kn env (def,typ,proj,poly,univs,inline_code,ctx)
str " used but not declared:" ++
fnl () ++ pr_sequence Id.print (List.rev l) ++ str ".")) in
let sort evn l =
- List.filter (fun (id,_,_) ->
- List.exists (fun (id',_,_) -> Names.Id.equal id id') l)
+ List.filter (fun decl ->
+ let id = get_id decl in
+ List.exists (Names.Id.equal id % get_id) l)
(named_context env) in
(* We try to postpone the computation of used section variables *)
let hyps, def =
- let context_ids = List.map pi1 (named_context env) in
+ let context_ids = List.map get_id (named_context env) in
match ctx with
| None when not (List.is_empty context_ids) ->
(* No declared section vars, and non-empty section context:
@@ -472,7 +477,8 @@ let translate_local_def mb env id centry =
| Undef _ -> ()
| Def _ -> ()
| OpaqueDef lc ->
- let context_ids = List.map pi1 (named_context env) in
+ let open Context.Named.Declaration in
+ let context_ids = List.map get_id (named_context env) in
let ids_typ = global_vars_set env typ in
let ids_def = global_vars_set env
(Opaqueproof.force_proof (opaque_tables env) lc) in
diff --git a/kernel/typeops.ml b/kernel/typeops.ml
index f31cba0a8..eeb12a2b4 100644
--- a/kernel/typeops.ml
+++ b/kernel/typeops.ml
@@ -77,8 +77,9 @@ let judge_of_type u =
(*s Type of a de Bruijn index. *)
let judge_of_relative env n =
+ let open Context.Rel.Declaration in
try
- let (_,_,typ) = lookup_rel n env in
+ let typ = get_type (lookup_rel n env) in
{ uj_val = mkRel n;
uj_type = lift n typ }
with Not_found ->
@@ -98,18 +99,20 @@ let judge_of_variable env id =
variables of the current env.
Order does not have to be checked assuming that all names are distinct *)
let check_hyps_inclusion env c sign =
+ let open Context.Named.Declaration in
Context.Named.fold_outside
- (fun (id,b1,ty1) () ->
+ (fun d1 () ->
+ let id = get_id d1 in
try
- let (_,b2,ty2) = lookup_named id env in
- conv env ty2 ty1;
- (match b2,b1 with
- | None, None -> ()
- | None, Some _ ->
+ let d2 = lookup_named id env in
+ conv env (get_type d2) (get_type d1);
+ (match d2,d1 with
+ | LocalAssum _, LocalAssum _ -> ()
+ | LocalAssum _, LocalDef _ ->
(* This is wrong, because we don't know if the body is
needed or not for typechecking: *) ()
- | Some _, None -> raise NotConvertible
- | Some b2, Some b1 -> conv env b2 b1);
+ | LocalDef _, LocalAssum _ -> raise NotConvertible
+ | LocalDef (_,b2,_), LocalDef (_,b1,_) -> conv env b2 b1);
with Not_found | NotConvertible | Option.Heterogeneous ->
error_reference_variables env id c)
sign
@@ -124,9 +127,10 @@ let extract_level env p =
match kind_of_term c with Sort (Type u) -> Univ.Universe.level u | _ -> None
let extract_context_levels env l =
- let fold l (_, b, p) = match b with
- | None -> extract_level env p :: l
- | _ -> l
+ let open Context.Rel.Declaration in
+ let fold l = function
+ | LocalAssum (_,p) -> extract_level env p :: l
+ | LocalDef _ -> l
in
List.fold_left fold [] l
@@ -416,6 +420,7 @@ let type_fixpoint env lna lar vdefj =
Ind et Constructsi un jour cela devient des constructions
arbitraires et non plus des variables *)
let rec execute env cstr =
+ let open Context.Rel.Declaration in
match kind_of_term cstr with
(* Atomic terms *)
| Sort (Prop c) ->
@@ -458,13 +463,13 @@ let rec execute env cstr =
| Lambda (name,c1,c2) ->
let varj = execute_type env c1 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 c2 in
judge_of_abstraction env name varj j'
| Prod (name,c1,c2) ->
let varj = execute_type env c1 in
- let env1 = push_rel (name,None,varj.utj_val) env in
+ let env1 = push_rel (LocalAssum (name,varj.utj_val)) env in
let varj' = execute_type env1 c2 in
judge_of_product env name varj varj'
@@ -472,7 +477,7 @@ let rec execute env cstr =
let j1 = execute env c1 in
let j2 = execute_type env c2 in
let _ = judge_of_cast env j1 DEFAULTcast j2 in
- let env1 = push_rel (name,Some j1.uj_val,j2.utj_val) env in
+ let env1 = push_rel (LocalDef (name,j1.uj_val,j2.utj_val)) env in
let j' = execute env1 c3 in
judge_of_letin env name j1 j2 j'
@@ -550,10 +555,10 @@ let infer_v env cv =
let infer_local_decl env id = function
| LocalDef c ->
let j = infer env c in
- (Name id, Some j.uj_val, j.uj_type)
+ Context.Rel.Declaration.LocalDef (Name id, j.uj_val, j.uj_type)
| LocalAssum c ->
let j = infer env c in
- (Name id, None, assumption_of_judgment env j)
+ Context.Rel.Declaration.LocalAssum (Name id, assumption_of_judgment env j)
let infer_local_decls env decls =
let rec inferec env = function
diff --git a/kernel/vars.ml b/kernel/vars.ml
index 4554c6be1..b935ab6b9 100644
--- a/kernel/vars.ml
+++ b/kernel/vars.ml
@@ -8,6 +8,7 @@
open Names
open Esubst
+open Context.Rel.Declaration
(*********************)
(* Occurring *)
@@ -159,17 +160,17 @@ let substnl laml n c = substn_many (make_subst laml) n c
let substl laml c = substn_many (make_subst laml) 0 c
let subst1 lam c = substn_many [|make_substituend lam|] 0 c
-let substnl_decl laml k r = Context.Rel.Declaration.map (fun c -> substnl laml k c) r
-let substl_decl laml r = Context.Rel.Declaration.map (fun c -> substnl laml 0 c) r
-let subst1_decl lam r = Context.Rel.Declaration.map (fun c -> subst1 lam c) r
+let substnl_decl laml k r = map_constr (fun c -> substnl laml k c) r
+let substl_decl laml r = map_constr (fun c -> substnl laml 0 c) r
+let subst1_decl lam r = map_constr (fun c -> subst1 lam c) r
(* Build a substitution from an instance, inserting missing let-ins *)
let subst_of_rel_context_instance sign l =
let rec aux subst sign l =
match sign, l with
- | (_,None,_)::sign', a::args' -> aux (a::subst) sign' args'
- | (_,Some c,_)::sign', args' ->
+ | LocalAssum _ :: sign', a::args' -> aux (a::subst) sign' args'
+ | LocalDef (_,c,_)::sign', args' ->
aux (substl subst c :: subst) sign' args'
| [], [] -> subst
| _ -> Errors.anomaly (Pp.str "Instance and signature do not match")
diff --git a/lib/util.ml b/lib/util.ml
index b67539918..6e22d6be5 100644
--- a/lib/util.ml
+++ b/lib/util.ml
@@ -88,6 +88,7 @@ let matrix_transpose mat =
let identity x = x
let compose f g x = f (g x)
+let (%) = compose
let const x _ = x
diff --git a/lib/util.mli b/lib/util.mli
index 7923c65a3..bf342c1e6 100644
--- a/lib/util.mli
+++ b/lib/util.mli
@@ -84,6 +84,7 @@ val matrix_transpose : 'a list list -> 'a list list
val identity : 'a -> 'a
val compose : ('a -> 'b) -> ('c -> 'a) -> 'c -> 'b
+val (%) : ('a -> 'b) -> ('c -> 'a) -> 'c -> 'b
val const : 'a -> 'b -> 'a
val iterate : ('a -> 'a) -> int -> 'a -> 'a
val repeat : int -> ('a -> unit) -> 'a -> unit
diff --git a/library/decls.ml b/library/decls.ml
index cafdcd0ab..6e21880f1 100644
--- a/library/decls.ml
+++ b/library/decls.ml
@@ -46,16 +46,20 @@ let constant_kind kn = Cmap.find kn !csttab
(** Miscellaneous functions. *)
+open Context.Named.Declaration
+
let initialize_named_context_for_proof () =
let sign = Global.named_context () in
List.fold_right
- (fun (id,c,t as d) signv ->
- let d = if variable_opacity id then (id,None,t) else d in
+ (fun d signv ->
+ let id = get_id d in
+ let d = if variable_opacity id then LocalAssum (id, get_type d) else d in
Environ.push_named_context_val d signv) sign Environ.empty_named_context_val
let last_section_hyps dir =
Context.Named.fold_outside
- (fun (id,_,_) sec_ids ->
+ (fun d sec_ids ->
+ let id = get_id d in
try if DirPath.equal dir (variable_path id) then id::sec_ids else sec_ids
with Not_found -> sec_ids)
(Environ.named_context (Global.env()))
diff --git a/library/heads.ml b/library/heads.ml
index 8124d3474..4c9b78976 100644
--- a/library/heads.ml
+++ b/library/heads.ml
@@ -15,6 +15,7 @@ open Environ
open Globnames
open Libobject
open Lib
+open Context.Named.Declaration
(** Characterization of the head of a term *)
@@ -63,9 +64,9 @@ let kind_of_head env t =
(try on_subterm k l b (variable_head id)
with Not_found ->
(* a goal variable *)
- match pi2 (lookup_named id env) with
- | Some c -> aux k l c b
- | None -> NotImmediatelyComputableHead)
+ match lookup_named id env with
+ | LocalDef (_,c,_) -> aux k l c b
+ | LocalAssum _ -> NotImmediatelyComputableHead)
| Const (cst,_) ->
(try on_subterm k l b (constant_head cst)
with Not_found ->
@@ -132,8 +133,8 @@ let compute_head = function
| None -> RigidHead (RigidParameter cst)
| Some c -> kind_of_head env c)
| EvalVarRef id ->
- (match pi2 (Global.lookup_named id) with
- | Some c when not (Decls.variable_opacity id) ->
+ (match Global.lookup_named id with
+ | LocalDef (_,c,_) when not (Decls.variable_opacity id) ->
kind_of_head (Global.env()) c
| _ ->
RigidHead (RigidVar id))
diff --git a/library/impargs.ml b/library/impargs.ml
index f5f6a3eba..4e344a954 100644
--- a/library/impargs.ml
+++ b/library/impargs.ml
@@ -165,6 +165,7 @@ let update pos rig (na,st) =
(* modified is_rigid_reference with a truncated env *)
let is_flexible_reference env bound depth f =
+ let open Context.Named.Declaration in
match kind_of_term f with
| Rel n when n >= bound+depth -> (* inductive type *) false
| Rel n when n >= depth -> (* previous argument *) true
@@ -173,8 +174,7 @@ let is_flexible_reference env bound depth f =
let cb = Environ.lookup_constant kn env in
(match cb.const_body with Def _ -> true | _ -> false)
| Var id ->
- let (_, value, _) = Environ.lookup_named id env in
- begin match value with None -> false | _ -> true end
+ Environ.lookup_named id env |> is_local_def
| Ind _ | Construct _ -> false
| _ -> true
@@ -234,13 +234,14 @@ let find_displayed_name_in all avoid na (_,b as envnames_b) =
let compute_implicits_gen strict strongly_strict revpat contextual all env t =
let rigid = ref true in
+ let open Context.Rel.Declaration in
let rec aux env avoid n names t =
let t = whd_betadeltaiota env t in
match kind_of_term t with
| Prod (na,a,b) ->
let na',avoid' = find_displayed_name_in all avoid na (names,b) in
add_free_rels_until strict strongly_strict revpat n env a (Hyp (n+1))
- (aux (push_rel (na',None,a) env) avoid' (n+1) (na'::names) b)
+ (aux (push_rel (LocalAssum (na',a)) env) avoid' (n+1) (na'::names) b)
| _ ->
rigid := is_rigid_head t;
let names = List.rev names in
@@ -252,7 +253,7 @@ let compute_implicits_gen strict strongly_strict revpat contextual all env t =
match kind_of_term (whd_betadeltaiota env t) with
| Prod (na,a,b) ->
let na',avoid = find_displayed_name_in all [] na ([],b) in
- let v = aux (push_rel (na',None,a) env) avoid 1 [na'] b in
+ let v = aux (push_rel (LocalAssum (na',a)) env) avoid 1 [na'] b in
!rigid, Array.to_list v
| _ -> true, []
@@ -427,7 +428,7 @@ let compute_mib_implicits flags manual kn =
(Array.mapi (* No need to lift, arities contain no de Bruijn *)
(fun i mip ->
(** No need to care about constraints here *)
- (Name mip.mind_typename, None, Global.type_of_global_unsafe (IndRef (kn,i))))
+ Context.Rel.Declaration.LocalAssum (Name mip.mind_typename, Global.type_of_global_unsafe (IndRef (kn,i))))
mib.mind_packets) in
let env_ar = push_rel_context ar env in
let imps_one_inductive i mip =
@@ -449,8 +450,8 @@ let compute_all_mib_implicits flags manual kn =
let compute_var_implicits flags manual id =
let env = Global.env () in
- let (_,_,ty) = lookup_named id env in
- compute_semi_auto_implicits env flags manual ty
+ let open Context.Named.Declaration in
+ compute_semi_auto_implicits env flags manual (get_type (lookup_named id env))
(* Implicits of a global reference. *)
diff --git a/library/lib.ml b/library/lib.ml
index e4617cafb..f8bb6bac5 100644
--- a/library/lib.ml
+++ b/library/lib.ml
@@ -428,8 +428,10 @@ let add_section_context ctx =
sectab := (Context ctx :: vars,repl,abs)::sl
let extract_hyps (secs,ohyps) =
+ let open Context.Named.Declaration in
let rec aux = function
- | (Variable (id,impl,poly,ctx)::idl,(id',b,t)::hyps) when Names.Id.equal id id' ->
+ | (Variable (id,impl,poly,ctx)::idl, decl::hyps) when Names.Id.equal id (get_id decl) ->
+ let (id',b,t) = to_tuple decl in
let l, r = aux (idl,hyps) in
(id',impl,b,t) :: l, if poly then Univ.ContextSet.union r ctx else r
| (Variable (_,_,poly,ctx)::idl,hyps) ->
@@ -448,7 +450,10 @@ let instance_from_variable_context sign =
| [] -> [] in
Array.of_list (inst_rec sign)
-let named_of_variable_context ctx = List.map (fun (id,_,b,t) -> (id,b,t)) ctx
+let named_of_variable_context ctx = let open Context.Named.Declaration in
+ List.map (function id,_,None,t -> LocalAssum (id,t)
+ | id,_,Some b,t -> LocalDef (id,b,t))
+ ctx
let add_section_replacement f g poly hyps =
match !sectab with
diff --git a/plugins/cc/ccalgo.ml b/plugins/cc/ccalgo.ml
index bc3d9ed56..359157a4c 100644
--- a/plugins/cc/ccalgo.ml
+++ b/plugins/cc/ccalgo.ml
@@ -824,7 +824,7 @@ let __eps__ = Id.of_string "_eps_"
let new_state_var typ state =
let id = pf_get_new_id __eps__ state.gls in
let {it=gl ; sigma=sigma} = state.gls in
- let gls = Goal.V82.new_goal_with sigma gl [id,None,typ] in
+ let gls = Goal.V82.new_goal_with sigma gl [Context.Named.Declaration.LocalAssum (id,typ)] in
state.gls<- gls;
id
diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml
index 09d9cf019..e5b68af8e 100644
--- a/plugins/cc/cctac.ml
+++ b/plugins/cc/cctac.ml
@@ -23,6 +23,7 @@ open Pp
open Errors
open Util
open Proofview.Notations
+open Context.Rel.Declaration
let reference dir s = lazy (Coqlib.gen_reference "CC" dir s)
@@ -152,7 +153,7 @@ let rec quantified_atom_of_constr env sigma nrels term =
let patts=patterns_of_constr env sigma nrels atom in
`Nrule patts
else
- quantified_atom_of_constr (Environ.push_rel (id,None,atom) env) sigma (succ nrels) ff
+ quantified_atom_of_constr (Environ.push_rel (LocalAssum (id,atom)) env) sigma (succ nrels) ff
| _ ->
let patts=patterns_of_constr env sigma nrels term in
`Rule patts
@@ -167,7 +168,7 @@ let litteral_of_constr env sigma term=
else
begin
try
- quantified_atom_of_constr (Environ.push_rel (id,None,atom) env) sigma 1 ff
+ quantified_atom_of_constr (Environ.push_rel (LocalAssum (id,atom)) env) sigma 1 ff
with Not_found ->
`Other (decompose_term env sigma term)
end
@@ -188,7 +189,8 @@ let make_prb gls depth additionnal_terms =
let t = decompose_term env sigma c in
ignore (add_term state t)) additionnal_terms;
List.iter
- (fun (id,_,e) ->
+ (fun decl ->
+ let (id,_,e) = Context.Named.Declaration.to_tuple decl in
begin
let cid=mkVar id in
match litteral_of_constr env sigma e with
diff --git a/plugins/decl_mode/decl_interp.ml b/plugins/decl_mode/decl_interp.ml
index 2a44dca21..7cfca53c5 100644
--- a/plugins/decl_mode/decl_interp.ml
+++ b/plugins/decl_mode/decl_interp.ml
@@ -403,7 +403,7 @@ let interp_suffices_clause env sigma (hyps,cot)=
match hyp with
(Hprop st | Hvar st) ->
match st.st_label with
- Name id -> Environ.push_named (id,None,st.st_it) env0
+ Name id -> Environ.push_named (Context.Named.Declaration.LocalAssum (id,st.st_it)) env0
| _ -> env in
let nenv = List.fold_right push_one locvars env in
nenv,res
diff --git a/plugins/decl_mode/decl_mode.ml b/plugins/decl_mode/decl_mode.ml
index acee3d6c2..f9399d682 100644
--- a/plugins/decl_mode/decl_mode.ml
+++ b/plugins/decl_mode/decl_mode.ml
@@ -116,7 +116,7 @@ let get_top_stack pts =
let get_stack pts = Proof.get_at_focus proof_focus pts
let get_last env = match Environ.named_context env with
- | (id,_,_)::_ -> id
+ | decl :: _ -> Context.Named.Declaration.get_id decl
| [] -> error "no previous statement to use"
diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml
index f47b35541..22a646b3f 100644
--- a/plugins/decl_mode/decl_proof_instr.ml
+++ b/plugins/decl_mode/decl_proof_instr.ml
@@ -30,6 +30,7 @@ open Namegen
open Goptions
open Misctypes
open Sigma.Notations
+open Context.Named.Declaration
(* Strictness option *)
@@ -229,7 +230,8 @@ let close_previous_case pts =
(* automation *)
let filter_hyps f gls =
- let filter_aux (id,_,_) =
+ let filter_aux id =
+ let id = get_id id in
if f id then
tclIDTAC
else
@@ -331,11 +333,12 @@ let enstack_subsubgoals env se stack gls=
let rc,_ = Reduction.dest_prod env apptype in
let rec meta_aux last lenv = function
[] -> (last,lenv,[])
- | (nam,_,typ)::q ->
+ | decl::q ->
let nlast=succ last in
let (llast,holes,metas) =
meta_aux nlast (mkMeta nlast :: lenv) q in
- (llast,holes,(nlast,special_nf gls (substl lenv typ))::metas) in
+ let open Context.Rel.Declaration in
+ (llast,holes,(nlast,special_nf gls (substl lenv (get_type decl)))::metas) in
let (nlast,holes,nmetas) =
meta_aux se.se_last_meta [] (List.rev rc) in
let refiner = applist (appterm,List.rev holes) in
@@ -411,7 +414,7 @@ let concl_refiner metas body gls =
let _A = subst_meta subst typ in
let x = id_of_name_using_hdchar env _A Anonymous in
let _x = fresh_id avoid x gls in
- let nenv = Environ.push_named (_x,None,_A) env in
+ let nenv = Environ.push_named (LocalAssum (_x,_A)) env in
let asort = family_of_sort (Typing.sort_of nenv (ref evd) _A) in
let nsubst = (n,mkVar _x)::subst in
if List.is_empty rest then
@@ -606,7 +609,7 @@ let assume_tac hyps gls =
tclTHEN
(push_intro_tac
(fun id ->
- Proofview.V82.of_tactic (convert_hyp (id,None,st.st_it))) st.st_label))
+ Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,st.st_it)))) st.st_label))
hyps tclIDTAC gls
let assume_hyps_or_theses hyps gls =
@@ -616,7 +619,7 @@ let assume_hyps_or_theses hyps gls =
tclTHEN
(push_intro_tac
(fun id ->
- Proofview.V82.of_tactic (convert_hyp (id,None,c))) nam)
+ Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,c)))) nam)
| Hprop {st_label=nam;st_it=Thesis (tk)} ->
tclTHEN
(push_intro_tac
@@ -628,7 +631,7 @@ let assume_st hyps gls =
(fun st ->
tclTHEN
(push_intro_tac
- (fun id -> Proofview.V82.of_tactic (convert_hyp (id,None,st.st_it))) st.st_label))
+ (fun id -> Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,st.st_it)))) st.st_label))
hyps tclIDTAC gls
let assume_st_letin hyps gls =
@@ -637,7 +640,7 @@ let assume_st_letin hyps gls =
tclTHEN
(push_intro_tac
(fun id ->
- Proofview.V82.of_tactic (convert_hyp (id,Some (fst st.st_it),snd st.st_it))) st.st_label))
+ Proofview.V82.of_tactic (convert_hyp (LocalDef (id, fst st.st_it, snd st.st_it)))) st.st_label))
hyps tclIDTAC gls
(* suffices *)
@@ -731,7 +734,7 @@ let rec consider_match may_intro introduced available expected gls =
error "Not enough sub-hypotheses to match statements."
(* should tell which ones *)
| id::rest_ids,(Hvar st | Hprop st)::rest ->
- tclIFTHENELSE (Proofview.V82.of_tactic (convert_hyp (id,None,st.st_it)))
+ tclIFTHENELSE (Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,st.st_it))))
begin
match st.st_label with
Anonymous ->
@@ -799,8 +802,8 @@ let define_tac id args body gls =
let cast_tac id_or_thesis typ gls =
match id_or_thesis with
This id ->
- let (_,body,_) = pf_get_hyp gls id in
- Proofview.V82.of_tactic (convert_hyp (id,body,typ)) gls
+ let body = pf_get_hyp gls id |> get_value in
+ Proofview.V82.of_tactic (convert_hyp (of_tuple (id,body,typ))) gls
| Thesis (For _ ) ->
error "\"thesis for ...\" is not applicable here."
| Thesis Plain ->
diff --git a/plugins/derive/derive.ml b/plugins/derive/derive.ml
index a47dc5b2f..5d1551106 100644
--- a/plugins/derive/derive.ml
+++ b/plugins/derive/derive.ml
@@ -6,6 +6,8 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Context.Named.Declaration
+
let map_const_entry_body (f:Term.constr->Term.constr) (x:Safe_typing.private_constants Entries.const_entry_body)
: Safe_typing.private_constants Entries.const_entry_body =
Future.chain ~pure:true x begin fun ((b,ctx),fx) ->
@@ -32,7 +34,7 @@ let start_deriving f suchthat lemma =
let open Proofview in
TCons ( env , sigma , f_type_type , (fun sigma f_type ->
TCons ( env , sigma , f_type , (fun sigma ef ->
- let env' = Environ.push_named (f , (Some ef) , f_type) env in
+ let env' = Environ.push_named (LocalDef (f, ef, f_type)) env in
let evdref = ref sigma in
let suchthat = Constrintern.interp_type_evars env' evdref suchthat in
TCons ( env' , !evdref , suchthat , (fun sigma _ ->
diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml
index 38aef62e1..6c57bc2bb 100644
--- a/plugins/extraction/extraction.ml
+++ b/plugins/extraction/extraction.ml
@@ -25,6 +25,7 @@ open Globnames
open Miniml
open Table
open Mlutil
+open Context.Rel.Declaration
(*i*)
exception I of inductive_kind
@@ -74,7 +75,7 @@ type flag = info * scheme
let rec flag_of_type env t : flag =
let t = whd_betadeltaiota env none t in
match kind_of_term t with
- | Prod (x,t,c) -> flag_of_type (push_rel (x,None,t) env) c
+ | Prod (x,t,c) -> flag_of_type (push_rel (LocalAssum (x,t)) env) c
| Sort s when Sorts.is_prop s -> (Logic,TypeScheme)
| Sort _ -> (Info,TypeScheme)
| _ -> if (sort_of env t) == InProp then (Logic,Default) else (Info,Default)
@@ -247,7 +248,7 @@ let rec extract_type env db j c args =
| _ when sort_of env (applist (c, args)) == InProp -> Tdummy Kprop
| Rel n ->
(match lookup_rel n env with
- | (_,Some t,_) -> extract_type env db j (lift n t) args
+ | LocalDef (_,t,_) -> extract_type env db j (lift n t) args
| _ ->
(* Asks [db] a translation for [n]. *)
if n > List.length db then Tunknown
@@ -560,7 +561,7 @@ let rec extract_term env mle mlt c args =
put_magic_if magic (MLlam (id, d')))
| LetIn (n, c1, t1, c2) ->
let id = id_of_name n in
- let env' = push_rel (Name id, Some c1, t1) env in
+ let env' = push_rel (LocalDef (Name id, c1, t1)) env in
(* We directly push the args inside the [LetIn].
TODO: the opt_let_app flag is supposed to prevent that *)
let args' = List.map (lift 1) args in
@@ -835,7 +836,7 @@ and extract_fix env mle i (fi,ti,ci as recd) mlt =
let decomp_lams_eta_n n m env c t =
let rels = fst (splay_prod_n env none n t) in
- let rels = List.map (fun (id,_,c) -> (id,c)) rels in
+ let rels = List.map (fun (LocalAssum (id,c) | LocalDef (id,_,c)) -> (id,c)) rels in
let rels',c = decompose_lam c in
let d = n - m in
(* we'd better keep rels' as long as possible. *)
diff --git a/plugins/firstorder/formula.ml b/plugins/firstorder/formula.ml
index ae2d059fa..2ed436c6b 100644
--- a/plugins/firstorder/formula.ml
+++ b/plugins/firstorder/formula.ml
@@ -15,6 +15,7 @@ open Tacmach
open Util
open Declarations
open Globnames
+open Context.Rel.Declaration
let qflag=ref true
@@ -139,8 +140,8 @@ let build_atoms gl metagen side cciterm =
negative:= unsigned :: !negative
end;
let v = ind_hyps 0 i l gl in
- let g i _ (_,_,t) =
- build_rec env polarity (lift i t) in
+ let g i _ decl =
+ build_rec env polarity (lift i (get_type decl)) in
let f l =
List.fold_left_i g (1-(List.length l)) () l in
if polarity && (* we have a constant constructor *)
@@ -150,8 +151,8 @@ let build_atoms gl metagen side cciterm =
| Exists(i,l)->
let var=mkMeta (metagen true) in
let v =(ind_hyps 1 i l gl).(0) in
- let g i _ (_,_,t) =
- build_rec (var::env) polarity (lift i t) in
+ let g i _ decl =
+ build_rec (var::env) polarity (lift i (get_type decl)) in
List.fold_left_i g (2-(List.length l)) () v
| Forall(_,b)->
let var=mkMeta (metagen true) in
@@ -224,7 +225,7 @@ let build_formula side nam typ gl metagen=
| And(_,_,_) -> Rand
| Or(_,_,_) -> Ror
| Exists (i,l) ->
- let (_,_,d)=List.last (ind_hyps 0 i l gl).(0) in
+ let d = get_type (List.last (ind_hyps 0 i l gl).(0)) in
Rexists(m,d,trivial)
| Forall (_,a) -> Rforall
| Arrow (a,b) -> Rarrow in
diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml
index a717cc91e..797f43f2a 100644
--- a/plugins/firstorder/instances.ml
+++ b/plugins/firstorder/instances.ml
@@ -22,6 +22,7 @@ open Formula
open Sequent
open Names
open Misctypes
+open Context.Rel.Declaration
let compare_instance inst1 inst2=
match inst1,inst2 with
@@ -117,7 +118,7 @@ let mk_open_instance id idc gl m t=
if Int.equal n 0 then evmap, decls else
let nid=(fresh_id avoid var_id gl) in
let evmap, (c, _) = Evarutil.new_type_evar env evmap Evd.univ_flexible in
- let decl = (Name nid,None,c) in
+ let decl = LocalAssum (Name nid,c) in
aux (n-1) (nid::avoid) (Environ.push_rel decl env) evmap (decl::decls) in
let evmap, decls = aux m [] env evmap [] in
evmap, decls, revt
diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml
index e676a8a93..b0e8f7d25 100644
--- a/plugins/firstorder/rules.ml
+++ b/plugins/firstorder/rules.ml
@@ -19,6 +19,7 @@ open Formula
open Sequent
open Globnames
open Locus
+open Context.Named.Declaration
type seqtac= (Sequent.t -> tactic) -> Sequent.t -> tactic
@@ -34,12 +35,13 @@ let wrap n b continue seq gls=
if i<=0 then seq else
match nc with
[]->anomaly (Pp.str "Not the expected number of hyps")
- | ((id,_,typ) as nd)::q->
+ | nd::q->
+ let id = get_id nd in
if occur_var env id (pf_concl gls) ||
List.exists (occur_var_in_decl env id) ctx then
(aux (i-1) q (nd::ctx))
else
- add_formula Hyp (VarRef id) typ (aux (i-1) q (nd::ctx)) gls in
+ add_formula Hyp (VarRef id) (get_type nd) (aux (i-1) q (nd::ctx)) gls in
let seq1=aux n nc [] in
let seq2=if b then
add_formula Concl dummy_id (pf_concl gls) seq1 gls else seq1 in
diff --git a/plugins/fourier/fourierR.ml b/plugins/fourier/fourierR.ml
index 4c0aa6c75..8bc84608e 100644
--- a/plugins/fourier/fourierR.ml
+++ b/plugins/fourier/fourierR.ml
@@ -445,7 +445,11 @@ let is_ineq (h,t) =
;;
*)
-let list_of_sign s = List.map (fun (x,_,z)->(x,z)) s;;
+let list_of_sign s =
+ let open Context.Named.Declaration in
+ List.map (function LocalAssum (name, typ) -> name, typ
+ | LocalDef (name, _, typ) -> name, typ)
+ s;;
let mkAppL a =
let l = Array.to_list a in
diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml
index 4eab5f912..28982c37f 100644
--- a/plugins/funind/functional_principles_proofs.ml
+++ b/plugins/funind/functional_principles_proofs.ml
@@ -15,6 +15,7 @@ open Tactics
open Indfun_common
open Libnames
open Globnames
+open Context.Rel.Declaration
(* let msgnl = Pp.msgnl *)
@@ -304,11 +305,11 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type =
in
let new_type_of_hyp,ctxt_size,witness_fun =
List.fold_left_i
- (fun i (end_of_type,ctxt_size,witness_fun) ((x',b',t') as decl) ->
+ (fun i (end_of_type,ctxt_size,witness_fun) decl ->
try
let witness = Int.Map.find i sub in
- if not (Option.is_empty b') then anomaly (Pp.str "can not redefine a rel!");
- (Termops.pop end_of_type,ctxt_size,mkLetIn(x',witness,t',witness_fun))
+ if is_local_def decl then anomaly (Pp.str "can not redefine a rel!");
+ (Termops.pop end_of_type,ctxt_size,mkLetIn (get_name decl, witness, get_type decl, witness_fun))
with Not_found ->
(mkProd_or_LetIn decl end_of_type, ctxt_size + 1, mkLambda_or_LetIn decl witness_fun)
)
@@ -536,7 +537,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma =
(scan_type new_context new_t')
with Failure "NoChange" ->
(* Last thing todo : push the rel in the context and continue *)
- scan_type ((x,None,t_x)::context) t'
+ scan_type (LocalAssum (x,t_x) :: context) t'
end
end
else
@@ -736,7 +737,8 @@ let build_proof
tclTHEN
(Proofview.V82.of_tactic intro)
(fun g' ->
- let (id,_,_) = pf_last_hyp g' in
+ let open Context.Named.Declaration in
+ let id = pf_last_hyp g' |> get_id in
let new_term =
pf_nf_betaiota g'
(mkApp(dyn_infos.info,[|mkVar id|]))
@@ -921,7 +923,9 @@ let generalize_non_dep hyp g =
let env = Global.env () in
let hyp_typ = pf_unsafe_type_of g (mkVar hyp) in
let to_revert,_ =
- Environ.fold_named_context_reverse (fun (clear,keep) (hyp,_,_ as decl) ->
+ let open Context.Named.Declaration in
+ Environ.fold_named_context_reverse (fun (clear,keep) decl ->
+ let hyp = get_id decl in
if Id.List.mem hyp hyps
|| List.exists (Termops.occur_var_in_decl env hyp) keep
|| Termops.occur_var env hyp hyp_typ
@@ -936,7 +940,7 @@ let generalize_non_dep hyp g =
((* observe_tac "thin" *) (thin to_revert))
g
-let id_of_decl (na,_,_) = (Nameops.out_name na)
+let id_of_decl decl = Nameops.out_name (get_name decl)
let var_of_decl decl = mkVar (id_of_decl decl)
let revert idl =
tclTHEN
@@ -1044,7 +1048,8 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a
(
fun g' ->
let just_introduced = nLastDecls nb_intro_to_do g' in
- let just_introduced_id = List.map (fun (id,_,_) -> id) just_introduced in
+ let open Context.Named.Declaration in
+ let just_introduced_id = List.map get_id just_introduced in
tclTHEN (Proofview.V82.of_tactic (Equality.rewriteLR equation_lemma))
(revert just_introduced_id) g'
)
@@ -1069,11 +1074,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
(Name new_id)
)
in
- let fresh_decl =
- (fun (na,b,t) ->
- (fresh_id na,b,t)
- )
- in
+ let fresh_decl = map_name fresh_id in
let princ_info : elim_scheme =
{ princ_info with
params = List.map fresh_decl princ_info.params;
@@ -1120,11 +1121,11 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
)
in
observe (str "full_params := " ++
- prlist_with_sep spc (fun (na,_,_) -> Ppconstr.pr_id (Nameops.out_name na))
+ prlist_with_sep spc (fun decl -> Ppconstr.pr_id (Nameops.out_name (get_name decl)))
full_params
);
observe (str "princ_params := " ++
- prlist_with_sep spc (fun (na,_,_) -> Ppconstr.pr_id (Nameops.out_name na))
+ prlist_with_sep spc (fun decl -> Ppconstr.pr_id (Nameops.out_name (get_name decl)))
princ_params
);
observe (str "fbody_with_full_params := " ++
@@ -1165,7 +1166,8 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
in
let pte_to_fix,rev_info =
List.fold_left_i
- (fun i (acc_map,acc_info) (pte,_,_) ->
+ (fun i (acc_map,acc_info) decl ->
+ let pte = get_name decl in
let infos = info_array.(i) in
let type_args,_ = decompose_prod infos.types in
let nargs = List.length type_args in
@@ -1259,7 +1261,8 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
let args = nLastDecls nb_args g in
let fix_body = fix_info.body_with_param in
(* observe (str "fix_body := "++ pr_lconstr_env (pf_env gl) fix_body); *)
- let args_id = List.map (fun (id,_,_) -> id) args in
+ let open Context.Named.Declaration in
+ let args_id = List.map get_id args in
let dyn_infos =
{
nb_rec_hyps = -100;
@@ -1276,7 +1279,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
(do_replace evd
full_params
(fix_info.idx + List.length princ_params)
- (args_id@(List.map (fun (id,_,_) -> Nameops.out_name id ) princ_params))
+ (args_id@(List.map (fun decl -> Nameops.out_name (get_name decl)) princ_params))
(all_funs.(fix_info.num_in_block))
fix_info.num_in_block
all_funs
@@ -1317,8 +1320,9 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
[
tclDO nb_args (Proofview.V82.of_tactic intro);
(fun g -> (* replacement of the function by its body *)
- let args = nLastDecls nb_args g in
- let args_id = List.map (fun (id,_,_) -> id) args in
+ let args = nLastDecls nb_args g in
+ let open Context.Named.Declaration in
+ let args_id = List.map get_id args in
let dyn_infos =
{
nb_rec_hyps = -100;
@@ -1520,7 +1524,7 @@ let prove_principle_for_gen
avoid := new_id :: !avoid;
Name new_id
in
- let fresh_decl (na,b,t) = (fresh_id na,b,t) in
+ let fresh_decl = map_name fresh_id in
let princ_info : elim_scheme =
{ princ_info with
params = List.map fresh_decl princ_info.params;
@@ -1550,11 +1554,11 @@ let prove_principle_for_gen
in
let rec_arg_id =
match List.rev post_rec_arg with
- | (Name id,_,_)::_ -> id
+ | (LocalAssum (Name id,_) | LocalDef (Name id,_,_)) :: _ -> id
| _ -> assert false
in
(* observe (str "rec_arg_id := " ++ pr_lconstr (mkVar rec_arg_id)); *)
- let subst_constrs = List.map (fun (na,_,_) -> mkVar (Nameops.out_name na)) (pre_rec_arg@princ_info.params) in
+ let subst_constrs = List.map (fun decl -> mkVar (Nameops.out_name (get_name decl))) (pre_rec_arg@princ_info.params) in
let relation = substl subst_constrs relation in
let input_type = substl subst_constrs rec_arg_type in
let wf_thm_id = Nameops.out_name (fresh_id (Name (Id.of_string "wf_R"))) in
@@ -1582,7 +1586,7 @@ let prove_principle_for_gen
)
g
in
- let args_ids = List.map (fun (na,_,_) -> Nameops.out_name na) princ_info.args in
+ let args_ids = List.map (fun decl -> Nameops.out_name (get_name decl)) princ_info.args in
let lemma =
match !tcc_lemma_ref with
| None -> error "No tcc proof !!"
@@ -1629,7 +1633,7 @@ let prove_principle_for_gen
[
observe_tac "start_tac" start_tac;
h_intros
- (List.rev_map (fun (na,_,_) -> Nameops.out_name na)
+ (List.rev_map (fun decl -> Nameops.out_name (get_name decl))
(princ_info.args@princ_info.branches@princ_info.predicates@princ_info.params)
);
(* observe_tac "" *) Proofview.V82.of_tactic (assert_by
@@ -1667,7 +1671,7 @@ let prove_principle_for_gen
in
let acc_inv = lazy (mkApp(Lazy.force acc_inv, [|mkVar acc_rec_arg_id|])) in
let predicates_names =
- List.map (fun (na,_,_) -> Nameops.out_name na) princ_info.predicates
+ List.map (fun decl -> Nameops.out_name (get_name decl)) princ_info.predicates
in
let pte_info =
{ proving_tac =
@@ -1683,7 +1687,7 @@ let prove_principle_for_gen
is_mes acc_inv fix_id
(!tcc_list@(List.map
- (fun (na,_,_) -> (Nameops.out_name na))
+ (fun decl -> (Nameops.out_name (get_name decl)))
(princ_info.args@princ_info.params)
)@ ([acc_rec_arg_id])) eqs
)
@@ -1712,7 +1716,7 @@ let prove_principle_for_gen
(* observe_tac "instanciate_hyps_with_args" *)
(instanciate_hyps_with_args
make_proof
- (List.map (fun (na,_,_) -> Nameops.out_name na) princ_info.branches)
+ (List.map (fun decl -> Nameops.out_name (get_name decl)) princ_info.branches)
(List.rev args_ids)
)
gl'
diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml
index e2c3bbb97..7a933c04b 100644
--- a/plugins/funind/functional_principles_types.ml
+++ b/plugins/funind/functional_principles_types.ml
@@ -8,6 +8,7 @@ open Names
open Pp
open Entries
open Tactics
+open Context.Rel.Declaration
open Indfun_common
open Functional_principles_proofs
open Misctypes
@@ -32,11 +33,13 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
let rec change_predicates_names (avoid:Id.t list) (predicates:Context.Rel.t) : Context.Rel.t =
match predicates with
| [] -> []
- |(Name x,v,t)::predicates ->
- let id = Namegen.next_ident_away x avoid in
- Hashtbl.add tbl id x;
- (Name id,v,t)::(change_predicates_names (id::avoid) predicates)
- | (Anonymous,_,_)::_ -> anomaly (Pp.str "Anonymous property binder ")
+ | decl :: predicates ->
+ (match Context.Rel.Declaration.get_name decl with
+ | Name x ->
+ let id = Namegen.next_ident_away x avoid in
+ Hashtbl.add tbl id x;
+ set_name (Name id) decl :: change_predicates_names (id::avoid) predicates
+ | Anonymous -> anomaly (Pp.str "Anonymous property binder "))
in
let avoid = (Termops.ids_of_context env_with_params ) in
let princ_type_info =
@@ -46,15 +49,16 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
in
(* observe (str "starting princ_type := " ++ pr_lconstr_env env princ_type); *)
(* observe (str "princ_infos : " ++ pr_elim_scheme princ_type_info); *)
- let change_predicate_sort i (x,_,t) =
+ let change_predicate_sort i decl =
let new_sort = sorts.(i) in
- let args,_ = decompose_prod t in
+ let args,_ = decompose_prod (get_type decl) in
let real_args =
if princ_type_info.indarg_in_concl
then List.tl args
else args
in
- Nameops.out_name x,None,compose_prod real_args (mkSort new_sort)
+ Context.Named.Declaration.LocalAssum (Nameops.out_name (Context.Rel.Declaration.get_name decl),
+ compose_prod real_args (mkSort new_sort))
in
let new_predicates =
List.map_i
@@ -69,7 +73,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
| _ -> error "Not a valid predicate"
)
in
- let ptes_vars = List.map (fun (id,_,_) -> id) new_predicates in
+ let ptes_vars = List.map Context.Named.Declaration.get_id new_predicates in
let is_pte =
let set = List.fold_right Id.Set.add ptes_vars Id.Set.empty in
fun t ->
@@ -114,7 +118,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
| Rel n ->
begin
try match Environ.lookup_rel n env with
- | _,_,t when is_dom t -> raise Toberemoved
+ | LocalAssum (_,t) | LocalDef (_,_,t) when is_dom t -> raise Toberemoved
| _ -> pre_princ,[]
with Not_found -> assert false
end
@@ -159,7 +163,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
try
let new_t,binders_to_remove_from_t = compute_new_princ_type remove env t in
let new_x : Name.t = get_name (Termops.ids_of_context env) x in
- let new_env = Environ.push_rel (x,None,t) env in
+ let new_env = Environ.push_rel (LocalAssum (x,t)) env in
let new_b,binders_to_remove_from_b = compute_new_princ_type remove new_env b in
if List.exists (eq_constr (mkRel 1)) binders_to_remove_from_b
then (Termops.pop new_b), filter_map (eq_constr (mkRel 1)) Termops.pop binders_to_remove_from_b
@@ -188,7 +192,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
let new_t,binders_to_remove_from_t = compute_new_princ_type remove env t in
let new_v,binders_to_remove_from_v = compute_new_princ_type remove env v in
let new_x : Name.t = get_name (Termops.ids_of_context env) x in
- let new_env = Environ.push_rel (x,Some v,t) env in
+ let new_env = Environ.push_rel (LocalDef (x,v,t)) env in
let new_b,binders_to_remove_from_b = compute_new_princ_type remove new_env b in
if List.exists (eq_constr (mkRel 1)) binders_to_remove_from_b
then (Termops.pop new_b),filter_map (eq_constr (mkRel 1)) Termops.pop binders_to_remove_from_b
@@ -227,7 +231,8 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
in
it_mkProd_or_LetIn
(it_mkProd_or_LetIn
- pre_res (List.map (fun (id,t,b) -> Name(Hashtbl.find tbl id), t,b)
+ pre_res (List.map (function Context.Named.Declaration.LocalAssum (id,b) -> LocalAssum (Name (Hashtbl.find tbl id), b)
+ | Context.Named.Declaration.LocalDef (id,t,b) -> LocalDef (Name (Hashtbl.find tbl id), t, b))
new_predicates)
)
princ_type_info.params
@@ -235,10 +240,12 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
let change_property_sort evd toSort princ princName =
+ let open Context.Rel.Declaration in
let princ_info = compute_elim_sig princ in
- let change_sort_in_predicate (x,v,t) =
- (x,None,
- let args,ty = decompose_prod t in
+ let change_sort_in_predicate decl =
+ LocalAssum
+ (get_name decl,
+ let args,ty = decompose_prod (get_type decl) in
let s = destSort ty in
Global.add_constraints (Univ.enforce_leq (univ_of_sort toSort) (univ_of_sort s) Univ.Constraint.empty);
compose_prod args (mkSort toSort)
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index 80de8e764..8a0a1a064 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -335,15 +335,17 @@ let raw_push_named (na,raw_value,raw_typ) env =
| Name id ->
let value = Option.map (fun x-> fst (Pretyping.understand env (Evd.from_env env) x)) raw_value in
let typ,ctx = Pretyping.understand env (Evd.from_env env) ~expected_type:Pretyping.IsType raw_typ in
- Environ.push_named (id,value,typ) env
+ let open Context.Named.Declaration in
+ Environ.push_named (of_tuple (id,value,typ)) env
let add_pat_variables pat typ env : Environ.env =
let rec add_pat_variables env pat typ : Environ.env =
+ let open Context.Rel.Declaration in
observe (str "new rel env := " ++ Printer.pr_rel_context_of env (Evd.from_env env));
match pat with
- | PatVar(_,na) -> Environ.push_rel (na,None,typ) env
+ | PatVar(_,na) -> Environ.push_rel (LocalAssum (na,typ)) env
| PatCstr(_,c,patl,na) ->
let Inductiveops.IndType(indf,indargs) =
try Inductiveops.find_rectype env (Evd.from_env env) typ
@@ -351,15 +353,16 @@ let add_pat_variables pat typ env : Environ.env =
in
let constructors = Inductiveops.get_constructors env indf in
let constructor : Inductiveops.constructor_summary = List.find (fun cs -> eq_constructor c (fst cs.Inductiveops.cs_cstr)) (Array.to_list constructors) in
- let cs_args_types :types list = List.map (fun (_,_,t) -> t) constructor.Inductiveops.cs_args in
+ let cs_args_types :types list = List.map get_type constructor.Inductiveops.cs_args in
List.fold_left2 add_pat_variables env patl (List.rev cs_args_types)
in
let new_env = add_pat_variables env pat typ in
let res =
fst (
Context.Rel.fold_outside
- (fun (na,v,t) (env,ctxt) ->
- match na with
+ (fun decl (env,ctxt) ->
+ let _,v,t = Context.Rel.Declaration.to_tuple decl in
+ match Context.Rel.Declaration.get_name decl with
| Anonymous -> assert false
| Name id ->
let new_t = substl ctxt t in
@@ -370,7 +373,8 @@ let add_pat_variables pat typ env : Environ.env =
Option.fold_right (fun v _ -> str "old value := " ++ Printer.pr_lconstr v ++ fnl ()) v (mt ()) ++
Option.fold_right (fun v _ -> str "new value := " ++ Printer.pr_lconstr v ++ fnl ()) new_v (mt ())
);
- (Environ.push_named (id,new_v,new_t) env,mkVar id::ctxt)
+ let open Context.Named.Declaration in
+ (Environ.push_named (of_tuple (id,new_v,new_t)) env,mkVar id::ctxt)
)
(Environ.rel_context new_env)
~init:(env,[])
@@ -398,7 +402,8 @@ let rec pattern_to_term_and_type env typ = function
in
let constructors = Inductiveops.get_constructors env indf in
let constructor = List.find (fun cs -> eq_constructor (fst cs.Inductiveops.cs_cstr) constr) (Array.to_list constructors) in
- let cs_args_types :types list = List.map (fun (_,_,t) -> t) constructor.Inductiveops.cs_args in
+ let open Context.Rel.Declaration in
+ let cs_args_types :types list = List.map get_type constructor.Inductiveops.cs_args in
let _,cstl = Inductiveops.dest_ind_family indf in
let csta = Array.of_list cstl in
let implicit_args =
@@ -597,9 +602,10 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
let v_as_constr,ctx = Pretyping.understand env (Evd.from_env env) v in
let v_type = Typing.unsafe_type_of env (Evd.from_env env) v_as_constr in
let new_env =
+ let open Context.Named.Declaration in
match n with
Anonymous -> env
- | Name id -> Environ.push_named (id,Some v_as_constr,v_type) env
+ | Name id -> Environ.push_named (of_tuple (id,Some v_as_constr,v_type)) env
in
let b_res = build_entry_lc new_env funnames avoid b in
combine_results (combine_letin n) v_res b_res
@@ -875,7 +881,7 @@ exception Continue
*)
let rec rebuild_cons env nb_args relname args crossed_types depth rt =
observe (str "rebuilding : " ++ pr_glob_constr rt);
-
+ let open Context.Rel.Declaration in
match rt with
| GProd(_,n,k,t,b) ->
let not_free_in_t id = not (is_free_in id t) in
@@ -895,7 +901,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
mkGApp(mkGVar(mk_rel_id this_relname),args'@[res_rt])
in
let t',ctx = Pretyping.understand env (Evd.from_env env) new_t in
- let new_env = Environ.push_rel (n,None,t') env in
+ let new_env = Environ.push_rel (LocalAssum (n,t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
@@ -926,7 +932,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let subst_b =
if is_in_b then b else replace_var_by_term id rt b
in
- let new_env = Environ.push_rel (n,None,t') env in
+ let new_env = Environ.push_rel (LocalAssum (n,t')) env in
let new_b,id_to_exclude =
rebuild_cons
new_env
@@ -970,9 +976,8 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
(fun acc var_as_constr arg ->
if isRel var_as_constr
then
- let (na,_,_) =
- Environ.lookup_rel (destRel var_as_constr) env
- in
+ let open Context.Rel.Declaration in
+ let na = get_name (Environ.lookup_rel (destRel var_as_constr) env) in
match na with
| Anonymous -> acc
| Name id' ->
@@ -1010,7 +1015,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
in
let new_env =
let t',ctx = Pretyping.understand env (Evd.from_env env) eq' in
- Environ.push_rel (n,None,t') env
+ Environ.push_rel (LocalAssum (n,t')) env
in
let new_b,id_to_exclude =
rebuild_cons
@@ -1048,7 +1053,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
with Continue ->
observe (str "computing new type for prod : " ++ pr_glob_constr rt);
let t',ctx = Pretyping.understand env (Evd.from_env env) t in
- let new_env = Environ.push_rel (n,None,t') env in
+ let new_env = Environ.push_rel (LocalAssum (n,t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
@@ -1064,7 +1069,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
| _ ->
observe (str "computing new type for prod : " ++ pr_glob_constr rt);
let t',ctx = Pretyping.understand env (Evd.from_env env) t in
- let new_env = Environ.push_rel (n,None,t') env in
+ let new_env = Environ.push_rel (LocalAssum (n,t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
@@ -1085,7 +1090,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let t',ctx = Pretyping.understand env (Evd.from_env env) t in
match n with
| Name id ->
- let new_env = Environ.push_rel (n,None,t') env in
+ let new_env = Environ.push_rel (LocalAssum (n,t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
@@ -1108,7 +1113,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let t',ctx = Pretyping.understand env evd t in
let evd = Evd.from_ctx ctx in
let type_t' = Typing.unsafe_type_of env evd t' in
- let new_env = Environ.push_rel (n,Some t',type_t') env in
+ let new_env = Environ.push_rel (LocalDef (n,t',type_t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
@@ -1132,7 +1137,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
depth t
in
let t',ctx = Pretyping.understand env (Evd.from_env env) new_t in
- let new_env = Environ.push_rel (na,None,t') env in
+ let new_env = Environ.push_rel (LocalAssum (na,t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
@@ -1254,12 +1259,13 @@ let do_build_inductive
let relnames = Array.map mk_rel_id funnames in
let relnames_as_set = Array.fold_right Id.Set.add relnames Id.Set.empty in
(* Construction of the pseudo constructors *)
+ let open Context.Named.Declaration in
let evd,env =
Array.fold_right2
(fun id c (evd,env) ->
let evd,t = Typing.type_of env evd (mkConstU c) in
evd,
- Environ.push_named (id,None,t)
+ Environ.push_named (LocalAssum (id,t))
(* try *)
(* Typing.e_type_of env evd (mkConstU c) *)
(* with Not_found -> *)
@@ -1298,8 +1304,8 @@ let do_build_inductive
*)
let rel_arities = Array.mapi rel_arity funsargs in
Util.Array.fold_left2 (fun env rel_name rel_ar ->
- Environ.push_named (rel_name,None,
- fst (with_full_print (Constrintern.interp_constr env evd) rel_ar)) env) env relnames rel_arities
+ Environ.push_named (LocalAssum (rel_name,
+ fst (with_full_print (Constrintern.interp_constr env evd) rel_ar))) env) env relnames rel_arities
in
(* and of the real constructors*)
let constr i res =
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index d1e109825..86abf9e2e 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -1,3 +1,4 @@
+open Context.Rel.Declaration
open Errors
open Util
open Names
@@ -13,10 +14,10 @@ open Decl_kinds
open Sigma.Notations
let is_rec_info scheme_info =
- let test_branche min acc (_,_,br) =
+ let test_branche min acc decl =
acc || (
let new_branche =
- it_mkProd_or_LetIn mkProp (fst (decompose_prod_assum br)) in
+ it_mkProd_or_LetIn mkProp (fst (decompose_prod_assum (get_type decl))) in
let free_rels_in_br = Termops.free_rels new_branche in
let max = min + scheme_info.Tactics.npredicates in
Int.Set.exists (fun i -> i >= min && i< max) free_rels_in_br
@@ -153,7 +154,8 @@ let build_newrecursive
let evdref = ref (Evd.from_env env0) in
let _, (_, impls') = Constrintern.interp_context_evars env evdref bl in
let impl = Constrintern.compute_internalization_data env0 Constrintern.Recursive arity impls' in
- (Environ.push_named (recname,None,arity) env, Id.Map.add recname impl impls))
+ let open Context.Named.Declaration in
+ (Environ.push_named (LocalAssum (recname,arity)) env, Id.Map.add recname impl impls))
(env0,Constrintern.empty_internalization_env) lnameargsardef in
let recdef =
(* Declare local notations *)
diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml
index 0c9d3bb81..56bc4328d 100644
--- a/plugins/funind/invfun.ml
+++ b/plugins/funind/invfun.ml
@@ -5,6 +5,7 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+
open Tacexpr
open Declarations
open Errors
@@ -20,6 +21,7 @@ open Indfun_common
open Tacmach
open Misctypes
open Termops
+open Context.Rel.Declaration
(* Some pretty printing function for debugging purpose *)
@@ -134,18 +136,21 @@ let generate_type evd g_to_f f graph i =
let fun_ctxt,res_type =
match ctxt with
| [] | [_] -> anomaly (Pp.str "Not a valid context")
- | (_,_,res_type)::fun_ctxt -> fun_ctxt,res_type
+ | decl :: fun_ctxt -> fun_ctxt, get_type decl
in
let rec args_from_decl i accu = function
| [] -> accu
- | (_, Some _, _) :: l ->
+ | LocalDef _ :: l ->
args_from_decl (succ i) accu l
| _ :: l ->
let t = mkRel i in
args_from_decl (succ i) (t :: accu) l
in
(*i We need to name the vars [res] and [fv] i*)
- let filter = function (Name id,_,_) -> Some id | (Anonymous,_,_) -> None in
+ let filter = fun decl -> match get_name decl with
+ | Name id -> Some id
+ | Anonymous -> None
+ in
let named_ctxt = List.map_filter filter fun_ctxt in
let res_id = Namegen.next_ident_away_in_goal (Id.of_string "_res") named_ctxt in
let fv_id = Namegen.next_ident_away_in_goal (Id.of_string "fv") (res_id :: named_ctxt) in
@@ -171,12 +176,12 @@ let generate_type evd g_to_f f graph i =
\[\forall (x_1:t_1)\ldots(x_n:t_n), let fv := f x_1\ldots x_n in, forall res, \]
i*)
let pre_ctxt =
- (Name res_id,None,lift 1 res_type)::(Name fv_id,Some (mkApp(f,args_as_rels)),res_type)::fun_ctxt
+ LocalAssum (Name res_id, lift 1 res_type) :: LocalDef (Name fv_id, mkApp (f,args_as_rels), res_type) :: fun_ctxt
in
(*i and we can return the solution depending on which lemma type we are defining i*)
if g_to_f
- then (Anonymous,None,graph_applied)::pre_ctxt,(lift 1 res_eq_f_of_args),graph
- else (Anonymous,None,res_eq_f_of_args)::pre_ctxt,(lift 1 graph_applied),graph
+ then LocalAssum (Anonymous,graph_applied)::pre_ctxt,(lift 1 res_eq_f_of_args),graph
+ else LocalAssum (Anonymous,res_eq_f_of_args)::pre_ctxt,(lift 1 graph_applied),graph
(*
@@ -260,10 +265,10 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
(* and built the intro pattern for each of them *)
let intro_pats =
List.map
- (fun (_,_,br_type) ->
+ (fun decl ->
List.map
(fun id -> Loc.ghost, IntroNaming (IntroIdentifier id))
- (generate_fresh_id (Id.of_string "y") ids (List.length (fst (decompose_prod_assum br_type))))
+ (generate_fresh_id (Id.of_string "y") ids (List.length (fst (decompose_prod_assum (get_type decl)))))
)
branches
in
@@ -390,10 +395,10 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
(fun ((_,(ctxt,concl))) ->
match ctxt with
| [] | [_] | [_;_] -> anomaly (Pp.str "bad context")
- | hres::res::(x,_,t)::ctxt ->
+ | hres::res::decl::ctxt ->
let res = Termops.it_mkLambda_or_LetIn
(Termops.it_mkProd_or_LetIn concl [hres;res])
- ((x,None,t)::ctxt)
+ (LocalAssum (get_name decl, get_type decl) :: ctxt)
in
res
)
@@ -408,8 +413,8 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
let bindings =
let params_bindings,avoid =
List.fold_left2
- (fun (bindings,avoid) (x,_,_) p ->
- let id = Namegen.next_ident_away (Nameops.out_name x) avoid in
+ (fun (bindings,avoid) decl p ->
+ let id = Namegen.next_ident_away (Nameops.out_name (get_name decl)) avoid in
p::bindings,id::avoid
)
([],pf_ids_of_hyps g)
@@ -418,8 +423,8 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
in
let lemmas_bindings =
List.rev (fst (List.fold_left2
- (fun (bindings,avoid) (x,_,_) p ->
- let id = Namegen.next_ident_away (Nameops.out_name x) avoid in
+ (fun (bindings,avoid) decl p ->
+ let id = Namegen.next_ident_away (Nameops.out_name (get_name decl)) avoid in
(nf_zeta p)::bindings,id::avoid)
([],avoid)
princ_infos.predicates
@@ -455,9 +460,10 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
generalize every hypothesis which depends of [x] but [hyp]
*)
let generalize_dependent_of x hyp g =
+ let open Context.Named.Declaration in
tclMAP
(function
- | (id,None,t) when not (Id.equal id hyp) &&
+ | LocalAssum (id,t) when not (Id.equal id hyp) &&
(Termops.occur_var (pf_env g) x t) -> tclTHEN (Tactics.generalize [mkVar id]) (thin [id])
| _ -> tclIDTAC
)
@@ -663,10 +669,10 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
let branches = List.rev princ_infos.branches in
let intro_pats =
List.map
- (fun (_,_,br_type) ->
+ (fun decl ->
List.map
(fun id -> id)
- (generate_fresh_id (Id.of_string "y") ids (nb_prod br_type))
+ (generate_fresh_id (Id.of_string "y") ids (nb_prod (get_type decl)))
)
branches
in
diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml
index 57782dd71..c71d9a9ca 100644
--- a/plugins/funind/merge.ml
+++ b/plugins/funind/merge.ml
@@ -24,6 +24,7 @@ open Declarations
open Glob_term
open Glob_termops
open Decl_kinds
+open Context.Rel.Declaration
(** {1 Utilities} *)
@@ -134,9 +135,9 @@ let showind (id:Id.t) =
let cstrid = Constrintern.global_reference id in
let ind1,cstrlist = Inductiveops.find_inductive (Global.env()) Evd.empty cstrid in
let mib1,ib1 = Inductive.lookup_mind_specif (Global.env()) (fst ind1) in
- List.iter (fun (nm, optcstr, tp) ->
- print_string (string_of_name nm^":");
- prconstr tp; print_string "\n")
+ List.iter (fun decl ->
+ print_string (string_of_name (Context.Rel.Declaration.get_name decl) ^ ":");
+ prconstr (get_type decl); print_string "\n")
ib1.mind_arity_ctxt;
Printf.printf "arity :"; prconstr (Inductiveops.type_of_inductive (Global.env ()) ind1);
Array.iteri
@@ -459,11 +460,12 @@ let shift_linked_params mib1 mib2 (lnk1:linked_var array) (lnk2:linked_var array
let recprms2,otherprms2,args2,funresprms2 = bldprms (List.rev oib2.mind_arity_ctxt) mlnk2 in
let _ = prstr "\notherprms1:\n" in
let _ =
- List.iter (fun (x,_,y) -> prstr (string_of_name x^" : ");prconstr y;prstr "\n")
+ List.iter (fun decl -> prstr (string_of_name (get_name decl) ^ " : ");
+ prconstr (get_type decl); prstr "\n")
otherprms1 in
let _ = prstr "\notherprms2:\n" in
let _ =
- List.iter (fun (x,_,y) -> prstr (string_of_name x^" : ");prconstr y;prstr "\n")
+ List.iter (fun decl -> prstr (string_of_name (get_name decl) ^ " : "); prconstr (get_type decl); prstr "\n")
otherprms2 in
{
ident=id;
@@ -823,9 +825,11 @@ let merge_rec_params_and_arity prms1 prms2 shift (concl:constr) =
let concl = Constrextern.extern_constr false (Global.env()) Evd.empty concl in
let arity,_ =
List.fold_left
- (fun (acc,env) (nm,_,c) ->
+ (fun (acc,env) decl ->
+ let nm = Context.Rel.Declaration.get_name decl in
+ let c = get_type decl in
let typ = Constrextern.extern_constr false env Evd.empty c in
- let newenv = Environ.push_rel (nm,None,c) env in
+ let newenv = Environ.push_rel (LocalAssum (nm,c)) env in
CProdN (Loc.ghost, [[(Loc.ghost,nm)],Constrexpr_ops.default_binder_kind,typ] , acc) , newenv)
(concl,Global.env())
(shift.funresprms2 @ shift.funresprms1
@@ -852,10 +856,10 @@ let glob_constr_list_to_inductive_expr prms1 prms2 mib1 mib2 shift
let mkProd_reldecl (rdecl:Context.Rel.Declaration.t) (t2:glob_constr) =
match rdecl with
- | (nme,None,t) ->
+ | LocalAssum (nme,t) ->
let traw = Detyping.detype false [] (Global.env()) Evd.empty t in
GProd (Loc.ghost,nme,Explicit,traw,t2)
- | (_,Some _,_) -> assert false
+ | LocalDef _ -> assert false
(** [merge_inductive ind1 ind2 lnk] merges two graphs, linking
@@ -969,7 +973,7 @@ let funify_branches relinfo nfuns branch =
| Rel i -> let reali = i-shift in (reali>=0 && reali<relinfo.nbranches)
| _ -> false in
(* FIXME: *)
- (Anonymous,Some mkProp,mkProp)
+ LocalDef (Anonymous,mkProp,mkProp)
let relprinctype_to_funprinctype relprinctype nfuns =
diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml
index b09678341..09c5aa567 100644
--- a/plugins/funind/recdef.ml
+++ b/plugins/funind/recdef.ml
@@ -40,7 +40,7 @@ open Eauto
open Indfun_common
open Sigma.Notations
-
+open Context.Rel.Declaration
(* Ugly things which should not be here *)
@@ -181,7 +181,7 @@ let (value_f:constr list -> global_reference -> constr) =
)
in
let context = List.map
- (fun (x, c) -> Name x, None, c) (List.combine rev_x_id_l (List.rev al))
+ (fun (x, c) -> LocalAssum (Name x, c)) (List.combine rev_x_id_l (List.rev al))
in
let env = Environ.push_rel_context context (Global.env ()) in
let glob_body =
@@ -678,8 +678,10 @@ let mkDestructEq :
let hyps = pf_hyps g in
let to_revert =
Util.List.map_filter
- (fun (id, _, t) ->
- if Id.List.mem id not_on_hyp || not (Termops.occur_term expr t)
+ (fun decl ->
+ let open Context.Named.Declaration in
+ let id = get_id decl in
+ if Id.List.mem id not_on_hyp || not (Termops.occur_term expr (get_type decl))
then None else Some id) hyps in
let to_revert_constr = List.rev_map mkVar to_revert in
let type_of_expr = pf_unsafe_type_of g expr in
@@ -1253,7 +1255,7 @@ let clear_goals =
then Termops.pop b'
else if b' == b then t
else mkProd(na,t',b')
- | _ -> map_constr clear_goal t
+ | _ -> Term.map_constr clear_goal t
in
List.map clear_goal
@@ -1489,7 +1491,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
let env = Global.env() in
let evd = ref (Evd.from_env env) in
let function_type = interp_type_evars env evd type_of_f in
- let env = push_named (function_name,None,function_type) env in
+ let env = push_named (Context.Named.Declaration.LocalAssum (function_name,function_type)) env in
(* Pp.msgnl (str "function type := " ++ Printer.pr_lconstr function_type); *)
let ty = interp_type_evars env evd ~impls:rec_impls eq in
let evm, nf = Evarutil.nf_evars_and_universes !evd in
@@ -1497,7 +1499,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
let function_type = nf function_type in
(* Pp.msgnl (str "lemma type := " ++ Printer.pr_lconstr equation_lemma_type ++ fnl ()); *)
let res_vars,eq' = decompose_prod equation_lemma_type in
- let env_eq' = Environ.push_rel_context (List.map (fun (x,y) -> (x,None,y)) res_vars) env in
+ let env_eq' = Environ.push_rel_context (List.map (fun (x,y) -> LocalAssum (x,y)) res_vars) env in
let eq' = nf_zeta env_eq' eq' in
let res =
(* Pp.msgnl (str "res_var :=" ++ Printer.pr_lconstr_env (push_rel_context (List.map (function (x,t) -> (x,None,t)) res_vars) env) eq'); *)
@@ -1515,7 +1517,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
let functional_id = add_suffix function_name "_F" in
let term_id = add_suffix function_name "_terminate" in
let functional_ref = declare_fun functional_id (IsDefinition Decl_kinds.Definition) ~ctx:(snd (Evd.universe_context evm)) res in
- let env_with_pre_rec_args = push_rel_context(List.map (function (x,t) -> (x,None,t)) pre_rec_args) env in
+ let env_with_pre_rec_args = push_rel_context(List.map (function (x,t) -> LocalAssum (x,t)) pre_rec_args) env in
let relation =
fst (*FIXME*)(interp_constr
env_with_pre_rec_args
diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml
index 7e38109d6..b740649e9 100644
--- a/plugins/omega/coq_omega.ml
+++ b/plugins/omega/coq_omega.ml
@@ -28,6 +28,7 @@ open Nametab
open Contradiction
open Misctypes
open Proofview.Notations
+open Context.Named.Declaration
module OmegaSolver = Omega.MakeOmegaSolver (Bigint)
open OmegaSolver
@@ -1695,25 +1696,26 @@ let destructure_hyps =
let pf_nf = Tacmach.New.of_old pf_nf gl in
let rec loop = function
| [] -> (Tacticals.New.tclTHEN nat_inject coq_omega)
- | (i,body,t)::lit ->
+ | decl::lit ->
+ let (i,_,t) = to_tuple decl in
begin try match destructurate_prop t with
| Kapp(False,[]) -> elim_id i
| Kapp((Zle|Zge|Zgt|Zlt|Zne),[t1;t2]) -> loop lit
| Kapp(Or,[t1;t2]) ->
(Tacticals.New.tclTHENS
(elim_id i)
- [ onClearedName i (fun i -> (loop ((i,None,t1)::lit)));
- onClearedName i (fun i -> (loop ((i,None,t2)::lit))) ])
+ [ onClearedName i (fun i -> (loop (LocalAssum (i,t1)::lit)));
+ onClearedName i (fun i -> (loop (LocalAssum (i,t2)::lit))) ])
| Kapp(And,[t1;t2]) ->
Tacticals.New.tclTHEN
(elim_id i)
(onClearedName2 i (fun i1 i2 ->
- loop ((i1,None,t1)::(i2,None,t2)::lit)))
+ loop (LocalAssum (i1,t1) :: LocalAssum (i2,t2) :: lit)))
| Kapp(Iff,[t1;t2]) ->
Tacticals.New.tclTHEN
(elim_id i)
(onClearedName2 i (fun i1 i2 ->
- loop ((i1,None,mkArrow t1 t2)::(i2,None,mkArrow t2 t1)::lit)))
+ loop (LocalAssum (i1,mkArrow t1 t2) :: LocalAssum (i2,mkArrow t2 t1) :: lit)))
| Kimp(t1,t2) ->
(* t1 and t2 might be in Type rather than Prop.
For t1, the decidability check will ensure being Prop. *)
@@ -1724,7 +1726,7 @@ let destructure_hyps =
Proofview.V82.tactic (generalize_tac [mkApp (Lazy.force coq_imp_simp,
[| t1; t2; d1; mkVar i|])]);
(onClearedName i (fun i ->
- (loop ((i,None,mk_or (mk_not t1) t2)::lit))))
+ (loop (LocalAssum (i,mk_or (mk_not t1) t2) :: lit))))
]
else
loop lit
@@ -1735,7 +1737,7 @@ let destructure_hyps =
Proofview.V82.tactic (generalize_tac
[mkApp (Lazy.force coq_not_or,[| t1; t2; mkVar i |])]);
(onClearedName i (fun i ->
- (loop ((i,None,mk_and (mk_not t1) (mk_not t2)):: lit))))
+ (loop (LocalAssum (i,mk_and (mk_not t1) (mk_not t2)) :: lit))))
]
| Kapp(And,[t1;t2]) ->
let d1 = decidability t1 in
@@ -1744,7 +1746,7 @@ let destructure_hyps =
[mkApp (Lazy.force coq_not_and,
[| t1; t2; d1; mkVar i |])]);
(onClearedName i (fun i ->
- (loop ((i,None,mk_or (mk_not t1) (mk_not t2))::lit))))
+ (loop (LocalAssum (i,mk_or (mk_not t1) (mk_not t2)) :: lit))))
]
| Kapp(Iff,[t1;t2]) ->
let d1 = decidability t1 in
@@ -1754,9 +1756,8 @@ let destructure_hyps =
[mkApp (Lazy.force coq_not_iff,
[| t1; t2; d1; d2; mkVar i |])]);
(onClearedName i (fun i ->
- (loop ((i,None,
- mk_or (mk_and t1 (mk_not t2))
- (mk_and (mk_not t1) t2))::lit))))
+ (loop (LocalAssum (i, mk_or (mk_and t1 (mk_not t2))
+ (mk_and (mk_not t1) t2)) :: lit))))
]
| Kimp(t1,t2) ->
(* t2 must be in Prop otherwise ~(t1->t2) wouldn't be ok.
@@ -1767,14 +1768,14 @@ let destructure_hyps =
[mkApp (Lazy.force coq_not_imp,
[| t1; t2; d1; mkVar i |])]);
(onClearedName i (fun i ->
- (loop ((i,None,mk_and t1 (mk_not t2)) :: lit))))
+ (loop (LocalAssum (i,mk_and t1 (mk_not t2)) :: lit))))
]
| Kapp(Not,[t]) ->
let d = decidability t in
Tacticals.New.tclTHENLIST [
Proofview.V82.tactic (generalize_tac
[mkApp (Lazy.force coq_not_not, [| t; d; mkVar i |])]);
- (onClearedName i (fun i -> (loop ((i,None,t)::lit))))
+ (onClearedName i (fun i -> (loop (LocalAssum (i,t) :: lit))))
]
| Kapp(op,[t1;t2]) ->
(try
@@ -1807,15 +1808,13 @@ let destructure_hyps =
match destructurate_type (pf_nf typ) with
| Kapp(Nat,_) ->
(Tacticals.New.tclTHEN
- (convert_hyp_no_check
- (i,body,
- (mkApp (Lazy.force coq_neq, [| t1;t2|]))))
+ (convert_hyp_no_check (set_type (mkApp (Lazy.force coq_neq, [| t1;t2|]))
+ decl))
(loop lit))
| Kapp(Z,_) ->
(Tacticals.New.tclTHEN
- (convert_hyp_no_check
- (i,body,
- (mkApp (Lazy.force coq_Zne, [| t1;t2|]))))
+ (convert_hyp_no_check (set_type (mkApp (Lazy.force coq_Zne, [| t1;t2|]))
+ decl))
(loop lit))
| _ -> loop lit
end
diff --git a/plugins/rtauto/refl_tauto.ml b/plugins/rtauto/refl_tauto.ml
index 9c22b5adb..2f3a3e551 100644
--- a/plugins/rtauto/refl_tauto.ml
+++ b/plugins/rtauto/refl_tauto.ml
@@ -13,6 +13,7 @@ open Util
open Term
open Tacmach
open Proof_search
+open Context.Named.Declaration
let force count lazc = incr count;Lazy.force lazc
@@ -128,9 +129,9 @@ let rec make_form atom_env gls term =
let rec make_hyps atom_env gls lenv = function
[] -> []
- | (_,Some body,typ)::rest ->
+ | LocalDef (_,body,typ)::rest ->
make_hyps atom_env gls (typ::body::lenv) rest
- | (id,None,typ)::rest ->
+ | LocalAssum (id,typ)::rest ->
let hrec=
make_hyps atom_env gls (typ::lenv) rest in
if List.exists (Termops.dependent (mkVar id)) lenv ||
diff --git a/plugins/rtauto/refl_tauto.mli b/plugins/rtauto/refl_tauto.mli
index c9e591bbd..9a14ac6c7 100644
--- a/plugins/rtauto/refl_tauto.mli
+++ b/plugins/rtauto/refl_tauto.mli
@@ -18,7 +18,7 @@ val make_hyps :
atom_env ->
Proof_type.goal Tacmach.sigma ->
Term.types list ->
- (Names.Id.t * Term.types option * Term.types) list ->
+ Context.Named.t ->
(Names.Id.t * Proof_search.form) list
val rtauto_tac : Proof_type.tactic
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<depth then i else i+n in
- Abstract (i, Context.Rel.Declaration.map (liftn n depth) d)
+ Abstract (i, map_constr (liftn n depth) d)
::(liftn_tomatch_stack n (depth+1) rest)
let lift_tomatch_stack n = liftn_tomatch_stack n 1
@@ -695,7 +697,7 @@ let merge_name get_name obj = function
let merge_names get_name = List.map2 (merge_name get_name)
let get_names env sign eqns =
- let names1 = List.make (List.length sign) Anonymous in
+ let names1 = List.make (Context.Rel.length sign) Anonymous in
(* If any, we prefer names used in pats, from top to bottom *)
let names2,aliasname =
List.fold_right
@@ -713,7 +715,7 @@ let get_names env sign eqns =
(fun (l,avoid) d na ->
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 =
diff --git a/printing/prettyp.ml b/printing/prettyp.ml
index b448df337..b7b1d67f0 100644
--- a/printing/prettyp.ml
+++ b/printing/prettyp.ml
@@ -35,7 +35,7 @@ type object_pr = {
print_syntactic_def : kernel_name -> std_ppcmds;
print_module : bool -> Names.module_path -> std_ppcmds;
print_modtype : module_path -> std_ppcmds;
- print_named_decl : Id.t * constr option * types -> std_ppcmds;
+ print_named_decl : Context.Named.Declaration.t -> std_ppcmds;
print_library_entry : bool -> (object_name * Lib.node) -> std_ppcmds option;
print_context : bool -> int option -> Lib.library_segment -> std_ppcmds;
print_typed_value_in_env : Environ.env -> Evd.evar_map -> Term.constr * Term.types -> Pp.std_ppcmds;
@@ -132,7 +132,8 @@ let print_renames_list prefix l =
let need_expansion impl ref =
let typ = Global.type_of_global_unsafe ref in
let ctx = prod_assum typ in
- let nprods = List.count (fun (_,b,_) -> Option.is_empty b) ctx in
+ let open Context.Rel.Declaration in
+ let nprods = List.count is_local_assum ctx in
not (List.is_empty impl) && List.length impl >= nprods &&
let _,lastimpl = List.chop nprods impl in
List.exists is_status_implicit lastimpl
@@ -168,8 +169,10 @@ type opacity =
| FullyOpaque
| TransparentMaybeOpacified of Conv_oracle.level
-let opacity env = function
- | VarRef v when not (Option.is_empty (pi2 (Environ.lookup_named v env))) ->
+let opacity env =
+ let open Context.Named.Declaration in
+ function
+ | VarRef v when is_local_def (Environ.lookup_named v env) ->
Some(TransparentMaybeOpacified
(Conv_oracle.get_strategy (Environ.oracle env) (VarKey v)))
| ConstRef cst ->
@@ -440,11 +443,13 @@ let print_named_def name body typ =
let print_named_assum name typ =
str "*** [" ++ str name ++ str " : " ++ pr_ltype typ ++ str "]"
-let gallina_print_named_decl (id,c,typ) =
- let s = Id.to_string id in
- match c with
- | Some body -> print_named_def s body typ
- | None -> print_named_assum s typ
+let gallina_print_named_decl =
+ let open Context.Named.Declaration in
+ function
+ | LocalAssum (id, typ) ->
+ print_named_assum (Id.to_string id) typ
+ | LocalDef (id, body, typ) ->
+ print_named_def (Id.to_string id) body typ
let assumptions_for_print lna =
List.fold_right (fun na env -> add_name na env) lna empty_names_context
@@ -721,8 +726,8 @@ let print_any_name = function
try (* Var locale de but, pas var de section... donc pas d'implicits *)
let dir,str = repr_qualid qid in
if not (DirPath.is_empty dir) then raise Not_found;
- let (_,c,typ) = Global.lookup_named str in
- (print_named_decl (str,c,typ))
+ let open Context.Named.Declaration in
+ str |> Global.lookup_named |> set_id str |> print_named_decl
with Not_found ->
errorlabstrm
"print_name" (pr_qualid qid ++ spc () ++ str "not a defined object.")
@@ -750,8 +755,8 @@ let print_opaque_name qid =
let ty = Universes.unsafe_type_of_global gr in
print_typed_value (mkConstruct cstr, ty)
| VarRef id ->
- let (_,c,ty) = lookup_named id env in
- print_named_decl (id,c,ty)
+ let open Context.Named.Declaration in
+ lookup_named id env |> set_id id |> print_named_decl
let print_about_any loc k =
match k with
diff --git a/printing/prettyp.mli b/printing/prettyp.mli
index 6f3556adb..0eab15579 100644
--- a/printing/prettyp.mli
+++ b/printing/prettyp.mli
@@ -66,7 +66,7 @@ type object_pr = {
print_syntactic_def : kernel_name -> std_ppcmds;
print_module : bool -> Names.module_path -> std_ppcmds;
print_modtype : module_path -> std_ppcmds;
- print_named_decl : Id.t * constr option * types -> std_ppcmds;
+ print_named_decl : Context.Named.Declaration.t -> std_ppcmds;
print_library_entry : bool -> (object_name * Lib.node) -> std_ppcmds option;
print_context : bool -> int option -> Lib.library_segment -> std_ppcmds;
print_typed_value_in_env : Environ.env -> Evd.evar_map -> Term.constr * Term.types -> Pp.std_ppcmds;
diff --git a/printing/printer.ml b/printing/printer.ml
index 93850e41f..5f4371f2d 100644
--- a/printing/printer.ml
+++ b/printing/printer.ml
@@ -262,16 +262,19 @@ let pr_var_decl_skel pr_id env sigma (id,c,typ) =
let ptyp = (str" : " ++ pt) in
(pr_id id ++ hov 0 (pbody ++ ptyp))
-let pr_var_decl env sigma (id,c,typ) =
- pr_var_decl_skel pr_id env sigma (id,c,typ)
+let pr_var_decl env sigma d =
+ pr_var_decl_skel pr_id env sigma (Context.Named.Declaration.to_tuple d)
let pr_var_list_decl env sigma (l,c,typ) =
hov 0 (pr_var_decl_skel (fun ids -> prlist_with_sep pr_comma pr_id ids) env sigma (l,c,typ))
-let pr_rel_decl env sigma (na,c,typ) =
- let pbody = match c with
- | None -> mt ()
- | Some c ->
+let pr_rel_decl env sigma decl =
+ let open Context.Rel.Declaration in
+ let na = get_name decl in
+ let typ = get_type decl in
+ let pbody = match decl with
+ | LocalAssum _ -> mt ()
+ | LocalDef (_,c,_) ->
(* Force evaluation *)
let pb = pr_lconstr_env env sigma c in
let pb = if isCast c then surround pb else pb in
@@ -420,7 +423,8 @@ let pr_evgl_sign sigma evi =
| None -> [], []
| Some f -> List.filter2 (fun b c -> not b) f (evar_context evi)
in
- let ids = List.rev_map pi1 l in
+ let open Context.Named.Declaration in
+ let ids = List.rev_map get_id l in
let warn =
if List.is_empty ids then mt () else
(str "(" ++ prlist_with_sep pr_comma pr_id ids ++ str " cannot be used)")
diff --git a/proofs/goal.ml b/proofs/goal.ml
index 1dd5be0e7..84ffee23c 100644
--- a/proofs/goal.ml
+++ b/proofs/goal.ml
@@ -9,6 +9,7 @@
open Util
open Pp
open Term
+open Context.Named.Declaration
(* This module implements the abstract interface to goals *)
(* A general invariant of the module, is that a goal whose associated
@@ -73,7 +74,7 @@ module V82 = struct
let (evars, evk) = Evarutil.new_pure_evar_full evars evi in
let evars = Evd.restore_future_goals evars prev_future_goals prev_principal_goal in
let ctxt = Environ.named_context_of_val hyps in
- let inst = Array.map_of_list (fun (id, _, _) -> mkVar id) ctxt in
+ let inst = Array.map_of_list (mkVar % get_id) ctxt in
let ev = Term.mkEvar (evk,inst) in
(evk, ev, evars)
@@ -139,7 +140,7 @@ module V82 = struct
let env = env sigma gl in
let genv = Global.env () in
let is_proof_var decl =
- try ignore (Environ.lookup_named (Util.pi1 decl) genv); false
+ try ignore (Environ.lookup_named (get_id decl) genv); false
with Not_found -> true in
Environ.fold_named_context_reverse (fun t decl ->
if is_proof_var decl then
diff --git a/proofs/logic.ml b/proofs/logic.ml
index 99e32db04..09f308abe 100644
--- a/proofs/logic.ml
+++ b/proofs/logic.ml
@@ -22,6 +22,7 @@ open Proof_type
open Type_errors
open Retyping
open Misctypes
+open Context.Named.Declaration
type refiner_error =
@@ -160,7 +161,8 @@ let reorder_context env sign ord =
| _ ->
(match ctxt_head with
| [] -> error_no_such_hypothesis (List.hd ord)
- | (x,_,_ as d) :: ctxt ->
+ | d :: ctxt ->
+ let x = get_id d in
if Id.Set.mem x expected then
step ord (Id.Set.remove x expected)
ctxt (push_item x d moved_hyps) ctxt_tail
@@ -175,7 +177,8 @@ let reorder_val_context env sign ord =
-let check_decl_position env sign (x,_,_ as d) =
+let check_decl_position env sign d =
+ let x = get_id d in
let needed = global_vars_set_of_decl env d in
let deps = dependency_closure env (named_context_of_val sign) needed in
if Id.List.mem x deps then
@@ -200,16 +203,17 @@ let move_location_eq m1 m2 = match m1, m2 with
let rec get_hyp_after h = function
| [] -> error_no_such_hypothesis h
- | (hyp,_,_) :: right ->
- if Id.equal hyp h then
- match right with (id,_,_)::_ -> MoveBefore id | [] -> MoveFirst
+ | d :: right ->
+ if Id.equal (get_id d) h then
+ match right with d' ::_ -> MoveBefore (get_id d') | [] -> MoveFirst
else
get_hyp_after h right
let split_sign hfrom hto l =
let rec splitrec left toleft = function
| [] -> error_no_such_hypothesis hfrom
- | (hyp,c,typ) as d :: right ->
+ | d :: right ->
+ let hyp,_,typ = to_tuple d in
if Id.equal hyp hfrom then
(left,right,d, toleft || move_location_eq hto MoveLast)
else
@@ -227,27 +231,28 @@ let hyp_of_move_location = function
| MoveBefore id -> id
| _ -> assert false
-let move_hyp toleft (left,(idfrom,_,_ as declfrom),right) hto =
+let move_hyp toleft (left,declfrom,right) hto =
let env = Global.env() in
- let test_dep (hyp,c,typ as d) (hyp2,c,typ2 as d2) =
+ let test_dep d d2 =
if toleft
- then occur_var_in_decl env hyp2 d
- else occur_var_in_decl env hyp d2
+ then occur_var_in_decl env (get_id d2) d
+ else occur_var_in_decl env (get_id d) d2
in
let rec moverec first middle = function
| [] ->
if match hto with MoveFirst | MoveLast -> false | _ -> true then
error_no_such_hypothesis (hyp_of_move_location hto);
List.rev first @ List.rev middle
- | (hyp,_,_) :: _ as right when move_location_eq hto (MoveBefore hyp) ->
+ | d :: _ as right when move_location_eq hto (MoveBefore (get_id d)) ->
List.rev first @ List.rev middle @ right
- | (hyp,_,_) as d :: right ->
+ | d :: right ->
+ let hyp = get_id d in
let (first',middle') =
if List.exists (test_dep d) middle then
if not (move_location_eq hto (MoveAfter hyp)) then
(first, d::middle)
else
- errorlabstrm "move_hyp" (str "Cannot move " ++ pr_id idfrom ++
+ errorlabstrm "move_hyp" (str "Cannot move " ++ pr_id (get_id declfrom) ++
Miscprint.pr_move_location pr_id hto ++
str (if toleft then ": it occurs in " else ": it depends on ")
++ pr_id hyp ++ str ".")
@@ -483,12 +488,14 @@ and mk_casegoals sigma goal goalacc p c =
(acc'',lbrty,conclty,sigma,p',c')
-let convert_hyp check sign sigma (id,b,bt as d) =
+let convert_hyp check sign sigma d =
+ let id,b,bt = to_tuple d in
let env = Global.env() in
let reorder = ref [] in
let sign' =
apply_to_hyp sign id
- (fun _ (_,c,ct) _ ->
+ (fun _ d' _ ->
+ let _,c,ct = to_tuple d' in
let env = Global.env_of_context sign in
if check && not (is_conv env sigma bt ct) then
errorlabstrm "Logic.convert_hyp"
@@ -522,14 +529,14 @@ let prim_refiner r sigma goal =
if replace then
let nexthyp = get_hyp_after id (named_context_of_val sign) in
let sign,t,cl,sigma = clear_hyps2 env sigma (Id.Set.singleton id) sign t cl in
- move_hyp false ([],(id,None,t),named_context_of_val sign)
+ move_hyp false ([], LocalAssum (id,t),named_context_of_val sign)
nexthyp,
t,cl,sigma
else
(if !check && mem_named_context id (named_context_of_val sign) then
errorlabstrm "Logic.prim_refiner"
(str "Variable " ++ pr_id id ++ str " is already declared.");
- push_named_context_val (id,None,t) sign,t,cl,sigma) in
+ push_named_context_val (LocalAssum (id,t)) sign,t,cl,sigma) in
let (sg2,ev2,sigma) =
Goal.V82.mk_goal sigma sign cl (Goal.V82.extra sigma goal) in
let oterm = Term.mkNamedLetIn id ev1 t ev2 in
@@ -546,7 +553,8 @@ let prim_refiner r sigma goal =
with Not_found ->
error "Cannot do a fixpoint on a non inductive type."
else
- check_ind (push_rel (na,None,c1) env) (k-1) b
+ let open Context.Rel.Declaration in
+ check_ind (push_rel (LocalAssum (na,c1)) env) (k-1) b
| _ -> error "Not enough products."
in
let ((sp,_),u) = check_ind env n cl in
@@ -560,7 +568,7 @@ let prim_refiner r sigma goal =
if !check && mem_named_context f (named_context_of_val sign) then
errorlabstrm "Logic.prim_refiner"
(str "Name " ++ pr_id f ++ str " already used in the environment");
- mk_sign (push_named_context_val (f,None,ar) sign) oth
+ mk_sign (push_named_context_val (LocalAssum (f,ar)) sign) oth
| [] ->
Evd.Monad.List.map (fun (_,_,c) sigma ->
let gl,ev,sig' =
@@ -584,7 +592,8 @@ let prim_refiner r sigma goal =
let rec check_is_coind env cl =
let b = whd_betadeltaiota env sigma cl in
match kind_of_term b with
- | Prod (na,c1,b) -> check_is_coind (push_rel (na,None,c1) env) b
+ | Prod (na,c1,b) -> let open Context.Rel.Declaration in
+ check_is_coind (push_rel (LocalAssum (na,c1)) env) b
| _ ->
try
let _ = find_coinductive env sigma b in ()
@@ -601,7 +610,7 @@ let prim_refiner r sigma goal =
error "Name already used in the environment.")
with
| Not_found ->
- mk_sign (push_named_context_val (f,None,ar) sign) oth)
+ mk_sign (push_named_context_val (LocalAssum (f,ar)) sign) oth)
| [] ->
Evd.Monad.List.map (fun (_,c) sigma ->
let gl,ev,sigma =
diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml
index fc33e9a65..403a36141 100644
--- a/proofs/proof_global.ml
+++ b/proofs/proof_global.ml
@@ -267,18 +267,19 @@ let _ = Goptions.declare_bool_option
Goptions.optwrite = (fun b -> proof_using_auto_clear := b) }
let set_used_variables l =
+ let open Context.Named.Declaration in
let env = Global.env () in
let ids = List.fold_right Id.Set.add l Id.Set.empty in
let ctx = Environ.keep_hyps env ids in
let ctx_set =
- List.fold_right Id.Set.add (List.map pi1 ctx) Id.Set.empty in
+ List.fold_right Id.Set.add (List.map get_id ctx) Id.Set.empty in
let vars_of = Environ.global_vars_set in
let aux env entry (ctx, all_safe, to_clear as orig) =
match entry with
- | (x,None,_) ->
+ | LocalAssum (x,_) ->
if Id.Set.mem x all_safe then orig
else (ctx, all_safe, (Loc.ghost,x)::to_clear)
- | (x,Some bo, ty) as decl ->
+ | LocalDef (x,bo, ty) as decl ->
if Id.Set.mem x all_safe then orig else
let vars = Id.Set.union (vars_of env bo) (vars_of env ty) in
if Id.Set.subset vars all_safe
diff --git a/proofs/proof_using.ml b/proofs/proof_using.ml
index a69645b11..681a7fa1a 100644
--- a/proofs/proof_using.ml
+++ b/proofs/proof_using.ml
@@ -10,6 +10,7 @@ open Names
open Environ
open Util
open Vernacexpr
+open Context.Named.Declaration
let to_string e =
let rec aux = function
@@ -33,7 +34,8 @@ let in_nameset =
let rec close_fwd e s =
let s' =
- List.fold_left (fun s (id,b,ty) ->
+ List.fold_left (fun s decl ->
+ let (id,b,ty) = Context.Named.Declaration.to_tuple decl in
let vb = Option.(default Id.Set.empty (map (global_vars_set e) b)) in
let vty = global_vars_set e ty in
let vbty = Id.Set.union vb vty in
@@ -61,13 +63,13 @@ and set_of_id env ty id =
Id.Set.union (global_vars_set env ty) acc)
Id.Set.empty ty
else if Id.to_string id = "All" then
- List.fold_right Id.Set.add (List.map pi1 (named_context env)) Id.Set.empty
+ List.fold_right Id.Set.add (List.map get_id (named_context env)) Id.Set.empty
else if CList.mem_assoc_f Id.equal id !known_names then
process_expr env (CList.assoc_f Id.equal id !known_names) []
else Id.Set.singleton id
and full_set env =
- List.fold_right Id.Set.add (List.map pi1 (named_context env)) Id.Set.empty
+ List.fold_right Id.Set.add (List.map get_id (named_context env)) Id.Set.empty
let process_expr env e ty =
let ty_expr = SsSingl(Loc.ghost, Id.of_string "Type") in
diff --git a/proofs/proofview.ml b/proofs/proofview.ml
index ff8effcda..ebce25d29 100644
--- a/proofs/proofview.ml
+++ b/proofs/proofview.ml
@@ -17,6 +17,7 @@ open Pp
open Util
open Proofview_monad
open Sigma.Notations
+open Context.Named.Declaration
(** Main state of tactics *)
type proofview = Proofview_monad.proofview
@@ -750,9 +751,15 @@ module Progress = struct
let eq_named_context_val sigma1 sigma2 ctx1 ctx2 =
let open Environ in
let c1 = named_context_of_val ctx1 and c2 = named_context_of_val ctx2 in
- let eq_named_declaration (i1, c1, t1) (i2, c2, t2) =
- Names.Id.equal i1 i2 && Option.equal (eq_constr sigma1 sigma2) c1 c2
- && (eq_constr sigma1 sigma2) t1 t2
+ let eq_named_declaration d1 d2 =
+ match d1, d2 with
+ | LocalAssum (i1,t1), LocalAssum (i2,t2) ->
+ Names.Id.equal i1 i2 && eq_constr sigma1 sigma2 t1 t2
+ | LocalDef (i1,c1,t1), LocalDef (i2,c2,t2) ->
+ Names.Id.equal i1 i2 && eq_constr sigma1 sigma2 c1 c2
+ && eq_constr sigma1 sigma2 t1 t2
+ | _ ->
+ false
in List.equal eq_named_declaration c1 c2
let eq_evar_body sigma1 sigma2 b1 b2 =
@@ -1075,12 +1082,13 @@ struct
let typecheck_evar ev env sigma =
let info = Evd.find sigma ev in
(** Typecheck the hypotheses. *)
- let type_hyp (sigma, env) (na, body, t as decl) =
+ let type_hyp (sigma, env) decl =
+ let t = get_type decl in
let evdref = ref sigma in
let _ = Typing.sort_of env evdref t in
- let () = match body with
- | None -> ()
- | Some body -> Typing.check env evdref body t
+ let () = match decl with
+ | LocalAssum _ -> ()
+ | LocalDef (_,body,_) -> Typing.check env evdref body t
in
(!evdref, Environ.push_named decl env)
in
diff --git a/proofs/refiner.ml b/proofs/refiner.ml
index 8d6bdf6ae..186525e15 100644
--- a/proofs/refiner.ml
+++ b/proofs/refiner.ml
@@ -13,7 +13,7 @@ open Evd
open Environ
open Proof_type
open Logic
-
+open Context.Named.Declaration
let sig_it x = x.it
let project x = x.sigma
@@ -202,7 +202,7 @@ let tclSHOWHYPS (tac : tactic) (goal: Goal.goal Evd.sigma)
let { it = gls; sigma = sigma; } = rslt in
let hyps:Context.Named.t list =
List.map (fun gl -> pf_hyps { it = gl; sigma=sigma; }) gls in
- let cmp (i1, c1, t1) (i2, c2, t2) = Names.Id.equal i1 i2 in
+ let cmp d1 d2 = Names.Id.equal (get_id d1) (get_id d2) in
let newhyps =
List.map
(fun hypl -> List.subtract cmp hypl oldhyps)
@@ -215,7 +215,7 @@ let tclSHOWHYPS (tac : tactic) (goal: Goal.goal Evd.sigma)
List.fold_left
(fun acc lh -> acc ^ (if !frst then (frst:=false;"") else " | ")
^ (List.fold_left
- (fun acc (nm,_,_) -> (Names.Id.to_string nm) ^ " " ^ acc)
+ (fun acc d -> (Names.Id.to_string (get_id d)) ^ " " ^ acc)
"" lh))
"" newhyps in
pp (str (emacs_str "<infoH>")
diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml
index a1ebacea8..429bd2774 100644
--- a/proofs/tacmach.ml
+++ b/proofs/tacmach.ml
@@ -18,6 +18,7 @@ open Tacred
open Proof_type
open Logic
open Refiner
+open Context.Named.Declaration
let re_sig it gc = { it = it; sigma = gc; }
@@ -40,9 +41,11 @@ let pf_hyps = Refiner.pf_hyps
let pf_concl gls = Goal.V82.concl (project gls) (sig_it gls)
let pf_hyps_types gls =
let sign = Environ.named_context (pf_env gls) in
- List.map (fun (id,_,x) -> (id, x)) sign
+ List.map (function LocalAssum (id,x)
+ | LocalDef (id,_,x) -> id, x)
+ sign
-let pf_nth_hyp_id gls n = let (id,c,t) = List.nth (pf_hyps gls) (n-1) in id
+let pf_nth_hyp_id gls n = List.nth (pf_hyps gls) (n-1) |> get_id
let pf_last_hyp gl = List.hd (pf_hyps gl)
@@ -53,8 +56,7 @@ let pf_get_hyp gls id =
raise (RefinerError (NoSuchHyp id))
let pf_get_hyp_typ gls id =
- let (_,_,ty)= (pf_get_hyp gls id) in
- ty
+ pf_get_hyp gls id |> get_type
let pf_ids_of_hyps gls = ids_of_named_context (pf_hyps gls)
@@ -204,13 +206,14 @@ module New = struct
sign
let pf_get_hyp_typ id gl =
- let (_,_,ty) = pf_get_hyp id gl in
- ty
+ pf_get_hyp id gl |> get_type
let pf_hyps_types gl =
let env = Proofview.Goal.env gl in
let sign = Environ.named_context env in
- List.map (fun (id,_,x) -> (id, x)) sign
+ List.map (function LocalAssum (id,x)
+ | LocalDef (id,_,x) -> id, x)
+ sign
let pf_last_hyp gl =
let hyps = Proofview.Goal.hyps gl in
diff --git a/stm/lemmas.ml b/stm/lemmas.ml
index 845f83a40..ac54028eb 100644
--- a/stm/lemmas.ml
+++ b/stm/lemmas.ml
@@ -31,6 +31,7 @@ open Reductionops
open Constrexpr
open Constrintern
open Impargs
+open Context.Rel.Declaration
type 'a declaration_hook = Decl_kinds.locality -> Globnames.global_reference -> 'a
let mk_hook hook = hook
@@ -44,7 +45,8 @@ let call_hook fix_exn hook l c =
let retrieve_first_recthm = function
| VarRef id ->
- (pi2 (Global.lookup_named id),variable_opacity id)
+ let open Context.Named.Declaration in
+ (get_value (Global.lookup_named id),variable_opacity id)
| ConstRef cst ->
let cb = Global.lookup_constant cst in
(Global.body_of_constant_body cb, is_opaque cb)
@@ -107,11 +109,12 @@ let find_mutually_recursive_statements thms =
(fun env c -> fst (whd_betadeltaiota_stack env Evd.empty c))
(Global.env()) hyps in
let ind_hyps =
- List.flatten (List.map_i (fun i (_,b,t) ->
+ List.flatten (List.map_i (fun i decl ->
+ let t = get_type decl in
match kind_of_term t with
| Ind ((kn,_ as ind),u) when
let mind = Global.lookup_mind kn in
- mind.mind_finite <> Decl_kinds.CoFinite && Option.is_empty b ->
+ mind.mind_finite <> Decl_kinds.CoFinite && is_local_assum decl ->
[ind,x,i]
| _ ->
[]) 0 (List.rev whnf_hyp_hds)) in
@@ -450,7 +453,7 @@ let start_proof_com kind thms hook =
let impls, ((env, ctx), imps) = interp_context_evars env0 evdref bl in
let t', imps' = interp_type_evars_impls ~impls env evdref t in
evdref := solve_remaining_evars all_and_fail_flags env !evdref (Evd.empty,!evdref);
- let ids = List.map pi1 ctx in
+ let ids = List.map get_name ctx in
(compute_proof_name (pi1 kind) sopt,
(nf_evar !evdref (it_mkProd_or_LetIn t' ctx),
(ids, imps @ lift_implicits (List.length ids) imps'),
diff --git a/stm/stm.ml b/stm/stm.ml
index e8b500a62..1503c0f8a 100644
--- a/stm/stm.ml
+++ b/stm/stm.ml
@@ -1477,10 +1477,8 @@ end = struct (* {{{ *)
let g = Evd.find sigma0 r_goal in
if not (
Evarutil.is_ground_term sigma0 Evd.(evar_concl g) &&
- List.for_all (fun (_,bo,ty) ->
- Evarutil.is_ground_term sigma0 ty &&
- Option.cata (Evarutil.is_ground_term sigma0) true bo)
- Evd.(evar_context g))
+ List.for_all (Context.Named.Declaration.for_all (Evarutil.is_ground_term sigma0))
+ Evd.(evar_context g))
then
Errors.errorlabstrm "Stm" (strbrk("the par: goal selector supports ground "^
"goals only"))
diff --git a/tactics/auto.ml b/tactics/auto.ml
index 6caebf6c4..86b71999b 100644
--- a/tactics/auto.ml
+++ b/tactics/auto.ml
@@ -324,7 +324,7 @@ let rec trivial_fail_db dbg mod_delta db_list local_db =
let env = Proofview.Goal.env gl in
let nf c = Evarutil.nf_evar sigma c in
let decl = Tacmach.New.pf_last_hyp (Proofview.Goal.assume gl) in
- let hyp = Context.Named.Declaration.map nf decl in
+ let hyp = Context.Named.Declaration.map_constr nf decl in
let hintl = make_resolve_hyp env sigma hyp
in trivial_fail_db dbg mod_delta db_list
(Hint_db.add_list env sigma hintl local_db)
diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml
index 40c0f7f9b..ea598b61c 100644
--- a/tactics/autorewrite.ml
+++ b/tactics/autorewrite.ml
@@ -133,7 +133,7 @@ let autorewrite_multi_in ?(conds=Naive) idl tac_main lbas =
fun dir cstr tac gl ->
let last_hyp_id =
match Tacmach.pf_hyps gl with
- (last_hyp_id,_,_)::_ -> last_hyp_id
+ d :: _ -> Context.Named.Declaration.get_id d
| _ -> (* even the hypothesis id is missing *)
raise (Logic.RefinerError (Logic.NoSuchHyp !id))
in
@@ -142,7 +142,8 @@ let autorewrite_multi_in ?(conds=Naive) idl tac_main lbas =
match gls with
g::_ ->
(match Environ.named_context_of_val (Goal.V82.hyps gl'.Evd.sigma g) with
- (lastid,_,_)::_ ->
+ d ::_ ->
+ let lastid = Context.Named.Declaration.get_id d in
if not (Id.equal last_hyp_id lastid) then
begin
let gl'' =
diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml
index c9b2c7cfd..7c05befdd 100644
--- a/tactics/class_tactics.ml
+++ b/tactics/class_tactics.ml
@@ -302,8 +302,10 @@ type ('a,'b) optionk2 =
| Nonek2 of failure
| Somek2 of 'a * 'b * ('a,'b) optionk2 fk
-let make_resolve_hyp env sigma st flags only_classes pri (id, _, cty) =
- let cty = Evarutil.nf_evar sigma cty in
+let make_resolve_hyp env sigma st flags only_classes pri decl =
+ let open Context.Named.Declaration in
+ let id = get_id decl in
+ let cty = Evarutil.nf_evar sigma (get_type decl) in
let rec iscl env ty =
let ctx, ar = decompose_prod_assum ty in
match kind_of_term (fst (decompose_app ar)) with
@@ -345,9 +347,10 @@ let make_hints g st only_classes sign =
List.fold_left
(fun (paths, hints) hyp ->
let consider =
- try let (_, b, t) = Global.lookup_named (pi1 hyp) in
+ let open Context.Named.Declaration in
+ try let t = Global.lookup_named (get_id hyp) |> get_type in
(* Section variable, reindex only if the type changed *)
- not (Term.eq_constr t (pi3 hyp))
+ not (Term.eq_constr t (get_type hyp))
with Not_found -> true
in
if consider then
diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml
index c4a23f686..ab6fb37fd 100644
--- a/tactics/contradiction.ml
+++ b/tactics/contradiction.ml
@@ -15,6 +15,7 @@ open Reductionops
open Misctypes
open Sigma.Notations
open Proofview.Notations
+open Context.Named.Declaration
(* Absurd *)
@@ -47,7 +48,7 @@ let absurd c = absurd c
let filter_hyp f tac =
let rec seek = function
| [] -> Proofview.tclZERO Not_found
- | (id,_,t)::rest when f t -> tac id
+ | d::rest when f (get_type d) -> tac (get_id d)
| _::rest -> seek rest in
Proofview.Goal.enter { enter = begin fun gl ->
let hyps = Proofview.Goal.hyps (Proofview.Goal.assume gl) in
@@ -60,8 +61,9 @@ let contradiction_context =
let env = Proofview.Goal.env gl in
let rec seek_neg l = match l with
| [] -> Tacticals.New.tclZEROMSG (Pp.str"No such contradiction")
- | (id,_,typ)::rest ->
- let typ = nf_evar sigma typ in
+ | d :: rest ->
+ let id = get_id d in
+ let typ = nf_evar sigma (get_type d) in
let typ = whd_betadeltaiota env sigma typ in
if is_empty_type typ then
simplest_elim (mkVar id)
diff --git a/tactics/elim.ml b/tactics/elim.ml
index 7767affcc..d441074f6 100644
--- a/tactics/elim.ml
+++ b/tactics/elim.ml
@@ -16,6 +16,7 @@ open Tacmach.New
open Tacticals.New
open Tactics
open Proofview.Notations
+open Context.Named.Declaration
(* Supposed to be called without as clause *)
let introElimAssumsThen tac ba =
@@ -137,7 +138,8 @@ let induction_trailer abs_i abs_j bargs =
in
let (hyps,_) =
List.fold_left
- (fun (bring_ids,leave_ids) (cid,_,_ as d) ->
+ (fun (bring_ids,leave_ids) d ->
+ let cid = get_id d in
if not (List.mem cid leave_ids)
then (d::bring_ids,leave_ids)
else (bring_ids,cid::leave_ids))
diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml
index e0bea7770..a03489c80 100644
--- a/tactics/eqschemes.ml
+++ b/tactics/eqschemes.ml
@@ -58,6 +58,7 @@ open Inductiveops
open Ind_tables
open Indrec
open Sigma.Notations
+open Context.Rel.Declaration
let hid = Id.of_string "H"
let xid = Id.of_string "X"
@@ -104,7 +105,7 @@ let get_sym_eq_data env (ind,u) =
error "Not an inductive type with a single constructor.";
let arityctxt = Vars.subst_instance_context u mip.mind_arity_ctxt in
let realsign,_ = List.chop mip.mind_nrealdecls arityctxt in
- if List.exists (fun (_,b,_) -> not (Option.is_empty b)) realsign then
+ if List.exists is_local_def realsign then
error "Inductive equalities with local definitions in arity not supported.";
let constrsign,ccl = decompose_prod_assum mip.mind_nf_lc.(0) in
let _,constrargs = decompose_app ccl in
@@ -139,7 +140,7 @@ let get_non_sym_eq_data env (ind,u) =
error "Not an inductive type with a single constructor.";
let arityctxt = Vars.subst_instance_context u mip.mind_arity_ctxt in
let realsign,_ = List.chop mip.mind_nrealdecls arityctxt in
- if List.exists (fun (_,b,_) -> not (Option.is_empty b)) realsign then
+ if List.exists is_local_def realsign then
error "Inductive equalities with local definitions in arity not supported";
let constrsign,ccl = decompose_prod_assum mip.mind_nf_lc.(0) in
let _,constrargs = decompose_app ccl in
@@ -173,7 +174,7 @@ let build_sym_scheme env ind =
let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in
let applied_ind = build_dependent_inductive indu specif in
let realsign_ind =
- name_context env ((Name varH,None,applied_ind)::realsign) in
+ name_context env ((LocalAssum (Name varH,applied_ind))::realsign) in
let ci = make_case_info (Global.env()) ind RegularStyle in
let c =
(my_it_mkLambda_or_LetIn paramsctxt
@@ -232,7 +233,7 @@ let build_sym_involutive_scheme env ind =
(Context.Rel.to_extended_vect (nrealargs+1) mib.mind_params_ctxt)
(rel_vect (nrealargs+1) nrealargs)) in
let realsign_ind =
- name_context env ((Name varH,None,applied_ind)::realsign) in
+ name_context env ((LocalAssum (Name varH,applied_ind))::realsign) in
let ci = make_case_info (Global.env()) ind RegularStyle in
let c =
(my_it_mkLambda_or_LetIn paramsctxt
@@ -352,9 +353,9 @@ let build_l2r_rew_scheme dep env ind kind =
rel_vect 0 nrealargs]) in
let realsign_P = lift_rel_context nrealargs realsign in
let realsign_ind_P =
- name_context env ((Name varH,None,applied_ind_P)::realsign_P) in
+ name_context env ((LocalAssum (Name varH,applied_ind_P))::realsign_P) in
let realsign_ind_G =
- name_context env ((Name varH,None,applied_ind_G)::
+ name_context env ((LocalAssum (Name varH,applied_ind_G))::
lift_rel_context (nrealargs+3) realsign) in
let applied_sym_C n =
mkApp(sym,
@@ -465,9 +466,9 @@ let build_l2r_forward_rew_scheme dep env ind kind =
rel_vect (2*nrealargs+1) nrealargs]) in
let realsign_P n = lift_rel_context (nrealargs*n+n) realsign in
let realsign_ind =
- name_context env ((Name varH,None,applied_ind)::realsign) in
+ name_context env ((LocalAssum (Name varH,applied_ind))::realsign) in
let realsign_ind_P n aP =
- name_context env ((Name varH,None,aP)::realsign_P n) in
+ name_context env ((LocalAssum (Name varH,aP))::realsign_P n) in
let s, ctx' = Universes.fresh_sort_in_family (Global.env ()) kind in
let ctx = Univ.ContextSet.union ctx ctx' in
let s = mkSort s in
@@ -545,7 +546,7 @@ let build_r2l_forward_rew_scheme dep env ind kind =
let varP = fresh env (Id.of_string "P") in
let applied_ind = build_dependent_inductive indu specif in
let realsign_ind =
- name_context env ((Name varH,None,applied_ind)::realsign) in
+ name_context env ((LocalAssum (Name varH,applied_ind))::realsign) in
let s, ctx' = Universes.fresh_sort_in_family (Global.env ()) kind in
let ctx = Univ.ContextSet.union ctx ctx' in
let s = mkSort s in
@@ -599,9 +600,9 @@ let fix_r2l_forward_rew_scheme (c, ctx') =
| hp :: p :: ind :: indargs ->
let c' =
my_it_mkLambda_or_LetIn indargs
- (mkLambda_or_LetIn (Context.Rel.Declaration.map (liftn (-1) 1) p)
- (mkLambda_or_LetIn (Context.Rel.Declaration.map (liftn (-1) 2) hp)
- (mkLambda_or_LetIn (Context.Rel.Declaration.map (lift 2) ind)
+ (mkLambda_or_LetIn (map_constr (liftn (-1) 1) p)
+ (mkLambda_or_LetIn (map_constr (liftn (-1) 2) hp)
+ (mkLambda_or_LetIn (map_constr (lift 2) ind)
(Reductionops.whd_beta Evd.empty
(applist (c,
Context.Rel.to_extended_list 3 indargs @ [mkRel 1;mkRel 3;mkRel 2]))))))
@@ -737,10 +738,10 @@ let build_congr env (eq,refl,ctx) ind =
let arityctxt = Vars.subst_instance_context u mip.mind_arity_ctxt in
let paramsctxt = Vars.subst_instance_context u mib.mind_params_ctxt in
let realsign,_ = List.chop mip.mind_nrealdecls arityctxt in
- if List.exists (fun (_,b,_) -> not (Option.is_empty b)) realsign then
+ if List.exists is_local_def realsign then
error "Inductive equalities with local definitions in arity not supported.";
let env_with_arity = push_rel_context arityctxt env in
- let (_,_,ty) = lookup_rel (mip.mind_nrealargs - i + 1) env_with_arity in
+ let ty = get_type (lookup_rel (mip.mind_nrealargs - i + 1) env_with_arity) in
let constrsign,ccl = decompose_prod_assum mip.mind_nf_lc.(0) in
let _,constrargs = decompose_app ccl in
if Int.equal (Context.Rel.length constrsign) (Context.Rel.length mib.mind_params_ctxt) then
diff --git a/tactics/equality.ml b/tactics/equality.ml
index 1e814e861..d27dcd82a 100644
--- a/tactics/equality.ml
+++ b/tactics/equality.ml
@@ -43,6 +43,7 @@ open Misctypes
open Sigma.Notations
open Proofview.Notations
open Unification
+open Context.Named.Declaration
(* Options *)
@@ -960,7 +961,7 @@ let apply_on_clause (f,t) clause =
let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn sort =
let e = next_ident_away eq_baseid (ids_of_context env) in
- let e_env = push_named (e,None,t) env in
+ let e_env = push_named (Context.Named.Declaration.LocalAssum (e,t)) env in
let discriminator =
build_discriminator e_env sigma dirn (mkVar e) sort cpath in
let sigma,(pf, absurd_term), eff =
@@ -1064,7 +1065,7 @@ let make_tuple env sigma (rterm,rty) lind =
assert (dependent (mkRel lind) rty);
let sigdata = find_sigma_data env (get_sort_of env sigma rty) in
let sigma, a = type_of ~refresh:true env sigma (mkRel lind) in
- let (na,_,_) = lookup_rel lind env in
+ let na = Context.Rel.Declaration.get_name (lookup_rel lind env) in
(* We move [lind] to [1] and lift other rels > [lind] by 1 *)
let rty = lift (1-lind) (liftn lind (lind+1) rty) in
(* Now [lind] is [mkRel 1] and we abstract on (na:a) *)
@@ -1335,7 +1336,7 @@ let simplify_args env sigma t =
let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac =
let e = next_ident_away eq_baseid (ids_of_context env) in
- let e_env = push_named (e, None,t) env in
+ let e_env = push_named (LocalAssum (e,t)) env in
let evdref = ref sigma in
let filter (cpath, t1', t2') =
try
@@ -1612,14 +1613,14 @@ let unfold_body x =
Proofview.Goal.enter { enter = begin fun gl ->
(** We normalize the given hypothesis immediately. *)
let hyps = Proofview.Goal.hyps (Proofview.Goal.assume gl) in
- let (_, xval, _) = Context.Named.lookup x hyps in
+ let xval = Context.Named.lookup x hyps |> get_value in
let xval = match xval with
| None -> errorlabstrm "unfold_body"
(pr_id x ++ str" is not a defined hypothesis.")
| Some xval -> pf_nf_evar gl xval
in
afterHyp x begin fun aft ->
- let hl = List.fold_right (fun (y,yval,_) cl -> (y,InHyp) :: cl) aft [] in
+ let hl = List.fold_right (fun d cl -> (get_id d, InHyp) :: cl) aft [] in
let xvar = mkVar x in
let rfun _ _ c = replace_term xvar xval c in
let reducth h = Proofview.V82.tactic (fun gl -> reduct_in_hyp rfun h gl) in
@@ -1636,9 +1637,10 @@ let restrict_to_eq_and_identity eq = (* compatibility *)
exception FoundHyp of (Id.t * constr * bool)
(* tests whether hyp [c] is [x = t] or [t = x], [x] not occurring in [t] *)
-let is_eq_x gl x (id,_,c) =
+let is_eq_x gl x d =
+ let id = get_id d in
try
- let c = pf_nf_evar gl c in
+ let c = pf_nf_evar gl (get_type d) in
let (_,lhs,rhs) = pi3 (find_eq_data_decompose gl c) in
if (Term.eq_constr x lhs) && not (occur_term x rhs) then raise (FoundHyp (id,rhs,true));
if (Term.eq_constr x rhs) && not (occur_term x lhs) then raise (FoundHyp (id,lhs,false))
@@ -1655,11 +1657,12 @@ let subst_one dep_proof_ok x (hyp,rhs,dir) =
let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in
(* The set of hypotheses using x *)
let dephyps =
- List.rev (snd (List.fold_right (fun (id,b,_ as dcl) (deps,allhyps) ->
+ List.rev (snd (List.fold_right (fun dcl (deps,allhyps) ->
+ let id = get_id dcl in
if not (Id.equal id hyp)
&& List.exists (fun y -> occur_var_in_decl env y dcl) deps
then
- ((if b = None then deps else id::deps), id::allhyps)
+ ((if is_local_assum dcl then deps else id::deps), id::allhyps)
else
(deps,allhyps))
hyps
@@ -1683,7 +1686,7 @@ let subst_one dep_proof_ok x (hyp,rhs,dir) =
let subst_one_var dep_proof_ok x =
Proofview.Goal.enter { enter = begin fun gl ->
let gl = Proofview.Goal.assume gl in
- let (_,xval,_) = pf_get_hyp x gl in
+ let xval = pf_get_hyp x gl |> get_value in
(* If x has a body, simply replace x with body and clear x *)
if not (Option.is_empty xval) then tclTHEN (unfold_body x) (clear [x]) else
(* x is a variable: *)
@@ -1742,14 +1745,14 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () =
let gl = Proofview.Goal.assume gl in
let env = Proofview.Goal.env gl in
let find_eq_data_decompose = find_eq_data_decompose gl in
- let test (hyp,_,c) =
+ let test decl =
try
- let lbeq,u,(_,x,y) = find_eq_data_decompose c in
+ let lbeq,u,(_,x,y) = find_eq_data_decompose (get_type decl) in
let eq = Universes.constr_of_global_univ (lbeq.eq,u) in
if flags.only_leibniz then restrict_to_eq_and_identity eq;
match kind_of_term x, kind_of_term y with
| Var z, _ | _, Var z when not (is_evaluable env (EvalVarRef z)) ->
- Some hyp
+ Some (get_id decl)
| _ ->
None
with Constr_matching.PatternMatchingFailure -> None
@@ -1763,7 +1766,7 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () =
Proofview.Goal.enter { enter = begin fun gl ->
let gl = Proofview.Goal.assume gl in
let find_eq_data_decompose = find_eq_data_decompose gl in
- let (_,_,c) = pf_get_hyp hyp gl in
+ let c = pf_get_hyp hyp gl |> get_type in
let _,_,(_,x,y) = find_eq_data_decompose c in
(* J.F.: added to prevent failure on goal containing x=x as an hyp *)
if Term.eq_constr x y then Proofview.tclUNIT () else
@@ -1831,10 +1834,11 @@ let cond_eq_term c t gl =
let rewrite_assumption_cond cond_eq_term cl =
let rec arec hyps gl = match hyps with
| [] -> error "No such assumption."
- | (id,_,t) ::rest ->
+ | hyp ::rest ->
+ let id = get_id hyp in
begin
try
- let dir = cond_eq_term t gl in
+ let dir = cond_eq_term (get_type hyp) gl in
general_rewrite_clause dir false (mkVar id,NoBindings) cl
with | Failure _ | UserError _ -> arec rest gl
end
diff --git a/tactics/evar_tactics.ml b/tactics/evar_tactics.ml
index 97b5ba0cc..f443837a4 100644
--- a/tactics/evar_tactics.ml
+++ b/tactics/evar_tactics.ml
@@ -16,6 +16,7 @@ open Evd
open Locus
open Sigma.Notations
open Proofview.Notations
+open Context.Named.Declaration
(* The instantiate tactic *)
@@ -43,14 +44,14 @@ let instantiate_tac n c ido =
match hloc with
InHyp ->
(match decl with
- (_,None,typ) -> evar_list typ
+ | LocalAssum (_,typ) -> evar_list typ
| _ -> error
"Please be more specific: in type or value?")
| InHypTypeOnly ->
- let (_, _, typ) = decl in evar_list typ
+ evar_list (get_type decl)
| InHypValueOnly ->
(match decl with
- (_,Some body,_) -> evar_list body
+ | LocalDef (_,body,_) -> evar_list body
| _ -> error "Not a defined hypothesis.") in
if List.length evl < n then
error "Not enough uninstantiated existential variables.";
diff --git a/tactics/hints.ml b/tactics/hints.ml
index c99e591fe..730da147a 100644
--- a/tactics/hints.ml
+++ b/tactics/hints.ml
@@ -34,6 +34,7 @@ open Tacred
open Printer
open Vernacexpr
open Sigma.Notations
+open Context.Named.Declaration
(****************************************)
(* General functions *)
@@ -727,11 +728,12 @@ let make_resolves env sigma flags pri poly ?name cr =
ents
(* used to add an hypothesis to the local hint database *)
-let make_resolve_hyp env sigma (hname,_,htyp) =
+let make_resolve_hyp env sigma decl =
+ let hname = get_id decl in
try
[make_apply_entry env sigma (true, true, false) None false
~name:(PathHints [VarRef hname])
- (mkVar hname, htyp, Univ.ContextSet.empty)]
+ (mkVar hname, get_type decl, Univ.ContextSet.empty)]
with
| Failure _ -> []
| e when Logic.catchable_exception e -> anomaly (Pp.str "make_resolve_hyp")
@@ -1061,7 +1063,7 @@ let prepare_hint check (poly,local) env init (sigma,c) =
(* Not clever enough to construct dependency graph of evars *)
error "Not clever enough to deal with evars dependent in other evars.";
raise (Found (c,t))
- | _ -> iter_constr find_next_evar c in
+ | _ -> Constr.iter find_next_evar c in
let rec iter c =
try find_next_evar c; c
with Found (evar,t) ->
diff --git a/tactics/hipattern.ml4 b/tactics/hipattern.ml4
index 29d848ca1..bcec90f80 100644
--- a/tactics/hipattern.ml4
+++ b/tactics/hipattern.ml4
@@ -19,6 +19,7 @@ open Constr_matching
open Coqlib
open Declarations
open Tacmach.New
+open Context.Rel.Declaration
(* I implemented the following functions which test whether a term t
is an inductive but non-recursive type, a general conjuction, a
@@ -101,13 +102,16 @@ let match_with_one_constructor style onlybinary allow_rec t =
(decompose_prod_n_assum mib.mind_nparams mip.mind_nf_lc.(0)))) in
if
List.for_all
- (fun (_,b,c) -> Option.is_empty b && isRel c && Int.equal (destRel c) mib.mind_nparams) ctx
+ (fun decl -> let c = get_type decl in
+ is_local_assum decl &&
+ isRel c &&
+ Int.equal (destRel c) mib.mind_nparams) ctx
then
Some (hdapp,args)
else None
else
let ctyp = prod_applist mip.mind_nf_lc.(0) args in
- let cargs = List.map pi3 ((prod_assum ctyp)) in
+ let cargs = List.map get_type (prod_assum ctyp) in
if not (is_lax_conjunction style) || has_nodep_prod ctyp then
(* Record or non strict conjunction *)
Some (hdapp,List.rev cargs)
@@ -152,7 +156,7 @@ let is_tuple t =
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)
+ | [LocalAssum (_,c)] -> isRel c && Int.equal (destRel c) (n - i)
| _ -> false) 0 lc
let match_with_disjunction ?(strict=false) ?(onlybinary=false) t =
diff --git a/tactics/inv.ml b/tactics/inv.ml
index ded1e8076..9bfbbc41b 100644
--- a/tactics/inv.ml
+++ b/tactics/inv.ml
@@ -28,6 +28,7 @@ open Misctypes
open Tacexpr
open Sigma.Notations
open Proofview.Notations
+open Context.Named.Declaration
let clear hyps = Proofview.V82.tactic (clear hyps)
@@ -181,9 +182,9 @@ let make_inv_predicate env evd indf realargs id status concl =
let dependent_hyps env id idlist gl =
let rec dep_rec =function
| [] -> []
- | (id1,_,_)::l ->
+ | d::l ->
(* Update the type of id1: it may have been subject to rewriting *)
- let d = pf_get_hyp id1 gl in
+ let d = pf_get_hyp (get_id d) gl in
if occur_var_in_decl env id d
then d :: dep_rec l
else dep_rec l
@@ -192,8 +193,8 @@ let dependent_hyps env id idlist gl =
let split_dep_and_nodep hyps gl =
List.fold_right
- (fun (id,_,_ as d) (l1,l2) ->
- if var_occurs_in_pf gl id then (d::l1,l2) else (l1,d::l2))
+ (fun d (l1,l2) ->
+ if var_occurs_in_pf gl (get_id d) then (d::l1,l2) else (l1,d::l2))
hyps ([],[])
(* Computation of dids is late; must have been done in rewrite_equations*)
@@ -296,8 +297,8 @@ let get_names (allow_conj,issimple) (loc, pat as x) = match pat with
error "Discarding pattern not allowed for inversion equations."
| IntroAction (IntroRewrite _) ->
error "Rewriting pattern not allowed for inversion equations."
- | IntroAction (IntroOrAndPattern (IntroAndPattern [] | IntroOrPattern [[]])) when allow_conj -> (None, [])
- | IntroAction (IntroOrAndPattern (IntroAndPattern ((_,IntroNaming (IntroIdentifier id)) :: _ as l) | IntroOrPattern [(_,IntroNaming (IntroIdentifier id)) :: _ as l ]))
+ | IntroAction (IntroOrAndPattern (IntroAndPattern [])) when allow_conj -> (None, [])
+ | IntroAction (IntroOrAndPattern (IntroAndPattern ((_,IntroNaming (IntroIdentifier id)) :: _ as l)))
when allow_conj -> (Some id,l)
| IntroAction (IntroOrAndPattern (IntroAndPattern _)) ->
if issimple then
@@ -384,7 +385,7 @@ let rewrite_equations as_mode othin neqns names ba =
Proofview.Goal.nf_enter { enter = begin fun gl ->
let (depids,nodepids) = split_dep_and_nodep ba.Tacticals.assums gl in
let first_eq = ref MoveLast in
- let avoid = if as_mode then List.map pi1 nodepids else [] in
+ let avoid = if as_mode then List.map get_id nodepids else [] in
match othin with
| Some thin ->
tclTHENLIST
@@ -399,11 +400,11 @@ let rewrite_equations as_mode othin neqns names ba =
(onLastHypId (fun id ->
tclTRY (projectAndApply as_mode thin avoid id first_eq names depids)))))
names;
- tclMAP (fun (id,_,_) -> tclIDTAC >>= fun () -> (* delay for [first_eq]. *)
- let idopt = if as_mode then Some id else None in
+ tclMAP (fun d -> tclIDTAC >>= fun () -> (* delay for [first_eq]. *)
+ let idopt = if as_mode then Some (get_id d) else None in
intro_move idopt (if thin then MoveLast else !first_eq))
nodepids;
- (tclMAP (fun (id,_,_) -> tclTRY (clear [id])) depids)]
+ (tclMAP (fun d -> tclTRY (clear [get_id d])) depids)]
| None ->
(* simple inversion *)
if as_mode then
diff --git a/tactics/leminv.ml b/tactics/leminv.ml
index cdf38ae46..70782ec64 100644
--- a/tactics/leminv.ml
+++ b/tactics/leminv.ml
@@ -27,6 +27,7 @@ open Tacticals.New
open Tactics
open Decl_kinds
open Proofview.Notations
+open Context.Named.Declaration
let no_inductive_inconstr env sigma constr =
(str "Cannot recognize an inductive predicate in " ++
@@ -117,11 +118,11 @@ let rec add_prods_sign env sigma t =
| Prod (na,c1,b) ->
let id = id_of_name_using_hdchar env t na in
let b'= subst1 (mkVar id) b in
- add_prods_sign (push_named (id,None,c1) env) sigma b'
+ add_prods_sign (push_named (LocalAssum (id,c1)) env) sigma b'
| LetIn (na,c1,t1,b) ->
let id = id_of_name_using_hdchar env t na in
let b'= subst1 (mkVar id) b in
- add_prods_sign (push_named (id,Some c1,t1) env) sigma b'
+ add_prods_sign (push_named (LocalDef (id,c1,t1)) env) sigma b'
| _ -> (env,t)
(* [dep_option] indicates whether the inversion lemma is dependent or not.
@@ -154,7 +155,8 @@ let compute_first_inversion_scheme env sigma ind sort dep_option =
let ivars = global_vars env i in
let revargs,ownsign =
fold_named_context
- (fun env (id,_,_ as d) (revargs,hyps) ->
+ (fun env d (revargs,hyps) ->
+ let id = get_id d in
if Id.List.mem id ivars then
((mkVar id)::revargs, Context.Named.add d hyps)
else
@@ -166,7 +168,7 @@ let compute_first_inversion_scheme env sigma ind sort dep_option =
(pty,goal)
in
let npty = nf_betadeltaiota env sigma pty in
- let extenv = push_named (p,None,npty) env in
+ let extenv = push_named (LocalAssum (p,npty)) env in
extenv, goal
(* [inversion_scheme sign I]
@@ -203,8 +205,8 @@ let inversion_scheme env sigma t sort dep_option inv_op =
let global_named_context = Global.named_context () in
let ownSign = ref begin
fold_named_context
- (fun env (id,_,_ as d) sign ->
- if mem_named_context id global_named_context then sign
+ (fun env d sign ->
+ if mem_named_context (get_id d) global_named_context then sign
else Context.Named.add d sign)
invEnv ~init:Context.Named.empty
end in
@@ -217,9 +219,9 @@ let inversion_scheme env sigma t sort dep_option inv_op =
let h = next_ident_away (Id.of_string "H") !avoid in
let ty,inst = Evarutil.generalize_evar_over_rels sigma (e,args) in
avoid := h::!avoid;
- ownSign := Context.Named.add (h,None,ty) !ownSign;
+ ownSign := Context.Named.add (LocalAssum (h,ty)) !ownSign;
applist (mkVar h, inst)
- | _ -> map_constr fill_holes c
+ | _ -> Constr.map fill_holes c
in
let c = fill_holes pfterm in
(* warning: side-effect on ownSign *)
diff --git a/tactics/rewrite.ml b/tactics/rewrite.ml
index 29002af9e..b39e34fc1 100644
--- a/tactics/rewrite.ml
+++ b/tactics/rewrite.ml
@@ -36,6 +36,7 @@ open Termops
open Libnames
open Sigma.Notations
open Proofview.Notations
+open Context.Named.Declaration
(** Typeclass-based generalized rewriting. *)
@@ -134,6 +135,7 @@ module GlobalBindings (M : sig
val arrow : evars -> evars * constr
end) = struct
open M
+ open Context.Rel.Declaration
let relation : evars -> evars * constr = find_global (fst relation) (snd relation)
let reflexive_type = find_global relation_classes "Reflexive"
@@ -219,8 +221,8 @@ end) = struct
let evars, newarg = app_poly env evars respectful [| ty ; b' ; relty ; arg |] in
evars, mkProd(na, ty, b), newarg, (ty, Some relty) :: cstrs
else
- let (evars, b, arg, cstrs) =
- aux (Environ.push_rel (na, None, ty) env) evars b cstrs
+ let (evars, b, arg, cstrs) =
+ aux (Environ.push_rel (LocalAssum (na, ty)) env) evars b cstrs
in
let ty = Reductionops.nf_betaiota (goalevars evars) ty in
let pred = mkLambda (na, ty, b) in
@@ -318,7 +320,7 @@ end) = struct
let evars, rb = aux evars env b' (pred n) in
app_poly env evars pointwise_relation [| ty; b'; rb |]
else
- let evars, rb = aux evars (Environ.push_rel (na, None, ty) env) b (pred n) in
+ let evars, rb = aux evars (Environ.push_rel (LocalAssum (na, ty)) env) b (pred n) in
app_poly env evars forall_relation
[| ty; mkLambda (na, ty, b); mkLambda (na, ty, rb) |]
| _ -> raise Not_found
@@ -469,6 +471,7 @@ let rec decompose_app_rel env evd t =
| _ -> error "Cannot find a relation to rewrite."
let decompose_applied_relation env sigma (c,l) =
+ let open Context.Rel.Declaration in
let ctype = Retyping.get_type_of env sigma c in
let find_rel ty =
let sigma, cl = Clenv.make_evar_clause env sigma ty in
@@ -491,7 +494,7 @@ let decompose_applied_relation env sigma (c,l) =
| Some c -> c
| None ->
let ctx,t' = Reductionops.splay_prod env sigma ctype in (* Search for underlying eq *)
- match find_rel (it_mkProd_or_LetIn t' (List.map (fun (n,t) -> n, None, t) ctx)) with
+ match find_rel (it_mkProd_or_LetIn t' (List.map (fun (n,t) -> LocalAssum (n, t)) ctx)) with
| Some c -> c
| None -> error "Cannot find an homogeneous relation to rewrite."
@@ -766,9 +769,9 @@ let resolve_morphism env avoid oldt m ?(fnewt=fun x -> x) args args' (b,cstr) ev
else TypeGlobal.do_subrelation, TypeGlobal.apply_subrelation
in
Environ.push_named
- (Id.of_string "do_subrelation",
- Some (snd (app_poly_sort b env evars dosub [||])),
- snd (app_poly_nocheck env evars appsub [||]))
+ (LocalDef (Id.of_string "do_subrelation",
+ snd (app_poly_sort b env evars dosub [||]),
+ snd (app_poly_nocheck env evars appsub [||])))
env
in
let evars, morph = new_cstr_evar evars env' app in
@@ -1110,8 +1113,9 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy =
(* | _ -> b') *)
| Lambda (n, t, b) when flags.under_lambdas ->
- let n' = name_app (fun id -> Tactics.fresh_id_in_env unfresh id env) n in
- let env' = Environ.push_rel (n', None, t) env in
+ let n' = name_app (fun id -> Tactics.fresh_id_in_env unfresh id env) n in
+ let open Context.Rel.Declaration in
+ let env' = Environ.push_rel (LocalAssum (n', t)) env in
let bty = Retyping.get_type_of env' (goalevars evars) b in
let unlift = if prop then PropGlobal.unlift_cstr else TypeGlobal.unlift_cstr in
let state, b' = s.strategy { state ; env = env' ; unfresh ;
@@ -1495,8 +1499,8 @@ let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : resul
(** Insert a declaration after the last declaration it depends on *)
let rec insert_dependent env decl accu hyps = match hyps with
| [] -> List.rev_append accu [decl]
-| (id, _, _ as ndecl) :: rem ->
- if occur_var_in_decl env id decl then
+| ndecl :: rem ->
+ if occur_var_in_decl env (get_id ndecl) decl then
List.rev_append accu (decl :: hyps)
else
insert_dependent env decl (ndecl :: accu) rem
@@ -1506,16 +1510,19 @@ let assert_replacing id newt tac =
let concl = Proofview.Goal.concl gl in
let env = Proofview.Goal.env gl in
let ctx = Environ.named_context env in
- let after, before = List.split_when (fun (n, b, t) -> Id.equal n id) ctx in
+ let after, before = List.split_when (Id.equal id % get_id) ctx in
let nc = match before with
| [] -> assert false
- | (id, b, _) :: rem -> insert_dependent env (id, None, newt) [] after @ rem
+ | d :: rem -> insert_dependent env (LocalAssum (get_id d, newt)) [] after @ rem
in
let env' = Environ.reset_with_named_context (val_of_named_context nc) env in
Proofview.Refine.refine ~unsafe:false { run = begin fun sigma ->
let Sigma (ev, sigma, p) = Evarutil.new_evar env' sigma concl in
let Sigma (ev', sigma, q) = Evarutil.new_evar env sigma newt in
- let map (n, _, _) = if Id.equal n id then ev' else mkVar n in
+ let map d =
+ let n = get_id d in
+ if Id.equal n id then ev' else mkVar n
+ in
let (e, _) = destEvar ev in
Sigma (mkEvar (e, Array.map_of_list map nc), sigma, p +> q)
end }
@@ -1543,7 +1550,7 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause =
assert_replacing id newt tac
| Some id, None ->
Proofview.Unsafe.tclEVARS undef <*>
- convert_hyp_no_check (id, None, newt)
+ convert_hyp_no_check (LocalAssum (id, newt))
| None, Some p ->
Proofview.Unsafe.tclEVARS undef <*>
Proofview.Goal.enter { enter = begin fun gl ->
@@ -2053,7 +2060,8 @@ let setoid_proof ty fn fallback =
try
let rel, _, _ = decompose_app_rel env sigma concl in
let evm = sigma in
- let car = pi3 (List.hd (fst (Reduction.dest_prod env (Typing.unsafe_type_of env evm rel)))) in
+ let open Context.Rel.Declaration in
+ let car = get_type (List.hd (fst (Reduction.dest_prod env (Typing.unsafe_type_of env evm rel)))) in
(try init_setoid () with _ -> raise Not_found);
fn env sigma car rel
with e -> Proofview.tclZERO e
diff --git a/tactics/tacinterp.ml b/tactics/tacinterp.ml
index 1112da4a0..30a9071fd 100644
--- a/tactics/tacinterp.ml
+++ b/tactics/tacinterp.ml
@@ -42,6 +42,7 @@ open Tacintern
open Taccoerce
open Sigma.Notations
open Proofview.Notations
+open Context.Named.Declaration
let has_type : type a. Val.t -> a typed_abstract_argument_type -> bool = fun v wit ->
let Val.Dyn (t, _) = v in
@@ -444,14 +445,13 @@ let interp_reference ist env sigma = function
try try_interp_ltac_var (coerce_to_reference env) ist (Some (env,sigma)) (loc, id)
with Not_found ->
try
- let (v, _, _) = Environ.lookup_named id env in
- VarRef v
+ VarRef (get_id (Environ.lookup_named id env))
with Not_found -> error_global_not_found_loc loc (qualid_of_ident id)
let try_interp_evaluable env (loc, id) =
let v = Environ.lookup_named id env in
match v with
- | (_, Some _, _) -> EvalVarRef id
+ | LocalDef _ -> EvalVarRef id
| _ -> error_not_evaluable (VarRef id)
let interp_evaluable ist env sigma = function
diff --git a/tactics/tactic_matching.ml b/tactics/tactic_matching.ml
index 80786058d..2144b75e7 100644
--- a/tactics/tactic_matching.ml
+++ b/tactics/tactic_matching.ml
@@ -11,6 +11,7 @@
open Names
open Tacexpr
+open Context.Named.Declaration
(** [t] is the type of matching successes. It ultimately contains a
{!Tacexpr.glob_tactic_expr} representing the left-hand side of the
@@ -278,9 +279,10 @@ module PatternMatching (E:StaticEnvironment) = struct
[hyps]. Tries the hypotheses in order. For each success returns
the name of the matched hypothesis. *)
let hyp_match_type hypname pat hyps =
- pick hyps >>= fun (id,b,hyp) ->
- let refresh = not (Option.is_empty b) in
- pattern_match_term refresh pat hyp () <*>
+ pick hyps >>= fun decl ->
+ let id = get_id decl in
+ let refresh = is_local_def decl in
+ pattern_match_term refresh pat (get_type decl) () <*>
put_terms (id_map_try_add_name hypname (Term.mkVar id) empty_term_subst) <*>
return id
@@ -290,12 +292,12 @@ module PatternMatching (E:StaticEnvironment) = struct
success returns the name of the matched hypothesis. *)
let hyp_match_body_and_type hypname bodypat typepat hyps =
pick hyps >>= function
- | (id,Some body,hyp) ->
+ | LocalDef (id,body,hyp) ->
pattern_match_term false bodypat body () <*>
pattern_match_term true typepat hyp () <*>
put_terms (id_map_try_add_name hypname (Term.mkVar id) empty_term_subst) <*>
return id
- | (id,None,hyp) -> fail
+ | LocalAssum (id,hyp) -> fail
(** [hyp_match pat hyps] dispatches to
{!hyp_match_type} or {!hyp_match_body_and_type} depending on whether
@@ -317,7 +319,7 @@ module PatternMatching (E:StaticEnvironment) = struct
(* spiwack: alternatively it is possible to return the list
with the matched hypothesis removed directly in
[hyp_match]. *)
- let select_matched_hyp (id,_,_) = Id.equal id matched_hyp in
+ let select_matched_hyp decl = Id.equal (get_id decl) matched_hyp in
let hyps = CList.remove_first select_matched_hyp hyps in
hyp_pattern_list_match pats hyps lhs
| [] -> return lhs
diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml
index d79de4913..7f904a561 100644
--- a/tactics/tacticals.ml
+++ b/tactics/tacticals.ml
@@ -16,6 +16,7 @@ open Declarations
open Tacmach
open Clenv
open Sigma.Notations
+open Context.Named.Declaration
(************************************************************************)
(* Tacticals re-exported from the Refiner module *)
@@ -69,7 +70,7 @@ let nthDecl m gl =
try List.nth (pf_hyps gl) (m-1)
with Failure _ -> error "No such assumption."
-let nthHypId m gl = pi1 (nthDecl m gl)
+let nthHypId m gl = nthDecl m gl |> get_id
let nthHyp m gl = mkVar (nthHypId m gl)
let lastDecl gl = nthDecl 1 gl
@@ -80,7 +81,7 @@ let nLastDecls n gl =
try List.firstn n (pf_hyps gl)
with Failure _ -> error "Not enough hypotheses in the goal."
-let nLastHypsId n gl = List.map pi1 (nLastDecls n gl)
+let nLastHypsId n gl = List.map get_id (nLastDecls n gl)
let nLastHyps n gl = List.map mkVar (nLastHypsId n gl)
let onNthDecl m tac gl = tac (nthDecl m gl) gl
@@ -98,7 +99,7 @@ let onNLastHypsId n tac = onHyps (nLastHypsId n) tac
let onNLastHyps n tac = onHyps (nLastHyps n) tac
let afterHyp id gl =
- fst (List.split_when (fun (hyp,_,_) -> Id.equal hyp id) (pf_hyps gl))
+ fst (List.split_when (Id.equal id % get_id) (pf_hyps gl))
(***************************************)
(* Clause Tacticals *)
@@ -552,8 +553,7 @@ module New = struct
let nthHypId m gl =
(** We only use [id] *)
let gl = Proofview.Goal.assume gl in
- let (id,_,_) = nthDecl m gl in
- id
+ nthDecl m gl |> get_id
let nthHyp m gl =
mkVar (nthHypId m gl)
@@ -585,7 +585,7 @@ module New = struct
let afterHyp id tac =
Proofview.Goal.nf_enter { enter = begin fun gl ->
let hyps = Proofview.Goal.hyps gl in
- let rem, _ = List.split_when (fun (hyp,_,_) -> Id.equal hyp id) hyps in
+ let rem, _ = List.split_when (Id.equal id % get_id) hyps in
tac rem
end }
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index aeb3726a0..8f30df5c0 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -161,19 +161,21 @@ let _ =
(** This tactic creates a partial proof realizing the introduction rule, but
does not check anything. *)
-let unsafe_intro env store (id, c, t) b =
+let unsafe_intro env store decl b =
+ let open Context.Named.Declaration in
Proofview.Refine.refine ~unsafe:true { run = begin fun sigma ->
let sigma = Sigma.to_evar_map sigma in
let ctx = named_context_val env in
- let nctx = push_named_context_val (id, c, t) ctx in
- let inst = List.map (fun (id, _, _) -> mkVar id) (named_context env) in
+ let nctx = push_named_context_val decl ctx in
+ let inst = List.map (mkVar % get_id) (named_context env) in
let ninst = mkRel 1 :: inst in
- let nb = subst1 (mkVar id) b in
+ let nb = subst1 (mkVar (get_id decl)) b in
let sigma, ev = new_evar_instance nctx sigma nb ~principal:true ~store ninst in
- Sigma.Unsafe.of_pair (mkNamedLambda_or_LetIn (id, c, t) ev, sigma)
+ Sigma.Unsafe.of_pair (mkNamedLambda_or_LetIn decl ev, sigma)
end }
let introduction ?(check=true) id =
+ let open Context.Named.Declaration in
Proofview.Goal.enter { enter = begin fun gl ->
let gl = Proofview.Goal.assume gl in
let concl = Proofview.Goal.concl gl in
@@ -186,8 +188,8 @@ let introduction ?(check=true) id =
(str "Variable " ++ pr_id id ++ str " is already declared.")
in
match kind_of_term (whd_evar sigma concl) with
- | Prod (_, t, b) -> unsafe_intro env store (id, None, t) b
- | LetIn (_, c, t, b) -> unsafe_intro env store (id, Some c, t) b
+ | Prod (_, t, b) -> unsafe_intro env store (LocalAssum (id, t)) b
+ | LetIn (_, c, t, b) -> unsafe_intro env store (LocalDef (id, c, t)) b
| _ -> raise (RefinerError IntroNeedsProduct)
end }
@@ -296,6 +298,7 @@ let move_hyp id dest gl = Tacmach.move_hyp id dest gl
(* Renaming hypotheses *)
let rename_hyp repl =
+ let open Context.Named.Declaration in
let fold accu (src, dst) = match accu with
| None -> None
| Some (srcs, dsts) ->
@@ -317,7 +320,7 @@ let rename_hyp repl =
let concl = Proofview.Goal.concl gl in
let store = Proofview.Goal.extra gl in
(** Check that we do not mess variables *)
- let fold accu (id, _, _) = Id.Set.add id accu in
+ let fold accu decl = Id.Set.add (get_id decl) accu in
let vars = List.fold_left fold Id.Set.empty hyps in
let () =
if not (Id.Set.subset src vars) then
@@ -335,14 +338,14 @@ let rename_hyp repl =
let make_subst (src, dst) = (src, mkVar dst) in
let subst = List.map make_subst repl in
let subst c = Vars.replace_vars subst c in
- let map (id, body, t) =
- let id = try List.assoc_f Id.equal id repl with Not_found -> id in
- (id, Option.map subst body, subst t)
+ let map decl =
+ decl |> map_id (fun id -> try List.assoc_f Id.equal id repl with Not_found -> id)
+ |> map_constr subst
in
let nhyps = List.map map hyps in
let nconcl = subst concl in
let nctx = Environ.val_of_named_context nhyps in
- let instance = List.map (fun (id, _, _) -> mkVar id) hyps in
+ let instance = List.map (mkVar % get_id) hyps in
Proofview.Refine.refine ~unsafe:true { run = begin fun sigma ->
let sigma = Sigma.to_evar_map sigma in
let (sigma, c) = Evarutil.new_evar_instance nctx sigma nconcl ~store instance in
@@ -370,11 +373,13 @@ let id_of_name_with_default id = function
let default_id_of_sort s =
if Sorts.is_small s then default_small_ident else default_type_ident
-let default_id env sigma = function
- | (name,None,t) ->
+let default_id env sigma decl =
+ let open Context.Rel.Declaration in
+ match decl with
+ | LocalAssum (name,t) ->
let dft = default_id_of_sort (Retyping.get_sort_of env sigma t) in
id_of_name_with_default dft name
- | (name,Some b,_) -> id_of_name_using_hdchar env b name
+ | LocalDef (name,b,_) -> id_of_name_using_hdchar env b name
(* Non primitive introduction tactics are treated by intro_then_gen
There is possibly renaming, with possibly names to avoid and
@@ -409,8 +414,9 @@ let find_name mayrepl decl naming gl = match naming with
(**************************************************************)
let assert_before_then_gen b naming t tac =
+ let open Context.Rel.Declaration in
Proofview.Goal.enter { enter = begin fun gl ->
- let id = find_name b (Anonymous,None,t) naming gl in
+ let id = find_name b (LocalAssum (Anonymous,t)) naming gl in
Tacticals.New.tclTHENLAST
(Proofview.V82.tactic
(fun gl ->
@@ -427,8 +433,9 @@ let assert_before na = assert_before_gen false (naming_of_name na)
let assert_before_replacing id = assert_before_gen true (NamingMustBe (dloc,id))
let assert_after_then_gen b naming t tac =
+ let open Context.Rel.Declaration in
Proofview.Goal.enter { enter = begin fun gl ->
- let id = find_name b (Anonymous,None,t) naming gl in
+ let id = find_name b (LocalAssum (Anonymous,t)) naming gl in
Tacticals.New.tclTHENFIRST
(Proofview.V82.tactic
(fun gl ->
@@ -472,17 +479,18 @@ let cofix ido gl = match ido with
type tactic_reduction = env -> evar_map -> constr -> constr
-let pf_reduce_decl redfun where (id,c,ty) gl =
+let pf_reduce_decl redfun where decl gl =
+ let open Context.Named.Declaration in
let redfun' = Tacmach.pf_reduce redfun gl in
- match c with
- | None ->
+ match decl with
+ | LocalAssum (id,ty) ->
if where == InHypValueOnly then
errorlabstrm "" (pr_id id ++ str " has no value.");
- (id,None,redfun' ty)
- | Some b ->
+ LocalAssum (id,redfun' ty)
+ | LocalDef (id,b,ty) ->
let b' = if where != InHypTypeOnly then redfun' b else b in
let ty' = if where != InHypValueOnly then redfun' ty else ty in
- (id,Some b',ty')
+ LocalDef (id,b',ty')
(* Possibly equip a reduction with the occurrences mentioned in an
occurrence clause *)
@@ -568,19 +576,20 @@ let reduct_option ?(check=false) redfun = function
(** Tactic reduction modulo evars (for universes essentially) *)
-let pf_e_reduce_decl redfun where (id,c,ty) gl =
+let pf_e_reduce_decl redfun where decl gl =
+ let open Context.Named.Declaration in
let sigma = project gl in
let redfun = redfun (pf_env gl) in
- match c with
- | None ->
+ match decl with
+ | LocalAssum (id,ty) ->
if where == InHypValueOnly then
errorlabstrm "" (pr_id id ++ str " has no value.");
let sigma, ty' = redfun sigma ty in
- sigma, (id,None,ty')
- | Some b ->
+ sigma, LocalAssum (id,ty')
+ | LocalDef (id,b,ty) ->
let sigma, b' = if where != InHypTypeOnly then redfun sigma b else sigma, b in
let sigma, ty' = if where != InHypValueOnly then redfun sigma ty else sigma, ty in
- sigma, (id,Some b',ty')
+ sigma, LocalDef (id,b',ty')
let e_reduct_in_concl (redfun,sty) gl =
Proofview.V82.of_tactic
@@ -609,21 +618,22 @@ let e_change_in_concl (redfun,sty) =
Sigma.Unsafe.of_pair (convert_concl_no_check c sty, sigma)
end }
-let e_pf_change_decl (redfun : bool -> e_reduction_function) where (id,c,ty) env sigma =
- match c with
- | None ->
+let e_pf_change_decl (redfun : bool -> e_reduction_function) where decl env sigma =
+ let open Context.Named.Declaration in
+ match decl with
+ | LocalAssum (id,ty) ->
if where == InHypValueOnly then
errorlabstrm "" (pr_id id ++ str " has no value.");
let sigma',ty' = redfun false env sigma ty in
- sigma', (id,None,ty')
- | Some b ->
+ sigma', LocalAssum (id,ty')
+ | LocalDef (id,b,ty) ->
let sigma',b' =
if where != InHypTypeOnly then redfun true env sigma b else sigma, b
in
let sigma',ty' =
if where != InHypValueOnly then redfun false env sigma' ty else sigma', ty
in
- sigma', (id,Some b',ty')
+ sigma', LocalDef (id,b',ty')
let e_change_in_hyp redfun (id,where) =
Proofview.Goal.s_enter { s_enter = begin fun gl ->
@@ -767,10 +777,9 @@ let unfold_constr = function
let find_intro_names ctxt gl =
let _, res = List.fold_right
(fun decl acc ->
- let wantedname,x,typdecl = decl in
let env,idl = acc in
let name = fresh_id idl (default_id env gl.sigma decl) gl in
- let newenv = push_rel (wantedname,x,typdecl) env in
+ let newenv = push_rel decl env in
(newenv,(name::idl)))
ctxt (pf_env gl , []) in
List.rev res
@@ -782,15 +791,16 @@ let build_intro_tac id dest tac = match dest with
Proofview.V82.tactic (move_hyp id dest); tac id]
let rec intro_then_gen name_flag move_flag force_flag dep_flag tac =
+ let open Context.Rel.Declaration in
Proofview.Goal.enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in
let concl = nf_evar (Tacmach.New.project gl) concl in
match kind_of_term concl with
| Prod (name,t,u) when not dep_flag || (dependent (mkRel 1) u) ->
- let name = find_name false (name,None,t) name_flag gl in
+ let name = find_name false (LocalAssum (name,t)) name_flag gl in
build_intro_tac name move_flag tac
| LetIn (name,b,t,u) when not dep_flag || (dependent (mkRel 1) u) ->
- let name = find_name false (name,Some b,t) name_flag gl in
+ let name = find_name false (LocalDef (name,b,t)) name_flag gl in
build_intro_tac name move_flag tac
| _ ->
begin if not force_flag then Proofview.tclZERO (RefinerError IntroNeedsProduct)
@@ -853,21 +863,24 @@ let intro_forthcoming_then_gen name_flag move_flag dep_flag n bound tac =
aux n []
let get_next_hyp_position id gl =
+ let open Context.Named.Declaration in
let rec aux = function
| [] -> raise (RefinerError (NoSuchHyp id))
- | (hyp,_,_) :: right ->
- if Id.equal hyp id then
- match right with (id,_,_)::_ -> MoveBefore id | [] -> MoveLast
+ | decl :: right ->
+ if Id.equal (get_id decl) id then
+ match right with decl::_ -> MoveBefore (get_id decl) | [] -> MoveLast
else
aux right
in
aux (Proofview.Goal.hyps (Proofview.Goal.assume gl))
let get_previous_hyp_position id gl =
+ let open Context.Named.Declaration in
let rec aux dest = function
| [] -> raise (RefinerError (NoSuchHyp id))
- | (hyp,_,_) :: right ->
- if Id.equal hyp id then dest else aux (MoveAfter hyp) right
+ | decl :: right ->
+ let hyp = get_id decl in
+ if Id.equal hyp id then dest else aux (MoveAfter hyp) right
in
aux MoveLast (Proofview.Goal.hyps (Proofview.Goal.assume gl))
@@ -1146,6 +1159,7 @@ let index_of_ind_arg t =
in aux None 0 t
let enforce_prop_bound_names rename tac =
+ let open Context.Rel.Declaration in
match rename with
| Some (isrec,nn) when Namegen.use_h_based_elimination_names () ->
(* Rename dependent arguments in Prop with name "H" *)
@@ -1165,11 +1179,11 @@ let enforce_prop_bound_names rename tac =
Name (add_suffix Namegen.default_prop_ident s)
else
na in
- mkProd (na,t,aux (push_rel (na,None,t) env) sigma (i-1) t')
+ mkProd (na,t,aux (push_rel (LocalAssum (na,t)) env) sigma (i-1) t')
| Prod (Anonymous,t,t') ->
- mkProd (Anonymous,t,aux (push_rel (Anonymous,None,t) env) sigma (i-1) t')
+ mkProd (Anonymous,t,aux (push_rel (LocalAssum (Anonymous,t)) env) sigma (i-1) t')
| LetIn (na,c,t,t') ->
- mkLetIn (na,c,t,aux (push_rel (na,Some c,t) env) sigma (i-1) t')
+ mkLetIn (na,c,t,aux (push_rel (LocalDef (na,c,t)) env) sigma (i-1) t')
| _ -> print_int i; Pp.msg (print_constr t); assert false in
let rename_branch i =
Proofview.Goal.nf_enter { enter = begin fun gl ->
@@ -1391,11 +1405,13 @@ type conjunction_status =
| NotADefinedRecordUseScheme of constr
let make_projection env sigma params cstr sign elim i n c u =
+ let open Context.Rel.Declaration in
let elim = match elim with
| NotADefinedRecordUseScheme elim ->
(* bugs: goes from right to left when i increases! *)
- let (na,b,t) = List.nth cstr.cs_args i in
- let b = match b with None -> mkRel (i+1) | Some b -> b in
+ let decl = List.nth cstr.cs_args i in
+ let t = get_type decl in
+ let b = match decl with LocalAssum _ -> mkRel (i+1) | LocalDef (_,b,_) -> b in
let branch = it_mkLambda_or_LetIn b cstr.cs_args in
if
(* excludes dependent projection types *)
@@ -1651,6 +1667,7 @@ let apply_in_once_main flags innerclause env sigma (d,lbind) =
let apply_in_once sidecond_first with_delta with_destruct with_evars naming
id (clear_flag,(loc,(d,lbind))) tac =
+ let open Context.Rel.Declaration in
Proofview.Goal.nf_enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
@@ -1658,7 +1675,7 @@ let apply_in_once sidecond_first with_delta with_destruct with_evars naming
if with_delta then default_unify_flags () else default_no_delta_unify_flags () in
let t' = Tacmach.New.pf_get_hyp_typ id gl in
let innerclause = mk_clenv_from_env env sigma (Some 0) (mkVar id,t') in
- let targetid = find_name true (Anonymous,None,t') naming gl in
+ let targetid = find_name true (LocalAssum (Anonymous,t')) naming gl in
let rec aux idstoclear with_destruct c =
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
@@ -1770,13 +1787,15 @@ let exact_proof c gl =
in tclTHEN (tclEVARUNIVCONTEXT ctx) (Tacmach.refine_no_check c) gl
let assumption =
+ let open Context.Named.Declaration in
let rec arec gl only_eq = function
| [] ->
if only_eq then
let hyps = Proofview.Goal.hyps gl in
arec gl false hyps
else Tacticals.New.tclZEROMSG (str "No such assumption.")
- | (id, c, t)::rest ->
+ | decl::rest ->
+ let t = get_type decl in
let concl = Proofview.Goal.concl gl in
let sigma = Tacmach.New.project gl in
let (sigma, is_same_type) =
@@ -1787,7 +1806,7 @@ let assumption =
in
if is_same_type then
(Proofview.Unsafe.tclEVARS sigma) <*>
- Proofview.Refine.refine ~unsafe:true { run = fun h -> Sigma.here (mkVar id) h }
+ Proofview.Refine.refine ~unsafe:true { run = fun h -> Sigma.here (mkVar (get_id decl)) h }
else arec gl only_eq rest
in
let assumption_tac = { enter = begin fun gl ->
@@ -1822,40 +1841,43 @@ let check_is_type env ty msg =
with e when Errors.noncritical e ->
msg e
-let check_decl env (_, c, ty) msg =
+let check_decl env decl msg =
+ let open Context.Named.Declaration in
+ let ty = get_type decl in
Proofview.tclEVARMAP >>= fun sigma ->
let evdref = ref sigma in
try
let _ = Typing.sort_of env evdref ty in
- let _ = match c with
- | None -> ()
- | Some c -> Typing.check env evdref c ty
+ let _ = match decl with
+ | LocalAssum _ -> ()
+ | LocalDef (_,c,_) -> Typing.check env evdref c ty
in
Proofview.Unsafe.tclEVARS !evdref
with e when Errors.noncritical e ->
msg e
let clear_body ids =
+ let open Context.Named.Declaration in
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in
let ctx = named_context env in
- let map (id, body, t as decl) = match body with
- | None ->
+ let map = function
+ | LocalAssum (id,t) as decl ->
let () = if List.mem_f Id.equal id ids then
errorlabstrm "" (str "Hypothesis " ++ pr_id id ++ str " is not a local definition")
in
decl
- | Some _ ->
- if List.mem_f Id.equal id ids then (id, None, t) else decl
+ | LocalDef (id,_,t) as decl ->
+ if List.mem_f Id.equal id ids then LocalAssum (id, t) else decl
in
let ctx = List.map map ctx in
let base_env = reset_context env in
let env = push_named_context ctx base_env in
let check_hyps =
- let check env (id, _, _ as decl) =
+ let check env decl =
let msg _ = Tacticals.New.tclZEROMSG
- (str "Hypothesis " ++ pr_id id ++ on_the_bodies ids)
+ (str "Hypothesis " ++ pr_id (get_id decl) ++ on_the_bodies ids)
in
check_decl env decl msg <*> Proofview.tclUNIT (push_named decl env)
in
@@ -1897,11 +1919,13 @@ let rec intros_clearing = function
(* Keeping only a few hypotheses *)
let keep hyps =
+ let open Context.Named.Declaration in
Proofview.Goal.nf_enter { enter = begin fun gl ->
Proofview.tclENV >>= fun env ->
let ccl = Proofview.Goal.concl gl in
let cl,_ =
- fold_named_context_reverse (fun (clear,keep) (hyp,_,_ as decl) ->
+ fold_named_context_reverse (fun (clear,keep) decl ->
+ let hyp = get_id decl in
if Id.List.mem hyp hyps
|| List.exists (occur_var_in_decl env hyp) keep
|| occur_var env hyp ccl
@@ -2442,20 +2466,24 @@ let letin_tac_gen with_eq (id,depdecls,lastlhyp,ccl,c) ty =
end }
let insert_before decls lasthyp env =
+ let open Context.Named.Declaration in
match lasthyp with
| None -> push_named_context decls env
| Some id ->
Environ.fold_named_context
- (fun _ (id',_,_ as d) env ->
- let env = if Id.equal id id' then push_named_context decls env else env in
+ (fun _ d env ->
+ let env = if Id.equal id (get_id d) then push_named_context decls env else env in
push_named d env)
~init:(reset_context env) env
(* unsafe *)
let mkletin_goal env sigma store with_eq dep (id,lastlhyp,ccl,c) ty =
- let body = if dep then Some c else None in
+ let open Context.Named.Declaration in
let t = match ty with Some t -> t | _ -> typ_of env sigma c in
+ let decl = if dep then LocalDef (id,c,t)
+ else LocalAssum (id,t)
+ in
match with_eq with
| Some (lr,(loc,ido)) ->
let heq = match ido with
@@ -2471,11 +2499,11 @@ let mkletin_goal env sigma store with_eq dep (id,lastlhyp,ccl,c) ty =
let Sigma (refl, sigma, q) = Sigma.fresh_global env sigma eqdata.refl in
let eq = applist (eq,args) in
let refl = applist (refl, [t;mkVar id]) in
- let newenv = insert_before [heq,None,eq;id,body,t] lastlhyp env in
+ let newenv = insert_before [LocalAssum (heq,eq); decl] lastlhyp env in
let Sigma (x, sigma, r) = new_evar newenv sigma ~principal:true ~store ccl in
Sigma (mkNamedLetIn id c t (mkNamedLetIn heq refl eq x), sigma, p +> q +> r)
| None ->
- let newenv = insert_before [id,body,t] lastlhyp env in
+ let newenv = insert_before [decl] lastlhyp env in
let Sigma (x, sigma, p) = new_evar newenv sigma ~principal:true ~store ccl in
Sigma (mkNamedLetIn id c t x, sigma, p)
@@ -2557,12 +2585,17 @@ let generalized_name c t ids cl = function
but only those at [occs] in [T] *)
let generalize_goal_gen env sigma ids i ((occs,c,b),na) t cl =
+ let open Context.Rel.Declaration in
let decls,cl = decompose_prod_n_assum i cl in
let dummy_prod = it_mkProd_or_LetIn mkProp decls in
let newdecls,_ = decompose_prod_n_assum i (subst_term_gen eq_constr_nounivs c dummy_prod) in
let cl',sigma' = subst_closed_term_occ env sigma (AtOccs occs) c (it_mkProd_or_LetIn cl newdecls) in
let na = generalized_name c t ids cl' na in
- mkProd_or_LetIn (na,b,t) cl', sigma'
+ let decl = match b with
+ | None -> LocalAssum (na,t)
+ | Some b -> LocalDef (na,b,t)
+ in
+ mkProd_or_LetIn decl cl', sigma'
let generalize_goal gl i ((occs,c,b),na as o) (cl,sigma) =
let env = Tacmach.pf_env gl in
@@ -2571,18 +2604,19 @@ let generalize_goal gl i ((occs,c,b),na as o) (cl,sigma) =
generalize_goal_gen env sigma ids i o t cl
let generalize_dep ?(with_let=false) c gl =
+ let open Context.Named.Declaration in
let env = pf_env gl in
let sign = pf_hyps gl in
let init_ids = ids_of_named_context (Global.named_context()) in
- let seek d toquant =
- if List.exists (fun (id,_,_) -> occur_var_in_decl env id d) toquant
+ let seek (d:Context.Named.Declaration.t) (toquant:Context.Named.t) =
+ if List.exists (fun d' -> occur_var_in_decl env (get_id d') d) toquant
|| dependent_in_decl c d then
d::toquant
else
toquant in
let to_quantify = Context.Named.fold_outside seek sign ~init:[] in
let to_quantify_rev = List.rev to_quantify in
- let qhyps = List.map (fun (id,_,_) -> id) to_quantify_rev in
+ let qhyps = List.map get_id to_quantify_rev in
let tothin = List.filter (fun id -> not (Id.List.mem id init_ids)) qhyps in
let tothin' =
match kind_of_term c with
@@ -2594,7 +2628,7 @@ let generalize_dep ?(with_let=false) c gl =
let body =
if with_let then
match kind_of_term c with
- | Var id -> pi2 (Tacmach.pf_get_hyp gl id)
+ | Var id -> Tacmach.pf_get_hyp gl id |> get_value
| _ -> None
else None
in
@@ -2720,14 +2754,15 @@ let specialize (c,lbind) =
(* The two following functions should already exist, but found nowhere *)
(* Unfolds x by its definition everywhere *)
let unfold_body x gl =
+ let open Context.Named.Declaration in
let hyps = pf_hyps gl in
let xval =
match Context.Named.lookup x hyps with
- (_,Some xval,_) -> xval
+ | LocalDef (_,xval,_) -> xval
| _ -> errorlabstrm "unfold_body"
(pr_id x ++ str" is not a defined hypothesis.") in
let aft = afterHyp x gl in
- let hl = List.fold_right (fun (y,yval,_) cl -> (y,InHyp) :: cl) aft [] in
+ let hl = List.fold_right (fun decl cl -> (get_id decl, InHyp) :: cl) aft [] in
let xvar = mkVar x in
let rfun _ _ c = replace_term xvar xval c in
tclTHENLIST
@@ -3041,6 +3076,7 @@ exception Shunt of Id.t move_location
let cook_sign hyp0_opt inhyps indvars env =
(* First phase from L to R: get [toclear], [decldep] and [statuslist]
for the hypotheses before (= more ancient than) hyp0 (see above) *)
+ let open Context.Named.Declaration in
let toclear = ref [] in
let avoid = ref [] in
let decldeps = ref [] in
@@ -3049,7 +3085,8 @@ let cook_sign hyp0_opt inhyps indvars env =
let lstatus = ref [] in
let before = ref true in
let maindep = ref false in
- let seek_deps env (hyp,_,_ as decl) rhyp =
+ let seek_deps env decl rhyp =
+ let hyp = get_id decl in
if (match hyp0_opt with Some hyp0 -> Id.equal hyp hyp0 | _ -> false)
then begin
before:=false;
@@ -3068,7 +3105,7 @@ let cook_sign hyp0_opt inhyps indvars env =
in
let depother = List.is_empty inhyps &&
(List.exists (fun id -> occur_var_in_decl env id decl) indvars ||
- List.exists (fun (id,_,_) -> occur_var_in_decl env id decl) !decldeps)
+ List.exists (fun decl' -> occur_var_in_decl env (get_id decl') decl) !decldeps)
in
if not (List.is_empty inhyps) && Id.List.mem hyp inhyps
|| dephyp0 || depother
@@ -3090,7 +3127,8 @@ let cook_sign hyp0_opt inhyps indvars env =
in
let _ = fold_named_context seek_deps env ~init:MoveFirst in
(* 2nd phase from R to L: get left hyp of [hyp0] and [lhyps] *)
- let compute_lstatus lhyp (hyp,_,_) =
+ let compute_lstatus lhyp decl =
+ let hyp = get_id decl in
if (match hyp0_opt with Some hyp0 -> Id.equal hyp hyp0 | _ -> false) then
raise (Shunt lhyp);
if Id.List.mem hyp !ldeps then begin
@@ -3280,6 +3318,7 @@ let mk_term_eq env sigma ty t ty' t' =
mkHEq ty t ty' t', mkHRefl ty' t'
let make_abstract_generalize env id typ concl dep ctx body c eqs args refls =
+ let open Context.Rel.Declaration in
Proofview.Refine.refine { run = begin fun sigma ->
let eqslen = List.length eqs in
(* Abstract by the "generalized" hypothesis equality proof if necessary. *)
@@ -3291,9 +3330,13 @@ let make_abstract_generalize env id typ concl dep ctx body c eqs args refls =
in
(* Abstract by equalities *)
let eqs = lift_togethern 1 eqs in (* lift together and past genarg *)
- let abseqs = it_mkProd_or_LetIn (lift eqslen abshypeq) (List.map (fun x -> (Anonymous, None, x)) eqs) in
+ let abseqs = it_mkProd_or_LetIn (lift eqslen abshypeq) (List.map (fun x -> LocalAssum (Anonymous, x)) eqs) in
+ let decl = match body with
+ | None -> LocalAssum (Name id, c)
+ | Some body -> LocalDef (Name id, body, c)
+ in
(* Abstract by the "generalized" hypothesis. *)
- let genarg = mkProd_or_LetIn (Name id, body, c) abseqs in
+ let genarg = mkProd_or_LetIn decl abseqs in
(* Abstract by the extension of the context *)
let genctyp = it_mkProd_or_LetIn genarg ctx in
(* The goal will become this product. *)
@@ -3309,11 +3352,13 @@ let make_abstract_generalize env id typ concl dep ctx body c eqs args refls =
end }
let hyps_of_vars env sign nogen hyps =
+ let open Context.Named.Declaration in
if Id.Set.is_empty hyps then []
else
let (_,lh) =
Context.Named.fold_inside
- (fun (hs,hl) (x,_,_ as d) ->
+ (fun (hs,hl) d ->
+ let x = get_id d in
if Id.Set.mem x nogen then (hs,hl)
else if Id.Set.mem x hs then (hs,x::hl)
else
@@ -3342,11 +3387,12 @@ let linear vars args =
true
with Seen -> false
-let is_defined_variable env id = match lookup_named id env with
-| (_, None, _) -> false
-| (_, Some _, _) -> true
+let is_defined_variable env id =
+ let open Context.Named.Declaration in
+ lookup_named id env |> is_local_def
let abstract_args gl generalize_vars dep id defined f args =
+ let open Context.Rel.Declaration in
let sigma = ref (Tacmach.project gl) in
let env = Tacmach.pf_env gl in
let concl = Tacmach.pf_concl gl in
@@ -3363,9 +3409,10 @@ let abstract_args gl generalize_vars dep id defined f args =
eqs are not lifted w.r.t. each other yet. (* will be needed when going to dependent indexes *)
*)
let aux (prod, ctx, ctxenv, c, args, eqs, refls, nongenvars, vars, env) arg =
- let (name, _, ty), arity =
+ let name, ty, arity =
let rel, c = Reductionops.splay_prod_n env !sigma 1 prod in
- List.hd rel, c
+ let decl = List.hd rel in
+ get_name decl, get_type decl, c
in
let argty = Tacmach.pf_unsafe_type_of gl arg in
let sigma', ty = Evarsolve.refresh_universes (Some true) env !sigma ty in
@@ -3379,7 +3426,7 @@ let abstract_args gl generalize_vars dep id defined f args =
Id.Set.add id nongenvars, Id.Set.remove id vars, env)
| _ ->
let name = get_id name in
- let decl = (Name name, None, ty) in
+ let decl = LocalAssum (Name name, ty) in
let ctx = decl :: ctx in
let c' = mkApp (lift 1 c, [|mkRel 1|]) in
let args = arg :: args in
@@ -3430,15 +3477,15 @@ let abstract_args gl generalize_vars dep id defined f args =
else None
let abstract_generalize ?(generalize_vars=true) ?(force_dep=false) id =
+ let open Context.Named.Declaration in
Proofview.Goal.nf_enter { enter = begin fun gl ->
Coqlib.check_required_library Coqlib.jmeq_module_name;
let (f, args, def, id, oldid) =
let oldid = Tacmach.New.pf_get_new_id id gl in
- let (_, b, t) = Tacmach.New.pf_get_hyp id gl in
- match b with
- | None -> let f, args = decompose_app t in
+ match Tacmach.New.pf_get_hyp id gl with
+ | LocalAssum (_,t) -> let f, args = decompose_app t in
(f, args, false, id, oldid)
- | Some t ->
+ | LocalDef (_,t,_) ->
let f, args = decompose_app t in
(f, args, true, id, oldid)
in
@@ -3473,6 +3520,7 @@ let rec compare_upto_variables x y =
else compare_constr compare_upto_variables x y
let specialize_eqs id gl =
+ let open Context.Rel.Declaration in
let env = Tacmach.pf_env gl in
let ty = Tacmach.pf_get_hyp_typ gl id in
let evars = ref (project gl) in
@@ -3501,15 +3549,14 @@ let specialize_eqs id gl =
if in_eqs then acc, in_eqs, ctx, ty
else
let e = e_new_evar (push_rel_context ctx env) evars t in
- aux false ((na, Some e, t) :: ctx) (mkApp (lift 1 acc, [| mkRel 1 |])) b)
+ aux false (LocalDef (na,e,t) :: ctx) (mkApp (lift 1 acc, [| mkRel 1 |])) b)
| t -> acc, in_eqs, ctx, ty
in
let acc, worked, ctx, ty = aux false [] (mkVar id) ty in
let ctx' = nf_rel_context_evar !evars ctx in
- let ctx'' = List.map (fun (n,b,t as decl) ->
- match b with
- | Some k when isEvar k -> (n,None,t)
- | b -> decl) ctx'
+ let ctx'' = List.map (function
+ | LocalDef (n,k,t) when isEvar k -> LocalAssum (n,t)
+ | decl -> decl) ctx'
in
let ty' = it_mkProd_or_LetIn ty ctx'' in
let acc' = it_mkLambda_or_LetIn acc ctx'' in
@@ -3543,18 +3590,19 @@ let occur_rel n c =
We also return the conclusion.
*)
let decompose_paramspred_branch_args elimt =
- let rec cut_noccur elimt acc2 : Context.Rel.t * Context.Rel.t * types =
+ let open Context.Rel.Declaration in
+ let rec cut_noccur elimt acc2 =
match kind_of_term elimt with
| Prod(nme,tpe,elimt') ->
let hd_tpe,_ = decompose_app ((strip_prod_assum tpe)) in
if not (occur_rel 1 elimt') && isRel hd_tpe
- then cut_noccur elimt' ((nme,None,tpe)::acc2)
+ then cut_noccur elimt' (LocalAssum (nme,tpe)::acc2)
else let acc3,ccl = decompose_prod_assum elimt in acc2 , acc3 , ccl
| App(_, _) | Rel _ -> acc2 , [] , elimt
| _ -> error_ind_scheme "" in
- let rec cut_occur elimt acc1 : Context.Rel.t * Context.Rel.t * Context.Rel.t * types =
+ let rec cut_occur elimt acc1 =
match kind_of_term elimt with
- | Prod(nme,tpe,c) when occur_rel 1 c -> cut_occur c ((nme,None,tpe)::acc1)
+ | Prod(nme,tpe,c) when occur_rel 1 c -> cut_occur c (LocalAssum (nme,tpe)::acc1)
| Prod(nme,tpe,c) -> let acc2,acc3,ccl = cut_noccur elimt [] in acc1,acc2,acc3,ccl
| App(_, _) | Rel _ -> acc1,[],[],elimt
| _ -> error_ind_scheme "" in
@@ -3596,6 +3644,7 @@ let exchange_hd_app subst_hd t =
- finish to fill in the elim_scheme: indarg/farg/args and finally indref. *)
let compute_elim_sig ?elimc elimt =
+ let open Context.Rel.Declaration in
let params_preds,branches,args_indargs,conclusion =
decompose_paramspred_branch_args elimt in
@@ -3629,8 +3678,8 @@ let compute_elim_sig ?elimc elimt =
(* 3- Look at last arg: is it the indarg? *)
ignore (
match List.hd args_indargs with
- | hiname,Some _,hi -> error_ind_scheme ""
- | hiname,None,hi ->
+ | LocalDef (hiname,_,hi) -> error_ind_scheme ""
+ | LocalAssum (hiname,hi) ->
let hi_ind, hi_args = decompose_app hi in
let hi_is_ind = (* hi est d'un type globalisable *)
match kind_of_term hi_ind with
@@ -3654,24 +3703,25 @@ let compute_elim_sig ?elimc elimt =
with Exit -> (* Ending by computing indref: *)
match !res.indarg with
| None -> !res (* No indref *)
- | Some ( _,Some _,_) -> error_ind_scheme ""
- | Some ( _,None,ind) ->
+ | Some (LocalDef _) -> error_ind_scheme ""
+ | Some (LocalAssum (_,ind)) ->
let indhd,indargs = decompose_app ind in
try {!res with indref = Some (global_of_constr indhd) }
with e when Errors.noncritical e ->
error "Cannot find the inductive type of the inductive scheme."
let compute_scheme_signature scheme names_info ind_type_guess =
+ let open Context.Rel.Declaration in
let f,l = decompose_app scheme.concl in
(* Vérifier que les arguments de Qi sont bien les xi. *)
let cond, check_concl =
match scheme.indarg with
- | Some (_,Some _,_) ->
+ | Some (LocalDef _) ->
error "Strange letin, cannot recognize an induction scheme."
| None -> (* Non standard scheme *)
let cond hd = Term.eq_constr hd ind_type_guess && not scheme.farg_in_concl
in (cond, fun _ _ -> ())
- | Some ( _,None,ind) -> (* Standard scheme from an inductive type *)
+ | Some (LocalAssum (_,ind)) -> (* Standard scheme from an inductive type *)
let indhd,indargs = decompose_app ind in
let cond hd = Term.eq_constr hd indhd in
let check_concl is_pred p =
@@ -3703,7 +3753,7 @@ let compute_scheme_signature scheme names_info ind_type_guess =
in
let rec find_branches p lbrch =
match lbrch with
- | (_,None,t)::brs ->
+ | LocalAssum (_,t) :: brs ->
(try
let lchck_brch = check_branch p t in
let n = List.fold_left
@@ -3716,7 +3766,7 @@ let compute_scheme_signature scheme names_info ind_type_guess =
lchck_brch in
(avoid,namesign) :: find_branches (p+1) brs
with Exit-> error_ind_scheme "the branches of")
- | (_,Some _,_)::_ -> error_ind_scheme "the branches of"
+ | LocalDef _ :: _ -> error_ind_scheme "the branches of"
| [] -> check_concl is_pred p; []
in
Array.of_list (find_branches 0 (List.rev scheme.branches))
@@ -3797,13 +3847,15 @@ let is_functional_induction elimc gl =
(* Wait the last moment to guess the eliminator so as to know if we
need a dependent one or not *)
-let get_eliminator elim dep s gl = match elim with
+let get_eliminator elim dep s gl =
+ let open Context.Rel.Declaration in
+ match elim with
| ElimUsing (elim,indsign) ->
Tacmach.New.project gl, (* bugged, should be computed *) true, elim, indsign
| ElimOver (isrec,id) ->
let evd, (elimc,elimt),_ as elims = guess_elim isrec dep s id gl in
let _, (l, s) = compute_elim_signature elims id in
- let branchlengthes = List.map (fun (_,b,c) -> assert (b=None); pi1 (decompose_prod_letin c)) (List.rev s.branches) in
+ let branchlengthes = List.map (fun d -> assert (is_local_assum d); pi1 (decompose_prod_letin (get_type d))) (List.rev s.branches) in
evd, isrec, ({elimindex = None; elimbody = elimc; elimrename = Some (isrec,Array.of_list branchlengthes)}, elimt), l
(* Instantiate all meta variables of elimclause using lid, some elts
@@ -3864,6 +3916,7 @@ let induction_tac with_evars params indvars elim gl =
induction applies with the induction hypotheses *)
let apply_induction_in_context hyp0 inhyps elim indvars names induct_tac =
+ let open Context.Named.Declaration in
Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let env = Proofview.Goal.env gl in
@@ -3876,7 +3929,7 @@ let apply_induction_in_context hyp0 inhyps elim indvars names induct_tac =
let s = Retyping.get_sort_family_of env sigma tmpcl in
let deps_cstr =
List.fold_left
- (fun a (id,b,_) -> if Option.is_empty b then (mkVar id)::a else a) [] deps in
+ (fun a decl -> if is_local_assum decl then (mkVar (get_id decl))::a else a) [] deps in
let (sigma, isrec, elim, indsign) = get_eliminator elim dep s (Proofview.Goal.assume gl) in
let branchletsigns =
let f (_,is_not_let,_,_) = is_not_let in
@@ -3956,6 +4009,7 @@ let induction_without_atomization isrec with_evars elim names lid =
(* assume that no occurrences are selected *)
let clear_unselected_context id inhyps cls gl =
+ let open Context.Named.Declaration in
if occur_var (Tacmach.pf_env gl) id (Tacmach.pf_concl gl) &&
cls.concl_occs == NoOccurrences
then errorlabstrm ""
@@ -3963,7 +4017,8 @@ let clear_unselected_context id inhyps cls gl =
++ str ".");
match cls.onhyps with
| Some hyps ->
- let to_erase (id',_,_ as d) =
+ let to_erase d =
+ let id' = get_id d in
if Id.List.mem id' inhyps then (* if selected, do not erase *) None
else
(* erase if not selected and dependent on id or selected hyps *)
@@ -4536,39 +4591,45 @@ let intros_transitivity n = Tacticals.New.tclTHEN intros (transitivity_gen n)
is solved by tac *)
(** d1 is the section variable in the global context, d2 in the goal context *)
-let interpretable_as_section_decl evd d1 d2 = match d2,d1 with
- | (_,Some _,_), (_,None,_) -> false
- | (_,Some b1,t1), (_,Some b2,t2) ->
+let interpretable_as_section_decl evd d1 d2 =
+ let open Context.Named.Declaration in
+ match d2, d1 with
+ | LocalDef _, LocalAssum _ -> false
+ | LocalDef (_,b1,t1), LocalDef (_,b2,t2) ->
e_eq_constr_univs evd b1 b2 && e_eq_constr_univs evd t1 t2
- | (_,None,t1), (_,_,t2) -> e_eq_constr_univs evd t1 t2
+ | LocalAssum (_,t1), d2 -> e_eq_constr_univs evd t1 (get_type d2)
let rec decompose len c t accu =
+ let open Context.Rel.Declaration in
if len = 0 then (c, t, accu)
else match kind_of_term c, kind_of_term t with
| Lambda (na, u, c), Prod (_, _, t) ->
- decompose (pred len) c t ((na, None, u) :: accu)
+ decompose (pred len) c t (LocalAssum (na, u) :: accu)
| LetIn (na, b, u, c), LetIn (_, _, _, t) ->
- decompose (pred len) c t ((na, Some b, u) :: accu)
+ decompose (pred len) c t (LocalDef (na, b, u) :: accu)
| _ -> assert false
-let rec shrink ctx sign c t accu = match ctx, sign with
-| [], [] -> (c, t, accu)
-| p :: ctx, (id, _, _) :: sign ->
- if noccurn 1 c then
- let c = subst1 mkProp c in
- let t = subst1 mkProp t in
- shrink ctx sign c t accu
- else
- let c = mkLambda_or_LetIn p c in
- let t = mkProd_or_LetIn p t in
- let accu = match p with
- | (_, None, _) -> mkVar id :: accu
- | (_, Some _, _) -> accu
+let rec shrink ctx sign c t accu =
+ let open Context.Rel.Declaration in
+ match ctx, sign with
+ | [], [] -> (c, t, accu)
+ | p :: ctx, decl :: sign ->
+ if noccurn 1 c then
+ let c = subst1 mkProp c in
+ let t = subst1 mkProp t in
+ shrink ctx sign c t accu
+ else
+ let c = mkLambda_or_LetIn p c in
+ let t = mkProd_or_LetIn p t in
+ let accu = if is_local_assum p then let open Context.Named.Declaration in
+ mkVar (get_id decl) :: accu
+ else accu
in
shrink ctx sign c t accu
| _ -> assert false
let shrink_entry sign const =
+ let open Context.Named.Declaration in
let open Entries in
let typ = match const.const_entry_type with
| None -> assert false
@@ -4589,6 +4650,7 @@ let abstract_subproof id gk tac =
let open Tacticals.New in
let open Tacmach.New in
let open Proofview.Notations in
+ let open Context.Named.Declaration in
Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let current_sign = Global.named_context()
@@ -4597,7 +4659,8 @@ let abstract_subproof id gk tac =
let evdref = ref sigma in
let sign,secsign =
List.fold_right
- (fun (id,_,_ as d) (s1,s2) ->
+ (fun d (s1,s2) ->
+ let id = get_id d in
if mem_named_context id current_sign &&
interpretable_as_section_decl evdref (Context.Named.lookup id current_sign) d
then (s1,push_named_context_val d s2)
diff --git a/toplevel/assumptions.ml b/toplevel/assumptions.ml
index cddc55515..b29ceb78b 100644
--- a/toplevel/assumptions.ml
+++ b/toplevel/assumptions.ml
@@ -23,6 +23,7 @@ open Declarations
open Mod_subst
open Globnames
open Printer
+open Context.Named.Declaration
(** For a constant c in a module sealed by an interface (M:T and
not M<:T), [Global.lookup_constant] may return a [constant_body]
@@ -145,7 +146,7 @@ let push (r : Context.Rel.Declaration.t) (ctx : Context.Rel.t) = r :: ctx
let rec traverse current ctx accu t = match kind_of_term t with
| Var id ->
- let body () = match Global.lookup_named id with (_, body, _) -> body in
+ let body () = Global.lookup_named id |> get_value in
traverse_object accu body (VarRef id)
| Const (kn, _) ->
let body () = Global.body_of_constant_body (lookup_constant kn) in
@@ -208,8 +209,8 @@ let assumptions ?(add_opaque=false) ?(add_transparent=false) st gr t =
let (_, graph, ax2ty) = traverse (label_of gr) t in
let fold obj _ accu = match obj with
| VarRef id ->
- let (_, body, t) = Global.lookup_named id in
- if Option.is_empty body then ContextObjectMap.add (Variable id) t accu
+ let decl = Global.lookup_named id in
+ if is_local_assum decl then ContextObjectMap.add (Variable id) t accu
else accu
| ConstRef kn ->
let cb = lookup_constant kn in
diff --git a/toplevel/auto_ind_decl.ml b/toplevel/auto_ind_decl.ml
index f0c7a3961..0755f8bcf 100644
--- a/toplevel/auto_ind_decl.ml
+++ b/toplevel/auto_ind_decl.ml
@@ -25,6 +25,7 @@ open Tactics
open Ind_tables
open Misctypes
open Proofview.Notations
+open Context.Rel.Declaration
let out_punivs = Univ.out_punivs
@@ -146,17 +147,17 @@ let build_beq_scheme mode kn =
) ext_rel_list in
let eq_input = List.fold_left2
- ( fun a b (n,_,_) -> (* mkLambda(n,b,a) ) *)
+ ( fun a b decl -> (* mkLambda(n,b,a) ) *)
(* here I leave the Naming thingy so that the type of
the function is more readable for the user *)
- mkNamedLambda (eqName n) b a )
+ mkNamedLambda (eqName (get_name decl)) b a )
c (List.rev eqs_typ) lnamesparrec
in
- List.fold_left (fun a (n,_,t) ->(* mkLambda(n,t,a)) eq_input rel_list *)
+ List.fold_left (fun a decl ->(* mkLambda(n,t,a)) eq_input rel_list *)
(* Same here , hoping the auto renaming will do something good ;) *)
mkNamedLambda
- (match n with Name s -> s | Anonymous -> Id.of_string "A")
- t a) eq_input lnamesparrec
+ (match get_name decl with Name s -> s | Anonymous -> Id.of_string "A")
+ (get_type decl) a) eq_input lnamesparrec
in
let make_one_eq cur =
let u = Univ.Instance.empty in
@@ -248,7 +249,7 @@ let build_beq_scheme mode kn =
| 0 -> Lazy.force tt
| _ -> let eqs = Array.make nb_cstr_args (Lazy.force tt) in
for ndx = 0 to nb_cstr_args-1 do
- let _,_,cc = List.nth constrsi.(i).cs_args ndx in
+ let cc = get_type (List.nth constrsi.(i).cs_args ndx) in
let eqA, eff' = compute_A_equality rel_list
nparrec
(nparrec+3+2*nb_cstr_args)
@@ -267,14 +268,14 @@ let build_beq_scheme mode kn =
(Array.sub eqs 1 (nb_cstr_args - 1))
)
in
- (List.fold_left (fun a (p,q,r) -> mkLambda (p,r,a)) cc
+ (List.fold_left (fun a decl -> mkLambda (get_name decl, get_type decl, a)) cc
(constrsj.(j).cs_args)
)
- else ar2.(j) <- (List.fold_left (fun a (p,q,r) ->
- mkLambda (p,r,a)) (Lazy.force ff) (constrsj.(j).cs_args) )
+ else ar2.(j) <- (List.fold_left (fun a decl ->
+ mkLambda (get_name decl, get_type decl, a)) (Lazy.force ff) (constrsj.(j).cs_args) )
done;
- ar.(i) <- (List.fold_left (fun a (p,q,r) -> mkLambda (p,r,a))
+ ar.(i) <- (List.fold_left (fun a decl -> mkLambda (get_name decl, get_type decl, a))
(mkCase (ci,do_predicate rel_list nb_cstr_args,
mkVar (Id.of_string "Y") ,ar2))
(constrsi.(i).cs_args))
@@ -487,8 +488,8 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt =
create, from a list of ids [i1,i2,...,in] the list
[(in,eq_in,in_bl,in_al),,...,(i1,eq_i1,i1_bl_i1_al )]
*)
-let list_id l = List.fold_left ( fun a (n,_,t) -> let s' =
- match n with
+let list_id l = List.fold_left ( fun a decl -> let s' =
+ match get_name decl with
Name s -> Id.to_string s
| Anonymous -> "A" in
(Id.of_string s',Id.of_string ("eq_"^s'),
@@ -535,9 +536,9 @@ let compute_bl_goal ind lnamesparrec nparrec =
let eq_input = List.fold_left2 ( fun a (s,seq,_,_) b ->
mkNamedProd seq b a
) bl_input (List.rev list_id) (List.rev eqs_typ) in
- List.fold_left (fun a (n,_,t) -> mkNamedProd
- (match n with Name s -> s | Anonymous -> Id.of_string "A")
- t a) eq_input lnamesparrec
+ List.fold_left (fun a decl -> mkNamedProd
+ (match get_name decl with Name s -> s | Anonymous -> Id.of_string "A")
+ (get_type decl) a) eq_input lnamesparrec
in
let n = Id.of_string "x" and
m = Id.of_string "y" in
@@ -678,9 +679,9 @@ let compute_lb_goal ind lnamesparrec nparrec =
let eq_input = List.fold_left2 ( fun a (s,seq,_,_) b ->
mkNamedProd seq b a
) lb_input (List.rev list_id) (List.rev eqs_typ) in
- List.fold_left (fun a (n,_,t) -> mkNamedProd
- (match n with Name s -> s | Anonymous -> Id.of_string "A")
- t a) eq_input lnamesparrec
+ List.fold_left (fun a decl -> mkNamedProd
+ (match (get_name decl) with Name s -> s | Anonymous -> Id.of_string "A")
+ (get_type decl) a) eq_input lnamesparrec
in
let n = Id.of_string "x" and
m = Id.of_string "y" in
@@ -819,9 +820,9 @@ let compute_dec_goal ind lnamesparrec nparrec =
let eq_input = List.fold_left2 ( fun a (s,seq,_,_) b ->
mkNamedProd seq b a
) bl_input (List.rev list_id) (List.rev eqs_typ) in
- List.fold_left (fun a (n,_,t) -> mkNamedProd
- (match n with Name s -> s | Anonymous -> Id.of_string "A")
- t a) eq_input lnamesparrec
+ List.fold_left (fun a decl -> mkNamedProd
+ (match get_name decl with Name s -> s | Anonymous -> Id.of_string "A")
+ (get_type decl) a) eq_input lnamesparrec
in
let n = Id.of_string "x" and
m = Id.of_string "y" in
diff --git a/toplevel/classes.ml b/toplevel/classes.ml
index 6bb047af0..2089bc944 100644
--- a/toplevel/classes.ml
+++ b/toplevel/classes.ml
@@ -21,6 +21,7 @@ open Globnames
open Constrintern
open Constrexpr
open Sigma.Notations
+open Context.Rel.Declaration
(*i*)
open Decl_kinds
@@ -75,14 +76,14 @@ let mismatched_props env n m = mismatched_ctx_inst env Properties n m
let type_ctx_instance evars env ctx inst subst =
let rec aux (subst, instctx) l = function
- (na, b, t) :: ctx ->
- let t' = substl subst t in
+ decl :: ctx ->
+ let t' = substl subst (get_type decl) in
let c', l =
- match b with
- | None -> interp_casted_constr_evars env evars (List.hd l) t', List.tl l
- | Some b -> substl subst b, l
+ match decl with
+ | LocalAssum _ -> interp_casted_constr_evars env evars (List.hd l) t', List.tl l
+ | LocalDef (_,b,_) -> substl subst b, l
in
- let d = na, Some c', t' in
+ let d = get_name decl, Some c', t' in
aux (c' :: subst, d :: instctx) l ctx
| [] -> subst
in aux (subst, []) inst (List.rev ctx)
@@ -131,7 +132,7 @@ let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) pro
match bk with
| Implicit ->
Implicit_quantifiers.implicit_application Id.Set.empty ~allow_partial:false
- (fun avoid (clname, (id, _, t)) ->
+ (fun avoid (clname, _) ->
match clname with
| Some (cl, b) ->
let t = CHole (Loc.ghost, None, Misctypes.IntroAnonymous, None) in
@@ -154,10 +155,11 @@ let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) pro
let k, args = Typeclasses.dest_class_app (push_rel_context ctx'' env) c in
let cl, u = Typeclasses.typeclass_univ_instance k in
let _, args =
- List.fold_right (fun (na, b, t) (args, args') ->
- match b with
- | None -> (List.tl args, List.hd args :: args')
- | Some b -> (args, substl args' b :: args'))
+ List.fold_right (fun decl (args, args') ->
+ let open Context.Rel.Declaration in
+ match decl with
+ | LocalAssum _ -> (List.tl args, List.hd args :: args')
+ | LocalDef (_,b,_) -> (args, substl args' b :: args'))
(snd cl.cl_context) (args, [])
in
cl, u, c', ctx', ctx, len, imps, args
@@ -180,7 +182,7 @@ let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) pro
if abstract then
begin
let subst = List.fold_left2
- (fun subst' s (_, b, _) -> if Option.is_empty b then s :: subst' else subst')
+ (fun subst' s decl -> if is_local_assum decl then s :: subst' else subst')
[] subst (snd k.cl_context)
in
let (_, ty_constr) = instance_constructor (k,u) subst in
@@ -224,10 +226,10 @@ let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) pro
in
let props, rest =
List.fold_left
- (fun (props, rest) (id,b,_) ->
- if Option.is_empty b then
+ (fun (props, rest) decl ->
+ if is_local_assum decl then
try
- let is_id (id', _) = match id, get_id id' with
+ let is_id (id', _) = match get_name decl, get_id id' with
| Name id, (_, id') -> Id.equal id id'
| Anonymous, _ -> false
in
@@ -261,7 +263,7 @@ let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) pro
None, termtype
| Some (Inl subst) ->
let subst = List.fold_left2
- (fun subst' s (_, b, _) -> if Option.is_empty b then s :: subst' else subst')
+ (fun subst' s decl -> if is_local_assum decl then s :: subst' else subst')
[] subst (k.cl_props @ snd k.cl_context)
in
let (app, ty_constr) = instance_constructor (k,u) subst in
@@ -344,9 +346,11 @@ let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) pro
let named_of_rel_context l =
let acc, ctx =
List.fold_right
- (fun (na, b, t) (subst, ctx) ->
- let id = match na with Anonymous -> invalid_arg "named_of_rel_context" | Name id -> id in
- let d = (id, Option.map (substl subst) b, substl subst t) in
+ (fun decl (subst, ctx) ->
+ let id = match get_name decl with Anonymous -> invalid_arg "named_of_rel_context" | Name id -> id in
+ let d = match decl with
+ | LocalAssum (_,t) -> id, None, substl subst t
+ | LocalDef (_,b,t) -> id, Some (substl subst b), substl subst t in
(mkVar id :: subst, d :: ctx))
l ([], [])
in ctx
@@ -358,7 +362,7 @@ let context poly l =
let subst = Evarutil.evd_comb0 Evarutil.nf_evars_and_universes evars in
let fullctx = Context.Rel.map subst fullctx in
let ce t = Evarutil.check_evars env Evd.empty !evars t in
- let () = List.iter (fun (n, b, t) -> Option.iter ce b; ce t) fullctx in
+ let () = List.iter (fun decl -> Context.Rel.Declaration.iter_constr ce decl) fullctx in
let ctx =
try named_of_rel_context fullctx
with e when Errors.noncritical e ->
diff --git a/toplevel/command.ml b/toplevel/command.ml
index 18b2b1444..02f29b155 100644
--- a/toplevel/command.ml
+++ b/toplevel/command.ml
@@ -36,6 +36,7 @@ open Evarconv
open Indschemes
open Misctypes
open Vernacexpr
+open Context.Rel.Declaration
let do_universe poly l = Declare.do_universe poly l
let do_constraint poly l = Declare.do_constraint poly l
@@ -44,9 +45,9 @@ let rec under_binders env sigma f n c =
if Int.equal n 0 then snd (f env sigma c) else
match kind_of_term c with
| Lambda (x,t,c) ->
- mkLambda (x,t,under_binders (push_rel (x,None,t) env) sigma f (n-1) c)
+ mkLambda (x,t,under_binders (push_rel (LocalAssum (x,t)) env) sigma f (n-1) c)
| LetIn (x,b,t,c) ->
- mkLetIn (x,b,t,under_binders (push_rel (x,Some b,t) env) sigma f (n-1) c)
+ mkLetIn (x,b,t,under_binders (push_rel (LocalDef (x,b,t)) env) sigma f (n-1) c)
| _ -> assert false
let rec complete_conclusion a cs = function
@@ -259,6 +260,7 @@ let declare_assumptions idl is_coe k (c,ctx) pl imps impl_is_on nl =
List.rev refs, status
let do_assumptions_unbound_univs (_, poly, _ as kind) nl l =
+ let open Context.Named.Declaration in
let env = Global.env () in
let evdref = ref (Evd.from_env env) in
let l =
@@ -273,7 +275,7 @@ let do_assumptions_unbound_univs (_, poly, _ as kind) nl l =
let _,l = List.fold_map (fun (env,ienv) (is_coe,(idl,c)) ->
let (t,ctx),imps = interp_assumption evdref env ienv [] c in
let env =
- push_named_context (List.map (fun (_,id) -> (id,None,t)) idl) env in
+ push_named_context (List.map (fun (_,id) -> LocalAssum (id,t)) idl) env in
let ienv = List.fold_right (fun (_,id) ienv ->
let impls = compute_internalization_data env Variable t imps in
Id.Map.add id impls ienv) idl ienv in
@@ -335,7 +337,7 @@ let do_assumptions kind nl l = match l with
(* 3b| Mutual inductive definitions *)
let push_types env idl tl =
- List.fold_left2 (fun env id t -> Environ.push_rel (Name id,None,t) env)
+ List.fold_left2 (fun env id t -> Environ.push_rel (LocalAssum (Name id,t)) env)
env idl tl
type structured_one_inductive_expr = {
@@ -378,8 +380,8 @@ let mk_mltype_data evdref env assums arity indname =
(is_ml_type,indname,assums)
let prepare_param = function
- | (na,None,t) -> out_name na, LocalAssum t
- | (na,Some b,_) -> out_name na, LocalDef b
+ | LocalAssum (na,t) -> out_name na, Entries.LocalAssum t
+ | LocalDef (na,b,_) -> out_name na, Entries.LocalDef b
(** Make the arity conclusion flexible to avoid generating an upper bound universe now,
only if the universe does not appear anywhere else.
@@ -433,12 +435,12 @@ let interp_cstrs evdref env impls mldata arity ind =
let sign_level env evd sign =
fst (List.fold_right
- (fun (_,b,t as d) (lev,env) ->
- match b with
- | Some _ -> (lev, push_rel d env)
- | None ->
+ (fun d (lev,env) ->
+ match d with
+ | LocalDef _ -> lev, push_rel d env
+ | LocalAssum _ ->
let s = destSort (Reduction.whd_betadeltaiota env
- (nf_evar evd (Retyping.get_type_of env evd t)))
+ (nf_evar evd (Retyping.get_type_of env evd (get_type d))))
in
let u = univ_of_sort s in
(Univ.sup u lev, push_rel d env))
@@ -449,7 +451,7 @@ let sup_list min = List.fold_left Univ.sup min
let extract_level env evd min tys =
let sorts = List.map (fun ty ->
let ctx, concl = Reduction.dest_prod_assum env ty in
- sign_level env evd ((Anonymous, None, concl) :: ctx)) tys
+ sign_level env evd (LocalAssum (Anonymous, concl) :: ctx)) tys
in sup_list min sorts
let is_flexible_sort evd u =
@@ -555,8 +557,8 @@ let interp_mutual_inductive (paramsl,indl) notations poly prv finite =
let indnames = List.map (fun ind -> ind.ind_name) indl in
(* Names of parameters as arguments of the inductive type (defs removed) *)
- let assums = List.filter(fun (_,b,_) -> Option.is_empty b) ctx_params in
- let params = List.map (fun (na,_,_) -> out_name na) assums in
+ let assums = List.filter is_local_assum ctx_params in
+ let params = List.map (fun decl -> out_name (get_name decl)) assums in
(* Interpret the arities *)
let arities = List.map (interp_ind_arity env_params evdref) indl in
@@ -876,12 +878,13 @@ let lt_ref = make_qref "Init.Peano.lt"
let rec telescope = function
| [] -> assert false
- | [(n, None, t)] -> t, [n, Some (mkRel 1), t], mkRel 1
- | (n, None, t) :: tl ->
+ | [LocalAssum (n, t)] -> t, [LocalDef (n, mkRel 1, t)], mkRel 1
+ | LocalAssum (n, t) :: tl ->
let ty, tys, (k, constr) =
List.fold_left
- (fun (ty, tys, (k, constr)) (n, b, t) ->
- let pred = mkLambda (n, t, ty) in
+ (fun (ty, tys, (k, constr)) decl ->
+ let t = get_type decl in
+ let pred = mkLambda (get_name decl, t, ty) in
let ty = Universes.constr_of_global (Lazy.force sigT).typ in
let intro = Universes.constr_of_global (Lazy.force sigT).intro in
let sigty = mkApp (ty, [|t; pred|]) in
@@ -890,21 +893,21 @@ let rec telescope = function
(t, [], (2, mkRel 1)) tl
in
let (last, subst) = List.fold_right2
- (fun pred (n, b, t) (prev, subst) ->
+ (fun pred decl (prev, subst) ->
+ let t = get_type decl in
let p1 = Universes.constr_of_global (Lazy.force sigT).proj1 in
let p2 = Universes.constr_of_global (Lazy.force sigT).proj2 in
let proj1 = applistc p1 [t; pred; prev] in
let proj2 = applistc p2 [t; pred; prev] in
- (lift 1 proj2, (n, Some proj1, t) :: subst))
+ (lift 1 proj2, LocalDef (get_name decl, proj1, t) :: subst))
(List.rev tys) tl (mkRel 1, [])
- in ty, ((n, Some last, t) :: subst), constr
+ in ty, (LocalDef (n, last, t) :: subst), constr
- | (n, Some b, t) :: tl -> let ty, subst, term = telescope tl in
- ty, ((n, Some b, t) :: subst), lift 1 term
+ | LocalDef (n, b, t) :: tl -> let ty, subst, term = telescope tl in
+ ty, (LocalDef (n, b, t) :: subst), lift 1 term
let nf_evar_context sigma ctx =
- List.map (fun (n, b, t) ->
- (n, Option.map (Evarutil.nf_evar sigma) b, Evarutil.nf_evar sigma t)) ctx
+ List.map (map_constr (Evarutil.nf_evar sigma)) ctx
let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
Coqlib.check_required_library ["Coq";"Program";"Wf"];
@@ -918,7 +921,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
let full_arity = it_mkProd_or_LetIn top_arity binders_rel in
let argtyp, letbinders, make = telescope binders_rel in
let argname = Id.of_string "recarg" in
- let arg = (Name argname, None, argtyp) in
+ let arg = LocalAssum (Name argname, argtyp) in
let binders = letbinders @ [arg] in
let binders_env = push_rel_context binders_rel env in
let rel, _ = interp_constr_evars_impls env evdref r in
@@ -933,7 +936,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
try
let ctx, ar = Reductionops.splay_prod_n env !evdref 2 relty in
match ctx, kind_of_term ar with
- | [(_, None, t); (_, None, u)], Sort (Prop Null)
+ | [LocalAssum (_,t); LocalAssum (_,u)], Sort (Prop Null)
when Reductionops.is_conv env !evdref t u -> t
| _, _ -> error ()
with e when Errors.noncritical e -> error ()
@@ -953,9 +956,9 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
in
let wf_proof = mkApp (delayed_force well_founded, [| argtyp ; wf_rel |]) in
let argid' = Id.of_string (Id.to_string argname ^ "'") in
- let wfarg len = (Name argid', None,
- mkSubset (Name argid') argtyp
- (wf_rel_fun (mkRel 1) (mkRel (len + 1))))
+ let wfarg len = LocalAssum (Name argid',
+ mkSubset (Name argid') argtyp
+ (wf_rel_fun (mkRel 1) (mkRel (len + 1))))
in
let intern_bl = wfarg 1 :: [arg] in
let _intern_env = push_rel_context intern_bl env in
@@ -969,22 +972,22 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
(* substitute the projection of wfarg for something,
now intern_arity is in wfarg :: arg *)
let intern_fun_arity_prod = it_mkProd_or_LetIn intern_arity [wfarg 1] in
- let intern_fun_binder = (Name (add_suffix recname "'"), None, intern_fun_arity_prod) in
+ let intern_fun_binder = LocalAssum (Name (add_suffix recname "'"), intern_fun_arity_prod) in
let curry_fun =
let wfpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel (2 * len + 4))) in
let intro = (*FIXME*)Universes.constr_of_global (delayed_force build_sigma).Coqlib.intro in
let arg = mkApp (intro, [| argtyp; wfpred; lift 1 make; mkRel 1 |]) in
let app = mkApp (mkRel (2 * len + 2 (* recproof + orig binders + current binders *)), [| arg |]) in
let rcurry = mkApp (rel, [| measure; lift len measure |]) in
- let lam = (Name (Id.of_string "recproof"), None, rcurry) in
+ let lam = LocalAssum (Name (Id.of_string "recproof"), rcurry) in
let body = it_mkLambda_or_LetIn app (lam :: binders_rel) in
let ty = it_mkProd_or_LetIn (lift 1 top_arity) (lam :: binders_rel) in
- (Name recname, Some body, ty)
+ LocalDef (Name recname, body, ty)
in
let fun_bl = intern_fun_binder :: [arg] in
let lift_lets = Termops.lift_rel_context 1 letbinders in
let intern_body =
- let ctx = (Name recname, None, pi3 curry_fun) :: binders_rel in
+ let ctx = LocalAssum (Name recname, get_type curry_fun) :: binders_rel in
let (r, l, impls, scopes) =
Constrintern.compute_internalization_data env
Constrintern.Recursive full_arity impls
@@ -1046,6 +1049,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
evars_typ ctx evars ~hook)
let interp_recursive isfix fixl notations =
+ let open Context.Named.Declaration in
let env = Global.env() in
let fixnames = List.map (fun fix -> fix.fix_name) fixl in
@@ -1081,8 +1085,8 @@ let interp_recursive isfix fixl notations =
Typing.solve_evars env evdref app
with e when Errors.noncritical e -> t
in
- (id,None,fixprot) :: env'
- else (id,None,t) :: env')
+ LocalAssum (id,fixprot) :: env'
+ else LocalAssum (id,t) :: env')
[] fixnames fixtypes
in
let env_rec = push_named_context rec_sign env in
@@ -1104,7 +1108,7 @@ let interp_recursive isfix fixl notations =
let evd, nf = nf_evars_and_universes evd in
let fixdefs = List.map (Option.map nf) fixdefs in
let fixtypes = List.map nf fixtypes in
- let fixctxnames = List.map (fun (_,ctx) -> List.map pi1 ctx) fixctxs in
+ let fixctxnames = List.map (fun (_,ctx) -> List.map get_name ctx) fixctxs in
(* Build the fix declaration block *)
(env,rec_sign,all_universes,evd), (fixnames,fixdefs,fixtypes), List.combine3 fixctxnames fiximps fixannots
diff --git a/toplevel/discharge.ml b/toplevel/discharge.ml
index 61a6e1094..5fa51e06e 100644
--- a/toplevel/discharge.ml
+++ b/toplevel/discharge.ml
@@ -14,14 +14,16 @@ open Vars
open Entries
open Declarations
open Cooking
+open Context.Rel.Declaration
(********************************)
(* Discharging mutual inductive *)
-let detype_param = function
- | (Name id,None,p) -> id, LocalAssum p
- | (Name id,Some p,_) -> id, LocalDef p
- | (Anonymous,_,_) -> anomaly (Pp.str "Unnamed inductive local variable")
+let detype_param =
+ function
+ | LocalAssum (Name id, p) -> id, Entries.LocalAssum p
+ | LocalDef (Name id, p,_) -> id, Entries.LocalDef p
+ | _ -> anomaly (Pp.str "Unnamed inductive local variable")
(* Replace
@@ -52,7 +54,7 @@ let abstract_inductive hyps nparams inds =
(* To be sure to be the same as before, should probably be moved to process_inductive *)
let params' = let (_,arity,_,_,_) = List.hd inds' in
let (params,_) = decompose_prod_n_assum nparams' arity in
- List.map detype_param params
+ List.map detype_param params
in
let ind'' =
List.map
diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml
index 4a5f14917..4a145481f 100644
--- a/toplevel/himsg.ml
+++ b/toplevel/himsg.ml
@@ -23,22 +23,26 @@ open Cases
open Logic
open Printer
open Evd
+open Context.Rel.Declaration
(* This simplifies the typing context of Cases clauses *)
(* hope it does not disturb other typing contexts *)
let contract env lc =
let l = ref [] in
- let contract_context (na,c,t) env =
- match c with
- | Some c' when isRel c' ->
+ let contract_context decl env =
+ match decl with
+ | LocalDef (_,c',_) when isRel c' ->
l := (Vars.substl !l c') :: !l;
env
| _ ->
- let t' = Vars.substl !l t in
- let c' = Option.map (Vars.substl !l) c in
- let na' = named_hd env t' na in
+ let t' = Vars.substl !l (get_type decl) in
+ let c' = Option.map (Vars.substl !l) (get_value decl) in
+ let na' = named_hd env t' (get_name decl) in
l := (mkRel 1) :: List.map (Vars.lift 1) !l;
- push_rel (na',c',t') env in
+ match c' with
+ | None -> push_rel (LocalAssum (na',t')) env
+ | Some c' -> push_rel (LocalDef (na',c',t')) env
+ in
let env = process_rel_context contract_context env in
(env, List.map (Vars.substl !l) lc)
@@ -136,9 +140,9 @@ let pr_explicit env sigma t1 t2 = pr_explicit_aux env sigma t1 t2 explicit_flags
let pr_db env i =
try
- match lookup_rel i env with
- Name id, _, _ -> pr_id id
- | Anonymous, _, _ -> str "<>"
+ match lookup_rel i env |> get_name with
+ | Name id -> pr_id id
+ | Anonymous -> str "<>"
with Not_found -> str "UNBOUND_REL_" ++ int i
let explain_unbound_rel env sigma n =
diff --git a/toplevel/indschemes.ml b/toplevel/indschemes.ml
index c4ac0e411..251d14af7 100644
--- a/toplevel/indschemes.ml
+++ b/toplevel/indschemes.ml
@@ -38,6 +38,7 @@ open Ind_tables
open Auto_ind_decl
open Eqschemes
open Elimschemes
+open Context.Rel.Declaration
(* Flags governing automatic synthesis of schemes *)
@@ -463,7 +464,7 @@ let build_combined_scheme env schemes =
in
let ctx, _ =
list_split_rev_at prods
- (List.rev_map (fun (x, y) -> x, None, y) ctx) in
+ (List.rev_map (fun (x, y) -> LocalAssum (x, y)) ctx) in
let typ = it_mkProd_wo_LetIn concl_typ ctx in
let body = it_mkLambda_or_LetIn concl_bod ctx in
(body, typ)
diff --git a/toplevel/obligations.ml b/toplevel/obligations.ml
index 0ea9f959f..0e8d224e4 100644
--- a/toplevel/obligations.ml
+++ b/toplevel/obligations.ml
@@ -54,7 +54,8 @@ type oblinfo =
(** Substitute evar references in t using De Bruijn indices,
where n binders were passed through. *)
-let subst_evar_constr evs n idf t =
+let subst_evar_constr evs n idf t =
+ let open Context.Named.Declaration in
let seen = ref Int.Set.empty in
let transparent = ref Id.Set.empty in
let evar_info id = List.assoc_f Evar.equal id evs in
@@ -78,9 +79,9 @@ let subst_evar_constr evs n idf t =
let args =
let rec aux hyps args acc =
match hyps, args with
- ((_, None, _) :: tlh), (c :: tla) ->
+ (LocalAssum _ :: tlh), (c :: tla) ->
aux tlh tla ((substrec (depth, fixrels) c) :: acc)
- | ((_, Some _, _) :: tlh), (_ :: tla) ->
+ | (LocalDef _ :: tlh), (_ :: tla) ->
aux tlh tla acc
| [], [] -> acc
| _, _ -> acc (*failwith "subst_evars: invalid argument"*)
@@ -116,22 +117,23 @@ let subst_vars acc n t =
Changes evars and hypothesis references to variable references.
*)
let etype_of_evar evs hyps concl =
+ let open Context.Named.Declaration in
let rec aux acc n = function
- (id, copt, t) :: tl ->
- let t', s, trans = subst_evar_constr evs n mkVar t in
+ decl :: tl ->
+ let t', s, trans = subst_evar_constr evs n mkVar (get_type decl) in
let t'' = subst_vars acc 0 t' in
- let rest, s', trans' = aux (id :: acc) (succ n) tl in
+ let rest, s', trans' = aux (get_id decl :: acc) (succ n) tl in
let s' = Int.Set.union s s' in
let trans' = Id.Set.union trans trans' in
- (match copt with
- Some c ->
+ (match decl with
+ | LocalDef (id,c,_) ->
let c', s'', trans'' = subst_evar_constr evs n mkVar c in
let c' = subst_vars acc 0 c' in
- mkNamedProd_or_LetIn (id, Some c', t'') rest,
+ mkNamedProd_or_LetIn (LocalDef (id, c', t'')) rest,
Int.Set.union s'' s',
Id.Set.union trans'' trans'
- | None ->
- mkNamedProd_or_LetIn (id, None, t'') rest, s', trans')
+ | LocalAssum (id,_) ->
+ mkNamedProd_or_LetIn (LocalAssum (id, t'')) rest, s', trans')
| [] ->
let t', s, trans = subst_evar_constr evs n mkVar concl in
subst_vars acc 0 t', s, trans
@@ -589,15 +591,16 @@ let declare_mutual_definition l =
Lemmas.call_hook fix_exn first.prg_hook local gr first.prg_ctx;
List.iter progmap_remove l; kn
-let shrink_body c =
+let shrink_body c =
+ let open Context.Rel.Declaration in
let ctx, b = decompose_lam_assum c in
let b', n, args =
- List.fold_left (fun (b, i, args) (n, u, t) ->
+ List.fold_left (fun (b, i, args) decl ->
if noccurn 1 b then
subst1 mkProp b, succ i, args
else
- let args = if Option.is_empty u then mkRel i :: args else args in
- mkLambda_or_LetIn (n, u, t) b, succ i, args)
+ let args = if is_local_assum decl then mkRel i :: args else args in
+ mkLambda_or_LetIn decl b, succ i, args)
(b, 1, []) ctx
in ctx, b', Array.of_list args
diff --git a/toplevel/record.ml b/toplevel/record.ml
index 7ae203034..facc8b75d 100644
--- a/toplevel/record.ml
+++ b/toplevel/record.ml
@@ -24,6 +24,7 @@ open Type_errors
open Constrexpr
open Constrexpr_ops
open Goptions
+open Context.Rel.Declaration
(********** definition d'un record (structure) **************)
@@ -68,16 +69,19 @@ let interp_fields_evars env evars impls_env nots l =
| Anonymous -> impls
| Name id -> Id.Map.add id (compute_internalization_data env Constrintern.Method t' impl) impls
in
- let d = (i,b',t') in
+ let d = match b' with
+ | None -> LocalAssum (i,t')
+ | Some b' -> LocalDef (i,b',t')
+ in
List.iter (Metasyntax.set_notation_for_interpretation impls) no;
(push_rel d env, impl :: uimpls, d::params, impls))
(env, [], [], impls_env) nots l
let compute_constructor_level evars env l =
- List.fold_right (fun (n,b,t as d) (env, univ) ->
+ List.fold_right (fun d (env, univ) ->
let univ =
- if b = None then
- let s = Retyping.get_sort_of env evars t in
+ if is_local_assum d then
+ let s = Retyping.get_sort_of env evars (get_type d) in
Univ.sup (univ_of_sort s) univ
else univ
in (push_rel d env, univ))
@@ -122,7 +126,7 @@ let typecheck_params_and_fields def id pl t ps nots fs =
mkSort (Evarutil.evd_comb0 (Evd.new_sort_variable uvarkind) evars), false
in
let fullarity = it_mkProd_or_LetIn t' newps in
- let env_ar = push_rel_context newps (push_rel (Name id,None,fullarity) env0) in
+ let env_ar = push_rel_context newps (push_rel (LocalAssum (Name id,fullarity)) env0) in
let env2,impls,newfs,data =
interp_fields_evars env_ar evars impls_env nots (binders_of_decls fs)
in
@@ -150,17 +154,17 @@ let typecheck_params_and_fields def id pl t ps nots fs =
let newps = Context.Rel.map nf newps in
let newfs = Context.Rel.map nf newfs in
let ce t = Evarutil.check_evars env0 Evd.empty evars t in
- List.iter (fun (n, b, t) -> Option.iter ce b; ce t) (List.rev newps);
- List.iter (fun (n, b, t) -> Option.iter ce b; ce t) (List.rev newfs);
+ List.iter (iter_constr ce) (List.rev newps);
+ List.iter (iter_constr ce) (List.rev newfs);
Evd.universe_context ?names:pl evars, nf arity, template, imps, newps, impls, newfs
-let degenerate_decl (na,b,t) =
- let id = match na with
+let degenerate_decl decl =
+ let id = match get_name decl with
| Name id -> id
| Anonymous -> anomaly (Pp.str "Unnamed record variable") in
- match b with
- | None -> (id, LocalAssum t)
- | Some b -> (id, LocalDef b)
+ match decl with
+ | LocalAssum (_,t) -> (id, Entries.LocalAssum t)
+ | LocalDef (_,b,_) -> (id, Entries.LocalDef b)
type record_error =
| MissingProj of Id.t * Id.t list
@@ -264,23 +268,25 @@ let declare_projections indsp ?(kind=StructureComponent) binder_name coers field
in
let (_,_,kinds,sp_projs,_) =
List.fold_left3
- (fun (nfi,i,kinds,sp_projs,subst) coe (fi,optci,ti) impls ->
+ (fun (nfi,i,kinds,sp_projs,subst) coe decl impls ->
+ let fi = get_name decl in
+ let ti = get_type decl in
let (sp_projs,i,subst) =
match fi with
| Anonymous ->
(None::sp_projs,i,NoProjection fi::subst)
| Name fid -> try
let kn, term =
- if optci = None && primitive then
+ if is_local_assum decl && primitive then
(** Already defined in the kernel silently *)
let kn = destConstRef (Nametab.locate (Libnames.qualid_of_ident fid)) in
Declare.definition_message fid;
kn, mkProj (Projection.make kn false,mkRel 1)
else
let ccl = subst_projection fid subst ti in
- let body = match optci with
- | Some ci -> subst_projection fid subst ci
- | None ->
+ let body = match decl with
+ | LocalDef (_,ci,_) -> subst_projection fid subst ci
+ | LocalAssum _ ->
(* [ccl] is defined in context [params;x:rp] *)
(* [ccl'] is defined in context [params;x:rp;x:rp] *)
let ccl' = liftn 1 2 ccl in
@@ -322,28 +328,28 @@ let declare_projections indsp ?(kind=StructureComponent) binder_name coers field
let cl = Class.class_of_global (IndRef indsp) in
Class.try_add_new_coercion_with_source refi ~local:false poly ~source:cl
end;
- let i = if Option.is_empty optci then i+1 else i in
+ let i = if is_local_assum decl then i+1 else i in
(Some kn::sp_projs, i, Projection term::subst)
with NotDefinable why ->
warning_or_error coe indsp why;
(None::sp_projs,i,NoProjection fi::subst) in
- (nfi-1,i,(fi, Option.is_empty optci)::kinds,sp_projs,subst))
+ (nfi-1,i,(fi, is_local_assum decl)::kinds,sp_projs,subst))
(List.length fields,0,[],[],[]) coers (List.rev fields) (List.rev fieldimpls)
in (kinds,sp_projs)
let structure_signature ctx =
let rec deps_to_evar evm l =
match l with [] -> Evd.empty
- | [(_,_,typ)] ->
+ | [decl] ->
let env = Environ.empty_named_context_val in
- let (evm, _) = Evarutil.new_pure_evar env evm typ in
+ let (evm, _) = Evarutil.new_pure_evar env evm (get_type decl) in
evm
- | (_,_,typ)::tl ->
+ | decl::tl ->
let env = Environ.empty_named_context_val in
- let (evm, ev) = Evarutil.new_pure_evar env evm typ in
+ let (evm, ev) = Evarutil.new_pure_evar env evm (get_type decl) in
let new_tl = Util.List.map_i
- (fun pos (n,c,t) -> n,c,
- Termops.replace_term (mkRel pos) (mkEvar(ev,[||])) t) 1 tl in
+ (fun pos decl ->
+ map_type (fun t -> Termops.replace_term (mkRel pos) (mkEvar(ev,[||])) t) decl) 1 tl in
deps_to_evar evm new_tl in
deps_to_evar Evd.empty (List.rev ctx)
@@ -391,7 +397,7 @@ let implicits_of_context ctx =
| Name n -> Some n
| Anonymous -> None
in ExplByPos (i, explname), (true, true, true))
- 1 (List.rev (Anonymous :: (List.map pi1 ctx)))
+ 1 (List.rev (Anonymous :: (List.map get_name ctx)))
let declare_class finite def poly ctx id idbuild paramimpls params arity
template fieldimpls fields ?(kind=StructureComponent) is_coe coers priorities sign =
@@ -404,7 +410,7 @@ let declare_class finite def poly ctx id idbuild paramimpls params arity
let binder_name = Namegen.next_ident_away (snd id) (Termops.ids_of_context (Global.env())) in
let impl, projs =
match fields with
- | [(Name proj_name, _, field)] when def ->
+ | [LocalAssum (Name proj_name, field) | LocalDef (Name proj_name, _, field)] when def ->
let class_body = it_mkLambda_or_LetIn field params in
let _class_type = it_mkProd_or_LetIn arity params in
let class_entry =
@@ -445,13 +451,13 @@ let declare_class finite def poly ctx id idbuild paramimpls params arity
if b then Backward, pri else Forward, pri) coe)
coers priorities
in
- let l = List.map3 (fun (id, _, _) b y -> (id, b, y))
+ let l = List.map3 (fun decl b y -> get_name decl, b, y)
(List.rev fields) coers (Recordops.lookup_projections ind)
in IndRef ind, l
in
let ctx_context =
- List.map (fun (na, b, t) ->
- match Typeclasses.class_of_constr t with
+ List.map (fun decl ->
+ match Typeclasses.class_of_constr (get_type decl) with
| Some (_, ((cl,_), _)) -> Some (cl.cl_impl, true)
| None -> None)
params, params
@@ -473,7 +479,7 @@ let add_constant_class cst =
let tc =
{ cl_impl = ConstRef cst;
cl_context = (List.map (const None) ctx, ctx);
- cl_props = [(Anonymous, None, arity)];
+ cl_props = [LocalAssum (Anonymous, arity)];
cl_projs = [];
cl_strict = !typeclasses_strict;
cl_unique = !typeclasses_unique
@@ -487,8 +493,8 @@ let add_inductive_class ind =
let ctx = oneind.mind_arity_ctxt in
let inst = Univ.UContext.instance mind.mind_universes in
let map = function
- | (_, Some _, _) -> None
- | (_, None, t) -> Some (lazy t)
+ | LocalDef _ -> None
+ | LocalAssum (_, t) -> Some (lazy t)
in
let args = List.map_filter map ctx in
let ty = Inductive.type_of_inductive_knowing_parameters
@@ -498,7 +504,7 @@ let add_inductive_class ind =
in
{ cl_impl = IndRef ind;
cl_context = List.map (const None) ctx, ctx;
- cl_props = [Anonymous, None, ty];
+ cl_props = [LocalAssum (Anonymous, ty)];
cl_projs = [];
cl_strict = !typeclasses_strict;
cl_unique = !typeclasses_unique }
diff --git a/toplevel/search.ml b/toplevel/search.ml
index 89e0eb88a..646e2e08a 100644
--- a/toplevel/search.ml
+++ b/toplevel/search.ml
@@ -67,7 +67,9 @@ let iter_constructors indsp u fn env nconstr =
fn (ConstructRef (indsp, i)) env typ
done
-let iter_named_context_name_type f = List.iter (fun (nme,_,typ) -> f nme typ)
+let iter_named_context_name_type f =
+ let open Context.Named.Declaration in
+ List.iter (fun decl -> f (get_id decl) (get_type decl))
(* General search over hypothesis of a goal *)
let iter_hypothesis glnum (fn : global_reference -> env -> constr -> unit) =
@@ -79,12 +81,13 @@ let iter_hypothesis glnum (fn : global_reference -> env -> constr -> unit) =
(* General search over declarations *)
let iter_declarations (fn : global_reference -> env -> constr -> unit) =
+ let open Context.Named.Declaration in
let env = Global.env () in
let iter_obj (sp, kn) lobj = match object_tag lobj with
| "VARIABLE" ->
begin try
- let (id, _, typ) = Global.lookup_named (basename sp) in
- fn (VarRef id) env typ
+ let decl = Global.lookup_named (basename sp) in
+ fn (VarRef (get_id decl)) env (get_type decl)
with Not_found -> (* we are in a section *) () end
| "CONSTANT" ->
let cst = Global.constant_of_delta_kn kn in
diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml
index 0f81943e2..8b3e11848 100644
--- a/toplevel/vernacentries.ml
+++ b/toplevel/vernacentries.ml
@@ -871,13 +871,14 @@ let vernac_set_end_tac tac =
(* TO DO verifier s'il faut pas mettre exist s | TacId s ici*)
let vernac_set_used_variables e =
+ let open Context.Named.Declaration in
let env = Global.env () in
let tys =
List.map snd (Proof.initial_goals (Proof_global.give_me_the_proof ())) in
let l = Proof_using.process_expr env e tys in
let vars = Environ.named_context env in
List.iter (fun id ->
- if not (List.exists (fun (id',_,_) -> Id.equal id id') vars) then
+ if not (List.exists (Id.equal id % get_id) vars) then
errorlabstrm "vernac_set_used_variables"
(str "Unknown variable: " ++ pr_id id))
l;
@@ -1568,6 +1569,7 @@ exception NoHyp
We only print the type and a small statement to this comes from the
goal. Precondition: there must be at least one current goal. *)
let print_about_hyp_globs ref_or_by_not glnumopt =
+ let open Context.Named.Declaration in
try
let gl,id =
match glnumopt,ref_or_by_not with
@@ -1580,11 +1582,11 @@ let print_about_hyp_globs ref_or_by_not glnumopt =
(str "No such goal: " ++ int n ++ str "."))
| _ , _ -> raise NoHyp in
let hyps = pf_hyps gl in
- let (id,bdyopt,typ) = Context.Named.lookup id hyps in
- let natureofid = match bdyopt with
- | None -> "Hypothesis"
- | Some bdy ->"Constant (let in)" in
- v 0 (pr_id id ++ str":" ++ pr_constr typ ++ fnl() ++ fnl()
+ let decl = Context.Named.lookup id hyps in
+ let natureofid = match decl with
+ | LocalAssum _ -> "Hypothesis"
+ | LocalDef (_,bdy,_) ->"Constant (let in)" in
+ v 0 (pr_id id ++ str":" ++ pr_constr (get_type decl) ++ fnl() ++ fnl()
++ str natureofid ++ str " of the goal context.")
with (* fallback to globals *)
| NoHyp | Not_found -> print_about ref_or_by_not