aboutsummaryrefslogtreecommitdiffhomepage
path: root/vernac
diff options
context:
space:
mode:
Diffstat (limited to 'vernac')
-rw-r--r--vernac/assumptions.ml10
-rw-r--r--vernac/assumptions.mli10
-rw-r--r--vernac/auto_ind_decl.ml18
-rw-r--r--vernac/auto_ind_decl.mli10
-rw-r--r--vernac/class.ml10
-rw-r--r--vernac/class.mli10
-rw-r--r--vernac/classes.ml50
-rw-r--r--vernac/classes.mli10
-rw-r--r--vernac/comAssumption.ml20
-rw-r--r--vernac/comAssumption.mli12
-rw-r--r--vernac/comDefinition.ml12
-rw-r--r--vernac/comDefinition.mli14
-rw-r--r--vernac/comFixpoint.ml11
-rw-r--r--vernac/comFixpoint.mli12
-rw-r--r--vernac/comInductive.ml34
-rw-r--r--vernac/comInductive.mli12
-rw-r--r--vernac/comProgramFixpoint.ml4
-rw-r--r--vernac/declareDef.ml10
-rw-r--r--vernac/declareDef.mli10
-rw-r--r--vernac/explainErr.ml10
-rw-r--r--vernac/explainErr.mli10
-rw-r--r--vernac/himsg.ml10
-rw-r--r--vernac/himsg.mli10
-rw-r--r--vernac/indschemes.ml19
-rw-r--r--vernac/indschemes.mli10
-rw-r--r--vernac/lemmas.ml10
-rw-r--r--vernac/lemmas.mli10
-rw-r--r--vernac/locality.ml10
-rw-r--r--vernac/locality.mli10
-rw-r--r--vernac/metasyntax.ml10
-rw-r--r--vernac/metasyntax.mli10
-rw-r--r--vernac/mltop.ml10
-rw-r--r--vernac/mltop.mli10
-rw-r--r--vernac/obligations.ml8
-rw-r--r--vernac/obligations.mli10
-rw-r--r--vernac/proof_using.ml10
-rw-r--r--vernac/proof_using.mli10
-rw-r--r--vernac/record.ml16
-rw-r--r--vernac/record.mli10
-rw-r--r--vernac/search.ml10
-rw-r--r--vernac/search.mli10
-rw-r--r--vernac/topfmt.ml10
-rw-r--r--vernac/topfmt.mli10
-rw-r--r--vernac/vernacentries.ml52
-rw-r--r--vernac/vernacentries.mli10
-rw-r--r--vernac/vernacinterp.ml10
-rw-r--r--vernac/vernacinterp.mli10
-rw-r--r--vernac/vernacprop.ml10
-rw-r--r--vernac/vernacprop.mli10
-rw-r--r--vernac/vernacstate.ml10
-rw-r--r--vernac/vernacstate.mli10
51 files changed, 382 insertions, 272 deletions
diff --git a/vernac/assumptions.ml b/vernac/assumptions.ml
index d22024568..45ccf7276 100644
--- a/vernac/assumptions.ml
+++ b/vernac/assumptions.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* The following definitions are used by the function
diff --git a/vernac/assumptions.mli b/vernac/assumptions.mli
index afe932ead..7e13f8f28 100644
--- a/vernac/assumptions.mli
+++ b/vernac/assumptions.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Names
diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml
index ec6b62ee2..2879feba7 100644
--- a/vernac/auto_ind_decl.ml
+++ b/vernac/auto_ind_decl.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* This file is about the automatic generation of schemes about
@@ -321,7 +323,7 @@ let build_beq_scheme mode kn =
raise NoDecidabilityCoInductive;
let fix = mkFix (((Array.make nb_ind 0),i),(names,types,cores)) in
create_input fix),
- Evd.make_evar_universe_context (Global.env ()) None),
+ UState.make (Global.universes ())),
!eff
let beq_scheme_kind = declare_mutual_scheme_object "_beq" build_beq_scheme
@@ -669,7 +671,7 @@ let make_bl_scheme mode mind =
let lnonparrec,lnamesparrec = (* TODO subst *)
context_chop (nparams-nparrec) mib.mind_params_ctxt in
let bl_goal, eff = compute_bl_goal ind lnamesparrec nparrec in
- let ctx = Evd.make_evar_universe_context (Global.env ()) None in
+ let ctx = UState.make (Global.universes ()) in
let side_eff = side_effect_of_mode mode in
let bl_goal = EConstr.of_constr bl_goal in
let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx bl_goal
@@ -793,7 +795,7 @@ let make_lb_scheme mode mind =
let lnonparrec,lnamesparrec =
context_chop (nparams-nparrec) mib.mind_params_ctxt in
let lb_goal, eff = compute_lb_goal ind lnamesparrec nparrec in
- let ctx = Evd.make_evar_universe_context (Global.env ()) None in
+ let ctx = UState.make (Global.universes ()) in
let side_eff = side_effect_of_mode mode in
let lb_goal = EConstr.of_constr lb_goal in
let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx lb_goal
@@ -963,7 +965,7 @@ let make_eq_decidability mode mind =
let nparams = mib.mind_nparams in
let nparrec = mib.mind_nparams_rec in
let u = Univ.Instance.empty in
- let ctx = Evd.make_evar_universe_context (Global.env ()) None in
+ let ctx = UState.make (Global.universes ()) in
let lnonparrec,lnamesparrec =
context_chop (nparams-nparrec) mib.mind_params_ctxt in
let side_eff = side_effect_of_mode mode in
diff --git a/vernac/auto_ind_decl.mli b/vernac/auto_ind_decl.mli
index d841cca11..5cc783df7 100644
--- a/vernac/auto_ind_decl.mli
+++ b/vernac/auto_ind_decl.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Names
diff --git a/vernac/class.ml b/vernac/class.ml
index 943da8fa8..cc676af1b 100644
--- a/vernac/class.ml
+++ b/vernac/class.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open CErrors
diff --git a/vernac/class.mli b/vernac/class.mli
index 29486073b..33d31fe1f 100644
--- a/vernac/class.mli
+++ b/vernac/class.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Names
diff --git a/vernac/classes.ml b/vernac/classes.ml
index 695be74bb..192cc8a55 100644
--- a/vernac/classes.ml
+++ b/vernac/classes.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(*i*)
@@ -52,7 +54,7 @@ let _ =
let open Vernacexpr in
{ info with hint_pattern =
Option.map
- (Constrintern.intern_constr_pattern (Global.env()))
+ (Constrintern.intern_constr_pattern (Global.env()) Evd.(from_env Global.(env())))
info.hint_pattern } in
Flags.silently (fun () ->
Hints.add_hints local [typeclasses_db]
@@ -334,7 +336,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
let init_refine =
Tacticals.New.tclTHENLIST [
Refine.refine ~typecheck:false (fun sigma -> (sigma,EConstr.of_constr (Option.get term)));
- Proofview.Unsafe.tclNEWGOALS gls;
+ Proofview.Unsafe.tclNEWGOALS (CList.map Proofview.with_empty_state gls);
Tactics.New.reduce_after_refine;
]
in
@@ -372,16 +374,34 @@ let context poly l =
with e when CErrors.noncritical e ->
user_err Pp.(str "Anonymous variables not allowed in contexts.")
in
- let uctx = ref (Evd.universe_context_set sigma) in
+ let univs =
+ let uctx = Evd.universe_context_set sigma in
+ match ctx with
+ | [] -> assert false
+ | [_] ->
+ if poly
+ then Polymorphic_const_entry (Univ.ContextSet.to_context uctx)
+ else Monomorphic_const_entry uctx
+ | _::_::_ ->
+ if Lib.sections_are_opened ()
+ then
+ begin
+ Declare.declare_universe_context poly uctx;
+ if poly then Polymorphic_const_entry Univ.UContext.empty
+ else Monomorphic_const_entry Univ.ContextSet.empty
+ end
+ else if poly
+ then Polymorphic_const_entry (Univ.ContextSet.to_context uctx)
+ else
+ begin
+ Declare.declare_universe_context poly uctx;
+ Monomorphic_const_entry Univ.ContextSet.empty
+ end
+ in
let fn status (id, b, t) =
let b, t = Option.map (to_constr sigma) b, to_constr sigma t in
if Lib.is_modtype () && not (Lib.sections_are_opened ()) then
(* Declare the universe context once *)
- let univs = if poly
- then Polymorphic_const_entry (Univ.ContextSet.to_context !uctx)
- else Monomorphic_const_entry !uctx
- in
- let () = uctx := Univ.ContextSet.empty in
let decl = match b with
| None ->
(ParameterEntry (None,(t,univs),None), IsAssumption Logical)
@@ -403,10 +423,6 @@ let context poly l =
in
let impl = List.exists test impls in
let decl = (Discharge, poly, Definitional) in
- let univs = if poly
- then Polymorphic_const_entry (Univ.ContextSet.to_context !uctx)
- else Monomorphic_const_entry !uctx
- in
let nstatus = match b with
| None ->
pi3 (ComAssumption.declare_assumption false decl (t, univs) Universes.empty_binders [] impl
@@ -420,6 +436,4 @@ let context poly l =
in
status && nstatus
in
- if Lib.sections_are_opened () then
- Declare.declare_universe_context poly !uctx;
List.fold_left fn true (List.rev ctx)
diff --git a/vernac/classes.mli b/vernac/classes.mli
index d47c6a6f8..0342c840e 100644
--- a/vernac/classes.mli
+++ b/vernac/classes.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Names
diff --git a/vernac/comAssumption.ml b/vernac/comAssumption.ml
index 7e5b941ad..6a590758f 100644
--- a/vernac/comAssumption.ml
+++ b/vernac/comAssumption.ml
@@ -1,16 +1,17 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Pp
open CErrors
open Util
open Vars
-open Environ
open Declare
open Names
open Globnames
@@ -87,7 +88,6 @@ match local with
let interp_assumption sigma env impls bl c =
let c = mkCProdN ?loc:(local_binders_loc bl) bl c in
let sigma, (ty, impls) = interp_type_evars_impls env sigma ~impls c in
- let ty = EConstr.Unsafe.to_constr ty in
sigma, (ty, impls)
(* When monomorphic the universe constraints are declared with the first declaration only. *)
@@ -151,17 +151,17 @@ let do_assumptions kind nl l =
let (sigma,_,_),l = List.fold_left_map (fun (sigma,env,ienv) (is_coe,(idl,c)) ->
let sigma,(t,imps) = interp_assumption sigma env ienv [] c in
let env =
- push_named_context (List.map (fun {CAst.v=id} -> LocalAssum (id,t)) idl) env in
+ EConstr.push_named_context (List.map (fun {CAst.v=id} -> LocalAssum (id,t)) idl) env in
let ienv = List.fold_right (fun {CAst.v=id} ienv ->
- let impls = compute_internalization_data env Variable t imps in
+ let impls = compute_internalization_data env sigma Variable t imps in
Id.Map.add id impls ienv) idl ienv in
((sigma,env,ienv),((is_coe,idl),t,imps)))
(sigma,env,empty_internalization_env) l
in
let sigma = solve_remaining_evars all_and_fail_flags env sigma Evd.empty in
(* The universe constraints come from the whole telescope. *)
- let sigma = Evd.nf_constraints sigma in
- let nf_evar c = EConstr.to_constr sigma (EConstr.of_constr c) in
+ let sigma = Evd.minimize_universes sigma in
+ let nf_evar c = EConstr.to_constr sigma c in
let uvars, l = List.fold_left_map (fun uvars (coe,t,imps) ->
let t = nf_evar t in
let uvars = Univ.LSet.union uvars (Univops.universes_of_constr env t) in
diff --git a/vernac/comAssumption.mli b/vernac/comAssumption.mli
index 0491638c9..56e324376 100644
--- a/vernac/comAssumption.mli
+++ b/vernac/comAssumption.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Names
@@ -17,7 +19,7 @@ open Decl_kinds
(** {6 Parameters/Assumptions} *)
val do_assumptions : locality * polymorphic * assumption_object_kind ->
- Vernacexpr.inline -> (Vernacexpr.ident_decl list * constr_expr) with_coercion list -> bool
+ Vernacexpr.inline -> (ident_decl list * constr_expr) with_coercion list -> bool
(************************************************************************)
(** Internal API *)
diff --git a/vernac/comDefinition.ml b/vernac/comDefinition.ml
index d376696f7..b18a60a1f 100644
--- a/vernac/comDefinition.ml
+++ b/vernac/comDefinition.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2018 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Pp
@@ -83,7 +85,7 @@ let interp_definition pl bl poly red_option c ctypopt =
evd, c, imps1@Impargs.lift_implicits (Context.Rel.nhyps ctx) impsty, Some ty
in
(* universe minimization *)
- let evd = Evd.nf_constraints evd in
+ let evd = Evd.minimize_universes evd in
(* Substitute evars and universes, and add parameters.
Note: in program mode some evars may remain. *)
let ctx = List.map (EConstr.to_rel_decl evd) ctx in
diff --git a/vernac/comDefinition.mli b/vernac/comDefinition.mli
index 4a65c1e91..6f81c4575 100644
--- a/vernac/comDefinition.mli
+++ b/vernac/comDefinition.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2018 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Names
@@ -15,7 +17,7 @@ open Constrexpr
(** {6 Definitions/Let} *)
val do_definition : program_mode:bool ->
- Id.t -> definition_kind -> Vernacexpr.universe_decl_expr option ->
+ Id.t -> definition_kind -> universe_decl_expr option ->
local_binder_expr list -> red_expr option -> constr_expr ->
constr_expr option -> unit Lemmas.declaration_hook -> unit
@@ -25,6 +27,6 @@ val do_definition : program_mode:bool ->
(** Not used anywhere. *)
val interp_definition :
- Vernacexpr.universe_decl_expr option -> local_binder_expr list -> polymorphic -> red_expr option -> constr_expr ->
+ universe_decl_expr option -> local_binder_expr list -> polymorphic -> red_expr option -> constr_expr ->
constr_expr option -> Safe_typing.private_constants definition_entry * Evd.evar_map *
Univdecls.universe_decl * Impargs.manual_implicits
diff --git a/vernac/comFixpoint.ml b/vernac/comFixpoint.ml
index 489f299a2..a794c2db0 100644
--- a/vernac/comFixpoint.ml
+++ b/vernac/comFixpoint.ml
@@ -14,7 +14,6 @@ open Pretyping
open Evarutil
open Evarconv
open Misctypes
-open Vernacexpr
module RelDecl = Context.Rel.Declaration
@@ -212,8 +211,7 @@ let interp_recursive ~program_mode ~cofix fixl notations =
let env_rec = push_named_context rec_sign env in
(* Get interpretation metadatas *)
- let fixtypes = List.map EConstr.Unsafe.to_constr fixtypes in
- let impls = compute_internalization_env env Recursive fixnames fixtypes fiximps in
+ let impls = compute_internalization_env env sigma Recursive fixnames fixtypes fiximps in
(* Interp bodies with rollback because temp use of notations/implicit *)
let sigma, fixdefs =
@@ -226,10 +224,9 @@ let interp_recursive ~program_mode ~cofix fixl notations =
(* Instantiate evars and check all are resolved *)
let sigma = solve_unif_constraints_with_heuristics env_rec sigma in
- let sigma, nf = nf_evars_and_universes sigma in
- let fixdefs = List.map (fun c -> Option.map EConstr.Unsafe.to_constr c) fixdefs in
- let fixdefs = List.map (Option.map nf) fixdefs in
- let fixtypes = List.map nf fixtypes in
+ let sigma, _ = nf_evars_and_universes sigma in
+ let fixdefs = List.map (fun c -> Option.map EConstr.(to_constr sigma) c) fixdefs in
+ let fixtypes = List.map EConstr.(to_constr sigma) fixtypes in
let fixctxs = List.map (fun (_,ctx) -> ctx) fixctxs in
(* Build the fix declaration block *)
diff --git a/vernac/comFixpoint.mli b/vernac/comFixpoint.mli
index 2926e30e5..36c2993af 100644
--- a/vernac/comFixpoint.mli
+++ b/vernac/comFixpoint.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2018 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Names
@@ -30,7 +32,7 @@ val do_cofixpoint :
type structured_fixpoint_expr = {
fix_name : Id.t;
- fix_univs : Vernacexpr.universe_decl_expr option;
+ fix_univs : Constrexpr.universe_decl_expr option;
fix_annot : Misctypes.lident option;
fix_binders : local_binder_expr list;
fix_body : constr_expr option;
diff --git a/vernac/comInductive.ml b/vernac/comInductive.ml
index c650e9e40..c59286d1a 100644
--- a/vernac/comInductive.ml
+++ b/vernac/comInductive.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Pp
@@ -11,7 +13,6 @@ open CErrors
open Sorts
open Util
open Constr
-open Termops
open Environ
open Declare
open Names
@@ -51,12 +52,12 @@ let rec complete_conclusion a cs = CAst.map_with_loc (fun ?loc -> function
)
let push_types env idl tl =
- List.fold_left2 (fun env id t -> Environ.push_rel (LocalAssum (Name id,t)) env)
+ List.fold_left2 (fun env id t -> EConstr.push_rel (LocalAssum (Name id,t)) env)
env idl tl
type structured_one_inductive_expr = {
ind_name : Id.t;
- ind_univs : Vernacexpr.universe_decl_expr option;
+ ind_univs : universe_decl_expr option;
ind_arity : constr_expr;
ind_lc : (Id.t * constr_expr) list
}
@@ -90,7 +91,7 @@ let check_all_names_different indl =
| _ -> raise (InductiveError (SameNamesOverlap l))
let mk_mltype_data sigma env assums arity indname =
- let is_ml_type = is_sort env sigma (EConstr.of_constr arity) in
+ let is_ml_type = is_sort env sigma arity in
(is_ml_type,indname,assums)
let prepare_param = function
@@ -130,14 +131,13 @@ let is_impredicative env u =
u = Prop Null || (is_impredicative_set env && u = Prop Pos)
let interp_ind_arity env sigma ind =
- let c = intern_gen IsType env ind.ind_arity in
+ let c = intern_gen IsType env sigma ind.ind_arity in
let impls = Implicit_quantifiers.implicits_of_glob_constr ~with_products:true c in
let sigma,t = understand_tcc env sigma ~expected_type:IsType c in
let pseudo_poly = check_anonymous_type c in
let () = if not (Reductionops.is_arity env sigma t) then
user_err ?loc:(constr_loc ind.ind_arity) (str "Not an arity")
in
- let t = EConstr.Unsafe.to_constr t in
sigma, (t, pseudo_poly, impls)
let interp_cstrs env sigma impls mldata arity ind =
@@ -272,7 +272,6 @@ let interp_mutual_inductive (paramsl,indl) notations cum poly prv finite =
let sigma, (impls, ((env_params, ctx_params), userimpls)) =
interp_context_evars env0 sigma paramsl
in
- let ctx_params = List.map (fun d -> map_rel_decl EConstr.Unsafe.to_constr d) ctx_params in
let indnames = List.map (fun ind -> ind.ind_name) indl in
(* Names of parameters as arguments of the inductive type (defs removed) *)
@@ -282,16 +281,16 @@ let interp_mutual_inductive (paramsl,indl) notations cum poly prv finite =
(* Interpret the arities *)
let sigma, arities = List.fold_left_map (fun sigma -> interp_ind_arity env_params sigma) sigma indl in
- let fullarities = List.map (fun (c, _, _) -> Term.it_mkProd_or_LetIn c ctx_params) arities in
+ let fullarities = List.map (fun (c, _, _) -> EConstr.it_mkProd_or_LetIn c ctx_params) arities in
let env_ar = push_types env0 indnames fullarities in
- let env_ar_params = push_rel_context ctx_params env_ar in
+ let env_ar_params = EConstr.push_rel_context ctx_params env_ar in
(* Compute interpretation metadatas *)
let indimpls = List.map (fun (_, _, impls) -> userimpls @
lift_implicits (Context.Rel.nhyps ctx_params) impls) arities in
let arities = List.map pi1 arities and aritypoly = List.map pi2 arities in
- let impls = compute_internalization_env env0 ~impls (Inductive (params,true)) indnames fullarities indimpls in
- let ntn_impls = compute_internalization_env env0 (Inductive (params,true)) indnames fullarities indimpls in
+ let impls = compute_internalization_env env0 sigma ~impls (Inductive (params,true)) indnames fullarities indimpls in
+ let ntn_impls = compute_internalization_env env0 sigma (Inductive (params,true)) indnames fullarities indimpls in
let mldatas = List.map2 (mk_mltype_data sigma env_params params) arities indnames in
let sigma, constructors =
@@ -306,15 +305,14 @@ let interp_mutual_inductive (paramsl,indl) notations cum poly prv finite =
let sigma = solve_remaining_evars all_and_fail_flags env_params sigma Evd.empty in
(* Compute renewed arities *)
let sigma, nf = nf_evars_and_universes sigma in
- let arities = List.map nf arities in
let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map nf cl,impsl)) constructors in
+ let arities = List.map EConstr.(to_constr sigma) arities in
let sigma = List.fold_left2 (fun sigma ty poly -> make_conclusion_flexible sigma ty poly) sigma arities aritypoly in
let sigma, arities = inductive_levels env_ar_params sigma poly arities constructors in
let sigma, nf' = nf_evars_and_universes sigma in
- let nf x = nf' (nf x) in
let arities = List.map nf' arities in
let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map nf' cl,impsl)) constructors in
- let ctx_params = Context.Rel.map nf ctx_params in
+ let ctx_params = List.map Termops.(map_rel_decl (EConstr.to_constr sigma)) ctx_params in
let uctx = Evd.check_univ_decl ~poly sigma decl in
List.iter (fun c -> check_evars env_params Evd.empty sigma (EConstr.of_constr c)) arities;
Context.Rel.iter (fun c -> check_evars env0 Evd.empty sigma (EConstr.of_constr c)) ctx_params;
diff --git a/vernac/comInductive.mli b/vernac/comInductive.mli
index 82ea131e1..833935724 100644
--- a/vernac/comInductive.mli
+++ b/vernac/comInductive.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Names
@@ -45,7 +47,7 @@ val declare_mutual_inductive_with_eliminations :
type structured_one_inductive_expr = {
ind_name : Id.t;
- ind_univs : Vernacexpr.universe_decl_expr option;
+ ind_univs : universe_decl_expr option;
ind_arity : constr_expr;
ind_lc : (Id.t * constr_expr) list
}
diff --git a/vernac/comProgramFixpoint.ml b/vernac/comProgramFixpoint.ml
index af34f8b29..bd7ee0978 100644
--- a/vernac/comProgramFixpoint.ml
+++ b/vernac/comProgramFixpoint.ml
@@ -171,8 +171,8 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
let sigma, intern_body =
let ctx = LocalAssum (Name recname, get_type curry_fun) :: binders_rel in
let (r, l, impls, scopes) =
- Constrintern.compute_internalization_data env
- Constrintern.Recursive (EConstr.Unsafe.to_constr full_arity) impls
+ Constrintern.compute_internalization_data env sigma
+ Constrintern.Recursive full_arity impls
in
let newimpls = Id.Map.singleton recname
(r, l, impls @ [(Some (Id.of_string "recproof", Impargs.Manual, (true, false)))],
diff --git a/vernac/declareDef.ml b/vernac/declareDef.ml
index dfac78c04..77177dfa4 100644
--- a/vernac/declareDef.ml
+++ b/vernac/declareDef.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Decl_kinds
diff --git a/vernac/declareDef.mli b/vernac/declareDef.mli
index 55f7c7861..010874e23 100644
--- a/vernac/declareDef.mli
+++ b/vernac/declareDef.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Decl_kinds
diff --git a/vernac/explainErr.ml b/vernac/explainErr.ml
index fc3495796..f9167f969 100644
--- a/vernac/explainErr.ml
+++ b/vernac/explainErr.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Pp
diff --git a/vernac/explainErr.mli b/vernac/explainErr.mli
index 0cbd71fa4..b54912a14 100644
--- a/vernac/explainErr.mli
+++ b/vernac/explainErr.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(** Toplevel Exception *)
diff --git a/vernac/himsg.ml b/vernac/himsg.ml
index f00c1e604..c839ed0c7 100644
--- a/vernac/himsg.ml
+++ b/vernac/himsg.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Pp
diff --git a/vernac/himsg.mli b/vernac/himsg.mli
index 8945ebadb..0e20d18c6 100644
--- a/vernac/himsg.mli
+++ b/vernac/himsg.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Indtypes
diff --git a/vernac/indschemes.ml b/vernac/indschemes.ml
index 447c5085b..27587416b 100644
--- a/vernac/indschemes.ml
+++ b/vernac/indschemes.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* Created by Hugo Herbelin from contents related to inductive schemes
@@ -59,13 +61,6 @@ let _ =
optkey = ["Nonrecursive";"Elimination";"Schemes"];
optread = (fun () -> !bifinite_elim_flag) ;
optwrite = (fun b -> bifinite_elim_flag := b) }
-let _ =
- declare_bool_option
- { optdepr = true; (* compatibility 2014-09-03*)
- optname = "automatic declaration of induction schemes for non-recursive types";
- optkey = ["Record";"Elimination";"Schemes"];
- optread = (fun () -> !bifinite_elim_flag) ;
- optwrite = (fun b -> bifinite_elim_flag := b) }
let case_flag = ref false
let _ =
@@ -386,7 +381,7 @@ let do_mutual_induction_scheme lnamedepindsort =
| None ->
let _, ctx = Global.type_of_global_in_context env0 (IndRef ind) in
let u, ctx = Universes.fresh_instance_from ctx None in
- let evd = Evd.from_ctx (Evd.evar_universe_context_of ctx) in
+ let evd = Evd.from_ctx (UState.of_context_set ctx) in
evd, (ind,u), Some u
| Some ui -> evd, (ind, ui), inst
in
diff --git a/vernac/indschemes.mli b/vernac/indschemes.mli
index 8658d85f6..bd4249cac 100644
--- a/vernac/indschemes.mli
+++ b/vernac/indschemes.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Names
diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml
index 7661fff6d..30dd6ec74 100644
--- a/vernac/lemmas.ml
+++ b/vernac/lemmas.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* Created by Hugo Herbelin from contents related to lemma proofs in
diff --git a/vernac/lemmas.mli b/vernac/lemmas.mli
index 126dcd8b0..ad4c278e0 100644
--- a/vernac/lemmas.mli
+++ b/vernac/lemmas.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Names
diff --git a/vernac/locality.ml b/vernac/locality.ml
index 87b411625..21be73b39 100644
--- a/vernac/locality.ml
+++ b/vernac/locality.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Decl_kinds
diff --git a/vernac/locality.mli b/vernac/locality.mli
index 922538b23..3c63c8211 100644
--- a/vernac/locality.mli
+++ b/vernac/locality.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(** * Managing locality *)
diff --git a/vernac/metasyntax.ml b/vernac/metasyntax.ml
index f63206216..a0baca62b 100644
--- a/vernac/metasyntax.ml
+++ b/vernac/metasyntax.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Pp
diff --git a/vernac/metasyntax.mli b/vernac/metasyntax.mli
index 7740604c3..a6c12e089 100644
--- a/vernac/metasyntax.mli
+++ b/vernac/metasyntax.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Names
diff --git a/vernac/mltop.ml b/vernac/mltop.ml
index 053b9d070..343b0925d 100644
--- a/vernac/mltop.ml
+++ b/vernac/mltop.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open CErrors
diff --git a/vernac/mltop.mli b/vernac/mltop.mli
index e44a7c243..da195f4fc 100644
--- a/vernac/mltop.mli
+++ b/vernac/mltop.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(** {5 Toplevel management} *)
diff --git a/vernac/obligations.ml b/vernac/obligations.ml
index e4bcbc4bb..4f16e1cf6 100644
--- a/vernac/obligations.ml
+++ b/vernac/obligations.ml
@@ -472,7 +472,7 @@ let subst_body expand prg =
let declare_definition prg =
let body, typ = subst_body true prg in
let nf = Universes.nf_evars_and_universes_opt_subst (fun x -> None)
- (Evd.evar_universe_context_subst prg.prg_ctx) in
+ (UState.subst prg.prg_ctx) in
let opaque = prg.prg_opaque in
let fix_exn = Hook.get get_fix_exn () in
let typ = nf typ in
@@ -816,13 +816,13 @@ let solve_by_tac name evi t poly ctx =
let id = name in
let concl = EConstr.of_constr evi.evar_concl in
(* spiwack: the status is dropped. *)
- let (entry,_,ctx') = Pfedit.build_constant_by_tactic
+ let (entry,_,ctx') = Pfedit.build_constant_by_tactic
id ~goal_kind:(goal_kind poly) ctx evi.evar_hyps concl (Tacticals.New.tclCOMPLETE t) in
let env = Global.env () in
let entry = Safe_typing.inline_private_constants_in_definition_entry env entry in
let body, () = Future.force entry.const_entry_body in
let ctx' = Evd.merge_context_set ~sideff:true Evd.univ_rigid (Evd.from_ctx ctx') (snd body) in
- Inductiveops.control_only_guard (Global.env ()) (fst body);
+ Inductiveops.control_only_guard env ctx' (EConstr.of_constr (fst body));
(fst body), entry.const_entry_type, Evd.evar_universe_context ctx'
let obligation_terminator name num guard hook auto pf =
@@ -838,7 +838,7 @@ let obligation_terminator name num guard hook auto pf =
let (body, cstr), () = Future.force entry.Entries.const_entry_body in
let sigma = Evd.from_ctx uctx in
let sigma = Evd.merge_context_set ~sideff:true Evd.univ_rigid sigma cstr in
- Inductiveops.control_only_guard (Global.env ()) body;
+ Inductiveops.control_only_guard (Global.env ()) sigma (EConstr.of_constr body);
(** Declare the obligation ourselves and drop the hook *)
let prg = get_info (ProgMap.find name !from_prg) in
(** Ensure universes are substituted properly in body and type *)
diff --git a/vernac/obligations.mli b/vernac/obligations.mli
index 0ec127152..cc2cacd86 100644
--- a/vernac/obligations.mli
+++ b/vernac/obligations.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Environ
diff --git a/vernac/proof_using.ml b/vernac/proof_using.ml
index 8422baf57..f8b085f3e 100644
--- a/vernac/proof_using.ml
+++ b/vernac/proof_using.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Names
diff --git a/vernac/proof_using.mli b/vernac/proof_using.mli
index f63c8e242..7d1110aaa 100644
--- a/vernac/proof_using.mli
+++ b/vernac/proof_using.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(** Utility code for section variables handling in Proof using... *)
diff --git a/vernac/record.ml b/vernac/record.ml
index 1140e3d37..e21f53f55 100644
--- a/vernac/record.ml
+++ b/vernac/record.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Pp
@@ -70,7 +72,7 @@ let interp_fields_evars env sigma impls_env nots l =
let impls =
match i with
| Anonymous -> impls
- | Name id -> Id.Map.add id (compute_internalization_data env Constrintern.Method (EConstr.to_constr sigma t') impl) impls
+ | Name id -> Id.Map.add id (compute_internalization_data env sigma Constrintern.Method t' impl) impls
in
let d = match b' with
| None -> LocalAssum (i,t')
@@ -122,7 +124,7 @@ let typecheck_params_and_fields finite def id poly pl t ps nots fs =
match t with
| { CAst.v = CSort (Misctypes.GType []) } -> true | _ -> false in
let sigma, s = interp_type_evars env sigma ~impls:empty_internalization_env t in
- let sred = Reductionops.whd_all env sigma s in
+ let sred = Reductionops.whd_allnolet env sigma s in
(match EConstr.kind sigma sred with
| Sort s' ->
let s' = EConstr.ESorts.kind sigma s' in
@@ -145,7 +147,7 @@ let typecheck_params_and_fields finite def id poly pl t ps nots fs =
let assums = List.filter is_local_assum newps in
let params = List.map (RelDecl.get_name %> Name.get_id) assums in
let ty = Inductive (params,(finite != Declarations.BiFinite)) in
- let impls_env = compute_internalization_env env0 ~impls:impls_env ty [id] [EConstr.to_constr sigma arity] [imps] in
+ let impls_env = compute_internalization_env env0 sigma ~impls:impls_env ty [id] [arity] [imps] in
let env2,sigma,impls,newfs,data =
interp_fields_evars env_ar sigma impls_env nots (binders_of_decls fs)
in
diff --git a/vernac/record.mli b/vernac/record.mli
index e0a4b8fdd..992da2aa5 100644
--- a/vernac/record.mli
+++ b/vernac/record.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Names
diff --git a/vernac/search.ml b/vernac/search.ml
index 6da6a0c2d..a2a4fb40f 100644
--- a/vernac/search.ml
+++ b/vernac/search.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Pp
diff --git a/vernac/search.mli b/vernac/search.mli
index 2eda3980a..a1fb7ed3e 100644
--- a/vernac/search.mli
+++ b/vernac/search.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Names
diff --git a/vernac/topfmt.ml b/vernac/topfmt.ml
index 1ad7ead72..4e4077f42 100644
--- a/vernac/topfmt.ml
+++ b/vernac/topfmt.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Pp
diff --git a/vernac/topfmt.mli b/vernac/topfmt.mli
index afe76f6f8..2fdefc6fc 100644
--- a/vernac/topfmt.mli
+++ b/vernac/topfmt.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(** Console printing options *)
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index 4613100fc..5eae4fd7f 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* Concrete syntax of the mathematical vernacular MV V2.6 *)
@@ -70,7 +72,7 @@ let show_top_evars () =
let show_universes () =
let pfts = Proof_global.give_me_the_proof () in
let gls,_,_,_,sigma = Proof.proof pfts in
- let ctx = Evd.universe_context_set (Evd.nf_constraints sigma) in
+ let ctx = Evd.universe_context_set (Evd.minimize_universes sigma) in
Termops.pr_evar_universe_context (Evd.evar_universe_context sigma) ++ fnl () ++
str "Normalized constraints: " ++ Univ.pr_universe_context_set (Termops.pr_evd_level sigma) ctx
@@ -803,7 +805,14 @@ let vernac_end_segment ({v=id} as lid) =
(* Libraries *)
+let warn_require_in_section =
+ let name = "require-in-section" in
+ let category = "deprecated" in
+ CWarnings.create ~name ~category
+ (fun () -> strbrk "Use of “Require” inside a section is deprecated.")
+
let vernac_require from import qidl =
+ if Lib.sections_are_opened () then warn_require_in_section ();
let qidl = List.map qualid_of_reference qidl in
let root = match from with
| None -> None
@@ -1025,7 +1034,9 @@ let vernac_arguments ~atts reference args more_implicits nargs_for_red flags =
let sr = smart_global reference in
let inf_names =
let ty, _ = Global.type_of_global_in_context (Global.env ()) sr in
- Impargs.compute_implicits_names (Global.env ()) ty
+ let env = Global.env () in
+ let sigma = Evd.from_env env in
+ Impargs.compute_implicits_names env sigma (EConstr.of_constr ty)
in
let prev_names =
try Arguments_renaming.arguments_names sr with Not_found -> inf_names
@@ -1253,7 +1264,7 @@ let vernac_reserve bl =
let env = Global.env() in
let sigma = Evd.from_env env in
let t,ctx = Constrintern.interp_type env sigma c in
- let t = Detyping.detype Detyping.Now false Id.Set.empty env (Evd.from_ctx ctx) (EConstr.of_constr t) in
+ let t = Detyping.detype Detyping.Now false Id.Set.empty env (Evd.from_ctx ctx) t in
let t,_ = Notation_ops.notation_constr_of_glob_constr (default_env ()) t in
Reserve.declare_reserved_type idl t)
in List.iter sb_decl bl
@@ -1627,6 +1638,7 @@ let vernac_global_check c =
let senv = Global.safe_env() in
let uctx = UState.context_set uctx in
let senv = Safe_typing.push_context_set false uctx senv in
+ let c = EConstr.to_constr sigma c in
let j = Safe_typing.typing senv c in
let env = Safe_typing.env_of_safe_env senv in
print_safe_judgment env sigma j ++
@@ -1747,10 +1759,10 @@ let interp_search_restriction = function
open Search
-let interp_search_about_item env =
+let interp_search_about_item env sigma =
function
| SearchSubPattern pat ->
- let _,pat = intern_constr_pattern env pat in
+ let _,pat = intern_constr_pattern env sigma pat in
GlobSearchSubPattern pat
| SearchString (s,None) when Id.is_valid s ->
GlobSearchString s
@@ -1796,13 +1808,13 @@ let vernac_search ~atts s gopt r =
(* if goal selector is given and wrong, then let exceptions be raised. *)
| Some g -> snd (Pfedit.get_goal_context g) , Some g
in
- let get_pattern c = snd (intern_constr_pattern env c) in
+ let get_pattern c = snd (intern_constr_pattern env Evd.(from_env env) c) in
let pr_search ref env c =
let pr = pr_global ref in
let pp = if !search_output_name_only
then pr
else begin
- let pc = pr_lconstr_env env Evd.empty c in
+ let pc = pr_lconstr_env env Evd.(from_env env) c in
hov 2 (pr ++ str":" ++ spc () ++ pc)
end
in Feedback.msg_notice pp
@@ -1815,7 +1827,8 @@ let vernac_search ~atts s gopt r =
| SearchHead c ->
(Search.search_by_head gopt (get_pattern c) r |> Search.prioritize_search) pr_search
| SearchAbout sl ->
- (Search.search_about gopt (List.map (on_snd (interp_search_about_item env)) sl) r |> Search.prioritize_search) pr_search
+ (Search.search_about gopt (List.map (on_snd (interp_search_about_item env Evd.(from_env env))) sl) r |>
+ Search.prioritize_search) pr_search
let vernac_locate = function
| LocateAny (AN qid) -> print_located_qualid qid
@@ -1910,8 +1923,7 @@ let vernac_check_guard () =
let message =
try
let { Evd.it=gl ; sigma=sigma } = Proof.V82.top_goal pts in
- Inductiveops.control_only_guard (Goal.V82.env sigma gl)
- (EConstr.Unsafe.to_constr pfterm);
+ Inductiveops.control_only_guard (Goal.V82.env sigma gl) sigma pfterm;
(str "The condition holds up to here")
with UserError(_,s) ->
(str ("Condition violated: ") ++s)
@@ -1926,6 +1938,8 @@ exception End_of_input
without a considerable amount of refactoring.
*)
let vernac_load interp fname =
+ if Proof_global.there_are_pending_proofs () then
+ CErrors.user_err Pp.(str "Load is not supported inside proofs.");
let interp x =
let proof_mode = Proof_global.get_default_proof_mode_name () [@ocaml.warning "-3"] in
Proof_global.activate_proof_mode proof_mode [@ocaml.warning "-3"];
@@ -1942,8 +1956,13 @@ let vernac_load interp fname =
let longfname = Loadpath.locate_file fname in
let in_chan = open_utf8_file_in longfname in
Pcoq.Gram.parsable ~file:(Loc.InFile longfname) (Stream.of_channel in_chan) in
- try while true do interp (snd (parse_sentence input)) done
- with End_of_input -> ()
+ begin
+ try while true do interp (snd (parse_sentence input)) done
+ with End_of_input -> ()
+ end;
+ (* If Load left a proof open, we fail too. *)
+ if Proof_global.there_are_pending_proofs () then
+ CErrors.user_err Pp.(str "Files processed by Load cannot leave open proofs.")
(* "locality" is the prefix "Local" attribute, while the "local" component
* is the outdated/deprecated "Local" attribute of some vernacular commands
@@ -1954,6 +1973,7 @@ let interp ?proof ~atts ~st c =
vernac_pperr_endline (fun () -> str "interpreting: " ++ Ppvernac.pr_vernac_expr c);
match c with
+ (* Loading a file requires access to the control interpreter *)
| VernacLoad _ -> assert false
(* The STM should handle that, but LOAD bypasses the STM... *)
diff --git a/vernac/vernacentries.mli b/vernac/vernacentries.mli
index e99a62fe6..13ecaf37b 100644
--- a/vernac/vernacentries.mli
+++ b/vernac/vernacentries.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Misctypes
diff --git a/vernac/vernacinterp.ml b/vernac/vernacinterp.ml
index c40ca27db..1f2d2e4b4 100644
--- a/vernac/vernacinterp.ml
+++ b/vernac/vernacinterp.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Util
diff --git a/vernac/vernacinterp.mli b/vernac/vernacinterp.mli
index c5e610f89..935cacf77 100644
--- a/vernac/vernacinterp.mli
+++ b/vernac/vernacinterp.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(** Interpretation of extended vernac phrases. *)
diff --git a/vernac/vernacprop.ml b/vernac/vernacprop.ml
index 172a20b7a..44a7a9b15 100644
--- a/vernac/vernacprop.ml
+++ b/vernac/vernacprop.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* We define some high-level properties of vernacular commands, used
diff --git a/vernac/vernacprop.mli b/vernac/vernacprop.mli
index df739f96a..8296a039f 100644
--- a/vernac/vernacprop.mli
+++ b/vernac/vernacprop.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* We define some high-level properties of vernacular commands, used
diff --git a/vernac/vernacstate.ml b/vernac/vernacstate.ml
index 4980333b5..aa8bcdc32 100644
--- a/vernac/vernacstate.ml
+++ b/vernac/vernacstate.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
type t = {
diff --git a/vernac/vernacstate.mli b/vernac/vernacstate.mli
index 3ed27ddb7..b4d478d12 100644
--- a/vernac/vernacstate.mli
+++ b/vernac/vernacstate.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
type t = {