aboutsummaryrefslogtreecommitdiff
path: root/kernel
diff options
context:
space:
mode:
Diffstat (limited to 'kernel')
-rw-r--r--kernel/cbytegen.ml35
-rw-r--r--kernel/cemitcodes.ml8
-rw-r--r--kernel/closure.ml217
-rw-r--r--kernel/closure.mli30
-rw-r--r--kernel/constr.ml238
-rw-r--r--kernel/constr.mli53
-rw-r--r--kernel/conv_oracle.mli8
-rw-r--r--kernel/cooking.ml110
-rw-r--r--kernel/cooking.mli3
-rw-r--r--kernel/declarations.mli44
-rw-r--r--kernel/declareops.ml220
-rw-r--r--kernel/declareops.mli2
-rw-r--r--kernel/entries.mli12
-rw-r--r--kernel/environ.ml174
-rw-r--r--kernel/environ.mli40
-rw-r--r--kernel/fast_typeops.ml475
-rw-r--r--kernel/fast_typeops.mli28
-rw-r--r--kernel/indtypes.ml377
-rw-r--r--kernel/indtypes.mli11
-rw-r--r--kernel/inductive.ml259
-rw-r--r--kernel/inductive.mli49
-rw-r--r--kernel/kernel.mllib2
-rw-r--r--kernel/mod_subst.ml51
-rw-r--r--kernel/mod_subst.mli19
-rw-r--r--kernel/mod_typing.ml35
-rw-r--r--kernel/modops.ml4
-rw-r--r--kernel/names.ml25
-rw-r--r--kernel/names.mli13
-rw-r--r--kernel/nativecode.ml81
-rw-r--r--kernel/nativeconv.ml8
-rw-r--r--kernel/nativeinstr.mli1
-rw-r--r--kernel/nativelambda.ml24
-rw-r--r--kernel/nativelambda.mli3
-rw-r--r--kernel/nativevalues.ml4
-rw-r--r--kernel/nativevalues.mli2
-rw-r--r--kernel/opaqueproof.ml11
-rw-r--r--kernel/opaqueproof.mli8
-rw-r--r--kernel/pre_env.ml2
-rw-r--r--kernel/pre_env.mli2
-rw-r--r--kernel/reduction.ml338
-rw-r--r--kernel/reduction.mli39
-rw-r--r--kernel/safe_typing.ml93
-rw-r--r--kernel/safe_typing.mli30
-rw-r--r--kernel/sorts.ml32
-rw-r--r--kernel/sorts.mli5
-rw-r--r--kernel/subtyping.ml67
-rw-r--r--kernel/term.ml39
-rw-r--r--kernel/term.mli51
-rw-r--r--kernel/term_typing.ml132
-rw-r--r--kernel/term_typing.mli8
-rw-r--r--kernel/type_errors.ml17
-rw-r--r--kernel/type_errors.mli13
-rw-r--r--kernel/typeops.ml363
-rw-r--r--kernel/typeops.mli56
-rw-r--r--kernel/univ.ml1822
-rw-r--r--kernel/univ.mli321
-rw-r--r--kernel/vars.ml86
-rw-r--r--kernel/vars.mli14
-rw-r--r--kernel/vconv.ml30
59 files changed, 4673 insertions, 1571 deletions
diff --git a/kernel/cbytegen.ml b/kernel/cbytegen.ml
index d0da84623f..894f887101 100644
--- a/kernel/cbytegen.ml
+++ b/kernel/cbytegen.ml
@@ -353,7 +353,7 @@ let rec str_const c =
| App(f,args) ->
begin
match kind_of_term f with
- | Construct((kn,j),i) ->
+ | Construct(((kn,j),i),u) ->
begin
let oib = lookup_mind kn !global_env in
let oip = oib.mind_packets.(j) in
@@ -422,8 +422,8 @@ let rec str_const c =
end
| _ -> Bconstr c
end
- | Ind ind -> Bstrconst (Const_ind ind)
- | Construct ((kn,j),i) ->
+ | Ind (ind,u) -> Bstrconst (Const_ind ind)
+ | Construct (((kn,j),i),u) ->
begin
(* spiwack: tries first to apply the run-time compilation
behavior of the constructor, as in 2/ above *)
@@ -487,11 +487,11 @@ let rec compile_fv reloc l sz cont =
(* Compiling constants *)
let rec get_allias env kn =
- let tps = (lookup_constant kn env).const_body_code in
- match Cemitcodes.force tps with
- | BCallias kn' -> get_allias env kn'
- | _ -> kn
-
+ let cb = lookup_constant kn env in
+ let tps = cb.const_body_code in
+ (match Cemitcodes.force tps with
+ | BCallias kn' -> get_allias env kn'
+ | _ -> kn)
(* Compiling expressions *)
@@ -499,12 +499,19 @@ let rec compile_constr reloc c sz cont =
match kind_of_term c with
| Meta _ -> invalid_arg "Cbytegen.compile_constr : Meta"
| Evar _ -> invalid_arg "Cbytegen.compile_constr : Evar"
+ | Proj (p,c) ->
+ (* compile_const reloc p [|c|] sz cont *)
+ let cb = lookup_constant p !global_env in
+ (* TODO: better representation of projections *)
+ let pb = Option.get cb.const_proj in
+ let args = Array.make pb.proj_npars mkProp in
+ compile_const reloc p Univ.Instance.empty (Array.append args [|c|]) sz cont
| Cast(c,_,_) -> compile_constr reloc c sz cont
| Rel i -> pos_rel i reloc sz :: cont
| Var id -> pos_named id reloc :: cont
- | Const kn -> compile_const reloc kn [||] sz cont
+ | Const (kn,u) -> compile_const reloc kn u [||] sz cont
| Sort _ | Ind _ | Construct _ ->
compile_str_cst reloc (str_const c) sz cont
@@ -531,7 +538,7 @@ let rec compile_constr reloc c sz cont =
begin
match kind_of_term f with
| Construct _ -> compile_str_cst reloc (str_const c) sz cont
- | Const kn -> compile_const reloc kn args sz cont
+ | Const (kn,u) -> compile_const reloc kn u args sz cont
| _ -> comp_app compile_constr compile_constr reloc f args sz cont
end
| Fix ((rec_args,init),(_,type_bodies,rec_bodies)) ->
@@ -682,14 +689,14 @@ and compile_str_cst reloc sc sz cont =
(* spiwack : compilation of constants with their arguments.
Makes a special treatment with 31-bit integer addition *)
and compile_const =
- fun reloc-> fun kn -> fun args -> fun sz -> fun cont ->
+ fun reloc-> fun kn u -> fun args -> fun sz -> fun cont ->
let nargs = Array.length args in
(* spiwack: checks if there is a specific way to compile the constant
if there is not, Not_found is raised, and the function
falls back on its normal behavior *)
try
Retroknowledge.get_vm_compiling_info (!global_env).retroknowledge
- (mkConst kn) reloc args sz cont
+ (mkConstU (kn,u)) reloc args sz cont
with Not_found ->
if Int.equal nargs 0 then
Kgetglobal (get_allias !global_env kn) :: cont
@@ -723,7 +730,7 @@ let compile_constant_body env = function
match kind_of_term body with
| Const kn' ->
(* we use the canonical name of the constant*)
- let con= constant_of_kn (canonical_con kn') in
+ let con= constant_of_kn (canonical_con (Univ.out_punivs kn')) in
BCallias (get_allias env con)
| _ ->
let res = compile env body in
@@ -751,7 +758,7 @@ let compile_structured_int31 fc args =
Const_b0
(Array.fold_left
(fun temp_i -> fun t -> match kind_of_term t with
- | Construct (_,d) -> 2*temp_i+d-1
+ | Construct ((_,d),_) -> 2*temp_i+d-1
| _ -> raise NotClosed)
0 args
)
diff --git a/kernel/cemitcodes.ml b/kernel/cemitcodes.ml
index 2b9ca425f9..2de8ef2bfa 100644
--- a/kernel/cemitcodes.ml
+++ b/kernel/cemitcodes.ml
@@ -320,16 +320,16 @@ let rec subst_strcst s sc =
match sc with
| Const_sorts _ | Const_b0 _ -> sc
| Const_bn(tag,args) -> Const_bn(tag,Array.map (subst_strcst s) args)
- | Const_ind(ind) -> let kn,i = ind in Const_ind((subst_ind s kn, i))
+ | Const_ind(ind) -> let kn,i = ind in Const_ind((subst_mind s kn, i))
let subst_patch s (ri,pos) =
match ri with
| Reloc_annot a ->
let (kn,i) = a.ci.ci_ind in
- let ci = {a.ci with ci_ind = (subst_ind s kn,i)} in
+ let ci = {a.ci with ci_ind = (subst_mind s kn,i)} in
(Reloc_annot {a with ci = ci},pos)
| Reloc_const sc -> (Reloc_const (subst_strcst s sc), pos)
- | Reloc_getglobal kn -> (Reloc_getglobal (fst (subst_con s kn)), pos)
+ | Reloc_getglobal kn -> (Reloc_getglobal (fst (subst_con_kn s kn)), pos)
let subst_to_patch s (code,pl,fv) =
code,List.rev_map (subst_patch s) pl,fv
@@ -341,7 +341,7 @@ type body_code =
let subst_body_code s = function
| BCdefined tp -> BCdefined (subst_to_patch s tp)
- | BCallias kn -> BCallias (fst (subst_con s kn))
+ | BCallias kn -> BCallias (fst (subst_con_kn s kn))
| BCconstant -> BCconstant
type to_patch_substituted = body_code substituted
diff --git a/kernel/closure.ml b/kernel/closure.ml
index 7b94ecfb85..fd3ab525e2 100644
--- a/kernel/closure.ml
+++ b/kernel/closure.ml
@@ -206,32 +206,39 @@ let unfold_red kn =
* instantiations (cbv or lazy) are.
*)
-type table_key = id_key
+type table_key = constant puniverses tableKey
+let eq_pconstant_key (c,u) (c',u') =
+ eq_constant_key c c' && Univ.Instance.eq u u'
+
module IdKeyHash =
struct
- type t = id_key
- let equal = Names.eq_id_key
open Hashset.Combine
+ type t = table_key
+ let equal = Names.eq_table_key eq_pconstant_key
let hash = function
- | ConstKey c -> combinesmall 1 (Constant.UserOrd.hash c)
+ | ConstKey (c, _) -> combinesmall 1 (Constant.UserOrd.hash c)
| VarKey id -> combinesmall 2 (Id.hash id)
| RelKey i -> combinesmall 3 (Int.hash i)
end
module KeyTable = Hashtbl.Make(IdKeyHash)
-let eq_table_key = Names.eq_id_key
+let eq_table_key = IdKeyHash.equal
-type 'a infos = {
- i_flags : reds;
+type 'a infos_cache = {
i_repr : 'a infos -> constr -> 'a;
i_env : env;
i_sigma : existential -> constr option;
i_rels : constr option array;
i_tab : 'a KeyTable.t }
+and 'a infos = {
+ i_flags : reds;
+ i_cache : 'a infos_cache }
+
let info_flags info = info.i_flags
+let info_env info = info.i_cache.i_env
let rec assoc_defined id = function
| [] -> raise Not_found
@@ -239,34 +246,34 @@ let rec assoc_defined id = function
| (id', Some c, _) :: ctxt ->
if Id.equal id id' then c else assoc_defined id ctxt
-let ref_value_cache info ref =
+let ref_value_cache ({i_cache = cache} as infos) ref =
try
- Some (KeyTable.find info.i_tab ref)
+ Some (KeyTable.find cache.i_tab ref)
with Not_found ->
try
let body =
match ref with
| RelKey n ->
- let len = Array.length info.i_rels in
+ let len = Array.length cache.i_rels in
let i = n - 1 in
let () = if i < 0 || len <= i then raise Not_found in
- begin match Array.unsafe_get info.i_rels i with
+ begin match Array.unsafe_get cache.i_rels i with
| None -> raise Not_found
| Some t -> lift n t
end
- | VarKey id -> assoc_defined id (named_context info.i_env)
- | ConstKey cst -> constant_value info.i_env cst
+ | VarKey id -> assoc_defined id (named_context cache.i_env)
+ | ConstKey cst -> constant_value_in cache.i_env cst
in
- let v = info.i_repr info body in
- KeyTable.add info.i_tab ref v;
+ let v = cache.i_repr infos body in
+ KeyTable.add cache.i_tab ref v;
Some v
with
| Not_found (* List.assoc *)
| NotEvaluableConst _ (* Const *)
-> None
-let evar_value info ev =
- info.i_sigma ev
+let evar_value cache ev =
+ cache.i_sigma ev
let defined_rels flags env =
(* if red_local_const (snd flags) then*)
@@ -282,12 +289,13 @@ let defined_rels flags env =
(* else (0,[])*)
let create mk_cl flgs env evars =
- { i_flags = flgs;
- i_repr = mk_cl;
- i_env = env;
- i_sigma = evars;
- i_rels = defined_rels flgs env;
- i_tab = KeyTable.create 17 }
+ let cache =
+ { i_repr = mk_cl;
+ i_env = env;
+ i_sigma = evars;
+ i_rels = defined_rels flgs env;
+ i_tab = KeyTable.create 17 }
+ in { i_flags = flgs; i_cache = cache }
(**********************************************************************)
@@ -327,9 +335,10 @@ and fterm =
| FAtom of constr (* Metas and Sorts *)
| FCast of fconstr * cast_kind * fconstr
| FFlex of table_key
- | FInd of inductive
- | FConstruct of constructor
+ | FInd of pinductive
+ | FConstruct of pconstructor
| FApp of fconstr * fconstr array
+ | FProj of constant * fconstr
| FFix of fixpoint * fconstr subs
| FCoFix of cofixpoint * fconstr subs
| FCases of case_info * fconstr * fconstr * fconstr array
@@ -362,6 +371,7 @@ let update v1 no t =
type stack_member =
| Zapp of fconstr array
| Zcase of case_info * fconstr * fconstr array
+ | Zproj of int * int * constant
| Zfix of fconstr * stack
| Zshift of int
| Zupdate of fconstr
@@ -494,6 +504,9 @@ let rec compact_constr (lg, subs as s) c k =
let (s, f') = compact_constr s f k in
let (s, v') = compact_vect s v k in
if f==f' && v==v' then s, c else s, mkApp(f',v')
+ | Proj (p,t) ->
+ let (s, t') = compact_constr s t k in
+ if t'==t then s, c else s, mkProj (p,t')
| Lambda(n,a,b) ->
let (s, a') = compact_constr s a k in
let (s, b') = compact_constr s b (k+1) in
@@ -559,7 +572,7 @@ let mk_clos e t =
| Meta _ | Sort _ -> { norm = Norm; term = FAtom t }
| Ind kn -> { norm = Norm; term = FInd kn }
| Construct kn -> { norm = Cstr; term = FConstruct kn }
- | (CoFix _|Lambda _|Fix _|Prod _|Evar _|App _|Case _|Cast _|LetIn _) ->
+ | (CoFix _|Lambda _|Fix _|Prod _|Evar _|App _|Case _|Cast _|LetIn _|Proj _) ->
{norm = Red; term = FCLOS(t,e)}
let mk_clos_vect env v = CArray.Fun1.map mk_clos env v
@@ -578,6 +591,9 @@ let mk_clos_deep clos_fun env t =
| App (f,v) ->
{ norm = Red;
term = FApp (clos_fun env f, CArray.Fun1.map clos_fun env v) }
+ | Proj (p,c) ->
+ { norm = Red;
+ term = FProj (p, clos_fun env c) }
| Case (ci,p,c,v) ->
{ norm = Red;
term = FCases (ci, clos_fun env p, clos_fun env c,
@@ -609,9 +625,9 @@ let rec to_constr constr_fun lfts v =
| FAtom c -> exliftn lfts c
| FCast (a,k,b) ->
mkCast (constr_fun lfts a, k, constr_fun lfts b)
- | FFlex (ConstKey op) -> mkConst op
- | FInd op -> mkInd op
- | FConstruct op -> mkConstruct op
+ | FFlex (ConstKey op) -> mkConstU op
+ | FInd op -> mkIndU op
+ | FConstruct op -> mkConstructU op
| FCases (ci,p,c,ve) ->
mkCase (ci, constr_fun lfts p,
constr_fun lfts c,
@@ -633,6 +649,9 @@ let rec to_constr constr_fun lfts v =
| FApp (f,ve) ->
mkApp (constr_fun lfts f,
CArray.Fun1.map constr_fun lfts ve)
+ | FProj (p,c) ->
+ mkProj (p,constr_fun lfts c)
+
| FLambda _ ->
let (na,ty,bd) = destFLambda mk_clos2 v in
mkLambda (na, constr_fun lfts ty,
@@ -688,6 +707,8 @@ let rec zip m stk rem = match stk with
| Zcase(ci,p,br)::s ->
let t = FCases(ci, p, m, br) in
zip {norm=neutr m.norm; term=t} s rem
+| Zproj (i,j,cst) :: s ->
+ zip {norm=neutr m.norm; term=FProj (cst,m)} s rem
| Zfix(fx,par)::s ->
zip fx par ((Zapp [|m|] :: s) :: rem)
| Zshift(n)::s ->
@@ -774,7 +795,7 @@ let rec get_args n tys f e stk =
(* Eta expansion: add a reference to implicit surrounding lambda at end of stack *)
let rec eta_expand_stack = function
- | (Zapp _ | Zfix _ | Zcase _ | Zshift _ | Zupdate _ as e) :: s ->
+ | (Zapp _ | Zfix _ | Zcase _ | Zshift _ | Zupdate _ | Zproj _ as e) :: s ->
e :: eta_expand_stack s
| [] ->
[Zshift 1; Zapp [|{norm=Norm; term= FRel 1}|]]
@@ -808,6 +829,64 @@ let rec drop_parameters depth n argstk =
| _ -> assert false
(* strip_update_shift_app only produces Zapp and Zshift items *)
+let rec get_parameters depth n argstk =
+ match argstk with
+ Zapp args::s ->
+ let q = Array.length args in
+ if n > q then Array.append args (get_parameters depth (n-q) s)
+ else if Int.equal n q then [||]
+ else Array.sub args 0 n
+ | Zshift(k)::s ->
+ get_parameters (depth-k) n s
+ | [] -> (* we know that n < stack_args_size(argstk) (if well-typed term) *)
+ if Int.equal n 0 then [||]
+ else raise Not_found (* Trying to eta-expand a partial application..., should do
+ eta expansion first? *)
+ | _ -> assert false
+ (* strip_update_shift_app only produces Zapp and Zshift items *)
+
+let eta_expand_ind_stack env lft (ind,u) m s (lft, h) =
+ let mib = lookup_mind (fst ind) env in
+ match mib.Declarations.mind_record with
+ | None -> raise Not_found
+ | Some (exp,_) ->
+ let pars = mib.Declarations.mind_nparams in
+ let h' = fapp_stack h in
+ let (depth, args, _) = strip_update_shift_app m s in
+ let paramargs = get_parameters depth pars args in
+ let subs = subs_cons (Array.append paramargs [|h'|], subs_id 0) in
+ let fexp = mk_clos subs exp in
+ (lft, (fexp, []))
+
+let eta_expand_ind_stacks env ind m s h =
+ let mib = lookup_mind (fst ind) env in
+ match mib.Declarations.mind_record with
+ | Some (exp,projs) when Array.length projs > 0 ->
+ let pars = mib.Declarations.mind_nparams in
+ let h' = fapp_stack h in
+ let (depth, args, _) = strip_update_shift_app m s in
+ let primitive = Environ.is_projection projs.(0) env in
+ if primitive then
+ let s' = drop_parameters depth pars args in
+ (* Construct, pars1 .. parsm :: arg1...argn :: s ~= (t, []) ->
+ arg1..argn :: s ~= (proj1 t...projn t) s
+ *)
+ let hstack = Array.map (fun p -> { norm = Red;
+ term = FProj (p, h') }) projs in
+ s', [Zapp hstack]
+ else raise Not_found (* disallow eta-exp for non-primitive records *)
+ | _ -> raise Not_found
+
+let rec project_nth_arg n argstk =
+ match argstk with
+ | Zapp args :: s ->
+ let q = Array.length args in
+ if n >= q then project_nth_arg (n - q) s
+ else (* n < q *) args.(n)
+ | _ -> assert false
+ (* After drop_parameters we have a purely applicative stack *)
+
+
(* Iota reduction: expansion of a fixpoint.
* Given a fixpoint and a substitution, returns the corresponding
* fixpoint body, and the substitution in which it should be
@@ -832,39 +911,48 @@ let contract_fix_vect fix =
in
(subs_cons(Array.init nfix make_body, env), thisbody)
-
(*********************************************************************)
(* A machine that inspects the head of a term until it finds an
atom or a subterm that may produce a redex (abstraction,
constructor, cofix, letin, constant), or a neutral term (product,
inductive) *)
-let rec knh m stk =
+let rec knh info m stk =
match m.term with
- | FLIFT(k,a) -> knh a (zshift k stk)
- | FCLOS(t,e) -> knht e t (zupdate m stk)
+ | FLIFT(k,a) -> knh info a (zshift k stk)
+ | FCLOS(t,e) -> knht info e t (zupdate m stk)
| FLOCKED -> assert false
- | FApp(a,b) -> knh a (append_stack b (zupdate m stk))
- | FCases(ci,p,t,br) -> knh t (Zcase(ci,p,br)::zupdate m stk)
+ | FApp(a,b) -> knh info a (append_stack b (zupdate m stk))
+ | FCases(ci,p,t,br) -> knh info t (Zcase(ci,p,br)::zupdate m stk)
| FFix(((ri,n),(_,_,_)),_) ->
(match get_nth_arg m ri.(n) stk with
- (Some(pars,arg),stk') -> knh arg (Zfix(m,pars)::stk')
+ (Some(pars,arg),stk') -> knh info arg (Zfix(m,pars)::stk')
| (None, stk') -> (m,stk'))
- | FCast(t,_,_) -> knh t stk
+ | FCast(t,_,_) -> knh info t stk
+ | FProj (p,c) ->
+ if red_set info.i_flags (fCONST 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, p)
+ :: zupdate m stk))
+ else (m,stk)
+
(* cases where knh stops *)
| (FFlex _|FLetIn _|FConstruct _|FEvar _|
FCoFix _|FLambda _|FRel _|FAtom _|FInd _|FProd _) ->
(m, stk)
(* The same for pure terms *)
-and knht e t stk =
+and knht info e t stk =
match kind_of_term t with
| App(a,b) ->
- knht e a (append_stack (mk_clos_vect e b) stk)
+ knht info e a (append_stack (mk_clos_vect e b) stk)
| Case(ci,p,t,br) ->
- knht e t (Zcase(ci, mk_clos e p, mk_clos_vect e br)::stk)
- | Fix _ -> knh (mk_clos2 e t) stk
- | Cast(a,_,_) -> knht e a stk
- | Rel n -> knh (clos_rel e n) stk
+ knht info e t (Zcase(ci, mk_clos e p, mk_clos_vect e br)::stk)
+ | Fix _ -> knh info (mk_clos2 e t) stk
+ | Cast(a,_,_) -> knht info e a stk
+ | Rel n -> knh info (clos_rel e n) stk
+ | Proj (p,c) -> knh info (mk_clos2 e t) stk
| (Lambda _|Prod _|Construct _|CoFix _|Ind _|
LetIn _|Const _|Var _|Evar _|Meta _|Sort _) ->
(mk_clos2 e t, stk)
@@ -879,8 +967,8 @@ let rec knr info m stk =
(match get_args n tys f e stk with
Inl e', s -> knit info e' f s
| Inr lam, s -> (lam,s))
- | FFlex(ConstKey kn) when red_set info.i_flags (fCONST kn) ->
- (match ref_value_cache info (ConstKey kn) with
+ | FFlex(ConstKey (kn,_ as c)) when red_set info.i_flags (fCONST kn) ->
+ (match ref_value_cache info (ConstKey c) with
Some v -> kni info v stk
| None -> (set_norm m; (m,stk)))
| FFlex(VarKey id) when red_set info.i_flags (fVAR id) ->
@@ -891,7 +979,7 @@ let rec knr info m stk =
(match ref_value_cache info (RelKey k) with
Some v -> kni info v stk
| None -> (set_norm m; (m,stk)))
- | FConstruct(ind,c) when red_set info.i_flags fIOTA ->
+ | FConstruct((ind,c),u) when red_set info.i_flags fIOTA ->
(match strip_update_shift_app m stk with
(depth, args, Zcase(ci,_,br)::s) ->
assert (ci.ci_npar>=0);
@@ -902,6 +990,10 @@ let rec knr info m stk =
let stk' = par @ append_stack [|rarg|] s in
let (fxe,fxbd) = contract_fix_vect fx.term in
knit info fxe fxbd stk'
+ | (depth, args, Zproj (n, m, cst)::s) ->
+ let rargs = drop_parameters depth n args in
+ let rarg = project_nth_arg m rargs in
+ kni info rarg s
| (_,args,s) -> (m,args@s))
| FCoFix _ when red_set info.i_flags fIOTA ->
(match strip_update_shift_app m stk with
@@ -912,17 +1004,17 @@ let rec knr info m stk =
| FLetIn (_,v,_,bd,e) when red_set info.i_flags fZETA ->
knit info (subs_cons([|v|],e)) bd stk
| FEvar(ev,env) ->
- (match evar_value info ev with
+ (match evar_value info.i_cache ev with
Some c -> knit info env c stk
| None -> (m,stk))
| _ -> (m,stk)
(* Computes the weak head normal form of a term *)
and kni info m stk =
- let (hm,s) = knh m stk in
+ let (hm,s) = knh info m stk in
knr info hm s
and knit info e t stk =
- let (ht,s) = knht e t stk in
+ let (ht,s) = knht info e t stk in
knr info ht s
let kh info v stk = fapp_stack(kni info v stk)
@@ -937,6 +1029,9 @@ let rec zip_term zfun m stk =
| Zcase(ci,p,br)::s ->
let t = mkCase(ci, zfun p, m, Array.map zfun br) in
zip_term zfun t s
+ | Zproj(_,_,p)::s ->
+ let t = mkProj (p, m) in
+ zip_term zfun t s
| Zfix(fx,par)::s ->
let h = mkApp(zip_term zfun (zfun fx) par,[|m|]) in
zip_term zfun h s
@@ -985,6 +1080,8 @@ and norm_head info m =
mkFix(n,(na, CArray.Fun1.map kl info ftys, CArray.Fun1.map kl info fbds))
| FEvar((i,args),env) ->
mkEvar(i, Array.map (fun a -> kl info (mk_clos env a)) args)
+ | FProj (p,c) ->
+ mkProj (p, kl info c)
| t -> term_of_fconstr m
(* Initialization and then normalization *)
@@ -1009,6 +1106,20 @@ type clos_infos = fconstr infos
let create_clos_infos ?(evars=fun _ -> None) flgs env =
create (fun _ -> inject) flgs env evars
-let oracle_of_infos { i_env } = Environ.oracle i_env
-
-let unfold_reference = ref_value_cache
+let oracle_of_infos infos = Environ.oracle infos.i_cache.i_env
+
+let infos_with_reds infos reds =
+ { infos with i_flags = reds }
+
+let unfold_reference info key =
+ match key with
+ | ConstKey (kn,_) ->
+ if red_set info.i_flags (fCONST kn) then
+ ref_value_cache info key
+ else None
+ | VarKey i ->
+ if red_set info.i_flags (fVAR i) then
+ ref_value_cache info key
+ else None
+ | _ -> ref_value_cache info key
+
diff --git a/kernel/closure.mli b/kernel/closure.mli
index 19baedf276..ee35e7d49b 100644
--- a/kernel/closure.mli
+++ b/kernel/closure.mli
@@ -80,14 +80,20 @@ val unfold_side_red : reds
val unfold_red : evaluable_global_reference -> reds
(***********************************************************************)
-type table_key = id_key
+type table_key = constant puniverses tableKey
+
+type 'a infos_cache
+type 'a infos = {
+ i_flags : reds;
+ i_cache : 'a infos_cache }
-type 'a infos
val ref_value_cache: 'a infos -> table_key -> 'a option
-val info_flags: 'a infos -> reds
val create: ('a infos -> constr -> 'a) -> reds -> env ->
(existential -> constr option) -> 'a infos
-val evar_value : 'a infos -> existential -> constr option
+val evar_value : 'a infos_cache -> existential -> constr option
+
+val info_env : 'a infos -> env
+val info_flags: 'a infos -> reds
(***********************************************************************
s Lazy reduction. *)
@@ -104,9 +110,10 @@ type fterm =
| FAtom of constr (** Metas and Sorts *)
| FCast of fconstr * cast_kind * fconstr
| FFlex of table_key
- | FInd of inductive
- | FConstruct of constructor
+ | FInd of inductive puniverses
+ | FConstruct of constructor puniverses
| FApp of fconstr * fconstr array
+ | FProj of constant * fconstr
| FFix of fixpoint * fconstr subs
| FCoFix of cofixpoint * fconstr subs
| FCases of case_info * fconstr * fconstr * fconstr array
@@ -126,6 +133,7 @@ type fterm =
type stack_member =
| Zapp of fconstr array
| Zcase of case_info * fconstr * fconstr array
+ | Zproj of int * int * constant
| Zfix of fconstr * stack
| Zshift of int
| Zupdate of fconstr
@@ -159,11 +167,13 @@ val destFLambda :
(fconstr subs -> constr -> fconstr) -> fconstr -> Name.t * fconstr * fconstr
(** Global and local constant cache *)
-type clos_infos
+type clos_infos = fconstr infos
val create_clos_infos :
?evars:(existential->constr option) -> reds -> env -> clos_infos
val oracle_of_infos : clos_infos -> Conv_oracle.oracle
+val infos_with_reds : clos_infos -> reds -> clos_infos
+
(** Reduction function *)
(** [norm_val] is for strong normalization *)
@@ -177,6 +187,12 @@ val whd_val : clos_infos -> fconstr -> constr
val whd_stack :
clos_infos -> fconstr -> stack -> fconstr * stack
+val eta_expand_ind_stack : env -> lift -> pinductive -> fconstr -> stack ->
+ (lift * (fconstr * stack)) -> lift * (fconstr * stack)
+
+val eta_expand_ind_stacks : env -> inductive -> fconstr -> stack ->
+ (fconstr * stack) -> stack * stack
+
(** Conversion auxiliary functions to do step by step normalisation *)
(** [unfold_reference] unfolds references in a [fconstr] *)
diff --git a/kernel/constr.ml b/kernel/constr.ml
index e9e21d30d9..89c138a085 100644
--- a/kernel/constr.ml
+++ b/kernel/constr.ml
@@ -25,7 +25,7 @@
open Util
open Names
-
+open Univ
type existential_key = Evar.t
type metavariable = int
@@ -61,6 +61,10 @@ type ('constr, 'types) pfixpoint =
(int array * int) * ('constr, 'types) prec_declaration
type ('constr, 'types) pcofixpoint =
int * ('constr, 'types) prec_declaration
+type 'a puniverses = 'a Univ.puniverses
+type pconstant = constant puniverses
+type pinductive = inductive puniverses
+type pconstructor = constructor puniverses
(* [Var] is used for named variables and [Rel] for variables as
de Bruijn indices. *)
@@ -75,13 +79,13 @@ type ('constr, 'types) kind_of_term =
| Lambda of Name.t * 'types * 'constr
| LetIn of Name.t * 'constr * 'types * 'constr
| App of 'constr * 'constr array
- | Const of constant
- | Ind of inductive
- | Construct of constructor
+ | Const of pconstant
+ | Ind of pinductive
+ | Construct of pconstructor
| Case of case_info * 'constr * 'constr * 'constr array
| Fix of ('constr, 'types) pfixpoint
| CoFix of ('constr, 'types) pcofixpoint
-
+ | Proj of constant * 'constr
(* constr is the fixpoint of the previous type. Requires option
-rectypes of the Caml compiler to be set *)
type t = (t,t) kind_of_term
@@ -139,19 +143,29 @@ let mkApp (f, a) =
| App (g, cl) -> App (g, Array.append cl a)
| _ -> App (f, a)
+let out_punivs (a, _) = a
+let map_puniverses f (x,u) = (f x, u)
+let in_punivs a = (a, Univ.Instance.empty)
+
(* Constructs a constant *)
-let mkConst c = Const c
+let mkConst c = Const (in_punivs c)
+let mkConstU c = Const c
+
+(* Constructs an applied projection *)
+let mkProj (p,c) = Proj (p,c)
(* Constructs an existential variable *)
let mkEvar e = Evar e
(* Constructs the ith (co)inductive type of the block named kn *)
-let mkInd m = Ind m
+let mkInd m = Ind (in_punivs m)
+let mkIndU m = Ind m
(* Constructs the jth constructor of the ith (co)inductive type of the
- block named kn. The array of terms correspond to the variables
- introduced in the section *)
-let mkConstruct c = Construct c
+ block named kn. *)
+let mkConstruct c = Construct (in_punivs c)
+let mkConstructU c = Construct c
+let mkConstructUi ((ind,u),i) = Construct ((ind,i),u)
(* Constructs the term <p>Case c of c1 | c2 .. | cn end *)
let mkCase (ci, p, c, ac) = Case (ci, p, c, ac)
@@ -225,6 +239,7 @@ let fold f acc c = match kind c with
| Lambda (_,t,c) -> f (f acc t) c
| LetIn (_,b,t,c) -> f (f (f acc b) t) c
| App (c,l) -> Array.fold_left f (f acc c) l
+ | Proj (p,c) -> f acc c
| Evar (_,l) -> Array.fold_left f acc l
| Case (_,p,c,bl) -> Array.fold_left f (f (f acc p) c) bl
| Fix (_,(lna,tl,bl)) ->
@@ -244,6 +259,7 @@ let iter f c = match kind c with
| Lambda (_,t,c) -> f t; f c
| LetIn (_,b,t,c) -> f b; f t; f c
| App (c,l) -> f c; Array.iter f l
+ | Proj (p,c) -> f c
| Evar (_,l) -> Array.iter f l
| Case (_,p,c,bl) -> f p; f c; Array.iter f bl
| Fix (_,(_,tl,bl)) -> Array.iter f tl; Array.iter f bl
@@ -265,6 +281,7 @@ let iter_with_binders g f n c = match kind c with
| App (c,l) -> f n c; CArray.Fun1.iter f n l
| Evar (_,l) -> CArray.Fun1.iter f n l
| Case (_,p,c,bl) -> f n p; f n c; CArray.Fun1.iter f n bl
+ | Proj (p,c) -> f n c
| Fix (_,(_,tl,bl)) ->
CArray.Fun1.iter f n tl;
CArray.Fun1.iter f (iterate g (Array.length tl) n) bl
@@ -305,6 +322,10 @@ let map f c = match kind c with
let l' = Array.smartmap f l in
if b'==b && l'==l then c
else mkApp (b', l')
+ | Proj (p,t) ->
+ let t' = f t in
+ if t' == t then c
+ else mkProj (p, t')
| Evar (e,l) ->
let l' = Array.smartmap f l in
if l'==l then c
@@ -413,6 +434,10 @@ let map_with_binders g f l c0 = match kind c0 with
let al' = CArray.Fun1.smartmap f l al in
if c' == c && al' == al then c0
else mkApp (c', al')
+ | Proj (p, t) ->
+ let t' = f l t in
+ if t' == t then c0
+ else mkProj (p, t')
| Evar (e, al) ->
let al' = CArray.Fun1.smartmap f l al in
if al' == al then c0
@@ -435,13 +460,13 @@ let map_with_binders g f l c0 = match kind c0 with
let bl' = CArray.Fun1.smartmap f l' bl in
mkCoFix (ln,(lna,tl',bl'))
-(* [compare f c1 c2] compare [c1] and [c2] using [f] to compare
- the immediate subterms of [c1] of [c2] if needed; Cast's,
+(* [compare_head_gen u s f c1 c2] compare [c1] and [c2] using [f] to compare
+ the immediate subterms of [c1] of [c2] if needed, [u] to compare universe
+ instances and [s] to compare sorts; Cast's,
application associativity, binders name and Cases annotations are
not taken into account *)
-
-let compare_head f t1 t2 =
+let compare_head_gen eq_universes eq_sorts f t1 t2 =
match kind t1, kind t2 with
| Rel n1, Rel n2 -> Int.equal n1 n2
| Meta m1, Meta m2 -> Int.equal m1 m2
@@ -458,9 +483,10 @@ let compare_head f t1 t2 =
Int.equal (Array.length l1) (Array.length l2) &&
f c1 c2 && Array.equal f l1 l2
| Evar (e1,l1), Evar (e2,l2) -> Evar.equal e1 e2 && Array.equal f l1 l2
- | Const c1, Const c2 -> eq_constant c1 c2
- | Ind c1, Ind c2 -> eq_ind c1 c2
- | Construct c1, Construct c2 -> eq_constructor c1 c2
+ | Proj (p1,c1), Proj (p2,c2) -> eq_constant p1 p2 && f c1 c2
+ | Const (c1,u1), Const (c2,u2) -> eq_constant c1 c2 && eq_universes true u1 u2
+ | Ind (c1,u1), Ind (c2,u2) -> eq_ind c1 c2 && eq_universes false u1 u2
+ | Construct (c1,u1), Construct (c2,u2) -> eq_constructor c1 c2 && eq_universes false u1 u2
| Case (_,p1,c1,bl1), Case (_,p2,c2,bl2) ->
f p1 p2 && f c1 c2 && Array.equal f bl1 bl2
| Fix ((ln1, i1),(_,tl1,bl1)), Fix ((ln2, i2),(_,tl2,bl2)) ->
@@ -470,6 +496,44 @@ let compare_head f t1 t2 =
Int.equal ln1 ln2 && Array.equal f tl1 tl2 && Array.equal f bl1 bl2
| _ -> false
+let compare_head = compare_head_gen (fun _ -> Univ.Instance.eq) Sorts.equal
+
+(* [compare_head_gen_leq u s sl eq leq c1 c2] compare [c1] and [c2] using [eq] to compare
+ the immediate subterms of [c1] of [c2] for conversion if needed, [leq] for cumulativity,
+ [u] to compare universe instances and [s] to compare sorts; Cast's,
+ application associativity, binders name and Cases annotations are
+ not taken into account *)
+
+let compare_head_gen_leq eq_universes eq_sorts leq_sorts eq leq t1 t2 =
+ match kind t1, kind t2 with
+ | Rel n1, Rel n2 -> Int.equal n1 n2
+ | Meta m1, Meta m2 -> Int.equal m1 m2
+ | Var id1, Var id2 -> Int.equal (id_ord id1 id2) 0
+ | Sort s1, Sort s2 -> leq_sorts s1 s2
+ | Cast (c1,_,_), _ -> leq c1 t2
+ | _, Cast (c2,_,_) -> leq t1 c2
+ | Prod (_,t1,c1), Prod (_,t2,c2) -> eq t1 t2 && leq c1 c2
+ | Lambda (_,t1,c1), Lambda (_,t2,c2) -> eq t1 t2 && eq c1 c2
+ | LetIn (_,b1,t1,c1), LetIn (_,b2,t2,c2) -> eq b1 b2 && eq t1 t2 && leq c1 c2
+ | App (Cast(c1, _, _),l1), _ -> leq (mkApp (c1,l1)) t2
+ | _, App (Cast (c2, _, _),l2) -> leq t1 (mkApp (c2,l2))
+ | App (c1,l1), App (c2,l2) ->
+ Int.equal (Array.length l1) (Array.length l2) &&
+ eq c1 c2 && Array.equal eq l1 l2
+ | Proj (p1,c1), Proj (p2,c2) -> eq_constant p1 p2 && eq c1 c2
+ | Evar (e1,l1), Evar (e2,l2) -> Evar.equal e1 e2 && Array.equal eq l1 l2
+ | Const (c1,u1), Const (c2,u2) -> eq_constant c1 c2 && eq_universes true u1 u2
+ | Ind (c1,u1), Ind (c2,u2) -> eq_ind c1 c2 && eq_universes false u1 u2
+ | Construct (c1,u1), Construct (c2,u2) -> eq_constructor c1 c2 && eq_universes false u1 u2
+ | Case (_,p1,c1,bl1), Case (_,p2,c2,bl2) ->
+ eq p1 p2 && eq c1 c2 && Array.equal eq bl1 bl2
+ | Fix ((ln1, i1),(_,tl1,bl1)), Fix ((ln2, i2),(_,tl2,bl2)) ->
+ Int.equal i1 i2 && Array.equal Int.equal ln1 ln2
+ && Array.equal eq tl1 tl2 && Array.equal eq bl1 bl2
+ | CoFix(ln1,(_,tl1,bl1)), CoFix(ln2,(_,tl2,bl2)) ->
+ Int.equal ln1 ln2 && Array.equal eq tl1 tl2 && Array.equal eq bl1 bl2
+ | _ -> false
+
(*******************************)
(* alpha conversion functions *)
(*******************************)
@@ -477,10 +541,81 @@ let compare_head f t1 t2 =
(* alpha conversion : ignore print names and casts *)
let rec eq_constr m n =
- (m == n) || compare_head eq_constr m n
+ (m == n) || compare_head_gen (fun _ -> Univ.Instance.eq) Sorts.equal eq_constr m n
+
+(** Strict equality of universe instances. *)
+let compare_constr = compare_head_gen (fun _ -> Univ.Instance.eq) Sorts.equal
let equal m n = eq_constr m n (* to avoid tracing a recursive fun *)
+let eq_constr_univs univs m n =
+ if m == n then true
+ else
+ let eq_universes _ = Univ.Instance.check_eq univs in
+ let eq_sorts s1 s2 = Univ.check_eq univs (Sorts.univ_of_sort s1) (Sorts.univ_of_sort s2) in
+ let rec eq_constr' m n =
+ m == n || compare_head_gen eq_universes eq_sorts eq_constr' m n
+ in compare_head_gen eq_universes eq_sorts eq_constr' m n
+
+let leq_constr_univs univs m n =
+ if m == n then true
+ else
+ let eq_universes _ = Univ.Instance.check_eq univs in
+ let eq_sorts s1 s2 = Univ.check_eq univs (Sorts.univ_of_sort s1) (Sorts.univ_of_sort s2) in
+ let leq_sorts s1 s2 = Univ.check_leq univs (Sorts.univ_of_sort s1) (Sorts.univ_of_sort s2) in
+ let rec eq_constr' m n =
+ m == n || compare_head_gen eq_universes eq_sorts eq_constr' m n
+ in
+ let rec compare_leq m n =
+ compare_head_gen_leq eq_universes eq_sorts leq_sorts eq_constr' leq_constr' m n
+ and leq_constr' m n = m == n || compare_leq m n in
+ compare_leq m n
+
+let eq_constr_universes m n =
+ if m == n then true, UniverseConstraints.empty
+ else
+ let cstrs = ref UniverseConstraints.empty in
+ let eq_universes strict l l' =
+ cstrs := Univ.enforce_eq_instances_univs strict l l' !cstrs; true in
+ let eq_sorts s1 s2 =
+ cstrs := Univ.UniverseConstraints.add
+ (Sorts.univ_of_sort s1, Univ.UEq, Sorts.univ_of_sort s2) !cstrs;
+ true
+ in
+ let rec eq_constr' m n =
+ m == n || compare_head_gen eq_universes eq_sorts eq_constr' m n
+ in
+ let res = compare_head_gen eq_universes eq_sorts eq_constr' m n in
+ res, !cstrs
+
+let leq_constr_universes m n =
+ if m == n then true, UniverseConstraints.empty
+ else
+ let cstrs = ref UniverseConstraints.empty in
+ let eq_universes strict l l' =
+ cstrs := Univ.enforce_eq_instances_univs strict l l' !cstrs; true in
+ let eq_sorts s1 s2 =
+ cstrs := Univ.UniverseConstraints.add
+ (Sorts.univ_of_sort s1,Univ.UEq,Sorts.univ_of_sort s2) !cstrs; true
+ in
+ let leq_sorts s1 s2 =
+ cstrs := Univ.UniverseConstraints.add
+ (Sorts.univ_of_sort s1,Univ.ULe,Sorts.univ_of_sort s2) !cstrs; true
+ in
+ let rec eq_constr' m n =
+ m == n || compare_head_gen eq_universes eq_sorts eq_constr' m n
+ in
+ let rec compare_leq m n =
+ compare_head_gen_leq eq_universes eq_sorts leq_sorts eq_constr' leq_constr' m n
+ and leq_constr' m n = m == n || compare_leq m n in
+ let res = compare_leq m n in
+ res, !cstrs
+
+let always_true _ _ = true
+
+let rec eq_constr_nounivs m n =
+ (m == n) || compare_head_gen (fun _ -> always_true) always_true eq_constr_nounivs m n
+
(** We only use this function over blocks! *)
let tag t = Obj.tag (Obj.repr t)
@@ -509,11 +644,12 @@ let constr_ord_int f t1 t2 =
| App (Cast(c1,_,_),l1), _ -> f (mkApp (c1,l1)) t2
| _, App (Cast(c2, _,_),l2) -> f t1 (mkApp (c2,l2))
| App (c1,l1), App (c2,l2) -> (f =? (Array.compare f)) c1 c2 l1 l2
+ | Proj (p1,c1), Proj (p2,c2) -> (con_ord =? f) p1 p2 c1 c2
| Evar (e1,l1), Evar (e2,l2) ->
(Evar.compare =? (Array.compare f)) e1 e2 l1 l2
- | Const c1, Const c2 -> con_ord c1 c2
- | Ind ind1, Ind ind2 -> ind_ord ind1 ind2
- | Construct ct1, Construct ct2 -> constructor_ord ct1 ct2
+ | Const (c1,u1), Const (c2,u2) -> con_ord c1 c2
+ | Ind (ind1, u1), Ind (ind2, u2) -> ind_ord ind1 ind2
+ | Construct (ct1,u1), Construct (ct2,u2) -> constructor_ord ct1 ct2
| Case (_,p1,c1,bl1), Case (_,p2,c2,bl2) ->
((f =? f) ==? (Array.compare f)) p1 p2 c1 c2 bl1 bl2
| Fix (ln1,(_,tl1,bl1)), Fix (ln2,(_,tl2,bl2)) ->
@@ -587,12 +723,14 @@ let hasheq t1 t2 =
| Lambda (n1,t1,c1), Lambda (n2,t2,c2) -> n1 == n2 && t1 == t2 && c1 == c2
| LetIn (n1,b1,t1,c1), LetIn (n2,b2,t2,c2) ->
n1 == n2 && b1 == b2 && t1 == t2 && c1 == c2
+ | Proj (c1,t1), Proj (c2,t2) -> c1 == c2 && t1 == t2
| App (c1,l1), App (c2,l2) -> c1 == c2 && array_eqeq l1 l2
| Evar (e1,l1), Evar (e2,l2) -> Evar.equal e1 e2 && array_eqeq l1 l2
- | Const c1, Const c2 -> c1 == c2
- | Ind (sp1,i1), Ind (sp2,i2) -> sp1 == sp2 && Int.equal i1 i2
- | Construct ((sp1,i1),j1), Construct ((sp2,i2),j2) ->
- sp1 == sp2 && Int.equal i1 i2 && Int.equal j1 j2
+ | Const (c1,u1), Const (c2,u2) -> c1 == c2 && Univ.Instance.eqeq u1 u2
+ | Ind ((sp1,i1),u1), Ind ((sp2,i2),u2) ->
+ sp1 == sp2 && Int.equal i1 i2 && Univ.Instance.eqeq u1 u2
+ | Construct (((sp1,i1),j1),u1), Construct (((sp2,i2),j2),u2) ->
+ sp1 == sp2 && Int.equal i1 i2 && Int.equal j1 j2 && Univ.Instance.eqeq u1 u2
| Case (ci1,p1,c1,bl1), Case (ci2,p2,c2,bl2) ->
ci1 == ci2 && p1 == p2 && c1 == c2 && array_eqeq bl1 bl2
| Fix ((ln1, i1),(lna1,tl1,bl1)), Fix ((ln2, i2),(lna2,tl2,bl2)) ->
@@ -631,6 +769,8 @@ let hash_cast_kind = function
| DEFAULTcast -> 2
| REVERTcast -> 3
+let hash_instance = Univ.Instance.hcons
+
(* [hashcons hash_consing_functions constr] computes an hash-consed
representation for [constr] using [hash_consing_functions] on
leaves. *)
@@ -665,12 +805,16 @@ let hashcons (sh_sort,sh_ci,sh_construct,sh_ind,sh_con,sh_na,sh_id) =
| Evar (e,l) ->
let l, hl = hash_term_array l in
(Evar (e,l), combinesmall 8 (combine (Evar.hash e) hl))
- | Const c ->
- (Const (sh_con c), combinesmall 9 (Constant.hash c))
- | Ind ind ->
- (Ind (sh_ind ind), combinesmall 10 (ind_hash ind))
- | Construct c ->
- (Construct (sh_construct c), combinesmall 11 (constructor_hash c))
+ | Proj (p,c) ->
+ let c, hc = sh_rec c in
+ let p' = sh_con p in
+ (Proj (p', c), combinesmall 17 (Hashtbl.hash p')) (** FIXME *)
+ | Const (c,u) ->
+ (Const (sh_con c, hash_instance u), combinesmall 9 (Constant.hash c))
+ | Ind (ind, u) ->
+ (Ind (sh_ind ind, hash_instance u), combinesmall 10 (ind_hash ind))
+ | Construct (c, u) ->
+ (Construct (sh_construct c, hash_instance u), combinesmall 11 (constructor_hash c))
| Case (ci,p,c,bl) ->
let p, hp = sh_rec p
and c, hc = sh_rec c in
@@ -742,13 +886,15 @@ let rec hash t =
| App (Cast(c, _, _),l) -> hash (mkApp (c,l))
| App (c,l) ->
combinesmall 7 (combine (hash_term_array l) (hash c))
+ | Proj (p,c) ->
+ combinesmall 17 (combine (Hashtbl.hash p) (hash c))
| Evar (e,l) ->
combinesmall 8 (combine (Evar.hash e) (hash_term_array l))
- | Const c ->
+ | Const (c, _) ->
combinesmall 9 (Constant.hash c)
- | Ind ind ->
+ | Ind (ind, _) ->
combinesmall 10 (ind_hash ind)
- | Construct c ->
+ | Construct (c, _) ->
combinesmall 11 (constructor_hash c)
| Case (_ , p, c, bl) ->
combinesmall 12 (combine3 (hash c) (hash p) (hash_term_array bl))
@@ -799,8 +945,32 @@ module Hcaseinfo = Hashcons.Make(CaseinfoHash)
let case_info_hash = CaseinfoHash.hash
+module Hsorts =
+ Hashcons.Make(
+ struct
+ open Sorts
+
+ type t = Sorts.t
+ type u = universe -> universe
+ let hashcons huniv = function
+ Prop c -> Prop c
+ | Type u -> Type (huniv u)
+ let equal s1 s2 =
+ s1 == s2 ||
+ match (s1,s2) with
+ (Prop c1, Prop c2) -> c1 == c2
+ | (Type u1, Type u2) -> u1 == u2
+ |_ -> false
+ let hash = Hashtbl.hash
+ end)
+
+let hcons_sorts = Hashcons.simple_hcons Hsorts.generate hcons_univ
let hcons_caseinfo = Hashcons.simple_hcons Hcaseinfo.generate hcons_ind
+let hcons_pconstruct (c,u) = (hcons_construct c, Univ.Instance.hcons u)
+let hcons_pind (i,u) = (hcons_ind i, Univ.Instance.hcons u)
+let hcons_pcon (c,u) = (hcons_con c, Univ.Instance.hcons u)
+
let hcons =
hashcons
(Sorts.hcons,
diff --git a/kernel/constr.mli b/kernel/constr.mli
index 82a2de0945..be62502570 100644
--- a/kernel/constr.mli
+++ b/kernel/constr.mli
@@ -8,6 +8,14 @@
open Names
+(** {6 Value under universe substitution } *)
+type 'a puniverses = 'a Univ.puniverses
+
+(** {6 Simply type aliases } *)
+type pconstant = constant puniverses
+type pinductive = inductive puniverses
+type pconstructor = constructor puniverses
+
(** {6 Existential variables } *)
type existential_key = Evar.t
@@ -88,20 +96,26 @@ val mkLetIn : Name.t * constr * types * constr -> constr
{% $(f~t_1~\dots~t_n)$ %}. *)
val mkApp : constr * constr array -> constr
-(** Constructs a constant
- The array of terms correspond to the variables introduced in the section *)
+val map_puniverses : ('a -> 'b) -> 'a puniverses -> 'b puniverses
+
+(** Constructs a constant *)
val mkConst : constant -> constr
+val mkConstU : pconstant -> constr
+
+(** Constructs a projection application *)
+val mkProj : (constant * constr) -> constr
(** Inductive types *)
-(** Constructs the ith (co)inductive type of the block named kn
- The array of terms correspond to the variables introduced in the section *)
+(** Constructs the ith (co)inductive type of the block named kn *)
val mkInd : inductive -> constr
+val mkIndU : pinductive -> constr
(** Constructs the jth constructor of the ith (co)inductive type of the
- block named kn. The array of terms correspond to the variables
- introduced in the section *)
+ block named kn. *)
val mkConstruct : constructor -> constr
+val mkConstructU : pconstructor -> constr
+val mkConstructUi : pinductive * int -> constr
(** Constructs a destructor of inductive type.
@@ -170,12 +184,13 @@ type ('constr, 'types) kind_of_term =
| Lambda of Name.t * 'types * 'constr
| LetIn of Name.t * 'constr * 'types * 'constr
| App of 'constr * 'constr array
- | Const of constant
- | Ind of inductive
- | Construct of constructor
+ | Const of constant puniverses
+ | Ind of inductive puniverses
+ | Construct of constructor puniverses
| Case of case_info * 'constr * 'constr * 'constr array
| Fix of ('constr, 'types) pfixpoint
| CoFix of ('constr, 'types) pcofixpoint
+ | Proj of constant * 'constr
(** User view of [constr]. For [App], it is ensured there is at
least one argument and the function is not itself an applicative
@@ -187,6 +202,26 @@ val kind : constr -> (constr, types) kind_of_term
and application grouping *)
val equal : constr -> constr -> bool
+(** [eq_constr_univs a b] [true, c] if [a] equals [b] modulo alpha, casts,
+ application grouping and the universe equalities in [c]. *)
+val eq_constr_univs : constr Univ.check_function
+
+(** [leq_constr_univs a b] [true, c] if [a] is convertible to [b] modulo
+ alpha, casts, application grouping and the universe inequalities in [c]. *)
+val leq_constr_univs : constr Univ.check_function
+
+(** [eq_constr_universes a b] [true, c] if [a] equals [b] modulo alpha, casts,
+ application grouping and the universe equalities in [c]. *)
+val eq_constr_universes : constr -> constr -> bool Univ.universe_constrained
+
+(** [leq_constr_universes a b] [true, c] if [a] is convertible to [b] modulo
+ alpha, casts, application grouping and the universe inequalities in [c]. *)
+val leq_constr_universes : constr -> constr -> bool Univ.universe_constrained
+
+(** [eq_constr_univs a b] [true, c] if [a] equals [b] modulo alpha, casts,
+ application grouping and ignoring universe instances. *)
+val eq_constr_nounivs : constr -> constr -> bool
+
(** Total ordering compatible with [equal] *)
val compare : constr -> constr -> int
diff --git a/kernel/conv_oracle.mli b/kernel/conv_oracle.mli
index cd8cd2cf7a..ae501ce872 100644
--- a/kernel/conv_oracle.mli
+++ b/kernel/conv_oracle.mli
@@ -16,7 +16,7 @@ val empty : oracle
If [oracle_order kn1 kn2] is true, then unfold kn1 first.
Note: the oracle does not introduce incompleteness, it only
tries to postpone unfolding of "opaque" constants. *)
-val oracle_order : oracle -> bool -> 'a tableKey -> 'a tableKey -> bool
+val oracle_order : oracle -> bool -> constant tableKey -> constant tableKey -> bool
(** Priority for the expansion of constant in the conversion test.
* Higher levels means that the expansion is less prioritary.
@@ -29,14 +29,14 @@ val transparent : level
(** Check whether a level is transparent *)
val is_transparent : level -> bool
-val get_strategy : oracle -> 'a tableKey -> level
+val get_strategy : oracle -> constant tableKey -> level
(** Sets the level of a constant.
* Level of RelKey constant cannot be set. *)
-val set_strategy : oracle -> 'a tableKey -> level -> oracle
+val set_strategy : oracle -> constant tableKey -> level -> oracle
(** Fold over the non-transparent levels of the oracle. Order unspecified. *)
-val fold_strategy : (unit tableKey -> level -> 'a -> 'a) -> oracle -> 'a -> 'a
+val fold_strategy : (constant tableKey -> level -> 'a -> 'a) -> oracle -> 'a -> 'a
val get_transp_state : oracle -> transparent_state
diff --git a/kernel/cooking.ml b/kernel/cooking.ml
index dbe188bd44..4bae6de66a 100644
--- a/kernel/cooking.ml
+++ b/kernel/cooking.ml
@@ -19,6 +19,7 @@ open Names
open Term
open Declarations
open Environ
+open Univ
(*s Cooking the constants. *)
@@ -56,27 +57,36 @@ end
module RefTable = Hashtbl.Make(RefHash)
+let instantiate_my_gr gr u =
+ match gr with
+ | ConstRef c -> mkConstU (c, u)
+ | IndRef i -> mkIndU (i, u)
+ | ConstructRef c -> mkConstructU (c, u)
+
let share cache r (cstl,knl) =
try RefTable.find cache r
with Not_found ->
- let f,l =
+ let f,(u,l) =
match r with
| IndRef (kn,i) ->
- mkInd (pop_mind kn,i), Mindmap.find kn knl
+ IndRef (pop_mind kn,i), Mindmap.find kn knl
| ConstructRef ((kn,i),j) ->
- mkConstruct ((pop_mind kn,i),j), Mindmap.find kn knl
+ ConstructRef ((pop_mind kn,i),j), Mindmap.find kn knl
| ConstRef cst ->
- mkConst (pop_con cst), Cmap.find cst cstl in
- let c = mkApp (f, Array.map mkVar l) in
+ ConstRef (pop_con cst), Cmap.find cst cstl in
+ let c = (f, (u, Array.map mkVar l)) in
RefTable.add cache r c;
c
+let share_univs cache r u l =
+ let r', (u', args) = share cache r l in
+ mkApp (instantiate_my_gr r' (Instance.append u' u), args)
+
let update_case_info cache ci modlist =
try
let ind, n =
- match kind_of_term (share cache (IndRef ci.ci_ind) modlist) with
- | App (f,l) -> (destInd f, Array.length l)
- | Ind ind -> ind, 0
+ match share cache (IndRef ci.ci_ind) modlist with
+ | (IndRef f,(u,l)) -> (f, Array.length l)
| _ -> assert false in
{ ci with ci_ind = ind; ci_npar = ci.ci_npar + n }
with Not_found ->
@@ -86,31 +96,43 @@ let is_empty_modlist (cm, mm) =
Cmap.is_empty cm && Mindmap.is_empty mm
let expmod_constr cache modlist c =
- let share = share cache in
+ let share_univs = share_univs cache in
let update_case_info = update_case_info cache in
let rec substrec c =
match kind_of_term c with
| Case (ci,p,t,br) ->
map_constr substrec (mkCase (update_case_info ci modlist,p,t,br))
- | Ind ind ->
+ | Ind (ind,u) ->
(try
- share (IndRef ind) modlist
+ share_univs (IndRef ind) u modlist
with
| Not_found -> map_constr substrec c)
- | Construct cstr ->
+ | Construct (cstr,u) ->
(try
- share (ConstructRef cstr) modlist
+ share_univs (ConstructRef cstr) u modlist
with
| Not_found -> map_constr substrec c)
- | Const cst ->
+ | Const (cst,u) ->
(try
- share (ConstRef cst) modlist
+ share_univs (ConstRef cst) u modlist
with
| Not_found -> map_constr substrec c)
+ | Proj (p, c') ->
+ (try
+ let p' = share_univs (ConstRef p) Univ.Instance.empty modlist in
+ match kind_of_term p' with
+ | Const (p',_) -> mkProj (p', substrec c')
+ | App (f, args) ->
+ (match kind_of_term f with
+ | Const (p', _) -> mkProj (p', substrec c')
+ | _ -> assert false)
+ | _ -> assert false
+ with Not_found -> map_constr substrec c)
+
| _ -> map_constr substrec c
in
@@ -127,7 +149,8 @@ type recipe = { from : constant_body; info : Opaqueproof.cooking_info }
type inline = bool
type result =
- constant_def * constant_type * Univ.constraints * inline
+ constant_def * constant_type * projection_body option *
+ bool * constant_universes * inline
* Context.section_context option
let on_body ml hy f = function
@@ -142,15 +165,17 @@ let constr_of_def = function
| Def cs -> Mod_subst.force_constr cs
| OpaqueDef lc -> Opaqueproof.force_proof lc
+
let cook_constr { Opaqueproof.modlist ; abstract } c =
let cache = RefTable.create 13 in
- let hyps = Context.map_named_context (expmod_constr cache modlist) abstract in
+ let hyps = Context.map_named_context (expmod_constr cache modlist) (fst abstract) in
abstract_constant_body (expmod_constr cache modlist c) hyps
let cook_constant env { from = cb; info = { Opaqueproof.modlist; abstract } } =
let cache = RefTable.create 13 in
+ let abstract, abs_ctx = abstract in
let hyps = Context.map_named_context (expmod_constr cache modlist) abstract in
- let body = on_body modlist hyps
+ let body = on_body modlist (hyps, abs_ctx)
(fun c -> abstract_constant_body (expmod_constr cache modlist c) hyps)
cb.const_body
in
@@ -158,18 +183,43 @@ let cook_constant env { from = cb; info = { Opaqueproof.modlist; abstract } } =
Context.fold_named_context (fun (h,_,_) hyps ->
List.filter (fun (id,_,_) -> not (Id.equal id h)) hyps)
hyps ~init:cb.const_hyps in
- let typ = match cb.const_type with
- | NonPolymorphicType t ->
- let typ =
- abstract_constant_type (expmod_constr cache modlist t) hyps in
- NonPolymorphicType typ
- | PolymorphicArity (ctx,s) ->
- let t = mkArity (ctx,Type s.poly_level) in
- let typ =
- abstract_constant_type (expmod_constr cache modlist t) hyps in
- let j = make_judge (constr_of_def body) typ in
- Typeops.make_polymorphic_if_constant_for_ind env j
+
+ (* let typ = match cb.const_type with *)
+ (* | NonPolymorphicType t -> *)
+ (* let typ = *)
+ (* abstract_constant_type (expmod_constr cache r.d_modlist t) hyps in *)
+ (* NonPolymorphicType typ *)
+ (* | PolymorphicArity (ctx,s) -> *)
+ (* let t = mkArity (ctx,Type s.poly_level) in *)
+ (* let typ = *)
+ (* abstract_constant_type (expmod_constr cache r.d_modlist t) hyps in *)
+ (* let j = make_judge (constr_of_def body) typ in *)
+ (* Typeops.make_polymorphic_if_constant_for_ind env j *)
+ (* in *)
+ let typ =
+ abstract_constant_type (expmod_constr cache modlist cb.const_type) hyps
in
- (body, typ, cb.const_constraints, cb.const_inline_code, Some const_hyps)
+ let projection pb =
+ let c' = abstract_constant_body (expmod_constr cache modlist pb.proj_body) hyps in
+ let ((mind, _), _), n' =
+ try
+ let c' = share_univs cache (IndRef (pb.proj_ind,0)) Univ.Instance.empty modlist in
+ match kind_of_term c' with
+ | App (f,l) -> (destInd f, Array.length l)
+ | Ind ind -> ind, 0
+ | _ -> assert false
+ with Not_found -> (((pb.proj_ind,0),Univ.Instance.empty), 0)
+ in
+ let ctx, ty' = decompose_prod_n (n' + pb.proj_npars + 1) typ in
+ { proj_ind = mind; proj_npars = pb.proj_npars + n'; proj_arg = pb.proj_arg;
+ proj_type = ty'; proj_body = c' }
+ in
+ let univs = Future.from_val (UContext.union abs_ctx (Future.force cb.const_universes)) in
+ (body, typ, Option.map projection cb.const_proj,
+ cb.const_polymorphic, univs, cb.const_inline_code,
+ Some const_hyps)
+
+(* let cook_constant_key = Profile.declare_profile "cook_constant" *)
+(* let cook_constant = Profile.profile2 cook_constant_key cook_constant *)
let expmod_constr modlist c = expmod_constr (RefTable.create 13) modlist c
diff --git a/kernel/cooking.mli b/kernel/cooking.mli
index 030e888294..489360af7f 100644
--- a/kernel/cooking.mli
+++ b/kernel/cooking.mli
@@ -17,7 +17,8 @@ type recipe = { from : constant_body; info : Opaqueproof.cooking_info }
type inline = bool
type result =
- constant_def * constant_type * Univ.constraints * inline
+ constant_def * constant_type * projection_body option *
+ bool * constant_universes * inline
* Context.section_context option
val cook_constant : env -> recipe -> result
diff --git a/kernel/declarations.mli b/kernel/declarations.mli
index 1e94e243c0..f3cb41f329 100644
--- a/kernel/declarations.mli
+++ b/kernel/declarations.mli
@@ -18,14 +18,7 @@ type engagement = ImpredicativeSet
(** {6 Representation of constants (Definition/Axiom) } *)
-type polymorphic_arity = {
- poly_param_levels : Univ.universe option list;
- poly_level : Univ.universe;
-}
-
-type constant_type =
- | NonPolymorphicType of types
- | PolymorphicArity of rel_context * polymorphic_arity
+type constant_type = types
(** Inlining level of parameters at functor applications.
None means no inlining *)
@@ -35,11 +28,24 @@ type inline = int option
(** A constant can have no body (axiom/parameter), or a
transparent body, or an opaque one *)
+(** Projections are a particular kind of constant:
+ always transparent. *)
+
+type projection_body = {
+ proj_ind : mutual_inductive;
+ proj_npars : int;
+ proj_arg : int;
+ proj_type : types; (* Type under params *)
+ proj_body : constr; (* For compatibility, the match version *)
+}
+
type constant_def =
| Undef of inline
| Def of constr Mod_subst.substituted
| OpaqueDef of Opaqueproof.opaque
+type constant_universes = Univ.universe_context Future.computation
+
(* some contraints are in constant_constraints, some other may be in
* the OpaueDef *)
type constant_body = {
@@ -47,7 +53,9 @@ type constant_body = {
const_body : constant_def;
const_type : constant_type;
const_body_code : Cemitcodes.to_patch_substituted;
- const_constraints : Univ.constraints;
+ const_polymorphic : bool; (** Is it polymorphic or not *)
+ const_universes : constant_universes;
+ const_proj : projection_body option;
const_inline_code : bool }
type side_effect =
@@ -71,15 +79,11 @@ type wf_paths = recarg Rtree.t
v}
*)
-type monomorphic_inductive_arity = {
- mind_user_arity : constr;
+type inductive_arity = {
+ mind_user_arity : types;
mind_sort : sorts;
}
-type inductive_arity =
-| Monomorphic of monomorphic_inductive_arity
-| Polymorphic of polymorphic_arity
-
type one_inductive_body = {
(** {8 Primitive datas } *)
@@ -87,7 +91,7 @@ type one_inductive_body = {
mind_arity_ctxt : rel_context; (** Arity context of [Ii] with parameters: [forall params, Ui] *)
- mind_arity : inductive_arity; (** Arity sort and original user arity if monomorphic *)
+ mind_arity : inductive_arity; (** Arity sort and original user arity *)
mind_consnames : Id.t array; (** Names of the constructors: [cij] *)
@@ -129,7 +133,9 @@ type mutual_inductive_body = {
mind_packets : one_inductive_body array; (** The component of the mutual inductive block *)
- mind_record : bool; (** Whether the inductive type has been declared as a record *)
+ mind_record : (constr * constant array) option;
+ (** Whether the inductive type has been declared as a record,
+ In that case we get its canonical eta-expansion and list of projections. *)
mind_finite : bool; (** Whether the type is inductive or coinductive *)
@@ -143,7 +149,9 @@ type mutual_inductive_body = {
mind_params_ctxt : rel_context; (** The context of parameters (includes let-in declaration) *)
- mind_constraints : Univ.constraints; (** Universes constraints enforced by the inductive declaration *)
+ mind_polymorphic : bool; (** Is it polymorphic or not *)
+
+ mind_universes : Univ.universe_context; (** Local universe variables and constraints *)
}
diff --git a/kernel/declareops.ml b/kernel/declareops.ml
index 1b67de0eaa..0e4b80495e 100644
--- a/kernel/declareops.ml
+++ b/kernel/declareops.ml
@@ -20,8 +20,9 @@ let body_of_constant cb = match cb.const_body with
| Def c -> Some (force_constr c)
| OpaqueDef o -> Some (Opaqueproof.force_proof o)
-let constraints_of_constant cb = Univ.union_constraints cb.const_constraints
- (match cb.const_body with
+let constraints_of_constant cb = Univ.Constraint.union
+ (Univ.UContext.constraints (Future.force cb.const_universes))
+ (match cb.const_body with
| Undef _ -> Univ.empty_constraint
| Def c -> Univ.empty_constraint
| OpaqueDef o -> Opaqueproof.force_constraints o)
@@ -43,36 +44,48 @@ let subst_rel_declaration sub (id,copt,t as x) =
let subst_rel_context sub = List.smartmap (subst_rel_declaration sub)
-let subst_const_type sub arity = match arity with
- | NonPolymorphicType s ->
- let s' = subst_mps sub s in
- if s==s' then arity else NonPolymorphicType s'
- | PolymorphicArity (ctx,s) ->
- let ctx' = subst_rel_context sub ctx in
- if ctx==ctx' then arity else PolymorphicArity (ctx',s)
+(* let subst_const_type sub arity = match arity with *)
+(* | NonPolymorphicType s -> *)
+(* let s' = subst_mps sub s in *)
+(* if s==s' then arity else NonPolymorphicType s' *)
+(* | PolymorphicArity (ctx,s) -> *)
+(* let ctx' = subst_rel_context sub ctx in *)
+(* if ctx==ctx' then arity else PolymorphicArity (ctx',s) *)
+
+let subst_const_type sub arity =
+ if is_empty_subst sub then arity
+ else subst_mps sub arity
(** No need here to check for physical equality after substitution,
at least for Def due to the delayed substitution [subst_constr_subst]. *)
-
let subst_const_def sub def = match def with
| Undef _ -> def
| Def c -> Def (subst_constr sub c)
| OpaqueDef o -> OpaqueDef (Opaqueproof.subst_opaque sub o)
+let subst_const_proj sub pb =
+ { pb with proj_ind = subst_mind sub pb.proj_ind;
+ proj_type = subst_mps sub pb.proj_type;
+ proj_body = subst_const_type sub pb.proj_body }
+
let subst_const_body sub cb =
assert (List.is_empty cb.const_hyps); (* we're outside sections *)
if is_empty_subst sub then cb
else
let body' = subst_const_def sub cb.const_body in
let type' = subst_const_type sub cb.const_type in
- if body' == cb.const_body && type' == cb.const_type then cb
+ let proj' = Option.smartmap (subst_const_proj sub) cb.const_proj in
+ if body' == cb.const_body && type' == cb.const_type
+ && proj' == cb.const_proj then cb
else
{ const_hyps = [];
const_body = body';
const_type = type';
+ const_proj = proj';
const_body_code =
Cemitcodes.subst_to_patch_subst sub cb.const_body_code;
- const_constraints = cb.const_constraints;
+ const_polymorphic = cb.const_polymorphic;
+ const_universes = cb.const_universes;
const_inline_code = cb.const_inline_code }
(** {7 Hash-consing of constants } *)
@@ -89,16 +102,7 @@ let hcons_rel_decl ((n,oc,t) as d) =
let hcons_rel_context l = List.smartmap hcons_rel_decl l
-let hcons_polyarity ar =
- { poly_param_levels =
- List.smartmap (Option.smartmap Univ.hcons_univ) ar.poly_param_levels;
- poly_level = Univ.hcons_univ ar.poly_level }
-
-let hcons_const_type = function
- | NonPolymorphicType t ->
- NonPolymorphicType (Term.hcons_constr t)
- | PolymorphicArity (ctx,s) ->
- PolymorphicArity (hcons_rel_context ctx, hcons_polyarity s)
+let hcons_const_type t = Term.hcons_constr t
let hcons_const_def = function
| Undef inl -> Undef inl
@@ -111,7 +115,11 @@ let hcons_const_body cb =
{ cb with
const_body = hcons_const_def cb.const_body;
const_type = hcons_const_type cb.const_type;
- const_constraints = Univ.hcons_constraints cb.const_constraints; }
+ const_universes =
+ if Future.is_val cb.const_universes then
+ Future.from_val
+ (Univ.hcons_universe_context (Future.force cb.const_universes))
+ else (* FIXME: Why not? *) cb.const_universes }
(** {6 Inductive types } *)
@@ -124,10 +132,10 @@ let eq_recarg r1 r2 = match r1, r2 with
let subst_recarg sub r = match r with
| Norec -> r
| Mrec (kn,i) ->
- let kn' = subst_ind sub kn in
+ let kn' = subst_mind sub kn in
if kn==kn' then r else Mrec (kn',i)
| Imbr (kn,i) ->
- let kn' = subst_ind sub kn in
+ let kn' = subst_mind sub kn in
if kn==kn' then r else Imbr (kn',i)
let mk_norec = Rtree.mk_node Norec [||]
@@ -156,63 +164,108 @@ let subst_wf_paths sub p = Rtree.smartmap (subst_recarg sub) p
(** {7 Substitution of inductive declarations } *)
-let subst_indarity sub ar = match ar with
- | Monomorphic s ->
- let uar' = subst_mps sub s.mind_user_arity in
- if uar' == s.mind_user_arity then ar
- else Monomorphic { mind_user_arity = uar'; mind_sort = s.mind_sort }
- | Polymorphic _ -> ar
-
-let subst_mind_packet sub mip =
- let { mind_nf_lc = nf;
- mind_user_lc = user;
- mind_arity_ctxt = ctxt;
- mind_arity = ar;
- mind_recargs = ra } = mip
- in
- let nf' = Array.smartmap (subst_mps sub) nf in
- let user' =
- (* maintain sharing of [mind_user_lc] and [mind_nf_lc] *)
- if user==nf then nf'
- else Array.smartmap (subst_mps sub) user
- in
- let ctxt' = subst_rel_context sub ctxt in
- let ar' = subst_indarity sub ar in
- let ra' = subst_wf_paths sub ra in
- if nf==nf' && user==user' && ctxt==ctxt' && ar==ar' && ra==ra'
- then mip
- else
- { mip with
- mind_nf_lc = nf';
- mind_user_lc = user';
- mind_arity_ctxt = ctxt';
- mind_arity = ar';
- mind_recargs = ra' }
-
-let subst_mind sub mib =
- assert (List.is_empty mib.mind_hyps); (* we're outside sections *)
- if is_empty_subst sub then mib
- else
- let params = mib.mind_params_ctxt in
- let params' = Context.map_rel_context (subst_mps sub) params in
- let packets = mib.mind_packets in
- let packets' = Array.smartmap (subst_mind_packet sub) packets in
- if params==params' && packets==packets' then mib
- else
- { mib with
- mind_params_ctxt = params';
- mind_packets = packets' }
-
-(** {6 Hash-consing of inductive declarations } *)
-
-(** Just as for constants, this hash-consing is quite partial *)
-
-let hcons_indarity = function
- | Monomorphic a ->
- Monomorphic
- { mind_user_arity = Term.hcons_constr a.mind_user_arity;
- mind_sort = Term.hcons_sorts a.mind_sort }
- | Polymorphic a -> Polymorphic (hcons_polyarity a)
+(* OLD POLYMORPHISM *)
+(* let subst_indarity sub ar = match ar with *)
+(* | Monomorphic s -> *)
+(* let uar' = subst_mps sub s.mind_user_arity in *)
+(* if uar' == s.mind_user_arity then ar *)
+(* else Monomorphic { mind_user_arity = uar'; mind_sort = s.mind_sort } *)
+(* | Polymorphic _ -> ar *)
+
+(* let subst_mind_packet sub mip = *)
+(* let { mind_nf_lc = nf; *)
+(* mind_user_lc = user; *)
+(* mind_arity_ctxt = ctxt; *)
+(* mind_arity = ar; *)
+(* mind_recargs = ra } = mip *)
+(* in *)
+(* let nf' = Array.smartmap (subst_mps sub) nf in *)
+(* let user' = *)
+(* (\* maintain sharing of [mind_user_lc] and [mind_nf_lc] *\) *)
+(* if user==nf then nf' *)
+(* else Array.smartmap (subst_mps sub) user *)
+(* in *)
+(* let ctxt' = subst_rel_context sub ctxt in *)
+(* let ar' = subst_indarity sub ar in *)
+(* let ra' = subst_wf_paths sub ra in *)
+(* if nf==nf' && user==user' && ctxt==ctxt' && ar==ar' && ra==ra' *)
+(* then mip *)
+(* else *)
+(* { mip with *)
+(* mind_nf_lc = nf'; *)
+(* mind_user_lc = user'; *)
+(* mind_arity_ctxt = ctxt'; *)
+(* mind_arity = ar'; *)
+(* mind_recargs = ra' } *)
+
+(* let subst_mind sub mib = *)
+(* assert (List.is_empty mib.mind_hyps); (\* we're outside sections *\) *)
+(* if is_empty_subst sub then mib *)
+(* else *)
+(* let params = mib.mind_params_ctxt in *)
+(* let params' = Context.map_rel_context (subst_mps sub) params in *)
+(* let packets = mib.mind_packets in *)
+(* let packets' = Array.smartmap (subst_mind_packet sub) packets in *)
+(* if params==params' && packets==packets' then mib *)
+(* else *)
+(* { mib with *)
+(* mind_params_ctxt = params'; *)
+(* mind_packets = packets'; *)
+(* mind_native_name = ref NotLinked } *)
+
+(* (\** {6 Hash-consing of inductive declarations } *\) *)
+
+(* (\** Just as for constants, this hash-consing is quite partial *\) *)
+
+(* let hcons_indarity = function *)
+(* | Monomorphic a -> *)
+(* Monomorphic *)
+(* { mind_user_arity = Term.hcons_constr a.mind_user_arity; *)
+(* mind_sort = Term.hcons_sorts a.mind_sort } *)
+(* | Polymorphic a -> Polymorphic (hcons_polyarity a) *)
+
+(** Substitution of inductive declarations *)
+
+let subst_indarity sub s =
+ { mind_user_arity = subst_mps sub s.mind_user_arity;
+ mind_sort = s.mind_sort;
+ }
+
+let subst_mind_packet sub mbp =
+ { mind_consnames = mbp.mind_consnames;
+ mind_consnrealdecls = mbp.mind_consnrealdecls;
+ mind_consnrealargs = mbp.mind_consnrealargs;
+ mind_typename = mbp.mind_typename;
+ mind_nf_lc = Array.smartmap (subst_mps sub) mbp.mind_nf_lc;
+ mind_arity_ctxt = subst_rel_context sub mbp.mind_arity_ctxt;
+ mind_arity = subst_indarity sub mbp.mind_arity;
+ mind_user_lc = Array.smartmap (subst_mps sub) mbp.mind_user_lc;
+ mind_nrealargs = mbp.mind_nrealargs;
+ mind_nrealargs_ctxt = mbp.mind_nrealargs_ctxt;
+ mind_kelim = mbp.mind_kelim;
+ mind_recargs = subst_wf_paths sub mbp.mind_recargs (*wf_paths*);
+ mind_nb_constant = mbp.mind_nb_constant;
+ mind_nb_args = mbp.mind_nb_args;
+ mind_reloc_tbl = mbp.mind_reloc_tbl }
+
+let subst_mind_body sub mib =
+ { mind_record = mib.mind_record ;
+ mind_finite = mib.mind_finite ;
+ mind_ntypes = mib.mind_ntypes ;
+ mind_hyps = (match mib.mind_hyps with [] -> [] | _ -> assert false);
+ mind_nparams = mib.mind_nparams;
+ mind_nparams_rec = mib.mind_nparams_rec;
+ mind_params_ctxt =
+ Context.map_rel_context (subst_mps sub) mib.mind_params_ctxt;
+ mind_packets = Array.smartmap (subst_mind_packet sub) mib.mind_packets ;
+ mind_polymorphic = mib.mind_polymorphic;
+ mind_universes = mib.mind_universes }
+
+(** Hash-consing of inductive declarations *)
+
+let hcons_indarity a =
+ { mind_user_arity = Term.hcons_constr a.mind_user_arity;
+ mind_sort = Term.hcons_sorts a.mind_sort }
let hcons_mind_packet oib =
let user = Array.smartmap Term.hcons_types oib.mind_user_lc in
@@ -231,11 +284,12 @@ let hcons_mind mib =
{ mib with
mind_packets = Array.smartmap hcons_mind_packet mib.mind_packets;
mind_params_ctxt = hcons_rel_context mib.mind_params_ctxt;
- mind_constraints = Univ.hcons_constraints mib.mind_constraints }
+ mind_universes = Univ.hcons_universe_context mib.mind_universes }
(** {6 Stm machinery } *)
let join_constant_body cb =
+ ignore(Future.join cb.const_universes);
match cb.const_body with
| OpaqueDef o -> Opaqueproof.join_opaque o
| _ -> ()
diff --git a/kernel/declareops.mli b/kernel/declareops.mli
index 800b167abd..0c5f3584e8 100644
--- a/kernel/declareops.mli
+++ b/kernel/declareops.mli
@@ -57,7 +57,7 @@ val recarg_length : wf_paths -> int -> int
val subst_wf_paths : substitution -> wf_paths -> wf_paths
-val subst_mind : substitution -> mutual_inductive_body -> mutual_inductive_body
+val subst_mind_body : substitution -> mutual_inductive_body -> mutual_inductive_body
val join_constant_body : constant_body -> unit
diff --git a/kernel/entries.mli b/kernel/entries.mli
index 73efc73727..24e029bc09 100644
--- a/kernel/entries.mli
+++ b/kernel/entries.mli
@@ -44,12 +44,16 @@ type mutual_inductive_entry = {
mind_entry_record : bool;
mind_entry_finite : bool;
mind_entry_params : (Id.t * local_entry) list;
- mind_entry_inds : one_inductive_entry list }
+ mind_entry_inds : one_inductive_entry list;
+ mind_entry_polymorphic : bool;
+ mind_entry_universes : Univ.universe_context }
(** {6 Constants (Definition/Axiom) } *)
type proof_output = constr * Declareops.side_effects
type const_entry_body = proof_output Future.computation
+type projection = mutual_inductive * int * int * types
+
type definition_entry = {
const_entry_body : const_entry_body;
(* List of sectoin variables *)
@@ -57,12 +61,16 @@ type definition_entry = {
(* State id on which the completion of type checking is reported *)
const_entry_feedback : Stateid.t option;
const_entry_type : types option;
+ const_entry_polymorphic : bool;
+ const_entry_universes : Univ.universe_context;
+ const_entry_proj : projection option;
const_entry_opaque : bool;
const_entry_inline_code : bool }
type inline = int option (* inlining level, None for no inlining *)
-type parameter_entry = Context.section_context option * types * inline
+type parameter_entry =
+ Context.section_context option * bool * types Univ.in_universe_context * inline
type constant_entry =
| DefinitionEntry of definition_entry
diff --git a/kernel/environ.ml b/kernel/environ.ml
index d306599ad8..323d6fcea6 100644
--- a/kernel/environ.ml
+++ b/kernel/environ.ml
@@ -26,7 +26,6 @@ open Names
open Term
open Context
open Vars
-open Univ
open Declarations
open Pre_env
@@ -46,6 +45,12 @@ let empty_named_context_val = empty_named_context_val
let empty_env = empty_env
let engagement env = env.env_stratification.env_engagement
+
+let is_impredicative_set env =
+ match engagement env with
+ | Some ImpredicativeSet -> true
+ | _ -> false
+
let universes env = env.env_stratification.env_universes
let named_context env = env.env_named_context
let named_context_val env = env.env_named_context,env.env_named_vals
@@ -160,6 +165,30 @@ let fold_named_context f env ~init =
let fold_named_context_reverse f ~init env =
Context.fold_named_context_reverse f ~init:init (named_context env)
+
+(* Universe constraints *)
+
+let add_constraints c env =
+ if Univ.Constraint.is_empty c then
+ env
+ else
+ let s = env.env_stratification in
+ { env with env_stratification =
+ { s with env_universes = Univ.merge_constraints c s.env_universes } }
+
+let check_constraints c env =
+ Univ.check_constraints c env.env_stratification.env_universes
+
+let set_engagement c env = (* Unsafe *)
+ { env with env_stratification =
+ { env.env_stratification with env_engagement = Some c } }
+
+let push_constraints_to_env (_,univs) env =
+ add_constraints univs env
+
+let push_context ctx env = add_constraints (Univ.UContext.constraints ctx) env
+let push_context_set ctx env = add_constraints (Univ.ContextSet.constraints ctx) env
+
(* Global constants *)
let lookup_constant = lookup_constant
@@ -177,30 +206,113 @@ let add_constant_key kn cb linkinfo env =
let add_constant kn cb env =
add_constant_key kn cb (no_link_info ()) env
+let universes_of cb =
+ Future.force cb.const_universes
+
+let universes_and_subst_of cb u =
+ let univs = universes_of cb in
+ let subst = Univ.make_universe_subst u univs in
+ subst, Univ.instantiate_univ_context subst univs
+
(* constant_type gives the type of a constant *)
-let constant_type env kn =
+let constant_type env (kn,u) =
let cb = lookup_constant kn env in
- cb.const_type
+ if cb.const_polymorphic then
+ let subst, csts = universes_and_subst_of cb u in
+ (subst_univs_constr subst cb.const_type, csts)
+ else cb.const_type, Univ.Constraint.empty
-type const_evaluation_result = NoBody | Opaque
+let constant_type_in_ctx env kn =
+ let cb = lookup_constant kn env in
+ cb.const_type, Future.force cb.const_universes
+
+let constant_context env kn =
+ let cb = lookup_constant kn env in
+ if cb.const_polymorphic then Future.force cb.const_universes
+ else Univ.UContext.empty
+
+type const_evaluation_result = NoBody | Opaque | IsProj
exception NotEvaluableConst of const_evaluation_result
-let constant_value env kn =
+let constant_value env (kn,u) =
+ let cb = lookup_constant kn env in
+ if cb.const_proj = None then
+ match cb.const_body with
+ | Def l_body ->
+ if cb.const_polymorphic then
+ let subst, csts = universes_and_subst_of cb u in
+ (subst_univs_constr subst (Mod_subst.force_constr l_body), csts)
+ else Mod_subst.force_constr l_body, Univ.Constraint.empty
+ | OpaqueDef _ -> raise (NotEvaluableConst Opaque)
+ | Undef _ -> raise (NotEvaluableConst NoBody)
+ else raise (NotEvaluableConst IsProj)
+
+let constant_opt_value env cst =
+ try Some (constant_value env cst)
+ with NotEvaluableConst _ -> None
+
+let constant_value_and_type env (kn, u) =
+ let cb = lookup_constant kn env in
+ if cb.const_polymorphic then
+ let subst, cst = universes_and_subst_of cb u in
+ let b' = match cb.const_body with
+ | Def l_body -> Some (subst_univs_constr subst (Mod_subst.force_constr l_body))
+ | OpaqueDef _ -> None
+ | Undef _ -> None
+ in b', subst_univs_constr subst cb.const_type, cst
+ else
+ let b' = match cb.const_body with
+ | Def l_body -> Some (Mod_subst.force_constr l_body)
+ | OpaqueDef _ -> None
+ | Undef _ -> None
+ in b', cb.const_type, Univ.Constraint.empty
+
+(* These functions should be called under the invariant that [env]
+ already contains the constraints corresponding to the constant
+ application. *)
+
+(* constant_type gives the type of a constant *)
+let constant_type_in env (kn,u) =
+ let cb = lookup_constant kn env in
+ if cb.const_polymorphic then
+ let subst = Univ.make_universe_subst u (Future.force cb.const_universes) in
+ subst_univs_constr subst cb.const_type
+ else cb.const_type
+
+let constant_value_in env (kn,u) =
let cb = lookup_constant kn env in
match cb.const_body with
- | Def l_body -> Mod_subst.force_constr l_body
+ | Def l_body ->
+ let b = Mod_subst.force_constr l_body in
+ if cb.const_polymorphic then
+ let subst = Univ.make_universe_subst u (Future.force cb.const_universes) in
+ subst_univs_constr subst b
+ else b
| OpaqueDef _ -> raise (NotEvaluableConst Opaque)
| Undef _ -> raise (NotEvaluableConst NoBody)
-let constant_opt_value env cst =
- try Some (constant_value env cst)
+let constant_opt_value_in env cst =
+ try Some (constant_value_in env cst)
with NotEvaluableConst _ -> None
(* A global const is evaluable if it is defined and not opaque *)
-let evaluable_constant cst env =
- try let _ = constant_value env cst in true
- with NotEvaluableConst _ -> false
+let evaluable_constant kn env =
+ let cb = lookup_constant kn env in
+ match cb.const_body with
+ | Def _ -> true
+ | OpaqueDef _ -> false
+ | Undef _ -> false
+
+let lookup_projection cst env =
+ match (lookup_constant cst env).const_proj with
+ | Some pb -> pb
+ | None -> anomaly (Pp.str "lookup_projection: constant is not a projection")
+
+let is_projection cst env =
+ match (lookup_constant cst env).const_proj with
+ | Some _ -> true
+ | None -> false
(* Mutual Inductives *)
let lookup_mind = lookup_mind
@@ -215,21 +327,10 @@ let add_mind_key kn mind_key env =
let add_mind kn mib env =
let li = no_link_info () in add_mind_key kn (mib, li) env
-(* Universe constraints *)
-
-let add_constraints c env =
- if is_empty_constraint c then
- env
- else
- let s = env.env_stratification in
- { env with env_stratification =
- { s with env_universes = merge_constraints c s.env_universes } }
+(* Lookup of section variables *)
-let set_engagement c env = (* Unsafe *)
- { env with env_stratification =
- { env.env_stratification with env_engagement = Some c } }
+let constant_body_hyps cb = cb.const_hyps
-(* Lookup of section variables *)
let lookup_constant_variables c env =
let cmap = lookup_constant c env in
Context.vars_of_named_context cmap.const_hyps
@@ -246,9 +347,10 @@ let lookup_constructor_variables (ind,_) env =
let vars_of_global env constr =
match kind_of_term constr with
Var id -> Id.Set.singleton id
- | Const kn -> lookup_constant_variables kn env
- | Ind ind -> lookup_inductive_variables ind env
- | Construct cstr -> lookup_constructor_variables cstr env
+ | Const (kn, _) -> lookup_constant_variables kn env
+ | Ind (ind, _) -> lookup_inductive_variables ind env
+ | Construct (cstr, _) -> lookup_constructor_variables cstr env
+ (** FIXME: is Proj missing? *)
| _ -> raise Not_found
let global_vars_set env constr =
@@ -423,7 +525,7 @@ let unregister env field =
is abstract, and that the only function which add elements to the
retroknowledge is Environ.register which enforces this shape *)
(match kind_of_term (retroknowledge find env field) with
- | Ind i31t -> let i31c = mkConstruct (i31t, 1) in
+ | Ind i31t -> let i31c = mkConstructUi (i31t, 1) in
{env with retroknowledge =
remove (retroknowledge clear_info env i31c) field}
| _ -> assert false)
@@ -487,7 +589,7 @@ let register =
let add_int31_before_match rk grp v =
let rk = add_vm_before_match_info rk v Cbytegen.int31_escape_before_match in
match kind_of_term (Retroknowledge.find rk (KInt31 (grp,Int31Bits))) with
- | Ind i31bit_type ->
+ | Ind (i31bit_type,_) ->
add_native_before_match_info rk v (Nativelambda.before_match_int31 i31bit_type)
| _ ->
anomaly ~label:"Environ.register" (Pp.str "Int31Bits should be an inductive type")
@@ -498,13 +600,13 @@ fun env field value ->
operators to the reactive retroknowledge. *)
let add_int31_binop_from_const op prim =
match kind_of_term value with
- | Const kn -> retroknowledge add_int31_op env value 2
+ | Const (kn,_) -> retroknowledge add_int31_op env value 2
op prim kn
| _ -> anomaly ~label:"Environ.register" (Pp.str "should be a constant")
in
let add_int31_unop_from_const op prim =
match kind_of_term value with
- | Const kn -> retroknowledge add_int31_op env value 1
+ | Const (kn,_) -> retroknowledge add_int31_op env value 1
op prim kn
| _ -> anomaly ~label:"Environ.register" (Pp.str "should be a constant")
in
@@ -516,9 +618,9 @@ fun env field value ->
match field with
| KInt31 (grp, Int31Type) ->
(match kind_of_term (Retroknowledge.find rk (KInt31 (grp,Int31Bits))) with
- | Ind i31bit_type ->
+ | Ind (i31bit_type,_) ->
(match kind_of_term value with
- | Ind i31t ->
+ | Ind (i31t,_) ->
Retroknowledge.add_vm_decompile_constant_info rk
value (constr_of_int31 i31t i31bit_type)
| _ -> anomaly ~label:"Environ.register" (Pp.str "should be an inductive type"))
@@ -530,7 +632,7 @@ fun env field value ->
match field with
| KInt31 (grp, Int31Type) ->
let i31c = match kind_of_term value with
- | Ind i31t -> mkConstruct (i31t, 1)
+ | Ind i31t -> mkConstructUi (i31t, 1)
| _ -> anomaly ~label:"Environ.register" (Pp.str "should be an inductive type")
in
add_int31_decompilation_from_type
@@ -554,7 +656,7 @@ fun env field value ->
Primitives.Int31mulc
| KInt31 (_, Int31Div21) -> (* this is a ternary operation *)
(match kind_of_term value with
- | Const kn ->
+ | Const (kn,u) ->
retroknowledge add_int31_op env value 3
Cbytecodes.Kdiv21int31 Primitives.Int31div21 kn
| _ -> anomaly ~label:"Environ.register" (Pp.str "should be a constant"))
@@ -562,7 +664,7 @@ fun env field value ->
Primitives.Int31div
| KInt31 (_, Int31AddMulDiv) -> (* this is a ternary operation *)
(match kind_of_term value with
- | Const kn ->
+ | Const (kn,u) ->
retroknowledge add_int31_op env value 3
Cbytecodes.Kaddmuldivint31 Primitives.Int31addmuldiv kn
| _ -> anomaly ~label:"Environ.register" (Pp.str "should be a constant"))
diff --git a/kernel/environ.mli b/kernel/environ.mli
index 652cd59bf1..fb5d79718c 100644
--- a/kernel/environ.mli
+++ b/kernel/environ.mli
@@ -10,6 +10,7 @@ open Names
open Term
open Context
open Declarations
+open Univ
(** Unsafe environments. We define here a datatype for environments.
Since typing is not yet defined, it is not possible to check the
@@ -47,6 +48,7 @@ val named_context_val : env -> named_context_val
val engagement : env -> engagement option
+val is_impredicative_set : env -> bool
(** is the local context empty *)
val empty_context : env -> bool
@@ -125,17 +127,36 @@ val add_constant_key : constant -> constant_body -> Pre_env.link_info ref ->
val lookup_constant : constant -> env -> constant_body
val evaluable_constant : constant -> env -> bool
+val lookup_projection : Names.projection -> env -> projection_body
+val is_projection : constant -> env -> bool
+
(** {6 ... } *)
(** [constant_value env c] raises [NotEvaluableConst Opaque] if
[c] is opaque and [NotEvaluableConst NoBody] if it has no
- body and [Not_found] if it does not exist in [env] *)
+ body and [NotEvaluableConst IsProj] if [c] is a projection
+ and [Not_found] if it does not exist in [env] *)
-type const_evaluation_result = NoBody | Opaque
+type const_evaluation_result = NoBody | Opaque | IsProj
exception NotEvaluableConst of const_evaluation_result
-val constant_value : env -> constant -> constr
-val constant_type : env -> constant -> constant_type
-val constant_opt_value : env -> constant -> constr option
+val constant_value : env -> constant puniverses -> constr constrained
+val constant_type : env -> constant puniverses -> types constrained
+val constant_type_in_ctx : env -> constant -> types Univ.in_universe_context
+
+val constant_opt_value : env -> constant puniverses -> (constr * Univ.constraints) option
+val constant_value_and_type : env -> constant puniverses ->
+ types option * constr * Univ.constraints
+(** The universe context associated to the constant, empty if not
+ polymorphic *)
+val constant_context : env -> constant -> Univ.universe_context
+
+(* These functions should be called under the invariant that [env]
+ already contains the constraints corresponding to the constant
+ application. *)
+val constant_value_in : env -> constant puniverses -> constr
+val constant_type_in : env -> constant puniverses -> types
+val constant_opt_value_in : env -> constant puniverses -> constr option
+
(** {5 Inductive types } *)
val add_mind_key : mutual_inductive -> Pre_env.mind_key -> env -> env
@@ -157,8 +178,17 @@ val lookup_modtype : module_path -> env -> module_type_body
(** {5 Universe constraints } *)
+(** Add universe constraints to the environment.
+ @raises UniverseInconsistency
+*)
val add_constraints : Univ.constraints -> env -> env
+(** Check constraints are satifiable in the environment. *)
+val check_constraints : Univ.constraints -> env -> bool
+val push_context : Univ.universe_context -> env -> env
+val push_context_set : Univ.universe_context_set -> env -> env
+val push_constraints_to_env : 'a Univ.constrained -> env -> env
+
val set_engagement : engagement -> env -> env
(** {6 Sets of referred section variables }
diff --git a/kernel/fast_typeops.ml b/kernel/fast_typeops.ml
new file mode 100644
index 0000000000..b441e02a36
--- /dev/null
+++ b/kernel/fast_typeops.ml
@@ -0,0 +1,475 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Errors
+open Util
+open Names
+open Univ
+open Term
+open Vars
+open Context
+open Declarations
+open Environ
+open Entries
+open Reduction
+open Inductive
+open Type_errors
+
+let conv_leq l2r env x y = default_conv CUMUL ~l2r env x y
+
+let conv_leq_vecti env v1 v2 =
+ Array.fold_left2_i
+ (fun i _ t1 t2 ->
+ try conv_leq false env t1 t2
+ with NotConvertible -> raise (NotConvertibleVect i))
+ ()
+ v1
+ v2
+
+let check_constraints cst env =
+ if Environ.check_constraints cst env then ()
+ else error_unsatisfied_constraints env cst
+
+(* This should be a type (a priori without intension to be an assumption) *)
+let type_judgment env c t =
+ match kind_of_term(whd_betadeltaiota env t) with
+ | Sort s -> {utj_val = c; utj_type = s }
+ | _ -> error_not_type env (make_judge c t)
+
+let check_type env c t =
+ match kind_of_term(whd_betadeltaiota env t) with
+ | Sort s -> s
+ | _ -> error_not_type env (make_judge c t)
+
+(* This should be a type intended to be assumed. The error message is *)
+(* not as useful as for [type_judgment]. *)
+let assumption_of_judgment env t ty =
+ try let _ = check_type env t ty in t
+ with TypeError _ ->
+ error_assumption env (make_judge t ty)
+
+(************************************************)
+(* Incremental typing rules: builds a typing judgement given the *)
+(* judgements for the subterms. *)
+
+(*s Type of sorts *)
+
+(* Prop and Set *)
+
+let judge_of_prop = mkSort type1_sort
+let judge_of_set = judge_of_prop
+
+let judge_of_prop_contents _ = judge_of_prop
+
+(* Type of Type(i). *)
+
+let judge_of_type u =
+ let uu = Universe.super u in
+ mkType uu
+
+(*s Type of a de Bruijn index. *)
+
+let judge_of_relative env n =
+ try
+ let (_,_,typ) = lookup_rel n env in
+ lift n typ
+ with Not_found ->
+ error_unbound_rel env n
+
+(* Type of variables *)
+let judge_of_variable env id =
+ try named_type id env
+ with Not_found ->
+ error_unbound_var env id
+
+(* Management of context of variables. *)
+
+(* Checks if a context of variables can be instantiated by the
+ variables of the current env *)
+(* TODO: check order? *)
+let check_hyps_inclusion env f c sign =
+ Context.fold_named_context
+ (fun (id,_,ty1) () ->
+ try
+ let ty2 = named_type id env in
+ if not (eq_constr ty2 ty1) then raise Exit
+ with Not_found | Exit ->
+ error_reference_variables env id (f c))
+ sign
+ ~init:()
+
+(* Instantiation of terms on real arguments. *)
+
+(* Make a type polymorphic if an arity *)
+
+(* Type of constants *)
+
+let type_of_constant env cst = constant_type env cst
+let type_of_constant_in env cst = constant_type_in env cst
+let type_of_constant_knowing_parameters env t _ = t
+
+let judge_of_constant env (kn,u as cst) =
+ let cb = lookup_constant kn env in
+ let () = check_hyps_inclusion env mkConstU cst cb.const_hyps in
+ let ty, cu = type_of_constant env cst in
+ let () = check_constraints cu env in
+ ty
+
+let type_of_projection env (cst,u) =
+ let cb = lookup_constant cst env in
+ match cb.const_proj with
+ | Some pb ->
+ if cb.const_polymorphic then
+ let mib,_ = lookup_mind_specif env (pb.proj_ind,0) in
+ let subst = make_inductive_subst mib u in
+ Vars.subst_univs_constr subst pb.proj_type
+ else pb.proj_type
+ | None -> raise (Invalid_argument "type_of_projection: not a projection")
+
+
+(* Type of a lambda-abstraction. *)
+
+(* [judge_of_abstraction env name var j] implements the rule
+
+ env, name:typ |- j.uj_val:j.uj_type env, |- (name:typ)j.uj_type : s
+ -----------------------------------------------------------------------
+ env |- [name:typ]j.uj_val : (name:typ)j.uj_type
+
+ Since all products are defined in the Calculus of Inductive Constructions
+ and no upper constraint exists on the sort $s$, we don't need to compute $s$
+*)
+
+let judge_of_abstraction env name var ty =
+ mkProd (name, var, ty)
+
+(* Type of let-in. *)
+
+let judge_of_letin env name defj typj j =
+ subst1 defj j
+
+(* Type of an application. *)
+
+let make_judgev c t =
+ Array.map2 make_judge c t
+
+let judge_of_apply env func funt argsv argstv =
+ let len = Array.length argsv in
+ let rec apply_rec i typ =
+ if Int.equal i len then typ
+ else
+ (match kind_of_term (whd_betadeltaiota env typ) with
+ | Prod (_,c1,c2) ->
+ let arg = argsv.(i) and argt = argstv.(i) in
+ (try
+ let () = conv_leq false env argt c1 in
+ apply_rec (i+1) (subst1 arg c2)
+ with NotConvertible ->
+ error_cant_apply_bad_type env
+ (i+1,c1,argt)
+ (make_judge func funt)
+ (make_judgev argsv argstv))
+
+ | _ ->
+ error_cant_apply_not_functional env
+ (make_judge func funt)
+ (make_judgev argsv argstv))
+ in apply_rec 0 funt
+
+(* Type of product *)
+
+let sort_of_product env domsort rangsort =
+ match (domsort, rangsort) with
+ (* Product rule (s,Prop,Prop) *)
+ | (_, Prop Null) -> rangsort
+ (* Product rule (Prop/Set,Set,Set) *)
+ | (Prop _, Prop Pos) -> rangsort
+ (* Product rule (Type,Set,?) *)
+ | (Type u1, Prop Pos) ->
+ begin match engagement env with
+ | Some ImpredicativeSet ->
+ (* Rule is (Type,Set,Set) in the Set-impredicative calculus *)
+ rangsort
+ | _ ->
+ (* Rule is (Type_i,Set,Type_i) in the Set-predicative calculus *)
+ Type (Universe.sup Universe.type0 u1)
+ end
+ (* Product rule (Prop,Type_i,Type_i) *)
+ | (Prop Pos, Type u2) -> Type (Universe.sup Universe.type0 u2)
+ (* Product rule (Prop,Type_i,Type_i) *)
+ | (Prop Null, Type _) -> rangsort
+ (* Product rule (Type_i,Type_i,Type_i) *)
+ | (Type u1, Type u2) -> Type (Universe.sup u1 u2)
+
+(* [judge_of_product env name (typ1,s1) (typ2,s2)] implements the rule
+
+ env |- typ1:s1 env, name:typ1 |- typ2 : s2
+ -------------------------------------------------------------------------
+ s' >= (s1,s2), env |- (name:typ)j.uj_val : s'
+
+ where j.uj_type is convertible to a sort s2
+*)
+let judge_of_product env name s1 s2 =
+ let s = sort_of_product env s1 s2 in
+ mkSort s
+
+(* Type of a type cast *)
+
+(* [judge_of_cast env (c,typ1) (typ2,s)] implements the rule
+
+ env |- c:typ1 env |- typ2:s env |- typ1 <= typ2
+ ---------------------------------------------------------------------
+ env |- c:typ2
+*)
+
+let judge_of_cast env c ct k expected_type =
+ try
+ match k with
+ | VMcast ->
+ vm_conv CUMUL env ct expected_type
+ | DEFAULTcast ->
+ default_conv ~l2r:false CUMUL env ct expected_type
+ | REVERTcast ->
+ default_conv ~l2r:true CUMUL env ct expected_type
+ | NATIVEcast ->
+ let sigma = Nativelambda.empty_evars in
+ native_conv CUMUL sigma env ct expected_type
+ with NotConvertible ->
+ error_actual_type env (make_judge c ct) expected_type
+
+(* Inductive types. *)
+
+(* The type is parametric over the uniform parameters whose conclusion
+ is in Type; to enforce the internal constraints between the
+ parameters and the instances of Type occurring in the type of the
+ constructors, we use the level variables _statically_ assigned to
+ the conclusions of the parameters as mediators: e.g. if a parameter
+ has conclusion Type(alpha), static constraints of the form alpha<=v
+ exist between alpha and the Type's occurring in the constructor
+ types; when the parameters is finally instantiated by a term of
+ conclusion Type(u), then the constraints u<=alpha is computed in
+ the App case of execute; from this constraints, the expected
+ dynamic constraints of the form u<=v are enforced *)
+
+(* let judge_of_inductive_knowing_parameters env ind jl = *)
+(* let c = mkInd ind in *)
+(* let (mib,mip) = lookup_mind_specif env ind in *)
+(* check_args env c mib.mind_hyps; *)
+(* let paramstyp = Array.map (fun j -> j.uj_type) jl in *)
+(* let t = in *)
+(* make_judge c t *)
+
+let judge_of_inductive env (ind,u as indu) =
+ let (mib,mip) = lookup_mind_specif env ind in
+ check_hyps_inclusion env mkIndU indu mib.mind_hyps;
+ let t,cst = Inductive.constrained_type_of_inductive env ((mib,mip),u) in
+ check_constraints cst env;
+ t
+
+(* Constructors. *)
+
+let judge_of_constructor env (c,u as cu) =
+ let _ =
+ let ((kn,_),_) = c in
+ let mib = lookup_mind kn env in
+ check_hyps_inclusion env mkConstructU cu mib.mind_hyps in
+ let specif = lookup_mind_specif env (inductive_of_constructor c) in
+ let t,cst = constrained_type_of_constructor cu specif in
+ let () = check_constraints cst env in
+ t
+
+(* Case. *)
+
+let check_branch_types env (ind,u) c ct lft explft =
+ try conv_leq_vecti env lft explft
+ with
+ NotConvertibleVect i ->
+ error_ill_formed_branch env c ((ind,i+1),u) lft.(i) explft.(i)
+ | Invalid_argument _ ->
+ error_number_branches env (make_judge c ct) (Array.length explft)
+
+let judge_of_case env ci p pt c ct lf lft =
+ let (pind, _ as indspec) =
+ try find_rectype env ct
+ with Not_found -> error_case_not_inductive env (make_judge c ct) in
+ let _ = check_case_info env pind ci in
+ let (bty,rslty) =
+ type_case_branches env indspec (make_judge p pt) c in
+ let () = check_branch_types env pind c ct lft bty in
+ rslty
+
+let judge_of_projection env p c ct =
+ let pb = lookup_projection p env in
+ let (ind,u), args =
+ try find_rectype env ct
+ with Not_found -> error_case_not_inductive env (make_judge c ct)
+ in
+ assert(eq_mind pb.proj_ind (fst ind));
+ let usubst = make_inductive_subst (fst (lookup_mind_specif env ind)) u in
+ let ty = Vars.subst_univs_constr usubst pb.Declarations.proj_type in
+ substl (c :: List.rev args) ty
+
+
+(* Fixpoints. *)
+
+(* Checks the type of a general (co)fixpoint, i.e. without checking *)
+(* the specific guard condition. *)
+
+let type_fixpoint env lna lar vdef vdeft =
+ let lt = Array.length vdeft in
+ assert (Int.equal (Array.length lar) lt);
+ try
+ conv_leq_vecti env vdeft (Array.map (fun ty -> lift lt ty) lar)
+ with NotConvertibleVect i ->
+ error_ill_typed_rec_body env i lna (make_judgev vdef vdeft) lar
+
+(************************************************************************)
+(************************************************************************)
+
+(* The typing machine. *)
+ (* ATTENTION : faudra faire le typage du contexte des Const,
+ Ind et Constructsi un jour cela devient des constructions
+ arbitraires et non plus des variables *)
+let rec execute env cstr =
+ match kind_of_term cstr with
+ (* Atomic terms *)
+ | Sort (Prop c) ->
+ judge_of_prop_contents c
+
+ | Sort (Type u) ->
+ judge_of_type u
+
+ | Rel n ->
+ judge_of_relative env n
+
+ | Var id ->
+ judge_of_variable env id
+
+ | Const c ->
+ judge_of_constant env c
+
+ | Proj (p, c) ->
+ let ct = execute env c in
+ judge_of_projection env p c ct
+
+ (* Lambda calculus operators *)
+ | App (f,args) ->
+ let argst = execute_array env args in
+ let ft = execute env f in
+ judge_of_apply env f ft args argst
+
+ | Lambda (name,c1,c2) ->
+ let _ = execute_is_type env c1 in
+ let env1 = push_rel (name,None,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 vars' = execute_is_type env1 c2 in
+ judge_of_product env name vars vars'
+
+ | LetIn (name,c1,c2,c3) ->
+ 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 c3t = execute env1 c3 in
+ subst1 c1 c3t
+
+ | Cast (c,k,t) ->
+ let ct = execute env c in
+ let _ts = execute_type env t in
+ let _ = judge_of_cast env c ct k t in
+ t
+
+ (* Inductive types *)
+ | Ind ind ->
+ judge_of_inductive env ind
+
+ | Construct c ->
+ judge_of_constructor env c
+
+ | Case (ci,p,c,lf) ->
+ let ct = execute env c in
+ let pt = execute env p in
+ let lft = execute_array env lf in
+ judge_of_case env ci p pt c ct lf lft
+
+ | Fix ((vn,i as vni),recdef) ->
+ let (fix_ty,recdef') = execute_recdef env recdef i in
+ let fix = (vni,recdef') in
+ check_fix env fix; fix_ty
+
+ | CoFix (i,recdef) ->
+ let (fix_ty,recdef') = execute_recdef env recdef i in
+ let cofix = (i,recdef') in
+ check_cofix env cofix; fix_ty
+
+ (* Partial proofs: unsupported by the kernel *)
+ | Meta _ ->
+ anomaly (Pp.str "the kernel does not support metavariables")
+
+ | Evar _ ->
+ anomaly (Pp.str "the kernel does not support existential variables")
+
+and execute_is_type env constr =
+ let t = execute env constr in
+ check_type env constr t
+
+and execute_type env constr =
+ let t = execute env constr in
+ type_judgment env constr t
+
+and execute_recdef env (names,lar,vdef) i =
+ let lart = execute_array env lar in
+ let lara = Array.map2 (assumption_of_judgment env) lar lart in
+ let env1 = push_rec_types (names,lara,vdef) env in
+ let vdeft = execute_array env1 vdef in
+ let () = type_fixpoint env1 names lara vdef vdeft in
+ (lara.(i),(names,lara,vdef))
+
+and execute_array env = Array.map (execute env)
+
+(* Derived functions *)
+let infer env constr =
+ let t = execute env constr in
+ make_judge constr t
+
+(* let infer_key = Profile.declare_profile "Fast_infer" *)
+(* let infer = Profile.profile2 infer_key infer *)
+
+let infer_type env constr =
+ execute_type env constr
+
+let infer_v env cv =
+ let jv = execute_array env cv in
+ make_judgev cv jv
+
+(* Typing of several terms. *)
+
+let infer_local_decl env id = function
+ | LocalDef c ->
+ let t = execute env c in
+ (Name id, Some c, t)
+ | LocalAssum c ->
+ let t = execute env c in
+ (Name id, None, assumption_of_judgment env c t)
+
+let infer_local_decls env decls =
+ let rec inferec env = function
+ | (id, d) :: l ->
+ let (env, l) = inferec env l in
+ let d = infer_local_decl env id d in
+ (push_rel d env, add_rel_decl d l)
+ | [] -> (env, empty_rel_context) in
+ inferec env decls
+
+(* Exported typing functions *)
+
+let typing env c = infer env c
diff --git a/kernel/fast_typeops.mli b/kernel/fast_typeops.mli
new file mode 100644
index 0000000000..7ff5577cb3
--- /dev/null
+++ b/kernel/fast_typeops.mli
@@ -0,0 +1,28 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Univ
+open Term
+open Context
+open Environ
+open Entries
+open Declarations
+
+(** {6 Typing functions (not yet tagged as safe) }
+
+ They return unsafe judgments that are "in context" of a set of
+ (local) universe variables (the ones that appear in the term)
+ and associated constraints. In case of polymorphic definitions,
+ these variables and constraints will be generalized.
+ *)
+
+
+val infer : env -> constr -> unsafe_judgment
+val infer_v : env -> constr array -> unsafe_judgment array
+val infer_type : env -> types -> unsafe_type_judgment
diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml
index 2defb66f4b..0ac6a4e4a9 100644
--- a/kernel/indtypes.ml
+++ b/kernel/indtypes.ml
@@ -20,6 +20,15 @@ open Environ
open Reduction
open Typeops
open Entries
+open Pp
+
+(* Tell if indices (aka real arguments) contribute to size of inductive type *)
+(* If yes, this is compatible with the univalent model *)
+
+let indices_matter = ref false
+
+let enforce_indices_matter () = indices_matter := true
+let is_indices_matter () = !indices_matter
(* Same as noccur_between but may perform reductions.
Could be refined more... *)
@@ -107,26 +116,22 @@ let is_logic_constr infos = List.for_all (fun (logic,small) -> logic) infos
*)
let is_unit constrsinfos =
match constrsinfos with (* One info = One constructor *)
- | [constrinfos] -> is_logic_constr constrinfos
+ | [level] -> is_type0m_univ level
| [] -> (* type without constructors *) true
| _ -> false
-let rec infos_and_sort env t =
- let t = whd_betadeltaiota env t in
- match kind_of_term t with
- | Prod (name,c1,c2) ->
- let (varj,_) = infer_type env c1 in
+let infos_and_sort env ctx t =
+ let rec aux env ctx t max =
+ let t = whd_betadeltaiota env t in
+ 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 logic = is_logic_type varj in
- let small = Term.is_small varj.utj_type in
- (logic,small) :: (infos_and_sort env1 c2)
- | _ when is_constructor_head t -> []
- | _ -> (* don't fail if not positive, it is tested later *) []
-
-let small_unit constrsinfos =
- let issmall = List.for_all is_small constrsinfos
- and isunit = is_unit constrsinfos in
- issmall, isunit
+ let max = Universe.sup max (univ_of_sort varj.utj_type) in
+ aux env1 ctx c2 max
+ | _ when is_constructor_head t -> max
+ | _ -> (* don't fail if not positive, it is tested later *) max
+ in aux env ctx t Universe.type0m
(* Computing the levels of polymorphic inductive types
@@ -148,40 +153,52 @@ let small_unit constrsinfos =
w1,w2,w3 <= u3
*)
-let extract_level (_,_,_,lc,lev) =
+let extract_level (_,_,lc,(_,lev)) =
(* Enforce that the level is not in Prop if more than one constructor *)
- if Array.length lc >= 2 then sup type0_univ lev else lev
+ (* if Array.length lc >= 2 then sup type0_univ lev else lev *)
+ lev
let inductive_levels arities inds =
- let levels = Array.map pi3 arities in
let cstrs_levels = Array.map extract_level inds in
(* Take the transitive closure of the system of constructors *)
(* level constraints and remove the recursive dependencies *)
- solve_constraints_system levels cstrs_levels
+ cstrs_levels
(* This (re)computes informations relevant to extraction and the sort of an
arity or type constructor; we do not to recompute universes constraints *)
-let constraint_list_union =
- List.fold_left union_constraints empty_constraint
+let context_set_list_union =
+ List.fold_left ContextSet.union ContextSet.empty
-let infer_constructor_packet env_ar_par params lc =
+let infer_constructor_packet env_ar_par ctx params lc =
(* type-check the constructors *)
- let jlc,cstl = List.split (List.map (infer_type env_ar_par) lc) in
- let cst = constraint_list_union cstl in
+ let jlc = List.map (infer_type env_ar_par) lc in
let jlc = Array.of_list jlc in
(* generalize the constructor over the parameters *)
let lc'' = Array.map (fun j -> it_mkProd_or_LetIn j.utj_val params) jlc in
- (* compute the max of the sorts of the products of the constructor type *)
- let level = max_inductive_sort (Array.map (fun j -> j.utj_type) jlc) in
- (* compute *)
- let info = small_unit (List.map (infos_and_sort env_ar_par) lc) in
-
- (info,lc'',level,cst)
+ (* compute the max of the sorts of the products of the constructors types *)
+ let levels = List.map (infos_and_sort env_ar_par ctx) lc in
+ let level = List.fold_left (fun max l -> Universe.sup max l) Universe.type0m levels in
+ (lc'',(is_unit levels,level))
+
+(* If indices matter *)
+let cumulate_arity_large_levels env sign =
+ fst (List.fold_right
+ (fun (_,_,t as d) (lev,env) ->
+ let tj = infer_type env t in
+ let u = univ_of_sort tj.utj_type in
+ (Universe.sup u lev, push_rel d env))
+ sign (Universe.type0m,env))
+
+let is_impredicative env u =
+ is_type0m_univ u || (is_type0_univ u && engagement env = Some ImpredicativeSet)
(* Type-check an inductive definition. Does not check positivity
conditions. *)
-let typecheck_inductive env mie =
+(* TODO check that we don't overgeneralize construcors/inductive arities with
+ universes that are absent from them. Is it possible?
+*)
+let typecheck_inductive env ctx mie =
let () = match mie.mind_entry_inds with
| [] -> anomaly (Pp.str "empty inductive types declaration")
| _ -> ()
@@ -189,116 +206,103 @@ let typecheck_inductive env mie =
(* Check unicity of names *)
mind_check_names mie;
(* Params are typed-checked here *)
- let env_params, params, cst1 = infer_local_decls env mie.mind_entry_params in
+ let env' = push_context ctx env in
+ let (env_params, params) = infer_local_decls env' mie.mind_entry_params in
(* We first type arity of each inductive definition *)
(* This allows to build the environment of arities and to share *)
(* the set of constraints *)
- let cst, env_arities, rev_arity_list =
+ let env_arities, rev_arity_list =
List.fold_left
- (fun (cst,env_ar,l) ind ->
+ (fun (env_ar,l) ind ->
(* Arities (without params) are typed-checked here *)
- let arity, cst2 = infer_type env_params ind.mind_entry_arity in
+ let arity =
+ if isArity ind.mind_entry_arity then
+ let (ctx,s) = destArity ind.mind_entry_arity in
+ match s with
+ | Type u when Univ.universe_level u = None ->
+ (** We have an algebraic universe as the conclusion of the arity,
+ typecheck the dummy Π ctx, Prop and do a special case for the conclusion.
+ *)
+ let proparity = infer_type env_params (mkArity (ctx, prop_sort)) in
+ let (cctx, _) = destArity proparity.utj_val in
+ (* Any universe is well-formed, we don't need to check [s] here *)
+ mkArity (cctx, s)
+ | _ -> let arity = infer_type env_params ind.mind_entry_arity in
+ arity.utj_val
+ else let arity = infer_type env_params ind.mind_entry_arity in
+ arity.utj_val
+ in
+ let (sign, deflev) = dest_arity env_params arity in
+ let inflev =
+ (* The level of the inductive includes levels of indices if
+ in indices_matter mode *)
+ if !indices_matter
+ then Some (cumulate_arity_large_levels env_params sign)
+ else None
+ in
(* We do not need to generate the universe of full_arity; if
later, after the validation of the inductive definition,
full_arity is used as argument or subject to cast, an
upper universe will be generated *)
- let full_arity = it_mkProd_or_LetIn arity.utj_val params in
- let cst = union_constraints cst cst2 in
+ 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)
- (add_constraints cst2 env_ar) in
- let lev =
- (* Decide that if the conclusion is not explicitly Type *)
- (* then the inductive type is not polymorphic *)
- match kind_of_term ((strip_prod_assum arity.utj_val)) with
- | Sort (Type u) -> Some u
- | _ -> None in
- (cst,env_ar',(id,full_arity,lev)::l))
- (cst1,env,[])
+ push_rel (Name id, None, full_arity) env_ar in
+ (* (add_constraints cst2 env_ar) in *)
+ (env_ar', (id,full_arity,sign @ params,deflev,inflev)::l))
+ (env',[])
mie.mind_entry_inds in
let arity_list = List.rev rev_arity_list in
(* builds the typing context "Gamma, I1:A1, ... In:An, params" *)
- let env_ar_par =
- push_rel_context params (add_constraints cst1 env_arities) in
+ let env_ar_par = push_rel_context params env_arities in
(* Now, we type the constructors (without params) *)
- let inds,cst =
+ let inds =
List.fold_right2
- (fun ind arity_data (inds,cst) ->
- let (info,lc',cstrs_univ,cst') =
- infer_constructor_packet env_ar_par params ind.mind_entry_lc in
+ (fun ind arity_data inds ->
+ let (lc',cstrs_univ) =
+ infer_constructor_packet env_ar_par ContextSet.empty
+ params ind.mind_entry_lc in
let consnames = ind.mind_entry_consnames in
- let ind' = (arity_data,consnames,info,lc',cstrs_univ) in
- (ind'::inds, union_constraints cst cst'))
+ let ind' = (arity_data,consnames,lc',cstrs_univ) in
+ ind'::inds)
mie.mind_entry_inds
arity_list
- ([],cst) in
+ ([]) in
let inds = Array.of_list inds in
- let arities = Array.of_list arity_list in
- let has_some_univ u = function
- | Some v when Universe.equal u v -> true
- | _ -> false
- in
- let remove_some_univ u = function
- | Some v when Universe.equal u v -> None
- | x -> x
- in
- let fold l (_, b, p) = match b with
- | None ->
- (* Parameter contributes to polymorphism only if explicit Type *)
- let c = strip_prod_assum p in
- (* Add Type levels to the ordered list of parameters contributing to *)
- (* polymorphism unless there is aliasing (i.e. non distinct levels) *)
- begin match kind_of_term c with
- | Sort (Type u) ->
- if List.exists (has_some_univ u) l then
- None :: List.map (remove_some_univ u) l
- else
- Some u :: l
- | _ ->
- None :: l
- end
- | _ -> l
- in
- let param_ccls = List.fold_left fold [] params in
(* Compute/check the sorts of the inductive types *)
- let ind_min_levels = inductive_levels arities inds in
- let inds, cst =
- Array.fold_map2' (fun ((id,full_arity,ar_level),cn,info,lc,_) lev cst ->
- let sign, s =
- try dest_arity env full_arity
- with NotArity -> raise (InductiveError (NotAnArity (env, full_arity)))
+
+ let inds =
+ Array.map (fun ((id,full_arity,sign,def_level,inf_level),cn,lc,(is_unit,clev)) ->
+ let defu = Term.univ_of_sort def_level in
+ let infu =
+ (** Inferred level, with parameters and constructors. *)
+ match inf_level with
+ | Some alev -> Universe.sup clev alev
+ | None -> clev
+ in
+ let is_natural =
+ check_leq (universes env') infu defu &&
+ not (is_type0m_univ defu && not is_unit)
in
- let status,cst = match s with
- | Type u when ar_level != None (* Explicitly polymorphic *)
- && no_upper_constraints u cst ->
- (* The polymorphic level is a function of the level of the *)
- (* conclusions of the parameters *)
- (* We enforce [u >= lev] in case [lev] has a strict upper *)
- (* constraints over [u] *)
- Inr (param_ccls, lev), enforce_leq lev u cst
- | Type u (* Not an explicit occurrence of Type *) ->
- Inl (info,full_arity,s), enforce_leq lev u cst
- | Prop Pos when
- begin match engagement env with
- | Some ImpredicativeSet -> false
- | _ -> true
- end ->
- (* Predicative set: check that the content is indeed predicative *)
- if not (is_type0m_univ lev) && not (is_type0_univ lev) then
- raise (InductiveError LargeNonPropInductiveNotInType);
- Inl (info,full_arity,s), cst
- | Prop _ ->
- Inl (info,full_arity,s), cst in
- (id,cn,lc,(sign,status)),cst)
- inds ind_min_levels cst in
-
- (env_arities, params, inds, cst)
+ let _ =
+ (** Impredicative sort, always allow *)
+ if is_impredicative env defu then ()
+ else (** Predicative case: the inferred level must be lower or equal to the
+ declared level. *)
+ if not is_natural then
+ anomaly ~label:"check_inductive"
+ (Pp.str"Incorrect universe " ++
+ Universe.pr defu ++ Pp.str " declared for inductive type, inferred level is "
+ ++ Universe.pr infu)
+ in
+ (id,cn,lc,(sign,(not is_natural,full_arity,defu))))
+ inds
+ in (env_arities, params, inds)
(************************************************************************)
(************************************************************************)
@@ -387,7 +391,7 @@ if Int.equal nmr 0 then 0 else
in find 0 (n-1) (lpar,List.rev hyps)
let lambda_implicit_lift n a =
- let level = UniverseLevel.make (DirPath.make [Id.of_string "implicit"]) 0 in
+ let level = Level.make (DirPath.make [Id.of_string "implicit"]) 0 in
let implicit_sort = mkType (Universe.make level) in
let lambda_implicit a = mkLambda (Anonymous, implicit_sort, a) in
iterate lambda_implicit n (lift n a)
@@ -413,12 +417,13 @@ let abstract_mind_lc env ntyps npars lc =
let ienv_push_var (env, n, ntypes, lra) (x,a,ra) =
(push_rel (x,None,a) env, n+1, ntypes, (Norec,ra)::lra)
-let ienv_push_inductive (env, n, ntypes, ra_env) (mi,lpar) =
+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 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 (type_of_inductive env specif) lpar) env in
+ hnf_prod_applist env ty 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
@@ -476,7 +481,7 @@ let check_positivity_one (env,_,ntypes,_ as ienv) hyps (_,i as ind) nargs lcname
else failwith_non_pos_list n ntypes (x::largs)
(* accesses to the environment are not factorised, but is it worth? *)
- and check_positive_nested (env,n,ntypes,ra_env as ienv) nmr (mi, largs) =
+ and check_positive_nested (env,n,ntypes,ra_env as ienv) nmr ((mi,u), largs) =
let (mib,mip) = lookup_mind_specif env mi in
let auxnpar = mib.mind_nparams_rec in
let nonrecpar = mib.mind_nparams - auxnpar in
@@ -495,7 +500,7 @@ let check_positivity_one (env,_,ntypes,_ as ienv) hyps (_,i as ind) nargs lcname
let auxlcvect = abstract_mind_lc env auxntyp auxnpar mip.mind_nf_lc in
(* Extends the environment with a variable corresponding to
the inductive def *)
- let (env',_,_,_ as ienv') = ienv_push_inductive ienv (mi,lpar) in
+ let (env',_,_,_ as ienv') = ienv_push_inductive ienv ((mi,u),lpar) in
(* Parameters expressed in env' *)
let lpar' = List.map (lift auxntyp) lpar in
let irecargs_nmr =
@@ -586,40 +591,72 @@ let all_sorts = [InProp;InSet;InType]
let small_sorts = [InProp;InSet]
let logical_sorts = [InProp]
-let allowed_sorts issmall isunit s =
- match family_of_sort s with
- (* Type: all elimination allowed *)
- | InType -> all_sorts
-
- (* Small Set is predicative: all elimination allowed *)
- | InSet when issmall -> all_sorts
-
- (* Large Set is necessarily impredicative: forbids large elimination *)
- | InSet -> small_sorts
-
- (* Unitary/empty Prop: elimination to all sorts are realizable *)
- (* unless the type is large. If it is large, forbids large elimination *)
- (* which otherwise allows to simulate the inconsistent system Type:Type *)
- | InProp when isunit -> if issmall then all_sorts else small_sorts
-
- (* Other propositions: elimination only to Prop *)
- | InProp -> logical_sorts
+let allowed_sorts is_smashed s =
+ if not is_smashed
+ then (** Naturally in the defined sort.
+ If [s] is Prop, it must be small and unitary.
+ Unsmashed, predicative Type and Set: all elimination allowed
+ as well. *)
+ all_sorts
+ else
+ match family_of_sort s with
+ (* Type: all elimination allowed: above and below *)
+ | InType -> all_sorts
+ (* Smashed Set is necessarily impredicative: forbids large elimination *)
+ | InSet -> small_sorts
+ (* Smashed to Prop, no informative eliminations allowed *)
+ | InProp -> logical_sorts
+
+(* Previous comment: *)
+(* Unitary/empty Prop: elimination to all sorts are realizable *)
+(* unless the type is large. If it is large, forbids large elimination *)
+(* which otherwise allows to simulate the inconsistent system Type:Type. *)
+(* -> this is now handled by is_smashed: *)
+(* - all_sorts in case of small, unitary Prop (not smashed) *)
+(* - logical_sorts in case of large, unitary Prop (smashed) *)
let fold_inductive_blocks f =
- let concl = function
- | Inr _ -> mkSet (* dummy *)
- | Inl (_,ar,_) -> ar
- in
- Array.fold_left (fun acc (_,_,lc,(arsign,ar)) ->
- f (Array.fold_left f acc lc) (it_mkProd_or_LetIn (concl ar) arsign))
+ Array.fold_left (fun acc (_,_,lc,(arsign,ar)) ->
+ f (Array.fold_left f acc lc) (it_mkProd_or_LetIn (pi2 ar) arsign))
let used_section_variables env inds =
let ids = fold_inductive_blocks
(fun l c -> Id.Set.union (Environ.global_vars_set env c) l)
Id.Set.empty inds in
keep_hyps env ids
-
-let build_inductive env env_ar params isrecord isfinite inds nmr recargs cst =
+let lift_decl n d =
+ map_rel_declaration (lift n) d
+
+let rel_vect n m = Array.init m (fun i -> mkRel(n+m-i))
+let rel_list n m = Array.to_list (rel_vect n m)
+let rel_appvect n m = rel_vect n (List.length m)
+
+exception UndefinableExpansion
+
+(** From a rel context describing the constructor arguments,
+ build an expansion function.
+ The term built is expecting to be substituted first by
+ a substitution of the form [params, x : ind params] *)
+let compute_expansion ((kn, _ as ind), u) params ctx =
+ let mp, dp, l = repr_mind kn in
+ let make_proj id = Constant.make1 (KerName.make mp dp (Label.of_id id)) in
+ let rec projections acc (na, b, t) =
+ match b with
+ | Some c -> acc
+ | None ->
+ match na with
+ | Name id -> make_proj id :: acc
+ | Anonymous -> raise UndefinableExpansion
+ in
+ let projs = List.fold_left projections [] ctx in
+ let projarr = Array.of_list projs in
+ let exp =
+ mkApp (mkConstructU ((ind, 1),u),
+ Array.append (rel_appvect 1 params)
+ (Array.map (fun p -> mkProj (p, mkRel 1)) projarr))
+ in exp, projarr
+
+let build_inductive env p ctx env_ar params kn isrecord isfinite inds nmr recargs =
let ntypes = Array.length inds in
(* Compute the set of used section variables *)
let hyps = used_section_variables env inds in
@@ -637,18 +674,13 @@ let build_inductive env env_ar params isrecord isfinite inds nmr recargs cst =
Array.map (fun (d,_) -> rel_context_nhyps d - rel_context_nhyps params)
splayed_lc in
(* Elimination sorts *)
- let arkind,kelim = match ar_kind with
- | Inr (param_levels,lev) ->
- Polymorphic {
- poly_param_levels = param_levels;
- poly_level = lev;
- }, all_sorts
- | Inl ((issmall,isunit),ar,s) ->
- let kelim = allowed_sorts issmall isunit s in
- Monomorphic {
- mind_user_arity = ar;
- mind_sort = s;
- }, kelim in
+ let arkind,kelim =
+ let (info,ar,defs) = ar_kind in
+ let s = sort_of_univ defs in
+ let kelim = allowed_sorts info s in
+ { mind_user_arity = ar;
+ mind_sort = s;
+ }, kelim in
(* Assigning VM tags to constructors *)
let nconst, nblock = ref 0, ref 0 in
let transf num =
@@ -681,6 +713,19 @@ let build_inductive env env_ar params isrecord isfinite inds nmr recargs cst =
mind_reloc_tbl = rtbl;
} in
let packets = Array.map2 build_one_packet inds recargs in
+ let isrecord =
+ let pkt = packets.(0) in
+ if isrecord (* || (Array.length pkt.mind_consnames = 1 && *)
+ (* inductive_sort_family pkt <> InProp) *) then
+ let rctx, _ = decompose_prod_assum pkt.mind_nf_lc.(0) in
+ let u = if p then Univ.UContext.instance ctx else Univ.Instance.empty in
+ try
+ let exp = compute_expansion ((kn, 0), u) params
+ (List.firstn pkt.mind_consnrealdecls.(0) rctx)
+ in Some exp
+ with UndefinableExpansion -> None
+ else None
+ in
(* Build the mutual inductive *)
{ mind_record = isrecord;
mind_ntypes = ntypes;
@@ -690,7 +735,8 @@ let build_inductive env env_ar params isrecord isfinite inds nmr recargs cst =
mind_nparams_rec = nmr;
mind_params_ctxt = params;
mind_packets = packets;
- mind_constraints = cst
+ mind_polymorphic = p;
+ mind_universes = ctx;
}
(************************************************************************)
@@ -698,9 +744,14 @@ let build_inductive env env_ar params isrecord isfinite inds nmr recargs cst =
let check_inductive env kn mie =
(* First type-check the inductive definition *)
- let (env_ar, params, inds, cst) = typecheck_inductive env mie in
+ let env = push_context mie.mind_entry_universes env in
+ let (env_ar, params, inds) =
+ typecheck_inductive env mie.mind_entry_universes mie
+ in
(* Then check positivity conditions *)
let (nmr,recargs) = check_positivity kn env_ar params inds in
(* Build the inductive packets *)
- build_inductive env env_ar params mie.mind_entry_record mie.mind_entry_finite
- inds nmr recargs cst
+ build_inductive env mie.mind_entry_polymorphic
+ mie.mind_entry_universes
+ env_ar params kn mie.mind_entry_record mie.mind_entry_finite
+ inds nmr recargs
diff --git a/kernel/indtypes.mli b/kernel/indtypes.mli
index 016a1a5b55..8730a30457 100644
--- a/kernel/indtypes.mli
+++ b/kernel/indtypes.mli
@@ -34,5 +34,12 @@ exception InductiveError of inductive_error
(** The following function does checks on inductive declarations. *)
-val check_inductive :
- env -> mutual_inductive -> mutual_inductive_entry -> mutual_inductive_body
+val check_inductive : env -> mutual_inductive -> mutual_inductive_entry -> mutual_inductive_body
+
+(** The following enforces a system compatible with the univalent model *)
+
+val enforce_indices_matter : unit -> unit
+val is_indices_matter : unit -> bool
+
+val compute_expansion : pinductive ->
+ Context.rel_context -> Context.rel_context -> (constr * constant array)
diff --git a/kernel/inductive.ml b/kernel/inductive.ml
index 2b2caaf3bf..e57b0b4ad4 100644
--- a/kernel/inductive.ml
+++ b/kernel/inductive.ml
@@ -19,6 +19,9 @@ open Environ
open Reduction
open Type_errors
+type pinductive = inductive puniverses
+type pconstructor = constructor puniverses
+
type mind_specif = mutual_inductive_body * one_inductive_body
(* raise Not_found if not an inductive type *)
@@ -38,31 +41,55 @@ let find_inductive env c =
let (t, l) = decompose_app (whd_betadeltaiota env c) in
match kind_of_term t with
| Ind ind
- when (fst (lookup_mind_specif env ind)).mind_finite -> (ind, l)
+ when (fst (lookup_mind_specif env (out_punivs ind))).mind_finite -> (ind, l)
| _ -> raise Not_found
let find_coinductive env c =
let (t, l) = decompose_app (whd_betadeltaiota env c) in
match kind_of_term t with
| Ind ind
- when not (fst (lookup_mind_specif env ind)).mind_finite -> (ind, l)
+ when not (fst (lookup_mind_specif env (out_punivs ind))).mind_finite -> (ind, l)
| _ -> raise Not_found
let inductive_params (mib,_) = mib.mind_nparams
+let make_inductive_subst mib u =
+ if mib.mind_polymorphic then
+ make_universe_subst u mib.mind_universes
+ else Univ.empty_subst
+
+let inductive_params_ctxt (mib,u) =
+ let subst = make_inductive_subst mib u in
+ Vars.subst_univs_context subst mib.mind_params_ctxt
+
+let inductive_instance mib =
+ if mib.mind_polymorphic then
+ UContext.instance mib.mind_universes
+ else Instance.empty
+
+let inductive_context mib =
+ if mib.mind_polymorphic then
+ mib.mind_universes
+ else UContext.empty
+
+let instantiate_inductive_constraints mib subst =
+ if mib.mind_polymorphic then
+ instantiate_univ_context subst mib.mind_universes
+ else Constraint.empty
+
(************************************************************************)
(* Build the substitution that replaces Rels by the appropriate *)
(* inductives *)
-let ind_subst mind mib =
+let ind_subst mind mib u =
let ntypes = mib.mind_ntypes in
- let make_Ik k = mkInd (mind,ntypes-k-1) in
+ let make_Ik k = mkIndU ((mind,ntypes-k-1),u) in
List.init ntypes make_Ik
(* Instantiate inductives in constructor type *)
-let constructor_instantiate mind mib c =
- let s = ind_subst mind mib in
- substl s c
+let constructor_instantiate mind u subst mib c =
+ let s = ind_subst mind mib u in
+ substl s (subst_univs_constr subst c)
let instantiate_params full t args sign =
let fail () =
@@ -81,13 +108,16 @@ let instantiate_params full t args sign =
let () = if not (List.is_empty rem_args) then fail () in
substl subs ty
-let full_inductive_instantiate mib params sign =
+let full_inductive_instantiate mib u params sign =
let dummy = prop_sort in
let t = mkArity (sign,dummy) in
- fst (destArity (instantiate_params true t params mib.mind_params_ctxt))
+ let subst = make_inductive_subst mib u in
+ let ar = fst (destArity (instantiate_params true t params mib.mind_params_ctxt)) in
+ Vars.subst_univs_context subst ar
-let full_constructor_instantiate ((mind,_),(mib,_),params) =
- let inst_ind = constructor_instantiate mind mib in
+let full_constructor_instantiate ((mind,_),u,(mib,_),params) =
+ let subst = make_inductive_subst mib u in
+ let inst_ind = constructor_instantiate mind u subst mib in
(fun t ->
instantiate_params true (inst_ind t) params mib.mind_params_ctxt)
@@ -119,122 +149,85 @@ Remark: Set (predicative) is encoded as Type(0)
let sort_as_univ = function
| Type u -> u
-| Prop Null -> type0m_univ
-| Prop Pos -> type0_univ
+| Prop Null -> Universe.type0m
+| Prop Pos -> Universe.type0
let cons_subst u su subst =
try
- (u, sup su (List.assoc_f Universe.equal u subst)) ::
- List.remove_assoc_f Universe.equal u subst
+ (u, Universe.sup su (List.assoc_f Universe.eq u subst)) ::
+ List.remove_assoc_f Universe.eq u subst
with Not_found -> (u, su) :: subst
-let actualize_decl_level env lev t =
- let sign,s = dest_arity env t in
- mkArity (sign,lev)
-
-let polymorphism_on_non_applied_parameters = false
-
-(* Bind expected levels of parameters to actual levels *)
-(* Propagate the new levels in the signature *)
-let rec make_subst env = function
- | (_,Some _,_ as t)::sign, exp, args ->
- let ctx,subst = make_subst env (sign, exp, args) in
- t::ctx, subst
- | d::sign, None::exp, args ->
- let args = match args with _::args -> args | [] -> [] in
- let ctx,subst = make_subst env (sign, exp, args) in
- d::ctx, subst
- | d::sign, Some u::exp, a::args ->
- (* We recover the level of the argument, but we don't change the *)
- (* level in the corresponding type in the arity; this level in the *)
- (* arity is a global level which, at typing time, will be enforce *)
- (* to be greater than the level of the argument; this is probably *)
- (* a useless extra constraint *)
- let s = sort_as_univ (snd (dest_arity env (Lazy.force a))) in
- let ctx,subst = make_subst env (sign, exp, args) in
- d::ctx, cons_subst u s subst
- | (na,None,t as d)::sign, Some u::exp, [] ->
- (* No more argument here: we instantiate the type with a fresh level *)
- (* which is first propagated to the corresponding premise in the arity *)
- (* (actualize_decl_level), then to the conclusion of the arity (via *)
- (* the substitution) *)
- let ctx,subst = make_subst env (sign, exp, []) in
- if polymorphism_on_non_applied_parameters then
- let s = fresh_local_univ () in
- let t = actualize_decl_level env (Type s) t in
- (na,None,t)::ctx, cons_subst u s subst
- else
- d::ctx, subst
- | sign, [], _ ->
- (* Uniform parameters are exhausted *)
- sign,[]
- | [], _, _ ->
- assert false
-
-let instantiate_universes env ctx ar argsorts =
- let args = Array.to_list argsorts in
- let ctx,subst = make_subst env (ctx,ar.poly_param_levels,args) in
- let level = subst_large_constraints subst ar.poly_level in
- ctx,
- (* Singleton type not containing types are interpretable in Prop *)
- if is_type0m_univ level then prop_sort
- (* Non singleton type not containing types are interpretable in Set *)
- else if is_type0_univ level then set_sort
- (* This is a Type with constraints *)
- else Type level
-
exception SingletonInductiveBecomesProp of Id.t
-let type_of_inductive_knowing_parameters ?(polyprop=true) env mip paramtyps =
- match mip.mind_arity with
- | Monomorphic s ->
- s.mind_user_arity
- | Polymorphic ar ->
- let ctx = List.rev mip.mind_arity_ctxt in
- let ctx,s = instantiate_universes env ctx ar paramtyps in
- (* The Ocaml extraction cannot handle (yet?) "Prop-polymorphism", i.e.
- the situation where a non-Prop singleton inductive becomes Prop
- when applied to Prop params *)
- if not polyprop && not (is_type0m_univ ar.poly_level) && is_prop_sort s
- then raise (SingletonInductiveBecomesProp mip.mind_typename);
- mkArity (List.rev ctx,s)
-
-(* Type of a (non applied) inductive type *)
-
-let type_of_inductive env (_,mip) =
- type_of_inductive_knowing_parameters env mip [||]
+(* Type of an inductive type *)
+
+let type_of_inductive_gen env ((mib,mip),u) =
+ let subst = make_inductive_subst mib u in
+ (subst_univs_constr subst mip.mind_arity.mind_user_arity, subst)
+
+let type_of_inductive env pind =
+ fst (type_of_inductive_gen env pind)
+
+let constrained_type_of_inductive env ((mib,mip),u as pind) =
+ let ty, subst = type_of_inductive_gen env pind in
+ let cst = instantiate_inductive_constraints mib subst in
+ (ty, cst)
+
+let type_of_inductive_knowing_parameters env ?(polyprop=false) mip args =
+ type_of_inductive env mip
(* The max of an array of universes *)
let cumulate_constructor_univ u = function
| Prop Null -> u
- | Prop Pos -> sup type0_univ u
- | Type u' -> sup u u'
+ | Prop Pos -> Universe.sup Universe.type0 u
+ | Type u' -> Universe.sup u u'
let max_inductive_sort =
- Array.fold_left cumulate_constructor_univ type0m_univ
+ Array.fold_left cumulate_constructor_univ Universe.type0m
(************************************************************************)
(* Type of a constructor *)
-let type_of_constructor cstr (mib,mip) =
+let type_of_constructor_subst cstr u subst (mib,mip) =
let ind = inductive_of_constructor cstr in
let specif = mip.mind_user_lc in
let i = index_of_constructor cstr in
let nconstr = Array.length mip.mind_consnames in
if i > nconstr then error "Not enough constructors in the type.";
- constructor_instantiate (fst ind) mib specif.(i-1)
+ let c = constructor_instantiate (fst ind) u subst mib specif.(i-1) in
+ c
+
+let type_of_constructor_gen (cstr,u) (mib,mip as mspec) =
+ let subst = make_inductive_subst mib u in
+ type_of_constructor_subst cstr u subst mspec, subst
-let arities_of_specif kn (mib,mip) =
+let type_of_constructor cstru mspec =
+ fst (type_of_constructor_gen cstru mspec)
+
+let type_of_constructor_in_ctx cstr (mib,mip as mspec) =
+ let u = UContext.instance mib.mind_universes in
+ let c = type_of_constructor_gen (cstr, u) mspec in
+ (fst c, mib.mind_universes)
+
+let constrained_type_of_constructor (cstr,u as cstru) (mib,mip as ind) =
+ let ty, subst = type_of_constructor_gen cstru ind in
+ let cst = instantiate_inductive_constraints mib subst in
+ (ty, cst)
+
+let arities_of_specif (kn,u) (mib,mip) =
let specif = mip.mind_nf_lc in
- Array.map (constructor_instantiate kn mib) specif
+ let subst = make_inductive_subst mib u in
+ Array.map (constructor_instantiate kn u subst mib) specif
let arities_of_constructors ind specif =
- arities_of_specif (fst ind) specif
+ arities_of_specif (fst (fst ind), snd ind) specif
-let type_of_constructors ind (mib,mip) =
+let type_of_constructors (ind,u) (mib,mip) =
let specif = mip.mind_user_lc in
- Array.map (constructor_instantiate (fst ind) mib) specif
+ let subst = make_inductive_subst mib u in
+ Array.map (constructor_instantiate (fst ind) u subst mib) specif
(************************************************************************)
@@ -255,16 +248,14 @@ let local_rels ctxt =
(* Get type of inductive, with parameters instantiated *)
let inductive_sort_family mip =
- match mip.mind_arity with
- | Monomorphic s -> family_of_sort s.mind_sort
- | Polymorphic _ -> InType
+ family_of_sort mip.mind_arity.mind_sort
let mind_arity mip =
mip.mind_arity_ctxt, inductive_sort_family mip
-let get_instantiated_arity (mib,mip) params =
+let get_instantiated_arity (ind,u) (mib,mip) params =
let sign, s = mind_arity mip in
- full_inductive_instantiate mib params sign, s
+ full_inductive_instantiate mib u params sign, s
let elim_sorts (_,mip) = mip.mind_kelim
@@ -279,7 +270,7 @@ let extended_rel_list n hyps =
let build_dependent_inductive ind (_,mip) params =
let realargs,_ = List.chop mip.mind_nrealargs_ctxt mip.mind_arity_ctxt in
applist
- (mkInd ind,
+ (mkIndU ind,
List.map (lift mip.mind_nrealargs_ctxt) params
@ extended_rel_list 0 realargs)
@@ -295,15 +286,15 @@ let check_allowed_sort ksort specif =
raise (LocalArity (Some(ksort,s,error_elim_explain ksort s)))
let is_correct_arity env c pj ind specif params =
- let arsign,_ = get_instantiated_arity specif params in
- let rec srec env pt ar u =
+ let arsign,_ = get_instantiated_arity ind specif params in
+ 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' ->
- let univ =
+ let () =
try conv env a1 a1'
with NotConvertible -> raise (LocalArity None) in
- srec (push_rel (na1,None,a1) env) t ar' (union_constraints u univ)
+ srec (push_rel (na1,None,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
@@ -311,17 +302,16 @@ let is_correct_arity env c pj ind specif params =
| Sort s -> family_of_sort s
| _ -> raise (LocalArity None) in
let dep_ind = build_dependent_inductive ind specif params in
- let univ =
+ let _ =
try conv env a1 dep_ind
with NotConvertible -> raise (LocalArity None) in
- check_allowed_sort ksort specif;
- union_constraints u univ
+ check_allowed_sort ksort specif
| _, (_,Some _,_ as d)::ar' ->
- srec (push_rel d env) (lift 1 pt') ar' u
+ srec (push_rel d env) (lift 1 pt') ar'
| _ ->
raise (LocalArity None)
in
- try srec env pj.uj_type (List.rev arsign) empty_constraint
+ try srec env pj.uj_type (List.rev arsign)
with LocalArity kinds ->
error_elim_arity env ind (elim_sorts specif) c pj kinds
@@ -331,16 +321,16 @@ let is_correct_arity env c pj ind specif params =
(* [p] is the predicate, [i] is the constructor number (starting from 0),
and [cty] is the type of the constructor (params not instantiated) *)
-let build_branches_type ind (_,mip as specif) params p =
+let build_branches_type (ind,u) (_,mip as specif) params p =
let build_one_branch i cty =
- let typi = full_constructor_instantiate (ind,specif,params) cty in
+ let typi = full_constructor_instantiate (ind,u,specif,params) cty in
let (args,ccl) = decompose_prod_assum typi in
let nargs = rel_context_length args in
let (_,allargs) = decompose_app ccl in
let (lparams,vargs) = List.chop (inductive_params specif) allargs in
let cargs =
let cstr = ith_constructor_of_inductive ind (i+1) in
- let dep_cstr = applist (mkConstruct cstr,lparams@(local_rels args)) in
+ let dep_cstr = applist (mkConstructU (cstr,u),lparams@(local_rels args)) in
vargs @ [dep_cstr] in
let base = beta_appvect (lift nargs p) (Array.of_list cargs) in
it_mkProd_or_LetIn base args in
@@ -348,31 +338,31 @@ let build_branches_type ind (_,mip as specif) params p =
(* [p] is the predicate, [c] is the match object, [realargs] is the
list of real args of the inductive type *)
-let build_case_type n p c realargs =
- whd_betaiota (betazeta_appvect (n+1) p (Array.of_list (realargs@[c])))
+let build_case_type env n p c realargs =
+ whd_betaiota env (betazeta_appvect (n+1) p (Array.of_list (realargs@[c])))
-let type_case_branches env (ind,largs) pj c =
- let specif = lookup_mind_specif env ind in
+let type_case_branches env (pind,largs) pj c =
+ let specif = lookup_mind_specif env (fst pind) in
let nparams = inductive_params specif in
let (params,realargs) = List.chop nparams largs in
let p = pj.uj_val in
- let univ = is_correct_arity env c pj ind specif params in
- let lc = build_branches_type ind specif params p in
- let ty = build_case_type (snd specif).mind_nrealargs_ctxt p c realargs in
- (lc, ty, univ)
+ let () = is_correct_arity env c pj pind specif params in
+ let lc = build_branches_type pind specif params p in
+ let ty = build_case_type env (snd specif).mind_nrealargs_ctxt p c realargs in
+ (lc, ty)
(************************************************************************)
(* Checking the case annotation is relevent *)
-let check_case_info env indsp ci =
+let check_case_info env (indsp,u) ci =
let (mib,mip) = lookup_mind_specif env indsp in
if
not (eq_ind indsp ci.ci_ind) ||
not (Int.equal mib.mind_nparams ci.ci_npar) ||
not (Array.equal Int.equal mip.mind_consnrealdecls ci.ci_cstr_ndecls) ||
not (Array.equal Int.equal mip.mind_consnrealargs ci.ci_cstr_nargs)
- then raise (TypeError(env,WrongCaseInfo(indsp,ci)))
+ then raise (TypeError(env,WrongCaseInfo((indsp,u),ci)))
(************************************************************************)
(************************************************************************)
@@ -450,7 +440,7 @@ type guard_env =
genv : subterm_spec Lazy.t list;
}
-let make_renv env recarg (kn,tyi) =
+let make_renv env recarg ((kn,tyi),u) =
let mib = Environ.lookup_mind kn env in
let mind_recvec =
Array.map (fun mip -> mip.mind_recargs) mib.mind_packets in
@@ -552,7 +542,6 @@ let rec subterm_specif renv stack t =
let f,l = decompose_app (whd_betadeltaiota renv.env t) in
match kind_of_term f with
| Rel k -> subterm_var k renv
-
| Case (ci,p,c,lbr) ->
let stack' = push_stack_closures renv l stack in
if not (check_inductive_codomain renv.env p) then Not_subterm
@@ -581,7 +570,7 @@ let rec subterm_specif renv stack t =
with Not_found -> None in
(match oind with
None -> Not_subterm (* happens if fix is polymorphic *)
- | Some ind ->
+ | Some (ind, _) ->
let nbfix = Array.length typarray in
let recargs = lookup_subterms renv.env ind in
(* pushing the fixpoints *)
@@ -668,7 +657,7 @@ let check_one_fix renv recpos def =
(* if [t] does not make recursive calls, it is guarded: *)
if noccur_with_meta renv.rel_min nfi t then ()
else
- let (f,l) = decompose_app (whd_betaiotazeta t) in
+ let (f,l) = decompose_app (whd_betaiotazeta renv.env t) in
match kind_of_term f with
| Rel p ->
(* Test if [p] is a fixpoint (recursive call) *)
@@ -739,11 +728,11 @@ let check_one_fix renv recpos def =
else check_rec_call renv' [] body)
bodies
- | Const kn ->
+ | Const (kn,u as cu) ->
if evaluable_constant kn renv.env then
try List.iter (check_rec_call renv []) l
with (FixGuardError _ ) ->
- let value = (applist(constant_value renv.env kn, l)) in
+ let value = (applist(constant_value_in renv.env cu, l)) in
check_rec_call renv stack value
else List.iter (check_rec_call renv []) l
@@ -785,6 +774,8 @@ let check_one_fix renv recpos def =
| (Evar _ | Meta _) -> ()
| (App _ | LetIn _ | Cast _) -> assert false (* beta zeta reduction *)
+
+ | Proj (p, c) -> check_rec_call renv [] c
and check_nested_fix_body renv decr recArgsDecrArg body =
if Int.equal decr 0 then
@@ -888,7 +879,7 @@ let check_one_cofix env nbfix def deftype =
else if not(List.for_all (noccur_with_meta n nbfix) args) then
raise (CoFixGuardError (env,NestedRecursiveOccurrences))
- | Construct (_,i as cstr_kn) ->
+ | Construct ((_,i as cstr_kn),u) ->
let lra = vlra.(i-1) in
let mI = inductive_of_constructor cstr_kn in
let (mib,mip) = lookup_mind_specif env mI in
@@ -947,7 +938,7 @@ let check_one_cofix env nbfix def deftype =
| _ -> raise (CoFixGuardError (env,NotGuardedForm t)) in
- let (mind, _) = codomain_is_coind env deftype in
+ let ((mind, _),_) = codomain_is_coind env deftype in
let vlra = lookup_subterms env mind in
check_rec_call env false 1 (dest_subterms vlra) def
diff --git a/kernel/inductive.mli b/kernel/inductive.mli
index d9841085e1..c4a7452f04 100644
--- a/kernel/inductive.mli
+++ b/kernel/inductive.mli
@@ -7,9 +7,9 @@
(************************************************************************)
open Names
-open Univ
open Term
open Context
+open Univ
open Declarations
open Environ
@@ -21,9 +21,9 @@ open Environ
only a coinductive type.
They raise [Not_found] if not convertible to a recursive type. *)
-val find_rectype : env -> types -> inductive * constr list
-val find_inductive : env -> types -> inductive * constr list
-val find_coinductive : env -> types -> inductive * constr list
+val find_rectype : env -> types -> pinductive * constr list
+val find_inductive : env -> types -> pinductive * constr list
+val find_coinductive : env -> types -> pinductive * constr list
type mind_specif = mutual_inductive_body * one_inductive_body
@@ -33,23 +33,38 @@ type mind_specif = mutual_inductive_body * one_inductive_body
val lookup_mind_specif : env -> inductive -> mind_specif
(** {6 Functions to build standard types related to inductive } *)
-val ind_subst : mutual_inductive -> mutual_inductive_body -> constr list
+val ind_subst : mutual_inductive -> mutual_inductive_body -> universe_instance -> constr list
+
+val make_inductive_subst : mutual_inductive_body -> universe_instance -> universe_subst
+
+val inductive_instance : mutual_inductive_body -> universe_instance
+val inductive_context : mutual_inductive_body -> universe_context
+val inductive_params_ctxt : mutual_inductive_body puniverses -> rel_context
+
+val instantiate_inductive_constraints : mutual_inductive_body -> universe_subst -> constraints
-val type_of_inductive : env -> mind_specif -> types
+val constrained_type_of_inductive : env -> mind_specif puniverses -> types constrained
+
+val type_of_inductive : env -> mind_specif puniverses -> types
+
+val type_of_inductive_knowing_parameters : env -> ?polyprop:bool -> mind_specif puniverses -> types Lazy.t array -> types
val elim_sorts : mind_specif -> sorts_family list
(** Return type as quoted by the user *)
-val type_of_constructor : constructor -> mind_specif -> types
+
+val constrained_type_of_constructor : pconstructor -> mind_specif -> types constrained
+val type_of_constructor : pconstructor -> mind_specif -> types
+val type_of_constructor_in_ctx : constructor -> mind_specif -> types in_universe_context
(** Return constructor types in normal form *)
-val arities_of_constructors : inductive -> mind_specif -> types array
+val arities_of_constructors : pinductive -> mind_specif -> types array
(** Return constructor types in user form *)
-val type_of_constructors : inductive -> mind_specif -> types array
+val type_of_constructors : pinductive -> mind_specif -> types array
(** Transforms inductive specification into types (in nf) *)
-val arities_of_specif : mutual_inductive -> mind_specif -> types array
+val arities_of_specif : mutual_inductive puniverses -> mind_specif -> types array
val inductive_params : mind_specif -> int
@@ -61,11 +76,11 @@ val inductive_params : mind_specif -> int
the universe constraints generated.
*)
val type_case_branches :
- env -> inductive * constr list -> unsafe_judgment -> constr
- -> types array * types * constraints
+ env -> pinductive * constr list -> unsafe_judgment -> constr
+ -> types array * types
val build_branches_type :
- inductive -> mutual_inductive_body * one_inductive_body ->
+ pinductive -> mutual_inductive_body * one_inductive_body ->
constr list -> constr -> types array
(** Return the arity of an inductive type *)
@@ -75,7 +90,7 @@ val inductive_sort_family : one_inductive_body -> sorts_family
(** Check a [case_info] actually correspond to a Case expression on the
given inductive type. *)
-val check_case_info : env -> inductive -> case_info -> unit
+val check_case_info : env -> pinductive -> case_info -> unit
(** {6 Guard conditions for fix and cofix-points. } *)
val check_fix : env -> fixpoint -> unit
@@ -92,14 +107,8 @@ val check_cofix : env -> cofixpoint -> unit
exception SingletonInductiveBecomesProp of Id.t
-val type_of_inductive_knowing_parameters : ?polyprop:bool ->
- env -> one_inductive_body -> types Lazy.t array -> types
-
val max_inductive_sort : sorts array -> universe
-val instantiate_universes : env -> rel_context ->
- polymorphic_arity -> types Lazy.t array -> rel_context * sorts
-
(** {6 Debug} *)
type size = Large | Strict
diff --git a/kernel/kernel.mllib b/kernel/kernel.mllib
index 0d0adf9a7d..29fe887d75 100644
--- a/kernel/kernel.mllib
+++ b/kernel/kernel.mllib
@@ -32,6 +32,7 @@ Type_errors
Modops
Inductive
Typeops
+Fast_typeops
Indtypes
Cooking
Term_typing
@@ -39,7 +40,6 @@ Subtyping
Mod_typing
Nativelibrary
Safe_typing
-
Vm
Csymtable
Vconv
diff --git a/kernel/mod_subst.ml b/kernel/mod_subst.ml
index 9589c0656a..cfe46152ef 100644
--- a/kernel/mod_subst.ml
+++ b/kernel/mod_subst.ml
@@ -271,7 +271,7 @@ let progress f x ~orelse =
let y = f x in
if y != x then y else orelse
-let subst_ind sub mind =
+let subst_mind sub mind =
let mpu,dir,l = MutInd.repr3 mind in
let mpc = KerName.modpath (MutInd.canonical mind) in
try
@@ -284,7 +284,14 @@ let subst_ind sub mind =
MutInd.make knu knc'
with No_subst -> mind
-let subst_con0 sub cst =
+let subst_ind sub (ind,i as indi) =
+ let ind' = subst_mind sub ind in
+ if ind' == ind then indi else ind',i
+
+let subst_pind sub (ind,u) =
+ (subst_ind sub ind, u)
+
+let subst_con0 sub (cst,u) =
let mpu,dir,l = Constant.repr3 cst in
let mpc = KerName.modpath (Constant.canonical cst) in
let mpu,mpc,resolve,user = subst_dual_mp sub mpu mpc in
@@ -299,11 +306,28 @@ let subst_con0 sub cst =
progress (kn_of_delta resolve) (if user then knu else knc) ~orelse:knc
in
let cst' = Constant.make knu knc' in
- cst', mkConst cst'
+ cst', mkConstU (cst',u)
let subst_con sub cst =
try subst_con0 sub cst
- with No_subst -> cst, mkConst cst
+ with No_subst -> fst cst, mkConstU cst
+
+let subst_con_kn sub con =
+ subst_con sub (con,Univ.Instance.empty)
+
+let subst_pcon sub (con,u as pcon) =
+ try let con', can = subst_con0 sub pcon in
+ con',u
+ with No_subst -> pcon
+
+let subst_pcon_term sub (con,u as pcon) =
+ try let con', can = subst_con0 sub pcon in
+ (con',u), can
+ with No_subst -> pcon, mkConstU pcon
+
+let subst_constant sub con =
+ try fst (subst_con0 sub (con,Univ.Instance.empty))
+ with No_subst -> con
(* Here the semantics is completely unclear.
What does "Hint Unfold t" means when "t" is a parameter?
@@ -312,18 +336,25 @@ let subst_con sub cst =
interpretation (i.e. an evaluable reference is never expanded). *)
let subst_evaluable_reference subst = function
| EvalVarRef id -> EvalVarRef id
- | EvalConstRef kn -> EvalConstRef (fst (subst_con subst kn))
+ | EvalConstRef kn -> EvalConstRef (subst_constant subst kn)
let rec map_kn f f' c =
let func = map_kn f f' in
match kind_of_term c with
| Const kn -> (try snd (f' kn) with No_subst -> c)
- | Ind (kn,i) ->
+ | Proj (kn,t) ->
+ let kn' = try fst (f' (kn,Univ.Instance.empty))
+ with No_subst -> kn
+ in
+ let t' = func t in
+ if kn' == kn && t' == t then c
+ else mkProj (kn', t')
+ | Ind ((kn,i),u) ->
let kn' = f kn in
- if kn'==kn then c else mkInd (kn',i)
- | Construct ((kn,i),j) ->
+ if kn'==kn then c else mkIndU ((kn',i),u)
+ | Construct (((kn,i),j),u) ->
let kn' = f kn in
- if kn'==kn then c else mkConstruct ((kn',i),j)
+ if kn'==kn then c else mkConstructU (((kn',i),j),u)
| Case (ci,p,ct,l) ->
let ci_ind =
let (kn,i) = ci.ci_ind in
@@ -382,7 +413,7 @@ let rec map_kn f f' c =
let subst_mps sub c =
if is_empty_subst sub then c
- else map_kn (subst_ind sub) (subst_con0 sub) c
+ else map_kn (subst_mind sub) (subst_con0 sub) c
let rec replace_mp_in_mp mpfrom mpto mp =
match mp with
diff --git a/kernel/mod_subst.mli b/kernel/mod_subst.mli
index 34f10b31ab..5a913a9067 100644
--- a/kernel/mod_subst.mli
+++ b/kernel/mod_subst.mli
@@ -118,15 +118,32 @@ val debug_pr_delta : delta_resolver -> Pp.std_ppcmds
val subst_mp :
substitution -> module_path -> module_path
-val subst_ind :
+val subst_mind :
substitution -> mutual_inductive -> mutual_inductive
+val subst_ind :
+ substitution -> inductive -> inductive
+
+val subst_pind : substitution -> pinductive -> pinductive
+
val subst_kn :
substitution -> kernel_name -> kernel_name
val subst_con :
+ substitution -> pconstant -> constant * constr
+
+val subst_pcon :
+ substitution -> pconstant -> pconstant
+
+val subst_pcon_term :
+ substitution -> pconstant -> pconstant * constr
+
+val subst_con_kn :
substitution -> constant -> constant * constr
+val subst_constant :
+ substitution -> constant -> constant
+
(** Here the semantics is completely unclear.
What does "Hint Unfold t" means when "t" is a parameter?
Does the user mean "Unfold X.t" or does she mean "Unfold y"
diff --git a/kernel/mod_typing.ml b/kernel/mod_typing.ml
index 6c0f1b060c..b20fe96710 100644
--- a/kernel/mod_typing.ml
+++ b/kernel/mod_typing.ml
@@ -52,7 +52,7 @@ let rec rebuild_mp mp l =
| []-> mp
| i::r -> rebuild_mp (MPdot(mp,Label.of_id i)) r
-let (+++) = Univ.union_constraints
+let (+++) = Univ.Constraint.union
let rec check_with_def env struc (idl,c) mp equiv =
let lab,idl = match idl with
@@ -72,24 +72,31 @@ let rec check_with_def env struc (idl,c) mp equiv =
(* In the spirit of subtyping.check_constant, we accept
any implementations of parameters and opaques terms,
as long as they have the right type *)
+ let env' = Environ.add_constraints
+ (Univ.UContext.constraints (Future.force cb.const_universes)) env' in
let c',cst = match cb.const_body with
| Undef _ | OpaqueDef _ ->
- let j,cst1 = Typeops.infer env' c in
+ let j = Typeops.infer env' c in
let typ = Typeops.type_of_constant_type env' cb.const_type in
- let cst2 = Reduction.conv_leq env' j.uj_type typ in
- let cst = cb.const_constraints +++ cst1 +++ cst2 in
- j.uj_val, cst
+ let cst = Reduction.infer_conv_leq env' (Environ.universes env')
+ j.uj_type typ in
+ j.uj_val,cst
| Def cs ->
- let cst1 = Reduction.conv env' c (Mod_subst.force_constr cs) in
- let cst = cb.const_constraints +++ cst1 in
- c, cst
+ let cst = Reduction.infer_conv env' (Environ.universes env') c
+ (Mod_subst.force_constr cs) in
+ let cst =
+ if cb.const_polymorphic then cst
+ else (*FIXME MS: computed above *)
+ (Univ.UContext.constraints (Future.force cb.const_universes)) +++ cst
+ in
+ c, cst
in
let def = Def (Mod_subst.from_val c') in
let cb' =
{ cb with
const_body = def;
- const_body_code = Cemitcodes.from_val (compile_constant_body env' def);
- const_constraints = cst }
+ const_body_code = Cemitcodes.from_val (compile_constant_body env' def) }
+ (* const_universes = Future.from_val cst } *)
in
before@(lab,SFBconst(cb'))::after, c', cst
else
@@ -185,7 +192,7 @@ let rec check_with_mod env struc (idl,mp1) mp equiv =
| Algebraic (NoFunctor (MEident mp0)) ->
let mpnew = rebuild_mp mp0 idl in
check_modpath_equiv env' mpnew mp;
- before@(lab,spec)::after, equiv, Univ.empty_constraint
+ before@(lab,spec)::after, equiv, Univ.Constraint.empty
| _ -> error_generative_module_expected lab
end
with
@@ -229,7 +236,7 @@ let rec translate_mse env mpo inl = function
let mtb = lookup_modtype mp1 env in
mtb.typ_expr, mtb.typ_delta
in
- sign,Some (MEident mp1),reso,Univ.empty_constraint
+ sign,Some (MEident mp1),reso,Univ.Constraint.empty
|MEapply (fe,mp1) ->
translate_apply env inl (translate_mse env mpo inl fe) mp1 (mk_alg_app mpo)
|MEwith(me, with_decl) ->
@@ -297,7 +304,7 @@ let finalize_module env mp (sign,alg,reso,cst) restype = match restype with
typ_mp = mp;
typ_expr = sign;
typ_expr_alg = None;
- typ_constraints = Univ.empty_constraint;
+ typ_constraints = Univ.Constraint.empty;
typ_delta = reso } in
let cst' = Subtyping.check_subtypes env auto_mtb res_mtb in
let impl = match alg with Some e -> Algebraic e | None -> Struct sign in
@@ -322,7 +329,7 @@ let rec translate_mse_incl env mp inl = function
|MEident mp1 ->
let mb = strengthen_and_subst_mb (lookup_module mp1 env) mp true in
let sign = clean_bounded_mod_expr mb.mod_type in
- sign,None,mb.mod_delta,Univ.empty_constraint
+ sign,None,mb.mod_delta,Univ.Constraint.empty
|MEapply (fe,arg) ->
let ftrans = translate_mse_incl env mp inl fe in
translate_apply env inl ftrans arg (fun _ _ -> None)
diff --git a/kernel/modops.ml b/kernel/modops.ml
index 6d0e919f8f..093ee70247 100644
--- a/kernel/modops.ml
+++ b/kernel/modops.ml
@@ -210,7 +210,7 @@ and subst_structure sub do_delta sign =
let cb' = subst_const_body sub cb in
if cb==cb' then orig else (l,SFBconst cb')
|SFBmind mib ->
- let mib' = subst_mind sub mib in
+ let mib' = subst_mind_body sub mib in
if mib==mib' then orig else (l,SFBmind mib')
|SFBmodule mb ->
let mb' = subst_module sub do_delta mb in
@@ -460,7 +460,7 @@ and strengthen_and_subst_struct str subst mp_from mp_to alias incl reso =
because reso' contains mp_to maps to reso(mp_from) *)
reso', item'::rest'
| (l,SFBmind mib) :: rest ->
- let item' = l,SFBmind (subst_mind subst mib) in
+ let item' = l,SFBmind (subst_mind_body subst mib) in
let reso',rest' =
strengthen_and_subst_struct rest subst mp_from mp_to alias incl reso
in
diff --git a/kernel/names.ml b/kernel/names.ml
index ef0e812ed7..c76d95937f 100644
--- a/kernel/names.ml
+++ b/kernel/names.ml
@@ -309,6 +309,11 @@ module ModPath = struct
let initial = MPfile DirPath.initial
+ let rec dp = function
+ | MPfile sl -> sl
+ | MPbound (_,_,dp) -> dp
+ | MPdot (mp,l) -> dp mp
+
module Self_Hashcons = struct
type t = module_path
type u = (DirPath.t -> DirPath.t) * (MBId.t -> MBId.t) *
@@ -424,7 +429,6 @@ module KerName = struct
let hcons =
Hashcons.simple_hcons HashKN.generate
(ModPath.hcons,DirPath.hcons,String.hcons)
-
end
module KNmap = HMap.Make(KerName)
@@ -567,6 +571,7 @@ let constr_modpath (ind,_) = ind_modpath ind
let ith_mutual_inductive (mind, _) i = (mind, i)
let ith_constructor_of_inductive ind i = (ind, i)
+let ith_constructor_of_pinductive (ind,u) i = ((ind,i),u)
let inductive_of_constructor (ind, i) = ind
let index_of_constructor (ind, i) = i
@@ -663,8 +668,7 @@ let hcons_mind = Hashcons.simple_hcons MutInd.HashKP.generate KerName.hcons
let hcons_ind = Hashcons.simple_hcons Hind.generate hcons_mind
let hcons_construct = Hashcons.simple_hcons Hconstruct.generate hcons_ind
-
-(*******)
+(*****************)
type transparent_state = Id.Pred.t * Cpred.t
@@ -674,25 +678,26 @@ let var_full_transparent_state = (Id.Pred.full, Cpred.empty)
let cst_full_transparent_state = (Id.Pred.empty, Cpred.full)
type 'a tableKey =
- | ConstKey of Constant.t
+ | ConstKey of 'a
| VarKey of Id.t
- | RelKey of 'a
-
+ | RelKey of Int.t
type inv_rel_key = int (* index in the [rel_context] part of environment
starting by the end, {\em inverse}
of de Bruijn indice *)
-type id_key = inv_rel_key tableKey
+type id_key = Constant.t tableKey
-let eq_id_key ik1 ik2 =
+let eq_table_key f ik1 ik2 =
if ik1 == ik2 then true
else match ik1,ik2 with
- | ConstKey c1, ConstKey c2 -> Constant.UserOrd.equal c1 c2
+ | ConstKey c1, ConstKey c2 -> f c1 c2
| VarKey id1, VarKey id2 -> Id.equal id1 id2
| RelKey k1, RelKey k2 -> Int.equal k1 k2
| _ -> false
+let eq_id_key = eq_table_key Constant.UserOrd.equal
+
let eq_con_chk = Constant.UserOrd.equal
let eq_mind_chk = MutInd.UserOrd.equal
let eq_ind_chk (kn1,i1) (kn2,i2) = Int.equal i1 i2 && eq_mind_chk kn1 kn2
@@ -777,6 +782,7 @@ let kn_ord = KerName.compare
(** Compatibility layer for [Constant] *)
type constant = Constant.t
+type projection = constant
let constant_of_kn = Constant.make1
let constant_of_kn_equiv = Constant.make
@@ -787,6 +793,7 @@ let user_con = Constant.user
let con_label = Constant.label
let con_modpath = Constant.modpath
let eq_constant = Constant.equal
+let eq_constant_key = Constant.UserOrd.equal
let con_ord = Constant.CanOrd.compare
let con_user_ord = Constant.UserOrd.compare
let string_of_con = Constant.to_string
diff --git a/kernel/names.mli b/kernel/names.mli
index db973ed3a4..49a838ae54 100644
--- a/kernel/names.mli
+++ b/kernel/names.mli
@@ -216,6 +216,8 @@ sig
val initial : t
(** Name of the toplevel structure ([= MPfile initial_dir]) *)
+ val dp : t -> DirPath.t
+
end
module MPset : Set.S with type elt = ModPath.t
@@ -440,10 +442,11 @@ val hcons_construct : constructor -> constructor
(******)
type 'a tableKey =
- | ConstKey of Constant.t
+ | ConstKey of 'a
| VarKey of Id.t
- | RelKey of 'a
+ | RelKey of Int.t
+(** Sets of names *)
type transparent_state = Id.Pred.t * Cpred.t
val empty_transparent_state : transparent_state
@@ -455,8 +458,10 @@ type inv_rel_key = int (** index in the [rel_context] part of environment
starting by the end, {e inverse}
of de Bruijn indice *)
-type id_key = inv_rel_key tableKey
+type id_key = Constant.t tableKey
+val eq_table_key : ('a -> 'a -> bool) -> 'a tableKey -> 'a tableKey -> bool
+val eq_constant_key : Constant.t -> Constant.t -> bool
val eq_id_key : id_key -> id_key -> bool
(** equalities on constant and inductive names (for the checker) *)
@@ -629,6 +634,8 @@ val kn_ord : kernel_name -> kernel_name -> int
type constant = Constant.t
(** @deprecated Alias type *)
+type projection = constant
+
val constant_of_kn_equiv : KerName.t -> KerName.t -> constant
(** @deprecated Same as [Constant.make] *)
diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml
index 1f6effba63..bd659a471f 100644
--- a/kernel/nativecode.ml
+++ b/kernel/nativecode.ml
@@ -48,6 +48,7 @@ type gname =
| Gind of string * inductive (* prefix, inductive name *)
| Gconstruct of string * constructor (* prefix, constructor name *)
| Gconstant of string * constant (* prefix, constant name *)
+ | Gproj of string * constant (* prefix, constant name *)
| Gcase of label option * int
| Gpred of label option * int
| Gfixtype of label option * int
@@ -95,6 +96,7 @@ let gname_hash gn = match gn with
| Ginternal s -> combinesmall 9 (String.hash s)
| Grel i -> combinesmall 10 (Int.hash i)
| Gnamed id -> combinesmall 11 (Id.hash id)
+| Gproj (s, p) -> combinesmall 12 (combine (String.hash s) (Constant.hash p))
let case_ctr = ref (-1)
@@ -253,6 +255,7 @@ type primitive =
| Mk_cofix of int
| Mk_rel of int
| Mk_var of identifier
+ | Mk_proj
| Is_accu
| Is_int
| Cast_accu
@@ -298,6 +301,8 @@ let eq_primitive p1 p2 =
| Force_cofix, Force_cofix -> true
| Mk_meta, Mk_meta -> true
| Mk_evar, Mk_evar -> true
+ | Mk_proj, Mk_proj -> true
+
| _ -> false
let primitive_hash = function
@@ -344,6 +349,7 @@ let primitive_hash = function
| Coq_primitive (prim, None) -> combinesmall 36 (Primitives.hash prim)
| Coq_primitive (prim, Some (prefix,kn)) ->
combinesmall 37 (combine3 (String.hash prefix) (Constant.hash kn) (Primitives.hash prim))
+ | Mk_proj -> 38
type mllambda =
| MLlocal of lname
@@ -1002,6 +1008,7 @@ let compile_prim decl cond paux =
| Lapp(f,args) ->
MLapp(ml_of_lam env l f, Array.map (ml_of_lam env l) args)
| Lconst (prefix,c) -> MLglobal(Gconstant (prefix,c))
+ | Lproj (prefix,c) -> MLglobal(Gproj (prefix,c))
| Lprim _ ->
let decl,cond,paux = extract_prim (ml_of_lam env l) t in
compile_prim decl cond paux
@@ -1461,6 +1468,8 @@ let string_of_gname g =
Format.sprintf "%sconstruct_%s_%i_%i" prefix (string_of_mind mind) i (j-1)
| Gconstant (prefix, c) ->
Format.sprintf "%sconst_%s" prefix (string_of_con c)
+ | Gproj (prefix, c) ->
+ Format.sprintf "%sproj_%s" prefix (string_of_con c)
| Gcase (l,i) ->
Format.sprintf "case_%s_%i" (string_of_label_def l) i
| Gpred (l,i) ->
@@ -1518,12 +1527,12 @@ let pp_mllam fmt l =
| MLif(t,l1,l2) ->
Format.fprintf fmt "@[(if %a then@\n %a@\nelse@\n %a)@]"
pp_mllam t pp_mllam l1 pp_mllam l2
- | MLmatch (asw, c, accu_br, br) ->
- let mind,i = asw.asw_ind in
- let prefix = asw.asw_prefix in
- let accu = Format.sprintf "%sAccu_%s_%i" prefix (string_of_mind mind) i in
- Format.fprintf fmt
- "@[begin match Obj.magic (%a) with@\n| %s _ ->@\n %a@\n%aend@]"
+ | MLmatch (annot, c, accu_br, br) ->
+ let mind,i = annot.asw_ind in
+ let prefix = annot.asw_prefix in
+ let accu = Format.sprintf "%sAccu_%s_%i" prefix (string_of_mind mind) i in
+ Format.fprintf fmt
+ "@[begin match Obj.magic (%a) with@\n| %s _ ->@\n %a@\n%aend@]"
pp_mllam c accu pp_mllam accu_br (pp_branches prefix) br
| MLconstruct(prefix,c,args) ->
@@ -1626,6 +1635,7 @@ let pp_mllam fmt l =
| Mk_rel i -> Format.fprintf fmt "mk_rel_accu %i" i
| Mk_var id ->
Format.fprintf fmt "mk_var_accu (Names.id_of_string \"%s\")" (string_of_id id)
+ | Mk_proj -> Format.fprintf fmt "mk_proj_accu"
| Is_accu -> Format.fprintf fmt "is_accu"
| Is_int -> Format.fprintf fmt "is_int"
| Cast_accu -> Format.fprintf fmt "cast_accu"
@@ -1758,9 +1768,11 @@ and compile_named env sigma auxdefs id =
| None ->
Glet(Gnamed id, MLprimitive (Mk_var id))::auxdefs
-let compile_constant env sigma prefix ~interactive con body =
- match body with
- | Def t ->
+let compile_constant env sigma prefix ~interactive con cb =
+ match cb.const_proj with
+ | None ->
+ begin match cb.const_body with
+ | Def t ->
let t = Mod_subst.force_constr t in
let code = lambda_of_constr env sigma t in
if !Flags.debug then Pp.msg_debug (Pp.str "Generated lambda code");
@@ -1778,11 +1790,42 @@ let compile_constant env sigma prefix ~interactive con body =
in
if !Flags.debug then Pp.msg_debug (Pp.str "Optimized mllambda code");
code, name
- | _ ->
+ | _ ->
+ let i = push_symbol (SymbConst con) in
+ [Glet(Gconstant ("",con), MLapp (MLprimitive Mk_const, [|get_const_code i|]))],
+ if interactive then LinkedInteractive prefix
+ else Linked prefix
+ end
+ | Some pb ->
+ let mind = pb.proj_ind in
+ let ind = (mind,0) in
+ let mib = lookup_mind mind env in
+ let oib = mib.mind_packets.(0) in
+ let tbl = oib.mind_reloc_tbl in
+ (* Building info *)
+ let prefix = get_mind_prefix env mind in
+ let ci = { ci_ind = ind; ci_npar = mib.mind_nparams;
+ ci_cstr_nargs = [|0|];
+ ci_cstr_ndecls = [||] (*FIXME*);
+ ci_pp_info = { ind_nargs = 0; style = RegularStyle } } in
+ let asw = { asw_ind = ind; asw_prefix = prefix; asw_ci = ci;
+ asw_reloc = tbl; asw_finite = true } in
+ let c_uid = fresh_lname Anonymous in
+ let _, arity = tbl.(0) in
+ let ci_uid = fresh_lname Anonymous in
+ let cargs = Array.init arity
+ (fun i -> if Int.equal i pb.proj_arg then Some ci_uid else None)
+ in
let i = push_symbol (SymbConst con) in
- [Glet(Gconstant ("",con), MLapp (MLprimitive Mk_const, [|get_const_code i|]))],
- if interactive then LinkedInteractive prefix
- else Linked prefix
+ let accu = MLapp (MLprimitive Cast_accu, [|MLlocal c_uid|]) in
+ let accu_br = MLapp (MLprimitive Mk_proj, [|get_const_code i;accu|]) in
+ let code = MLmatch(asw,MLlocal c_uid,accu_br,[|[((ind,1),cargs)],MLlocal ci_uid|]) in
+ let gn = Gproj ("",con) in
+ let fargs = Array.init (pb.proj_npars + 1) (fun _ -> fresh_lname Anonymous) in
+ let arg = fargs.(pb.proj_npars) in
+ Glet(Gconstant ("",con), mkMLlam fargs (MLapp (MLglobal gn, [|MLlocal
+ arg|])))::
+ [Glet(gn, mkMLlam [|c_uid|] code)], Linked prefix
let loaded_native_files = ref ([] : string list)
@@ -1858,8 +1901,8 @@ let compile_mind_deps env prefix ~interactive
reverse order, as well as linking information updates *)
let rec compile_deps env sigma prefix ~interactive init t =
match kind_of_term t with
- | Ind (mind,_) -> compile_mind_deps env prefix ~interactive init mind
- | Const c ->
+ | Ind ((mind,_),u) -> compile_mind_deps env prefix ~interactive init mind
+ | Const (c,u) ->
let c = get_allias env c in
let cb,(nameref,_) = lookup_constant_key c env in
let (_, (_, const_updates)) = init in
@@ -1873,12 +1916,14 @@ let rec compile_deps env sigma prefix ~interactive init t =
| _ -> init
in
let code, name =
- compile_constant env sigma prefix ~interactive c cb.const_body
+ compile_constant env sigma prefix ~interactive c cb
in
let comp_stack = code@comp_stack in
let const_updates = Cmap_env.add c (nameref, name) const_updates in
comp_stack, (mind_updates, const_updates)
- | Construct ((mind,_),_) -> compile_mind_deps env prefix ~interactive init mind
+ | Construct (((mind,_),_),u) -> compile_mind_deps env prefix ~interactive init mind
+ | Proj (p,c) ->
+ compile_deps env sigma prefix ~interactive init (mkApp (mkConst p, [|c|]))
| Case (ci, p, c, ac) ->
let mind = fst ci.ci_ind in
let init = compile_mind_deps env prefix ~interactive init mind in
@@ -1888,7 +1933,7 @@ let rec compile_deps env sigma prefix ~interactive init t =
let compile_constant_field env prefix con acc cb =
let (gl, _) =
compile_constant ~interactive:false env empty_evars prefix
- con cb.const_body
+ con cb
in
gl@acc
diff --git a/kernel/nativeconv.ml b/kernel/nativeconv.ml
index 82786df64c..766e6513c7 100644
--- a/kernel/nativeconv.ml
+++ b/kernel/nativeconv.ml
@@ -70,7 +70,7 @@ and conv_atom pb lvl a1 a2 cu =
if not (eq_constant c1 c2) then raise NotConvertible;
cu
| Asort s1, Asort s2 ->
- sort_cmp pb s1 s2 cu
+ ignore(sort_cmp_universes pb s1 s2 (cu, None)); cu
| Avar id1, Avar id2 ->
if not (Id.equal id1 id2) then raise NotConvertible;
cu
@@ -131,9 +131,9 @@ let native_conv pb sigma env t1 t2 =
vm_conv pb env t1 t2
end
else
- let env = Environ.pre_env env in
+ let penv = Environ.pre_env env in
let ml_filename, prefix = get_ml_filename () in
- let code, upds = mk_conv_code env sigma prefix t1 t2 in
+ let code, upds = mk_conv_code penv sigma prefix t1 t2 in
match compile ml_filename code with
| (0,fn) ->
begin
@@ -144,7 +144,7 @@ let native_conv pb sigma env t1 t2 =
let time_info = Format.sprintf "Evaluation done in %.5f@." (t1 -. t0) in
if !Flags.debug then Pp.msg_debug (Pp.str time_info);
(* TODO change 0 when we can have deBruijn *)
- conv_val pb 0 !rt1 !rt2 empty_constraint
+ ignore(conv_val pb 0 !rt1 !rt2 (Environ.universes env))
end
| _ -> anomaly (Pp.str "Compilation failure")
diff --git a/kernel/nativeinstr.mli b/kernel/nativeinstr.mli
index 13d61841f9..7d1bf0d192 100644
--- a/kernel/nativeinstr.mli
+++ b/kernel/nativeinstr.mli
@@ -29,6 +29,7 @@ and lambda =
| Llet of name * lambda * lambda
| Lapp of lambda * lambda array
| Lconst of prefix * constant
+ | Lproj of prefix * constant (* prefix, projection name *)
| Lprim of prefix * constant * Primitives.t * lambda array
| Lcase of annot_sw * lambda * lambda * lam_branches
(* annotations, term being matched, accu, branches *)
diff --git a/kernel/nativelambda.ml b/kernel/nativelambda.ml
index 8ea28ddff1..16ca444e3a 100644
--- a/kernel/nativelambda.ml
+++ b/kernel/nativelambda.ml
@@ -79,12 +79,12 @@ let get_const_prefix env c =
| NotLinked -> ""
| Linked s -> s
| LinkedInteractive s -> s
-
+
(* A generic map function *)
let map_lam_with_binders g f n lam =
match lam with
- | Lrel _ | Lvar _ | Lconst _ | Luint _ | Lval _ | Lsort _ | Lind _
+ | Lrel _ | Lvar _ | Lconst _ | Lproj _ | Luint _ | Lval _ | Lsort _ | Lind _
| Lconstruct _ | Llazy | Lforce | Lmeta _ | Levar _ -> lam
| Lprod(dom,codom) ->
let dom' = f n dom in
@@ -183,7 +183,7 @@ let lam_subst_args subst args =
let can_subst lam =
match lam with
- | Lrel _ | Lvar _ | Lconst _ | Lval _ | Lsort _ | Lind _ | Llam _
+ | Lrel _ | Lvar _ | Lconst _ | Lproj _ | Lval _ | Lsort _ | Lind _ | Llam _
| Lconstruct _ | Lmeta _ | Levar _ -> true
| _ -> false
@@ -257,6 +257,7 @@ and simplify_app substf f substa args =
let args = Array.append
(lam_subst_args substf args') (lam_subst_args substa args) in
simplify_app substf f subst_id args
+ (* TODO | Lproj -> simplify if the argument is known or a known global *)
| _ -> mkLapp (simplify substf f) (simplify_args substa args)
and simplify_args subst args = Array.smartmap (simplify subst) args
@@ -290,7 +291,7 @@ let rec occurence k kind lam =
if Int.equal n k then
if kind then false else raise Not_found
else kind
- | Lvar _ | Lconst _ | Luint _ | Lval _ | Lsort _ | Lind _
+ | Lvar _ | Lconst _ | Lproj _ | Luint _ | Lval _ | Lsort _ | Lind _
| Lconstruct _ | Llazy | Lforce | Lmeta _ | Levar _ -> kind
| Lprod(dom, codom) ->
occurence k (occurence k kind dom) codom
@@ -504,7 +505,7 @@ let is_lazy prefix t =
match kind_of_term t with
| App (f,args) ->
begin match kind_of_term f with
- | Construct c ->
+ | Construct (c,_) ->
let entry = mkInd (fst c) in
(try
let _ =
@@ -552,7 +553,7 @@ let rec lambda_of_constr env sigma c =
| Sort s -> Lsort s
- | Ind ind ->
+ | Ind (ind,u) ->
let prefix = get_mind_prefix !global_env (fst ind) in
Lind (prefix, ind)
@@ -584,6 +585,9 @@ let rec lambda_of_constr env sigma c =
| Construct _ -> lambda_of_app env sigma c empty_args
+ | Proj (p, c) ->
+ mkLapp (Lproj (get_const_prefix !global_env p, p)) [|lambda_of_constr env sigma c|]
+
| Case(ci,t,a,branches) ->
let (mind,i as ind) = ci.ci_ind in
let mib = lookup_mind mind !global_env in
@@ -642,7 +646,7 @@ let rec lambda_of_constr env sigma c =
and lambda_of_app env sigma f args =
match kind_of_term f with
- | Const kn ->
+ | Const (kn,u) ->
let kn = get_allias !global_env kn in
let cb = lookup_constant kn !global_env in
(try
@@ -654,7 +658,7 @@ and lambda_of_app env sigma f args =
f args
with Not_found ->
begin match cb.const_body with
- | Def csubst ->
+ | Def csubst -> (* TODO optimize if f is a proj and argument is known *)
if cb.const_inline_code then
lambda_of_app env sigma (Mod_subst.force_constr csubst) args
else
@@ -669,7 +673,7 @@ and lambda_of_app env sigma f args =
let prefix = get_const_prefix !global_env kn in
mkLapp (Lconst (prefix, kn)) (lambda_of_args env sigma 0 args)
end)
- | Construct c ->
+ | Construct (c,u) ->
let tag, nparams, arity = Renv.get_construct_info env c in
let expected = nparams + arity in
let nargs = Array.length args in
@@ -737,7 +741,7 @@ let compile_static_int31 fc args =
Luint (UintVal
(Uint31.of_int (Array.fold_left
(fun temp_i -> fun t -> match kind_of_term t with
- | Construct (_,d) -> 2*temp_i+d-1
+ | Construct ((_,d),_) -> 2*temp_i+d-1
| _ -> raise NotClosed)
0 args)))
diff --git a/kernel/nativelambda.mli b/kernel/nativelambda.mli
index a2763626c0..33a0dacf65 100644
--- a/kernel/nativelambda.mli
+++ b/kernel/nativelambda.mli
@@ -12,7 +12,6 @@ open Nativevalues
open Nativeinstr
(** This file defines the lambda code generation phase of the native compiler *)
-
type evars =
{ evars_val : existential -> constr option;
evars_typ : existential -> types;
@@ -26,6 +25,8 @@ val decompose_Llam_Llet : lambda -> (Names.name * lambda option) array * lambda
val is_lazy : prefix -> constr -> bool
val mk_lazy : lambda -> lambda
+val get_mind_prefix : env -> mutual_inductive -> string
+
val get_allias : env -> constant -> constant
val lambda_of_constr : env -> evars -> Constr.constr -> lambda
diff --git a/kernel/nativevalues.ml b/kernel/nativevalues.ml
index 043f06e26e..d88d5d25d3 100644
--- a/kernel/nativevalues.ml
+++ b/kernel/nativevalues.ml
@@ -60,6 +60,7 @@ type atom =
| Aprod of name * t * (t -> t)
| Ameta of metavariable * t
| Aevar of existential * t
+ | Aproj of constant * accumulator
let accumulate_tag = 0
@@ -128,6 +129,9 @@ let mk_meta_accu mv ty =
let mk_evar_accu ev ty =
mk_accu (Aevar (ev,ty))
+let mk_proj_accu kn c =
+ mk_accu (Aproj (kn,c))
+
let atom_of_accu (k:accumulator) =
(Obj.magic (Obj.field (Obj.magic k) 2) : atom)
diff --git a/kernel/nativevalues.mli b/kernel/nativevalues.mli
index 4fbf493cc9..32079c8d0c 100644
--- a/kernel/nativevalues.mli
+++ b/kernel/nativevalues.mli
@@ -52,6 +52,7 @@ type atom =
| Aprod of name * t * (t -> t)
| Ameta of metavariable * t
| Aevar of existential * t
+ | Aproj of constant * accumulator
(* Constructors *)
@@ -68,6 +69,7 @@ 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_proj_accu : constant -> accumulator -> t
val upd_cofix : t -> t -> unit
val force_cofix : t -> t
val mk_const : tag -> t
diff --git a/kernel/opaqueproof.ml b/kernel/opaqueproof.ml
index 102dcf99f3..673b12b2ca 100644
--- a/kernel/opaqueproof.ml
+++ b/kernel/opaqueproof.ml
@@ -7,11 +7,16 @@
(************************************************************************)
open Names
+open Univ
open Term
open Mod_subst
-type work_list = Id.t array Cmap.t * Id.t array Mindmap.t
-type cooking_info = { modlist : work_list; abstract : Context.named_context }
+type work_list = (Instance.t * Id.t array) Cmap.t *
+ (Instance.t * Id.t array) Mindmap.t
+
+type cooking_info = {
+ modlist : work_list;
+ abstract : Context.named_context in_universe_context }
type proofterm = (constr * Univ.constraints) Future.computation
type opaque =
| Indirect of substitution list * DirPath.t * int (* subst, lib, index *)
@@ -94,7 +99,7 @@ let force_constraints = function
| NoIndirect(_,cu) -> snd(Future.force cu)
| Indirect (_,dp,i) ->
match !get_univ dp i with
- | None -> Univ.empty_constraint
+ | None -> Univ.Constraint.empty
| Some u -> Future.force u
let get_constraints = function
diff --git a/kernel/opaqueproof.mli b/kernel/opaqueproof.mli
index 957889aa98..71f4917541 100644
--- a/kernel/opaqueproof.mli
+++ b/kernel/opaqueproof.mli
@@ -38,8 +38,12 @@ val get_constraints : opaque -> Univ.constraints Future.computation option
val subst_opaque : substitution -> opaque -> opaque
val iter_direct_opaque : (constr -> unit) -> opaque -> opaque
-type work_list = Id.t array Cmap.t * Id.t array Mindmap.t
-type cooking_info = { modlist : work_list; abstract : Context.named_context }
+type work_list = (Univ.Instance.t * Id.t array) Cmap.t *
+ (Univ.Instance.t * Id.t array) Mindmap.t
+
+type cooking_info = {
+ modlist : work_list;
+ abstract : Context.named_context Univ.in_universe_context }
(* The type has two caveats:
1) cook_constr is defined after
diff --git a/kernel/pre_env.ml b/kernel/pre_env.ml
index b655887d70..ba9f302334 100644
--- a/kernel/pre_env.ml
+++ b/kernel/pre_env.ml
@@ -124,7 +124,7 @@ let env_of_rel n env =
let push_named_context_val d (ctxt,vals) =
let id,_,_ = d in
let rval = ref VKnone in
- Context.add_named_decl d ctxt, (id,rval)::vals
+ add_named_decl d ctxt, (id,rval)::vals
let push_named d env =
(* if not (env.env_rel_context = []) then raise (ASSERT env.env_rel_context);
diff --git a/kernel/pre_env.mli b/kernel/pre_env.mli
index 964d709cff..74a5fb1aed 100644
--- a/kernel/pre_env.mli
+++ b/kernel/pre_env.mli
@@ -7,10 +7,10 @@
(************************************************************************)
open Names
-open Univ
open Term
open Context
open Declarations
+open Univ
(** The type of environments. *)
diff --git a/kernel/reduction.ml b/kernel/reduction.ml
index 5397e42f9e..63bd406817 100644
--- a/kernel/reduction.ml
+++ b/kernel/reduction.ml
@@ -26,11 +26,11 @@ open Environ
open Closure
open Esubst
-let unfold_reference ((ids, csts), infos) k =
+let conv_key k =
match k with
- | VarKey id when not (Id.Pred.mem id ids) -> None
- | ConstKey cst when not (Cpred.mem cst csts) -> None
- | _ -> unfold_reference infos k
+ | VarKey id -> VarKey id
+ | ConstKey (cst,_) -> ConstKey cst
+ | RelKey n -> RelKey n
let rec is_empty_stack = function
[] -> true
@@ -58,6 +58,8 @@ let compare_stack_shape stk1 stk2 =
| (_, (Zupdate _|Zshift _)::s2) -> compare_rec bal stk1 s2
| (Zapp l1::s1, _) -> compare_rec (bal+Array.length l1) s1 stk2
| (_, Zapp l2::s2) -> compare_rec (bal-Array.length l2) stk1 s2
+ | (Zproj (n1,m1,p1)::s1, Zproj (n2,m2,p2)::s2) ->
+ Int.equal bal 0 && compare_rec 0 s1 s2
| (Zcase(c1,_,_)::s1, Zcase(c2,_,_)::s2) ->
Int.equal bal 0 (* && c1.ci_ind = c2.ci_ind *) && compare_rec 0 s1 s2
| (Zfix(_,a1)::s1, Zfix(_,a2)::s2) ->
@@ -67,6 +69,7 @@ let compare_stack_shape stk1 stk2 =
type lft_constr_stack_elt =
Zlapp of (lift * fconstr) array
+ | Zlproj of constant * lift
| Zlfix of (lift * fconstr) * lft_constr_stack
| Zlcase of case_info * lift * fconstr * fconstr array
and lft_constr_stack = lft_constr_stack_elt list
@@ -85,6 +88,8 @@ let pure_stack lfts stk =
| (Zshift n,(l,pstk)) -> (el_shft n l, pstk)
| (Zapp a, (l,pstk)) ->
(l,zlapp (Array.map (fun t -> (l,t)) a) pstk)
+ | (Zproj (n,m,c), (l,pstk)) ->
+ (l, Zlproj (c,l)::pstk)
| (Zfix(fx,a),(l,pstk)) ->
let (lfx,pa) = pure_rec l a in
(l, Zlfix((lfx,fx),pa)::pstk)
@@ -96,17 +101,17 @@ let pure_stack lfts stk =
(* Reduction Functions *)
(****************************************************************************)
-let whd_betaiota t =
- whd_val (create_clos_infos betaiota empty_env) (inject t)
+let whd_betaiota env t =
+ whd_val (create_clos_infos betaiota env) (inject t)
-let nf_betaiota t =
- norm_val (create_clos_infos betaiota empty_env) (inject t)
+let nf_betaiota env t =
+ norm_val (create_clos_infos betaiota env) (inject t)
-let whd_betaiotazeta x =
+let whd_betaiotazeta env x =
match kind_of_term x with
| (Sort _|Var _|Meta _|Evar _|Const _|Ind _|Construct _|
Prod _|Lambda _|Fix _|CoFix _) -> x
- | _ -> whd_val (create_clos_infos betaiotazeta empty_env) (inject x)
+ | _ -> whd_val (create_clos_infos betaiotazeta env) (inject x)
let whd_betadeltaiota env t =
match kind_of_term t with
@@ -143,12 +148,31 @@ let betazeta_appvect n c v =
(********************************************************************)
(* Conversion utility functions *)
-type 'a conversion_function = env -> 'a -> 'a -> Univ.constraints
-type 'a trans_conversion_function = transparent_state -> env -> 'a -> 'a -> Univ.constraints
+type 'a conversion_function = env -> 'a -> 'a -> unit
+type 'a trans_conversion_function = Names.transparent_state -> 'a conversion_function
+type 'a universe_conversion_function = env -> Univ.universes -> 'a -> 'a -> unit
+type 'a trans_universe_conversion_function =
+ Names.transparent_state -> 'a universe_conversion_function
+type 'a infer_conversion_function = env -> Univ.universes -> 'a -> 'a -> Univ.constraints
+
+type conv_universes = Univ.universes * Univ.constraints option
exception NotConvertible
exception NotConvertibleVect of int
+let convert_universes (univs,cstrs as cuniv) u u' =
+ if Univ.Instance.check_eq univs u u' then cuniv
+ else
+ (match cstrs with
+ | None -> raise NotConvertible
+ | Some cstrs -> (univs, Some (Univ.enforce_eq_instances u u' cstrs)))
+
+let conv_table_key k1 k2 cuniv =
+ match k1, k2 with
+ | ConstKey (cst, u), ConstKey (cst', u') when eq_constant_key cst cst' ->
+ convert_universes cuniv u u'
+ | _ -> raise NotConvertible
+
let compare_stacks f fmind lft1 stk1 lft2 stk2 cuniv =
let rec cmp_rec pstk1 pstk2 cuniv =
match (pstk1,pstk2) with
@@ -156,6 +180,10 @@ let compare_stacks f fmind lft1 stk1 lft2 stk2 cuniv =
let cu1 = cmp_rec s1 s2 cuniv in
(match (z1,z2) with
| (Zlapp a1,Zlapp a2) -> Array.fold_right2 f a1 a2 cu1
+ | (Zlproj (c1,l1),Zlproj (c2,l2)) ->
+ if not (eq_constant c1 c2) then
+ raise NotConvertible
+ else cu1
| (Zlfix(fx1,a1),Zlfix(fx2,a2)) ->
let cu2 = f fx1 fx2 cu1 in
cmp_rec a1 a2 cu2
@@ -184,34 +212,64 @@ type conv_pb =
| CUMUL
let is_cumul = function CUMUL -> true | CONV -> false
-
-let sort_cmp pb s0 s1 cuniv =
+let is_pos = function Pos -> true | Null -> false
+
+(* let sort_cmp env pb s0 s1 cuniv = *)
+(* match (s0,s1) with *)
+(* | (Prop c1, Prop c2) when is_cumul pb -> *)
+(* begin match c1, c2 with *)
+(* | Null, _ | _, Pos -> cuniv (\* Prop <= Set *\) *)
+(* | _ -> raise NotConvertible *)
+(* end *)
+(* | (Prop c1, Prop c2) -> *)
+(* if c1 == c2 then cuniv else raise NotConvertible *)
+(* | (Prop c1, Type u) when is_cumul pb -> *)
+(* enforce_leq (if is_pos c1 then Universe.type0 else Universe.type0m) u cuniv *)
+(* | (Type u, Prop c) when is_cumul pb -> *)
+(* enforce_leq u (if is_pos c then Universe.type0 else Universe.type0m) cuniv *)
+(* | (Type u1, Type u2) -> *)
+(* (match pb with *)
+(* | CONV -> Univ.enforce_eq u1 u2 cuniv *)
+(* | CUMUL -> enforce_leq u1 u2 cuniv) *)
+(* | (_, _) -> raise NotConvertible *)
+
+(* let conv_sort env s0 s1 = sort_cmp CONV s0 s1 Constraint.empty *)
+(* let conv_sort_leq env s0 s1 = sort_cmp CUMUL s0 s1 Constraint.empty *)
+
+let check_eq (univs, cstrs as cuniv) u u' =
+ match cstrs with
+ | None -> if check_eq univs u u' then cuniv else raise NotConvertible
+ | Some cstrs -> univs, Some (Univ.enforce_eq u u' cstrs)
+
+let check_leq (univs, cstrs as cuniv) u u' =
+ match cstrs with
+ | None -> if check_leq univs u u' then cuniv else raise NotConvertible
+ | Some cstrs -> univs, Some (Univ.enforce_leq u u' cstrs)
+
+let sort_cmp_universes pb s0 s1 univs =
+ let dir = if is_cumul pb then check_leq univs else check_eq univs in
match (s0,s1) with
| (Prop c1, Prop c2) when is_cumul pb ->
begin match c1, c2 with
- | Null, _ | _, Pos -> cuniv (* Prop <= Set *)
+ | Null, _ | _, Pos -> univs (* Prop <= Set *)
| _ -> raise NotConvertible
end
- | (Prop c1, Prop c2) ->
- if c1 == c2 then cuniv else raise NotConvertible
- | (Prop c1, Type u) when is_cumul pb -> assert (is_univ_variable u); cuniv
- | (Type u1, Type u2) ->
- assert (is_univ_variable u2);
- (match pb with
- | CONV -> enforce_eq u1 u2 cuniv
- | CUMUL -> enforce_leq u1 u2 cuniv)
- | (_, _) -> raise NotConvertible
+ | (Prop c1, Prop c2) -> if c1 == c2 then univs else raise NotConvertible
+ | (Prop c1, Type u) -> dir (univ_of_sort s0) u
+ | (Type u, Prop c) -> dir u (univ_of_sort s1)
+ | (Type u1, Type u2) -> dir u1 u2
+(* let sort_cmp _ _ _ cuniv = cuniv *)
-let conv_sort env s0 s1 = sort_cmp CONV s0 s1 empty_constraint
-
-let conv_sort_leq env s0 s1 = sort_cmp CUMUL s0 s1 empty_constraint
+(* let conv_sort env s0 s1 = sort_cmp CONV s0 s1 empty_constraint *)
+(* let conv_sort_leq env s0 s1 = sort_cmp CUMUL s0 s1 empty_constraint *)
let rec no_arg_available = function
| [] -> true
| Zupdate _ :: stk -> no_arg_available stk
| Zshift _ :: stk -> no_arg_available stk
| Zapp v :: stk -> Int.equal (Array.length v) 0 && no_arg_available stk
+ | Zproj _ :: _ -> true
| Zcase _ :: _ -> true
| Zfix _ :: _ -> true
@@ -223,6 +281,7 @@ let rec no_nth_arg_available n = function
let k = Array.length v in
if n >= k then no_nth_arg_available (n-k) stk
else false
+ | Zproj _ :: _ -> true
| Zcase _ :: _ -> true
| Zfix _ :: _ -> true
@@ -231,6 +290,7 @@ let rec no_case_available = function
| Zupdate _ :: stk -> no_case_available stk
| Zshift _ :: stk -> no_case_available stk
| Zapp _ :: stk -> no_case_available stk
+ | Zproj (_,_,p) :: _ -> false
| Zcase _ :: _ -> false
| Zfix _ :: _ -> true
@@ -241,7 +301,7 @@ let in_whnf (t,stk) =
| FConstruct _ -> no_case_available stk
| FCoFix _ -> no_case_available stk
| FFix(((ri,n),(_,_,_)),_) -> no_nth_arg_available ri.(n) stk
- | (FFlex _ | FProd _ | FEvar _ | FInd _ | FAtom _ | FRel _) -> true
+ | (FFlex _ | FProd _ | FEvar _ | FInd _ | FAtom _ | FRel _ | FProj _) -> true
| FLOCKED -> assert false
let steps = ref 0
@@ -253,6 +313,15 @@ let slave_process =
| _ -> f := (fun () -> false); !f ()) in
fun () -> !f ()
+let unfold_projection infos p c =
+ if RedFlags.red_set infos.i_flags (RedFlags.fCONST 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, 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
@@ -266,9 +335,10 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
steps := 0;
end;
(* First head reduce both terms *)
+ let whd = whd_stack (infos_with_reds infos betaiotazeta) in
let rec whd_both (t1,stk1) (t2,stk2) =
- let st1' = whd_stack (snd infos) t1 stk1 in
- let st2' = whd_stack (snd infos) t2 stk2 in
+ let st1' = whd t1 stk1 in
+ let st2' = whd t2 stk2 in
(* Now, whd_stack on term2 might have modified st1 (due to sharing),
and st1 might not be in whnf anymore. If so, we iterate ccnv. *)
if in_whnf st1' then (st1',st2') else whd_both st1' st2' in
@@ -284,7 +354,7 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
| (Sort s1, Sort s2) ->
if not (is_empty_stack v1 && is_empty_stack v2) then
anomaly (Pp.str "conversion was given ill-typed terms (Sort)");
- sort_cmp cv_pb s1 s2 cuniv
+ sort_cmp_universes cv_pb s1 s2 cuniv
| (Meta n, Meta m) ->
if Int.equal n m
then convert_stacks l2r infos lft1 lft2 v1 v2 cuniv
@@ -292,10 +362,10 @@ 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 u1 = convert_stacks l2r infos lft1 lft2 v1 v2 cuniv 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)
- (Array.map (mk_clos env2) args2) u1
+ (Array.map (mk_clos env2) args2) cuniv
else raise NotConvertible
(* 2 index known to be bound to no constant *)
@@ -307,28 +377,59 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
(* 2 constants, 2 local defined vars or 2 defined rels *)
| (FFlex fl1, FFlex fl2) ->
(try (* try first intensional equality *)
- if eq_table_key fl1 fl2
- then convert_stacks l2r infos lft1 lft2 v1 v2 cuniv
- else raise NotConvertible
+ if eq_table_key fl1 fl2 then convert_stacks l2r infos lft1 lft2 v1 v2 cuniv
+ else
+ (let cuniv = conv_table_key fl1 fl2 cuniv in
+ convert_stacks l2r infos lft1 lft2 v1 v2 cuniv)
with NotConvertible ->
(* else the oracle tells which constant is to be expanded *)
let (app1,app2) =
- if Conv_oracle.oracle_order (Closure.oracle_of_infos (snd infos)) l2r fl1 fl2 then
+ if Conv_oracle.oracle_order (Closure.oracle_of_infos infos) l2r (conv_key fl1) (conv_key fl2) then
match unfold_reference infos fl1 with
- | Some def1 -> ((lft1, whd_stack (snd infos) def1 v1), appr2)
+ | Some def1 -> ((lft1, whd def1 v1), appr2)
| None ->
(match unfold_reference infos fl2 with
- | Some def2 -> (appr1, (lft2, whd_stack (snd infos) def2 v2))
+ | Some def2 -> (appr1, (lft2, whd def2 v2))
| None -> raise NotConvertible)
else
match unfold_reference infos fl2 with
- | Some def2 -> (appr1, (lft2, whd_stack (snd infos) def2 v2))
+ | Some def2 -> (appr1, (lft2, whd def2 v2))
| None ->
(match unfold_reference infos fl1 with
- | Some def1 -> ((lft1, whd_stack (snd infos) def1 v1), appr2)
+ | Some def1 -> ((lft1, whd def1 v1), appr2)
| None -> raise NotConvertible) in
eqappr cv_pb l2r infos app1 app2 cuniv)
+ | (FProj (p1,c1), FProj (p2, c2)) ->
+ (* 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, whd def1 (s1 :: v1)) appr2 cuniv
+ | None ->
+ match unfold_projection infos p2 c2 with
+ | Some (def2,s2) ->
+ eqappr cv_pb l2r infos appr1 (lft2, whd def2 (s2 :: v2)) cuniv
+ | None ->
+ if eq_constant p1 p2 && compare_stack_shape v1 v2 then
+ 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 *)
+ raise NotConvertible)
+
+ | (FProj (p1,c1), t2) ->
+ (match unfold_projection infos p1 c1 with
+ | Some (def1,s1) ->
+ eqappr cv_pb l2r infos (lft1, whd def1 (s1 :: v1)) appr2 cuniv
+ | None -> raise NotConvertible)
+
+ | (_, FProj (p2,c2)) ->
+ (match unfold_projection infos p2 c2 with
+ | Some (def2,s2) ->
+ eqappr cv_pb l2r infos appr1 (lft2, whd def2 (s2 :: v2)) cuniv
+ | None -> raise NotConvertible)
+
(* other constructors *)
| (FLambda _, FLambda _) ->
(* Inconsistency: we tolerate that v1, v2 contain shift and update but
@@ -337,15 +438,15 @@ 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 u1 = ccnv CONV l2r infos el1 el2 ty1 ty2 cuniv in
- ccnv CONV l2r infos (el_lift el1) (el_lift el2) bd1 bd2 u1
+ 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
| (FProd (_,c1,c2), FProd (_,c'1,c'2)) ->
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 u1 = 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 u1
+ 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
(* Eta-expansion on the fly *)
| (FLambda _, _) ->
@@ -368,30 +469,63 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
(el_lift lft1, (hd1, eta_expand_stack v1)) (el_lift lft2, (bd2, [])) cuniv
(* only one constant, defined var or defined rel *)
- | (FFlex fl1, _) ->
+ | (FFlex fl1, c2) ->
(match unfold_reference infos fl1 with
| Some def1 ->
- eqappr cv_pb l2r infos (lft1, whd_stack (snd infos) def1 v1) appr2 cuniv
- | None -> raise NotConvertible)
- | (_, FFlex fl2) ->
+ eqappr cv_pb l2r infos (lft1, whd def1 v1) appr2 cuniv
+ | None ->
+ match c2 with
+ | FConstruct ((ind2,j2),u2) ->
+ (try
+ let v2, v1 =
+ eta_expand_ind_stacks (info_env infos) ind2 hd2 v2 (snd appr1)
+ in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv
+ with Not_found -> raise NotConvertible)
+ | _ -> raise NotConvertible)
+
+ | (c1, FFlex fl2) ->
(match unfold_reference infos fl2 with
| Some def2 ->
- eqappr cv_pb l2r infos appr1 (lft2, whd_stack (snd infos) def2 v2) cuniv
- | None -> raise NotConvertible)
-
+ eqappr cv_pb l2r infos appr1 (lft2, whd def2 v2) cuniv
+ | None ->
+ match c1 with
+ | FConstruct ((ind1,j1),u1) ->
+ (try let v1, v2 =
+ eta_expand_ind_stacks (info_env infos) ind1 hd1 v1 (snd appr2)
+ in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv
+ with Not_found -> raise NotConvertible)
+ | _ -> raise NotConvertible)
+
(* Inductive types: MutInd MutConstruct Fix Cofix *)
- | (FInd ind1, FInd ind2) ->
+ | (FInd (ind1,u1), FInd (ind2,u2)) ->
if eq_ind ind1 ind2
then
- convert_stacks l2r infos lft1 lft2 v1 v2 cuniv
+ (let cuniv = convert_universes cuniv u1 u2 in
+ convert_stacks l2r infos lft1 lft2 v1 v2 cuniv)
else raise NotConvertible
- | (FConstruct (ind1,j1), FConstruct (ind2,j2)) ->
+ | (FConstruct ((ind1,j1),u1), FConstruct ((ind2,j2),u2)) ->
if Int.equal j1 j2 && eq_ind ind1 ind2
then
- convert_stacks l2r infos lft1 lft2 v1 v2 cuniv
+ (let cuniv = convert_universes cuniv u1 u2 in
+ convert_stacks l2r infos lft1 lft2 v1 v2 cuniv)
else raise NotConvertible
+
+ (* Eta expansion of records *)
+ | (FConstruct ((ind1,j1),u1), _) ->
+ (try
+ let v1, v2 =
+ eta_expand_ind_stacks (info_env infos) ind1 hd1 v1 (snd appr2)
+ in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv
+ with Not_found -> raise NotConvertible)
+
+ | (_, FConstruct ((ind2,j2),u2)) ->
+ (try
+ let v2, v1 =
+ eta_expand_ind_stacks (info_env infos) ind2 hd2 v2 (snd appr1)
+ in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv
+ with Not_found -> raise NotConvertible)
| (FFix (((op1, i1),(_,tys1,cl1)),e1), FFix(((op2, i2),(_,tys2,cl2)),e2)) ->
if Int.equal i1 i2 && Array.equal Int.equal op1 op2
@@ -401,11 +535,11 @@ 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 u1 = convert_vect l2r infos el1 el2 fty1 fty2 cuniv in
- let u2 =
+ let cuniv = convert_vect l2r infos el1 el2 fty1 fty2 cuniv in
+ let cuniv =
convert_vect l2r infos
- (el_liftn n el1) (el_liftn n el2) fcl1 fcl2 u1 in
- convert_stacks l2r infos lft1 lft2 v1 v2 u2
+ (el_liftn n el1) (el_liftn n el2) fcl1 fcl2 cuniv in
+ convert_stacks l2r infos lft1 lft2 v1 v2 cuniv
else raise NotConvertible
| (FCoFix ((op1,(_,tys1,cl1)),e1), FCoFix((op2,(_,tys2,cl2)),e2)) ->
@@ -416,11 +550,11 @@ 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 u1 = convert_vect l2r infos el1 el2 fty1 fty2 cuniv in
- let u2 =
+ let cuniv = convert_vect l2r infos el1 el2 fty1 fty2 cuniv in
+ let cuniv =
convert_vect l2r infos
- (el_liftn n el1) (el_liftn n el2) fcl1 fcl2 u1 in
- convert_stacks l2r infos lft1 lft2 v1 v2 u2
+ (el_liftn n el1) (el_liftn n el2) fcl1 fcl2 cuniv in
+ convert_stacks l2r infos lft1 lft2 v1 v2 cuniv
else raise NotConvertible
(* Should not happen because both (hd1,v1) and (hd2,v2) are in whnf *)
@@ -433,7 +567,7 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
and convert_stacks l2r infos lft1 lft2 stk1 stk2 cuniv =
compare_stacks
- (fun (l1,t1) (l2,t2) c -> ccnv CONV l2r infos l1 l2 t1 t2 c)
+ (fun (l1,t1) (l2,t2) c -> ccnv CONV l2r infos l1 l2 t1 t2 cuniv)
(eq_ind)
lft1 stk1 lft2 stk2 cuniv
@@ -442,26 +576,45 @@ and convert_vect l2r infos lft1 lft2 v1 v2 cuniv =
let lv2 = Array.length v2 in
if Int.equal lv1 lv2
then
- let rec fold n univ =
- if n >= lv1 then univ
+ let rec fold n cuniv =
+ if n >= lv1 then cuniv
else
- let u1 = ccnv CONV l2r infos lft1 lft2 v1.(n) v2.(n) univ in
- fold (n+1) u1 in
+ let cuniv = ccnv CONV l2r infos lft1 lft2 v1.(n) v2.(n) cuniv in
+ fold (n+1) cuniv in
fold 0 cuniv
else raise NotConvertible
-let clos_fconv trans cv_pb l2r evars env t1 t2 =
- let infos = trans, create_clos_infos ~evars betaiotazeta env in
- ccnv cv_pb l2r infos el_id el_id (inject t1) (inject t2) empty_constraint
+let clos_fconv trans cv_pb l2r evars env univs t1 t2 =
+ let reds = Closure.RedFlags.red_add_transparent betaiotazeta trans in
+ let infos = create_clos_infos ~evars reds env in
+ ccnv cv_pb l2r infos el_id el_id (inject t1) (inject t2) univs
-let trans_fconv reds cv_pb l2r evars env t1 t2 =
- if eq_constr t1 t2 then empty_constraint
- else clos_fconv reds cv_pb l2r evars env t1 t2
+let trans_fconv_universes reds cv_pb l2r evars env univs t1 t2 =
+ let b =
+ if cv_pb = CUMUL then leq_constr_univs univs t1 t2
+ else eq_constr_univs univs t1 t2
+ in
+ if b then ()
+ else
+ let _ = clos_fconv reds cv_pb l2r evars env (univs, None) t1 t2 in
+ ()
+
+(* Profiling *)
+(* let trans_fconv_universes_key = Profile.declare_profile "trans_fconv_universes" *)
+(* let trans_fconv_universes = Profile.profile8 trans_fconv_universes_key trans_fconv_universes *)
+
+let trans_fconv reds cv_pb l2r evars env =
+ trans_fconv_universes reds cv_pb l2r evars env (universes env)
let trans_conv_cmp ?(l2r=false) conv reds = trans_fconv reds conv l2r (fun _->None)
let trans_conv ?(l2r=false) ?(evars=fun _->None) reds = trans_fconv reds CONV l2r evars
let trans_conv_leq ?(l2r=false) ?(evars=fun _->None) reds = trans_fconv reds CUMUL l2r evars
+let trans_conv_universes ?(l2r=false) ?(evars=fun _->None) reds =
+ trans_fconv_universes reds CONV l2r evars
+let trans_conv_leq_universes ?(l2r=false) ?(evars=fun _->None) reds =
+ trans_fconv_universes reds CUMUL l2r evars
+
let fconv = trans_fconv (Id.Pred.full, Cpred.full)
let conv_cmp ?(l2r=false) cv_pb = fconv cv_pb l2r (fun _->None)
@@ -470,22 +623,43 @@ let conv_leq ?(l2r=false) ?(evars=fun _->None) = fconv CUMUL l2r evars
let conv_leq_vecti ?(l2r=false) ?(evars=fun _->None) env v1 v2 =
Array.fold_left2_i
- (fun i c t1 t2 ->
- let c' =
- try conv_leq ~l2r ~evars env t1 t2
- with NotConvertible -> raise (NotConvertibleVect i) in
- union_constraints c c')
- empty_constraint
+ (fun i _ t1 t2 ->
+ try conv_leq ~l2r ~evars env t1 t2
+ with NotConvertible -> raise (NotConvertibleVect i))
+ ()
v1
v2
+let infer_conv_universes cv_pb l2r evars reds env univs t1 t2 =
+ let b =
+ if cv_pb = CUMUL then leq_constr_univs univs t1 t2
+ else eq_constr_univs univs t1 t2
+ in
+ if b then Constraint.empty
+ else
+ let (u, cstrs) =
+ clos_fconv reds cv_pb l2r evars env (univs, Some Constraint.empty) t1 t2
+ in Option.get cstrs
+
+(* Profiling *)
+(* let infer_conv_universes_key = Profile.declare_profile "infer_conv_universes" *)
+(* let infer_conv_universes = Profile.profile8 infer_conv_universes_key infer_conv_universes *)
+
+let infer_conv ?(l2r=false) ?(evars=fun _ -> None) ?(ts=full_transparent_state)
+ env univs t1 t2 =
+ infer_conv_universes CONV l2r evars ts env univs t1 t2
+
+let infer_conv_leq ?(l2r=false) ?(evars=fun _ -> None) ?(ts=full_transparent_state)
+ env univs t1 t2 =
+ infer_conv_universes CUMUL l2r evars ts env univs t1 t2
+
(* option for conversion *)
let nat_conv = ref (fun cv_pb sigma ->
fconv cv_pb false (sigma.Nativelambda.evars_val))
let set_nat_conv f = nat_conv := f
let native_conv cv_pb sigma env t1 t2 =
- if eq_constr t1 t2 then empty_constraint
+ if eq_constr t1 t2 then ()
else begin
let t1 = (it_mkLambda_or_LetIn t1 (rel_context env)) in
let t2 = (it_mkLambda_or_LetIn t2 (rel_context env)) in
diff --git a/kernel/reduction.mli b/kernel/reduction.mli
index 7c0607cc4d..b9bd41f289 100644
--- a/kernel/reduction.mli
+++ b/kernel/reduction.mli
@@ -13,28 +13,39 @@ open Environ
(***********************************************************************
s Reduction functions *)
-val whd_betaiotazeta : constr -> constr
+val whd_betaiotazeta : env -> constr -> constr
val whd_betadeltaiota : env -> constr -> constr
val whd_betadeltaiota_nolet : env -> constr -> constr
-val whd_betaiota : constr -> constr
-val nf_betaiota : constr -> constr
+val whd_betaiota : env -> constr -> constr
+val nf_betaiota : env -> constr -> constr
(***********************************************************************
s conversion functions *)
exception NotConvertible
exception NotConvertibleVect of int
-type 'a conversion_function = env -> 'a -> 'a -> Univ.constraints
-type 'a trans_conversion_function = Names.transparent_state -> env -> 'a -> 'a -> Univ.constraints
+
+type conv_universes = Univ.universes * Univ.constraints option
+
+type 'a conversion_function = env -> 'a -> 'a -> unit
+type 'a trans_conversion_function = Names.transparent_state -> 'a conversion_function
+type 'a universe_conversion_function = env -> Univ.universes -> 'a -> 'a -> unit
+type 'a trans_universe_conversion_function =
+ Names.transparent_state -> 'a universe_conversion_function
+
+type 'a infer_conversion_function = env -> Univ.universes -> 'a -> 'a -> Univ.constraints
type conv_pb = CONV | CUMUL
-val sort_cmp :
- conv_pb -> sorts -> sorts -> Univ.constraints -> Univ.constraints
+val sort_cmp_universes :
+ conv_pb -> sorts -> sorts -> conv_universes -> conv_universes
-val conv_sort : sorts conversion_function
-val conv_sort_leq : sorts conversion_function
+(* val sort_cmp : *)
+(* conv_pb -> sorts -> sorts -> Univ.constraints -> Univ.constraints *)
+
+(* val conv_sort : sorts conversion_function *)
+(* val conv_sort_leq : sorts conversion_function *)
val trans_conv_cmp : ?l2r:bool -> conv_pb -> constr trans_conversion_function
val trans_conv :
@@ -42,6 +53,11 @@ val trans_conv :
val trans_conv_leq :
?l2r:bool -> ?evars:(existential->constr option) -> types trans_conversion_function
+val trans_conv_universes :
+ ?l2r:bool -> ?evars:(existential->constr option) -> constr trans_universe_conversion_function
+val trans_conv_leq_universes :
+ ?l2r:bool -> ?evars:(existential->constr option) -> types trans_universe_conversion_function
+
val conv_cmp : ?l2r:bool -> conv_pb -> constr conversion_function
val conv :
?l2r:bool -> ?evars:(existential->constr option) -> constr conversion_function
@@ -50,6 +66,11 @@ val conv_leq :
val conv_leq_vecti :
?l2r:bool -> ?evars:(existential->constr option) -> types array conversion_function
+val infer_conv : ?l2r:bool -> ?evars:(existential->constr option) ->
+ ?ts:Names.transparent_state -> constr infer_conversion_function
+val infer_conv_leq : ?l2r:bool -> ?evars:(existential->constr option) ->
+ ?ts:Names.transparent_state -> types infer_conversion_function
+
(** option for conversion *)
val set_vm_conv : (conv_pb -> types conversion_function) -> unit
val vm_conv : conv_pb -> types conversion_function
diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml
index c89766fb9b..093797fc05 100644
--- a/kernel/safe_typing.ml
+++ b/kernel/safe_typing.ml
@@ -139,7 +139,7 @@ let empty_environment =
modlabels = Label.Set.empty;
objlabels = Label.Set.empty;
future_cst = [];
- univ = Univ.empty_constraint;
+ univ = Univ.Constraint.empty;
engagement = None;
imports = [];
loads = [];
@@ -197,7 +197,10 @@ let add_constraints cst senv =
| Now cst ->
{ senv with
env = Environ.add_constraints cst senv.env;
- univ = Univ.union_constraints cst senv.univ }
+ univ = Univ.Constraint.union cst senv.univ }
+
+let push_context_set ctx = add_constraints (Now (Univ.ContextSet.constraints ctx))
+let push_context ctx = add_constraints (Now (Univ.UContext.constraints ctx))
let is_curmod_library senv =
match senv.modvariant with LIBRARY -> true | _ -> false
@@ -291,22 +294,22 @@ let safe_push_named (id,_,_ as d) env =
with Not_found -> () in
Environ.push_named d env
+
let push_named_def (id,de) senv =
- let (c,typ,cst) = Term_typing.translate_local_def senv.env id de in
- let c,cst' = match c with
- | Def c -> Mod_subst.force_constr c, Univ.empty_constraint
- | OpaqueDef o -> Opaqueproof.force_proof o, Opaqueproof.force_constraints o
+ let c,typ,univs = Term_typing.translate_local_def senv.env id de in
+ let c = match c with
+ | Def c -> Mod_subst.force_constr c
+ | OpaqueDef o -> Opaqueproof.force_proof o
| _ -> assert false in
- let senv = add_constraints (Now cst') senv in
- let senv' = add_constraints (Now cst) senv in
+ let senv' = push_context de.Entries.const_entry_universes senv in
let env'' = safe_push_named (id,Some c,typ) senv'.env in
- (Univ.union_constraints cst cst', {senv' with env=env''})
+ {senv' with env=env''}
-let push_named_assum (id,t) senv =
- let (t,cst) = Term_typing.translate_local_assum senv.env t in
- let senv' = add_constraints (Now cst) senv in
+let push_named_assum ((id,t),ctx) senv =
+ let senv' = push_context_set 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
- (cst, {senv' with env=env''})
+ {senv' with env=env''}
(** {6 Insertion of new declarations to current environment } *)
@@ -324,20 +327,35 @@ let labels_of_mib mib =
Array.iter visit_mip mib.mind_packets;
get ()
-let constraints_of_sfb = function
- | SFBmind mib -> [Now mib.mind_constraints]
- | SFBmodtype mtb -> [Now mtb.typ_constraints]
- | SFBmodule mb -> [Now mb.mod_constraints]
- | SFBconst cb -> [Now cb.const_constraints] @
- match cb.const_body with
- | (Undef _ | Def _) -> []
- | OpaqueDef lc ->
- match Opaqueproof.get_constraints lc with
- | None -> []
- | Some fc ->
- match Future.peek_val fc with
- | None -> [Later fc]
- | Some c -> [Now c]
+let globalize_constant_universes cb =
+ if cb.const_polymorphic then
+ Now Univ.Constraint.empty
+ else
+ (match Future.peek_val cb.const_universes with
+ | Some c -> Now (Univ.UContext.constraints c)
+ | None -> Later (Future.chain ~pure:true cb.const_universes Univ.UContext.constraints))
+
+let globalize_mind_universes mb =
+ if mb.mind_polymorphic then
+ Now Univ.Constraint.empty
+ else
+ Now (Univ.UContext.constraints mb.mind_universes)
+
+let constraints_of_sfb sfb =
+ match sfb with
+ | SFBconst cb -> globalize_constant_universes cb
+ | SFBmind mib -> globalize_mind_universes mib
+ | SFBmodtype mtb -> Now mtb.typ_constraints
+ | SFBmodule mb -> Now mb.mod_constraints
+
+(* let add_constraints cst senv = *)
+(* { senv with *)
+(* env = Environ.add_constraints cst senv.env; *)
+(* univ = Univ.Constraint.union cst senv.univ } *)
+
+(* let next_universe senv = *)
+(* let univ = senv.max_univ in *)
+(* univ + 1, { senv with max_univ = univ + 1 } *)
(** A generic function for adding a new field in a same environment.
It also performs the corresponding [add_constraints]. *)
@@ -358,7 +376,8 @@ let add_field ((l,sfb) as field) gn senv =
| SFBmodule _ | SFBmodtype _ ->
check_modlabel l senv; (Label.Set.singleton l, Label.Set.empty)
in
- let senv = List.fold_right add_constraints (constraints_of_sfb sfb) senv in
+ let cst = constraints_of_sfb sfb in
+ let senv = add_constraints cst senv in
let env' = match sfb, gn with
| SFBconst cb, C con -> Environ.add_constant con cb senv.env
| SFBmind mib, I mind -> Environ.add_mind mind mib senv.env
@@ -377,7 +396,6 @@ let add_field ((l,sfb) as field) gn senv =
let update_resolver f senv = { senv with modresolver = f senv.modresolver }
(** Insertion of constants and parameters in environment *)
-
type global_declaration =
| ConstantEntry of Entries.constant_entry
| GlobalRecipe of Cooking.recipe
@@ -548,8 +566,8 @@ let propagate_senv newdef newenv newresolver senv oldsenv =
modlabels = Label.Set.add (fst newdef) oldsenv.modlabels;
univ =
List.fold_left (fun acc cst ->
- Univ.union_constraints acc (Future.force cst))
- (Univ.union_constraints senv.univ oldsenv.univ)
+ Univ.Constraint.union acc (Future.force cst))
+ (Univ.Constraint.union senv.univ oldsenv.univ)
now_cst;
future_cst = later_cst @ oldsenv.future_cst;
(* engagement is propagated to the upper level *)
@@ -571,7 +589,7 @@ let end_module l restype senv =
let senv'=
propagate_loads { senv with
env = newenv;
- univ = Univ.union_constraints senv.univ mb.mod_constraints} in
+ univ = Univ.Constraint.union senv.univ mb.mod_constraints} in
let newenv = Environ.add_constraints mb.mod_constraints senv'.env in
let newenv = Modops.add_module mb newenv in
let newresolver =
@@ -637,7 +655,7 @@ let add_include me is_module inl senv =
{ typ_mp = mp_sup;
typ_expr = NoFunctor (List.rev senv.revstruct);
typ_expr_alg = None;
- typ_constraints = Univ.empty_constraint;
+ typ_constraints = Univ.Constraint.empty;
typ_delta = senv.modresolver } in
compute_sign sign mtb resolver senv
in
@@ -672,6 +690,10 @@ type compiled_library = {
type native_library = Nativecode.global list
+(** FIXME: MS: remove?*)
+let current_modpath senv = senv.modpath
+let current_dirpath senv = Names.ModPath.dp (current_modpath senv)
+
let start_library dir senv =
check_initial senv;
assert (not (DirPath.is_empty dir));
@@ -747,10 +769,7 @@ type judgment = Environ.unsafe_judgment
let j_val j = j.Environ.uj_val
let j_type j = j.Environ.uj_type
-let safe_infer senv = Typeops.infer (env_of_senv senv)
-
-let typing senv = Typeops.typing (env_of_senv senv)
-
+let typing senv = Typeops.infer (env_of_senv senv)
(** {6 Retroknowledge / native compiler } *)
diff --git a/kernel/safe_typing.mli b/kernel/safe_typing.mli
index d70d7d8be2..ad2148ead9 100644
--- a/kernel/safe_typing.mli
+++ b/kernel/safe_typing.mli
@@ -55,9 +55,9 @@ val join_safe_environment : safe_environment -> safe_environment
(** Insertion of local declarations (Local or Variables) *)
val push_named_assum :
- Id.t * Term.types -> Univ.constraints safe_transformer
+ (Id.t * Term.types) Univ.in_universe_context_set -> safe_transformer0
val push_named_def :
- Id.t * Entries.definition_entry -> Univ.constraints safe_transformer
+ Id.t * Entries.definition_entry -> safe_transformer0
(** Insertion of global axioms or definitions *)
@@ -85,10 +85,19 @@ val add_modtype :
(** Adding universe constraints *)
-val add_constraints : Univ.constraints -> safe_transformer0
+val push_context_set :
+ Univ.universe_context_set -> safe_transformer0
-(** Setting the Set-impredicative engagement *)
+val push_context :
+ Univ.universe_context -> safe_transformer0
+val add_constraints :
+ Univ.constraints -> safe_transformer0
+
+(* (\** Generator of universes *\) *)
+(* val next_universe : int safe_transformer *)
+
+(** Settin the strongly constructive or classical logical engagement *)
val set_engagement : Declarations.engagement -> safe_transformer0
(** {6 Interactive module functions } *)
@@ -113,6 +122,10 @@ val add_include :
Entries.module_struct_entry -> bool -> Declarations.inline ->
Mod_subst.delta_resolver safe_transformer
+val current_modpath : safe_environment -> module_path
+
+val current_dirpath : safe_environment -> dir_path
+
(** {6 Libraries : loading and saving compilation units } *)
type compiled_library
@@ -137,12 +150,7 @@ type judgment
val j_val : judgment -> Term.constr
val j_type : judgment -> Term.constr
-(** The safe typing of a term returns a typing judgment and some universe
- constraints (to be added to the environment for the judgment to
- hold). It is guaranteed that the constraints are satisfiable.
- *)
-val safe_infer : safe_environment -> Term.constr -> judgment * Univ.constraints
-
+(** The safe typing of a term returns a typing judgment. *)
val typing : safe_environment -> Term.constr -> judgment
(** {6 Queries } *)
@@ -164,4 +172,4 @@ val register :
val register_inline : constant -> safe_transformer0
val set_strategy :
- safe_environment -> 'a Names.tableKey -> Conv_oracle.level -> safe_environment
+ safe_environment -> Names.constant Names.tableKey -> Conv_oracle.level -> safe_environment
diff --git a/kernel/sorts.ml b/kernel/sorts.ml
index 03f1cd2653..3ebd06dd87 100644
--- a/kernel/sorts.ml
+++ b/kernel/sorts.ml
@@ -20,6 +20,16 @@ let prop = Prop Null
let set = Prop Pos
let type1 = Type type1_univ
+let univ_of_sort = function
+ | Type u -> u
+ | Prop Pos -> Universe.type0
+ | Prop Null -> Universe.type0m
+
+let sort_of_univ u =
+ if is_type0m_univ u then Prop Null
+ else if is_type0_univ u then Prop Pos
+ else Type u
+
let compare s1 s2 =
if s1 == s2 then 0 else
match s1, s2 with
@@ -36,8 +46,16 @@ let compare s1 s2 =
let equal s1 s2 = Int.equal (compare s1 s2) 0
let is_prop = function
-| Prop Null -> true
-| _ -> false
+ | Prop Null -> true
+ | _ -> false
+
+let is_set = function
+ | Prop Pos -> true
+ | _ -> false
+
+let is_small = function
+ | Prop _ -> true
+ | Type u -> is_small_univ u
let family = function
| Prop Null -> InProp
@@ -56,7 +74,7 @@ let hash = function
in
combinesmall 1 h
| Type u ->
- let h = Universe.hash u in
+ let h = Hashtbl.hash u in (** FIXME *)
combinesmall 2 h
module List = struct
@@ -70,14 +88,18 @@ module Hsorts =
type _t = t
type t = _t
type u = universe -> universe
+
let hashcons huniv = function
- | Type u -> Type (huniv u)
+ | Type u as c ->
+ let u' = huniv u in
+ if u' == u then c else Type u'
| s -> s
let equal s1 s2 = match (s1,s2) with
| (Prop c1, Prop c2) -> c1 == c2
| (Type u1, Type u2) -> u1 == u2
|_ -> false
- let hash = hash
+
+ let hash = Hashtbl.hash (** FIXME *)
end)
let hcons = Hashcons.simple_hcons Hsorts.generate hcons_univ
diff --git a/kernel/sorts.mli b/kernel/sorts.mli
index 2750145f11..ff7d138d64 100644
--- a/kernel/sorts.mli
+++ b/kernel/sorts.mli
@@ -24,7 +24,9 @@ val equal : t -> t -> bool
val compare : t -> t -> int
val hash : t -> int
+val is_set : t -> bool
val is_prop : t -> bool
+val is_small : t -> bool
val family : t -> family
val hcons : t -> t
@@ -35,3 +37,6 @@ module List : sig
val mem : family -> family list -> bool
val intersect : family list -> family list -> family list
end
+
+val univ_of_sort : t -> Univ.universe
+val sort_of_univ : Univ.universe -> t
diff --git a/kernel/subtyping.ml b/kernel/subtyping.ml
index af4468981e..2c093939ae 100644
--- a/kernel/subtyping.ml
+++ b/kernel/subtyping.ml
@@ -80,10 +80,8 @@ let make_labmap mp list =
let check_conv_error error why cst f env a1 a2 =
- try
- union_constraints cst (f env a1 a2)
- with
- NotConvertible -> error why
+ try Constraint.union cst (f env (Environ.universes env) a1 a2)
+ with NotConvertible -> error why
(* for now we do not allow reorderings *)
@@ -94,10 +92,15 @@ let check_inductive cst env mp1 l info1 mp2 mib2 spec2 subst1 subst2 reso1 reso2
let check_conv why cst f = check_conv_error error why cst f in
let mib1 =
match info1 with
- | IndType ((_,0), mib) -> Declareops.subst_mind subst1 mib
+ | IndType ((_,0), mib) -> Declareops.subst_mind_body subst1 mib
| _ -> error (InductiveFieldExpected mib2)
in
- let mib2 = Declareops.subst_mind subst2 mib2 in
+ let u =
+ if mib1.mind_polymorphic then
+ UContext.instance mib1.mind_universes
+ else Instance.empty
+ in
+ let mib2 = Declareops.subst_mind_body subst2 mib2 in
let check_inductive_type cst name env t1 t2 =
(* Due to sort-polymorphism in inductive types, the conclusions of
@@ -131,7 +134,7 @@ let check_inductive cst env mp1 l info1 mp2 mib2 spec2 subst1 subst2 reso1 reso2
error (NotConvertibleInductiveField name)
| _ -> (s1, s2) in
check_conv (NotConvertibleInductiveField name)
- cst conv_leq env (mkArity (ctx1,s1)) (mkArity (ctx2,s2))
+ cst infer_conv_leq env (mkArity (ctx1,s1)) (mkArity (ctx2,s2))
in
let check_packet cst p1 p2 =
@@ -149,18 +152,20 @@ let check_inductive cst env mp1 l info1 mp2 mib2 spec2 subst1 subst2 reso1 reso2
(* nparams done *)
(* params_ctxt done because part of the inductive types *)
(* Don't check the sort of the type if polymorphic *)
- let cst = check_inductive_type cst p2.mind_typename env (type_of_inductive env (mib1,p1)) (type_of_inductive env (mib2,p2))
- in
+ let ty1, cst1 = constrained_type_of_inductive env ((mib1,p1),u) in
+ let ty2, cst2 = constrained_type_of_inductive env ((mib2,p2),u) in
+ let cst = Constraint.union cst1 (Constraint.union cst2 cst) in
+ let cst = check_inductive_type cst p2.mind_typename env ty1 ty2 in
cst
in
let mind = mind_of_kn kn1 in
let check_cons_types i cst p1 p2 =
Array.fold_left3
- (fun cst id t1 t2 -> check_conv (NotConvertibleConstructorField id) cst conv env t1 t2)
+ (fun cst id t1 t2 -> check_conv (NotConvertibleConstructorField id) cst infer_conv env t1 t2)
cst
p2.mind_consnames
- (arities_of_specif mind (mib1,p1))
- (arities_of_specif mind (mib2,p2))
+ (arities_of_specif (mind,u) (mib1,p1))
+ (arities_of_specif (mind,u) (mib2,p2))
in
let check f test why = if not (test (f mib1) (f mib2)) then error (why (f mib2)) in
check (fun mib -> mib.mind_finite) (==) (fun x -> FiniteInductiveFieldExpected x);
@@ -180,13 +185,13 @@ let check_inductive cst env mp1 l info1 mp2 mib2 spec2 subst1 subst2 reso1 reso2
let kn2' = kn_of_delta reso2 kn2 in
if KerName.equal kn2 kn2' ||
MutInd.equal (mind_of_delta_kn reso1 kn1)
- (subst_ind subst2 (MutInd.make kn2 kn2'))
+ (subst_mind subst2 (MutInd.make kn2 kn2'))
then ()
else error NotEqualInductiveAliases
end;
(* we check that records and their field names are preserved. *)
- check (fun mib -> mib.mind_record) (==) (fun x -> RecordFieldExpected x);
- if mib1.mind_record then begin
+ check (fun mib -> mib.mind_record <> None) (==) (fun x -> RecordFieldExpected x);
+ if mib1.mind_record <> None then begin
let rec names_prod_letin t = match kind_of_term t with
| Prod(n,_,t) -> n::(names_prod_letin t)
| LetIn(n,_,_,t) -> n::(names_prod_letin t)
@@ -264,17 +269,16 @@ let check_constant cst env mp1 l info1 cb2 spec2 subst1 subst2 =
t1,t2
else
(t1,t2) in
- check_conv err cst conv_leq env t1 t2
+ check_conv err cst infer_conv_leq env t1 t2
in
-
match info1 with
| Constant cb1 ->
let () = assert (List.is_empty cb1.const_hyps && List.is_empty cb2.const_hyps) in
let cb1 = Declareops.subst_const_body subst1 cb1 in
let cb2 = Declareops.subst_const_body subst2 cb2 in
(* Start by checking types*)
- let typ1 = Typeops.type_of_constant_type env cb1.const_type in
- let typ2 = Typeops.type_of_constant_type env cb2.const_type in
+ let typ1 = cb1.const_type in
+ let typ2 = cb2.const_type in
let cst = check_type cst env typ1 typ2 in
(* Now we check the bodies:
- A transparent constant can only be implemented by a compatible
@@ -292,7 +296,7 @@ let check_constant cst env mp1 l info1 cb2 spec2 subst1 subst2 =
Anyway [check_conv] will handle that afterwards. *)
let c1 = Mod_subst.force_constr lc1 in
let c2 = Mod_subst.force_constr lc2 in
- check_conv NotConvertibleBodyField cst conv env c1 c2))
+ check_conv NotConvertibleBodyField cst infer_conv env c1 c2))
| IndType ((kn,i),mind1) ->
ignore (Errors.error (
"The kernel does not recognize yet that a parameter can be " ^
@@ -301,10 +305,14 @@ let check_constant cst env mp1 l info1 cb2 spec2 subst1 subst2 =
"name."));
let () = assert (List.is_empty mind1.mind_hyps && List.is_empty cb2.const_hyps) in
if Declareops.constant_has_body cb2 then error DefinitionFieldExpected;
- let arity1 = type_of_inductive env (mind1,mind1.mind_packets.(i)) in
- let typ2 = Typeops.type_of_constant_type env cb2.const_type in
+ let u1 = inductive_instance mind1 in
+ let arity1,cst1 = constrained_type_of_inductive env
+ ((mind1,mind1.mind_packets.(i)),u1) in
+ let cst2 = UContext.constraints (Future.force cb2.const_universes) in
+ let typ2 = cb2.const_type in
+ let cst = Constraint.union cst (Constraint.union cst1 cst2) in
let error = NotConvertibleTypeField (env, arity1, typ2) in
- check_conv error cst conv_leq env arity1 typ2
+ check_conv error cst infer_conv_leq env arity1 typ2
| IndConstr (((kn,i),j) as cstr,mind1) ->
ignore (Errors.error (
"The kernel does not recognize yet that a parameter can be " ^
@@ -313,10 +321,13 @@ let check_constant cst env mp1 l info1 cb2 spec2 subst1 subst2 =
"name."));
let () = assert (List.is_empty mind1.mind_hyps && List.is_empty cb2.const_hyps) in
if Declareops.constant_has_body cb2 then error DefinitionFieldExpected;
- let ty1 = type_of_constructor cstr (mind1,mind1.mind_packets.(i)) in
- let ty2 = Typeops.type_of_constant_type env cb2.const_type in
+ let u1 = inductive_instance mind1 in
+ let ty1,cst1 = constrained_type_of_constructor (cstr,u1) (mind1,mind1.mind_packets.(i)) in
+ let cst2 = UContext.constraints (Future.force cb2.const_universes) in
+ let ty2 = cb2.const_type in
+ let cst = Constraint.union cst (Constraint.union cst1 cst2) in
let error = NotConvertibleTypeField (env, ty1, ty2) in
- check_conv error cst conv env ty1 ty2
+ check_conv error cst infer_conv env ty1 ty2
let rec check_modules cst env msb1 msb2 subst1 subst2 =
let mty1 = module_type_of_module msb1 in
@@ -368,7 +379,7 @@ and check_modtypes cst env mtb1 mtb2 subst1 subst2 equiv =
mtb2.typ_mp list2 mtb1.typ_mp list1 subst2 subst1
mtb2.typ_delta mtb1.typ_delta
in
- Univ.union_constraints cst1 cst2
+ Univ.Constraint.union cst1 cst2
else
check_signatures cst env
mtb1.typ_mp list1 mtb2.typ_mp list2 subst1 subst2
@@ -398,7 +409,7 @@ and check_modtypes cst env mtb1 mtb2 subst1 subst2 equiv =
let check_subtypes env sup super =
let env = add_module_type sup.typ_mp sup env in
- check_modtypes empty_constraint env
+ check_modtypes Univ.Constraint.empty env
(strengthen sup sup.typ_mp) super empty_subst
(map_mp super.typ_mp sup.typ_mp sup.typ_delta) false
diff --git a/kernel/term.ml b/kernel/term.ml
index 24fe6d9622..b85c525d1f 100644
--- a/kernel/term.ml
+++ b/kernel/term.ml
@@ -32,7 +32,6 @@ type types = Constr.t
(** Same as [constr], for documentation purposes. *)
type existential_key = Constr.existential_key
-
type existential = Constr.existential
type metavariable = Constr.metavariable
@@ -54,6 +53,10 @@ type case_info = Constr.case_info =
type cast_kind = Constr.cast_kind =
VMcast | NATIVEcast | DEFAULTcast | REVERTcast
+(********************************************************************)
+(* Constructions as implemented *)
+(********************************************************************)
+
type rec_declaration = Constr.rec_declaration
type fixpoint = Constr.fixpoint
type cofixpoint = Constr.cofixpoint
@@ -62,6 +65,12 @@ type ('constr, 'types) prec_declaration =
('constr, 'types) Constr.prec_declaration
type ('constr, 'types) pfixpoint = ('constr, 'types) Constr.pfixpoint
type ('constr, 'types) pcofixpoint = ('constr, 'types) Constr.pcofixpoint
+type 'a puniverses = 'a Univ.puniverses
+
+(** Simply type aliases *)
+type pconstant = constant puniverses
+type pinductive = inductive puniverses
+type pconstructor = constructor puniverses
type ('constr, 'types) kind_of_term = ('constr, 'types) Constr.kind_of_term =
| Rel of int
@@ -74,12 +83,13 @@ type ('constr, 'types) kind_of_term = ('constr, 'types) Constr.kind_of_term =
| Lambda of Name.t * 'types * 'constr
| LetIn of Name.t * 'constr * 'types * 'constr
| App of 'constr * 'constr array
- | Const of constant
- | Ind of inductive
- | Construct of constructor
+ | Const of pconstant
+ | Ind of pinductive
+ | Construct of pconstructor
| Case of case_info * 'constr * 'constr * 'constr array
| Fix of ('constr, 'types) pfixpoint
| CoFix of ('constr, 'types) pcofixpoint
+ | Proj of constant * 'constr
type values = Constr.values
@@ -93,6 +103,8 @@ let type1_sort = Sorts.type1
let sorts_ord = Sorts.compare
let is_prop_sort = Sorts.is_prop
let family_of_sort = Sorts.family
+let univ_of_sort = Sorts.univ_of_sort
+let sort_of_univ = Sorts.sort_of_univ
(** {6 Term constructors. } *)
@@ -110,8 +122,13 @@ let mkLambda = Constr.mkLambda
let mkLetIn = Constr.mkLetIn
let mkApp = Constr.mkApp
let mkConst = Constr.mkConst
+let mkProj = Constr.mkProj
let mkInd = Constr.mkInd
let mkConstruct = Constr.mkConstruct
+let mkConstU = Constr.mkConstU
+let mkIndU = Constr.mkIndU
+let mkConstructU = Constr.mkConstructU
+let mkConstructUi = Constr.mkConstructUi
let mkCase = Constr.mkCase
let mkFix = Constr.mkFix
let mkCoFix = Constr.mkCoFix
@@ -121,9 +138,16 @@ let mkCoFix = Constr.mkCoFix
(**********************************************************************)
let eq_constr = Constr.equal
+let eq_constr_univs = Constr.eq_constr_univs
+let leq_constr_univs = Constr.leq_constr_univs
+let eq_constr_universes = Constr.eq_constr_universes
+let leq_constr_universes = Constr.leq_constr_universes
+let eq_constr_nounivs = Constr.eq_constr_nounivs
+
let kind_of_term = Constr.kind
let constr_ord = Constr.compare
let fold_constr = Constr.fold
+let map_puniverses = Constr.map_puniverses
let map_constr = Constr.map
let map_constr_with_binders = Constr.map_with_binders
let iter_constr = Constr.iter
@@ -195,9 +219,7 @@ let rec is_Type c = match kind_of_term c with
| Cast (c,_,_) -> is_Type c
| _ -> false
-let is_small = function
- | Prop _ -> true
- | _ -> false
+let is_small = Sorts.is_small
let iskind c = isprop c || is_Type c
@@ -649,6 +671,7 @@ let kind_of_type t = match kind_of_term t with
| Prod (na,t,c) -> ProdType (na, t, c)
| LetIn (na,b,t,c) -> LetInType (na, b, t, c)
| App (c,l) -> AtomicType (c, l)
- | (Rel _ | Meta _ | Var _ | Evar _ | Const _ | Case _ | Fix _ | CoFix _ | Ind _)
+ | (Rel _ | Meta _ | Var _ | Evar _ | Const _
+ | Proj _ | Case _ | Fix _ | CoFix _ | Ind _)
-> AtomicType (t,[||])
| (Lambda _ | Construct _) -> failwith "Not a type"
diff --git a/kernel/term.mli b/kernel/term.mli
index f2f5e8495d..2d3df6e1e8 100644
--- a/kernel/term.mli
+++ b/kernel/term.mli
@@ -24,6 +24,13 @@ type sorts = Sorts.t =
type sorts_family = Sorts.family = InProp | InSet | InType
+type 'a puniverses = 'a Univ.puniverses
+
+(** Simply type aliases *)
+type pconstant = constant puniverses
+type pinductive = inductive puniverses
+type pconstructor = constructor puniverses
+
type constr = Constr.constr
(** Alias types, for compatibility. *)
@@ -73,12 +80,13 @@ type ('constr, 'types) kind_of_term = ('constr, 'types) Constr.kind_of_term =
| Lambda of Name.t * 'types * 'constr
| LetIn of Name.t * 'constr * 'types * 'constr
| App of 'constr * 'constr array
- | Const of constant
- | Ind of inductive
- | Construct of constructor
+ | Const of constant puniverses
+ | Ind of inductive puniverses
+ | Construct of constructor puniverses
| Case of case_info * 'constr * 'constr * 'constr array
| Fix of ('constr, 'types) pfixpoint
| CoFix of ('constr, 'types) pcofixpoint
+ | Proj of constant * 'constr
type values = Constr.values
@@ -157,16 +165,16 @@ val decompose_app : constr -> constr * constr list
val decompose_appvect : constr -> constr * constr array
(** Destructs a constant *)
-val destConst : constr -> constant
+val destConst : constr -> constant puniverses
(** Destructs an existential variable *)
val destEvar : constr -> existential
(** Destructs a (co)inductive type *)
-val destInd : constr -> inductive
+val destInd : constr -> inductive puniverses
(** Destructs a constructor *)
-val destConstruct : constr -> constructor
+val destConstruct : constr -> constructor puniverses
(** Destructs a [match c as x in I args return P with ... |
Ci(...yij...) => ti | ... end] (or [let (..y1i..) := c as x in I args
@@ -397,8 +405,13 @@ val mkLambda : Name.t * types * constr -> constr
val mkLetIn : Name.t * constr * types * constr -> constr
val mkApp : constr * constr array -> constr
val mkConst : constant -> constr
+val mkProj : (constant * constr) -> constr
val mkInd : inductive -> constr
val mkConstruct : constructor -> constr
+val mkConstU : constant puniverses -> constr
+val mkIndU : inductive puniverses -> constr
+val mkConstructU : constructor puniverses -> constr
+val mkConstructUi : (pinductive * int) -> constr
val mkCase : case_info * constr * constr * constr array -> constr
val mkFix : fixpoint -> constr
val mkCoFix : cofixpoint -> constr
@@ -408,6 +421,26 @@ val mkCoFix : cofixpoint -> constr
val eq_constr : constr -> constr -> bool
(** Alias for [Constr.equal] *)
+(** [eq_constr_univs a b] [true, c] if [a] equals [b] modulo alpha, casts,
+ application grouping and the universe equalities in [c]. *)
+val eq_constr_univs : constr Univ.check_function
+
+(** [leq_constr_univs a b] [true, c] if [a] is convertible to [b] modulo
+ alpha, casts, application grouping and the universe inequalities in [c]. *)
+val leq_constr_univs : constr Univ.check_function
+
+(** [eq_constr_universes a b] [true, c] if [a] equals [b] modulo alpha, casts,
+ application grouping and the universe equalities in [c]. *)
+val eq_constr_universes : constr -> constr -> bool Univ.universe_constrained
+
+(** [leq_constr_universes a b] [true, c] if [a] is convertible to [b] modulo
+ alpha, casts, application grouping and the universe inequalities in [c]. *)
+val leq_constr_universes : constr -> constr -> bool Univ.universe_constrained
+
+(** [eq_constr_univs a b] [true, c] if [a] equals [b] modulo alpha, casts,
+ application grouping and ignoring universe instances. *)
+val eq_constr_nounivs : constr -> constr -> bool
+
val kind_of_term : constr -> (constr, types) kind_of_term
(** Alias for [Constr.kind] *)
@@ -424,6 +457,10 @@ val map_constr_with_binders :
('a -> 'a) -> ('a -> constr -> constr) -> 'a -> constr -> constr
(** Alias for [Constr.map_with_binders] *)
+val map_puniverses : ('a -> 'b) -> 'a puniverses -> 'b puniverses
+val univ_of_sort : sorts -> Univ.universe
+val sort_of_univ : Univ.universe -> sorts
+
val iter_constr : (constr -> unit) -> constr -> unit
(** Alias for [Constr.iter] *)
@@ -437,6 +474,8 @@ val compare_constr : (constr -> constr -> bool) -> constr -> constr -> bool
val hash_constr : constr -> int
(** Alias for [Constr.hash] *)
+(*********************************************************************)
+
val hcons_sorts : sorts -> sorts
(** Alias for [Constr.hashcons_sorts] *)
diff --git a/kernel/term_typing.ml b/kernel/term_typing.ml
index a084504dcb..9aa688fc71 100644
--- a/kernel/term_typing.ml
+++ b/kernel/term_typing.ml
@@ -22,29 +22,35 @@ open Declarations
open Environ
open Entries
open Typeops
+open Fast_typeops
-let constrain_type env j cst1 = function
- | `None ->
- make_polymorphic_if_constant_for_ind env j, cst1
+let debug = false
+let prerr_endline =
+ if debug then prerr_endline else fun _ -> ()
+
+let constrain_type env j poly = function
+ | `None -> j.uj_type
| `Some t ->
- let (tj,cst2) = infer_type env t in
- let (_,cst3) = judge_of_cast env j DEFAULTcast tj in
- assert (eq_constr t tj.utj_val);
- let cstrs = union_constraints (union_constraints cst1 cst2) cst3 in
- NonPolymorphicType t, cstrs
- | `SomeWJ (t, tj, cst2) ->
- let (_,cst3) = judge_of_cast env j DEFAULTcast tj in
- assert (eq_constr t tj.utj_val);
- let cstrs = union_constraints (union_constraints cst1 cst2) cst3 in
- NonPolymorphicType t, cstrs
+ let tj = infer_type env t in
+ let _ = judge_of_cast env j DEFAULTcast tj in
+ assert (eq_constr t tj.utj_val);
+ t
+ | `SomeWJ (t, tj) ->
+ let tj = infer_type env t in
+ let _ = judge_of_cast env j DEFAULTcast tj in
+ assert (eq_constr t tj.utj_val);
+ t
let map_option_typ = function None -> `None | Some x -> `Some x
-let translate_local_assum env t =
- let (j,cst) = infer env t in
- let t = Typeops.assumption_of_judgment env j in
- (t,cst)
-
+let local_constrain_type env j = function
+ | None ->
+ j.uj_type
+ | Some t ->
+ let tj = infer_type env t in
+ let _ = judge_of_cast env j DEFAULTcast tj in
+ assert (eq_constr t tj.utj_val);
+ t
(* Insertion of constants and parameters in environment. *)
@@ -59,19 +65,19 @@ let handle_side_effects env body side_eff =
if name.[i] == '.' || name.[i] == '#' then name.[i] <- '_' done;
Name (id_of_string name) in
let rec sub c i x = match kind_of_term x with
- | Const c' when eq_constant c c' -> mkRel i
+ | Const (c', _) when eq_constant c c' -> mkRel i
| _ -> map_constr_with_binders ((+) 1) (fun i x -> sub c i x) i x in
let fix_body (c,cb) t =
match cb.const_body with
| Undef _ -> assert false
| Def b ->
let b = Mod_subst.force_constr b in
- let b_ty = Typeops.type_of_constant_type env cb.const_type in
+ let b_ty = cb.const_type in
let t = sub c 1 (Vars.lift 1 t) in
mkLetIn (cname c, b, b_ty, t)
| OpaqueDef b ->
let b = Opaqueproof.force_proof b in
- let b_ty = Typeops.type_of_constant_type env cb.const_type in
+ let b_ty = cb.const_type in
let t = sub c 1 (Vars.lift 1 t) in
mkApp (mkLambda (cname c, b_ty, t), [|b|]) in
List.fold_right fix_body cbl t
@@ -86,46 +92,50 @@ let hcons_j j =
let feedback_completion_typecheck =
Option.iter (fun state_id -> Pp.feedback ~state_id Interface.Complete)
-let infer_declaration env dcl =
+let infer_declaration env kn dcl =
match dcl with
- | ParameterEntry (ctx,t,nl) ->
- let j, cst = infer env t in
+ | ParameterEntry (ctx,poly,(t,uctx),nl) ->
+ let env = push_context uctx env in
+ let j = infer env t in
let t = hcons_constr (Typeops.assumption_of_judgment env j) in
- Undef nl, NonPolymorphicType t, cst, false, ctx
+ Undef nl, t, None, poly, Future.from_val uctx, false, ctx
| DefinitionEntry ({ const_entry_type = Some typ;
const_entry_opaque = true } as c) ->
+ let env = push_context c.const_entry_universes env in
let { const_entry_body = body; const_entry_feedback = feedback_id } = c in
- let tyj, tycst = infer_type env typ in
+ let tyj = infer_type env typ in
let proofterm =
Future.chain ~greedy:true ~pure:true body (fun (body, side_eff) ->
let body = handle_side_effects env body side_eff in
- let j, cst = infer env body in
+ let j = infer env body in
let j = hcons_j j in
- let _typ, cst = constrain_type env j cst (`SomeWJ (typ,tyj,tycst)) in
+ let _typ = constrain_type env j c.const_entry_polymorphic (`SomeWJ (typ,tyj)) in
feedback_completion_typecheck feedback_id;
- j.uj_val, cst) in
+ j.uj_val, Univ.empty_constraint) in
let def = OpaqueDef (Opaqueproof.create proofterm) in
- let typ = NonPolymorphicType typ in
- def, typ, tycst, c.const_entry_inline_code, c.const_entry_secctx
+ def, typ, None, c.const_entry_polymorphic, Future.from_val c.const_entry_universes,
+ c.const_entry_inline_code, c.const_entry_secctx
| DefinitionEntry c ->
+ let env = push_context c.const_entry_universes env in
let { const_entry_type = typ; const_entry_opaque = opaque } = c in
let { const_entry_body = body; const_entry_feedback = feedback_id } = c in
let body, side_eff = Future.join body in
let body = handle_side_effects env body side_eff in
- let j, cst = infer env body in
+ let j = infer env body in
let j = hcons_j j in
- let typ, cst = constrain_type env j cst (map_option_typ typ) in
+ let typ = constrain_type env j c.const_entry_polymorphic (map_option_typ typ) in
feedback_completion_typecheck feedback_id;
let def = Def (Mod_subst.from_val j.uj_val) in
- def, typ, cst, c.const_entry_inline_code, c.const_entry_secctx
+ def, typ, None, c.const_entry_polymorphic,
+ Future.from_val c.const_entry_universes, c.const_entry_inline_code, c.const_entry_secctx
-let global_vars_set_constant_type env = function
- | NonPolymorphicType t -> global_vars_set env t
- | PolymorphicArity (ctx,_) ->
- Context.fold_rel_context
- (fold_rel_declaration
- (fun t c -> Id.Set.union (global_vars_set env t) c))
- ctx ~init:Id.Set.empty
+(* let global_vars_set_constant_type env = function *)
+(* | NonPolymorphicType t -> global_vars_set env t *)
+(* | PolymorphicArity (ctx,_) -> *)
+(* Context.fold_rel_context *)
+(* (fold_rel_declaration *)
+(* (fun t c -> Id.Set.union (global_vars_set env t) c)) *)
+(* ctx ~init:Id.Set.empty *)
let record_aux env s1 s2 =
let v =
@@ -137,7 +147,7 @@ let record_aux env s1 s2 =
let suggest_proof_using = ref (fun _ _ _ _ _ -> ())
let set_suggest_proof_using f = suggest_proof_using := f
-let build_constant_declaration kn env (def,typ,cst,inline_code,ctx) =
+let build_constant_declaration kn env (def,typ,proj,poly,univs,inline_code,ctx) =
let check declared inferred =
let mk_set l = List.fold_right Id.Set.add (List.map pi1 l) Id.Set.empty in
let inferred_set, declared_set = mk_set inferred, mk_set declared in
@@ -152,12 +162,14 @@ let build_constant_declaration kn env (def,typ,cst,inline_code,ctx) =
| None when not (List.is_empty context_ids) ->
(* No declared section vars, and non-empty section context:
we must look at the body NOW, if any *)
- let ids_typ = global_vars_set_constant_type env typ in
+ let ids_typ = global_vars_set env typ in
let ids_def = match def with
| Undef _ -> Idset.empty
| Def cs -> global_vars_set env (Mod_subst.force_constr cs)
| OpaqueDef lc ->
let vars = global_vars_set env (Opaqueproof.force_proof lc) in
+ (* we force so that cst are added to the env immediately after *)
+ ignore(Future.join univs);
!suggest_proof_using kn env vars ids_typ context_ids;
if !Flags.compilation_mode = Flags.BuildVo then
record_aux env ids_typ vars;
@@ -174,38 +186,50 @@ let build_constant_declaration kn env (def,typ,cst,inline_code,ctx) =
match def with
| Undef _ as x -> x (* nothing to check *)
| Def cs as x ->
- let ids_typ = global_vars_set_constant_type env typ in
+ let ids_typ = global_vars_set env typ in
let ids_def = global_vars_set env (Mod_subst.force_constr cs) in
let inferred = keep_hyps env (Idset.union ids_typ ids_def) in
check declared inferred;
x
| OpaqueDef lc -> (* In this case we can postpone the check *)
OpaqueDef (Opaqueproof.iter_direct_opaque (fun c ->
- let ids_typ = global_vars_set_constant_type env typ in
+ let ids_typ = global_vars_set env typ in
let ids_def = global_vars_set env c in
let inferred = keep_hyps env (Idset.union ids_typ ids_def) in
check declared inferred) lc) in
+ let tps =
+ match proj with
+ | None -> Cemitcodes.from_val (compile_constant_body env def)
+ | Some pb ->
+ Cemitcodes.from_val (compile_constant_body env (Def (Mod_subst.from_val pb.proj_body)))
+ in
{ const_hyps = hyps;
const_body = def;
const_type = typ;
- const_body_code = Cemitcodes.from_val (compile_constant_body env def);
- const_constraints = cst;
+ const_proj = proj;
+ const_body_code = tps;
+ const_polymorphic = poly;
+ const_universes = univs;
const_inline_code = inline_code }
+
(*s Global and local constant declaration. *)
let translate_constant env kn ce =
- build_constant_declaration kn env (infer_declaration env ce)
+ build_constant_declaration kn env (infer_declaration env (Some kn) ce)
+
+let translate_local_assum env t =
+ let j = infer env t in
+ let t = Typeops.assumption_of_judgment env j in
+ t
let translate_recipe env kn r =
- let def,typ,cst,inline_code,hyps = Cooking.cook_constant env r in
- build_constant_declaration kn env (def,typ,cst,inline_code,hyps)
+ build_constant_declaration kn env (Cooking.cook_constant env r)
let translate_local_def env id centry =
- let def,typ,cst,inline_code,ctx =
- infer_declaration env (DefinitionEntry centry) in
- let typ = type_of_constant_type env typ in
- def, typ, cst
+ let def,typ,proj,poly,univs,inline_code,ctx =
+ infer_declaration env None (DefinitionEntry centry) in
+ def, typ, univs
(* Insertion of inductive types. *)
diff --git a/kernel/term_typing.mli b/kernel/term_typing.mli
index b1c336ad9a..a2a35492e3 100644
--- a/kernel/term_typing.mli
+++ b/kernel/term_typing.mli
@@ -14,9 +14,9 @@ open Declarations
open Entries
val translate_local_def : env -> Id.t -> definition_entry ->
- constant_def * types * Univ.constraints
+ constant_def * types * constant_universes
-val translate_local_assum : env -> types -> types * constraints
+val translate_local_assum : env -> types -> types
(* returns the same definition_entry but with side effects turned into
* let-ins or beta redexes. it is meant to get a term out of a not yet
@@ -32,7 +32,9 @@ val translate_recipe : env -> constant -> Cooking.recipe -> constant_body
(** Internal functions, mentioned here for debug purpose only *)
-val infer_declaration : env -> constant_entry -> Cooking.result
+val infer_declaration : env -> constant option ->
+ constant_entry -> Cooking.result
+
val build_constant_declaration :
constant -> env -> Cooking.result -> constant_body
diff --git a/kernel/type_errors.ml b/kernel/type_errors.ml
index 42b93dd375..30dcbafe6b 100644
--- a/kernel/type_errors.ml
+++ b/kernel/type_errors.ml
@@ -42,12 +42,12 @@ type type_error =
| NotAType of unsafe_judgment
| BadAssumption of unsafe_judgment
| ReferenceVariables of identifier * constr
- | ElimArity of inductive * sorts_family list * constr * unsafe_judgment
+ | ElimArity of pinductive * sorts_family list * constr * unsafe_judgment
* (sorts_family * sorts_family * arity_error) option
| CaseNotInductive of unsafe_judgment
- | WrongCaseInfo of inductive * case_info
+ | WrongCaseInfo of pinductive * case_info
| NumberBranches of unsafe_judgment * int
- | IllFormedBranch of constr * constructor * constr * constr
+ | IllFormedBranch of constr * pconstructor * constr * constr
| Generalization of (Name.t * types) * unsafe_judgment
| ActualType of unsafe_judgment * types
| CantApplyBadType of
@@ -56,11 +56,12 @@ type type_error =
| IllFormedRecBody of guard_error * Name.t array * int * env * unsafe_judgment array
| IllTypedRecBody of
int * Name.t array * unsafe_judgment array * types array
+ | UnsatisfiedConstraints of Univ.constraints
exception TypeError of env * type_error
-let nfj {uj_val=c;uj_type=ct} =
- {uj_val=c;uj_type=nf_betaiota ct}
+let nfj env {uj_val=c;uj_type=ct} =
+ {uj_val=c;uj_type=nf_betaiota env ct}
let error_unbound_rel env n =
raise (TypeError (env, UnboundRel n))
@@ -84,11 +85,11 @@ let error_case_not_inductive env j =
raise (TypeError (env, CaseNotInductive j))
let error_number_branches env cj expn =
- raise (TypeError (env, NumberBranches (nfj cj,expn)))
+ raise (TypeError (env, NumberBranches (nfj env cj,expn)))
let error_ill_formed_branch env c i actty expty =
raise (TypeError (env,
- IllFormedBranch (c,i,nf_betaiota actty, nf_betaiota expty)))
+ IllFormedBranch (c,i,nf_betaiota env actty, nf_betaiota env expty)))
let error_generalization env nvar c =
raise (TypeError (env, Generalization (nvar,c)))
@@ -114,3 +115,5 @@ let error_elim_explain kp ki =
| InType, InSet -> StrongEliminationOnNonSmallType (* if Set impredicative *)
| _ -> WrongArity
+let error_unsatisfied_constraints env c =
+ raise (TypeError (env, UnsatisfiedConstraints c))
diff --git a/kernel/type_errors.mli b/kernel/type_errors.mli
index b9d8efbcde..09310b42b2 100644
--- a/kernel/type_errors.mli
+++ b/kernel/type_errors.mli
@@ -43,12 +43,12 @@ type type_error =
| NotAType of unsafe_judgment
| BadAssumption of unsafe_judgment
| ReferenceVariables of identifier * constr
- | ElimArity of inductive * sorts_family list * constr * unsafe_judgment
+ | ElimArity of pinductive * sorts_family list * constr * unsafe_judgment
* (sorts_family * sorts_family * arity_error) option
| CaseNotInductive of unsafe_judgment
- | WrongCaseInfo of inductive * case_info
+ | WrongCaseInfo of pinductive * case_info
| NumberBranches of unsafe_judgment * int
- | IllFormedBranch of constr * constructor * constr * constr
+ | IllFormedBranch of constr * pconstructor * constr * constr
| Generalization of (Name.t * types) * unsafe_judgment
| ActualType of unsafe_judgment * types
| CantApplyBadType of
@@ -57,6 +57,7 @@ type type_error =
| IllFormedRecBody of guard_error * Name.t array * int * env * unsafe_judgment array
| IllTypedRecBody of
int * Name.t array * unsafe_judgment array * types array
+ | UnsatisfiedConstraints of Univ.constraints
exception TypeError of env * type_error
@@ -71,14 +72,14 @@ val error_assumption : env -> unsafe_judgment -> 'a
val error_reference_variables : env -> identifier -> constr -> 'a
val error_elim_arity :
- env -> inductive -> sorts_family list -> constr -> unsafe_judgment ->
+ env -> pinductive -> sorts_family list -> constr -> unsafe_judgment ->
(sorts_family * sorts_family * arity_error) option -> 'a
val error_case_not_inductive : env -> unsafe_judgment -> 'a
val error_number_branches : env -> unsafe_judgment -> int -> 'a
-val error_ill_formed_branch : env -> constr -> constructor -> constr -> constr -> 'a
+val error_ill_formed_branch : env -> constr -> pconstructor -> constr -> constr -> 'a
val error_generalization : env -> Name.t * types -> unsafe_judgment -> 'a
@@ -98,3 +99,5 @@ val error_ill_typed_rec_body :
env -> int -> Name.t array -> unsafe_judgment array -> types array -> 'a
val error_elim_explain : sorts_family -> sorts_family -> arity_error
+
+val error_unsatisfied_constraints : env -> Univ.constraints -> 'a
diff --git a/kernel/typeops.ml b/kernel/typeops.ml
index 3400d8ce62..09fd4cc7f5 100644
--- a/kernel/typeops.ml
+++ b/kernel/typeops.ml
@@ -20,19 +20,21 @@ open Reduction
open Inductive
open Type_errors
-let conv_leq l2r = default_conv CUMUL ~l2r
+let conv_leq l2r env x y = default_conv CUMUL ~l2r env x y
let conv_leq_vecti env v1 v2 =
Array.fold_left2_i
- (fun i c t1 t2 ->
- let c' =
- try default_conv CUMUL env t1 t2
- with NotConvertible -> raise (NotConvertibleVect i) in
- union_constraints c c')
- empty_constraint
+ (fun i _ t1 t2 ->
+ try conv_leq false env t1 t2
+ with NotConvertible -> raise (NotConvertibleVect i))
+ ()
v1
v2
+let check_constraints cst env =
+ if Environ.check_constraints cst env then ()
+ else error_unsatisfied_constraints env cst
+
(* This should be a type (a priori without intension to be an assumption) *)
let type_judgment env j =
match kind_of_term(whd_betadeltaiota env j.uj_type) with
@@ -69,9 +71,9 @@ let judge_of_prop_contents = function
(* Type of Type(i). *)
let judge_of_type u =
- let uu = super u in
- { uj_val = mkType u;
- uj_type = mkType uu }
+ let uu = Universe.super u in
+ { uj_val = mkType u;
+ uj_type = mkType uu }
(*s Type of a de Bruijn index. *)
@@ -111,53 +113,32 @@ let check_hyps_inclusion env c sign =
(* Make a type polymorphic if an arity *)
-let extract_level env p =
- let _,c = dest_prod_assum env p in
- match kind_of_term c with Sort (Type u) -> Some u | _ -> None
-
-let extract_context_levels env l =
- let fold l (_, b, p) = match b with
- | None -> extract_level env p :: l
- | _ -> l
- in
- List.fold_left fold [] l
-
-let make_polymorphic_if_constant_for_ind env {uj_val = c; uj_type = t} =
- let params, ccl = dest_prod_assum env t in
- match kind_of_term ccl with
- | Sort (Type u) when isInd (fst (decompose_app (whd_betadeltaiota env c))) ->
- let param_ccls = extract_context_levels env params in
- let s = { poly_param_levels = param_ccls; poly_level = u} in
- PolymorphicArity (params,s)
- | _ ->
- NonPolymorphicType t
-
(* Type of constants *)
-let type_of_constant_knowing_parameters env t paramtyps =
- match t with
- | NonPolymorphicType t -> t
- | PolymorphicArity (sign,ar) ->
- let ctx = List.rev sign in
- let ctx,s = instantiate_universes env ctx ar paramtyps in
- mkArity (List.rev ctx,s)
-
-let type_of_constant_type env t =
- type_of_constant_knowing_parameters env t [||]
+let type_of_constant env cst = constant_type env cst
+let type_of_constant_in env cst = constant_type_in env cst
+let type_of_constant_knowing_parameters env t _ = t
+let type_of_constant_type env cst = cst
-let type_of_constant env cst =
- type_of_constant_type env (constant_type env cst)
+let judge_of_constant env (kn,u as cst) =
+ let c = mkConstU cst in
+ let cb = lookup_constant kn env in
+ let _ = check_hyps_inclusion env c cb.const_hyps in
+ let ty, cu = type_of_constant env cst in
+ let _ = Environ.check_constraints cu env in
+ (make_judge c ty)
-let judge_of_constant_knowing_parameters env cst jl =
- let c = mkConst cst in
+let type_of_projection env (cst,u) =
let cb = lookup_constant cst env in
- let _ = check_hyps_inclusion env c cb.const_hyps in
- let paramstyp = Array.map (fun j -> lazy j.uj_type) jl in
- let t = type_of_constant_knowing_parameters env cb.const_type paramstyp in
- make_judge c t
+ match cb.const_proj with
+ | Some pb ->
+ if cb.const_polymorphic then
+ let mib,_ = lookup_mind_specif env (pb.proj_ind,0) in
+ let subst = make_inductive_subst mib u in
+ Vars.subst_univs_constr subst pb.proj_type
+ else pb.proj_type
+ | None -> raise (Invalid_argument "type_of_projection: not a projection")
-let judge_of_constant env cst =
- judge_of_constant_knowing_parameters env cst [||]
(* Type of a lambda-abstraction. *)
@@ -184,26 +165,27 @@ let judge_of_letin env name defj typj j =
(* Type of an application. *)
let judge_of_apply env funj argjv =
- let len = Array.length argjv in
- let rec apply_rec n typ cst =
- if len <= n then
- { uj_val = mkApp (j_val funj, Array.map j_val argjv); uj_type = typ },
- cst
- else
- let hj = Array.unsafe_get argjv n in
- match kind_of_term (whd_betadeltaiota env typ) with
- | Prod (_,c1,c2) ->
- let c =
- try conv_leq false env hj.uj_type c1
- with NotConvertible ->
- error_cant_apply_bad_type env (n + 1, c1, hj.uj_type) funj argjv
- in
- let cst' = union_constraints cst c in
- apply_rec (n+1) (subst1 hj.uj_val c2) cst'
- | _ ->
- error_cant_apply_not_functional env funj argjv
+ let rec apply_rec n typ = function
+ | [] ->
+ { uj_val = mkApp (j_val funj, Array.map j_val argjv);
+ uj_type = typ }
+ | hj::restjl ->
+ (match kind_of_term (whd_betadeltaiota env typ) with
+ | Prod (_,c1,c2) ->
+ (try
+ let () = conv_leq false env hj.uj_type c1 in
+ apply_rec (n+1) (subst1 hj.uj_val c2) restjl
+ with NotConvertible ->
+ error_cant_apply_bad_type env
+ (n,c1, hj.uj_type)
+ funj argjv)
+
+ | _ ->
+ error_cant_apply_not_functional env funj argjv)
in
- apply_rec 0 funj.uj_type empty_constraint
+ apply_rec 1
+ funj.uj_type
+ (Array.to_list argjv)
(* Type of product *)
@@ -221,14 +203,14 @@ let sort_of_product env domsort rangsort =
rangsort
| _ ->
(* Rule is (Type_i,Set,Type_i) in the Set-predicative calculus *)
- Type (sup u1 type0_univ)
+ Type (Universe.sup Universe.type0 u1)
end
(* Product rule (Prop,Type_i,Type_i) *)
- | (Prop Pos, Type u2) -> Type (sup type0_univ u2)
+ | (Prop Pos, Type u2) -> Type (Universe.sup Universe.type0 u2)
(* Product rule (Prop,Type_i,Type_i) *)
| (Prop Null, Type _) -> rangsort
(* Product rule (Type_i,Type_i,Type_i) *)
- | (Type u1, Type u2) -> Type (sup u1 u2)
+ | (Type u1, Type u2) -> Type (Universe.sup u1 u2)
(* [judge_of_product env name (typ1,s1) (typ2,s2)] implements the rule
@@ -262,18 +244,17 @@ let judge_of_cast env cj k tj =
vm_conv CUMUL env cj.uj_type expected_type
| DEFAULTcast ->
mkCast (cj.uj_val, k, expected_type),
- conv_leq false env cj.uj_type expected_type
+ default_conv ~l2r:false CUMUL env cj.uj_type expected_type
| REVERTcast ->
cj.uj_val,
- conv_leq true env cj.uj_type expected_type
+ default_conv ~l2r:true CUMUL env cj.uj_type expected_type
| NATIVEcast ->
let sigma = Nativelambda.empty_evars in
mkCast (cj.uj_val, k, expected_type),
native_conv CUMUL sigma env cj.uj_type expected_type
in
- { uj_val = c;
- uj_type = expected_type },
- cst
+ { uj_val = c;
+ uj_type = expected_type }
with NotConvertible ->
error_actual_type env cj expected_type
@@ -291,50 +272,70 @@ let judge_of_cast env cj k tj =
the App case of execute; from this constraints, the expected
dynamic constraints of the form u<=v are enforced *)
-let judge_of_inductive_knowing_parameters env ind jl =
- let c = mkInd ind in
+(* let judge_of_inductive_knowing_parameters env ind jl = *)
+(* let c = mkInd ind in *)
+(* let (mib,mip) = lookup_mind_specif env ind in *)
+(* check_args env c mib.mind_hyps; *)
+(* let paramstyp = Array.map (fun j -> j.uj_type) jl in *)
+(* let t = in *)
+(* make_judge c t *)
+
+let judge_of_inductive env (ind,u as indu) =
+ let c = mkIndU indu in
let (mib,mip) = lookup_mind_specif env ind in
check_hyps_inclusion env c mib.mind_hyps;
- let paramstyp = Array.map (fun j -> lazy j.uj_type) jl in
- let t = Inductive.type_of_inductive_knowing_parameters env mip paramstyp in
- make_judge c t
-
-let judge_of_inductive env ind =
- judge_of_inductive_knowing_parameters env ind [||]
+ let t,cst = Inductive.constrained_type_of_inductive env ((mib,mip),u) in
+ check_constraints cst env;
+ (make_judge c t)
(* Constructors. *)
-let judge_of_constructor env c =
- let constr = mkConstruct c in
+let judge_of_constructor env (c,u as cu) =
+ let constr = mkConstructU cu in
let _ =
let ((kn,_),_) = c in
let mib = lookup_mind kn env in
check_hyps_inclusion env constr mib.mind_hyps in
let specif = lookup_mind_specif env (inductive_of_constructor c) in
- make_judge constr (type_of_constructor c specif)
+ let t,cst = constrained_type_of_constructor cu specif in
+ let () = check_constraints cst env in
+ (make_judge constr t)
(* Case. *)
-let check_branch_types env ind cj (lfj,explft) =
+let check_branch_types env (ind,u) cj (lfj,explft) =
try conv_leq_vecti env (Array.map j_type lfj) explft
with
NotConvertibleVect i ->
- error_ill_formed_branch env cj.uj_val (ind,i+1) lfj.(i).uj_type explft.(i)
+ error_ill_formed_branch env cj.uj_val ((ind,i+1),u) lfj.(i).uj_type explft.(i)
| Invalid_argument _ ->
error_number_branches env cj (Array.length explft)
let judge_of_case env ci pj cj lfj =
- let indspec =
+ let (pind, _ as indspec) =
try find_rectype env cj.uj_type
with Not_found -> error_case_not_inductive env cj in
- let _ = check_case_info env (fst indspec) ci in
- let (bty,rslty,univ) =
+ let _ = check_case_info env pind ci in
+ let (bty,rslty) =
type_case_branches env indspec pj cj.uj_val in
- let univ' = check_branch_types env (fst indspec) cj (lfj,bty) in
+ let () = check_branch_types env pind cj (lfj,bty) in
({ uj_val = mkCase (ci, (*nf_betaiota*) pj.uj_val, cj.uj_val,
Array.map j_val lfj);
- uj_type = rslty },
- union_constraints univ univ')
+ uj_type = rslty })
+
+let judge_of_projection env p cj =
+ let pb = lookup_projection p env in
+ let (ind,u), args =
+ try find_rectype env cj.uj_type
+ with Not_found -> error_case_not_inductive env cj
+ in
+ assert(eq_mind pb.proj_ind (fst ind));
+ let usubst = make_inductive_subst (fst (lookup_mind_specif env ind)) u in
+ let ty = Vars.subst_univs_constr usubst pb.Declarations.proj_type in
+ let ty = substl (cj.uj_val :: List.rev args) ty in
+ (* TODO: Universe polymorphism for projections *)
+ {uj_val = mkProj (p,cj.uj_val);
+ uj_type = ty}
(* Fixpoints. *)
@@ -352,104 +353,88 @@ let type_fixpoint env lna lar vdefj =
(************************************************************************)
(************************************************************************)
-(* This combinator adds the universe constraints both in the local
- graph and in the universes of the environment. This is to ensure
- that the infered local graph is satisfiable. *)
-let univ_combinator (cst,univ) (j,c') =
- (j,(union_constraints cst c', merge_constraints c' univ))
-
(* The typing machine. *)
(* ATTENTION : faudra faire le typage du contexte des Const,
Ind et Constructsi un jour cela devient des constructions
arbitraires et non plus des variables *)
-let rec execute env cstr cu =
+let rec execute env cstr =
match kind_of_term cstr with
(* Atomic terms *)
| Sort (Prop c) ->
- (judge_of_prop_contents c, cu)
-
+ judge_of_prop_contents c
+
| Sort (Type u) ->
- (judge_of_type u, cu)
+ judge_of_type u
| Rel n ->
- (judge_of_relative env n, cu)
+ judge_of_relative env n
| Var id ->
- (judge_of_variable env id, cu)
+ judge_of_variable env id
| Const c ->
- (judge_of_constant env c, cu)
+ judge_of_constant env c
+
+ | Proj (p, c) ->
+ let cj = execute env c in
+ judge_of_projection env p cj
(* Lambda calculus operators *)
| App (f,args) ->
- let (jl,cu1) = execute_array env args cu in
- let (j,cu2) =
- match kind_of_term f with
- | Ind ind ->
- (* Sort-polymorphism of inductive types *)
- judge_of_inductive_knowing_parameters env ind jl, cu1
- | Const cst ->
- (* Sort-polymorphism of constant *)
- judge_of_constant_knowing_parameters env cst jl, cu1
- | _ ->
- (* No sort-polymorphism *)
- execute env f cu1
- in
- univ_combinator cu2 (judge_of_apply env j jl)
+ let jl = execute_array env args in
+ let j = execute env f in
+ judge_of_apply env j jl
| Lambda (name,c1,c2) ->
- let (varj,cu1) = execute_type env c1 cu in
- let env1 = push_rel (name,None,varj.utj_val) env in
- let (j',cu2) = execute env1 c2 cu1 in
- (judge_of_abstraction env name varj j', cu2)
+ let varj = execute_type env c1 in
+ let env1 = push_rel (name,None,varj.utj_val) env in
+ let j' = execute env1 c2 in
+ judge_of_abstraction env name varj j'
| Prod (name,c1,c2) ->
- let (varj,cu1) = execute_type env c1 cu in
- let env1 = push_rel (name,None,varj.utj_val) env in
- let (varj',cu2) = execute_type env1 c2 cu1 in
- (judge_of_product env name varj varj', cu2)
+ let varj = execute_type env c1 in
+ let env1 = push_rel (name,None,varj.utj_val) env in
+ let varj' = execute_type env1 c2 in
+ judge_of_product env name varj varj'
| LetIn (name,c1,c2,c3) ->
- let (j1,cu1) = execute env c1 cu in
- let (j2,cu2) = execute_type env c2 cu1 in
- let (_,cu3) =
- univ_combinator cu2 (judge_of_cast env j1 DEFAULTcast j2) in
- let env1 = push_rel (name,Some j1.uj_val,j2.utj_val) env in
- let (j',cu4) = execute env1 c3 cu3 in
- (judge_of_letin env name j1 j2 j', cu4)
+ 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 j' = execute env1 c3 in
+ judge_of_letin env name j1 j2 j'
| Cast (c,k, t) ->
- let (cj,cu1) = execute env c cu in
- let (tj,cu2) = execute_type env t cu1 in
- univ_combinator cu2
- (judge_of_cast env cj k tj)
+ let cj = execute env c in
+ let tj = execute_type env t in
+ judge_of_cast env cj k tj
(* Inductive types *)
| Ind ind ->
- (judge_of_inductive env ind, cu)
+ judge_of_inductive env ind
| Construct c ->
- (judge_of_constructor env c, cu)
+ judge_of_constructor env c
| Case (ci,p,c,lf) ->
- let (cj,cu1) = execute env c cu in
- let (pj,cu2) = execute env p cu1 in
- let (lfj,cu3) = execute_array env lf cu2 in
- univ_combinator cu3
- (judge_of_case env ci pj cj lfj)
+ let cj = execute env c in
+ let pj = execute env p in
+ let lfj = execute_array env lf in
+ judge_of_case env ci pj cj lfj
| Fix ((vn,i as vni),recdef) ->
- let ((fix_ty,recdef'),cu1) = execute_recdef env recdef i cu in
- let fix = (vni,recdef') in
+ let (fix_ty,recdef') = execute_recdef env recdef i in
+ let fix = (vni,recdef') in
check_fix env fix;
- (make_judge (mkFix fix) fix_ty, cu1)
-
+ make_judge (mkFix fix) fix_ty
+
| CoFix (i,recdef) ->
- let ((fix_ty,recdef'),cu1) = execute_recdef env recdef i cu in
- let cofix = (i,recdef') in
+ let (fix_ty,recdef') = execute_recdef env recdef i in
+ let cofix = (i,recdef') in
check_cofix env cofix;
- (make_judge (mkCoFix cofix) fix_ty, cu1)
-
+ (make_judge (mkCoFix cofix) fix_ty)
+
(* Partial proofs: unsupported by the kernel *)
| Meta _ ->
anomaly (Pp.str "the kernel does not support metavariables")
@@ -457,61 +442,53 @@ let rec execute env cstr cu =
| Evar _ ->
anomaly (Pp.str "the kernel does not support existential variables")
-and execute_type env constr cu =
- let (j,cu1) = execute env constr cu in
- (type_judgment env j, cu1)
+and execute_type env constr =
+ let j = execute env constr in
+ type_judgment env j
-and execute_recdef env (names,lar,vdef) i cu =
- let (larj,cu1) = execute_array env lar cu in
+and execute_recdef env (names,lar,vdef) i =
+ let larj = execute_array env lar in
let lara = Array.map (assumption_of_judgment env) larj in
let env1 = push_rec_types (names,lara,vdef) env in
- let (vdefj,cu2) = execute_array env1 vdef cu1 in
+ let vdefj = execute_array env1 vdef in
let vdefv = Array.map j_val vdefj in
- let cst = type_fixpoint env1 names lara vdefj in
- univ_combinator cu2
- ((lara.(i),(names,lara,vdefv)),cst)
+ let () = type_fixpoint env1 names lara vdefj in
+ (lara.(i),(names,lara,vdefv))
-and execute_array env = Array.fold_map' (execute env)
+and execute_array env = Array.map (execute env)
(* Derived functions *)
let infer env constr =
- let (j,(cst,_)) =
- execute env constr (empty_constraint, universes env) in
- assert (eq_constr j.uj_val constr);
- (j, cst)
+ let j = execute env constr in
+ assert (eq_constr j.uj_val constr);
+ j
+
+(* let infer_key = Profile.declare_profile "infer" *)
+(* let infer = Profile.profile2 infer_key infer *)
let infer_type env constr =
- let (j,(cst,_)) =
- execute_type env constr (empty_constraint, universes env) in
- (j, cst)
+ let j = execute_type env constr in
+ j
let infer_v env cv =
- let (jv,(cst,_)) =
- execute_array env cv (empty_constraint, universes env) in
- (jv, cst)
+ let jv = execute_array env cv in
+ jv
(* Typing of several terms. *)
let infer_local_decl env id = function
| LocalDef c ->
- let (j,cst) = infer env c in
- (Name id, Some j.uj_val, j.uj_type), cst
+ let j = infer env c in
+ (Name id, Some j.uj_val, j.uj_type)
| LocalAssum c ->
- let (j,cst) = infer env c in
- (Name id, None, assumption_of_judgment env j), cst
+ let j = infer env c in
+ (Name id, None, assumption_of_judgment env j)
let infer_local_decls env decls =
let rec inferec env = function
| (id, d) :: l ->
- let env, l, cst1 = inferec env l in
- let d, cst2 = infer_local_decl env id d in
- push_rel d env, add_rel_decl d l, union_constraints cst1 cst2
- | [] -> env, empty_rel_context, empty_constraint in
+ let (env, l) = inferec env l in
+ let d = infer_local_decl env id d in
+ (push_rel d env, add_rel_decl d l)
+ | [] -> (env, empty_rel_context) in
inferec env decls
-
-(* Exported typing functions *)
-
-let typing env c =
- let (j,cst) = infer env c in
- let _ = add_constraints cst env in
- j
diff --git a/kernel/typeops.mli b/kernel/typeops.mli
index d6484e8233..6bc1ce496b 100644
--- a/kernel/typeops.mli
+++ b/kernel/typeops.mli
@@ -14,15 +14,21 @@ open Environ
open Entries
open Declarations
-(** {6 Typing functions (not yet tagged as safe) } *)
+(** {6 Typing functions (not yet tagged as safe) }
-val infer : env -> constr -> unsafe_judgment * constraints
-val infer_v : env -> constr array -> unsafe_judgment array * constraints
-val infer_type : env -> types -> unsafe_type_judgment * constraints
+ They return unsafe judgments that are "in context" of a set of
+ (local) universe variables (the ones that appear in the term)
+ and associated constraints. In case of polymorphic definitions,
+ these variables and constraints will be generalized.
+ *)
+
+
+val infer : env -> constr -> unsafe_judgment
+val infer_v : env -> constr array -> unsafe_judgment array
+val infer_type : env -> types -> unsafe_type_judgment
val infer_local_decls :
- env -> (Id.t * local_entry) list
- -> env * rel_context * constraints
+ env -> (Id.t * local_entry) list -> (env * rel_context)
(** {6 Basic operations of the typing machine. } *)
@@ -45,21 +51,25 @@ val judge_of_relative : env -> int -> unsafe_judgment
val judge_of_variable : env -> variable -> unsafe_judgment
(** {6 type of a constant } *)
-val judge_of_constant : env -> constant -> unsafe_judgment
+val judge_of_constant : env -> constant puniverses -> unsafe_judgment
+
+(* val judge_of_constant_knowing_parameters : *)
+(* env -> constant -> unsafe_judgment array -> unsafe_judgment *)
-val judge_of_constant_knowing_parameters :
- env -> constant -> unsafe_judgment array -> unsafe_judgment
+val judge_of_projection : env -> Names.projection -> unsafe_judgment -> unsafe_judgment
(** {6 Type of application. } *)
val judge_of_apply :
env -> unsafe_judgment -> unsafe_judgment array
- -> unsafe_judgment * constraints
+ -> unsafe_judgment
(** {6 Type of an abstraction. } *)
val judge_of_abstraction :
env -> Name.t -> unsafe_type_judgment -> unsafe_judgment
-> unsafe_judgment
+val sort_of_product : env -> sorts -> sorts -> sorts
+
(** {6 Type of a product. } *)
val judge_of_product :
env -> Name.t -> unsafe_type_judgment -> unsafe_type_judgment
@@ -73,37 +83,35 @@ val judge_of_letin :
(** {6 Type of a cast. } *)
val judge_of_cast :
env -> unsafe_judgment -> cast_kind -> unsafe_type_judgment ->
- unsafe_judgment * constraints
+ unsafe_judgment
(** {6 Inductive types. } *)
-val judge_of_inductive : env -> inductive -> unsafe_judgment
+val judge_of_inductive : env -> inductive puniverses -> unsafe_judgment
-val judge_of_inductive_knowing_parameters :
- env -> inductive -> unsafe_judgment array -> unsafe_judgment
+(* val judge_of_inductive_knowing_parameters : *)
+(* env -> inductive -> unsafe_judgment array -> unsafe_judgment *)
-val judge_of_constructor : env -> constructor -> unsafe_judgment
+val judge_of_constructor : env -> constructor puniverses -> unsafe_judgment
(** {6 Type of Cases. } *)
val judge_of_case : env -> case_info
-> unsafe_judgment -> unsafe_judgment -> unsafe_judgment array
- -> unsafe_judgment * constraints
+ -> unsafe_judgment
(** Typecheck general fixpoint (not checking guard conditions) *)
val type_fixpoint : env -> Name.t array -> types array
- -> unsafe_judgment array -> constraints
-
-(** Kernel safe typing but applicable to partial proofs *)
-val typing : env -> constr -> unsafe_judgment
+ -> unsafe_judgment array -> unit
-val type_of_constant : env -> constant -> types
+val type_of_constant : env -> constant puniverses -> types constrained
val type_of_constant_type : env -> constant_type -> types
+val type_of_projection : env -> Names.projection puniverses -> types
+
+val type_of_constant_in : env -> constant puniverses -> types
+
val type_of_constant_knowing_parameters :
env -> constant_type -> types Lazy.t array -> types
-(** Make a type polymorphic if an arity *)
-val make_polymorphic_if_constant_for_ind : env -> unsafe_judgment ->
- constant_type
diff --git a/kernel/univ.ml b/kernel/univ.ml
index 551d740434..8322a7e962 100644
--- a/kernel/univ.ml
+++ b/kernel/univ.ml
@@ -29,13 +29,56 @@ open Util
union-find algorithm. The assertions $<$ and $\le$ are represented by
adjacency lists *)
-module UniverseLevel = struct
+module Level = struct
open Names
type t =
+ | Prop
| Set
- | Level of int * DirPath.t
+ | Level of int * Names.DirPath.t
+ type _t = t
+
+ (* Hash-consing *)
+
+ module Hunivlevel =
+ Hashcons.Make(
+ struct
+ type t = _t
+ type u = Names.DirPath.t -> Names.DirPath.t
+ let hashcons hdir = function
+ | Prop as x -> x
+ | Set as x -> x
+ | Level (n,d) -> Level (n,hdir d)
+ let equal l1 l2 =
+ l1 == l2 ||
+ match l1,l2 with
+ | Prop, Prop -> true
+ | Set, Set -> true
+ | Level (n,d), Level (n',d') ->
+ n == n' && d == d'
+ | _ -> false
+ let hash = Hashtbl.hash
+ end)
+
+ let hcons = Hashcons.simple_hcons Hunivlevel.generate Names.DirPath.hcons
+
+ let make m n = hcons (Level (n, m))
+
+ let set = hcons Set
+ let prop = hcons Prop
+
+ let is_small = function
+ | Level _ -> false
+ | _ -> true
+
+ let is_prop = function
+ | Prop -> true
+ | _ -> false
+
+ let is_set = function
+ | Set -> true
+ | _ -> false
(* A specialized comparison function: we compare the [int] part first.
This way, most of the time, the [DirPath.t] part is not considered.
@@ -49,6 +92,9 @@ module UniverseLevel = struct
if u == v then 0
else
(match u,v with
+ | Prop,Prop -> 0
+ | Prop, _ -> -1
+ | _, Prop -> 1
| Set, Set -> 0
| Set, _ -> -1
| _, Set -> 1
@@ -63,27 +109,385 @@ module UniverseLevel = struct
Int.equal i1 i2 && DirPath.equal dp1 dp2
| _ -> false
- let hash = function
- | Set -> 0
- | Level (i, dp) ->
- Hashset.Combine.combine (Int.hash i) (DirPath.hash dp)
-
- let make m n = Level (n, m)
+ let eq u v = u == v
+ let leq u v = compare u v <= 0
let to_string = function
+ | Prop -> "Prop"
| Set -> "Set"
- | Level (n,d) -> DirPath.to_string d^"."^string_of_int n
+ | Level (n,d) -> Names.DirPath.to_string d^"."^string_of_int n
+
+ let pr u = str (to_string u)
+
+ let apart u v =
+ match u, v with
+ | Prop, Set | Set, Prop -> true
+ | _ -> false
+
end
-module UniverseLMap = Map.Make (UniverseLevel)
-module UniverseLSet = Set.Make (UniverseLevel)
+let pr_universe_level_list l =
+ prlist_with_sep spc Level.pr l
+
+module LSet = struct
+ module M = Set.Make (Level)
+ include M
+
+ let pr s =
+ str"{" ++ pr_universe_level_list (elements s) ++ str"}"
-type universe_level = UniverseLevel.t
+ let of_list l =
+ List.fold_left (fun acc x -> add x acc) empty l
-let compare_levels = UniverseLevel.compare
+ let of_array l =
+ Array.fold_left (fun acc x -> add x acc) empty l
+end
+
+module LMap = struct
+ module M = Map.Make (Level)
+ include M
+
+ let union l r =
+ merge (fun k l r ->
+ match l, r with
+ | Some _, _ -> l
+ | _, _ -> r) l r
+
+ let subst_union l r =
+ merge (fun k l r ->
+ match l, r with
+ | Some (Some _), _ -> l
+ | Some None, None -> l
+ | _, _ -> r) l r
+
+ let diff ext orig =
+ fold (fun u v acc ->
+ if mem u orig then acc
+ else add u v acc)
+ ext empty
+
+ let elements = bindings
+ let of_set s d =
+ LSet.fold (fun u -> add u d) s
+ empty
+
+ let of_list l =
+ List.fold_left (fun m (u, v) -> add u v m) empty l
+
+ let universes m =
+ fold (fun u _ acc -> LSet.add u acc) m LSet.empty
+
+ let pr f m =
+ h 0 (prlist_with_sep fnl (fun (u, v) ->
+ Level.pr u ++ f v) (elements m))
+
+ let find_opt t m =
+ try Some (find t m)
+ with Not_found -> None
+end
+
+type universe_level = Level.t
+
+module LList = struct
+ type t = Level.t list
+ type _t = t
+ module Huniverse_level_list =
+ Hashcons.Make(
+ struct
+ type t = _t
+ type u = universe_level -> universe_level
+ let hashcons huc s =
+ List.fold_right (fun x a -> huc x :: a) s []
+ let equal s s' = List.for_all2eq (==) s s'
+ let hash = Hashtbl.hash
+ end)
+
+ let hcons =
+ Hashcons.simple_hcons Huniverse_level_list.generate Level.hcons
+
+ let empty = hcons []
+ let eq l l' = l == l' ||
+ (try List.for_all2 Level.eq l l'
+ with Invalid_argument _ -> false)
+
+ let levels =
+ List.fold_left (fun s x -> LSet.add x s) LSet.empty
+
+end
+
+type universe_level_list = universe_level list
+
+type universe_level_subst_fn = universe_level -> universe_level
+
+type universe_set = LSet.t
+type 'a universe_map = 'a LMap.t
+
+let compare_levels = Level.compare
+let eq_levels = Level.eq
+
+module Hashconsing = struct
+ module Uid = struct
+ type t = int
+
+ let make_maker () =
+ let _id = ref ~-1 in
+ ((fun () -> incr _id;!_id),
+ (fun () -> !_id),
+ (fun i -> _id := i))
+
+ let dummy = -1
+
+ external to_int : t -> int = "%identity"
+
+
+ external of_int : int -> t= "%identity"
+ end
+
+ module Hcons = struct
+
+ module type SA =
+ sig
+ type data
+ type t
+ val make : data -> t
+ val node : t -> data
+ val hash : t -> int
+ val uid : t -> Uid.t
+ val equal : t -> t -> bool
+ val stats : unit -> unit
+ val init : unit -> unit
+ end
+
+ module type S =
+ sig
+
+ type data
+ type t = private { id : Uid.t;
+ key : int;
+ node : data }
+ val make : data -> t
+ val node : t -> data
+ val hash : t -> int
+ val uid : t -> Uid.t
+ val equal : t -> t -> bool
+ val stats : unit -> unit
+ val init : unit -> unit
+ end
+
+ module Make (H : Hashtbl.HashedType) : S with type data = H.t =
+ struct
+ let uid_make,uid_current,uid_set = Uid.make_maker()
+ type data = H.t
+ type t = { id : Uid.t;
+ key : int;
+ node : data }
+ let node t = t.node
+ let uid t = t.id
+ let hash t = t.key
+ let equal t1 t2 = t1 == t2
+ module WH = Weak.Make( struct
+ type _t = t
+ type t = _t
+ let hash = hash
+ let equal a b = a == b || H.equal a.node b.node
+ end)
+ let pool = WH.create 491
+
+ exception Found of Uid.t
+ let total_count = ref 0
+ let miss_count = ref 0
+ let init () =
+ total_count := 0;
+ miss_count := 0
+
+ let make x =
+ incr total_count;
+ let cell = { id = Uid.dummy; key = H.hash x; node = x } in
+ try
+ WH.find pool cell
+ with
+ | Not_found ->
+ let cell = { cell with id = uid_make(); } in
+ incr miss_count;
+ WH.add pool cell;
+ cell
+
+ exception Found of t
+
+ let stats () = ()
+ end
+ end
+ module HList = struct
+
+ module type S = sig
+ type elt
+ type 'a node = Nil | Cons of elt * 'a
+
+ module rec Node :
+ sig
+ include Hcons.S with type data = Data.t
+ end
+ and Data : sig
+ include Hashtbl.HashedType with type t = Node.t node
+ end
+ type data = Data.t
+ type t = Node.t
+ val hash : t -> int
+ val uid : t -> Uid.t
+ val make : data -> t
+ val equal : t -> t -> bool
+ val nil : t
+ val is_nil : t -> bool
+ val tip : elt -> t
+ val node : t -> t node
+ val cons : (* ?sorted:bool -> *) elt -> t -> t
+ val hd : t -> elt
+ val tl : t -> t
+ val fold : (elt -> 'a -> 'a) -> t -> 'a -> 'a
+ val map : (elt -> elt) -> t -> t
+ val iter : (elt -> 'a) -> t -> unit
+ val exists : (elt -> bool) -> t -> bool
+ val for_all : (elt -> bool) -> t -> bool
+ val rev : t -> t
+ val rev_map : (elt -> elt) -> t -> t
+ val length : t -> int
+ val mem : elt -> t -> bool
+ val remove : elt -> t -> t
+ val stats : unit -> unit
+ val init : unit -> unit
+ val to_list : t -> elt list
+ val compare : (elt -> elt -> int) -> t -> t -> int
+ end
+
+ module Make (H : Hcons.SA) : S with type elt = H.t =
+ struct
+ type elt = H.t
+ type 'a node = Nil | Cons of elt * 'a
+ module rec Node : Hcons.S with type data = Data.t = Hcons.Make (Data)
+ and Data : Hashtbl.HashedType with type t = Node.t node =
+ struct
+ type t = Node.t node
+ let equal x y =
+ match x,y with
+ | _,_ when x==y -> true
+ | Cons (a,aa), Cons(b,bb) -> (aa==bb) && (H.equal a b)
+ | _ -> false
+ let hash = function
+ | Nil -> 0
+ | Cons(a,aa) -> 17 + 65599 * (Uid.to_int (H.uid a)) + 491 * (Uid.to_int aa.Node.id)
+ end
+
+ type data = Data.t
+ type t = Node.t
+ let make = Node.make
+ let node x = x.Node.node
+ let hash x = x.Node.key
+ let equal = Node.equal
+ let uid x= x.Node.id
+ let nil = Node.make Nil
+ let stats = Node.stats
+ let init = Node.init
+
+ let is_nil =
+ function { Node.node = Nil } -> true | _ -> false
+
+ (* doing sorted insertion allows to make
+ better use of hash consing *)
+ let rec sorted_cons e l =
+ match l.Node.node with
+ | Nil -> Node.make (Cons(e, l))
+ | Cons (x, ll) ->
+ if H.uid e < H.uid x
+ then Node.make (Cons(e, l))
+ else Node.make (Cons(x, sorted_cons e ll))
+
+ let cons e l =
+ Node.make(Cons(e, l))
+
+ let tip e = Node.make (Cons(e, nil))
+
+ (* let cons ?(sorted=true) e l = *)
+ (* if sorted then sorted_cons e l else cons e l *)
+
+ let hd = function { Node.node = Cons(a,_) } -> a | _ -> failwith "hd"
+ let tl = function { Node.node = Cons(_,a) } -> a | _ -> failwith "tl"
+
+ let fold f l acc =
+ let rec loop acc l = match l.Node.node with
+ | Nil -> acc
+ | Cons (a, aa) -> loop (f a acc) aa
+ in
+ loop acc l
+
+ let map f l =
+ let rec loop l = match l.Node.node with
+ | Nil -> nil
+ | Cons(a, aa) -> cons (f a) (loop aa)
+ in
+ loop l
+
+ let iter f l =
+ let rec loop l = match l.Node.node with
+ | Nil -> ()
+ | Cons(a,aa) -> (f a);(loop aa)
+ in
+ loop l
+
+ let exists f l =
+ let rec loop l = match l.Node.node with
+ | Nil -> false
+ | Cons(a,aa) -> f a || loop aa
+ in
+ loop l
+
+ let for_all f l =
+ let rec loop l = match l.Node.node with
+ | Nil -> true
+ | Cons(a,aa) -> f a && loop aa
+ in
+ loop l
+
+ let to_list l =
+ let rec loop l = match l.Node.node with
+ | Nil -> []
+ | Cons(a,aa) -> a :: loop aa
+ in
+ loop l
+
+ let remove x l =
+ let rec loop l = match l.Node.node with
+ | Nil -> l
+ | Cons(a,aa) ->
+ if H.equal a x then aa
+ else cons a (loop aa)
+ in
+ loop l
+
+ let rev l = fold cons l nil
+ let rev_map f l = fold (fun x acc -> cons (f x) acc) l nil
+ let length l = fold (fun _ c -> c+1) l 0
+ let rec mem e l =
+ match l.Node.node with
+ | Nil -> false
+ | Cons (x, ll) -> x == e || mem e ll
+
+ let rec compare cmp l1 l2 =
+ if l1 == l2 then 0 else
+ match node l1, node l2 with
+ | Nil, Nil -> 0
+ | _, Nil -> 1
+ | Nil, _ -> -1
+ | Cons (x1,l1), Cons(x2,l2) ->
+ (match cmp x1 x2 with
+ | 0 -> compare cmp l1 l2
+ | c -> c)
+
+ end
+ end
+end
(* An algebraic universe [universe] is either a universe variable
- [UniverseLevel.t] or a formal universe known to be greater than some
+ [Level.t] or a formal universe known to be greater than some
universe variables and strictly greater than some (other) universe
variables
@@ -96,158 +500,346 @@ let compare_levels = UniverseLevel.compare
module Universe =
struct
- type t =
- | Atom of UniverseLevel.t
- | Max of UniverseLevel.t list * UniverseLevel.t list
+ (* Invariants: non empty, sorted and without duplicates *)
+
+ module Expr =
+ struct
+ type t = Level.t * int
+ type _t = t
+
+ (* Hashing of expressions *)
+ module ExprHash =
+ struct
+ type t = _t
+ type u = Level.t -> Level.t
+ let hashcons hdir (b,n as x) =
+ let b' = hdir b in
+ if b' == b then x else (b',n)
+ let equal l1 l2 =
+ l1 == l2 ||
+ match l1,l2 with
+ | (b,n), (b',n') -> b == b' && n == n'
+ let hash = Hashtbl.hash
+
+ end
+
+ module HExpr =
+ struct
+
+ include Hashcons.Make(ExprHash)
+
+ type data = t
+ type node = t
+
+ let make =
+ Hashcons.simple_hcons generate Level.hcons
+ external node : node -> data = "%identity"
+ let hash = ExprHash.hash
+ let uid = hash
+ let equal x y = x == y
+ let stats _ = ()
+ let init _ = ()
+ end
+
+ let hcons = HExpr.make
+
+ let make l = hcons (l, 0)
+
+ let compare u v =
+ if u == v then 0
+ else
+ let (x, n) = u and (x', n') = v in
+ if Int.equal n n' then Level.compare x x'
+ else n - n'
+
+ let prop = make Level.prop
+ let set = make Level.set
+ let type1 = hcons (Level.set, 1)
+
+ let is_prop = function
+ | (l,0) -> Level.is_prop l
+ | _ -> false
+
+ let is_set = function
+ | (l,0) -> Level.is_set l
+ | _ -> false
+
+ let is_type1 = function
+ | (l,1) -> Level.is_set l
+ | _ -> false
+
+ let is_small = function
+ | (l, 0) -> Level.is_small l
+ | _ -> false
+
+ (* let eq (u,n) (v,n') = *)
+ (* Int.equal n n' && Level.eq u v *)
+ let eq x y = x == y
+
+ let leq (u,n) (v,n') =
+ let cmp = Level.compare u v in
+ if Int.equal cmp 0 then n <= n'
+ else if n <= n' then
+ (Level.is_prop u && Level.is_small v)
+ else false
+
+ let successor (u,n) =
+ if Level.is_prop u then type1
+ else hcons (u, n + 1)
+
+ let addn k (u,n as x) =
+ if k = 0 then x
+ else if Level.is_prop u then
+ hcons (Level.set,n+k)
+ else hcons (u,n+k)
+
+ let super (u,n as x) (v,n' as y) =
+ let cmp = Level.compare u v in
+ if Int.equal cmp 0 then
+ if n < n' then Inl true
+ else Inl false
+ else if is_prop x then Inl true
+ else if is_prop y then Inl false
+ else Inr cmp
+
+ let to_string (v, n) =
+ if Int.equal n 0 then Level.to_string v
+ else Level.to_string v ^ "+" ^ string_of_int n
+
+ let pr x = str(to_string x)
+
+ let level = function
+ | (v,0) -> Some v
+ | _ -> None
+
+ let get_level (v,n) = v
+
+ let map f (v, n as x) =
+ let v' = f v in
+ if v' == v then x
+ else if Level.is_prop v' && n != 0 then
+ hcons (Level.set, n)
+ else hcons (v', n)
+
+ end
+
+ module Hunivelt = struct
+ let node x = x
+ let make x = x
+ end
+
+ (* module Hunivelt = Hashconsing.Hcons.Make( *)
+ (* struct *)
+ (* type t = Expr.t *)
+ (* let equal l1 l2 = *)
+ (* l1 == l2 || *)
+ (* match l1,l2 with *)
+ (* | (b,n), (b',n') -> b == b' && n == n' *)
+ (* let hash = Hashtbl.hash *)
+ (* end) *)
+
+ let compare_expr n m = Expr.compare (Hunivelt.node n) (Hunivelt.node m)
+ let pr_expr n = Expr.pr (Hunivelt.node n)
+
+ module Huniv = Hashconsing.HList.Make(Expr.HExpr)
+ type t = Huniv.t
+ open Huniv
+
+ let eq x y = x == y (* Huniv.equal *)
let compare u1 u2 =
- if u1 == u2 then 0 else
- match u1, u2 with
- | Atom l1, Atom l2 -> UniverseLevel.compare l1 l2
- | Max (lt1, le1), Max (lt2, le2) ->
- let c = List.compare UniverseLevel.compare lt1 lt2 in
- if Int.equal c 0 then
- List.compare UniverseLevel.compare le1 le2
- else c
- | Atom _, Max _ -> -1
- | Max _, Atom _ -> 1
-
- let equal u1 u2 = Int.equal (compare u1 u2) 0
-
- let make l = Atom l
-
- open Hashset.Combine
-
- let rec hash_list accu = function
- | [] -> 0
- | u :: us ->
- let accu = combine (UniverseLevel.hash u) accu in
- hash_list accu us
-
- let hash = function
- | Atom u -> combinesmall 1 (UniverseLevel.hash u)
- | Max (lt, le) ->
- let hlt = hash_list 0 lt in
- let hle = hash_list 0 le in
- combinesmall 2 (combine hlt hle)
+ if eq u1 u2 then 0 else
+ Huniv.compare compare_expr u1 u2
+
+ let hcons_unique = Huniv.make
+ let normalize x = x
+ (* let hcons_unique x = x *)
+ let hcons x = hcons_unique (normalize x)
+
+ let make l = Huniv.tip (Hunivelt.make (Expr.make l))
+ let tip x = Huniv.tip (Hunivelt.make x)
+
+ let equal_universes x y =
+ x == y
+(* then true *)
+(* else *)
+(* (\* Consider lists as sets, i.e. up to reordering, *)
+(* they are already without duplicates thanks to normalization. *\) *)
+(* CList.eq_set x' y' *)
+
+ let pr l = match node l with
+ | Cons (u, n) when is_nil n -> Expr.pr (Hunivelt.node u)
+ | _ ->
+ str "max(" ++ hov 0
+ (prlist_with_sep pr_comma Expr.pr (List.map Hunivelt.node (to_list l))) ++
+ str ")"
+
+ let atom l = match node l with
+ | Cons (l, n) when is_nil n -> Some l
+ | _ -> None
+
+ let level l = match node l with
+ | Cons (l, n) when is_nil n -> Expr.level (Hunivelt.node l)
+ | _ -> None
+
+ let levels l =
+ fold (fun x acc -> LSet.add (Expr.get_level (Hunivelt.node x)) acc) l LSet.empty
+
+ let is_small u =
+ match level (normalize u) with
+ | Some l -> Level.is_small l
+ | _ -> false
-end
+ (* The lower predicative level of the hierarchy that contains (impredicative)
+ Prop and singleton inductive types *)
+ let type0m = tip Expr.prop
-open Universe
+ (* The level of sets *)
+ let type0 = tip Expr.set
+
+ (* When typing [Prop] and [Set], there is no constraint on the level,
+ hence the definition of [type1_univ], the type of [Prop] *)
+ let type1 = tip (Expr.successor Expr.set)
+
+ let is_type0m u =
+ match level u with
+ | Some l -> Level.is_prop l
+ | _ -> false
+
+ let is_type0 u =
+ match level u with
+ | Some l -> Level.is_set l
+ | _ -> false
+
+ let is_type1 u =
+ match node u with
+ | Cons (l, n) when is_nil n -> Expr.is_type1 (Hunivelt.node l)
+ | _ -> false
+
+ (* Returns the formal universe that lies juste above the universe variable u.
+ Used to type the sort u. *)
+ let super l =
+ Huniv.map (fun x -> Hunivelt.make (Expr.successor (Hunivelt.node x))) l
+
+ let addn n l =
+ Huniv.map (fun x -> Hunivelt.make (Expr.addn n (Hunivelt.node x))) l
+
+ let rec merge_univs l1 l2 =
+ match node l1, node l2 with
+ | Nil, _ -> l2
+ | _, Nil -> l1
+ | Cons (h1, t1), Cons (h2, t2) ->
+ (match Expr.super (Hunivelt.node h1) (Hunivelt.node h2) with
+ | Inl true (* h1 < h2 *) -> merge_univs t1 l2
+ | Inl false -> merge_univs l1 t2
+ | Inr c ->
+ if c <= 0 (* h1 < h2 is name order *)
+ then cons h1 (merge_univs t1 l2)
+ else cons h2 (merge_univs l1 t2))
+
+ let sort u =
+ let rec aux a l =
+ match node l with
+ | Cons (b, l') ->
+ (match Expr.super (Hunivelt.node a) (Hunivelt.node b) with
+ | Inl false -> aux a l'
+ | Inl true -> l
+ | Inr c ->
+ if c <= 0 then cons a l
+ else cons b (aux a l'))
+ | Nil -> cons a l
+ in
+ fold (fun a acc -> aux a acc) u nil
+
+ (* Returns the formal universe that is greater than the universes u and v.
+ Used to type the products. *)
+ let sup x y = merge_univs x y
+
+ let of_list l =
+ List.fold_right
+ (fun x acc -> cons (Hunivelt.make x) acc)
+ l nil
+
+ let empty = nil
+ let is_empty n = is_nil n
+
+ let exists f l =
+ Huniv.exists (fun x -> f (Hunivelt.node x)) l
+
+ let for_all f l =
+ Huniv.for_all (fun x -> f (Hunivelt.node x)) l
+
+ let smartmap f l =
+ Huniv.map (fun x ->
+ let n = Hunivelt.node x in
+ let x' = f n in
+ if x' == n then x else Hunivelt.make x')
+ l
+
+end
type universe = Universe.t
-let universe_level = function
- | Atom l -> Some l
- | Max _ -> None
+open Universe
-let pr_uni_level u = str (UniverseLevel.to_string u)
+(* type universe_list = UList.t *)
+(* let pr_universe_list = UList.pr *)
-let pr_uni = function
- | Atom u ->
- pr_uni_level u
- | Max ([],[u]) ->
- str "(" ++ pr_uni_level u ++ str ")+1"
- | Max (gel,gtl) ->
- let opt_sep = match gel, gtl with
- | [], _ | _, [] -> mt ()
- | _ -> pr_comma ()
- in
- str "max(" ++ hov 0
- (prlist_with_sep pr_comma pr_uni_level gel ++ opt_sep ++
- prlist_with_sep pr_comma
- (fun x -> str "(" ++ pr_uni_level x ++ str ")+1") gtl) ++
- str ")"
-
-(* Returns the formal universe that lies juste above the universe variable u.
- Used to type the sort u. *)
-let super = function
- | Atom u ->
- Max ([],[u])
- | Max _ ->
- anomaly (str "Cannot take the successor of a non variable universe" ++ spc () ++
- str "(maybe a bugged tactic)")
-
-(* Returns the formal universe that is greater than the universes u and v.
- Used to type the products. *)
-let sup u v =
- match u,v with
- | Atom u, Atom v ->
- if UniverseLevel.equal u v then Atom u else Max ([u;v],[])
- | u, Max ([],[]) -> u
- | Max ([],[]), v -> v
- | Atom u, Max (gel,gtl) -> Max (List.add_set UniverseLevel.equal u gel,gtl)
- | Max (gel,gtl), Atom v -> Max (List.add_set UniverseLevel.equal v gel,gtl)
- | Max (gel,gtl), Max (gel',gtl') ->
- let gel'' = List.union UniverseLevel.equal gel gel' in
- let gtl'' = List.union UniverseLevel.equal gtl gtl' in
- Max (List.subtract UniverseLevel.equal gel'' gtl'',gtl'')
+let pr_uni = Universe.pr
+let is_small_univ = Universe.is_small
+
+let universe_level = Universe.level
(* Comparison on this type is pointer equality *)
type canonical_arc =
- { univ: UniverseLevel.t;
- lt: UniverseLevel.t list;
- le: UniverseLevel.t list;
- rank: int }
+ { univ: Level.t;
+ lt: Level.t list;
+ le: Level.t list;
+ rank : int}
let terminal u = {univ=u; lt=[]; le=[]; rank=0}
-(* A UniverseLevel.t is either an alias for another one, or a canonical one,
+(* A Level.t is either an alias for another one, or a canonical one,
for which we know the universes that are above *)
type univ_entry =
Canonical of canonical_arc
- | Equiv of UniverseLevel.t
+ | Equiv of Level.t
-type universes = univ_entry UniverseLMap.t
+type universes = univ_entry LMap.t
let enter_equiv_arc u v g =
- UniverseLMap.add u (Equiv v) g
+ LMap.add u (Equiv v) g
let enter_arc ca g =
- UniverseLMap.add ca.univ (Canonical ca) g
-
-(* The lower predicative level of the hierarchy that contains (impredicative)
- Prop and singleton inductive types *)
-let type0m_univ = Max ([],[])
+ LMap.add ca.univ (Canonical ca) g
-let is_type0m_univ = function
- | Max ([],[]) -> true
- | _ -> false
+let is_type0m_univ = Universe.is_type0m
(* The level of predicative Set *)
-let type0_univ = Atom UniverseLevel.Set
+let type0m_univ = Universe.type0m
+let type0_univ = Universe.type0
+let type1_univ = Universe.type1
-let is_type0_univ = function
- | Atom UniverseLevel.Set -> true
- | Max ([UniverseLevel.Set], []) -> msg_warning (str "Non canonical Set"); true
- | u -> false
+let sup = Universe.sup
+let super = Universe.super
-let is_univ_variable = function
- | Atom UniverseLevel.Set -> false
- | Atom _ -> true
- | _ -> false
+let is_type0_univ = Universe.is_type0
-(* When typing [Prop] and [Set], there is no constraint on the level,
- hence the definition of [type1_univ], the type of [Prop] *)
+let is_univ_variable l = Universe.level l != None
-let type1_univ = Max ([], [UniverseLevel.Set])
+(* Every Level.t has a unique canonical arc representative *)
-let initial_universes = UniverseLMap.empty
-let is_initial_universes = UniverseLMap.is_empty
-
-(* Every UniverseLevel.t has a unique canonical arc representative *)
-
-(* repr : universes -> UniverseLevel.t -> canonical_arc *)
+(* repr : universes -> Level.t -> canonical_arc *)
(* canonical representative : we follow the Equiv links *)
let repr g u =
let rec repr_rec u =
let a =
- try UniverseLMap.find u g
+ try LMap.find u g
with Not_found -> anomaly ~label:"Univ.repr"
- (str "Universe" ++ spc () ++ pr_uni_level u ++ spc () ++ str "undefined")
+ (str"Universe " ++ Level.pr u ++ str" undefined")
in
match a with
| Equiv v -> repr_rec v
@@ -262,7 +854,7 @@ let can g = List.map (repr g)
let safe_repr g u =
let rec safe_repr_rec u =
- match UniverseLMap.find u g with
+ match LMap.find u g with
| Equiv v -> safe_repr_rec v
| Canonical arc -> arc
in
@@ -286,7 +878,7 @@ let reprleq g arcu =
searchrec [] arcu.le
-(* between : UniverseLevel.t -> canonical_arc -> canonical_arc list *)
+(* between : Level.t -> canonical_arc -> canonical_arc list *)
(* between u v = { w | u<=w<=v, w canonical } *)
(* between is the most costly operation *)
@@ -320,6 +912,7 @@ let between g arcu arcv =
*)
type constraint_type = Lt | Le | Eq
+
type explanation = (constraint_type * universe) list
let constraint_type_ord c1 c2 = match c1, c2 with
@@ -335,10 +928,10 @@ let constraint_type_ord c1 c2 = match c1, c2 with
correspond to the universes in (direct) relation [rel] with it,
make a list of canonical universe, updating the relation with
the starting point (path stored in reverse order). *)
-let canp g (p:explanation) rel l : (canonical_arc * explanation) list =
- List.map (fun u -> (repr g u, (rel,Atom u)::p)) l
+let canp g (p:explanation Lazy.t) rel l : (canonical_arc * explanation Lazy.t) list =
+ List.map (fun u -> (repr g u, lazy ((rel,Universe.make u):: Lazy.force p))) l
-type order = EQ | LT of explanation | LE of explanation | NLE
+type order = EQ | LT of explanation Lazy.t | LE of explanation Lazy.t | NLE
(** [compare_neq] : is [arcv] in the transitive upward closure of [arcu] ?
@@ -375,14 +968,14 @@ let compare_neq strict g arcu arcv =
| [] -> cmp c (arc :: lt_done) le_done lt_todo le_todo
| u :: lt ->
let arc = repr g u in
- let p = (Lt, Atom u) :: p in
+ let p = lazy ((Lt, make u) :: Lazy.force p) in
if arc == arcv then
if strict then LT p else LE p
else find ((arc, p) :: lt_todo) lt le
end
| u :: le ->
let arc = repr g u in
- let p = (Le, Atom u) :: p in
+ let p = lazy ((Le, make u) :: Lazy.force p) in
if arc == arcv then
if strict then LT p else LE p
else find ((arc, p) :: lt_todo) lt le
@@ -402,21 +995,22 @@ let compare_neq strict g arcu arcv =
let rec find lt_todo lt = match lt with
| [] ->
let fold accu u =
- let node = (repr g u, (Le, Atom u) :: p) in
+ let p = lazy ((Le, make u) :: Lazy.force p) in
+ let node = (repr g u, p) in
node :: accu
in
let le_new = List.fold_left fold le_todo arc.le in
cmp c lt_done (arc :: le_done) lt_todo le_new
| u :: lt ->
let arc = repr g u in
- let p = (Lt, Atom u) :: p in
+ let p = lazy ((Lt, make u) :: Lazy.force p) in
if arc == arcv then
if strict then LT p else LE p
else find ((arc, p) :: lt_todo) lt
in
find [] arc.lt
in
- cmp NLE [] [] [] [arcu, []]
+ cmp NLE [] [] [] [arcu, Lazy.lazy_from_val []]
let compare g arcu arcv =
if arcu == arcv then EQ else compare_neq true g arcu arcv
@@ -456,50 +1050,80 @@ let check_smaller g strict u v =
if strict then
is_lt g arcu arcv
else
- arcu == snd (safe_repr g UniverseLevel.Set) || is_leq g arcu arcv
+ arcu == snd (safe_repr g Level.prop) || is_leq g arcu arcv
(** Then, checks on universes *)
-type check_function = universes -> universe -> universe -> bool
+type 'a check_function = universes -> 'a -> 'a -> bool
+
+(* let equiv_list cmp l1 l2 = *)
+(* let rec aux l1 l2 = *)
+(* match l1 with *)
+(* | [] -> l2 = [] *)
+(* | hd :: tl1 -> *)
+(* let rec aux' acc = function *)
+(* | hd' :: tl2 -> *)
+(* if cmp hd hd' then aux tl1 (acc @ tl2) *)
+(* else aux' (hd' :: acc) tl2 *)
+(* | [] -> false *)
+(* in aux' [] l2 *)
+(* in aux l1 l2 *)
let incl_list cmp l1 l2 =
- List.for_all (fun x1 -> List.exists (fun x2 -> cmp x1 x2) l2) l1
+ Huniv.for_all (fun x1 -> Huniv.exists (fun x2 -> cmp x1 x2) l2) l1
let compare_list cmp l1 l2 =
- (l1 == l2)
- || (incl_list cmp l1 l2 && incl_list cmp l2 l1)
+ (l1 == l2) || (* (equiv_list cmp l1 l2) *)
+ (incl_list cmp l1 l2 && incl_list cmp l2 l1)
+
+let check_equal_expr g x y =
+ x == y || (let (u, n) = Hunivelt.node x and (v, m) = Hunivelt.node y in
+ n = m && check_equal g u v)
(** [check_eq] is also used in [Evd.set_eq_sort],
hence [Evarconv] and [Unification]. In this case,
it seems that the Atom/Max case may occur,
hence a relaxed version. *)
-let gen_check_eq strict g u v =
- match u,v with
- | Atom ul, Atom vl -> check_equal g ul vl
- | Max(ule,ult), Max(vle,vlt) ->
- (* TODO: remove elements of lt in le! *)
- compare_list (check_equal g) ule vle &&
- compare_list (check_equal g) ult vlt
- | _ ->
- (* not complete! (Atom(u) = Max([u],[]) *)
- if strict then anomaly (str "check_eq")
- else false (* in non-strict mode, under-approximation *)
-
-let check_eq = gen_check_eq true
-let lax_check_eq = gen_check_eq false
+(* let gen_check_eq strict g u v = *)
+(* match u,v with *)
+(* | Atom ul, Atom vl -> check_equal g ul vl *)
+(* | Max(ule,ult), Max(vle,vlt) -> *)
+(* (\* TODO: remove elements of lt in le! *\) *)
+(* compare_list (check_equal g) ule vle && *)
+(* compare_list (check_equal g) ult vlt *)
+(* | _ -> *)
+(* (\* not complete! (Atom(u) = Max([u],[]) *\) *)
+(* if strict then anomaly (str "check_eq") *)
+(* else false (\* in non-strict mode, under-approximation *\) *)
+
+(* let check_eq = gen_check_eq true *)
+(* let lax_check_eq = gen_check_eq false *)
+let check_eq g u v =
+ compare_list (check_equal_expr g) u v
+let check_eq_level g u v = u == v || check_equal g u v
+let lax_check_eq = check_eq
+
+let check_smaller_expr g (u,n) (v,m) =
+ if n <= m then
+ check_smaller g false u v
+ else if n - m = 1 then
+ check_smaller g true u v
+ else false
+
+let exists_bigger g ul l =
+ Huniv.exists (fun ul' ->
+ check_smaller_expr g (Hunivelt.node ul) (Hunivelt.node ul')) l
let check_leq g u v =
- match u,v with
- | Atom ul, Atom vl -> check_smaller g false ul vl
- | Max(le,lt), Atom vl ->
- List.for_all (fun ul -> check_smaller g false ul vl) le &&
- List.for_all (fun ul -> check_smaller g true ul vl) lt
- | _ -> anomaly (str "check_leq")
+ u == v ||
+ match Universe.level u with
+ | Some l when Level.is_prop l -> true
+ | _ -> Huniv.for_all (fun ul -> exists_bigger g ul v) u
(** Enforcing new constraints : [setlt], [setleq], [merge], [merge_disc] *)
-(* setlt : UniverseLevel.t -> UniverseLevel.t -> reason -> unit *)
+(* setlt : Level.t -> Level.t -> reason -> unit *)
(* forces u > v *)
(* this is normally an update of u in g rather than a creation. *)
let setlt g arcu arcv =
@@ -512,7 +1136,7 @@ let setlt_if (g,arcu) v =
if is_lt g arcu arcv then g, arcu
else setlt g arcu arcv
-(* setleq : UniverseLevel.t -> UniverseLevel.t -> unit *)
+(* setleq : Level.t -> Level.t -> unit *)
(* forces u >= v *)
(* this is normally an update of u in g rather than a creation. *)
let setleq g arcu arcv =
@@ -526,7 +1150,7 @@ let setleq_if (g,arcu) v =
if is_leq g arcu arcv then g, arcu
else setleq g arcu arcv
-(* merge : UniverseLevel.t -> UniverseLevel.t -> unit *)
+(* merge : Level.t -> Level.t -> unit *)
(* we assume compare(u,v) = LE *)
(* merge u v forces u ~ v with repr u as canonical repr *)
let merge g arcu arcv =
@@ -559,7 +1183,7 @@ let merge g arcu arcv =
let g_arcu = List.fold_left setleq_if g_arcu w' in
fst g_arcu
-(* merge_disc : UniverseLevel.t -> UniverseLevel.t -> unit *)
+(* merge_disc : Level.t -> Level.t -> unit *)
(* we assume compare(u,v) = compare(v,u) = NLE *)
(* merge_disc u v forces u ~ v with repr u as canonical repr *)
let merge_disc g arc1 arc2 =
@@ -579,36 +1203,37 @@ let merge_disc g arc1 arc2 =
(* Universe inconsistency: error raised when trying to enforce a relation
that would create a cycle in the graph of universes. *)
-exception UniverseInconsistency of
- constraint_type * universe * universe * explanation
+type univ_inconsistency = constraint_type * universe * universe * explanation
+
+exception UniverseInconsistency of univ_inconsistency
let error_inconsistency o u v (p:explanation) =
- raise (UniverseInconsistency (o,Atom u,Atom v,p))
+ raise (UniverseInconsistency (o,make u,make v,p))
-(* enforce_univ_leq : UniverseLevel.t -> UniverseLevel.t -> unit *)
+(* enforce_univ_leq : Level.t -> Level.t -> unit *)
(* enforce_univ_leq u v will force u<=v if possible, will fail otherwise *)
let enforce_univ_leq u v g =
let g,arcu = safe_repr g u in
let g,arcv = safe_repr g v in
if is_leq g arcu arcv then g
else match compare g arcv arcu with
- | LT p -> error_inconsistency Le u v (List.rev p)
+ | LT p -> error_inconsistency Le u v (List.rev (Lazy.force p))
| LE _ -> merge g arcv arcu
| NLE -> fst (setleq g arcu arcv)
| EQ -> anomaly (Pp.str "Univ.compare")
-(* enforc_univ_eq : UniverseLevel.t -> UniverseLevel.t -> unit *)
+(* enforc_univ_eq : Level.t -> Level.t -> unit *)
(* enforc_univ_eq u v will force u=v if possible, will fail otherwise *)
let enforce_univ_eq u v g =
let g,arcu = safe_repr g u in
let g,arcv = safe_repr g v in
match compare g arcu arcv with
| EQ -> g
- | LT p -> error_inconsistency Eq u v (List.rev p)
+ | LT p -> error_inconsistency Eq v u (List.rev (Lazy.force p))
| LE _ -> merge g arcu arcv
| NLE ->
(match compare g arcv arcu with
- | LT p -> error_inconsistency Eq u v (List.rev p)
+ | LT p -> error_inconsistency Eq u v (List.rev (Lazy.force p))
| LE _ -> merge g arcv arcu
| NLE -> merge_disc g arcu arcv
| EQ -> anomaly (Pp.str "Univ.compare"))
@@ -620,16 +1245,20 @@ let enforce_univ_lt u v g =
match compare g arcu arcv with
| LT _ -> g
| LE _ -> fst (setlt g arcu arcv)
- | EQ -> error_inconsistency Lt u v [(Eq,Atom v)]
+ | EQ -> error_inconsistency Lt u v [(Eq,make v)]
| NLE ->
(match compare_neq false g arcv arcu with
NLE -> fst (setlt g arcu arcv)
| EQ -> anomaly (Pp.str "Univ.compare")
- | (LE p|LT p) -> error_inconsistency Lt u v (List.rev p))
+ | (LE p|LT p) -> error_inconsistency Lt u v (List.rev (Lazy.force p)))
-(* Constraints and sets of consrtaints. *)
+let empty_universes = LMap.empty
+let initial_universes = enforce_univ_lt Level.prop Level.set LMap.empty
+let is_initial_universes g = LMap.equal (==) g initial_universes
-type univ_constraint = UniverseLevel.t * constraint_type * UniverseLevel.t
+(* Constraints and sets of constraints. *)
+
+type univ_constraint = Level.t * constraint_type * Level.t
let enforce_constraint cst g =
match cst with
@@ -637,6 +1266,14 @@ let enforce_constraint cst g =
| (u,Le,v) -> enforce_univ_leq u v g
| (u,Eq,v) -> enforce_univ_eq u v g
+let pr_constraint_type op =
+ let op_str = match op with
+ | Lt -> " < "
+ | Le -> " <= "
+ | Eq -> " = "
+ in str op_str
+
+
module UConstraintOrd =
struct
type t = univ_constraint
@@ -644,51 +1281,566 @@ struct
let i = constraint_type_ord c c' in
if not (Int.equal i 0) then i
else
- let i' = UniverseLevel.compare u u' in
+ let i' = Level.compare u u' in
if not (Int.equal i' 0) then i'
- else UniverseLevel.compare v v'
+ else Level.compare v v'
end
-module Constraint = Set.Make(UConstraintOrd)
+module Constraint =
+struct
+ module S = Set.Make(UConstraintOrd)
+ include S
-type constraints = Constraint.t
+ let pr c =
+ fold (fun (u1,op,u2) pp_std ->
+ pp_std ++ Level.pr u1 ++ pr_constraint_type op ++
+ Level.pr u2 ++ fnl () ) c (str "")
+
+end
let empty_constraint = Constraint.empty
let is_empty_constraint = Constraint.is_empty
+let union_constraint = Constraint.union
let eq_constraint = Constraint.equal
-let union_constraints = Constraint.union
+type constraints = Constraint.t
+
+module Hconstraint =
+ Hashcons.Make(
+ struct
+ type t = univ_constraint
+ type u = universe_level -> universe_level
+ let hashcons hul (l1,k,l2) = (hul l1, k, hul l2)
+ let equal (l1,k,l2) (l1',k',l2') =
+ l1 == l1' && k == k' && l2 == l2'
+ let hash = Hashtbl.hash
+ end)
+
+module Hconstraints =
+ Hashcons.Make(
+ struct
+ type t = constraints
+ type u = univ_constraint -> univ_constraint
+ let hashcons huc s =
+ Constraint.fold (fun x -> Constraint.add (huc x)) s Constraint.empty
+ let equal s s' =
+ List.for_all2eq (==)
+ (Constraint.elements s)
+ (Constraint.elements s')
+ let hash = Hashtbl.hash
+ end)
+
+let hcons_constraint = Hashcons.simple_hcons Hconstraint.generate Level.hcons
+let hcons_constraints = Hashcons.simple_hcons Hconstraints.generate hcons_constraint
+
+type universe_constraint_type = ULe | UEq | ULub
+
+type universe_constraint = universe * universe_constraint_type * universe
+module UniverseConstraints = struct
+ module S = Set.Make(
+ struct
+ type t = universe_constraint
+
+ let compare_type c c' =
+ match c, c' with
+ | ULe, ULe -> 0
+ | ULe, _ -> -1
+ | _, ULe -> 1
+ | UEq, UEq -> 0
+ | UEq, _ -> -1
+ | ULub, ULub -> 0
+ | ULub, _ -> 1
+
+ let compare (u,c,v) (u',c',v') =
+ let i = compare_type c c' in
+ if Int.equal i 0 then
+ let i' = Universe.compare u u' in
+ if Int.equal i' 0 then Universe.compare v v'
+ else
+ if c != ULe && Universe.compare u v' = 0 && Universe.compare v u' = 0 then 0
+ else i'
+ else i
+ end)
+
+ include S
+
+ let add (l,d,r as cst) s =
+ if Universe.eq l r then s
+ else add cst s
+
+ let tr_dir = function
+ | ULe -> Le
+ | UEq -> Eq
+ | ULub -> Eq
+
+ let op_str = function ULe -> " <= " | UEq -> " = " | ULub -> " /\\ "
+
+ let pr c =
+ fold (fun (u1,op,u2) pp_std ->
+ pp_std ++ Universe.pr u1 ++ str (op_str op) ++
+ Universe.pr u2 ++ fnl ()) c (str "")
+
+ let equal x y =
+ x == y || equal x y
-type constraint_function =
- universe -> universe -> constraints -> constraints
+end
+
+type universe_constraints = UniverseConstraints.t
+type 'a universe_constrained = 'a * universe_constraints
+
+(** A value with universe constraints. *)
+type 'a constrained = 'a * constraints
+
+let level_subst_of f =
+ fun l ->
+ try let u = f l in
+ match Universe.level u with
+ | None -> l
+ | Some l -> l
+ with Not_found -> l
+
+module Instance = struct
+ type t = Level.t array
+
+ module HInstance =
+ Hashcons.Make(
+ struct
+ type _t = t
+ type t = _t
+ type u = Level.t -> Level.t
+ let hashcons huniv a =
+ CArray.smartmap huniv a
+
+ let equal t1 t2 =
+ t1 == t2 ||
+ (Int.equal (Array.length t1) (Array.length t2) &&
+ let rec aux i =
+ (Int.equal i (Array.length t1)) || (t1.(i) == t2.(i) && aux (i + 1))
+ in aux 0)
+
+ let hash = Hashtbl.hash
+ end)
+
+ let hcons_instance = Hashcons.simple_hcons HInstance.generate Level.hcons
+
+ let hcons = hcons_instance
+ let empty = [||]
+ let is_empty x = Int.equal (Array.length x) 0
+
+ let eq t u = t == u || CArray.for_all2 Level.eq t u
+
+ let of_array a = a
+ let to_array a = a
+
+ let eqeq t1 t2 =
+ t1 == t2 ||
+ (Int.equal (Array.length t1) (Array.length t2) &&
+ let rec aux i =
+ (Int.equal i (Array.length t1)) || (t1.(i) == t2.(i) && aux (i + 1))
+ in aux 0)
+
+ let subst_fn fn t = CArray.smartmap fn t
+ let subst s t = CArray.smartmap (fun x -> try LMap.find x s with Not_found -> x) t
+
+ let levels x = LSet.of_array x
+
+ let pr =
+ prvect_with_sep spc Level.pr
+
+ let append x y =
+ if Array.length x = 0 then y
+ else if Array.length y = 0 then x
+ else Array.append x y
+
+ let max_level i =
+ if Array.is_empty i then 0
+ else
+ let l = CArray.last i in
+ match l with
+ | Level.Level (i,_) -> i
+ | _ -> assert false
+
+ let check_eq g t1 t2 =
+ t1 == t2 ||
+ (Int.equal (Array.length t1) (Array.length t2) &&
+ let rec aux i =
+ (Int.equal i (Array.length t1)) || (check_eq_level g t1.(i) t2.(i) && aux (i + 1))
+ in aux 0)
+
+end
+
+type universe_instance = Instance.t
+
+type 'a puniverses = 'a * Instance.t
+let out_punivs (x, y) = x
+let in_punivs x = (x, Instance.empty)
+
+(** A context of universe levels with universe constraints,
+ representiong local universe variables and constraints *)
+
+module UContext =
+struct
+ type t = Instance.t constrained
+
+ let make x = x
+
+ (** Universe contexts (variables as a list) *)
+ let empty = (Instance.empty, Constraint.empty)
+ let is_empty (univs, cst) = Instance.is_empty univs && Constraint.is_empty cst
+
+ let pr (univs, cst as ctx) =
+ if is_empty ctx then mt() else
+ Instance.pr univs ++ str " |= " ++ v 1 (Constraint.pr cst)
+
+ let hcons (univs, cst) =
+ (Instance.hcons univs, hcons_constraints cst)
+
+ let instance (univs, cst) = univs
+ let constraints (univs, cst) = cst
+
+ let union (univs, cst) (univs', cst') =
+ Instance.append univs univs', Constraint.union cst cst'
+end
+
+type universe_context = UContext.t
+let hcons_universe_context = UContext.hcons
+
+(** A set of universes with universe constraints.
+ We linearize the set to a list after typechecking.
+ Beware, representation could change.
+*)
+
+module ContextSet =
+struct
+ type t = universe_set constrained
+
+ let empty = (LSet.empty, Constraint.empty)
+ let is_empty (univs, cst) = LSet.is_empty univs && Constraint.is_empty cst
+
+ let of_context (ctx,cst) =
+ (Instance.levels ctx, cst)
+
+ let of_set s = (s, Constraint.empty)
+ let singleton l = of_set (LSet.singleton l)
+ let of_instance i = of_set (Instance.levels i)
+
+ let union (univs, cst) (univs', cst') =
+ LSet.union univs univs', Constraint.union cst cst'
+
+ let diff (univs, cst) (univs', cst') =
+ LSet.diff univs univs', Constraint.diff cst cst'
+
+ let add_constraints (univs, cst) cst' =
+ univs, Constraint.union cst cst'
+
+ let add_universes univs ctx =
+ union (of_instance univs) ctx
+
+ let to_context (ctx, cst) =
+ (Array.of_list (LSet.elements ctx), cst)
+
+ let of_context (ctx, cst) =
+ (Instance.levels ctx, cst)
+
+ let pr (univs, cst as ctx) =
+ if is_empty ctx then mt() else
+ LSet.pr univs ++ str " |= " ++ v 1 (Constraint.pr cst)
+
+ let constraints (univs, cst) = cst
+ let levels (univs, cst) = univs
+
+end
+
+type universe_context_set = ContextSet.t
+
+(** A value in a universe context (resp. context set). *)
+type 'a in_universe_context = 'a * universe_context
+type 'a in_universe_context_set = 'a * universe_context_set
+
+(** A universe level substitution, note that no algebraic universes are
+ involved *)
+type universe_level_subst = universe_level universe_map
+
+(** A full substitution might involve algebraic universes *)
+type universe_subst = universe universe_map
+
+(** Pretty-printing *)
+let pr_constraints = Constraint.pr
+
+let pr_universe_context = UContext.pr
+
+let pr_universe_context_set = ContextSet.pr
+
+let pr_universe_subst =
+ LMap.pr (fun u -> str" := " ++ Universe.pr u ++ spc ())
+
+let pr_universe_level_subst =
+ LMap.pr (fun u -> str" := " ++ Level.pr u ++ spc ())
+
+let constraints_of (_, cst) = cst
+
+let constraint_depend (l,d,r) u =
+ Level.eq l u || Level.eq l r
+
+let constraint_depend_list (l,d,r) us =
+ List.mem l us || List.mem r us
+
+let constraints_depend cstr us =
+ Constraint.exists (fun c -> constraint_depend_list c us) cstr
+
+let remove_dangling_constraints dangling cst =
+ Constraint.fold (fun (l,d,r as cstr) cst' ->
+ if List.mem l dangling || List.mem r dangling then cst'
+ else
+ (** Unnecessary constraints Prop <= u *)
+ if Level.eq l Level.prop && d == Le then cst'
+ else Constraint.add cstr cst') cst Constraint.empty
+
+let check_context_subset (univs, cst) (univs', cst') =
+ let newunivs, dangling = List.partition (fun u -> LSet.mem u univs) (Array.to_list univs') in
+ (* Some universe variables that don't appear in the term
+ are still mentionned in the constraints. This is the
+ case for "fake" universe variables that correspond to +1s. *)
+ (* if not (CList.is_empty dangling) then *)
+ (* todo ("A non-empty set of inferred universes do not appear in the term or type"); *)
+ (* (not (constraints_depend cst' dangling));*)
+ (* TODO: check implication *)
+ (** Remove local universes that do not appear in any constraint, they
+ are really entirely parametric. *)
+ (* let newunivs, dangling' = List.partition (fun u -> constraints_depend cst [u]) newunivs in *)
+ let cst' = remove_dangling_constraints dangling cst in
+ Array.of_list newunivs, cst'
+
+(** Substitutions. *)
+
+let make_universe_subst inst (ctx, csts) =
+ try Array.fold_left2 (fun acc c i -> LMap.add c (Universe.make i) acc)
+ LMap.empty ctx inst
+ with Invalid_argument _ ->
+ anomaly (Pp.str "Mismatched instance and context when building universe substitution")
+
+let empty_subst = LMap.empty
+let is_empty_subst = LMap.is_empty
+
+let empty_level_subst = LMap.empty
+let is_empty_level_subst = LMap.is_empty
+
+(** Substitution functions *)
+
+(** With level to level substitutions. *)
+let subst_univs_level_level subst l =
+ try LMap.find l subst
+ with Not_found -> l
+
+let rec normalize_univs_level_level subst l =
+ try
+ let l' = LMap.find l subst in
+ normalize_univs_level_level subst l'
+ with Not_found -> l
+
+let subst_univs_level_fail subst l =
+ try match Universe.level (subst l) with
+ | Some l' -> l'
+ | None -> l
+ with Not_found -> l
+
+let rec subst_univs_level_universe subst u =
+ let u' = Universe.smartmap (Universe.Expr.map (subst_univs_level_level subst)) u in
+ if u == u' then u
+ else Universe.sort u'
+
+let subst_univs_level_constraint subst (u,d,v) =
+ let u' = subst_univs_level_level subst u
+ and v' = subst_univs_level_level subst v in
+ if d != Lt && Level.eq u' v' then None
+ else Some (u',d,v')
+
+let subst_univs_level_constraints subst csts =
+ Constraint.fold
+ (fun c -> Option.fold_right Constraint.add (subst_univs_level_constraint subst c))
+ csts Constraint.empty
+
+(* let subst_univs_level_constraint_key = Profile.declare_profile "subst_univs_level_constraint";; *)
+(* let subst_univs_level_constraint = *)
+(* Profile.profile2 subst_univs_level_constraint_key subst_univs_level_constraint *)
+
+(** With level to universe substitutions. *)
+type universe_subst_fn = universe_level -> universe
+
+let make_subst subst = fun l -> LMap.find l subst
+
+let subst_univs_level fn l =
+ try fn l
+ with Not_found -> make l
+
+let subst_univs_expr_opt fn (l,n) =
+ try Some (Universe.addn n (fn l))
+ with Not_found -> None
+
+let subst_univs_universe fn ul =
+ let subst, nosubst =
+ Universe.Huniv.fold (fun u (subst,nosubst) ->
+ match subst_univs_expr_opt fn (Hunivelt.node u) with
+ | Some a' -> (a' :: subst, nosubst)
+ | None -> (subst, u :: nosubst))
+ ul ([], [])
+ in
+ if CList.is_empty subst then ul
+ else
+ let substs =
+ List.fold_left Universe.merge_univs Universe.empty subst
+ in
+ List.fold_left (fun acc u -> Universe.merge_univs acc (Universe.Huniv.tip u))
+ substs nosubst
+
+let subst_univs_constraint fn (u,d,v) =
+ let u' = subst_univs_level fn u and v' = subst_univs_level fn v in
+ if d != Lt && Universe.eq u' v' then None
+ else Some (u',d,v')
+
+let subst_univs_universe_constraint fn (u,d,v) =
+ let u' = subst_univs_universe fn u and v' = subst_univs_universe fn v in
+ if Universe.eq u' v' then None
+ else Some (u',d,v')
+
+(** Constraint functions. *)
+
+type 'a constraint_function = 'a -> 'a -> constraints -> constraints
let constraint_add_leq v u c =
- (* We just discard trivial constraints like Set<=u or u<=u *)
- if UniverseLevel.equal v UniverseLevel.Set || UniverseLevel.equal v u then c
- else Constraint.add (v,Le,u) c
+ (* We just discard trivial constraints like u<=u *)
+ if Expr.eq v u then c
+ else
+ match v, u with
+ | (x,n), (y,m) ->
+ let j = m - n in
+ if j = -1 (* n = m+1, v+1 <= u <-> v < u *) then
+ Constraint.add (x,Lt,y) c
+ else if j <= -1 (* n = m+k, v+k <= u <-> v+(k-1) < u *) then
+ if Level.eq x y then (* u+(k+1) <= u *)
+ raise (UniverseInconsistency (Le, Universe.tip v, Universe.tip u, []))
+ else anomaly (Pp.str"Unable to handle arbitrary u+k <= v constraints")
+ else if j = 0 then
+ Constraint.add (x,Le,y) c
+ else (* j >= 1 *) (* m = n + k, u <= v+k *)
+ if Level.eq x y then c (* u <= u+k, trivial *)
+ else if Level.is_small x then c (* Prop,Set <= u+S k, trivial *)
+ else anomaly (Pp.str"Unable to handle arbitrary u <= v+k constraints")
+
+let check_univ_eq u v = Universe.eq u v
+
+let check_univ_leq_one u v = Universe.exists (Expr.leq u) v
+
+let check_univ_leq u v =
+ Universe.for_all (fun u -> check_univ_leq_one u v) u
let enforce_leq u v c =
- match u, v with
- | Atom u, Atom v -> constraint_add_leq u v c
- | Max (gel,gtl), Atom v ->
- let d = List.fold_right (fun u -> constraint_add_leq u v) gel c in
- List.fold_right (fun u -> Constraint.add (u,Lt,v)) gtl d
- | _ -> anomaly (Pp.str "A universe bound can only be a variable")
+ match Huniv.node v with
+ | Universe.Huniv.Cons (v, n) when Universe.is_empty n ->
+ Universe.Huniv.fold (fun u -> constraint_add_leq (Hunivelt.node u) (Hunivelt.node v)) u c
+ | _ -> anomaly (Pp.str"A universe bound can only be a variable")
+
+let enforce_leq u v c =
+ if check_univ_leq u v then c
+ else enforce_leq u v c
+
+let enforce_eq_level u v c =
+ (* We discard trivial constraints like u=u *)
+ if Level.eq u v then c
+ else if Level.apart u v then
+ error_inconsistency Eq u v []
+ else Constraint.add (u,Eq,v) c
let enforce_eq u v c =
- match (u,v) with
- | Atom u, Atom v ->
- (* We discard trivial constraints like u=u *)
- if UniverseLevel.equal u v then c else Constraint.add (u,Eq,v) c
+ match Universe.level u, Universe.level v with
+ | Some u, Some v -> enforce_eq_level u v c
| _ -> anomaly (Pp.str "A universe comparison can only happen between variables")
+let enforce_eq u v c =
+ if check_univ_eq u v then c
+ else enforce_eq u v c
+
+let enforce_leq_level u v c =
+ if Level.eq u v then c else Constraint.add (u,Le,v) c
+
+let enforce_eq_instances = CArray.fold_right2 enforce_eq_level
+
+type 'a universe_constraint_function = 'a -> 'a -> universe_constraints -> universe_constraints
+
+let enforce_eq_instances_univs strict t1 t2 c =
+ let d = if strict then ULub else UEq in
+ CArray.fold_right2 (fun x y -> UniverseConstraints.add (Universe.make x, d, Universe.make y))
+ t1 t2 c
+
let merge_constraints c g =
Constraint.fold enforce_constraint c g
+(* let merge_constraints_key = Profile.declare_profile "merge_constraints";; *)
+(* let merge_constraints = Profile.profile2 merge_constraints_key merge_constraints *)
+
+let check_constraint g (l,d,r) =
+ match d with
+ | Eq -> check_equal g l r
+ | Le -> check_smaller g false l r
+ | Lt -> check_smaller g true l r
+
+let check_constraints c g =
+ Constraint.for_all (check_constraint g) c
+
+(* let check_constraints_key = Profile.declare_profile "check_constraints";; *)
+(* let check_constraints = *)
+(* Profile.profile2 check_constraints_key check_constraints *)
+
+let enforce_univ_constraint (u,d,v) =
+ match d with
+ | Eq -> enforce_eq u v
+ | Le -> enforce_leq u v
+ | Lt -> enforce_leq (super u) v
+
+let subst_univs_constraints subst csts =
+ Constraint.fold
+ (fun c -> Option.fold_right enforce_univ_constraint (subst_univs_constraint subst c))
+ csts Constraint.empty
+
+(* let subst_univs_constraints_key = Profile.declare_profile "subst_univs_constraints";; *)
+(* let subst_univs_constraints = *)
+(* Profile.profile2 subst_univs_constraints_key subst_univs_constraints *)
+
+let subst_univs_universe_constraints subst csts =
+ UniverseConstraints.fold
+ (fun c -> Option.fold_right UniverseConstraints.add (subst_univs_universe_constraint subst c))
+ csts UniverseConstraints.empty
+
+(* let subst_univs_universe_constraints_key = Profile.declare_profile "subst_univs_universe_constraints";; *)
+(* let subst_univs_universe_constraints = *)
+(* Profile.profile2 subst_univs_universe_constraints_key subst_univs_universe_constraints *)
+
+(** Substitute instance inst for ctx in csts *)
+let instantiate_univ_context subst (_, csts) =
+ subst_univs_constraints (make_subst subst) csts
+
+let check_consistent_constraints (ctx,cstrs) cstrs' =
+ (* TODO *) ()
+
+let to_constraints g s =
+ let rec tr (x,d,y) acc =
+ let add l d l' acc = Constraint.add (l,UniverseConstraints.tr_dir d,l') acc in
+ match Universe.level x, d, Universe.level y with
+ | Some l, (ULe | UEq | ULub), Some l' -> add l d l' acc
+ | _, ULe, Some l' -> enforce_leq x y acc
+ | _, ULub, _ -> acc
+ | _, d, _ ->
+ let f = if d == ULe then check_leq else check_eq in
+ if f g x y then acc else
+ raise (Invalid_argument
+ "to_constraints: non-trivial algebraic constraint between universes")
+ in UniverseConstraints.fold tr s Constraint.empty
+
+
(* Normalization *)
let lookup_level u g =
- try Some (UniverseLMap.find u g) with Not_found -> None
+ try Some (LMap.find u g) with Not_found -> None
(** [normalize_universes g] returns a graph where all edges point
directly to the canonical representent of their target. The output
@@ -702,20 +1854,20 @@ let normalize_universes g =
| Some x -> x, cache
| None -> match Lazy.force arc with
| None ->
- u, UniverseLMap.add u u cache
+ u, LMap.add u u cache
| Some (Canonical {univ=v; lt=_; le=_}) ->
- v, UniverseLMap.add u v cache
+ v, LMap.add u v cache
| Some (Equiv v) ->
let v, cache = visit v (lazy (lookup_level v g)) cache in
- v, UniverseLMap.add u v cache
+ v, LMap.add u v cache
in
- let cache = UniverseLMap.fold
+ let cache = LMap.fold
(fun u arc cache -> snd (visit u (Lazy.lazy_from_val (Some arc)) cache))
- g UniverseLMap.empty
+ g LMap.empty
in
- let repr x = UniverseLMap.find x cache in
+ let repr x = LMap.find x cache in
let lrepr us = List.fold_left
- (fun e x -> UniverseLSet.add (repr x) e) UniverseLSet.empty us
+ (fun e x -> LSet.add (repr x) e) LSet.empty us
in
let canonicalize u = function
| Equiv _ -> Equiv (repr u)
@@ -723,24 +1875,24 @@ let normalize_universes g =
assert (u == v);
(* avoid duplicates and self-loops *)
let lt = lrepr lt and le = lrepr le in
- let le = UniverseLSet.filter
- (fun x -> x != u && not (UniverseLSet.mem x lt)) le
+ let le = LSet.filter
+ (fun x -> x != u && not (LSet.mem x lt)) le
in
- UniverseLSet.iter (fun x -> assert (x != u)) lt;
+ LSet.iter (fun x -> assert (x != u)) lt;
Canonical {
univ = v;
- lt = UniverseLSet.elements lt;
- le = UniverseLSet.elements le;
+ lt = LSet.elements lt;
+ le = LSet.elements le;
rank = rank
}
in
- UniverseLMap.mapi canonicalize g
+ LMap.mapi canonicalize g
(** [check_sorted g sorted]: [g] being a universe graph, [sorted]
being a map to levels, checks that all constraints in [g] are
satisfied in [sorted]. *)
let check_sorted g sorted =
- let get u = try UniverseLMap.find u sorted with
+ let get u = try LMap.find u sorted with
| Not_found -> assert false
in
let iter u arc =
@@ -751,7 +1903,7 @@ let check_sorted g sorted =
List.iter (fun v -> assert (lu <= get v)) le;
List.iter (fun v -> assert (lu < get v)) lt
in
- UniverseLMap.iter iter g
+ LMap.iter iter g
(**
Bellman-Ford algorithm with a few customizations:
@@ -765,7 +1917,7 @@ let bellman_ford bottom g =
| None -> ()
| Some _ -> assert false
in
- let ( << ) a b = match a, b with
+ let ( <? ) a b = match a, b with
| _, None -> true
| None, _ -> false
| Some x, Some y -> (x : int) < y
@@ -774,38 +1926,38 @@ let bellman_ford bottom g =
| Some x -> Some (x-y)
and push u x m = match x with
| None -> m
- | Some y -> UniverseLMap.add u y m
+ | Some y -> LMap.add u y m
in
let relax u v uv distances =
let x = lookup_level u distances ++ uv in
- if x << lookup_level v distances then push v x distances
+ if x <? lookup_level v distances then push v x distances
else distances
in
- let init = UniverseLMap.add bottom 0 UniverseLMap.empty in
- let vertices = UniverseLMap.fold (fun u arc res ->
- let res = UniverseLSet.add u res in
+ let init = LMap.add bottom 0 LMap.empty in
+ let vertices = LMap.fold (fun u arc res ->
+ let res = LSet.add u res in
match arc with
- | Equiv e -> UniverseLSet.add e res
+ | Equiv e -> LSet.add e res
| Canonical {univ=univ; lt=lt; le=le} ->
assert (u == univ);
- let add res v = UniverseLSet.add v res in
+ let add res v = LSet.add v res in
let res = List.fold_left add res le in
let res = List.fold_left add res lt in
- res) g UniverseLSet.empty
+ res) g LSet.empty
in
let g =
let node = Canonical {
univ = bottom;
lt = [];
- le = UniverseLSet.elements vertices;
+ le = LSet.elements vertices;
rank = 0
- } in UniverseLMap.add bottom node g
+ } in LMap.add bottom node g
in
let rec iter count accu =
if count <= 0 then
accu
else
- let accu = UniverseLMap.fold (fun u arc res -> match arc with
+ let accu = LMap.fold (fun u arc res -> match arc with
| Equiv e -> relax e u 0 (relax u e 0 res)
| Canonical {univ=univ; lt=lt; le=le} ->
assert (u == univ);
@@ -814,16 +1966,16 @@ let bellman_ford bottom g =
res) g accu
in iter (count-1) accu
in
- let distances = iter (UniverseLSet.cardinal vertices) init in
- let () = UniverseLMap.iter (fun u arc ->
+ let distances = iter (LSet.cardinal vertices) init in
+ let () = LMap.iter (fun u arc ->
let lu = lookup_level u distances in match arc with
| Equiv v ->
let lv = lookup_level v distances in
- assert (not (lu << lv) && not (lv << lu))
+ assert (not (lu <? lv) && not (lv <? lu))
| Canonical {univ=univ; lt=lt; le=le} ->
assert (u == univ);
- List.iter (fun v -> assert (not (lu ++ 0 << lookup_level v distances))) le;
- List.iter (fun v -> assert (not (lu ++ 1 << lookup_level v distances))) lt) g
+ List.iter (fun v -> assert (not (lu ++ 0 <? lookup_level v distances))) le;
+ List.iter (fun v -> assert (not (lu ++ 1 <? lookup_level v distances))) lt) g
in distances
(** [sort_universes g] builds a map from universes in [g] to natural
@@ -837,23 +1989,23 @@ let bellman_ford bottom g =
let sort_universes orig =
let mp = Names.DirPath.make [Names.Id.of_string "Type"] in
let rec make_level accu g i =
- let type0 = UniverseLevel.Level (i, mp) in
+ let type0 = Level.make mp i in
let distances = bellman_ford type0 g in
- let accu, continue = UniverseLMap.fold (fun u x (accu, continue) ->
+ let accu, continue = LMap.fold (fun u x (accu, continue) ->
let continue = continue || x < 0 in
let accu =
- if Int.equal x 0 && u != type0 then UniverseLMap.add u i accu
+ if Int.equal x 0 && u != type0 then LMap.add u i accu
else accu
in accu, continue) distances (accu, false)
in
- let filter x = not (UniverseLMap.mem x accu) in
+ let filter x = not (LMap.mem x accu) in
let push g u =
- if UniverseLMap.mem u g then g else UniverseLMap.add u (Equiv u) g
+ if LMap.mem u g then g else LMap.add u (Equiv u) g
in
- let g = UniverseLMap.fold (fun u arc res -> match arc with
+ let g = LMap.fold (fun u arc res -> match arc with
| Equiv v as x ->
begin match filter u, filter v with
- | true, true -> UniverseLMap.add u x res
+ | true, true -> LMap.add u x res
| true, false -> push res u
| false, true -> push res v
| false, false -> res
@@ -863,24 +2015,24 @@ let sort_universes orig =
if filter u then
let lt = List.filter filter lt in
let le = List.filter filter le in
- UniverseLMap.add u (Canonical {univ=u; lt=lt; le=le; rank=r}) res
+ LMap.add u (Canonical {univ=u; lt=lt; le=le; rank=r}) res
else
let res = List.fold_left (fun g u -> if filter u then push g u else g) res lt in
let res = List.fold_left (fun g u -> if filter u then push g u else g) res le in
- res) g UniverseLMap.empty
+ res) g LMap.empty
in
if continue then make_level accu g (i+1) else i, accu
in
- let max, levels = make_level UniverseLMap.empty orig 0 in
+ let max, levels = make_level LMap.empty orig 0 in
(* defensively check that the result makes sense *)
check_sorted orig levels;
- let types = Array.init (max+1) (fun x -> UniverseLevel.Level (x, mp)) in
- let g = UniverseLMap.map (fun x -> Equiv types.(x)) levels in
+ let types = Array.init (max+1) (fun x -> Level.make mp x) in
+ let g = LMap.map (fun x -> Equiv types.(x)) levels in
let g =
let rec aux i g =
if i < max then
let u = types.(i) in
- let g = UniverseLMap.add u (Canonical {
+ let g = LMap.add u (Canonical {
univ = u;
le = [];
lt = [types.(i+1)];
@@ -893,26 +2045,19 @@ let sort_universes orig =
(**********************************************************************)
(* Tools for sort-polymorphic inductive types *)
-(* Temporary inductive type levels *)
-
-let fresh_local_univ, set_remote_fresh_local_univ =
- RemoteCounter.new_counter ~name:"local_univ" 0 ~incr:((+) 1)
- ~build:(fun n -> Atom (UniverseLevel.Level (n, Names.DirPath.empty)))
-
(* Miscellaneous functions to remove or test local univ assumed to
occur only in the le constraints *)
-let make_max = function
- | ([u],[]) -> Atom u
- | (le,lt) -> Max (le,lt)
-
-let remove_large_constraint u = function
- | Atom u' as x -> if UniverseLevel.equal u u' then Max ([],[]) else x
- | Max (le,lt) -> make_max (List.remove UniverseLevel.equal u le,lt)
+let remove_large_constraint u v min =
+ match Universe.level v with
+ | Some u' -> if Level.eq u u' then min else v
+ | None -> Huniv.remove (Hunivelt.make (Universe.Expr.make u)) v
-let is_direct_constraint u = function
- | Atom u' -> UniverseLevel.equal u u'
- | Max (le,lt) -> List.mem_f UniverseLevel.equal u le
+(* [is_direct_constraint u v] if level [u] is a member of universe [v] *)
+let is_direct_constraint u v =
+ match Universe.level v with
+ | Some u' -> Level.eq u u'
+ | None -> Huniv.mem (Hunivelt.make (Universe.Expr.make u)) v
(*
Solve a system of universe constraint of the form
@@ -932,29 +2077,31 @@ let is_direct_sort_constraint s v = match s with
| Some u -> is_direct_constraint u v
| None -> false
-let solve_constraints_system levels level_bounds =
+let solve_constraints_system levels level_bounds level_min =
let levels =
- Array.map (Option.map (function Atom u -> u | _ -> anomaly (Pp.str "expects Atom")))
+ Array.map (Option.map (fun u -> match level u with Some u -> u | _ -> anomaly (Pp.str"expects Atom")))
levels in
let v = Array.copy level_bounds in
let nind = Array.length v in
for i=0 to nind-1 do
for j=0 to nind-1 do
if not (Int.equal i j) && is_direct_sort_constraint levels.(j) v.(i) then
- v.(i) <- sup v.(i) level_bounds.(j)
+ (v.(i) <- Universe.sup v.(i) level_bounds.(j);
+ level_min.(i) <- Universe.sup level_min.(i) level_min.(j))
done;
for j=0 to nind-1 do
match levels.(j) with
- | Some u -> v.(i) <- remove_large_constraint u v.(i)
+ | Some u -> v.(i) <- remove_large_constraint u v.(i) level_min.(i)
| None -> ()
done
done;
v
let subst_large_constraint u u' v =
- match u with
- | Atom u ->
- if is_direct_constraint u v then sup u' (remove_large_constraint u v)
+ match level u with
+ | Some u ->
+ if is_direct_constraint u v then
+ Universe.sup u' (remove_large_constraint u v type0m_univ)
else v
| _ ->
anomaly (Pp.str "expect a universe level")
@@ -963,21 +2110,30 @@ let subst_large_constraints =
List.fold_right (fun (u,u') -> subst_large_constraint u u')
let no_upper_constraints u cst =
- match u with
- | Atom u ->
- let test (u1, _, _) = not (UniverseLevel.equal u1 u) in
+ match level u with
+ | Some u ->
+ let test (u1, _, _) =
+ not (Int.equal (Level.compare u1 u) 0) in
Constraint.for_all test cst
- | Max _ -> anomaly (Pp.str "no_upper_constraints")
+ | _ -> anomaly (Pp.str "no_upper_constraints")
(* Is u mentionned in v (or equals to v) ? *)
-let level_list_mem u l = List.mem_f UniverseLevel.equal u l
-
let univ_depends u v =
- match u, v with
- | Atom u, Atom v -> UniverseLevel.equal u v
- | Atom u, Max (gel,gtl) -> level_list_mem u gel || level_list_mem u gtl
- | _ -> anomaly (Pp.str "univ_depends given a non-atomic 1st arg")
+ match atom u with
+ | Some u -> Huniv.mem u v
+ | _ -> anomaly (Pp.str"univ_depends given a non-atomic 1st arg")
+
+let constraints_of_universes g =
+ let constraints_of u v acc =
+ match v with
+ | Canonical {univ=u; lt=lt; le=le} ->
+ let acc = List.fold_left (fun acc v -> Constraint.add (u,Lt,v) acc) acc lt in
+ let acc = List.fold_left (fun acc v -> Constraint.add (u,Le,v) acc) acc le in
+ acc
+ | Equiv v -> Constraint.add (u,Eq,v) acc
+ in
+ LMap.fold constraints_of g Constraint.empty
(* Pretty-printing *)
@@ -989,101 +2145,67 @@ let pr_arc = function
| [], _ | _, [] -> mt ()
| _ -> spc ()
in
- pr_uni_level u ++ str " " ++
+ Level.pr u ++ str " " ++
v 0
- (pr_sequence (fun v -> str "< " ++ pr_uni_level v) lt ++
+ (pr_sequence (fun v -> str "< " ++ Level.pr v) lt ++
opt_sep ++
- pr_sequence (fun v -> str "<= " ++ pr_uni_level v) le) ++
+ pr_sequence (fun v -> str "<= " ++ Level.pr v) le) ++
fnl ()
| u, Equiv v ->
- pr_uni_level u ++ str " = " ++ pr_uni_level v ++ fnl ()
+ Level.pr u ++ str " = " ++ Level.pr v ++ fnl ()
let pr_universes g =
- let graph = UniverseLMap.fold (fun u a l -> (u,a)::l) g [] in
+ let graph = LMap.fold (fun u a l -> (u,a)::l) g [] in
prlist pr_arc graph
-let pr_constraints c =
- Constraint.fold (fun (u1,op,u2) pp_std ->
- let op_str = match op with
- | Lt -> " < "
- | Le -> " <= "
- | Eq -> " = "
- in pp_std ++ pr_uni_level u1 ++ str op_str ++
- pr_uni_level u2 ++ fnl () ) c (str "")
-
(* Dumping constraints to a file *)
let dump_universes output g =
let dump_arc u = function
| Canonical {univ=u; lt=lt; le=le} ->
- let u_str = UniverseLevel.to_string u in
- List.iter (fun v -> output Lt u_str (UniverseLevel.to_string v)) lt;
- List.iter (fun v -> output Le u_str (UniverseLevel.to_string v)) le
+ let u_str = Level.to_string u in
+ List.iter (fun v -> output Lt u_str (Level.to_string v)) lt;
+ List.iter (fun v -> output Le u_str (Level.to_string v)) le
| Equiv v ->
- output Eq (UniverseLevel.to_string u) (UniverseLevel.to_string v)
+ output Eq (Level.to_string u) (Level.to_string v)
in
- UniverseLMap.iter dump_arc g
+ LMap.iter dump_arc g
-(* Hash-consing *)
-
-module Hunivlevel =
+module Huniverse_set =
Hashcons.Make(
struct
- type t = universe_level
- type u = Names.DirPath.t -> Names.DirPath.t
- let hashcons hdir = function
- | UniverseLevel.Set -> UniverseLevel.Set
- | UniverseLevel.Level (n,d) -> UniverseLevel.Level (n,hdir d)
- let equal l1 l2 =
- l1 == l2 ||
- match l1,l2 with
- | UniverseLevel.Set, UniverseLevel.Set -> true
- | UniverseLevel.Level (n,d), UniverseLevel.Level (n',d') ->
- n == n' && d == d'
- | _ -> false
- let hash = UniverseLevel.hash
- end)
-
-module Huniv =
- Hashcons.Make(
- struct
- type t = universe
+ type t = universe_set
type u = universe_level -> universe_level
- let hashcons hdir = function
- | Atom u -> Atom (hdir u)
- | Max (gel,gtl) -> Max (List.map hdir gel, List.map hdir gtl)
- let equal u v =
- u == v ||
- match u, v with
- | Atom u, Atom v -> u == v
- | Max (gel,gtl), Max (gel',gtl') ->
- (List.for_all2eq (==) gel gel') &&
- (List.for_all2eq (==) gtl gtl')
- | _ -> false
- let hash = Universe.hash
+ let hashcons huc s =
+ LSet.fold (fun x -> LSet.add (huc x)) s LSet.empty
+ let equal s s' =
+ LSet.equal s s'
+ let hash = Hashtbl.hash
end)
-let hcons_univlevel = Hashcons.simple_hcons Hunivlevel.generate Names.DirPath.hcons
-let hcons_univ = Hashcons.simple_hcons Huniv.generate hcons_univlevel
+let hcons_universe_set =
+ Hashcons.simple_hcons Huniverse_set.generate Level.hcons
-module Hconstraint =
- Hashcons.Make(
- struct
- type t = univ_constraint
- type u = universe_level -> universe_level
- let hashcons hul (l1,k,l2) = (hul l1, k, hul l2)
- let equal (l1,k,l2) (l1',k',l2') =
- l1 == l1' && k == k' && l2 == l2'
- let hash = Hashtbl.hash
- end)
+let hcons_universe_context_set (v, c) =
+ (hcons_universe_set v, hcons_constraints c)
-module UConstraintHash =
-struct
- type t = univ_constraint
- let hash = Hashtbl.hash
-end
-module Hconstraints = Set.Hashcons(UConstraintOrd)(UConstraintHash)
+let hcons_univlevel = Level.hcons
+let hcons_univ x = x (* Universe.hcons (Huniv.node x) *)
+let equal_universes = Universe.equal_universes
-let hcons_constraint = Hashcons.simple_hcons Hconstraint.generate hcons_univlevel
-let hcons_constraints = Hashcons.simple_hcons Hconstraints.generate hcons_constraint
+let explain_universe_inconsistency (o,u,v,p) =
+ let pr_rel = function
+ | Eq -> str"=" | Lt -> str"<" | Le -> str"<="
+ in
+ let reason = match p with
+ [] -> mt()
+ | _::_ ->
+ str " because" ++ spc() ++ pr_uni v ++
+ prlist (fun (r,v) -> spc() ++ pr_rel r ++ str" " ++ pr_uni v)
+ p ++
+ (if Universe.eq (snd (List.last p)) u then mt() else
+ (spc() ++ str "= " ++ pr_uni u))
+ in
+ str "Cannot enforce" ++ spc() ++ pr_uni u ++ spc() ++
+ pr_rel o ++ spc() ++ pr_uni v ++ reason ++ str")"
diff --git a/kernel/univ.mli b/kernel/univ.mli
index 04267de70d..9e7cc18b49 100644
--- a/kernel/univ.mli
+++ b/kernel/univ.mli
@@ -8,30 +8,67 @@
(** Universes. *)
-module UniverseLevel :
+module Level :
sig
type t
(** Type of universe levels. A universe level is essentially a unique name
that will be associated to constraints later on. *)
+ val set : t
+ val prop : t
+ val is_small : t -> bool
+
val compare : t -> t -> int
(** Comparison function *)
- val equal : t -> t -> bool
+ val eq : t -> t -> bool
(** Equality function *)
- val hash : t -> int
+ (* val hash : t -> int *)
(** Hash function *)
val make : Names.DirPath.t -> int -> t
(** Create a new universe level from a unique identifier and an associated
module path. *)
+ val pr : t -> Pp.std_ppcmds
end
-type universe_level = UniverseLevel.t
+type universe_level = Level.t
(** Alias name. *)
+module LSet :
+sig
+ include Set.S with type elt = universe_level
+
+ val pr : t -> Pp.std_ppcmds
+end
+
+type universe_set = LSet.t
+
+module LMap :
+sig
+ include Map.S with type key = universe_level
+
+ (** Favorizes the bindings in the first map. *)
+ val union : 'a t -> 'a t -> 'a t
+ val diff : 'a t -> 'a t -> 'a t
+
+ val subst_union : 'a option t -> 'a option t -> 'a option t
+
+ val elements : 'a t -> (universe_level * 'a) list
+ val of_list : (universe_level * 'a) list -> 'a t
+ val of_set : universe_set -> 'a -> 'a t
+ val mem : universe_level -> 'a t -> bool
+ val universes : 'a t -> universe_set
+
+ val find_opt : universe_level -> 'a t -> 'a option
+
+ val pr : ('a -> Pp.std_ppcmds) -> 'a t -> Pp.std_ppcmds
+end
+
+type 'a universe_map = 'a LMap.t
+
module Universe :
sig
type t
@@ -41,68 +78,260 @@ sig
val compare : t -> t -> int
(** Comparison function *)
- val equal : t -> t -> bool
+ val eq : t -> t -> bool
(** Equality function *)
- val hash : t -> int
+ (* val hash : t -> int *)
(** Hash function *)
- val make : UniverseLevel.t -> t
+ val make : Level.t -> t
(** Create a constraint-free universe out of a given level. *)
+ val pr : t -> Pp.std_ppcmds
+
+ val level : t -> Level.t option
+
+ val levels : t -> LSet.t
+
+ val normalize : t -> t
+
+ (** The type of a universe *)
+ val super : t -> t
+
+ (** The max of 2 universes *)
+ val sup : t -> t -> t
+
+ val type0m : t (** image of Prop in the universes hierarchy *)
+ val type0 : t (** image of Set in the universes hierarchy *)
+ val type1 : t (** the universe of the type of Prop/Set *)
end
type universe = Universe.t
-(** Alias name. *)
-module UniverseLSet : Set.S with type elt = universe_level
+(** Alias name. *)
+val pr_uni : universe -> Pp.std_ppcmds
+
(** The universes hierarchy: Type 0- = Prop <= Type 0 = Set <= Type 1 <= ...
Typing of universes: Type 0-, Type 0 : Type 1; Type i : Type (i+1) if i>0 *)
-
-val type0m_univ : universe (** image of Prop in the universes hierarchy *)
-val type0_univ : universe (** image of Set in the universes hierarchy *)
-val type1_univ : universe (** the universe of the type of Prop/Set *)
+val type0m_univ : universe
+val type0_univ : universe
+val type1_univ : universe
val is_type0_univ : universe -> bool
val is_type0m_univ : universe -> bool
val is_univ_variable : universe -> bool
+val is_small_univ : universe -> bool
-val universe_level : universe -> universe_level option
-
-(** The type of a universe *)
+val sup : universe -> universe -> universe
val super : universe -> universe
-(** The max of 2 universes *)
-val sup : universe -> universe -> universe
+val universe_level : universe -> universe_level option
+val compare_levels : universe_level -> universe_level -> int
+val eq_levels : universe_level -> universe_level -> bool
+
+(** Equality of formal universe expressions. *)
+val equal_universes : universe -> universe -> bool
(** {6 Graphs of universes. } *)
type universes
-type check_function = universes -> universe -> universe -> bool
-val check_leq : check_function
-val check_eq : check_function
-val lax_check_eq : check_function (* same, without anomaly *)
+type 'a check_function = universes -> 'a -> 'a -> bool
+val check_leq : universe check_function
+val check_eq : universe check_function
+val lax_check_eq : universe check_function (* same, without anomaly *)
(** The empty graph of universes *)
+val empty_universes : universes
+
+(** The initial graph of universes: Prop < Set *)
val initial_universes : universes
val is_initial_universes : universes -> bool
(** {6 Constraints. } *)
-type constraints
+type constraint_type = Lt | Le | Eq
+type univ_constraint = universe_level * constraint_type * universe_level
+
+module Constraint : sig
+ include Set.S with type elt = univ_constraint
+end
-val empty_constraint : constraints
-val union_constraints : constraints -> constraints -> constraints
+type constraints = Constraint.t
-val is_empty_constraint : constraints -> bool
+val empty_constraint : constraints
+val union_constraint : constraints -> constraints -> constraints
val eq_constraint : constraints -> constraints -> bool
-type constraint_function = universe -> universe -> constraints -> constraints
+type universe_constraint_type = ULe | UEq | ULub
+
+type universe_constraint = universe * universe_constraint_type * universe
+module UniverseConstraints : sig
+ include Set.S with type elt = universe_constraint
+
+ val pr : t -> Pp.std_ppcmds
+end
+
+type universe_constraints = UniverseConstraints.t
+type 'a universe_constrained = 'a * universe_constraints
+
+(** A value with universe constraints. *)
+type 'a constrained = 'a * constraints
+
+type universe_subst_fn = universe_level -> universe
+type universe_level_subst_fn = universe_level -> universe_level
+
+(** A full substitution, might involve algebraic universes *)
+type universe_subst = universe universe_map
+type universe_level_subst = universe_level universe_map
+
+val level_subst_of : universe_subst_fn -> universe_level_subst_fn
+
+module Instance :
+sig
+ type t
+
+ val hcons : t -> t
+ val empty : t
+ val is_empty : t -> bool
+
+ val eq : t -> t -> bool
+
+ val of_array : Level.t array -> t
+ val to_array : t -> Level.t array
+
+ (** Rely on physical equality of subterms only *)
+ val eqeq : t -> t -> bool
+
+ val subst_fn : universe_level_subst_fn -> t -> t
+ val subst : universe_level_subst -> t -> t
+
+ val pr : t -> Pp.std_ppcmds
+
+ val append : t -> t -> t
-val enforce_leq : constraint_function
-val enforce_eq : constraint_function
+ val levels : t -> LSet.t
+
+ val max_level : t -> int
+
+ val check_eq : t check_function
+
+end
+
+type universe_instance = Instance.t
+
+type 'a puniverses = 'a * universe_instance
+val out_punivs : 'a puniverses -> 'a
+val in_punivs : 'a -> 'a puniverses
+
+(** A list of universes with universe constraints,
+ representiong local universe variables and constraints *)
+
+module UContext :
+sig
+ type t
+
+ val make : Instance.t constrained -> t
+ val empty : t
+ val is_empty : t -> bool
+
+ val instance : t -> Instance.t
+ val constraints : t -> constraints
+
+ (** Keeps the order of the instances *)
+ val union : t -> t -> t
+
+end
+
+type universe_context = UContext.t
+
+(** Universe contexts (as sets) *)
+
+module ContextSet :
+sig
+ type t = universe_set constrained
+
+ val empty : t
+ val is_empty : t -> bool
+
+ val singleton : universe_level -> t
+ val of_instance : Instance.t -> t
+ val of_set : universe_set -> t
+
+ val union : t -> t -> t
+ val diff : t -> t -> t
+ val add_constraints : t -> constraints -> t
+ val add_universes : Instance.t -> t -> t
+
+ (** Arbitrary choice of linear order of the variables
+ and normalization of the constraints *)
+ val to_context : t -> universe_context
+ val of_context : universe_context -> t
+
+ val constraints : t -> constraints
+ val levels : t -> universe_set
+end
+
+(** A set of universes with universe constraints.
+ We linearize the set to a list after typechecking.
+ Beware, representation could change.
+*)
+type universe_context_set = ContextSet.t
+
+(** A value in a universe context (resp. context set). *)
+type 'a in_universe_context = 'a * universe_context
+type 'a in_universe_context_set = 'a * universe_context_set
+
+(** Constrained *)
+val constraints_of : 'a constrained -> constraints
+
+
+(** [check_context_subset s s'] checks that [s] is implied by [s'] as a set of constraints,
+ and shrinks [s'] to the set of variables declared in [s].
+. *)
+val check_context_subset : universe_context_set -> universe_context -> universe_context
+
+(** Make a universe level substitution: the list must match the context variables. *)
+val make_universe_subst : Instance.t -> universe_context -> universe_subst
+val empty_subst : universe_subst
+val is_empty_subst : universe_subst -> bool
+
+val empty_level_subst : universe_level_subst
+val is_empty_level_subst : universe_level_subst -> bool
+
+(** Get the instantiated graph. *)
+val instantiate_univ_context : universe_subst -> universe_context -> constraints
+
+(** Substitution of universes. *)
+val subst_univs_level_level : universe_level_subst -> universe_level -> universe_level
+val subst_univs_level_universe : universe_level_subst -> universe -> universe
+val subst_univs_level_constraints : universe_level_subst -> constraints -> constraints
+
+val normalize_univs_level_level : universe_level_subst -> universe_level -> universe_level
+
+val make_subst : universe_subst -> universe_subst_fn
+
+(* val subst_univs_level_fail : universe_subst_fn -> universe_level -> universe_level *)
+val subst_univs_level : universe_subst_fn -> universe_level -> universe
+val subst_univs_universe : universe_subst_fn -> universe -> universe
+val subst_univs_constraints : universe_subst_fn -> constraints -> constraints
+val subst_univs_universe_constraints : universe_subst_fn -> universe_constraints -> universe_constraints
+
+(** Raises universe inconsistency if not compatible. *)
+val check_consistent_constraints : universe_context_set -> constraints -> unit
+
+type 'a constraint_function = 'a -> 'a -> constraints -> constraints
+
+val enforce_leq : universe constraint_function
+val enforce_eq : universe constraint_function
+val enforce_eq_level : universe_level constraint_function
+val enforce_leq_level : universe_level constraint_function
+val enforce_eq_instances : universe_instance constraint_function
+
+type 'a universe_constraint_function = 'a -> 'a -> universe_constraints -> universe_constraints
+
+val enforce_eq_instances_univs : bool -> universe_instance universe_constraint_function
(** {6 ... } *)
(** Merge of constraints in a universes graph.
@@ -110,8 +339,6 @@ val enforce_eq : constraint_function
universes graph. It raises the exception [UniverseInconsistency] if the
constraints are not satisfiable. *)
-type constraint_type = Lt | Le | Eq
-
(** Type explanation is used to decorate error messages to provide
useful explanation why a given constraint is rejected. It is composed
of a path of universes and relation kinds [(r1,u1);..;(rn,un)] means
@@ -125,20 +352,26 @@ type constraint_type = Lt | Le | Eq
constraints...
*)
type explanation = (constraint_type * universe) list
+type univ_inconsistency = constraint_type * universe * universe * explanation
-exception UniverseInconsistency of
- constraint_type * universe * universe * explanation
+exception UniverseInconsistency of univ_inconsistency
+val enforce_constraint : univ_constraint -> universes -> universes
val merge_constraints : constraints -> universes -> universes
val normalize_universes : universes -> universes
val sort_universes : universes -> universes
-(** {6 Support for sort-polymorphic inductive types } *)
+val constraints_of_universes : universes -> constraints
+
+val to_constraints : universes -> universe_constraints -> constraints
+
+val check_constraint : universes -> univ_constraint -> bool
+val check_constraints : constraints -> universes -> bool
-val fresh_local_univ : unit -> universe
-val set_remote_fresh_local_univ : universe RemoteCounter.installer
-val solve_constraints_system : universe option array -> universe array ->
+(** {6 Support for sort-polymorphism } *)
+
+val solve_constraints_system : universe option array -> universe array -> universe array ->
universe array
val subst_large_constraint : universe -> universe -> universe -> universe
@@ -154,10 +387,15 @@ val univ_depends : universe -> universe -> bool
(** {6 Pretty-printing of universes. } *)
-val pr_uni_level : universe_level -> Pp.std_ppcmds
-val pr_uni : universe -> Pp.std_ppcmds
val pr_universes : universes -> Pp.std_ppcmds
+val pr_constraint_type : constraint_type -> Pp.std_ppcmds
val pr_constraints : constraints -> Pp.std_ppcmds
+(* val pr_universe_list : universe_list -> Pp.std_ppcmds *)
+val pr_universe_context : universe_context -> Pp.std_ppcmds
+val pr_universe_context_set : universe_context_set -> Pp.std_ppcmds
+val pr_universe_level_subst : universe_level_subst -> Pp.std_ppcmds
+val pr_universe_subst : universe_subst -> Pp.std_ppcmds
+val explain_universe_inconsistency : univ_inconsistency -> Pp.std_ppcmds
(** {6 Dumping to a file } *)
@@ -170,3 +408,8 @@ val dump_universes :
val hcons_univlevel : universe_level -> universe_level
val hcons_univ : universe -> universe
val hcons_constraints : constraints -> constraints
+val hcons_universe_set : universe_set -> universe_set
+val hcons_universe_context : universe_context -> universe_context
+val hcons_universe_context_set : universe_context_set -> universe_context_set
+
+(******)
diff --git a/kernel/vars.ml b/kernel/vars.ml
index f23d5fc2cf..3cff51ba6d 100644
--- a/kernel/vars.ml
+++ b/kernel/vars.ml
@@ -212,3 +212,89 @@ let substn_vars p vars c =
in replace_vars (List.rev subst) c
let subst_vars subst c = substn_vars 1 subst c
+
+(** Universe substitutions *)
+open Constr
+
+let subst_univs_puniverses subst =
+ if Univ.is_empty_level_subst subst then fun c -> c
+ else
+ let f = Univ.Instance.subst subst in
+ fun ((c, u) as x) -> let u' = f u in if u' == u then x else (c, u')
+
+let subst_univs_fn_puniverses fn =
+ let f = Univ.Instance.subst_fn fn in
+ fun ((c, u) as x) -> let u' = f u in if u' == u then x else (c, u')
+
+let subst_univs_fn_constr f c =
+ let changed = ref false in
+ let fu = Univ.subst_univs_universe f in
+ let fi = Univ.Instance.subst_fn (Univ.level_subst_of f) in
+ let rec aux t =
+ match kind t with
+ | Sort (Sorts.Type u) ->
+ let u' = fu u in
+ if u' == u then t else
+ (changed := true; mkSort (Sorts.sort_of_univ u'))
+ | Const (c, u) ->
+ let u' = fi u in
+ if u' == u then t
+ else (changed := true; mkConstU (c, u'))
+ | Ind (i, u) ->
+ let u' = fi u in
+ if u' == u then t
+ else (changed := true; mkIndU (i, u'))
+ | Construct (c, u) ->
+ let u' = fi u in
+ if u' == u then t
+ else (changed := true; mkConstructU (c, u'))
+ | _ -> map aux t
+ in
+ let c' = aux c in
+ if !changed then c' else c
+
+let subst_univs_constr subst c =
+ if Univ.is_empty_subst subst then c
+ else
+ let f = Univ.make_subst subst in
+ subst_univs_fn_constr f c
+
+(* let subst_univs_constr_key = Profile.declare_profile "subst_univs_constr" *)
+(* let subst_univs_constr = Profile.profile2 subst_univs_constr_key subst_univs_constr *)
+
+let subst_univs_level_constr subst c =
+ if Univ.is_empty_level_subst subst then c
+ else
+ let f = Univ.Instance.subst_fn (Univ.subst_univs_level_level subst) in
+ let changed = ref false in
+ let rec aux t =
+ match kind t with
+ | Const (c, u) ->
+ if Univ.Instance.is_empty u then t
+ else
+ let u' = f u in
+ if u' == u then t
+ else (changed := true; mkConstU (c, u'))
+ | Ind (i, u) ->
+ if Univ.Instance.is_empty u then t
+ else
+ let u' = f u in
+ if u' == u then t
+ else (changed := true; mkIndU (i, u'))
+ | Construct (c, u) ->
+ if Univ.Instance.is_empty u then t
+ else
+ let u' = f u in
+ if u' == u then t
+ else (changed := true; mkConstructU (c, u'))
+ | Sort (Sorts.Type u) ->
+ let u' = Univ.subst_univs_level_universe subst u in
+ if u' == u then t else
+ (changed := true; mkSort (Sorts.sort_of_univ u'))
+ | _ -> Constr.map aux t
+ in
+ let c' = aux c in
+ if !changed then c' else c
+
+let subst_univs_context s =
+ map_rel_context (subst_univs_constr s)
diff --git a/kernel/vars.mli b/kernel/vars.mli
index ef3381ab59..9d5d16d0c0 100644
--- a/kernel/vars.mli
+++ b/kernel/vars.mli
@@ -68,3 +68,17 @@ val subst_vars : Id.t list -> constr -> constr
(** [substn_vars n [id1;...;idn] t] substitute [VAR idj] by [Rel j+n-1] in [t]
if two names are identical, the one of least indice is kept *)
val substn_vars : int -> Id.t list -> constr -> constr
+
+(** {3 Substitution of universes} *)
+
+open Univ
+
+val subst_univs_fn_constr : universe_subst_fn -> constr -> constr
+val subst_univs_fn_puniverses : universe_level_subst_fn ->
+ 'a puniverses -> 'a puniverses
+
+val subst_univs_constr : universe_subst -> constr -> constr
+val subst_univs_puniverses : universe_level_subst -> 'a puniverses -> 'a puniverses
+val subst_univs_level_constr : universe_level_subst -> constr -> constr
+
+val subst_univs_context : Univ.universe_subst -> rel_context -> rel_context
diff --git a/kernel/vconv.ml b/kernel/vconv.ml
index 484ee2a50b..62ddeb182d 100644
--- a/kernel/vconv.ml
+++ b/kernel/vconv.ml
@@ -42,13 +42,15 @@ let conv_vect fconv vect1 vect2 cu =
let infos = ref (create_clos_infos betaiotazeta Environ.empty_env)
+let eq_table_key = Names.eq_table_key eq_constant
+
let rec conv_val pb k v1 v2 cu =
if v1 == v2 then cu
else conv_whd pb k (whd_val v1) (whd_val v2) cu
and conv_whd pb k whd1 whd2 cu =
match whd1, whd2 with
- | Vsort s1, Vsort s2 -> sort_cmp pb s1 s2 cu
+ | Vsort s1, Vsort s2 -> ignore(sort_cmp_universes pb s1 s2 (cu,None)); cu
| Vprod p1, Vprod p2 ->
let cu = conv_val CONV k (dom p1) (dom p2) cu in
conv_fun pb k (codom p1) (codom p2) cu
@@ -169,6 +171,13 @@ and conv_arguments k args1 args2 cu =
!rcu
else raise NotConvertible
+let rec eq_puniverses f (x,l1) (y,l2) cu =
+ if f x y then conv_universes l1 l2 cu
+ else raise NotConvertible
+
+and conv_universes l1 l2 cu =
+ if Univ.Instance.eq l1 l2 then cu else raise NotConvertible
+
let rec conv_eq pb t1 t2 cu =
if t1 == t2 then cu
else
@@ -179,7 +188,7 @@ let rec conv_eq pb t1 t2 cu =
if Int.equal m1 m2 then cu else raise NotConvertible
| Var id1, Var id2 ->
if Id.equal id1 id2 then cu else raise NotConvertible
- | Sort s1, Sort s2 -> sort_cmp pb s1 s2 cu
+ | Sort s1, Sort s2 -> ignore(sort_cmp_universes pb s1 s2 (cu,None)); cu
| Cast (c1,_,_), _ -> conv_eq pb c1 t2 cu
| _, Cast (c2,_,_) -> conv_eq pb t1 c2 cu
| Prod (_,t1,c1), Prod (_,t2,c2) ->
@@ -192,12 +201,13 @@ let rec conv_eq pb t1 t2 cu =
| Evar (e1,l1), Evar (e2,l2) ->
if Evar.equal e1 e2 then conv_eq_vect l1 l2 cu
else raise NotConvertible
- | Const c1, Const c2 ->
- if eq_constant c1 c2 then cu else raise NotConvertible
+ | Const c1, Const c2 -> eq_puniverses eq_constant c1 c2 cu
+ | Proj (p1,c1), Proj (p2,c2) ->
+ if eq_constant p1 p2 then conv_eq pb c1 c2 cu else raise NotConvertible
| Ind c1, Ind c2 ->
- if eq_ind c1 c2 then cu else raise NotConvertible
+ eq_puniverses eq_ind c1 c2 cu
| Construct c1, Construct c2 ->
- if eq_constructor c1 c2 then cu else raise NotConvertible
+ eq_puniverses eq_constructor c1 c2 cu
| Case (_,p1,c1,bl1), Case (_,p2,c2,bl2) ->
let pcu = conv_eq CONV p1 p2 cu in
let ccu = conv_eq CONV c1 c2 pcu in
@@ -221,14 +231,14 @@ and conv_eq_vect vt1 vt2 cu =
let vconv pb env t1 t2 =
infos := create_clos_infos betaiotazeta env;
- let cu =
- try conv_eq pb t1 t2 empty_constraint
+ let _cu =
+ try conv_eq pb t1 t2 (universes env)
with NotConvertible ->
let v1 = val_of_constr env t1 in
let v2 = val_of_constr env t2 in
- let cu = conv_val pb (nb_rel env) v1 v2 empty_constraint in
+ let cu = conv_val pb (nb_rel env) v1 v2 (universes env) in
cu
- in cu
+ in ()
let _ = Reduction.set_vm_conv vconv