(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * Copyright INRIA, CNRS and contributors *) (* [S.c]b * This corresponds to the usual strategy of weak reduction * FIXP(op,bd,S,args) is the fixpoint (Fix or Cofix) of bodies bd under * the bindings S, and then applied to args. Here again, * weak reduction. * CONSTR(c,args) is the constructor [c] applied to [args]. * PRIMITIVE(cop,args) represent a particial application of * a primitive, or a fully applied primitive * which does not reduce. * cop is the constr representing op. * *) type cbv_value = | VAL of int * constr | STACK of int * cbv_value * cbv_stack | CBN of constr * cbv_value subs | LAM of int * (Name.t Context.binder_annot * constr) list * constr * cbv_value subs | FIXP of fixpoint * cbv_value subs * cbv_value array | COFIXP of cofixpoint * cbv_value subs * cbv_value array | CONSTR of constructor Univ.puniverses * cbv_value array | PRIMITIVE of CPrimitives.t * pconstant * cbv_value array | ARRAY of Univ.Instance.t * cbv_value Parray.t * cbv_value (* type of terms with a hole. This hole can appear only under App or Case. * TOP means the term is considered without context * APP(v,stk) means the term is applied to v, and then the context stk * (v.0 is the first argument). * this corresponds to the application stack of the KAM. * The members of l are values: we evaluate arguments before calling the function. * CASE(t,br,pat,S,stk) means the term is in a case (which is himself in stk * t is the type of the case and br are the branches, all of them under * the subs S, pat is information on the patterns of the Case * (Weak reduction: we propagate the sub only when the selected branch * is determined) * PROJ(p,pb,stk) means the term is in a primitive projection p, itself in stk. * pb is the associated projection body * * Important remark: the APPs should be collapsed: * (APP (l,(APP ...))) forbidden *) and cbv_stack = | TOP | APP of cbv_value array * cbv_stack | CASE of Univ.Instance.t * constr array * case_return * case_branch array * Constr.case_invert * case_info * cbv_value subs * cbv_stack | PROJ of Projection.t * cbv_stack (* les vars pourraient etre des constr, cela permet de retarder les lift: utile ?? *) (* relocation of a value; used when a value stored in a context is expanded * in a larger context. e.g. [%k (S.t)](k+1) --> [^k]t (t is shifted of k) *) let rec shift_value n = function | VAL (k,t) -> VAL (k+n,t) | STACK(k,v,stk) -> STACK(k+n,v,stk) | CBN (t,s) -> CBN(t,subs_shft(n,s)) | LAM (nlams,ctxt,b,s) -> LAM (nlams,ctxt,b,subs_shft (n,s)) | FIXP (fix,s,args) -> FIXP (fix,subs_shft (n,s), Array.map (shift_value n) args) | COFIXP (cofix,s,args) -> COFIXP (cofix,subs_shft (n,s), Array.map (shift_value n) args) | CONSTR (c,args) -> CONSTR (c, Array.map (shift_value n) args) | PRIMITIVE(op,c,args) -> PRIMITIVE(op,c,Array.map (shift_value n) args) | ARRAY (u,t,ty) -> ARRAY(u, Parray.map (shift_value n) t, shift_value n ty) let shift_value n v = if Int.equal n 0 then v else shift_value n v (* Contracts a fixpoint: given a fixpoint and a bindings, * returns the corresponding fixpoint body, and the bindings in which * it should be evaluated: its first variables are the fixpoint bodies * (S, (fix Fi {F0 := T0 .. Fn-1 := Tn-1})) * -> (S. [S]F0 . [S]F1 ... . [S]Fn-1, Ti) *) let rec mk_fix_subs make_body n env i = if Int.equal i n then env else mk_fix_subs make_body n (subs_cons (make_body i) env) (i + 1) let contract_fixp env ((reci,i),(_,_,bds as bodies)) = let make_body j = FIXP(((reci,j),bodies), env, [||]) in let n = Array.length bds in mk_fix_subs make_body n env 0, bds.(i) let contract_cofixp env (i,(_,_,bds as bodies)) = let make_body j = COFIXP((j,bodies), env, [||]) in let n = Array.length bds in mk_fix_subs make_body n env 0, bds.(i) let make_constr_ref n k t = match k with | RelKey p -> mkRel (n+p) | VarKey id -> t | ConstKey cst -> t (* Adds an application list. Collapse APPs! *) let stack_app appl stack = if Int.equal (Array.length appl) 0 then stack else match stack with | APP(args,stk) -> APP(Array.append appl args,stk) | _ -> APP(appl, stack) let rec stack_concat stk1 stk2 = match stk1 with TOP -> stk2 | APP(v,stk1') -> APP(v,stack_concat stk1' stk2) | CASE(u,pms,c,b,iv,i,s,stk1') -> CASE(u,pms,c,b,iv,i,s,stack_concat stk1' stk2) | PROJ (p,stk1') -> PROJ (p,stack_concat stk1' stk2) (* merge stacks when there is no shifts in between *) let mkSTACK = function v, TOP -> v | STACK(0,v0,stk0), stk -> STACK(0,v0,stack_concat stk0 stk) | v,stk -> STACK(0,v,stk) type cbv_infos = { env : Environ.env; tab : (cbv_value, Empty.t) Declarations.constant_def KeyTable.t; reds : RedFlags.reds; sigma : Evd.evar_map } (* Change: zeta reduction cannot be avoided in CBV *) open RedFlags let red_set_ref flags = function | RelKey _ -> red_set flags fDELTA | VarKey id -> red_set flags (fVAR id) | ConstKey (sp,_) -> red_set flags (fCONST sp) (* Transfer application lists from a value to the stack * useful because fixpoints may be totally applied in several times. * On the other hand, irreductible atoms absorb the full stack. *) let strip_appl head stack = match head with | FIXP (fix,env,app) -> (FIXP(fix,env,[||]), stack_app app stack) | COFIXP (cofix,env,app) -> (COFIXP(cofix,env,[||]), stack_app app stack) | CONSTR (c,app) -> (CONSTR(c,[||]), stack_app app stack) | PRIMITIVE(op,c,app) -> (PRIMITIVE(op,c,[||]), stack_app app stack) | VAL _ | STACK _ | CBN _ | LAM _ | ARRAY _ -> (head, stack) (* Tests if fixpoint reduction is possible. *) let fixp_reducible flgs ((reci,i),_) stk = if red_set flgs fFIX then match stk with | APP(appl,_) -> Array.length appl > reci.(i) && (match appl.(reci.(i)) with CONSTR _ -> true | _ -> false) | _ -> false else false let cofixp_reducible flgs _ stk = if red_set flgs fCOFIX then match stk with | (CASE _ | PROJ _ | APP(_,CASE _) | APP(_,PROJ _)) -> true | _ -> false else false let debug_cbv = CDebug.create ~name:"Cbv" () (* Reduction of primitives *) open Primred module VNativeEntries = struct type elem = cbv_value type args = cbv_value array type evd = unit type uinstance = Univ.Instance.t let get = Array.get let get_int () e = match e with | VAL(_, ci) -> (match kind ci with | Int i -> i | _ -> raise Primred.NativeDestKO) | _ -> raise Primred.NativeDestKO let get_float () e = match e with | VAL(_, cf) -> (match kind cf with | Float f -> f | _ -> raise Primred.NativeDestKO) | _ -> raise Primred.NativeDestKO let get_parray () e = match e with | ARRAY(_u,t,_ty) -> t | _ -> raise Primred.NativeDestKO let mkInt env i = VAL(0, mkInt i) let mkFloat env f = VAL(0, mkFloat f) let mkBool env b = let (ct,cf) = get_bool_constructors env in CONSTR(Univ.in_punivs (if b then ct else cf), [||]) let int_ty env = VAL(0, mkConst @@ get_int_type env) let float_ty env = VAL(0, mkConst @@ get_float_type env) let mkCarry env b e = let (c0,c1) = get_carry_constructors env in CONSTR(Univ.in_punivs (if b then c1 else c0), [|int_ty env;e|]) let mkIntPair env e1 e2 = let int_ty = int_ty env in let c = get_pair_constructor env in CONSTR(Univ.in_punivs c, [|int_ty;int_ty;e1;e2|]) let mkFloatIntPair env f i = let float_ty = float_ty env in let int_ty = int_ty env in let c = get_pair_constructor env in CONSTR(Univ.in_punivs c, [|float_ty;int_ty;f;i|]) let mkLt env = let (_eq,lt,_gt) = get_cmp_constructors env in CONSTR(Univ.in_punivs lt, [||]) let mkEq env = let (eq,_lt,_gt) = get_cmp_constructors env in CONSTR(Univ.in_punivs eq, [||]) let mkGt env = let (_eq,_lt,gt) = get_cmp_constructors env in CONSTR(Univ.in_punivs gt, [||]) let mkFLt env = let (_eq,lt,_gt,_nc) = get_f_cmp_constructors env in CONSTR(Univ.in_punivs lt, [||]) let mkFEq env = let (eq,_lt,_gt,_nc) = get_f_cmp_constructors env in CONSTR(Univ.in_punivs eq, [||]) let mkFGt env = let (_eq,_lt,gt,_nc) = get_f_cmp_constructors env in CONSTR(Univ.in_punivs gt, [||]) let mkFNotComparable env = let (_eq,_lt,_gt,nc) = get_f_cmp_constructors env in CONSTR(Univ.in_punivs nc, [||]) let mkPNormal env = let (pNormal,_nNormal,_pSubn,_nSubn,_pZero,_nZero,_pInf,_nInf,_nan) = get_f_class_constructors env in CONSTR(Univ.in_punivs pNormal, [||]) let mkNNormal env = let (_pNormal,nNormal,_pSubn,_nSubn,_pZero,_nZero,_pInf,_nInf,_nan) = get_f_class_constructors env in CONSTR(Univ.in_punivs nNormal, [||]) let mkPSubn env = let (_pNormal,_nNormal,pSubn,_nSubn,_pZero,_nZero,_pInf,_nInf,_nan) = get_f_class_constructors env in CONSTR(Univ.in_punivs pSubn, [||]) let mkNSubn env = let (_pNormal,_nNormal,_pSubn,nSubn,_pZero,_nZero,_pInf,_nInf,_nan) = get_f_class_constructors env in CONSTR(Univ.in_punivs nSubn, [||]) let mkPZero env = let (_pNormal,_nNormal,_pSubn,_nSubn,pZero,_nZero,_pInf,_nInf,_nan) = get_f_class_constructors env in CONSTR(Univ.in_punivs pZero, [||]) let mkNZero env = let (_pNormal,_nNormal,_pSubn,_nSubn,_pZero,nZero,_pInf,_nInf,_nan) = get_f_class_constructors env in CONSTR(Univ.in_punivs nZero, [||]) let mkPInf env = let (_pNormal,_nNormal,_pSubn,_nSubn,_pZero,_nZero,pInf,_nInf,_nan) = get_f_class_constructors env in CONSTR(Univ.in_punivs pInf, [||]) let mkNInf env = let (_pNormal,_nNormal,_pSubn,_nSubn,_pZero,_nZero,_pInf,nInf,_nan) = get_f_class_constructors env in CONSTR(Univ.in_punivs nInf, [||]) let mkNaN env = let (_pNormal,_nNormal,_pSubn,_nSubn,_pZero,_nZero,_pInf,_nInf,nan) = get_f_class_constructors env in CONSTR(Univ.in_punivs nan, [||]) let mkArray env u t ty = ARRAY (u,t,ty) end module VredNative = RedNative(VNativeEntries) let debug_pr_key = function | ConstKey (sp,_) -> Names.Constant.print sp | VarKey id -> Names.Id.print id | RelKey n -> Pp.(str "REL_" ++ int n) let rec reify_stack t = function | TOP -> t | APP (args,st) -> reify_stack (mkApp(t,Array.map reify_value args)) st | CASE (u,pms,ty,br,iv,ci,env,st) -> reify_stack (mkCase (ci, u, pms, ty, iv, t,br)) st | PROJ (p, st) -> reify_stack (mkProj (p, t)) st and reify_value = function (* reduction under binders *) | VAL (n,t) -> lift n t | STACK (0,v,stk) -> reify_stack (reify_value v) stk | STACK (n,v,stk) -> lift n (reify_stack (reify_value v) stk) | CBN(t,env) -> apply_env env t | LAM (k,ctxt,b,env) -> apply_env env @@ List.fold_left (fun c (n,t) -> mkLambda (n, t, c)) b ctxt | FIXP ((lij,fix),env,args) -> let fix = mkFix (lij, fix) in mkApp (apply_env env fix, Array.map reify_value args) | COFIXP ((j,cofix),env,args) -> let cofix = mkCoFix (j, cofix) in mkApp (apply_env env cofix, Array.map reify_value args) | CONSTR (c,args) -> mkApp(mkConstructU c, Array.map reify_value args) | PRIMITIVE(op,c,args) -> mkApp(mkConstU c, Array.map reify_value args) | ARRAY (u,t,ty) -> let t, def = Parray.to_array t in mkArray(u, Array.map reify_value t, reify_value def, reify_value ty) and apply_env env t = match kind t with | Rel i -> begin match expand_rel i env with | Inl (k, v) -> reify_value (shift_value k v) | Inr (k,_) -> mkRel k end | _ -> map_with_binders subs_lift apply_env env t let rec strip_app = function | APP (args,st) -> APP (args,strip_app st) | s -> TOP let rec subs_consn v i n s = if Int.equal i n then s else subs_consn v (i + 1) n (subs_cons v.(i) s) (* TODO: share the common parts with EConstr *) let expand_branch env u pms (ind, i) br = let open Declarations in let nas, _br = br.(i - 1) in let (mib, mip) = Inductive.lookup_mind_specif env ind in let paramdecl = Vars.subst_instance_context u mib.mind_params_ctxt in let paramsubst = Vars.subst_of_rel_context_instance paramdecl (Array.to_list pms) in let subst = paramsubst @ Inductive.ind_subst (fst ind) mib u in let (ctx, _) = mip.mind_nf_lc.(i - 1) in let (ctx, _) = List.chop mip.mind_consnrealdecls.(i - 1) ctx in Inductive.instantiate_context u subst nas ctx let cbv_subst_of_rel_context_instance mkclos sign args env = let rec aux subst sign l = let open Context.Rel.Declaration in match sign, l with | LocalAssum _ :: sign', a::args' -> aux (subs_cons a subst) sign' args' | LocalDef (_,c,_)::sign', args' -> aux (subs_cons (mkclos subst c) subst) sign' args' | [], [] -> subst | _ -> CErrors.anomaly (Pp.str "Instance and signature do not match.") in aux env (List.rev sign) (Array.to_list args) (* The main recursive functions * * Go under applications and cases/projections (pushed in the stack), * expand head constants or substitued de Bruijn, and try to a make a * constructor, a lambda or a fixp appear in the head. If not, it is a value * and is completely computed here. The head redexes are NOT reduced: * the function returns the pair of a cbv_value and its stack. * * Invariant: if the result of norm_head is CONSTR or (CO)FIXP, it last * argument is []. Because we must put all the applied terms in the * stack. *) let rec norm_head info env t stack = (* no reduction under binders *) match kind t with (* stack grows (remove casts) *) | App (head,args) -> (* Applied terms are normalized immediately; they could be computed when getting out of the stack *) let nargs = Array.map (cbv_stack_term info TOP env) args in norm_head info env head (stack_app nargs stack) | Case (ci,u,pms,p,iv,c,v) -> norm_head info env c (CASE(u,pms,p,v,iv,ci,env,stack)) | Cast (ct,_,_) -> norm_head info env ct stack | Proj (p, c) -> let p' = if red_set info.reds (fCONST (Projection.constant p)) && red_set info.reds fBETA then Projection.unfold p else p in norm_head info env c (PROJ (p', stack)) (* constants, axioms * the first pattern is CRUCIAL, n=0 happens very often: * when reducing closed terms, n is always 0 *) | Rel i -> (match expand_rel i env with | Inl (0,v) -> strip_appl v stack | Inl (n,v) -> strip_appl (shift_value n v) stack | Inr (n,None) -> (VAL(0, mkRel n), stack) | Inr (n,Some p) -> norm_head_ref (n-p) info env stack (RelKey p) t) | Var id -> norm_head_ref 0 info env stack (VarKey id) t | Const sp -> Reductionops.reduction_effect_hook info.env info.sigma (fst sp) (lazy (reify_stack t (strip_app stack))); norm_head_ref 0 info env stack (ConstKey sp) t | LetIn (_, b, _, c) -> (* zeta means letin are contracted; delta without zeta means we *) (* allow bindings but leave let's in place *) if red_set info.reds fZETA then (* New rule: for Cbv, Delta does not apply to locally bound variables or red_set info.reds fDELTA *) let env' = subs_cons (cbv_stack_term info TOP env b) env in norm_head info env' c stack else (CBN(t,env), stack) (* Should we consider a commutative cut ? *) | Evar ev -> (match Reductionops.safe_evar_value info.sigma ev with Some c -> norm_head info env c stack | None -> let e, xs = ev in let xs' = List.map (apply_env env) xs in (VAL(0, mkEvar (e,xs')), stack)) (* non-neutral cases *) | Lambda _ -> let ctxt,b = Term.decompose_lam t in (LAM(List.length ctxt, List.rev ctxt,b,env), stack) | Fix fix -> (FIXP(fix,env,[||]), stack) | CoFix cofix -> (COFIXP(cofix,env,[||]), stack) | Construct c -> (CONSTR(c, [||]), stack) | Array(u,t,def,ty) -> let ty = cbv_stack_term info TOP env ty in let len = Array.length t in let t = Parray.init (Uint63.of_int len) (fun i -> cbv_stack_term info TOP env t.(i)) (cbv_stack_term info TOP env def) in (ARRAY (u,t,ty), stack) (* neutral cases *) | (Sort _ | Meta _ | Ind _ | Int _ | Float _) -> (VAL(0, t), stack) | Prod _ -> (CBN(t,env), stack) and norm_head_ref k info env stack normt t = if red_set_ref info.reds normt then match cbv_value_cache info normt with | Declarations.Def body -> debug_cbv (fun () -> Pp.(str "Unfolding " ++ debug_pr_key normt)); strip_appl (shift_value k body) stack | Declarations.Primitive op -> let c = match normt with | ConstKey c -> c | RelKey _ | VarKey _ -> assert false in (PRIMITIVE(op,c,[||]),stack) | Declarations.OpaqueDef _ | Declarations.Undef _ -> debug_cbv (fun () -> Pp.(str "Not unfolding " ++ debug_pr_key normt)); (VAL(0,make_constr_ref k normt t),stack) else begin debug_cbv (fun () -> Pp.(str "Not unfolding " ++ debug_pr_key normt)); (VAL(0,make_constr_ref k normt t),stack) end (* cbv_stack_term performs weak reduction on constr t under the subs * env, with context stack, i.e. ([env]t stack). First computes weak * head normal form of t and checks if a redex appears with the stack. * If so, recursive call to reach the real head normal form. If not, * we build a value. *) and cbv_stack_term info stack env t = cbv_stack_value info env (norm_head info env t stack) and cbv_stack_value info env = function (* a lambda meets an application -> BETA *) | (LAM (nlams,ctxt,b,env), APP (args, stk)) when red_set info.reds fBETA -> let nargs = Array.length args in if nargs == nlams then cbv_stack_term info stk (subs_consn args 0 nargs env) b else if nlams < nargs then let env' = subs_consn args 0 nlams env in let eargs = Array.sub args nlams (nargs-nlams) in cbv_stack_term info (APP(eargs,stk)) env' b else let ctxt' = List.skipn nargs ctxt in LAM(nlams-nargs,ctxt', b, subs_consn args 0 nargs env) (* a Fix applied enough -> IOTA *) | (FIXP(fix,env,[||]), stk) when fixp_reducible info.reds fix stk -> let (envf,redfix) = contract_fixp env fix in cbv_stack_term info stk envf redfix (* constructor guard satisfied or Cofix in a Case -> IOTA *) | (COFIXP(cofix,env,[||]), stk) when cofixp_reducible info.reds cofix stk-> let (envf,redfix) = contract_cofixp env cofix in cbv_stack_term info stk envf redfix (* constructor in a Case -> IOTA *) | (CONSTR(((sp,n),_),[||]), APP(args,CASE(u,pms,_p,br,iv,ci,env,stk))) when red_set info.reds fMATCH -> let nargs = Array.length args - ci.ci_npar in let cargs = Array.sub args ci.ci_npar nargs in let env = if (Int.equal ci.ci_cstr_ndecls.(n - 1) ci.ci_cstr_nargs.(n - 1)) then (* no lets *) subs_consn cargs 0 nargs env else let mkclos env c = cbv_stack_term info TOP env c in let ctx = expand_branch info.env u pms (sp, n) br in cbv_subst_of_rel_context_instance mkclos ctx cargs env in cbv_stack_term info stk env (snd br.(n-1)) (* constructor of arity 0 in a Case -> IOTA *) | (CONSTR(((sp, n), _),[||]), CASE(u,pms,_,br,_,ci,env,stk)) when red_set info.reds fMATCH -> let env = if (Int.equal ci.ci_cstr_ndecls.(n - 1) ci.ci_cstr_nargs.(n - 1)) then (* no lets *) env else let mkclos env c = cbv_stack_term info TOP env c in let ctx = expand_branch info.env u pms (sp, n) br in cbv_subst_of_rel_context_instance mkclos ctx [||] env in cbv_stack_term info stk env (snd br.(n-1)) (* constructor in a Projection -> IOTA *) | (CONSTR(((sp,n),u),[||]), APP(args,PROJ(p,stk))) when red_set info.reds fMATCH && Projection.unfolded p -> let arg = args.(Projection.npars p + Projection.arg p) in cbv_stack_value info env (strip_appl arg stk) (* may be reduced later by application *) | (FIXP(fix,env,[||]), APP(appl,TOP)) -> FIXP(fix,env,appl) | (COFIXP(cofix,env,[||]), APP(appl,TOP)) -> COFIXP(cofix,env,appl) | (CONSTR(c,[||]), APP(appl,TOP)) -> CONSTR(c,appl) (* primitive apply to arguments *) | (PRIMITIVE(op,(_,u as c),[||]), APP(appl,stk)) -> let nargs = CPrimitives.arity op in let len = Array.length appl in if nargs <= len then let args = if len = nargs then appl else Array.sub appl 0 nargs in let stk = if nargs < len then stack_app (Array.sub appl nargs (len - nargs)) stk else stk in match VredNative.red_prim info.env () op u args with | Some (CONSTR (c, args)) -> (* args must be moved to the stack to allow future reductions *) cbv_stack_value info env (CONSTR(c, [||]), stack_app args stk) | Some v -> cbv_stack_value info env (v,stk) | None -> mkSTACK(PRIMITIVE(op,c,args), stk) else (* partial application *) (assert (stk = TOP); PRIMITIVE(op,c,appl)) (* definitely a value *) | (head,stk) -> mkSTACK(head, stk) and cbv_value_cache info ref = try KeyTable.find info.tab ref with Not_found -> let v = try let body = match ref with | RelKey n -> let open Context.Rel.Declaration in begin match Environ.lookup_rel n info.env with | LocalDef (_, c, _) -> lift n c | LocalAssum _ -> raise Not_found end | VarKey id -> let open Context.Named.Declaration in begin match Environ.lookup_named id info.env with | LocalDef (_, c, _) -> c | LocalAssum _ -> raise Not_found end | ConstKey cst -> Environ.constant_value_in info.env cst in let v = cbv_stack_term info TOP (subs_id 0) body in Declarations.Def v with | Environ.NotEvaluableConst (Environ.IsPrimitive (_u,op)) -> Declarations.Primitive op | Not_found | Environ.NotEvaluableConst _ -> Declarations.Undef None in KeyTable.add info.tab ref v; v (* When we are sure t will never produce a redex with its stack, we * normalize (even under binders) the applied terms and we build the * final term *) let rec apply_stack info t = function | TOP -> t | APP (args,st) -> apply_stack info (mkApp(t,Array.map (cbv_norm_value info) args)) st | CASE (u,pms,ty,br,iv,ci,env,st) -> (* FIXME: Prevent this expansion by caching whether an inductive contains let-bindings *) let (_, ty, _, _, br) = Inductive.expand_case info.env (ci, u, pms, ty, iv, mkProp, br) in let ty = let (_, mip) = Inductive.lookup_mind_specif info.env ci.ci_ind in Term.decompose_lam_n_decls (mip.Declarations.mind_nrealdecls + 1) ty in let mk_br c n = Term.decompose_lam_n_decls n c in let br = Array.map2 mk_br br ci.ci_cstr_ndecls in let map_ctx (nas, c) = let open Context.Rel.Declaration in let fold decl e = match decl with | LocalAssum _ -> subs_lift e | LocalDef (_, b, _) -> let b = cbv_stack_term info TOP e b in (* The let-binding persists, so we have to shift *) subs_shft (1, subs_cons b e) in let env = List.fold_right fold nas env in let nas = Array.of_list (List.rev_map get_annot nas) in (nas, cbv_norm_term info env c) in apply_stack info (mkCase (ci, u, Array.map (cbv_norm_term info env) pms, map_ctx ty, iv, t, Array.map map_ctx br)) st | PROJ (p, st) -> apply_stack info (mkProj (p, t)) st (* performs the reduction on a constr, and returns a constr *) and cbv_norm_term info env t = (* reduction under binders *) cbv_norm_value info (cbv_stack_term info TOP env t) (* reduction of a cbv_value to a constr *) and cbv_norm_value info = function (* reduction under binders *) | VAL (n,t) -> lift n t | STACK (0,v,stk) -> apply_stack info (cbv_norm_value info v) stk | STACK (n,v,stk) -> lift n (apply_stack info (cbv_norm_value info v) stk) | CBN(t,env) -> Constr.map_with_binders subs_lift (cbv_norm_term info) env t | LAM (n,ctxt,b,env) -> let nctxt = List.map_i (fun i (x,ty) -> (x,cbv_norm_term info (subs_liftn i env) ty)) 0 ctxt in Term.compose_lam (List.rev nctxt) (cbv_norm_term info (subs_liftn n env) b) | FIXP ((lij,(names,lty,bds)),env,args) -> mkApp (mkFix (lij, (names, Array.map (cbv_norm_term info env) lty, Array.map (cbv_norm_term info (subs_liftn (Array.length lty) env)) bds)), Array.map (cbv_norm_value info) args) | COFIXP ((j,(names,lty,bds)),env,args) -> mkApp (mkCoFix (j, (names,Array.map (cbv_norm_term info env) lty, Array.map (cbv_norm_term info (subs_liftn (Array.length lty) env)) bds)), Array.map (cbv_norm_value info) args) | CONSTR (c,args) -> mkApp(mkConstructU c, Array.map (cbv_norm_value info) args) | PRIMITIVE(op,c,args) -> mkApp(mkConstU c,Array.map (cbv_norm_value info) args) | ARRAY (u,t,ty) -> let ty = cbv_norm_value info ty in let t, def = Parray.to_array t in let def = cbv_norm_value info def in mkArray(u, Array.map (cbv_norm_value info) t, def, ty) (* with profiling *) let cbv_norm infos constr = let constr = EConstr.Unsafe.to_constr constr in EConstr.of_constr (with_stats (lazy (cbv_norm_term infos (subs_id 0) constr))) (* constant bodies are normalized at the first expansion *) let create_cbv_infos reds env sigma = { tab = KeyTable.create 91; reds; env; sigma }