aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr>2016-06-16 15:26:07 +0200
committerGravatar Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr>2016-06-16 15:26:50 +0200
commit568aa9dff652d420e66cda7914d4bc265bb276e7 (patch)
treec493eaaa87636e304f5788136a5fd1c255816821
parentbce318b6d991587773ef2fb18c83de8d24bc4a5f (diff)
parent2d4701b4d1bdb0fb4f64dec9ffbd9ad90506ba26 (diff)
Merge PR #79: Let the kernel assume that a (co-)inductive type is positive.
-rw-r--r--checker/cic.mli15
-rw-r--r--checker/values.ml11
-rw-r--r--kernel/declarations.mli16
-rw-r--r--kernel/declareops.ml7
-rw-r--r--kernel/entries.mli5
-rw-r--r--kernel/fast_typeops.ml69
-rw-r--r--kernel/fast_typeops.mli7
-rw-r--r--kernel/indtypes.ml48
-rw-r--r--kernel/inductive.ml54
-rw-r--r--kernel/inductive.mli7
-rw-r--r--kernel/safe_typing.ml20
-rw-r--r--kernel/safe_typing.mli4
-rw-r--r--kernel/term_typing.ml45
-rw-r--r--kernel/term_typing.mli12
-rw-r--r--kernel/typeops.ml4
-rw-r--r--library/declare.ml25
-rw-r--r--library/declare.mli2
-rw-r--r--library/global.ml4
-rw-r--r--library/global.mli4
-rw-r--r--plugins/funind/glob_term_to_relation.ml2
-rw-r--r--plugins/funind/indfun.ml2
-rw-r--r--plugins/funind/merge.ml2
-rw-r--r--pretyping/inductiveops.ml4
-rw-r--r--pretyping/pretyping.ml25
-rw-r--r--pretyping/pretyping.mli2
-rw-r--r--pretyping/typing.ml4
-rw-r--r--printing/printer.ml39
-rw-r--r--printing/printer.mli14
-rw-r--r--stm/lemmas.ml2
-rw-r--r--toplevel/assumptions.ml18
-rw-r--r--toplevel/command.ml64
-rw-r--r--toplevel/command.mli11
-rw-r--r--toplevel/discharge.ml3
-rw-r--r--toplevel/indschemes.ml2
-rw-r--r--toplevel/obligations.ml3
-rw-r--r--toplevel/record.ml18
-rw-r--r--toplevel/record.mli5
-rw-r--r--toplevel/vernacentries.ml34
38 files changed, 387 insertions, 226 deletions
diff --git a/checker/cic.mli b/checker/cic.mli
index 00ac2f56c..469cf8d4c 100644
--- a/checker/cic.mli
+++ b/checker/cic.mli
@@ -212,6 +212,15 @@ type constant_def =
type constant_universes = Univ.universe_context
+(** The [typing_flags] are instructions to the type-checker which
+ modify its behaviour. The typing flags used in the type-checking
+ of a constant are tracked in their {!constant_body} so that they
+ can be displayed to the user. *)
+type typing_flags = {
+ check_guarded : bool; (** If [false] then fixed points and co-fixed
+ points are assumed to be total. *)
+}
+
type constant_body = {
const_hyps : section_context; (** New: younger hyp at top *)
const_body : constant_def;
@@ -220,7 +229,9 @@ type constant_body = {
const_polymorphic : bool; (** Is it polymorphic or not *)
const_universes : constant_universes;
const_proj : projection_body option;
- const_inline_code : bool }
+ const_inline_code : bool;
+ const_typing_flags : typing_flags;
+}
(** {6 Representation of mutual inductive types } *)
@@ -316,6 +327,8 @@ type mutual_inductive_body = {
mind_private : bool option; (** allow pattern-matching: Some true ok, Some false blocked *)
+ mind_checked_positive : bool; (** [false] when the mutual-inductive was assumed to be well-founded, bypassing the positivity checker. *)
+
(** {8 Data for native compilation } *)
mind_native_name : native_name ref; (** status of the code (linked or not, and where) *)
diff --git a/checker/values.ml b/checker/values.ml
index 19cbb5060..dd29f6fbe 100644
--- a/checker/values.ml
+++ b/checker/values.ml
@@ -13,7 +13,7 @@
To ensure this file is up-to-date, 'make' now compares the md5 of cic.mli
with a copy we maintain here:
-MD5 9f7fd499f812b6548a55f7067e6a9d06 checker/cic.mli
+MD5 7d7963269852d32324e10aa77beb938d checker/cic.mli
*)
@@ -213,6 +213,9 @@ let v_projbody =
v_tuple "proj_eta" [|v_constr;v_constr|];
v_constr|]
+let v_typing_flags =
+ v_tuple "typing_flags" [|v_bool|]
+
let v_cb = v_tuple "constant_body"
[|v_section_ctxt;
v_cst_def;
@@ -221,7 +224,8 @@ let v_cb = v_tuple "constant_body"
v_bool;
v_context;
Opt v_projbody;
- v_bool|]
+ v_bool;
+ v_typing_flags|]
let v_recarg = v_sum "recarg" 1 (* Norec *)
[|[|v_ind|] (* Mrec *);[|v_ind|] (* Imbr *)|]
@@ -270,7 +274,8 @@ let v_ind_pack = v_tuple "mutual_inductive_body"
v_rctxt;
v_bool;
v_context;
- Opt v_bool|]
+ Opt v_bool;
+ v_bool|]
let v_with =
Sum ("with_declaration_body",0,
diff --git a/kernel/declarations.mli b/kernel/declarations.mli
index 1b77d5b7c..8b42a90e4 100644
--- a/kernel/declarations.mli
+++ b/kernel/declarations.mli
@@ -66,6 +66,15 @@ type constant_def =
type constant_universes = Univ.universe_context
+(** The [typing_flags] are instructions to the type-checker which
+ modify its behaviour. The typing flags used in the type-checking
+ of a constant are tracked in their {!constant_body} so that they
+ can be displayed to the user. *)
+type typing_flags = {
+ check_guarded : bool; (** If [false] then fixed points and co-fixed
+ points are assumed to be total. *)
+}
+
(* some contraints are in constant_constraints, some other may be in
* the OpaueDef *)
type constant_body = {
@@ -76,7 +85,11 @@ type constant_body = {
const_polymorphic : bool; (** Is it polymorphic or not *)
const_universes : constant_universes;
const_proj : projection_body option;
- const_inline_code : bool }
+ const_inline_code : bool;
+ const_typing_flags : typing_flags; (** The typing options which
+ were used for
+ type-checking. *)
+}
(** {6 Representation of mutual inductive types in the kernel } *)
@@ -178,6 +191,7 @@ type mutual_inductive_body = {
mind_private : bool option; (** allow pattern-matching: Some true ok, Some false blocked *)
+ mind_checked_positive : bool; (** [false] when the mutual-inductive was assumed to be well-founded, bypassing the positivity checker. *)
}
(** {6 Module declarations } *)
diff --git a/kernel/declareops.ml b/kernel/declareops.ml
index a09a8b786..78e2f386e 100644
--- a/kernel/declareops.ml
+++ b/kernel/declareops.ml
@@ -131,7 +131,8 @@ let subst_const_body sub cb =
Option.map (Cemitcodes.subst_to_patch_subst sub) cb.const_body_code;
const_polymorphic = cb.const_polymorphic;
const_universes = cb.const_universes;
- const_inline_code = cb.const_inline_code }
+ const_inline_code = cb.const_inline_code;
+ const_typing_flags = cb.const_typing_flags }
(** {7 Hash-consing of constants } *)
@@ -254,7 +255,9 @@ let subst_mind_body sub mib =
mind_packets = Array.smartmap (subst_mind_packet sub) mib.mind_packets ;
mind_polymorphic = mib.mind_polymorphic;
mind_universes = mib.mind_universes;
- mind_private = mib.mind_private }
+ mind_private = mib.mind_private;
+ mind_checked_positive = mib.mind_checked_positive;
+ }
let inductive_instance mib =
if mib.mind_polymorphic then
diff --git a/kernel/entries.mli b/kernel/entries.mli
index d07ca2103..8b29e3abd 100644
--- a/kernel/entries.mli
+++ b/kernel/entries.mli
@@ -51,7 +51,10 @@ type mutual_inductive_entry = {
mind_entry_inds : one_inductive_entry list;
mind_entry_polymorphic : bool;
mind_entry_universes : Univ.universe_context;
- mind_entry_private : bool option }
+ mind_entry_private : bool option;
+ mind_entry_check_positivity : bool;
+ (** [false] if positivity is to be assumed. *)
+}
(** {6 Constants (Definition/Axiom) } *)
type 'a proof_output = constr Univ.in_universe_context_set * 'a
diff --git a/kernel/fast_typeops.ml b/kernel/fast_typeops.ml
index 7f4ba8ecb..35c162cf3 100644
--- a/kernel/fast_typeops.ml
+++ b/kernel/fast_typeops.ml
@@ -327,7 +327,7 @@ let type_fixpoint env lna lar vdef vdeft =
(* 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 =
+let rec execute ~flags env cstr =
let open Context.Rel.Declaration in
match kind_of_term cstr with
(* Atomic terms *)
@@ -347,12 +347,12 @@ let rec execute env cstr =
judge_of_constant env c
| Proj (p, c) ->
- let ct = execute env c in
+ let ct = execute ~flags env c in
judge_of_projection env p c ct
(* Lambda calculus operators *)
| App (f,args) ->
- let argst = execute_array env args in
+ let argst = execute_array ~flags env args in
let ft =
match kind_of_term f with
| Ind ind when Environ.template_polymorphic_pind ind env ->
@@ -365,7 +365,7 @@ let rec execute env cstr =
judge_of_constant_knowing_parameters env cst args
| _ ->
(* Full or no sort-polymorphism *)
- execute env f
+ execute ~flags env f
in
judge_of_apply env f ft args argst
@@ -373,25 +373,25 @@ let rec execute env cstr =
| Lambda (name,c1,c2) ->
let _ = execute_is_type env c1 in
let env1 = push_rel (LocalAssum (name,c1)) env in
- let c2t = execute env1 c2 in
+ let c2t = execute ~flags env1 c2 in
judge_of_abstraction env name c1 c2t
| Prod (name,c1,c2) ->
- let vars = execute_is_type env c1 in
+ let vars = execute_is_type ~flags env c1 in
let env1 = push_rel (LocalAssum (name,c1)) env in
- let vars' = execute_is_type env1 c2 in
+ let vars' = execute_is_type ~flags env1 c2 in
judge_of_product env name vars vars'
| LetIn (name,c1,c2,c3) ->
- let c1t = execute env c1 in
+ let c1t = execute ~flags env c1 in
let _c2s = execute_is_type env c2 in
let _ = judge_of_cast env c1 c1t DEFAULTcast c2 in
let env1 = push_rel (LocalDef (name,c1,c2)) env in
- let c3t = execute env1 c3 in
+ let c3t = execute ~flags env1 c3 in
subst1 c1 c3t
| Cast (c,k,t) ->
- let ct = execute env c in
+ let ct = execute ~flags env c in
let _ts = execute_type env t in
let _ = judge_of_cast env c ct k t in
t
@@ -404,20 +404,20 @@ let rec execute env cstr =
judge_of_constructor env c
| Case (ci,p,c,lf) ->
- let ct = execute env c in
- let pt = execute env p in
- let lft = execute_array env lf in
+ let ct = execute ~flags env c in
+ let pt = execute ~flags env p in
+ let lft = execute_array ~flags env lf in
judge_of_case env ci p pt c ct lf lft
| Fix ((vn,i as vni),recdef) ->
- let (fix_ty,recdef') = execute_recdef env recdef i in
+ let (fix_ty,recdef') = execute_recdef ~flags env recdef i in
let fix = (vni,recdef') in
- check_fix env fix; fix_ty
+ check_fix env ~flags fix; fix_ty
| CoFix (i,recdef) ->
- let (fix_ty,recdef') = execute_recdef env recdef i in
+ let (fix_ty,recdef') = execute_recdef ~flags env recdef i in
let cofix = (i,recdef') in
- check_cofix env cofix; fix_ty
+ check_cofix env ~flags cofix; fix_ty
(* Partial proofs: unsupported by the kernel *)
| Meta _ ->
@@ -426,38 +426,41 @@ let rec execute env cstr =
| Evar _ ->
anomaly (Pp.str "the kernel does not support existential variables")
-and execute_is_type env constr =
- let t = execute env constr in
+and execute_is_type ~flags env constr =
+ let t = execute ~flags env constr in
check_type env constr t
-and execute_type env constr =
- let t = execute env constr in
+and execute_type ~flags env constr =
+ let t = execute ~flags env constr in
type_judgment env constr t
-and execute_recdef env (names,lar,vdef) i =
- let lart = execute_array env lar in
+and execute_recdef ~flags env (names,lar,vdef) i =
+ let lart = execute_array ~flags env lar in
let lara = Array.map2 (assumption_of_judgment env) lar lart in
let env1 = push_rec_types (names,lara,vdef) env in
- let vdeft = execute_array env1 vdef in
+ let vdeft = execute_array ~flags env1 vdef in
let () = type_fixpoint env1 names lara vdef vdeft in
(lara.(i),(names,lara,vdef))
-and execute_array env = Array.map (execute env)
+and execute_array ~flags env = Array.map (execute ~flags env)
(* Derived functions *)
-let infer env constr =
- let t = execute env constr in
+let infer ~flags env constr =
+ let t = execute ~flags env constr in
make_judge constr t
let infer =
if Flags.profile then
let infer_key = Profile.declare_profile "Fast_infer" in
- Profile.profile2 infer_key infer
- else infer
+ Profile.profile3 infer_key (fun a b c -> infer ~flags:a b c)
+ else (fun a b c -> infer ~flags:a b c)
-let infer_type env constr =
- execute_type env constr
+(* Restores the labels of [infer] lost to profiling. *)
+let infer ~flags env t = infer flags env t
-let infer_v env cv =
- let jv = execute_array env cv in
+let infer_type ~flags env constr =
+ execute_type ~flags env constr
+
+let infer_v ~flags env cv =
+ let jv = execute_array ~flags env cv in
make_judgev cv jv
diff --git a/kernel/fast_typeops.mli b/kernel/fast_typeops.mli
index 05d52b2d3..45a603038 100644
--- a/kernel/fast_typeops.mli
+++ b/kernel/fast_typeops.mli
@@ -8,6 +8,7 @@
open Term
open Environ
+open Declarations
(** {6 Typing functions (not yet tagged as safe) }
@@ -18,6 +19,6 @@ open Environ
*)
-val infer : env -> constr -> unsafe_judgment
-val infer_v : env -> constr array -> unsafe_judgment array
-val infer_type : env -> types -> unsafe_type_judgment
+val infer : flags:typing_flags -> env -> constr -> unsafe_judgment
+val infer_v : flags:typing_flags -> env -> constr array -> unsafe_judgment array
+val infer_type : flags:typing_flags -> env -> types -> unsafe_type_judgment
diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml
index edb758f07..b74788d21 100644
--- a/kernel/indtypes.ml
+++ b/kernel/indtypes.ml
@@ -483,8 +483,12 @@ let array_min nmr a = if Int.equal nmr 0 then 0 else
for use by the guard condition (terms at these positions are
considered sub-terms) as well as the number of of non-uniform
arguments (used to generate induction schemes, so a priori less
- relevant to the kernel). *)
-let check_positivity_one recursive (env,_,ntypes,_ as ienv) paramsctxt (_,i as ind) nnonrecargs lcnames indlc =
+ relevant to the kernel).
+
+ If [chkpos] is [false] then positivity is assumed, and
+ [check_positivity_one] computes the subterms occurrences in a
+ best-effort fashion. *)
+let check_positivity_one ~chkpos recursive (env,_,ntypes,_ as ienv) paramsctxt (_,i as ind) nnonrecargs lcnames indlc =
let nparamsctxt = Context.Rel.length paramsctxt in
let nmr = Context.Rel.nhyps paramsctxt in
(** Positivity of one argument [c] of a constructor (i.e. the
@@ -501,9 +505,12 @@ let check_positivity_one recursive (env,_,ntypes,_ as ienv) paramsctxt (_,i as i
recursive call. Occurrences in the right-hand side of
the product must be strictly positive.*)
(match weaker_noccur_between env n ntypes b with
- None -> failwith_non_pos_list n ntypes [b]
+ | None when chkpos ->
+ failwith_non_pos_list n ntypes [b]
+ | None ->
+ check_pos (ienv_push_var ienv (na, b, mk_norec)) nmr d
| Some b ->
- check_pos (ienv_push_var ienv (na, b, mk_norec)) nmr d)
+ check_pos (ienv_push_var ienv (na, b, mk_norec)) nmr d)
| Rel k ->
(try let (ra,rarg) = List.nth ra_env (k-1) in
let largs = List.map (whd_betadeltaiota env) largs in
@@ -515,7 +522,8 @@ let check_positivity_one recursive (env,_,ntypes,_ as ienv) paramsctxt (_,i as i
(** The case where one of the inductives of the mutually
inductive block occurs as an argument of another is not
known to be safe. So Coq rejects it. *)
- if not (List.for_all (noccur_between n ntypes) largs)
+ if chkpos &&
+ not (List.for_all (noccur_between n ntypes) largs)
then failwith_non_pos_list n ntypes largs
else (nmr1,rarg)
with Failure _ | Invalid_argument _ -> (nmr,mk_norec))
@@ -530,8 +538,9 @@ let check_positivity_one recursive (env,_,ntypes,_ as ienv) paramsctxt (_,i as i
(** If an inductive of the mutually inductive block
appears in any other way, then the positivy check gives
up. *)
- if noccur_between n ntypes x &&
- List.for_all (noccur_between n ntypes) largs
+ if not chkpos ||
+ (noccur_between n ntypes x &&
+ List.for_all (noccur_between n ntypes) largs)
then (nmr,mk_norec)
else failwith_non_pos_list n ntypes (x::largs)
@@ -553,7 +562,7 @@ let check_positivity_one recursive (env,_,ntypes,_ as ienv) paramsctxt (_,i as i
(** Inductives of the inductive block being defined are only
allowed to appear nested in the parameters of another inductive
type. Not in the proper indices. *)
- if not (List.for_all (noccur_between n ntypes) auxnonrecargs) then
+ if chkpos && not (List.for_all (noccur_between n ntypes) auxnonrecargs) then
failwith_non_pos_list n ntypes auxnonrecargs;
(* Nested mutual inductive types are not supported *)
let auxntyp = mib.mind_ntypes in
@@ -613,7 +622,8 @@ let check_positivity_one recursive (env,_,ntypes,_ as ienv) paramsctxt (_,i as i
| _ -> raise (IllFormedInd (LocalNotConstructor(paramsctxt,nnonrecargs)))
end
else
- if not (List.for_all (noccur_between n ntypes) largs)
+ if chkpos &&
+ not (List.for_all (noccur_between n ntypes) largs)
then failwith_non_pos_list n ntypes largs
in
(nmr, List.rev lrec)
@@ -633,9 +643,13 @@ let check_positivity_one recursive (env,_,ntypes,_ as ienv) paramsctxt (_,i as i
and nmr' = array_min nmr irecargs_nmr
in (nmr', mk_paths (Mrec ind) irecargs)
-(** [check_positivity kn env_ar paramsctxt inds] checks that the mutually
- inductive block [inds] is strictly positive. *)
-let check_positivity kn env_ar_par paramsctxt finite inds =
+(** [check_positivity ~chkpos kn env_ar paramsctxt inds] checks that the mutually
+ inductive block [inds] is strictly positive.
+
+ If [chkpos] is [false] then positivity is assumed, and
+ [check_positivity_one] computes the subterms occurrences in a
+ best-effort fashion. *)
+let check_positivity ~chkpos kn env_ar_par paramsctxt finite inds =
let ntypes = Array.length inds in
let recursive = finite != Decl_kinds.BiFinite in
let rc = Array.mapi (fun j t -> (Mrec (kn,j),t)) (Rtree.mk_rec_calls ntypes) in
@@ -647,7 +661,7 @@ let check_positivity kn env_ar_par paramsctxt finite inds =
List.init nparamsctxt (fun _ -> (Norec,mk_norec)) @ ra_env_ar in
let ienv = (env_ar_par, 1+nparamsctxt, ntypes, ra_env_ar_par) in
let nnonrecargs = Context.Rel.nhyps sign - nmr in
- check_positivity_one recursive ienv paramsctxt (kn,i) nnonrecargs lcnames lc
+ check_positivity_one ~chkpos recursive ienv paramsctxt (kn,i) nnonrecargs lcnames lc
in
let irecargs_nmr = Array.mapi check_one inds in
let irecargs = Array.map snd irecargs_nmr
@@ -802,7 +816,7 @@ let compute_projections ((kn, _ as ind), u as indu) n x nparamargs params
Array.of_list (List.rev kns),
Array.of_list (List.rev pbs)
-let build_inductive env p prv ctx env_ar paramsctxt kn isrecord isfinite inds nmr recargs =
+let build_inductive env p prv ctx env_ar paramsctxt kn isrecord isfinite inds nmr recargs is_checked =
let ntypes = Array.length inds in
(* Compute the set of used section variables *)
let hyps = used_section_variables env inds in
@@ -908,6 +922,7 @@ let build_inductive env p prv ctx env_ar paramsctxt kn isrecord isfinite inds nm
mind_polymorphic = p;
mind_universes = ctx;
mind_private = prv;
+ mind_checked_positive = is_checked;
}
(************************************************************************)
@@ -916,10 +931,11 @@ let build_inductive env p prv ctx env_ar paramsctxt kn isrecord isfinite inds nm
let check_inductive env kn mie =
(* First type-check the inductive definition *)
let (env_ar, env_ar_par, paramsctxt, inds) = typecheck_inductive env mie in
+ let chkpos = mie.mind_entry_check_positivity in
(* Then check positivity conditions *)
- let (nmr,recargs) = check_positivity kn env_ar_par paramsctxt mie.mind_entry_finite inds in
+ let (nmr,recargs) = check_positivity ~chkpos kn env_ar_par paramsctxt mie.mind_entry_finite inds in
(* Build the inductive packets *)
build_inductive env mie.mind_entry_polymorphic mie.mind_entry_private
mie.mind_entry_universes
env_ar paramsctxt kn mie.mind_entry_record mie.mind_entry_finite
- inds nmr recargs
+ inds nmr recargs chkpos
diff --git a/kernel/inductive.ml b/kernel/inductive.ml
index 499cbf0df..24bdaa5c4 100644
--- a/kernel/inductive.ml
+++ b/kernel/inductive.ml
@@ -1067,21 +1067,24 @@ let inductive_of_mutfix env ((nvect,bodynum),(names,types,bodies as recdef)) =
(Array.map fst rv, Array.map snd rv)
-let check_fix env ((nvect,_),(names,_,bodies as recdef) as fix) =
- let (minds, rdef) = inductive_of_mutfix env fix in
- let get_tree (kn,i) =
- let mib = Environ.lookup_mind kn env in
- mib.mind_packets.(i).mind_recargs
- in
- let trees = Array.map (fun (mind,_) -> get_tree mind) minds in
- for i = 0 to Array.length bodies - 1 do
- let (fenv,body) = rdef.(i) in
- let renv = make_renv fenv nvect.(i) trees.(i) in
- try check_one_fix renv nvect trees body
- with FixGuardError (fixenv,err) ->
- error_ill_formed_rec_body fixenv err names i
- (push_rec_types recdef env) (judgment_of_fixpoint recdef)
- done
+let check_fix env ~flags ((nvect,_),(names,_,bodies as recdef) as fix) =
+ if flags.check_guarded then
+ let (minds, rdef) = inductive_of_mutfix env fix in
+ let get_tree (kn,i) =
+ let mib = Environ.lookup_mind kn env in
+ mib.mind_packets.(i).mind_recargs
+ in
+ let trees = Array.map (fun (mind,_) -> get_tree mind) minds in
+ for i = 0 to Array.length bodies - 1 do
+ let (fenv,body) = rdef.(i) in
+ let renv = make_renv fenv nvect.(i) trees.(i) in
+ try check_one_fix renv nvect trees body
+ with FixGuardError (fixenv,err) ->
+ error_ill_formed_rec_body fixenv err names i
+ (push_rec_types recdef env) (judgment_of_fixpoint recdef)
+ done
+ else
+ ()
(*
let cfkey = Profile.declare_profile "check_fix";;
@@ -1192,12 +1195,15 @@ let check_one_cofix env nbfix def deftype =
(* 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 check_one_cofix fixenv nbfix bodies.(i) types.(i)
- with CoFixGuardError (errenv,err) ->
- error_ill_formed_rec_body errenv err names i
- fixenv (judgment_of_fixpoint recdef)
- done
+let check_cofix env ~flags (bodynum,(names,types,bodies as recdef)) =
+ if flags.check_guarded then
+ let nbfix = Array.length bodies in
+ for i = 0 to nbfix-1 do
+ let fixenv = push_rec_types recdef env in
+ try check_one_cofix fixenv nbfix bodies.(i) types.(i)
+ with CoFixGuardError (errenv,err) ->
+ error_ill_formed_rec_body errenv err names i
+ fixenv (judgment_of_fixpoint recdef)
+ done
+ else
+ ()
diff --git a/kernel/inductive.mli b/kernel/inductive.mli
index c0d18bc6e..25a557472 100644
--- a/kernel/inductive.mli
+++ b/kernel/inductive.mli
@@ -94,8 +94,11 @@ val inductive_sort_family : one_inductive_body -> sorts_family
val check_case_info : env -> pinductive -> case_info -> unit
(** {6 Guard conditions for fix and cofix-points. } *)
-val check_fix : env -> fixpoint -> unit
-val check_cofix : env -> cofixpoint -> unit
+
+(** When [chk] is false, the guard condition is not actually
+ checked. *)
+val check_fix : env -> flags:typing_flags -> fixpoint -> unit
+val check_cofix : env -> flags:typing_flags -> cofixpoint -> unit
(** {6 Support for sort-polymorphic inductive types } *)
diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml
index ce05190b6..72d6ee518 100644
--- a/kernel/safe_typing.ml
+++ b/kernel/safe_typing.ml
@@ -373,8 +373,8 @@ let safe_push_named d env =
Environ.push_named d env
-let push_named_def (id,de) senv =
- let c,typ,univs = Term_typing.translate_local_def senv.revstruct senv.env id de in
+let push_named_def ~flags (id,de) senv =
+ let c,typ,univs = Term_typing.translate_local_def ~flags senv.revstruct senv.env id de in
let poly = de.Entries.const_entry_polymorphic in
let univs = Univ.ContextSet.of_context univs in
let c, univs = match c with
@@ -388,9 +388,9 @@ let push_named_def (id,de) senv =
let env'' = safe_push_named (LocalDef (id,c,typ)) senv'.env in
univs, {senv' with env=env''}
-let push_named_assum ((id,t,poly),ctx) senv =
+let push_named_assum ~flags ((id,t,poly),ctx) senv =
let senv' = push_context_set poly ctx senv in
- let t = Term_typing.translate_local_assum senv'.env t in
+ let t = Term_typing.translate_local_assum ~flags senv'.env t in
let env'' = safe_push_named (LocalAssum (id,t)) senv'.env in
{senv' with env=env''}
@@ -479,7 +479,7 @@ let update_resolver f senv = { senv with modresolver = f senv.modresolver }
(** Insertion of constants and parameters in environment *)
type global_declaration =
- | ConstantEntry of bool * private_constants Entries.constant_entry
+ | ConstantEntry of bool * private_constants Entries.constant_entry * Declarations.typing_flags
| GlobalRecipe of Cooking.recipe
type exported_private_constant =
@@ -512,10 +512,10 @@ let add_constant dir l decl senv =
let no_section = DirPath.is_empty dir in
let seff_to_export, decl =
match decl with
- | ConstantEntry (true, ce) ->
+ | ConstantEntry (true, ce, flags) ->
let exports, ce =
- Term_typing.export_side_effects senv.revstruct senv.env ce in
- exports, ConstantEntry (false, ce)
+ Term_typing.export_side_effects ~flags senv.revstruct senv.env ce in
+ exports, ConstantEntry (false, ce, flags)
| _ -> [], decl
in
let senv =
@@ -524,8 +524,8 @@ let add_constant dir l decl senv =
let senv =
let cb =
match decl with
- | ConstantEntry (export_seff,ce) ->
- Term_typing.translate_constant senv.revstruct senv.env kn ce
+ | ConstantEntry (export_seff,ce,flags) ->
+ Term_typing.translate_constant ~flags senv.revstruct senv.env kn ce
| GlobalRecipe r ->
let cb = Term_typing.translate_recipe senv.env kn r in
if no_section then Declareops.hcons_const_body cb else cb in
diff --git a/kernel/safe_typing.mli b/kernel/safe_typing.mli
index 71dac321f..b614a368c 100644
--- a/kernel/safe_typing.mli
+++ b/kernel/safe_typing.mli
@@ -77,19 +77,21 @@ val is_joined_environment : safe_environment -> bool
(** Insertion of local declarations (Local or Variables) *)
val push_named_assum :
+ flags:Declarations.typing_flags ->
(Id.t * Term.types * bool (* polymorphic *))
Univ.in_universe_context_set -> safe_transformer0
(** Returns the full universe context necessary to typecheck the definition
(futures are forced) *)
val push_named_def :
+ flags:Declarations.typing_flags ->
Id.t * private_constants Entries.definition_entry -> Univ.universe_context_set safe_transformer
(** Insertion of global axioms or definitions *)
type global_declaration =
(* bool: export private constants *)
- | ConstantEntry of bool * private_constants Entries.constant_entry
+ | ConstantEntry of bool * private_constants Entries.constant_entry * Declarations.typing_flags
| GlobalRecipe of Cooking.recipe
type exported_private_constant =
diff --git a/kernel/term_typing.ml b/kernel/term_typing.ml
index 6bfd2457a..8d74dc390 100644
--- a/kernel/term_typing.ml
+++ b/kernel/term_typing.ml
@@ -22,18 +22,18 @@ open Entries
open Typeops
open Fast_typeops
-let constrain_type env j poly subst = function
+let constrain_type ~flags env j poly subst = function
| `None ->
if not poly then (* Old-style polymorphism *)
make_polymorphic_if_constant_for_ind env j
else RegularArity (Vars.subst_univs_level_constr subst j.uj_type)
| `Some t ->
- let tj = infer_type env t in
+ let tj = infer_type ~flags env t in
let _ = judge_of_cast env j DEFAULTcast tj in
assert (eq_constr t tj.utj_val);
RegularArity (Vars.subst_univs_level_constr subst t)
| `SomeWJ (t, tj) ->
- let tj = infer_type env t in
+ let tj = infer_type ~flags env t in
let _ = judge_of_cast env j DEFAULTcast tj in
assert (eq_constr t tj.utj_val);
RegularArity (Vars.subst_univs_level_constr subst t)
@@ -171,11 +171,11 @@ let feedback_completion_typecheck =
Option.iter (fun state_id ->
feedback ~id:(State state_id) Feedback.Complete)
-let infer_declaration ~trust env kn dcl =
+let infer_declaration ~flags ~trust env kn dcl =
match dcl with
| ParameterEntry (ctx,poly,(t,uctx),nl) ->
let env = push_context ~strict:(not poly) uctx env in
- let j = infer env t in
+ let j = infer ~flags env t in
let abstract = poly && not (Option.is_empty kn) in
let usubst, univs = Univ.abstract_universes abstract uctx in
let c = Typeops.assumption_of_judgment env j in
@@ -196,7 +196,7 @@ let infer_declaration ~trust env kn dcl =
let env' = push_context_set uctx env in
let j =
let body,env',ectx = skip_trusted_seff valid_signatures body env' in
- let j = infer env' body in
+ let j = infer ~flags env' body in
unzip ectx j in
let j = hcons_j j in
let subst = Univ.LMap.empty in
@@ -220,8 +220,8 @@ let infer_declaration ~trust env kn dcl =
let abstract = c.const_entry_polymorphic && not (Option.is_empty kn) in
let usubst, univs =
Univ.abstract_universes abstract (Univ.ContextSet.to_context ctx) in
- let j = infer env body in
- let typ = constrain_type env j c.const_entry_polymorphic usubst (map_option_typ typ) in
+ let j = infer ~flags env body in
+ let typ = constrain_type ~flags env j c.const_entry_polymorphic usubst (map_option_typ typ) in
let def = hcons_constr (Vars.subst_univs_level_constr usubst j.uj_val) in
let def =
if opaque then OpaqueDef (Opaqueproof.create (Future.from_val (def, Univ.ContextSet.empty)))
@@ -268,7 +268,7 @@ let record_aux env s_ty s_bo suggested_expr =
let suggest_proof_using = ref (fun _ _ _ _ _ -> "")
let set_suggest_proof_using f = suggest_proof_using := f
-let build_constant_declaration kn env (def,typ,proj,poly,univs,inline_code,ctx) =
+let build_constant_declaration ~flags kn env (def,typ,proj,poly,univs,inline_code,ctx) =
let open Context.Named.Declaration in
let check declared inferred =
let mk_set l = List.fold_right Id.Set.add (List.map get_id l) Id.Set.empty in
@@ -352,7 +352,9 @@ let build_constant_declaration kn env (def,typ,proj,poly,univs,inline_code,ctx)
const_body_code = None;
const_polymorphic = poly;
const_universes = univs;
- const_inline_code = inline_code }
+ const_inline_code = inline_code;
+ const_typing_flags = flags;
+ }
in
let env = add_constant kn cb env in
compile_constant_body env comp_univs def
@@ -365,13 +367,14 @@ let build_constant_declaration kn env (def,typ,proj,poly,univs,inline_code,ctx)
const_body_code = tps;
const_polymorphic = poly;
const_universes = univs;
- const_inline_code = inline_code }
+ const_inline_code = inline_code;
+ const_typing_flags = flags }
(*s Global and local constant declaration. *)
-let translate_constant mb env kn ce =
- build_constant_declaration kn env
- (infer_declaration ~trust:mb env (Some kn) ce)
+let translate_constant ~flags mb env kn ce =
+ build_constant_declaration ~flags kn env
+ (infer_declaration ~flags ~trust:mb env (Some kn) ce)
let constant_entry_of_side_effect cb u =
let pt =
@@ -406,7 +409,7 @@ type side_effect_role =
type exported_side_effect =
constant * constant_body * side_effects constant_entry * side_effect_role
-let export_side_effects mb env ce =
+let export_side_effects ~flags mb env ce =
match ce with
| ParameterEntry _ | ProjectionEntry _ -> [], ce
| DefinitionEntry c ->
@@ -447,7 +450,7 @@ let export_side_effects mb env ce =
let env, cbs =
List.fold_left (fun (env,cbs) (kn, ocb, u, r) ->
let ce = constant_entry_of_side_effect ocb u in
- let cb = translate_constant mb env kn ce in
+ let cb = translate_constant ~flags mb env kn ce in
(push_seff env (kn, cb,`Nothing, Subproof),(kn,cb,ce,r) :: cbs))
(env,[]) cbs in
translate_seff sl rest (cbs @ acc) env
@@ -462,17 +465,17 @@ let export_side_effects mb env ce =
translate_seff trusted seff [] env
;;
-let translate_local_assum env t =
- let j = infer env t in
+let translate_local_assum ~flags env t =
+ let j = infer ~flags env t in
let t = Typeops.assumption_of_judgment env j in
t
let translate_recipe env kn r =
- build_constant_declaration kn env (Cooking.cook_constant env r)
+ build_constant_declaration ~flags:{check_guarded=true} kn env (Cooking.cook_constant env r)
-let translate_local_def mb env id centry =
+let translate_local_def ~flags mb env id centry =
let def,typ,proj,poly,univs,inline_code,ctx =
- infer_declaration ~trust:mb env None (DefinitionEntry centry) in
+ infer_declaration ~flags ~trust:mb env None (DefinitionEntry centry) in
let typ = type_of_constant_type env typ in
if ctx = None && !Flags.compilation_mode = Flags.BuildVo then begin
match def with
diff --git a/kernel/term_typing.mli b/kernel/term_typing.mli
index fcd95576c..b635f2d31 100644
--- a/kernel/term_typing.mli
+++ b/kernel/term_typing.mli
@@ -12,10 +12,10 @@ open Environ
open Declarations
open Entries
-val translate_local_def : structure_body -> env -> Id.t -> side_effects definition_entry ->
+val translate_local_def : flags:typing_flags -> structure_body -> env -> Id.t -> side_effects definition_entry ->
constant_def * types * constant_universes
-val translate_local_assum : env -> types -> types
+val translate_local_assum : flags:typing_flags -> env -> types -> types
val mk_pure_proof : constr -> side_effects proof_output
@@ -32,7 +32,7 @@ val inline_entry_side_effects :
val uniq_seff : side_effects -> side_effects
val translate_constant :
- structure_body -> env -> constant -> side_effects constant_entry ->
+ flags:typing_flags -> structure_body -> env -> constant -> side_effects constant_entry ->
constant_body
type side_effect_role =
@@ -47,7 +47,7 @@ type exported_side_effect =
* be pushed in the safe_env by safe typing. The main constant entry
* needs to be translated as usual after this step. *)
val export_side_effects :
- structure_body -> env -> side_effects constant_entry ->
+ flags:typing_flags -> structure_body -> env -> side_effects constant_entry ->
exported_side_effect list * side_effects constant_entry
val constant_entry_of_side_effect :
@@ -60,11 +60,11 @@ val translate_recipe : env -> constant -> Cooking.recipe -> constant_body
(** Internal functions, mentioned here for debug purpose only *)
-val infer_declaration : trust:structure_body -> env -> constant option ->
+val infer_declaration : flags:typing_flags -> trust:structure_body -> env -> constant option ->
side_effects constant_entry -> Cooking.result
val build_constant_declaration :
- constant -> env -> Cooking.result -> constant_body
+ flags:typing_flags -> constant -> env -> Cooking.result -> constant_body
val set_suggest_proof_using :
(string -> env -> Id.Set.t -> Id.Set.t -> Id.t list -> string) -> unit
diff --git a/kernel/typeops.ml b/kernel/typeops.ml
index 0ea68e2bc..a94a049df 100644
--- a/kernel/typeops.ml
+++ b/kernel/typeops.ml
@@ -500,13 +500,13 @@ let rec execute env cstr =
| Fix ((vn,i as vni),recdef) ->
let (fix_ty,recdef') = execute_recdef env recdef i in
let fix = (vni,recdef') in
- check_fix env fix;
+ check_fix ~flags:{check_guarded=true} env fix;
make_judge (mkFix fix) fix_ty
| CoFix (i,recdef) ->
let (fix_ty,recdef') = execute_recdef env recdef i in
let cofix = (i,recdef') in
- check_cofix env cofix;
+ check_cofix ~flags:{check_guarded=true} env cofix;
(make_judge (mkCoFix cofix) fix_ty)
(* Partial proofs: unsupported by the kernel *)
diff --git a/library/declare.ml b/library/declare.ml
index 84284fd18..9ec299bed 100644
--- a/library/declare.ml
+++ b/library/declare.ml
@@ -58,11 +58,11 @@ let cache_variable ((sp,_),o) =
let impl,opaq,poly,ctx = match d with (* Fails if not well-typed *)
| SectionLocalAssum ((ty,ctx),poly,impl) ->
- let () = Global.push_named_assum ((id,ty,poly),ctx) in
+ let () = Global.push_named_assum ~flags:{check_guarded=true} ((id,ty,poly),ctx) in
let impl = if impl then Implicit else Explicit in
impl, true, poly, ctx
| SectionLocalDef (de) ->
- let univs = Global.push_named_def (id,de) in
+ let univs = Global.push_named_def ~flags:{check_guarded=true} (id,de) in
Explicit, de.const_entry_opaque,
de.const_entry_polymorphic, univs in
Nametab.push (Nametab.Until 1) (restrict_path 0 sp) (VarRef id);
@@ -180,7 +180,7 @@ let discharge_constant ((sp, kn), obj) =
(* Hack to reduce the size of .vo: we keep only what load/open needs *)
let dummy_constant_entry =
ConstantEntry
- (false, ParameterEntry (None,false,(mkProp,Univ.UContext.empty),None))
+ (false, ParameterEntry (None,false,(mkProp,Univ.UContext.empty),None), {check_guarded=true})
let dummy_constant cst = {
cst_decl = dummy_constant_entry;
@@ -205,7 +205,7 @@ let (inConstant, outConstant : (constant_obj -> obj) * (obj -> constant_obj)) =
let declare_scheme = ref (fun _ _ -> assert false)
let set_declare_scheme f = declare_scheme := f
-let declare_constant_common id cst =
+let declare_constant_common ~flags id cst =
let update_tables c =
(* Printf.eprintf "tables: %s\n%!" (Names.Constant.to_string c); *)
declare_constant_implicits c;
@@ -216,7 +216,7 @@ let declare_constant_common id cst =
List.iter (fun (c,ce,role) ->
(* handling of private_constants just exported *)
let o = inConstant {
- cst_decl = ConstantEntry (false, ce);
+ cst_decl = ConstantEntry (false, ce, flags);
cst_hyps = [] ;
cst_kind = IsProof Theorem;
cst_locl = false;
@@ -246,7 +246,7 @@ let definition_entry ?fix_exn ?(opaque=false) ?(inline=false) ?types
const_entry_feedback = None;
const_entry_inline_code = inline}
-let declare_constant ?(internal = UserIndividualRequest) ?(local = false) id ?(export_seff=false) (cd, kind) =
+let declare_constant ?(flags={check_guarded=true}) ?(internal = UserIndividualRequest) ?(local = false) id ?(export_seff=false) (cd, kind) =
let export = (* We deal with side effects *)
match cd with
| DefinitionEntry de when
@@ -259,24 +259,24 @@ let declare_constant ?(internal = UserIndividualRequest) ?(local = false) id ?(e
| _ -> false
in
let cst = {
- cst_decl = ConstantEntry (export,cd);
+ cst_decl = ConstantEntry (export,cd,flags);
cst_hyps = [] ;
cst_kind = kind;
cst_locl = local;
cst_exported = [];
cst_was_seff = false;
} in
- let kn = declare_constant_common id cst in
+ let kn = declare_constant_common id cst ~flags in
let () = if_xml (Hook.get f_xml_declare_constant) (internal, kn) in
kn
-let declare_definition ?(internal=UserIndividualRequest)
+let declare_definition ?flags ?(internal=UserIndividualRequest)
?(opaque=false) ?(kind=Decl_kinds.Definition) ?(local = false)
?(poly=false) id ?types (body,ctx) =
let cb =
definition_entry ?types ~poly ~univs:(Univ.ContextSet.to_context ctx) ~opaque body
in
- declare_constant ~internal ~local id
+ declare_constant ?flags ~internal ~local id
(Entries.DefinitionEntry cb, Decl_kinds.IsDefinition kind)
(** Declaration of inductive blocks *)
@@ -353,7 +353,8 @@ let dummy_inductive_entry (_,m) = ([],{
mind_entry_inds = List.map dummy_one_inductive_entry m.mind_entry_inds;
mind_entry_polymorphic = false;
mind_entry_universes = Univ.UContext.empty;
- mind_entry_private = None })
+ mind_entry_private = None;
+ mind_entry_check_positivity = true; })
type inductive_obj = Dischargedhypsmap.discharged_hyps * mutual_inductive_entry
@@ -373,7 +374,7 @@ let declare_projections mind =
Array.iteri (fun i kn ->
let id = Label.to_id (Constant.label kn) in
let entry = {proj_entry_ind = mind; proj_entry_arg = i} in
- let kn' = declare_constant id (ProjectionEntry entry,
+ let kn' = declare_constant ~flags:{check_guarded=true} id (ProjectionEntry entry,
IsDefinition StructureComponent)
in
assert(eq_constant kn kn')) kns; true,true
diff --git a/library/declare.mli b/library/declare.mli
index 8dd24d278..3ba63b5a6 100644
--- a/library/declare.mli
+++ b/library/declare.mli
@@ -54,9 +54,11 @@ val definition_entry : ?fix_exn:Future.fix_exn ->
?eff:Safe_typing.private_constants -> constr -> Safe_typing.private_constants definition_entry
val declare_constant :
+ ?flags:Declarations.typing_flags -> (** default [check_guarded=true] *)
?internal:internal_flag -> ?local:bool -> Id.t -> ?export_seff:bool -> constant_declaration -> constant
val declare_definition :
+ ?flags:Declarations.typing_flags -> (** default [check_guarded=true] *)
?internal:internal_flag -> ?opaque:bool -> ?kind:definition_object_kind ->
?local:bool -> ?poly:polymorphic -> Id.t -> ?types:constr ->
constr Univ.in_universe_context_set -> constant
diff --git a/library/global.ml b/library/global.ml
index 2398e92b0..f4ee62b6e 100644
--- a/library/global.ml
+++ b/library/global.ml
@@ -77,8 +77,8 @@ let globalize_with_summary fs f =
let i2l = Label.of_id
-let push_named_assum a = globalize0 (Safe_typing.push_named_assum a)
-let push_named_def d = globalize (Safe_typing.push_named_def d)
+let push_named_assum ~flags a = globalize0 (Safe_typing.push_named_assum ~flags a)
+let push_named_def ~flags d = globalize (Safe_typing.push_named_def ~flags d)
let add_constraints c = globalize0 (Safe_typing.add_constraints c)
let push_context_set b c = globalize0 (Safe_typing.push_context_set b c)
let push_context b c = globalize0 (Safe_typing.push_context b c)
diff --git a/library/global.mli b/library/global.mli
index bf653307c..7c6cecb4e 100644
--- a/library/global.mli
+++ b/library/global.mli
@@ -30,8 +30,8 @@ val set_engagement : Declarations.engagement -> unit
(** Variables, Local definitions, constants, inductive types *)
-val push_named_assum : (Id.t * Constr.types * bool) Univ.in_universe_context_set -> unit
-val push_named_def : (Id.t * Safe_typing.private_constants Entries.definition_entry) -> Univ.universe_context_set
+val push_named_assum : flags:Declarations.typing_flags -> (Id.t * Constr.types * bool) Univ.in_universe_context_set -> unit
+val push_named_def : flags:Declarations.typing_flags -> (Id.t * Safe_typing.private_constants Entries.definition_entry) -> Univ.universe_context_set
val add_constant :
DirPath.t -> Id.t -> Safe_typing.global_declaration ->
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index c424fe122..f69c4d821 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -1432,7 +1432,7 @@ let do_build_inductive
(* in *)
let _time2 = System.get_time () in
try
- with_full_print (Flags.silently (Command.do_mutual_inductive rel_inds (Flags.is_universe_polymorphism ()) false)) Decl_kinds.Finite
+ with_full_print (Flags.silently (Command.do_mutual_inductive true rel_inds (Flags.is_universe_polymorphism ()) false)) Decl_kinds.Finite
with
| UserError(s,msg) as e ->
let _time3 = System.get_time () in
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index 0cacb003d..d8340dddb 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -400,7 +400,7 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp
in
evd,List.rev rev_pconstants
| _ ->
- Command.do_fixpoint Global (Flags.is_universe_polymorphism ()) fixpoint_exprl;
+ Command.do_fixpoint ~flags:{Declarations.check_guarded=true} Global (Flags.is_universe_polymorphism ()) fixpoint_exprl;
let evd,rev_pconstants =
List.fold_left
(fun (evd,l) ((((_,fname),_),_,_,_,_),_) ->
diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml
index 99a165044..a78eb1af7 100644
--- a/plugins/funind/merge.ml
+++ b/plugins/funind/merge.ml
@@ -887,7 +887,7 @@ let merge_inductive (ind1: inductive) (ind2: inductive)
let indexpr = glob_constr_list_to_inductive_expr prms1 prms2 mib1 mib2 shift_prm rawlist in
(* Declare inductive *)
let indl,_,_ = Command.extract_mutual_inductive_declaration_components [(indexpr,[])] in
- let mie,pl,impls = Command.interp_mutual_inductive indl []
+ let mie,pl,impls = Command.interp_mutual_inductive true indl []
false (*FIXMEnon-poly *) false (* means not private *) Decl_kinds.Finite (* means: not coinductive *) in
(* Declare the mutual inductive block with its associated schemes *)
ignore (Command.declare_mutual_inductive_with_eliminations mie pl impls)
diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml
index 80f1988a9..0b488308a 100644
--- a/pretyping/inductiveops.ml
+++ b/pretyping/inductiveops.ml
@@ -592,9 +592,9 @@ let type_of_projection_knowing_arg env sigma p c ty =
let control_only_guard env c =
let check_fix_cofix e c = match kind_of_term c with
| CoFix (_,(_,_,_) as cofix) ->
- Inductive.check_cofix e cofix
+ Inductive.check_cofix ~flags:{check_guarded=true} e cofix
| Fix (_,(_,_,_) as fix) ->
- Inductive.check_fix e fix
+ Inductive.check_fix ~flags:{check_guarded=true} e fix
| _ -> ()
in
let rec iter env c =
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index 378294c98..c894d96a7 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -72,14 +72,17 @@ open Inductiveops
exception Found of int array
-let search_guard loc env possible_indexes fixdefs =
+(* spiwack: I chose [tflags] rather than [flags], like in the rest of
+ the code, for the argument name to avoid interference with the
+ argument for [inference_flags] also used in this module. *)
+let search_guard ~tflags loc env possible_indexes fixdefs =
(* Standard situation with only one possibility for each fix. *)
(* We treat it separately in order to get proper error msg. *)
let is_singleton = function [_] -> true | _ -> false in
if List.for_all is_singleton possible_indexes then
let indexes = Array.of_list (List.map List.hd possible_indexes) in
let fix = ((indexes, 0),fixdefs) in
- (try check_fix env fix
+ (try check_fix env ~flags:tflags fix
with reraise ->
let (e, info) = Errors.push reraise in
let info = Loc.add_loc info loc in
@@ -91,8 +94,14 @@ let search_guard loc env possible_indexes fixdefs =
List.iter
(fun l ->
let indexes = Array.of_list l in
- let fix = ((indexes, 0),fixdefs) in
- try check_fix env fix; raise (Found indexes)
+ let fix = ((indexes, 0),fixdefs) in
+ (* spiwack: We search for a unspecified structural
+ argument under the assumption that we need to check the
+ guardedness condition (otherwise the first inductive argument
+ will be chosen). A more robust solution may be to raise an
+ error when totality is assumed but the strutural argument is
+ not specified. *)
+ try check_fix env ~flags:{Declarations.check_guarded=true} fix; raise (Found indexes)
with TypeError _ -> ())
(List.combinations possible_indexes);
let errmsg = "Cannot guess decreasing argument of fix." in
@@ -606,11 +615,15 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) env evdref (lvar : ltac_
vn)
in
let fixdecls = (names,ftys,fdefs) in
- let indexes = search_guard loc env possible_indexes fixdecls in
+ let indexes =
+ search_guard
+ ~tflags:{Declarations.check_guarded=true}
+ loc env possible_indexes fixdecls
+ in
make_judge (mkFix ((indexes,i),fixdecls)) ftys.(i)
| GCoFix i ->
let cofix = (i,(names,ftys,fdefs)) in
- (try check_cofix env cofix
+ (try check_cofix env ~flags:{Declarations.check_guarded=true} cofix
with reraise ->
let (e, info) = Errors.push reraise in
let info = Loc.add_loc info loc in
diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli
index 824bb11aa..2c02b4a21 100644
--- a/pretyping/pretyping.mli
+++ b/pretyping/pretyping.mli
@@ -22,7 +22,7 @@ open Misctypes
(** An auxiliary function for searching for fixpoint guard indexes *)
-val search_guard :
+val search_guard : tflags:Declarations.typing_flags ->
Loc.t -> env -> int list list -> rec_declaration -> int array
type typing_constraint = OfType of types | IsType | WithoutTypeConstraint
diff --git a/pretyping/typing.ml b/pretyping/typing.ml
index 52afa7f83..598dd16d0 100644
--- a/pretyping/typing.ml
+++ b/pretyping/typing.ml
@@ -189,13 +189,13 @@ let rec execute env evdref cstr =
| Fix ((vn,i as vni),recdef) ->
let (_,tys,_ as recdef') = execute_recdef env evdref recdef in
let fix = (vni,recdef') in
- check_fix env fix;
+ check_fix env ~flags:{Declarations.check_guarded=true} fix;
make_judge (mkFix fix) tys.(i)
| CoFix (i,recdef) ->
let (_,tys,_ as recdef') = execute_recdef env evdref recdef in
let cofix = (i,recdef') in
- check_cofix env cofix;
+ check_cofix env ~flags:{Declarations.check_guarded=true} cofix;
make_judge (mkCoFix cofix) tys.(i)
| Sort (Prop c) ->
diff --git a/printing/printer.ml b/printing/printer.ml
index cc8da4097..2357ca0ea 100644
--- a/printing/printer.ml
+++ b/printing/printer.ml
@@ -697,9 +697,14 @@ let prterm = pr_lconstr
(* Printer function for sets of Assumptions.assumptions.
It is used primarily by the Print Assumptions command. *)
+type axiom =
+ | Constant of constant (* An axiom or a constant. *)
+ | Positive of MutInd.t (* A mutually inductive definition which has been assumed positive. *)
+ | Guarded of constant (* a constant whose (co)fixpoints have been assumed to be guarded *)
+
type context_object =
| Variable of Id.t (* A section variable or a Let definition *)
- | Axiom of constant * (Label.t * Context.Rel.t * types) list
+ | Axiom of axiom * (Label.t * Context.Rel.t * types) list
| Opaque of constant (* An opaque constant. *)
| Transparent of constant
@@ -707,12 +712,25 @@ type context_object =
module OrderedContextObject =
struct
type t = context_object
+
+ let compare_axiom x y =
+ match x,y with
+ | Constant k1 , Constant k2 ->
+ con_ord k1 k2
+ | Positive m1 , Positive m2 ->
+ MutInd.CanOrd.compare m1 m2
+ | Guarded k1 , Guarded k2 ->
+ con_ord k1 k2
+ | _ , Constant _ -> 1
+ | _ , Positive _ -> 1
+ | _ -> -1
+
let compare x y =
match x , y with
| Variable i1 , Variable i2 -> Id.compare i1 i2
| Variable _ , _ -> -1
| _ , Variable _ -> 1
- | Axiom (k1,_) , Axiom (k2, _) -> con_ord k1 k2
+ | Axiom (k1,_) , Axiom (k2, _) -> compare_axiom k1 k2
| Axiom _ , _ -> -1
| _ , Axiom _ -> 1
| Opaque k1 , Opaque k2 -> con_ord k1 k2
@@ -745,17 +763,26 @@ let pr_assumptionset env s =
try str " " ++ pr_ltype_env env sigma typ
with e when Errors.noncritical e -> mt ()
in
+ let pr_axiom env ax typ =
+ match ax with
+ | Constant kn ->
+ safe_pr_constant env kn ++ safe_pr_ltype typ
+ | Positive m ->
+ hov 2 (MutInd.print m ++ spc () ++ strbrk"is positive.")
+ | Guarded kn ->
+ hov 2 (safe_pr_constant env kn ++ spc () ++ strbrk"is positive.")
+ in
let fold t typ accu =
let (v, a, o, tr) = accu in
match t with
| Variable id ->
let var = pr_id id ++ str " : " ++ pr_ltype typ in
(var :: v, a, o, tr)
- | Axiom (kn,[]) ->
- let ax = safe_pr_constant env kn ++ safe_pr_ltype typ in
+ | Axiom (axiom, []) ->
+ let ax = pr_axiom env axiom typ in
(v, ax :: a, o, tr)
- | Axiom (kn,l) ->
- let ax = safe_pr_constant env kn ++ safe_pr_ltype typ ++
+ | Axiom (axiom,l) ->
+ let ax = pr_axiom env axiom typ ++
cut() ++
prlist_with_sep cut (fun (lbl, ctx, ty) ->
str " used in " ++ pr_label lbl ++
diff --git a/printing/printer.mli b/printing/printer.mli
index 70993bb72..695ab33b2 100644
--- a/printing/printer.mli
+++ b/printing/printer.mli
@@ -161,12 +161,16 @@ val prterm : constr -> std_ppcmds (** = pr_lconstr *)
(** Declarations for the "Print Assumption" command *)
+type axiom =
+ | Constant of constant (* An axiom or a constant. *)
+ | Positive of MutInd.t (* A mutually inductive definition which has been assumed positive. *)
+ | Guarded of constant (* a constant whose (co)fixpoints have been assumed to be guarded *)
+
type context_object =
- | Variable of Id.t (** A section variable or a Let definition *)
- (** An axiom and the type it inhabits (if an axiom of the empty type) *)
- | Axiom of constant * (Label.t * Context.Rel.t * types) list
- | Opaque of constant (** An opaque constant. *)
- | Transparent of constant (** A transparent constant *)
+ | Variable of Id.t (* A section variable or a Let definition *)
+ | Axiom of axiom * (Label.t * Context.Rel.t * types) list
+ | Opaque of constant (* An opaque constant. *)
+ | Transparent of constant
module ContextObjectSet : Set.S with type elt = context_object
module ContextObjectMap : CMap.ExtS
diff --git a/stm/lemmas.ml b/stm/lemmas.ml
index 0a63a3a0f..b84b1265e 100644
--- a/stm/lemmas.ml
+++ b/stm/lemmas.ml
@@ -79,7 +79,7 @@ let adjust_guardness_conditions const = function
List.fold_left (fun e (_,c,cb,_) -> add c cb e) env l)
env (Safe_typing.side_effects_of_private_constants eff) in
let indexes =
- search_guard Loc.ghost env
+ search_guard ~tflags:{Declarations.check_guarded=true} Loc.ghost env
possible_indexes fixdecls in
(mkFix ((indexes,0),fixdecls), ctx), eff
| _ -> (body, ctx), eff) }
diff --git a/toplevel/assumptions.ml b/toplevel/assumptions.ml
index 1802b2d36..c05c5f6c2 100644
--- a/toplevel/assumptions.ml
+++ b/toplevel/assumptions.ml
@@ -286,11 +286,17 @@ let assumptions ?(add_opaque=false) ?(add_transparent=false) st gr t =
if is_local_assum decl then ContextObjectMap.add (Variable id) t accu
else accu
| ConstRef kn ->
- let cb = lookup_constant kn in
+ let cb = lookup_constant kn in
+ let accu =
+ if cb.const_typing_flags.check_guarded then accu
+ else
+ let l = try Refmap_env.find obj ax2ty with Not_found -> [] in
+ ContextObjectMap.add (Axiom (Guarded kn, l)) Constr.mkProp accu
+ in
if not (Declareops.constant_has_body cb) then
let t = type_of_constant cb in
let l = try Refmap_env.find obj ax2ty with Not_found -> [] in
- ContextObjectMap.add (Axiom (kn,l)) t accu
+ ContextObjectMap.add (Axiom (Constant kn,l)) t accu
else if add_opaque && (Declareops.is_opaque cb || not (Cpred.mem kn knst)) then
let t = type_of_constant cb in
ContextObjectMap.add (Opaque kn) t accu
@@ -299,6 +305,12 @@ let assumptions ?(add_opaque=false) ?(add_transparent=false) st gr t =
ContextObjectMap.add (Transparent kn) t accu
else
accu
- | IndRef _ | ConstructRef _ -> accu
+ | IndRef (m,_) | ConstructRef ((m,_),_) ->
+ let mind = Global.lookup_mind m in
+ if mind.mind_checked_positive then
+ accu
+ else
+ let l = try Refmap_env.find obj ax2ty with Not_found -> [] in
+ ContextObjectMap.add (Axiom (Positive m, l)) Constr.mkProp accu
in
Refmap_env.fold fold graph ContextObjectMap.empty
diff --git a/toplevel/command.ml b/toplevel/command.ml
index f0f678e08..6f2dd1bf1 100644
--- a/toplevel/command.ml
+++ b/toplevel/command.ml
@@ -145,9 +145,9 @@ let get_locality id = function
| Local -> true
| Global -> false
-let declare_global_definition ident ce local k pl imps =
+let declare_global_definition ~flags ident ce local k pl imps =
let local = get_locality ident local in
- let kn = declare_constant ident ~local (DefinitionEntry ce, IsDefinition k) in
+ let kn = declare_constant ~flags ident ~local (DefinitionEntry ce, IsDefinition k) in
let gr = ConstRef kn in
let () = maybe_declare_manual_implicits false gr imps in
let () = Universes.register_universe_binders gr pl in
@@ -158,7 +158,7 @@ let declare_definition_hook = ref ignore
let set_declare_definition_hook = (:=) declare_definition_hook
let get_declare_definition_hook () = !declare_definition_hook
-let declare_definition ident (local, p, k) ce pl imps hook =
+let declare_definition ~flags ident (local, p, k) ce pl imps hook =
let fix_exn = Future.fix_exn_of ce.const_entry_body in
let () = !declare_definition_hook ce in
let r = match local with
@@ -175,11 +175,11 @@ let declare_definition ident (local, p, k) ce pl imps hook =
in
gr
| Discharge | Local | Global ->
- declare_global_definition ident ce local k pl imps in
+ declare_global_definition ~flags ident ce local k pl imps in
Lemmas.call_hook fix_exn hook local r
let _ = Obligations.declare_definition_ref :=
- (fun i k c imps hook -> declare_definition i k c [] imps hook)
+ (fun i k c imps hook -> declare_definition ~flags:{Declarations.check_guarded=true} i k c [] imps hook)
let do_definition ident k pl bl red_option c ctypopt hook =
let (ce, evd, pl', imps as def) =
@@ -203,7 +203,7 @@ let do_definition ident k pl bl red_option c ctypopt hook =
ignore(Obligations.add_definition
ident ~term:c cty ctx ?pl ~implicits:imps ~kind:k ~hook obls)
else let ce = check_definition def in
- ignore(declare_definition ident k ce pl' imps
+ ignore(declare_definition ~flags:{Declarations.check_guarded=true} ident k ce pl' imps
(Lemmas.mk_hook
(fun l r -> Lemmas.call_hook (fun exn -> exn) hook l r;r)))
@@ -549,7 +549,7 @@ let check_param = function
| LocalRawAssum (nas, Default _, _) -> List.iter check_named nas
| LocalRawAssum (nas, Generalized _, _) -> ()
-let interp_mutual_inductive (paramsl,indl) notations poly prv finite =
+let interp_mutual_inductive chk (paramsl,indl) notations poly prv finite =
check_all_names_different indl;
List.iter check_param paramsl;
let env0 = Global.env() in
@@ -630,7 +630,8 @@ let interp_mutual_inductive (paramsl,indl) notations poly prv finite =
mind_entry_inds = entries;
mind_entry_polymorphic = poly;
mind_entry_private = if prv then Some false else None;
- mind_entry_universes = uctx },
+ mind_entry_universes = uctx;
+ mind_entry_check_positivity = chk; },
pl, impls
(* Very syntactical equality *)
@@ -715,17 +716,20 @@ type one_inductive_impls =
Impargs.manual_explicitation list (* for inds *)*
Impargs.manual_explicitation list list (* for constrs *)
-let do_mutual_inductive indl poly prv finite =
+let do_mutual_inductive chk indl poly prv finite =
let indl,coes,ntns = extract_mutual_inductive_declaration_components indl in
(* Interpret the types *)
- let mie,pl,impls = interp_mutual_inductive indl ntns poly prv finite in
+ let mie,pl,impls = interp_mutual_inductive chk indl ntns poly prv finite in
(* Declare the mutual inductive block with its associated schemes *)
ignore (declare_mutual_inductive_with_eliminations mie pl impls);
(* Declare the possible notations of inductive types *)
List.iter Metasyntax.add_notation_interpretation ntns;
(* Declare the coercions *)
- List.iter (fun qid -> Class.try_add_new_coercion (locate qid) false poly) coes
-
+ List.iter (fun qid -> Class.try_add_new_coercion (locate qid) false poly) coes;
+ (* If [chk] is [false] (i.e. positivity is assumed) declares itself
+ as unsafe. *)
+ if not chk then Feedback.feedback Feedback.AddedAxiom else ()
+
(* 3c| Fixpoints and co-fixpoints *)
(* An (unoptimized) function that maps preorders to partial orders...
@@ -829,12 +833,12 @@ let interp_fix_body env_rec evdref impls (_,ctx) fix ccl =
let build_fix_type (_,ctx) ccl = it_mkProd_or_LetIn ccl ctx
-let declare_fix ?(opaque = false) (_,poly,_ as kind) pl ctx f ((def,_),eff) t imps =
+let declare_fix ~flags ?(opaque = false) (_,poly,_ as kind) pl ctx f ((def,_),eff) t imps =
let ce = definition_entry ~opaque ~types:t ~poly ~univs:ctx ~eff def in
- declare_definition f kind ce pl imps (Lemmas.mk_hook (fun _ r -> r))
+ declare_definition ~flags f kind ce pl imps (Lemmas.mk_hook (fun _ r -> r))
let _ = Obligations.declare_fix_ref :=
- (fun ?opaque k ctx f d t imps -> declare_fix ?opaque k [] ctx f d t imps)
+ (fun ?opaque k ctx f d t imps -> declare_fix ~flags:{Declarations.check_guarded=true} ?opaque k [] ctx f d t imps)
let prepare_recursive_declaration fixnames fixtypes fixdefs =
let defs = List.map (subst_vars (List.rev fixnames)) fixdefs in
@@ -1135,7 +1139,7 @@ let interp_cofixpoint l ntns =
check_recursive false env evd fix;
(fix,pl,Evd.evar_universe_context evd,info)
-let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) indexes ntns =
+let declare_fixpoint ~flags local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) indexes ntns =
if List.exists Option.is_empty fixdefs then
(* Some bodies to define by proof *)
let thms =
@@ -1155,7 +1159,7 @@ let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ind
let fixdefs = List.map Option.get fixdefs in
let fixdecls = prepare_recursive_declaration fixnames fixtypes fixdefs in
let env = Global.env() in
- let indexes = search_guard Loc.ghost env indexes fixdecls in
+ let indexes = search_guard ~tflags:flags Loc.ghost env indexes fixdecls in
let fiximps = List.map (fun (n,r,p) -> r) fiximps in
let vars = Universes.universes_of_constr (mkFix ((indexes,0),fixdecls)) in
let fixdecls =
@@ -1164,7 +1168,7 @@ let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ind
let evd = Evd.restrict_universe_context evd vars in
let fixdecls = List.map Safe_typing.mk_pure_proof fixdecls in
let pl, ctx = Evd.universe_context ?names:pl evd in
- ignore (List.map4 (declare_fix (local, poly, Fixpoint) pl ctx)
+ ignore (List.map4 (declare_fix ~flags (local, poly, Fixpoint) pl ctx)
fixnames fixdecls fixtypes fiximps);
(* Declare the recursive definitions *)
fixpoint_message (Some indexes) fixnames;
@@ -1172,7 +1176,7 @@ let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ind
(* Declare notations *)
List.iter Metasyntax.add_notation_interpretation ntns
-let declare_cofixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ntns =
+let declare_cofixpoint ~flags local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ntns =
if List.exists Option.is_empty fixdefs then
(* Some bodies to define by proof *)
let thms =
@@ -1198,7 +1202,7 @@ let declare_cofixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) n
let evd = Evd.from_ctx ctx in
let evd = Evd.restrict_universe_context evd vars in
let pl, ctx = Evd.universe_context ?names:pl evd in
- ignore (List.map4 (declare_fix (local, poly, CoFixpoint) pl ctx)
+ ignore (List.map4 (declare_fix ~flags (local, poly, CoFixpoint) pl ctx)
fixnames fixdecls fixtypes fiximps);
(* Declare the recursive definitions *)
cofixpoint_message fixnames
@@ -1264,8 +1268,14 @@ let do_program_recursive local p fixkind fixl ntns =
Array.of_list (List.map (subst_vars (List.rev fixnames)) fixdefs)
in
let indexes =
- Pretyping.search_guard Loc.ghost (Global.env ()) possible_indexes fixdecls in
- List.iteri (fun i _ -> Inductive.check_fix env ((indexes,i),fixdecls)) fixl
+ Pretyping.search_guard
+ ~tflags:{Declarations.check_guarded=true}
+ Loc.ghost (Global.env ()) possible_indexes fixdecls in
+ List.iteri (fun i _ ->
+ Inductive.check_fix env
+ ~flags:{Declarations.check_guarded=true}
+ ((indexes,i),fixdecls))
+ fixl
end in
let ctx = Evd.evar_universe_context evd in
let kind = match fixkind with
@@ -1299,19 +1309,21 @@ let do_program_fixpoint local poly l =
errorlabstrm "do_program_fixpoint"
(str "Well-founded fixpoints not allowed in mutually recursive blocks")
-let do_fixpoint local poly l =
+let do_fixpoint ~flags local poly l =
if Flags.is_program_mode () then do_program_fixpoint local poly l
else
let fixl, ntns = extract_fixpoint_components true l in
let (_, _, _, info as fix) = interp_fixpoint fixl ntns in
let possible_indexes =
List.map compute_possible_guardness_evidences info in
- declare_fixpoint local poly fix possible_indexes ntns
+ declare_fixpoint ~flags local poly fix possible_indexes ntns;
+ if not flags.Declarations.check_guarded then Feedback.feedback Feedback.AddedAxiom else ()
-let do_cofixpoint local poly l =
+let do_cofixpoint ~flags local poly l =
let fixl,ntns = extract_cofixpoint_components l in
if Flags.is_program_mode () then
do_program_recursive local poly Obligations.IsCoFixpoint fixl ntns
else
let cofix = interp_cofixpoint fixl ntns in
- declare_cofixpoint local poly cofix ntns
+ declare_cofixpoint ~flags local poly cofix ntns;
+ if not flags.Declarations.check_guarded then Feedback.feedback Feedback.AddedAxiom else ()
diff --git a/toplevel/command.mli b/toplevel/command.mli
index b97cb487d..2d27552a1 100644
--- a/toplevel/command.mli
+++ b/toplevel/command.mli
@@ -36,7 +36,7 @@ val interp_definition :
constr_expr option -> Safe_typing.private_constants definition_entry * Evd.evar_map *
Universes.universe_binders * Impargs.manual_implicits
-val declare_definition : Id.t -> definition_kind ->
+val declare_definition : flags:Declarations.typing_flags -> Id.t -> definition_kind ->
Safe_typing.private_constants definition_entry -> Universes.universe_binders -> Impargs.manual_implicits ->
Globnames.global_reference Lemmas.declaration_hook -> Globnames.global_reference
@@ -91,6 +91,7 @@ type one_inductive_impls =
Impargs.manual_implicits list (** for constrs *)
val interp_mutual_inductive :
+ bool -> (* if [false], then positivity is assumed *)
structured_inductive_expr -> decl_notation list -> polymorphic ->
private_flag -> Decl_kinds.recursivity_kind ->
mutual_inductive_entry * Universes.universe_binders * one_inductive_impls list
@@ -105,6 +106,7 @@ val declare_mutual_inductive_with_eliminations :
(** Entry points for the vernacular commands Inductive and CoInductive *)
val do_mutual_inductive :
+ bool -> (* if [false], then positivity is assumed *)
(one_inductive_expr * decl_notation list) list -> polymorphic ->
private_flag -> Decl_kinds.recursivity_kind -> unit
@@ -148,12 +150,13 @@ val interp_cofixpoint :
(** Registering fixpoints and cofixpoints in the environment *)
val declare_fixpoint :
+ flags:Declarations.typing_flags ->
locality -> polymorphic ->
recursive_preentry * lident list option * Evd.evar_universe_context *
(Name.t list * Impargs.manual_implicits * int option) list ->
lemma_possible_guards -> decl_notation list -> unit
-val declare_cofixpoint : locality -> polymorphic ->
+val declare_cofixpoint : flags:Declarations.typing_flags -> locality -> polymorphic ->
recursive_preentry * lident list option * Evd.evar_universe_context *
(Name.t list * Impargs.manual_implicits * int option) list ->
decl_notation list -> unit
@@ -161,14 +164,16 @@ val declare_cofixpoint : locality -> polymorphic ->
(** Entry points for the vernacular commands Fixpoint and CoFixpoint *)
val do_fixpoint :
+ flags:Declarations.typing_flags -> (* When [false], assume guarded. *)
locality -> polymorphic -> (fixpoint_expr * decl_notation list) list -> unit
val do_cofixpoint :
+ flags:Declarations.typing_flags -> (* When [false], assume guarded. *)
locality -> polymorphic -> (cofixpoint_expr * decl_notation list) list -> unit
(** Utils *)
val check_mutuality : Environ.env -> bool -> (Id.t * types) list -> unit
-val declare_fix : ?opaque:bool -> definition_kind -> Universes.universe_binders -> Univ.universe_context -> Id.t ->
+val declare_fix : flags:Declarations.typing_flags -> ?opaque:bool -> definition_kind -> Universes.universe_binders -> Univ.universe_context -> Id.t ->
Safe_typing.private_constants Entries.proof_output -> types -> Impargs.manual_implicits -> global_reference
diff --git a/toplevel/discharge.ml b/toplevel/discharge.ml
index 6b267283a..74361eb1c 100644
--- a/toplevel/discharge.ml
+++ b/toplevel/discharge.ml
@@ -116,5 +116,6 @@ let process_inductive (sechyps,abs_ctx) modlist mib =
mind_entry_inds = inds';
mind_entry_polymorphic = mib.mind_polymorphic;
mind_entry_private = mib.mind_private;
- mind_entry_universes = univs
+ mind_entry_universes = univs;
+ mind_entry_check_positivity = mib.mind_checked_positive;
}
diff --git a/toplevel/indschemes.ml b/toplevel/indschemes.ml
index f995a390c..a43758e3c 100644
--- a/toplevel/indschemes.ml
+++ b/toplevel/indschemes.ml
@@ -497,7 +497,7 @@ let map_inductive_block f kn n = for i=0 to n-1 do f (kn,i) done
let declare_default_schemes kn =
let mib = Global.lookup_mind kn in
let n = Array.length mib.mind_packets in
- if !elim_flag && (mib.mind_finite <> BiFinite || !bifinite_elim_flag) then
+ if !elim_flag && (mib.mind_finite <> BiFinite || !bifinite_elim_flag) && mib.mind_checked_positive then
declare_induction_schemes kn;
if !case_flag then map_inductive_block declare_one_case_analysis_scheme kn n;
if is_eq_flag() then try_declare_beq_scheme kn;
diff --git a/toplevel/obligations.ml b/toplevel/obligations.ml
index dbc3ccaac..00fd97210 100644
--- a/toplevel/obligations.ml
+++ b/toplevel/obligations.ml
@@ -567,7 +567,8 @@ let declare_mutual_definition l =
List.map3 compute_possible_guardness_evidences
wfl fixdefs fixtypes in
let indexes =
- Pretyping.search_guard Loc.ghost (Global.env())
+ Pretyping.search_guard ~tflags:{Declarations.check_guarded=true}
+ Loc.ghost (Global.env())
possible_indexes fixdecls in
Some indexes,
List.map_i (fun i _ ->
diff --git a/toplevel/record.ml b/toplevel/record.ml
index 0c3bd953c..214d44d83 100644
--- a/toplevel/record.ml
+++ b/toplevel/record.ml
@@ -361,7 +361,7 @@ let structure_signature ctx =
open Typeclasses
-let declare_structure finite poly ctx id idbuild paramimpls params arity template
+let declare_structure chk finite poly ctx id idbuild paramimpls params arity template
fieldimpls fields ?(kind=StructureComponent) ?name is_coe coers sign =
let nparams = List.length params and nfields = List.length fields in
let args = Context.Rel.to_extended_list nfields params in
@@ -386,7 +386,8 @@ let declare_structure finite poly ctx id idbuild paramimpls params arity templat
mind_entry_inds = [mie_ind];
mind_entry_polymorphic = poly;
mind_entry_private = None;
- mind_entry_universes = ctx } in
+ mind_entry_universes = ctx;
+ mind_entry_check_positivity = chk; } in
let kn = Command.declare_mutual_inductive_with_eliminations mie [] [(paramimpls,[])] in
let rsp = (kn,0) in (* This is ind path of idstruc *)
let cstr = (rsp,1) in
@@ -405,7 +406,7 @@ let implicits_of_context ctx =
in ExplByPos (i, explname), (true, true, true))
1 (List.rev (Anonymous :: (List.map get_name ctx)))
-let declare_class finite def poly ctx id idbuild paramimpls params arity
+let declare_class chk finite def poly ctx id idbuild paramimpls params arity
template fieldimpls fields ?(kind=StructureComponent) is_coe coers priorities sign =
let fieldimpls =
(* Make the class implicit in the projections, and the params if applicable. *)
@@ -448,7 +449,7 @@ let declare_class finite def poly ctx id idbuild paramimpls params arity
in
cref, [Name proj_name, sub, Some proj_cst]
| _ ->
- let ind = declare_structure BiFinite poly ctx (snd id) idbuild paramimpls
+ let ind = declare_structure chk BiFinite poly ctx (snd id) idbuild paramimpls
params arity template fieldimpls fields
~kind:Method ~name:binder_name false (List.map (fun _ -> false) fields) sign
in
@@ -527,8 +528,9 @@ open Vernacexpr
(* [fs] corresponds to fields and [ps] to parameters; [coers] is a
list telling if the corresponding fields must me declared as coercions
- or subinstances *)
-let definition_structure (kind,poly,finite,(is_coe,((loc,idstruc),pl)),ps,cfs,idbuild,s) =
+ or subinstances. When [chk] is false positivity is
+ assumed. *)
+let definition_structure chk (kind,poly,finite,(is_coe,((loc,idstruc),pl)),ps,cfs,idbuild,s) =
let cfs,notations = List.split cfs in
let cfs,priorities = List.split cfs in
let coers,fs = List.split cfs in
@@ -549,14 +551,14 @@ let definition_structure (kind,poly,finite,(is_coe,((loc,idstruc),pl)),ps,cfs,id
let sign = structure_signature (fields@params) in
let gr = match kind with
| Class def ->
- let gr = declare_class finite def poly ctx (loc,idstruc) idbuild
+ let gr = declare_class chk finite def poly ctx (loc,idstruc) idbuild
implpars params arity template implfs fields is_coe coers priorities sign in
gr
| _ ->
let implfs = List.map
(fun impls -> implpars @ Impargs.lift_implicits
(succ (List.length params)) impls) implfs in
- let ind = declare_structure finite poly ctx idstruc
+ let ind = declare_structure chk finite poly ctx idstruc
idbuild implpars params arity template implfs
fields is_coe (List.map (fun coe -> not (Option.is_empty coe)) coers) sign in
IndRef ind
diff --git a/toplevel/record.mli b/toplevel/record.mli
index 26eb3378b..525326237 100644
--- a/toplevel/record.mli
+++ b/toplevel/record.mli
@@ -24,7 +24,9 @@ val declare_projections :
coercion_flag list -> manual_explicitation list list -> Context.Rel.t ->
(Name.t * bool) list * constant option list
-val declare_structure : Decl_kinds.recursivity_kind ->
+val declare_structure :
+ bool -> (** check positivity? *)
+ Decl_kinds.recursivity_kind ->
bool (** polymorphic?*) -> Univ.universe_context ->
Id.t -> Id.t ->
manual_explicitation list -> Context.Rel.t -> (** params *) constr -> (** arity *)
@@ -37,6 +39,7 @@ val declare_structure : Decl_kinds.recursivity_kind ->
inductive
val definition_structure :
+ bool -> (** check positivity? *)
inductive_kind * Decl_kinds.polymorphic * Decl_kinds.recursivity_kind * plident with_coercion * local_binder list *
(local_decl_expr with_instance with_priority with_notation) list *
Id.t * constr_expr option -> global_reference
diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml
index 222b7d3df..0b4dc0d18 100644
--- a/toplevel/vernacentries.ml
+++ b/toplevel/vernacentries.ml
@@ -519,7 +519,7 @@ let vernac_assumption locality poly (local, kind) l nl =
let status = do_assumptions kind nl l in
if not status then Feedback.feedback Feedback.AddedAxiom
-let vernac_record k poly finite struc binders sort nameopt cfs =
+let vernac_record chk k poly finite struc binders sort nameopt cfs =
let const = match nameopt with
| None -> add_prefix "Build_" (snd (fst (snd struc)))
| Some (_,id as lid) ->
@@ -530,9 +530,14 @@ let vernac_record k poly finite struc binders sort nameopt cfs =
match x with
| Vernacexpr.AssumExpr ((loc, Name id), _) -> Dumpglob.dump_definition (loc,id) false "proj"
| _ -> ()) cfs);
- ignore(Record.definition_structure (k,poly,finite,struc,binders,cfs,const,sort))
-
-let vernac_inductive poly lo finite indl =
+ ignore(Record.definition_structure chk (k,poly,finite,struc,binders,cfs,const,sort))
+
+(** When [chk] is false, positivity is assumed. When [poly] is true
+ the type is declared polymorphic. When [lo] is true, then the type
+ is declared private (as per the [Private] keyword). [finite]
+ indicates whether the type is inductive, co-inductive or
+ neither. *)
+let vernac_inductive chk poly lo finite indl =
if Dumpglob.dump () then
List.iter (fun (((coe,(lid,_)), _, _, _, cstrs), _) ->
match cstrs with
@@ -548,14 +553,14 @@ let vernac_inductive poly lo finite indl =
| [ (_ , _ , _ ,Variant, RecordDecl _),_ ] ->
Errors.error "The Variant keyword cannot be used to define a record type. Use Record instead."
| [ ( id , bl , c , b, RecordDecl (oc,fs) ), [] ] ->
- vernac_record (match b with Class true -> Class false | _ -> b)
+ vernac_record chk (match b with Class true -> Class false | _ -> b)
poly finite id bl c oc fs
| [ ( id , bl , c , Class true, Constructors [l]), _ ] ->
let f =
let (coe, ((loc, id), ce)) = l in
let coe' = if coe then Some true else None in
(((coe', AssumExpr ((loc, Name id), ce)), None), [])
- in vernac_record (Class true) poly finite id bl c None [f]
+ in vernac_record chk (Class true) poly finite id bl c None [f]
| [ ( id , bl , c , Class true, _), _ ] ->
Errors.error "Definitional classes must have a single method"
| [ ( id , bl , c , Class false, Constructors _), _ ] ->
@@ -569,19 +574,19 @@ let vernac_inductive poly lo finite indl =
| _ -> Errors.error "Cannot handle mutually (co)inductive records."
in
let indl = List.map unpack indl in
- do_mutual_inductive indl poly lo finite
+ do_mutual_inductive chk indl poly lo finite
-let vernac_fixpoint locality poly local l =
+let vernac_fixpoint ~flags locality poly local l =
let local = enforce_locality_exp locality local in
if Dumpglob.dump () then
List.iter (fun (((lid,_), _, _, _, _), _) -> Dumpglob.dump_definition lid false "def") l;
- do_fixpoint local poly l
+ do_fixpoint ~flags local poly l
-let vernac_cofixpoint locality poly local l =
+let vernac_cofixpoint ~flags locality poly local l =
let local = enforce_locality_exp locality local in
if Dumpglob.dump () then
List.iter (fun (((lid,_), _, _, _), _) -> Dumpglob.dump_definition lid false "def") l;
- do_cofixpoint local poly l
+ do_cofixpoint ~flags local poly l
let vernac_scheme l =
if Dumpglob.dump () then
@@ -1726,6 +1731,7 @@ let vernac_load interp fname =
try while true do interp (snd (parse_sentence input)) done
with End_of_input -> ()
+let all_checks = { Declarations.check_guarded = true }
(* "locality" is the prefix "Local" attribute, while the "local" component
* is the outdated/deprecated "Local" attribute of some vernacular commands
@@ -1763,9 +1769,9 @@ let interp ?proof ~loc locality poly c =
| VernacEndProof e -> vernac_end_proof ?proof e
| VernacExactProof c -> vernac_exact_proof c
| VernacAssumption (stre,nl,l) -> vernac_assumption locality poly stre l nl
- | VernacInductive (priv,finite,l) -> vernac_inductive poly priv finite l
- | VernacFixpoint (local, l) -> vernac_fixpoint locality poly local l
- | VernacCoFixpoint (local, l) -> vernac_cofixpoint locality poly local l
+ | VernacInductive (priv,finite,l) -> vernac_inductive true poly priv finite l
+ | VernacFixpoint (local, l) -> vernac_fixpoint locality ~flags:all_checks poly local l
+ | VernacCoFixpoint (local, l) -> vernac_cofixpoint ~flags:all_checks locality poly local l
| VernacScheme l -> vernac_scheme l
| VernacCombinedScheme (id, l) -> vernac_combined_scheme id l
| VernacUniverse l -> vernac_universe loc poly l