From 118d24281bc62bb7ff503abee56f156545eb9eea Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Sat, 10 Mar 2018 23:54:14 +0100 Subject: [api] Remove deprecated object from `Term` We remove most of what was deprecated in `Term`. Now, `intf` and `kernel` are almost deprecation-free, tho I am not very convinced about the whole `Term -> Constr` renaming but I'm afraid there is no way back. Inconsistencies with the constructor policy (see #6440) remain along the code-base and I'm afraid I don't see a plan to reconcile them. The `Sorts` deprecation is hard to finalize, opening `Sorts` is not a good idea as someone added a `List` module inside it. --- dev/vm_printers.ml | 1 + engine/namegen.ml | 1 + interp/constrextern.ml | 9 +- interp/constrintern.ml | 5 +- interp/notation_ops.ml | 1 + kernel/cemitcodes.ml | 2 +- kernel/term.ml | 256 +++--------------------- kernel/term.mli | 396 +------------------------------------ library/decl_kinds.ml | 11 -- library/keys.ml | 4 +- library/misctypes.ml | 10 - parsing/g_constr.ml4 | 1 + plugins/btauto/refl_btauto.ml | 26 ++- plugins/firstorder/unify.ml | 2 +- plugins/funind/glob_termops.ml | 3 +- plugins/funind/indfun_common.ml | 2 +- plugins/ltac/evar_tactics.ml | 2 +- plugins/ltac/extratactics.ml4 | 2 +- plugins/ltac/g_auto.ml4 | 3 +- plugins/ltac/tauto.ml | 2 +- plugins/micromega/coq_micromega.ml | 36 ++-- plugins/omega/coq_omega.ml | 2 +- plugins/romega/refl_omega.ml | 5 +- plugins/ssr/ssrcommon.ml | 8 +- plugins/ssr/ssrelim.ml | 1 + plugins/ssr/ssripats.ml | 19 +- plugins/ssr/ssrparser.ml4 | 4 +- plugins/ssr/ssrtacticals.ml | 7 +- pretyping/coercion.ml | 1 + pretyping/constr_matching.ml | 1 + pretyping/detyping.ml | 1 + pretyping/glob_ops.ml | 1 + pretyping/pretyping.ml | 1 + pretyping/tacred.ml | 2 +- pretyping/typing.ml | 1 + printing/ppconstr.ml | 1 + proofs/clenv.ml | 2 +- proofs/clenvtac.ml | 2 +- proofs/redexpr.ml | 2 +- tactics/btermdn.ml | 2 +- tactics/class_tactics.ml | 1 + tactics/contradiction.ml | 2 +- tactics/eauto.ml | 2 +- tactics/eqdecide.ml | 2 +- tactics/equality.ml | 1 + tactics/hints.ml | 2 +- tactics/hipattern.ml | 2 +- tactics/inv.ml | 1 + tactics/leminv.ml | 2 +- tactics/tacticals.ml | 2 +- tactics/tactics.ml | 4 +- tactics/term_dnet.ml | 2 +- vernac/vernacexpr.ml | 11 -- 53 files changed, 141 insertions(+), 731 deletions(-) diff --git a/dev/vm_printers.ml b/dev/vm_printers.ml index 2ddf927d9..16917586f 100644 --- a/dev/vm_printers.ml +++ b/dev/vm_printers.ml @@ -1,5 +1,6 @@ open Format open Term +open Constr open Names open Cbytecodes open Cemitcodes diff --git a/engine/namegen.ml b/engine/namegen.ml index d66b77b57..c069ec5a0 100644 --- a/engine/namegen.ml +++ b/engine/namegen.ml @@ -17,6 +17,7 @@ open Util open Names open Term +open Constr open Environ open EConstr open Vars diff --git a/interp/constrextern.ml b/interp/constrextern.ml index 7792eff66..86f6ce9ae 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -28,7 +28,6 @@ open Pattern open Nametab open Notation open Detyping -open Misctypes open Decl_kinds module NamedDecl = Context.Named.Declaration @@ -931,7 +930,7 @@ and sub_extern inctx (_,scopes) = extern inctx (None,scopes) and factorize_prod scopes vars na bk aty c = let store, get = set_temporary_memory () in match na, DAst.get c with - | Name id, GCases (LetPatternStyle, None, [(e,(Anonymous,None))],(_::_ as eqns)) + | Name id, GCases (Constr.LetPatternStyle, None, [(e,(Anonymous,None))],(_::_ as eqns)) when is_gvar id e && List.length (store (factorize_eqns eqns)) = 1 -> (match get () with | [{CAst.v=(ids,disj_of_patl,b)}] -> @@ -959,7 +958,7 @@ and factorize_prod scopes vars na bk aty c = and factorize_lambda inctx scopes vars na bk aty c = let store, get = set_temporary_memory () in match na, DAst.get c with - | Name id, GCases (LetPatternStyle, None, [(e,(Anonymous,None))],(_::_ as eqns)) + | Name id, GCases (Constr.LetPatternStyle, None, [(e,(Anonymous,None))],(_::_ as eqns)) when is_gvar id e && List.length (store (factorize_eqns eqns)) = 1 -> (match get () with | [{CAst.v=(ids,disj_of_patl,b)}] -> @@ -1209,7 +1208,7 @@ let rec glob_of_pat avoid env sigma pat = DAst.make @@ match pat with | PIf (c,b1,b2) -> GIf (glob_of_pat avoid env sigma c, (Anonymous,None), glob_of_pat avoid env sigma b1, glob_of_pat avoid env sigma b2) - | PCase ({cip_style=LetStyle; cip_ind_tags=None},PMeta None,tm,[(0,n,b)]) -> + | PCase ({cip_style=Constr.LetStyle; cip_ind_tags=None},PMeta None,tm,[(0,n,b)]) -> let nal,b = it_destRLambda_or_LetIn_names n (glob_of_pat avoid env sigma b) in GLetTuple (nal,(Anonymous,None),glob_of_pat avoid env sigma tm,b) | PCase (info,p,tm,bl) -> @@ -1228,7 +1227,7 @@ let rec glob_of_pat avoid env sigma pat = DAst.make @@ match pat with return_type_of_predicate ind nargs (glob_of_pat avoid env sigma p) | _ -> anomaly (Pp.str "PCase with non-trivial predicate but unknown inductive.") in - GCases (RegularStyle,rtn,[glob_of_pat avoid env sigma tm,indnames],mat) + GCases (Constr.RegularStyle,rtn,[glob_of_pat avoid env sigma tm,indnames],mat) | PFix ((ln,i),(lna,tl,bl)) -> let def_avoid, def_env, lfi = Array.fold_left diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 48f15f897..3a88284e4 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -14,6 +14,7 @@ open Util open CAst open Names open Nameops +open Constr open Namegen open Libnames open Globnames @@ -525,7 +526,7 @@ let rec expand_binders ?loc mk bl c = let tm = DAst.make ?loc (GVar id) in (* Distribute the disjunctive patterns over the shared right-hand side *) let eqnl = List.map (fun pat -> CAst.make ?loc (ids,[pat],c)) disjpat in - let c = DAst.make ?loc @@ GCases (Misctypes.LetPatternStyle, None, [tm,(Anonymous,None)], eqnl) in + let c = DAst.make ?loc @@ GCases (LetPatternStyle, None, [tm,(Anonymous,None)], eqnl) in expand_binders ?loc mk bl (mk ?loc (Name id,Explicit,ty) c) (**********************************************************************) @@ -1965,7 +1966,7 @@ let internalize globalenv env pattern_mode (_, ntnvars as lvar) c = if List.for_all (irrefutable globalenv) thepats then [] else [CAst.make @@ ([],List.make (List.length thepats) (DAst.make @@ PatVar Anonymous), (* "|_,..,_" *) DAst.make @@ GHole(Evar_kinds.ImpossibleCase,Misctypes.IntroAnonymous,None))] (* "=> _" *) in - Some (DAst.make @@ GCases(Term.RegularStyle,sub_rtn,sub_tms,main_sub_eqn::catch_all_sub_eqn)) + Some (DAst.make @@ GCases(RegularStyle,sub_rtn,sub_tms,main_sub_eqn::catch_all_sub_eqn)) in let eqns' = List.map (intern_eqn (List.length tms) env) eqns in DAst.make ?loc @@ diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index 448881dcf..b0480aa70 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -13,6 +13,7 @@ open CErrors open Util open Names open Nameops +open Constr open Globnames open Decl_kinds open Misctypes diff --git a/kernel/cemitcodes.ml b/kernel/cemitcodes.ml index cea09c510..f4e6d45c2 100644 --- a/kernel/cemitcodes.ml +++ b/kernel/cemitcodes.ml @@ -13,7 +13,7 @@ (* Extension: Arnaud Spiwack (support for native arithmetic), May 2005 *) open Names -open Term +open Constr open Cbytecodes open Copcodes open Mod_subst diff --git a/kernel/term.ml b/kernel/term.ml index e1affb1c0..b44e038e9 100644 --- a/kernel/term.ml +++ b/kernel/term.ml @@ -15,219 +15,17 @@ open Names open Vars open Constr -(**********************************************************************) -(** Redeclaration of types from module Constr *) -(**********************************************************************) - +(* Deprecated *) type contents = Sorts.contents = Pos | Null - -type sorts = Sorts.t = - | Prop of contents (** Prop and Set *) - | Type of Univ.Universe.t (** Type *) +[@@ocaml.deprecated "Alias for Sorts.contents"] type sorts_family = Sorts.family = InProp | InSet | InType +[@@ocaml.deprecated "Alias for Sorts.family"] -type constr = Constr.t -(** Alias types, for compatibility. *) - -type types = Constr.t -(** Same as [constr], for documentation purposes. *) - -type existential_key = Evar.t -type existential = Constr.existential - -type metavariable = Constr.metavariable - -type case_style = Constr.case_style = - LetStyle | IfStyle | LetPatternStyle | MatchStyle | RegularStyle - -type case_printing = Constr.case_printing = - { ind_tags : bool list; cstr_tags : bool list array; style : case_style } - -type case_info = Constr.case_info = - { ci_ind : inductive; - ci_npar : int; - ci_cstr_ndecls : int array; - ci_cstr_nargs : int array; - ci_pp_info : case_printing - } - -type cast_kind = Constr.cast_kind = - VMcast | NATIVEcast | DEFAULTcast | REVERTcast - -(********************************************************************) -(* Constructions as implemented *) -(********************************************************************) - -type rec_declaration = Constr.rec_declaration -type fixpoint = Constr.fixpoint -type cofixpoint = Constr.cofixpoint -type 'constr pexistential = 'constr Constr.pexistential -type ('constr, 'types) prec_declaration = - ('constr, 'types) Constr.prec_declaration -type ('constr, 'types) pfixpoint = ('constr, 'types) Constr.pfixpoint -type ('constr, 'types) pcofixpoint = ('constr, 'types) Constr.pcofixpoint -type 'a puniverses = 'a Univ.puniverses - -(** Simply type aliases *) -type pconstant = Constant.t puniverses -type pinductive = inductive puniverses -type pconstructor = constructor puniverses - -type ('constr, 'types, 'sort, 'univs) kind_of_term = - ('constr, 'types, 'sort, 'univs) Constr.kind_of_term = - | Rel of int - | Var of Id.t - | Meta of metavariable - | Evar of 'constr pexistential - | Sort of 'sort - | Cast of 'constr * cast_kind * 'types - | Prod of Name.t * 'types * 'types - | Lambda of Name.t * 'types * 'constr - | LetIn of Name.t * 'constr * 'types * 'constr - | App of 'constr * 'constr array - | Const of (Constant.t * 'univs) - | Ind of (inductive * 'univs) - | Construct of (constructor * 'univs) - | Case of case_info * 'constr * 'constr * 'constr array - | Fix of ('constr, 'types) pfixpoint - | CoFix of ('constr, 'types) pcofixpoint - | Proj of Projection.t * 'constr - -type values = Vmvalues.values - -(**********************************************************************) -(** Redeclaration of functions from module Constr *) -(**********************************************************************) - -let set_sort = Sorts.set -let prop_sort = Sorts.prop -let type1_sort = Sorts.type1 -let sorts_ord = Sorts.compare -let is_prop_sort = Sorts.is_prop -let family_of_sort = Sorts.family -let univ_of_sort = Sorts.univ_of_sort -let sort_of_univ = Sorts.sort_of_univ - -(** {6 Term constructors. } *) - -let mkRel = Constr.mkRel -let mkVar = Constr.mkVar -let mkMeta = Constr.mkMeta -let mkEvar = Constr.mkEvar -let mkSort = Constr.mkSort -let mkProp = Constr.mkProp -let mkSet = Constr.mkSet -let mkType = Constr.mkType -let mkCast = Constr.mkCast -let mkProd = Constr.mkProd -let mkLambda = Constr.mkLambda -let mkLetIn = Constr.mkLetIn -let mkApp = Constr.mkApp -let mkConst = Constr.mkConst -let mkProj = Constr.mkProj -let mkInd = Constr.mkInd -let mkConstruct = Constr.mkConstruct -let mkConstU = Constr.mkConstU -let mkIndU = Constr.mkIndU -let mkConstructU = Constr.mkConstructU -let mkConstructUi = Constr.mkConstructUi -let mkCase = Constr.mkCase -let mkFix = Constr.mkFix -let mkCoFix = Constr.mkCoFix - -(**********************************************************************) -(** Aliases of functions from module Constr *) -(**********************************************************************) - -let eq_constr = Constr.equal -let eq_constr_univs = Constr.eq_constr_univs -let leq_constr_univs = Constr.leq_constr_univs -let eq_constr_nounivs = Constr.eq_constr_nounivs - -let kind_of_term = Constr.kind -let compare = Constr.compare -let constr_ord = compare -let fold_constr = Constr.fold -let map_puniverses = Constr.map_puniverses -let map_constr = Constr.map -let map_constr_with_binders = Constr.map_with_binders -let iter_constr = Constr.iter -let iter_constr_with_binders = Constr.iter_with_binders -let compare_constr = Constr.compare_head -let hash_constr = Constr.hash -let hcons_sorts = Sorts.hcons -let hcons_constr = Constr.hcons -let hcons_types = Constr.hcons - -(**********************************************************************) -(** HERE BEGINS THE INTERESTING STUFF *) -(**********************************************************************) - -(**********************************************************************) -(* Non primitive term destructors *) -(**********************************************************************) - -exception DestKO = DestKO -(* Destructs a de Bruijn index *) -let destRel = destRel -let destMeta = destRel -let isMeta = isMeta -let destVar = destVar -let isSort = isSort -let destSort = destSort -let isprop = isprop -let is_Prop = is_Prop -let is_Set = is_Set -let is_Type = is_Type -let is_small = is_small -let iskind = iskind -let isEvar = isEvar -let isEvar_or_Meta = isEvar_or_Meta -let destCast = destCast -let isCast = isCast -let isRel = isRel -let isRelN = isRelN -let isVar = isVar -let isVarId = isVarId -let isInd = isInd -let destProd = destProd -let isProd = isProd -let destLambda = destLambda -let isLambda = isLambda -let destLetIn = destLetIn -let isLetIn = isLetIn -let destApp = destApp -let destApplication = destApp -let isApp = isApp -let destConst = destConst -let isConst = isConst -let destEvar = destEvar -let destInd = destInd -let destConstruct = destConstruct -let isConstruct = isConstruct -let destCase = destCase -let isCase = isCase -let isProj = isProj -let destProj = destProj -let destFix = destFix -let isFix = isFix -let destCoFix = destCoFix -let isCoFix = isCoFix - -(******************************************************************) -(* Flattening and unflattening of embedded applications and casts *) -(******************************************************************) - -let decompose_app c = - match kind_of_term c with - | App (f,cl) -> (f, Array.to_list cl) - | _ -> (c,[]) - -let decompose_appvect c = - match kind_of_term c with - | App (f,cl) -> (f, cl) - | _ -> (c,[||]) +type sorts = Sorts.t = + | Prop of Sorts.contents (** Prop and Set *) + | Type of Univ.Universe.t (** Type *) +[@@ocaml.deprecated "Alias for Sorts.t"] (****************************************************************************) (* Functions for dealing with constr terms *) @@ -321,7 +119,7 @@ let rec to_lambda n prod = if Int.equal n 0 then prod else - match kind_of_term prod with + match kind prod with | Prod (na,ty,bd) -> mkLambda (na,ty,to_lambda (n-1) bd) | Cast (c,_,_) -> to_lambda n c | _ -> user_err ~hdr:"to_lambda" (mt ()) @@ -330,7 +128,7 @@ let rec to_prod n lam = if Int.equal n 0 then lam else - match kind_of_term lam with + match kind lam with | Lambda (na,ty,bd) -> mkProd (na,ty,to_prod (n-1) bd) | Cast (c,_,_) -> to_prod n c | _ -> user_err ~hdr:"to_prod" (mt ()) @@ -342,7 +140,7 @@ let it_mkLambda_or_LetIn = List.fold_left (fun c d -> mkLambda_or_LetIn d c) let lambda_applist c l = let rec app subst c l = - match kind_of_term c, l with + match kind c, l with | Lambda(_,_,c), arg::l -> app (arg::subst) c l | _, [] -> substl subst c | _ -> anomaly (Pp.str "Not enough lambda's.") in @@ -355,7 +153,7 @@ let lambda_applist_assum n c l = if Int.equal n 0 then if l == [] then substl subst t else anomaly (Pp.str "Too many arguments.") - else match kind_of_term t, l with + else match kind t, l with | Lambda(_,_,c), arg::l -> app (n-1) (arg::subst) c l | LetIn(_,b,_,c), _ -> app (n-1) (substl subst b::subst) c l | _, [] -> anomaly (Pp.str "Not enough arguments.") @@ -367,7 +165,7 @@ let lambda_appvect_assum n c v = lambda_applist_assum n c (Array.to_list v) (* prod_applist T [ a1 ; ... ; an ] -> (T a1 ... an) *) let prod_applist c l = let rec app subst c l = - match kind_of_term c, l with + match kind c, l with | Prod(_,_,c), arg::l -> app (arg::subst) c l | _, [] -> substl subst c | _ -> anomaly (Pp.str "Not enough prod's.") in @@ -381,7 +179,7 @@ let prod_applist_assum n c l = if Int.equal n 0 then if l == [] then substl subst t else anomaly (Pp.str "Too many arguments.") - else match kind_of_term t, l with + else match kind t, l with | Prod(_,_,c), arg::l -> app (n-1) (arg::subst) c l | LetIn(_,b,_,c), _ -> app (n-1) (substl subst b::subst) c l | _, [] -> anomaly (Pp.str "Not enough arguments.") @@ -397,7 +195,7 @@ let prod_appvect_assum n c v = prod_applist_assum n c (Array.to_list v) (* Transforms a product term (x1:T1)..(xn:Tn)T into the pair ([(xn,Tn);...;(x1,T1)],T), where T is not a product *) let decompose_prod = - let rec prodec_rec l c = match kind_of_term c with + let rec prodec_rec l c = match kind c with | Prod (x,t,c) -> prodec_rec ((x,t)::l) c | Cast (c,_,_) -> prodec_rec l c | _ -> l,c @@ -407,7 +205,7 @@ let decompose_prod = (* Transforms a lambda term [x1:T1]..[xn:Tn]T into the pair ([(xn,Tn);...;(x1,T1)],T), where T is not a lambda *) let decompose_lam = - let rec lamdec_rec l c = match kind_of_term c with + let rec lamdec_rec l c = match kind c with | Lambda (x,t,c) -> lamdec_rec ((x,t)::l) c | Cast (c,_,_) -> lamdec_rec l c | _ -> l,c @@ -420,7 +218,7 @@ let decompose_prod_n n = if n < 0 then user_err (str "decompose_prod_n: integer parameter must be positive"); let rec prodec_rec l n c = if Int.equal n 0 then l,c - else match kind_of_term c with + else match kind c with | Prod (x,t,c) -> prodec_rec ((x,t)::l) (n-1) c | Cast (c,_,_) -> prodec_rec l n c | _ -> user_err (str "decompose_prod_n: not enough products") @@ -433,7 +231,7 @@ let decompose_lam_n n = if n < 0 then user_err (str "decompose_lam_n: integer parameter must be positive"); let rec lamdec_rec l n c = if Int.equal n 0 then l,c - else match kind_of_term c with + else match kind c with | Lambda (x,t,c) -> lamdec_rec ((x,t)::l) (n-1) c | Cast (c,_,_) -> lamdec_rec l n c | _ -> user_err (str "decompose_lam_n: not enough abstractions") @@ -445,7 +243,7 @@ let decompose_lam_n n = let decompose_prod_assum = let open Context.Rel.Declaration in let rec prodec_rec l c = - match kind_of_term c with + match kind c with | Prod (x,t,c) -> prodec_rec (Context.Rel.add (LocalAssum (x,t)) l) c | LetIn (x,b,t,c) -> prodec_rec (Context.Rel.add (LocalDef (x,b,t)) l) c | Cast (c,_,_) -> prodec_rec l c @@ -458,7 +256,7 @@ let decompose_prod_assum = let decompose_lam_assum = let rec lamdec_rec l c = let open Context.Rel.Declaration in - match kind_of_term c with + match kind c with | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (LocalAssum (x,t)) l) c | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (LocalDef (x,b,t)) l) c | Cast (c,_,_) -> lamdec_rec l c @@ -477,7 +275,7 @@ let decompose_prod_n_assum n = if Int.equal n 0 then l,c else let open Context.Rel.Declaration in - match kind_of_term c with + match kind c with | Prod (x,t,c) -> prodec_rec (Context.Rel.add (LocalAssum (x,t)) l) (n-1) c | LetIn (x,b,t,c) -> prodec_rec (Context.Rel.add (LocalDef (x,b,t)) l) (n-1) c | Cast (c,_,_) -> prodec_rec l n c @@ -498,7 +296,7 @@ let decompose_lam_n_assum n = if Int.equal n 0 then l,c else let open Context.Rel.Declaration in - match kind_of_term c with + match kind c with | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (LocalAssum (x,t)) l) (n-1) c | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (LocalDef (x,b,t)) l) n c | Cast (c,_,_) -> lamdec_rec l n c @@ -514,7 +312,7 @@ let decompose_lam_n_decls n = if Int.equal n 0 then l,c else let open Context.Rel.Declaration in - match kind_of_term c with + match kind c with | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (LocalAssum (x,t)) l) (n-1) c | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (LocalDef (x,b,t)) l) (n-1) c | Cast (c,_,_) -> lamdec_rec l n c @@ -541,12 +339,12 @@ let strip_lam_n n t = snd (decompose_lam_n n t) Such a term can canonically be seen as the pair of a context of types and of a sort *) -type arity = Context.Rel.t * sorts +type arity = Context.Rel.t * Sorts.t let destArity = let open Context.Rel.Declaration in let rec prodec_rec l c = - match kind_of_term c with + match kind c with | Prod (x,t,c) -> prodec_rec (LocalAssum (x,t) :: l) c | LetIn (x,b,t,c) -> prodec_rec (LocalDef (x,b,t) :: l) c | Cast (c,_,_) -> prodec_rec l c @@ -558,7 +356,7 @@ let destArity = let mkArity (sign,s) = it_mkProd_or_LetIn (mkSort s) sign let rec isArity c = - match kind_of_term c with + match kind c with | Prod (_,_,c) -> isArity c | LetIn (_,b,_,c) -> isArity (subst1 b c) | Cast (c,_,_) -> isArity c @@ -569,13 +367,13 @@ let rec isArity c = (* Experimental, used in Presburger contrib *) type ('constr, 'types) kind_of_type = - | SortType of sorts + | SortType of Sorts.t | CastType of 'types * 'types | ProdType of Name.t * 'types * 'types | LetInType of Name.t * 'constr * 'types * 'types | AtomicType of 'constr * 'constr array -let kind_of_type t = match kind_of_term t with +let kind_of_type t = match kind t with | Sort s -> SortType s | Cast (c,_,t) -> CastType (c, t) | Prod (na,t,c) -> ProdType (na, t, c) diff --git a/kernel/term.mli b/kernel/term.mli index ee84dcb2b..f651d1a58 100644 --- a/kernel/term.mli +++ b/kernel/term.mli @@ -11,166 +11,6 @@ open Names open Constr -(** {5 Redeclaration of types from module Constr and Sorts} - - This reexports constructors of inductive types defined in module [Constr], - for compatibility purposes. Refer to this module for further info. - -*) - -exception DestKO -[@@ocaml.deprecated "Alias for [Constr.DestKO]"] - -(** {5 Simple term case analysis. } *) -val isRel : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isRel]"] -val isRelN : int -> constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isRelN]"] -val isVar : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isVar]"] -val isVarId : Id.t -> constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isVarId]"] -val isInd : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isInd]"] -val isEvar : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isEvar]"] -val isMeta : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isMeta]"] -val isEvar_or_Meta : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isEvar_or_Meta]"] -val isSort : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isSort]"] -val isCast : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isCast]"] -val isApp : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isApp]"] -val isLambda : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isLambda]"] -val isLetIn : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isletIn]"] -val isProd : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isProp]"] -val isConst : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isConst]"] -val isConstruct : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isConstruct]"] -val isFix : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isFix]"] -val isCoFix : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isCoFix]"] -val isCase : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isCase]"] -val isProj : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isProj]"] - -val is_Prop : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.is_Prop]"] -val is_Set : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.is_Set]"] -val isprop : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.isprop]"] -val is_Type : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.is_Type]"] -val iskind : constr -> bool -[@@ocaml.deprecated "Alias for [Constr.is_kind]"] -val is_small : Sorts.t -> bool -[@@ocaml.deprecated "Alias for [Constr.is_small]"] - - -(** {5 Term destructors } *) -(** Destructor operations are partial functions and - @raise DestKO if the term has not the expected form. *) - -(** Destructs a de Bruijn index *) -val destRel : constr -> int -[@@ocaml.deprecated "Alias for [Constr.destRel]"] - -(** Destructs an existential variable *) -val destMeta : constr -> metavariable -[@@ocaml.deprecated "Alias for [Constr.destMeta]"] - -(** Destructs a variable *) -val destVar : constr -> Id.t -[@@ocaml.deprecated "Alias for [Constr.destVar]"] - -(** Destructs a sort. [is_Prop] recognizes the sort {% \textsf{%}Prop{% }%}, whether - [isprop] recognizes both {% \textsf{%}Prop{% }%} and {% \textsf{%}Set{% }%}. *) -val destSort : constr -> Sorts.t -[@@ocaml.deprecated "Alias for [Constr.destSort]"] - -(** Destructs a casted term *) -val destCast : constr -> constr * cast_kind * constr -[@@ocaml.deprecated "Alias for [Constr.destCast]"] - -(** Destructs the product {% $ %}(x:t_1)t_2{% $ %} *) -val destProd : types -> Name.t * types * types -[@@ocaml.deprecated "Alias for [Constr.destProd]"] - -(** Destructs the abstraction {% $ %}[x:t_1]t_2{% $ %} *) -val destLambda : constr -> Name.t * types * constr -[@@ocaml.deprecated "Alias for [Constr.destLambda]"] - -(** Destructs the let {% $ %}[x:=b:t_1]t_2{% $ %} *) -val destLetIn : constr -> Name.t * constr * types * constr -[@@ocaml.deprecated "Alias for [Constr.destLetIn]"] - -(** Destructs an application *) -val destApp : constr -> constr * constr array -[@@ocaml.deprecated "Alias for [Constr.destApp]"] - -(** Obsolete synonym of destApp *) -val destApplication : constr -> constr * constr array -[@@ocaml.deprecated "Alias for [Constr.destApplication]"] - -(** Decompose any term as an applicative term; the list of args can be empty *) -val decompose_app : constr -> constr * constr list -[@@ocaml.deprecated "Alias for [Constr.decompose_app]"] - -(** Same as [decompose_app], but returns an array. *) -val decompose_appvect : constr -> constr * constr array -[@@ocaml.deprecated "Alias for [Constr.decompose_appvect]"] - -(** Destructs a constant *) -val destConst : constr -> Constant.t Univ.puniverses -[@@ocaml.deprecated "Alias for [Constr.destConst]"] - -(** Destructs an existential variable *) -val destEvar : constr -> existential -[@@ocaml.deprecated "Alias for [Constr.destEvar]"] - -(** Destructs a (co)inductive type *) -val destInd : constr -> inductive Univ.puniverses -[@@ocaml.deprecated "Alias for [Constr.destInd]"] - -(** Destructs a constructor *) -val destConstruct : constr -> constructor Univ.puniverses -[@@ocaml.deprecated "Alias for [Constr.destConstruct]"] - -(** Destructs a [match c as x in I args return P with ... | -Ci(...yij...) => ti | ... end] (or [let (..y1i..) := c as x in I args -return P in t1], or [if c then t1 else t2]) -@return [(info,c,fun args x => P,[|...|fun yij => ti| ...|])] -where [info] is pretty-printing information *) -val destCase : constr -> case_info * constr * constr * constr array -[@@ocaml.deprecated "Alias for [Constr.destCase]"] - -(** Destructs a projection *) -val destProj : constr -> Projection.t * constr -[@@ocaml.deprecated "Alias for [Constr.destProj]"] - -(** Destructs the {% $ %}i{% $ %}th function of the block - [Fixpoint f{_ 1} ctx{_ 1} = b{_ 1} - with f{_ 2} ctx{_ 2} = b{_ 2} - ... - with f{_ n} ctx{_ n} = b{_ n}], - where the length of the {% $ %}j{% $ %}th context is {% $ %}ij{% $ %}. -*) -val destFix : constr -> fixpoint -[@@ocaml.deprecated "Alias for [Constr.destFix]"] - -val destCoFix : constr -> cofixpoint -[@@ocaml.deprecated "Alias for [Constr.destCoFix]"] - (** {5 Derived constructors} *) (** non-dependent product [t1 -> t2], an alias for @@ -349,242 +189,14 @@ type ('constr, 'types) kind_of_type = val kind_of_type : types -> (constr, types) kind_of_type -(** {5 Redeclaration of stuff from module [Sorts]} *) - -val set_sort : Sorts.t -[@@ocaml.deprecated "Alias for Sorts.set"] - -val prop_sort : Sorts.t -[@@ocaml.deprecated "Alias for Sorts.prop"] - -val type1_sort : Sorts.t -[@@ocaml.deprecated "Alias for Sorts.type1"] - -val sorts_ord : Sorts.t -> Sorts.t -> int -[@@ocaml.deprecated "Alias for Sorts.compare"] - -val is_prop_sort : Sorts.t -> bool -[@@ocaml.deprecated "Alias for Sorts.is_prop"] - -val family_of_sort : Sorts.t -> Sorts.family -[@@ocaml.deprecated "Alias for Sorts.family"] - -(** {5 Redeclaration of stuff from module [Constr]} - - See module [Constr] for further info. *) - -(** {6 Term constructors. } *) - -val mkRel : int -> constr -[@@ocaml.deprecated "Alias for Constr.mkRel"] -val mkVar : Id.t -> constr -[@@ocaml.deprecated "Alias for Constr.mkVar"] -val mkMeta : metavariable -> constr -[@@ocaml.deprecated "Alias for Constr.mkMeta"] -val mkEvar : existential -> constr -[@@ocaml.deprecated "Alias for Constr.mkEvar"] -val mkSort : Sorts.t -> types -[@@ocaml.deprecated "Alias for Constr.mkSort"] -val mkProp : types -[@@ocaml.deprecated "Alias for Constr.mkProp"] -val mkSet : types -[@@ocaml.deprecated "Alias for Constr.mkSet"] -val mkType : Univ.Universe.t -> types -[@@ocaml.deprecated "Alias for Constr.mkType"] -val mkCast : constr * cast_kind * constr -> constr -[@@ocaml.deprecated "Alias for Constr"] -val mkProd : Name.t * types * types -> types -[@@ocaml.deprecated "Alias for Constr"] -val mkLambda : Name.t * types * constr -> constr -[@@ocaml.deprecated "Alias for Constr"] -val mkLetIn : Name.t * constr * types * constr -> constr -[@@ocaml.deprecated "Alias for Constr"] -val mkApp : constr * constr array -> constr -[@@ocaml.deprecated "Alias for Constr"] -val mkConst : Constant.t -> constr -[@@ocaml.deprecated "Alias for Constr"] -val mkProj : Projection.t * constr -> constr -[@@ocaml.deprecated "Alias for Constr"] -val mkInd : inductive -> constr -[@@ocaml.deprecated "Alias for Constr"] -val mkConstruct : constructor -> constr -[@@ocaml.deprecated "Alias for Constr"] -val mkConstU : Constant.t Univ.puniverses -> constr -[@@ocaml.deprecated "Alias for Constr"] -val mkIndU : inductive Univ.puniverses -> constr -[@@ocaml.deprecated "Alias for Constr"] -val mkConstructU : constructor Univ.puniverses -> constr -[@@ocaml.deprecated "Alias for Constr"] -val mkConstructUi : (pinductive * int) -> constr -[@@ocaml.deprecated "Alias for Constr"] -val mkCase : case_info * constr * constr * constr array -> constr -[@@ocaml.deprecated "Alias for Constr.mkCase"] -val mkFix : fixpoint -> constr -[@@ocaml.deprecated "Alias for Constr.mkFix"] -val mkCoFix : cofixpoint -> constr -[@@ocaml.deprecated "Alias for Constr.mkCoFix"] - -(** {6 Aliases} *) - -val eq_constr : constr -> constr -> bool -[@@ocaml.deprecated "Alias for Constr.equal"] - -(** [eq_constr_univs u a b] is [true] if [a] equals [b] modulo alpha, casts, - application grouping and the universe constraints in [u]. *) -val eq_constr_univs : constr UGraph.check_function -[@@ocaml.deprecated "Alias for Constr.eq_constr_univs"] - -(** [leq_constr_univs u a b] is [true] if [a] is convertible to [b] modulo - alpha, casts, application grouping and the universe constraints in [u]. *) -val leq_constr_univs : constr UGraph.check_function -[@@ocaml.deprecated "Alias for Constr.leq_constr_univs"] - -(** [eq_constr_univs a b] [true, c] if [a] equals [b] modulo alpha, casts, - application grouping and ignoring universe instances. *) -val eq_constr_nounivs : constr -> constr -> bool -[@@ocaml.deprecated "Alias for Constr.qe_constr_nounivs"] - -val kind_of_term : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term -[@@ocaml.deprecated "Alias for Constr.kind"] - -val compare : constr -> constr -> int -[@@ocaml.deprecated "Alias for [Constr.compare]"] - -val constr_ord : constr -> constr -> int -[@@ocaml.deprecated "Alias for [Term.compare]"] - -val fold_constr : ('a -> constr -> 'a) -> 'a -> constr -> 'a -[@@ocaml.deprecated "Alias for [Constr.fold]"] - -val map_constr : (constr -> constr) -> constr -> constr -[@@ocaml.deprecated "Alias for [Constr.map]"] - -val map_constr_with_binders : - ('a -> 'a) -> ('a -> constr -> constr) -> 'a -> constr -> constr -[@@ocaml.deprecated "Alias for [Constr.map_with_binders]"] - -val map_puniverses : ('a -> 'b) -> 'a Univ.puniverses -> 'b Univ.puniverses -[@@ocaml.deprecated "Alias for [Constr.map_puniverses]"] -val univ_of_sort : Sorts.t -> Univ.Universe.t -[@@ocaml.deprecated "Alias for [Sorts.univ_of_sort]"] -val sort_of_univ : Univ.Universe.t -> Sorts.t -[@@ocaml.deprecated "Alias for [Sorts.sort_of_univ]"] - -val iter_constr : (constr -> unit) -> constr -> unit -[@@ocaml.deprecated "Alias for [Constr.iter]"] - -val iter_constr_with_binders : - ('a -> 'a) -> ('a -> constr -> unit) -> 'a -> constr -> unit -[@@ocaml.deprecated "Alias for [Constr.iter_with_binders]"] - -val compare_constr : (int -> constr -> constr -> bool) -> int -> constr -> constr -> bool -[@@ocaml.deprecated "Alias for [Constr.compare_head]"] - -type constr = Constr.constr -[@@ocaml.deprecated "Alias for Constr.t"] - -(** Alias types, for compatibility. *) - -type types = Constr.types -[@@ocaml.deprecated "Alias for Constr.types"] - +(* Deprecated *) type contents = Sorts.contents = Pos | Null [@@ocaml.deprecated "Alias for Sorts.contents"] +type sorts_family = Sorts.family = InProp | InSet | InType +[@@ocaml.deprecated "Alias for Sorts.family"] + type sorts = Sorts.t = | Prop of Sorts.contents (** Prop and Set *) | Type of Univ.Universe.t (** Type *) [@@ocaml.deprecated "Alias for Sorts.t"] - -type sorts_family = Sorts.family = InProp | InSet | InType -[@@ocaml.deprecated "Alias for Sorts.family"] - -type 'a puniverses = 'a Univ.puniverses -[@@ocaml.deprecated "Alias for Constr.puniverses"] - -(** Simply type aliases *) -type pconstant = Constr.pconstant -[@@ocaml.deprecated "Alias for Constr.pconstant"] -type pinductive = Constr.pinductive -[@@ocaml.deprecated "Alias for Constr.pinductive"] -type pconstructor = Constr.pconstructor -[@@ocaml.deprecated "Alias for Constr.pconstructor"] -type existential_key = Evar.t -[@@ocaml.deprecated "Alias for Evar.t"] -type existential = Constr.existential -[@@ocaml.deprecated "Alias for Constr.existential"] -type metavariable = Constr.metavariable -[@@ocaml.deprecated "Alias for Constr.metavariable"] - -type case_style = Constr.case_style = - LetStyle | IfStyle | LetPatternStyle | MatchStyle | RegularStyle -[@@ocaml.deprecated "Alias for Constr.case_style"] - -type case_printing = Constr.case_printing = - { ind_tags : bool list; cstr_tags : bool list array; style : Constr.case_style } -[@@ocaml.deprecated "Alias for Constr.case_printing"] - -type case_info = Constr.case_info = - { ci_ind : inductive; - ci_npar : int; - ci_cstr_ndecls : int array; - ci_cstr_nargs : int array; - ci_pp_info : Constr.case_printing - } -[@@ocaml.deprecated "Alias for Constr.case_info"] - -type cast_kind = Constr.cast_kind = - VMcast | NATIVEcast | DEFAULTcast | REVERTcast -[@@ocaml.deprecated "Alias for Constr.cast_kind"] - -type rec_declaration = Constr.rec_declaration -[@@ocaml.deprecated "Alias for Constr.rec_declaration"] -type fixpoint = Constr.fixpoint -[@@ocaml.deprecated "Alias for Constr.fixpoint"] -type cofixpoint = Constr.cofixpoint -[@@ocaml.deprecated "Alias for Constr.cofixpoint"] -type 'constr pexistential = 'constr Constr.pexistential -[@@ocaml.deprecated "Alias for Constr.pexistential"] -type ('constr, 'types) prec_declaration = - ('constr, 'types) Constr.prec_declaration -[@@ocaml.deprecated "Alias for Constr.prec_declaration"] -type ('constr, 'types) pfixpoint = ('constr, 'types) Constr.pfixpoint -[@@ocaml.deprecated "Alias for Constr.pfixpoint"] -type ('constr, 'types) pcofixpoint = ('constr, 'types) Constr.pcofixpoint -[@@ocaml.deprecated "Alias for Constr.pcofixpoint"] - -type ('constr, 'types, 'sort, 'univs) kind_of_term = - ('constr, 'types, 'sort, 'univs) Constr.kind_of_term = - | Rel of int - | Var of Id.t - | Meta of Constr.metavariable - | Evar of 'constr Constr.pexistential - | Sort of 'sort - | Cast of 'constr * Constr.cast_kind * 'types - | Prod of Name.t * 'types * 'types - | Lambda of Name.t * 'types * 'constr - | LetIn of Name.t * 'constr * 'types * 'constr - | App of 'constr * 'constr array - | Const of (Constant.t * 'univs) - | Ind of (inductive * 'univs) - | Construct of (constructor * 'univs) - | Case of Constr.case_info * 'constr * 'constr * 'constr array - | Fix of ('constr, 'types) Constr.pfixpoint - | CoFix of ('constr, 'types) Constr.pcofixpoint - | Proj of Projection.t * 'constr -[@@ocaml.deprecated "Alias for Constr.kind_of_term"] - -type values = Vmvalues.values -[@@ocaml.deprecated "Alias for Vmvalues.values"] - -val hash_constr : Constr.constr -> int -[@@ocaml.deprecated "Alias for Constr.hash"] - -val hcons_sorts : Sorts.t -> Sorts.t -[@@ocaml.deprecated "Alias for [Sorts.hcons]"] - -val hcons_constr : Constr.constr -> Constr.constr -[@@ocaml.deprecated "Alias for [Constr.hcons]"] - -val hcons_types : Constr.types -> Constr.types -[@@ocaml.deprecated "Alias for [Constr.hcons]"] diff --git a/library/decl_kinds.ml b/library/decl_kinds.ml index 0d3285311..c1a673edf 100644 --- a/library/decl_kinds.ml +++ b/library/decl_kinds.ml @@ -74,14 +74,3 @@ type logical_kind = | IsAssumption of assumption_object_kind | IsDefinition of definition_object_kind | IsProof of theorem_kind - -(** Recursive power of type declarations *) - -type recursivity_kind = Declarations.recursivity_kind = - | Finite (** = inductive *) - [@ocaml.deprecated "Please use [Declarations.Finite"] - | CoFinite (** = coinductive *) - [@ocaml.deprecated "Please use [Declarations.CoFinite"] - | BiFinite (** = non-recursive, like in "Record" definitions *) - [@ocaml.deprecated "Please use [Declarations.BiFinite"] -[@@ocaml.deprecated "Please use [Declarations.recursivity_kind"] diff --git a/library/keys.ml b/library/keys.ml index 89363455d..3cadcb647 100644 --- a/library/keys.ml +++ b/library/keys.ml @@ -11,9 +11,9 @@ (** Keys for unification and indexing *) open Names -open Term -open Globnames +open Constr open Libobject +open Globnames type key = | KGlob of GlobRef.t diff --git a/library/misctypes.ml b/library/misctypes.ml index b5d30559d..a3c887045 100644 --- a/library/misctypes.ml +++ b/library/misctypes.ml @@ -54,16 +54,6 @@ type 'id move_location = type existential_key = Evar.t -(** Case style, shared with Term *) - -type case_style = Constr.case_style = - | LetStyle - | IfStyle - | LetPatternStyle - | MatchStyle - | RegularStyle (** infer printing form from number of constructor *) -[@@ocaml.deprecated "Alias for Constr.case_style"] - (** Casts *) type 'a cast_type = diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4 index a03ef268d..f8af79cd7 100644 --- a/parsing/g_constr.ml4 +++ b/parsing/g_constr.ml4 @@ -9,6 +9,7 @@ (************************************************************************) open Names +open Constr open Libnames open Glob_term open Constrexpr diff --git a/plugins/btauto/refl_btauto.ml b/plugins/btauto/refl_btauto.ml index 7f98ed427..c2bc8c079 100644 --- a/plugins/btauto/refl_btauto.ml +++ b/plugins/btauto/refl_btauto.ml @@ -1,3 +1,15 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* + | App (head, args) -> if head === andb && Array.length args = 2 then Andb (aux args.(0), aux args.(1)) else if head === orb && Array.length args = 2 then @@ -116,9 +128,9 @@ module Bool = struct else if head === negb && Array.length args = 1 then Negb (aux args.(0)) else Var (Env.add env c) - | Term.Case (info, r, arg, pats) -> + | Case (info, r, arg, pats) -> let is_bool = - let i = info.Term.ci_ind in + let i = info.ci_ind in Names.eq_ind i (Lazy.force ind) in if is_bool then @@ -176,9 +188,9 @@ module Btauto = struct let _, var = Tacmach.pf_reduction_of_red_expr gl (Genredexpr.CbvVm None) var in let var = EConstr.Unsafe.to_constr var in let rec to_list l = match decomp_term (Tacmach.project gl) l with - | Term.App (c, _) + | App (c, _) when c === (Lazy.force CoqList._nil) -> [] - | Term.App (c, [|_; h; t|]) + | App (c, [|_; h; t|]) when c === (Lazy.force CoqList._cons) -> if h === (Lazy.force Bool.trueb) then (true :: to_list t) else if h === (Lazy.force Bool.falseb) then (false :: to_list t) @@ -218,7 +230,7 @@ module Btauto = struct let concl = EConstr.Unsafe.to_constr concl in let t = decomp_term (Tacmach.New.project gl) concl in match t with - | Term.App (c, [|typ; p; _|]) when c === eq -> + | App (c, [|typ; p; _|]) when c === eq -> (* should be an equality [@eq poly ?p (Cst false)] *) let tac = Tacticals.New.tclORELSE0 Tactics.reflexivity (Proofview.V82.tactic (print_counterexample p env)) in tac @@ -236,7 +248,7 @@ module Btauto = struct let bool = Lazy.force Bool.typ in let t = decomp_term sigma concl in match t with - | Term.App (c, [|typ; tl; tr|]) + | App (c, [|typ; tl; tr|]) when typ === bool && c === eq -> let env = Env.empty () in let fl = Bool.quote env sigma tl in diff --git a/plugins/firstorder/unify.ml b/plugins/firstorder/unify.ml index b869c04a2..06f56d06e 100644 --- a/plugins/firstorder/unify.ml +++ b/plugins/firstorder/unify.ml @@ -9,7 +9,7 @@ (************************************************************************) open Util -open Term +open Constr open EConstr open Vars open Termops diff --git a/plugins/funind/glob_termops.ml b/plugins/funind/glob_termops.ml index ae238b846..bb1587507 100644 --- a/plugins/funind/glob_termops.ml +++ b/plugins/funind/glob_termops.ml @@ -1,4 +1,5 @@ open Pp +open Constr open Glob_term open CErrors open Util @@ -16,7 +17,7 @@ let mkGApp(rt,rtl) = DAst.make @@ GApp(rt,rtl) let mkGLambda(n,t,b) = DAst.make @@ GLambda(n,Explicit,t,b) let mkGProd(n,t,b) = DAst.make @@ GProd(n,Explicit,t,b) let mkGLetIn(n,b,t,c) = DAst.make @@ GLetIn(n,b,t,c) -let mkGCases(rto,l,brl) = DAst.make @@ GCases(Term.RegularStyle,rto,l,brl) +let mkGCases(rto,l,brl) = DAst.make @@ GCases(RegularStyle,rto,l,brl) let mkGHole () = DAst.make @@ GHole(Evar_kinds.BinderType Anonymous,Misctypes.IntroAnonymous,None) (* diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml index b0c9ff8fc..c6faa142a 100644 --- a/plugins/funind/indfun_common.ml +++ b/plugins/funind/indfun_common.ml @@ -109,7 +109,7 @@ let const_of_id id = let def_of_const t = match Constr.kind t with - Term.Const sp -> + Const sp -> (try (match Environ.constant_opt_value_in (Global.env()) sp with | Some c -> c | _ -> assert false) diff --git a/plugins/ltac/evar_tactics.ml b/plugins/ltac/evar_tactics.ml index 5463893ad..ea8dcf57d 100644 --- a/plugins/ltac/evar_tactics.ml +++ b/plugins/ltac/evar_tactics.ml @@ -10,7 +10,7 @@ open Util open Names -open Term +open Constr open CErrors open Evar_refiner open Tacmach diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 index 17011f206..c5254b37c 100644 --- a/plugins/ltac/extratactics.ml4 +++ b/plugins/ltac/extratactics.ml4 @@ -9,6 +9,7 @@ (************************************************************************) open Pp +open Constr open Genarg open Stdarg open Tacarg @@ -286,7 +287,6 @@ END (**********************************************************************) (* Hint Resolve *) -open Term open EConstr open Vars open Coqlib diff --git a/plugins/ltac/g_auto.ml4 b/plugins/ltac/g_auto.ml4 index 07047d016..642e52155 100644 --- a/plugins/ltac/g_auto.ml4 +++ b/plugins/ltac/g_auto.ml4 @@ -9,6 +9,7 @@ (************************************************************************) open Pp +open Constr open Genarg open Stdarg open Pcoq.Prim @@ -169,7 +170,7 @@ END TACTIC EXTEND convert_concl_no_check -| ["convert_concl_no_check" constr(x) ] -> [ Tactics.convert_concl_no_check x Term.DEFAULTcast ] +| ["convert_concl_no_check" constr(x) ] -> [ Tactics.convert_concl_no_check x DEFAULTcast ] END let pr_pre_hints_path_atom _ _ _ = Hints.pp_hints_path_atom Libnames.pr_reference diff --git a/plugins/ltac/tauto.ml b/plugins/ltac/tauto.ml index a51c09ca4..8eeb8903e 100644 --- a/plugins/ltac/tauto.ml +++ b/plugins/ltac/tauto.ml @@ -8,7 +8,7 @@ (* * (see LICENSE file for the text of the license) *) (************************************************************************) -open Term +open Constr open EConstr open Hipattern open Names diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml index 4c0357dd8..c7abd58b0 100644 --- a/plugins/micromega/coq_micromega.ml +++ b/plugins/micromega/coq_micromega.ml @@ -19,10 +19,10 @@ (************************************************************************) open Pp -open Mutils -open Goptions open Names open Constr +open Goptions +open Mutils (** * Debug flag @@ -601,10 +601,10 @@ struct let get_left_construct sigma term = match EConstr.kind sigma term with - | Term.Construct((_,i),_) -> (i,[| |]) - | Term.App(l,rst) -> + | Construct((_,i),_) -> (i,[| |]) + | App(l,rst) -> (match EConstr.kind sigma l with - | Term.Construct((_,i),_) -> (i,rst) + | Construct((_,i),_) -> (i,rst) | _ -> raise ParseError ) | _ -> raise ParseError @@ -688,7 +688,7 @@ struct let parse_q sigma term = match EConstr.kind sigma term with - | Term.App(c, args) -> if EConstr.eq_constr sigma c (Lazy.force coq_Qmake) then + | App(c, args) -> if EConstr.eq_constr sigma c (Lazy.force coq_Qmake) then {Mc.qnum = parse_z sigma args.(0) ; Mc.qden = parse_positive sigma args.(1) } else raise ParseError | _ -> raise ParseError @@ -904,8 +904,8 @@ struct let parse_zop gl (op,args) = let sigma = gl.sigma in match EConstr.kind sigma op with - | Term.Const (x,_) -> (assoc_const sigma op zop_table, args.(0) , args.(1)) - | Term.Ind((n,0),_) -> + | Const (x,_) -> (assoc_const sigma op zop_table, args.(0) , args.(1)) + | Ind((n,0),_) -> if EConstr.eq_constr sigma op (Lazy.force coq_Eq) && is_convertible gl args.(0) (Lazy.force coq_Z) then (Mc.OpEq, args.(1), args.(2)) else raise ParseError @@ -914,8 +914,8 @@ struct let parse_rop gl (op,args) = let sigma = gl.sigma in match EConstr.kind sigma op with - | Term.Const (x,_) -> (assoc_const sigma op rop_table, args.(0) , args.(1)) - | Term.Ind((n,0),_) -> + | Const (x,_) -> (assoc_const sigma op rop_table, args.(0) , args.(1)) + | Ind((n,0),_) -> if EConstr.eq_constr sigma op (Lazy.force coq_Eq) && is_convertible gl args.(0) (Lazy.force coq_R) then (Mc.OpEq, args.(1), args.(2)) else raise ParseError @@ -926,7 +926,7 @@ struct let is_constant sigma t = (* This is an approx *) match EConstr.kind sigma t with - | Term.Construct(i,_) -> true + | Construct(i,_) -> true | _ -> false type 'a op = @@ -1011,10 +1011,10 @@ struct try (Mc.PEc (parse_constant term) , env) with ParseError -> match EConstr.kind sigma term with - | Term.App(t,args) -> + | App(t,args) -> ( match EConstr.kind sigma t with - | Term.Const c -> + | Const c -> ( match assoc_ops sigma t ops_spec with | Binop f -> combine env f (args.(0),args.(1)) | Opp -> let (expr,env) = parse_expr env args.(0) in @@ -1077,13 +1077,13 @@ struct let rec rconstant sigma term = match EConstr.kind sigma term with - | Term.Const x -> + | Const x -> if EConstr.eq_constr sigma term (Lazy.force coq_R0) then Mc.C0 else if EConstr.eq_constr sigma term (Lazy.force coq_R1) then Mc.C1 else raise ParseError - | Term.App(op,args) -> + | App(op,args) -> begin try (* the evaluation order is important in the following *) @@ -1153,7 +1153,7 @@ struct if debug then Feedback.msg_debug (Pp.str "parse_arith: " ++ Printer.pr_leconstr_env gl.env sigma cstr ++ fnl ()); match EConstr.kind sigma cstr with - | Term.App(op,args) -> + | App(op,args) -> let (op,lhs,rhs) = parse_op gl (op,args) in let (e1,env) = parse_expr sigma env lhs in let (e2,env) = parse_expr sigma env rhs in @@ -1208,7 +1208,7 @@ struct let rec xparse_formula env tg term = match EConstr.kind sigma term with - | Term.App(l,rst) -> + | App(l,rst) -> (match rst with | [|a;b|] when EConstr.eq_constr sigma l (Lazy.force coq_and) -> let f,env,tg = xparse_formula env tg a in @@ -1225,7 +1225,7 @@ struct let g,env,tg = xparse_formula env tg b in mkformula_binary mkIff term f g,env,tg | _ -> parse_atom env tg term) - | Term.Prod(typ,a,b) when EConstr.Vars.noccurn sigma 1 b -> + | Prod(typ,a,b) when EConstr.Vars.noccurn sigma 1 b -> let f,env,tg = xparse_formula env tg a in let g,env,tg = xparse_formula env tg b in mkformula_binary mkI term f g,env,tg diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index 3594c8765..c615cf278 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -18,8 +18,8 @@ open CErrors open Util open Names +open Constr open Nameops -open Term open EConstr open Tacticals.New open Tacmach.New diff --git a/plugins/romega/refl_omega.ml b/plugins/romega/refl_omega.ml index d18249784..e60348065 100644 --- a/plugins/romega/refl_omega.ml +++ b/plugins/romega/refl_omega.ml @@ -8,6 +8,7 @@ open Pp open Util +open Constr open Const_omega module OmegaSolver = Omega_plugin.Omega.MakeOmegaSolver (Bigint) open OmegaSolver @@ -1036,13 +1037,13 @@ let resolution unsafe sigma env (reified_concl,reified_hyps) systems_list = let decompose_tactic = decompose_tree env context solution_tree in Tactics.generalize (l_generalize_arg @ l_reified_hypnames) >> - Tactics.convert_concl_no_check reified Term.DEFAULTcast >> + Tactics.convert_concl_no_check reified DEFAULTcast >> Tactics.apply (app coq_do_omega [|decompose_tactic|]) >> show_goal >> (if unsafe then (* Trust the produced term. Faster, but might fail later at Qed. Also handy when debugging, e.g. via a Show Proof after romega. *) - Tactics.convert_concl_no_check (Lazy.force coq_True) Term.VMcast + Tactics.convert_concl_no_check (Lazy.force coq_True) VMcast else Tactics.normalise_vm_in_concl) >> Tactics.apply (Lazy.force coq_I) diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml index c0026616d..3f6503e73 100644 --- a/plugins/ssr/ssrcommon.ml +++ b/plugins/ssr/ssrcommon.ml @@ -423,12 +423,12 @@ let mk_anon_id t gl_ids = (set s i (Char.chr (Char.code (get s i) + 1)); s) in Id.of_bytes (loop (n - 1)) -let convert_concl_no_check t = Tactics.convert_concl_no_check t Term.DEFAULTcast -let convert_concl t = Tactics.convert_concl t Term.DEFAULTcast +let convert_concl_no_check t = Tactics.convert_concl_no_check t DEFAULTcast +let convert_concl t = Tactics.convert_concl t DEFAULTcast let rename_hd_prod orig_name_ref gl = match EConstr.kind (project gl) (pf_concl gl) with - | Term.Prod(_,src,tgt) -> + | Prod(_,src,tgt) -> Proofview.V82.of_tactic (convert_concl_no_check (EConstr.mkProd (!orig_name_ref,src,tgt))) gl | _ -> CErrors.anomaly (str "gentac creates no product") @@ -1446,7 +1446,7 @@ let tclINTRO_ANON = tclINTRO ~id:None ~conclusion:return let tclRENAME_HD_PROD name = Goal.enter begin fun gl -> let convert_concl_no_check t = - Tactics.convert_concl_no_check t Term.DEFAULTcast in + Tactics.convert_concl_no_check t DEFAULTcast in let concl = Goal.concl gl in let sigma = Goal.sigma gl in match EConstr.kind sigma concl with diff --git a/plugins/ssr/ssrelim.ml b/plugins/ssr/ssrelim.ml index 87d107d65..83b4d6562 100644 --- a/plugins/ssr/ssrelim.ml +++ b/plugins/ssr/ssrelim.ml @@ -14,6 +14,7 @@ open Util open Names open Printer open Term +open Constr open Termops open Globnames open Misctypes diff --git a/plugins/ssr/ssripats.ml b/plugins/ssr/ssripats.ml index b397c5531..8207bc11e 100644 --- a/plugins/ssr/ssripats.ml +++ b/plugins/ssr/ssripats.ml @@ -12,6 +12,7 @@ open Ssrmatching_plugin open Util open Names +open Constr open Proofview open Proofview.Notations @@ -90,11 +91,11 @@ open State (** Warning: unlike [nb_deps_assums], it does not perform reduction *) let rec nb_assums cur env sigma t = match EConstr.kind sigma t with - | Term.Prod(name,ty,body) -> + | Prod(name,ty,body) -> nb_assums (cur+1) env sigma body - | Term.LetIn(name,ty,t1,t2) -> + | LetIn(name,ty,t1,t2) -> nb_assums (cur+1) env sigma t2 - | Term.Cast(t,_,_) -> + | Cast(t,_,_) -> nb_assums cur env sigma t | _ -> cur let nb_assums = nb_assums 0 @@ -556,7 +557,7 @@ let rec eqmoveipats eqpat = function let ssrsmovetac = Goal.enter begin fun g -> let sigma, concl = Goal.(sigma g, concl g) in match EConstr.kind sigma concl with - | Term.Prod _ | Term.LetIn _ -> tclUNIT () + | Prod _ | LetIn _ -> tclUNIT () | _ -> Tactics.hnf_in_concl end @@ -594,8 +595,8 @@ let rec is_Evar_or_CastedMeta sigma x = let occur_existential_or_casted_meta sigma c = let rec occrec c = match EConstr.kind sigma c with - | Term.Evar _ -> raise Not_found - | Term.Cast (m,_,_) when EConstr.isMeta sigma m -> raise Not_found + | Evar _ -> raise Not_found + | Cast (m,_,_) when EConstr.isMeta sigma m -> raise Not_found | _ -> EConstr.iter sigma occrec c in try occrec c; false @@ -625,7 +626,7 @@ let tacFIND_ABSTRACT_PROOF check_lock abstract_n = let sigma, env = Goal.(sigma g, env g) in let l = Evd.fold_undefined (fun e ei l -> match EConstr.kind sigma ei.Evd.evar_concl with - | Term.App(hd, [|ty; n; lock|]) + | App(hd, [|ty; n; lock|]) when (not check_lock || (occur_existential_or_casted_meta sigma ty && is_Evar_or_CastedMeta sigma lock)) && @@ -654,8 +655,8 @@ let ssrabstract dgens = let sigma, env, concl = Goal.(sigma g, env g, concl g) in let t = args_id.(0) in match EConstr.kind sigma t with - | (Term.Evar _ | Term.Meta _) -> Ssrcommon.tacUNIFY concl t <*> tclUNIT id - | Term.Cast(m,_,_) + | (Evar _ | Meta _) -> Ssrcommon.tacUNIFY concl t <*> tclUNIT id + | Cast(m,_,_) when EConstr.isEvar sigma m || EConstr.isMeta sigma m -> Ssrcommon.tacUNIFY concl t <*> tclUNIT id | _ -> diff --git a/plugins/ssr/ssrparser.ml4 b/plugins/ssr/ssrparser.ml4 index 5f3967440..138b42e54 100644 --- a/plugins/ssr/ssrparser.ml4 +++ b/plugins/ssr/ssrparser.ml4 @@ -10,6 +10,7 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) +let _vmcast = Constr.VMcast open Names open Pp open Pcoq @@ -17,7 +18,6 @@ open Ltac_plugin open Genarg open Stdarg open Tacarg -open Term open Libnames open Tactics open Tacmach @@ -1938,7 +1938,7 @@ END let vmexacttac pf = Goal.nf_enter begin fun gl -> - exact_no_check (EConstr.mkCast (pf, VMcast, Tacmach.New.pf_concl gl)) + exact_no_check (EConstr.mkCast (pf, _vmcast, Tacmach.New.pf_concl gl)) end TACTIC EXTEND ssrexact diff --git a/plugins/ssr/ssrtacticals.ml b/plugins/ssr/ssrtacticals.ml index 937e68b06..372ae86bd 100644 --- a/plugins/ssr/ssrtacticals.ml +++ b/plugins/ssr/ssrtacticals.ml @@ -11,6 +11,7 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) open Names +open Constr open Termops open Tacmach open Misctypes @@ -103,10 +104,10 @@ let endclausestac id_map clseq gl_id cl0 gl = | ids, dc' -> forced && ids = [] && (not hide_goal || dc' = [] && c_hidden) in let rec unmark c = match EConstr.kind (project gl) c with - | Term.Var id when hidden_clseq clseq && id = gl_id -> cl0 - | Term.Prod (Name id, t, c') when List.mem_assoc id id_map -> + | Var id when hidden_clseq clseq && id = gl_id -> cl0 + | Prod (Name id, t, c') when List.mem_assoc id id_map -> EConstr.mkProd (Name (orig_id id), unmark t, unmark c') - | Term.LetIn (Name id, v, t, c') when List.mem_assoc id id_map -> + | LetIn (Name id, v, t, c') when List.mem_assoc id id_map -> EConstr.mkLetIn (Name (orig_id id), unmark v, unmark t, unmark c') | _ -> EConstr.map (project gl) unmark c in let utac hyp = diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index c9c2445a7..bf9e37aa7 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -20,6 +20,7 @@ open CErrors open Util open Names open Term +open Constr open Environ open EConstr open Vars diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index 0ff6a330f..22da5315f 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -13,6 +13,7 @@ open Pp open CErrors open Util open Names +open Constr open Globnames open Termops open Term diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index fc398df9a..e6cfe1f76 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -14,6 +14,7 @@ open Pp open CErrors open Util open Names +open Constr open Term open EConstr open Vars diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml index 5056c0457..63618c918 100644 --- a/pretyping/glob_ops.ml +++ b/pretyping/glob_ops.ml @@ -11,6 +11,7 @@ open Util open CAst open Names +open Constr open Nameops open Globnames open Misctypes diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index de72f9427..92f87ab95 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -28,6 +28,7 @@ open CErrors open Util open Names open Evd +open Constr open Term open Termops open Environ diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 5a47acd22..40c4cfaa4 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -12,7 +12,7 @@ open Pp open CErrors open Util open Names -open Term +open Constr open Libnames open Globnames open Termops diff --git a/pretyping/typing.ml b/pretyping/typing.ml index 68f9610d1..bffe36eea 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -14,6 +14,7 @@ open Pp open CErrors open Util open Term +open Constr open Environ open EConstr open Vars diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml index 60268c9de..2a5f38697 100644 --- a/printing/ppconstr.ml +++ b/printing/ppconstr.ml @@ -15,6 +15,7 @@ open Pp open CAst open Names open Nameops +open Constr open Libnames open Pputils open Ppextend diff --git a/proofs/clenv.ml b/proofs/clenv.ml index aeaf16723..450fcddfd 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -13,8 +13,8 @@ open CErrors open Util open Names open Nameops -open Term open Termops +open Constr open Namegen open Environ open Evd diff --git a/proofs/clenvtac.ml b/proofs/clenvtac.ml index 209104ac3..38ed63c23 100644 --- a/proofs/clenvtac.ml +++ b/proofs/clenvtac.ml @@ -10,7 +10,7 @@ open Util open Names -open Term +open Constr open Termops open Evd open EConstr diff --git a/proofs/redexpr.ml b/proofs/redexpr.ml index a75711bae..f9e7bbfac 100644 --- a/proofs/redexpr.ml +++ b/proofs/redexpr.ml @@ -12,7 +12,7 @@ open Pp open CErrors open Util open Names -open Term +open Constr open EConstr open Declarations open Globnames diff --git a/tactics/btermdn.ml b/tactics/btermdn.ml index 8f50b0aa2..aca7f6c65 100644 --- a/tactics/btermdn.ml +++ b/tactics/btermdn.ml @@ -9,7 +9,7 @@ (************************************************************************) open Util -open Term +open Constr open EConstr open Names open Pattern diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index 3e08c6d87..28eaf9f28 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -18,6 +18,7 @@ open CErrors open Util open Names open Term +open Constr open Termops open EConstr open Tacmach diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml index c285f21e7..b92bc75bc 100644 --- a/tactics/contradiction.ml +++ b/tactics/contradiction.ml @@ -8,7 +8,7 @@ (* * (see LICENSE file for the text of the license) *) (************************************************************************) -open Term +open Constr open EConstr open Hipattern open Tactics diff --git a/tactics/eauto.ml b/tactics/eauto.ml index 3df9e3f82..80d07c5c0 100644 --- a/tactics/eauto.ml +++ b/tactics/eauto.ml @@ -12,7 +12,7 @@ open Pp open CErrors open Util open Names -open Term +open Constr open Termops open EConstr open Proof_type diff --git a/tactics/eqdecide.ml b/tactics/eqdecide.ml index b0deeed17..176701d99 100644 --- a/tactics/eqdecide.ml +++ b/tactics/eqdecide.ml @@ -17,7 +17,7 @@ open Util open Names open Namegen -open Term +open Constr open EConstr open Declarations open Tactics diff --git a/tactics/equality.ml b/tactics/equality.ml index 8904cd170..f9e06391a 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -15,6 +15,7 @@ open Util open Names open Nameops open Term +open Constr open Termops open EConstr open Vars diff --git a/tactics/hints.ml b/tactics/hints.ml index 7b5be4c1c..7b943b373 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -12,7 +12,7 @@ open Pp open Util open CErrors open Names -open Term +open Constr open Evd open EConstr open Vars diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index b8f1ed720..5d264058a 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -12,7 +12,7 @@ open Pp open CErrors open Util open Names -open Term +open Constr open Termops open EConstr open Inductiveops diff --git a/tactics/inv.ml b/tactics/inv.ml index b346ed223..28cfd57a2 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -14,6 +14,7 @@ open Util open Names open Term open Termops +open Constr open EConstr open Vars open Namegen diff --git a/tactics/leminv.ml b/tactics/leminv.ml index a4cdc1592..f47e6b2cd 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -12,9 +12,9 @@ open Pp open CErrors open Util open Names -open Term open Termops open Environ +open Constr open EConstr open Vars open Namegen diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index 6c7db26c7..732d06f8a 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -509,7 +509,7 @@ module New = struct match Evd.evar_body evi with | Evd.Evar_empty -> Some (evk,evi) | Evd.Evar_defined c -> match Constr.kind (EConstr.Unsafe.to_constr c) with - | Term.Evar (evk,l) -> is_undefined_up_to_restriction sigma evk + | Evar (evk,l) -> is_undefined_up_to_restriction sigma evk | _ -> (* We make the assumption that there is no way to refine an evar remaining after typing from the initial term given to diff --git a/tactics/tactics.ml b/tactics/tactics.ml index bb57e2bf4..58c62af85 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -1910,8 +1910,8 @@ let cast_no_check cast c = exact_no_check (mkCast (c, cast, concl)) end -let vm_cast_no_check c = cast_no_check Term.VMcast c -let native_cast_no_check c = cast_no_check Term.NATIVEcast c +let vm_cast_no_check c = cast_no_check VMcast c +let native_cast_no_check c = cast_no_check NATIVEcast c let exact_proof c = let open Tacmach.New in diff --git a/tactics/term_dnet.ml b/tactics/term_dnet.ml index 611799990..8bdcc6321 100644 --- a/tactics/term_dnet.ml +++ b/tactics/term_dnet.ml @@ -290,7 +290,7 @@ struct | Const (c,u) -> Term (DRef (ConstRef c)) | Ind (i,u) -> Term (DRef (IndRef i)) | Construct (c,u)-> Term (DRef (ConstructRef c)) - | Term.Meta _ -> assert false + | Meta _ -> assert false | Evar (i,_) -> let meta = try Evar.Map.find i !metas diff --git a/vernac/vernacexpr.ml b/vernac/vernacexpr.ml index cf0da4de2..291bd7a26 100644 --- a/vernac/vernacexpr.ml +++ b/vernac/vernacexpr.ml @@ -520,14 +520,3 @@ type vernac_when = | VtNow | VtLater type vernac_classification = vernac_type * vernac_when - - -(** Deprecated stuff *) -type universe_decl_expr = Constrexpr.universe_decl_expr -[@@ocaml.deprecated "alias of Constrexpr.universe_decl_expr"] - -type ident_decl = Constrexpr.ident_decl -[@@ocaml.deprecated "alias of Constrexpr.ident_decl"] - -type name_decl = Constrexpr.name_decl -[@@ocaml.deprecated "alias of Constrexpr.name_decl"] -- cgit v1.2.3