aboutsummaryrefslogtreecommitdiffhomepage
path: root/kernel/reduction.ml
diff options
context:
space:
mode:
Diffstat (limited to 'kernel/reduction.ml')
-rw-r--r--kernel/reduction.ml260
1 files changed, 109 insertions, 151 deletions
diff --git a/kernel/reduction.ml b/kernel/reduction.ml
index f26ab77a4..68f53c355 100644
--- a/kernel/reduction.ml
+++ b/kernel/reduction.ml
@@ -200,13 +200,9 @@ let is_cumul = function CUMUL -> true | CONV -> false
type 'a universe_compare =
{ (* Might raise NotConvertible *)
- compare : env -> conv_pb -> Sorts.t -> Sorts.t -> 'a -> 'a;
+ compare_sorts : env -> conv_pb -> Sorts.t -> Sorts.t -> 'a -> 'a;
compare_instances: flex:bool -> Univ.Instance.t -> Univ.Instance.t -> 'a -> 'a;
- conv_inductives : conv_pb -> (Declarations.mutual_inductive_body * int) -> Univ.Instance.t -> int ->
- Univ.Instance.t -> int -> 'a -> 'a;
- conv_constructors : (Declarations.mutual_inductive_body * int * int) ->
- Univ.Instance.t -> int -> Univ.Instance.t -> int -> 'a -> 'a;
- }
+ compare_cumul_instances : Univ.Constraint.t -> 'a -> 'a }
type 'a universe_state = 'a * 'a universe_compare
@@ -215,18 +211,68 @@ type ('a,'b) generic_conversion_function = env -> 'b universe_state -> 'a -> 'a
type 'a infer_conversion_function = env -> UGraph.t -> 'a -> 'a -> Univ.Constraint.t
let sort_cmp_universes env pb s0 s1 (u, check) =
- (check.compare env pb s0 s1 u, check)
+ (check.compare_sorts env pb s0 s1 u, check)
(* [flex] should be true for constants, false for inductive types and
constructors. *)
let convert_instances ~flex u u' (s, check) =
(check.compare_instances ~flex u u' s, check)
-
-let convert_inductives cv_pb ind u1 sv1 u2 sv2 (s, check) =
- (check.conv_inductives cv_pb ind u1 sv1 u2 sv2 s, check)
-let convert_constructors cons u1 sv1 u2 sv2 (s, check) =
- (check.conv_constructors cons u1 sv1 u2 sv2 s, check)
+let get_cumulativity_constraints cv_pb cumi u u' =
+ match cv_pb with
+ | CONV ->
+ Univ.ACumulativityInfo.eq_constraints cumi u u' Univ.Constraint.empty
+ | CUMUL ->
+ Univ.ACumulativityInfo.leq_constraints cumi u u' Univ.Constraint.empty
+
+let inductive_cumulativity_arguments (mind,ind) =
+ mind.Declarations.mind_nparams +
+ mind.Declarations.mind_packets.(ind).Declarations.mind_nrealargs
+
+let convert_inductives_gen cmp_instances cmp_cumul cv_pb (mind,ind) nargs u1 u2 s =
+ match mind.Declarations.mind_universes with
+ | Declarations.Monomorphic_ind _ ->
+ assert (Univ.Instance.length u1 = 0 && Univ.Instance.length u2 = 0);
+ s
+ | Declarations.Polymorphic_ind _ ->
+ cmp_instances u1 u2 s
+ | Declarations.Cumulative_ind cumi ->
+ let num_param_arity = inductive_cumulativity_arguments (mind,ind) in
+ if not (Int.equal num_param_arity nargs) then
+ cmp_instances u1 u2 s
+ else
+ let csts = get_cumulativity_constraints cv_pb cumi u1 u2 in
+ cmp_cumul csts s
+
+let convert_inductives cv_pb ind nargs u1 u2 (s, check) =
+ convert_inductives_gen (check.compare_instances ~flex:false) check.compare_cumul_instances
+ cv_pb ind nargs u1 u2 s, check
+
+let constructor_cumulativity_arguments (mind, ind, ctor) =
+ let nparamsctxt =
+ mind.Declarations.mind_nparams +
+ mind.Declarations.mind_packets.(ind).Declarations.mind_nrealargs
+ (* Context.Rel.length mind.Declarations.mind_params_ctxt *) in
+ nparamsctxt + mind.Declarations.mind_packets.(ind).Declarations.mind_consnrealargs.(ctor - 1)
+
+let convert_constructors_gen cmp_instances cmp_cumul (mind, ind, cns) nargs u1 u2 s =
+ match mind.Declarations.mind_universes with
+ | Declarations.Monomorphic_ind _ ->
+ assert (Univ.Instance.length u1 = 0 && Univ.Instance.length u2 = 0);
+ s
+ | Declarations.Polymorphic_ind _ ->
+ cmp_instances u1 u2 s
+ | Declarations.Cumulative_ind cumi ->
+ let num_cnstr_args = constructor_cumulativity_arguments (mind,ind,cns) in
+ if not (Int.equal num_cnstr_args nargs) then
+ cmp_instances u1 u2 s
+ else
+ let csts = get_cumulativity_constraints CONV cumi u1 u2 in
+ cmp_cumul csts s
+
+let convert_constructors ctor nargs u1 u2 (s, check) =
+ convert_constructors_gen (check.compare_instances ~flex:false) check.compare_cumul_instances
+ ctor nargs u1 u2 s, check
let conv_table_key infos k1 k2 cuniv =
if k1 == k2 then cuniv else
@@ -325,9 +371,8 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
if in_whnf st1' then (st1',st2') else whd_both st1' st2' in
let ((hd1,v1),(hd2,v2)) = whd_both st1 st2 in
let appr1 = (lft1,(hd1,v1)) and appr2 = (lft2,(hd2,v2)) in
- (* compute the lifts that apply to the head of the term (hd1 and hd2) *)
- let el1 = el_stack lft1 v1 in
- let el2 = el_stack lft2 v2 in
+ (** We delay the computation of the lifts that apply to the head of the term
+ with [el_stack] inside the branches where they are actually used. *)
match (fterm_of hd1, fterm_of hd2) with
(* case of leaves *)
| (FAtom a1, FAtom a2) ->
@@ -343,6 +388,8 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
| _ -> raise NotConvertible)
| (FEvar ((ev1,args1),env1), FEvar ((ev2,args2),env2)) ->
if Evar.equal ev1 ev2 then
+ let el1 = el_stack lft1 v1 in
+ let el2 = el_stack lft2 v2 in
let cuniv = convert_stacks l2r infos lft1 lft2 v1 v2 cuniv in
convert_vect l2r infos el1 el2
(Array.map (mk_clos env1) args1)
@@ -351,6 +398,8 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
(* 2 index known to be bound to no constant *)
| (FRel n, FRel m) ->
+ let el1 = el_stack lft1 v1 in
+ let el2 = el_stack lft2 v2 in
if Int.equal (reloc_rel n el1) (reloc_rel m el2)
then convert_stacks l2r infos lft1 lft2 v1 v2 cuniv
else raise NotConvertible
@@ -395,6 +444,8 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
| None ->
if Constant.equal (Projection.constant p1) (Projection.constant p2)
&& compare_stack_shape v1 v2 then
+ let el1 = el_stack lft1 v1 in
+ let el2 = el_stack lft2 v2 in
let u1 = ccnv CONV l2r infos el1 el2 c1 c2 cuniv in
convert_stacks l2r infos lft1 lft2 v1 v2 u1
else (* Two projections in WHNF: unfold *)
@@ -434,6 +485,8 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
anomaly (Pp.str "conversion was given ill-typed terms (FLambda).");
let (_,ty1,bd1) = destFLambda mk_clos hd1 in
let (_,ty2,bd2) = destFLambda mk_clos hd2 in
+ let el1 = el_stack lft1 v1 in
+ let el2 = el_stack lft2 v2 in
let cuniv = ccnv CONV l2r infos el1 el2 ty1 ty2 cuniv in
ccnv CONV l2r infos (el_lift el1) (el_lift el2) bd1 bd2 cuniv
@@ -441,6 +494,8 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
if not (is_empty_stack v1 && is_empty_stack v2) then
anomaly (Pp.str "conversion was given ill-typed terms (FProd).");
(* Luo's system *)
+ let el1 = el_stack lft1 v1 in
+ let el2 = el_stack lft2 v2 in
let cuniv = ccnv CONV l2r infos el1 el2 c1 c'1 cuniv in
ccnv cv_pb l2r infos (el_lift el1) (el_lift el2) c2 c'2 cuniv
@@ -509,15 +564,12 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
convert_stacks l2r infos lft1 lft2 v1 v2 cuniv
else
let mind = Environ.lookup_mind (fst ind1) (info_env infos) in
- let cuniv =
- match mind.Declarations.mind_universes with
- | Declarations.Monomorphic_ind _ | Declarations.Polymorphic_ind _ ->
- convert_instances ~flex:false u1 u2 cuniv
- | Declarations.Cumulative_ind cumi ->
- convert_inductives cv_pb (mind, snd ind1) u1 (CClosure.stack_args_size v1)
- u2 (CClosure.stack_args_size v2) cuniv
- in
- convert_stacks l2r infos lft1 lft2 v1 v2 cuniv
+ let nargs = CClosure.stack_args_size v1 in
+ if not (Int.equal nargs (CClosure.stack_args_size v2))
+ then raise NotConvertible
+ else
+ let cuniv = convert_inductives cv_pb (mind, snd ind1) nargs u1 u2 cuniv in
+ convert_stacks l2r infos lft1 lft2 v1 v2 cuniv
else raise NotConvertible
| (FConstruct ((ind1,j1),u1), FConstruct ((ind2,j2),u2)) ->
@@ -527,16 +579,12 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
convert_stacks l2r infos lft1 lft2 v1 v2 cuniv
else
let mind = Environ.lookup_mind (fst ind1) (info_env infos) in
- let cuniv =
- match mind.Declarations.mind_universes with
- | Declarations.Monomorphic_ind _ | Declarations.Polymorphic_ind _ ->
- convert_instances ~flex:false u1 u2 cuniv
- | Declarations.Cumulative_ind _ ->
- convert_constructors
- (mind, snd ind1, j1) u1 (CClosure.stack_args_size v1)
- u2 (CClosure.stack_args_size v2) cuniv
- in
- convert_stacks l2r infos lft1 lft2 v1 v2 cuniv
+ let nargs = CClosure.stack_args_size v1 in
+ if not (Int.equal nargs (CClosure.stack_args_size v2))
+ then raise NotConvertible
+ else
+ let cuniv = convert_constructors (mind, snd ind1, j1) nargs u1 u2 cuniv in
+ convert_stacks l2r infos lft1 lft2 v1 v2 cuniv
else raise NotConvertible
(* Eta expansion of records *)
@@ -562,6 +610,8 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
let fty2 = Array.map (mk_clos e2) tys2 in
let fcl1 = Array.map (mk_clos (subs_liftn n e1)) cl1 in
let fcl2 = Array.map (mk_clos (subs_liftn n e2)) cl2 in
+ let el1 = el_stack lft1 v1 in
+ let el2 = el_stack lft2 v2 in
let cuniv = convert_vect l2r infos el1 el2 fty1 fty2 cuniv in
let cuniv =
convert_vect l2r infos
@@ -577,6 +627,8 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
let fty2 = Array.map (mk_clos e2) tys2 in
let fcl1 = Array.map (mk_clos (subs_liftn n e1)) cl1 in
let fcl2 = Array.map (mk_clos (subs_liftn n e2)) cl2 in
+ let el1 = el_stack lft1 v1 in
+ let el2 = el_stack lft2 v2 in
let cuniv = convert_vect l2r infos el1 el2 fty1 fty2 cuniv in
let cuniv =
convert_vect l2r infos
@@ -653,84 +705,14 @@ let check_convert_instances ~flex u u' univs =
else raise NotConvertible
(* general conversion and inference functions *)
-let infer_check_conv_inductives
- infer_check_convert_instances
- infer_check_inductive_instances
- cv_pb (mind, ind) u1 sv1 u2 sv2 univs =
- match mind.Declarations.mind_universes with
- | Declarations.Monomorphic_ind _ | Declarations.Polymorphic_ind _ ->
- infer_check_convert_instances ~flex:false u1 u2 univs
- | Declarations.Cumulative_ind cumi ->
- let num_param_arity =
- mind.Declarations.mind_nparams + mind.Declarations.mind_packets.(ind).Declarations.mind_nrealargs
- in
- if not (num_param_arity = sv1 && num_param_arity = sv2) then
- infer_check_convert_instances ~flex:false u1 u2 univs
- else
- infer_check_inductive_instances cv_pb cumi u1 u2 univs
-
-let infer_check_conv_constructors
- infer_check_convert_instances
- infer_check_inductive_instances
- (mind, ind, cns) u1 sv1 u2 sv2 univs =
- match mind.Declarations.mind_universes with
- | Declarations.Monomorphic_ind _ | Declarations.Polymorphic_ind _ ->
- infer_check_convert_instances ~flex:false u1 u2 univs
- | Declarations.Cumulative_ind cumi ->
- let num_cnstr_args =
- let nparamsctxt =
- mind.Declarations.mind_nparams + mind.Declarations.mind_packets.(ind).Declarations.mind_nrealargs
- (* Context.Rel.length mind.Declarations.mind_params_ctxt *) in
- nparamsctxt + mind.Declarations.mind_packets.(ind).Declarations.mind_consnrealargs.(cns - 1)
- in
- if not (num_cnstr_args = sv1 && num_cnstr_args = sv2) then
- infer_check_convert_instances ~flex:false u1 u2 univs
- else
- infer_check_inductive_instances CONV cumi u1 u2 univs
-
-let check_inductive_instances cv_pb cumi u u' univs =
- let length_ind_instance =
- Univ.AUContext.size (Univ.ACumulativityInfo.univ_context cumi)
- in
- let ind_subtypctx = Univ.ACumulativityInfo.subtyp_context cumi in
- if not ((length_ind_instance = Univ.Instance.length u) &&
- (length_ind_instance = Univ.Instance.length u')) then
- anomaly (Pp.str "Invalid inductive subtyping encountered!")
- else
- let comp_cst =
- let comp_subst = (Univ.Instance.append u u') in
- Univ.AUContext.instantiate comp_subst ind_subtypctx
- in
- let comp_cst =
- match cv_pb with
- CONV ->
- let comp_cst' =
- let comp_subst = (Univ.Instance.append u' u) in
- Univ.AUContext.instantiate comp_subst ind_subtypctx
- in
- Univ.Constraint.union comp_cst comp_cst'
- | CUMUL -> comp_cst
- in
- if (UGraph.check_constraints comp_cst univs) then univs
- else raise NotConvertible
-
-let check_conv_inductives cv_pb ind u1 sv1 u2 sv2 univs =
- infer_check_conv_inductives
- check_convert_instances
- check_inductive_instances
- cv_pb ind u1 sv1 u2 sv2 univs
-
-let check_conv_constructors cns u1 sv1 u2 sv2 univs =
- infer_check_conv_constructors
- check_convert_instances
- check_inductive_instances
- cns u1 sv1 u2 sv2 univs
+let check_inductive_instances csts univs =
+ if (UGraph.check_constraints csts univs) then univs
+ else raise NotConvertible
let checked_universes =
- { compare = checked_sort_cmp_universes;
+ { compare_sorts = checked_sort_cmp_universes;
compare_instances = check_convert_instances;
- conv_inductives = check_conv_inductives;
- conv_constructors = check_conv_constructors}
+ compare_cumul_instances = check_inductive_instances; }
let infer_eq (univs, cstrs as cuniv) u u' =
if UGraph.check_eq univs u u' then cuniv
@@ -773,49 +755,13 @@ let infer_convert_instances ~flex u u' (univs,cstrs) =
else Univ.enforce_eq_instances u u' cstrs
in (univs, cstrs')
-let infer_inductive_instances cv_pb cumi u u' (univs, cstrs) =
- let length_ind_instance =
- Univ.AUContext.size (Univ.ACumulativityInfo.univ_context cumi)
- in
- let ind_subtypctx = Univ.ACumulativityInfo.subtyp_context cumi in
- if not ((length_ind_instance = Univ.Instance.length u) &&
- (length_ind_instance = Univ.Instance.length u')) then
- anomaly (Pp.str "Invalid inductive subtyping encountered!")
- else
- let comp_cst =
- let comp_subst = (Univ.Instance.append u u') in
- Univ.AUContext.instantiate comp_subst ind_subtypctx
- in
- let comp_cst =
- match cv_pb with
- CONV ->
- let comp_cst' =
- let comp_subst = (Univ.Instance.append u' u) in
- Univ.AUContext.instantiate comp_subst ind_subtypctx
- in
- Univ.Constraint.union comp_cst comp_cst'
- | CUMUL -> comp_cst
- in
- (univs, Univ.Constraint.union cstrs comp_cst)
-
-
-let infer_conv_inductives cv_pb ind u1 sv1 u2 sv2 univs =
- infer_check_conv_inductives
- infer_convert_instances
- infer_inductive_instances
- cv_pb ind u1 sv1 u2 sv2 univs
-
-let infer_conv_constructors cns u1 sv1 u2 sv2 univs =
- infer_check_conv_constructors
- infer_convert_instances
- infer_inductive_instances
- cns u1 sv1 u2 sv2 univs
-
-let inferred_universes : (UGraph.t * Univ.Constraint.t) universe_compare =
- { compare = infer_cmp_universes;
+let infer_inductive_instances csts (univs,csts') =
+ (univs, Univ.Constraint.union csts csts')
+
+let inferred_universes : (UGraph.t * Univ.Constraint.t) universe_compare =
+ { compare_sorts = infer_cmp_universes;
compare_instances = infer_convert_instances;
- conv_inductives = infer_conv_inductives;
- conv_constructors = infer_conv_constructors}
+ compare_cumul_instances = infer_inductive_instances; }
let gen_conv cv_pb l2r reds env evars univs t1 t2 =
let b =
@@ -936,6 +882,18 @@ let hnf_prod_app env t n =
let hnf_prod_applist env t nl =
List.fold_left (hnf_prod_app env) t nl
+let hnf_prod_applist_assum env n c l =
+ let rec app n subst t l =
+ if Int.equal n 0 then
+ if l == [] then substl subst t
+ else anomaly (Pp.str "Too many arguments.")
+ else match kind (whd_allnolet env t), l with
+ | Prod(_,_,c), arg::l -> app (n-1) (arg::subst) c l
+ | LetIn(_,b,_,c), _ -> app (n-1) (substl subst b::subst) c l
+ | _, [] -> anomaly (Pp.str "Not enough arguments.")
+ | _ -> anomaly (Pp.str "Not enough prod/let's.") in
+ app n [] c l
+
(* Dealing with arities *)
let dest_prod env =