aboutsummaryrefslogtreecommitdiffhomepage
path: root/plugins
diff options
context:
space:
mode:
Diffstat (limited to 'plugins')
-rw-r--r--plugins/cc/ccalgo.ml48
-rw-r--r--plugins/cc/cctac.ml2
-rw-r--r--plugins/extraction/common.mli3
-rw-r--r--plugins/extraction/extract_env.ml2
-rw-r--r--plugins/extraction/extract_env.mli3
-rw-r--r--plugins/extraction/miniml.ml25
-rw-r--r--plugins/extraction/miniml.mli25
-rw-r--r--plugins/extraction/mlutil.ml16
-rw-r--r--plugins/extraction/mlutil.mli5
-rw-r--r--plugins/extraction/modutil.ml2
-rw-r--r--plugins/extraction/modutil.mli7
-rw-r--r--plugins/extraction/table.ml8
-rw-r--r--plugins/extraction/table.mli69
-rw-r--r--plugins/firstorder/formula.ml2
-rw-r--r--plugins/firstorder/formula.mli8
-rw-r--r--plugins/firstorder/instances.ml2
-rw-r--r--plugins/firstorder/instances.mli4
-rw-r--r--plugins/firstorder/rules.ml2
-rw-r--r--plugins/firstorder/rules.mli9
-rw-r--r--plugins/firstorder/sequent.ml28
-rw-r--r--plugins/firstorder/sequent.mli22
-rw-r--r--plugins/fourier/fourierR.ml2
-rw-r--r--plugins/funind/functional_principles_proofs.ml3
-rw-r--r--plugins/funind/functional_principles_types.ml16
-rw-r--r--plugins/funind/glob_term_to_relation.ml4
-rw-r--r--plugins/funind/glob_termops.ml11
-rw-r--r--plugins/funind/glob_termops.mli2
-rw-r--r--plugins/funind/indfun.ml28
-rw-r--r--plugins/funind/indfun.mli3
-rw-r--r--plugins/funind/indfun_common.mli6
-rw-r--r--plugins/funind/invfun.ml8
-rw-r--r--plugins/funind/invfun.mli2
-rw-r--r--plugins/funind/recdef.ml20
-rw-r--r--plugins/ltac/extratactics.ml44
-rw-r--r--plugins/ltac/g_ltac.ml43
-rw-r--r--plugins/ltac/pptactic.ml1
-rw-r--r--plugins/ltac/pptactic.mli2
-rw-r--r--plugins/ltac/rewrite.ml8
-rw-r--r--plugins/ltac/taccoerce.mli2
-rw-r--r--plugins/ltac/tacexpr.ml5
-rw-r--r--plugins/ltac/tacexpr.mli5
-rw-r--r--plugins/ltac/tacinterp.ml3
-rw-r--r--plugins/ltac/tactic_debug.ml6
-rw-r--r--plugins/setoid_ring/Algebra_syntax.v9
-rw-r--r--plugins/setoid_ring/Integral_domain.v10
-rw-r--r--plugins/setoid_ring/RealField.v10
-rw-r--r--plugins/setoid_ring/Ring_tac.v10
-rw-r--r--plugins/setoid_ring/Rings_Q.v10
-rw-r--r--plugins/setoid_ring/Rings_R.v10
-rw-r--r--plugins/setoid_ring/Rings_Z.v10
-rw-r--r--plugins/setoid_ring/newring.ml10
-rw-r--r--plugins/setoid_ring/newring.mli3
-rw-r--r--plugins/ssr/ssrcommon.ml27
-rw-r--r--plugins/ssr/ssrcommon.mli4
-rw-r--r--plugins/ssr/ssrelim.ml4
-rw-r--r--plugins/ssr/ssrequality.ml21
-rw-r--r--plugins/ssr/ssripats.ml23
-rw-r--r--plugins/ssr/ssrview.ml4
-rw-r--r--plugins/ssrmatching/ssrmatching.ml424
-rw-r--r--plugins/syntax/ascii_syntax.ml2
-rw-r--r--plugins/syntax/int31_syntax.ml2
-rw-r--r--plugins/syntax/nat_syntax.ml9
-rw-r--r--plugins/syntax/r_syntax.ml6
-rw-r--r--plugins/syntax/string_syntax.ml5
-rw-r--r--plugins/syntax/z_syntax.ml8
65 files changed, 360 insertions, 297 deletions
diff --git a/plugins/cc/ccalgo.ml b/plugins/cc/ccalgo.ml
index 5a4818926..8e53a044d 100644
--- a/plugins/cc/ccalgo.ml
+++ b/plugins/cc/ccalgo.ml
@@ -9,7 +9,7 @@
(************************************************************************)
(* This file implements the basic congruence-closure algorithm by *)
-(* Downey,Sethi and Tarjan. *)
+(* Downey, Sethi and Tarjan. *)
(* Plus some e-matching and constructor handling by P. Corbineau *)
open CErrors
@@ -18,7 +18,6 @@ open Names
open Sorts
open Constr
open Vars
-open Evd
open Goptions
open Tacmach
open Util
@@ -272,7 +271,8 @@ type state =
mutable rew_depth:int;
mutable changed:bool;
by_type: Int.Set.t Typehash.t;
- mutable gls:Goal.goal Evd.sigma}
+ mutable env:Environ.env;
+ sigma:Evd.evar_map}
let dummy_node =
{
@@ -307,7 +307,8 @@ let empty depth gls:state =
rew_depth=depth;
by_type=Constrhash.create init_size;
changed=false;
- gls=gls
+ env=pf_env gls;
+ sigma=project gls
}
let forest state = state.uf
@@ -426,7 +427,7 @@ let cc_product s1 s2 =
mkLambda(_B_,mkSort(s2),_body_))
let rec constr_of_term = function
- Symb s-> applist_projection s []
+ Symb s-> s
| Product(s1,s2) -> cc_product s1 s2
| Eps id -> mkVar id
| Constructor cinfo -> mkConstructU cinfo.ci_constr
@@ -434,25 +435,7 @@ let rec constr_of_term = function
make_app [(constr_of_term s2)] s1
and make_app l=function
Appli (s1,s2)->make_app ((constr_of_term s2)::l) s1
- | other ->
- applist_proj other l
-and applist_proj c l =
- match c with
- | Symb s -> applist_projection s l
- | _ -> Term.applistc (constr_of_term c) l
-and applist_projection c l =
- match Constr.kind c with
- | Const c when Environ.is_projection (fst c) (Global.env()) ->
- let p = Projection.make (fst c) false in
- (match l with
- | [] -> (* Expand the projection *)
- let ty = Typeops.type_of_constant_in (Global.env ()) c in (* FIXME constraints *)
- let pb = Environ.lookup_projection p (Global.env()) in
- let ctx,_ = Term.decompose_prod_n_assum (pb.Declarations.proj_npars + 1) ty in
- Term.it_mkLambda_or_LetIn (mkProj(p,mkRel 1)) ctx
- | hd :: tl ->
- Term.applistc (mkProj (p, hd)) tl)
- | _ -> Term.applistc c l
+ | other -> Term.applist (constr_of_term other,l)
let rec canonize_name sigma c =
let c = EConstr.Unsafe.to_constr c in
@@ -511,8 +494,8 @@ let rec add_term state t=
Not_found ->
let b=next uf in
let trm = constr_of_term t in
- let typ = pf_unsafe_type_of state.gls (EConstr.of_constr trm) in
- let typ = canonize_name (project state.gls) typ in
+ let typ = Typing.unsafe_type_of state.env state.sigma (EConstr.of_constr trm) in
+ let typ = canonize_name state.sigma typ in
let new_node=
match t with
Symb _ | Product (_,_) ->
@@ -820,11 +803,10 @@ let one_step state =
let __eps__ = Id.of_string "_eps_"
let new_state_var typ state =
- let id = pf_get_new_id __eps__ state.gls in
- let {it=gl ; sigma=sigma} = state.gls in
- let gls = Goal.V82.new_goal_with sigma gl [Context.Named.Declaration.LocalAssum (id,typ)] in
- state.gls<- gls;
- id
+ let ids = Environ.ids_of_named_context_val (Environ.named_context_val state.env) in
+ let id = Namegen.next_ident_away __eps__ ids in
+ state.env<- EConstr.push_named (Context.Named.Declaration.LocalAssum (id,typ)) state.env;
+ id
let complete_one_class state i=
match (get_representative state.uf i).inductive_status with
@@ -832,9 +814,9 @@ let complete_one_class state i=
let rec app t typ n =
if n<=0 then t else
let _,etyp,rest= destProd typ in
- let id = new_state_var etyp state in
+ let id = new_state_var (EConstr.of_constr etyp) state in
app (Appli(t,Eps id)) (substl [mkVar id] rest) (n-1) in
- let _c = pf_unsafe_type_of state.gls
+ let _c = Typing.unsafe_type_of state.env state.sigma
(EConstr.of_constr (constr_of_term (term state.uf pac.cnode))) in
let _c = EConstr.Unsafe.to_constr _c in
let _args =
diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml
index d19817e74..c4db49cd3 100644
--- a/plugins/cc/cctac.ml
+++ b/plugins/cc/cctac.ml
@@ -90,7 +90,7 @@ let rec decompose_term env sigma t=
decompose_term env sigma c
| _ ->
let t = Termops.strip_outer_cast sigma t in
- if closed0 sigma t then Symb (EConstr.to_constr sigma t) else raise Not_found
+ if closed0 sigma t then Symb (EConstr.to_constr ~abort_on_undefined_evars:false sigma t) else raise Not_found
(* decompose equality in members and type *)
open Termops
diff --git a/plugins/extraction/common.mli b/plugins/extraction/common.mli
index 78545c8bd..07237d750 100644
--- a/plugins/extraction/common.mli
+++ b/plugins/extraction/common.mli
@@ -9,7 +9,6 @@
(************************************************************************)
open Names
-open Globnames
open Miniml
(** By default, in module Format, you can do horizontal placing of blocks
@@ -54,7 +53,7 @@ val opened_libraries : unit -> ModPath.t list
type kind = Term | Type | Cons | Mod
-val pp_global : kind -> global_reference -> string
+val pp_global : kind -> GlobRef.t -> string
val pp_module : ModPath.t -> string
val top_visible_mp : unit -> ModPath.t
diff --git a/plugins/extraction/extract_env.ml b/plugins/extraction/extract_env.ml
index 397cb2920..bebd27e11 100644
--- a/plugins/extraction/extract_env.ml
+++ b/plugins/extraction/extract_env.ml
@@ -79,7 +79,7 @@ module type VISIT = sig
(* Add reference / ... in the visit lists.
These functions silently add the mp of their arg in the mp list *)
- val add_ref : global_reference -> unit
+ val add_ref : GlobRef.t -> unit
val add_kn : KerName.t -> unit
val add_decl_deps : ml_decl -> unit
val add_spec_deps : ml_spec -> unit
diff --git a/plugins/extraction/extract_env.mli b/plugins/extraction/extract_env.mli
index 591d3bb86..77f1fb5ef 100644
--- a/plugins/extraction/extract_env.mli
+++ b/plugins/extraction/extract_env.mli
@@ -12,7 +12,6 @@
open Names
open Libnames
-open Globnames
val simple_extraction : reference -> unit
val full_extraction : string option -> reference list -> unit
@@ -26,7 +25,7 @@ val extract_and_compile : reference list -> unit
(* For debug / external output via coqtop.byte + Drop : *)
val mono_environment :
- global_reference list -> ModPath.t list -> Miniml.ml_structure
+ GlobRef.t list -> ModPath.t list -> Miniml.ml_structure
(* Used by the Relation Extraction plugin *)
diff --git a/plugins/extraction/miniml.ml b/plugins/extraction/miniml.ml
index e1e49d926..ce920ad6a 100644
--- a/plugins/extraction/miniml.ml
+++ b/plugins/extraction/miniml.ml
@@ -11,7 +11,6 @@
(*s Target language for extraction: a core ML called MiniML. *)
open Names
-open Globnames
(* The [signature] type is used to know how many arguments a CIC
object expects, and what these arguments will become in the ML
@@ -26,7 +25,7 @@ open Globnames
type kill_reason =
| Ktype
| Kprop
- | Kimplicit of global_reference * int (* n-th arg of a cst or construct *)
+ | Kimplicit of GlobRef.t * int (* n-th arg of a cst or construct *)
type sign = Keep | Kill of kill_reason
@@ -39,7 +38,7 @@ type signature = sign list
type ml_type =
| Tarr of ml_type * ml_type
- | Tglob of global_reference * ml_type list
+ | Tglob of GlobRef.t * ml_type list
| Tvar of int
| Tvar' of int (* same as Tvar, used to avoid clash *)
| Tmeta of ml_meta (* used during ML type reconstruction *)
@@ -60,7 +59,7 @@ type inductive_kind =
| Singleton
| Coinductive
| Standard
- | Record of global_reference option list (* None for anonymous field *)
+ | Record of GlobRef.t option list (* None for anonymous field *)
(* A [ml_ind_packet] is the miniml counterpart of a [one_inductive_body].
If the inductive is logical ([ip_logical = false]), then all other fields
@@ -118,8 +117,8 @@ and ml_ast =
| MLapp of ml_ast * ml_ast list
| MLlam of ml_ident * ml_ast
| MLletin of ml_ident * ml_ast * ml_ast
- | MLglob of global_reference
- | MLcons of ml_type * global_reference * ml_ast list
+ | MLglob of GlobRef.t
+ | MLcons of ml_type * GlobRef.t * ml_ast list
| MLtuple of ml_ast list
| MLcase of ml_type * ml_ast * ml_branch array
| MLfix of int * Id.t array * ml_ast array
@@ -129,24 +128,24 @@ and ml_ast =
| MLmagic of ml_ast
and ml_pattern =
- | Pcons of global_reference * ml_pattern list
+ | Pcons of GlobRef.t * ml_pattern list
| Ptuple of ml_pattern list
| Prel of int (** Cf. the idents in the branch. [Prel 1] is the last one. *)
| Pwild
- | Pusual of global_reference (** Shortcut for Pcons (r,[Prel n;...;Prel 1]) **)
+ | Pusual of GlobRef.t (** Shortcut for Pcons (r,[Prel n;...;Prel 1]) **)
(*s ML declarations. *)
type ml_decl =
| Dind of MutInd.t * ml_ind
- | Dtype of global_reference * Id.t list * ml_type
- | Dterm of global_reference * ml_ast * ml_type
- | Dfix of global_reference array * ml_ast array * ml_type array
+ | Dtype of GlobRef.t * Id.t list * ml_type
+ | Dterm of GlobRef.t * ml_ast * ml_type
+ | Dfix of GlobRef.t array * ml_ast array * ml_type array
type ml_spec =
| Sind of MutInd.t * ml_ind
- | Stype of global_reference * Id.t list * ml_type option
- | Sval of global_reference * ml_type
+ | Stype of GlobRef.t * Id.t list * ml_type option
+ | Sval of GlobRef.t * ml_type
type ml_specif =
| Spec of ml_spec
diff --git a/plugins/extraction/miniml.mli b/plugins/extraction/miniml.mli
index e1e49d926..ce920ad6a 100644
--- a/plugins/extraction/miniml.mli
+++ b/plugins/extraction/miniml.mli
@@ -11,7 +11,6 @@
(*s Target language for extraction: a core ML called MiniML. *)
open Names
-open Globnames
(* The [signature] type is used to know how many arguments a CIC
object expects, and what these arguments will become in the ML
@@ -26,7 +25,7 @@ open Globnames
type kill_reason =
| Ktype
| Kprop
- | Kimplicit of global_reference * int (* n-th arg of a cst or construct *)
+ | Kimplicit of GlobRef.t * int (* n-th arg of a cst or construct *)
type sign = Keep | Kill of kill_reason
@@ -39,7 +38,7 @@ type signature = sign list
type ml_type =
| Tarr of ml_type * ml_type
- | Tglob of global_reference * ml_type list
+ | Tglob of GlobRef.t * ml_type list
| Tvar of int
| Tvar' of int (* same as Tvar, used to avoid clash *)
| Tmeta of ml_meta (* used during ML type reconstruction *)
@@ -60,7 +59,7 @@ type inductive_kind =
| Singleton
| Coinductive
| Standard
- | Record of global_reference option list (* None for anonymous field *)
+ | Record of GlobRef.t option list (* None for anonymous field *)
(* A [ml_ind_packet] is the miniml counterpart of a [one_inductive_body].
If the inductive is logical ([ip_logical = false]), then all other fields
@@ -118,8 +117,8 @@ and ml_ast =
| MLapp of ml_ast * ml_ast list
| MLlam of ml_ident * ml_ast
| MLletin of ml_ident * ml_ast * ml_ast
- | MLglob of global_reference
- | MLcons of ml_type * global_reference * ml_ast list
+ | MLglob of GlobRef.t
+ | MLcons of ml_type * GlobRef.t * ml_ast list
| MLtuple of ml_ast list
| MLcase of ml_type * ml_ast * ml_branch array
| MLfix of int * Id.t array * ml_ast array
@@ -129,24 +128,24 @@ and ml_ast =
| MLmagic of ml_ast
and ml_pattern =
- | Pcons of global_reference * ml_pattern list
+ | Pcons of GlobRef.t * ml_pattern list
| Ptuple of ml_pattern list
| Prel of int (** Cf. the idents in the branch. [Prel 1] is the last one. *)
| Pwild
- | Pusual of global_reference (** Shortcut for Pcons (r,[Prel n;...;Prel 1]) **)
+ | Pusual of GlobRef.t (** Shortcut for Pcons (r,[Prel n;...;Prel 1]) **)
(*s ML declarations. *)
type ml_decl =
| Dind of MutInd.t * ml_ind
- | Dtype of global_reference * Id.t list * ml_type
- | Dterm of global_reference * ml_ast * ml_type
- | Dfix of global_reference array * ml_ast array * ml_type array
+ | Dtype of GlobRef.t * Id.t list * ml_type
+ | Dterm of GlobRef.t * ml_ast * ml_type
+ | Dfix of GlobRef.t array * ml_ast array * ml_type array
type ml_spec =
| Sind of MutInd.t * ml_ind
- | Stype of global_reference * Id.t list * ml_type option
- | Sval of global_reference * ml_type
+ | Stype of GlobRef.t * Id.t list * ml_type option
+ | Sval of GlobRef.t * ml_type
type ml_specif =
| Spec of ml_spec
diff --git a/plugins/extraction/mlutil.ml b/plugins/extraction/mlutil.ml
index 0656d487a..0901acc7d 100644
--- a/plugins/extraction/mlutil.ml
+++ b/plugins/extraction/mlutil.ml
@@ -59,7 +59,7 @@ let rec eq_ml_type t1 t2 = match t1, t2 with
| Tarr (tl1, tr1), Tarr (tl2, tr2) ->
eq_ml_type tl1 tl2 && eq_ml_type tr1 tr2
| Tglob (gr1, t1), Tglob (gr2, t2) ->
- eq_gr gr1 gr2 && List.equal eq_ml_type t1 t2
+ GlobRef.equal gr1 gr2 && List.equal eq_ml_type t1 t2
| Tvar i1, Tvar i2 -> Int.equal i1 i2
| Tvar' i1, Tvar' i2 -> Int.equal i1 i2
| Tmeta m1, Tmeta m2 -> eq_ml_meta m1 m2
@@ -120,7 +120,7 @@ let rec mgu = function
| None -> m.contents <- Some t)
| Tarr(a, b), Tarr(a', b') ->
mgu (a, a'); mgu (b, b')
- | Tglob (r,l), Tglob (r',l') when Globnames.eq_gr r r' ->
+ | Tglob (r,l), Tglob (r',l') when GlobRef.equal r r' ->
List.iter mgu (List.combine l l')
| Tdummy _, Tdummy _ -> ()
| Tvar i, Tvar j when Int.equal i j -> ()
@@ -270,7 +270,7 @@ let rec var2var' = function
| Tglob (r,l) -> Tglob (r, List.map var2var' l)
| a -> a
-type abbrev_map = global_reference -> ml_type option
+type abbrev_map = GlobRef.t -> ml_type option
(*s Delta-reduction of type constants everywhere in a ML type [t].
[env] is a function of type [ml_type_env]. *)
@@ -381,9 +381,9 @@ let rec eq_ml_ast t1 t2 = match t1, t2 with
eq_ml_ident na1 na2 && eq_ml_ast t1 t2
| MLletin (na1, c1, t1), MLletin (na2, c2, t2) ->
eq_ml_ident na1 na2 && eq_ml_ast c1 c2 && eq_ml_ast t1 t2
-| MLglob gr1, MLglob gr2 -> eq_gr gr1 gr2
+| MLglob gr1, MLglob gr2 -> GlobRef.equal gr1 gr2
| MLcons (t1, gr1, c1), MLcons (t2, gr2, c2) ->
- eq_ml_type t1 t2 && eq_gr gr1 gr2 && List.equal eq_ml_ast c1 c2
+ eq_ml_type t1 t2 && GlobRef.equal gr1 gr2 && List.equal eq_ml_ast c1 c2
| MLtuple t1, MLtuple t2 ->
List.equal eq_ml_ast t1 t2
| MLcase (t1, c1, p1), MLcase (t2, c2, p2) ->
@@ -398,13 +398,13 @@ let rec eq_ml_ast t1 t2 = match t1, t2 with
and eq_ml_pattern p1 p2 = match p1, p2 with
| Pcons (gr1, p1), Pcons (gr2, p2) ->
- eq_gr gr1 gr2 && List.equal eq_ml_pattern p1 p2
+ GlobRef.equal gr1 gr2 && List.equal eq_ml_pattern p1 p2
| Ptuple p1, Ptuple p2 ->
List.equal eq_ml_pattern p1 p2
| Prel i1, Prel i2 ->
Int.equal i1 i2
| Pwild, Pwild -> true
-| Pusual gr1, Pusual gr2 -> eq_gr gr1 gr2
+| Pusual gr1, Pusual gr2 -> GlobRef.equal gr1 gr2
| _ -> false
and eq_ml_branch (id1, p1, t1) (id2, p2, t2) =
@@ -984,7 +984,7 @@ let rec iota_red i lift br ((typ,r,a) as cons) =
if i >= Array.length br then raise Impossible;
let (ids,p,c) = br.(i) in
match p with
- | Pusual r' | Pcons (r',_) when not (Globnames.eq_gr r' r) -> iota_red (i+1) lift br cons
+ | Pusual r' | Pcons (r',_) when not (GlobRef.equal r' r) -> iota_red (i+1) lift br cons
| Pusual r' ->
let c = named_lams (List.rev ids) c in
let c = ast_lift lift c
diff --git a/plugins/extraction/mlutil.mli b/plugins/extraction/mlutil.mli
index 55a1ee893..d23fdb3d5 100644
--- a/plugins/extraction/mlutil.mli
+++ b/plugins/extraction/mlutil.mli
@@ -9,7 +9,6 @@
(************************************************************************)
open Names
-open Globnames
open Miniml
open Table
@@ -59,7 +58,7 @@ val type_recomp : ml_type list * ml_type -> ml_type
val var2var' : ml_type -> ml_type
-type abbrev_map = global_reference -> ml_type option
+type abbrev_map = GlobRef.t -> ml_type option
val type_expand : abbrev_map -> ml_type -> ml_type
val type_simpl : ml_type -> ml_type
@@ -117,7 +116,7 @@ val dump_unused_vars : ml_ast -> ml_ast
val normalize : ml_ast -> ml_ast
val optimize_fix : ml_ast -> ml_ast
-val inline : global_reference -> ml_ast -> bool
+val inline : GlobRef.t -> ml_ast -> bool
val is_basic_pattern : ml_pattern -> bool
val has_deep_pattern : ml_branch array -> bool
diff --git a/plugins/extraction/modutil.ml b/plugins/extraction/modutil.ml
index f33a59edf..b398bc07a 100644
--- a/plugins/extraction/modutil.ml
+++ b/plugins/extraction/modutil.ml
@@ -76,7 +76,7 @@ let struct_iter do_decl do_spec do_mp s =
(*s Apply some fonctions upon all references in [ml_type], [ml_ast],
[ml_decl], [ml_spec] and [ml_structure]. *)
-type do_ref = global_reference -> unit
+type do_ref = GlobRef.t -> unit
let record_iter_references do_term = function
| Record l -> List.iter (Option.iter do_term) l
diff --git a/plugins/extraction/modutil.mli b/plugins/extraction/modutil.mli
index 6a81f2705..f45773f09 100644
--- a/plugins/extraction/modutil.mli
+++ b/plugins/extraction/modutil.mli
@@ -9,7 +9,6 @@
(************************************************************************)
open Names
-open Globnames
open Miniml
(*s Functions upon ML modules. *)
@@ -17,7 +16,7 @@ open Miniml
val struct_ast_search : (ml_ast -> bool) -> ml_structure -> bool
val struct_type_search : (ml_type -> bool) -> ml_structure -> bool
-type do_ref = global_reference -> unit
+type do_ref = GlobRef.t -> unit
val type_iter_references : do_ref -> ml_type -> unit
val ast_iter_references : do_ref -> do_ref -> do_ref -> ml_ast -> unit
@@ -30,7 +29,7 @@ val mtyp_of_mexpr : ml_module_expr -> ml_module_type
val msid_of_mt : ml_module_type -> ModPath.t
-val get_decl_in_structure : global_reference -> ml_structure -> ml_decl
+val get_decl_in_structure : GlobRef.t -> ml_structure -> ml_decl
(* Some transformations of ML terms. [optimize_struct] simplify
all beta redexes (when the argument does not occur, it is just
@@ -39,5 +38,5 @@ val get_decl_in_structure : global_reference -> ml_structure -> ml_decl
optimizations. The first argument is the list of objects we want to appear.
*)
-val optimize_struct : global_reference list * ModPath.t list ->
+val optimize_struct : GlobRef.t list * ModPath.t list ->
ml_structure -> ml_structure
diff --git a/plugins/extraction/table.ml b/plugins/extraction/table.ml
index 54c6d9d72..c3f4cfe65 100644
--- a/plugins/extraction/table.ml
+++ b/plugins/extraction/table.ml
@@ -652,7 +652,7 @@ let add_inline_entries b l =
(* Registration of operations for rollback. *)
-let inline_extraction : bool * global_reference list -> obj =
+let inline_extraction : bool * GlobRef.t list -> obj =
declare_object
{(default_object "Extraction Inline") with
cache_function = (fun (_,(b,l)) -> add_inline_entries b l);
@@ -736,7 +736,7 @@ let add_implicits r l =
(* Registration of operations for rollback. *)
-let implicit_extraction : global_reference * int_or_id list -> obj =
+let implicit_extraction : GlobRef.t * int_or_id list -> obj =
declare_object
{(default_object "Extraction Implicit") with
cache_function = (fun (_,(r,l)) -> add_implicits r l);
@@ -857,7 +857,7 @@ let find_custom_match pv =
(* Registration of operations for rollback. *)
-let in_customs : global_reference * string list * string -> obj =
+let in_customs : GlobRef.t * string list * string -> obj =
declare_object
{(default_object "ML extractions") with
cache_function = (fun (_,(r,ids,s)) -> add_custom r ids s);
@@ -867,7 +867,7 @@ let in_customs : global_reference * string list * string -> obj =
(fun (s,(r,ids,str)) -> (fst (subst_global s r), ids, str))
}
-let in_custom_matchs : global_reference * string -> obj =
+let in_custom_matchs : GlobRef.t * string -> obj =
declare_object
{(default_object "ML extractions custom matchs") with
cache_function = (fun (_,(r,s)) -> add_custom_match r s);
diff --git a/plugins/extraction/table.mli b/plugins/extraction/table.mli
index 906dfd96e..5bf944434 100644
--- a/plugins/extraction/table.mli
+++ b/plugins/extraction/table.mli
@@ -10,31 +10,30 @@
open Names
open Libnames
-open Globnames
open Miniml
open Declarations
-module Refset' : CSig.SetS with type elt = global_reference
-module Refmap' : CSig.MapS with type key = global_reference
+module Refset' : CSig.SetS with type elt = GlobRef.t
+module Refmap' : CSig.MapS with type key = GlobRef.t
-val safe_basename_of_global : global_reference -> Id.t
+val safe_basename_of_global : GlobRef.t -> Id.t
(*s Warning and Error messages. *)
val warning_axioms : unit -> unit
val warning_opaques : bool -> unit
-val warning_ambiguous_name : ?loc:Loc.t -> qualid * ModPath.t * global_reference -> unit
+val warning_ambiguous_name : ?loc:Loc.t -> qualid * ModPath.t * GlobRef.t -> unit
val warning_id : string -> unit
-val error_axiom_scheme : global_reference -> int -> 'a
-val error_constant : global_reference -> 'a
-val error_inductive : global_reference -> 'a
+val error_axiom_scheme : GlobRef.t -> int -> 'a
+val error_constant : GlobRef.t -> 'a
+val error_inductive : GlobRef.t -> 'a
val error_nb_cons : unit -> 'a
val error_module_clash : ModPath.t -> ModPath.t -> 'a
val error_no_module_expr : ModPath.t -> 'a
-val error_singleton_become_prop : Id.t -> global_reference option -> 'a
+val error_singleton_become_prop : Id.t -> GlobRef.t option -> 'a
val error_unknown_module : qualid -> 'a
val error_scheme : unit -> 'a
-val error_not_visible : global_reference -> 'a
+val error_not_visible : GlobRef.t -> 'a
val error_MPfile_as_mod : ModPath.t -> bool -> 'a
val check_inside_module : unit -> unit
val check_inside_section : unit -> unit
@@ -44,12 +43,12 @@ val err_or_warn_remaining_implicit : kill_reason -> unit
val info_file : string -> unit
-(*s utilities about [module_path] and [kernel_names] and [global_reference] *)
+(*s utilities about [module_path] and [kernel_names] and [GlobRef.t] *)
-val occur_kn_in_ref : MutInd.t -> global_reference -> bool
-val repr_of_r : global_reference -> ModPath.t * DirPath.t * Label.t
-val modpath_of_r : global_reference -> ModPath.t
-val label_of_r : global_reference -> Label.t
+val occur_kn_in_ref : MutInd.t -> GlobRef.t -> bool
+val repr_of_r : GlobRef.t -> ModPath.t * DirPath.t * Label.t
+val modpath_of_r : GlobRef.t -> ModPath.t
+val label_of_r : GlobRef.t -> Label.t
val base_mp : ModPath.t -> ModPath.t
val is_modfile : ModPath.t -> bool
val string_of_modfile : ModPath.t -> string
@@ -61,7 +60,7 @@ val prefixes_mp : ModPath.t -> MPset.t
val common_prefix_from_list :
ModPath.t -> ModPath.t list -> ModPath.t option
val get_nth_label_mp : int -> ModPath.t -> Label.t
-val labels_of_ref : global_reference -> ModPath.t * Label.t list
+val labels_of_ref : GlobRef.t -> ModPath.t * Label.t list
(*s Some table-related operations *)
@@ -83,27 +82,27 @@ val add_ind : MutInd.t -> mutual_inductive_body -> ml_ind -> unit
val lookup_ind : MutInd.t -> mutual_inductive_body -> ml_ind option
val add_inductive_kind : MutInd.t -> inductive_kind -> unit
-val is_coinductive : global_reference -> bool
+val is_coinductive : GlobRef.t -> bool
val is_coinductive_type : ml_type -> bool
(* What are the fields of a record (empty for a non-record) *)
val get_record_fields :
- global_reference -> global_reference option list
-val record_fields_of_type : ml_type -> global_reference option list
+ GlobRef.t -> GlobRef.t option list
+val record_fields_of_type : ml_type -> GlobRef.t option list
val add_recursors : Environ.env -> MutInd.t -> unit
-val is_recursor : global_reference -> bool
+val is_recursor : GlobRef.t -> bool
val add_projection : int -> Constant.t -> inductive -> unit
-val is_projection : global_reference -> bool
-val projection_arity : global_reference -> int
-val projection_info : global_reference -> inductive * int (* arity *)
+val is_projection : GlobRef.t -> bool
+val projection_arity : GlobRef.t -> int
+val projection_info : GlobRef.t -> inductive * int (* arity *)
-val add_info_axiom : global_reference -> unit
-val remove_info_axiom : global_reference -> unit
-val add_log_axiom : global_reference -> unit
+val add_info_axiom : GlobRef.t -> unit
+val remove_info_axiom : GlobRef.t -> unit
+val add_log_axiom : GlobRef.t -> unit
-val add_opaque : global_reference -> unit
-val remove_opaque : global_reference -> unit
+val add_opaque : GlobRef.t -> unit
+val remove_opaque : GlobRef.t -> unit
val reset_tables : unit -> unit
@@ -172,22 +171,22 @@ val is_extrcompute : unit -> bool
(*s Table for custom inlining *)
-val to_inline : global_reference -> bool
-val to_keep : global_reference -> bool
+val to_inline : GlobRef.t -> bool
+val to_keep : GlobRef.t -> bool
(*s Table for implicits arguments *)
-val implicits_of_global : global_reference -> Int.Set.t
+val implicits_of_global : GlobRef.t -> Int.Set.t
(*s Table for user-given custom ML extractions. *)
(* UGLY HACK: registration of a function defined in [extraction.ml] *)
val type_scheme_nb_args_hook : (Environ.env -> Constr.t -> int) Hook.t
-val is_custom : global_reference -> bool
-val is_inline_custom : global_reference -> bool
-val find_custom : global_reference -> string
-val find_type_custom : global_reference -> string list * string
+val is_custom : GlobRef.t -> bool
+val is_inline_custom : GlobRef.t -> bool
+val find_custom : GlobRef.t -> string
+val find_type_custom : GlobRef.t -> string list * string
val is_custom_match : ml_branch array -> bool
val find_custom_match : ml_branch array -> string
diff --git a/plugins/firstorder/formula.ml b/plugins/firstorder/formula.ml
index 047fc9fbf..a60a966ce 100644
--- a/plugins/firstorder/formula.ml
+++ b/plugins/firstorder/formula.ml
@@ -211,7 +211,7 @@ type left_pattern=
| Lexists of pinductive
| LA of constr*left_arrow_pattern
-type t={id:global_reference;
+type t={id:GlobRef.t;
constr:constr;
pat:(left_pattern,right_pattern) sum;
atoms:atoms}
diff --git a/plugins/firstorder/formula.mli b/plugins/firstorder/formula.mli
index 2962d9230..e2c6f1c4b 100644
--- a/plugins/firstorder/formula.mli
+++ b/plugins/firstorder/formula.mli
@@ -8,9 +8,9 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
+open Names
open Constr
open EConstr
-open Globnames
val qflag : bool ref
@@ -35,7 +35,7 @@ type atoms = {positive:constr list;negative:constr list}
type side = Hyp | Concl | Hint
-val dummy_id: global_reference
+val dummy_id: GlobRef.t
val build_atoms : Environ.env -> Evd.evar_map -> counter ->
side -> constr -> bool * atoms
@@ -65,13 +65,13 @@ type left_pattern=
| Lexists of pinductive
| LA of constr*left_arrow_pattern
-type t={id: global_reference;
+type t={id: GlobRef.t;
constr: constr;
pat: (left_pattern,right_pattern) sum;
atoms: atoms}
(*exception Is_atom of constr*)
-val build_formula : Environ.env -> Evd.evar_map -> side -> global_reference -> types ->
+val build_formula : Environ.env -> Evd.evar_map -> side -> GlobRef.t -> types ->
counter -> (t,types) sum
diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml
index e8c0b927d..22a3e1f67 100644
--- a/plugins/firstorder/instances.ml
+++ b/plugins/firstorder/instances.ml
@@ -43,7 +43,7 @@ let compare_gr id1 id2 =
module OrderedInstance=
struct
- type t=instance * Globnames.global_reference
+ type t=instance * GlobRef.t
let compare (inst1,id1) (inst2,id2)=
(compare_instance =? compare_gr) inst2 inst1 id2 id1
(* we want a __decreasing__ total order *)
diff --git a/plugins/firstorder/instances.mli b/plugins/firstorder/instances.mli
index 61786ffdc..9f9ade3aa 100644
--- a/plugins/firstorder/instances.mli
+++ b/plugins/firstorder/instances.mli
@@ -8,13 +8,13 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
-open Globnames
+open Names
open Rules
val collect_quantified : Evd.evar_map -> Sequent.t -> Formula.t list * Sequent.t
val give_instances : Evd.evar_map -> Formula.t list -> Sequent.t ->
- (Unify.instance * global_reference) list
+ (Unify.instance * GlobRef.t) list
val quantified_tac : Formula.t list -> seqtac with_backtracking
diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml
index cfcd65619..2d7a3e37b 100644
--- a/plugins/firstorder/rules.ml
+++ b/plugins/firstorder/rules.ml
@@ -29,7 +29,7 @@ type tactic = unit Proofview.tactic
type seqtac= (Sequent.t -> tactic) -> Sequent.t -> tactic
-type lseqtac= global_reference -> seqtac
+type lseqtac= GlobRef.t -> seqtac
type 'a with_backtracking = tactic -> 'a
diff --git a/plugins/firstorder/rules.mli b/plugins/firstorder/rules.mli
index 859388b30..924c26790 100644
--- a/plugins/firstorder/rules.mli
+++ b/plugins/firstorder/rules.mli
@@ -11,21 +11,20 @@
open Names
open Constr
open EConstr
-open Globnames
type tactic = unit Proofview.tactic
type seqtac= (Sequent.t -> tactic) -> Sequent.t -> tactic
-type lseqtac= global_reference -> seqtac
+type lseqtac= GlobRef.t -> seqtac
type 'a with_backtracking = tactic -> 'a
val wrap : int -> bool -> seqtac
-val basename_of_global: global_reference -> Id.t
+val basename_of_global: GlobRef.t -> Id.t
-val clear_global: global_reference -> tactic
+val clear_global: GlobRef.t -> tactic
val axiom_tac : constr -> Sequent.t -> tactic
@@ -41,7 +40,7 @@ val left_and_tac : pinductive -> lseqtac with_backtracking
val left_or_tac : pinductive -> lseqtac with_backtracking
-val left_false_tac : global_reference -> tactic
+val left_false_tac : GlobRef.t -> tactic
val ll_ind_tac : pinductive -> constr list -> lseqtac with_backtracking
diff --git a/plugins/firstorder/sequent.ml b/plugins/firstorder/sequent.ml
index 285991797..0c752d4a4 100644
--- a/plugins/firstorder/sequent.ml
+++ b/plugins/firstorder/sequent.ml
@@ -8,13 +8,13 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
-open EConstr
-open CErrors
open Util
+open Pp
+open CErrors
+open Names
+open EConstr
open Formula
open Unify
-open Globnames
-open Pp
let newcnt ()=
let cnt=ref (-1) in
@@ -56,7 +56,7 @@ struct
(priority e1.pat) - (priority e2.pat)
end
-type h_item = global_reference * (int*Constr.t) option
+type h_item = GlobRef.t * (int*Constr.t) option
module Hitem=
struct
@@ -77,17 +77,17 @@ module CM=Map.Make(Constr)
module History=Set.Make(Hitem)
let cm_add sigma typ nam cm=
- let typ = EConstr.to_constr sigma typ in
+ let typ = EConstr.to_constr ~abort_on_undefined_evars:false sigma typ in
try
let l=CM.find typ cm in CM.add typ (nam::l) cm
with
Not_found->CM.add typ [nam] cm
let cm_remove sigma typ nam cm=
- let typ = EConstr.to_constr sigma typ in
+ let typ = EConstr.to_constr ~abort_on_undefined_evars:false sigma typ in
try
let l=CM.find typ cm in
- let l0=List.filter (fun id-> not (Globnames.eq_gr id nam)) l in
+ let l0=List.filter (fun id-> not (GlobRef.equal id nam)) l in
match l0 with
[]->CM.remove typ cm
| _ ->CM.add typ l0 cm
@@ -97,7 +97,7 @@ module HP=Heap.Functional(OrderedFormula)
type t=
{redexes:HP.t;
- context:(global_reference list) CM.t;
+ context:(GlobRef.t list) CM.t;
latoms:constr list;
gl:types;
glatom:constr option;
@@ -117,7 +117,7 @@ let lookup sigma item seq=
let p (id2,o)=
match o with
None -> false
- | Some (m2, t2)-> Globnames.eq_gr id id2 && m2>m && more_general sigma (m2, EConstr.of_constr t2) (m, EConstr.of_constr t) in
+ | Some (m2, t2)-> GlobRef.equal id id2 && m2>m && more_general sigma (m2, EConstr.of_constr t2) (m, EConstr.of_constr t) in
History.exists p seq.history
let add_formula env sigma side nam t seq =
@@ -152,7 +152,7 @@ let re_add_formula_list sigma lf seq=
redexes=List.fold_right HP.add lf seq.redexes;
context=List.fold_right do_one lf seq.context}
-let find_left sigma t seq=List.hd (CM.find (EConstr.to_constr sigma t) seq.context)
+let find_left sigma t seq=List.hd (CM.find (EConstr.to_constr ~abort_on_undefined_evars:false sigma t) seq.context)
(*let rev_left seq=
try
@@ -187,9 +187,9 @@ let empty_seq depth=
let expand_constructor_hints =
List.map_append (function
- | IndRef ind ->
+ | GlobRef.IndRef ind ->
List.init (Inductiveops.nconstructors ind)
- (fun i -> ConstructRef (ind,i+1))
+ (fun i -> GlobRef.ConstructRef (ind,i+1))
| gr ->
[gr])
@@ -197,7 +197,7 @@ let extend_with_ref_list env sigma l seq =
let l = expand_constructor_hints l in
let f gr (seq, sigma) =
let sigma, c = Evd.fresh_global env sigma gr in
- let sigma, typ= Typing.type_of env sigma (EConstr.of_constr c) in
+ let sigma, typ= Typing.type_of env sigma c in
(add_formula env sigma Hyp gr typ seq, sigma) in
List.fold_right f l (seq, sigma)
diff --git a/plugins/firstorder/sequent.mli b/plugins/firstorder/sequent.mli
index c4ed3e21f..709b278ec 100644
--- a/plugins/firstorder/sequent.mli
+++ b/plugins/firstorder/sequent.mli
@@ -8,26 +8,26 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
+open Names
open EConstr
open Formula
-open Globnames
module CM: CSig.MapS with type key=Constr.t
-type h_item = global_reference * (int*Constr.t) option
+type h_item = GlobRef.t * (int*Constr.t) option
module History: Set.S with type elt = h_item
-val cm_add : Evd.evar_map -> constr -> global_reference -> global_reference list CM.t ->
- global_reference list CM.t
+val cm_add : Evd.evar_map -> constr -> GlobRef.t -> GlobRef.t list CM.t ->
+ GlobRef.t list CM.t
-val cm_remove : Evd.evar_map -> constr -> global_reference -> global_reference list CM.t ->
- global_reference list CM.t
+val cm_remove : Evd.evar_map -> constr -> GlobRef.t -> GlobRef.t list CM.t ->
+ GlobRef.t list CM.t
module HP: Heap.S with type elt=Formula.t
type t = {redexes:HP.t;
- context: global_reference list CM.t;
+ context: GlobRef.t list CM.t;
latoms:constr list;
gl:types;
glatom:constr option;
@@ -41,20 +41,20 @@ val record: h_item -> t -> t
val lookup: Evd.evar_map -> h_item -> t -> bool
-val add_formula : Environ.env -> Evd.evar_map -> side -> global_reference -> constr -> t -> t
+val add_formula : Environ.env -> Evd.evar_map -> side -> GlobRef.t -> constr -> t -> t
val re_add_formula_list : Evd.evar_map -> Formula.t list -> t -> t
-val find_left : Evd.evar_map -> constr -> t -> global_reference
+val find_left : Evd.evar_map -> constr -> t -> GlobRef.t
val take_formula : Evd.evar_map -> t -> Formula.t * t
val empty_seq : int -> t
-val extend_with_ref_list : Environ.env -> Evd.evar_map -> global_reference list ->
+val extend_with_ref_list : Environ.env -> Evd.evar_map -> GlobRef.t list ->
t -> t * Evd.evar_map
val extend_with_auto_hints : Environ.env -> Evd.evar_map -> Hints.hint_db_name list ->
t -> t * Evd.evar_map
-val print_cmap: global_reference list CM.t -> Pp.t
+val print_cmap: GlobRef.t list CM.t -> Pp.t
diff --git a/plugins/fourier/fourierR.ml b/plugins/fourier/fourierR.ml
index b1c003de2..0ea70c19f 100644
--- a/plugins/fourier/fourierR.ml
+++ b/plugins/fourier/fourierR.ml
@@ -227,7 +227,7 @@ let ineq1_of_constr (h,t) =
hstrict=false}]
|_-> raise NoIneq)
| Ind ((kn,i),_) ->
- if not (eq_gr (IndRef(kn,i)) Coqlib.glob_eq) then raise NoIneq;
+ if not (GlobRef.equal (IndRef(kn,i)) Coqlib.glob_eq) then raise NoIneq;
let t0= args.(0) in
let t1= args.(1) in
let t2= args.(2) in
diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml
index d04887a48..8da0e1c4f 100644
--- a/plugins/funind/functional_principles_proofs.ml
+++ b/plugins/funind/functional_principles_proofs.ml
@@ -1050,8 +1050,7 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a
(Global.env ()) !evd
(Constrintern.locate_reference (qualid_of_ident equation_lemma_id))
in
- let res = EConstr.of_constr res in
- evd:=evd';
+ evd:=evd';
let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd res in
res
in
diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml
index 7a9bbd92c..04a23cdb9 100644
--- a/plugins/funind/functional_principles_types.ml
+++ b/plugins/funind/functional_principles_types.ml
@@ -266,7 +266,7 @@ let change_property_sort evd toSort princ princName =
(Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident princName)) in
let init =
let nargs = (princ_info.nparams + (List.length princ_info.predicates)) in
- mkApp(princName_as_constr,
+ mkApp(EConstr.Unsafe.to_constr princName_as_constr,
Array.init nargs
(fun i -> mkRel (nargs - i )))
in
@@ -628,16 +628,22 @@ let build_scheme fas =
user_err ~hdr:"FunInd.build_scheme"
(str "Cannot find " ++ Libnames.pr_reference f)
in
- let evd',f = Evd.fresh_global (Global.env ()) !evd f_as_constant in
+ let evd',f = Evd.fresh_global (Global.env ()) !evd f_as_constant in
let _ = evd := evd' in
- let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd (EConstr.of_constr f) in
- (destConst f,sort)
- )
+ let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd f in
+ let c, u =
+ try EConstr.destConst !evd f
+ with DestKO ->
+ user_err Pp.(pr_econstr_env (Global.env ()) !evd f ++spc () ++ str "should be the named of a globally defined function")
+ in
+ (c, EConstr.EInstance.kind !evd u), sort
+ )
fas
) in
let bodies_types =
make_scheme evd pconstants
in
+
List.iter2
(fun (princ_id,_,_) def_entry ->
ignore
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index 319b410df..3ba3bafa4 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -885,7 +885,7 @@ let is_res r = match DAst.get r with
| _ -> false
let is_gr c gr = match DAst.get c with
-| GRef (r, _) -> Globnames.eq_gr r gr
+| GRef (r, _) -> GlobRef.equal r gr
| _ -> false
let is_gvar c = match DAst.get c with
@@ -894,7 +894,7 @@ let is_gvar c = match DAst.get c with
let same_raw_term rt1 rt2 =
match DAst.get rt1, DAst.get rt2 with
- | GRef(r1,_), GRef (r2,_) -> Globnames.eq_gr r1 r2
+ | GRef(r1,_), GRef (r2,_) -> GlobRef.equal r1 r2
| GHole _, GHole _ -> true
| _ -> false
let decompose_raw_eq lhs rhs =
diff --git a/plugins/funind/glob_termops.ml b/plugins/funind/glob_termops.ml
index 40ea40b6b..ae238b846 100644
--- a/plugins/funind/glob_termops.ml
+++ b/plugins/funind/glob_termops.ml
@@ -563,7 +563,8 @@ let resolve_and_replace_implicits ?(flags=Pretyping.all_and_fail_flags) ?(expect
(* FIXME : JF (30/03/2017) I'm not completely sure to have split understand as needed.
If someone knows how to prevent solved existantial removal in understand, please do not hesitate to change the computation of [ctx] here *)
let ctx,_,_ = Pretyping.ise_pretype_gen flags env sigma Glob_ops.empty_lvar expected_type rt in
- let ctx, f = Evarutil.nf_evars_and_universes ctx in
+ let ctx = Evd.minimize_universes ctx in
+ let f c = EConstr.of_constr (Evarutil.nf_evars_universes ctx (EConstr.Unsafe.to_constr c)) in
(* then we map [rt] to replace the implicit holes by their values *)
let rec change rt =
@@ -575,7 +576,7 @@ If someone knows how to prevent solved existantial removal in understand, pleas
(fun _ evi _ ->
match evi.evar_source with
| (loc_evi,ImplicitArg(gr_evi,p_evi,b_evi)) ->
- if Globnames.eq_gr grk gr_evi && pk=p_evi && bk=b_evi && rt.CAst.loc = loc_evi
+ if GlobRef.equal grk gr_evi && pk=p_evi && bk=b_evi && rt.CAst.loc = loc_evi
then raise (Found evi)
| _ -> ()
)
@@ -586,8 +587,8 @@ If someone knows how to prevent solved existantial removal in understand, pleas
with Found evi -> (* we found the evar corresponding to this hole *)
match evi.evar_body with
| Evar_defined c ->
- (* we just have to lift the solution in glob_term *)
- Detyping.detype Detyping.Now false Id.Set.empty env ctx (EConstr.of_constr (f c))
+ (* we just have to lift the solution in glob_term *)
+ Detyping.detype Detyping.Now false Id.Set.empty env ctx (f c)
| Evar_empty -> rt (* the hole was not solved : we do nothing *)
)
| (GHole(BinderType na,_,_)) -> (* we only want to deal with implicit arguments *)
@@ -609,7 +610,7 @@ If someone knows how to prevent solved existantial removal in understand, pleas
match evi.evar_body with
| Evar_defined c ->
(* we just have to lift the solution in glob_term *)
- Detyping.detype Detyping.Now false Id.Set.empty env ctx (EConstr.of_constr (f c))
+ Detyping.detype Detyping.Now false Id.Set.empty env ctx (f c)
| Evar_empty -> rt (* the hole was not solved : we d when falseo nothing *)
in
res
diff --git a/plugins/funind/glob_termops.mli b/plugins/funind/glob_termops.mli
index 7088ae596..481a8be3b 100644
--- a/plugins/funind/glob_termops.mli
+++ b/plugins/funind/glob_termops.mli
@@ -13,7 +13,7 @@ val pattern_to_term : cases_pattern -> glob_constr
Some basic functions to rebuild glob_constr
In each of them the location is Util.Loc.ghost
*)
-val mkGRef : Globnames.global_reference -> glob_constr
+val mkGRef : GlobRef.t -> glob_constr
val mkGVar : Id.t -> glob_constr
val mkGApp : glob_constr*(glob_constr list) -> glob_constr
val mkGLambda : Name.t * glob_constr * glob_constr -> glob_constr
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index 9c350483b..7df57b577 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -77,8 +77,7 @@ let functional_induction with_clean c princl pat =
user_err (str "Cannot find induction principle for "
++ Printer.pr_leconstr_env (Tacmach.pf_env g) sigma (mkConst c') )
in
- let princ = EConstr.of_constr princ in
- (princ,NoBindings,Tacmach.pf_unsafe_type_of g' princ,g')
+ (princ,NoBindings,Tacmach.pf_unsafe_type_of g' princ,g')
| _ -> raise (UserError(None,str "functional induction must be used with a function" ))
end
| Some ((princ,binding)) ->
@@ -91,10 +90,19 @@ let functional_induction with_clean c princl pat =
if princ_infos.Tactics.farg_in_concl
then [c] else []
in
+ if List.length args + List.length c_list = 0
+ then user_err Pp.(str "Cannot recognize a valid functional scheme" );
let encoded_pat_as_patlist =
- List.make (List.length args + List.length c_list - 1) None @ [pat] in
- List.map2 (fun c pat -> ((None,Ltac_plugin.Tacexpr.ElimOnConstr (fun env sigma -> (sigma,(c,NoBindings)) )),(None,pat),None))
- (args@c_list) encoded_pat_as_patlist
+ List.make (List.length args + List.length c_list - 1) None @ [pat]
+ in
+ List.map2
+ (fun c pat ->
+ ((None,
+ Ltac_plugin.Tacexpr.ElimOnConstr (fun env sigma -> (sigma,(c,NoBindings)))),
+ (None,pat),
+ None))
+ (args@c_list)
+ encoded_pat_as_patlist
in
let princ' = Some (princ,bindings) in
let princ_vars =
@@ -252,7 +260,6 @@ let derive_inversion fix_names =
let evd,c =
Evd.fresh_global
(Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident id)) in
- let c = EConstr.of_constr c in
let (cst, u) = destConst evd c in
evd, (cst, EInstance.kind evd u) :: l
)
@@ -274,8 +281,7 @@ let derive_inversion fix_names =
(Global.env ()) evd
(Constrintern.locate_reference (Libnames.qualid_of_ident (mk_rel_id id)))
in
- let id = EConstr.of_constr id in
- evd,(fst (destInd evd id))::l
+ evd,(fst (destInd evd id))::l
)
fix_names
(evd',[])
@@ -379,7 +385,7 @@ let generate_principle (evd:Evd.evar_map ref) pconstants on_error
let evd = ref (Evd.from_env env) in
let evd',uprinc = Evd.fresh_global env !evd princ in
let _ = evd := evd' in
- let princ_type = Typing.e_type_of ~refresh:true env evd (EConstr.of_constr uprinc) in
+ let princ_type = Typing.e_type_of ~refresh:true env evd uprinc in
let princ_type = EConstr.Unsafe.to_constr princ_type in
Functional_principles_types.generate_functional_principle
evd
@@ -416,7 +422,6 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp
let evd,c =
Evd.fresh_global
(Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident fname)) in
- let c = EConstr.of_constr c in
let (cst, u) = destConst evd c in
let u = EInstance.kind evd u in
evd,((cst, u) :: l)
@@ -433,7 +438,6 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp
let evd,c =
Evd.fresh_global
(Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident fname)) in
- let c = EConstr.of_constr c in
let (cst, u) = destConst evd c in
let u = EInstance.kind evd u in
evd,((cst, u) :: l)
@@ -842,7 +846,7 @@ let rec get_args b t : Constrexpr.local_binder_expr list *
| _ -> [],b,t
-let make_graph (f_ref:global_reference) =
+let make_graph (f_ref : GlobRef.t) =
let c,c_body =
match f_ref with
| ConstRef c ->
diff --git a/plugins/funind/indfun.mli b/plugins/funind/indfun.mli
index dcc1c2ea6..24304e361 100644
--- a/plugins/funind/indfun.mli
+++ b/plugins/funind/indfun.mli
@@ -1,3 +1,4 @@
+open Names
open Misctypes
val warn_cannot_define_graph : ?loc:Loc.t -> Pp.t * Pp.t -> unit
@@ -18,4 +19,4 @@ val functional_induction :
Goal.goal Evd.sigma -> Goal.goal list Evd.sigma
-val make_graph : Globnames.global_reference -> unit
+val make_graph : GlobRef.t -> unit
diff --git a/plugins/funind/indfun_common.mli b/plugins/funind/indfun_common.mli
index 5cc7163aa..346b21ef2 100644
--- a/plugins/funind/indfun_common.mli
+++ b/plugins/funind/indfun_common.mli
@@ -41,7 +41,7 @@ val chop_rprod_n : int -> Glob_term.glob_constr ->
val def_of_const : Constr.t -> Constr.t
val eq : EConstr.constr Lazy.t
val refl_equal : EConstr.constr Lazy.t
-val const_of_id: Id.t -> Globnames.global_reference(* constantyes *)
+val const_of_id: Id.t -> GlobRef.t(* constantyes *)
val jmeq : unit -> EConstr.constr
val jmeq_refl : unit -> EConstr.constr
@@ -107,11 +107,11 @@ val h_intros: Names.Id.t list -> Tacmach.tactic
val h_id : Names.Id.t
val hrec_id : Names.Id.t
val acc_inv_id : EConstr.constr Util.delayed
-val ltof_ref : Globnames.global_reference Util.delayed
+val ltof_ref : GlobRef.t Util.delayed
val well_founded_ltof : EConstr.constr Util.delayed
val acc_rel : EConstr.constr Util.delayed
val well_founded : EConstr.constr Util.delayed
-val evaluable_of_global_reference : Globnames.global_reference -> Names.evaluable_global_reference
+val evaluable_of_global_reference : GlobRef.t -> Names.evaluable_global_reference
val list_rewrite : bool -> (EConstr.constr*bool) list -> Tacmach.tactic
val decompose_lam_n : Evd.evar_map -> int -> EConstr.t ->
diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml
index ae84eaa93..28e85268a 100644
--- a/plugins/funind/invfun.ml
+++ b/plugins/funind/invfun.ml
@@ -102,7 +102,6 @@ let generate_type evd g_to_f f graph i =
let evd',graph =
Evd.fresh_global (Global.env ()) !evd (Globnames.IndRef (fst (destInd !evd graph)))
in
- let graph = EConstr.of_constr graph in
evd:=evd';
let graph_arity = Typing.e_type_of (Global.env ()) evd graph in
let ctxt,_ = decompose_prod_assum !evd graph_arity in
@@ -172,7 +171,6 @@ let find_induction_principle evd f =
| None -> raise Not_found
| Some rect_lemma ->
let evd',rect_lemma = Evd.fresh_global (Global.env ()) !evd (Globnames.ConstRef rect_lemma) in
- let rect_lemma = EConstr.of_constr rect_lemma in
let evd',typ = Typing.type_of ~refresh:true (Global.env ()) evd' rect_lemma in
evd:=evd';
rect_lemma,typ
@@ -823,8 +821,7 @@ let derive_correctness make_scheme (funs: pconstant list) (graphs:inductive list
(* let lem_cst = fst (destConst (Constrintern.global_reference lem_id)) in *)
let _,lem_cst_constr = Evd.fresh_global
(Global.env ()) !evd (Constrintern.locate_reference (Libnames.qualid_of_ident lem_id)) in
- let lem_cst_constr = EConstr.of_constr lem_cst_constr in
- let (lem_cst,_) = destConst !evd lem_cst_constr in
+ let (lem_cst,_) = destConst !evd lem_cst_constr in
update_Function {finfo with correctness_lemma = Some lem_cst};
)
@@ -884,8 +881,7 @@ let derive_correctness make_scheme (funs: pconstant list) (graphs:inductive list
let finfo = find_Function_infos (fst f_as_constant) in
let _,lem_cst_constr = Evd.fresh_global
(Global.env ()) !evd (Constrintern.locate_reference (Libnames.qualid_of_ident lem_id)) in
- let lem_cst_constr = EConstr.of_constr lem_cst_constr in
- let (lem_cst,_) = destConst !evd lem_cst_constr in
+ let (lem_cst,_) = destConst !evd lem_cst_constr in
update_Function {finfo with completeness_lemma = Some lem_cst}
)
funs)
diff --git a/plugins/funind/invfun.mli b/plugins/funind/invfun.mli
index ad306ab25..9151fd0e2 100644
--- a/plugins/funind/invfun.mli
+++ b/plugins/funind/invfun.mli
@@ -10,7 +10,7 @@
val invfun :
Misctypes.quantified_hypothesis ->
- Globnames.global_reference option ->
+ Names.GlobRef.t option ->
Evar.t Evd.sigma -> Evar.t list Evd.sigma
val derive_correctness :
(Evd.evar_map ref ->
diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml
index fb9ae64bf..2a3a85fcc 100644
--- a/plugins/funind/recdef.ml
+++ b/plugins/funind/recdef.ml
@@ -181,7 +181,7 @@ let simpl_iter clause =
clause
(* Others ugly things ... *)
-let (value_f: Constr.t list -> global_reference -> Constr.t) =
+let (value_f: Constr.t list -> GlobRef.t -> Constr.t) =
let open Term in
let open Constr in
fun al fterm ->
@@ -215,7 +215,7 @@ let (value_f: Constr.t list -> global_reference -> Constr.t) =
let body = EConstr.Unsafe.to_constr body in
it_mkLambda_or_LetIn body context
-let (declare_f : Id.t -> logical_kind -> Constr.t list -> global_reference -> global_reference) =
+let (declare_f : Id.t -> logical_kind -> Constr.t list -> GlobRef.t -> GlobRef.t) =
fun f_id kind input_type fterm_ref ->
declare_fun f_id kind (value_f input_type fterm_ref);;
@@ -356,7 +356,7 @@ type 'a infos =
f_id : Id.t; (* function name *)
f_constr : constr; (* function term *)
f_terminate : constr; (* termination proof term *)
- func : global_reference; (* functional reference *)
+ func : GlobRef.t; (* functional reference *)
info : 'a;
is_main_branch : bool; (* on the main branch or on a matched expression *)
is_final : bool; (* final first order term or not *)
@@ -1456,7 +1456,7 @@ let com_terminate
-let start_equation (f:global_reference) (term_f:global_reference)
+let start_equation (f:GlobRef.t) (term_f:GlobRef.t)
(cont_tactic:Id.t list -> tactic) g =
let sigma = project g in
let ids = pf_ids_of_hyps g in
@@ -1473,7 +1473,7 @@ let start_equation (f:global_reference) (term_f:global_reference)
observe_tac (str "prove_eq") (cont_tactic x)]) g;;
let (com_eqn : int -> Id.t ->
- global_reference -> global_reference -> global_reference
+ GlobRef.t -> GlobRef.t -> GlobRef.t
-> Constr.t -> unit) =
fun nb_arg eq_name functional_ref f_ref terminate_ref equation_lemma_type ->
let open CVars in
@@ -1533,14 +1533,12 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
let env = Global.env() in
let evd = Evd.from_env env in
let evd, function_type = interp_type_evars env evd type_of_f in
- let function_type = EConstr.Unsafe.to_constr function_type in
- let env = push_named (Context.Named.Declaration.LocalAssum (function_name,function_type)) env in
+ let env = EConstr.push_named (Context.Named.Declaration.LocalAssum (function_name,function_type)) env in
(* Pp.msgnl (str "function type := " ++ Printer.pr_lconstr function_type); *)
let evd, ty = interp_type_evars env evd ~impls:rec_impls eq in
- let ty = EConstr.Unsafe.to_constr ty in
- let evd, nf = Evarutil.nf_evars_and_universes evd in
- let equation_lemma_type = nf_betaiotazeta (EConstr.of_constr (nf ty)) in
- let function_type = nf function_type in
+ let evd = Evd.minimize_universes evd in
+ let equation_lemma_type = nf_betaiotazeta (Evarutil.nf_evar evd ty) in
+ let function_type = EConstr.to_constr ~abort_on_undefined_evars:false evd function_type in
let equation_lemma_type = EConstr.Unsafe.to_constr equation_lemma_type in
(* Pp.msgnl (str "lemma type := " ++ Printer.pr_lconstr equation_lemma_type ++ fnl ()); *)
let res_vars,eq' = decompose_prod equation_lemma_type in
diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4
index 2e90ce90c..797dfbe23 100644
--- a/plugins/ltac/extratactics.ml4
+++ b/plugins/ltac/extratactics.ml4
@@ -296,7 +296,6 @@ let project_hint ~poly pri l2r r =
let env = Global.env() in
let sigma = Evd.from_env env in
let sigma, c = Evd.fresh_global env sigma gr in
- let c = EConstr.of_constr c in
let t = Retyping.get_type_of env sigma c in
let t =
Tacred.reduce_to_quantified_ref env sigma (Lazy.force coq_iff_ref) t in
@@ -307,7 +306,6 @@ let project_hint ~poly pri l2r r =
let p =
if l2r then build_coq_iff_left_proj () else build_coq_iff_right_proj () in
let sigma, p = Evd.fresh_global env sigma p in
- let p = EConstr.of_constr p in
let c = Reductionops.whd_beta sigma (mkApp (c, Context.Rel.to_extended_vect mkRel 0 sign)) in
let c = it_mkLambda_or_LetIn
(mkApp (p,[|mkArrow a (lift 1 b);mkArrow b (lift 1 a);c|])) sign in
@@ -317,7 +315,7 @@ let project_hint ~poly pri l2r r =
let ctx = Evd.const_univ_entry ~poly sigma in
let c = EConstr.to_constr sigma c in
let c = Declare.declare_definition ~internal:Declare.InternalTacticRequest id (c,ctx) in
- let info = {Vernacexpr.hint_priority = pri; hint_pattern = None} in
+ let info = {Typeclasses.hint_priority = pri; hint_pattern = None} in
(info,false,true,Hints.PathAny, Hints.IsGlobRef (Globnames.ConstRef c))
let add_hints_iff ~atts l2r lc n bl =
diff --git a/plugins/ltac/g_ltac.ml4 b/plugins/ltac/g_ltac.ml4
index 0c42a8bb2..4857beffa 100644
--- a/plugins/ltac/g_ltac.ml4
+++ b/plugins/ltac/g_ltac.ml4
@@ -325,6 +325,7 @@ GEXTEND Gram
;
toplevel_selector:
[ [ sel = selector_body; ":" -> sel
+ | "!"; ":" -> SelectAlreadyFocused
| IDENT "all"; ":" -> SelectAll ] ]
;
tactic_mode:
@@ -415,7 +416,7 @@ let is_explicit_terminator = function TacSolve _ -> true | _ -> false
VERNAC tactic_mode EXTEND VernacSolve
| [ - ltac_selector_opt(g) ltac_info_opt(n) tactic(t) ltac_use_default(def) ] =>
[ classify_as_proofstep ] -> [
- let g = Option.default (Proof_bullet.get_default_goal_selector ()) g in
+ let g = Option.default (Goal_select.get_default_goal_selector ()) g in
vernac_solve g n t def
]
| [ - "par" ":" ltac_info_opt(n) tactic(t) ltac_use_default(def) ] =>
diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml
index 11bb7a234..bd02d85d5 100644
--- a/plugins/ltac/pptactic.ml
+++ b/plugins/ltac/pptactic.ml
@@ -515,6 +515,7 @@ let string_of_genarg_arg (ArgumentType arg) =
else int i ++ str "-" ++ int j
let pr_goal_selector toplevel = function
+ | SelectAlreadyFocused -> str "!:"
| SelectNth i -> int i ++ str ":"
| SelectList l -> prlist_with_sep (fun () -> str ", ") pr_range_selector l ++ str ":"
| SelectId id -> str "[" ++ Id.print id ++ str "]:"
diff --git a/plugins/ltac/pptactic.mli b/plugins/ltac/pptactic.mli
index aea00c240..799a52cc8 100644
--- a/plugins/ltac/pptactic.mli
+++ b/plugins/ltac/pptactic.mli
@@ -84,7 +84,7 @@ type pp_tactic = {
pptac_prods : grammar_terminals;
}
-val pr_goal_selector : toplevel:bool -> Vernacexpr.goal_selector -> Pp.t
+val pr_goal_selector : toplevel:bool -> Goal_select.t -> Pp.t
val declare_notation_tactic_pprule : KerName.t -> pp_tactic -> unit
diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml
index d32a2faef..9eb55aa5e 100644
--- a/plugins/ltac/rewrite.ml
+++ b/plugins/ltac/rewrite.ml
@@ -428,7 +428,8 @@ let split_head = function
| [] -> assert(false)
let eq_pb (ty, env, x, y as pb) (ty', env', x', y' as pb') =
- pb == pb' || (ty == ty' && Constr.equal x x' && Constr.equal y y')
+ let equal x y = Constr.equal (EConstr.Unsafe.to_constr x) (EConstr.Unsafe.to_constr y) in
+ pb == pb' || (ty == ty' && equal x x' && equal y y')
let problem_inclusion x y =
List.for_all (fun pb -> List.exists (fun pb' -> eq_pb pb pb') y) x
@@ -626,9 +627,9 @@ let solve_remaining_by env sigma holes by =
(** Evar should not be defined, but just in case *)
| Some evi ->
let env = Environ.reset_with_named_context evi.evar_hyps env in
- let ty = EConstr.of_constr evi.evar_concl in
+ let ty = evi.evar_concl in
let c, sigma = Pfedit.refine_by_tactic env sigma ty solve_tac in
- Evd.define evk c sigma
+ Evd.define evk (EConstr.of_constr c) sigma
in
List.fold_left solve sigma indep
@@ -1862,7 +1863,6 @@ let declare_projection n instance_id r =
let env = Global.env () in
let sigma = Evd.from_env env in
let sigma,c = Evd.fresh_global env sigma r in
- let c = EConstr.of_constr c in
let ty = Retyping.get_type_of env sigma c in
let term = proper_projection sigma c ty in
let sigma, typ = Typing.type_of env sigma term in
diff --git a/plugins/ltac/taccoerce.mli b/plugins/ltac/taccoerce.mli
index 1fa5e3c07..5185217cd 100644
--- a/plugins/ltac/taccoerce.mli
+++ b/plugins/ltac/taccoerce.mli
@@ -80,7 +80,7 @@ val coerce_to_hyp : Environ.env -> Evd.evar_map -> Value.t -> Id.t
val coerce_to_hyp_list : Environ.env -> Evd.evar_map -> Value.t -> Id.t list
-val coerce_to_reference : Environ.env -> Evd.evar_map -> Value.t -> Globnames.global_reference
+val coerce_to_reference : Environ.env -> Evd.evar_map -> Value.t -> GlobRef.t
val coerce_to_quantified_hypothesis : Evd.evar_map -> Value.t -> quantified_hypothesis
diff --git a/plugins/ltac/tacexpr.ml b/plugins/ltac/tacexpr.ml
index 3baa475ab..17f5e5d41 100644
--- a/plugins/ltac/tacexpr.ml
+++ b/plugins/ltac/tacexpr.ml
@@ -35,7 +35,8 @@ type advanced_flag = bool (* true = advanced false = basic *)
type letin_flag = bool (* true = use local def false = use Leibniz *)
type clear_flag = bool option (* true = clear hyp, false = keep hyp, None = use default *)
-type goal_selector = Vernacexpr.goal_selector =
+type goal_selector = Goal_select.t =
+ | SelectAlreadyFocused
| SelectNth of int
| SelectList of (int * int) list
| SelectId of Id.t
@@ -269,7 +270,7 @@ and 'a gen_tactic_expr =
('p,'a gen_tactic_expr) match_rule list
| TacFun of 'a gen_tactic_fun_ast
| TacArg of 'a gen_tactic_arg located
- | TacSelect of Vernacexpr.goal_selector * 'a gen_tactic_expr
+ | TacSelect of Goal_select.t * 'a gen_tactic_expr
(* For ML extensions *)
| TacML of (ml_tactic_entry * 'a gen_tactic_arg list) Loc.located
(* For syntax extensions *)
diff --git a/plugins/ltac/tacexpr.mli b/plugins/ltac/tacexpr.mli
index 3baa475ab..17f5e5d41 100644
--- a/plugins/ltac/tacexpr.mli
+++ b/plugins/ltac/tacexpr.mli
@@ -35,7 +35,8 @@ type advanced_flag = bool (* true = advanced false = basic *)
type letin_flag = bool (* true = use local def false = use Leibniz *)
type clear_flag = bool option (* true = clear hyp, false = keep hyp, None = use default *)
-type goal_selector = Vernacexpr.goal_selector =
+type goal_selector = Goal_select.t =
+ | SelectAlreadyFocused
| SelectNth of int
| SelectList of (int * int) list
| SelectId of Id.t
@@ -269,7 +270,7 @@ and 'a gen_tactic_expr =
('p,'a gen_tactic_expr) match_rule list
| TacFun of 'a gen_tactic_fun_ast
| TacArg of 'a gen_tactic_arg located
- | TacSelect of Vernacexpr.goal_selector * 'a gen_tactic_expr
+ | TacSelect of Goal_select.t * 'a gen_tactic_expr
(* For ML extensions *)
| TacML of (ml_tactic_entry * 'a gen_tactic_arg list) Loc.located
(* For syntax extensions *)
diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml
index 6a4bf577b..84049d4ed 100644
--- a/plugins/ltac/tacinterp.ml
+++ b/plugins/ltac/tacinterp.ml
@@ -2010,7 +2010,8 @@ let interp_redexp env sigma r =
let _ =
let eval lfun env sigma ty tac =
- let ist = { lfun = lfun; extra = TacStore.empty; } in
+ let extra = TacStore.set TacStore.empty f_debug (get_debug ()) in
+ let ist = { lfun = lfun; extra; } in
let tac = interp_tactic ist tac in
let (c, sigma) = Pfedit.refine_by_tactic env sigma ty tac in
(EConstr.of_constr c, sigma)
diff --git a/plugins/ltac/tactic_debug.ml b/plugins/ltac/tactic_debug.ml
index 57a11d947..105b5c59a 100644
--- a/plugins/ltac/tactic_debug.ml
+++ b/plugins/ltac/tactic_debug.ml
@@ -399,8 +399,6 @@ let skip_extensions trace =
| [] -> [] in
List.rev (aux (List.rev trace))
-let finer_loc loc1 loc2 = Loc.merge_opt loc1 loc2 = loc2
-
let extract_ltac_trace ?loc trace =
let trace = skip_extensions trace in
let (tloc,c),tail = List.sep_last trace in
@@ -408,7 +406,7 @@ let extract_ltac_trace ?loc trace =
(* We entered a user-defined tactic,
we display the trace with location of the call *)
let msg = hov 0 (explain_ltac_call_trace c tail loc ++ fnl()) in
- (if finer_loc loc tloc then loc else tloc), Some msg
+ (if Loc.finer loc tloc then loc else tloc), Some msg
else
(* We entered a primitive tactic, we don't display trace but
report on the finest location *)
@@ -417,7 +415,7 @@ let extract_ltac_trace ?loc trace =
let rec aux best_loc = function
| (loc,_)::tail ->
if Option.is_empty best_loc ||
- not (Option.is_empty loc) && finer_loc loc best_loc
+ not (Option.is_empty loc) && Loc.finer loc best_loc
then
aux loc tail
else
diff --git a/plugins/setoid_ring/Algebra_syntax.v b/plugins/setoid_ring/Algebra_syntax.v
index e896554ea..1204bbd2e 100644
--- a/plugins/setoid_ring/Algebra_syntax.v
+++ b/plugins/setoid_ring/Algebra_syntax.v
@@ -1,3 +1,12 @@
+(************************************************************************)
+(* * 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 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
Class Zero (A : Type) := zero : A.
Notation "0" := zero.
diff --git a/plugins/setoid_ring/Integral_domain.v b/plugins/setoid_ring/Integral_domain.v
index 0c16fe1a3..98407cb6d 100644
--- a/plugins/setoid_ring/Integral_domain.v
+++ b/plugins/setoid_ring/Integral_domain.v
@@ -1,3 +1,13 @@
+(************************************************************************)
+(* * 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 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
Require Export Cring.
diff --git a/plugins/setoid_ring/RealField.v b/plugins/setoid_ring/RealField.v
index facd2e062..38bc58a65 100644
--- a/plugins/setoid_ring/RealField.v
+++ b/plugins/setoid_ring/RealField.v
@@ -1,3 +1,13 @@
+(************************************************************************)
+(* * 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 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
Require Import Nnat.
Require Import ArithRing.
Require Export Ring Field.
diff --git a/plugins/setoid_ring/Ring_tac.v b/plugins/setoid_ring/Ring_tac.v
index 36d1e7c54..e8efb362e 100644
--- a/plugins/setoid_ring/Ring_tac.v
+++ b/plugins/setoid_ring/Ring_tac.v
@@ -1,3 +1,13 @@
+(************************************************************************)
+(* * 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 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
Set Implicit Arguments.
Require Import Setoid.
Require Import BinPos.
diff --git a/plugins/setoid_ring/Rings_Q.v b/plugins/setoid_ring/Rings_Q.v
index fd7654713..ae91ee166 100644
--- a/plugins/setoid_ring/Rings_Q.v
+++ b/plugins/setoid_ring/Rings_Q.v
@@ -1,3 +1,13 @@
+(************************************************************************)
+(* * 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 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
Require Export Cring.
Require Export Integral_domain.
diff --git a/plugins/setoid_ring/Rings_R.v b/plugins/setoid_ring/Rings_R.v
index fd219c235..901b36ed3 100644
--- a/plugins/setoid_ring/Rings_R.v
+++ b/plugins/setoid_ring/Rings_R.v
@@ -1,3 +1,13 @@
+(************************************************************************)
+(* * 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 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
Require Export Cring.
Require Export Integral_domain.
diff --git a/plugins/setoid_ring/Rings_Z.v b/plugins/setoid_ring/Rings_Z.v
index 605a23a98..75e77ab6e 100644
--- a/plugins/setoid_ring/Rings_Z.v
+++ b/plugins/setoid_ring/Rings_Z.v
@@ -1,3 +1,13 @@
+(************************************************************************)
+(* * 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 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
Require Export Cring.
Require Export Integral_domain.
Require Export Ncring_initial.
diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml
index 99bb8440c..5facf2a80 100644
--- a/plugins/setoid_ring/newring.ml
+++ b/plugins/setoid_ring/newring.ml
@@ -78,7 +78,7 @@ and mk_clos_app_but f_map subs f args n =
| None -> mk_atom (mkApp (f, args))
let interp_map l t =
- try Some(List.assoc_f eq_gr t l) with Not_found -> None
+ try Some(List.assoc_f GlobRef.equal t l) with Not_found -> None
let protect_maps = ref String.Map.empty
let add_map s m = protect_maps := String.Map.add s m !protect_maps
@@ -186,8 +186,8 @@ let dummy_goal env sigma =
Goal.V82.mk_goal sigma (named_context_val env) EConstr.mkProp Evd.Store.empty in
{Evd.it = gl; Evd.sigma = sigma}
-let constr_of v = match Value.to_constr v with
- | Some c -> EConstr.Unsafe.to_constr c
+let constr_of evd v = match Value.to_constr v with
+ | Some c -> EConstr.to_constr evd c
| None -> failwith "Ring.exec_tactic: anomaly"
let tactic_res = ref [||]
@@ -221,8 +221,8 @@ let exec_tactic env evd n f args =
(** Evaluate the whole result *)
let gl = dummy_goal env evd in
let gls = Proofview.V82.of_tactic (Tacinterp.eval_tactic_ist ist (ltac_call f (args@[getter]))) gl in
- let evd, nf = Evarutil.nf_evars_and_universes (Refiner.project gls) in
- let nf c = nf (constr_of c) in
+ let evd = Evd.minimize_universes (Refiner.project gls) in
+ let nf c = constr_of evd c in
Array.map nf !tactic_res, Evd.universe_context_set evd
let stdlib_modules =
diff --git a/plugins/setoid_ring/newring.mli b/plugins/setoid_ring/newring.mli
index 1d1557b12..0e056a472 100644
--- a/plugins/setoid_ring/newring.mli
+++ b/plugins/setoid_ring/newring.mli
@@ -11,7 +11,6 @@
open Names
open EConstr
open Libnames
-open Globnames
open Constrexpr
open Newring_ast
@@ -19,7 +18,7 @@ val protect_tac_in : string -> Id.t -> unit Proofview.tactic
val protect_tac : string -> unit Proofview.tactic
-val closed_term : EConstr.constr -> global_reference list -> unit Proofview.tactic
+val closed_term : EConstr.constr -> GlobRef.t list -> unit Proofview.tactic
val add_theory :
Id.t ->
diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml
index d5118da4c..e9e045a53 100644
--- a/plugins/ssr/ssrcommon.ml
+++ b/plugins/ssr/ssrcommon.ml
@@ -504,16 +504,17 @@ let nf_evar sigma t =
EConstr.Unsafe.to_constr (Evarutil.nf_evar sigma (EConstr.of_constr t))
let pf_abs_evars2 gl rigid (sigma, c0) =
- let c0 = EConstr.to_constr sigma c0 in
+ let c0 = EConstr.to_constr ~abort_on_undefined_evars:false sigma c0 in
let sigma0, ucst = project gl, Evd.evar_universe_context sigma in
let nenv = env_size (pf_env gl) in
let abs_evar n k =
let evi = Evd.find sigma k in
- let dc = CList.firstn n (evar_filtered_context evi) in
+ let concl = EConstr.Unsafe.to_constr evi.evar_concl in
+ let dc = EConstr.Unsafe.to_named_context (CList.firstn n (evar_filtered_context evi)) in
let abs_dc c = function
| NamedDecl.LocalDef (x,b,t) -> mkNamedLetIn x b t (mkArrow t c)
| NamedDecl.LocalAssum (x,t) -> mkNamedProd x t c in
- let t = Context.Named.fold_inside abs_dc ~init:evi.evar_concl dc in
+ let t = Context.Named.fold_inside abs_dc ~init:concl dc in
nf_evar sigma t in
let rec put evlist c = match Constr.kind c with
| Evar (k, a) ->
@@ -569,11 +570,12 @@ let pf_abs_evars_pirrel gl (sigma, c0) =
let nenv = env_size (pf_env gl) in
let abs_evar n k =
let evi = Evd.find sigma k in
- let dc = CList.firstn n (evar_filtered_context evi) in
+ let concl = EConstr.Unsafe.to_constr evi.evar_concl in
+ let dc = EConstr.Unsafe.to_named_context (CList.firstn n (evar_filtered_context evi)) in
let abs_dc c = function
| NamedDecl.LocalDef (x,b,t) -> mkNamedLetIn x b t (mkArrow t c)
| NamedDecl.LocalAssum (x,t) -> mkNamedProd x t c in
- let t = Context.Named.fold_inside abs_dc ~init:evi.evar_concl dc in
+ let t = Context.Named.fold_inside abs_dc ~init:concl dc in
nf_evar sigma0 (nf_evar sigma t) in
let rec put evlist c = match Constr.kind c with
| Evar (k, a) ->
@@ -581,7 +583,7 @@ let pf_abs_evars_pirrel gl (sigma, c0) =
let n = max 0 (Array.length a - nenv) in
let k_ty =
Retyping.get_sort_family_of
- (pf_env gl) sigma (EConstr.of_constr (Evd.evar_concl (Evd.find sigma k))) in
+ (pf_env gl) sigma (Evd.evar_concl (Evd.find sigma k)) in
let is_prop = k_ty = InProp in
let t = abs_evar n k in (k, (n, t, is_prop)) :: put evlist t
| _ -> Constr.fold put evlist c in
@@ -746,7 +748,7 @@ let pf_mkSsrConst name gl =
let pf_fresh_global name gl =
let sigma, env, it = project gl, pf_env gl, sig_it gl in
let sigma,t = Evd.fresh_global env sigma name in
- t, re_sig it sigma
+ EConstr.Unsafe.to_constr t, re_sig it sigma
let mkProt t c gl =
let prot, gl = pf_mkSsrConst "protect_term" gl in
@@ -800,8 +802,11 @@ let rec is_name_in_ipats name = function
List.exists (function SsrHyp(_,id) -> id = name) clr
|| is_name_in_ipats name tl
| IPatId id :: tl -> id = name || is_name_in_ipats name tl
- | (IPatCase l | IPatDispatch l) :: tl -> List.exists (is_name_in_ipats name) l || is_name_in_ipats name tl
- | _ :: tl -> is_name_in_ipats name tl
+ | IPatAbstractVars ids :: tl ->
+ CList.mem_f Id.equal name ids || is_name_in_ipats name tl
+ | (IPatCase l | IPatDispatch l | IPatInj l) :: tl ->
+ List.exists (is_name_in_ipats name) l || is_name_in_ipats name tl
+ | (IPatView _ | IPatAnon _ | IPatSimpl _ | IPatRewrite _ | IPatTac _ | IPatNoop) :: tl -> is_name_in_ipats name tl
| [] -> false
let view_error s gv =
@@ -980,7 +985,7 @@ let applyn ~with_evars ?beta ?(with_shelve=false) n t gl =
if not (EConstr.Vars.closed0 sigma ty) then
raise dependent_apply_error;
let m = Evarutil.new_meta () in
- loop (meta_declare m (EConstr.Unsafe.to_constr ty) sigma) bo ((EConstr.mkMeta m)::args) (n-1)
+ loop (meta_declare m ty sigma) bo ((EConstr.mkMeta m)::args) (n-1)
| _ -> assert false
in loop sigma t [] n in
pp(lazy(str"Refiner.refiner " ++ Printer.pr_econstr_env (pf_env gl) (project gl) t));
@@ -1500,7 +1505,7 @@ let tclOPTION o d =
let tacIS_INJECTION_CASE ?ty t = begin
tclOPTION ty (tacTYPEOF t) >>= fun ty ->
tacREDUCE_TO_QUANTIFIED_IND ty >>= fun ((mind,_),_) ->
- tclUNIT (Globnames.eq_gr (Globnames.IndRef mind) (Coqlib.build_coq_eq ()))
+ tclUNIT (GlobRef.equal (GlobRef.IndRef mind) (Coqlib.build_coq_eq ()))
end
let tclWITHTOP tac = Goal.enter begin fun gl ->
diff --git a/plugins/ssr/ssrcommon.mli b/plugins/ssr/ssrcommon.mli
index 2b8f1d540..9ba23467e 100644
--- a/plugins/ssr/ssrcommon.mli
+++ b/plugins/ssr/ssrcommon.mli
@@ -212,7 +212,7 @@ val pf_abs_prod :
EConstr.t -> Goal.goal Evd.sigma * EConstr.types
val mkSsrRRef : string -> Glob_term.glob_constr * 'a option
-val mkSsrRef : string -> Globnames.global_reference
+val mkSsrRef : string -> GlobRef.t
val mkSsrConst :
string ->
env -> evar_map -> evar_map * EConstr.t
@@ -224,7 +224,7 @@ val new_wild_id : tac_ctx -> Names.Id.t * tac_ctx
val pf_fresh_global :
- Globnames.global_reference ->
+ GlobRef.t ->
Goal.goal Evd.sigma ->
Constr.constr * Goal.goal Evd.sigma
diff --git a/plugins/ssr/ssrelim.ml b/plugins/ssr/ssrelim.ml
index 717657a24..87d107d65 100644
--- a/plugins/ssr/ssrelim.ml
+++ b/plugins/ssr/ssrelim.ml
@@ -356,7 +356,7 @@ let ssrelim ?(ind=ref None) ?(is_case=false) deps what ?elim eqid elim_intro_tac
let ev = List.fold_left Evar.Set.union Evar.Set.empty patterns_ev in
let ty_ev = Evar.Set.fold (fun i e ->
let ex = i in
- let i_ty = EConstr.of_constr (Evd.evar_concl (Evd.find (project gl) ex)) in
+ let i_ty = Evd.evar_concl (Evd.find (project gl) ex) in
Evar.Set.union e (evars_of_term i_ty))
ev Evar.Set.empty in
let inter = Evar.Set.inter ev ty_ev in
@@ -418,7 +418,7 @@ let injectl2rtac sigma c = match EConstr.kind sigma c with
let is_injection_case c gl =
let gl, cty = pfe_type_of gl c in
let (mind,_), _ = pf_reduce_to_quantified_ind gl cty in
- eq_gr (IndRef mind) (Coqlib.build_coq_eq ())
+ GlobRef.equal (IndRef mind) (Coqlib.build_coq_eq ())
let perform_injection c gl =
let gl, cty = pfe_type_of gl c in
diff --git a/plugins/ssr/ssrequality.ml b/plugins/ssr/ssrequality.ml
index 57635edac..7d7655d29 100644
--- a/plugins/ssr/ssrequality.ml
+++ b/plugins/ssr/ssrequality.ml
@@ -276,7 +276,7 @@ let unfoldintac occ rdx t (kt,_) gl =
let foldtac occ rdx ft gl =
let sigma0, concl0, env0 = project gl, pf_concl gl, pf_env gl in
let sigma, t = ft in
- let t = EConstr.to_constr sigma t in
+ let t = EConstr.to_constr ~abort_on_undefined_evars:false sigma t in
let fold, conclude = match rdx with
| Some (_, (In_T _ | In_X_In_T _)) | None ->
let ise = Evd.create_evar_defs sigma in
@@ -359,7 +359,7 @@ let pirrel_rewrite pred rdx rdx_ty new_rdx dir (sigma, c) c_ty gl =
let evs = Evar.Set.elements (Evarutil.undefined_evars_of_term sigma t) in
let open_evs = List.filter (fun k ->
Sorts.InProp <> Retyping.get_sort_family_of
- env sigma (EConstr.of_constr (Evd.evar_concl (Evd.find sigma k))))
+ env sigma (Evd.evar_concl (Evd.find sigma k)))
evs in
if open_evs <> [] then Some name else None)
(List.combine (Array.to_list args) names)
@@ -369,8 +369,8 @@ let pirrel_rewrite pred rdx rdx_ty new_rdx dir (sigma, c) c_ty gl =
;;
let is_construct_ref sigma c r =
- EConstr.isConstruct sigma c && eq_gr (ConstructRef (fst(EConstr.destConstruct sigma c))) r
-let is_ind_ref sigma c r = EConstr.isInd sigma c && eq_gr (IndRef (fst(EConstr.destInd sigma c))) r
+ EConstr.isConstruct sigma c && GlobRef.equal (ConstructRef (fst(EConstr.destConstruct sigma c))) r
+let is_ind_ref sigma c r = EConstr.isInd sigma c && GlobRef.equal (IndRef (fst(EConstr.destInd sigma c))) r
let rwcltac cl rdx dir sr gl =
let n, r_n,_, ucst = pf_abs_evars gl sr in
@@ -478,10 +478,10 @@ let rwprocess_rule dir rule gl =
| _ -> let ra = Array.append a [|r|] in
function 1 ->
let sigma, pi1 = Evd.fresh_global env sigma coq_prod.Coqlib.proj1 in
- EConstr.mkApp (EConstr.of_constr pi1, ra), sigma
+ EConstr.mkApp (pi1, ra), sigma
| _ ->
let sigma, pi2 = Evd.fresh_global env sigma coq_prod.Coqlib.proj2 in
- EConstr.mkApp (EConstr.of_constr pi2, ra), sigma in
+ EConstr.mkApp (pi2, ra), sigma in
if EConstr.eq_constr sigma a.(0) (EConstr.of_constr (Universes.constr_of_global @@ Coqlib.build_coq_True ())) then
let s, sigma = sr sigma 2 in
loop (converse_dir d) sigma s a.(1) rs 0
@@ -557,7 +557,7 @@ let rwrxtac occ rdx_pat dir rule gl =
let rpat env sigma0 (sigma, pats) (d, r, lhs, rhs) =
let sigma, pat =
let rw_progress rhs t evd = rw_progress rhs (EConstr.of_constr t) evd in
- mk_tpattern env sigma0 (sigma,EConstr.to_constr sigma r) (rw_progress rhs) d (EConstr.to_constr sigma lhs) in
+ mk_tpattern env sigma0 (sigma, EConstr.to_constr ~abort_on_undefined_evars:false sigma r) (rw_progress rhs) d (EConstr.to_constr ~abort_on_undefined_evars:false sigma lhs) in
sigma, pats @ [pat] in
let rpats = List.fold_left (rpat env0 sigma0) (r_sigma,[]) rules in
let find_R, end_R = mk_tpattern_matcher sigma0 occ ~upats_origin rpats in
@@ -567,7 +567,7 @@ let rwrxtac occ rdx_pat dir rule gl =
let r = ref None in
(fun env c _ h -> do_once r (fun () -> find_rule (EConstr.of_constr c), c); mkRel h),
(fun concl -> closed0_check concl e gl;
- let (d,(ev,ctx,c)) , x = assert_done r in (d,(ev,ctx, EConstr.to_constr ev c)) , x) in
+ let (d,(ev,ctx,c)) , x = assert_done r in (d,(ev,ctx, EConstr.to_constr ~abort_on_undefined_evars:false ev c)) , x) in
let concl0 = EConstr.Unsafe.to_constr concl0 in
let concl = eval_pattern env0 sigma0 concl0 rdx_pat occ find_R in
let (d, r), rdx = conclude concl in
@@ -589,7 +589,10 @@ let ssrinstancesofrule ist dir arg gl =
let rpat env sigma0 (sigma, pats) (d, r, lhs, rhs) =
let sigma, pat =
let rw_progress rhs t evd = rw_progress rhs (EConstr.of_constr t) evd in
- mk_tpattern env sigma0 (sigma,EConstr.to_constr sigma r) (rw_progress rhs) d (EConstr.to_constr sigma lhs) in
+ mk_tpattern env sigma0
+ (sigma,EConstr.to_constr ~abort_on_undefined_evars:false sigma r)
+ (rw_progress rhs) d
+ (EConstr.to_constr ~abort_on_undefined_evars:false sigma lhs) in
sigma, pats @ [pat] in
let rpats = List.fold_left (rpat env0 sigma0) (r_sigma,[]) rules in
mk_tpattern_matcher ~all_instances:true ~raise_NoMatch:true sigma0 None ~upats_origin rpats in
diff --git a/plugins/ssr/ssripats.ml b/plugins/ssr/ssripats.ml
index 7897cb170..b397c5531 100644
--- a/plugins/ssr/ssripats.ml
+++ b/plugins/ssr/ssripats.ml
@@ -218,15 +218,16 @@ let rec ipat_tac1 future_ipats ipat : unit tactic =
Ssrview.tclIPAT_VIEWS ~views:l
~conclusion:(fun ~to_clear:clr -> intro_clear clr future_ipats)
| IPatDispatch ipatss ->
- tclEXTEND (List.map ipat_tac ipatss) (tclUNIT ()) []
+ tclEXTEND (List.map (ipat_tac future_ipats) ipatss) (tclUNIT ()) []
| IPatId id -> Ssrcommon.tclINTRO_ID id
| IPatCase ipatss ->
- tclIORPAT (Ssrcommon.tclWITHTOP tac_case) ipatss
+ tclIORPAT (Ssrcommon.tclWITHTOP tac_case) future_ipats ipatss
| IPatInj ipatss ->
tclIORPAT (Ssrcommon.tclWITHTOP
- (fun t -> V82.tactic ~nf_evars:false (Ssrelim.perform_injection t))) ipatss
+ (fun t -> V82.tactic ~nf_evars:false (Ssrelim.perform_injection t)))
+ future_ipats ipatss
| IPatAnon Drop -> intro_drop
| IPatAnon One -> Ssrcommon.tclINTRO_ANON
@@ -254,17 +255,17 @@ let rec ipat_tac1 future_ipats ipat : unit tactic =
| IPatTac t -> t
-and ipat_tac pl : unit tactic =
+and ipat_tac future_ipats pl : unit tactic =
match pl with
| [] -> tclUNIT ()
| pat :: pl ->
- Ssrcommon.tcl0G (tclLOG pat (ipat_tac1 pl)) <*>
+ Ssrcommon.tcl0G (tclLOG pat (ipat_tac1 (pl @ future_ipats))) <*>
isTICK pat <*>
- ipat_tac pl
+ ipat_tac future_ipats pl
-and tclIORPAT tac = function
+and tclIORPAT tac future_ipats = function
| [[]] -> tac
- | p -> Tacticals.New.tclTHENS tac (List.map ipat_tac p)
+ | p -> Tacticals.New.tclTHENS tac (List.map (ipat_tac future_ipats) p)
let split_at_first_case ipats =
let rec loop acc = function
@@ -285,7 +286,7 @@ let main ?eqtac ~first_case_is_dispatch ipats =
let case = ssr_exception first_case_is_dispatch case in
let case = option_to_list case in
let eqtac = option_to_list (Option.map (fun x -> IPatTac x) eqtac) in
- Ssrcommon.tcl0G (ipat_tac (ip_before @ case @ eqtac @ ip_after) <*> intro_end)
+ Ssrcommon.tcl0G (ipat_tac [] (ip_before @ case @ eqtac @ ip_after) <*> intro_end)
end (* }}} *)
@@ -418,7 +419,7 @@ let tclLAST_GEN ~to_ind ((oclr, occ), t) conclusion = tclINDEPENDENTL begin
Goal.enter_one begin fun g ->
let pat = Ssrmatching.interp_cpattern sigma0 t None in
let cl0, env, sigma, hyps = Goal.(concl g, env g, sigma g, hyps g) in
- let cl = EConstr.to_constr sigma cl0 in
+ let cl = EConstr.to_constr ~abort_on_undefined_evars:false sigma cl0 in
let (c, ucst), cl =
try Ssrmatching.fill_occ_pattern ~raise_NoMatch:true env sigma cl pat occ 1
with Ssrmatching.NoMatch -> Ssrmatching.redex_of_pattern env pat, cl in
@@ -623,7 +624,7 @@ let tacFIND_ABSTRACT_PROOF check_lock abstract_n =
Goal.enter_one ~__LOC__ begin fun g ->
let sigma, env = Goal.(sigma g, env g) in
let l = Evd.fold_undefined (fun e ei l ->
- match EConstr.kind sigma (EConstr.of_constr ei.Evd.evar_concl) with
+ match EConstr.kind sigma ei.Evd.evar_concl with
| Term.App(hd, [|ty; n; lock|])
when (not check_lock ||
(occur_existential_or_casted_meta sigma ty &&
diff --git a/plugins/ssr/ssrview.ml b/plugins/ssr/ssrview.ml
index aa614fbc1..fc50b24a6 100644
--- a/plugins/ssr/ssrview.ml
+++ b/plugins/ssr/ssrview.ml
@@ -254,7 +254,7 @@ let finalize_view s0 ?(simple_types=true) p =
Goal.enter_one ~__LOC__ begin fun g ->
let env = Goal.env g in
let sigma = Goal.sigma g in
- let evars_of_p = Evd.evars_of_term (EConstr.to_constr sigma p) in
+ let evars_of_p = Evd.evars_of_term (EConstr.to_constr ~abort_on_undefined_evars:false sigma p) in
let filter x _ = Evar.Set.mem x evars_of_p in
let sigma = Typeclasses.resolve_typeclasses ~fail:false ~filter env sigma in
let p = Reductionops.nf_evar sigma p in
@@ -265,7 +265,7 @@ Goal.enter_one ~__LOC__ begin fun g ->
List.fold_left (fun l k ->
if Evd.is_defined sigma k then
let bo = get_body Evd.(evar_body (find sigma k)) in
- k :: l @ Evar.Set.elements (evars_of_econstr sigma bo)
+ k :: l @ Evar.Set.elements (evars_of_econstr sigma (EConstr.Unsafe.to_constr bo))
else l
) [] s in
let und0 = (* Unassigned evars in the initial goal *)
diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4
index 2ba6acc03..a10437a63 100644
--- a/plugins/ssrmatching/ssrmatching.ml4
+++ b/plugins/ssrmatching/ssrmatching.ml4
@@ -283,7 +283,7 @@ exception NoProgress
(* comparison can be much faster than the HO one. *)
let unif_EQ env sigma p c =
- let evars = existential_opt_value sigma, Evd.universes sigma in
+ let evars = existential_opt_value0 sigma, Evd.universes sigma in
try let _ = Reduction.conv env p ~evars c in true with _ -> false
let unif_EQ_args env sigma pa a =
@@ -337,7 +337,7 @@ let nf_open_term sigma0 ise c =
let s = ise and s' = ref sigma0 in
let rec nf c' = match kind c' with
| Evar ex ->
- begin try nf (existential_value s ex) with _ ->
+ begin try nf (existential_value0 s ex) with _ ->
let k, a = ex in let a' = Array.map nf a in
if not (Evd.mem !s' k) then
s' := Evd.add !s' k (Evarutil.nf_evar_info s (Evd.find s k));
@@ -347,7 +347,9 @@ let nf_open_term sigma0 ise c =
let copy_def k evi () =
if evar_body evi != Evd.Evar_empty then () else
match Evd.evar_body (Evd.find s k) with
- | Evar_defined c' -> s' := Evd.define k (nf c') !s'
+ | Evar_defined c' ->
+ let c' = EConstr.of_constr (nf (EConstr.Unsafe.to_constr c')) in
+ s' := Evd.define k c' !s'
| _ -> () in
let c' = nf c in let _ = Evd.fold copy_def sigma0 () in
!s', Evd.evar_universe_context s, EConstr.of_constr c'
@@ -446,7 +448,7 @@ let evars_for_FO ~hack env sigma0 (ise0:evar_map) c0 =
let nenv = env_size env + if hack then 1 else 0 in
let rec put c = match kind c with
| Evar (k, a as ex) ->
- begin try put (existential_value !sigma ex)
+ begin try put (existential_value0 !sigma ex)
with NotInstantiatedEvar ->
if Evd.mem sigma0 k then map put c else
let evi = Evd.find !sigma k in
@@ -457,11 +459,13 @@ let evars_for_FO ~hack env sigma0 (ise0:evar_map) c0 =
| Context.Named.Declaration.LocalAssum (x, t) ->
mkVar x :: d, mkNamedProd x (put t) c in
let a, t =
- Context.Named.fold_inside abs_dc ~init:([], (put evi.evar_concl)) dc in
+ Context.Named.fold_inside abs_dc
+ ~init:([], (put @@ EConstr.Unsafe.to_constr evi.evar_concl))
+ (EConstr.Unsafe.to_named_context dc) in
let m = Evarutil.new_meta () in
- ise := meta_declare m t !ise;
- sigma := Evd.define k (applistc (mkMeta m) a) !sigma;
- put (existential_value !sigma ex)
+ ise := meta_declare m (EConstr.of_constr t) !ise;
+ sigma := Evd.define k (EConstr.of_constr (applistc (mkMeta m) a)) !sigma;
+ put (existential_value0 !sigma ex)
end
| _ -> map put c in
let c1 = put c0 in !ise, c1
@@ -541,7 +545,7 @@ let splay_app ise =
| App (f, a') -> loop f (Array.append a' a)
| Cast (c', _, _) -> loop c' a
| Evar ex ->
- (try loop (existential_value ise ex) a with _ -> c, a)
+ (try loop (existential_value0 ise ex) a with _ -> c, a)
| _ -> c, a in
fun c -> match kind c with
| App (f, a) -> loop f a
@@ -1255,7 +1259,7 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst =
let fs sigma x = nf_evar sigma x in
let pop_evar sigma e p =
let { Evd.evar_body = e_body } as e_def = Evd.find sigma e in
- let e_body = match e_body with Evar_defined c -> c
+ let e_body = match e_body with Evar_defined c -> EConstr.Unsafe.to_constr c
| _ -> errorstrm (str "Matching the pattern " ++ pr_constr_env env0 sigma0 p ++
str " did not instantiate ?" ++ int (Evar.repr e) ++ spc () ++
str "Does the variable bound by the \"in\" construct occur "++
diff --git a/plugins/syntax/ascii_syntax.ml b/plugins/syntax/ascii_syntax.ml
index acb297ddf..47a59ba63 100644
--- a/plugins/syntax/ascii_syntax.ml
+++ b/plugins/syntax/ascii_syntax.ml
@@ -28,7 +28,7 @@ let make_kn dir id = Globnames.encode_mind (make_dir dir) (Id.of_string id)
let make_path dir id = Libnames.make_path (make_dir dir) (Id.of_string id)
let is_gr c gr = match DAst.get c with
-| GRef (r, _) -> Globnames.eq_gr r gr
+| GRef (r, _) -> GlobRef.equal r gr
| _ -> false
let ascii_module = ["Coq";"Strings";"Ascii"]
diff --git a/plugins/syntax/int31_syntax.ml b/plugins/syntax/int31_syntax.ml
index 5529ea700..f10f98e23 100644
--- a/plugins/syntax/int31_syntax.ml
+++ b/plugins/syntax/int31_syntax.ml
@@ -26,7 +26,7 @@ let make_dir l = DirPath.make (List.rev_map Id.of_string l)
let make_path dir id = Libnames.make_path (make_dir dir) (Id.of_string id)
let is_gr c gr = match DAst.get c with
-| GRef (r, _) -> Globnames.eq_gr r gr
+| GRef (r, _) -> GlobRef.equal r gr
| _ -> false
let make_mind mp id = Names.MutInd.make2 mp (Label.make id)
diff --git a/plugins/syntax/nat_syntax.ml b/plugins/syntax/nat_syntax.ml
index ad8b54d4d..e158e0b51 100644
--- a/plugins/syntax/nat_syntax.ml
+++ b/plugins/syntax/nat_syntax.ml
@@ -16,11 +16,12 @@ let () = Mltop.add_known_module __coq_plugin_name
(* This file defines the printer for natural numbers in [nat] *)
(*i*)
+open Pp
+open CErrors
+open Names
open Glob_term
open Bigint
open Coqlib
-open Pp
-open CErrors
(*i*)
(**********************************************************************)
@@ -61,10 +62,10 @@ exception Non_closed_number
let rec int_of_nat x = DAst.with_val (function
| GApp (r, [a]) ->
begin match DAst.get r with
- | GRef (s,_) when Globnames.eq_gr s glob_S -> add_1 (int_of_nat a)
+ | GRef (s,_) when GlobRef.equal s glob_S -> add_1 (int_of_nat a)
| _ -> raise Non_closed_number
end
- | GRef (z,_) when Globnames.eq_gr z glob_O -> zero
+ | GRef (z,_) when GlobRef.equal z glob_O -> zero
| _ -> raise Non_closed_number
) x
diff --git a/plugins/syntax/r_syntax.ml b/plugins/syntax/r_syntax.ml
index 372e8ff30..94aa14335 100644
--- a/plugins/syntax/r_syntax.ml
+++ b/plugins/syntax/r_syntax.ml
@@ -30,7 +30,7 @@ let make_dir l = DirPath.make (List.rev_map Id.of_string l)
let make_path dir id = Libnames.make_path (make_dir dir) (Id.of_string id)
let is_gr c gr = match DAst.get c with
-| GRef (r, _) -> Globnames.eq_gr r gr
+| GRef (r, _) -> GlobRef.equal r gr
| _ -> false
let positive_path = make_path binnums "positive"
@@ -66,7 +66,7 @@ let pos_of_bignat ?loc x =
let rec bignat_of_pos c = match DAst.get c with
| GApp (r, [a]) when is_gr r glob_xO -> mult_2(bignat_of_pos a)
| GApp (r, [a]) when is_gr r glob_xI -> add_1(mult_2(bignat_of_pos a))
- | GRef (a, _) when Globnames.eq_gr a glob_xH -> Bigint.one
+ | GRef (a, _) when GlobRef.equal a glob_xH -> Bigint.one
| _ -> raise Non_closed_number
(**********************************************************************)
@@ -98,7 +98,7 @@ let z_of_int ?loc n =
let bigint_of_z c = match DAst.get c with
| GApp (r,[a]) when is_gr r glob_POS -> bignat_of_pos a
| GApp (r,[a]) when is_gr r glob_NEG -> Bigint.neg (bignat_of_pos a)
- | GRef (a, _) when Globnames.eq_gr a glob_ZERO -> Bigint.zero
+ | GRef (a, _) when GlobRef.equal a glob_ZERO -> Bigint.zero
| _ -> raise Non_closed_number
(**********************************************************************)
diff --git a/plugins/syntax/string_syntax.ml b/plugins/syntax/string_syntax.ml
index 2421cc12f..c22869f4d 100644
--- a/plugins/syntax/string_syntax.ml
+++ b/plugins/syntax/string_syntax.ml
@@ -8,6 +8,7 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
+open Names
open Globnames
open Ascii_syntax_plugin.Ascii_syntax
open Glob_term
@@ -34,7 +35,7 @@ let glob_String = lazy (make_reference "String")
let glob_EmptyString = lazy (make_reference "EmptyString")
let is_gr c gr = match DAst.get c with
-| GRef (r, _) -> Globnames.eq_gr r gr
+| GRef (r, _) -> GlobRef.equal r gr
| _ -> false
open Lazy
@@ -55,7 +56,7 @@ let uninterp_string (AnyGlobConstr r) =
(match uninterp_ascii a with
| Some c -> Buffer.add_char b (Char.chr c); aux s
| _ -> raise Non_closed_string)
- | GRef (z,_) when eq_gr z (force glob_EmptyString) ->
+ | GRef (z,_) when GlobRef.equal z (force glob_EmptyString) ->
Some (Buffer.contents b)
| _ ->
raise Non_closed_string
diff --git a/plugins/syntax/z_syntax.ml b/plugins/syntax/z_syntax.ml
index d5300e474..09fe8bf70 100644
--- a/plugins/syntax/z_syntax.ml
+++ b/plugins/syntax/z_syntax.ml
@@ -71,13 +71,13 @@ let interp_positive ?loc n =
(**********************************************************************)
let is_gr c gr = match DAst.get c with
-| GRef (r, _) -> Globnames.eq_gr r gr
+| GRef (r, _) -> GlobRef.equal r gr
| _ -> false
let rec bignat_of_pos x = DAst.with_val (function
| GApp (r ,[a]) when is_gr r glob_xO -> mult_2(bignat_of_pos a)
| GApp (r ,[a]) when is_gr r glob_xI -> add_1(mult_2(bignat_of_pos a))
- | GRef (a, _) when Globnames.eq_gr a glob_xH -> Bigint.one
+ | GRef (a, _) when GlobRef.equal a glob_xH -> Bigint.one
| _ -> raise Non_closed_number
) x
@@ -132,7 +132,7 @@ let n_of_int ?loc n =
let bignat_of_n n = DAst.with_val (function
| GApp (r, [a]) when is_gr r glob_Npos -> bignat_of_pos a
- | GRef (a,_) when Globnames.eq_gr a glob_N0 -> Bigint.zero
+ | GRef (a,_) when GlobRef.equal a glob_N0 -> Bigint.zero
| _ -> raise Non_closed_number
) n
@@ -180,7 +180,7 @@ let z_of_int ?loc n =
let bigint_of_z z = DAst.with_val (function
| GApp (r, [a]) when is_gr r glob_POS -> bignat_of_pos a
| GApp (r, [a]) when is_gr r glob_NEG -> Bigint.neg (bignat_of_pos a)
- | GRef (a, _) when Globnames.eq_gr a glob_ZERO -> Bigint.zero
+ | GRef (a, _) when GlobRef.equal a glob_ZERO -> Bigint.zero
| _ -> raise Non_closed_number
) z