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.ml54
-rw-r--r--vernac/classes.mli10
-rw-r--r--vernac/comAssumption.ml30
-rw-r--r--vernac/comAssumption.mli14
-rw-r--r--vernac/comDefinition.ml12
-rw-r--r--vernac/comDefinition.mli14
-rw-r--r--vernac/comFixpoint.ml19
-rw-r--r--vernac/comFixpoint.mli14
-rw-r--r--vernac/comInductive.ml44
-rw-r--r--vernac/comInductive.mli12
-rw-r--r--vernac/comProgramFixpoint.ml12
-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.ml43
-rw-r--r--vernac/indschemes.mli17
-rw-r--r--vernac/lemmas.ml16
-rw-r--r--vernac/lemmas.mli10
-rw-r--r--vernac/locality.ml10
-rw-r--r--vernac/locality.mli10
-rw-r--r--vernac/metasyntax.ml32
-rw-r--r--vernac/metasyntax.mli11
-rw-r--r--vernac/mltop.ml10
-rw-r--r--vernac/mltop.mli10
-rw-r--r--vernac/obligations.ml14
-rw-r--r--vernac/obligations.mli14
-rw-r--r--vernac/proof_using.ml12
-rw-r--r--vernac/proof_using.mli10
-rw-r--r--vernac/record.ml33
-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.ml123
-rw-r--r--vernac/vernacentries.mli10
-rw-r--r--vernac/vernacinterp.ml10
-rw-r--r--vernac/vernacinterp.mli10
-rw-r--r--vernac/vernacprop.ml22
-rw-r--r--vernac/vernacprop.mli10
-rw-r--r--vernac/vernacstate.ml10
-rw-r--r--vernac/vernacstate.mli10
51 files changed, 495 insertions, 375 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 4a2dba859..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]
@@ -132,7 +134,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance)
~program_mode poly ctx (instid, bk, cl) props ?(generalize=true)
?(tac:unit Proofview.tactic option) ?hook pri =
let env = Global.env() in
- let ((loc, instid), pl) = instid in
+ let ({CAst.loc;v=instid}, pl) = instid in
let sigma, decl = Univdecls.interp_univ_decl_opt env pl in
let tclass, ids =
match bk with
@@ -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,14 +423,10 @@ 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
- Vernacexpr.NoInline (Loc.tag id))
+ Vernacexpr.NoInline (CAst.make id))
| Some b ->
let decl = (Discharge, poly, Definition) in
let entry = Declare.definition_entry ~univs ~types:t b in
@@ -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 5d7adb24a..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
@@ -42,7 +43,7 @@ let should_axiom_into_instance = function
true
| Global | Local -> !axiom_into_instance
-let declare_assumption is_coe (local,p,kind) (c,ctx) pl imps impl nl (_,ident) =
+let declare_assumption is_coe (local,p,kind) (c,ctx) pl imps impl nl {CAst.v=ident} =
match local with
| Discharge when Lib.sections_are_opened () ->
let ctx = match ctx with
@@ -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. *)
@@ -109,7 +109,7 @@ let declare_assumptions idl is_coe k (c,uctx) pl imps nl =
let maybe_error_many_udecls = function
- | ((loc,id), Some _) ->
+ | ({CAst.loc;v=id}, Some _) ->
user_err ?loc ~hdr:"many_universe_declarations"
Pp.(str "When declaring multiple axioms in one command, " ++
str "only the first is allowed a universe binder " ++
@@ -126,7 +126,7 @@ let process_assumptions_udecls kind l =
in
let () = match kind, udecl with
| (Discharge, _, _), Some _ when Lib.sections_are_opened () ->
- let loc = fst first_id in
+ let loc = first_id.CAst.loc in
let msg = Pp.str "Section variables cannot be polymorphic." in
user_err ?loc msg
| _ -> ()
@@ -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 (_,id) -> LocalAssum (id,t)) idl) env in
- let ienv = List.fold_right (fun (_,id) ienv ->
- let impls = compute_internalization_data env Variable t imps 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 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
@@ -175,7 +175,7 @@ let do_assumptions kind nl l =
let t = replace_vars subst t in
let refs, status' = declare_assumptions idl is_coe kind (t,uctx) ubinders imps nl in
let subst' = List.map2
- (fun (_,id) (c,u) -> (id, Universes.constr_of_global_univ (c,u)))
+ (fun {CAst.v=id} (c,u) -> (id, Universes.constr_of_global_univ (c,u)))
idl refs
in
subst'@subst, status' && status, next_uctx uctx)
diff --git a/vernac/comAssumption.mli b/vernac/comAssumption.mli
index 2fa156abd..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 *)
@@ -30,5 +32,5 @@ val do_assumptions : locality * polymorphic * assumption_object_kind ->
val declare_assumption : coercion_flag -> assumption_kind ->
types in_constant_universes_entry ->
Universes.universe_binders -> Impargs.manual_implicits ->
- bool (** implicit *) -> Vernacexpr.inline -> variable Loc.located ->
+ bool (** implicit *) -> Vernacexpr.inline -> variable CAst.t ->
global_reference * Univ.Instance.t * bool
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 d648e293a..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
@@ -102,7 +101,7 @@ let check_mutuality env evd isfix fixl =
type structured_fixpoint_expr = {
fix_name : Id.t;
fix_univs : universe_decl_expr option;
- fix_annot : Id.t Loc.located option;
+ fix_annot : lident option;
fix_binders : local_binder_expr list;
fix_body : constr_expr option;
fix_type : constr_expr
@@ -175,7 +174,7 @@ let interp_recursive ~program_mode ~cofix fixl notations =
| x , None -> x
| Some ls , Some us ->
let lsu = ls.univdecl_instance and usu = us.univdecl_instance in
- if not (CList.for_all2eq (fun x y -> Id.equal (snd x) (snd y)) lsu usu) then
+ if not (CList.for_all2eq (fun x y -> Id.equal x.CAst.v y.CAst.v) lsu usu) then
user_err Pp.(str "(co)-recursive definitions should all have the same universe binders");
Some us) fixl None in
let sigma, decl = Univdecls.interp_univ_decl_opt env all_universes in
@@ -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 *)
@@ -323,7 +320,7 @@ let extract_decreasing_argument limit = function
let extract_fixpoint_components limit l =
let fixl, ntnl = List.split l in
- let fixl = List.map (fun (((_,id),pl),ann,bl,typ,def) ->
+ let fixl = List.map (fun (({CAst.v=id},pl),ann,bl,typ,def) ->
let ann = extract_decreasing_argument limit ann in
{fix_name = id; fix_annot = ann; fix_univs = pl;
fix_binders = bl; fix_body = def; fix_type = typ}) fixl in
@@ -331,7 +328,7 @@ let extract_fixpoint_components limit l =
let extract_cofixpoint_components l =
let fixl, ntnl = List.split l in
- List.map (fun (((_,id),pl),bl,typ,def) ->
+ List.map (fun (({CAst.v=id},pl),bl,typ,def) ->
{fix_name = id; fix_annot = None; fix_univs = pl;
fix_binders = bl; fix_body = def; fix_type = typ}) fixl,
List.flatten ntnl
diff --git a/vernac/comFixpoint.mli b/vernac/comFixpoint.mli
index 2c84bd84d..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,8 +32,8 @@ val do_cofixpoint :
type structured_fixpoint_expr = {
fix_name : Id.t;
- fix_univs : Vernacexpr.universe_decl_expr option;
- fix_annot : Id.t Loc.located option;
+ fix_univs : Constrexpr.universe_decl_expr option;
+ fix_annot : Misctypes.lident option;
fix_binders : local_binder_expr list;
fix_body : constr_expr option;
fix_type : constr_expr
diff --git a/vernac/comInductive.ml b/vernac/comInductive.ml
index 41474fc63..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 =
@@ -249,7 +249,7 @@ let inductive_levels env evd poly arities inds =
(evd,[]) (Array.to_list levels') destarities sizes
in evd, List.rev arities
-let check_named (loc, na) = match na with
+let check_named {CAst.loc;v=na} = match na with
| Name _ -> ()
| Anonymous ->
let msg = str "Parameters must be named." in
@@ -260,7 +260,7 @@ let check_param = function
| CLocalDef (na, _, _) -> check_named na
| CLocalAssum (nas, Default _, _) -> List.iter check_named nas
| CLocalAssum (nas, Generalized _, _) -> ()
-| CLocalPattern (loc,_) ->
+| CLocalPattern {CAst.loc} ->
Loc.raise ?loc (Stream.Error "pattern with quote not allowed here.")
let interp_mutual_inductive (paramsl,indl) notations cum poly prv finite =
@@ -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;
@@ -364,7 +362,7 @@ let eq_local_binders bl1 bl2 =
List.equal local_binder_eq bl1 bl2
let extract_coercions indl =
- let mkqid (_,((_,id),_)) = qualid_of_ident id in
+ let mkqid (_,({CAst.v=id},_)) = qualid_of_ident id in
let extract lc = List.filter (fun (iscoe,_) -> iscoe) lc in
List.map mkqid (List.flatten(List.map (fun (_,_,_,lc) -> extract lc) indl))
@@ -378,10 +376,10 @@ let extract_params indl =
params
let extract_inductive indl =
- List.map (fun (((_,indname),pl),_,ar,lc) -> {
+ List.map (fun (({CAst.v=indname},pl),_,ar,lc) -> {
ind_name = indname; ind_univs = pl;
ind_arity = Option.cata (fun x -> x) (CAst.make @@ CSort (GType [])) ar;
- ind_lc = List.map (fun (_,((_,id),t)) -> (id,t)) lc
+ ind_lc = List.map (fun (_,({CAst.v=id},t)) -> (id,t)) lc
}) indl
let extract_mutual_inductive_declaration_components indl =
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 a9a91e304..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)))],
@@ -298,16 +298,16 @@ let do_program_recursive local poly fixkind fixl ntns =
let do_program_fixpoint local poly l =
let g = List.map (fun ((_,wf,_,_,_),_) -> wf) l in
match g, l with
- | [(n, CWfRec r)], [((((_,id),pl),_,bl,typ,def),ntn)] ->
+ | [(n, CWfRec r)], [((({CAst.v=id},pl),_,bl,typ,def),ntn)] ->
let recarg =
match n with
- | Some n -> mkIdentC (snd n)
+ | Some n -> mkIdentC n.CAst.v
| None ->
user_err ~hdr:"do_program_fixpoint"
(str "Recursive argument required for well-founded fixpoints")
in build_wellfounded (id, pl, n, bl, typ, out_def def) poly r recarg ntn
- | [(n, CMeasureRec (m, r))], [((((_,id),pl),_,bl,typ,def),ntn)] ->
+ | [(n, CMeasureRec (m, r))], [((({CAst.v=id},pl),_,bl,typ,def),ntn)] ->
build_wellfounded (id, pl, n, bl, typ, out_def def) poly
(Option.default (CAst.make @@ CRef (lt_ref,None)) r) m ntn
@@ -322,7 +322,7 @@ let do_program_fixpoint local poly l =
let extract_cofixpoint_components l =
let fixl, ntnl = List.split l in
- List.map (fun (((_,id),pl),bl,typ,def) ->
+ List.map (fun (({CAst.v=id},pl),bl,typ,def) ->
{fix_name = id; fix_annot = None; fix_univs = pl;
fix_binders = bl; fix_body = def; fix_type = typ}) fixl,
List.flatten ntnl
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 2f4c95aff..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 _ =
@@ -367,17 +362,16 @@ requested
| InType -> recs ^ "t_nodep")
) in
let newid = add_suffix (basename_of_global (IndRef ind)) suffix in
- let newref = Loc.tag newid in
+ let newref = CAst.make newid in
((newref,isdep,ind,z)::l1),l2
in
match t with
| CaseScheme (x,y,z) -> names "_case" "_case" x y z
| InductionScheme (x,y,z) -> names "_ind" "_rec" x y z
| EqualityScheme x -> l1,((None,smart_global_inductive x)::l2)
-
let do_mutual_induction_scheme lnamedepindsort =
- let lrecnames = List.map (fun ((_,f),_,_,_) -> f) lnamedepindsort
+ let lrecnames = List.map (fun ({CAst.v},_,_,_) -> v) lnamedepindsort
and env0 = Global.env() in
let sigma, lrecspec, _ =
List.fold_right
@@ -387,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
@@ -416,7 +410,7 @@ let get_common_underlying_mutual_inductive = function
then user_err Pp.(str "A type occurs twice");
mind,
List.map_filter
- (function (Some id,(_,i)) -> Some (i,snd id) | (None,_) -> None) all
+ (function (Some id,(_,i)) -> Some (i,id.CAst.v) | (None,_) -> None) all
let do_scheme l =
let ischeme,escheme = split_scheme l in
@@ -450,7 +444,7 @@ let fold_left' f = function
let mk_coq_and sigma = Evarutil.new_global sigma (Coqlib.build_coq_and ())
let mk_coq_conj sigma = Evarutil.new_global sigma (Coqlib.build_coq_conj ())
-
+
let build_combined_scheme env schemes =
let evdref = ref (Evd.from_env env) in
let defs = List.map (fun cst ->
@@ -492,18 +486,19 @@ let build_combined_scheme env schemes =
(!evdref, body, typ)
let do_combined_scheme name schemes =
+ let open CAst in
let csts =
- List.map (fun x ->
- let refe = Ident x in
- let qualid = qualid_of_reference refe in
- try Nametab.locate_constant (snd qualid)
- with Not_found -> user_err Pp.(pr_qualid (snd qualid) ++ str " is not declared."))
+ List.map (fun {CAst.loc;v} ->
+ let refe = Ident (Loc.tag ?loc v) in
+ let qualid = qualid_of_reference refe in
+ try Nametab.locate_constant (snd qualid)
+ with Not_found -> user_err Pp.(pr_qualid (snd qualid) ++ str " is not declared."))
schemes
in
let sigma,body,typ = build_combined_scheme (Global.env ()) csts in
let proof_output = Future.from_val ((body,Univ.ContextSet.empty),Safe_typing.empty_private_constants) in
- ignore (define (snd name) UserIndividualRequest sigma proof_output (Some typ));
- fixpoint_message None [snd name]
+ ignore (define name.v UserIndividualRequest sigma proof_output (Some typ));
+ fixpoint_message None [name.v]
(**********************************************************************)
diff --git a/vernac/indschemes.mli b/vernac/indschemes.mli
index 4b31389ab..bd4249cac 100644
--- a/vernac/indschemes.mli
+++ b/vernac/indschemes.mli
@@ -1,12 +1,13 @@
(************************************************************************)
-(* 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 Loc
open Names
open Constr
open Environ
@@ -31,17 +32,17 @@ val declare_rewriting_schemes : inductive -> unit
(** Mutual Minimality/Induction scheme *)
val do_mutual_induction_scheme :
- (Id.t located * bool * inductive * Sorts.family) list -> unit
+ (Misctypes.lident * bool * inductive * Sorts.family) list -> unit
(** Main calls to interpret the Scheme command *)
-val do_scheme : (Id.t located option * scheme) list -> unit
+val do_scheme : (Misctypes.lident option * scheme) list -> unit
(** Combine a list of schemes into a conjunction of them *)
val build_combined_scheme : env -> Constant.t list -> Evd.evar_map * constr * types
-val do_combined_scheme : Id.t located -> Id.t located list -> unit
+val do_combined_scheme : Misctypes.lident -> Misctypes.lident list -> unit
(** Hook called at each inductive type definition *)
diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml
index 57436784b..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
@@ -214,7 +216,7 @@ let fresh_name_for_anonymous_theorem () =
let avoid = Id.Set.of_list (Proof_global.get_all_proof_names ()) in
next_global_ident_away default_thm_id avoid
-let check_name_freshness locality (loc,id) : unit =
+let check_name_freshness locality {CAst.loc;v=id} : unit =
(* We check existence here: it's a bit late at Qed time *)
if Nametab.exists_cci (Lib.make_path id) || is_section_variable id ||
locality == Global && Nametab.exists_cci (Lib.make_path_except_section id)
@@ -339,7 +341,7 @@ let universe_proof_terminator compute_guard hook =
(hook (Some (proof.Proof_global.universes))) is_opaque in
begin match idopt with
| None -> save_named ~export_seff proof
- | Some (_,id) -> save_anonymous ~export_seff proof id
+ | Some { CAst.v = id } -> save_anonymous ~export_seff proof id
end
end
@@ -444,7 +446,7 @@ let start_proof_com ?inference_hook kind thms hook =
let ids = List.map RelDecl.get_name ctx in
check_name_freshness (pi1 kind) id;
(* XXX: The nf_evar is critical !! *)
- evd, (snd id,
+ evd, (id.CAst.v,
(Evarutil.nf_evar evd (EConstr.it_mkProd_or_LetIn t' ctx),
(ids, imps @ lift_implicits (Context.Rel.nhyps ctx) imps'))))
evd thms 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 bcffe640c..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
@@ -89,7 +91,7 @@ let pr_grammar = function
(* Parse a format (every terminal starting with a letter or a single
quote (except a single quote alone) must be quoted) *)
-let parse_format ((loc, str) : lstring) =
+let parse_format ({CAst.loc;v=str} : Misctypes.lstring) =
let len = String.length str in
(* TODO: update the line of the location when the string contains newlines *)
let make_loc i j = Option.map (Loc.shift_loc (i+1) (j-len)) loc in
@@ -789,7 +791,7 @@ type notation_modifier = {
only_parsing : bool;
only_printing : bool;
compat : Flags.compat_version option;
- format : string Loc.located option;
+ format : Misctypes.lstring option;
extra : (string * string) list;
}
@@ -847,7 +849,7 @@ let interp_modifiers modl = let open NotationMods in
| SetFormat ("text",s) :: l ->
if not (Option.is_empty acc.format) then user_err Pp.(str "A format is given more than once.");
interp { acc with format = Some s; } l
- | SetFormat (k,(_,s)) :: l ->
+ | SetFormat (k,{CAst.v=s}) :: l ->
interp { acc with extra = (k,s)::acc.extra; } l
in interp default modl
@@ -1101,7 +1103,7 @@ module SynData = struct
only_parsing : bool;
only_printing : bool;
compat : Flags.compat_version option;
- format : string Loc.located option;
+ format : Misctypes.lstring option;
extra : (string * string) list;
(* XXX: Callback to printing, must remove *)
@@ -1400,7 +1402,7 @@ let add_notation_interpretation_core local df env ?(impls=empty_internalization_
(* Notations without interpretation (Reserved Notation) *)
-let add_syntax_extension local ((loc,df),mods) = let open SynData in
+let add_syntax_extension local ({CAst.loc;v=df},mods) = let open SynData in
let psd = compute_pure_syntax_data df mods in
let sy_rules = make_syntax_rules {psd with compat = None} in
Flags.if_verbose (List.iter (fun (f,x) -> f x)) psd.msgs;
@@ -1408,11 +1410,11 @@ let add_syntax_extension local ((loc,df),mods) = let open SynData in
(* Notations with only interpretation *)
-let add_notation_interpretation env ((loc,df),c,sc) =
+let add_notation_interpretation env ({CAst.loc;v=df},c,sc) =
let df' = add_notation_interpretation_core false df env c sc false false None in
Dumpglob.dump_notation (loc,df') sc true
-let set_notation_for_interpretation env impls ((_,df),c,sc) =
+let set_notation_for_interpretation env impls ({CAst.v=df},c,sc) =
(try ignore
(Flags.silently (fun () -> add_notation_interpretation_core false df env ~impls c sc false false None) ());
with NoSyntaxRule ->
@@ -1421,7 +1423,7 @@ let set_notation_for_interpretation env impls ((_,df),c,sc) =
(* Main entry point *)
-let add_notation local env c ((loc,df),modifiers) sc =
+let add_notation local env c ({CAst.loc;v=df},modifiers) sc =
let df' =
if no_syntax_modifiers modifiers then
(* No syntax data: try to rely on a previously declared rule *)
@@ -1448,13 +1450,13 @@ let add_notation_extra_printing_rule df k v =
let inject_var x = CAst.make @@ CRef (Ident (Loc.tag @@ Id.of_string x),None)
-let add_infix local env ((loc,inf),modifiers) pr sc =
+let add_infix local env ({CAst.loc;v=inf},modifiers) pr sc =
check_infix_modifiers modifiers;
(* check the precedence *)
let metas = [inject_var "x"; inject_var "y"] in
let c = mkAppC (pr,metas) in
- let df = "x "^(quote_notation_token inf)^" y" in
- add_notation local env c ((loc,df),modifiers) sc
+ let df = CAst.make ?loc @@ "x "^(quote_notation_token inf)^" y" in
+ add_notation local env c (df,modifiers) sc
(**********************************************************************)
(* Delimiters and classes bound to scopes *)
diff --git a/vernac/metasyntax.mli b/vernac/metasyntax.mli
index 9137f7a7e..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
@@ -12,6 +14,7 @@ open Notation
open Constrexpr
open Notation_term
open Environ
+open Misctypes
val add_token_obj : string -> unit
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 58e4b00fc..4f16e1cf6 100644
--- a/vernac/obligations.ml
+++ b/vernac/obligations.ml
@@ -295,10 +295,10 @@ type obligation =
type obligations = (obligation array * int)
type fixpoint_kind =
- | IsFixpoint of (Id.t Loc.located option * Constrexpr.recursion_order_expr) list
+ | IsFixpoint of (Misctypes.lident option * Constrexpr.recursion_order_expr) list
| IsCoFixpoint
-type notations = (Vernacexpr.lstring * Constrexpr.constr_expr * Notation_term.scope_name option) list
+type notations = (Misctypes.lstring * Constrexpr.constr_expr * Notation_term.scope_name option) list
type program_info_aux = {
prg_name: Id.t;
@@ -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
@@ -500,7 +500,7 @@ let rec lam_index n t acc =
let compute_possible_guardness_evidences (n,_) fixbody fixtype =
match n with
- | Some (loc, n) -> [lam_index n fixbody 0]
+ | Some { CAst.loc; v = n } -> [lam_index n fixbody 0]
| None ->
(* If recursive argument was not given by user, we try all args.
An earlier approach was to look only for inductive arguments,
@@ -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 bdc97d48c..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
@@ -61,10 +63,10 @@ val add_definition : Names.Id.t -> ?term:constr -> types ->
?hook:(UState.t -> unit) Lemmas.declaration_hook -> ?opaque:bool -> obligation_info -> progress
type notations =
- (Vernacexpr.lstring * Constrexpr.constr_expr * Notation_term.scope_name option) list
+ (Misctypes.lstring * Constrexpr.constr_expr * Notation_term.scope_name option) list
type fixpoint_kind =
- | IsFixpoint of (Id.t Loc.located option * Constrexpr.recursion_order_expr) list
+ | IsFixpoint of (Misctypes.lident option * Constrexpr.recursion_order_expr) list
| IsCoFixpoint
val add_mutual_definitions :
diff --git a/vernac/proof_using.ml b/vernac/proof_using.ml
index ffe99cfd8..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
@@ -51,7 +53,7 @@ let rec process_expr env e ty =
let rec aux = function
| SsEmpty -> Id.Set.empty
| SsType -> set_of_type env ty
- | SsSingl (_,id) -> set_of_id env id
+ | SsSingl { CAst.v = id } -> set_of_id env id
| SsUnion(e1,e2) -> Id.Set.union (aux e1) (aux e2)
| SsSubstr(e1,e2) -> Id.Set.diff (aux e1) (aux e2)
| SsCompl e -> Id.Set.diff (full_set env) (aux e)
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 16b95a5ef..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
@@ -62,7 +64,7 @@ let _ =
let interp_fields_evars env sigma impls_env nots l =
List.fold_left2
- (fun (env, sigma, uimpls, params, impls) no ((loc, i), b, t) ->
+ (fun (env, sigma, uimpls, params, impls) no ({CAst.loc;v=i}, b, t) ->
let sigma, (t', impl) = interp_type_evars_impls env sigma ~impls t in
let sigma, b' =
Option.cata (fun x -> on_snd (fun x -> Some (fst x)) @@
@@ -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')
@@ -92,8 +94,9 @@ let compute_constructor_level evars env l =
let binder_of_decl = function
| Vernacexpr.AssumExpr(n,t) -> (n,None,t)
- | Vernacexpr.DefExpr(n,c,t) -> (n,Some c, match t with Some c -> c
- | None -> CAst.make ?loc:(fst n) @@ CHole (None, Misctypes.IntroAnonymous, None))
+ | Vernacexpr.DefExpr(n,c,t) ->
+ (n,Some c, match t with Some c -> c
+ | None -> CAst.make ?loc:n.CAst.loc @@ CHole (None, Misctypes.IntroAnonymous, None))
let binders_of_decls = List.map binder_of_decl
@@ -101,7 +104,7 @@ let typecheck_params_and_fields finite def id poly pl t ps nots fs =
let env0 = Global.env () in
let sigma, decl = Univdecls.interp_univ_decl_opt env0 pl in
let _ =
- let error bk (loc, name) =
+ let error bk {CAst.loc; v=name} =
match bk, name with
| Default _, Anonymous ->
user_err ?loc ~hdr:"record" (str "Record parameters must be named")
@@ -110,7 +113,7 @@ let typecheck_params_and_fields finite def id poly pl t ps nots fs =
List.iter
(function CLocalDef (b, _, _) -> error default_binder_kind b
| CLocalAssum (ls, bk, ce) -> List.iter (error bk) ls
- | CLocalPattern (loc,(_,_)) ->
+ | CLocalPattern {CAst.loc} ->
Loc.raise ?loc (Stream.Error "pattern with quote not allowed in record parameters.")) ps
in
let sigma, (impls_env, ((env1,newps), imps)) = interp_context_evars env0 sigma ps in
@@ -121,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
@@ -144,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
@@ -571,13 +574,13 @@ 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,cum,poly,finite,(is_coe,((loc,idstruc),pl)),ps,cfs,idbuild,s) =
+let definition_structure (kind,cum,poly,finite,(is_coe,({CAst.loc;v=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
let extract_name acc = function
- Vernacexpr.AssumExpr((_,Name id),_) -> id::acc
- | Vernacexpr.DefExpr ((_,Name id),_,_) -> id::acc
+ Vernacexpr.AssumExpr({CAst.v=Name id},_) -> id::acc
+ | Vernacexpr.DefExpr ({CAst.v=Name id},_,_) -> id::acc
| _ -> acc in
let allnames = idstruc::(List.fold_left extract_name [] fs) in
let () = match List.duplicates Id.equal allnames with
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 bab5c5ae9..4c9b41b21 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -1,15 +1,18 @@
(************************************************************************)
-(* 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 *)
open Pp
open CErrors
+open CAst
open Util
open Names
open Nameops
@@ -69,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
@@ -471,13 +474,13 @@ let vernac_definition_hook p = function
| SubClass -> Class.add_subclass_hook p
| _ -> no_hook
-let vernac_definition ~atts discharge kind ((loc, id), pl) def =
+let vernac_definition ~atts discharge kind ({loc;v=id}, pl) def =
let local = enforce_locality_exp atts.locality discharge in
let hook = vernac_definition_hook atts.polymorphic kind in
let () =
match id with
| Anonymous -> ()
- | Name n -> let lid = (loc, n) in
+ | Name n -> let lid = CAst.make ?loc n in
match local with
| Discharge -> Dumpglob.dump_definition lid true "var"
| Local | Global -> Dumpglob.dump_definition lid false "def"
@@ -491,7 +494,7 @@ let vernac_definition ~atts discharge kind ((loc, id), pl) def =
(match def with
| ProveBody (bl,t) -> (* local binders, typ *)
start_proof_and_print (local, atts.polymorphic, DefinitionBody kind)
- [((loc, name), pl), (bl, t)] hook
+ [(CAst.make ?loc name, pl), (bl, t)] hook
| DefineBody (bl,red_option,c,typ_opt) ->
let red_option = match red_option with
| None -> None
@@ -546,14 +549,14 @@ let should_treat_as_cumulative cum poly =
let vernac_record cum k poly finite struc binders sort nameopt cfs =
let is_cumulative = should_treat_as_cumulative cum poly in
let const = match nameopt with
- | None -> add_prefix "Build_" (snd (fst (snd struc)))
- | Some (_,id as lid) ->
+ | None -> add_prefix "Build_" (fst (snd struc)).v
+ | Some ({v=id} as lid) ->
Dumpglob.dump_definition lid false "constr"; id in
if Dumpglob.dump () then (
Dumpglob.dump_definition (fst (snd struc)) false "rec";
List.iter (fun (((_, x), _), _) ->
match x with
- | Vernacexpr.AssumExpr ((loc, Name id), _) -> Dumpglob.dump_definition (loc,id) false "proj"
+ | Vernacexpr.AssumExpr ({loc;v=Name id}, _) -> Dumpglob.dump_definition (make ?loc id) false "proj"
| _ -> ()) cfs);
ignore(Record.definition_structure (k,is_cumulative,poly,finite,struc,binders,cfs,const,sort))
@@ -582,9 +585,9 @@ let vernac_inductive ~atts cum lo finite indl =
atts.polymorphic finite id bl c oc fs
| [ ( id , bl , c , Class _, Constructors [l]), [] ] ->
let f =
- let (coe, ((loc, id), ce)) = l in
+ let (coe, ({loc;v=id}, ce)) = l in
let coe' = if coe then Some true else None in
- (((coe', AssumExpr ((loc, Name id), ce)), None), [])
+ (((coe', AssumExpr ((make ?loc @@ Name id), ce)), None), [])
in vernac_record cum (Class true) atts.polymorphic finite id bl c None [f]
| [ ( _ , _, _, Class _, Constructors _), [] ] ->
user_err Pp.(str "Inductive classes not supported")
@@ -637,7 +640,7 @@ let vernac_scheme l =
let vernac_combined_scheme lid l =
if Dumpglob.dump () then
(Dumpglob.dump_definition lid false "def";
- List.iter (fun lid -> dump_global (Misctypes.AN (Ident lid))) l);
+ List.iter (fun {loc;v=id} -> dump_global (Misctypes.AN (Ident (Loc.tag ?loc id)))) l);
Indschemes.do_combined_scheme lid l
let vernac_universe ~atts l =
@@ -660,7 +663,7 @@ let vernac_constraint ~atts l =
let vernac_import export refl =
Library.import_module export (List.map qualid_of_reference refl)
-let vernac_declare_module export (loc, id) binders_ast mty_ast =
+let vernac_declare_module export {loc;v=id} binders_ast mty_ast =
(* We check the state of the system (in section, in module type)
and what module information is supplied *)
if Lib.sections_are_opened () then
@@ -678,7 +681,7 @@ let vernac_declare_module export (loc, id) binders_ast mty_ast =
Flags.if_verbose Feedback.msg_info (str "Module " ++ Id.print id ++ str " is declared");
Option.iter (fun export -> vernac_import export [Ident (Loc.tag id)]) export
-let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l =
+let vernac_define_module export {loc;v=id} (binders_ast : module_binder list) mty_ast_o mexpr_ast_l =
(* We check the state of the system (in section, in module type)
and what module information is supplied *)
if Lib.sections_are_opened () then
@@ -689,7 +692,7 @@ let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l =
let binders_ast,argsexport =
List.fold_right
(fun (export,idl,ty) (args,argsexport) ->
- (idl,ty)::args, (List.map (fun (_,i) -> export,i)idl)@argsexport) binders_ast
+ (idl,ty)::args, (List.map (fun {v=i} -> export,i)idl)@argsexport) binders_ast
([],[]) in
let mp =
Declaremods.start_module Modintern.interp_module_ast
@@ -719,13 +722,13 @@ let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l =
Option.iter (fun export -> vernac_import export [Ident (Loc.tag id)])
export
-let vernac_end_module export (loc,id as lid) =
+let vernac_end_module export {loc;v=id} =
let mp = Declaremods.end_module () in
Dumpglob.dump_modref ?loc mp "mod";
Flags.if_verbose Feedback.msg_info (str "Module " ++ Id.print id ++ str " is defined");
- Option.iter (fun export -> vernac_import export [Ident lid]) export
+ Option.iter (fun export -> vernac_import export [Ident (Loc.tag ?loc id)]) export
-let vernac_declare_module_type (loc,id) binders_ast mty_sign mty_ast_l =
+let vernac_declare_module_type {loc;v=id} binders_ast mty_sign mty_ast_l =
if Lib.sections_are_opened () then
user_err Pp.(str "Modules and Module Types are not allowed inside sections.");
@@ -735,7 +738,7 @@ let vernac_declare_module_type (loc,id) binders_ast mty_sign mty_ast_l =
let binders_ast,argsexport =
List.fold_right
(fun (export,idl,ty) (args,argsexport) ->
- (idl,ty)::args, (List.map (fun (_,i) -> export,i)idl)@argsexport) binders_ast
+ (idl,ty)::args, (List.map (fun {v=i} -> export,i)idl)@argsexport) binders_ast
([],[]) in
let mp =
@@ -765,7 +768,7 @@ let vernac_declare_module_type (loc,id) binders_ast mty_sign mty_ast_l =
Flags.if_verbose Feedback.msg_info
(str "Module Type " ++ Id.print id ++ str " is defined")
-let vernac_end_modtype (loc,id) =
+let vernac_end_modtype {loc;v=id} =
let mp = Declaremods.end_modtype () in
Dumpglob.dump_modref ?loc mp "modtype";
Flags.if_verbose Feedback.msg_info (str "Module Type " ++ Id.print id ++ str " is defined")
@@ -778,21 +781,21 @@ let vernac_include l =
(* Sections *)
-let vernac_begin_section (_, id as lid) =
+let vernac_begin_section ({v=id} as lid) =
Proof_global.check_no_pending_proof ();
Dumpglob.dump_definition lid true "sec";
Lib.open_section id
-let vernac_end_section (loc,_) =
+let vernac_end_section {CAst.loc} =
Dumpglob.dump_reference ?loc
(DirPath.to_string (Lib.current_dirpath true)) "<>" "sec";
Lib.close_section ()
-let vernac_name_sec_hyp (_,id) set = Proof_using.name_set id set
+let vernac_name_sec_hyp {v=id} set = Proof_using.name_set id set
(* Dispatcher of the "End" command *)
-let vernac_end_segment (_,id as lid) =
+let vernac_end_segment ({v=id} as lid) =
Proof_global.check_no_pending_proof ();
match Lib.find_opening_node id with
| Lib.OpenedModule (false,export,_,_) -> vernac_end_module export lid
@@ -802,7 +805,14 @@ let vernac_end_segment (_,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
@@ -975,7 +985,7 @@ let vernac_hints ~atts lb h =
let vernac_syntactic_definition ~atts lid x y =
Dumpglob.dump_definition lid false "syndef";
let local = enforce_module_locality atts.locality in
- Metasyntax.add_syntactic_definition (Global.env()) (snd lid) x local y
+ Metasyntax.add_syntactic_definition (Global.env()) lid.v x local y
let vernac_declare_implicits ~atts r l =
let local = make_section_locality atts.locality in
@@ -1024,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
@@ -1211,7 +1223,7 @@ let vernac_arguments ~atts reference args more_implicits nargs_for_red flags =
end;
if scopes_specified || clear_scopes_flag then begin
- let scopes = List.map (Option.map (fun (loc,k) ->
+ let scopes = List.map (Option.map (fun {loc;v=k} ->
try ignore (Notation.find_scope k); k
with UserError _ ->
Notation.find_delimiters_scope ?loc k)) scopes
@@ -1252,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
@@ -1468,6 +1480,14 @@ let _ =
let _ =
declare_bool_option
{ optdepr = false;
+ optname = "dumping VM lambda code after compilation";
+ optkey = ["Dump";"Lambda"];
+ optread = Flags.get_dump_lambda;
+ optwrite = Flags.set_dump_lambda }
+
+let _ =
+ declare_bool_option
+ { optdepr = false;
optname = "explicitly parsing implicit arguments";
optkey = ["Parsing";"Explicit"];
optread = (fun () -> !Constrintern.parsing_explicit);
@@ -1618,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 ++
@@ -1738,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
@@ -1787,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
@@ -1806,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
@@ -1824,7 +1846,7 @@ let vernac_locate = function
let vernac_register id r =
if Proof_global.there_are_pending_proofs () then
user_err Pp.(str "Cannot register a primitive while in proof editing mode.");
- let kn = Constrintern.global_reference (snd id) in
+ let kn = Constrintern.global_reference id.v in
if not (isConstRef kn) then
user_err Pp.(str "Register inline: a constant is expected");
match r with
@@ -1901,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)
@@ -1917,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"];
@@ -1933,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
@@ -1945,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... *)
@@ -2015,7 +2044,7 @@ let interp ?proof ~atts ~st c =
| VernacImport (export,qidl) -> vernac_import export qidl
| VernacCanonical qid -> vernac_canonical qid
| VernacCoercion (r,s,t) -> vernac_coercion ~atts r s t
- | VernacIdentityCoercion ((_,id),s,t) ->
+ | VernacIdentityCoercion ({v=id},s,t) ->
vernac_identity_coercion ~atts id s t
(* Type classes *)
@@ -2233,12 +2262,12 @@ let interp ?(verbosely=true) ?proof ~st (loc,c) =
aux ~polymorphism ~atts v
| VernacFail v -> with_fail st true (fun () -> control v)
| VernacTimeout (n,v) ->
- current_timeout := Some n;
- control v
- | VernacRedirect (s, (_,v)) ->
- Topfmt.with_output_to_file s control v
- | VernacTime (batch, (_loc,v)) ->
- System.with_time ~batch control v;
+ current_timeout := Some n;
+ control v
+ | VernacRedirect (s, {v}) ->
+ Topfmt.with_output_to_file s control v
+ | VernacTime (batch, {v}) ->
+ System.with_time ~batch control v;
and aux ~polymorphism ~atts : _ -> unit =
function
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 3932d1c7b..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
@@ -13,15 +15,15 @@ open Vernacexpr
let rec under_control = function
| VernacExpr (_, c) -> c
- | VernacRedirect (_,(_,c))
- | VernacTime (_,(_,c))
+ | VernacRedirect (_,{CAst.v=c})
+ | VernacTime (_,{CAst.v=c})
| VernacFail c
| VernacTimeout (_,c) -> under_control c
let rec has_Fail = function
| VernacExpr _ -> false
- | VernacRedirect (_,(_,c))
- | VernacTime (_,(_,c))
+ | VernacRedirect (_,{CAst.v=c})
+ | VernacTime (_,{CAst.v=c})
| VernacTimeout (_,c) -> has_Fail c
| VernacFail _ -> true
@@ -38,8 +40,8 @@ let is_navigation_vernac c =
is_navigation_vernac_expr (under_control c)
let rec is_deep_navigation_vernac = function
- | VernacTime (_,(_,c)) -> is_deep_navigation_vernac c
- | VernacRedirect (_, (_,c))
+ | VernacTime (_,{CAst.v=c}) -> is_deep_navigation_vernac c
+ | VernacRedirect (_, {CAst.v=c})
| VernacTimeout (_,c) | VernacFail c -> is_navigation_vernac c
| VernacExpr _ -> false
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 = {