aboutsummaryrefslogtreecommitdiffhomepage
path: root/kernel
diff options
context:
space:
mode:
Diffstat (limited to 'kernel')
-rw-r--r--kernel/cClosure.ml20
-rw-r--r--kernel/cClosure.mli1
-rw-r--r--kernel/nativecode.ml17
-rw-r--r--kernel/nativecode.mli2
-rw-r--r--kernel/nativeconv.ml11
-rw-r--r--kernel/nativeinstr.mli2
-rw-r--r--kernel/nativelambda.ml5
-rw-r--r--kernel/nativevalues.ml18
-rw-r--r--kernel/nativevalues.mli6
-rw-r--r--kernel/reduction.ml35
10 files changed, 55 insertions, 62 deletions
diff --git a/kernel/cClosure.ml b/kernel/cClosure.ml
index 4fd274ae1..219ea5b24 100644
--- a/kernel/cClosure.ml
+++ b/kernel/cClosure.ml
@@ -850,6 +850,14 @@ let contract_fix_vect fix =
in
(subs_cons(Array.init nfix make_body, env), thisbody)
+let unfold_projection info p =
+ if red_projection info.i_flags p
+ then
+ let open Declarations in
+ let pb = lookup_projection p (info_env info) in
+ Some (Zproj (pb.proj_npars, pb.proj_arg, Projection.constant p))
+ else None
+
(*********************************************************************)
(* A machine that inspects the head of a term until it finds an
atom or a subterm that may produce a redex (abstraction,
@@ -868,15 +876,9 @@ let rec knh info m stk =
| (None, stk') -> (m,stk'))
| FCast(t,_,_) -> knh info t stk
| FProj (p,c) ->
- let unf = Projection.unfolded p in
- if unf || red_set info.i_flags (fCONST (Projection.constant p)) then
- (match try Some (lookup_projection p (info_env info)) with Not_found -> None with
- | None -> (m, stk)
- | Some pb ->
- knh info c (Zproj (pb.Declarations.proj_npars, pb.Declarations.proj_arg,
- Projection.constant p)
- :: zupdate m stk))
- else (m,stk)
+ (match unfold_projection info p with
+ | None -> (m, stk)
+ | Some s -> knh info c (s :: zupdate m stk))
(* cases where knh stops *)
| (FFlex _|FLetIn _|FConstruct _|FEvar _|
diff --git a/kernel/cClosure.mli b/kernel/cClosure.mli
index 41db0af75..c43fc4623 100644
--- a/kernel/cClosure.mli
+++ b/kernel/cClosure.mli
@@ -166,6 +166,7 @@ val stack_tail : int -> stack -> stack
val stack_nth : stack -> int -> fconstr
val zip_term : (fconstr -> constr) -> constr -> stack -> constr
val eta_expand_stack : stack -> stack
+val unfold_projection : 'a infos -> Projection.t -> stack_member option
(** To lazy reduce a constr, create a [clos_infos] with
[create_clos_infos], inject the term to reduce with [inject]; then use
diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml
index 613b2f2ec..8fa254053 100644
--- a/kernel/nativecode.ml
+++ b/kernel/nativecode.ml
@@ -148,7 +148,7 @@ type symbol =
| SymbMatch of annot_sw
| SymbInd of inductive
| SymbMeta of metavariable
- | SymbEvar of existential
+ | SymbEvar of Evar.t
| SymbLevel of Univ.Level.t
let dummy_symb = SymbValue (dummy_value ())
@@ -162,8 +162,7 @@ let eq_symbol sy1 sy2 =
| SymbMatch sw1, SymbMatch sw2 -> eq_annot_sw sw1 sw2
| SymbInd ind1, SymbInd ind2 -> eq_ind ind1 ind2
| SymbMeta m1, SymbMeta m2 -> Int.equal m1 m2
- | SymbEvar (evk1,args1), SymbEvar (evk2,args2) ->
- Evar.equal evk1 evk2 && Array.for_all2 Constr.equal args1 args2
+ | SymbEvar evk1, SymbEvar evk2 -> Evar.equal evk1 evk2
| SymbLevel l1, SymbLevel l2 -> Univ.Level.equal l1 l2
| _, _ -> false
@@ -176,10 +175,7 @@ let hash_symbol symb =
| SymbMatch sw -> combinesmall 5 (hash_annot_sw sw)
| SymbInd ind -> combinesmall 6 (ind_hash ind)
| SymbMeta m -> combinesmall 7 m
- | SymbEvar (evk,args) ->
- let evh = Evar.hash evk in
- let hl = Array.fold_left (fun h t -> combine h (Constr.hash t)) evh args in
- combinesmall 8 hl
+ | SymbEvar evk -> combinesmall 8 (Evar.hash evk)
| SymbLevel l -> combinesmall 9 (Univ.Level.hash l)
module HashedTypeSymbol = struct
@@ -1047,11 +1043,12 @@ let ml_of_instance instance u =
let tyn = fresh_lname Anonymous in
let i = push_symbol (SymbMeta mv) in
MLapp(MLprimitive Mk_meta, [|get_meta_code i; MLlocal tyn|])
- | Levar(ev,ty) ->
+ | Levar(evk,ty,args) ->
let tyn = fresh_lname Anonymous in
- let i = push_symbol (SymbEvar ev) in
+ let i = push_symbol (SymbEvar evk) in
+ let args = MLarray(Array.map (ml_of_lam env l) args) in
MLlet(tyn, ml_of_lam env l ty,
- MLapp(MLprimitive Mk_evar, [|get_evar_code i;MLlocal tyn|]))
+ MLapp(MLprimitive Mk_evar, [|get_evar_code i;MLlocal tyn; args|]))
| Lprod(dom,codom) ->
let dom = ml_of_lam env l dom in
let codom = ml_of_lam env l codom in
diff --git a/kernel/nativecode.mli b/kernel/nativecode.mli
index d08f49095..7d20054f7 100644
--- a/kernel/nativecode.mli
+++ b/kernel/nativecode.mli
@@ -44,7 +44,7 @@ val get_ind : symbols -> int -> inductive
val get_meta : symbols -> int -> metavariable
-val get_evar : symbols -> int -> existential
+val get_evar : symbols -> int -> Evar.t
val get_level : symbols -> int -> Univ.Level.t
diff --git a/kernel/nativeconv.ml b/kernel/nativeconv.ml
index 9f9102f7d..bfa982136 100644
--- a/kernel/nativeconv.ml
+++ b/kernel/nativeconv.ml
@@ -54,13 +54,18 @@ and conv_accu env pb lvl k1 k2 cu =
conv_atom env pb lvl (atom_of_accu k1) (atom_of_accu k2) cu
else
let cu = conv_atom env pb lvl (atom_of_accu k1) (atom_of_accu k2) cu in
- List.fold_right2 (conv_val env CONV lvl) (args_of_accu k1) (args_of_accu k2) cu
+ Array.fold_right2 (conv_val env CONV lvl) (args_of_accu k1) (args_of_accu k2) cu
and conv_atom env pb lvl a1 a2 cu =
if a1 == a2 then cu
else
match a1, a2 with
- | Ameta _, _ | _, Ameta _ | Aevar _, _ | _, Aevar _ -> assert false
+ | Ameta (m1,_), Ameta (m2,_) ->
+ if Int.equal m1 m2 then cu else raise NotConvertible
+ | Aevar (ev1,_,args1), Aevar (ev2,_,args2) ->
+ if Evar.equal ev1 ev2 then
+ Array.fold_right2 (conv_val env CONV lvl) args1 args2 cu
+ else raise NotConvertible
| Arel i1, Arel i2 ->
if Int.equal i1 i2 then cu else raise NotConvertible
| Aind (ind1,u1), Aind (ind2,u2) ->
@@ -112,7 +117,7 @@ and conv_atom env pb lvl a1 a2 cu =
else conv_accu env CONV lvl ac1 ac2 cu
| Arel _, _ | Aind _, _ | Aconstant _, _ | Asort _, _ | Avar _, _
| Acase _, _ | Afix _, _ | Acofix _, _ | Acofixe _, _ | Aprod _, _
- | Aproj _, _ -> raise NotConvertible
+ | Aproj _, _ | Ameta _, _ | Aevar _, _ -> raise NotConvertible
(* Precondition length t1 = length f1 = length f2 = length t2 *)
and conv_fix env lvl t1 f1 t2 f2 cu =
diff --git a/kernel/nativeinstr.mli b/kernel/nativeinstr.mli
index 928283a4d..48ad88444 100644
--- a/kernel/nativeinstr.mli
+++ b/kernel/nativeinstr.mli
@@ -23,7 +23,7 @@ and lambda =
| Lrel of Name.t * int
| Lvar of Id.t
| Lmeta of metavariable * lambda (* type *)
- | Levar of existential * lambda (* type *)
+ | Levar of Evar.t * lambda (* type *) * lambda array (* arguments *)
| Lprod of lambda * lambda
| Llam of Name.t array * lambda
| Llet of Name.t * lambda * lambda
diff --git a/kernel/nativelambda.ml b/kernel/nativelambda.ml
index b333b0fb9..29b3e59da 100644
--- a/kernel/nativelambda.ml
+++ b/kernel/nativelambda.ml
@@ -453,11 +453,12 @@ let rec lambda_of_constr env sigma c =
let ty = meta_type sigma mv in
Lmeta (mv, lambda_of_constr env sigma ty)
- | Evar ev ->
+ | Evar (evk,args as ev) ->
(match evar_value sigma ev with
| None ->
let ty = evar_type sigma ev in
- Levar(ev, lambda_of_constr env sigma ty)
+ let args = Array.map (lambda_of_constr env sigma) args in
+ Levar(evk, lambda_of_constr env sigma ty, args)
| Some t -> lambda_of_constr env sigma t)
| Cast (c, _, _) -> lambda_of_constr env sigma c
diff --git a/kernel/nativevalues.ml b/kernel/nativevalues.ml
index ae66362ca..3d47b1672 100644
--- a/kernel/nativevalues.ml
+++ b/kernel/nativevalues.ml
@@ -61,7 +61,7 @@ type atom =
| Acofixe of t array * t array * int * t
| Aprod of Name.t * t * (t -> t)
| Ameta of metavariable * t
- | Aevar of existential * t
+ | Aevar of Evar.t * t * t array
| Aproj of Constant.t * accumulator
let accumulate_tag = 0
@@ -132,8 +132,8 @@ let mk_prod_accu s dom codom =
let mk_meta_accu mv ty =
mk_accu (Ameta (mv,ty))
-let mk_evar_accu ev ty =
- mk_accu (Aevar (ev,ty))
+let mk_evar_accu ev ty args =
+ mk_accu (Aevar (ev,ty,args))
let mk_proj_accu kn c =
mk_accu (Aproj (kn,c))
@@ -153,8 +153,7 @@ let accu_nargs (k:accumulator) =
let args_of_accu (k:accumulator) =
let nargs = accu_nargs k in
let f i = (Obj.magic (Obj.field (Obj.magic k) (nargs-i+2)) : t) in
- let t = Array.init nargs f in
- Array.to_list t
+ Array.init nargs f
let is_accu x =
let o = Obj.repr x in
@@ -179,11 +178,10 @@ let force_cofix (cofix : t) =
let atom = atom_of_accu accu in
match atom with
| Acofix(typ,norm,pos,f) ->
- let f = ref f in
- let args = List.rev (args_of_accu accu) in
- List.iter (fun x -> f := !f x) args;
- let v = !f (Obj.magic ()) in
- set_atom_of_accu accu (Acofixe(typ,norm,pos,v));
+ let args = args_of_accu accu in
+ let f = Array.fold_right (fun arg f -> f arg) args f in
+ let v = f (Obj.magic ()) in
+ set_atom_of_accu accu (Acofixe(typ,norm,pos,v));
v
| Acofixe(_,_,_,v) -> v
| _ -> cofix
diff --git a/kernel/nativevalues.mli b/kernel/nativevalues.mli
index 18b877745..993842740 100644
--- a/kernel/nativevalues.mli
+++ b/kernel/nativevalues.mli
@@ -51,7 +51,7 @@ type atom =
| Acofixe of t array * t array * int * t
| Aprod of Name.t * t * (t -> t)
| Ameta of metavariable * t
- | Aevar of existential * t
+ | Aevar of Evar.t * t (* type *) * t array (* arguments *)
| Aproj of Constant.t * accumulator
(* Constructors *)
@@ -68,7 +68,7 @@ val mk_prod_accu : Name.t -> t -> t -> t
val mk_fix_accu : rec_pos -> int -> t array -> t array -> t
val mk_cofix_accu : int -> t array -> t array -> t
val mk_meta_accu : metavariable -> t
-val mk_evar_accu : existential -> t -> t
+val mk_evar_accu : Evar.t -> t -> t array -> t
val mk_proj_accu : Constant.t -> accumulator -> t
val upd_cofix : t -> t -> unit
val force_cofix : t -> t
@@ -84,7 +84,7 @@ val napply : t -> t array -> t
val dummy_value : unit -> t
val atom_of_accu : accumulator -> atom
-val args_of_accu : accumulator -> t list
+val args_of_accu : accumulator -> t array
val accu_nargs : accumulator -> int
val cast_accu : t -> accumulator
diff --git a/kernel/reduction.ml b/kernel/reduction.ml
index 6104f56c6..1724f210d 100644
--- a/kernel/reduction.ml
+++ b/kernel/reduction.ml
@@ -308,17 +308,6 @@ let in_whnf (t,stk) =
| (FFlex _ | FProd _ | FEvar _ | FInd _ | FAtom _ | FRel _ | FProj _) -> true
| FLOCKED -> assert false
-let unfold_projection infos p c =
- let unf = Projection.unfolded p in
- if unf || RedFlags.red_set infos.i_flags (RedFlags.fCONST (Projection.constant p)) then
- (match try Some (lookup_projection p (info_env infos)) with Not_found -> None with
- | Some pb ->
- let s = Zproj (pb.Declarations.proj_npars, pb.Declarations.proj_arg,
- Projection.constant p) in
- Some (c, s)
- | None -> None)
- else None
-
(* Conversion between [lft1]term1 and [lft2]term2 *)
let rec ccnv cv_pb l2r infos lft1 lft2 term1 term2 cuniv =
eqappr cv_pb l2r infos (lft1, (term1,[])) (lft2, (term2,[])) cuniv
@@ -396,13 +385,13 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
(* Projections: prefer unfolding to first-order unification,
which will happen naturally if the terms c1, c2 are not in constructor
form *)
- (match unfold_projection infos p1 c1 with
- | Some (def1,s1) ->
- eqappr cv_pb l2r infos (lft1, (def1, (s1 :: v1))) appr2 cuniv
+ (match unfold_projection infos p1 with
+ | Some s1 ->
+ eqappr cv_pb l2r infos (lft1, (c1, (s1 :: v1))) appr2 cuniv
| None ->
- match unfold_projection infos p2 c2 with
- | Some (def2,s2) ->
- eqappr cv_pb l2r infos appr1 (lft2, (def2, (s2 :: v2))) cuniv
+ match unfold_projection infos p2 with
+ | Some s2 ->
+ eqappr cv_pb l2r infos appr1 (lft2, (c2, (s2 :: v2))) cuniv
| None ->
if Constant.equal (Projection.constant p1) (Projection.constant p2)
&& compare_stack_shape v1 v2 then
@@ -412,9 +401,9 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
raise NotConvertible)
| (FProj (p1,c1), t2) ->
- (match unfold_projection infos p1 c1 with
- | Some (def1,s1) ->
- eqappr cv_pb l2r infos (lft1, (def1, (s1 :: v1))) appr2 cuniv
+ (match unfold_projection infos p1 with
+ | Some s1 ->
+ eqappr cv_pb l2r infos (lft1, (c1, (s1 :: v1))) appr2 cuniv
| None ->
(match t2 with
| FFlex fl2 ->
@@ -425,9 +414,9 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
| _ -> raise NotConvertible))
| (t1, FProj (p2,c2)) ->
- (match unfold_projection infos p2 c2 with
- | Some (def2,s2) ->
- eqappr cv_pb l2r infos appr1 (lft2, (def2, (s2 :: v2))) cuniv
+ (match unfold_projection infos p2 with
+ | Some s2 ->
+ eqappr cv_pb l2r infos appr1 (lft2, (c2, (s2 :: v2))) cuniv
| None ->
(match t1 with
| FFlex fl1 ->