aboutsummaryrefslogtreecommitdiffhomepage
path: root/kernel
diff options
context:
space:
mode:
authorGravatar barras <barras@85f007b7-540e-0410-9357-904b9bb8a0f7>2001-11-05 16:48:30 +0000
committerGravatar barras <barras@85f007b7-540e-0410-9357-904b9bb8a0f7>2001-11-05 16:48:30 +0000
commitb91f60aab99980b604dc379b4ca62f152315c841 (patch)
treecd1948fc5156988dd74d94ef4abb3e4ac77e3de8 /kernel
parent2ff72589e5c90a25b315922b5ba2d7c11698adef (diff)
GROS COMMIT:
- reduction du noyau (variables existentielles, fonctions auxiliaires pour inventer des noms, etc. deplacees hors de kernel/) - changement de noms de constructeurs des constr (suppression de "Is" et "Mut") git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@2158 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'kernel')
-rw-r--r--kernel/closure.ml200
-rw-r--r--kernel/closure.mli53
-rw-r--r--kernel/cooking.ml65
-rw-r--r--kernel/declarations.ml90
-rw-r--r--kernel/declarations.mli69
-rw-r--r--kernel/environ.ml380
-rw-r--r--kernel/environ.mli153
-rw-r--r--kernel/evd.ml74
-rw-r--r--kernel/evd.mli57
-rw-r--r--kernel/indtypes.ml256
-rw-r--r--kernel/indtypes.mli40
-rw-r--r--kernel/inductive.ml966
-rw-r--r--kernel/inductive.mli243
-rw-r--r--kernel/instantiate.ml147
-rw-r--r--kernel/instantiate.mli63
-rw-r--r--kernel/names.ml246
-rw-r--r--kernel/names.mli78
-rw-r--r--kernel/reduction.ml727
-rw-r--r--kernel/reduction.mli200
-rw-r--r--kernel/safe_typing.ml435
-rw-r--r--kernel/safe_typing.mli70
-rw-r--r--kernel/sign.ml82
-rw-r--r--kernel/sign.mli78
-rw-r--r--kernel/term.ml1124
-rw-r--r--kernel/term.mli476
-rw-r--r--kernel/type_errors.ml70
-rw-r--r--kernel/type_errors.mli44
-rw-r--r--kernel/typeops.ml1146
-rw-r--r--kernel/typeops.mli104
-rw-r--r--kernel/univ.ml2
-rw-r--r--kernel/univ.mli10
31 files changed, 2496 insertions, 5252 deletions
diff --git a/kernel/closure.ml b/kernel/closure.ml
index 283b60e28..4bb9c941f 100644
--- a/kernel/closure.ml
+++ b/kernel/closure.ml
@@ -13,9 +13,7 @@ open Pp
open Term
open Names
open Environ
-open Instantiate
open Univ
-open Evd
open Esubst
@@ -67,7 +65,6 @@ module type RedFlagsSig = sig
type reds
type red_kind
val fBETA : red_kind
- val fEVAR : red_kind
val fDELTA : red_kind
val fIOTA : red_kind
val fZETA : red_kind
@@ -96,11 +93,10 @@ module RedFlags = (struct
r_evar : bool;
r_iota : bool }
- type red_kind = BETA | DELTA | EVAR | IOTA | ZETA
+ type red_kind = BETA | DELTA | IOTA | ZETA
| CONST of constant | VAR of identifier
let fBETA = BETA
let fDELTA = DELTA
- let fEVAR = EVAR
let fIOTA = IOTA
let fZETA = ZETA
let fCONST sp = CONST sp
@@ -120,7 +116,6 @@ module RedFlags = (struct
let (l1,l2) = red.r_const in
{ red with r_const = l1, Sppred.add sp l2 }
| IOTA -> { red with r_iota = true }
- | EVAR -> { red with r_evar = true }
| ZETA -> { red with r_zeta = true }
| VAR id ->
let (l1,l2) = red.r_const in
@@ -133,7 +128,6 @@ module RedFlags = (struct
let (l1,l2) = red.r_const in
{ red with r_const = l1, Sppred.remove sp l2 }
| IOTA -> { red with r_iota = false }
- | EVAR -> { red with r_evar = false }
| ZETA -> { red with r_zeta = false }
| VAR id ->
let (l1,l2) = red.r_const in
@@ -155,7 +149,6 @@ module RedFlags = (struct
let c = Idpred.mem id l in
incr_cnt c delta
| ZETA -> incr_cnt red.r_zeta zeta
- | EVAR -> incr_cnt red.r_zeta evar
| IOTA -> incr_cnt red.r_iota iota
| DELTA -> (* Used for Rel/Var defined in context *)
incr_cnt red.r_delta delta
@@ -174,7 +167,8 @@ end : RedFlagsSig)
open RedFlags
-let betadeltaiota_red = mkflags [fBETA;fDELTA;fZETA;fEVAR;fIOTA]
+let betadeltaiota_red = mkflags [fBETA;fDELTA;fZETA;fIOTA]
+let betadeltaiotanolet_red = mkflags [fBETA;fDELTA;fIOTA]
let betaiota_red = mkflags [fBETA;fIOTA]
let beta_red = mkflags [fBETA]
let betaiotazeta_red = mkflags [fBETA;fIOTA;fZETA]
@@ -248,7 +242,7 @@ let unfold_red sp =
a LetIn expression is Letin reduction *)
type red_kind =
- BETA | DELTA | ZETA | EVAR | IOTA
+ BETA | DELTA | ZETA | IOTA
| CONST of constant_path list | CONSTBUT of constant_path list
| VAR of identifier | VARBUT of identifier
@@ -270,7 +264,6 @@ let rec red_add red = function
{ red with r_const = true, list_union cl l1, l2;
r_zeta = true; r_evar = true })
| IOTA -> { red with r_iota = true }
- | EVAR -> { red with r_evar = true }
| ZETA -> { red with r_zeta = true }
| VAR id ->
(match red.r_const with
@@ -331,6 +324,7 @@ let no_flag = (UNIFORM,no_red)
let beta = (UNIFORM,beta_red)
let betaiota = (UNIFORM,betaiota_red)
let betadeltaiota = (UNIFORM,betadeltaiota_red)
+let betadeltaiotanolet = (UNIFORM,betadeltaiotanolet_red)
let hnf_flags = (SIMPL,betaiotazeta_red)
let unfold_flags sp = (UNIFORM, unfold_red sp)
@@ -362,7 +356,6 @@ let red_under (md,r) rk =
* mapped to constr. 'a infos implements a cache for constants and
* abstractions, storing a representation (of type 'a) of the body of
* this constant or abstraction.
- * * i_evc is the set of constraints for existential variables
* * i_tab is the cache table of the results
* * i_repr is the function to get the representation from the current
* state of the cache and the body of the constant. The result
@@ -379,20 +372,19 @@ let red_under (md,r) rk =
* instantiations (cbv or lazy) are.
*)
-type 'a table_key =
- | ConstBinding of constant
- | EvarBinding of existential
- | VarBinding of identifier
- | FarRelBinding of int
+type table_key =
+ | ConstKey of constant
+ | VarKey of identifier
+ | FarRelKey of int
+ (* FarRel: index in the rel_context part of _initial_ environment *)
-type ('a, 'b) infos = {
+type 'a infos = {
i_flags : flags;
- i_repr : ('a, 'b) infos -> constr -> 'a;
+ i_repr : 'a infos -> constr -> 'a;
i_env : env;
- i_evc : 'b evar_map;
i_rels : int * (int * constr) list;
i_vars : (identifier * constr) list;
- i_tab : ('a table_key, 'a) Hashtbl.t }
+ i_tab : (table_key, 'a) Hashtbl.t }
let info_flags info = info.i_flags
@@ -403,18 +395,16 @@ let ref_value_cache info ref =
try
let body =
match ref with
- | FarRelBinding n ->
+ | FarRelKey n ->
let (s,l) = info.i_rels in lift n (List.assoc (s-n) l)
- | VarBinding id -> List.assoc id info.i_vars
- | EvarBinding evc -> existential_value info.i_evc evc
- | ConstBinding cst -> constant_value info.i_env cst
+ | VarKey id -> List.assoc id info.i_vars
+ | ConstKey cst -> constant_value info.i_env cst
in
let v = info.i_repr info body in
Hashtbl.add info.i_tab ref v;
Some v
with
| Not_found (* List.assoc *)
- | NotInstantiatedEvar (* Evar *)
| NotEvaluableConst _ (* Const *)
-> None
@@ -438,11 +428,10 @@ let defined_rels flags env =
env (0,[])
(* else (0,[])*)
-let create mk_cl flgs env sigma =
+let create mk_cl flgs env =
{ i_flags = flgs;
i_repr = mk_cl;
i_env = env;
- i_evc = sigma;
i_rels = defined_rels flgs env;
i_vars = defined_vars flgs env;
i_tab = Hashtbl.create 17 }
@@ -549,7 +538,7 @@ let rec stack_nth s p = match s with
(* Lazy reduction: the one used in kernel operations *)
(* type of shared terms. fconstr and frterm are mutually recursive.
- * Clone of the Generic.term structure, but completely mutable, and
+ * Clone of the constr structure, but completely mutable, and
* annotated with booleans (true when we noticed that the term is
* normal and neutral) FLIFT is a delayed shift; allows sharing
* between 2 lifted copies of a given term FCLOS is a delayed
@@ -565,7 +554,7 @@ and fterm =
| FRel of int
| FAtom of constr
| FCast of fconstr * fconstr
- | FFlex of freference
+ | FFlex of table_key
| FInd of inductive
| FConstruct of constructor
| FApp of fconstr * fconstr array
@@ -577,17 +566,11 @@ and fterm =
| FLambda of name * fconstr * fconstr * constr * fconstr subs
| FProd of name * fconstr * fconstr * constr * fconstr subs
| FLetIn of name * fconstr * fconstr * fconstr * constr * fconstr subs
+ | FEvar of existential_key * fconstr array
| FLIFT of int * fconstr
| FCLOS of constr * fconstr subs
| FLOCKED
-and freference =
- (* only vars as args of FConst ... exploited for caching *)
- | FConst of constant
- | FEvar of existential_key * fconstr array
- | FVar of identifier
- | FFarRel of int (* index in the rel_context part of _initial_ environment *)
-
let fterm_of v = v.term
let set_whnf v = if v.norm = Red then v.norm <- Whnf
let set_cstr v = if v.norm = Red then v.norm <- Cstr
@@ -646,19 +629,22 @@ let clos_rel e i =
| Inl(n,mt) -> lift_fconstr n mt
| Inr(k,None) -> {norm=Red; term= FRel k}
| Inr(k,Some p) ->
- lift_fconstr (k-p) {norm=Norm;term=FFlex(FFarRel p)}
+ lift_fconstr (k-p) {norm=Norm;term=FFlex(FarRelKey p)}
(* Optimization: do not enclose variables in a closure.
Makes variable access much faster *)
-let mk_clos e t =
+let rec mk_clos e t =
match kind_of_term t with
- | IsRel i -> clos_rel e i
- | IsVar x -> { norm = Red; term = FFlex (FVar x) }
- | IsMeta _ | IsSort _ -> { norm = Norm; term = FAtom t }
- | (IsMutInd _|IsMutConstruct _|IsFix _|IsCoFix _
- |IsLambda _|IsProd _) ->
+ | Rel i -> clos_rel e i
+ | Var x -> { norm = Red; term = FFlex (VarKey x) }
+ | Meta _ | Sort _ -> { norm = Norm; term = FAtom t }
+ | Ind sp -> { norm = Norm; term = FInd sp }
+ | Construct sp -> { norm = Cstr; term = FConstruct sp }
+ | Evar (ev,args) ->
+ { norm = Cstr; term = FEvar (ev,Array.map (mk_clos e) args) }
+ | (Fix _|CoFix _|Lambda _|Prod _) ->
{norm = Cstr; term = FCLOS(t,e)}
- | (IsApp _|IsMutCase _|IsCast _|IsConst _|IsEvar _|IsLetIn _) ->
+ | (App _|Case _|Cast _|Const _|LetIn _) ->
{norm = Red; term = FCLOS(t,e)}
let mk_clos_vect env v = Array.map (mk_clos env) v
@@ -669,55 +655,46 @@ let mk_clos_vect env v = Array.map (mk_clos env) v
Could be used insted of mk_clos. *)
let mk_clos_deep clos_fun env t =
match kind_of_term t with
- | IsRel i -> clos_rel env i
- | (IsVar _|IsMeta _ | IsSort _) -> mk_clos env t
- | IsCast (a,b) ->
+ | (Rel _|Ind _|Construct _|Var _|Meta _ | Sort _|Evar _) ->
+ mk_clos env t
+ | Cast (a,b) ->
{ norm = Red;
term = FCast (clos_fun env a, clos_fun env b)}
- | IsApp (f,v) ->
+ | App (f,v) ->
{ norm = Red;
term = FApp (clos_fun env f, Array.map (clos_fun env) v) }
- | IsMutInd sp ->
- { norm = Norm; term = FInd sp }
- | IsMutConstruct sp ->
- { norm = Norm; term = FConstruct sp }
- | IsConst sp ->
+ | Const sp ->
{ norm = Red;
- term = FFlex (FConst sp) }
- | IsEvar (n,v) ->
- { norm = Red;
- term = FFlex (FEvar (n, Array.map (clos_fun env) v)) }
-
- | IsMutCase (ci,p,c,v) ->
+ term = FFlex (ConstKey sp) }
+ | Case (ci,p,c,v) ->
{ norm = Red;
term = FCases (ci, clos_fun env p, clos_fun env c,
Array.map (clos_fun env) v) }
- | IsFix (op,(lna,tys,bds)) ->
+ | Fix (op,(lna,tys,bds)) ->
let env' = subs_liftn (Array.length bds) env in
{ norm = Cstr;
term = FFix
(op,(lna, Array.map (clos_fun env) tys,
Array.map (clos_fun env') bds),
bds, env) }
- | IsCoFix (op,(lna,tys,bds)) ->
+ | CoFix (op,(lna,tys,bds)) ->
let env' = subs_liftn (Array.length bds) env in
{ norm = Cstr;
term = FCoFix
(op,(lna, Array.map (clos_fun env) tys,
Array.map (clos_fun env') bds),
bds, env) }
-
- | IsLambda (n,t,c) ->
+ | Lambda (n,t,c) ->
{ norm = Cstr;
term = FLambda (n, clos_fun env t,
clos_fun (subs_lift env) c,
c, env) }
- | IsProd (n,t,c) ->
+ | Prod (n,t,c) ->
{ norm = Cstr;
term = FProd (n, clos_fun env t,
clos_fun (subs_lift env) c,
c, env) }
- | IsLetIn (n,b,t,c) ->
+ | LetIn (n,b,t,c) ->
{ norm = Red;
term = FLetIn (n, clos_fun env b, clos_fun env t,
clos_fun (subs_lift env) c,
@@ -727,24 +704,22 @@ let mk_clos_deep clos_fun env t =
let rec to_constr constr_fun lfts v =
match v.term with
| FRel i -> mkRel (reloc_rel i lfts)
- | FFlex (FFarRel p) -> mkRel (reloc_rel p lfts)
- | FFlex (FVar x) -> mkVar x
+ | FFlex (FarRelKey p) -> mkRel (reloc_rel p lfts)
+ | FFlex (VarKey x) -> mkVar x
| FAtom c ->
(match kind_of_term c with
- | IsSort s -> mkSort s
- | IsMeta m -> mkMeta m
+ | Sort s -> mkSort s
+ | Meta m -> mkMeta m
| _ -> assert false)
| FCast (a,b) ->
mkCast (constr_fun lfts a, constr_fun lfts b)
- | FFlex (FConst op) -> mkConst op
- | FFlex (FEvar (n,args)) ->
- mkEvar (n, Array.map (constr_fun lfts) args)
- | FInd op -> mkMutInd op
- | FConstruct op -> mkMutConstruct op
+ | FFlex (ConstKey op) -> mkConst op
+ | FInd op -> mkInd op
+ | FConstruct op -> mkConstruct op
| FCases (ci,p,c,ve) ->
- mkMutCase (ci, constr_fun lfts p,
- constr_fun lfts c,
- Array.map (constr_fun lfts) ve)
+ mkCase (ci, constr_fun lfts p,
+ constr_fun lfts c,
+ Array.map (constr_fun lfts) ve)
| FFix (op,(lna,tys,bds),_,_) ->
let lfts' = el_liftn (Array.length bds) lfts in
mkFix (op, (lna, Array.map (constr_fun lfts) tys,
@@ -766,6 +741,7 @@ let rec to_constr constr_fun lfts v =
mkLetIn (n, constr_fun lfts b,
constr_fun lfts t,
constr_fun (el_lift lfts) c)
+ | FEvar (ev,args) -> mkEvar(ev,Array.map (constr_fun lfts) args)
| FLIFT (k,a) -> to_constr constr_fun (el_shft k lfts) a
| FCLOS (t,env) ->
let fr = mk_clos_deep mk_clos env t in
@@ -952,23 +928,23 @@ let rec knh m stk =
| (None, stk') -> (m,stk'))
| FCast(t,_) -> knh t stk
(* cases where knh stops *)
- | (FFlex _|FLetIn _) -> (m, stk)
- | (FRel _|FAtom _) -> (set_norm m; (m, stk))
- | (FLambda _|FConstruct _|FCoFix _|FInd _|FProd _) ->
+ | (FFlex _|FLetIn _|FInd _|FConstruct _|FEvar _) -> (m, stk)
+ | (FRel _|FAtom _|FInd _) -> (set_norm m; (m, stk))
+ | (FLambda _|FCoFix _|FProd _) ->
(set_whnf m; (m, stk))
(* The same for pure terms *)
and knht e t stk =
match kind_of_term t with
- | IsApp(a,b) ->
+ | App(a,b) ->
knht e a (append_stack (mk_clos_vect e b) stk)
- | IsMutCase(ci,p,t,br) ->
+ | Case(ci,p,t,br) ->
knht e t (Zcase(ci, mk_clos e p, mk_clos_vect e br)::stk)
- | IsFix _ -> knh (mk_clos_deep mk_clos e t) stk
- | IsCast(a,b) -> knht e a stk
- | IsRel n -> knh (clos_rel e n) stk
- | (IsLambda _|IsProd _|IsMutConstruct _|IsCoFix _|IsMutInd _|
- IsLetIn _|IsConst _|IsVar _|IsEvar _|IsMeta _|IsSort _) ->
+ | Fix _ -> knh (mk_clos_deep mk_clos e t) stk
+ | Cast(a,b) -> knht e a stk
+ | Rel n -> knh (clos_rel e n) stk
+ | (Lambda _|Prod _|Construct _|CoFix _|Ind _|
+ LetIn _|Const _|Var _|Evar _|Meta _|Sort _) ->
(mk_clos_deep mk_clos e t, stk)
@@ -981,30 +957,23 @@ let rec knr info m stk =
(match get_arg m stk with
(Some(depth,arg),s) -> knit info (subs_shift_cons(depth,e,arg)) f s
| (None,s) -> (m,s))
- | FFlex(FConst sp) when can_red info stk (fCONST sp) ->
- (match ref_value_cache info (ConstBinding sp) with
+ | FFlex(ConstKey sp) when can_red info stk (fCONST sp) ->
+ (match ref_value_cache info (ConstKey sp) with
Some v -> kni info v stk
| None -> (set_norm m; (m,stk)))
- | FFlex(FEvar (n,args)) when can_red info stk fEVAR ->
-(* In the case of evars, if it is not defined, then we do not set the
- flag to Norm, because it may be instantiated later on *)
- let evar = (n, Array.map term_of_fconstr args) in
- (match ref_value_cache info (EvarBinding evar) with
- Some v -> kni info v stk
- | None -> (m,stk))
- | FFlex(FVar id) when can_red info stk (fVAR id) ->
- (match ref_value_cache info (VarBinding id) with
+ | FFlex(VarKey id) when can_red info stk (fVAR id) ->
+ (match ref_value_cache info (VarKey id) with
Some v -> kni info v stk
| None -> (set_norm m; (m,stk)))
- | FFlex(FFarRel k) when can_red info stk fDELTA ->
- (match ref_value_cache info (FarRelBinding k) with
+ | FFlex(FarRelKey k) when can_red info stk fDELTA ->
+ (match ref_value_cache info (FarRelKey k) with
Some v -> kni info v stk
| None -> (set_norm m; (m,stk)))
| FConstruct(ind,c) when can_red info stk fIOTA ->
(match strip_update_shift_app m stk with
- (depth, args, Zcase(((*cn*) npar,_),_,br)::s) ->
- assert (npar>=0);
- let rargs = drop_parameters depth npar args in
+ (depth, args, Zcase(ci,_,br)::s) ->
+ assert (ci.ci_npar>=0);
+ let rargs = drop_parameters depth ci.ci_npar args in
kni info br.(c-1) (rargs@s)
| (_, cargs, Zfix(fx,par)::s) ->
let rarg = fapp_stack(m,cargs) in
@@ -1014,7 +983,7 @@ let rec knr info m stk =
| (_,args,s) -> (m,args@s))
| FCoFix _ when can_red info stk fIOTA ->
(match strip_update_shift_app m stk with
- (_, args, ((Zcase((cn,_),_,br)::_) as stk')) ->
+ (_, args, ((Zcase _::_) as stk')) ->
let efx = contract_fix_vect m.term in
kni info efx (args@stk')
| (_,args,s) -> (m,args@s))
@@ -1060,8 +1029,7 @@ and down_then_up info m stk =
| FCoFix(n,(na,ftys,fbds),bds,e) ->
FCoFix(n,(na, Array.map (kl info) ftys,
Array.map (kl info) fbds),bds,e)
- | FFlex(FEvar(i,args)) ->
- FFlex(FEvar(i, Array.map (kl info) args))
+ | FEvar(i,args) -> FEvar(i, Array.map (kl info) args)
| t -> t in
{norm=Norm;term=nt} in
(* Precondition: m.norm = Norm *)
@@ -1081,18 +1049,12 @@ let norm_val info v =
let inject = mk_clos (ESID 0)
(* cache of constants: the body is computed only when needed. *)
-type 'a clos_infos = (fconstr, 'a) infos
-
-let create_clos_infos flgs env sigma =
- create (fun _ -> inject) flgs env sigma
-
-let unfold_reference info = function
- | FConst op -> ref_value_cache info (ConstBinding op)
- | FEvar (n,v) ->
- let evar = (n, Array.map (norm_val info) v) in
- ref_value_cache info (EvarBinding evar)
- | FVar id -> ref_value_cache info (VarBinding id)
- | FFarRel p -> ref_value_cache info (FarRelBinding p)
+type clos_infos = fconstr infos
+
+let create_clos_infos flgs env =
+ create (fun _ -> inject) flgs env
+
+let unfold_reference = ref_value_cache
(* Head normal form. *)
diff --git a/kernel/closure.mli b/kernel/closure.mli
index 16de949af..4abd866c3 100644
--- a/kernel/closure.mli
+++ b/kernel/closure.mli
@@ -12,7 +12,6 @@
open Pp
open Names
open Term
-open Evd
open Environ
open Esubst
(*i*)
@@ -48,7 +47,6 @@ module type RedFlagsSig = sig
of Constbut/Varbut should be unfolded (there may be several such
Constbut/Varbut *)
val fBETA : red_kind
- val fEVAR : red_kind
val fDELTA : red_kind
val fIOTA : red_kind
val fZETA : red_kind
@@ -84,6 +82,7 @@ val beta_red : reds
val betaiota_red : reds
val betadeltaiota_red : reds
val betaiotazeta_red : reds
+val betadeltaiotanolet_red : reds
(*s Reduction function specification. *)
@@ -105,25 +104,24 @@ val no_flag : flags
val beta : flags
val betaiota : flags
val betadeltaiota : flags
+val betadeltaiotanolet : flags
val hnf_flags : flags
val unfold_flags : evaluable_global_reference -> flags
(***********************************************************************)
-type 'a table_key =
- | ConstBinding of constant
- | EvarBinding of existential
- | VarBinding of identifier
- | FarRelBinding of int
+type table_key =
+ | ConstKey of constant
+ | VarKey of identifier
+ | FarRelKey of int
+ (* FarRel: index in the rel_context part of _initial_ environment *)
-type ('a,'b) infos
-val ref_value_cache: ('a,'b) infos -> 'a table_key -> 'a option
-val info_flags: ('a,'b) infos -> flags
-val infos_under: ('a,'b) infos -> ('a,'b) infos
-val create:
- (('a,'b) infos -> constr -> 'a) ->
- flags -> env -> 'b evar_map -> ('a,'b) infos
+type 'a infos
+val ref_value_cache: 'a infos -> table_key -> 'a option
+val info_flags: 'a infos -> flags
+val infos_under: 'a infos -> 'a infos
+val create: ('a infos -> constr -> 'a) -> flags -> env -> 'a infos
(***********************************************************************)
(*s A [stack] is a context of arguments, arguments are pushed by
@@ -165,7 +163,7 @@ type fterm =
| FRel of int
| FAtom of constr
| FCast of fconstr * fconstr
- | FFlex of freference
+ | FFlex of table_key
| FInd of inductive
| FConstruct of constructor
| FApp of fconstr * fconstr array
@@ -177,16 +175,11 @@ type fterm =
| FLambda of name * fconstr * fconstr * constr * fconstr subs
| FProd of name * fconstr * fconstr * constr * fconstr subs
| FLetIn of name * fconstr * fconstr * fconstr * constr * fconstr subs
+ | FEvar of existential_key * fconstr array
| FLIFT of int * fconstr
| FCLOS of constr * fconstr subs
| FLOCKED
-and freference =
- | FConst of constant
- | FEvar of existential_key * fconstr array
- | FVar of identifier
- | FFarRel of int
-
(* To lazy reduce a constr, create a ['a clos_infos] with
[create_cbv_infos], inject the term to reduce with [inject]; then use
@@ -197,28 +190,28 @@ val fterm_of : fconstr -> fterm
val term_of_fconstr : fconstr -> constr
(* Global and local constant cache *)
-type 'a clos_infos
-val create_clos_infos : flags -> env -> 'a evar_map -> 'a clos_infos
+type clos_infos
+val create_clos_infos : flags -> env -> clos_infos
(* Reduction function *)
(* [norm_val] is for strong normalization *)
-val norm_val : 'a clos_infos -> fconstr -> constr
+val norm_val : clos_infos -> fconstr -> constr
(* [whd_val] is for weak head normalization *)
-val whd_val : 'a clos_infos -> fconstr -> constr
+val whd_val : clos_infos -> fconstr -> constr
(* Conversion auxiliary functions to do step by step normalisation *)
(* [fhnf] and [fnf_apply] are for weak head normalization but staying
in [fconstr] world to perform step by step weak head normalization *)
-val fhnf: 'a clos_infos -> fconstr -> int * fconstr * fconstr stack
-val fhnf_apply : 'a clos_infos ->
+val fhnf: clos_infos -> fconstr -> int * fconstr * fconstr stack
+val fhnf_apply : clos_infos ->
int -> fconstr -> fconstr stack -> int * fconstr * fconstr stack
(* [unfold_reference] unfolds references in a [fconstr] *)
-val unfold_reference : 'a clos_infos -> freference -> fconstr option
+val unfold_reference : clos_infos -> table_key -> fconstr option
(***********************************************************************)
(*i This is for lazy debug *)
@@ -232,9 +225,9 @@ val mk_clos_deep :
(fconstr subs -> constr -> fconstr) ->
fconstr subs -> constr -> fconstr
-val knr: 'a clos_infos -> fconstr -> fconstr stack ->
+val knr: clos_infos -> fconstr -> fconstr stack ->
fconstr * fconstr stack
-val kl: 'a clos_infos -> fconstr -> fconstr
+val kl: clos_infos -> fconstr -> fconstr
val to_constr :
(lift -> fconstr -> constr) -> lift -> fconstr -> constr
diff --git a/kernel/cooking.ml b/kernel/cooking.ml
index 79420e040..4fb1663d0 100644
--- a/kernel/cooking.ml
+++ b/kernel/cooking.ml
@@ -14,7 +14,6 @@ open Names
open Term
open Sign
open Declarations
-open Instantiate
open Environ
open Reduction
@@ -47,61 +46,59 @@ let failure () =
let modify_opers replfun (constl,indl,cstrl) =
let rec substrec c =
- let op, cl = splay_constr c in
- let cl' = Array.map substrec cl in
- match op with
- | OpMutCase (n,(spi,a,b,c,d) as oper) ->
+ let c' = map_constr substrec c in
+ match kind_of_term c' with
+ | Case (ci,p,t,br) ->
(try
- match List.assoc spi indl with
- | DO_ABSTRACT (spi',abs_vars) ->
- let n' = Array.length abs_vars + n in
- gather_constr (OpMutCase (n',(spi',a,b,c,d)),cl')
+ match List.assoc ci.ci_ind indl with
+ | DO_ABSTRACT (ind,abs_vars) ->
+ let n' = Array.length abs_vars + ci.ci_npar in
+ let ci' = { ci with
+ ci_ind = ind;
+ ci_npar = n' } in
+ mkCase (ci',p,t,br)
| _ -> raise Not_found
with
- | Not_found -> gather_constr (op,cl'))
+ | Not_found -> c')
- | OpMutInd spi ->
- assert (Array.length cl=0);
+ | Ind spi ->
(try
(match List.assoc spi indl with
| NOT_OCCUR -> failure ()
| DO_ABSTRACT (oper',abs_vars) ->
- mkApp (mkMutInd oper', abs_vars)
+ mkApp (mkInd oper', abs_vars)
| DO_REPLACE _ -> assert false)
with
- | Not_found -> mkMutInd spi)
+ | Not_found -> c')
- | OpMutConstruct spi ->
- assert (Array.length cl=0);
+ | Construct spi ->
(try
(match List.assoc spi cstrl with
| NOT_OCCUR -> failure ()
| DO_ABSTRACT (oper',abs_vars) ->
- mkApp (mkMutConstruct oper', abs_vars)
+ mkApp (mkConstruct oper', abs_vars)
| DO_REPLACE _ -> assert false)
with
- | Not_found -> mkMutConstruct spi)
+ | Not_found -> c')
- | OpConst sp ->
- assert (Array.length cl=0);
+ | Const sp ->
(try
(match List.assoc sp constl with
| NOT_OCCUR -> failure ()
| DO_ABSTRACT (oper',abs_vars) ->
mkApp (mkConst oper', abs_vars)
- | DO_REPLACE cb -> substrec (replfun sp cb cl'))
+ | DO_REPLACE cb -> substrec (replfun (sp,cb)))
with
- | Not_found -> mkConst sp)
+ | Not_found -> c')
- | _ -> gather_constr (op, cl')
+ | _ -> c'
in
if (constl,indl,cstrl) = ([],[],[]) then fun x -> x else substrec
let expmod_constr modlist c =
- let sigma = Evd.empty in
let simpfun =
if modlist = ([],[],[]) then fun x -> x else nf_betaiota in
- let expfun sp cb args =
+ let expfun (sp,cb) =
if cb.const_opaque then
errorlabstrm "expmod_constr"
[< 'sTR"Cannot unfold the value of ";
@@ -110,14 +107,12 @@ let expmod_constr modlist c =
'sTR"and then require that theorems which use them"; 'sPC;
'sTR"be transparent" >];
match cb.const_body with
- | Some body ->
- instantiate_constr cb.const_hyps body (Array.to_list args)
+ | Some body -> body
| None -> assert false
in
- let c' =
- modify_opers expfun modlist c in
+ let c' = modify_opers expfun modlist c in
match kind_of_term c' with
- | IsCast (val_0,typ) -> mkCast (simpfun val_0,simpfun typ)
+ | Cast (value,typ) -> mkCast (simpfun value,simpfun typ)
| _ -> simpfun c'
let expmod_type modlist c =
@@ -141,7 +136,13 @@ let cook_constant env r =
let cb = r.d_from in
let typ = expmod_type r.d_modlist cb.const_type in
let body = option_app (expmod_constr r.d_modlist) cb.const_body in
- let hyps = List.map (fun (sp,c,t) -> (basename sp,c,t)) cb.const_hyps in
- let hyps = map_named_context (expmod_constr r.d_modlist) hyps in
+ let hyps =
+ Sign.fold_named_context
+ (fun d ctxt ->
+ Sign.add_named_decl
+ (map_named_declaration (expmod_constr r.d_modlist) d)
+ ctxt)
+ cb.const_hyps
+ empty_named_context in
let body,typ = abstract_constant r.d_abstract hyps (body,typ) in
(body, typ, cb.const_constraints, cb.const_opaque)
diff --git a/kernel/declarations.ml b/kernel/declarations.ml
index 47f65d4a3..a9b8737bb 100644
--- a/kernel/declarations.ml
+++ b/kernel/declarations.ml
@@ -6,99 +6,61 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id$ *)
+(*i $Id$ i*)
+(*i*)
open Names
open Univ
open Term
open Sign
+(*i*)
-(* Constant entries *)
+(* This module defines the types of global declarations. This includes
+ global constants/axioms and mutual inductive definitions *)
+
+(*s Constants (internal representation) (Definition/Axiom) *)
type constant_body = {
- const_kind : path_kind;
+ const_hyps : section_context; (* New: younger hyp at top *)
const_body : constr option;
const_type : types;
- const_hyps : section_context;
const_constraints : constraints;
const_opaque : bool }
-let is_defined cb =
- match cb.const_body with Some _ -> true | _ -> false
-
-let is_opaque cb = cb.const_opaque
-
-(*s Global and local constant declaration. *)
-
-type constant_entry = {
- const_entry_body : constr;
- const_entry_type : constr option;
- const_entry_opaque : bool }
-
-type local_entry =
- | LocalDef of constr
- | LocalAssum of constr
-
-(* Inductive entries *)
+(*s Inductive types (internal representation with redundant
+ information). *)
type recarg =
| Param of int
| Norec
| Mrec of int
- | Imbr of inductive * recarg list
+ | Imbr of inductive * (recarg list)
+
+(* [mind_typename] is the name of the inductive; [mind_arity] is
+ the arity generalized over global parameters; [mind_lc] is the list
+ of types of constructors generalized over global parameters and
+ relative to the global context enriched with the arities of the
+ inductives *)
type one_inductive_body = {
- mind_consnames : identifier array;
mind_typename : identifier;
- mind_nf_lc : types array;
+ mind_nparams : int;
+ mind_params_ctxt : rel_context;
+ mind_nrealargs : int;
mind_nf_arity : types;
- (* lc and arity as given by user if not in nf; useful e.g. for Ensemble.v *)
- mind_user_lc : types array option;
- mind_user_arity : types option;
+ mind_user_arity : types;
mind_sort : sorts;
- mind_nrealargs : int;
mind_kelim : sorts_family list;
+ mind_consnames : identifier array;
+ mind_nf_lc : types array; (* constrs and arity with pre-expanded ccl *)
+ mind_user_lc : types array;
mind_listrec : (recarg list) array;
- mind_finite : bool;
- mind_nparams : int;
- mind_params_ctxt : rel_context }
+ }
type mutual_inductive_body = {
- mind_kind : path_kind;
+ mind_finite : bool;
mind_ntypes : int;
mind_hyps : section_context;
mind_packets : one_inductive_body array;
mind_constraints : constraints;
mind_singl : constr option }
-
-let mind_type_finite mib i = mib.mind_packets.(i).mind_finite
-
-let mind_user_lc mip = match mip.mind_user_lc with
- | None -> mip.mind_nf_lc
- | Some lc -> lc
-
-let mind_user_arity mip = match mip.mind_user_arity with
- | None -> mip.mind_nf_arity
- | Some a -> a
-
-(*s Declaration. *)
-
-type one_inductive_entry = {
- mind_entry_nparams : int;
- mind_entry_params : (identifier * local_entry) list;
- mind_entry_typename : identifier;
- mind_entry_arity : constr;
- mind_entry_consnames : identifier list;
- mind_entry_lc : constr list }
-
-type mutual_inductive_entry = {
- mind_entry_finite : bool;
- mind_entry_inds : one_inductive_entry list }
-
-let mind_nth_type_packet mib n = mib.mind_packets.(n)
-
-let mind_arities_context mib =
- Array.to_list
- (Array.map (* No need to lift, arities contain no de Bruijn *)
- (fun mip -> (Name mip.mind_typename, None, mind_user_arity mip))
- mib.mind_packets)
diff --git a/kernel/declarations.mli b/kernel/declarations.mli
index 735f6f141..a9b8737bb 100644
--- a/kernel/declarations.mli
+++ b/kernel/declarations.mli
@@ -21,27 +21,14 @@ open Sign
(*s Constants (internal representation) (Definition/Axiom) *)
type constant_body = {
- const_kind : path_kind;
+ const_hyps : section_context; (* New: younger hyp at top *)
const_body : constr option;
const_type : types;
- const_hyps : section_context; (* New: younger hyp at top *)
const_constraints : constraints;
const_opaque : bool }
-val is_defined : constant_body -> bool
-
-(*s Global and local constant declaration. *)
-
-type constant_entry = {
- const_entry_body : constr;
- const_entry_type : constr option;
- const_entry_opaque : bool }
-
-type local_entry =
- | LocalDef of constr
- | LocalAssum of constr
-
-(*s Inductive types (internal representation). *)
+(*s Inductive types (internal representation with redundant
+ information). *)
type recarg =
| Param of int
@@ -56,56 +43,24 @@ type recarg =
inductives *)
type one_inductive_body = {
- mind_consnames : identifier array;
mind_typename : identifier;
- mind_nf_lc : types array; (* constrs and arity with pre-expanded ccl *)
+ mind_nparams : int;
+ mind_params_ctxt : rel_context;
+ mind_nrealargs : int;
mind_nf_arity : types;
- mind_user_lc : types array option;
- mind_user_arity : types option;
+ mind_user_arity : types;
mind_sort : sorts;
- mind_nrealargs : int;
mind_kelim : sorts_family list;
+ mind_consnames : identifier array;
+ mind_nf_lc : types array; (* constrs and arity with pre-expanded ccl *)
+ mind_user_lc : types array;
mind_listrec : (recarg list) array;
- mind_finite : bool;
- mind_nparams : int;
- mind_params_ctxt : rel_context }
+ }
type mutual_inductive_body = {
- mind_kind : path_kind;
+ mind_finite : bool;
mind_ntypes : int;
mind_hyps : section_context;
mind_packets : one_inductive_body array;
mind_constraints : constraints;
mind_singl : constr option }
-
-val mind_type_finite : mutual_inductive_body -> int -> bool
-val mind_user_lc : one_inductive_body -> types array
-val mind_user_arity : one_inductive_body -> types
-val mind_nth_type_packet : mutual_inductive_body -> int -> one_inductive_body
-
-val mind_arities_context : mutual_inductive_body -> rel_declaration list
-
-(*s Declaration of inductive types. *)
-
-(* Assume the following definition in concrete syntax:
-\begin{verbatim}
-Inductive I1 [x1:X1;...;xn:Xn] : A1 := c11 : T11 | ... | c1n1 : T1n1
-...
-with Ip [x1:X1;...;xn:Xn] : Ap := cp1 : Tp1 | ... | cpnp : Tpnp.
-\end{verbatim}
-then, in $i^{th}$ block, [mind_entry_params] is [[xn:Xn;...;x1:X1]];
-[mind_entry_arity] is [Ai], defined in context [[[x1:X1;...;xn:Xn]];
-[mind_entry_lc] is [Ti1;...;Tini], defined in context [[A'1;...;A'p;x1:X1;...;xn:Xn]] where [A'i] is [Ai] generalized over [[x1:X1;...;xn:Xn]].
-*)
-
-type one_inductive_entry = {
- mind_entry_nparams : int;
- mind_entry_params : (identifier * local_entry) list;
- mind_entry_typename : identifier;
- mind_entry_arity : constr;
- mind_entry_consnames : identifier list;
- mind_entry_lc : constr list }
-
-type mutual_inductive_entry = {
- mind_entry_finite : bool;
- mind_entry_inds : one_inductive_entry list }
diff --git a/kernel/environ.ml b/kernel/environ.ml
index 98f54337f..757fa34b0 100644
--- a/kernel/environ.ml
+++ b/kernel/environ.ml
@@ -8,7 +8,6 @@
(* $Id$ *)
-open Pp
open Util
open Names
open Sign
@@ -30,64 +29,55 @@ type globals = {
env_locals : (global * section_path) list;
env_imports : compilation_unit_name list }
-type context = {
- env_named_context : named_context;
- env_rel_context : rel_context }
-
type env = {
- env_context : context;
- env_globals : globals;
- env_universes : universes }
-
-let empty_context = {
- env_named_context = empty_named_context;
- env_rel_context = empty_rel_context }
+ env_globals : globals;
+ env_named_context : named_context;
+ env_rel_context : rel_context;
+ env_universes : universes }
let empty_env = {
- env_context = empty_context;
env_globals = {
env_constants = Spmap.empty;
env_inductives = Spmap.empty;
env_locals = [];
env_imports = [] };
+ env_named_context = empty_named_context;
+ env_rel_context = empty_rel_context;
env_universes = initial_universes }
let universes env = env.env_universes
-let context env = env.env_context
-let named_context env = env.env_context.env_named_context
-let rel_context env = env.env_context.env_rel_context
+let named_context env = env.env_named_context
+let rel_context env = env.env_rel_context
(* Construction functions. *)
-let map_context f env =
- let context = env.env_context in
- { env with
- env_context = {
- context with
- env_named_context = map_named_context f context.env_named_context ;
- env_rel_context = map_rel_context f context.env_rel_context } }
-
let named_context_app f env =
{ env with
- env_context = { env.env_context with
- env_named_context = f env.env_context.env_named_context } }
-
-let change_hyps = named_context_app
+ env_named_context = f env.env_named_context }
let push_named_decl d = named_context_app (add_named_decl d)
-let push_named_def def = named_context_app (add_named_def def)
-let push_named_assum decl = named_context_app (add_named_assum decl)
+let push_named_assum (id,ty) = push_named_decl (id,None,ty)
let pop_named_decl id = named_context_app (pop_named_decl id)
let rel_context_app f env =
{ env with
- env_context = { env.env_context with
- env_rel_context = f env.env_context.env_rel_context } }
+ env_rel_context = f env.env_rel_context }
let reset_context env =
{ env with
- env_context = { env_named_context = empty_named_context;
- env_rel_context = empty_rel_context} }
+ env_named_context = empty_named_context;
+ env_rel_context = empty_rel_context }
+
+let reset_with_named_context ctxt env =
+ { env with
+ env_named_context = ctxt;
+ env_rel_context = empty_rel_context }
+
+let reset_rel_context env =
+ { env with
+ env_rel_context = empty_rel_context }
+
+
let fold_named_context f env a =
snd (Sign.fold_named_context
@@ -97,33 +87,9 @@ let fold_named_context f env a =
let fold_named_context_reverse f a env =
Sign.fold_named_context_reverse f a (named_context env)
-let process_named_context f env =
- Sign.fold_named_context
- (fun d env -> f env d) (named_context env) (reset_context env)
-
-let process_named_context_both_sides f env =
- fold_named_context_both_sides f (named_context env) (reset_context env)
-
let push_rel d = rel_context_app (add_rel_decl d)
-let push_rel_def def = rel_context_app (add_rel_def def)
-let push_rel_assum decl = rel_context_app (add_rel_assum decl)
-let push_rels ctxt = rel_context_app (concat_rel_context ctxt)
-let push_rels_assum decl env =
- rel_context_app (List.fold_right add_rel_assum decl) env
-
-
-let push_rel_context_to_named_context env =
- let sign0 = env.env_context.env_named_context in
- let (subst,_,sign) =
- List.fold_right
- (fun (na,c,t) (subst,avoid,sign) ->
- let na = if na = Anonymous then Name(id_of_string"_") else na in
- let id = next_name_away na avoid in
- ((mkVar id)::subst,id::avoid,
- add_named_decl (id,option_app (substl subst) c,type_app (substl subst) t)
- sign))
- env.env_context.env_rel_context ([],ids_of_named_context sign0,sign0)
- in subst, (named_context_app (fun _ -> sign) env)
+let push_rel_context ctxt = fold_rel_context push_rel ctxt
+let push_rel_assum (id,ty) = push_rel (id,None,ty)
let push_rec_types (lna,typarray,_) env =
let ctxt =
@@ -131,40 +97,11 @@ let push_rec_types (lna,typarray,_) env =
Array.fold_left
(fun e assum -> push_rel_assum assum e) env ctxt
-let push_named_rec_types (lna,typarray,_) env =
- let ctxt =
- array_map2_i
- (fun i na t ->
- match na with
- | Name id -> (id, type_app (lift i) t)
- | Anonymous -> anomaly "Fix declarations must be named")
- lna typarray in
- Array.fold_left
- (fun e assum -> push_named_assum assum e) env ctxt
-
-let reset_rel_context env =
- { env with
- env_context = { env_named_context = env.env_context.env_named_context;
- env_rel_context = empty_rel_context} }
-
let fold_rel_context f env a =
snd (List.fold_right
(fun d (env,e) -> (push_rel d env, f env d e))
(rel_context env) (reset_rel_context env,a))
-let process_rel_context f env =
- List.fold_right (fun d env -> f env d)
- (rel_context env) (reset_rel_context env)
-
-let instantiate_named_context = instantiate_sign
-
-let ids_of_context env =
- (ids_of_rel_context env.env_context.env_rel_context)
- @ (ids_of_named_context env.env_context.env_named_context)
-
-let names_of_rel_context env =
- names_of_rel_context env.env_context.env_rel_context
-
let set_universes g env =
if env.env_universes == g then env else { env with env_universes = g }
@@ -193,20 +130,12 @@ let add_mind sp mib env =
{ env with env_globals = new_globals }
(* Access functions. *)
-
-let lookup_named_type id env =
- lookup_id_type id env.env_context.env_named_context
-
-let lookup_named_value id env =
- lookup_id_value id env.env_context.env_named_context
-let lookup_named id env = lookup_id id env.env_context.env_named_context
+let lookup_rel n env =
+ Sign.lookup_rel n env.env_rel_context
-let lookup_rel_type n env =
- Sign.lookup_rel_type n env.env_context.env_rel_context
-
-let lookup_rel_value n env =
- Sign.lookup_rel_value n env.env_context.env_rel_context
+let lookup_named id env =
+ Sign.lookup_named id env.env_named_context
let lookup_constant sp env =
Spmap.find sp env.env_globals.env_constants
@@ -214,15 +143,14 @@ let lookup_constant sp env =
let lookup_mind sp env =
Spmap.find sp env.env_globals.env_inductives
-
(* Lookup of section variables *)
let lookup_constant_variables c env =
let cmap = lookup_constant c env in
- Sign.instance_from_section_context cmap.const_hyps
+ Sign.instance_from_named_context cmap.const_hyps
let lookup_inductive_variables (sp,i) env =
let mis = lookup_mind sp env in
- Sign.instance_from_section_context mis.mind_hyps
+ Sign.instance_from_named_context mis.mind_hyps
let lookup_constructor_variables (ind,_) env =
lookup_inductive_variables ind env
@@ -231,28 +159,18 @@ let lookup_constructor_variables (ind,_) env =
let vars_of_global env constr =
match kind_of_term constr with
- IsVar id -> [id]
- | IsConst sp ->
+ Var id -> [id]
+ | Const sp ->
List.map destVar
(Array.to_list (lookup_constant_variables sp env))
- | IsMutInd ind ->
+ | Ind ind ->
List.map destVar
(Array.to_list (lookup_inductive_variables ind env))
- | IsMutConstruct cstr ->
+ | Construct cstr ->
List.map destVar
(Array.to_list (lookup_constructor_variables cstr env))
| _ -> []
-let rec global_varsl env l constr =
- let l = vars_of_global env constr @ l in
- fold_constr (global_varsl env) l constr
-
-let global_vars env = global_varsl env []
-
-let global_vars_decl env = function
- | (_, None, t) -> global_vars env t
- | (_, Some c, t) -> (global_vars env c)@(global_vars env t)
-
let global_vars_set env constr =
let rec filtrec acc c =
let vl = vars_of_global env c in
@@ -261,32 +179,12 @@ let global_vars_set env constr =
in
filtrec Idset.empty constr
-
-exception Occur
-
-let occur_in_global env id constr =
- let vars = vars_of_global env constr in
- if List.mem id vars then raise Occur
-
-let occur_var env s c =
- let rec occur_rec c =
- occur_in_global env s c;
- iter_constr occur_rec c
- in
- try occur_rec c; false with Occur -> true
-
-let occur_var_in_decl env hyp (_,c,typ) =
- match c with
- | None -> occur_var env hyp (body_of_type typ)
- | Some body ->
- occur_var env hyp (body_of_type typ) ||
- occur_var env hyp body
-
-(* [keep_hyps sign ids] keeps the part of the signature [sign] which
+(* [keep_hyps env ids] keeps the part of the section context of [env] which
contains the variables of the set [ids], and recursively the variables
contained in the types of the needed variables. *)
-let rec keep_hyps env needed = function
+let keep_hyps env needed =
+ let rec keep_rec needed = function
| (id,copt,t as d) ::sign when Idset.mem id needed ->
let globc =
match copt with
@@ -295,170 +193,63 @@ let rec keep_hyps env needed = function
let needed' =
Idset.union (global_vars_set env (body_of_type t))
(Idset.union globc needed) in
- d :: (keep_hyps env needed' sign)
- | _::sign -> keep_hyps env needed sign
- | [] -> []
-
-(* This renames bound variables with fresh and distinct names *)
-(* in such a way that the printer doe not generate new names *)
-(* and therefore that printed names are the intern names *)
-(* In this way, tactics such as Induction works well *)
-
-let rec rename_bound_var env l c =
- match kind_of_term c with
- | IsProd (Name s,c1,c2) ->
- if dependent (mkRel 1) c2 then
- let s' = next_ident_away s (global_vars env c2@l) in
- let env' = push_rel (Name s',None,c1) env in
- mkProd (Name s', c1, rename_bound_var env' (s'::l) c2)
- else
- let env' = push_rel (Name s,None,c1) env in
- mkProd (Name s, c1, rename_bound_var env' l c2)
- | IsProd (Anonymous,c1,c2) ->
- let env' = push_rel (Anonymous,None,c1) env in
- mkProd (Anonymous, c1, rename_bound_var env' l c2)
- | IsCast (c,t) -> mkCast (rename_bound_var env l c, t)
- | x -> c
-
-(* First character of a constr *)
-
-let lowercase_first_char id = String.lowercase (first_char id)
-
-(* id_of_global gives the name of the given sort oper *)
-let sp_of_global env = function
- | VarRef sp -> sp
- | ConstRef sp -> sp
- | IndRef (sp,tyi) ->
- (* Does not work with extracted inductive types when the first
- inductive is logic : if tyi=0 then basename sp else *)
- let mib = lookup_mind sp env in
- let mip = mind_nth_type_packet mib tyi in
- make_path (dirpath sp) mip.mind_typename CCI
- | ConstructRef ((sp,tyi),i) ->
- let mib = lookup_mind sp env in
- let mip = mind_nth_type_packet mib tyi in
- assert (i <= Array.length mip.mind_consnames && i > 0);
- make_path (dirpath sp) mip.mind_consnames.(i-1) CCI
-
-let id_of_global env ref = basename (sp_of_global env ref)
-
-let hdchar env c =
- let rec hdrec k c =
- match kind_of_term c with
- | IsProd (_,_,c) -> hdrec (k+1) c
- | IsLambda (_,_,c) -> hdrec (k+1) c
- | IsLetIn (_,_,_,c) -> hdrec (k+1) c
- | IsCast (c,_) -> hdrec k c
- | IsApp (f,l) -> hdrec k f
- | IsConst sp ->
- let c = lowercase_first_char (basename sp) in
- if c = "?" then "y" else c
- | IsMutInd ((sp,i) as x) ->
- if i=0 then
- lowercase_first_char (basename sp)
- else
- lowercase_first_char (id_of_global env (IndRef x))
- | IsMutConstruct ((sp,i) as x) ->
- lowercase_first_char (id_of_global env (ConstructRef x))
- | IsVar id -> lowercase_first_char id
- | IsSort s -> sort_hdchar s
- | IsRel n ->
- (if n<=k then "p" (* the initial term is flexible product/function *)
- else
- try match lookup_rel_type (n-k) env with
- | Name id,_ -> lowercase_first_char id
- | Anonymous,t -> hdrec 0 (lift (n-k) (body_of_type t))
- with Not_found -> "y")
- | IsFix ((_,i),(lna,_,_)) ->
- let id = match lna.(i) with Name id -> id | _ -> assert false in
- lowercase_first_char id
- | IsCoFix (i,(lna,_,_)) ->
- let id = match lna.(i) with Name id -> id | _ -> assert false in
- lowercase_first_char id
- | IsMeta _|IsEvar _|IsMutCase (_, _, _, _) -> "y"
- in
- hdrec 0 c
-
-let id_of_name_using_hdchar env a = function
- | Anonymous -> id_of_string (hdchar env a)
- | Name id -> id
-
-let named_hd env a = function
- | Anonymous -> Name (id_of_string (hdchar env a))
- | x -> x
-
-let named_hd_type env a = named_hd env (body_of_type a)
-
-let prod_name env (n,a,b) = mkProd (named_hd_type env a n, a, b)
-let lambda_name env (n,a,b) = mkLambda (named_hd_type env a n, a, b)
-
-let prod_create env (a,b) = mkProd (named_hd_type env a Anonymous, a, b)
-let lambda_create env (a,b) = mkLambda (named_hd_type env a Anonymous, a, b)
-
-let name_assumption env (na,c,t) =
- match c with
- | None -> (named_hd_type env t na, None, t)
- | Some body -> (named_hd env body na, c, t)
-
-let mkProd_or_LetIn_name env b d = mkProd_or_LetIn (name_assumption env d) b
-let mkLambda_or_LetIn_name env b d = mkLambda_or_LetIn (name_assumption env d)b
-
-let name_context env hyps =
- snd
- (List.fold_left
- (fun (env,hyps) d ->
- let d' = name_assumption env d in (push_rel d' env, d' :: hyps))
- (env,[]) (List.rev hyps))
-
-let it_mkProd_wo_LetIn = List.fold_left (fun c d -> mkProd_wo_LetIn d c)
-let it_mkProd_or_LetIn = List.fold_left (fun c d -> mkProd_or_LetIn d c)
-let it_mkLambda_or_LetIn = List.fold_left (fun c d -> mkLambda_or_LetIn d c)
+ d :: (keep_rec needed' sign)
+ | _::sign -> keep_rec needed sign
+ | [] -> [] in
+ keep_rec needed (named_context env)
-let it_mkProd_or_LetIn_name env b hyps =
- it_mkProd_or_LetIn b (name_context env hyps)
-
-let it_mkLambda_or_LetIn_name env b hyps =
- it_mkLambda_or_LetIn b (name_context env hyps)
-
-let it_mkNamedProd_or_LetIn = it_named_context_quantifier mkNamedProd_or_LetIn
-let it_mkNamedLambda_or_LetIn = it_named_context_quantifier mkNamedLambda_or_LetIn
-
-let it_mkNamedProd_wo_LetIn = it_named_context_quantifier mkNamedProd_wo_LetIn
-
-let make_all_name_different env =
- let avoid = ref (ids_of_named_context (named_context env)) in
- process_rel_context
- (fun newenv (na,c,t) ->
- let id = next_name_away na !avoid in
- avoid := id::!avoid;
- push_rel (Name id,c,t) newenv)
- env
(* Constants *)
-let defined_constant env sp = is_defined (lookup_constant sp env)
-
+let defined_constant env sp =
+ match (lookup_constant sp env).const_body with
+ Some _ -> true
+ | None -> false
+
let opaque_constant env sp = (lookup_constant sp env).const_opaque
(* A global const is evaluable if it is defined and not opaque *)
let evaluable_constant env sp =
- try
- defined_constant env sp && not (opaque_constant env sp)
+ try defined_constant env sp
with Not_found ->
false
(* A local const is evaluable if it is defined and not opaque *)
let evaluable_named_decl env id =
try
- lookup_named_value id env <> None
+ match lookup_named id env with
+ (_,Some _,_) -> true
+ | _ -> false
with Not_found ->
false
let evaluable_rel_decl env n =
- try
- lookup_rel_value n env <> None
+ try
+ match lookup_rel n env with
+ (_,Some _,_) -> true
+ | _ -> false
with Not_found ->
false
+(* constant_type gives the type of a constant *)
+let constant_type env sp =
+ let cb = lookup_constant sp env in
+ cb.const_type
+
+type const_evaluation_result = NoBody | Opaque
+
+exception NotEvaluableConst of const_evaluation_result
+
+let constant_value env sp =
+ let cb = lookup_constant sp env in
+ if cb.const_opaque then raise (NotEvaluableConst Opaque);
+ match cb.const_body with
+ | Some body -> body
+ | None -> raise (NotEvaluableConst NoBody)
+
+let constant_opt_value env cst =
+ try Some (constant_value env cst)
+ with NotEvaluableConst _ -> None
+
(*s Modules (i.e. compiled environments). *)
type compiled_env = {
@@ -498,8 +289,7 @@ let import_constraints g sp cst =
try
merge_constraints cst g
with UniverseInconsistency ->
- errorlabstrm "import_constraints"
- [< 'sTR "Universe Inconsistency during import of"; 'sPC; pr_sp sp >]
+ error "import_constraints: Universe Inconsistency during import"
let import cenv env =
check_imports env cenv.cenv_needed;
@@ -526,16 +316,14 @@ type unsafe_judgment = {
uj_val : constr;
uj_type : types }
+let make_judge v tj =
+ { uj_val = v;
+ uj_type = tj }
+
+let j_val j = j.uj_val
+let j_type j = body_of_type j.uj_type
+
type unsafe_type_judgment = {
utj_val : constr;
utj_type : sorts }
-(*s Memory use of an environment. *)
-
-open Printf
-
-let mem env =
- let glb = env.env_globals in
- h 0 [< 'sTR (sprintf "%dk (cst = %dk / ind = %dk / unv = %dk)"
- (size_kb env) (size_kb glb.env_constants)
- (size_kb glb.env_inductives) (size_kb env.env_universes)) >]
diff --git a/kernel/environ.mli b/kernel/environ.mli
index 761f196c0..83915157a 100644
--- a/kernel/environ.mli
+++ b/kernel/environ.mli
@@ -21,173 +21,92 @@ open Sign
informations added in environments, and that is why we speak here
of ``unsafe'' environments. *)
-type context
type env
-val empty_context : context
val empty_env : env
val universes : env -> universes
-val context : env -> context
val rel_context : env -> rel_context
val named_context : env -> named_context
(* This forgets named and rel contexts *)
val reset_context : env -> env
+(* This forgets rel context and sets a new named context *)
+val reset_with_named_context : named_context -> env -> env
(*s This concerns only local vars referred by names [named_context] *)
val push_named_decl : named_declaration -> env -> env
-val push_named_assum : identifier * types -> env -> env
-val push_named_def : identifier * constr * types -> env -> env
-val change_hyps : (named_context -> named_context) -> env -> env
-val instantiate_named_context : named_context -> constr list -> (identifier * constr) list
val pop_named_decl : identifier -> env -> env
(*s This concerns only local vars referred by indice [rel_context] *)
val push_rel : rel_declaration -> env -> env
-val push_rel_assum : name * types -> env -> env
-val push_rel_def : name * constr * types -> env -> env
-val push_rels : rel_context -> env -> env
-val push_rels_assum : (name * types) list -> env -> env
-val names_of_rel_context : env -> names_context
-
-(*s Returns also the substitution to be applied to rel's *)
-val push_rel_context_to_named_context : env -> constr list * env
+val push_rel_context : rel_context -> env -> env
(*s Push the types of a (co-)fixpoint to [rel_context] *)
val push_rec_types : rec_declaration -> env -> env
-(*s Push the types of a (co-)fixpoint to [named_context] *)
-val push_named_rec_types : rec_declaration -> env -> env
-
-(* Gives identifiers in [named_context] and [rel_context] *)
-val ids_of_context : env -> identifier list
-val map_context : (constr -> constr) -> env -> env
-
-(*s Recurrence on [named_context] *)
-val fold_named_context : (env -> named_declaration -> 'a -> 'a) -> env -> 'a -> 'a
-val process_named_context : (env -> named_declaration -> env) -> env -> env
+(*s Recurrence on [named_context]: older declarations processed first *)
+val fold_named_context :
+ (env -> named_declaration -> 'a -> 'a) -> env -> 'a -> 'a
(* Recurrence on [named_context] starting from younger decl *)
-val fold_named_context_reverse : ('a -> named_declaration -> 'a) -> 'a -> env -> 'a
-
-(* [process_named_context_both_sides f env] iters [f] on the named
- declarations of [env] taking as argument both the initial segment, the
- middle value and the tail segment *)
-val process_named_context_both_sides :
- (env -> named_declaration -> named_context -> env) -> env -> env
+val fold_named_context_reverse :
+ ('a -> named_declaration -> 'a) -> 'a -> env -> 'a
(*s Recurrence on [rel_context] *)
val fold_rel_context : (env -> rel_declaration -> 'a -> 'a) -> env -> 'a -> 'a
-val process_rel_context : (env -> rel_declaration -> env) -> env -> env
(*s add entries to environment *)
val set_universes : universes -> env -> env
val add_constraints : constraints -> env -> env
val add_constant :
- section_path -> constant_body -> env -> env
+ constant -> constant_body -> env -> env
val add_mind :
section_path -> mutual_inductive_body -> env -> env
-(*s Looks up in environment *)
-
-(* Looks up in the context of local vars referred by names ([named_context]) *)
-(* raises [Not_found] if the identifier is not found *)
-val lookup_named_type : identifier -> env -> types
-val lookup_named_value : identifier -> env -> constr option
-val lookup_named : identifier -> env -> constr option * types
+(*s Lookups in environment *)
(* Looks up in the context of local vars referred by indice ([rel_context]) *)
(* raises [Not_found] if the index points out of the context *)
-val lookup_rel_type : int -> env -> name * types
-val lookup_rel_value : int -> env -> constr option
+val lookup_rel : int -> env -> rel_declaration
+
+(* Looks up in the context of local vars referred by names ([named_context]) *)
+(* raises [Not_found] if the identifier is not found *)
+val lookup_named : variable -> env -> named_declaration
(* Looks up in the context of global constant names *)
(* raises [Not_found] if the required path is not found *)
val lookup_constant : constant -> env -> constant_body
+(*s [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] *)
+type const_evaluation_result = NoBody | Opaque
+exception NotEvaluableConst of const_evaluation_result
+
+val constant_value : env -> constant -> constr
+val constant_type : env -> constant -> types
+val constant_opt_value : env -> constant -> constr option
+
(* Looks up in the context of global inductive names *)
(* raises [Not_found] if the required path is not found *)
val lookup_mind : section_path -> env -> mutual_inductive_body
-(* Looks up the array of section variables used by a global (constant,
- inductive or constructor). *)
-val lookup_constant_variables : constant -> env -> constr array
-val lookup_inductive_variables : inductive -> env -> constr array
-val lookup_constructor_variables : constructor -> env -> constr array
-
-(*s Miscellanous *)
-
-val sp_of_global : env -> global_reference -> section_path
-
-val id_of_global : env -> global_reference -> identifier
-
-val make_all_name_different : env -> env
-
-(*s Functions creating names for anonymous names *)
-
-val id_of_name_using_hdchar : env -> constr -> name -> identifier
-(* [named_hd env t na] just returns [na] is it defined, otherwise it
- creates a name built from [t] (e.g. ["n"] if [t] is [nat]) *)
-
-val named_hd : env -> constr -> name -> name
-
-(* [lambda_name env (na,t,c)] builds [[[x:t]c] where [x] is created
- using [named_hd] if [na] is [Anonymous]; [prod_name env (na,t,c)]
- works similarly but build a product; for [it_lambda_name env c
- sign] and [it_prod_name env c sign], more recent types should come
- first in [sign]; none of these functions substitute named
- variables in [c] by de Bruijn indices *)
-
-val lambda_name : env -> name * types * constr -> constr
-val prod_name : env -> name * types * constr -> constr
-
-val mkLambda_or_LetIn_name : env -> constr -> rel_declaration -> constr
-val mkProd_or_LetIn_name : env -> constr -> rel_declaration -> constr
-
-val it_mkLambda_or_LetIn_name : env -> constr -> rel_context -> constr
-val it_mkProd_or_LetIn_name : env -> constr -> rel_context -> constr
-
-val it_mkProd_wo_LetIn : constr -> rel_context -> constr
-val it_mkLambda_or_LetIn : constr -> rel_context -> constr
-val it_mkProd_or_LetIn : constr -> rel_context -> constr
-
-val it_mkNamedLambda_or_LetIn : constr -> named_context -> constr
-val it_mkNamedProd_or_LetIn : constr -> named_context -> constr
-val it_mkNamedProd_wo_LetIn : constr -> named_context -> constr
-
-(* [lambda_create env (t,c)] builds [[x:t]c] where [x] is a name built
- from [t]; [prod_create env (t,c)] builds [(x:t)c] where [x] is a
- name built from [t] *)
+(* [global_vars_set c] returns the list of [id]'s occurring as [VAR
+ id] in [c] *)
+val global_vars_set : env -> constr -> Idset.t
+(* the constr must be an atomic construction *)
+val vars_of_global : env -> constr -> identifier list
-val lambda_create : env -> types * constr -> constr
-val prod_create : env -> types * constr -> constr
+val keep_hyps : env -> Idset.t -> section_context
val defined_constant : env -> constant -> bool
val evaluable_constant : env -> constant -> bool
-val evaluable_named_decl : env -> identifier -> bool
+val evaluable_named_decl : env -> variable -> bool
val evaluable_rel_decl : env -> int -> bool
-(*s Ocurrence of section variables. *)
-(* [(occur_var id c)] returns [true] if variable [id] occurs free
- in c, [false] otherwise *)
-val occur_var : env -> identifier -> constr -> bool
-val occur_var_in_decl : env -> identifier -> named_declaration -> bool
-
-(* [global_vars c] returns the list of [id]'s occurring as [VAR id] in [c] *)
-val global_vars : env -> constr -> identifier list
-
-(* [global_vars_decl d] returns the list of [id]'s occurring as [VAR
- id] in declaration [d] (type and body if any) *)
-val global_vars_decl : env -> named_declaration -> identifier list
-val global_vars_set : env -> constr -> Idset.t
-
-val keep_hyps : env -> Idset.t -> named_context -> named_context
-
-val rename_bound_var : env -> identifier list -> constr -> constr
-
(*s Modules. *)
type compiled_env
@@ -203,10 +122,10 @@ type unsafe_judgment = {
uj_val : constr;
uj_type : types }
+val make_judge : constr -> types -> unsafe_judgment
+val j_val : unsafe_judgment -> constr
+val j_type : unsafe_judgment -> types
+
type unsafe_type_judgment = {
utj_val : constr;
utj_type : sorts }
-
-(*s Displays the memory use of an environment. *)
-
-val mem : env -> Pp.std_ppcmds
diff --git a/kernel/evd.ml b/kernel/evd.ml
deleted file mode 100644
index a80f21b52..000000000
--- a/kernel/evd.ml
+++ /dev/null
@@ -1,74 +0,0 @@
-(***********************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
-(* \VV/ *************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(***********************************************************************)
-
-(* $Id$ *)
-
-open Util
-open Names
-open Term
-open Sign
-
-(* The type of mappings for existential variables *)
-
-type evar = int
-
-type evar_body =
- | Evar_empty
- | Evar_defined of constr
-
-type 'a evar_info = {
- evar_concl : constr;
- evar_hyps : named_context;
- evar_body : evar_body;
- evar_info : 'a option }
-
-type 'a evar_map = 'a evar_info Intmap.t
-
-let empty = Intmap.empty
-
-let to_list evc = Intmap.fold (fun ev x acc -> (ev,x)::acc) evc []
-let dom evc = Intmap.fold (fun ev _ acc -> ev::acc) evc []
-let map evc k = Intmap.find k evc
-let rmv evc k = Intmap.remove k evc
-let remap evc k i = Intmap.add k i evc
-let in_dom evc k = Intmap.mem k evc
-
-let add evd ev newinfo = Intmap.add ev newinfo evd
-
-let define evd ev body =
- let oldinfo = map evd ev in
- let newinfo =
- { evar_concl = oldinfo.evar_concl;
- evar_hyps = oldinfo.evar_hyps;
- evar_body = Evar_defined body;
- evar_info = oldinfo.evar_info }
- in
- match oldinfo.evar_body with
- | Evar_empty -> Intmap.add ev newinfo evd
- | _ -> anomaly "cannot define an isevar twice"
-
-(* The list of non-instantiated existential declarations *)
-
-let non_instantiated sigma =
- let listev = to_list sigma in
- List.fold_left
- (fun l ((ev,evd) as d) ->
- if evd.evar_body = Evar_empty then (d::l) else l)
- [] listev
-
-let is_evar sigma ev = in_dom sigma ev
-
-let is_defined sigma ev =
- let info = map sigma ev in
- not (info.evar_body = Evar_empty)
-
-let evar_body ev = ev.evar_body
-
-let id_of_existential ev =
- id_of_string ("?" ^ string_of_int ev)
-
diff --git a/kernel/evd.mli b/kernel/evd.mli
deleted file mode 100644
index f6192c7e5..000000000
--- a/kernel/evd.mli
+++ /dev/null
@@ -1,57 +0,0 @@
-(***********************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
-(* \VV/ *************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(***********************************************************************)
-
-(*i $Id$ i*)
-
-(*i*)
-open Names
-open Term
-open Sign
-(*i*)
-
-(* The type of mappings for existential variables.
- The keys are integers and the associated information is a record
- containing the type of the evar ([evar_concl]), the context under which
- it was introduced ([evar_hyps]) and its definition ([evar_body]).
- [evar_info] is used to add any other kind of information. *)
-
-type evar = int
-
-type evar_body =
- | Evar_empty
- | Evar_defined of constr
-
-type 'a evar_info = {
- evar_concl : constr;
- evar_hyps : named_context;
- evar_body : evar_body;
- evar_info : 'a option }
-
-type 'a evar_map
-
-val empty : 'a evar_map
-
-val add : 'a evar_map -> evar -> 'a evar_info -> 'a evar_map
-
-val dom : 'a evar_map -> evar list
-val map : 'a evar_map -> evar -> 'a evar_info
-val rmv : 'a evar_map -> evar -> 'a evar_map
-val remap : 'a evar_map -> evar -> 'a evar_info -> 'a evar_map
-val in_dom : 'a evar_map -> evar -> bool
-val to_list : 'a evar_map -> (evar * 'a evar_info) list
-
-val define : 'a evar_map -> evar -> constr -> 'a evar_map
-
-val non_instantiated : 'a evar_map -> (evar * 'a evar_info) list
-val is_evar : 'a evar_map -> evar -> bool
-
-val is_defined : 'a evar_map -> evar -> bool
-
-val evar_body : 'a evar_info -> evar_body
-
-val id_of_existential : evar -> identifier
diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml
index 33a26c800..1255e9787 100644
--- a/kernel/indtypes.ml
+++ b/kernel/indtypes.ml
@@ -10,24 +10,34 @@
open Util
open Names
+open Univ
open Term
open Declarations
open Inductive
open Sign
open Environ
-open Instantiate
open Reduction
open Typeops
-(* In the following, each time an [evar_map] is required, then [Evd.empty]
- is given, since inductive types are typed in an environment without
- existentials. *)
-
(* [check_constructors_names id s cl] checks that all the constructors names
appearing in [l] are not present in the set [s], and returns the new set
of names. The name [id] is the name of the current inductive type, used
when reporting the error. *)
+(*s Declaration. *)
+
+type one_inductive_entry = {
+ mind_entry_nparams : int;
+ mind_entry_params : (identifier * local_entry) list;
+ mind_entry_typename : identifier;
+ mind_entry_arity : constr;
+ mind_entry_consnames : identifier list;
+ mind_entry_lc : constr list }
+
+type mutual_inductive_entry = {
+ mind_entry_finite : bool;
+ mind_entry_inds : one_inductive_entry list }
+
(***********************************************************************)
(* Various well-formedness check for inductive declarations *)
@@ -85,7 +95,7 @@ let mind_extract_params = decompose_prod_n_assum
let mind_check_arities env mie =
let check_arity id c =
- if not (is_arity env Evd.empty c) then
+ if not (is_arity env c) then
raise (InductiveError (NotAnArity id))
in
List.iter
@@ -98,6 +108,143 @@ let mind_check_wellformed env mie =
mind_check_arities env mie
(***********************************************************************)
+(***********************************************************************)
+
+(* Typing the arities and constructor types *)
+
+let is_info_arity env c =
+ match dest_arity env c with
+ | (_,Prop Null) -> false
+ | (_,Prop Pos) -> true
+ | (_,Type _) -> true
+
+let is_info_type env t =
+ let s = t.utj_type in
+ if s = mk_Set then true
+ else if s = mk_Prop then false
+ else
+ try is_info_arity env t.utj_val
+ with UserError _ -> true
+
+(* [infos] is a sequence of pair [islogic,issmall] for each type in
+ the product of a constructor or arity *)
+
+let is_small infos = List.for_all (fun (logic,small) -> small) infos
+let is_logic_constr infos = List.for_all (fun (logic,small) -> logic) infos
+let is_logic_arity infos =
+ List.for_all (fun (logic,small) -> logic || small) infos
+
+let is_unit arinfos constrsinfos =
+ match constrsinfos with (* One info = One constructor *)
+ | [constrinfos] -> is_logic_constr constrinfos && is_logic_arity arinfos
+ | _ -> false
+
+let rec infos_and_sort env t =
+ match kind_of_term t with
+ | Prod (name,c1,c2) ->
+ let (varj,_) = infer_type env c1 in
+ let env1 = Environ.push_rel (name,None,varj.utj_val) env in
+ let logic = not (is_info_type env varj) in
+ let small = Term.is_small varj.utj_type in
+ (logic,small) :: (infos_and_sort env1 c2)
+ | Cast (c,_) -> infos_and_sort env c
+ | _ -> []
+
+let small_unit constrsinfos (env_ar_par,short_arity) =
+ let issmall = List.for_all is_small constrsinfos in
+ let arinfos = infos_and_sort env_ar_par short_arity in
+ let isunit = is_unit arinfos constrsinfos in
+ issmall, isunit
+
+(* This (re)computes informations relevant to extraction and the sort of an
+ arity or type constructor; we do not to recompute universes constraints *)
+
+(* [smax] is the max of the sorts of the products of the constructor type *)
+
+let enforce_type_constructor arsort smax cst =
+ match smax, arsort with
+ | Type uc, Type ua -> enforce_geq ua uc cst
+ | _,_ -> cst
+
+let type_one_constructor env_ar_par params arsort c =
+ let infos = infos_and_sort env_ar_par c in
+
+ (* Each constructor is typed-checked here *)
+ let (j,cst) = infer_type env_ar_par c in
+ let full_cstr_type = it_mkProd_or_LetIn j.utj_val params in
+
+ (* If the arity is at some level Type arsort, then the sort of the
+ constructor must be below arsort; here we consider constructors with the
+ global parameters (which add a priori more constraints on their sort) *)
+ let cst2 = enforce_type_constructor arsort j.utj_type cst in
+
+ (infos, full_cstr_type, cst2)
+
+let infer_constructor_packet env_ar params short_arity arsort vc =
+ let env_ar_par = push_rel_context params env_ar in
+ let (constrsinfos,jlc,cst) =
+ List.fold_right
+ (fun c (infosl,l,cst) ->
+ let (infos,ct,cst') =
+ type_one_constructor env_ar_par params arsort c in
+ (infos::infosl,ct::l, Constraint.union cst cst'))
+ vc
+ ([],[],Constraint.empty) in
+ let vc' = Array.of_list jlc in
+ let issmall,isunit = small_unit constrsinfos (env_ar_par,short_arity) in
+ (issmall,isunit,vc', cst)
+
+let type_inductive env mie =
+ (* We first type params and arity of each inductive definition *)
+ (* This allows to build the environment of arities and to share *)
+ (* the set of constraints *)
+ let cst, arities, rev_params_arity_list =
+ List.fold_left
+ (fun (cst,arities,l) ind ->
+ (* Params are typed-checked here *)
+ let params = ind.mind_entry_params in
+ let env_params, params, cst1 =
+ infer_local_decls env params in
+ (* Arities (without params) are typed-checked here *)
+ let arity, cst2 =
+ infer_type env_params ind.mind_entry_arity 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 id = ind.mind_entry_typename in
+ let full_arity = it_mkProd_or_LetIn arity.utj_val params in
+ Constraint.union cst (Constraint.union cst1 cst2),
+ Sign.add_rel_decl (Name id, None, full_arity) arities,
+ (params, id, full_arity, arity.utj_val)::l)
+ (Constraint.empty,empty_rel_context,[])
+ mie.mind_entry_inds in
+
+ let env_arities = push_rel_context arities env in
+
+ let params_arity_list = List.rev rev_params_arity_list in
+
+ (* Now, we type the constructors (without params) *)
+ let inds,cst =
+ List.fold_right2
+ (fun ind (params,id,full_arity,short_arity) (inds,cst) ->
+ let (_,arsort) = dest_arity env full_arity in
+ let lc = ind.mind_entry_lc in
+ let (issmall,isunit,lc',cst') =
+ infer_constructor_packet env_arities params short_arity arsort lc
+ in
+ let nparams = ind.mind_entry_nparams in
+ let consnames = ind.mind_entry_consnames in
+ let ind' = (params,nparams,id,full_arity,consnames,issmall,isunit,lc')
+ in
+ (ind'::inds, Constraint.union cst cst'))
+ mie.mind_entry_inds
+ params_arity_list
+ ([],cst) in
+ (env_arities, inds, cst)
+
+(***********************************************************************)
+(***********************************************************************)
let allowed_sorts issmall isunit = function
| Type _ ->
@@ -118,7 +265,7 @@ exception IllFormedInd of ill_formed_ind
let explain_ind_err ntyp env0 nbpar c err =
let (lpar,c') = mind_extract_params nbpar c in
- let env = push_rels lpar env0 in
+ let env = push_rel_context lpar env0 in
match err with
| LocalNonPos kt ->
raise (InductiveError (NonPos (env,c',mkRel (kt+nbpar))))
@@ -150,8 +297,8 @@ let check_correct_par env hyps nparams ntypes n l largs =
| [] -> ()
| (_,Some _,_)::hyps -> check k (index+1) hyps
| _::hyps ->
- match kind_of_term (whd_betadeltaiotaeta env Evd.empty lpar.(k)) with
- | IsRel w when w = index -> check (k-1) (index+1) hyps
+ match kind_of_term (whd_betadeltaiota env lpar.(k)) with
+ | Rel w when w = index -> check (k-1) (index+1) hyps
| _ -> raise (IllFormedInd (LocalNonPar (k+1,l)))
in check (nparams-1) (n-nhyps) hyps;
if not (array_for_all (noccur_between n ntypes) largs') then
@@ -166,20 +313,20 @@ let abstract_mind_lc env ntyps npars lc =
list_tabulate
(function i -> lambda_implicit_lift npars (mkRel (i+1))) ntyps
in
- Array.map (compose nf_beta (substl make_abs)) lc
+ Array.map (substl make_abs) lc
let listrec_mconstr env ntypes hyps nparams i indlc =
let nhyps = List.length hyps in
(* check the inductive types occur positively in [c] *)
let rec check_pos env n c =
- let x,largs = whd_betadeltaiota_stack env Evd.empty c in
+ let x,largs = decompose_app (whd_betadeltaiota env c) in
match kind_of_term x with
- | IsProd (na,b,d) ->
+ | Prod (na,b,d) ->
assert (largs = []);
if not (noccur_between n ntypes b) then
raise (IllFormedInd (LocalNonPos n));
- check_pos (push_rel_assum (na, b) env) (n+1) d
- | IsRel k ->
+ check_pos (push_rel (na, None, b) env) (n+1) d
+ | Rel k ->
if k >= n && k<n+ntypes then begin
check_correct_par env hyps nparams ntypes n (k-n+1) largs;
Mrec(n+ntypes-k-1)
@@ -189,7 +336,7 @@ let listrec_mconstr env ntypes hyps nparams i indlc =
else Norec
else
raise (IllFormedInd (LocalNonPos n))
- | IsMutInd ind_sp ->
+ | Ind ind_sp ->
if List.for_all (noccur_between n ntypes) largs
then Norec
else Imbr(ind_sp,imbr_positive env n ind_sp largs)
@@ -199,27 +346,29 @@ let listrec_mconstr env ntypes hyps nparams i indlc =
then Norec
else raise (IllFormedInd (LocalNonPos n))
+ (* accesses to the environment are not factorised, but does it worth
+ it? *)
and imbr_positive env n mi largs =
- let mispeci = lookup_mind_specif mi env in
- let auxnpar = mis_nparams mispeci in
+ let (mib,mip) = lookup_mind_specif env mi in
+ let auxnpar = mip.mind_nparams in
let (lpar,auxlargs) = list_chop auxnpar largs in
if not (List.for_all (noccur_between n ntypes) auxlargs) then
raise (IllFormedInd (LocalNonPos n));
- let auxlc = mis_nf_lc mispeci
- and auxntyp = mis_ntypes mispeci in
+ let auxlc = arities_of_constructors env mi in
+ let auxntyp = mib.mind_ntypes in
if auxntyp <> 1 then raise (IllFormedInd (LocalNonPos n));
let lrecargs = List.map (check_weak_pos env n) lpar in
(* The abstract imbricated inductive type with parameters substituted *)
let auxlcvect = abstract_mind_lc env auxntyp auxnpar auxlc in
let newidx = n + auxntyp in
(* Extends the environment with a variable corresponding to the inductive def *)
- let env' = push_rel_assum (Anonymous,mis_arity mispeci) env in
+ let env' = push_rel (Anonymous,None,type_of_inductive env mi) env in
let _ =
(* fails if the inductive type occurs non positively *)
(* when substituted *)
Array.map
(function c ->
- let c' = hnf_prod_applist env Evd.empty c
+ let c' = hnf_prod_applist env c
(List.map (lift auxntyp) lpar) in
check_construct env' false newidx c')
auxlcvect
@@ -240,16 +389,16 @@ let listrec_mconstr env ntypes hyps nparams i indlc =
Abstractions may occur in imbricated recursive ocurrences, but I am
not sure if they make sense in a form of constructor. This is why I
chose to duplicated the code. Eduardo 13/7/99. *)
- (* Since Lambda can no longer occur after a product or a MutInd,
+ (* Since Lambda can no longer occur after a product or a Ind,
I have branched the remaining cases on check_pos. HH 28/1/00 *)
and check_weak_pos env n c =
- let x = whd_betadeltaiota env Evd.empty c in
+ let x = whd_betadeltaiota env c in
match kind_of_term x with
(* The extra case *)
- | IsLambda (na,b,d) ->
+ | Lambda (na,b,d) ->
if noccur_between n ntypes b
- then check_weak_pos (push_rel_assum (na,b) env) (n+1) d
+ then check_weak_pos (push_rel (na,None,b) env) (n+1) d
else raise (IllFormedInd (LocalNonPos n))
(******************)
| _ -> check_pos env n x
@@ -260,29 +409,29 @@ let listrec_mconstr env ntypes hyps nparams i indlc =
and check_construct env check_head =
let rec check_constr_rec env lrec n c =
- let x,largs = whd_betadeltaiota_stack env Evd.empty c in
+ let x,largs = decompose_app (whd_betadeltaiota env c) in
match kind_of_term x with
- | IsProd (na,b,d) ->
+ | Prod (na,b,d) ->
assert (largs = []);
let recarg = check_pos env n b in
- check_constr_rec (push_rel_assum (na, b) env)
+ check_constr_rec (push_rel (na, None, b) env)
(recarg::lrec) (n+1) d
(* LetIn's must be free of occurrence of the inductive types and
they do not contribute to recargs *)
- | IsLetIn (na,b,t,d) ->
+ | LetIn (na,b,t,d) ->
assert (largs = []);
if not (noccur_between n ntypes b & noccur_between n ntypes t) then
- check_constr_rec (push_rel_def (na,b, b) env)
+ check_constr_rec (push_rel (na,Some b, b) env)
lrec n (subst1 b d)
else
let recarg = check_pos env n b in
- check_constr_rec (push_rel_def (na,b, b) env)
+ check_constr_rec (push_rel (na,Some b, b) env)
lrec (n+1) d
| hd ->
if check_head then
- if hd = IsRel (n+ntypes-i) then
+ if hd = Rel (n+ntypes-i) then
check_correct_par env hyps nparams ntypes n (ntypes-i+1) largs
else
raise (IllFormedInd LocalNotConstructor)
@@ -296,7 +445,7 @@ let listrec_mconstr env ntypes hyps nparams i indlc =
(fun c ->
let c = body_of_type c in
let sign, rawc = mind_extract_params nhyps c in
- let env' = push_rels sign env in
+ let env' = push_rel_context sign env in
try
check_construct env' true (1+nhyps) rawc
with IllFormedInd err ->
@@ -317,19 +466,19 @@ let abstract_inductive ntypes hyps (par,np,id,arity,cnames,issmall,isunit,lc) =
let nhyps = List.length hyps in
let nparams = Array.length args in (* nparams = nhyps - nb(letin) *)
let new_refs =
- list_tabulate (fun k -> appvect(mkRel (k+nhyps+1),args)) ntypes in
+ list_tabulate (fun k -> mkApp (mkRel (k+nhyps+1),args)) ntypes in
let abs_constructor b = it_mkNamedProd_or_LetIn (substl new_refs b) hyps in
let lc' = Array.map abs_constructor lc in
let arity' = it_mkNamedProd_or_LetIn arity hyps in
let par' = push_named_to_rel_context hyps par in
(par',np+nparams,id,arity',cnames,issmall,isunit,lc')
-let cci_inductive locals env env_ar kind finite inds cst =
+let cci_inductive env env_ar finite inds cst =
let ntypes = List.length inds in
let ids =
List.fold_left
(fun acc (_,_,_,ar,_,_,_,lc) ->
- Idset.union (global_vars_set env (body_of_type ar))
+ Idset.union (Environ.global_vars_set env (body_of_type ar))
(Array.fold_left
(fun acc c ->
Idset.union (global_vars_set env (body_of_type c)) acc)
@@ -337,41 +486,46 @@ let cci_inductive locals env env_ar kind finite inds cst =
lc))
Idset.empty inds
in
- let hyps = keep_hyps env ids (named_context env) in
+ let hyps = keep_hyps env ids in
let one_packet i (params,nparams,id,ar,cnames,issmall,isunit,lc) =
let recargs = listrec_mconstr env_ar ntypes params nparams i lc in
let isunit = isunit && ntypes = 1 && (not (is_recursive [0] recargs)) in
- let (ar_sign,ar_sort) = splay_arity env Evd.empty (body_of_type ar) in
+ let (ar_sign,ar_sort) = dest_arity env ar in
- let nf_ar,user_ar =
- if isArity (body_of_type ar) then ar,None
- else (prod_it (mkSort ar_sort) ar_sign, Some ar) in
+ let nf_ar =
+ if isArity (body_of_type ar) then ar
+ else it_mkProd_or_LetIn (mkSort ar_sort) ar_sign in
let kelim = allowed_sorts issmall isunit ar_sort in
- let lc_bodies = Array.map body_of_type lc in
- let splayed_lc = Array.map (splay_prod_assum env_ar Evd.empty) lc_bodies in
+ let splayed_lc = Array.map (dest_prod_assum env_ar) lc in
let nf_lc =
array_map2 (fun (d,b) c -> it_mkProd_or_LetIn b d) splayed_lc lc in
- let nf_lc,user_lc = if nf_lc = lc then lc,None else nf_lc, Some lc in
+ let nf_lc = if nf_lc = lc then lc else nf_lc in
{ mind_consnames = Array.of_list cnames;
mind_typename = id;
- mind_user_lc = user_lc;
+ mind_user_lc = lc;
mind_nf_lc = nf_lc;
- mind_user_arity = user_ar;
+ mind_user_arity = ar;
mind_nf_arity = nf_ar;
mind_nrealargs = List.length ar_sign - nparams;
mind_sort = ar_sort;
mind_kelim = kelim;
mind_listrec = recargs;
- mind_finite = finite;
mind_nparams = nparams;
mind_params_ctxt = params }
in
- let sp_hyps = List.map (fun (id,b,t) -> (List.assoc id locals,b,t)) hyps in
let packets = Array.of_list (list_map_i one_packet 1 inds) in
- { mind_kind = kind;
- mind_ntypes = ntypes;
- mind_hyps = sp_hyps;
+ { mind_ntypes = ntypes;
+ mind_finite = finite;
+ mind_hyps = hyps;
mind_packets = packets;
mind_constraints = cst;
mind_singl = None }
+
+(***********************************************************************)
+(***********************************************************************)
+
+let check_inductive env mie =
+ mind_check_wellformed env mie;
+ let (env_arities, inds, cst) = type_inductive env mie in
+ cci_inductive env env_arities mie.mind_entry_finite inds cst
diff --git a/kernel/indtypes.mli b/kernel/indtypes.mli
index 93bfb5454..7e803b11e 100644
--- a/kernel/indtypes.mli
+++ b/kernel/indtypes.mli
@@ -14,6 +14,7 @@ open Univ
open Term
open Declarations
open Environ
+open Typeops
(*i*)
@@ -37,21 +38,32 @@ type inductive_error =
exception InductiveError of inductive_error
-(*s The following function does checks on inductive declarations. *)
+(*s Declaration of inductive types. *)
+
+(* Assume the following definition in concrete syntax:
+\begin{verbatim}
+Inductive I1 [x1:X1;...;xn:Xn] : A1 := c11 : T11 | ... | c1n1 : T1n1
+...
+with Ip [x1:X1;...;xn:Xn] : Ap := cp1 : Tp1 | ... | cpnp : Tpnp.
+\end{verbatim}
+then, in $i^{th}$ block, [mind_entry_params] is [[xn:Xn;...;x1:X1]];
+[mind_entry_arity] is [Ai], defined in context [[[x1:X1;...;xn:Xn]];
+[mind_entry_lc] is [Ti1;...;Tini], defined in context [[A'1;...;A'p;x1:X1;...;xn:Xn]] where [A'i] is [Ai] generalized over [[x1:X1;...;xn:Xn]].
+*)
-(* [mind_check_wellformed env mie] checks that the types declared for
- all the inductive types are arities. It checks also that
- constructor and inductive names altogether are distinct. It raises
- an exception [InductiveError _] if [mie] is not well-formed *)
+type one_inductive_entry = {
+ mind_entry_nparams : int;
+ mind_entry_params : (identifier * local_entry) list;
+ mind_entry_typename : identifier;
+ mind_entry_arity : constr;
+ mind_entry_consnames : identifier list;
+ mind_entry_lc : constr list }
-val mind_check_wellformed : env -> mutual_inductive_entry -> unit
+type mutual_inductive_entry = {
+ mind_entry_finite : bool;
+ mind_entry_inds : one_inductive_entry list }
-(* [cci_inductive] checks positivity and builds an inductive body *)
+(*s The following function does checks on inductive declarations. *)
-val cci_inductive :
- (identifier * variable) list -> env -> env -> path_kind -> bool ->
- (Sign.rel_context * int * identifier * types *
- identifier list * bool * bool * types array)
- list ->
- constraints ->
- mutual_inductive_body
+val check_inductive :
+ env -> mutual_inductive_entry -> mutual_inductive_body
diff --git a/kernel/inductive.ml b/kernel/inductive.ml
index 6cd04f76f..06219f084 100644
--- a/kernel/inductive.ml
+++ b/kernel/inductive.ml
@@ -16,6 +16,44 @@ open Sign
open Declarations
open Environ
open Reduction
+open Type_errors
+
+exception Induc
+
+(* raise Induc if not an inductive type *)
+let lookup_mind_specif env (sp,tyi) =
+ let mib =
+ try Environ.lookup_mind sp env
+ with Not_found -> raise Induc in
+ if tyi >= Array.length mib.mind_packets then
+ error "Inductive.lookup_mind_specif: invalid inductive index";
+ (mib, mib.mind_packets.(tyi))
+
+let lookup_recargs env ind =
+ let (mib,mip) = lookup_mind_specif env ind in
+ Array.map (fun mip -> mip.mind_listrec) mib.mind_packets
+
+let find_rectype env c =
+ let (t, l) = decompose_app (whd_betadeltaiota env c) in
+ match kind_of_term t with
+ | Ind ind -> (ind, l)
+ | _ -> raise Induc
+
+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)
+ | _ -> raise Induc
+
+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)
+ | _ -> raise Induc
+
+(***********************************************************************)
type inductive_instance = {
mis_sp : section_path;
@@ -23,189 +61,95 @@ type inductive_instance = {
mis_tyi : int;
mis_mip : one_inductive_body }
-
-let build_mis (sp,tyi) mib =
- { mis_sp = sp; mis_mib = mib; mis_tyi = tyi;
- mis_mip = mind_nth_type_packet mib tyi }
-
-let mis_ntypes mis = mis.mis_mib.mind_ntypes
-let mis_nparams mis = mis.mis_mip.mind_nparams
-
-let mis_index mis = mis.mis_tyi
-
let mis_nconstr mis = Array.length (mis.mis_mip.mind_consnames)
-let mis_nrealargs mis = mis.mis_mip.mind_nrealargs
-let mis_kelim mis = mis.mis_mip.mind_kelim
-let mis_recargs mis =
- Array.map (fun mip -> mip.mind_listrec) mis.mis_mib.mind_packets
-let mis_recarg mis = mis.mis_mip.mind_listrec
-let mis_typename mis = mis.mis_mip.mind_typename
-let mis_typepath mis =
- make_path (dirpath mis.mis_sp) mis.mis_mip.mind_typename CCI
-let mis_consnames mis = mis.mis_mip.mind_consnames
-let mis_conspaths mis =
- let dir = dirpath mis.mis_sp in
- Array.map (fun id -> make_path dir id CCI) mis.mis_mip.mind_consnames
let mis_inductive mis = (mis.mis_sp,mis.mis_tyi)
-let mis_finite mis = mis.mis_mip.mind_finite
-
-let mis_typed_nf_lc mis =
- let sign = mis.mis_mib.mind_hyps in
- mis.mis_mip.mind_nf_lc
-
-let mis_nf_lc mis = Array.map body_of_type (mis_typed_nf_lc mis)
-
-let mis_user_lc mis =
- let sign = mis.mis_mib.mind_hyps in
- (mind_user_lc mis.mis_mip)
-
-(* gives the vector of constructors and of
- types of constructors of an inductive definition
- correctly instanciated *)
-
-let mis_type_mconstructs mispec =
- let specif = Array.map body_of_type (mis_user_lc mispec)
- and ntypes = mis_ntypes mispec
- and nconstr = mis_nconstr mispec in
- let make_Ik k = mkMutInd (mispec.mis_sp,ntypes-k-1)
- and make_Ck k =
- mkMutConstruct ((mispec.mis_sp,mispec.mis_tyi),k+1) in
- (Array.init nconstr make_Ck,
- Array.map (substl (list_tabulate make_Ik ntypes)) specif)
-
-let mis_nf_constructor_type i mispec =
- let specif = mis_nf_lc mispec
- and ntypes = mis_ntypes mispec
- and nconstr = mis_nconstr mispec in
- let make_Ik k = mkMutInd (mispec.mis_sp,ntypes-k-1) in
- if i > nconstr then error "Not enough constructors in the type";
- substl (list_tabulate make_Ik ntypes) specif.(i-1)
-
-let mis_constructor_type i mispec =
- let specif = mis_user_lc mispec
- and ntypes = mis_ntypes mispec
- and nconstr = mis_nconstr mispec in
- let make_Ik k = mkMutInd (mispec.mis_sp,ntypes-k-1) in
- if i > nconstr then error "Not enough constructors in the type";
- substl (list_tabulate make_Ik ntypes) specif.(i-1)
-
-let mis_arity mis =
- let hyps = mis.mis_mib.mind_hyps in
- mind_user_arity mis.mis_mip
-
-let mis_nf_arity mis =
- let hyps = mis.mis_mib.mind_hyps in
- mis.mis_mip.mind_nf_arity
-let mis_params_ctxt mis = mis.mis_mip.mind_params_ctxt
-(*
- let paramsign,_ =
- decompose_prod_n_assum mis.mis_mip.mind_nparams
- (body_of_type (mis_nf_arity mis))
- in paramsign
-*)
+let lookup_mind_instance (sp,tyi) env =
+ let (mib,mip) = lookup_mind_specif env (sp,tyi) in
+ { mis_sp = sp; mis_mib = mib; mis_tyi = tyi; mis_mip = mip }
-let mis_sort mispec = mispec.mis_mip.mind_sort
+(* Build the substitution that replaces Rels by the appropriate *)
+(* inductives *)
+let ind_subst mispec =
+ let ntypes = mispec.mis_mib.mind_ntypes in
+ let make_Ik k = mkInd (mispec.mis_sp,ntypes-k-1) in
+ (list_tabulate make_Ik ntypes)
-(* [inductive_family] = [inductive_instance] applied to global parameters *)
-type inductive_family = IndFamily of inductive_instance * constr list
+(* Instantiate both section variables and inductives *)
+let constructor_instantiate mispec =
+ let s = ind_subst mispec in
+ substl s
-type inductive_type = IndType of inductive_family * constr list
+(* Instantiate the parameters of the inductive type *)
+let instantiate_params t args sign =
+ let rec inst s t = function
+ | ((_,None,_)::ctxt,a::args) ->
+ (match kind_of_term t with
+ | Prod(_,_,t) -> inst (a::s) t (ctxt,args)
+ | _ -> anomaly"instantiate_params: type, ctxt and args mismatch")
+ | ((_,(Some b),_)::ctxt,args) ->
+ (match kind_of_term t with
+ | LetIn(_,_,_,t) -> inst ((substl s b)::s) t (ctxt,args)
+ | _ -> anomaly"instantiate_params: type, ctxt and args mismatch")
+ | [], [] -> substl s t
+ | _ -> anomaly"instantiate_params: type, ctxt and args mismatch"
+ in inst [] t (List.rev sign,args)
-let liftn_inductive_family n d (IndFamily (mis,params)) =
- IndFamily (mis, List.map (liftn n d) params)
-let lift_inductive_family n = liftn_inductive_family n 1
+let full_inductive_instantiate (mispec,params) t =
+ instantiate_params t params mispec.mis_mip.mind_params_ctxt
-let liftn_inductive_type n d (IndType (indf, realargs)) =
- IndType (liftn_inductive_family n d indf, List.map (liftn n d) realargs)
-let lift_inductive_type n = liftn_inductive_type n 1
+let full_constructor_instantiate (mispec,params) =
+ let inst_ind = constructor_instantiate mispec in
+ (fun t ->
+ instantiate_params (inst_ind t) params mispec.mis_mip.mind_params_ctxt)
-let substnl_ind_family l n (IndFamily (mis,params)) =
- IndFamily (mis, List.map (substnl l n) params)
+(***********************************************************************)
+(***********************************************************************)
-let substnl_ind_type l n (IndType (indf,realargs)) =
- IndType (substnl_ind_family l n indf, List.map (substnl l n) realargs)
+(* Functions to build standard types related to inductive *)
-let make_ind_family (mis, params) = IndFamily (mis,params)
-let dest_ind_family (IndFamily (mis,params)) = (mis,params)
+(* Type of an inductive type *)
-let make_ind_type (indf, realargs) = IndType (indf,realargs)
-let dest_ind_type (IndType (indf,realargs)) = (indf,realargs)
+let type_of_inductive env i =
+ let mis = lookup_mind_instance i env in
+ let hyps = mis.mis_mib.mind_hyps in
+ mis.mis_mip.mind_user_arity
-let mkAppliedInd (IndType (IndFamily (mis,params), realargs)) =
- applist (mkMutInd (mis_inductive mis),params@realargs)
+(* The same, with parameters instantiated *)
+let get_arity (mispec,params as indf) =
+ let arity = mispec.mis_mip.mind_nf_arity in
+ destArity (full_inductive_instantiate indf arity)
-let mis_is_recursive_subset listind mis =
- let rec one_is_rec rvec =
- List.exists
- (function
- | Mrec i -> List.mem i listind
- | Imbr(_,lvec) -> one_is_rec lvec
- | Norec -> false
- | Param _ -> false) rvec
- in
- array_exists one_is_rec (mis_recarg mis)
+(***********************************************************************)
+(* Type of a constructor *)
+
+let type_of_constructor env cstr =
+ let ind = inductive_of_constructor cstr in
+ let mispec = lookup_mind_instance ind env in
+ let specif = mispec.mis_mip.mind_user_lc in
+ let i = index_of_constructor cstr in
+ let nconstr = mis_nconstr mispec in
+ if i > nconstr then error "Not enough constructors in the type";
+ constructor_instantiate mispec specif.(i-1)
-let mis_is_recursive mis =
- mis_is_recursive_subset (interval 0 ((mis_ntypes mis)-1)) mis
+let arities_of_constructors env ind =
+ let mispec = lookup_mind_instance ind env in
+ let specif = mispec.mis_mip.mind_user_lc in
+ Array.map (constructor_instantiate mispec) specif
-(* Annotation for cases *)
-let make_case_info mis style pats_source =
-(* let constr_lengths = Array.map List.length (mis_recarg mis) in*)
- let indsp = (mis.mis_sp,mis.mis_tyi) in
- let print_info =
- (indsp,mis_consnames mis,mis.mis_mip.mind_nrealargs,style,pats_source) in
- ((*constr_lengths*) mis_nparams mis,print_info)
-let make_default_case_info mis =
- make_case_info mis None (Array.init (mis_nconstr mis) (fun _ -> RegularPat))
+(* gives the vector of constructors and of
+ types of constructors of an inductive definition
+ correctly instanciated *)
+let mis_nf_constructor_type i mispec =
+ let nconstr = mis_nconstr mispec in
+ if i > nconstr then error "Not enough constructors in the type";
+ constructor_instantiate mispec mispec.mis_mip.mind_nf_lc.(i-1)
+
(*s Useful functions *)
-let inductive_of_constructor (ind_sp,i) = ind_sp
-let index_of_constructor (ind_sp,i) = i
-let ith_constructor_of_inductive ind_sp i = (ind_sp,i)
-
-exception Induc
-
-let extract_mrectype t =
- let (t, l) = whd_stack t in
- match kind_of_term t with
- | IsMutInd ind -> (ind, l)
- | _ -> raise Induc
-
-let find_mrectype env sigma c =
- let (t, l) = whd_betadeltaiota_stack env sigma c in
- match kind_of_term t with
- | IsMutInd ind -> (ind, l)
- | _ -> raise Induc
-
-let find_inductive env sigma c =
- let (t, l) = whd_betadeltaiota_stack env sigma c in
- match kind_of_term t with
- | IsMutInd ((sp,i) as ind)
- when mind_type_finite (lookup_mind sp env) i -> (ind, l)
- | _ -> raise Induc
-
-let find_coinductive env sigma c =
- let (t, l) = whd_betadeltaiota_stack env sigma c in
- match kind_of_term t with
- | IsMutInd ((sp,i) as ind)
- when not (mind_type_finite (lookup_mind sp env) i) -> (ind, l)
- | _ -> raise Induc
-
-(* raise Induc if not an inductive type *)
-let lookup_mind_specif ((sp,tyi) as ind) env =
- build_mis ind (lookup_mind sp env)
-
-let find_rectype env sigma ty =
- let (mind,largs) = find_mrectype env sigma ty in
- let mispec = lookup_mind_specif mind env in
- let nparams = mis_nparams mispec in
- let (params,realargs) = list_chop nparams largs in
- make_ind_type (make_ind_family (mispec,params),realargs)
-
type constructor_summary = {
cs_cstr : constructor;
cs_params : constr list;
@@ -214,63 +158,24 @@ type constructor_summary = {
cs_concl_realargs : constr array
}
-let lift_constructor n cs = {
- cs_cstr = cs.cs_cstr;
- cs_params = List.map (lift n) cs.cs_params;
- cs_nargs = cs.cs_nargs;
- cs_args = lift_rel_context n cs.cs_args;
- cs_concl_realargs = Array.map (liftn n (cs.cs_nargs+1)) cs.cs_concl_realargs
-}
-
-let instantiate_params t args sign =
- let rec inst s t = function
- | ((_,None,_)::ctxt,a::args) ->
- (match kind_of_term t with
- | IsProd(_,_,t) -> inst (a::s) t (ctxt,args)
- | _ -> anomaly"instantiate_params: type, ctxt and args mismatch")
- | ((_,(Some b),_)::ctxt,args) ->
- (match kind_of_term t with
- | IsLetIn(_,_,_,t) -> inst ((substl s b)::s) t (ctxt,args)
- | _ -> anomaly"instantiate_params: type, ctxt and args mismatch")
- | [], [] -> substl s t
- | _ -> anomaly"instantiate_params: type, ctxt and args mismatch"
- in inst [] t (List.rev sign,args)
-
-let get_constructor_type (IndFamily (mispec,params)) j =
- assert (j <= mis_nconstr mispec);
- let typi = mis_constructor_type j mispec in
- instantiate_params typi params (mis_params_ctxt mispec)
-
-let get_constructors_types (IndFamily (mispec,params) as indf) =
- Array.init (mis_nconstr mispec) (fun j -> get_constructor_type indf (j+1))
-
-let get_constructor (IndFamily (mispec,params) as indf) j =
- assert (j <= mis_nconstr mispec);
- let typi = mis_nf_constructor_type j mispec in
- let typi = instantiate_params typi params (mis_params_ctxt mispec) in
+let process_constructor ((mispec,params) as indf) j typi =
+ let typi = full_constructor_instantiate indf typi in
let (args,ccl) = decompose_prod_assum typi in
- let (_,allargs) = whd_stack ccl in
- let (_,vargs) = list_chop (mis_nparams mispec) allargs in
- { cs_cstr = ith_constructor_of_inductive (mis_inductive mispec) j;
+ let (_,allargs) = decompose_app ccl in
+ let (_,vargs) = list_chop mispec.mis_mip.mind_nparams allargs in
+ { cs_cstr = ith_constructor_of_inductive (mis_inductive mispec) (j+1);
cs_params = params;
cs_nargs = rel_context_length args;
cs_args = args;
cs_concl_realargs = Array.of_list vargs }
-let get_constructors (IndFamily (mispec,params) as indf) =
- Array.init (mis_nconstr mispec) (fun j -> get_constructor indf (j+1))
-
-let get_arity_type (IndFamily (mispec,params)) =
- let arity = body_of_type (mis_arity mispec) in
-(* instantiate_params arity params (mis_params_ctxt mispec) *)
- prod_applist arity params
+let get_constructors ((mispec,params) as indf) =
+ let constr_tys = mispec.mis_mip.mind_nf_lc in
+ Array.mapi (process_constructor indf) constr_tys
-let get_arity (IndFamily (mispec,params)) =
- let arity = body_of_type (mis_nf_arity mispec) in
-(* instantiate_params arity params (mis_params_ctxt mispec) *)
- destArity (prod_applist arity params)
+(***********************************************************************)
-(* Functions to build standard types related to inductive *)
+(* Type of case branches *)
let local_rels =
let rec relrec acc n = function (* more recent arg in front *)
@@ -281,34 +186,627 @@ let local_rels =
let build_dependent_constructor cs =
applist
- (mkMutConstruct cs.cs_cstr,
+ (mkConstruct cs.cs_cstr,
(List.map (lift cs.cs_nargs) cs.cs_params)@(local_rels cs.cs_args))
-let build_dependent_inductive (IndFamily (mis, params) as indf) =
+let build_dependent_inductive ((mis, params) as indf) =
let arsign,_ = get_arity indf in
- let nrealargs = mis_nrealargs mis in
+ let nrealargs = mis.mis_mip.mind_nrealargs in
applist
- (mkMutInd (mis_inductive mis),
+ (mkInd (mis_inductive mis),
(List.map (lift nrealargs) params)@(local_rels arsign))
-(* builds the arity of an elimination predicate in sort [s] *)
-let make_arity env dep indf s =
- let (arsign,_) = get_arity indf in
- if dep then
- (* We need names everywhere *)
- it_mkProd_or_LetIn_name env
- (mkArrow (build_dependent_inductive indf) (mkSort s)) arsign
- else
- (* No need to enforce names *)
- it_mkProd_or_LetIn (mkSort s) arsign
-
(* [p] is the predicate and [cs] a constructor summary *)
-let build_branch_type env dep p cs =
- let base = appvect (lift cs.cs_nargs p, cs.cs_concl_realargs) in
- if dep then
- it_mkProd_or_LetIn_name env
- (applist (base,[build_dependent_constructor cs]))
- cs.cs_args
- else
- it_mkProd_or_LetIn base cs.cs_args
+let build_branch_type dep p cs =
+ let args =
+ if dep then
+ Array.append cs.cs_concl_realargs [|build_dependent_constructor cs|]
+ else
+ cs.cs_concl_realargs in
+ let base = beta_appvect (lift cs.cs_nargs p) args in
+ it_mkProd_or_LetIn base cs.cs_args
+
+
+let is_info_arity env c =
+ match dest_arity env c with
+ | (_,Prop Null) -> false
+ | (_,Prop Pos) -> true
+ | (_,Type _) -> true
+
+let error_elim_expln env kp ki =
+ if is_info_arity env kp && not (is_info_arity env ki) then
+ "non-informative objects may not construct informative ones."
+ else
+ match (kind_of_term kp,kind_of_term ki) with
+ | Sort (Type _), Sort (Prop _) ->
+ "strong elimination on non-small inductive types leads to paradoxes."
+ | _ -> "wrong arity"
+
+exception Arity of (constr * constr * string) option
+
+
+let is_correct_arity env kelim (c,pj) indf t =
+ let rec srec (pt,t) u =
+ let pt' = whd_betadeltaiota env pt in
+ let t' = whd_betadeltaiota env t in
+ match kind_of_term pt', kind_of_term t' with
+ | Prod (_,a1,a2), Prod (_,a1',a2') ->
+ let univ =
+ try conv env a1 a1'
+ with NotConvertible -> raise (Arity None) in
+ srec (a2,a2') (Constraint.union u univ)
+ | Prod (_,a1,a2), _ ->
+ let k = whd_betadeltaiota env a2 in
+ let ksort = match kind_of_term k with
+ | Sort s -> family_of_sort s
+ | _ -> raise (Arity None) in
+ let ind = build_dependent_inductive indf in
+ let univ =
+ try conv env a1 ind
+ with NotConvertible -> raise (Arity None) in
+ if List.exists ((=) ksort) kelim then
+ ((true,k), Constraint.union u univ)
+ else
+ raise (Arity (Some(k,t',error_elim_expln env k t')))
+ | k, Prod (_,_,_) ->
+ raise (Arity None)
+ | k, ki ->
+ let ksort = match k with
+ | Sort s -> family_of_sort s
+ | _ -> raise (Arity None) in
+ if List.exists ((=) ksort) kelim then
+ (false, pt'), u
+ else
+ raise (Arity (Some(pt',t',error_elim_expln env pt' t')))
+ in
+ try srec (pj.uj_type,t) Constraint.empty
+ with Arity kinds ->
+ let create_sort = function
+ | InProp -> mkProp
+ | InSet -> mkSet
+ | InType -> mkType (Univ.new_univ ()) in
+ let listarity = List.map create_sort kelim
+(* let listarity =
+ (List.map (fun s -> make_arity env true indf (create_sort s)) kelim)
+ @(List.map (fun s -> make_arity env false indf (create_sort s)) kelim)*)
+ in
+ let ind = mis_inductive (fst indf) in
+ error_elim_arity env ind listarity c pj kinds
+
+
+let find_case_dep_nparams env (c,pj) (ind,params) =
+ let indf = lookup_mind_instance ind env in
+ let kelim = indf.mis_mip.mind_kelim in
+ let arsign,s = get_arity (indf,params) in
+ let glob_t = it_mkProd_or_LetIn (mkSort s) arsign in
+ let ((dep,_),univ) =
+ is_correct_arity env kelim (c,pj) (indf,params) glob_t in
+ (dep,univ)
+
+
+let type_case_branches env (mind,largs) pj c =
+ let mispec = lookup_mind_instance mind env in
+ let nparams = mispec.mis_mip.mind_nparams in
+ let (params,realargs) = list_chop nparams largs in
+ let indf = (mispec,params) in
+ let p = pj.uj_val in
+ let (dep,univ) = find_case_dep_nparams env (c,pj) (mind,params) in
+ let constructs = get_constructors indf in
+ let lc = Array.map (build_branch_type dep p) constructs in
+ let args = if dep then realargs@[c] else realargs in
+ (lc, beta_appvect p (Array.of_list args), univ)
+
+
+let check_case_info env indsp ci =
+ let (mib,mip) = lookup_mind_specif env indsp in
+ if
+ (indsp <> ci.ci_ind) or
+ (mip.mind_nparams <> ci.ci_npar)
+ then raise (TypeError(env,WrongCaseInfo(indsp,ci)))
+
+(***********************************************************************)
+(***********************************************************************)
+
+(* Guard conditions for fix and cofix-points *)
+
+(* Check if t is a subterm of Rel n, and gives its specification,
+ assuming lst already gives index of
+ subterms with corresponding specifications of recursive arguments *)
+
+(* A powerful notion of subterm *)
+
+let find_sorted_assoc p =
+ let rec findrec = function
+ | (a,ta)::l ->
+ if a < p then findrec l else if a = p then ta else raise Not_found
+ | _ -> raise Not_found
+ in
+ findrec
+
+let map_lift_fst_n m = List.map (function (n,t)->(n+m,t))
+let map_lift_fst = map_lift_fst_n 1
+
+let rec instantiate_recarg sp lrc ra =
+ match ra with
+ | Mrec(j) -> Imbr((sp,j),lrc)
+ | Imbr(ind_sp,l) -> Imbr(ind_sp, List.map (instantiate_recarg sp lrc) l)
+ | Norec -> Norec
+ | Param(k) -> List.nth lrc k
+
+(* To each inductive definition corresponds an array describing the
+ structure of recursive arguments for each constructor, we call it
+ the recursive spec of the type (it has type recargs vect). For
+ checking the guard, we start from the decreasing argument (Rel n)
+ with its recursive spec. During checking the guardness condition,
+ we collect patterns variables corresponding to subterms of n, each
+ of them with its recursive spec. They are organised in a list lst
+ of type (int * recargs) list which is sorted with respect to the
+ first argument.
+*)
+
+(*
+ f is a function of type
+ env -> int -> (int * recargs) list -> constr -> 'a
+
+ c is a branch of an inductive definition corresponding to the spec
+ lrec. mind_recvec is the recursive spec of the inductive
+ definition of the decreasing argument n.
+
+ check_term env mind_recvec f n lst (lrec,c) will pass the lambdas
+ of c corresponding to pattern variables and collect possibly new
+ subterms variables and apply f to the body of the branch with the
+ correct env and decreasing arg.
+*)
+
+let check_term env mind_recvec f =
+ let rec crec env n lst (lrec,c) =
+ let c' = strip_outer_cast c in
+ match lrec, kind_of_term c' with
+ (ra::lr,Lambda (x,a,b)) ->
+ let lst' = map_lift_fst lst
+ and env' = push_rel (x,None,a) env
+ and n'=n+1
+ in begin match ra with
+ Mrec(i) -> crec env' n' ((1,mind_recvec.(i))::lst') (lr,b)
+ | Imbr((sp,i) as ind_sp,lrc) ->
+ let sprecargs = lookup_recargs env ind_sp in
+ let lc = Array.map
+ (List.map (instantiate_recarg sp lrc)) sprecargs.(i)
+ in crec env' n' ((1,lc)::lst') (lr,b)
+ | _ -> crec env' n' lst' (lr,b) end
+ | (_,_) -> f env n lst c'
+ in crec env
+
+(* c is supposed to be in beta-delta-iota head normal form *)
+
+let is_inst_var k c =
+ match kind_of_term (fst (decompose_app c)) with
+ | Rel n -> n=k
+ | _ -> false
+
+(*
+ is_subterm_specif env lcx mind_recvec n lst c
+
+ n is the principal arg and has recursive spec lcx, lst is the list
+ of subterms of n with spec. is_subterm_specif should test if c is
+ a subterm of n and fails with Not_found if not. In case it is, it
+ should send its recursive specification. This recursive spec
+ should be the same size as the number of constructors of the type
+ of c. A problem occurs when c is built by contradiction. In that
+ case no spec is given.
+*)
+let is_subterm_specif env lcx mind_recvec =
+ let rec crec env n lst c =
+ let f,l = decompose_app (whd_betadeltaiota env c) in
+ match kind_of_term f with
+ | Rel k -> Some (find_sorted_assoc k lst)
+
+ | Case ( _,_,c,br) ->
+ if Array.length br = 0 then None
+
+ else
+ let def = Array.create (Array.length br) []
+ in let lcv =
+ (try
+ if is_inst_var n c then lcx
+ else match crec env n lst c with Some lr -> lr | None -> def
+ with Not_found -> def)
+ in
+ assert (Array.length br = Array.length lcv);
+ let stl =
+ array_map2
+ (fun lc a ->
+ check_term env mind_recvec crec n lst (lc,a)) lcv br
+ in let stl0 = stl.(0) in
+ if array_for_all (fun st -> st=stl0) stl then stl0
+ else None
+
+ | Fix ((recindxs,i),(_,typarray,bodies as recdef)) ->
+ let nbfix = Array.length typarray in
+ let decrArg = recindxs.(i) in
+ let theBody = bodies.(i) in
+ let sign,strippedBody = decompose_lam_n_assum (decrArg+1) theBody in
+ let nbOfAbst = nbfix+decrArg+1 in
+(* when proving that the fixpoint f(x)=e is less than n, it is enough
+ to prove that e is less than n assuming f is less than n
+ furthermore when f is applied to a term which is strictly less than
+ n, one may assume that x itself is strictly less than n
+*)
+ let newlst =
+ let lst' = (nbOfAbst,lcx) :: (map_lift_fst_n nbOfAbst lst) in
+ if List.length l < (decrArg+1) then lst'
+ else
+ let theDecrArg = List.nth l decrArg in
+ try
+ match crec env n lst theDecrArg with
+ (Some recArgsDecrArg) -> (1,recArgsDecrArg) :: lst'
+ | None -> lst'
+ with Not_found -> lst' in
+ let env' = push_rec_types recdef env in
+ let env'' = push_rel_context sign env' in
+ crec env'' (n+nbOfAbst) newlst strippedBody
+
+ | Lambda (x,a,b) when l=[] ->
+ let lst' = map_lift_fst lst in
+ crec (push_rel (x, None, a) env) (n+1) lst' b
+
+ (*** Experimental change *************************)
+ | Meta _ -> None
+ | _ -> raise Not_found
+ in
+ crec env
+
+let spec_subterm_strict env lcx mind_recvec n lst c nb =
+ try match is_subterm_specif env lcx mind_recvec n lst c
+ with Some lr -> lr | None -> Array.create nb []
+ with Not_found -> Array.create nb []
+
+let spec_subterm_large env lcx mind_recvec n lst c nb =
+ if is_inst_var n c then lcx
+ else spec_subterm_strict env lcx mind_recvec n lst c nb
+
+
+let is_subterm env lcx mind_recvec n lst c =
+ try
+ let _ = is_subterm_specif env lcx mind_recvec n lst c in true
+ with Not_found ->
+ false
+
+(***********************************************************************)
+
+exception FixGuardError of guard_error
+
+(* Auxiliary function: it checks a condition f depending on a deBrujin
+ index for a certain number of abstractions *)
+
+let rec check_subterm_rec_meta env vectn k def =
+ (let nfi = Array.length vectn in
+ (* check fi does not appear in the k+1 first abstractions,
+ gives the type of the k+1-eme abstraction *)
+ let rec check_occur env n def =
+ match kind_of_term (strip_outer_cast def) with
+ | Lambda (x,a,b) ->
+ if noccur_with_meta n nfi a then
+ let env' = push_rel (x, None, a) env in
+ if n = k+1 then (env', lift 1 a, b)
+ else check_occur env' (n+1) b
+ else
+ anomaly "check_subterm_rec_meta: Bad occurrence of recursive call"
+ | _ -> raise (FixGuardError NotEnoughAbstractionInFixBody) in
+ let (env',c,d) = check_occur env 1 def in
+ let ((sp,tyi) as mind, largs) =
+ try find_inductive env' c
+ with Induc -> raise (FixGuardError RecursionNotOnInductiveType) in
+ let mind_recvec = lookup_recargs env' (sp,tyi) in
+ let lcx = mind_recvec.(tyi) in
+ (* n = decreasing argument in the definition;
+ lst = a mapping var |-> recargs
+ t = the term to be checked
+ *)
+ let rec check_rec_call env n lst t =
+ (* n gives the index of the recursive variable *)
+ (noccur_with_meta (n+k+1) nfi t) or
+ (* no recursive call in the term *)
+ (let f,l = hnf_stack env t in
+ match kind_of_term f with
+ | Rel p ->
+ if n+k+1 <= p & p < n+k+nfi+1 then
+ (* recursive call *)
+ let glob = nfi+n+k-p in (* the index of the recursive call *)
+ let np = vectn.(glob) in (* the decreasing arg of the rec call *)
+ if List.length l > np then
+ (match list_chop np l with
+ (la,(z::lrest)) ->
+ if (is_subterm env lcx mind_recvec n lst z)
+ then List.for_all (check_rec_call env n lst) (la@lrest)
+ else raise (FixGuardError RecursionOnIllegalTerm)
+ | _ -> assert false)
+ else raise (FixGuardError NotEnoughArgumentsForFixCall)
+ else List.for_all (check_rec_call env n lst) l
+
+ | Case (ci,p,c_0,lrest) ->
+ let lc = spec_subterm_large env lcx mind_recvec n lst c_0
+ (Array.length lrest)
+ in
+ (array_for_all2
+ (fun c0 a ->
+ check_term env mind_recvec check_rec_call n lst (c0,a))
+ lc lrest)
+ && (List.for_all (check_rec_call env n lst) (c_0::p::l))
+
+ (* Enables to traverse Fixpoint definitions in a more intelligent
+ way, ie, the rule :
+
+ if - g = Fix g/1 := [y1:T1]...[yp:Tp]e &
+ - f is guarded with respect to the set of pattern variables S
+ in a1 ... am &
+ - f is guarded with respect to the set of pattern variables S
+ in T1 ... Tp &
+ - ap is a sub-term of the formal argument of f &
+ - f is guarded with respect to the set of pattern variables S+{yp}
+ in e
+ then f is guarded with respect to S in (g a1 ... am).
+
+ Eduardo 7/9/98 *)
+
+ | Fix ((recindxs,i),(_,typarray,bodies as recdef)) ->
+ (List.for_all (check_rec_call env n lst) l) &&
+ (array_for_all (check_rec_call env n lst) typarray) &&
+ let nbfix = Array.length typarray in
+ let decrArg = recindxs.(i)
+ and env' = push_rec_types recdef env
+ and n' = n+nbfix
+ and lst' = map_lift_fst_n nbfix lst
+ in
+ if (List.length l < (decrArg+1)) then
+ array_for_all (check_rec_call env' n' lst') bodies
+ else
+ let theDecrArg = List.nth l decrArg in
+ (try
+ match
+ is_subterm_specif env lcx mind_recvec n lst theDecrArg
+ with
+ Some recArgsDecrArg ->
+ let theBody = bodies.(i) in
+ check_rec_call_fix_body
+ env' n' lst' (decrArg+1) recArgsDecrArg theBody
+ | None ->
+ array_for_all (check_rec_call env' n' lst') bodies
+ with Not_found ->
+ array_for_all (check_rec_call env' n' lst') bodies)
+
+ | Cast (a,b) ->
+ (check_rec_call env n lst a) &&
+ (check_rec_call env n lst b) &&
+ (List.for_all (check_rec_call env n lst) l)
+
+ | Lambda (x,a,b) ->
+ (check_rec_call env n lst a) &&
+ (check_rec_call (push_rel (x, None, a) env)
+ (n+1) (map_lift_fst lst) b) &&
+ (List.for_all (check_rec_call env n lst) l)
+
+ | Prod (x,a,b) ->
+ (check_rec_call env n lst a) &&
+ (check_rec_call (push_rel (x, None, a) env)
+ (n+1) (map_lift_fst lst) b) &&
+ (List.for_all (check_rec_call env n lst) l)
+
+ | LetIn (x,a,b,c) ->
+ anomaly "check_rec_call: should have been reduced"
+
+ | Ind _ ->
+ (List.for_all (check_rec_call env n lst) l)
+
+ | Construct _ ->
+ (List.for_all (check_rec_call env n lst) l)
+
+ | Const sp as c ->
+ (try
+ (List.for_all (check_rec_call env n lst) l)
+ with (FixGuardError _ ) as e
+ -> if evaluable_constant env sp then
+ check_rec_call env n lst (whd_betadeltaiota env t)
+ else raise e)
+
+ | App (f,la) ->
+ (check_rec_call env n lst f) &&
+ (array_for_all (check_rec_call env n lst) la) &&
+ (List.for_all (check_rec_call env n lst) l)
+
+ | CoFix (i,(_,typarray,bodies as recdef)) ->
+ let nbfix = Array.length typarray in
+ let env' = push_rec_types recdef env in
+ (array_for_all (check_rec_call env n lst) typarray) &&
+ (List.for_all (check_rec_call env n lst) l) &&
+ (array_for_all
+ (check_rec_call env' (n+nbfix) (map_lift_fst_n nbfix lst))
+ bodies)
+
+ | Evar (_,la) ->
+ (array_for_all (check_rec_call env n lst) la) &&
+ (List.for_all (check_rec_call env n lst) l)
+
+ | Meta _ -> true
+
+ | Var _ | Sort _ -> List.for_all (check_rec_call env n lst) l
+ )
+
+ and check_rec_call_fix_body env n lst decr recArgsDecrArg body =
+ if decr = 0 then
+ check_rec_call env n ((1,recArgsDecrArg)::lst) body
+ else
+ match kind_of_term body with
+ | Lambda (x,a,b) ->
+ (check_rec_call env n lst a) &
+ (check_rec_call_fix_body
+ (push_rel (x, None, a) env) (n+1)
+ (map_lift_fst lst) (decr-1) recArgsDecrArg b)
+ | _ -> anomaly "Not enough abstractions in fix body"
+
+ in
+ check_rec_call env' 1 [] d)
+
+(* vargs is supposed to be built from A1;..Ak;[f1]..[fk][|d1;..;dk|]
+and vdeft is [|t1;..;tk|] such that f1:A1,..,fk:Ak |- di:ti
+nvect is [|n1;..;nk|] which gives for each recursive definition
+the inductive-decreasing index
+the function checks the convertibility of ti with Ai *)
+
+let check_fix env ((nvect,bodynum),(names,types,bodies as recdef)) =
+ let nbfix = Array.length bodies in
+ if nbfix = 0
+ or Array.length nvect <> nbfix
+ or Array.length types <> nbfix
+ or Array.length names <> nbfix
+ or bodynum < 0
+ or bodynum >= nbfix
+ then anomaly "Ill-formed fix term";
+ for i = 0 to nbfix - 1 do
+ let fixenv = push_rec_types recdef env in
+ if nvect.(i) < 0 then anomaly "negative recarg position";
+ try
+ let _ = check_subterm_rec_meta fixenv nvect nvect.(i) bodies.(i)
+ in ()
+ with FixGuardError err ->
+ error_ill_formed_rec_body fixenv err names i bodies
+ done
+
+(*
+let cfkey = Profile.declare_profile "check_fix";;
+let check_fix env fix = Profile.profile3 cfkey check_fix env fix;;
+*)
+
+(***********************************************************************)
+(* Co-fixpoints. *)
+
+exception CoFixGuardError of guard_error
+
+let anomaly_ill_typed () =
+ anomaly "check_guard_rec_meta: too many arguments applied to constructor"
+
+
+let check_guard_rec_meta env nbfix def deftype =
+ let rec codomain_is_coind env c =
+ let b = whd_betadeltaiota env (strip_outer_cast c) in
+ match kind_of_term b with
+ | Prod (x,a,b) ->
+ codomain_is_coind (push_rel (x, None, a) env) b
+ | _ ->
+ try
+ find_coinductive env b
+ with Induc ->
+ raise (CoFixGuardError (CodomainNotInductiveType b))
+ in
+ let (mind, _) = codomain_is_coind env deftype in
+ let (sp,tyi) = mind in
+ let lvlra = lookup_recargs env (sp,tyi) in
+ let vlra = lvlra.(tyi) in
+ let rec check_rec_call env alreadygrd n vlra t =
+ if noccur_with_meta n nbfix t then
+ true
+ else
+ let c,args = decompose_app (whd_betadeltaiota env t) in
+ match kind_of_term c with
+ | Meta _ -> true
+
+ | Rel p ->
+ if n <= p && p < n+nbfix then
+ (* recursive call *)
+ if alreadygrd then
+ if List.for_all (noccur_with_meta n nbfix) args then
+ true
+ else
+ raise (CoFixGuardError NestedRecursiveOccurrences)
+ else
+ raise (CoFixGuardError (UnguardedRecursiveCall t))
+ else
+ error "check_guard_rec_meta: ???" (* ??? *)
+
+ | Construct (_,i as cstr_sp) ->
+ let lra =vlra.(i-1) in
+ let mI = inductive_of_constructor cstr_sp in
+ let (mib,mip) = lookup_mind_specif env mI in
+ let _,realargs = list_chop mip.mind_nparams args in
+ let rec process_args_of_constr l lra =
+ match l with
+ | [] -> true
+ | t::lr ->
+ (match lra with
+ | [] -> anomaly_ill_typed ()
+ | (Mrec i)::lrar ->
+ let newvlra = lvlra.(i) in
+ (check_rec_call env true n newvlra t) &&
+ (process_args_of_constr lr lrar)
+
+ | (Imbr((sp,i) as ind_sp,lrc)::lrar) ->
+ let sprecargs = lookup_recargs env ind_sp in
+ let lc = (Array.map
+ (List.map
+ (instantiate_recarg sp lrc))
+ sprecargs.(i))
+ in (check_rec_call env true n lc t) &
+ (process_args_of_constr lr lrar)
+
+ | _::lrar ->
+ if (noccur_with_meta n nbfix t)
+ then (process_args_of_constr lr lrar)
+ else raise (CoFixGuardError
+ (RecCallInNonRecArgOfConstructor t)))
+ in (process_args_of_constr realargs lra)
+
+
+ | Lambda (x,a,b) ->
+ assert (args = []);
+ if (noccur_with_meta n nbfix a) then
+ check_rec_call (push_rel (x, None, a) env)
+ alreadygrd (n+1) vlra b
+ else
+ raise (CoFixGuardError (RecCallInTypeOfAbstraction t))
+
+ | CoFix (j,(_,varit,vdefs as recdef)) ->
+ if (List.for_all (noccur_with_meta n nbfix) args)
+ then
+ let nbfix = Array.length vdefs in
+ if (array_for_all (noccur_with_meta n nbfix) varit) then
+ let env' = push_rec_types recdef env in
+ (array_for_all
+ (check_rec_call env' alreadygrd (n+1) vlra) vdefs)
+ &&
+ (List.for_all (check_rec_call env alreadygrd (n+1) vlra) args)
+ else
+ raise (CoFixGuardError (RecCallInTypeOfDef c))
+ else
+ raise (CoFixGuardError (UnguardedRecursiveCall c))
+
+ | Case (_,p,tm,vrest) ->
+ if (noccur_with_meta n nbfix p) then
+ if (noccur_with_meta n nbfix tm) then
+ if (List.for_all (noccur_with_meta n nbfix) args) then
+ (array_for_all (check_rec_call env alreadygrd n vlra) vrest)
+ else
+ raise (CoFixGuardError (RecCallInCaseFun c))
+ else
+ raise (CoFixGuardError (RecCallInCaseArg c))
+ else
+ raise (CoFixGuardError (RecCallInCasePred c))
+
+ | _ -> raise (CoFixGuardError NotGuardedForm)
+
+ in
+ check_rec_call env false 1 vlra def
+
+(* The function which checks that the whole block of definitions
+ satisfies the guarded condition *)
+
+let check_cofix env (bodynum,(names,types,bodies as recdef)) =
+ let nbfix = Array.length bodies in
+ for i = 0 to nbfix-1 do
+ let fixenv = push_rec_types recdef env in
+ try
+ let _ = check_guard_rec_meta fixenv nbfix bodies.(i) types.(i)
+ in ()
+ with CoFixGuardError err ->
+ error_ill_formed_rec_body fixenv err names i bodies
+ done
diff --git a/kernel/inductive.mli b/kernel/inductive.mli
index 2aee7f420..dbaf36788 100644
--- a/kernel/inductive.mli
+++ b/kernel/inductive.mli
@@ -12,216 +12,63 @@
open Names
open Univ
open Term
-open Sign
open Declarations
open Environ
-open Evd
(*i*)
-(*s Inductives are accessible at several stages:
-
-A [mutual_inductive_body] contains all information about a
-declaration of mutual (co-)inductive types. These informations are
-closed (they depend on no free variables) and an instance of them
-corresponds to a [mutual_inductive_instance =
-mutual_inductive_body * constr list]. One inductive type in an
-instanciated packet corresponds to an [inductive_instance =
-mutual_inductive_instance * int]. Applying global parameters to an
-[inductive_instance] gives an [inductive_family = inductive_instance *
-constr list]. Finally, applying real parameters gives an
-[inductive_type = inductive_family * constr list]. At each level
-corresponds various appropriated functions *)
-
-type inductive_instance (* ex-[mind_specif] *)
-
-val build_mis : inductive -> mutual_inductive_body -> inductive_instance
-
-val mis_index : inductive_instance -> int
-val mis_ntypes : inductive_instance -> int
-val mis_nconstr : inductive_instance -> int
-val mis_nparams : inductive_instance -> int
-val mis_nrealargs : inductive_instance -> int
-val mis_kelim : inductive_instance -> sorts_family list
-val mis_recargs : inductive_instance -> (recarg list) array array
-val mis_recarg : inductive_instance -> (recarg list) array
-val mis_typename : inductive_instance -> identifier
-val mis_typepath : inductive_instance -> section_path
-val mis_is_recursive_subset : int list -> inductive_instance -> bool
-val mis_is_recursive : inductive_instance -> bool
-val mis_consnames : inductive_instance -> identifier array
-val mis_conspaths : inductive_instance -> section_path array
-val mis_inductive : inductive_instance -> inductive
-val mis_arity : inductive_instance -> types
-val mis_nf_arity : inductive_instance -> types
-val mis_params_ctxt : inductive_instance -> rel_context
-val mis_sort : inductive_instance -> sorts
-val mis_constructor_type : int -> inductive_instance -> types
-val mis_finite : inductive_instance -> bool
-
-(* The ccl of constructor is pre-normalised in the following functions *)
-val mis_nf_lc : inductive_instance -> constr array
-
-(*s [inductive_family] = [inductive_instance] applied to global parameters *)
-type inductive_family = IndFamily of inductive_instance * constr list
-
-val make_ind_family : inductive_instance * constr list -> inductive_family
-val dest_ind_family : inductive_family -> inductive_instance * constr list
-
-val liftn_inductive_family :
- int -> int -> inductive_family -> inductive_family
-val lift_inductive_family :
- int -> inductive_family -> inductive_family
-
-(*s [inductive_type] = [inductive_family] applied to ``real'' parameters *)
-type inductive_type = IndType of inductive_family * constr list
-
-val make_ind_type : inductive_family * constr list -> inductive_type
-val dest_ind_type : inductive_type -> inductive_family * constr list
-
-val mkAppliedInd : inductive_type -> constr
-
-val liftn_inductive_type : int -> int -> inductive_type -> inductive_type
-val lift_inductive_type : int -> inductive_type -> inductive_type
-val substnl_ind_type : constr list -> int -> inductive_type -> inductive_type
-
-(*s A [constructor] is an [inductive] + an index; the following functions
- destructs and builds [constructor] *)
-val inductive_of_constructor : constructor -> inductive
-val index_of_constructor : constructor -> int
-val ith_constructor_of_inductive : inductive -> int -> constructor
-
-(*s This type gathers useful informations about some instance of a constructor
- relatively to some implicit context (the current one)
-
- If [cs_cstr] is a constructor in [(I p1...pm a1...an)] then
- [cs_params] is [p1...pm] and the type of [MutConstruct(cs_cstr)
- p1...pn] is [(cs_args)(I p1...pm cs_concl_realargs)] where [cs_args]
- and [cs_params] are relative to the current env and [cs_concl_realargs]
- is relative to the current env enriched by [cs_args]
-*)
-
-type constructor_summary = {
- cs_cstr : constructor;
- cs_params : constr list;
- cs_nargs : int;
- cs_args : rel_context;
- cs_concl_realargs : constr array
-}
-
-val lift_constructor : int -> constructor_summary -> constructor_summary
+exception Induc
-(*s Functions to build standard types related to inductive *)
+(*s Extracting an inductive type from a constructions *)
-(* This builds [(ci params (Rel 1)...(Rel ci_nargs))] which is the argument
- of a dependent predicate in a Cases branch *)
-val build_dependent_constructor : constructor_summary -> constr
+(* [find_m*type env sigma c] coerce [c] to an recursive type (I args).
+ [find_rectype], [find_inductive] and [find_coinductive]
+ respectively accepts any recursive type, only an inductive type and
+ only a coinductive type.
+ They raise [Induc] if not convertible to a recursive type. *)
-(* This builds [(I params (Rel 1)...(Rel nrealargs))] which is the type of
- the constructor argument of a dependent predicate in a cases branch *)
-val build_dependent_inductive : inductive_family -> constr
+val find_rectype : env -> constr -> inductive * constr list
+val find_inductive : env -> constr -> inductive * constr list
+val find_coinductive : env -> constr -> inductive * constr list
-(* if the arity for some inductive family [indf] associated to [(I
- params)] is [(x1:A1)...(xn:An)sort'] then [make_arity env sigma dep
- indf k] builds [(x1:A1)...(xn:An)sort] which is the arity of an
- elimination predicate on sort [k]; if [dep=true] then it rather
- builds [(x1:A1)...(xn:An)(I params x1...xn)->sort] *)
-val make_arity : env -> bool -> inductive_family -> sorts -> constr
+(*s Fetching information in the environment about an inductive type.
+ Raises Induc if the inductive type is not found. *)
+val lookup_mind_specif :
+ env -> inductive -> mutual_inductive_body * one_inductive_body
-(* [build_branch_type env dep p cs] builds the type of the branch
- associated to constructor [cs] in a Case with elimination predicate
- [p]; if [dep=true], the predicate is assumed dependent *)
-val build_branch_type : env -> bool -> constr -> constructor_summary -> constr
+(*s Functions to build standard types related to inductive *)
+val type_of_inductive : env -> inductive -> types
-(*s Extracting an inductive type from a constructions *)
+(* Return type as quoted by the user *)
+val type_of_constructor : env -> constructor -> types
-exception Induc
+(* Return constructor types in normal form *)
+val arities_of_constructors : env -> inductive -> types array
-(* [extract_mrectype c] assumes [c] is syntactically an inductive type
- applied to arguments then it returns its components; if not an
- inductive type, it raises [Induc] *)
-val extract_mrectype : constr -> inductive * constr list
-(* [find_m*type env sigma c] coerce [c] to an recursive type (I args).
- [find_rectype], [find_inductive] and [find_coinductive]
- respectively accepts any recursive type, only an inductive type and
- only a coinductive type.
- They raise [Induc] if not convertible to a recursive type. *)
+exception Arity of (constr * constr * string) option
+
+(* [type_case_branches env (I,args) (p:A) c] computes useful types
+ about the following Cases expression:
+ <p>Cases (c :: (I args)) of b1..bn end
+ It computes the type of every branch (pattern variables are
+ introduced by products), the type for the whole expression, and
+ the universe constraints generated.
+ *)
+val type_case_branches :
+ env -> inductive * constr list -> unsafe_judgment -> constr
+ -> types array * types * constraints
+
+(* Check a case_info actually correspond to a Case expression on the
+ given inductive type. *)
+val check_case_info : env -> inductive -> case_info -> unit
+
+(*s Guard conditions for fix and cofix-points. *)
+val check_fix : env -> fixpoint -> unit
+val check_cofix : env -> cofixpoint -> unit
-val find_mrectype : env -> 'a evar_map -> constr -> inductive * constr list
-val find_inductive : env -> 'a evar_map -> constr -> inductive * constr list
-val find_coinductive : env -> 'a evar_map -> constr -> inductive * constr list
-
-val lookup_mind_specif : inductive -> env -> inductive_instance
-
-(* [find_rectype env sigma t] builds an [inductive_type] or raises
- [Induc] if [t] is not a (co-)inductive type; The result is relative to
- [env] and [sigma] *)
-
-val find_rectype : env -> 'a evar_map -> constr -> inductive_type
-
-(* [get_constructors_types indf] returns the array of the types of
- constructors of the inductive\_family [indf], i.e. the types are
- instantiated by the parameters of the family (the type may be not
- in canonical form -- e.g. cf sets library) *)
-
-val get_constructors_types : inductive_family -> types array
-val get_constructor_type : inductive_family -> int -> types
-
-(* [get_constructors indf] build an array of [constructor_summary]
- from some inductive type already analysed as an [inductive_family];
- global parameters are already instanciated in the constructor
- types; the resulting summaries are valid in the environment where
- [indf] is valid; the names of the products of the constructors types
- are not renamed when [Anonymous] *)
-
-val get_constructors : inductive_family -> constructor_summary array
-val get_constructor : inductive_family -> int -> constructor_summary
-
-(* [get_arity_type indf] returns the type of the arity of the
- inductive family described by [indf]; global parameters are already
- instanciated (but the type may be not in canonical form -- e.g. cf
- sets library); the products signature is relative to the
- environment definition of [indf]; the names of the products of the
- constructors types are not renamed when [Anonymous]; [get_arity
- indf] does the same but normalises and decomposes it as an arity *)
-
-val get_arity_type : inductive_family -> types
-val get_arity : inductive_family -> arity
-
-(* [get_arity_type indf] returns the type of the arity of the inductive
- family described by [indf]; global parameters are already instanciated *)
-
-
-
-(* Examples: assume
-
-\begin{verbatim}
-Inductive listn [A:Set] : nat -> Set :=
- niln : (listn A O)
-| consn : (n:nat)A->(listn A n)->(listn A (S n)).
-\end{verbatim}
-
-has been defined. Then in some env containing ['x:nat'],
-\begin{quote}
-[find_rectype env sigma (listn bool (S x))] returns [IndType (indf, '(S x)')]
-\end{quote}
-where [indf = IndFamily ('listn',['bool'])].
-
-Then, [get_constructors indf] returns
-\begin{quote}
-[[| { cs_cstr = 'niln'; cs_params = 'bool'; cs_nargs = 0;
- cs_args = []; cs_concl_realargs = [|O|]};
- { cs_cstr = 'consn'; cs_params = 'bool'; cs_nargs = 3;
- cs_args = [(Anonymous,'(listn A n)'),(Anonymous,'A'),(Name n,'nat')];
- cs_concl_realargs = [|'(S n)'|]} |]]
-\end{quote}
-and [get_arity indf] returns [[(Anonymous,'nat')],'Set'].
-\smallskip
-*)
-
-(*s [Cases] info *)
-
-val make_case_info : inductive_instance -> case_style option
- -> pattern_source array -> case_info
-val make_default_case_info : inductive_instance -> case_info
+(********************)
+(* TODO: remove (used in pretyping only...) *)
+val find_case_dep_nparams :
+ env -> constr * unsafe_judgment -> inductive * constr list ->
+ bool * constraints
diff --git a/kernel/instantiate.ml b/kernel/instantiate.ml
deleted file mode 100644
index 0191b6391..000000000
--- a/kernel/instantiate.ml
+++ /dev/null
@@ -1,147 +0,0 @@
-(***********************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
-(* \VV/ *************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(***********************************************************************)
-
-(* $Id$ *)
-
-open Pp
-open Util
-open Names
-open Term
-open Sign
-open Evd
-open Declarations
-open Environ
-
-let is_id_inst inst =
- let is_id (id,c) = match kind_of_term c with
- | IsVar id' -> id = id'
- | _ -> false
- in
- List.for_all is_id inst
-
-let instantiate sign c args =
- let inst = instantiate_named_context sign args in
- if is_id_inst inst then
- c
- else
- replace_vars inst c
-
-(* Vérifier que les instances des let-in sont compatibles ?? *)
-let instantiate_sign_including_let sign args =
- let rec instrec = function
- | ((id,b,_) :: sign, c::args) -> (id,c) :: (instrec (sign,args))
- | ([],[]) -> []
- | ([],_) | (_,[]) ->
- anomaly "Signature and its instance do not match"
- in
- instrec (sign,args)
-
-let instantiate_evar sign c args =
- let inst = instantiate_sign_including_let sign args in
- if is_id_inst inst then
- c
- else
- replace_vars inst c
-
-let instantiate_constr sign c args =
- let sign = List.map (fun (sp,b,t) -> (basename sp,b,t)) sign in
- instantiate sign c args
-
-let instantiate_type sign tty args =
- type_app (fun c -> instantiate_constr sign c args) tty
-
-(* Constants. *)
-
-(* constant_type gives the type of a constant *)
-let constant_type env sigma sp =
- let cb = lookup_constant sp env in
- cb.const_type
-
-type const_evaluation_result = NoBody | Opaque
-
-exception NotEvaluableConst of const_evaluation_result
-
-let constant_value env sp =
- let cb = lookup_constant sp env in
- if cb.const_opaque then raise (NotEvaluableConst Opaque);
- match cb.const_body with
- | Some body -> body
- | None -> raise (NotEvaluableConst NoBody)
-
-let constant_opt_value env cst =
- try Some (constant_value env cst)
- with NotEvaluableConst _ -> None
-
-(* Existentials. *)
-
-let name_of_existential n = id_of_string ("?" ^ string_of_int n)
-
-let existential_type sigma (n,args) =
- let info = Evd.map sigma n in
- let hyps = info.evar_hyps in
- (* TODO: check args [this comment was in Typeops] *)
- instantiate_evar hyps info.evar_concl (Array.to_list args)
-
-exception NotInstantiatedEvar
-
-let existential_value sigma (n,args) =
- let info = Evd.map sigma n in
- let hyps = info.evar_hyps in
- match evar_body info with
- | Evar_defined c ->
- instantiate_evar hyps c (Array.to_list args)
- | Evar_empty ->
- raise NotInstantiatedEvar
-
-let existential_opt_value sigma ev =
- try Some (existential_value sigma ev)
- with NotInstantiatedEvar -> None
-
-
-type evaluable_reference =
- | EvalConst of constant
- | EvalVar of identifier
- | EvalRel of int
- | EvalEvar of existential
-
-let mkEvalRef = function
- | EvalConst cst -> mkConst cst
- | EvalVar id -> mkVar id
- | EvalRel n -> mkRel n
- | EvalEvar ev -> mkEvar ev
-
-let isEvalRef c = match kind_of_term c with
- | IsConst _ | IsVar _ | IsRel _ | IsEvar _ -> true
- | _ -> false
-
-let destEvalRef c = match kind_of_term c with
- | IsConst cst -> EvalConst cst
- | IsVar id -> EvalVar id
- | IsRel n -> EvalRel n
- | IsEvar ev -> EvalEvar ev
- | _ -> anomaly "Not an evaluable reference"
-
-let evaluable_reference sigma env = function
- | EvalConst sp -> evaluable_constant env sp
- | EvalVar id -> evaluable_named_decl env id
- | EvalRel n -> evaluable_rel_decl env n
- | EvalEvar (ev,_) -> Evd.is_defined sigma ev
-
-let reference_opt_value sigma env = function
- | EvalConst cst -> constant_opt_value env cst
- | EvalVar id -> lookup_named_value id env
- | EvalRel n -> lookup_rel_value n env
- | EvalEvar ev -> existential_opt_value sigma ev
-
-exception NotEvaluable
-let reference_value sigma env c =
- match reference_opt_value sigma env c with
- | None -> raise NotEvaluable
- | Some d -> d
-
-
diff --git a/kernel/instantiate.mli b/kernel/instantiate.mli
deleted file mode 100644
index 14a4746ee..000000000
--- a/kernel/instantiate.mli
+++ /dev/null
@@ -1,63 +0,0 @@
-(***********************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
-(* \VV/ *************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(***********************************************************************)
-
-(*i $Id$ i*)
-
-(*i*)
-open Names
-open Term
-open Evd
-open Sign
-open Environ
-(*i*)
-
-(* Instantiation of constants and inductives on their real arguments. *)
-
-val instantiate_constr :
- section_context -> constr -> constr list -> constr
-
-val instantiate_type :
- section_context -> types -> constr list -> types
-
-(*s [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] *)
-
-type const_evaluation_result = NoBody | Opaque
-exception NotEvaluableConst of const_evaluation_result
-
-val constant_value : env -> constant -> constr
-val constant_type : env -> 'a evar_map -> constant -> types
-val constant_opt_value : env -> constant -> constr option
-
-(*s [existential_value sigma ev] raises [NotInstantiatedEvar] if [ev] has
-no body and [Not_found] if it does not exist in [sigma] *)
-
-exception NotInstantiatedEvar
-val existential_value : 'a evar_map -> existential -> constr
-val existential_type : 'a evar_map -> existential -> constr
-val existential_opt_value : 'a evar_map -> existential -> constr option
-
-type evaluable_reference =
- | EvalConst of constant
- | EvalVar of identifier
- | EvalRel of int
- | EvalEvar of existential
-
-val destEvalRef : constr -> evaluable_reference
-val mkEvalRef : evaluable_reference -> constr
-val isEvalRef : constr -> bool
-
-val evaluable_reference : 'a evar_map -> env -> evaluable_reference -> bool
-
-val reference_opt_value :
- 'a evar_map -> env -> evaluable_reference -> constr option
-
-(* This may raise NotEvaluable *)
-exception NotEvaluable
-val reference_value : 'a evar_map -> env -> evaluable_reference -> constr
diff --git a/kernel/names.ml b/kernel/names.ml
index f2fe3be86..b91c6b08c 100644
--- a/kernel/names.ml
+++ b/kernel/names.ml
@@ -13,119 +13,12 @@ open Util
(*s Identifiers *)
-(* Utilities *)
-
-let code_of_0 = Char.code '0'
-let code_of_9 = Char.code '9'
-
-(* Identifiers *)
-
type identifier = string
-let cut_ident s =
- let slen = String.length s in
- (* [n'] is the position of the first non nullary digit *)
- let rec numpart n n' =
- if n = 0 then
- failwith
- ("The string " ^ s ^ " is not an identifier: it contains only digits")
- else
- let c = Char.code (String.get s (n-1)) in
- if c = code_of_0 && n <> slen then
- numpart (n-1) n'
- else if code_of_0 <= c && c <= code_of_9 then
- numpart (n-1) (n-1)
- else
- n'
- in
- numpart slen slen
-
-let repr_ident s =
- let slen = String.length s in
- let numstart = cut_ident s in
- if numstart = slen then
- (s, None)
- else
- (String.sub s 0 numstart,
- Some (int_of_string (String.sub s numstart (slen - numstart))))
-
-let make_ident sa = function
- | Some n ->
- let c = Char.code (String.get sa (String.length sa -1)) in
- if c < code_of_0 or c > code_of_9 then sa ^ (string_of_int n)
- else sa ^ "_" ^ (string_of_int n)
- | None -> String.copy sa
-
-let first_char id =
- assert (id <> "");
- String.make 1 id.[0]
-
let id_ord = Pervasives.compare
-(* Rem: semantics is a bit different, if an ident starts with toto00 then
- after successive renamings it comes to toto09, then it goes on with toto10 *)
-let lift_subscript id =
- let len = String.length id in
- let rec add carrypos =
- let c = id.[carrypos] in
- if is_digit c then
- if c = '9' then begin
- assert (carrypos>0);
- add (carrypos-1)
- end
- else begin
- let newid = String.copy id in
- String.fill newid (carrypos+1) (len-1-carrypos) '0';
- newid.[carrypos] <- Char.chr (Char.code c + 1);
- newid
- end
- else begin
- let newid = id^"0" in
- if carrypos < len-1 then begin
- String.fill newid (carrypos+1) (len-1-carrypos) '0';
- newid.[carrypos+1] <- '1'
- end;
- newid
- end
- in add (len-1)
-
-let has_subscript id = is_digit (id.[String.length id - 1])
-
-let forget_subscript id =
- let len = String.length id in
- let numstart = cut_ident id in
- let newid = String.make (numstart+1) '0' in
- String.blit id 0 newid 0 numstart;
- newid
-
-(* This checks that a string is acceptable as an ident, i.e. starts
- with a letter and contains only letters, digits or "'" *)
-
-let check_ident_suffix i l s =
- for i=1 to l-1 do
- let c = String.get s i in
- if not (is_letter c or is_digit c or c = '\'' or c = '_' or c = '@') then
- error
- ("Character "^(String.sub s i 1)^" is not allowed in identifier "^s)
- done
-
-let check_ident s =
- let l = String.length s in
- if l = 0 then error "The empty string is not an identifier";
- let c = String.get s 0 in
- if (is_letter c) or c = '_' or c = '$' or c = '?'
- then check_ident_suffix 1 l s
- else error (s^": an identifier should start with a letter")
-
-let is_ident s = try check_ident s; true with _ -> false
-
-let check_suffix s = check_ident_suffix 0 (String.length s) s
-
-let add_suffix id s = check_suffix s; id^s
-let add_prefix s id = check_ident s; s^id
-
let string_of_id id = String.copy id
-let id_of_string s = check_ident s; String.copy s
+let id_of_string s = String.copy s
(* Hash-consing of identifier *)
module Hident = Hashcons.Make(
@@ -147,65 +40,14 @@ module Idset = Set.Make(IdOrdered)
module Idmap = Map.Make(IdOrdered)
module Idpred = Predicate.Make(IdOrdered)
-let atompart_of_id id = fst (repr_ident id)
-let index_of_id id = snd (repr_ident id)
let pr_id id = [< 'sTR (string_of_id id) >]
let wildcard = id_of_string "_"
-(* Fresh names *)
-
-let lift_ident = lift_subscript
-
-let next_ident_away id avoid =
- if List.mem id avoid then
- let id0 = if not (has_subscript id) then id else
- (* Ce serait sans doute mieux avec quelque chose inspiré de
- *** make_ident id (Some 0) *** mais ça brise la compatibilité... *)
- forget_subscript id in
- let rec name_rec id =
- if List.mem id avoid then name_rec (lift_ident id) else id in
- name_rec id0
- else id
-
-let next_ident_away_from id avoid =
- let rec name_rec id =
- if List.mem id avoid then name_rec (lift_ident id) else id in
- name_rec id
-
(* Names *)
type name = Name of identifier | Anonymous
-let next_name_away_with_default default name l =
- match name with
- | Name str -> next_ident_away str l
- | Anonymous -> next_ident_away (id_of_string default) l
-
-let next_name_away name l =
- match name with
- | Name str -> next_ident_away str l
- | Anonymous -> id_of_string "_"
-
-let out_name = function
- | Name id -> id
- | Anonymous -> anomaly "out_name: expects a defined name"
-
-(* Kinds *)
-
-type path_kind = CCI | FW | OBJ
-
-let string_of_kind = function
- | CCI -> "cci"
- | FW -> "fw"
- | OBJ -> "obj"
-
-let kind_of_string = function
- | "cci" -> CCI
- | "fw" -> FW
- | "obj" -> OBJ
- | _ -> invalid_arg "kind_of_string"
-
(*s Directory paths = section names paths *)
let parse_fields s =
let len = String.length s in
@@ -234,81 +76,38 @@ module ModIdOrdered =
module ModIdmap = Map.Make(ModIdOrdered)
-(* These are the only functions which depend on how a dirpath is encoded *)
-let make_dirpath x = List.rev x
-let repr_dirpath x = List.rev x
-let rev_repr_dirpath x = x
-
-let dirpath_prefix = function
- | [] -> anomaly "dirpath_prefix: empty dirpath"
- | _::l -> l
-
-let split_dirpath = function
- | [] -> failwith "Empty"
- | d::b -> (b,d)
-
-let extend_dirpath d id = id::d
-let add_dirpath_prefix id d = d@[id]
-
-let is_dirpath_prefix_of d1 d2 = list_prefix_of (List.rev d1) (List.rev d2)
-(**)
-
-let is_empty_dirpath d = (d = [])
-
-let dirpath_of_string s =
- try
- let sl,s = parse_fields s in
- make_dirpath (sl @ [s])
- with
- | Invalid_argument _ -> invalid_arg "dirpath_of_string"
+let make_dirpath x = x
+let repr_dirpath x = x
let string_of_dirpath = function
| [] -> "<empty>"
- | sl -> String.concat "." (List.map string_of_id (repr_dirpath sl))
+ | sl ->
+ String.concat "." (List.map string_of_id (List.rev sl))
let pr_dirpath sl = [< 'sTR (string_of_dirpath sl) >]
-let default_module_name = id_of_string "Top"
-let default_module = make_dirpath [default_module_name]
-
(*s Section paths are absolute names *)
type section_path = {
dirpath : dir_path ;
- basename : identifier ;
- kind : path_kind }
-
-let make_path pa id k = { dirpath = pa; basename = id; kind = k }
-let repr_path { dirpath = pa; basename = id; kind = k} = (pa,id,k)
+ basename : identifier }
-let kind_of_path sp = sp.kind
-let basename sp = sp.basename
-let dirpath sp = sp.dirpath
+let make_path pa id = { dirpath = pa; basename = id }
+let repr_path { dirpath = pa; basename = id } = (pa,id)
(* parsing and printing of section paths *)
let string_of_path sp =
- let (sl,id,k) = repr_path sp in
+ let (sl,id) = repr_path sp in
if sl = [] then string_of_id id
else (string_of_dirpath sl) ^ "." ^ (string_of_id id)
-let path_of_string s =
- try
- let sl,s = parse_fields s in
- make_path (make_dirpath sl) s CCI
- with
- | Invalid_argument _ -> invalid_arg "path_of_string"
-
let pr_sp sp = [< 'sTR (string_of_path sp) >]
let sp_ord sp1 sp2 =
- let (p1,id1,k) = repr_path sp1
- and (p2,id2,k') = repr_path sp2 in
- let ck = compare k k' in
- if ck = 0 then
- let p_bit = compare p1 p2 in
- if p_bit = 0 then id_ord id1 id2 else p_bit
- else
- ck
+ let (p1,id1) = repr_path sp1
+ and (p2,id2) = repr_path sp2 in
+ let p_bit = compare p1 p2 in
+ if p_bit = 0 then id_ord id1 id2 else p_bit
module SpOrdered =
struct
@@ -323,17 +122,16 @@ module Spmap = Map.Make(SpOrdered)
(*s********************************************************************)
(* type of global reference *)
-type variable = section_path
+type variable = identifier
type constant = section_path
type inductive = section_path * int
type constructor = inductive * int
type mutual_inductive = section_path
-type global_reference =
- | VarRef of section_path
- | ConstRef of constant
- | IndRef of inductive
- | ConstructRef of constructor
+let ith_mutual_inductive (sp,_) i = (sp,i)
+let ith_constructor_of_inductive ind_sp i = (ind_sp,i)
+let inductive_of_constructor (ind_sp,i) = ind_sp
+let index_of_constructor (ind_sp,i) = i
(* Hash-consing of name objects *)
module Hname = Hashcons.Make(
@@ -366,12 +164,10 @@ module Hsp = Hashcons.Make(
type u = identifier -> identifier
let hash_sub hident sp =
{ dirpath = List.map hident sp.dirpath;
- basename = hident sp.basename;
- kind = sp.kind }
+ basename = hident sp.basename }
let equal sp1 sp2 =
- (sp1.basename == sp2.basename) && (sp1.kind = sp2.kind)
- && (List.length sp1.dirpath = List.length sp2.dirpath)
- && List.for_all2 (==) sp1.dirpath sp2.dirpath
+ (List.length sp1.dirpath = List.length sp2.dirpath) &&
+ (List.for_all2 (==) sp1.dirpath sp2.dirpath)
let hash = Hashtbl.hash
end)
diff --git a/kernel/names.mli b/kernel/names.mli
index 478b1c8e4..3aac8c40b 100644
--- a/kernel/names.mli
+++ b/kernel/names.mli
@@ -16,99 +16,43 @@ open Pp
type identifier
type name = Name of identifier | Anonymous
-
-(* Constructor of an identifier;
- [make_ident] builds an identifier from a string and an optional index; if
- the string ends by a digit, a ["_"] is inserted *)
-val make_ident : string -> int option -> identifier
-
-(* Some destructors of an identifier *)
-val atompart_of_id : identifier -> string
-val first_char : identifier -> string
-val index_of_id : identifier -> int option
-
(* Parsing and printing of identifiers *)
val string_of_id : identifier -> string
val id_of_string : string -> identifier
val pr_id : identifier -> std_ppcmds
-(* This is the identifier ["_"] *)
-val wildcard : identifier
-
-(* Deriving ident from other idents *)
-val add_suffix : identifier -> string -> identifier
-val add_prefix : string -> identifier -> identifier
-
(* Identifiers sets and maps *)
module Idset : Set.S with type elt = identifier
module Idpred : Predicate.S with type elt = identifier
module Idmap : Map.S with type key = identifier
-val lift_ident : identifier -> identifier
-val next_ident_away_from : identifier -> identifier list -> identifier
-val next_ident_away : identifier -> identifier list -> identifier
-val next_name_away : name -> identifier list -> identifier
-val next_name_away_with_default :
- string -> name -> identifier list -> identifier
-
-(* [out_name na] raises an anomaly if [na] is [Anonymous] *)
-val out_name : name -> identifier
-
-(*s [path_kind] is currently degenerated, [FW] is not used *)
-type path_kind = CCI | FW | OBJ
-
-(* parsing and printing of path kinds *)
-val string_of_kind : path_kind -> string
-val kind_of_string : string -> path_kind
-
(*s Directory paths = section names paths *)
type module_ident = identifier
-type dir_path (*= module_ident list*)
+type dir_path
module ModIdmap : Map.S with type key = module_ident
+(* Inner modules idents on top of list *)
val make_dirpath : module_ident list -> dir_path
val repr_dirpath : dir_path -> module_ident list
-val rev_repr_dirpath : dir_path -> module_ident list
-val is_empty_dirpath : dir_path -> bool
-
-(* Give the immediate prefix of a [dir_path] *)
-val dirpath_prefix : dir_path -> dir_path
-
-(* Give the immediate prefix and basename of a [dir_path] *)
-val split_dirpath : dir_path -> dir_path * identifier
-
-val extend_dirpath : dir_path -> module_ident -> dir_path
-val add_dirpath_prefix : module_ident -> dir_path -> dir_path
(* Printing of directory paths as ["coq_root.module.submodule"] *)
val string_of_dirpath : dir_path -> string
val pr_dirpath : dir_path -> std_ppcmds
-val default_module : dir_path
(*s Section paths are {\em absolute} names *)
type section_path
(* Constructors of [section_path] *)
-val make_path : dir_path -> identifier -> path_kind -> section_path
+val make_path : dir_path -> identifier -> section_path
(* Destructors of [section_path] *)
-val repr_path : section_path -> dir_path * identifier * path_kind
-val dirpath : section_path -> dir_path
-val basename : section_path -> identifier
-val kind_of_path : section_path -> path_kind
+val repr_path : section_path -> dir_path * identifier
(* Parsing and printing of section path as ["coq_root.module.id"] *)
-val path_of_string : string -> section_path
val string_of_path : section_path -> string
val pr_sp : section_path -> std_ppcmds
-val dirpath_of_string : string -> dir_path
-
-val sp_ord : section_path -> section_path -> int
-
-(* [is_dirpath_prefix p1 p2=true] if [p1] is a prefix of or is equal to [p2] *)
-val is_dirpath_prefix_of : dir_path -> dir_path -> bool
module Spset : Set.S with type elt = section_path
module Sppred : Predicate.S with type elt = section_path
@@ -117,17 +61,19 @@ module Spmap : Map.S with type key = section_path
(*s********************************************************************)
(* type of global reference *)
-type variable = section_path
+type variable = identifier
type constant = section_path
+(* Beware: first inductive has index 0 *)
type inductive = section_path * int
+(* Beware: first constructor has index 1 *)
type constructor = inductive * int
type mutual_inductive = section_path
-type global_reference =
- | VarRef of section_path
- | ConstRef of constant
- | IndRef of inductive
- | ConstructRef of constructor
+val ith_mutual_inductive : inductive -> int -> inductive
+
+val ith_constructor_of_inductive : inductive -> int -> constructor
+val inductive_of_constructor : constructor -> inductive
+val index_of_constructor : constructor -> int
(* Hash-consing *)
val hcons_names : unit ->
diff --git a/kernel/reduction.ml b/kernel/reduction.ml
index 734187a9c..10ce90291 100644
--- a/kernel/reduction.ml
+++ b/kernel/reduction.ml
@@ -8,445 +8,64 @@
(* $Id$ *)
-open Pp
open Util
open Names
open Term
open Univ
-open Evd
open Declarations
open Environ
-open Instantiate
open Closure
open Esubst
-exception Elimconst
-
-(* The type of (machine) states (= lambda-bar-calculus' cuts) *)
-type state = constr * constr stack
-
-type 'a contextual_reduction_function = env -> 'a evar_map -> constr -> constr
-type 'a reduction_function = 'a contextual_reduction_function
-type local_reduction_function = constr -> constr
-
-type 'a contextual_stack_reduction_function =
- env -> 'a evar_map -> constr -> constr * constr list
-type 'a stack_reduction_function = 'a contextual_stack_reduction_function
-type local_stack_reduction_function = constr -> constr * constr list
-
-type 'a contextual_state_reduction_function =
- env -> 'a evar_map -> state -> state
-type 'a state_reduction_function = 'a contextual_state_reduction_function
-type local_state_reduction_function = state -> state
-
-(*************************************)
-(*** Reduction Functions Operators ***)
-(*************************************)
-
-let rec whd_state (x, stack as s) =
- match kind_of_term x with
- | IsApp (f,cl) -> whd_state (f, append_stack cl stack)
- | IsCast (c,_) -> whd_state (c, stack)
- | _ -> s
+(****************************************************************************)
+(* Reduction Functions *)
+(****************************************************************************)
-let appterm_of_stack (f,s) = (f,list_of_stack s)
+let nf_betaiota t =
+ norm_val (create_clos_infos betaiota empty_env) (inject t)
-let whd_stack x = appterm_of_stack (whd_state (x, empty_stack))
-let whd_castapp_stack = whd_stack
+let hnf_stack env x =
+ decompose_app
+ (norm_val (create_clos_infos hnf_flags env) (inject x))
-let stack_reduction_of_reduction red_fun env sigma s =
- let t = red_fun env sigma (app_stack s) in
- whd_stack t
+let whd_betadeltaiota env t =
+ whd_val (create_clos_infos betadeltaiota env) (inject t)
-let strong whdfun env sigma t =
- let rec strongrec env t =
- map_constr_with_full_binders push_rel strongrec env (whdfun env sigma t) in
- strongrec env t
+let whd_betadeltaiota_nolet env t =
+ whd_val (create_clos_infos betadeltaiotanolet env) (inject t)
-let local_strong whdfun =
- let rec strongrec t = map_constr strongrec (whdfun t) in
- strongrec
+(* Beta *)
-let rec strong_prodspine redfun c =
- let x = redfun c in
- match kind_of_term x with
- | IsProd (na,a,b) -> mkProd (na,a,strong_prodspine redfun b)
- | _ -> x
+let beta_appvect c v =
+ let rec stacklam env t stack =
+ match (decomp_stack stack,kind_of_term t) with
+ | Some (h,stacktl), Lambda (_,_,c) -> stacklam (h::env) c stacktl
+ | _ -> app_stack (substl env t, stack) in
+ stacklam [] c (append_stack v empty_stack)
-(****************************************************************************)
-(* Reduction Functions *)
-(****************************************************************************)
+(* pseudo-reduction rule:
+ * [hnf_prod_app env s (Prod(_,B)) N --> B[N]
+ * with an HNF on the first argument to produce a product.
+ * if this does not work, then we use the string S as part of our
+ * error message. *)
-(* lazy reduction functions. The infos must be created for each term *)
-let clos_norm_flags flgs env sigma t =
- norm_val (create_clos_infos flgs env sigma) (inject t)
-
-let nf_beta = clos_norm_flags beta empty_env Evd.empty
-let nf_betaiota = clos_norm_flags betaiota empty_env Evd.empty
-let nf_betadeltaiota env sigma = clos_norm_flags betadeltaiota env sigma
-
-(* lazy weak head reduction functions *)
-let whd_flags flgs env sigma t =
- whd_val (create_clos_infos flgs env sigma) (inject t)
-
-(*************************************)
-(*** Reduction using substitutions ***)
-(*************************************)
-
-(* This signature is very similar to Closure.RedFlagsSig except there
- is eta but no per-constant unfolding *)
-
-module type RedFlagsSig = sig
- type flags
- type flag
- val fbeta : flag
- val fevar : flag
- val fdelta : flag
- val feta : flag
- val fiota : flag
- val fzeta : flag
- val mkflags : flag list -> flags
- val red_beta : flags -> bool
- val red_delta : flags -> bool
- val red_evar : flags -> bool
- val red_eta : flags -> bool
- val red_iota : flags -> bool
- val red_zeta : flags -> bool
-end
-
-(* Naive Implementation
-module RedFlags = (struct
- type flag = BETA | DELTA | EVAR | IOTA | ZETA | ETA
- type flags = flag list
- let fbeta = BETA
- let fdelta = DELTA
- let fevar = EVAR
- let fiota = IOTA
- let fzeta = ZETA
- let feta = ETA
- let mkflags l = l
- let red_beta = List.mem BETA
- let red_delta = List.mem DELTA
- let red_evar = List.mem EVAR
- let red_eta = List.mem ETA
- let red_iota = List.mem IOTA
- let red_zeta = List.mem ZETA
-end : RedFlagsSig)
-*)
+let hnf_prod_app env t n =
+ match kind_of_term (whd_betadeltaiota env t) with
+ | Prod (_,_,b) -> subst1 n b
+ | _ -> anomaly "hnf_prod_app: Need a product"
-(* Compact Implementation *)
-module RedFlags = (struct
- type flag = int
- type flags = int
- let fbeta = 1
- let fdelta = 2
- let fevar = 4
- let feta = 8
- let fiota = 16
- let fzeta = 32
- let mkflags = List.fold_left (lor) 0
- let red_beta f = f land fbeta <> 0
- let red_delta f = f land fdelta <> 0
- let red_evar f = f land fevar <> 0
- let red_eta f = f land feta <> 0
- let red_iota f = f land fiota <> 0
- let red_zeta f = f land fzeta <> 0
-end : RedFlagsSig)
-
-open RedFlags
-
-(* Local *)
-let beta = mkflags [fbeta]
-let betaevar = mkflags [fevar; fbeta]
-let betaiota = mkflags [fiota; fbeta]
-let betaiotazeta = mkflags [fiota; fbeta;fzeta]
-
-(* Contextual *)
-let delta = mkflags [fdelta;fevar]
-let betadelta = mkflags [fbeta;fdelta;fzeta;fevar]
-let betadeltaeta = mkflags [fbeta;fdelta;fzeta;fevar;feta]
-let betadeltaiota = mkflags [fbeta;fdelta;fzeta;fevar;fiota]
-let betadeltaiota_nolet = mkflags [fbeta;fdelta;fevar;fiota]
-let betadeltaiotaeta = mkflags [fbeta;fdelta;fzeta;fevar;fiota;feta]
-let betaiotaevar = mkflags [fbeta;fiota;fevar]
-let betaetalet = mkflags [fbeta;feta;fzeta]
-
-(* Beta Reduction tools *)
-
-let rec stacklam recfun env t stack =
- match (decomp_stack stack,kind_of_term t) with
- | Some (h,stacktl), IsLambda (_,_,c) -> stacklam recfun (h::env) c stacktl
- | _ -> recfun (substl env t, stack)
-
-let beta_applist (c,l) =
- stacklam app_stack [] c (append_stack (Array.of_list l) empty_stack)
-
-(* Iota reduction tools *)
-
-type 'a miota_args = {
- mP : constr; (* the result type *)
- mconstr : constr; (* the constructor *)
- mci : case_info; (* special info to re-build pattern *)
- mcargs : 'a list; (* the constructor's arguments *)
- mlf : 'a array } (* the branch code vector *)
-
-let reducible_mind_case c = match kind_of_term c with
- | IsMutConstruct _ | IsCoFix _ -> true
- | _ -> false
-
-let contract_cofix (bodynum,(types,names,bodies as typedbodies)) =
- let nbodies = Array.length bodies in
- let make_Fi j = mkCoFix (nbodies-j-1,typedbodies) in
- substl (list_tabulate make_Fi nbodies) bodies.(bodynum)
-
-let reduce_mind_case mia =
- match kind_of_term mia.mconstr with
- | IsMutConstruct (ind_sp,i as cstr_sp) ->
-(* let ncargs = (fst mia.mci).(i-1) in*)
- let real_cargs = snd (list_chop (fst mia.mci) mia.mcargs) in
- applist (mia.mlf.(i-1),real_cargs)
- | IsCoFix cofix ->
- let cofix_def = contract_cofix cofix in
- mkMutCase (mia.mci, mia.mP, applist(cofix_def,mia.mcargs), mia.mlf)
- | _ -> assert false
-
-(* contracts fix==FIX[nl;i](A1...Ak;[F1...Fk]{B1....Bk}) to produce
- Bi[Fj --> FIX[nl;j](A1...Ak;[F1...Fk]{B1...Bk})] *)
-
-let contract_fix ((recindices,bodynum),(types,names,bodies as typedbodies)) =
- let nbodies = Array.length recindices in
- let make_Fi j = mkFix ((recindices,nbodies-j-1),typedbodies) in
- substl (list_tabulate make_Fi nbodies) bodies.(bodynum)
-
-let fix_recarg ((recindices,bodynum),_) stack =
- assert (0 <= bodynum & bodynum < Array.length recindices);
- let recargnum = Array.get recindices bodynum in
- try
- Some (recargnum, stack_nth stack recargnum)
- with Not_found ->
- None
-
-type fix_reduction_result = NotReducible | Reduced of state
-
-let reduce_fix whdfun fix stack =
- match fix_recarg fix stack with
- | None -> NotReducible
- | Some (recargnum,recarg) ->
- let (recarg'hd,_ as recarg') = whdfun (recarg, empty_stack) in
- let stack' = stack_assign stack recargnum (app_stack recarg') in
- (match kind_of_term recarg'hd with
- | IsMutConstruct _ -> Reduced (contract_fix fix, stack')
- | _ -> NotReducible)
-
-(* Generic reduction function *)
-
-(* Y avait un commentaire pour whd_betadeltaiota :
-
- NB : Cette fonction alloue peu c'est l'appel
- ``let (c,cargs) = whfun (recarg, empty_stack)''
- -------------------
- qui coute cher *)
-
-let rec whd_state_gen flags env sigma =
- let rec whrec (x, stack as s) =
- match kind_of_term x with
- | IsRel n when red_delta flags ->
- (match lookup_rel_value n env with
- | Some body -> whrec (lift n body, stack)
- | None -> s)
- | IsVar id when red_delta flags ->
- (match lookup_named_value id env with
- | Some body -> whrec (body, stack)
- | None -> s)
- | IsEvar ev when red_evar flags ->
- (match existential_opt_value sigma ev with
- | Some body -> whrec (body, stack)
- | None -> s)
- | IsConst const when red_delta flags ->
- (match constant_opt_value env const with
- | Some body -> whrec (body, stack)
- | None -> s)
- | IsLetIn (_,b,_,c) when red_zeta flags -> stacklam whrec [b] c stack
- | IsCast (c,_) -> whrec (c, stack)
- | IsApp (f,cl) -> whrec (f, append_stack cl stack)
- | IsLambda (na,t,c) ->
- (match decomp_stack stack with
- | Some (a,m) when red_beta flags -> stacklam whrec [a] c m
- | None when red_eta flags ->
- let env' = push_rel_assum (na,t) env in
- let whrec' = whd_state_gen flags env' sigma in
- (match kind_of_term (app_stack (whrec' (c, empty_stack))) with
- | IsApp (f,cl) ->
- let napp = Array.length cl in
- if napp > 0 then
- let x', l' = whrec' (array_last cl, empty_stack) in
- match kind_of_term x', decomp_stack l' with
- | IsRel 1, None ->
- let lc = Array.sub cl 0 (napp-1) in
- let u = if napp=1 then f else appvect (f,lc) in
- if noccurn 1 u then (pop u,empty_stack) else s
- | _ -> s
- else s
- | _ -> s)
- | _ -> s)
-
- | IsMutCase (ci,p,d,lf) when red_iota flags ->
- let (c,cargs) = whrec (d, empty_stack) in
- if reducible_mind_case c then
- whrec (reduce_mind_case
- {mP=p; mconstr=c; mcargs=list_of_stack cargs;
- mci=ci; mlf=lf}, stack)
- else
- (mkMutCase (ci, p, app_stack (c,cargs), lf), stack)
-
- | IsFix fix when red_iota flags ->
- (match reduce_fix whrec fix stack with
- | Reduced s' -> whrec s'
- | NotReducible -> s)
-
- | x -> s
- in
- whrec
-
-let local_whd_state_gen flags =
- let rec whrec (x, stack as s) =
- match kind_of_term x with
- | IsLetIn (_,b,_,c) when red_zeta flags -> stacklam whrec [b] c stack
- | IsCast (c,_) -> whrec (c, stack)
- | IsApp (f,cl) -> whrec (f, append_stack cl stack)
- | IsLambda (_,_,c) ->
- (match decomp_stack stack with
- | Some (a,m) when red_beta flags -> stacklam whrec [a] c m
- | None when red_eta flags ->
- (match kind_of_term (app_stack (whrec (c, empty_stack))) with
- | IsApp (f,cl) ->
- let napp = Array.length cl in
- if napp > 0 then
- let x', l' = whrec (array_last cl, empty_stack) in
- match kind_of_term x', decomp_stack l' with
- | IsRel 1, None ->
- let lc = Array.sub cl 0 (napp-1) in
- let u = if napp=1 then f else appvect (f,lc) in
- if noccurn 1 u then (pop u,empty_stack) else s
- | _ -> s
- else s
- | _ -> s)
- | _ -> s)
-
- | IsMutCase (ci,p,d,lf) when red_iota flags ->
- let (c,cargs) = whrec (d, empty_stack) in
- if reducible_mind_case c then
- whrec (reduce_mind_case
- {mP=p; mconstr=c; mcargs=list_of_stack cargs;
- mci=ci; mlf=lf}, stack)
- else
- (mkMutCase (ci, p, app_stack (c,cargs), lf), stack)
-
- | IsFix fix when red_iota flags ->
- (match reduce_fix whrec fix stack with
- | Reduced s' -> whrec s'
- | NotReducible -> s)
-
- | x -> s
- in
- whrec
-
-(* 1. Beta Reduction Functions *)
-
-let whd_beta_state = local_whd_state_gen beta
-let whd_beta_stack x = appterm_of_stack (whd_beta_state (x, empty_stack))
-let whd_beta x = app_stack (whd_beta_state (x,empty_stack))
-
-(* Nouveau ! *)
-let whd_betaetalet_state = local_whd_state_gen betaetalet
-let whd_betaetalet_stack x =
- appterm_of_stack (whd_betaetalet_state (x, empty_stack))
-let whd_betaetalet x = app_stack (whd_betaetalet_state (x,empty_stack))
-
-(* 2. Delta Reduction Functions *)
-
-let whd_delta_state e = whd_state_gen delta e
-let whd_delta_stack env sigma x =
- appterm_of_stack (whd_delta_state env sigma (x, empty_stack))
-let whd_delta env sigma c =
- app_stack (whd_delta_state env sigma (c, empty_stack))
-
-let whd_betadelta_state e = whd_state_gen betadelta e
-let whd_betadelta_stack env sigma x =
- appterm_of_stack (whd_betadelta_state env sigma (x, empty_stack))
-let whd_betadelta env sigma c =
- app_stack (whd_betadelta_state env sigma (c, empty_stack))
-
-let whd_betaevar_state e = whd_state_gen betaevar e
-let whd_betaevar_stack env sigma c =
- appterm_of_stack (whd_betaevar_state env sigma (c, empty_stack))
-let whd_betaevar env sigma c =
- app_stack (whd_betaevar_state env sigma (c, empty_stack))
-
-
-let whd_betadeltaeta_state e = whd_state_gen betadeltaeta e
-let whd_betadeltaeta_stack env sigma x =
- appterm_of_stack (whd_betadeltaeta_state env sigma (x, empty_stack))
-let whd_betadeltaeta env sigma x =
- app_stack (whd_betadeltaeta_state env sigma (x, empty_stack))
-
-(* 3. Iota reduction Functions *)
-
-let whd_betaiota_state = local_whd_state_gen betaiota
-let whd_betaiota_stack x =
- appterm_of_stack (whd_betaiota_state (x, empty_stack))
-let whd_betaiota x =
- app_stack (whd_betaiota_state (x, empty_stack))
-
-let whd_betaiotazeta_state = local_whd_state_gen betaiotazeta
-let whd_betaiotazeta_stack x =
- appterm_of_stack (whd_betaiotazeta_state (x, empty_stack))
-let whd_betaiotazeta x =
- app_stack (whd_betaiotazeta_state (x, empty_stack))
-
-let whd_betaiotaevar_state e = whd_state_gen betaiotaevar e
-let whd_betaiotaevar_stack env sigma x =
- appterm_of_stack (whd_betaiotaevar_state env sigma (x, empty_stack))
-let whd_betaiotaevar env sigma x =
- app_stack (whd_betaiotaevar_state env sigma (x, empty_stack))
-
-let whd_betadeltaiota_state e = whd_state_gen betadeltaiota e
-let whd_betadeltaiota_stack env sigma x =
- appterm_of_stack (whd_betadeltaiota_state env sigma (x, empty_stack))
-let whd_betadeltaiota env sigma x =
- app_stack (whd_betadeltaiota_state env sigma (x, empty_stack))
-
-let whd_betadeltaiotaeta_state e = whd_state_gen betadeltaiotaeta e
-let whd_betadeltaiotaeta_stack env sigma x =
- appterm_of_stack (whd_betadeltaiotaeta_state env sigma (x, empty_stack))
-let whd_betadeltaiotaeta env sigma x =
- app_stack (whd_betadeltaiotaeta_state env sigma (x, empty_stack))
-
-let whd_betadeltaiota_nolet_state e = whd_state_gen betadeltaiota_nolet e
-let whd_betadeltaiota_nolet_stack env sigma x =
- appterm_of_stack (whd_betadeltaiota_nolet_state env sigma (x, empty_stack))
-let whd_betadeltaiota_nolet env sigma x =
- app_stack (whd_betadeltaiota_nolet_state env sigma (x, empty_stack))
+let hnf_prod_applist env t nl =
+ List.fold_left (hnf_prod_app env) t nl
(********************************************************************)
(* Conversion *)
(********************************************************************)
-(*
-let fkey = Profile.declare_profile "fhnf";;
-let fhnf info v = Profile.profile2 fkey fhnf info v;;
-
-let fakey = Profile.declare_profile "fhnf_apply";;
-let fhnf_apply info k h a = Profile.profile4 fakey fhnf_apply info k h a;;
-*)
-
-type 'a conversion_function =
- env -> 'a evar_map -> constr -> constr -> constraints
(* Conversion utility functions *)
-
-type conversion_test = constraints -> constraints
+type 'a conversion_function = env -> 'a -> 'a -> constraints
exception NotConvertible
+exception NotConvertibleVect of int
(* Convertibility of sorts *)
@@ -454,12 +73,6 @@ type conv_pb =
| CONV
| CUMUL
-let pb_is_equal pb = pb = CONV
-
-let pb_equal = function
- | CUMUL -> CONV
- | CONV -> CONV
-
let sort_cmp pb s0 s1 cuniv =
match (s0,s1) with
| (Prop c1, Prop c2) -> if c1 = c2 then cuniv else raise NotConvertible
@@ -473,13 +86,6 @@ let sort_cmp pb s0 s1 cuniv =
| CUMUL -> enforce_geq u2 u1 cuniv)
| (_, _) -> raise NotConvertible
-let base_sort_cmp pb s0 s1 =
- match (s0,s1) with
- | (Prop c1, Prop c2) -> c1 = c2
- | (Prop c1, Type u) -> pb = CUMUL
- | (Type u1, Type u2) -> true
- | (_, _) -> false
-
(* Conversion between [lft1]term1 and [lft2]term2 *)
let rec ccnv cv_pb infos lft1 lft2 term1 term2 cuniv =
eqappr cv_pb infos (lft1, fhnf infos term1) (lft2, fhnf infos term2) cuniv
@@ -494,15 +100,20 @@ and eqappr cv_pb infos appr1 appr2 cuniv =
(* case of leaves *)
| (FAtom a1, FAtom a2) ->
(match kind_of_term a1, kind_of_term a2 with
- | (IsSort s1, IsSort s2) ->
+ | (Sort s1, Sort s2) ->
if stack_args_size v1 = 0 && stack_args_size v2 = 0
then sort_cmp cv_pb s1 s2 cuniv
else raise NotConvertible
- | (IsMeta n, IsMeta m) ->
+ | (Meta n, Meta m) ->
if n=m
then convert_stacks infos lft1 lft2 v1 v2 cuniv
else raise NotConvertible
| _ -> raise NotConvertible)
+ | (FEvar (ev1,args1), FEvar (ev2,args2)) ->
+ if ev1=ev2 then
+ let u1 = convert_vect infos el1 el2 args1 args2 cuniv in
+ convert_stacks infos lft1 lft2 v1 v2 u1
+ else raise NotConvertible
(* 2 index known to be bound to no constant *)
| (FRel n, FRel m) ->
@@ -575,13 +186,14 @@ and eqappr cv_pb infos appr1 appr2 cuniv =
convert_stacks infos lft1 lft2 v1 v2 u3
| (FInd op1, FInd op2) ->
- if op1 = op2
- then convert_stacks infos lft1 lft2 v1 v2 cuniv
+ if op1 = op2 then
+ convert_stacks infos lft1 lft2 v1 v2 cuniv
else raise NotConvertible
| (FConstruct op1, FConstruct op2) ->
if op1 = op2
- then convert_stacks infos lft1 lft2 v1 v2 cuniv
+ then
+ convert_stacks infos lft1 lft2 v1 v2 cuniv
else raise NotConvertible
| (FFix (op1,(_,tys1,cl1),_,_), FFix(op2,(_,tys2,cl2),_,_)) ->
@@ -631,241 +243,80 @@ and convert_vect infos lft1 lft2 v1 v2 cuniv =
-let fconv cv_pb env sigma t1 t2 =
+let fconv cv_pb env t1 t2 =
if eq_constr t1 t2 then
Constraint.empty
else
- let infos = create_clos_infos hnf_flags env sigma in
+ let infos = create_clos_infos hnf_flags env in
ccnv cv_pb infos ELID ELID (inject t1) (inject t2)
Constraint.empty
let conv env = fconv CONV env
-let conv_leq env = fconv CUMUL env
-
-(*
-let convleqkey = Profile.declare_profile "conv_leq";;
-let conv_leq env sigma t1 t2 =
- Profile.profile4 convleqkey conv_leq env sigma t1 t2;;
-
-let convkey = Profile.declare_profile "conv";;
-let conv env sigma t1 t2 =
- Profile.profile4 convleqkey conv env sigma t1 t2;;
-*)
-
-let conv_forall2 f env sigma v1 v2 =
- array_fold_left2
- (fun c x y -> let c' = f env sigma x y in Constraint.union c c')
- Constraint.empty
- v1 v2
+let conv_leq env = fconv CUMUL env
-let conv_forall2_i f env sigma v1 v2 =
+let conv_leq_vecti env v1 v2 =
array_fold_left2_i
- (fun i c x y -> let c' = f i env sigma x y in Constraint.union c c')
+ (fun i c t1 t2 ->
+ let c' =
+ try conv_leq env t1 t2
+ with NotConvertible -> raise (NotConvertibleVect i) in
+ Constraint.union c c')
Constraint.empty
- v1 v2
+ v1
+ v2
-let test_conversion f env sigma x y =
- try let _ = f env sigma x y in true with NotConvertible -> false
+(*
+let convleqkey = Profile.declare_profile "Kernel_reduction.conv_leq";;
+let conv_leq env t1 t2 =
+ Profile.profile4 convleqkey conv_leq env t1 t2;;
-let is_conv env sigma = test_conversion conv env sigma
-let is_conv_leq env sigma = test_conversion conv_leq env sigma
-let is_fconv = function | CONV -> is_conv | CUMUL -> is_conv_leq
+let convkey = Profile.declare_profile "Kernel_reduction.conv";;
+let conv env t1 t2 =
+ Profile.profile4 convleqkey conv env t1 t2;;
+*)
(********************************************************************)
(* Special-Purpose Reduction *)
(********************************************************************)
-let whd_meta metamap c = match kind_of_term c with
- | IsMeta p -> (try List.assoc p metamap with Not_found -> c)
- | _ -> c
-
-(* Try to replace all metas. Does not replace metas in the metas' values
- * Differs from (strong whd_meta). *)
-let plain_instance s c =
- let rec irec u = match kind_of_term u with
- | IsMeta p -> (try List.assoc p s with Not_found -> u)
- | IsCast (m,_) when isMeta m ->
- (try List.assoc (destMeta m) s with Not_found -> u)
- | _ -> map_constr irec u
- in
- if s = [] then c else irec c
-
-(* Pourquoi ne fait-on pas nf_betaiota si s=[] ? *)
-let instance s c =
- if s = [] then c else local_strong whd_betaiota (plain_instance s c)
-
+(* Dealing with arities *)
-(* pseudo-reduction rule:
- * [hnf_prod_app env s (Prod(_,B)) N --> B[N]
- * with an HNF on the first argument to produce a product.
- * if this does not work, then we use the string S as part of our
- * error message. *)
-
-let hnf_prod_app env sigma t n =
- match kind_of_term (whd_betadeltaiota env sigma t) with
- | IsProd (_,_,b) -> subst1 n b
- | _ -> anomaly "hnf_prod_app: Need a product"
-
-let hnf_prod_appvect env sigma t nl =
- Array.fold_left (hnf_prod_app env sigma) t nl
-
-let hnf_prod_applist env sigma t nl =
- List.fold_left (hnf_prod_app env sigma) t nl
-
-let hnf_lam_app env sigma t n =
- match kind_of_term (whd_betadeltaiota env sigma t) with
- | IsLambda (_,_,b) -> subst1 n b
- | _ -> anomaly "hnf_lam_app: Need an abstraction"
-
-let hnf_lam_appvect env sigma t nl =
- Array.fold_left (hnf_lam_app env sigma) t nl
-
-let hnf_lam_applist env sigma t nl =
- List.fold_left (hnf_lam_app env sigma) t nl
-
-let splay_prod env sigma =
+let dest_prod env =
let rec decrec env m c =
- let t = whd_betadeltaiota env sigma c in
+ let t = whd_betadeltaiota env c in
match kind_of_term t with
- | IsProd (n,a,c0) ->
- decrec (push_rel_assum (n,a) env)
- ((n,a)::m) c0
+ | Prod (n,a,c0) ->
+ let d = (n,None,a) in
+ decrec (push_rel d env) (Sign.add_rel_decl d m) c0
| _ -> m,t
in
decrec env []
-let splay_prod_assum env sigma =
- let rec prodec_rec env l c =
- let t = whd_betadeltaiota_nolet env sigma c in
- match kind_of_term c with
- | IsProd (x,t,c) ->
- prodec_rec (push_rel_assum (x,t) env)
- (Sign.add_rel_assum (x, t) l) c
- | IsLetIn (x,b,t,c) ->
- prodec_rec (push_rel_def (x,b, t) env)
- (Sign.add_rel_def (x,b, t) l) c
- | IsCast (c,_) -> prodec_rec env l c
- | _ -> l,t
+(* The same but preserving lets *)
+let dest_prod_assum env =
+ let rec prodec_rec env l ty =
+ let rty = whd_betadeltaiota_nolet env ty in
+ match kind_of_term rty with
+ | Prod (x,t,c) ->
+ let d = (x,None,t) in
+ prodec_rec (push_rel d env) (Sign.add_rel_decl d l) c
+ | LetIn (x,b,t,c) ->
+ let d = (x,Some b,t) in
+ prodec_rec (push_rel d env) (Sign.add_rel_decl d l) c
+ | Cast (c,_) -> prodec_rec env l c
+ | _ -> l,rty
in
prodec_rec env Sign.empty_rel_context
-let splay_arity env sigma c =
- let l, c = splay_prod env sigma c in
+let dest_arity env c =
+ let l, c = dest_prod env c in
match kind_of_term c with
- | IsSort s -> l,s
+ | Sort s -> l,s
| _ -> error "not an arity"
-let sort_of_arity env c = snd (splay_arity env Evd.empty c)
-
-let decomp_n_prod env sigma n =
- let rec decrec env m ln c = if m = 0 then (ln,c) else
- match kind_of_term (whd_betadeltaiota env sigma c) with
- | IsProd (n,a,c0) ->
- decrec (push_rel_assum (n,a) env)
- (m-1) (Sign.add_rel_assum (n,a) ln) c0
- | _ -> error "decomp_n_prod: Not enough products"
- in
- decrec env n Sign.empty_rel_context
-
-(* One step of approximation *)
-
-let rec apprec env sigma s =
- let (t, stack as s) = whd_betaiota_state s in
- match kind_of_term t with
- | IsMutCase (ci,p,d,lf) ->
- let (cr,crargs) = whd_betadeltaiota_stack env sigma d in
- let rslt = mkMutCase (ci, p, applist (cr,crargs), lf) in
- if reducible_mind_case cr then
- apprec env sigma (rslt, stack)
- else
- s
- | IsFix fix ->
- (match reduce_fix (whd_betadeltaiota_state env sigma) fix stack with
- | Reduced s -> apprec env sigma s
- | NotReducible -> s)
- | _ -> s
-
-let hnf env sigma c = apprec env sigma (c, empty_stack)
-
-(* A reduction function like whd_betaiota but which keeps casts
- * and does not reduce redexes containing existential variables.
- * Used in Correctness.
- * Added by JCF, 29/1/98. *)
-
-let whd_programs_stack env sigma =
- let rec whrec (x, stack as s) =
- match kind_of_term x with
- | IsApp (f,cl) ->
- let n = Array.length cl - 1 in
- let c = cl.(n) in
- if occur_existential c then
- s
- else
- whrec (mkApp (f, Array.sub cl 0 n), append_stack [|c|] stack)
- | IsLetIn (_,b,_,c) ->
- if occur_existential b then
- s
- else
- stacklam whrec [b] c stack
- | IsLambda (_,_,c) ->
- (match decomp_stack stack with
- | None -> s
- | Some (a,m) -> stacklam whrec [a] c m)
- | IsMutCase (ci,p,d,lf) ->
- if occur_existential d then
- s
- else
- let (c,cargs) = whrec (d, empty_stack) in
- if reducible_mind_case c then
- whrec (reduce_mind_case
- {mP=p; mconstr=c; mcargs=list_of_stack cargs;
- mci=ci; mlf=lf}, stack)
- else
- (mkMutCase (ci, p, app_stack(c,cargs), lf), stack)
- | IsFix fix ->
- (match reduce_fix whrec fix stack with
- | Reduced s' -> whrec s'
- | NotReducible -> s)
- | _ -> s
- in
- whrec
-
-let whd_programs env sigma x =
- app_stack (whd_programs_stack env sigma (x, empty_stack))
+let is_arity env c =
+ try
+ let _ = dest_arity env c in
+ true
+ with UserError _ -> false
-exception IsType
-
-let find_conclusion env sigma =
- let rec decrec env c =
- let t = whd_betadeltaiota env sigma c in
- match kind_of_term t with
- | IsProd (x,t,c0) -> decrec (push_rel_assum (x,t) env) c0
- | IsLambda (x,t,c0) -> decrec (push_rel_assum (x,t) env) c0
- | t -> t
- in
- decrec env
-
-let is_arity env sigma c =
- match find_conclusion env sigma c with
- | IsSort _ -> true
- | _ -> false
-
-let info_arity env sigma c =
- match find_conclusion env sigma c with
- | IsSort (Prop Null) -> false
- | IsSort (Prop Pos) -> true
- | _ -> raise IsType
-
-let is_info_arity env sigma c =
- try (info_arity env sigma c) with IsType -> true
-
-let is_type_arity env sigma c =
- match find_conclusion env sigma c with
- | IsSort (Type _) -> true
- | _ -> false
-
-let is_info_type env sigma t =
- let s = t.utj_type in
- (s = Prop Pos) ||
- (s <> Prop Null &&
- try info_arity env sigma t.utj_val with IsType -> true)
diff --git a/kernel/reduction.mli b/kernel/reduction.mli
index 09d47fec9..d67b321e9 100644
--- a/kernel/reduction.mli
+++ b/kernel/reduction.mli
@@ -9,195 +9,39 @@
(*i $Id$ i*)
(*i*)
-open Names
open Term
-open Univ
-open Evd
open Environ
-open Closure
(*i*)
-(* Reduction Functions. *)
-
-exception Elimconst
-
-type state = constr * constr stack
-
-type 'a contextual_reduction_function = env -> 'a evar_map -> constr -> constr
-type 'a reduction_function = 'a contextual_reduction_function
-type local_reduction_function = constr -> constr
-
-type 'a contextual_stack_reduction_function =
- env -> 'a evar_map -> constr -> constr * constr list
-type 'a stack_reduction_function = 'a contextual_stack_reduction_function
-type local_stack_reduction_function = constr -> constr * constr list
-
-type 'a contextual_state_reduction_function =
- env -> 'a evar_map -> state -> state
-type 'a state_reduction_function = 'a contextual_state_reduction_function
-type local_state_reduction_function = state -> state
-
-(* Removes cast and put into applicative form *)
-val whd_stack : local_stack_reduction_function
-
-(* For compatibility: alias for whd\_stack *)
-val whd_castapp_stack : local_stack_reduction_function
-
-(*s Reduction Function Operators *)
-
-val strong : 'a reduction_function -> 'a reduction_function
-val local_strong : local_reduction_function -> local_reduction_function
-val strong_prodspine : local_reduction_function -> local_reduction_function
-(*i
-val stack_reduction_of_reduction :
- 'a reduction_function -> 'a state_reduction_function
-i*)
-val stacklam : (state -> 'a) -> constr list -> constr -> constr stack -> 'a
-
-(*s Generic Optimized Reduction Function using Closures *)
-
-val clos_norm_flags : Closure.flags -> 'a reduction_function
-(* Same as [(strong whd_beta[delta][iota])], but much faster on big terms *)
-val nf_beta : local_reduction_function
-val nf_betaiota : local_reduction_function
-val nf_betadeltaiota : 'a reduction_function
-
-(* Lazy strategy, weak head reduction *)
-val whd_beta : local_reduction_function
-val whd_betaiota : local_reduction_function
-val whd_betaiotazeta : local_reduction_function
-val whd_betadeltaiota : 'a contextual_reduction_function
-val whd_betadeltaiota_nolet : 'a contextual_reduction_function
-val whd_betaetalet : local_reduction_function
-
-val whd_beta_stack : local_stack_reduction_function
-val whd_betaiota_stack : local_stack_reduction_function
-val whd_betaiotazeta_stack : local_stack_reduction_function
-val whd_betadeltaiota_stack : 'a contextual_stack_reduction_function
-val whd_betadeltaiota_nolet_stack : 'a contextual_stack_reduction_function
-val whd_betaetalet_stack : local_stack_reduction_function
-
-val whd_beta_state : local_state_reduction_function
-val whd_betaiota_state : local_state_reduction_function
-val whd_betaiotazeta_state : local_state_reduction_function
-val whd_betadeltaiota_state : 'a contextual_state_reduction_function
-val whd_betadeltaiota_nolet_state : 'a contextual_state_reduction_function
-val whd_betaetalet_state : local_state_reduction_function
-
-(*s Head normal forms *)
-
-val whd_delta_stack : 'a stack_reduction_function
-val whd_delta_state : 'a state_reduction_function
-val whd_delta : 'a reduction_function
-val whd_betadelta_stack : 'a stack_reduction_function
-val whd_betadelta_state : 'a state_reduction_function
-val whd_betadelta : 'a reduction_function
-val whd_betaevar_stack : 'a stack_reduction_function
-val whd_betaevar_state : 'a state_reduction_function
-val whd_betaevar : 'a reduction_function
-val whd_betaiotaevar_stack : 'a stack_reduction_function
-val whd_betaiotaevar_state : 'a state_reduction_function
-val whd_betaiotaevar : 'a reduction_function
-val whd_betadeltaeta_stack : 'a stack_reduction_function
-val whd_betadeltaeta_state : 'a state_reduction_function
-val whd_betadeltaeta : 'a reduction_function
-val whd_betadeltaiotaeta_stack : 'a stack_reduction_function
-val whd_betadeltaiotaeta_state : 'a state_reduction_function
-val whd_betadeltaiotaeta : 'a reduction_function
-
-val beta_applist : constr * constr list -> constr
-
-val hnf_prod_app : env -> 'a evar_map -> constr -> constr -> constr
-val hnf_prod_appvect : env -> 'a evar_map -> constr -> constr array -> constr
-val hnf_prod_applist : env -> 'a evar_map -> constr -> constr list -> constr
-val hnf_lam_app : env -> 'a evar_map -> constr -> constr -> constr
-val hnf_lam_appvect : env -> 'a evar_map -> constr -> constr array -> constr
-val hnf_lam_applist : env -> 'a evar_map -> constr -> constr list -> constr
-
-val splay_prod : env -> 'a evar_map -> constr -> (name * constr) list * constr
-val splay_arity : env -> 'a evar_map -> constr -> (name * constr) list * sorts
-val sort_of_arity : env -> constr -> sorts
-val decomp_n_prod :
- env -> 'a evar_map -> int -> constr -> Sign.rel_context * constr
-val splay_prod_assum :
- env -> 'a evar_map -> constr -> Sign.rel_context * constr
-
-type 'a miota_args = {
- mP : constr; (* the result type *)
- mconstr : constr; (* the constructor *)
- mci : case_info; (* special info to re-build pattern *)
- mcargs : 'a list; (* the constructor's arguments *)
- mlf : 'a array } (* the branch code vector *)
-
-val reducible_mind_case : constr -> bool
-val reduce_mind_case : constr miota_args -> constr
-
-val is_arity : env -> 'a evar_map -> constr -> bool
-val is_info_type : env -> 'a evar_map -> unsafe_type_judgment -> bool
-val is_info_arity : env -> 'a evar_map -> constr -> bool
-(*i Pour l'extraction
-val is_type_arity : env -> 'a evar_map -> constr -> bool
-val is_info_cast_type : env -> 'a evar_map -> constr -> bool
-val contents_of_cast_type : env -> 'a evar_map -> constr -> contents
-i*)
-
-val whd_programs : 'a reduction_function
-
-(* [reduce_fix] contracts a fix redex if it is actually reducible *)
+(***********************************************************************)
+(*s Reduction functions *)
-type fix_reduction_result = NotReducible | Reduced of state
+val whd_betadeltaiota : env -> constr -> constr
+val whd_betadeltaiota_nolet : env -> constr -> constr
-val fix_recarg : fixpoint -> constr stack -> (int * constr) option
-val reduce_fix : local_state_reduction_function -> fixpoint
- -> constr stack -> fix_reduction_result
+val nf_betaiota : constr -> constr
+val hnf_stack : env -> constr -> constr * constr list
+val hnf_prod_applist : env -> types -> constr list -> types
-(*s Conversion Functions (uses closures, lazy strategy) *)
+(* Builds an application node, reducing beta redexes it may produce. *)
+val beta_appvect : constr -> constr array -> constr
-type conversion_test = constraints -> constraints
+(***********************************************************************)
+(*s conversion functions *)
exception NotConvertible
+exception NotConvertibleVect of int
+type 'a conversion_function = env -> 'a -> 'a -> Univ.constraints
-type conv_pb =
- | CONV
- | CUMUL
-
-val pb_is_equal : conv_pb -> bool
-val pb_equal : conv_pb -> conv_pb
-
-val sort_cmp : conv_pb -> sorts -> sorts -> conversion_test
-val base_sort_cmp : conv_pb -> sorts -> sorts -> bool
+val conv : constr conversion_function
+val conv_leq : types conversion_function
+val conv_leq_vecti : types array conversion_function
-type 'a conversion_function =
- env -> 'a evar_map -> constr -> constr -> constraints
-
-(* [fconv] has 2 instances: [conv = fconv CONV] i.e. conversion test, and
- [conv_leq = fconv CONV_LEQ] i.e. cumulativity test. *)
-
-val conv : 'a conversion_function
-val conv_leq : 'a conversion_function
-
-val conv_forall2 :
- 'a conversion_function -> env -> 'a evar_map -> constr array
- -> constr array -> constraints
-
-val conv_forall2_i :
- (int -> 'a conversion_function) -> env -> 'a evar_map
- -> constr array -> constr array -> constraints
-
-val is_conv : env -> 'a evar_map -> constr -> constr -> bool
-val is_conv_leq : env -> 'a evar_map -> constr -> constr -> bool
-val is_fconv : conv_pb -> env -> 'a evar_map -> constr -> constr -> bool
-
-(*s Special-Purpose Reduction Functions *)
-
-val whd_meta : (int * constr) list -> constr -> constr
-val plain_instance : (int * constr) list -> constr -> constr
-val instance : (int * constr) list -> constr -> constr
-
-(*s Obsolete Reduction Functions *)
+(***********************************************************************)
+(*s Recognizing products and arities modulo reduction *)
-(*i
-val hnf : env -> 'a evar_map -> constr -> constr * constr list
-i*)
-val apprec : 'a state_reduction_function
+val dest_prod : env -> types -> Sign.rel_context * types
+val dest_prod_assum : env -> types -> Sign.rel_context * types
+val dest_arity : env -> types -> arity
+val is_arity : env -> types -> bool
diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml
index a6ae51f89..c770e0237 100644
--- a/kernel/safe_typing.ml
+++ b/kernel/safe_typing.ml
@@ -8,7 +8,6 @@
(* $Id$ *)
-open Pp
open Util
open Names
open Univ
@@ -19,169 +18,16 @@ open Declarations
open Inductive
open Environ
open Type_errors
-open Typeops
open Indtypes
type judgment = unsafe_judgment
-
-let j_val j = j.uj_val
-let j_type j = body_of_type j.uj_type
-
-let vect_lift = Array.mapi lift
-let vect_lift_type = Array.mapi (fun i t -> type_app (lift i) t)
-
-(* The typing machine without information. *)
-
- (* ATTENTION : faudra faire le typage du contexte des Const,
- MutInd et MutConstructsi un jour cela devient des constructions
- arbitraires et non plus des variables *)
-
-let univ_combinator (cst,univ) (j,c') =
- (j,(Constraint.union cst c', merge_constraints c' univ))
-
-let rec execute env cstr cu =
- match kind_of_term cstr with
- | IsMeta _ ->
- anomaly "the kernel does not understand metas"
- | IsEvar _ ->
- anomaly "the kernel does not understand existential variables"
-
- | IsSort (Prop c) ->
- (judge_of_prop_contents c, cu)
-
- | IsSort (Type u) ->
- univ_combinator cu (judge_of_type u)
-
- | IsApp (f,args) ->
- let (j,cu1) = execute env f cu in
- let (jl,cu2) = execute_array env args cu1 in
- univ_combinator cu2
- (apply_rel_list env Evd.empty false (Array.to_list jl) j)
-
- | IsLambda (name,c1,c2) ->
- let (j,cu1) = execute env c1 cu in
- let var = assumption_of_judgment env Evd.empty j in
- let env1 = push_rel_assum (name,var) env in
- let (j',cu2) = execute env1 c2 cu1 in
- univ_combinator cu2 (abs_rel env1 Evd.empty name var j')
-
- | IsProd (name,c1,c2) ->
- let (j,cu1) = execute env c1 cu in
- let varj = type_judgment env Evd.empty j in
- let env1 = push_rel_assum (name,varj.utj_val) env in
- let (j',cu2) = execute env1 c2 cu1 in
- let varj' = type_judgment env Evd.empty j' in
- univ_combinator cu2
- (gen_rel env1 Evd.empty name varj varj')
-
- | IsLetIn (name,c1,c2,c3) ->
- let (j,cu1) = execute env (mkCast(c1,c2)) cu in
- let env1 = push_rel_def (name,j.uj_val,j.uj_type) env in
- let (j',cu2) = execute env1 c3 cu1 in
- univ_combinator cu2
- (judge_of_letin env1 Evd.empty name j j')
-
- | IsCast (c,t) ->
- let (cj,cu1) = execute env c cu in
- let (tj,cu2) = execute env t cu1 in
- let tj = assumption_of_judgment env Evd.empty tj in
- univ_combinator cu2
- (cast_rel env Evd.empty cj tj)
-
- | IsRel n ->
- (relative env n, cu)
-
- | IsVar id ->
- (make_judge cstr (lookup_named_type id env), cu)
-
- | IsConst c ->
- (make_judge cstr (type_of_constant env Evd.empty c), cu)
-
- (* Inductive types *)
- | IsMutInd ind ->
- (make_judge cstr (type_of_inductive env Evd.empty ind), cu)
-
- | IsMutConstruct c ->
- (make_judge cstr (type_of_constructor env Evd.empty c), cu)
-
- | IsMutCase (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 Evd.empty ci pj cj lfj)
-
- | IsFix ((vn,i as vni),recdef) ->
- if array_exists (fun n -> n < 0) vn then
- error "General Fixpoints not allowed";
- let ((_,tys,_ as recdef'),cu1) = execute_fix env recdef cu in
- let fix = (vni,recdef') in
- check_fix env Evd.empty fix;
- (make_judge (mkFix fix) tys.(i), cu1)
-
- | IsCoFix (i,recdef) ->
- let ((_,tys,_ as recdef'),cu1) = execute_fix env recdef cu in
- let cofix = (i,recdef') in
- check_cofix env Evd.empty cofix;
- (make_judge (mkCoFix cofix) tys.(i), cu1)
-
-and execute_fix env (names,lar,vdef) cu =
- let (larj,cu1) = execute_array env lar cu in
- let lara = Array.map (assumption_of_judgment env Evd.empty) larj in
- let env1 = push_rec_types (names,lara,vdef) env in
- let (vdefj,cu2) = execute_array env1 vdef cu1 in
- let vdefv = Array.map j_val vdefj in
- let cst = type_fixpoint env1 Evd.empty names lara vdefj in
- univ_combinator cu2 ((names,lara,vdefv),cst)
-
-and execute_array env v cu =
- let (jl,cu1) = execute_list env (Array.to_list v) cu in
- (Array.of_list jl, cu1)
-
-and execute_list env l cu =
- match l with
- | [] ->
- ([], cu)
- | c::r ->
- let (j,cu1) = execute env c cu in
- let (jr,cu2) = execute_list env r cu1 in
- (j::jr, cu2)
-
-(* The typed type of a judgment. *)
-
-let execute_type env constr cu =
- let (j,cu1) = execute env constr cu in
- (type_judgment env Evd.empty j, cu1)
+let j_val = j_val
+let j_type = j_type
(* Exported machines. *)
-let safe_infer env constr =
- let (j,(cst,_)) =
- execute env constr (Constraint.empty, universes env) in
- (j, cst)
-
-let safe_infer_type env constr =
- let (j,(cst,_)) =
- execute_type env constr (Constraint.empty, universes env) in
- (j, cst)
-
-(* Typing of several terms. *)
-
-let safe_infer_l env cl =
- let type_one (cst,l) c =
- let (j,cst') = safe_infer env c in
- (Constraint.union cst cst', j::l)
- in
- List.fold_left type_one (Constraint.empty,[]) cl
-
-let safe_infer_v env cv =
- let type_one (cst,l) c =
- let (j,cst') = safe_infer env c in
- (Constraint.union cst cst', j::l)
- in
- let cst',l = Array.fold_left type_one (Constraint.empty,[]) cv in
- (cst', Array.of_list l)
-
+let safe_infer = Typeops.infer
+let safe_infer_type = Typeops.infer_type
(*s Safe environments. *)
@@ -189,273 +35,107 @@ type safe_environment = env
let empty_environment = empty_env
-let universes = universes
-let context = context
-let named_context = named_context
-
-let lookup_named_type = lookup_named_type
-let lookup_rel_type = lookup_rel_type
-let lookup_named = lookup_named
-let lookup_constant = lookup_constant
-let lookup_mind = lookup_mind
-let lookup_mind_specif = lookup_mind_specif
-
(* Insertion of variables (named and de Bruijn'ed). They are now typed before
being added to the environment. *)
let push_rel_or_named_def push (id,b) env =
let (j,cst) = safe_infer env b in
let env' = add_constraints cst env in
- push (id,j.uj_val,j.uj_type) env'
+ let env'' = push (id,Some j.uj_val,j.uj_type) env' in
+ (cst,env'')
-let push_named_def = push_rel_or_named_def push_named_def
-let push_rel_def = push_rel_or_named_def push_rel_def
+let push_named_def = push_rel_or_named_def push_named_decl
+let push_rel_def = push_rel_or_named_def push_rel
let push_rel_or_named_assum push (id,t) env =
let (j,cst) = safe_infer env t in
+ let t = Typeops.assumption_of_judgment env j in
let env' = add_constraints cst env in
- let t = assumption_of_judgment env Evd.empty j in
- push (id,t) env'
-
-let push_named_assum = push_rel_or_named_assum push_named_assum
-let push_rel_assum = push_rel_or_named_assum push_rel_assum
-
-let check_and_push_named_def (id,b) env =
- let (j,cst) = safe_infer env b in
- let env' = add_constraints cst env in
- let env'' = Environ.push_named_def (id,j.uj_val,j.uj_type) env' in
- (Some j.uj_val,j.uj_type,cst),env''
+ let env'' = push (id,None,t) env' in
+ (cst,env'')
-let check_and_push_named_assum (id,t) env =
- let (j,cst) = safe_infer env t in
- let env' = add_constraints cst env in
- let t = assumption_of_judgment env Evd.empty j in
- let env'' = Environ.push_named_assum (id,t) env' in
- (None,t,cst),env''
+let push_named_assum = push_rel_or_named_assum push_named_decl
+let push_rel_assum d env = snd (push_rel_or_named_assum push_rel d env)
let push_rels_with_univ vars env =
List.fold_left (fun env nvar -> push_rel_assum nvar env) env vars
-let safe_infer_local_decl env id = function
- | LocalDef c ->
- let (j,cst) = safe_infer env c in
- (Name id, Some j.uj_val, j.uj_type), cst
- | LocalAssum c ->
- let (j,cst) = safe_infer env c in
- (Name id, None, assumption_of_judgment env Evd.empty j), cst
-
-let safe_infer_local_decls env decls =
- let rec inferec env = function
- | (id, d) :: l ->
- let env, l, cst1 = inferec env l in
- let d, cst2 = safe_infer_local_decl env id d in
- push_rel d env, d :: l, Constraint.union cst1 cst2
- | [] -> env, [], Constraint.empty in
- inferec env decls
-
(* Insertion of constants and parameters in environment. *)
-type global_declaration = Def of constr | Assum of constr
+type global_declaration = Def of constr * bool | Assum of constr
-let safe_infer_declaration env = function
- | Def c ->
+(* Definition always declared transparent *)
+let safe_infer_declaration env dcl =
+ match dcl with
+ | Def (c,op) ->
let (j,cst) = safe_infer env c in
- Some j.uj_val, j.uj_type, cst
+ Some j.uj_val, j.uj_type, cst, op
| Assum t ->
let (j,cst) = safe_infer env t in
- None, assumption_of_judgment env Evd.empty j, cst
+ None, Typeops.assumption_of_judgment env j, cst, false
-type local_names = (identifier * variable) list
-
-let add_global_declaration sp env locals (body,typ,cst) op =
+let add_global_declaration sp env (body,typ,cst,op) =
let env' = add_constraints cst env in
let ids = match body with
| None -> global_vars_set env typ
| Some b ->
Idset.union (global_vars_set env b) (global_vars_set env typ) in
- let hyps = keep_hyps env ids (named_context env) in
- let sp_hyps = List.map (fun (id,b,t) -> (List.assoc id locals, b, t)) hyps in
+ let hyps = keep_hyps env ids in
let cb = {
- const_kind = kind_of_path sp;
const_body = body;
const_type = typ;
- const_hyps = sp_hyps;
+ const_hyps = hyps;
const_constraints = cst;
- const_opaque = op }
- in
+ const_opaque = op } in
Environ.add_constant sp cb env'
-let add_parameter sp t locals env =
- add_global_declaration
- sp env locals (safe_infer_declaration env (Assum t)) false
+let add_parameter sp t env =
+ add_global_declaration sp env (safe_infer_declaration env (Assum t))
+
+(*s Global and local constant declaration. *)
-let add_constant sp ce locals env =
- let { const_entry_body = body;
- const_entry_type = typ;
- const_entry_opaque = op } = ce in
- let body' =
- match typ with
- | None -> body
- | Some ty -> mkCast (body, ty) in
- add_global_declaration
- sp env locals (safe_infer_declaration env (Def body')) op
+type constant_entry = {
+ const_entry_body : constr;
+ const_entry_type : types option;
+ const_entry_opaque : bool }
-let add_discharged_constant sp r locals env =
+let add_constant sp ce env =
+ let body =
+ match ce.const_entry_type with
+ | None -> ce.const_entry_body
+ | Some ty -> mkCast (ce.const_entry_body, ty) in
+ add_global_declaration sp env
+ (safe_infer_declaration env (Def (body, ce.const_entry_opaque)))
+
+let add_discharged_constant sp r env =
let (body,typ,cst,op) = Cooking.cook_constant env r in
- let env' = add_constraints cst env in
match body with
| None ->
- add_parameter sp typ locals (* Bricolage avant poubelle *) env'
+ add_parameter sp typ (* Bricolage avant poubelle *) env
| Some c ->
(* let c = hcons1_constr c in *)
- let ids =
- Idset.union (global_vars_set env c) (global_vars_set env typ) in
- let hyps = keep_hyps env ids (named_context env') in
- let sp_hyps =
- List.map (fun (id,b,t) -> (List.assoc id locals,b,t)) hyps in
+ let ids =
+ Idset.union (global_vars_set env c)
+ (global_vars_set env (body_of_type typ))
+ in
+ let hyps = keep_hyps env ids in
+ let env' = Environ.add_constraints cst env in
let cb =
- { const_kind = kind_of_path sp;
- const_body = Some c;
+ { const_body = Some c;
const_type = typ;
- const_hyps = sp_hyps;
+ const_hyps = hyps;
const_constraints = cst;
- const_opaque = op }
- in
+ const_opaque = op } in
Environ.add_constant sp cb env'
(* Insertion of inductive types. *)
-(* Only the case where at least s1 or s2 is a [Type] is taken into account *)
-let max_universe (s1,cst1) (s2,cst2) g =
- match s1,s2 with
- | Type u1, Type u2 ->
- let (u12,cst) = sup u1 u2 g in
- Type u12, Constraint.union cst (Constraint.union cst1 cst2)
- | Type u1, _ -> s1, cst1
- | _, _ -> s2, cst2
-
-(* This (re)computes informations relevant to extraction and the sort of an
- arity or type constructor; we do not to recompute universes constraints *)
-
-let rec infos_and_sort env t =
- match kind_of_term t with
- | IsProd (name,c1,c2) ->
- let (varj,_) = safe_infer_type env c1 in
- let env1 = Environ.push_rel_assum (name,varj.utj_val) env in
- let s1 = varj.utj_type in
- let logic = not (is_info_type env Evd.empty varj) in
- let small = is_small s1 in
- (logic,small) :: (infos_and_sort env1 c2)
- | IsCast (c,_) -> infos_and_sort env c
- | _ -> []
-
-(* [infos] is a sequence of pair [islogic,issmall] for each type in
- the product of a constructor or arity *)
+let add_mind sp mie env =
+ let mib = check_inductive env mie in
+ let cst = mib.mind_constraints in
+ Environ.add_mind sp mib (add_constraints cst env)
-let is_small infos = List.for_all (fun (logic,small) -> small) infos
-let is_logic_constr infos = List.for_all (fun (logic,small) -> logic) infos
-let is_logic_arity infos =
- List.for_all (fun (logic,small) -> logic || small) infos
-
-let is_unit arinfos constrsinfos =
- match constrsinfos with (* One info = One constructor *)
- | [constrinfos] -> is_logic_constr constrinfos && is_logic_arity arinfos
- | _ -> false
-
-let small_unit constrsinfos (env_ar_par,short_arity) =
- let issmall = List.for_all is_small constrsinfos in
- let arinfos = infos_and_sort env_ar_par short_arity in
- let isunit = is_unit arinfos constrsinfos in
- issmall, isunit
-
-(* [smax] is the max of the sorts of the products of the constructor type *)
-
-let enforce_type_constructor arsort smax cst =
- match smax, arsort with
- | Type uc, Type ua -> enforce_geq ua uc cst
- | _,_ -> cst
-
-let type_one_constructor env_ar_par params arsort c =
- let infos = infos_and_sort env_ar_par c in
-
- (* Each constructor is typed-checked here *)
- let (j,cst) = safe_infer_type env_ar_par c in
- let full_cstr_type = it_mkProd_or_LetIn j.utj_val params in
-
- (* If the arity is at some level Type arsort, then the sort of the
- constructor must be below arsort; here we consider constructors with the
- global parameters (which add a priori more constraints on their sort) *)
- let cst2 = enforce_type_constructor arsort j.utj_type cst in
-
- (infos, full_cstr_type, cst2)
-
-let infer_constructor_packet env_ar params short_arity arsort vc =
- let env_ar_par = push_rels params env_ar in
- let (constrsinfos,jlc,cst) =
- List.fold_right
- (fun c (infosl,l,cst) ->
- let (infos,ct,cst') =
- type_one_constructor env_ar_par params arsort c in
- (infos::infosl,ct::l, Constraint.union cst cst'))
- vc
- ([],[],Constraint.empty) in
- let vc' = Array.of_list jlc in
- let issmall,isunit = small_unit constrsinfos (env_ar_par,short_arity) in
- (issmall,isunit,vc', cst)
-
-let add_mind sp mie locals env =
- mind_check_wellformed env mie;
-
- (* We first type params and 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_params_arity_list =
- List.fold_left
- (fun (cst,env_arities,l) ind ->
- (* Params are typed-checked here *)
- let params = ind.mind_entry_params in
- let env_params, params, cst1 = safe_infer_local_decls env params in
- (* Arities (without params) are typed-checked here *)
- let arity, cst2 = safe_infer_type env_params ind.mind_entry_arity 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 id = ind.mind_entry_typename in
- let full_arity = it_mkProd_or_LetIn arity.utj_val params in
- Constraint.union cst (Constraint.union cst1 cst2),
- push_rel_assum (Name id, full_arity) env_arities,
- (params, id, full_arity, arity.utj_val)::l)
- (Constraint.empty,env,[])
- mie.mind_entry_inds in
-
- let params_arity_list = List.rev rev_params_arity_list in
-
- (* Now, we type the constructors (without params) *)
- let inds,cst =
- List.fold_right2
- (fun ind (params,id,full_arity,short_arity) (inds,cst) ->
- let arsort = sort_of_arity env full_arity in
- let lc = ind.mind_entry_lc in
- let (issmall,isunit,lc',cst') =
- infer_constructor_packet env_arities params short_arity arsort lc
- in
- let nparams = ind.mind_entry_nparams in
- let consnames = ind.mind_entry_consnames in
- let ind' = (params,nparams,id,full_arity,consnames,issmall,isunit,lc')
- in
- (ind'::inds, Constraint.union cst cst'))
- mie.mind_entry_inds
- params_arity_list
- ([],cst) in
-
- (* Finally, we build the inductive packet and push it to env *)
- let kind = kind_of_path sp in
- let mib = cci_inductive locals env env_arities kind mie.mind_entry_finite inds cst
- in
- add_mind sp mib (add_constraints cst env)
-
-let add_constraints = add_constraints
+let add_constraints = Environ.add_constraints
let rec pop_named_decls idl env =
match idl with
@@ -471,6 +151,5 @@ let env_of_safe_env e = e
let typing env c =
let (j,cst) = safe_infer env c in
+ let _ = add_constraints cst env in
j
-
-let typing_in_unsafe_env = typing
diff --git a/kernel/safe_typing.mli b/kernel/safe_typing.mli
index 23a970b49..5f6697b4e 100644
--- a/kernel/safe_typing.mli
+++ b/kernel/safe_typing.mli
@@ -9,15 +9,9 @@
(*i $Id$ i*)
(*i*)
-open Pp
open Names
open Term
-open Univ
-open Sign
open Declarations
-open Inductive
-open Environ
-open Typeops
(*i*)
(*s Safe environments. Since we are now able to type terms, we can define an
@@ -27,50 +21,47 @@ open Typeops
type safe_environment
-val empty_environment : safe_environment
+val env_of_safe_env : safe_environment -> Environ.env
-val universes : safe_environment -> universes
-val context : safe_environment -> context
-val named_context : safe_environment -> named_context
+val empty_environment : safe_environment
+(* Adding and removing local declarations (Local or Variables) *)
val push_named_assum :
- identifier * constr -> safe_environment -> safe_environment
+ identifier * types -> safe_environment ->
+ Univ.constraints * safe_environment
val push_named_def :
- identifier * constr -> safe_environment -> safe_environment
-
-val check_and_push_named_assum :
identifier * constr -> safe_environment ->
- (constr option * types * constraints) * safe_environment
-val check_and_push_named_def :
- identifier * constr -> safe_environment ->
- (constr option * types * constraints) * safe_environment
+ Univ.constraints * safe_environment
+val pop_named_decls : identifier list -> safe_environment -> safe_environment
-type local_names = (identifier * variable) list
+(* Adding global axioms or definitions *)
val add_parameter :
- section_path -> constr -> local_names -> safe_environment -> safe_environment
+ section_path -> constr -> safe_environment -> safe_environment
+
+(*s Global and local constant declaration. *)
+
+type constant_entry = {
+ const_entry_body : constr;
+ const_entry_type : types option;
+ const_entry_opaque : bool }
+
val add_constant :
- section_path -> constant_entry -> local_names ->
- safe_environment -> safe_environment
+ section_path -> constant_entry -> safe_environment -> safe_environment
val add_discharged_constant :
- section_path -> Cooking.recipe -> local_names -> safe_environment -> safe_environment
+ section_path -> Cooking.recipe -> safe_environment -> safe_environment
+(* Adding an inductive type *)
val add_mind :
- section_path -> mutual_inductive_entry -> local_names -> safe_environment
- -> safe_environment
-val add_constraints : constraints -> safe_environment -> safe_environment
-
-val pop_named_decls : identifier list -> safe_environment -> safe_environment
-
-val lookup_named : identifier -> safe_environment -> constr option * types
-val lookup_constant : section_path -> safe_environment -> constant_body
-val lookup_mind : section_path -> safe_environment -> mutual_inductive_body
-val lookup_mind_specif : inductive -> safe_environment -> inductive_instance
+ section_path -> Indtypes.mutual_inductive_entry -> safe_environment ->
+ safe_environment
-val export : safe_environment -> dir_path -> compiled_env
-val import : compiled_env -> safe_environment -> safe_environment
+(* Adding universe constraints *)
+val add_constraints : Univ.constraints -> safe_environment -> safe_environment
-val env_of_safe_env : safe_environment -> env
+(* Loading and saving a module *)
+val export : safe_environment -> dir_path -> Environ.compiled_env
+val import : Environ.compiled_env -> safe_environment -> safe_environment
(*s Typing judgments *)
@@ -80,9 +71,12 @@ type judgment
val j_val : judgment -> constr
val j_type : judgment -> constr
-val safe_infer : safe_environment -> constr -> judgment * constraints
+(* Safe typing of a term returning a typing judgment and 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 -> constr -> judgment * Univ.constraints
val typing : safe_environment -> constr -> judgment
-val typing_in_unsafe_env : env -> constr -> judgment
diff --git a/kernel/sign.ml b/kernel/sign.ml
index 0899cf5e6..c9da4ab65 100644
--- a/kernel/sign.ml
+++ b/kernel/sign.ml
@@ -26,17 +26,9 @@ type named_context = named_declaration list
let add_named_decl = add
let add_named_assum = add_decl
let add_named_def = add_def
-let rec lookup_id_type id = function
- | (id',c,t) :: _ when id=id' -> t
- | _ :: sign -> lookup_id_type id sign
- | [] -> raise Not_found
-let rec lookup_id_value id = function
- | (id',c,t) :: _ when id=id' -> c
- | _ :: sign -> lookup_id_value id sign
- | [] -> raise Not_found
-let rec lookup_id id = function
- | (id',c,t) :: _ when id=id' -> (c,t)
- | _ :: sign -> lookup_id id sign
+let rec lookup_named id = function
+ | (id',_,_ as decl) :: _ when id=id' -> decl
+ | _ :: sign -> lookup_named id sign
| [] -> raise Not_found
let empty_named_context = []
let pop_named_decl id = function
@@ -59,18 +51,13 @@ let fold_named_context_reverse = List.fold_left
let fold_named_context_both_sides = list_fold_right_and_left
let it_named_context_quantifier f = List.fold_left (fun c d -> f d c)
-(*s Signatures of ordered section variables *)
+let it_mkNamedProd_or_LetIn =
+ List.fold_left (fun c d -> mkNamedProd_or_LetIn d c)
+let it_mkNamedLambda_or_LetIn =
+ List.fold_left (fun c d -> mkNamedLambda_or_LetIn d c)
-type section_declaration = variable * constr option * constr
-type section_context = section_declaration list
-let instance_from_section_context sign =
- let rec inst_rec = function
- | (sp,None,_) :: sign -> mkVar (basename sp) :: inst_rec sign
- | _ :: sign -> inst_rec sign
- | [] -> [] in
- Array.of_list (inst_rec sign)
-let instance_from_section_context x =
- instance_from_section_context x
+(*s Signatures of ordered section variables *)
+type section_context = named_context
(*s Signatures of ordered optionally named variables, intended to be
accessed by de Bruijn indices *)
@@ -79,21 +66,20 @@ type rel_declaration = name * constr option * types
type rel_context = rel_declaration list
type rev_rel_context = rel_declaration list
+let fold_rel_context = List.fold_right
+let fold_rel_context_reverse = List.fold_left
+
+let it_mkProd_or_LetIn = List.fold_left (fun c d -> mkProd_or_LetIn d c)
+let it_mkLambda_or_LetIn = List.fold_left (fun c d -> mkLambda_or_LetIn d c)
+
let add_rel_decl = add
let add_rel_assum = add_decl
let add_rel_def = add_def
-let lookup_rel_type n sign =
+let lookup_rel n sign =
let rec lookrec = function
- | (1, (na,_,t) :: _) -> (na,t)
- | (n, _ :: sign) -> lookrec (n-1,sign)
- | (_, []) -> raise Not_found
- in
- lookrec (n,sign)
-let lookup_rel_value n sign =
- let rec lookrec = function
- | (1, (_,c,_) :: _) -> c
- | (n, _ :: sign ) -> lookrec (n-1,sign)
- | (_, []) -> raise Not_found
+ | (1, decl :: _) -> decl
+ | (n, _ :: sign) -> lookrec (n-1,sign)
+ | (_, []) -> raise Not_found
in
lookrec (n,sign)
let rec lookup_rel_id id sign =
@@ -127,10 +113,6 @@ let ids_of_rel_context sign =
(fun (na,_,_) l -> match na with Name id -> id::l | Anonymous -> l)
sign []
let names_of_rel_context = List.map (fun (na,_,_) -> na)
-let assums_of_rel_context sign =
- List.fold_right
- (fun (na,c,t) l -> match c with Some _ -> l | None -> (na,body_of_type t)::l)
- sign []
let map_rel_context = map
let push_named_to_rel_context hyps ctxt =
let rec push = function
@@ -157,7 +139,7 @@ let instantiate_sign sign args =
| ([],_) | (_,[]) ->
anomaly "Signature and its instance do not match"
in
- instrec (sign,args)
+ instrec (sign,Array.to_list args)
(*************************)
(* Names environments *)
@@ -185,9 +167,9 @@ let empty_names_context = []
let decompose_prod_assum =
let rec prodec_rec l c =
match kind_of_term c with
- | IsProd (x,t,c) -> prodec_rec (add_rel_assum (x,t) l) c
- | IsLetIn (x,b,t,c) -> prodec_rec (add_rel_def (x,b,t) l) c
- | IsCast (c,_) -> prodec_rec l c
+ | Prod (x,t,c) -> prodec_rec (add_rel_assum (x,t) l) c
+ | LetIn (x,b,t,c) -> prodec_rec (add_rel_def (x,b,t) l) c
+ | Cast (c,_) -> prodec_rec l c
| _ -> l,c
in
prodec_rec empty_rel_context
@@ -197,9 +179,9 @@ let decompose_prod_assum =
let decompose_lam_assum =
let rec lamdec_rec l c =
match kind_of_term c with
- | IsLambda (x,t,c) -> lamdec_rec (add_rel_assum (x,t) l) c
- | IsLetIn (x,b,t,c) -> lamdec_rec (add_rel_def (x,b,t) l) c
- | IsCast (c,_) -> lamdec_rec l c
+ | Lambda (x,t,c) -> lamdec_rec (add_rel_assum (x,t) l) c
+ | LetIn (x,b,t,c) -> lamdec_rec (add_rel_def (x,b,t) l) c
+ | Cast (c,_) -> lamdec_rec l c
| _ -> l,c
in
lamdec_rec empty_rel_context
@@ -212,10 +194,10 @@ let decompose_prod_n_assum n =
let rec prodec_rec l n c =
if n=0 then l,c
else match kind_of_term c with
- | IsProd (x,t,c) -> prodec_rec (add_rel_assum(x,t) l) (n-1) c
- | IsLetIn (x,b,t,c) ->
+ | Prod (x,t,c) -> prodec_rec (add_rel_assum(x,t) l) (n-1) c
+ | LetIn (x,b,t,c) ->
prodec_rec (add_rel_def (x,b,t) l) (n-1) c
- | IsCast (c,_) -> prodec_rec l n c
+ | Cast (c,_) -> prodec_rec l n c
| c -> error "decompose_prod_n_assum: not enough assumptions"
in
prodec_rec empty_rel_context n
@@ -228,10 +210,10 @@ let decompose_lam_n_assum n =
let rec lamdec_rec l n c =
if n=0 then l,c
else match kind_of_term c with
- | IsLambda (x,t,c) -> lamdec_rec (add_rel_assum (x,t) l) (n-1) c
- | IsLetIn (x,b,t,c) ->
+ | Lambda (x,t,c) -> lamdec_rec (add_rel_assum (x,t) l) (n-1) c
+ | LetIn (x,b,t,c) ->
lamdec_rec (add_rel_def (x,b,t) l) (n-1) c
- | IsCast (c,_) -> lamdec_rec l n c
+ | Cast (c,_) -> lamdec_rec l n c
| c -> error "decompose_lam_n_assum: not enough abstractions"
in
lamdec_rec empty_rel_context n
diff --git a/kernel/sign.mli b/kernel/sign.mli
index dd5aba6c6..d834e263a 100644
--- a/kernel/sign.mli
+++ b/kernel/sign.mli
@@ -16,38 +16,25 @@ open Term
(*s Signatures of ordered named declarations *)
type named_context = named_declaration list
+type section_context = named_context
-val add_named_decl :
- identifier * constr option * types -> named_context -> named_context
-val add_named_assum : identifier * types -> named_context -> named_context
-val add_named_def :
- identifier * constr * types -> named_context -> named_context
-val lookup_id : identifier -> named_context -> constr option * types
-val lookup_id_type : identifier -> named_context -> types
-val lookup_id_value : identifier -> named_context -> constr option
-val pop_named_decl : identifier -> named_context -> named_context
val empty_named_context : named_context
-val ids_of_named_context : named_context -> identifier list
-val map_named_context : (constr -> constr) -> named_context -> named_context
-val mem_named_context : identifier -> named_context -> bool
+val add_named_decl : named_declaration -> named_context -> named_context
+val pop_named_decl : identifier -> named_context -> named_context
+
+val lookup_named : identifier -> named_context -> named_declaration
+
+(*s Recurrence on [named_context]: older declarations processed first *)
val fold_named_context :
(named_declaration -> 'a -> 'a) -> named_context -> 'a -> 'a
+(* newer declarations first *)
val fold_named_context_reverse :
('a -> named_declaration -> 'a) -> 'a -> named_context -> 'a
-val fold_named_context_both_sides :
- ('a -> named_declaration -> named_context -> 'a) -> named_context -> 'a -> 'a
-val it_named_context_quantifier :
- (named_declaration -> constr -> constr) -> constr -> named_context -> constr
-val instantiate_sign :
- named_context -> constr list -> (identifier * constr) list
-val instance_from_named_context : named_context -> constr array
-
-(*s Signatures of ordered section variables *)
-
-type section_declaration = variable * constr option * constr
-type section_context = section_declaration list
-val instance_from_section_context : section_context -> constr array
+(*s Section-related auxiliary functions *)
+val instance_from_named_context : named_context -> constr array
+val instantiate_sign :
+ named_context -> constr array -> (identifier * constr) list
(*s Signatures of ordered optionally named variables, intended to be
accessed by de Bruijn indices *)
@@ -55,35 +42,28 @@ val instance_from_section_context : section_context -> constr array
(* In [rel_context], more recent declaration is on top *)
type rel_context = rel_declaration list
-(* In [rev_rel_context], older declaration is on top *)
-type rev_rel_context = rel_declaration list
-
-val add_rel_decl : (name * constr option * types) -> rel_context -> rel_context
-val add_rel_assum : (name * types) -> rel_context -> rel_context
-val add_rel_def : (name * constr * types) -> rel_context -> rel_context
-val lookup_rel_type : int -> rel_context -> name * types
-val lookup_rel_value : int -> rel_context -> constr option
-val lookup_rel_id : identifier -> rel_context -> int * types
val empty_rel_context : rel_context
+val add_rel_decl : rel_declaration -> rel_context -> rel_context
+
+val lookup_rel : int -> rel_context -> rel_declaration
val rel_context_length : rel_context -> int
-val lift_rel_context : int -> rel_context -> rel_context
-val lift_rev_rel_context : int -> rev_rel_context -> rev_rel_context
-val concat_rel_context : newer:rel_context -> older:rel_context -> rel_context
-val ids_of_rel_context : rel_context -> identifier list
-val assums_of_rel_context : rel_context -> (name * constr) list
-val map_rel_context : (constr -> constr) -> rel_context -> rel_context
+
val push_named_to_rel_context : named_context -> rel_context -> rel_context
-val reverse_rel_context : rel_context -> rev_rel_context
-(*s This is used to translate names into de Bruijn indices and
- vice-versa without to care about typing information *)
+(*s Recurrence on [rel_context]: older declarations processed first *)
+val fold_rel_context :
+ (rel_declaration -> 'a -> 'a) -> rel_context -> 'a -> 'a
+(* newer declarations first *)
+val fold_rel_context_reverse :
+ ('a -> rel_declaration -> 'a) -> 'a -> rel_context -> 'a
+
+(*s Term constructors *)
+
+val it_mkNamedLambda_or_LetIn : constr -> named_context -> constr
+val it_mkNamedProd_or_LetIn : constr -> named_context -> constr
-type names_context
-val add_name : name -> names_context -> names_context
-val lookup_name_of_rel : int -> names_context -> name
-val lookup_rel_of_name : identifier -> names_context -> int
-val names_of_rel_context : rel_context -> names_context
-val empty_names_context : names_context
+val it_mkLambda_or_LetIn : constr -> rel_context -> constr
+val it_mkProd_or_LetIn : constr -> rel_context -> constr
(*s Term destructors *)
diff --git a/kernel/term.ml b/kernel/term.ml
index 96a4d0d38..652c4d3c3 100644
--- a/kernel/term.ml
+++ b/kernel/term.ml
@@ -24,9 +24,15 @@ type existential_key = int
type pattern_source = DefaultPat of int | RegularPat
type case_style = PrintLet | PrintIf | PrintCases
type case_printing =
- inductive * identifier array * int
- * case_style option * pattern_source array
-type case_info = int * case_printing
+ { cnames : identifier array;
+ ind_nargs : int; (* number of real args of the inductive type *)
+ style : case_style option;
+ source : pattern_source array }
+type case_info =
+ { ci_ind : inductive;
+ ci_npar : int;
+ ci_pp_info : case_printing (* not interpreted by the kernel *)
+ }
(* Sorts. *)
@@ -39,12 +45,6 @@ type sorts =
let mk_Set = Prop Pos
let mk_Prop = Prop Null
-let print_sort = function
- | Prop Pos -> [< 'sTR "Set" >]
- | Prop Null -> [< 'sTR "Prop" >]
-(* | Type _ -> [< 'sTR "Type" >] *)
- | Type u -> [< 'sTR "Type("; pr_uni u; 'sTR ")" >]
-
type sorts_family = InProp | InSet | InType
let new_sort_in_family = function
@@ -76,22 +76,22 @@ type fixpoint = (int array * int) * rec_declaration
type cofixpoint = int * rec_declaration
type kind_of_term =
- | IsRel of int
- | IsMeta of int
- | IsVar of identifier
- | IsSort of sorts
- | IsCast of constr * constr
- | IsProd of name * constr * constr
- | IsLambda of name * constr * constr
- | IsLetIn of name * constr * constr * constr
- | IsApp of constr * constr array
- | IsEvar of existential
- | IsConst of constant
- | IsMutInd of inductive
- | IsMutConstruct of constructor
- | IsMutCase of case_info * constr * constr * constr array
- | IsFix of fixpoint
- | IsCoFix of cofixpoint
+ | Rel of int
+ | Meta of int
+ | Var of identifier
+ | Sort of sorts
+ | Cast of constr * constr
+ | Prod of name * constr * constr
+ | Lambda of name * constr * constr
+ | LetIn of name * constr * constr * constr
+ | App of constr * constr array
+ | Evar of existential
+ | Const of constant
+ | Ind of inductive
+ | Construct of constructor
+ | Case of case_info * constr * constr * constr array
+ | Fix of fixpoint
+ | CoFix of cofixpoint
val mkRel : int -> constr
val mkMeta : int -> constr
@@ -126,42 +126,38 @@ module Internal : InternalSig =
struct
*)
-module Polymorph =
-struct
(* [constr array] is an instance matching definitional [named_context] in
the same order (i.e. last argument first) *)
-type 'constr existential = existential_key * 'constr array
-type ('constr, 'types) rec_declaration =
+type 'constr pexistential = existential_key * 'constr array
+type ('constr, 'types) prec_declaration =
name array * 'types array * 'constr array
-type ('constr, 'types) fixpoint =
- (int array * int) * ('constr, 'types) rec_declaration
-type ('constr, 'types) cofixpoint =
- int * ('constr, 'types) rec_declaration
+type ('constr, 'types) pfixpoint =
+ (int array * int) * ('constr, 'types) prec_declaration
+type ('constr, 'types) pcofixpoint =
+ int * ('constr, 'types) prec_declaration
-(* [IsVar] is used for named variables and [IsRel] for variables as
+(* [Var] is used for named variables and [Rel] for variables as
de Bruijn indices. *)
-
-end
-open Polymorph
-
type ('constr, 'types) kind_of_term =
- | IsRel of int
- | IsMeta of int
- | IsVar of identifier
- | IsSort of sorts
- | IsCast of 'constr * 'constr
- | IsProd of name * 'types * 'constr
- | IsLambda of name * 'types * 'constr
- | IsLetIn of name * 'constr * 'types * 'constr
- | IsApp of 'constr * 'constr array
- | IsEvar of 'constr existential
- | IsConst of constant
- | IsMutInd of inductive
- | IsMutConstruct of constructor
- | IsMutCase of case_info * 'constr * 'constr * 'constr array
- | IsFix of ('constr, 'types) fixpoint
- | IsCoFix of ('constr, 'types) cofixpoint
-
+ | Rel of int
+ | Var of identifier
+ | Meta of int
+ | Evar of 'constr pexistential
+ | Sort of sorts
+ | Cast of 'constr * 'constr
+ | Prod of name * 'types * 'types
+ | Lambda of name * 'types * 'constr
+ | LetIn of name * 'constr * 'types * 'constr
+ | App of 'constr * 'constr array
+ | Const of constant
+ | Ind of inductive
+ | Construct of constructor
+ | Case of case_info * 'constr * 'constr * 'constr array
+ | Fix of ('constr, 'types) pfixpoint
+ | CoFix of ('constr, 'types) pcofixpoint
+
+(* constr is the fixpoint of the previous type. Requires option
+ -rectypes of the Caml compiler to be set *)
type constr = (constr,constr) kind_of_term
type existential = existential_key * constr array
@@ -175,28 +171,28 @@ type cofixpoint = int * rec_declaration
let comp_term t1 t2 =
match t1, t2 with
- | IsRel n1, IsRel n2 -> n1 = n2
- | IsMeta m1, IsMeta m2 -> m1 = m2
- | IsVar id1, IsVar id2 -> id1 == id2
- | IsSort s1, IsSort s2 -> s1 == s2
- | IsCast (c1,t1), IsCast (c2,t2) -> c1 == c2 & t1 == t2
- | IsProd (n1,t1,c1), IsProd (n2,t2,c2) -> n1 == n2 & t1 == t2 & c1 == c2
- | IsLambda (n1,t1,c1), IsLambda (n2,t2,c2) -> n1 == n2 & t1 == t2 & c1 == c2
- | IsLetIn (n1,b1,t1,c1), IsLetIn (n2,b2,t2,c2) ->
+ | Rel n1, Rel n2 -> n1 = n2
+ | Meta m1, Meta m2 -> m1 = m2
+ | Var id1, Var id2 -> id1 == id2
+ | Sort s1, Sort s2 -> s1 == s2
+ | Cast (c1,t1), Cast (c2,t2) -> c1 == c2 & t1 == t2
+ | Prod (n1,t1,c1), Prod (n2,t2,c2) -> n1 == n2 & t1 == t2 & c1 == c2
+ | 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
- | IsApp (c1,l1), IsApp (c2,l2) -> c1 == c2 & array_for_all2 (==) l1 l2
- | IsEvar (e1,l1), IsEvar (e2,l2) -> e1 = e2 & array_for_all2 (==) l1 l2
- | IsConst c1, IsConst c2 -> c1 == c2
- | IsMutInd c1, IsMutInd c2 -> c1 == c2
- | IsMutConstruct c1, IsMutConstruct c2 -> c1 == c2
- | IsMutCase (ci1,p1,c1,bl1), IsMutCase (ci2,p2,c2,bl2) ->
+ | App (c1,l1), App (c2,l2) -> c1 == c2 & array_for_all2 (==) l1 l2
+ | Evar (e1,l1), Evar (e2,l2) -> e1 = e2 & array_for_all2 (==) l1 l2
+ | Const c1, Const c2 -> c1 == c2
+ | Ind c1, Ind c2 -> c1 == c2
+ | Construct c1, Construct c2 -> c1 == c2
+ | Case (ci1,p1,c1,bl1), Case (ci2,p2,c2,bl2) ->
ci1 == ci2 & p1 == p2 & c1 == c2 & array_for_all2 (==) bl1 bl2
- | IsFix (ln1,(lna1,tl1,bl1)), IsFix (ln2,(lna2,tl2,bl2)) ->
+ | Fix (ln1,(lna1,tl1,bl1)), Fix (ln2,(lna2,tl2,bl2)) ->
ln1 = ln2
& array_for_all2 (==) lna1 lna2
& array_for_all2 (==) tl1 tl2
& array_for_all2 (==) bl1 bl2
- | IsCoFix(ln1,(lna1,tl1,bl1)), IsCoFix(ln2,(lna2,tl2,bl2)) ->
+ | CoFix(ln1,(lna1,tl1,bl1)), CoFix(ln2,(lna2,tl2,bl2)) ->
ln1 = ln2
& array_for_all2 (==) lna1 lna2
& array_for_all2 (==) tl1 tl2
@@ -205,26 +201,26 @@ let comp_term t1 t2 =
let hash_term (sh_rec,(sh_sort,sh_sp,sh_na,sh_id)) t =
match t with
- | IsRel _ | IsMeta _ -> t
- | IsVar x -> IsVar (sh_id x)
- | IsSort s -> IsSort (sh_sort s)
- | IsCast (c,t) -> IsCast (sh_rec c, sh_rec t)
- | IsProd (na,t,c) -> IsProd (sh_na na, sh_rec t, sh_rec c)
- | IsLambda (na,t,c) -> IsLambda (sh_na na, sh_rec t, sh_rec c)
- | IsLetIn (na,b,t,c) -> IsLetIn (sh_na na, sh_rec b, sh_rec t, sh_rec c)
- | IsApp (c,l) -> IsApp (sh_rec c, Array.map sh_rec l)
- | IsEvar (e,l) -> IsEvar (e, Array.map sh_rec l)
- | IsConst c -> IsConst (sh_sp c)
- | IsMutInd (sp,i) -> IsMutInd (sh_sp sp,i)
- | IsMutConstruct ((sp,i),j) -> IsMutConstruct ((sh_sp sp,i),j)
- | IsMutCase (ci,p,c,bl) -> (* TO DO: extract ind_sp *)
- IsMutCase (ci, sh_rec p, sh_rec c, Array.map sh_rec bl)
- | IsFix (ln,(lna,tl,bl)) ->
- IsFix (ln,(Array.map sh_na lna,
+ | Rel _ | Meta _ -> t
+ | Var x -> Var (sh_id x)
+ | Sort s -> Sort (sh_sort s)
+ | Cast (c,t) -> Cast (sh_rec c, sh_rec t)
+ | Prod (na,t,c) -> Prod (sh_na na, sh_rec t, sh_rec c)
+ | Lambda (na,t,c) -> Lambda (sh_na na, sh_rec t, sh_rec c)
+ | LetIn (na,b,t,c) -> LetIn (sh_na na, sh_rec b, sh_rec t, sh_rec c)
+ | App (c,l) -> App (sh_rec c, Array.map sh_rec l)
+ | Evar (e,l) -> Evar (e, Array.map sh_rec l)
+ | Const c -> Const (sh_sp c)
+ | Ind (sp,i) -> Ind (sh_sp sp,i)
+ | Construct ((sp,i),j) -> Construct ((sh_sp sp,i),j)
+ | Case (ci,p,c,bl) -> (* TO DO: extract ind_sp *)
+ Case (ci, sh_rec p, sh_rec c, Array.map sh_rec bl)
+ | Fix (ln,(lna,tl,bl)) ->
+ Fix (ln,(Array.map sh_na lna,
Array.map sh_rec tl,
Array.map sh_rec bl))
- | IsCoFix(ln,(lna,tl,bl)) ->
- IsCoFix (ln,(Array.map sh_na lna,
+ | CoFix(ln,(lna,tl,bl)) ->
+ CoFix (ln,(Array.map sh_na lna,
Array.map sh_rec tl,
Array.map sh_rec bl))
@@ -244,43 +240,36 @@ let hcons_term (hsorts,hsp,hname,hident) =
Hashcons.recursive_hcons Hconstr.f (hsorts,hsp,hname,hident)
(* Constructs a DeBrujin index with number n *)
-let mkRel n = IsRel n
-
-let r = ref None
+let rels =
+ [|Rel 1;Rel 2;Rel 3;Rel 4;Rel 5;Rel 6;Rel 7; Rel 8;
+ Rel 9;Rel 10;Rel 11;Rel 12;Rel 13;Rel 14;Rel 15; Rel 16|]
-let mkRel n =
- let rels = match !r with
- | None -> let a =
- [|mkRel 1;mkRel 2;mkRel 3;mkRel 4;mkRel 5;mkRel 6;mkRel 7; mkRel 8;
- mkRel 9;mkRel 10;mkRel 11;mkRel 12;mkRel 13;mkRel 14;mkRel 15; mkRel 16|]
- in r := Some a; a
- | Some a -> a in
- if 0<n & n<=16 then rels.(n-1) else mkRel n
+let mkRel n = if 0<n & n<=16 then rels.(n-1) else Rel n
(* Constructs an existential variable named "?n" *)
-let mkMeta n = IsMeta n
+let mkMeta n = Meta n
(* Constructs a Variable named id *)
-let mkVar id = IsVar id
+let mkVar id = Var id
(* Construct a type *)
-let mkSort s = IsSort s
+let mkSort s = Sort s
(* Constructs the term t1::t2, i.e. the term t1 casted with the type t2 *)
(* (that means t2 is declared as the type of t1) *)
let mkCast (t1,t2) =
match t1 with
- | IsCast (t,_) -> IsCast (t,t2)
- | _ -> IsCast (t1,t2)
+ | Cast (t,_) -> Cast (t,t2)
+ | _ -> Cast (t1,t2)
(* Constructs the product (x:t1)t2 *)
-let mkProd (x,t1,t2) = IsProd (x,t1,t2)
+let mkProd (x,t1,t2) = Prod (x,t1,t2)
(* Constructs the abstraction [x:t1]t2 *)
-let mkLambda (x,t1,t2) = IsLambda (x,t1,t2)
+let mkLambda (x,t1,t2) = Lambda (x,t1,t2)
(* Constructs [x=c_1:t]c_2 *)
-let mkLetIn (x,c1,t,c2) = IsLetIn (x,c1,t,c2)
+let mkLetIn (x,c1,t,c2) = LetIn (x,c1,t,c2)
(* If lt = [t1; ...; tn], constructs the application (t1 ... tn) *)
(* We ensure applicative terms have at least one argument and the
@@ -288,32 +277,32 @@ let mkLetIn (x,c1,t,c2) = IsLetIn (x,c1,t,c2)
let mkApp (f, a) =
if a=[||] then f else
match f with
- | IsApp (g, cl) -> IsApp (g, Array.append cl a)
- | _ -> IsApp (f, a)
+ | App (g, cl) -> App (g, Array.append cl a)
+ | _ -> App (f, a)
(* Constructs a constant *)
(* The array of terms correspond to the variables introduced in the section *)
-let mkConst c = IsConst c
+let mkConst c = Const c
(* Constructs an existential variable *)
-let mkEvar e = IsEvar e
+let mkEvar e = Evar e
(* Constructs the ith (co)inductive type of the block named sp *)
(* The array of terms correspond to the variables introduced in the section *)
-let mkMutInd m = IsMutInd m
+let mkInd m = Ind m
(* Constructs the jth constructor of the ith (co)inductive type of the
block named sp. The array of terms correspond to the variables
introduced in the section *)
-let mkMutConstruct c = IsMutConstruct c
+let mkConstruct c = Construct c
(* Constructs the term <p>Case c of c1 | c2 .. | cn end *)
-let mkMutCase (ci, p, c, ac) = IsMutCase (ci, p, c, ac)
+let mkCase (ci, p, c, ac) = Case (ci, p, c, ac)
-let mkFix fix = IsFix fix
+let mkFix fix = Fix fix
-let mkCoFix cofix = IsCoFix cofix
+let mkCoFix cofix = CoFix cofix
let kind_of_term c = c
@@ -341,7 +330,7 @@ open Internal
END of expected re-export of Internal module *)
-(* User view of [constr]. For [IsApp], it is ensured there is at
+(* User view of [constr]. For [App], it is ensured there is at
least one argument and the function is not itself an applicative
term *)
@@ -353,7 +342,7 @@ type hnftype =
| HnfSort of sorts
| HnfProd of name * constr * constr
| HnfAtom of constr
- | HnfMutInd of inductive * constr array
+ | HnfInd of inductive * constr array
(**********************************************************************)
(* Non primitive term destructors *)
@@ -364,48 +353,48 @@ type hnftype =
(* Destructs a DeBrujin index *)
let destRel c = match kind_of_term c with
- | IsRel n -> n
+ | Rel n -> n
| _ -> invalid_arg "destRel"
(* Destructs an existential variable *)
let destMeta c = match kind_of_term c with
- | IsMeta n -> n
+ | Meta n -> n
| _ -> invalid_arg "destMeta"
-let isMeta c = match kind_of_term c with IsMeta _ -> true | _ -> false
+let isMeta c = match kind_of_term c with Meta _ -> true | _ -> false
(* Destructs a variable *)
let destVar c = match kind_of_term c with
- | IsVar id -> id
+ | Var id -> id
| _ -> invalid_arg "destVar"
(* Destructs a type *)
let isSort c = match kind_of_term c with
- | IsSort s -> true
+ | Sort s -> true
| _ -> false
let destSort c = match kind_of_term c with
- | IsSort s -> s
+ | Sort s -> s
| _ -> invalid_arg "destSort"
let rec isprop c = match kind_of_term c with
- | IsSort (Prop _) -> true
- | IsCast (c,_) -> isprop c
+ | Sort (Prop _) -> true
+ | Cast (c,_) -> isprop c
| _ -> false
let rec is_Prop c = match kind_of_term c with
- | IsSort (Prop Null) -> true
- | IsCast (c,_) -> is_Prop c
+ | Sort (Prop Null) -> true
+ | Cast (c,_) -> is_Prop c
| _ -> false
let rec is_Set c = match kind_of_term c with
- | IsSort (Prop Pos) -> true
- | IsCast (c,_) -> is_Set c
+ | Sort (Prop Pos) -> true
+ | Cast (c,_) -> is_Set c
| _ -> false
let rec is_Type c = match kind_of_term c with
- | IsSort (Type _) -> true
- | IsCast (c,_) -> is_Type c
+ | Sort (Type _) -> true
+ | Cast (c,_) -> is_Type c
| _ -> false
let isType = function
@@ -422,79 +411,79 @@ let same_kind c1 c2 = (isprop c1 & isprop c2) or (is_Type c1 & is_Type c2)
(* Destructs a casted term *)
let destCast c = match kind_of_term c with
- | IsCast (t1, t2) -> (t1,t2)
+ | Cast (t1, t2) -> (t1,t2)
| _ -> invalid_arg "destCast"
-let isCast c = match kind_of_term c with IsCast (_,_) -> true | _ -> false
+let isCast c = match kind_of_term c with Cast (_,_) -> true | _ -> false
(* Tests if a de Bruijn index *)
-let isRel c = match kind_of_term c with IsRel _ -> true | _ -> false
+let isRel c = match kind_of_term c with Rel _ -> true | _ -> false
(* Tests if a variable *)
-let isVar c = match kind_of_term c with IsVar _ -> true | _ -> false
+let isVar c = match kind_of_term c with Var _ -> true | _ -> false
(* Destructs the product (x:t1)t2 *)
let destProd c = match kind_of_term c with
- | IsProd (x,t1,t2) -> (x,t1,t2)
+ | Prod (x,t1,t2) -> (x,t1,t2)
| _ -> invalid_arg "destProd"
(* Destructs the abstraction [x:t1]t2 *)
let destLambda c = match kind_of_term c with
- | IsLambda (x,t1,t2) -> (x,t1,t2)
+ | Lambda (x,t1,t2) -> (x,t1,t2)
| _ -> invalid_arg "destLambda"
(* Destructs the let [x:=b:t1]t2 *)
let destLetIn c = match kind_of_term c with
- | IsLetIn (x,b,t1,t2) -> (x,b,t1,t2)
+ | LetIn (x,b,t1,t2) -> (x,b,t1,t2)
| _ -> invalid_arg "destProd"
(* Destructs an application *)
let destApplication c = match kind_of_term c with
- | IsApp (f,a) -> (f, a)
+ | App (f,a) -> (f, a)
| _ -> invalid_arg "destApplication"
-let isApp c = match kind_of_term c with IsApp _ -> true | _ -> false
+let isApp c = match kind_of_term c with App _ -> true | _ -> false
(* Destructs a constant *)
let destConst c = match kind_of_term c with
- | IsConst sp -> sp
+ | Const sp -> sp
| _ -> invalid_arg "destConst"
-let isConst c = match kind_of_term c with IsConst _ -> true | _ -> false
+let isConst c = match kind_of_term c with Const _ -> true | _ -> false
(* Destructs an existential variable *)
let destEvar c = match kind_of_term c with
- | IsEvar (sp, a as r) -> r
+ | Evar (sp, a as r) -> r
| _ -> invalid_arg "destEvar"
let num_of_evar c = match kind_of_term c with
- | IsEvar (n, _) -> n
+ | Evar (n, _) -> n
| _ -> anomaly "num_of_evar called with bad args"
(* Destructs a (co)inductive type named sp *)
-let destMutInd c = match kind_of_term c with
- | IsMutInd (sp, a as r) -> r
- | _ -> invalid_arg "destMutInd"
+let destInd c = match kind_of_term c with
+ | Ind (sp, a as r) -> r
+ | _ -> invalid_arg "destInd"
(* Destructs a constructor *)
-let destMutConstruct c = match kind_of_term c with
- | IsMutConstruct (sp, a as r) -> r
+let destConstruct c = match kind_of_term c with
+ | Construct (sp, a as r) -> r
| _ -> invalid_arg "dest"
-let isMutConstruct c = match kind_of_term c with
- IsMutConstruct _ -> true | _ -> false
+let isConstruct c = match kind_of_term c with
+ Construct _ -> true | _ -> false
(* Destructs a term <p>Case c of lc1 | lc2 .. | lcn end *)
let destCase c = match kind_of_term c with
- | IsMutCase (ci,p,c,v) -> (ci,p,c,v)
+ | Case (ci,p,c,v) -> (ci,p,c,v)
| _ -> anomaly "destCase"
let destFix c = match kind_of_term c with
- | IsFix fix -> fix
+ | Fix fix -> fix
| _ -> invalid_arg "destFix"
let destCoFix c = match kind_of_term c with
- | IsCoFix cofix -> cofix
+ | CoFix cofix -> cofix
| _ -> invalid_arg "destCoFix"
(******************************************************************)
@@ -503,31 +492,31 @@ let destCoFix c = match kind_of_term c with
(* flattens application lists *)
let rec collapse_appl c = match kind_of_term c with
- | IsApp (f,cl) ->
+ | App (f,cl) ->
let rec collapse_rec f cl2 = match kind_of_term f with
- | IsApp (g,cl1) -> collapse_rec g (Array.append cl1 cl2)
- | IsCast (c,_) when isApp c -> collapse_rec c cl2
+ | App (g,cl1) -> collapse_rec g (Array.append cl1 cl2)
+ | Cast (c,_) when isApp c -> collapse_rec c cl2
| _ -> if cl2 = [||] then f else mkApp (f,cl2)
in
collapse_rec f cl
| _ -> c
-let rec decomp_app c =
+let rec decompose_app c =
match kind_of_term (collapse_appl c) with
- | IsApp (f,cl) -> (f, Array.to_list cl)
- | IsCast (c,t) -> decomp_app c
+ | App (f,cl) -> (f, Array.to_list cl)
+ | Cast (c,t) -> decompose_app c
| _ -> (c,[])
(* strips head casts and flattens head applications *)
let rec strip_head_cast c = match kind_of_term c with
- | IsApp (f,cl) ->
+ | App (f,cl) ->
let rec collapse_rec f cl2 = match kind_of_term f with
- | IsApp (g,cl1) -> collapse_rec g (Array.append cl1 cl2)
- | IsCast (c,_) -> collapse_rec c cl2
+ | App (g,cl1) -> collapse_rec g (Array.append cl1 cl2)
+ | Cast (c,_) -> collapse_rec c cl2
| _ -> if cl2 = [||] then f else mkApp (f,cl2)
in
collapse_rec f cl
- | IsCast (c,t) -> strip_head_cast c
+ | Cast (c,t) -> strip_head_cast c
| _ -> c
(****************************************************************************)
@@ -539,19 +528,19 @@ let rec strip_head_cast c = match kind_of_term c with
the usual representation of the constructions; it is not recursive *)
let fold_constr f acc c = match kind_of_term c with
- | (IsRel _ | IsMeta _ | IsVar _ | IsSort _ | IsConst _ | IsMutInd _
- | IsMutConstruct _) -> acc
- | IsCast (c,t) -> f (f acc c) t
- | IsProd (_,t,c) -> f (f acc t) c
- | IsLambda (_,t,c) -> f (f acc t) c
- | IsLetIn (_,b,t,c) -> f (f (f acc b) t) c
- | IsApp (c,l) -> Array.fold_left f (f acc c) l
- | IsEvar (_,l) -> Array.fold_left f acc l
- | IsMutCase (_,p,c,bl) -> Array.fold_left f (f (f acc p) c) bl
- | IsFix (_,(lna,tl,bl)) ->
+ | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
+ | Construct _) -> acc
+ | Cast (c,t) -> f (f acc c) t
+ | Prod (_,t,c) -> f (f acc t) c
+ | 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
+ | 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)) ->
let fd = array_map3 (fun na t b -> (na,t,b)) lna tl bl in
Array.fold_left (fun acc (na,t,b) -> f (f acc t) b) acc fd
- | IsCoFix (_,(lna,tl,bl)) ->
+ | CoFix (_,(lna,tl,bl)) ->
let fd = array_map3 (fun na t b -> (na,t,b)) lna tl bl in
Array.fold_left (fun acc (na,t,b) -> f (f acc t) b) acc fd
@@ -563,20 +552,20 @@ let fold_constr f acc c = match kind_of_term c with
each binder traversal; it is not recursive *)
let fold_constr_with_binders g f n acc c = match kind_of_term c with
- | (IsRel _ | IsMeta _ | IsVar _ | IsSort _ | IsConst _ | IsMutInd _
- | IsMutConstruct _) -> acc
- | IsCast (c,t) -> f n (f n acc c) t
- | IsProd (_,t,c) -> f (g n) (f n acc t) c
- | IsLambda (_,t,c) -> f (g n) (f n acc t) c
- | IsLetIn (_,b,t,c) -> f (g n) (f n (f n acc b) t) c
- | IsApp (c,l) -> Array.fold_left (f n) (f n acc c) l
- | IsEvar (_,l) -> Array.fold_left (f n) acc l
- | IsMutCase (_,p,c,bl) -> Array.fold_left (f n) (f n (f n acc p) c) bl
- | IsFix (_,(lna,tl,bl)) ->
+ | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
+ | Construct _) -> acc
+ | Cast (c,t) -> f n (f n acc c) t
+ | Prod (_,t,c) -> f (g n) (f n acc t) c
+ | Lambda (_,t,c) -> f (g n) (f n acc t) c
+ | LetIn (_,b,t,c) -> f (g n) (f n (f n acc b) t) c
+ | App (c,l) -> Array.fold_left (f n) (f n acc c) l
+ | Evar (_,l) -> Array.fold_left (f n) acc l
+ | Case (_,p,c,bl) -> Array.fold_left (f n) (f n (f n acc p) c) bl
+ | Fix (_,(lna,tl,bl)) ->
let n' = iterate g (Array.length tl) n in
let fd = array_map2 (fun t b -> (t,b)) tl bl in
Array.fold_left (fun acc (t,b) -> f n (f n' acc t) b) acc fd
- | IsCoFix (_,(lna,tl,bl)) ->
+ | CoFix (_,(lna,tl,bl)) ->
let n' = iterate g (Array.length tl) n in
let fd = array_map2 (fun t b -> (t,b)) tl bl in
Array.fold_left (fun acc (t,b) -> f n (f n' acc t) b) acc fd
@@ -586,17 +575,17 @@ let fold_constr_with_binders g f n acc c = match kind_of_term c with
not specified *)
let iter_constr f c = match kind_of_term c with
- | (IsRel _ | IsMeta _ | IsVar _ | IsSort _ | IsConst _ | IsMutInd _
- | IsMutConstruct _) -> ()
- | IsCast (c,t) -> f c; f t
- | IsProd (_,t,c) -> f t; f c
- | IsLambda (_,t,c) -> f t; f c
- | IsLetIn (_,b,t,c) -> f b; f t; f c
- | IsApp (c,l) -> f c; Array.iter f l
- | IsEvar (_,l) -> Array.iter f l
- | IsMutCase (_,p,c,bl) -> f p; f c; Array.iter f bl
- | IsFix (_,(_,tl,bl)) -> Array.iter f tl; Array.iter f bl
- | IsCoFix (_,(_,tl,bl)) -> Array.iter f tl; Array.iter f bl
+ | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
+ | Construct _) -> ()
+ | Cast (c,t) -> f c; f t
+ | Prod (_,t,c) -> f t; f c
+ | 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
+ | 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
+ | CoFix (_,(_,tl,bl)) -> Array.iter f tl; Array.iter f bl
(* [iter_constr_with_binders g f n c] iters [f n] on the immediate
subterms of [c]; it carries an extra data [n] (typically a lift
@@ -605,19 +594,19 @@ let iter_constr f c = match kind_of_term c with
subterms are processed is not specified *)
let iter_constr_with_binders g f n c = match kind_of_term c with
- | (IsRel _ | IsMeta _ | IsVar _ | IsSort _ | IsConst _ | IsMutInd _
- | IsMutConstruct _) -> ()
- | IsCast (c,t) -> f n c; f n t
- | IsProd (_,t,c) -> f n t; f (g n) c
- | IsLambda (_,t,c) -> f n t; f (g n) c
- | IsLetIn (_,b,t,c) -> f n b; f n t; f (g n) c
- | IsApp (c,l) -> f n c; Array.iter (f n) l
- | IsEvar (_,l) -> Array.iter (f n) l
- | IsMutCase (_,p,c,bl) -> f n p; f n c; Array.iter (f n) bl
- | IsFix (_,(_,tl,bl)) ->
+ | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
+ | Construct _) -> ()
+ | Cast (c,t) -> f n c; f n t
+ | Prod (_,t,c) -> f n t; f (g n) c
+ | Lambda (_,t,c) -> f n t; f (g n) c
+ | LetIn (_,b,t,c) -> f n b; f n t; f (g n) c
+ | App (c,l) -> f n c; Array.iter (f n) l
+ | Evar (_,l) -> Array.iter (f n) l
+ | Case (_,p,c,bl) -> f n p; f n c; Array.iter (f n) bl
+ | Fix (_,(_,tl,bl)) ->
Array.iter (f n) tl;
Array.iter (f (iterate g (Array.length tl) n)) bl
- | IsCoFix (_,(_,tl,bl)) ->
+ | CoFix (_,(_,tl,bl)) ->
Array.iter (f n) tl;
Array.iter (f (iterate g (Array.length tl) n)) bl
@@ -626,18 +615,18 @@ let iter_constr_with_binders g f n c = match kind_of_term c with
not specified *)
let map_constr f c = match kind_of_term c with
- | (IsRel _ | IsMeta _ | IsVar _ | IsSort _ | IsConst _ | IsMutInd _
- | IsMutConstruct _) -> c
- | IsCast (c,t) -> mkCast (f c, f t)
- | IsProd (na,t,c) -> mkProd (na, f t, f c)
- | IsLambda (na,t,c) -> mkLambda (na, f t, f c)
- | IsLetIn (na,b,t,c) -> mkLetIn (na, f b, f t, f c)
- | IsApp (c,l) -> mkApp (f c, Array.map f l)
- | IsEvar (e,l) -> mkEvar (e, Array.map f l)
- | IsMutCase (ci,p,c,bl) -> mkMutCase (ci, f p, f c, Array.map f bl)
- | IsFix (ln,(lna,tl,bl)) ->
+ | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
+ | Construct _) -> c
+ | Cast (c,t) -> mkCast (f c, f t)
+ | Prod (na,t,c) -> mkProd (na, f t, f c)
+ | Lambda (na,t,c) -> mkLambda (na, f t, f c)
+ | LetIn (na,b,t,c) -> mkLetIn (na, f b, f t, f c)
+ | App (c,l) -> mkApp (f c, Array.map f l)
+ | Evar (e,l) -> mkEvar (e, Array.map f l)
+ | Case (ci,p,c,bl) -> mkCase (ci, f p, f c, Array.map f bl)
+ | Fix (ln,(lna,tl,bl)) ->
mkFix (ln,(lna,Array.map f tl,Array.map f bl))
- | IsCoFix(ln,(lna,tl,bl)) ->
+ | CoFix(ln,(lna,tl,bl)) ->
mkCoFix (ln,(lna,Array.map f tl,Array.map f bl))
(* [map_constr_with_binders g f n c] maps [f n] on the immediate
@@ -647,118 +636,20 @@ let map_constr f c = match kind_of_term c with
subterms are processed is not specified *)
let map_constr_with_binders g f l c = match kind_of_term c with
- | (IsRel _ | IsMeta _ | IsVar _ | IsSort _ | IsConst _ | IsMutInd _
- | IsMutConstruct _) -> c
- | IsCast (c,t) -> mkCast (f l c, f l t)
- | IsProd (na,t,c) -> mkProd (na, f l t, f (g l) c)
- | IsLambda (na,t,c) -> mkLambda (na, f l t, f (g l) c)
- | IsLetIn (na,b,t,c) -> mkLetIn (na, f l b, f l t, f (g l) c)
- | IsApp (c,al) -> mkApp (f l c, Array.map (f l) al)
- | IsEvar (e,al) -> mkEvar (e, Array.map (f l) al)
- | IsMutCase (ci,p,c,bl) -> mkMutCase (ci, f l p, f l c, Array.map (f l) bl)
- | IsFix (ln,(lna,tl,bl)) ->
- let l' = iterate g (Array.length tl) l in
- mkFix (ln,(lna,Array.map (f l) tl,Array.map (f l') bl))
- | IsCoFix(ln,(lna,tl,bl)) ->
+ | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
+ | Construct _) -> c
+ | Cast (c,t) -> mkCast (f l c, f l t)
+ | Prod (na,t,c) -> mkProd (na, f l t, f (g l) c)
+ | Lambda (na,t,c) -> mkLambda (na, f l t, f (g l) c)
+ | LetIn (na,b,t,c) -> mkLetIn (na, f l b, f l t, f (g l) c)
+ | App (c,al) -> mkApp (f l c, Array.map (f l) al)
+ | Evar (e,al) -> mkEvar (e, Array.map (f l) al)
+ | Case (ci,p,c,bl) -> mkCase (ci, f l p, f l c, Array.map (f l) bl)
+ | Fix (ln,(lna,tl,bl)) ->
let l' = iterate g (Array.length tl) l in
- mkCoFix (ln,(lna,Array.map (f l) tl,Array.map (f l') bl))
-
-(* [map_constr_with_named_binders g f l c] maps [f l] on the immediate
- subterms of [c]; it carries an extra data [l] (typically a name
- list) which is processed by [g na] (which typically cons [na] to
- [l]) at each binder traversal (with name [na]); it is not recursive
- and the order with which subterms are processed is not specified *)
-
-let map_constr_with_named_binders g f l c = match kind_of_term c with
- | (IsRel _ | IsMeta _ | IsVar _ | IsSort _ | IsConst _ | IsMutInd _
- | IsMutConstruct _) -> c
- | IsCast (c,t) -> mkCast (f l c, f l t)
- | IsProd (na,t,c) -> mkProd (na, f l t, f (g na l) c)
- | IsLambda (na,t,c) -> mkLambda (na, f l t, f (g na l) c)
- | IsLetIn (na,b,t,c) -> mkLetIn (na, f l b, f l t, f (g na l) c)
- | IsApp (c,al) -> mkApp (f l c, Array.map (f l) al)
- | IsEvar (e,al) -> mkEvar (e, Array.map (f l) al)
- | IsMutCase (ci,p,c,bl) -> mkMutCase (ci, f l p, f l c, Array.map (f l) bl)
- | IsFix (ln,(lna,tl,bl)) ->
- let l' = Array.fold_left (fun l na -> g na l) l lna in
mkFix (ln,(lna,Array.map (f l) tl,Array.map (f l') bl))
- | IsCoFix(ln,(lna,tl,bl)) ->
- let l' = Array.fold_left (fun l na -> g na l) l lna in
- mkCoFix (ln,(lna,Array.map (f l) tl,Array.map (f l') bl))
-
-(* [map_constr_with_binders_left_to_right g f n c] maps [f n] on the
- immediate subterms of [c]; it carries an extra data [n] (typically
- a lift index) which is processed by [g] (which typically add 1 to
- [n]) at each binder traversal; the subterms are processed from left
- to right according to the usual representation of the constructions
- (this may matter if [f] does a side-effect); it is not recursive;
- in fact, the usual representation of the constructions is at the
- time being almost those of the ML representation (except for
- (co-)fixpoint) *)
-
-let array_map_left f a = (* Ocaml does not guarantee Array.map is LR *)
- let l = Array.length a in (* (even if so), then we rewrite it *)
- if l = 0 then [||] else begin
- let r = Array.create l (f a.(0)) in
- for i = 1 to l - 1 do
- r.(i) <- f a.(i)
- done;
- r
- end
-
-let array_map_left_pair f a g b =
- let l = Array.length a in
- if l = 0 then [||],[||] else begin
- let r = Array.create l (f a.(0)) in
- let s = Array.create l (g b.(0)) in
- for i = 1 to l - 1 do
- r.(i) <- f a.(i);
- s.(i) <- g b.(i)
- done;
- r, s
- end
-
-let map_constr_with_binders_left_to_right g f l c = match kind_of_term c with
- | (IsRel _ | IsMeta _ | IsVar _ | IsSort _ | IsConst _ | IsMutInd _
- | IsMutConstruct _) -> c
- | IsCast (c,t) -> let c' = f l c in mkCast (c', f l t)
- | IsProd (na,t,c) -> let t' = f l t in mkProd (na, t', f (g l) c)
- | IsLambda (na,t,c) -> let t' = f l t in mkLambda (na, t', f (g l) c)
- | IsLetIn (na,b,t,c) ->
- let b' = f l b in let t' = f l t in mkLetIn (na, b', t', f (g l) c)
- | IsApp (c,al) ->
- let c' = f l c in mkApp (c', array_map_left (f l) al)
- | IsEvar (e,al) -> mkEvar (e, array_map_left (f l) al)
- | IsMutCase (ci,p,c,bl) ->
- let p' = f l p in let c' = f l c in
- mkMutCase (ci, p', c', array_map_left (f l) bl)
- | IsFix (ln,(lna,tl,bl)) ->
+ | CoFix(ln,(lna,tl,bl)) ->
let l' = iterate g (Array.length tl) l in
- let (tl',bl') = array_map_left_pair (f l) tl (f l') bl in
- mkFix (ln,(lna,tl',bl'))
- | IsCoFix(ln,(lna,tl,bl)) ->
- let l' = iterate g (Array.length tl) l in
- let (tl',bl') = array_map_left_pair (f l) tl (f l') bl in
- mkCoFix (ln,(lna,tl',bl'))
-
-(* strong *)
-let map_constr_with_full_binders g f l c = match kind_of_term c with
- | (IsRel _ | IsMeta _ | IsVar _ | IsSort _ | IsConst _ | IsMutInd _
- | IsMutConstruct _) -> c
- | IsCast (c,t) -> mkCast (f l c, f l t)
- | IsProd (na,t,c) -> mkProd (na, f l t, f (g (na,None,t) l) c)
- | IsLambda (na,t,c) -> mkLambda (na, f l t, f (g (na,None,t) l) c)
- | IsLetIn (na,b,t,c) -> mkLetIn (na, f l b, f l t, f (g (na,Some b,t) l) c)
- | IsApp (c,al) -> mkApp (f l c, Array.map (f l) al)
- | IsEvar (e,al) -> mkEvar (e, Array.map (f l) al)
- | IsMutCase (ci,p,c,bl) -> mkMutCase (ci, f l p, f l c, Array.map (f l) bl)
- | IsFix (ln,(lna,tl,bl)) ->
- let l' =
- array_fold_left2 (fun l na t -> g (na,None,t) l) l lna tl in
- mkFix (ln,(lna,Array.map (f l) tl, Array.map (f l') bl))
- | IsCoFix(ln,(lna,tl,bl)) ->
- let l' =
- array_fold_left2 (fun l na t -> g (na,None,t) l) l lna tl in
mkCoFix (ln,(lna,Array.map (f l) tl,Array.map (f l') bl))
(* [compare_constr f c1 c2] compare [c1] and [c2] using [f] to compare
@@ -768,33 +659,33 @@ let map_constr_with_full_binders g f l c = match kind_of_term c with
let compare_constr f t1 t2 =
match kind_of_term t1, kind_of_term t2 with
- | IsRel n1, IsRel n2 -> n1 = n2
- | IsMeta m1, IsMeta m2 -> m1 = m2
- | IsVar id1, IsVar id2 -> id1 = id2
- | IsSort s1, IsSort s2 -> s1 = s2
- | IsCast (c1,_), _ -> f c1 t2
- | _, IsCast (c2,_) -> f t1 c2
- | IsProd (_,t1,c1), IsProd (_,t2,c2) -> f t1 t2 & f c1 c2
- | IsLambda (_,t1,c1), IsLambda (_,t2,c2) -> f t1 t2 & f c1 c2
- | IsLetIn (_,b1,t1,c1), IsLetIn (_,b2,t2,c2) -> f b1 b2 & f t1 t2 & f c1 c2
- | IsApp (c1,l1), IsApp (c2,l2) ->
+ | Rel n1, Rel n2 -> n1 = n2
+ | Meta m1, Meta m2 -> m1 = m2
+ | Var id1, Var id2 -> id1 = id2
+ | Sort s1, Sort s2 -> s1 = s2
+ | Cast (c1,_), _ -> f c1 t2
+ | _, Cast (c2,_) -> f t1 c2
+ | Prod (_,t1,c1), Prod (_,t2,c2) -> f t1 t2 & f c1 c2
+ | Lambda (_,t1,c1), Lambda (_,t2,c2) -> f t1 t2 & f c1 c2
+ | LetIn (_,b1,t1,c1), LetIn (_,b2,t2,c2) -> f b1 b2 & f t1 t2 & f c1 c2
+ | App (c1,l1), App (c2,l2) ->
if Array.length l1 = Array.length l2 then
f c1 c2 & array_for_all2 f l1 l2
else
- let (h1,l1) = decomp_app t1 in
- let (h2,l2) = decomp_app t2 in
+ let (h1,l1) = decompose_app t1 in
+ let (h2,l2) = decompose_app t2 in
if List.length l1 = List.length l2 then
f h1 h2 & List.for_all2 f l1 l2
else false
- | IsEvar (e1,l1), IsEvar (e2,l2) -> e1 = e2 & array_for_all2 f l1 l2
- | IsConst c1, IsConst c2 -> c1 = c2
- | IsMutInd c1, IsMutInd c2 -> c1 = c2
- | IsMutConstruct c1, IsMutConstruct c2 -> c1 = c2
- | IsMutCase (_,p1,c1,bl1), IsMutCase (_,p2,c2,bl2) ->
+ | Evar (e1,l1), Evar (e2,l2) -> e1 = e2 & array_for_all2 f l1 l2
+ | Const c1, Const c2 -> c1 = c2
+ | Ind c1, Ind c2 -> c1 = c2
+ | Construct c1, Construct c2 -> c1 = c2
+ | Case (_,p1,c1,bl1), Case (_,p2,c2,bl2) ->
f p1 p2 & f c1 c2 & array_for_all2 f bl1 bl2
- | IsFix (ln1,(_,tl1,bl1)), IsFix (ln2,(_,tl2,bl2)) ->
+ | Fix (ln1,(_,tl1,bl1)), Fix (ln2,(_,tl2,bl2)) ->
ln1 = ln2 & array_for_all2 f tl1 tl2 & array_for_all2 f bl1 bl2
- | IsCoFix(ln1,(_,tl1,bl1)), IsCoFix(ln2,(_,tl2,bl2)) ->
+ | CoFix(ln1,(_,tl1,bl1)), CoFix(ln2,(_,tl2,bl2)) ->
ln1 = ln2 & array_for_all2 f tl1 tl2 & array_for_all2 f bl1 bl2
| _ -> false
@@ -811,7 +702,9 @@ let body_of_type ty = ty
type named_declaration = identifier * constr option * types
type rel_declaration = name * constr option * types
-
+let map_named_declaration f = function
+ (id, Some v, ty) -> (id, Some (f v), f ty)
+ | (id, None, ty) -> (id, None, f ty)
(****************************************************************************)
(* Functions for dealing with constr terms *)
@@ -829,7 +722,7 @@ exception Occur
let closedn =
let rec closed_rec n c = match kind_of_term c with
- | IsRel m -> if m>n then raise FreeVar
+ | Rel m -> if m>n then raise FreeVar
| _ -> iter_constr_with_binders succ closed_rec n c
in
closed_rec
@@ -839,20 +732,11 @@ let closedn =
let closed0 term =
try closedn 0 term; true with FreeVar -> false
-(* returns the list of free debruijn indices in a term *)
-
-let free_rels m =
- let rec frec depth acc c = match kind_of_term c with
- | IsRel n -> if n >= depth then Intset.add (n-depth+1) acc else acc
- | _ -> fold_constr_with_binders succ frec depth acc c
- in
- frec 1 Intset.empty m
-
(* (noccurn n M) returns true iff (Rel n) does NOT occur in term M *)
let noccurn n term =
let rec occur_rec n c = match kind_of_term c with
- | IsRel m -> if m = n then raise Occur
+ | Rel m -> if m = n then raise Occur
| _ -> iter_constr_with_binders succ occur_rec n c
in
try occur_rec n term; true with Occur -> false
@@ -862,7 +746,7 @@ let noccurn n term =
let noccur_between n m term =
let rec occur_rec n c = match kind_of_term c with
- | IsRel(p) -> if n<=p && p<n+m then raise Occur
+ | Rel(p) -> if n<=p && p<n+m then raise Occur
| _ -> iter_constr_with_binders succ occur_rec n c
in
try occur_rec n term; true with Occur -> false
@@ -876,13 +760,13 @@ let noccur_between n m term =
let noccur_with_meta n m term =
let rec occur_rec n c = match kind_of_term c with
- | IsRel p -> if n<=p & p<n+m then raise Occur
- | IsApp(f,cl) ->
+ | Rel p -> if n<=p & p<n+m then raise Occur
+ | App(f,cl) ->
(match kind_of_term f with
- | IsCast (c,_) when isMeta c -> ()
- | IsMeta _ -> ()
+ | Cast (c,_) when isMeta c -> ()
+ | Meta _ -> ()
| _ -> iter_constr_with_binders succ occur_rec n c)
- | IsEvar (_, _) -> ()
+ | Evar (_, _) -> ()
| _ -> iter_constr_with_binders succ occur_rec n c
in
try (occur_rec n term; true) with Occur -> false
@@ -894,7 +778,7 @@ let noccur_with_meta n m term =
(* The generic lifting function *)
let rec exliftn el c = match kind_of_term c with
- | IsRel i -> mkRel(reloc_rel i el)
+ | Rel i -> mkRel(reloc_rel i el)
| _ -> map_constr_with_binders el_lift exliftn el c
(* Lifting the binding depth across k bindings *)
@@ -934,7 +818,7 @@ let make_substituend c = { sinfo=Unknown; sit=c }
let substn_many lamv n =
let lv = Array.length lamv in
let rec substrec depth c = match kind_of_term c with
- | IsRel k ->
+ | Rel k ->
if k<=depth then
c
else if k-depth <= lv then
@@ -976,7 +860,7 @@ let replace_vars var_alist =
List.map (fun (str,c) -> (str,make_substituend c)) var_alist in
let var_alist = thin_val var_alist in
let rec substrec n c = match kind_of_term c with
- | IsVar x ->
+ | Var x ->
(try lift_substituend n (List.assoc x var_alist)
with Not_found -> c)
| _ -> map_constr_with_binders succ substrec n c
@@ -1099,16 +983,16 @@ let mkEvar = mkEvar
(* Constructs the ith (co)inductive type of the block named sp *)
(* The array of terms correspond to the variables introduced in the section *)
-let mkMutInd = mkMutInd
+let mkInd = mkInd
(* Constructs the jth constructor of the ith (co)inductive type of the
block named sp. The array of terms correspond to the variables
introduced in the section *)
-let mkMutConstruct = mkMutConstruct
+let mkConstruct = mkConstruct
(* Constructs the term <p>Case c of c1 | c2 .. | cn end *)
-let mkMutCase = mkMutCase
-let mkMutCaseL (ci, p, c, ac) = mkMutCase (ci, p, c, Array.of_list ac)
+let mkCase = mkCase
+let mkCaseL (ci, p, c, ac) = mkCase (ci, p, c, Array.of_list ac)
(* If recindxs = [|i1,...in|]
funnames = [|f1,...fn|]
@@ -1151,17 +1035,17 @@ let implicit_sort = Type implicit_univ
let mkImplicit = mkSort implicit_sort
let rec strip_outer_cast c = match kind_of_term c with
- | IsCast (c,_) -> strip_outer_cast c
+ | Cast (c,_) -> strip_outer_cast c
| _ -> c
-(* Fonction spéciale qui laisse les cast clés sous les Fix ou les MutCase *)
+(* Fonction spéciale qui laisse les cast clés sous les Fix ou les Case *)
let under_outer_cast f c = match kind_of_term c with
- | IsCast (b,t) -> mkCast (f b,f t)
+ | Cast (b,t) -> mkCast (f b,f t)
| _ -> f c
let rec under_casts f c = match kind_of_term c with
- | IsCast (c,t) -> mkCast (under_casts f c, t)
+ | Cast (c,t) -> mkCast (under_casts f c, t)
| _ -> f c
(***************************)
@@ -1172,13 +1056,6 @@ let abs_implicit c = mkLambda (Anonymous, mkImplicit, c)
let lambda_implicit a = mkLambda (Name(id_of_string"y"), mkImplicit, a)
let lambda_implicit_lift n a = iterate lambda_implicit n (lift n a)
-
-(* prod_it b [xn:Tn;..;x1:T1] = (x1:T1)..(xn:Tn)b *)
-let prod_it = List.fold_left (fun c (n,t) -> mkProd (n, t, c))
-
-(* lam_it b [xn:Tn;..;x1:T1] = [x1:T1]..[xn:Tn]b *)
-let lam_it = List.fold_left (fun c (n,t) -> mkLambda (n, t, c))
-
(* prodn n [xn:Tn;..;x1:T1;Gamma] b = (x1:T1)..(xn:Tn)b *)
let prodn n env b =
let rec prodrec = function
@@ -1212,8 +1089,8 @@ let rec to_lambda n prod =
prod
else
match kind_of_term prod with
- | IsProd (na,ty,bd) -> mkLambda (na,ty,to_lambda (n-1) bd)
- | IsCast (c,_) -> to_lambda n c
+ | Prod (na,ty,bd) -> mkLambda (na,ty,to_lambda (n-1) bd)
+ | Cast (c,_) -> to_lambda n c
| _ -> errorlabstrm "to_lambda" [<>]
let rec to_prod n lam =
@@ -1221,8 +1098,8 @@ let rec to_prod n lam =
lam
else
match kind_of_term lam with
- | IsLambda (na,ty,bd) -> mkProd (na,ty,to_prod (n-1) bd)
- | IsCast (c,_) -> to_prod n c
+ | Lambda (na,ty,bd) -> mkProd (na,ty,to_prod (n-1) bd)
+ | Cast (c,_) -> to_prod n c
| _ -> errorlabstrm "to_prod" [<>]
(* pseudo-reduction rule:
@@ -1231,7 +1108,7 @@ let rec to_prod n lam =
let prod_app t n =
match kind_of_term (strip_outer_cast t) with
- | IsProd (_,_,b) -> subst1 n b
+ | Prod (_,_,b) -> subst1 n b
| _ ->
errorlabstrm "prod_app"
[< 'sTR"Needed a product, but didn't find one" ; 'fNL >]
@@ -1243,27 +1120,6 @@ let prod_appvect t nL = Array.fold_left prod_app t nL
(* prod_applist T [ a1 ; ... ; an ] -> (T a1 ... an) *)
let prod_applist t nL = List.fold_left prod_app t nL
-
-(* [Rel (n+m);...;Rel(n+1)] *)
-let rel_vect n m = Array.init m (fun i -> mkRel(n+m-i))
-
-let rel_list n m =
- let rec reln l p =
- if p>m then l else reln (mkRel(n+p)::l) (p+1)
- in
- reln [] 1
-
-(* Same as [rel_list] but takes a context as argument and skips let-ins *)
-let extended_rel_list n hyps =
- let rec reln l p = function
- | (_,None,_) :: hyps -> reln (mkRel (n+p) :: l) (p+1) hyps
- | (_,Some _,_) :: hyps -> reln l (p+1) hyps
- | [] -> l
- in
- reln [] 1 hyps
-
-let extended_rel_vect n hyps = Array.of_list (extended_rel_list n hyps)
-
(*********************************)
(* Other term destructors *)
(*********************************)
@@ -1275,43 +1131,37 @@ type arity = rel_declaration list * sorts
let destArity =
let rec prodec_rec l c =
match kind_of_term c with
- | IsProd (x,t,c) -> prodec_rec ((x,None,t)::l) c
- | IsLetIn (x,b,t,c) -> prodec_rec ((x,Some b,t)::l) c
- | IsCast (c,_) -> prodec_rec l c
- | IsSort s -> l,s
+ | Prod (x,t,c) -> prodec_rec ((x,None,t)::l) c
+ | LetIn (x,b,t,c) -> prodec_rec ((x,Some b,t)::l) c
+ | Cast (c,_) -> prodec_rec l c
+ | Sort s -> l,s
| _ -> anomaly "destArity: not an arity"
in
prodec_rec []
let rec isArity c =
match kind_of_term c with
- | IsProd (_,_,c) -> isArity c
- | IsCast (c,_) -> isArity c
- | IsSort _ -> true
+ | Prod (_,_,c) -> isArity c
+ | Cast (c,_) -> isArity c
+ | Sort _ -> true
| _ -> false
(* Transforms a product term (x1:T1)..(xn:Tn)T into the pair
([(xn,Tn);...;(x1,T1)],T), where T is not a product *)
let decompose_prod =
let rec prodec_rec l c = match kind_of_term c with
- | IsProd (x,t,c) -> prodec_rec ((x,t)::l) c
- | IsCast (c,_) -> prodec_rec l c
+ | Prod (x,t,c) -> prodec_rec ((x,t)::l) c
+ | Cast (c,_) -> prodec_rec l c
| _ -> l,c
in
prodec_rec []
-let rec hd_of_prod prod =
- match kind_of_term prod with
- | IsProd (n,c,t') -> hd_of_prod t'
- | IsCast (c,_) -> hd_of_prod c
- | _ -> prod
-
(* Transforms a lambda term [x1:T1]..[xn:Tn]T into the pair
([(xn,Tn);...;(x1,T1)],T), where T is not a lambda *)
let decompose_lam =
let rec lamdec_rec l c = match kind_of_term c with
- | IsLambda (x,t,c) -> lamdec_rec ((x,t)::l) c
- | IsCast (c,_) -> lamdec_rec l c
+ | Lambda (x,t,c) -> lamdec_rec ((x,t)::l) c
+ | Cast (c,_) -> lamdec_rec l c
| _ -> l,c
in
lamdec_rec []
@@ -1323,8 +1173,8 @@ let decompose_prod_n n =
let rec prodec_rec l n c =
if n=0 then l,c
else match kind_of_term c with
- | IsProd (x,t,c) -> prodec_rec ((x,t)::l) (n-1) c
- | IsCast (c,_) -> prodec_rec l n c
+ | Prod (x,t,c) -> prodec_rec ((x,t)::l) (n-1) c
+ | Cast (c,_) -> prodec_rec l n c
| _ -> error "decompose_prod_n: not enough products"
in
prodec_rec [] n
@@ -1336,8 +1186,8 @@ let decompose_lam_n n =
let rec lamdec_rec l n c =
if n=0 then l,c
else match kind_of_term c with
- | IsLambda (x,t,c) -> lamdec_rec ((x,t)::l) (n-1) c
- | IsCast (c,_) -> lamdec_rec l n c
+ | Lambda (x,t,c) -> lamdec_rec ((x,t)::l) (n-1) c
+ | Cast (c,_) -> lamdec_rec l n c
| _ -> error "decompose_lam_n: not enough abstractions"
in
lamdec_rec [] n
@@ -1346,8 +1196,8 @@ let decompose_lam_n n =
* gives n (casts are ignored) *)
let nb_lam =
let rec nbrec n c = match kind_of_term c with
- | IsLambda (_,_,c) -> nbrec (n+1) c
- | IsCast (c,_) -> nbrec n c
+ | Lambda (_,_,c) -> nbrec (n+1) c
+ | Cast (c,_) -> nbrec n c
| _ -> n
in
nbrec 0
@@ -1355,282 +1205,28 @@ let nb_lam =
(* similar to nb_lam, but gives the number of products instead *)
let nb_prod =
let rec nbrec n c = match kind_of_term c with
- | IsProd (_,_,c) -> nbrec (n+1) c
- | IsCast (c,_) -> nbrec n c
+ | Prod (_,_,c) -> nbrec (n+1) c
+ | Cast (c,_) -> nbrec n c
| _ -> n
in
nbrec 0
-(* Misc *)
-let sort_hdchar = function
- | Prop(_) -> "P"
- | Type(_) -> "T"
-
-(* Level comparison for information extraction : Prop <= Type *)
-let le_kind l m = (isprop l) or (is_Type m)
-
-let le_kind_implicit k1 k2 =
- (k1=mkImplicit) or (isprop k1) or (k2=mkImplicit) or (is_Type k2)
-
-
(* Rem: end of import from old module Generic *)
-(* Various occurs checks *)
-
-(* (occur_const s c) -> true if constant s occurs in c,
- * false otherwise *)
-let occur_const s c =
- let rec occur_rec c = match kind_of_term c with
- | IsConst sp when sp=s -> raise Occur
- | _ -> iter_constr occur_rec c
- in
- try occur_rec c; false with Occur -> true
-
-let occur_evar n c =
- let rec occur_rec c = match kind_of_term c with
- | IsEvar (sp,_) when sp=n -> raise Occur
- | _ -> iter_constr occur_rec c
- in
- try occur_rec c; false with Occur -> true
-
-(***************************************)
-(* alpha and eta conversion functions *)
-(***************************************)
+(*******************************)
+(* alpha conversion functions *)
+(*******************************)
(* alpha conversion : ignore print names and casts *)
let rec eq_constr m n =
- (m = n) or (* Rem: ocaml '=' includes '==' *)
+ (m==n) or
compare_constr eq_constr m n
let eq_constr m n = eq_constr m n (* to avoid tracing a recursive fun *)
-(* (dependent M N) is true iff M is eq_term with a subterm of N
- M is appropriately lifted through abstractions of N *)
-
-let dependent m t =
- let rec deprec m t =
- if (eq_constr m t) then
- raise Occur
- else
- iter_constr_with_binders (lift 1) deprec m t
- in
- try deprec m t; false with Occur -> true
-
-(* On reduit une serie d'eta-redex de tete ou rien du tout *)
-(* [x1:c1;...;xn:cn]@(f;a1...an;x1;...;xn) --> @(f;a1...an) *)
-(* Remplace 2 versions précédentes buggées *)
-
-let rec eta_reduce_head c =
- match kind_of_term c with
- | IsLambda (_,c1,c') ->
- (match kind_of_term (eta_reduce_head c') with
- | IsApp (f,cl) ->
- let lastn = (Array.length cl) - 1 in
- if lastn < 1 then anomaly "application without arguments"
- else
- (match kind_of_term cl.(lastn) with
- | IsRel 1 ->
- let c' =
- if lastn = 1 then f
- else mkApp (f, Array.sub cl 0 lastn)
- in
- if not (dependent (mkRel 1) c')
- then lift (-1) c'
- else c
- | _ -> c)
- | _ -> c)
- | _ -> c
-
-(* alpha-eta conversion : ignore print names and casts *)
-let eta_eq_constr =
- let rec aux t1 t2 =
- let t1 = eta_reduce_head (strip_head_cast t1)
- and t2 = eta_reduce_head (strip_head_cast t2) in
- t1=t2 or compare_constr aux t1 t2
- in aux
-
-
-(***************************)
-(* substitution functions *)
-(***************************)
-
-(* First utilities for avoiding telescope computation for subst_term *)
-
-let prefix_application (k,c) (t : constr) =
- let c' = strip_head_cast c and t' = strip_head_cast t in
- match kind_of_term c', kind_of_term t' with
- | IsApp (f1,cl1), IsApp (f2,cl2) ->
- let l1 = Array.length cl1
- and l2 = Array.length cl2 in
- if l1 <= l2
- && eq_constr c' (mkApp (f2, Array.sub cl2 0 l1)) then
- Some (mkApp (mkRel k, Array.sub cl2 l1 (l2 - l1)))
- else
- None
- | _ -> None
-
-let my_prefix_application (k,c) (by_c : constr) (t : constr) =
- let c' = strip_head_cast c and t' = strip_head_cast t in
- match kind_of_term c', kind_of_term t' with
- | IsApp (f1,cl1), IsApp (f2,cl2) ->
- let l1 = Array.length cl1
- and l2 = Array.length cl2 in
- if l1 <= l2
- && eq_constr c' (mkApp (f2, Array.sub cl2 0 l1)) then
- Some (mkApp ((lift k by_c), Array.sub cl2 l1 (l2 - l1)))
- else
- None
- | _ -> None
-
-let prefix_application_eta (k,c) (t : constr) =
- let c' = strip_head_cast c and t' = strip_head_cast t in
- match kind_of_term c', kind_of_term t' with
- | IsApp (f1,cl1), IsApp (f2,cl2) ->
- let l1 = Array.length cl1
- and l2 = Array.length cl2 in
- if l1 <= l2 &&
- eta_eq_constr c' (mkApp (f2, Array.sub cl2 0 l1)) then
- Some (mkApp (mkRel k, Array.sub cl2 l1 (l2 - l1)))
- else
- None
- | (_,_) -> None
-
-let sort_increasing_snd =
- Sort.list
- (fun (_,x) (_,y) -> match kind_of_term x, kind_of_term y with
- | IsRel m, IsRel n -> m < n
- | _ -> assert false)
-
-(* Recognizing occurrences of a given (closed) subterm in a term for Pattern :
- [subst_term c t] substitutes [(Rel 1)] for all occurrences of (closed)
- term [c] in a term [t] *)
-(*i Bizarre : si on cherche un sous terme clos, pourquoi le lifter ? i*)
-
-let subst_term_gen eq_fun c t =
- let rec substrec (k,c as kc) t =
- match prefix_application kc t with
- | Some x -> x
- | None ->
- (if eq_fun t c then mkRel k else match kind_of_term t with
- | IsConst _ | IsMutInd _ | IsMutConstruct _ -> t
- | _ ->
- map_constr_with_binders
- (fun (k,c) -> (k+1,lift 1 c))
- substrec kc t)
- in
- substrec (1,c) t
-
-(* Recognizing occurrences of a given (closed) subterm in a term :
- [replace_term c1 c2 t] substitutes [c2] for all occurrences of (closed)
- term [c1] in a term [t] *)
-(*i Meme remarque : a priori [c] n'est pas forcement clos i*)
-
-let replace_term_gen eq_fun c by_c in_t =
- let rec substrec (k,c as kc) t =
- match my_prefix_application kc by_c t with
- | Some x -> x
- | None ->
- (if eq_fun t c then (lift k by_c) else match kind_of_term t with
- | IsConst _ | IsMutInd _ | IsMutConstruct _ -> t
- | _ ->
- map_constr_with_binders
- (fun (k,c) -> (k+1,lift 1 c))
- substrec kc t)
- in
- substrec (0,c) in_t
-
-let subst_term = subst_term_gen eq_constr
-let subst_term_eta = subst_term_gen eta_eq_constr
-
-let replace_term = replace_term_gen eq_constr
-
-(* bl : (int,constr) Listmap.t = (int * constr) list *)
-(* c : constr *)
-(* for each binding (i,c_i) in bl, substitutes the metavar i by c_i in c *)
-(* Raises Not_found if c contains a meta that is not in the association list *)
-
-(* Bogué ? Pourquoi pas de lift en passant sous un lieur ?? *)
-(* Et puis meta doit fusionner avec Evar *)
-let rec subst_meta bl c =
- match kind_of_term c with
- | IsMeta i -> (try List.assoc i bl with Not_found -> c)
- | _ -> map_constr (subst_meta bl) c
-
-(* Substitute only a list of locations locs, the empty list is
- interpreted as substitute all, if 0 is in the list then no
- substitution is done. The list may contain only negative occurrences
- that will not be substituted. *)
-
-let subst_term_occ_gen locs occ c t =
- let maxocc = List.fold_right max locs 0 in
- let pos = ref occ in
- let check = ref true in
- let except = List.exists (fun n -> n<0) locs in
- if except & (List.exists (fun n -> n>=0) locs)
- then error "mixing of positive and negative occurences"
- else
- let rec substrec (k,c as kc) t =
- if (not except) & (!pos > maxocc) then t
- else
- if eq_constr t c then
- let r =
- if except then
- if List.mem (- !pos) locs then t else (mkRel k)
- else
- if List.mem !pos locs then (mkRel k) else t
- in incr pos; r
- else
- match kind_of_term t with
- | IsConst _ | IsMutConstruct _ | IsMutInd _ -> t
- | _ ->
- map_constr_with_binders_left_to_right
- (fun (k,c) -> (k+1,lift 1 c)) substrec kc t
- in
- let t' = substrec (1,c) t in
- (!pos, t')
-
-let subst_term_occ locs c t =
- if locs = [] then
- subst_term c t
- else if List.mem 0 locs then
- t
- else
- let (nbocc,t') = subst_term_occ_gen locs 1 c t in
- if List.exists (fun o -> o >= nbocc or o <= -nbocc) locs then
- errorlabstrm "subst_term_occ" [< 'sTR "Too few occurences" >];
- t'
-
-let subst_term_occ_decl locs c (id,bodyopt,typ as d) =
- match bodyopt with
- | None -> (id,None,subst_term_occ locs c typ)
- | Some body ->
- if locs = [] then
- (id,Some (subst_term c body),type_app (subst_term c) typ)
- else if List.mem 0 locs then
- d
- else
- let (nbocc,body') = subst_term_occ_gen locs 1 c body in
- let (nbocc',t') = type_app (subst_term_occ_gen locs nbocc c) typ in
- if List.exists (fun o -> o >= nbocc' or o <= -nbocc') locs then
- errorlabstrm "subst_term_occ_decl" [< 'sTR "Too few occurences" >];
- (id,Some body',t')
-
-(***************************)
-(* occurs check functions *)
-(***************************)
-
-let occur_meta c =
- let rec occrec c = match kind_of_term c with
- | IsMeta _ -> raise Occur
- | _ -> iter_constr occrec c
- in try occrec c; false with Occur -> true
-
-let occur_existential c =
- let rec occrec c = match kind_of_term c with
- | IsEvar _ -> raise Occur
- | _ -> iter_constr occrec c
- in try occrec c; false with Occur -> true
-
+(*******************)
+(* hash-consing *)
+(*******************)
module Htype =
Hashcons.Make(
@@ -1672,136 +1268,4 @@ let hcons_constr (hspcci,hdir,hname,hident,hstr) =
let htcci = Hashcons.simple_hcons Htype.f (hcci,hsortscci) in
(hcci,htcci)
-let hcons1_constr =
- let hnames = hcons_names () in
- let (hc,_) = hcons_constr hnames in
- hc
-
-let hcons1_types =
- let hnames = hcons_names () in
- let (_,ht) = hcons_constr hnames in
- ht
-
-(* Abstract decomposition of constr to deal with generic functions *)
-
-type fix_kind = RFix of (int array * int) | RCoFix of int
-
-type constr_operator =
- | OpMeta of int
- | OpSort of sorts
- | OpRel of int | OpVar of identifier
- | OpCast | OpProd of name | OpLambda of name | OpLetIn of name
- | OpApp | OpConst of constant
- | OpEvar of existential_key
- | OpMutInd of inductive
- | OpMutConstruct of constructor
- | OpMutCase of case_info
- | OpRec of fix_kind * name array
-
-let splay_constr c = match kind_of_term c with
- | IsRel n -> OpRel n, [||]
- | IsVar id -> OpVar id, [||]
- | IsMeta n -> OpMeta n, [||]
- | IsSort s -> OpSort s, [||]
- | IsCast (t1, t2) -> OpCast, [|t1;t2|]
- | IsProd (x, t1, t2) -> OpProd x, [|t1;t2|]
- | IsLambda (x, t1, t2) -> OpLambda x, [|t1;t2|]
- | IsLetIn (x, b, t1, t2) -> OpLetIn x, [|b;t1;t2|]
- | IsApp (f,a) -> OpApp, Array.append [|f|] a
- | IsConst sp -> OpConst sp,[||]
- | IsEvar (sp, a) -> OpEvar sp, a
- | IsMutInd ind_sp -> OpMutInd ind_sp,[||]
- | IsMutConstruct cstr_sp -> OpMutConstruct cstr_sp, [||]
- | IsMutCase (ci,p,c,bl) -> OpMutCase ci, Array.append [|p;c|] bl
- | IsFix (fi,(lna,tl,bl)) -> OpRec (RFix fi,lna), Array.append tl bl
- | IsCoFix (fi,(lna,tl,bl)) -> OpRec (RCoFix fi,lna), Array.append tl bl
-
-let gather_constr = function
- | OpRel n, [||] -> mkRel n
- | OpVar id, [||] -> mkVar id
- | OpMeta n, [||] -> mkMeta n
- | OpSort s, [||] -> mkSort s
- | OpCast, [|t1;t2|] -> mkCast (t1, t2)
- | OpProd x, [|t1;t2|] -> mkProd (x, t1, t2)
- | OpLambda x, [|t1;t2|] -> mkLambda (x, t1, t2)
- | OpLetIn x, [|b;t1;t2|] -> mkLetIn (x, b, t1, t2)
- | OpApp, v -> let f = v.(0) and a = array_tl v in mkApp (f, a)
- | OpConst sp, [||] -> mkConst sp
- | OpEvar sp, a -> mkEvar (sp, a)
- | OpMutInd ind_sp, [||] -> mkMutInd ind_sp
- | OpMutConstruct cstr_sp, [||] -> mkMutConstruct cstr_sp
- | OpMutCase ci, v ->
- let p = v.(0) and c = v.(1) and bl = Array.sub v 2 (Array.length v -2)
- in mkMutCase (ci, p, c, bl)
- | OpRec (RFix fi,na), a ->
- let n = Array.length a / 2 in
- mkFix (fi,(na, Array.sub a 0 n, Array.sub a n n))
- | OpRec (RCoFix i,na), a ->
- let n = Array.length a / 2 in
- mkCoFix (i,(na, Array.sub a 0 n, Array.sub a n n))
- | _ -> errorlabstrm "Term.gather_term" [< 'sTR "ill-formed splayed constr">]
-
-let splay_constr_with_binders c = match kind_of_term c with
- | IsRel n -> OpRel n, [], [||]
- | IsVar id -> OpVar id, [], [||]
- | IsMeta n -> OpMeta n, [], [||]
- | IsSort s -> OpSort s, [], [||]
- | IsCast (t1, t2) -> OpCast, [], [|t1;t2|]
- | IsProd (x, t1, t2) -> OpProd x, [x,None,t1], [|t2|]
- | IsLambda (x, t1, t2) -> OpLambda x, [x,None,t1], [|t2|]
- | IsLetIn (x, b, t1, t2) -> OpLetIn x, [x,Some b,t1], [|t2|]
- | IsApp (f,v) -> OpApp, [], Array.append [|f|] v
- | IsConst sp -> OpConst sp, [], [||]
- | IsEvar (sp, a) -> OpEvar sp, [], a
- | IsMutInd ind_sp -> OpMutInd ind_sp, [], [||]
- | IsMutConstruct cstr_sp -> OpMutConstruct cstr_sp, [], [||]
- | IsMutCase (ci,p,c,bl) ->
- let v = Array.append [|p;c|] bl in OpMutCase ci, [], v
- | IsFix (fi,(na,tl,bl)) ->
- let n = Array.length bl in
- let ctxt =
- Array.to_list
- (array_map2_i (fun i x t -> (x,None,lift i t)) na tl) in
- OpRec (RFix fi,na), ctxt, bl
- | IsCoFix (fi,(na,tl,bl)) ->
- let n = Array.length bl in
- let ctxt =
- Array.to_list
- (array_map2_i (fun i x t -> (x,None,lift i t)) na tl) in
- OpRec (RCoFix fi,na), ctxt, bl
-
-let gather_constr_with_binders = function
- | OpRel n, [], [||] -> mkRel n
- | OpVar id, [], [||] -> mkVar id
- | OpMeta n, [], [||] -> mkMeta n
- | OpSort s, [], [||] -> mkSort s
- | OpCast, [], [|t1;t2|] -> mkCast (t1, t2)
- | OpProd _, [x,None,t1], [|t2|] -> mkProd (x, t1, t2)
- | OpLambda _, [x,None,t1], [|t2|] -> mkLambda (x, t1, t2)
- | OpLetIn _, [x,Some b,t1], [|t2|] -> mkLetIn (x, b, t1, t2)
- | OpApp, [], v -> let f = v.(0) and a = array_tl v in mkApp (f, a)
- | OpConst sp, [], [||] -> mkConst sp
- | OpEvar sp, [], a -> mkEvar (sp, a)
- | OpMutInd ind_sp, [], [||] -> mkMutInd ind_sp
- | OpMutConstruct cstr_sp, [], [||] -> mkMutConstruct cstr_sp
- | OpMutCase ci, [], v ->
- let p = v.(0) and c = v.(1) and bl = Array.sub v 2 (Array.length v -2)
- in mkMutCase (ci, p, c, bl)
- | OpRec (RFix fi,na), ctxt, bl ->
- let tl =
- Array.mapi (fun i (_,_,t) -> lift (-i) t) (Array.of_list ctxt) in
- mkFix (fi,(na, tl, bl))
- | OpRec (RCoFix i,na), ctxt, bl ->
- let tl =
- Array.mapi (fun i (_,_,t) -> lift (-i) t) (Array.of_list ctxt) in
- mkCoFix (i,(na, tl, bl))
- | _ -> errorlabstrm "Term.gather_term" [< 'sTR "ill-formed splayed constr">]
-
-let generic_fold_left f acc bl tl =
- let acc =
- List.fold_left
- (fun acc (_,bo,t) ->
- match bo with
- | None -> f acc t
- | Some b -> f (f acc b) t) acc bl in
- Array.fold_left f acc tl
+let (hcons1_constr, hcons1_types) = hcons_constr (hcons_names())
diff --git a/kernel/term.mli b/kernel/term.mli
index 90b1dd807..0ce4f3d4a 100644
--- a/kernel/term.mli
+++ b/kernel/term.mli
@@ -1,4 +1,4 @@
-(***********************************************************************)
+(***********************Sppc************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
(* \VV/ *************************************************************)
@@ -9,8 +9,6 @@
(*i $Id$ i*)
(*i*)
-open Util
-open Pp
open Names
(*i*)
@@ -24,8 +22,7 @@ type sorts =
val mk_Set : sorts
val mk_Prop : sorts
-
-val print_sort : sorts -> std_ppcmds
+val type_0 : sorts
(*s The sorts family of CCI. *)
@@ -36,38 +33,33 @@ val new_sort_in_family : sorts_family -> sorts
(*s Useful types *)
+(*s Existential variables *)
type existential_key = int
+(*s Case annotation *)
type pattern_source = DefaultPat of int | RegularPat
type case_style = PrintLet | PrintIf | PrintCases
type case_printing =
- inductive * identifier array * int
- * case_style option * pattern_source array
+ { cnames : identifier array;
+ ind_nargs : int; (* number of real args of the inductive type *)
+ style : case_style option;
+ source : pattern_source array }
(* the integer is the number of real args, needed for reduction *)
-type case_info = int * case_printing
-
-(*s Concrete type for making pattern-matching. *)
-module Polymorph :
-sig
-(* [constr array] is an instance matching definitional [named_context] in
- the same order (i.e. last argument first) *)
-type 'constr existential = existential_key * 'constr array
-type ('constr, 'types) rec_declaration =
- name array * 'types array * 'constr array
-type ('constr, 'types) fixpoint =
- (int array * int) * ('constr, 'types) rec_declaration
-type ('constr, 'types) cofixpoint =
- int * ('constr, 'types) rec_declaration
-
-(* [IsVar] is used for named variables and [IsRel] for variables as
- de Bruijn indices. *)
-end
+type case_info =
+ { ci_ind : inductive;
+ ci_npar : int;
+ ci_pp_info : case_printing (* not interpreted by the kernel *)
+ }
(*s*******************************************************************)
(* The type of constructions *)
type constr
+(* [eq_constr a b] is true if [a] equals [b] modulo alpha, casts,
+ and application grouping *)
+val eq_constr : constr -> constr -> bool
+
(* [types] is the same as [constr] but is intended to be used where a
{\em type} in CCI sense is expected (Rem:plurial form since [type] is a
reserved ML keyword) *)
@@ -80,144 +72,79 @@ val type_app : (constr -> constr) -> types -> types
val body_of_type : types -> constr
-(*s A {\em declaration} has the form (name,body,type). It is either an
- {\em assumption} if [body=None] or a {\em definition} if
- [body=Some actualbody]. It is referred by {\em name} if [na] is an
- identifier or by {\em relative index} if [na] is not an identifier
- (in the latter case, [na] is of type [name] but just for printing
- purpose *)
-
-type named_declaration = identifier * constr option * types
-type rel_declaration = name * constr option * types
-
-type arity = rel_declaration list * sorts
-
(*s Functions for dealing with constr terms.
The following functions are intended to simplify and to uniform the
manipulation of terms. Some of these functions may be overlapped with
previous ones. *)
-open Polymorph
-type ('constr, 'types) kind_of_term =
- | IsRel of int
- | IsMeta of int
- | IsVar of identifier
- | IsSort of sorts
- | IsCast of 'constr * 'constr
- | IsProd of name * 'types * 'constr
- | IsLambda of name * 'types * 'constr
- | IsLetIn of name * 'constr * 'types * 'constr
- | IsApp of 'constr * 'constr array
- | IsEvar of 'constr existential
- | IsConst of constant
- | IsMutInd of inductive
- | IsMutConstruct of constructor
- | IsMutCase of case_info * 'constr * 'constr * 'constr array
- | IsFix of ('constr, 'types) fixpoint
- | IsCoFix of ('constr, 'types) cofixpoint
-
-type existential = existential_key * constr array
-type rec_declaration = name array * types array * constr array
-type fixpoint = (int array * int) * rec_declaration
-type cofixpoint = int * rec_declaration
-
-(* User view of [constr]. For [IsApp], it is ensured there is at
- least one argument and the function is not itself an applicative
- term *)
-
-val kind_of_term : constr -> (constr, types) kind_of_term
-
(*s Term constructors. *)
(* Constructs a DeBrujin index *)
val mkRel : int -> constr
-(* Constructs an existential variable named "?n" *)
-val mkMeta : int -> constr
-
(* Constructs a Variable *)
val mkVar : identifier -> constr
-(* Construct a type *)
+(* Constructs an metavariable named "?n" *)
+val mkMeta : int -> constr
+
+(* Constructs an existential variable *)
+type existential = existential_key * constr array
+val mkEvar : existential -> constr
+
+(* Construct a sort *)
val mkSort : sorts -> constr
val mkProp : constr
-val mkSet : constr
+val mkSet : constr
val mkType : Univ.universe -> constr
-val prop : sorts
-val spec : sorts
-(*val types : sorts *)
-val type_0 : sorts
-
-(* Construct an implicit (see implicit arguments in the RefMan).
- Used for extraction *)
-val mkImplicit : constr
-val implicit_sort : sorts
-(* Constructs the term $t_1::t2$, i.e. the term $t_1$ casted with the
+(* Constructs the term [t1::t2], i.e. the term $t_1$ casted with the
type $t_2$ (that means t2 is declared as the type of t1). *)
-val mkCast : constr * constr -> constr
+val mkCast : constr * types -> constr
-(* Constructs the product $(x:t_1)t_2$ *)
-val mkProd : name * types * constr -> constr
-val mkNamedProd : identifier -> constr -> constr -> constr
-val mkProd_string : string -> constr -> constr -> constr
-
-(* Constructs the product $(x:t_1)t_2$ *)
-val mkLetIn : name * constr * types * constr -> constr
-val mkNamedLetIn : identifier -> constr -> constr -> constr -> constr
-
-(* Constructs either [(x:t)c] or [[x=b:t]c] *)
-val mkProd_or_LetIn : rel_declaration -> constr -> constr
-val mkNamedProd_or_LetIn : named_declaration -> constr -> constr
-
-(* Constructs either [[x:t]c] or [[x=b:t]c] *)
-val mkLambda_or_LetIn : rel_declaration -> constr -> constr
-val mkNamedLambda_or_LetIn : named_declaration -> constr -> constr
-
-(* Constructs either [(x:t)c] or [c] where [x] is replaced by [b] *)
-val mkProd_wo_LetIn : rel_declaration -> constr -> constr
-val mkNamedProd_wo_LetIn : named_declaration -> constr -> constr
-
-(* non-dependant product $t_1 \rightarrow t_2$ *)
-val mkArrow : constr -> constr -> constr
+(* Constructs the product [(x:t1)t2] *)
+val mkProd : name * types * types -> constr
+val mkNamedProd : identifier -> types -> types -> constr
+(* non-dependant product $t_1 \rightarrow t_2$, an alias for
+ [(_:t1)t2]. Beware $t_2$ is NOT lifted.
+ Eg: A |- A->A is built by [(mkArrow (mkRel 0) (mkRel 1))] *)
+val mkArrow : types -> types -> constr
(* Constructs the abstraction $[x:t_1]t_2$ *)
val mkLambda : name * types * constr -> constr
-val mkNamedLambda : identifier -> constr -> constr -> constr
+val mkNamedLambda : identifier -> types -> constr -> constr
-(* [mkLambda_string s t c] constructs $[s:t]c$ *)
-val mkLambda_string : string -> constr -> constr -> constr
+(* Constructs the product [let x = t1 : t2 in t3] *)
+val mkLetIn : name * constr * types * constr -> constr
+val mkNamedLetIn : identifier -> constr -> types -> constr -> constr
(* [mkApp (f,[| t_1; ...; t_n |]] constructs the application
$(f~t_1~\dots~t_n)$. *)
val mkApp : constr * constr array -> constr
-val mkAppA : constr array -> constr
(* Constructs a constant *)
(* The array of terms correspond to the variables introduced in the section *)
val mkConst : constant -> constr
-(* Constructs an existential variable *)
-val mkEvar : existential -> constr
+(* Inductive types *)
(* Constructs the ith (co)inductive type of the block named sp *)
(* The array of terms correspond to the variables introduced in the section *)
-val mkMutInd : inductive -> constr
+val mkInd : inductive -> constr
(* Constructs the jth constructor of the ith (co)inductive type of the
block named sp. The array of terms correspond to the variables
introduced in the section *)
-val mkMutConstruct : constructor -> constr
+val mkConstruct : constructor -> constr
(* Constructs the term <p>Case c of c1 | c2 .. | cn end *)
-val mkMutCaseL : case_info * constr * constr * constr list -> constr
-val mkMutCase : case_info * constr * constr * constr array -> constr
+val mkCase : case_info * constr * constr * constr array -> constr
(* If [recindxs = [|i1,...in|]]
+ [funnames = [|f1,.....fn|]]
[typarray = [|t1,...tn|]]
- [funnames = [f1,.....fn]]
- [bodies = [b1,.....bn]]
- then [ mkFix ((recindxs,i),typarray, funnames, bodies) ]
+ [bodies = [|b1,.....bn|]]
+ then [ mkFix ((recindxs,i), funnames, typarray, bodies) ]
constructs the $i$th function of the block (counting from 0)
[Fixpoint f1 [ctx1] = b1
@@ -225,12 +152,14 @@ val mkMutCase : case_info * constr * constr * constr array -> constr
...
with fn [ctxn] = bn.]
- \noindent where the lenght of the $j$th context is $ij$.
+ \noindent where the length of the $j$th context is $ij$.
*)
+type rec_declaration = name array * types array * constr array
+type fixpoint = (int array * int) * rec_declaration
val mkFix : fixpoint -> constr
-(* If [typarray = [|t1,...tn|]]
- [funnames = [f1,.....fn]]
+(* If [funnames = [|f1,.....fn|]]
+ [typarray = [|t1,...tn|]]
[bodies = [b1,.....bn]] \par\noindent
then [mkCoFix (i, (typsarray, funnames, bodies))]
constructs the ith function of the block
@@ -240,22 +169,73 @@ val mkFix : fixpoint -> constr
...
with fn = bn.]
*)
+type cofixpoint = int * rec_declaration
val mkCoFix : cofixpoint -> constr
+
+(*s Concrete type for making pattern-matching. *)
+
+(* [constr array] is an instance matching definitional [named_context] in
+ the same order (i.e. last argument first) *)
+type 'constr pexistential = existential_key * 'constr array
+type ('constr, 'types) prec_declaration =
+ name array * 'types array * 'constr array
+type ('constr, 'types) pfixpoint =
+ (int array * int) * ('constr, 'types) prec_declaration
+type ('constr, 'types) pcofixpoint =
+ int * ('constr, 'types) prec_declaration
+
+type ('constr, 'types) kind_of_term =
+ | Rel of int
+ | Var of identifier
+ | Meta of int
+ | Evar of 'constr pexistential
+ | Sort of sorts
+ | Cast of 'constr * 'constr
+ | Prod of name * 'types * 'types
+ | Lambda of name * 'types * 'constr
+ | LetIn of name * 'constr * 'types * 'constr
+ | App of 'constr * 'constr array
+ | Const of constant
+ | Ind of inductive
+ | Construct of constructor
+ | Case of case_info * 'constr * 'constr * 'constr array
+ | Fix of ('constr, 'types) pfixpoint
+ | CoFix of ('constr, 'types) pcofixpoint
+
+(* User view of [constr]. For [App], it is ensured there is at
+ least one argument and the function is not itself an applicative
+ term *)
+
+val kind_of_term : constr -> (constr, types) kind_of_term
+
+(*s Simple term case analysis. *)
+
+val isRel : constr -> bool
+val isVar : constr -> bool
+val isMeta : constr -> bool
+val isSort : constr -> bool
+val isCast : constr -> bool
+val isApp : constr -> bool
+val isConst : constr -> bool
+val isConstruct : constr -> bool
+
+val is_Prop : constr -> bool
+val is_Set : constr -> bool
+val isprop : constr -> bool
+val is_Type : constr -> bool
+val iskind : constr -> bool
+val is_small : sorts -> bool
+
(*s Term destructors.
Destructor operations are partial functions and
raise [invalid_arg "dest*"] if the term has not the expected form. *)
-(* Destructs a term of the form $(x_1:T_1)..(x_n:T_n)s$ into the pair *)
-val destArity : constr -> arity
-val isArity : constr -> bool
-
(* Destructs a DeBrujin index *)
val destRel : constr -> int
(* Destructs an existential variable *)
val destMeta : constr -> int
-val isMeta : constr -> bool
(* Destructs a variable *)
val destVar : constr -> identifier
@@ -263,68 +243,35 @@ val destVar : constr -> identifier
(* Destructs a sort. [is_Prop] recognizes the sort \textsf{Prop}, whether
[isprop] recognizes both \textsf{Prop} and \textsf{Set}. *)
val destSort : constr -> sorts
-val is_Prop : constr -> bool
-val is_Set : constr -> bool
-val isprop : constr -> bool
-val is_Type : constr -> bool
-val iskind : constr -> bool
-val isSort : constr -> bool
-
-val isType : sorts -> bool
-val is_small : sorts -> bool (* true for \textsf{Prop} and \textsf{Set} *)
(* Destructs a casted term *)
-val destCast : constr -> constr * constr
-val isCast : constr -> bool
-
-(* Removes recursively the casts around a term i.e.
- [strip_outer_cast] (Cast (Cast ... (Cast c, t) ... ))] is [c]. *)
-val strip_outer_cast : constr -> constr
-
-(* Apply a function letting Casted types in place *)
-val under_casts : (constr -> constr) -> constr -> constr
-
-(* Tests if a de Bruijn index *)
-val isRel : constr -> bool
-
-(* Tests if a variable *)
-val isVar : constr -> bool
+val destCast : constr -> constr * types
(* Destructs the product $(x:t_1)t_2$ *)
-val destProd : constr -> name * constr * constr
-val hd_of_prod : constr -> constr
-(*i
-val hd_is_constructor : constr -> bool
-i*)
+val destProd : types -> name * types * types
(* Destructs the abstraction $[x:t_1]t_2$ *)
-val destLambda : constr -> name * constr * constr
+val destLambda : constr -> name * types * constr
(* Destructs the let $[x:=b:t_1]t_2$ *)
-val destLetIn : constr -> name * constr * constr * constr
+val destLetIn : constr -> name * constr * types * constr
(* Destructs an application *)
-val isApp : constr -> bool
-(*i
-val hd_app : constr -> constr
-val args_app : constr -> constr array
-i*)
val destApplication : constr -> constr * constr array
+(* ... removing casts *)
+val decompose_app : constr -> constr * constr list
(* Destructs a constant *)
val destConst : constr -> constant
-val isConst : constr -> bool
(* Destructs an existential variable *)
-val destEvar : constr -> existential_key * constr array
-val num_of_evar : constr -> existential_key
+val destEvar : constr -> existential
(* Destructs a (co)inductive type *)
-val destMutInd : constr -> inductive
+val destInd : constr -> inductive
(* Destructs a constructor *)
-val destMutConstruct : constr -> constructor
-val isMutConstruct : constr -> bool
+val destConstruct : constr -> constructor
(* Destructs a term <p>Case c of lc1 | lc2 .. | lcn end *)
val destCase : constr -> case_info * constr * constr * constr array
@@ -340,6 +287,30 @@ val destFix : constr -> fixpoint
val destCoFix : constr -> cofixpoint
+
+(*s A {\em declaration} has the form (name,body,type). It is either an
+ {\em assumption} if [body=None] or a {\em definition} if
+ [body=Some actualbody]. It is referred by {\em name} if [na] is an
+ identifier or by {\em relative index} if [na] is not an identifier
+ (in the latter case, [na] is of type [name] but just for printing
+ purpose *)
+
+type named_declaration = identifier * constr option * types
+type rel_declaration = name * constr option * types
+
+val map_named_declaration :
+ (constr -> constr) -> named_declaration -> named_declaration
+
+(* Constructs either [(x:t)c] or [[x=b:t]c] *)
+val mkProd_or_LetIn : rel_declaration -> types -> constr
+val mkNamedProd_or_LetIn : named_declaration -> types -> constr
+val mkNamedProd_wo_LetIn : named_declaration -> types -> constr
+
+(* Constructs either [[x:t]c] or [[x=b:t]c] *)
+val mkLambda_or_LetIn : rel_declaration -> constr -> constr
+val mkNamedLambda_or_LetIn : named_declaration -> constr -> constr
+
+
(*s Other term constructors. *)
val abs_implicit : constr -> constr
@@ -361,14 +332,6 @@ val prodn : int -> (name * constr) list -> constr -> constr
where $l = [(x_n,T_n);\dots;(x_1,T_1);Gamma]$ *)
val lamn : int -> (name * constr) list -> constr -> constr
-(* [prod_it b l] = $(x_1:T_1)..(x_n:T_n)b$
- where $l = [(x_n,T_n);\dots;(x_1,T_1)]$ *)
-val prod_it : constr -> (name * constr) list -> constr
-
-(* [lam_it b l] = $[x_1:T_1]..[x_n:T_n]b$
- where $l = [(x_n,T_n);\dots;(x_1,T_1)]$ *)
-val lam_it : constr -> (name * constr) list -> constr
-
(* [to_lambda n l]
= $[x_1:T_1]...[x_n:T_n](x_{n+1}:T_{n+1})...(x_{n+j}:T_{n+j})T$
where $l = (x_1:T_1)...(x_n:T_n)(x_{n+1}:T_{n+1})...(x_{n+j}:T_{n+j})T$ *)
@@ -381,6 +344,11 @@ val prod_applist : constr -> constr list -> constr
(*s Other term destructors. *)
+(* Destructs a term of the form $(x_1:T_1)..(x_n:T_n)s$ into the pair *)
+type arity = rel_declaration list * sorts
+val destArity : constr -> arity
+val isArity : constr -> bool
+
(* Transforms a product term $(x_1:T_1)..(x_n:T_n)T$ into the pair
$([(x_n,T_n);...;(x_1,T_1)],T)$, where $T$ is not a product.
It includes also local definitions *)
@@ -408,20 +376,16 @@ val nb_prod : constr -> int
(* flattens application lists *)
val collapse_appl : constr -> constr
-val decomp_app : constr -> constr * constr list
-
-(*s Misc functions on terms, sorts and conversion problems. *)
-(* Level comparison for information extraction : Prop <= Type *)
-val same_kind : constr -> constr -> bool
-val le_kind : constr -> constr -> bool
-val le_kind_implicit : constr -> constr -> bool
+(* Removes recursively the casts around a term i.e.
+ [strip_outer_cast] (Cast (Cast ... (Cast c, t) ... ))] is [c]. *)
+val strip_outer_cast : constr -> constr
-val sort_hdchar : sorts -> string
+(* Apply a function letting Casted types in place *)
+val under_casts : (constr -> constr) -> constr -> constr
-(* Generic functions *)
-val free_rels : constr -> Intset.t
+(*s Occur checks *)
(* [closed0 M] is true iff [M] is a (deBruijn) closed term *)
val closed0 : constr -> bool
@@ -439,6 +403,8 @@ val noccur_between : int -> int -> constr -> bool
context) (for existential variables, it is necessarily the case) *)
val noccur_with_meta : int -> int -> constr -> bool
+(*s Relocation and substitution *)
+
(* [liftn n k c] lifts by [n] indexes above [k] in [c] *)
val liftn : int -> int -> constr -> constr
@@ -469,91 +435,6 @@ val subst_vars : identifier list -> constr -> constr
if two names are identical, the one of least indice is keeped *)
val substn_vars : int -> identifier list -> constr -> constr
-(* [rel_list n m] and [rel_vect n m] compute [[Rel (n+m);...;Rel(n+1)]] *)
-val rel_vect : int -> int -> constr array
-val rel_list : int -> int -> constr list
-
-(*s [extended_rel_vect n hyps] and [extended_rel_list n hyps]
- generalizes [rel_vect] when local definitions may occur in parameters *)
-val extended_rel_vect : int -> rel_declaration list -> constr array
-val extended_rel_list : int -> rel_declaration list -> constr list
-
-(*s Occur check functions. *)
-
-val occur_meta : constr -> bool
-
-(*i Returns the maximum of metas. Returns -1 if there is no meta i*)
-(*i val max_occur_meta : constr -> int i*)
-
-val occur_existential : constr -> bool
-
-(* [(occur_const (s:section_path) c)] returns [true] if constant [s] occurs
- in c, [false] otherwise *)
-val occur_const : constant -> constr -> bool
-
-(* [(occur_evar ev c)] returns [true] if existential variable [ev] occurs
- in c, [false] otherwise *)
-val occur_evar : existential_key -> constr -> bool
-
-(* [dependent M N] is true iff M is eq\_constr with a subterm of N
- M is appropriately lifted through abstractions of N *)
-val dependent : constr -> constr -> bool
-
-(* strips head casts and flattens head applications *)
-val strip_head_cast : constr -> constr
-val eta_reduce_head : constr -> constr
-val eq_constr : constr -> constr -> bool
-val eta_eq_constr : constr -> constr -> bool
-
-(*s The following functions substitutes [what] by [Rel 1] in [where] *)
-val subst_term : what:constr -> where:constr -> constr
-val subst_term_occ : occs:int list -> what:constr -> where:constr -> constr
-val subst_term_occ_decl : occs:int list -> what:constr ->
- where:named_declaration -> named_declaration
-
-(* [replace_term c by_c in_t substitutes c by by_c in in_t *)
-val replace_term : constr -> constr -> constr -> constr
-
-(* [subst_meta bl c] substitutes the metavar $i$ by $c_i$ in [c]
- for each binding $(i,c_i)$ in [bl],
- and raises [Not_found] if [c] contains a meta that is not in the
- association list *)
-
-val subst_meta : (int * constr) list -> constr -> constr
-
-(*s Generic representation of constructions *)
-
-type fix_kind = RFix of (int array * int) | RCoFix of int
-
-type constr_operator =
- | OpMeta of int
- | OpSort of sorts
- | OpRel of int | OpVar of identifier
- | OpCast | OpProd of name | OpLambda of name | OpLetIn of name
- | OpApp | OpConst of constant
- | OpEvar of existential_key
- | OpMutInd of inductive
- | OpMutConstruct of constructor
- | OpMutCase of case_info
- | OpRec of fix_kind * name array
-
-
-val splay_constr : constr -> constr_operator * constr array
-val gather_constr : constr_operator * constr array -> constr
-(*i
-val splay_constr : ('a,'a)kind_of_term -> constr_operator * 'a array
-val gather_constr : constr_operator * 'a array -> ('a,'a) kind_of_term
-i*)
-val splay_constr_with_binders : constr ->
- constr_operator * rel_declaration list * constr array
-val gather_constr_with_binders :
- constr_operator * rel_declaration list * constr array
- -> constr
-
-val generic_fold_left :
- ('a -> constr -> 'a) -> 'a -> rel_declaration list
- -> constr array -> 'a
-
(*s Functionals working on the immediate subterm of a construction *)
(* [fold_constr f acc c] folds [f] on the immediate subterms of [c]
@@ -572,21 +453,6 @@ val fold_constr : ('a -> constr -> 'a) -> 'a -> constr -> 'a
val fold_constr_with_binders :
('a -> 'a) -> ('a -> 'b -> constr -> 'b) -> 'a -> 'b -> constr -> 'b
-(* [iter_constr f c] iters [f] on the immediate subterms of [c]; it is
- not recursive and the order with which subterms are processed is
- not specified *)
-
-val iter_constr : (constr -> unit) -> constr -> unit
-
-(* [iter_constr_with_binders g f n c] iters [f n] on the immediate
- subterms of [c]; it carries an extra data [n] (typically a lift
- index) which is processed by [g] (which typically add 1 to [n]) at
- each binder traversal; it is not recursive and the order with which
- subterms are processed is not specified *)
-
-val iter_constr_with_binders :
- ('a -> 'a) -> ('a -> constr -> unit) -> 'a -> constr -> unit
-
(* [map_constr f c] maps [f] on the immediate subterms of [c]; it is
not recursive and the order with which subterms are processed is
not specified *)
@@ -602,34 +468,14 @@ val map_constr : (constr -> constr) -> constr -> constr
val map_constr_with_binders :
('a -> 'a) -> ('a -> constr -> constr) -> 'a -> constr -> constr
-(* [map_constr_with_named_binders g f l c] maps [f l] on the immediate
- subterms of [c]; it carries an extra data [l] (typically a name
- list) which is processed by [g na] (which typically cons [na] to
- [l]) at each binder traversal (with name [na]); it is not recursive
- and the order with which subterms are processed is not specified *)
-
-val map_constr_with_named_binders :
- (name -> 'a -> 'a) -> ('a -> constr -> constr) -> 'a -> constr -> constr
-
-(* [map_constr_with_binders_left_to_right g f n c] maps [f n] on the
- immediate subterms of [c]; it carries an extra data [n] (typically
- a lift index) which is processed by [g] (which typically add 1 to
- [n]) at each binder traversal; the subterms are processed from left
- to right according to the usual representation of the constructions
- (this may matter if [f] does a side-effect); it is not recursive *)
-
-val map_constr_with_binders_left_to_right :
- ('a -> 'a) -> ('a -> constr -> constr) -> 'a -> constr -> constr
-
-(* [map_constr_with_named_binders g f l c] maps [f l] on the immediate
- subterms of [c]; it carries an extra data [l] (typically a name
- list) which is processed by [g na] (which typically cons [na] to
- [l]) at each binder traversal (with name [na]); it is not recursive
- and the order with which subterms are processed is not specified *)
+(* [iter_constr_with_binders g f n c] iters [f n] on the immediate
+ subterms of [c]; it carries an extra data [n] (typically a lift
+ index) which is processed by [g] (which typically add 1 to [n]) at
+ each binder traversal; it is not recursive and the order with which
+ subterms are processed is not specified *)
-val map_constr_with_full_binders :
- (rel_declaration -> 'a -> 'a) -> ('a -> constr -> constr) ->
- 'a -> constr -> constr
+val iter_constr_with_binders :
+ ('a -> 'a) -> ('a -> constr -> unit) -> 'a -> constr -> unit
(* [compare_constr f c1 c2] compare [c1] and [c2] using [f] to compare
the immediate subterms of [c1] of [c2] if needed; Cast's, binders
@@ -637,7 +483,7 @@ val map_constr_with_full_binders :
val compare_constr : (constr -> constr -> bool) -> constr -> constr -> bool
-(*s Hash-consing functions for constr. *)
+(*********************************************************************)
val hcons_constr:
(section_path -> section_path) *
diff --git a/kernel/type_errors.ml b/kernel/type_errors.ml
index 05b6e2675..169df5904 100644
--- a/kernel/type_errors.ml
+++ b/kernel/type_errors.ml
@@ -8,7 +8,6 @@
(* $Id$ *)
-open Pp
open Names
open Term
open Sign
@@ -38,68 +37,69 @@ type guard_error =
type type_error =
| UnboundRel of int
| NotAType of unsafe_judgment
- | BadAssumption of constr
- | ReferenceVariables of identifier
+ | BadAssumption of unsafe_judgment
+ | ReferenceVariables of constr
| ElimArity of inductive * constr list * constr * unsafe_judgment
* (constr * constr * string) option
| CaseNotInductive of unsafe_judgment
+ | WrongCaseInfo of inductive * case_info
| NumberBranches of unsafe_judgment * int
| IllFormedBranch of constr * int * constr * constr
| Generalization of (name * types) * unsafe_judgment
- | ActualType of constr * constr * constr
+ | ActualType of unsafe_judgment * types
| CantApplyBadType of (int * constr * constr)
- * unsafe_judgment * unsafe_judgment list
- | CantApplyNonFunctional of unsafe_judgment * unsafe_judgment list
+ * unsafe_judgment * unsafe_judgment array
+ | CantApplyNonFunctional of unsafe_judgment * unsafe_judgment array
| IllFormedRecBody of guard_error * name array * int * constr array
| IllTypedRecBody of int * name array * unsafe_judgment array
* types array
-exception TypeError of path_kind * env * type_error
+exception TypeError of env * type_error
let nfj {uj_val=c;uj_type=ct} =
{uj_val=c;uj_type=nf_betaiota ct}
-let error_unbound_rel k env n =
- raise (TypeError (k, env, UnboundRel n))
+let error_unbound_rel env n =
+ raise (TypeError (env, UnboundRel n))
-let error_not_type k env c =
- raise (TypeError (k, env, NotAType c))
+let error_not_type env j =
+ raise (TypeError (env, NotAType j))
-let error_assumption k env c =
- raise (TypeError (k, env, BadAssumption c))
+let error_assumption env j =
+ raise (TypeError (env, BadAssumption j))
-let error_reference_variables k env id =
- raise (TypeError (k, env, ReferenceVariables id))
+let error_reference_variables env id =
+ raise (TypeError (env, ReferenceVariables id))
-let error_elim_arity k env ind aritylst c pj okinds =
- raise (TypeError (k, env, ElimArity (ind,aritylst,c,pj,okinds)))
+let error_elim_arity env ind aritylst c pj okinds =
+ raise (TypeError (env, ElimArity (ind,aritylst,c,pj,okinds)))
-let error_case_not_inductive k env j =
- raise (TypeError (k, env, CaseNotInductive j))
+let error_case_not_inductive env j =
+ raise (TypeError (env, CaseNotInductive j))
-let error_number_branches k env cj expn =
- raise (TypeError (k, env, NumberBranches (nfj cj,expn)))
+let error_number_branches env cj expn =
+ raise (TypeError (env, NumberBranches (nfj cj,expn)))
-let error_ill_formed_branch k env c i actty expty =
- raise (TypeError (k, env,
+let error_ill_formed_branch env c i actty expty =
+ raise (TypeError (env,
IllFormedBranch (c,i,nf_betaiota actty, nf_betaiota expty)))
-let error_generalization k env nvar c =
- raise (TypeError (k, env, Generalization (nvar,c)))
+let error_generalization env nvar c =
+ raise (TypeError (env, Generalization (nvar,c)))
-let error_actual_type k env c actty expty =
- raise (TypeError (k, env, ActualType (c,actty,expty)))
+let error_actual_type env j expty =
+ raise (TypeError (env, ActualType (j,expty)))
-let error_cant_apply_not_functional k env rator randl =
- raise (TypeError (k, env, CantApplyNonFunctional (rator,randl)))
+let error_cant_apply_not_functional env rator randl =
+ raise (TypeError (env, CantApplyNonFunctional (rator,randl)))
-let error_cant_apply_bad_type k env t rator randl =
- raise(TypeError (k, env, CantApplyBadType (t,rator,randl)))
+let error_cant_apply_bad_type env t rator randl =
+ raise(TypeError (env, CantApplyBadType (t,rator,randl)))
-let error_ill_formed_rec_body k env why lna i vdefs =
- raise (TypeError (k, env, IllFormedRecBody (why,lna,i,vdefs)))
+let error_ill_formed_rec_body env why lna i vdefs =
+ raise (TypeError (env, IllFormedRecBody (why,lna,i,vdefs)))
-let error_ill_typed_rec_body k env i lna vdefj vargs =
- raise (TypeError (k, env, IllTypedRecBody (i,lna,vdefj,vargs)))
+let error_ill_typed_rec_body env i lna vdefj vargs =
+ raise (TypeError (env, IllTypedRecBody (i,lna,vdefj,vargs)))
diff --git a/kernel/type_errors.mli b/kernel/type_errors.mli
index 11729171b..c342ce892 100644
--- a/kernel/type_errors.mli
+++ b/kernel/type_errors.mli
@@ -9,7 +9,6 @@
(*i $Id$ i*)
(*i*)
-open Pp
open Names
open Term
open Sign
@@ -41,62 +40,63 @@ type guard_error =
type type_error =
| UnboundRel of int
| NotAType of unsafe_judgment
- | BadAssumption of constr
- | ReferenceVariables of identifier
+ | BadAssumption of unsafe_judgment
+ | ReferenceVariables of constr
| ElimArity of inductive * constr list * constr * unsafe_judgment
* (constr * constr * string) option
| CaseNotInductive of unsafe_judgment
+ | WrongCaseInfo of inductive * case_info
| NumberBranches of unsafe_judgment * int
| IllFormedBranch of constr * int * constr * constr
| Generalization of (name * types) * unsafe_judgment
- | ActualType of constr * constr * constr
+ | ActualType of unsafe_judgment * types
| CantApplyBadType of (int * constr * constr)
- * unsafe_judgment * unsafe_judgment list
- | CantApplyNonFunctional of unsafe_judgment * unsafe_judgment list
+ * unsafe_judgment * unsafe_judgment array
+ | CantApplyNonFunctional of unsafe_judgment * unsafe_judgment array
| IllFormedRecBody of guard_error * name array * int * constr array
| IllTypedRecBody of int * name array * unsafe_judgment array
* types array
-exception TypeError of path_kind * env * type_error
+exception TypeError of env * type_error
-val error_unbound_rel : path_kind -> env -> int -> 'a
+val error_unbound_rel : env -> int -> 'a
-val error_not_type : path_kind -> env -> unsafe_judgment -> 'a
+val error_not_type : env -> unsafe_judgment -> 'a
-val error_assumption : path_kind -> env -> constr -> 'a
+val error_assumption : env -> unsafe_judgment -> 'a
-val error_reference_variables : path_kind -> env -> identifier -> 'a
+val error_reference_variables : env -> constr -> 'a
val error_elim_arity :
- path_kind -> env -> inductive -> constr list -> constr
+ env -> inductive -> constr list -> constr
-> unsafe_judgment -> (constr * constr * string) option -> 'a
val error_case_not_inductive :
- path_kind -> env -> unsafe_judgment -> 'a
+ env -> unsafe_judgment -> 'a
val error_number_branches :
- path_kind -> env -> unsafe_judgment -> int -> 'a
+ env -> unsafe_judgment -> int -> 'a
val error_ill_formed_branch :
- path_kind -> env -> constr -> int -> constr -> constr -> 'a
+ env -> constr -> int -> constr -> constr -> 'a
val error_generalization :
- path_kind -> env -> name * types -> unsafe_judgment -> 'a
+ env -> name * types -> unsafe_judgment -> 'a
val error_actual_type :
- path_kind -> env -> constr -> constr -> constr -> 'a
+ env -> unsafe_judgment -> types -> 'a
val error_cant_apply_not_functional :
- path_kind -> env -> unsafe_judgment -> unsafe_judgment list -> 'a
+ env -> unsafe_judgment -> unsafe_judgment array -> 'a
val error_cant_apply_bad_type :
- path_kind -> env -> int * constr * constr ->
- unsafe_judgment -> unsafe_judgment list -> 'a
+ env -> int * constr * constr ->
+ unsafe_judgment -> unsafe_judgment array -> 'a
val error_ill_formed_rec_body :
- path_kind -> env -> guard_error -> name array -> int -> constr array -> 'a
+ env -> guard_error -> name array -> int -> constr array -> 'a
val error_ill_typed_rec_body :
- path_kind -> env -> int -> name array -> unsafe_judgment array
+ env -> int -> name array -> unsafe_judgment array
-> types array -> 'a
diff --git a/kernel/typeops.ml b/kernel/typeops.ml
index e8e8f35b9..a2c6fe686 100644
--- a/kernel/typeops.ml
+++ b/kernel/typeops.ml
@@ -8,7 +8,6 @@
(* $Id$ *)
-open Pp
open Util
open Names
open Univ
@@ -18,48 +17,42 @@ open Sign
open Environ
open Reduction
open Inductive
-
open Type_errors
-let make_judge v tj =
- { uj_val = v;
- uj_type = tj }
-
-let j_val j = j.uj_val
-(* This should be a type intended to be assumed *)
-let assumption_of_judgment env sigma j =
- match kind_of_term(whd_betadeltaiota env sigma (body_of_type j.uj_type)) with
- | IsSort s -> j.uj_val
- | _ -> error_assumption CCI env j.uj_val
+(* 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 (body_of_type j.uj_type)) with
+ | Sort s -> {utj_val = j.uj_val; utj_type = s }
+ | _ -> error_not_type env j
+
+(* 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 j =
+ try (type_judgment env j).utj_val
+ with TypeError _ ->
+ error_assumption env j
(*
let aojkey = Profile.declare_profile "assumption_of_judgment";;
-let assumption_of_judgment env sigma j
- = Profile.profile3 aojkey assumption_of_judgment env sigma j;;
+let assumption_of_judgment env j
+ = Profile.profile2 aojkey assumption_of_judgment env j;;
*)
-(* This should be a type (a priori without intension to be an assumption) *)
-let type_judgment env sigma j =
- match kind_of_term(whd_betadeltaiota env sigma (body_of_type j.uj_type)) with
- | IsSort s -> {utj_val = j.uj_val; utj_type = s }
- | _ -> error_not_type CCI env j
-
-
(************************************************)
(* Incremental typing rules: builds a typing judgement given the *)
(* judgements for the subterms. *)
-(* Type of sorts *)
+(*s Type of sorts *)
(* Prop and Set *)
let judge_of_prop =
- { uj_val = mkSort prop;
+ { uj_val = mkProp;
uj_type = mkSort type_0 }
let judge_of_set =
- { uj_val = mkSort spec;
+ { uj_val = mkSet;
uj_type = mkSort type_0 }
let judge_of_prop_contents = function
@@ -70,92 +63,84 @@ let judge_of_prop_contents = function
let judge_of_type u =
let (uu,c) = super u in
- { uj_val = mkSort (Type u);
- uj_type = mkSort (Type uu) },
+ { uj_val = mkType u;
+ uj_type = mkType uu },
c
-(*
-let type_of_sort c =
- match kind_of_term c with
- | IsSort (Type u) -> let (uu,cst) = super u in Type uu, cst
- | IsSort (Prop _) -> Type prop_univ, Constraint.empty
- | _ -> invalid_arg "type_of_sort"
-*)
-
-(* Type of a de Bruijn index. *)
+(*s Type of a de Bruijn index. *)
-let relative env n =
+let judge_of_relative env n =
try
- let (_,typ) = lookup_rel_type n env in
+ let (_,_,typ) = lookup_rel n env in
{ uj_val = mkRel n;
uj_type = type_app (lift n) typ }
with Not_found ->
- error_unbound_rel CCI env n
+ error_unbound_rel env n
(*
-let relativekey = Profile.declare_profile "relative";;
-let relative env sigma n = Profile.profile3 relativekey relative env sigma n;;
+let relativekey = Profile.declare_profile "judge_of_relative";;
+let judge_of_relative env n =
+ Profile.profile2 relativekey judge_of_relative env n;;
*)
(* Management of context of variables. *)
-(* Checks if a context of variable is included in another one. *)
-(*
-let rec hyps_inclusion env sigma sign1 sign2 =
- if sign1 = empty_named_context then true
- else
- let (id1,ty1) = hd_sign sign1 in
- let rec search sign2 =
- if sign2 = empty_named_context then false
- else
- let (id2,ty2) = hd_sign sign2 in
- if id1 = id2 then
- (is_conv env sigma (body_of_type ty1) (body_of_type ty2))
- &
- hyps_inclusion env sigma (tl_sign sign1) (tl_sign sign2)
- else
- search (tl_sign sign2)
- in
- search sign2
-*)
+(* Checks if a context of variable can be instanciated by the
+ variables of the current env *)
+(* TODO: check order? *)
+let rec check_hyps_inclusion env sign =
+ let env_sign = named_context env in
+ Sign.fold_named_context
+ (fun (id,_,ty1) () ->
+ let (_,_,ty2) = Sign.lookup_named id env_sign in
+ if not (eq_constr ty2 ty1) then
+ error "types do not match")
+ sign
+ ()
+
+
+let check_args env c hyps =
+ let hyps' = named_context env in
+ try check_hyps_inclusion env hyps
+ with UserError _ | Not_found ->
+ error_reference_variables env c
+
(* Checks if the given context of variables [hyps] is included in the
current context of [env]. *)
(*
-let check_hyps id env sigma hyps =
+let check_hyps id env hyps =
let hyps' = named_context env in
- if not (hyps_inclusion env sigma hyps hyps') then
- error_reference_variables CCI env id
+ if not (hyps_inclusion env hyps hyps') then
+ error_reference_variables env id
*)
(* Instantiation of terms on real arguments. *)
-let type_of_constant = Instantiate.constant_type
+(* Type of variables *)
+let judge_of_variable env id =
+ try
+ let (_,_,ty) = lookup_named id env in
+ make_judge (mkVar id) ty
+ with Not_found ->
+ error ("execute: variable " ^ (string_of_id id) ^ " not defined")
+
+(* Type of constants *)
+let judge_of_constant env cst =
+ let constr = mkConst cst in
+ let _ =
+ let ce = lookup_constant cst env in
+ check_args env constr ce.const_hyps in
+ make_judge constr (constant_type env cst)
(*
let tockey = Profile.declare_profile "type_of_constant";;
-let type_of_constant env sigma c
- = Profile.profile3 tockey type_of_constant env sigma c;;
+let type_of_constant env c
+ = Profile.profile3 tockey type_of_constant env c;;
*)
-(* Type of an existential variable. Not used in kernel. *)
-let type_of_existential env sigma ev =
- Instantiate.existential_type sigma ev
-
-
(* Type of a lambda-abstraction. *)
-let sort_of_product domsort rangsort g =
- match rangsort with
- (* Product rule (s,Prop,Prop) *)
- | Prop _ -> rangsort, Constraint.empty
- | Type u2 ->
- (match domsort with
- (* Product rule (Prop,Type_i,Type_i) *)
- | Prop _ -> rangsort, Constraint.empty
- (* Product rule (Type_i,Type_i,Type_i) *)
- | Type u1 -> let (u12,cst) = sup u1 u2 g in Type u12, cst)
-
-(* [abs_rel env sigma name var j] implements the rule
+(* [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
-----------------------------------------------------------------------
@@ -165,788 +150,335 @@ let sort_of_product domsort rangsort g =
and no upper constraint exists on the sort $s$, we don't need to compute $s$
*)
-let abs_rel env sigma name var j =
- { uj_val = mkLambda (name, var, j.uj_val);
- uj_type = mkProd (name, var, j.uj_type) },
- Constraint.empty
+let judge_of_abstraction env name var j =
+ { uj_val = mkLambda (name, var.utj_val, j.uj_val);
+ uj_type = mkProd (name, var.utj_val, j.uj_type) }
+
+(* Type of let-in. *)
-let judge_of_letin env sigma name defj j =
+let judge_of_letin env name defj j =
let v = match kind_of_term defj.uj_val with
- IsCast(c,t) -> c
+ Cast(c,t) -> c
| _ -> defj.uj_val in
- ({ uj_val = mkLetIn (name, v, defj.uj_type, j.uj_val) ;
- uj_type = type_app (subst1 v) j.uj_type },
- Constraint.empty)
-
-(* [gen_rel env sigma name (typ1,s1) (typ2,s2)] implements the rule
-
- env |- typ1:s1 env, name:typ |- typ2 : s2
- -------------------------------------------------------------------------
- s' >= (s1,s2), env |- (name:typ)j.uj_val : s'
-
- where j.uj_type is convertible to a sort s2
-*)
+ { uj_val = mkLetIn (name, v, defj.uj_type, j.uj_val) ;
+ uj_type = type_app (subst1 v) j.uj_type }
(* Type of an application. *)
-let apply_rel_list env sigma nocheck argjl funj =
+let judge_of_apply env funj argjv =
let rec apply_rec n typ cst = function
| [] ->
- { uj_val = applist (j_val funj, List.map j_val argjl);
- uj_type = type_app (fun _ -> typ) funj.uj_type },
+ { uj_val = mkApp (j_val funj, Array.map j_val argjv);
+ uj_type = typ },
cst
| hj::restjl ->
- match kind_of_term (whd_betadeltaiota env sigma typ) with
- | IsProd (_,c1,c2) ->
- if nocheck then
- apply_rec (n+1) (subst1 hj.uj_val c2) cst restjl
- else
- (try
- let c = conv_leq env sigma (body_of_type hj.uj_type) c1 in
- let cst' = Constraint.union cst c in
- apply_rec (n+1) (subst1 hj.uj_val c2) cst' restjl
- with NotConvertible ->
- error_cant_apply_bad_type CCI env
- (n,c1,body_of_type hj.uj_type)
- funj argjl)
+ (match kind_of_term (whd_betadeltaiota env typ) with
+ | Prod (_,c1,c2) ->
+ (try
+ let c = conv_leq env hj.uj_type c1 in
+ let cst' = Constraint.union cst c in
+ apply_rec (n+1) (subst1 hj.uj_val c2) cst' restjl
+ with NotConvertible ->
+ error_cant_apply_bad_type env
+ (n,c1,body_of_type hj.uj_type)
+ funj argjv)
| _ ->
- error_cant_apply_not_functional CCI env funj argjl
+ error_cant_apply_not_functional env funj argjv)
in
- apply_rec 1 (body_of_type funj.uj_type) Constraint.empty argjl
+ apply_rec 1
+ funj.uj_type
+ Constraint.empty
+ (Array.to_list argjv)
(*
-let applykey = Profile.declare_profile "apply_rel_list";;
-let apply_rel_list env sigma nocheck argjl funj
- = Profile.profile5 applykey apply_rel_list env sigma nocheck argjl funj;;
+let applykey = Profile.declare_profile "judge_of_apply";;
+let judge_of_apply env nocheck funj argjl
+ = Profile.profile5 applykey judge_of_apply env nocheck funj argjl;;
*)
+
(* Type of product *)
-let gen_rel env sigma name t1 t2 =
+
+let sort_of_product domsort rangsort g =
+ match rangsort with
+ (* Product rule (s,Prop,Prop) *)
+ | Prop _ -> rangsort, Constraint.empty
+ | Type u2 ->
+ (match domsort with
+ (* Product rule (Prop,Type_i,Type_i) *)
+ | Prop _ -> rangsort, Constraint.empty
+ (* Product rule (Type_i,Type_i,Type_i) *)
+ | Type u1 -> let (u12,cst) = sup u1 u2 g in Type u12, cst)
+
+(* [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 t1 t2 =
let (s,g) = sort_of_product t1.utj_type t2.utj_type (universes env) in
{ uj_val = mkProd (name, t1.utj_val, t2.utj_val);
uj_type = mkSort s },
g
-(* [cast_rel env sigma (typ1,s1) j] implements the rule
+(* Type of a type cast *)
+
+(* [judge_of_cast env (c,typ1) (typ2,s)] implements the rule
- env, sigma |- cj.uj_val:cj.uj_type cst, env, sigma |- cj.uj_type = t
+ env |- c:typ1 env |- typ2:s env |- typ1 <= typ2
---------------------------------------------------------------------
- cst, env, sigma |- cj.uj_val:t
+ env |- c:typ2
*)
-(* Type of casts *)
-let cast_rel env sigma cj t =
+let judge_of_cast env cj tj =
try
- let cst = conv_leq env sigma (body_of_type cj.uj_type) t in
+ let cst = conv_leq env cj.uj_type tj.utj_val in
{ uj_val = j_val cj;
- uj_type = t },
+ uj_type = tj.utj_val },
cst
with NotConvertible ->
- error_actual_type CCI env cj.uj_val (body_of_type cj.uj_type) t
+ error_actual_type env cj tj.utj_val
(* Inductive types. *)
-let type_of_inductive env sigma i =
- (* TODO: check args *)
- mis_arity (lookup_mind_specif i env)
+let judge_of_inductive env i =
+ let constr = mkInd i in
+ let _ =
+ let (sp,_) = i in
+ let mib = lookup_mind sp env in
+ check_args env constr mib.mind_hyps in
+ make_judge constr (type_of_inductive env i)
(*
-let toikey = Profile.declare_profile "type_of_inductive";;
-let type_of_inductive env sigma i
- = Profile.profile3 toikey type_of_inductive env sigma i;;
+let toikey = Profile.declare_profile "judge_of_inductive";;
+let judge_of_inductive env i
+ = Profile.profile2 toikey judge_of_inductive env i;;
*)
(* Constructors. *)
-let type_of_constructor env sigma cstr =
- mis_constructor_type
- (index_of_constructor cstr)
- (lookup_mind_specif (inductive_of_constructor cstr) env)
+let judge_of_constructor env c =
+ let constr = mkConstruct c in
+ let _ =
+ let ((sp,_),_) = c in
+ let mib = lookup_mind sp env in
+ check_args env constr mib.mind_hyps in
+ make_judge constr (type_of_constructor env c)
(*
-let tockey = Profile.declare_profile "type_of_constructor";;
-let type_of_constructor env sigma cstr
- = Profile.profile3 tockey type_of_constructor env sigma cstr;;
+let tockey = Profile.declare_profile "judge_of_constructor";;
+let judge_of_constructor env cstr
+ = Profile.profile2 tockey judge_of_constructor env cstr;;
*)
(* Case. *)
-let rec mysort_of_arity env sigma c =
- match kind_of_term (whd_betadeltaiota env sigma c) with
- | IsSort s -> s
- | IsProd(_,_,c2) -> mysort_of_arity env sigma c2
- | _ -> invalid_arg "mysort_of_arity"
-
-let error_elim_expln env sigma kp ki =
- if is_info_arity env sigma kp && not (is_info_arity env sigma ki) then
- "non-informative objects may not construct informative ones."
- else
- match (kind_of_term kp,kind_of_term ki) with
- | IsSort (Type _), IsSort (Prop _) ->
- "strong elimination on non-small inductive types leads to paradoxes."
- | _ -> "wrong arity"
-
-exception Arity of (constr * constr * string) option
-
-let is_correct_arity env sigma kelim (c,pj) indf t =
- let rec srec (pt,t) u =
- let pt' = whd_betadeltaiota env sigma pt in
- let t' = whd_betadeltaiota env sigma t in
- match kind_of_term pt', kind_of_term t' with
- | IsProd (_,a1,a2), IsProd (_,a1',a2') ->
- let univ =
- try conv env sigma a1 a1'
- with NotConvertible -> raise (Arity None) in
- srec (a2,a2') (Constraint.union u univ)
- | IsProd (_,a1,a2), _ ->
- let k = whd_betadeltaiota env sigma a2 in
- let ksort = match kind_of_term k with
- | IsSort s -> family_of_sort s
- | _ -> raise (Arity None) in
- let ind = build_dependent_inductive indf in
- let univ =
- try conv env sigma a1 ind
- with NotConvertible -> raise (Arity None) in
- if List.exists ((=) ksort) kelim then
- ((true,k), Constraint.union u univ)
- else
- raise (Arity (Some(k,t',error_elim_expln env sigma k t')))
- | k, IsProd (_,_,_) ->
- raise (Arity None)
- | k, ki ->
- let ksort = match k with
- | IsSort s -> family_of_sort s
- | _ -> raise (Arity None) in
- if List.exists ((=) ksort) kelim then
- (false, pt'), u
- else
- raise (Arity (Some(pt',t',error_elim_expln env sigma pt' t')))
- in
- try srec (pj.uj_type,t) Constraint.empty
- with Arity kinds ->
- let create_sort = function
- | InProp -> prop
- | InSet -> spec
- | InType -> Type (Univ.new_univ ()) in
- let listarity =
- (List.map (fun s -> make_arity env true indf (create_sort s)) kelim)
- @(List.map (fun s -> make_arity env false indf (create_sort s)) kelim)
- in
- let ind = mis_inductive (fst (dest_ind_family indf)) in
- error_elim_arity CCI env ind listarity c pj kinds
-
-
-let find_case_dep_nparams env sigma (c,pj) (IndFamily (mis,_) as indf) =
- let kelim = mis_kelim mis in
- let arsign,s = get_arity indf in
- let glob_t = it_mkProd_or_LetIn (mkSort s) arsign in
- let ((dep,_),univ) =
- is_correct_arity env sigma kelim (c,pj) indf glob_t in
- (dep,univ)
-
-(* type_case_branches type un <p>Case c of ... end
- IndType (indf,realargs) = type de c
- pt = sorte de p
- type_case_branches retourne (lb, lr); lb est le vecteur des types
- attendus dans les branches du Case; lr est le type attendu du resultat
- *)
-
-let type_case_branches env sigma (IndType (indf,realargs)) pj c =
- let p = pj.uj_val in
- let (dep,univ) = find_case_dep_nparams env sigma (c,pj) indf in
- let constructs = get_constructors indf in
- let lc = Array.map (build_branch_type env dep p) constructs in
- if dep then
- (lc, beta_applist (p,(realargs@[c])), univ)
- else
- (lc, beta_applist (p,realargs), univ)
-
-let check_branches_message env sigma cj (explft,lft) =
- let expn = Array.length explft and n = Array.length lft in
- if n<>expn then error_number_branches CCI env cj expn;
- let univ = ref Constraint.empty in
- (for i = 0 to n-1 do
- try
- univ := Constraint.union !univ
- (conv_leq env sigma lft.(i) (explft.(i)))
- with NotConvertible ->
- error_ill_formed_branch CCI env cj.uj_val i lft.(i) explft.(i)
- done;
- !univ)
-
-let nparams_of (IndType (IndFamily (mis,_),_)) = mis_nparams mis
-
-let judge_of_case env sigma (np,_ as ci) pj cj lfj =
- let lft = Array.map (fun j -> body_of_type j.uj_type) lfj in
+let check_branch_types env cj (lft,explft) =
+ try conv_leq_vecti env lft explft
+ with
+ NotConvertibleVect i ->
+ error_ill_formed_branch env cj.uj_val i lft.(i) explft.(i)
+ | Invalid_argument _ ->
+ error_number_branches env cj (Array.length explft)
+
+let judge_of_case env ci pj cj lfj =
let indspec =
- try find_rectype env sigma (body_of_type cj.uj_type)
- with Induc -> error_case_not_inductive CCI env cj in
- if np <> nparams_of indspec then
- anomaly "judge_of_case: wrong parameters number";
- let (bty,rslty,univ) = type_case_branches env sigma indspec pj cj.uj_val in
- let kind = mysort_of_arity env sigma (body_of_type pj.uj_type) in
- let univ' = check_branches_message env sigma cj (bty,lft) in
- ({ uj_val = mkMutCase (ci, (nf_betaiota pj.uj_val), cj.uj_val, Array.map j_val lfj);
+ try find_rectype env cj.uj_type
+ with Induc -> error_case_not_inductive env cj in
+ let _ = check_case_info env (fst indspec) ci in
+ let (bty,rslty,univ) =
+ type_case_branches env indspec pj cj.uj_val in
+ let (_,kind) = dest_arity env pj.uj_type in
+ let lft = Array.map j_type lfj in
+ let univ' = check_branch_types env cj (lft,bty) in
+ ({ uj_val = mkCase (ci, nf_betaiota pj.uj_val, cj.uj_val,
+ Array.map j_val lfj);
uj_type = rslty },
Constraint.union univ univ')
(*
-let tocasekey = Profile.declare_profile "type_of_case";;
-let type_of_case env sigma ci pj cj lfj
- = Profile.profile6 tocasekey type_of_case env sigma ci pj cj lfj;;
+let tocasekey = Profile.declare_profile "judge_of_case";;
+let judge_of_case env ci pj cj lfj
+ = Profile.profile6 tocasekey judge_of_case env ci pj cj lfj;;
*)
(* Fixpoints. *)
-(* Check if t is a subterm of Rel n, and gives its specification,
- assuming lst already gives index of
- subterms with corresponding specifications of recursive arguments *)
-
-(* A powerful notion of subterm *)
-
-let find_sorted_assoc p =
- let rec findrec = function
- | (a,ta)::l ->
- if a < p then findrec l else if a = p then ta else raise Not_found
- | _ -> raise Not_found
- in
- findrec
-
-let map_lift_fst_n m = List.map (function (n,t)->(n+m,t))
-let map_lift_fst = map_lift_fst_n 1
-
-let rec instantiate_recarg sp lrc ra =
- match ra with
- | Mrec(j) -> Imbr((sp,j),lrc)
- | Imbr(ind_sp,l) -> Imbr(ind_sp, List.map (instantiate_recarg sp lrc) l)
- | Norec -> Norec
- | Param(k) -> List.nth lrc k
-
-(* To each inductive definition corresponds an array describing the
- structure of recursive arguments for each constructor, we call it
- the recursive spec of the type (it has type recargs vect). For
- checking the guard, we start from the decreasing argument (Rel n)
- with its recursive spec. During checking the guardness condition,
- we collect patterns variables corresponding to subterms of n, each
- of them with its recursive spec. They are organised in a list lst
- of type (int * recargs) list which is sorted with respect to the
- first argument.
-*)
-
-(*
- f is a function of type
- env -> int -> (int * recargs) list -> constr -> 'a
-
- c is a branch of an inductive definition corresponding to the spec
- lrec. mind_recvec is the recursive spec of the inductive
- definition of the decreasing argument n.
-
- check_term env mind_recvec f n lst (lrec,c) will pass the lambdas
- of c corresponding to pattern variables and collect possibly new
- subterms variables and apply f to the body of the branch with the
- correct env and decreasing arg.
-*)
-
-let check_term env mind_recvec f =
- let rec crec env n lst (lrec,c) =
- let c' = strip_outer_cast c in
- match lrec, kind_of_term c' with
- (ra::lr,IsLambda (x,a,b)) ->
- let lst' = map_lift_fst lst
- and env' = push_rel_assum (x,a) env
- and n'=n+1
- in begin match ra with
- Mrec(i) -> crec env' n' ((1,mind_recvec.(i))::lst') (lr,b)
- | Imbr((sp,i) as ind_sp,lrc) ->
- let sprecargs =
- mis_recargs (lookup_mind_specif ind_sp env) in
- let lc = Array.map
- (List.map (instantiate_recarg sp lrc)) sprecargs.(i)
- in crec env' n' ((1,lc)::lst') (lr,b)
- | _ -> crec env' n' lst' (lr,b) end
- | (_,IsLetIn (x,c,a,b)) ->
- let env' = push_rel_def (x,c,a) env in
- crec env' (n+1) (map_lift_fst lst) (lrec,(subst1 c b))
- | (_,_) -> f env n lst c'
- in crec env
-
-(* c is supposed to be in beta-delta-iota head normal form *)
-
-let is_inst_var k c =
- match kind_of_term (fst (decomp_app c)) with
- | IsRel n -> n=k
- | _ -> false
-
-(*
- is_subterm_specif env sigma lcx mind_recvec n lst c
-
- n is the principal arg and has recursive spec lcx, lst is the list
- of subterms of n with spec. is_subterm_specif should test if c is
- a subterm of n and fails with Not_found if not. In case it is, it
- should send its recursive specification. This recursive spec
- should be the same size as the number of constructors of the type
- of c. A problem occurs when c is built by contradiction. In that
- case no spec is given.
-
-*)
-let is_subterm_specif env sigma lcx mind_recvec =
- let rec crec env n lst c =
- let f,l = whd_betadeltaiota_stack env sigma c in
- match kind_of_term f with
- | IsRel k -> Some (find_sorted_assoc k lst)
-
- | IsMutCase ( _,_,c,br) ->
- if Array.length br = 0 then None
-
- else
- let def = Array.create (Array.length br) []
- in let lcv =
- (try
- if is_inst_var n c then lcx
- else match crec env n lst c with Some lr -> lr | None -> def
- with Not_found -> def)
- in
- assert (Array.length br = Array.length lcv);
- let stl =
- array_map2
- (fun lc a ->
- check_term env mind_recvec crec n lst (lc,a)) lcv br
- in let stl0 = stl.(0) in
- if array_for_all (fun st -> st=stl0) stl then stl0
- else None
-
- | IsFix ((recindxs,i),(_,typarray,bodies as recdef)) ->
- let nbfix = Array.length typarray in
- let decrArg = recindxs.(i) in
- let theBody = bodies.(i) in
- let sign,strippedBody = decompose_lam_n_assum (decrArg+1) theBody in
- let nbOfAbst = nbfix+decrArg+1 in
-(* when proving that the fixpoint f(x)=e is less than n, it is enough
- to prove that e is less than n assuming f is less than n
- furthermore when f is applied to a term which is strictly less than
- n, one may assume that x itself is strictly less than n
-*)
- let newlst =
- let lst' = (nbOfAbst,lcx) :: (map_lift_fst_n nbOfAbst lst) in
- if List.length l < (decrArg+1) then lst'
- else let theDecrArg = List.nth l decrArg in
- try
- match crec env n lst theDecrArg with
- (Some recArgsDecrArg) -> (1,recArgsDecrArg) :: lst'
- | None -> lst'
- with Not_found -> lst'
- in let env' = push_rec_types recdef env in
- let env'' = push_rels sign env' in
- crec env'' (n+nbOfAbst) newlst strippedBody
-
- | IsLambda (x,a,b) when l=[] ->
- let lst' = map_lift_fst lst in
- crec (push_rel_assum (x, a) env) (n+1) lst' b
-
- (*** Experimental change *************************)
- | IsMeta _ -> None
- | _ -> raise Not_found
- in
- crec env
-
-let spec_subterm_strict env sigma lcx mind_recvec n lst c nb =
- try match is_subterm_specif env sigma lcx mind_recvec n lst c
- with Some lr -> lr | None -> Array.create nb []
- with Not_found -> Array.create nb []
+(* Checks the type of a general (co)fixpoint, i.e. without checking *)
+(* the specific guard condition. *)
-let spec_subterm_large env sigma lcx mind_recvec n lst c nb =
- if is_inst_var n c then lcx
- else spec_subterm_strict env sigma lcx mind_recvec n lst c nb
-
-
-let is_subterm env sigma lcx mind_recvec n lst c =
- try
- let _ = is_subterm_specif env sigma lcx mind_recvec n lst c in true
- with Not_found ->
- false
-
-
-exception FixGuardError of guard_error
-
-(* Auxiliary function: it checks a condition f depending on a deBrujin
- index for a certain number of abstractions *)
-
-let rec check_subterm_rec_meta env sigma vectn k def =
- (* If k<0, it is a general fixpoint *)
- (k < 0) or
- (let nfi = Array.length vectn in
- (* check fi does not appear in the k+1 first abstractions,
- gives the type of the k+1-eme abstraction *)
- let rec check_occur env n def =
- match kind_of_term (strip_outer_cast def) with
- | IsLambda (x,a,b) ->
- if noccur_with_meta n nfi a then
- let env' = push_rel_assum (x, a) env in
- if n = k+1 then (env', lift 1 a, b)
- else check_occur env' (n+1) b
- else
- anomaly "check_subterm_rec_meta: Bad occurrence of recursive call"
- | _ -> raise (FixGuardError NotEnoughAbstractionInFixBody) in
- let (env',c,d) = check_occur env 1 def in
- let ((sp,tyi) as mind, largs) =
- try find_inductive env' sigma c
- with Induc -> raise (FixGuardError RecursionNotOnInductiveType) in
- let mind_recvec = mis_recargs (lookup_mind_specif mind env') in
- let lcx = mind_recvec.(tyi) in
- (* n = decreasing argument in the definition;
- lst = a mapping var |-> recargs
- t = the term to be checked
- *)
- let rec check_rec_call env n lst t =
- (* n gives the index of the recursive variable *)
- (noccur_with_meta (n+k+1) nfi t) or
- (* no recursive call in the term *)
- (let f,l = whd_betaiotazeta_stack t in
- match kind_of_term f with
- | IsRel p ->
- if n+k+1 <= p & p < n+k+nfi+1 then
- (* recursive call *)
- let glob = nfi+n+k-p in (* the index of the recursive call *)
- let np = vectn.(glob) in (* the decreasing arg of the rec call *)
- if List.length l > np then
- (match list_chop np l with
- (la,(z::lrest)) ->
- if (is_subterm env sigma lcx mind_recvec n lst z)
- then List.for_all (check_rec_call env n lst) (la@lrest)
- else raise (FixGuardError RecursionOnIllegalTerm)
- | _ -> assert false)
- else raise (FixGuardError NotEnoughArgumentsForFixCall)
- else List.for_all (check_rec_call env n lst) l
-
- | IsMutCase (ci,p,c_0,lrest) ->
- let lc = spec_subterm_large env sigma lcx mind_recvec n lst c_0
- (Array.length lrest)
- in
- (array_for_all2
- (fun c0 a ->
- check_term env mind_recvec check_rec_call n lst (c0,a))
- lc lrest)
- && (List.for_all (check_rec_call env n lst) (c_0::p::l))
-
- (* Enables to traverse Fixpoint definitions in a more intelligent
- way, ie, the rule :
-
- if - g = Fix g/1 := [y1:T1]...[yp:Tp]e &
- - f is guarded with respect to the set of pattern variables S
- in a1 ... am &
- - f is guarded with respect to the set of pattern variables S
- in T1 ... Tp &
- - ap is a sub-term of the formal argument of f &
- - f is guarded with respect to the set of pattern variables S+{yp}
- in e
- then f is guarded with respect to S in (g a1 ... am).
-
- Eduardo 7/9/98 *)
-
- | IsFix ((recindxs,i),(_,typarray,bodies as recdef)) ->
- (List.for_all (check_rec_call env n lst) l) &&
- (array_for_all (check_rec_call env n lst) typarray) &&
- let nbfix = Array.length typarray in
- let decrArg = recindxs.(i)
- and env' = push_rec_types recdef env
- and n' = n+nbfix
- and lst' = map_lift_fst_n nbfix lst
- in
- if (List.length l < (decrArg+1)) then
- array_for_all (check_rec_call env' n' lst') bodies
- else
- let theDecrArg = List.nth l decrArg in
- (try
- match
- is_subterm_specif env sigma lcx mind_recvec n lst theDecrArg
- with
- Some recArgsDecrArg ->
- let theBody = bodies.(i) in
- check_rec_call_fix_body
- env' n' lst' (decrArg+1) recArgsDecrArg theBody
- | None ->
- array_for_all (check_rec_call env' n' lst') bodies
- with Not_found ->
- array_for_all (check_rec_call env' n' lst') bodies)
-
- | IsCast (a,b) ->
- (check_rec_call env n lst a) &&
- (check_rec_call env n lst b) &&
- (List.for_all (check_rec_call env n lst) l)
-
- | IsLambda (x,a,b) ->
- (check_rec_call env n lst a) &&
- (check_rec_call (push_rel_assum (x, a) env)
- (n+1) (map_lift_fst lst) b) &&
- (List.for_all (check_rec_call env n lst) l)
-
- | IsProd (x,a,b) ->
- (check_rec_call env n lst a) &&
- (check_rec_call (push_rel_assum (x, a) env)
- (n+1) (map_lift_fst lst) b) &&
- (List.for_all (check_rec_call env n lst) l)
-
- | IsLetIn (x,a,b,c) ->
- anomaly "check_rec_call: should have been reduced"
-
- | IsMutInd _ ->
- (List.for_all (check_rec_call env n lst) l)
-
- | IsMutConstruct _ ->
- (List.for_all (check_rec_call env n lst) l)
-
- | IsConst sp ->
- (try
- (List.for_all (check_rec_call env n lst) l)
- with (FixGuardError _ ) as e
- -> if evaluable_constant env sp then
- check_rec_call env n lst (whd_betadeltaiota env sigma t)
- else raise e)
-
- | IsApp (f,la) ->
- (check_rec_call env n lst f) &&
- (array_for_all (check_rec_call env n lst) la) &&
- (List.for_all (check_rec_call env n lst) l)
-
- | IsCoFix (i,(_,typarray,bodies as recdef)) ->
- let nbfix = Array.length typarray in
- let env' = push_rec_types recdef env in
- (array_for_all (check_rec_call env n lst) typarray) &&
- (List.for_all (check_rec_call env n lst) l) &&
- (array_for_all
- (check_rec_call env' (n+nbfix) (map_lift_fst_n nbfix lst))
- bodies)
-
- | IsEvar (_,la) ->
- (array_for_all (check_rec_call env n lst) la) &&
- (List.for_all (check_rec_call env n lst) l)
-
- | IsMeta _ -> true
-
- | IsVar _ | IsSort _ -> List.for_all (check_rec_call env n lst) l
- )
-
- and check_rec_call_fix_body env n lst decr recArgsDecrArg body =
- if decr = 0 then
- check_rec_call env n ((1,recArgsDecrArg)::lst) body
- else
- match kind_of_term body with
- | IsLambda (x,a,b) ->
- (check_rec_call env n lst a) &
- (check_rec_call_fix_body
- (push_rel_assum (x, a) env) (n+1)
- (map_lift_fst lst) (decr-1) recArgsDecrArg b)
- | _ -> anomaly "Not enough abstractions in fix body"
-
- in
- check_rec_call env' 1 [] d)
-
-(* vargs is supposed to be built from A1;..Ak;[f1]..[fk][|d1;..;dk|]
-and vdeft is [|t1;..;tk|] such that f1:A1,..,fk:Ak |- di:ti
-nvect is [|n1;..;nk|] which gives for each recursive definition
-the inductive-decreasing index
-the function checks the convertibility of ti with Ai *)
-
-let check_fix env sigma ((nvect,bodynum),(names,types,bodies as recdef)) =
- let nbfix = Array.length bodies in
- if nbfix = 0
- or Array.length nvect <> nbfix
- or Array.length types <> nbfix
- or Array.length names <> nbfix
- or bodynum < 0
- or bodynum >= nbfix
- then anomaly "Ill-formed fix term";
- for i = 0 to nbfix - 1 do
- let fixenv = push_rec_types recdef env in
- try
- let _ = check_subterm_rec_meta fixenv sigma nvect nvect.(i) bodies.(i)
- in ()
- with FixGuardError err ->
- error_ill_formed_rec_body CCI fixenv err names i bodies
- done
-
-(*
-let cfkey = Profile.declare_profile "check_fix";;
-let check_fix env sigma fix = Profile.profile3 cfkey check_fix env sigma fix;;
-*)
-
-(* Co-fixpoints. *)
-
-exception CoFixGuardError of guard_error
-
-let check_guard_rec_meta env sigma nbfix def deftype =
- let rec codomain_is_coind env c =
- let b = whd_betadeltaiota env sigma (strip_outer_cast c) in
- match kind_of_term b with
- | IsProd (x,a,b) ->
- codomain_is_coind (push_rel_assum (x, a) env) b
- | _ ->
- try
- find_coinductive env sigma b
- with Induc ->
- raise (CoFixGuardError (CodomainNotInductiveType b))
- in
- let (mind, _) = codomain_is_coind env deftype in
- let (sp,tyi) = mind in
- let lvlra = mis_recargs (lookup_mind_specif mind env) in
- let vlra = lvlra.(tyi) in
- let rec check_rec_call env alreadygrd n vlra t =
- if noccur_with_meta n nbfix t then
- true
- else
- let c,args = whd_betadeltaiota_stack env sigma t in
- match kind_of_term c with
- | IsMeta _ -> true
-
- | IsRel p ->
- if n <= p && p < n+nbfix then
- (* recursive call *)
- if alreadygrd then
- if List.for_all (noccur_with_meta n nbfix) args then
- true
- else
- raise (CoFixGuardError NestedRecursiveOccurrences)
- else
- raise (CoFixGuardError (UnguardedRecursiveCall t))
- else
- error "check_guard_rec_meta: ???" (* ??? *)
-
- | IsMutConstruct (_,i as cstr_sp) ->
- let lra =vlra.(i-1) in
- let mI = inductive_of_constructor cstr_sp in
- let mis = lookup_mind_specif mI env in
- let _,realargs = list_chop (mis_nparams mis) args in
- let rec process_args_of_constr l lra =
- match l with
- | [] -> true
- | t::lr ->
- (match lra with
- | [] ->
- anomalylabstrm "check_guard_rec_meta"
- [< 'sTR "a constructor with an empty list";
- 'sTR "of recargs is being applied" >]
- | (Mrec i)::lrar ->
- let newvlra = lvlra.(i) in
- (check_rec_call env true n newvlra t) &&
- (process_args_of_constr lr lrar)
-
- | (Imbr((sp,i) as ind_sp,lrc)::lrar) ->
- let mis = lookup_mind_specif ind_sp env in
- let sprecargs = mis_recargs mis in
- let lc = (Array.map
- (List.map
- (instantiate_recarg sp lrc))
- sprecargs.(i))
- in (check_rec_call env true n lc t) &
- (process_args_of_constr lr lrar)
-
- | _::lrar ->
- if (noccur_with_meta n nbfix t)
- then (process_args_of_constr lr lrar)
- else raise (CoFixGuardError
- (RecCallInNonRecArgOfConstructor t)))
- in (process_args_of_constr realargs lra)
-
-
- | IsLambda (x,a,b) ->
- assert (args = []);
- if (noccur_with_meta n nbfix a) then
- check_rec_call (push_rel_assum (x, a) env)
- alreadygrd (n+1) vlra b
- else
- raise (CoFixGuardError (RecCallInTypeOfAbstraction t))
-
- | IsCoFix (j,(_,varit,vdefs as recdef)) ->
- if (List.for_all (noccur_with_meta n nbfix) args)
- then
- let nbfix = Array.length vdefs in
- if (array_for_all (noccur_with_meta n nbfix) varit) then
- let env' = push_rec_types recdef env in
- (array_for_all
- (check_rec_call env' alreadygrd (n+1) vlra) vdefs)
- &&
- (List.for_all (check_rec_call env alreadygrd (n+1) vlra) args)
- else
- raise (CoFixGuardError (RecCallInTypeOfDef c))
- else
- raise (CoFixGuardError (UnguardedRecursiveCall c))
-
- | IsMutCase (_,p,tm,vrest) ->
- if (noccur_with_meta n nbfix p) then
- if (noccur_with_meta n nbfix tm) then
- if (List.for_all (noccur_with_meta n nbfix) args) then
- (array_for_all (check_rec_call env alreadygrd n vlra) vrest)
- else
- raise (CoFixGuardError (RecCallInCaseFun c))
- else
- raise (CoFixGuardError (RecCallInCaseArg c))
- else
- raise (CoFixGuardError (RecCallInCasePred c))
-
- | _ -> raise (CoFixGuardError NotGuardedForm)
-
- in
- check_rec_call env false 1 vlra def
-
-(* The function which checks that the whole block of definitions
- satisfies the guarded condition *)
-
-let check_cofix env sigma (bodynum,(names,types,bodies as recdef)) =
- let nbfix = Array.length bodies in
- for i = 0 to nbfix-1 do
- let fixenv = push_rec_types recdef env in
- try
- let _ = check_guard_rec_meta fixenv sigma nbfix bodies.(i) types.(i)
- in ()
- with CoFixGuardError err ->
- error_ill_formed_rec_body CCI fixenv err names i bodies
- done
-
-(* Checks the type of a (co)fixpoint.
- Fix and CoFix are typed the same way; only the guard condition differs. *)
-
-exception IllBranch of int
-
-let type_fixpoint env sigma lna lar vdefj =
+let type_fixpoint env lna lar vdefj =
let lt = Array.length vdefj in
assert (Array.length lar = lt);
try
- conv_forall2_i
- (fun i env sigma def ar ->
- try conv_leq env sigma def (lift lt ar)
- with NotConvertible -> raise (IllBranch i))
- env sigma
+ conv_leq_vecti env
(Array.map (fun j -> body_of_type j.uj_type) vdefj)
- (Array.map body_of_type lar)
- with IllBranch i ->
- error_ill_typed_rec_body CCI env i lna vdefj lar
-
-
-(* A function which checks that a term well typed verifies both
- syntaxic conditions *)
-
-let control_only_guard env sigma =
- let rec control_rec c = match kind_of_term c with
- | IsRel _ | IsVar _ -> ()
- | IsSort _ | IsMeta _ -> ()
- | IsMutInd _ -> ()
- | IsMutConstruct _ -> ()
- | IsConst _ -> ()
- | IsCoFix (_,(_,tys,bds) as cofix) ->
- check_cofix env sigma cofix;
- Array.iter control_rec tys;
- Array.iter control_rec bds;
- | IsFix (_,(_,tys,bds) as fix) ->
- check_fix env sigma fix;
- Array.iter control_rec tys;
- Array.iter control_rec bds;
- | IsMutCase(_,p,c,b) ->control_rec p;control_rec c;Array.iter control_rec b
- | IsEvar (_,cl) -> Array.iter control_rec cl
- | IsApp (_,cl) -> Array.iter control_rec cl
- | IsCast (c1,c2) -> control_rec c1; control_rec c2
- | IsProd (_,c1,c2) -> control_rec c1; control_rec c2
- | IsLambda (_,c1,c2) -> control_rec c1; control_rec c2
- | IsLetIn (_,c1,c2,c3) -> control_rec c1; control_rec c2; control_rec c3
- in
- control_rec
+ (Array.map (fun ty -> lift lt (body_of_type ty)) lar)
+ with NotConvertibleVect i ->
+ error_ill_typed_rec_body env i lna vdefj lar
+
+(***********************************************************************)
+(***********************************************************************)
+
+(* 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,(Constraint.union 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 =
+ match kind_of_term cstr with
+ (* Atomic terms *)
+ | Sort (Prop c) ->
+ (judge_of_prop_contents c, cu)
+
+ | Sort (Type u) ->
+ univ_combinator cu (judge_of_type u)
+
+ | Rel n ->
+ (judge_of_relative env n, cu)
+
+ | Var id ->
+ (judge_of_variable env id, cu)
+
+ | Const c ->
+ (judge_of_constant env c, cu)
+
+ (* Lambda calculus operators *)
+ | App (f,args) ->
+ let (j,cu1) = execute env f cu in
+ let (jl,cu2) = execute_array env args cu1 in
+ univ_combinator cu2
+ (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)
+
+ | 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
+ univ_combinator cu2
+ (judge_of_product env name varj varj')
+
+ | LetIn (name,c1,c2,c3) ->
+ let (j,cu1) = execute env (mkCast(c1,c2)) cu in
+ let env1 = push_rel (name,Some j.uj_val,j.uj_type) env in
+ let (j',cu2) = execute env1 c3 cu1 in
+ (judge_of_letin env name j j', cu2)
+
+ | Cast (c,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 tj)
+
+ (* Inductive types *)
+ | Ind ind ->
+ (judge_of_inductive env ind, cu)
+
+ | Construct c ->
+ (judge_of_constructor env c, cu)
+
+ | 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)
+
+ | Fix ((vn,i as vni),recdef) ->
+ let ((fix_ty,recdef'),cu1) = execute_recdef env recdef i cu in
+ let fix = (vni,recdef') in
+ check_fix env fix;
+ (make_judge (mkFix fix) fix_ty, cu1)
+
+ | CoFix (i,recdef) ->
+ let ((fix_ty,recdef'),cu1) = execute_recdef env recdef i cu in
+ let cofix = (i,recdef') in
+ check_cofix env cofix;
+ (make_judge (mkCoFix cofix) fix_ty, cu1)
+
+ (* Partial proofs: unsupported by the kernel *)
+ | Meta _ ->
+ anomaly "the kernel does not support metavariables"
+
+ | Evar _ ->
+ anomaly "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_recdef env (names,lar,vdef) i cu =
+ let (larj,cu1) = execute_array env lar cu 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 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)
+
+and execute_array env v cu =
+ let (jl,cu1) = execute_list env (Array.to_list v) cu in
+ (Array.of_list jl, cu1)
+
+and execute_list env l cu =
+ match l with
+ | [] ->
+ ([], cu)
+ | c::r ->
+ let (j,cu1) = execute env c cu in
+ let (jr,cu2) = execute_list env r cu1 in
+ (j::jr, cu2)
+
+(* Derived functions *)
+let infer env constr =
+ let (j,(cst,_)) =
+ execute env constr (Constraint.empty, universes env) in
+ (j, cst)
+
+let infer_type env constr =
+ let (j,(cst,_)) =
+ execute_type env constr (Constraint.empty, universes env) in
+ (j, cst)
+
+let infer_v env cv =
+ let (jv,(cst,_)) =
+ execute_array env cv (Constraint.empty, universes env) in
+ (jv, cst)
+
+(* Typing of several terms. *)
+
+type local_entry =
+ | LocalDef of constr
+ | LocalAssum of constr
+
+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
+ | LocalAssum c ->
+ let (j,cst) = infer env c in
+ (Name id, None, assumption_of_judgment env j), cst
+
+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, d :: l, Constraint.union cst1 cst2
+ | [] -> env, [], Constraint.empty in
+ inferec env decls
diff --git a/kernel/typeops.mli b/kernel/typeops.mli
index e4464fd89..24ffa47b1 100644
--- a/kernel/typeops.mli
+++ b/kernel/typeops.mli
@@ -10,97 +10,83 @@
(*i*)
open Names
-open Sign
open Univ
open Term
-open Evd
open Environ
+open Inductive
(*i*)
+(*s 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
-(* Basic operations of the typing machine. *)
+type local_entry =
+ | LocalDef of constr
+ | LocalAssum of constr
-val make_judge : constr -> types -> unsafe_judgment
+val infer_local_decls :
+ env -> (identifier * local_entry) list
+ -> env * Sign.rel_context * constraints
-val j_val : unsafe_judgment -> constr
+(*s Basic operations of the typing machine. *)
(* If [j] is the judgement $c:t$, then [assumption_of_judgement env j]
returns the type $c$, checking that $t$ is a sort. *)
-val assumption_of_judgment :
- env -> 'a evar_map -> unsafe_judgment -> types
-
-val type_judgment :
- env -> 'a evar_map -> unsafe_judgment -> unsafe_type_judgment
+val assumption_of_judgment : env -> unsafe_judgment -> types
+val type_judgment : env -> unsafe_judgment -> unsafe_type_judgment
(*s Type of sorts. *)
val judge_of_prop_contents : contents -> unsafe_judgment
+val judge_of_type : universe -> unsafe_judgment * constraints
-val judge_of_type : universe -> unsafe_judgment * constraints
+(*s Type of a bound variable. *)
+val judge_of_relative : env -> int -> unsafe_judgment
-(*s Type of atomic terms. *)
-val relative : env -> int -> unsafe_judgment
+(*s Type of variables *)
+val judge_of_variable : env -> identifier -> unsafe_judgment
-val type_of_constant : env -> 'a evar_map -> constant -> types
-
-val type_of_existential : env -> 'a evar_map -> existential -> types
-
-(*s Type of an abstraction. *)
-val abs_rel :
- env -> 'a evar_map -> name -> types -> unsafe_judgment
- -> unsafe_judgment * constraints
-
-(* s Type of a let in. *)
-val judge_of_letin :
- env -> 'a evar_map -> name -> unsafe_judgment -> unsafe_judgment
- -> unsafe_judgment * constraints
+(*s type of a constant *)
+val judge_of_constant : env -> constant -> unsafe_judgment
(*s Type of application. *)
-val apply_rel_list :
- env -> 'a evar_map -> bool -> unsafe_judgment list -> unsafe_judgment
+val judge_of_apply :
+ env -> unsafe_judgment -> unsafe_judgment array
-> unsafe_judgment * constraints
+(*s Type of an abstraction. *)
+val judge_of_abstraction :
+ env -> name -> unsafe_type_judgment -> unsafe_judgment
+ -> unsafe_judgment
+
(*s Type of a product. *)
-val gen_rel :
- env -> 'a evar_map -> name -> unsafe_type_judgment -> unsafe_type_judgment
+val judge_of_product :
+ env -> name -> unsafe_type_judgment -> unsafe_type_judgment
-> unsafe_judgment * constraints
-val sort_of_product : sorts -> sorts -> universes -> sorts * constraints
+(* s Type of a let in. *)
+val judge_of_letin :
+ env -> name -> unsafe_judgment -> unsafe_judgment
+ -> unsafe_judgment
(*s Type of a cast. *)
-val cast_rel :
- env -> 'a evar_map -> unsafe_judgment -> types
+val judge_of_cast :
+ env -> unsafe_judgment -> unsafe_type_judgment
-> unsafe_judgment * constraints
(*s Inductive types. *)
-open Inductive
-val type_of_inductive : env -> 'a evar_map -> inductive -> types
+val judge_of_inductive : env -> inductive -> unsafe_judgment
-val type_of_constructor : env -> 'a evar_map -> constructor -> types
+val judge_of_constructor : env -> constructor -> unsafe_judgment
(*s Type of Cases. *)
-val judge_of_case : env -> 'a evar_map -> case_info
- -> unsafe_judgment -> unsafe_judgment
- -> unsafe_judgment array -> unsafe_judgment * constraints
-
-val find_case_dep_nparams :
- env -> 'a evar_map -> constr * unsafe_judgment -> inductive_family
- -> bool * constraints
-
-val type_case_branches :
- env -> 'a evar_map -> Inductive.inductive_type -> unsafe_judgment
- -> constr -> types array * types * constraints
-
-(*s Type of fixpoints and guard condition. *)
-val check_fix : env -> 'a evar_map -> fixpoint -> unit
-val check_cofix : env -> 'a evar_map -> cofixpoint -> unit
-val type_fixpoint : env -> 'a evar_map -> name array -> types array
- -> unsafe_judgment array -> constraints
-
-val control_only_guard : env -> 'a evar_map -> constr -> unit
-
-(*i
-val hyps_inclusion : env -> 'a evar_map -> named_context -> named_context -> bool
-i*)
+val judge_of_case : env -> case_info
+ -> unsafe_judgment -> unsafe_judgment -> unsafe_judgment array
+ -> unsafe_judgment * constraints
+(* Typecheck general fixpoint (not checking guard conditions) *)
+val type_fixpoint : env -> name array -> types array
+ -> unsafe_judgment array -> constraints
diff --git a/kernel/univ.ml b/kernel/univ.ml
index a74ea74fb..b55b3ca6f 100644
--- a/kernel/univ.ml
+++ b/kernel/univ.ml
@@ -67,7 +67,7 @@ let implicit_univ =
{ u_mod = Names.make_dirpath [Names.id_of_string "implicit_univ"];
u_num = 0 }
-let current_module = ref Names.default_module
+let current_module = ref (Names.make_dirpath[Names.id_of_string"Top"])
let set_module m = current_module := m
diff --git a/kernel/univ.mli b/kernel/univ.mli
index da66f4aed..97dd6bdef 100644
--- a/kernel/univ.mli
+++ b/kernel/univ.mli
@@ -8,10 +8,6 @@
(*i $Id$ i*)
-(*i*)
-open Names
-(*i*)
-
(* Universes. *)
type universe
@@ -20,7 +16,7 @@ val implicit_univ : universe
val prop_univ : universe
-val set_module : dir_path -> unit
+val set_module : Names.dir_path -> unit
val new_univ : unit -> universe
@@ -32,9 +28,7 @@ val initial_universes : universes
(*s Constraints. *)
-type univ_constraint
-
-module Constraint : Set.S with type elt = univ_constraint
+module Constraint : Set.S
type constraints = Constraint.t