aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--API/API.mli7
-rw-r--r--dev/base_include1
-rw-r--r--interp/constrexpr_ops.ml475
-rw-r--r--interp/constrexpr_ops.mli37
-rw-r--r--interp/constrextern.ml7
-rw-r--r--interp/constrintern.ml1
-rw-r--r--interp/implicit_quantifiers.ml4
-rw-r--r--interp/interp.mllib6
-rw-r--r--interp/topconstr.ml300
-rw-r--r--interp/topconstr.mli41
-rw-r--r--parsing/egramcoq.ml2
-rw-r--r--parsing/g_vernac.ml46
-rw-r--r--printing/ppconstr.ml6
-rw-r--r--vernac/command.ml1
14 files changed, 462 insertions, 432 deletions
diff --git a/API/API.mli b/API/API.mli
index 704f1a356..44d02ce2b 100644
--- a/API/API.mli
+++ b/API/API.mli
@@ -4389,6 +4389,8 @@ sig
val default_binder_kind : Constrexpr.binder_kind
val mkLetInC : Names.Name.t Loc.located * Constrexpr.constr_expr * Constrexpr.constr_expr option * Constrexpr.constr_expr -> Constrexpr.constr_expr
val mkCProdN : ?loc:Loc.t -> Constrexpr.local_binder_expr list -> Constrexpr.constr_expr -> Constrexpr.constr_expr
+ val replace_vars_constr_expr :
+ Names.Id.t Names.Id.Map.t -> Constrexpr.constr_expr -> Constrexpr.constr_expr
end
module Notation_ops :
@@ -4443,8 +4445,11 @@ end
module Topconstr :
sig
+
val replace_vars_constr_expr :
- Names.Id.t Names.Id.Map.t -> Constrexpr.constr_expr -> Constrexpr.constr_expr
+ Names.Id.t Names.Id.Map.t -> Constrexpr.constr_expr -> Constrexpr.constr_expr
+ [@@ocaml.deprecated "use Constrexpr_ops.free_vars_of_constr_expr"]
+
end
module Constrintern :
diff --git a/dev/base_include b/dev/base_include
index f2912e112..0c2bdb9a2 100644
--- a/dev/base_include
+++ b/dev/base_include
@@ -130,7 +130,6 @@ open Reserve
open Syntax_def
open Constrexpr
open Constrexpr_ops
-open Topconstr
open Notation_term
open Notation_ops
open Prettyp
diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml
index 771c13734..737e86848 100644
--- a/interp/constrexpr_ops.ml
+++ b/interp/constrexpr_ops.ml
@@ -9,6 +9,7 @@
open Pp
open Util
open Names
+open Nameops
open Libnames
open Constrexpr
open Misctypes
@@ -72,22 +73,22 @@ let rec cases_pattern_expr_eq p1 p2 =
Option.equal (List.equal cases_pattern_expr_eq) a1 a2 &&
List.equal cases_pattern_expr_eq b1 b2
| CPatAtom(r1), CPatAtom(r2) ->
- Option.equal eq_reference r1 r2
+ Option.equal eq_reference r1 r2
| CPatOr a1, CPatOr a2 ->
- List.equal cases_pattern_expr_eq a1 a2
+ List.equal cases_pattern_expr_eq a1 a2
| CPatNotation (n1, s1, l1), CPatNotation (n2, s2, l2) ->
String.equal n1 n2 &&
cases_pattern_notation_substitution_eq s1 s2 &&
List.equal cases_pattern_expr_eq l1 l2
| CPatPrim i1, CPatPrim i2 ->
- prim_token_eq i1 i2
+ prim_token_eq i1 i2
| CPatRecord l1, CPatRecord l2 ->
- let equal (r1, e1) (r2, e2) =
- eq_reference r1 r2 && cases_pattern_expr_eq e1 e2
- in
- List.equal equal l1 l2
+ let equal (r1, e1) (r2, e2) =
+ eq_reference r1 r2 && cases_pattern_expr_eq e1 e2
+ in
+ List.equal equal l1 l2
| CPatDelimiters(s1,e1), CPatDelimiters(s2,e2) ->
- String.equal s1 s2 && cases_pattern_expr_eq e1 e2
+ String.equal s1 s2 && cases_pattern_expr_eq e1 e2
| _ -> false
and cases_pattern_notation_substitution_eq (s1, n1) (s2, n2) =
@@ -103,79 +104,79 @@ let eq_universes u1 u2 =
let rec constr_expr_eq e1 e2 =
if CAst.(e1.v == e2.v) then true
else match CAst.(e1.v, e2.v) with
- | CRef (r1,u1), CRef (r2,u2) -> eq_reference r1 r2 && eq_universes u1 u2
- | CFix(id1,fl1), CFix(id2,fl2) ->
+ | CRef (r1,u1), CRef (r2,u2) -> eq_reference r1 r2 && eq_universes u1 u2
+ | CFix(id1,fl1), CFix(id2,fl2) ->
eq_located Id.equal id1 id2 &&
List.equal fix_expr_eq fl1 fl2
- | CCoFix(id1,fl1), CCoFix(id2,fl2) ->
+ | CCoFix(id1,fl1), CCoFix(id2,fl2) ->
eq_located Id.equal id1 id2 &&
List.equal cofix_expr_eq fl1 fl2
- | CProdN(bl1,a1), CProdN(bl2,a2) ->
+ | CProdN(bl1,a1), CProdN(bl2,a2) ->
List.equal binder_expr_eq bl1 bl2 &&
constr_expr_eq a1 a2
- | CLambdaN(bl1,a1), CLambdaN(bl2,a2) ->
+ | CLambdaN(bl1,a1), CLambdaN(bl2,a2) ->
List.equal binder_expr_eq bl1 bl2 &&
constr_expr_eq a1 a2
- | CLetIn((_,na1),a1,t1,b1), CLetIn((_,na2),a2,t2,b2) ->
+ | CLetIn((_,na1),a1,t1,b1), CLetIn((_,na2),a2,t2,b2) ->
Name.equal na1 na2 &&
constr_expr_eq a1 a2 &&
Option.equal constr_expr_eq t1 t2 &&
constr_expr_eq b1 b2
- | CAppExpl((proj1,r1,_),al1), CAppExpl((proj2,r2,_),al2) ->
+ | CAppExpl((proj1,r1,_),al1), CAppExpl((proj2,r2,_),al2) ->
Option.equal Int.equal proj1 proj2 &&
eq_reference r1 r2 &&
List.equal constr_expr_eq al1 al2
- | CApp((proj1,e1),al1), CApp((proj2,e2),al2) ->
+ | CApp((proj1,e1),al1), CApp((proj2,e2),al2) ->
Option.equal Int.equal proj1 proj2 &&
constr_expr_eq e1 e2 &&
List.equal args_eq al1 al2
- | CRecord l1, CRecord l2 ->
- let field_eq (r1, e1) (r2, e2) =
- eq_reference r1 r2 && constr_expr_eq e1 e2
- in
- List.equal field_eq l1 l2
- | CCases(_,r1,a1,brl1), CCases(_,r2,a2,brl2) ->
+ | CRecord l1, CRecord l2 ->
+ let field_eq (r1, e1) (r2, e2) =
+ eq_reference r1 r2 && constr_expr_eq e1 e2
+ in
+ List.equal field_eq l1 l2
+ | CCases(_,r1,a1,brl1), CCases(_,r2,a2,brl2) ->
(** Don't care about the case_style *)
Option.equal constr_expr_eq r1 r2 &&
List.equal case_expr_eq a1 a2 &&
List.equal branch_expr_eq brl1 brl2
- | CLetTuple (n1, (m1, e1), t1, b1), CLetTuple (n2, (m2, e2), t2, b2) ->
- List.equal (eq_located Name.equal) n1 n2 &&
- Option.equal (eq_located Name.equal) m1 m2 &&
- Option.equal constr_expr_eq e1 e2 &&
- constr_expr_eq t1 t2 &&
- constr_expr_eq b1 b2
- | CIf (e1, (n1, r1), t1, f1), CIf (e2, (n2, r2), t2, f2) ->
- constr_expr_eq e1 e2 &&
- Option.equal (eq_located Name.equal) n1 n2 &&
- Option.equal constr_expr_eq r1 r2 &&
- constr_expr_eq t1 t2 &&
- constr_expr_eq f1 f2
- | CHole _, CHole _ -> true
- | CPatVar i1, CPatVar i2 ->
- Id.equal i1 i2
- | CEvar (id1, c1), CEvar (id2, c2) ->
- Id.equal id1 id2 && List.equal instance_eq c1 c2
- | CSort s1, CSort s2 ->
- Miscops.glob_sort_eq s1 s2
- | CCast(a1,(CastConv b1|CastVM b1)), CCast(a2,(CastConv b2|CastVM b2)) ->
+ | CLetTuple (n1, (m1, e1), t1, b1), CLetTuple (n2, (m2, e2), t2, b2) ->
+ List.equal (eq_located Name.equal) n1 n2 &&
+ Option.equal (eq_located Name.equal) m1 m2 &&
+ Option.equal constr_expr_eq e1 e2 &&
+ constr_expr_eq t1 t2 &&
+ constr_expr_eq b1 b2
+ | CIf (e1, (n1, r1), t1, f1), CIf (e2, (n2, r2), t2, f2) ->
+ constr_expr_eq e1 e2 &&
+ Option.equal (eq_located Name.equal) n1 n2 &&
+ Option.equal constr_expr_eq r1 r2 &&
+ constr_expr_eq t1 t2 &&
+ constr_expr_eq f1 f2
+ | CHole _, CHole _ -> true
+ | CPatVar i1, CPatVar i2 ->
+ Id.equal i1 i2
+ | CEvar (id1, c1), CEvar (id2, c2) ->
+ Id.equal id1 id2 && List.equal instance_eq c1 c2
+ | CSort s1, CSort s2 ->
+ Miscops.glob_sort_eq s1 s2
+ | CCast(a1,(CastConv b1|CastVM b1)), CCast(a2,(CastConv b2|CastVM b2)) ->
constr_expr_eq a1 a2 &&
constr_expr_eq b1 b2
- | CCast(a1,CastCoerce), CCast(a2, CastCoerce) ->
+ | CCast(a1,CastCoerce), CCast(a2, CastCoerce) ->
constr_expr_eq a1 a2
- | CNotation(n1, s1), CNotation(n2, s2) ->
+ | CNotation(n1, s1), CNotation(n2, s2) ->
String.equal n1 n2 &&
constr_notation_substitution_eq s1 s2
- | CPrim i1, CPrim i2 ->
- prim_token_eq i1 i2
- | CGeneralization (bk1, ak1, e1), CGeneralization (bk2, ak2, e2) ->
- binding_kind_eq bk1 bk2 &&
- Option.equal abstraction_kind_eq ak1 ak2 &&
- constr_expr_eq e1 e2
- | CDelimiters(s1,e1), CDelimiters(s2,e2) ->
- String.equal s1 s2 &&
- constr_expr_eq e1 e2
- | _ -> false
+ | CPrim i1, CPrim i2 ->
+ prim_token_eq i1 i2
+ | CGeneralization (bk1, ak1, e1), CGeneralization (bk2, ak2, e2) ->
+ binding_kind_eq bk1 bk2 &&
+ Option.equal abstraction_kind_eq ak1 ak2 &&
+ constr_expr_eq e1 e2
+ | CDelimiters(s1,e1), CDelimiters(s2,e2) ->
+ String.equal s1 s2 &&
+ constr_expr_eq e1 e2
+ | _ -> false
and args_eq (a1,e1) (a2,e2) =
Option.equal (eq_located explicitation_eq) e1 e2 &&
@@ -209,19 +210,19 @@ and cofix_expr_eq (id1,bl1,a1,b1) (id2,bl2,a2,b2) =
constr_expr_eq b1 b2
and recursion_order_expr_eq r1 r2 = match r1, r2 with
-| CStructRec, CStructRec -> true
-| CWfRec e1, CWfRec e2 -> constr_expr_eq e1 e2
-| CMeasureRec (e1, o1), CMeasureRec (e2, o2) ->
- constr_expr_eq e1 e2 && Option.equal constr_expr_eq o1 o2
-| _ -> false
+ | CStructRec, CStructRec -> true
+ | CWfRec e1, CWfRec e2 -> constr_expr_eq e1 e2
+ | CMeasureRec (e1, o1), CMeasureRec (e2, o2) ->
+ constr_expr_eq e1 e2 && Option.equal constr_expr_eq o1 o2
+ | _ -> false
and local_binder_eq l1 l2 = match l1, l2 with
-| CLocalDef (n1, e1, t1), CLocalDef (n2, e2, t2) ->
- eq_located Name.equal n1 n2 && constr_expr_eq e1 e2 && Option.equal constr_expr_eq t1 t2
-| CLocalAssum (n1, _, e1), CLocalAssum (n2, _, e2) ->
- (** Don't care about the [binder_kind] *)
- List.equal (eq_located Name.equal) n1 n2 && constr_expr_eq e1 e2
-| _ -> false
+ | CLocalDef (n1, e1, t1), CLocalDef (n2, e2, t2) ->
+ eq_located Name.equal n1 n2 && constr_expr_eq e1 e2 && Option.equal constr_expr_eq t1 t2
+ | CLocalAssum (n1, _, e1), CLocalAssum (n2, _, e2) ->
+ (** Don't care about the [binder_kind] *)
+ List.equal (eq_located Name.equal) n1 n2 && constr_expr_eq e1 e2
+ | _ -> false
and constr_notation_substitution_eq (e1, el1, bl1) (e2, el2, bl2) =
List.equal constr_expr_eq e1 e2 &&
@@ -245,6 +246,270 @@ let local_binders_loc bll = match bll with
| [] -> None
| h :: l -> Loc.merge_opt (local_binder_loc h) (local_binder_loc (List.last bll))
+(** Folds and maps *)
+
+(* Legacy functions *)
+let down_located f (_l, x) = f x
+let located_fold_left f x (_l, y) = f x y
+
+let is_constructor id =
+ try Globnames.isConstructRef
+ (Smartlocate.global_of_extended_global
+ (Nametab.locate_extended (qualid_of_ident id)))
+ with Not_found -> false
+
+let rec cases_pattern_fold_names f a pt = match CAst.(pt.v) with
+ | CPatRecord l ->
+ List.fold_left (fun acc (r, cp) -> cases_pattern_fold_names f acc cp) a l
+ | CPatAlias (pat,id) -> f id a
+ | CPatOr (patl) ->
+ List.fold_left (cases_pattern_fold_names f) a patl
+ | CPatCstr (_,patl1,patl2) ->
+ List.fold_left (cases_pattern_fold_names f)
+ (Option.fold_left (List.fold_left (cases_pattern_fold_names f)) a patl1) patl2
+ | CPatNotation (_,(patl,patll),patl') ->
+ List.fold_left (cases_pattern_fold_names f)
+ (List.fold_left (cases_pattern_fold_names f) a (patl@List.flatten patll)) patl'
+ | CPatDelimiters (_,pat) -> cases_pattern_fold_names f a pat
+ | CPatAtom (Some (Ident (_,id))) when not (is_constructor id) -> f id a
+ | CPatPrim _ | CPatAtom _ -> a
+ | CPatCast ({CAst.loc},_) ->
+ CErrors.user_err ?loc ~hdr:"cases_pattern_fold_names"
+ (Pp.strbrk "Casts are not supported here.")
+
+let ids_of_pattern =
+ cases_pattern_fold_names Id.Set.add Id.Set.empty
+
+let ids_of_pattern_list =
+ List.fold_left
+ (located_fold_left
+ (List.fold_left (cases_pattern_fold_names Id.Set.add)))
+ Id.Set.empty
+
+let ids_of_cases_indtype p =
+ cases_pattern_fold_names Id.Set.add Id.Set.empty p
+
+let ids_of_cases_tomatch tms =
+ List.fold_right
+ (fun (_, ona, indnal) l ->
+ Option.fold_right (fun t ids -> cases_pattern_fold_names Id.Set.add ids t)
+ indnal
+ (Option.fold_right (down_located (Name.fold_right Id.Set.add)) ona l))
+ tms Id.Set.empty
+
+let rec fold_constr_expr_binders g f n acc b = function
+ | (nal,bk,t)::l ->
+ let nal = snd (List.split nal) in
+ let n' = List.fold_right (Name.fold_right g) nal n in
+ f n (fold_constr_expr_binders g f n' acc b l) t
+ | [] ->
+ f n acc b
+
+let rec fold_local_binders g f n acc b = function
+ | CLocalAssum (nal,bk,t)::l ->
+ let nal = snd (List.split nal) in
+ let n' = List.fold_right (Name.fold_right g) nal n in
+ f n (fold_local_binders g f n' acc b l) t
+ | CLocalDef ((_,na),c,t)::l ->
+ Option.fold_left (f n) (f n (fold_local_binders g f (Name.fold_right g na n) acc b l) c) t
+ | CLocalPattern (_,(pat,t))::l ->
+ let acc = fold_local_binders g f (cases_pattern_fold_names g n pat) acc b l in
+ Option.fold_left (f n) acc t
+ | [] ->
+ f n acc b
+
+let fold_constr_expr_with_binders g f n acc = CAst.with_val (function
+ | CAppExpl ((_,_,_),l) -> List.fold_left (f n) acc l
+ | CApp ((_,t),l) -> List.fold_left (f n) (f n acc t) (List.map fst l)
+ | CProdN (l,b) | CLambdaN (l,b) -> fold_constr_expr_binders g f n acc b l
+ | CLetIn (na,a,t,b) ->
+ f (Name.fold_right g (snd na) n) (Option.fold_left (f n) (f n acc a) t) b
+ | CCast (a,(CastConv b|CastVM b|CastNative b)) -> f n (f n acc a) b
+ | CCast (a,CastCoerce) -> f n acc a
+ | CNotation (_,(l,ll,bll)) ->
+ (* The following is an approximation: we don't know exactly if
+ an ident is binding nor to which subterms bindings apply *)
+ let acc = List.fold_left (f n) acc (l@List.flatten ll) in
+ List.fold_left (fun acc bl -> fold_local_binders g f n acc (CAst.make @@ CHole (None,IntroAnonymous,None)) bl) acc bll
+ | CGeneralization (_,_,c) -> f n acc c
+ | CDelimiters (_,a) -> f n acc a
+ | CHole _ | CEvar _ | CPatVar _ | CSort _ | CPrim _ | CRef _ ->
+ acc
+ | CRecord l -> List.fold_left (fun acc (id, c) -> f n acc c) acc l
+ | CCases (sty,rtnpo,al,bl) ->
+ let ids = ids_of_cases_tomatch al in
+ let acc = Option.fold_left (f (Id.Set.fold g ids n)) acc rtnpo in
+ let acc = List.fold_left (f n) acc (List.map (fun (fst,_,_) -> fst) al) in
+ List.fold_right (fun (loc,(patl,rhs)) acc ->
+ let ids = ids_of_pattern_list patl in
+ f (Id.Set.fold g ids n) acc rhs) bl acc
+ | CLetTuple (nal,(ona,po),b,c) ->
+ let n' = List.fold_right (down_located (Name.fold_right g)) nal n in
+ f (Option.fold_right (down_located (Name.fold_right g)) ona n') (f n acc b) c
+ | CIf (c,(ona,po),b1,b2) ->
+ let acc = f n (f n (f n acc b1) b2) c in
+ Option.fold_left
+ (f (Option.fold_right (down_located (Name.fold_right g)) ona n)) acc po
+ | CFix (_,l) ->
+ let n' = List.fold_right (fun ((_,id),_,_,_,_) -> g id) l n in
+ List.fold_right (fun (_,(_,o),lb,t,c) acc ->
+ fold_local_binders g f n'
+ (fold_local_binders g f n acc t lb) c lb) l acc
+ | CCoFix (_,_) ->
+ Feedback.msg_warning (strbrk "Capture check in multiple binders not done"); acc
+ )
+
+let free_vars_of_constr_expr c =
+ let rec aux bdvars l = function
+ | { CAst.v = CRef (Ident (_,id),_) } -> if Id.List.mem id bdvars then l else Id.Set.add id l
+ | c -> fold_constr_expr_with_binders (fun a l -> a::l) aux bdvars l c
+ in aux [] Id.Set.empty c
+
+let occur_var_constr_expr id c = Id.Set.mem id (free_vars_of_constr_expr c)
+
+(* Used in correctness and interface *)
+let map_binder g e nal = List.fold_right (down_located (Name.fold_right g)) nal e
+
+let map_binders f g e bl =
+ (* TODO: avoid variable capture in [t] by some [na] in [List.tl nal] *)
+ let h (e,bl) (nal,bk,t) = (map_binder g e nal,(nal,bk,f e t)::bl) in
+ let (e,rbl) = List.fold_left h (e,[]) bl in
+ (e, List.rev rbl)
+
+let map_local_binders f g e bl =
+ (* TODO: avoid variable capture in [t] by some [na] in [List.tl nal] *)
+ let h (e,bl) = function
+ CLocalAssum(nal,k,ty) ->
+ (map_binder g e nal, CLocalAssum(nal,k,f e ty)::bl)
+ | CLocalDef((loc,na),c,ty) ->
+ (Name.fold_right g na e, CLocalDef((loc,na),f e c,Option.map (f e) ty)::bl)
+ | CLocalPattern (loc,(pat,t)) ->
+ let ids = ids_of_pattern pat in
+ (Id.Set.fold g ids e, CLocalPattern (loc,(pat,Option.map (f e) t))::bl) in
+ let (e,rbl) = List.fold_left h (e,[]) bl in
+ (e, List.rev rbl)
+
+let map_constr_expr_with_binders g f e = CAst.map (function
+ | CAppExpl (r,l) -> CAppExpl (r,List.map (f e) l)
+ | CApp ((p,a),l) ->
+ CApp ((p,f e a),List.map (fun (a,i) -> (f e a,i)) l)
+ | CProdN (bl,b) ->
+ let (e,bl) = map_binders f g e bl in CProdN (bl,f e b)
+ | CLambdaN (bl,b) ->
+ let (e,bl) = map_binders f g e bl in CLambdaN (bl,f e b)
+ | CLetIn (na,a,t,b) ->
+ CLetIn (na,f e a,Option.map (f e) t,f (Name.fold_right g (snd na) e) b)
+ | CCast (a,c) -> CCast (f e a, Miscops.map_cast_type (f e) c)
+ | CNotation (n,(l,ll,bll)) ->
+ (* This is an approximation because we don't know what binds what *)
+ CNotation (n,(List.map (f e) l,List.map (List.map (f e)) ll,
+ List.map (fun bl -> snd (map_local_binders f g e bl)) bll))
+ | CGeneralization (b,a,c) -> CGeneralization (b,a,f e c)
+ | CDelimiters (s,a) -> CDelimiters (s,f e a)
+ | CHole _ | CEvar _ | CPatVar _ | CSort _
+ | CPrim _ | CRef _ as x -> x
+ | CRecord l -> CRecord (List.map (fun (id, c) -> (id, f e c)) l)
+ | CCases (sty,rtnpo,a,bl) ->
+ let bl = List.map (fun (loc,(patl,rhs)) ->
+ let ids = ids_of_pattern_list patl in
+ (loc,(patl,f (Id.Set.fold g ids e) rhs))) bl in
+ let ids = ids_of_cases_tomatch a in
+ let po = Option.map (f (Id.Set.fold g ids e)) rtnpo in
+ CCases (sty, po, List.map (fun (tm,x,y) -> f e tm,x,y) a,bl)
+ | CLetTuple (nal,(ona,po),b,c) ->
+ let e' = List.fold_right (down_located (Name.fold_right g)) nal e in
+ let e'' = Option.fold_right (down_located (Name.fold_right g)) ona e in
+ CLetTuple (nal,(ona,Option.map (f e'') po),f e b,f e' c)
+ | CIf (c,(ona,po),b1,b2) ->
+ let e' = Option.fold_right (down_located (Name.fold_right g)) ona e in
+ CIf (f e c,(ona,Option.map (f e') po),f e b1,f e b2)
+ | CFix (id,dl) ->
+ CFix (id,List.map (fun (id,n,bl,t,d) ->
+ let (e',bl') = map_local_binders f g e bl in
+ let t' = f e' t in
+ (* Note: fix names should be inserted before the arguments... *)
+ let e'' = List.fold_left (fun e ((_,id),_,_,_,_) -> g id e) e' dl in
+ let d' = f e'' d in
+ (id,n,bl',t',d')) dl)
+ | CCoFix (id,dl) ->
+ CCoFix (id,List.map (fun (id,bl,t,d) ->
+ let (e',bl') = map_local_binders f g e bl in
+ let t' = f e' t in
+ let e'' = List.fold_left (fun e ((_,id),_,_,_) -> g id e) e' dl in
+ let d' = f e'' d in
+ (id,bl',t',d')) dl)
+ )
+
+(* Used in constrintern *)
+let rec replace_vars_constr_expr l = function
+ | { CAst.loc; v = CRef (Ident (loc_id,id),us) } as x ->
+ (try CAst.make ?loc @@ CRef (Ident (loc_id,Id.Map.find id l),us) with Not_found -> x)
+ | c -> map_constr_expr_with_binders Id.Map.remove
+ replace_vars_constr_expr l c
+
+(* Returns the ranges of locs of the notation that are not occupied by args *)
+(* and which are then occupied by proper symbols of the notation (or spaces) *)
+
+let locs_of_notation ?loc locs ntn =
+ let unloc loc = Option.cata Loc.unloc (0,0) loc in
+ let (bl, el) = unloc loc in
+ let locs = List.map unloc locs in
+ let rec aux pos = function
+ | [] -> if Int.equal pos el then [] else [(pos,el)]
+ | (ba,ea)::l -> if Int.equal pos ba then aux ea l else (pos,ba)::aux ea l
+ in aux bl (List.sort (fun l1 l2 -> fst l1 - fst l2) locs)
+
+let ntn_loc ?loc (args,argslist,binderslist) =
+ locs_of_notation ?loc
+ (List.map constr_loc (args@List.flatten argslist)@
+ List.map local_binders_loc binderslist)
+
+let patntn_loc ?loc (args,argslist) =
+ locs_of_notation ?loc
+ (List.map cases_pattern_expr_loc (args@List.flatten argslist))
+
+let error_invalid_pattern_notation ?loc () =
+ CErrors.user_err ?loc (str "Invalid notation for pattern.")
+
+(* Interpret the index of a recursion order annotation *)
+let split_at_annot bl na =
+ let names = List.map snd (names_of_local_assums bl) in
+ match na with
+ | None ->
+ begin match names with
+ | [] -> CErrors.user_err (Pp.str "A fixpoint needs at least one parameter.")
+ | _ -> ([], bl)
+ end
+ | Some (loc, id) ->
+ let rec aux acc = function
+ | CLocalAssum (bls, k, t) as x :: rest ->
+ let test (_, na) = match na with
+ | Name id' -> Id.equal id id'
+ | Anonymous -> false
+ in
+ let l, r = List.split_when test bls in
+ begin match r with
+ | [] -> aux (x :: acc) rest
+ | _ ->
+ let ans = match l with
+ | [] -> acc
+ | _ -> CLocalAssum (l, k, t) :: acc
+ in
+ (List.rev ans, CLocalAssum (r, k, t) :: rest)
+ end
+ | CLocalDef ((_,na),_,_) as x :: rest ->
+ if Name.equal (Name id) na then
+ CErrors.user_err ?loc
+ (Id.print id ++ str" must be a proper parameter and not a local definition.")
+ else
+ aux (x :: acc) rest
+ | CLocalPattern (_,_) :: rest ->
+ Loc.raise ?loc (Stream.Error "pattern with quote not allowed after fix")
+ | [] ->
+ CErrors.user_err ?loc
+ (str "No parameter named " ++ Id.print id ++ str".")
+ in aux [] bl
+
(** Pseudo-constructors *)
let mkIdentC id = CAst.make @@ CRef (Ident (Loc.tag id),None)
@@ -265,38 +530,40 @@ let add_name_in_env env n =
| Anonymous -> env
| Name id -> id :: env
-let (fresh_var, fresh_var_hook) = Hook.make ~default:(fun _ _ -> assert false) ()
+let fresh_var env c =
+ Namegen.next_ident_away (Id.of_string "pat")
+ (List.fold_left (fun accu id -> Id.Set.add id accu) (free_vars_of_constr_expr c) env)
let expand_binders ?loc mkC bl c =
let rec loop ?loc bl c =
match bl with
| [] -> ([], c)
| b :: bl ->
- match b with
- | CLocalDef ((loc1,_) as n, oty, b) ->
- let env, c = loop ?loc:(Loc.merge_opt loc1 loc) bl c in
- let env = add_name_in_env env n in
- (env, CAst.make ?loc @@ CLetIn (n,oty,b,c))
- | CLocalAssum ((loc1,_)::_ as nl, bk, t) ->
- let env, c = loop ?loc:(Loc.merge_opt loc1 loc) bl c in
- let env = List.fold_left add_name_in_env env nl in
- (env, mkC ?loc (nl,bk,t) c)
- | CLocalAssum ([],_,_) -> loop ?loc bl c
- | CLocalPattern (loc1, (p, ty)) ->
- let env, c = loop ?loc:(Loc.merge_opt loc1 loc) bl c in
- let ni = Hook.get fresh_var env c in
- let id = (loc1, Name ni) in
- let ty = match ty with
- | Some ty -> ty
- | None -> CAst.make ?loc:loc1 @@ CHole (None, IntroAnonymous, None)
- in
- let e = CAst.make @@ CRef (Libnames.Ident (loc1, ni), None) in
- let c = CAst.make ?loc @@
- CCases
- (LetPatternStyle, None, [(e,None,None)],
- [(Loc.tag ?loc:loc1 ([(loc1,[p])], c))])
- in
- (ni :: env, mkC ?loc ([id],Default Explicit,ty) c)
+ match b with
+ | CLocalDef ((loc1,_) as n, oty, b) ->
+ let env, c = loop ?loc:(Loc.merge_opt loc1 loc) bl c in
+ let env = add_name_in_env env n in
+ (env, CAst.make ?loc @@ CLetIn (n,oty,b,c))
+ | CLocalAssum ((loc1,_)::_ as nl, bk, t) ->
+ let env, c = loop ?loc:(Loc.merge_opt loc1 loc) bl c in
+ let env = List.fold_left add_name_in_env env nl in
+ (env, mkC ?loc (nl,bk,t) c)
+ | CLocalAssum ([],_,_) -> loop ?loc bl c
+ | CLocalPattern (loc1, (p, ty)) ->
+ let env, c = loop ?loc:(Loc.merge_opt loc1 loc) bl c in
+ let ni = fresh_var env c in
+ let id = (loc1, Name ni) in
+ let ty = match ty with
+ | Some ty -> ty
+ | None -> CAst.make ?loc:loc1 @@ CHole (None, IntroAnonymous, None)
+ in
+ let e = CAst.make @@ CRef (Libnames.Ident (loc1, ni), None) in
+ let c = CAst.make ?loc @@
+ CCases
+ (LetPatternStyle, None, [(e,None,None)],
+ [(Loc.tag ?loc:loc1 ([(loc1,[p])], c))])
+ in
+ (ni :: env, mkC ?loc ([id],Default Explicit,ty) c)
in
let (_, c) = loop ?loc bl c in
c
@@ -309,24 +576,34 @@ let mkCLambdaN ?loc bll c =
let mk ?loc b c = CAst.make ?loc @@ CLambdaN ([b],c) in
expand_binders ?loc mk bll c
-(* Deprecated *)
-let abstract_constr_expr c bl = mkCLambdaN ?loc:(local_binders_loc bl) bl c
-let prod_constr_expr c bl = mkCProdN ?loc:(local_binders_loc bl) bl c
-
let coerce_reference_to_id = function
| Ident (_,id) -> id
| Qualid (loc,_) ->
- CErrors.user_err ?loc ~hdr:"coerce_reference_to_id"
- (str "This expression should be a simple identifier.")
+ CErrors.user_err ?loc ~hdr:"coerce_reference_to_id"
+ (str "This expression should be a simple identifier.")
let coerce_to_id = function
| { CAst.v = CRef (Ident (loc,id),None) } -> (loc,id)
| { CAst.loc; _ } -> CErrors.user_err ?loc
- ~hdr:"coerce_to_id"
- (str "This expression should be a simple identifier.")
+ ~hdr:"coerce_to_id"
+ (str "This expression should be a simple identifier.")
let coerce_to_name = function
| { CAst.v = CRef (Ident (loc,id),None) } -> (loc,Name id)
| { CAst.loc; CAst.v = CHole (None,Misctypes.IntroAnonymous,None) } -> (loc,Anonymous)
| { CAst.loc; _ } -> CErrors.user_err ?loc ~hdr:"coerce_to_name"
(str "This expression should be a name.")
+
+let asymmetric_patterns = ref (false)
+let _ = Goptions.declare_bool_option {
+ Goptions.optdepr = false;
+ Goptions.optname = "no parameters in constructors";
+ Goptions.optkey = ["Asymmetric";"Patterns"];
+ Goptions.optread = (fun () -> !asymmetric_patterns);
+ Goptions.optwrite = (fun a -> asymmetric_patterns:=a);
+}
+
+(************************************************************************)
+(* Deprecated *)
+let abstract_constr_expr c bl = mkCLambdaN ?loc:(local_binders_loc bl) bl c
+let prod_constr_expr c bl = mkCProdN ?loc:(local_binders_loc bl) bl c
diff --git a/interp/constrexpr_ops.mli b/interp/constrexpr_ops.mli
index 7bd275e51..3ecb3d321 100644
--- a/interp/constrexpr_ops.mli
+++ b/interp/constrexpr_ops.mli
@@ -56,11 +56,11 @@ val mkCProdN : ?loc:Loc.t -> local_binder_expr list -> constr_expr -> constr_exp
(** @deprecated variant of mkCLambdaN *)
val abstract_constr_expr : constr_expr -> local_binder_expr list -> constr_expr
+[@@ocaml.deprecated "deprecated variant of mkCLambdaN"]
(** @deprecated variant of mkCProdN *)
val prod_constr_expr : constr_expr -> local_binder_expr list -> constr_expr
-
-val fresh_var_hook : (Names.Id.t list -> Constrexpr.constr_expr -> Names.Id.t) Hook.t
+[@@ocaml.deprecated "deprecated variant of mkCProdN"]
(** {6 Destructors}*)
@@ -83,3 +83,36 @@ val names_of_local_binders : local_binder_expr list -> Name.t located list
val names_of_local_assums : local_binder_expr list -> Name.t located list
(** Same as [names_of_local_binder_exprs], but does not take the [let] bindings into
account. *)
+
+(** { 6 Folds and maps } *)
+
+(** Used in typeclasses *)
+val fold_constr_expr_with_binders : (Id.t -> 'a -> 'a) ->
+ ('a -> 'b -> constr_expr -> 'b) -> 'a -> 'b -> constr_expr -> 'b
+
+(** Used in correctness and interface; absence of var capture not guaranteed
+ in pattern-matching clauses and in binders of the form [x,y:T(x)] *)
+
+val map_constr_expr_with_binders :
+ (Id.t -> 'a -> 'a) -> ('a -> constr_expr -> constr_expr) ->
+ 'a -> constr_expr -> constr_expr
+
+val replace_vars_constr_expr :
+ Id.t Id.Map.t -> constr_expr -> constr_expr
+
+(** Specific function for interning "in indtype" syntax of "match" *)
+val ids_of_cases_indtype : cases_pattern_expr -> Id.Set.t
+
+val free_vars_of_constr_expr : constr_expr -> Id.Set.t
+val occur_var_constr_expr : Id.t -> constr_expr -> bool
+
+val split_at_annot : local_binder_expr list -> Id.t located option -> local_binder_expr list * local_binder_expr list
+
+val ntn_loc : ?loc:Loc.t -> constr_notation_substitution -> string -> (int * int) list
+val patntn_loc : ?loc:Loc.t -> cases_pattern_notation_substitution -> string -> (int * int) list
+
+(** For cases pattern parsing errors *)
+val error_invalid_pattern_notation : ?loc:Loc.t -> unit -> 'a
+
+(** Placeholder for global option, should be moved to a parameter *)
+val asymmetric_patterns : bool ref
diff --git a/interp/constrextern.ml b/interp/constrextern.ml
index e1cf8f196..a92f97481 100644
--- a/interp/constrextern.ml
+++ b/interp/constrextern.ml
@@ -21,7 +21,6 @@ open CAst
open Constrexpr
open Constrexpr_ops
open Notation_ops
-open Topconstr
open Glob_term
open Glob_ops
open Pattern
@@ -424,7 +423,7 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat =
with
Not_found | No_match | Exit ->
let c = extern_reference ?loc Id.Set.empty (ConstructRef cstrsp) in
- if !Topconstr.asymmetric_patterns then
+ if !asymmetric_patterns then
if pattern_printable_in_both_syntax cstrsp
then CPatCstr (c, None, args)
else CPatCstr (c, Some (add_patt_for_params (fst cstrsp) args), [])
@@ -456,7 +455,7 @@ and apply_notation_to_pattern ?loc gr ((subst,substlist),(nb_to_drop,more_args))
List.map (extern_cases_pattern_in_scope subscope vars) c)
substlist in
let l2 = List.map (extern_cases_pattern_in_scope allscopes vars) more_args in
- let l2' = if !Topconstr.asymmetric_patterns || not (List.is_empty ll) then l2
+ let l2' = if !asymmetric_patterns || not (List.is_empty ll) then l2
else
match drop_implicits_in_patt gr nb_to_drop l2 with
|Some true_args -> true_args
@@ -472,7 +471,7 @@ and apply_notation_to_pattern ?loc gr ((subst,substlist),(nb_to_drop,more_args))
extern_cases_pattern_in_scope (scopt,scl@scopes) vars c)
subst in
let l2 = List.map (extern_cases_pattern_in_scope allscopes vars) more_args in
- let l2' = if !Topconstr.asymmetric_patterns then l2
+ let l2' = if !asymmetric_patterns then l2
else
match drop_implicits_in_patt gr (nb_to_drop + List.length l1) l2 with
|Some true_args -> true_args
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index c4e0ac500..ee3e11f8a 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -24,7 +24,6 @@ open Constrexpr
open Constrexpr_ops
open Notation_term
open Notation_ops
-open Topconstr
open Nametab
open Notation
open Inductiveops
diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml
index a5302b54d..519f2480b 100644
--- a/interp/implicit_quantifiers.ml
+++ b/interp/implicit_quantifiers.ml
@@ -94,8 +94,8 @@ let free_vars_of_constr_expr c ?(bound=Id.Set.empty) l =
let rec aux bdvars l c = match CAst.(c.v) with
| CRef (Ident (loc,id),_) -> found loc id bdvars l
| CNotation ("{ _ : _ | _ }", ({ CAst.v = CRef (Ident (_, id),_) } :: _, [], [])) when not (Id.Set.mem id bdvars) ->
- Topconstr.fold_constr_expr_with_binders (fun a l -> Id.Set.add a l) aux (Id.Set.add id bdvars) l c
- | _ -> Topconstr.fold_constr_expr_with_binders (fun a l -> Id.Set.add a l) aux bdvars l c
+ Constrexpr_ops.fold_constr_expr_with_binders (fun a l -> Id.Set.add a l) aux (Id.Set.add id bdvars) l c
+ | _ -> Constrexpr_ops.fold_constr_expr_with_binders (fun a l -> Id.Set.add a l) aux bdvars l c
in aux bound l c
let ids_of_names l =
diff --git a/interp/interp.mllib b/interp/interp.mllib
index e3500cfea..bb22cf468 100644
--- a/interp/interp.mllib
+++ b/interp/interp.mllib
@@ -1,13 +1,13 @@
Tactypes
Stdarg
Genintern
-Constrexpr_ops
Notation_ops
-Ppextend
Notation
-Dumpglob
Syntax_def
Smartlocate
+Constrexpr_ops
+Ppextend
+Dumpglob
Topconstr
Reserve
Impargs
diff --git a/interp/topconstr.ml b/interp/topconstr.ml
index c64d3aa26..ecfb766ff 100644
--- a/interp/topconstr.ml
+++ b/interp/topconstr.ml
@@ -6,294 +6,16 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i*)
-open Pp
-open CErrors
-open Util
-open Names
-open Nameops
-open Libnames
-open Misctypes
-open Constrexpr
open Constrexpr_ops
-(*i*)
-
-let asymmetric_patterns = ref (false)
-let _ = Goptions.declare_bool_option {
- Goptions.optdepr = false;
- Goptions.optname = "no parameters in constructors";
- Goptions.optkey = ["Asymmetric";"Patterns"];
- Goptions.optread = (fun () -> !asymmetric_patterns);
- Goptions.optwrite = (fun a -> asymmetric_patterns:=a);
-}
-
-(**********************************************************************)
-(* Miscellaneous *)
-
-let error_invalid_pattern_notation ?loc () =
- user_err ?loc (str "Invalid notation for pattern.")
-
-(* Legacy functions *)
-let down_located f (_l, x) = f x
-let located_fold_left f x (_l, y) = f x y
-
-(**********************************************************************)
-(* Functions on constr_expr *)
-
-let is_constructor id =
- try Globnames.isConstructRef
- (Smartlocate.global_of_extended_global
- (Nametab.locate_extended (qualid_of_ident id)))
- with Not_found -> false
-
-let rec cases_pattern_fold_names f a pt = match CAst.(pt.v) with
- | CPatRecord l ->
- List.fold_left (fun acc (r, cp) -> cases_pattern_fold_names f acc cp) a l
- | CPatAlias (pat,id) -> f id a
- | CPatOr (patl) ->
- List.fold_left (cases_pattern_fold_names f) a patl
- | CPatCstr (_,patl1,patl2) ->
- List.fold_left (cases_pattern_fold_names f)
- (Option.fold_left (List.fold_left (cases_pattern_fold_names f)) a patl1) patl2
- | CPatNotation (_,(patl,patll),patl') ->
- List.fold_left (cases_pattern_fold_names f)
- (List.fold_left (cases_pattern_fold_names f) a (patl@List.flatten patll)) patl'
- | CPatDelimiters (_,pat) -> cases_pattern_fold_names f a pat
- | CPatAtom (Some (Ident (_,id))) when not (is_constructor id) -> f id a
- | CPatPrim _ | CPatAtom _ -> a
- | CPatCast ({CAst.loc},_) ->
- CErrors.user_err ?loc ~hdr:"cases_pattern_fold_names"
- (Pp.strbrk "Casts are not supported here.")
-
-let ids_of_pattern =
- cases_pattern_fold_names Id.Set.add Id.Set.empty
-
-let ids_of_pattern_list =
- List.fold_left
- (located_fold_left
- (List.fold_left (cases_pattern_fold_names Id.Set.add)))
- Id.Set.empty
-
-let ids_of_cases_indtype p =
- cases_pattern_fold_names Id.Set.add Id.Set.empty p
-
-let ids_of_cases_tomatch tms =
- List.fold_right
- (fun (_, ona, indnal) l ->
- Option.fold_right (fun t ids -> cases_pattern_fold_names Id.Set.add ids t)
- indnal
- (Option.fold_right (down_located (Name.fold_right Id.Set.add)) ona l))
- tms Id.Set.empty
-
-let rec fold_constr_expr_binders g f n acc b = function
- | (nal,bk,t)::l ->
- let nal = snd (List.split nal) in
- let n' = List.fold_right (Name.fold_right g) nal n in
- f n (fold_constr_expr_binders g f n' acc b l) t
- | [] ->
- f n acc b
-
-let rec fold_local_binders g f n acc b = function
- | CLocalAssum (nal,bk,t)::l ->
- let nal = snd (List.split nal) in
- let n' = List.fold_right (Name.fold_right g) nal n in
- f n (fold_local_binders g f n' acc b l) t
- | CLocalDef ((_,na),c,t)::l ->
- Option.fold_left (f n) (f n (fold_local_binders g f (Name.fold_right g na n) acc b l) c) t
- | CLocalPattern (_,(pat,t))::l ->
- let acc = fold_local_binders g f (cases_pattern_fold_names g n pat) acc b l in
- Option.fold_left (f n) acc t
- | [] ->
- f n acc b
-
-let fold_constr_expr_with_binders g f n acc = CAst.with_val (function
- | CAppExpl ((_,_,_),l) -> List.fold_left (f n) acc l
- | CApp ((_,t),l) -> List.fold_left (f n) (f n acc t) (List.map fst l)
- | CProdN (l,b) | CLambdaN (l,b) -> fold_constr_expr_binders g f n acc b l
- | CLetIn (na,a,t,b) ->
- f (Name.fold_right g (snd na) n) (Option.fold_left (f n) (f n acc a) t) b
- | CCast (a,(CastConv b|CastVM b|CastNative b)) -> f n (f n acc a) b
- | CCast (a,CastCoerce) -> f n acc a
- | CNotation (_,(l,ll,bll)) ->
- (* The following is an approximation: we don't know exactly if
- an ident is binding nor to which subterms bindings apply *)
- let acc = List.fold_left (f n) acc (l@List.flatten ll) in
- List.fold_left (fun acc bl -> fold_local_binders g f n acc (CAst.make @@ CHole (None,IntroAnonymous,None)) bl) acc bll
- | CGeneralization (_,_,c) -> f n acc c
- | CDelimiters (_,a) -> f n acc a
- | CHole _ | CEvar _ | CPatVar _ | CSort _ | CPrim _ | CRef _ ->
- acc
- | CRecord l -> List.fold_left (fun acc (id, c) -> f n acc c) acc l
- | CCases (sty,rtnpo,al,bl) ->
- let ids = ids_of_cases_tomatch al in
- let acc = Option.fold_left (f (Id.Set.fold g ids n)) acc rtnpo in
- let acc = List.fold_left (f n) acc (List.map (fun (fst,_,_) -> fst) al) in
- List.fold_right (fun (loc,(patl,rhs)) acc ->
- let ids = ids_of_pattern_list patl in
- f (Id.Set.fold g ids n) acc rhs) bl acc
- | CLetTuple (nal,(ona,po),b,c) ->
- let n' = List.fold_right (down_located (Name.fold_right g)) nal n in
- f (Option.fold_right (down_located (Name.fold_right g)) ona n') (f n acc b) c
- | CIf (c,(ona,po),b1,b2) ->
- let acc = f n (f n (f n acc b1) b2) c in
- Option.fold_left
- (f (Option.fold_right (down_located (Name.fold_right g)) ona n)) acc po
- | CFix (_,l) ->
- let n' = List.fold_right (fun ((_,id),_,_,_,_) -> g id) l n in
- List.fold_right (fun (_,(_,o),lb,t,c) acc ->
- fold_local_binders g f n'
- (fold_local_binders g f n acc t lb) c lb) l acc
- | CCoFix (_,_) ->
- Feedback.msg_warning (strbrk "Capture check in multiple binders not done"); acc
- )
-
-let free_vars_of_constr_expr c =
- let rec aux bdvars l = function
- | { CAst.v = CRef (Ident (_,id),_) } -> if Id.List.mem id bdvars then l else Id.Set.add id l
- | c -> fold_constr_expr_with_binders (fun a l -> a::l) aux bdvars l c
- in aux [] Id.Set.empty c
-
-let occur_var_constr_expr id c = Id.Set.mem id (free_vars_of_constr_expr c)
-
-(* Interpret the index of a recursion order annotation *)
-
-let split_at_annot bl na =
- let names = List.map snd (names_of_local_assums bl) in
- match na with
- | None ->
- begin match names with
- | [] -> user_err (Pp.str "A fixpoint needs at least one parameter.")
- | _ -> ([], bl)
- end
- | Some (loc, id) ->
- let rec aux acc = function
- | CLocalAssum (bls, k, t) as x :: rest ->
- let test (_, na) = match na with
- | Name id' -> Id.equal id id'
- | Anonymous -> false
- in
- let l, r = List.split_when test bls in
- begin match r with
- | [] -> aux (x :: acc) rest
- | _ ->
- let ans = match l with
- | [] -> acc
- | _ -> CLocalAssum (l, k, t) :: acc
- in
- (List.rev ans, CLocalAssum (r, k, t) :: rest)
- end
- | CLocalDef ((_,na),_,_) as x :: rest ->
- if Name.equal (Name id) na then
- user_err ?loc
- (Id.print id ++ str" must be a proper parameter and not a local definition.")
- else
- aux (x :: acc) rest
- | CLocalPattern (_,_) :: rest ->
- Loc.raise ?loc (Stream.Error "pattern with quote not allowed after fix")
- | [] ->
- user_err ?loc
- (str "No parameter named " ++ Id.print id ++ str".")
- in aux [] bl
-
-(* Used in correctness and interface *)
-
-let map_binder g e nal = List.fold_right (down_located (Name.fold_right g)) nal e
-
-let map_binders f g e bl =
- (* TODO: avoid variable capture in [t] by some [na] in [List.tl nal] *)
- let h (e,bl) (nal,bk,t) = (map_binder g e nal,(nal,bk,f e t)::bl) in
- let (e,rbl) = List.fold_left h (e,[]) bl in
- (e, List.rev rbl)
-
-let map_local_binders f g e bl =
- (* TODO: avoid variable capture in [t] by some [na] in [List.tl nal] *)
- let h (e,bl) = function
- CLocalAssum(nal,k,ty) ->
- (map_binder g e nal, CLocalAssum(nal,k,f e ty)::bl)
- | CLocalDef((loc,na),c,ty) ->
- (Name.fold_right g na e, CLocalDef((loc,na),f e c,Option.map (f e) ty)::bl)
- | CLocalPattern (loc,(pat,t)) ->
- let ids = ids_of_pattern pat in
- (Id.Set.fold g ids e, CLocalPattern (loc,(pat,Option.map (f e) t))::bl) in
- let (e,rbl) = List.fold_left h (e,[]) bl in
- (e, List.rev rbl)
-
-let map_constr_expr_with_binders g f e = CAst.map (function
- | CAppExpl (r,l) -> CAppExpl (r,List.map (f e) l)
- | CApp ((p,a),l) ->
- CApp ((p,f e a),List.map (fun (a,i) -> (f e a,i)) l)
- | CProdN (bl,b) ->
- let (e,bl) = map_binders f g e bl in CProdN (bl,f e b)
- | CLambdaN (bl,b) ->
- let (e,bl) = map_binders f g e bl in CLambdaN (bl,f e b)
- | CLetIn (na,a,t,b) ->
- CLetIn (na,f e a,Option.map (f e) t,f (Name.fold_right g (snd na) e) b)
- | CCast (a,c) -> CCast (f e a, Miscops.map_cast_type (f e) c)
- | CNotation (n,(l,ll,bll)) ->
- (* This is an approximation because we don't know what binds what *)
- CNotation (n,(List.map (f e) l,List.map (List.map (f e)) ll,
- List.map (fun bl -> snd (map_local_binders f g e bl)) bll))
- | CGeneralization (b,a,c) -> CGeneralization (b,a,f e c)
- | CDelimiters (s,a) -> CDelimiters (s,f e a)
- | CHole _ | CEvar _ | CPatVar _ | CSort _
- | CPrim _ | CRef _ as x -> x
- | CRecord l -> CRecord (List.map (fun (id, c) -> (id, f e c)) l)
- | CCases (sty,rtnpo,a,bl) ->
- let bl = List.map (fun (loc,(patl,rhs)) ->
- let ids = ids_of_pattern_list patl in
- (loc,(patl,f (Id.Set.fold g ids e) rhs))) bl in
- let ids = ids_of_cases_tomatch a in
- let po = Option.map (f (Id.Set.fold g ids e)) rtnpo in
- CCases (sty, po, List.map (fun (tm,x,y) -> f e tm,x,y) a,bl)
- | CLetTuple (nal,(ona,po),b,c) ->
- let e' = List.fold_right (down_located (Name.fold_right g)) nal e in
- let e'' = Option.fold_right (down_located (Name.fold_right g)) ona e in
- CLetTuple (nal,(ona,Option.map (f e'') po),f e b,f e' c)
- | CIf (c,(ona,po),b1,b2) ->
- let e' = Option.fold_right (down_located (Name.fold_right g)) ona e in
- CIf (f e c,(ona,Option.map (f e') po),f e b1,f e b2)
- | CFix (id,dl) ->
- CFix (id,List.map (fun (id,n,bl,t,d) ->
- let (e',bl') = map_local_binders f g e bl in
- let t' = f e' t in
- (* Note: fix names should be inserted before the arguments... *)
- let e'' = List.fold_left (fun e ((_,id),_,_,_,_) -> g id e) e' dl in
- let d' = f e'' d in
- (id,n,bl',t',d')) dl)
- | CCoFix (id,dl) ->
- CCoFix (id,List.map (fun (id,bl,t,d) ->
- let (e',bl') = map_local_binders f g e bl in
- let t' = f e' t in
- let e'' = List.fold_left (fun e ((_,id),_,_,_) -> g id e) e' dl in
- let d' = f e'' d in
- (id,bl',t',d')) dl)
- )
-
-(* Used in constrintern *)
-let rec replace_vars_constr_expr l = function
- | { CAst.loc; v = CRef (Ident (loc_id,id),us) } as x ->
- (try CAst.make ?loc @@ CRef (Ident (loc_id,Id.Map.find id l),us) with Not_found -> x)
- | c -> map_constr_expr_with_binders Id.Map.remove
- replace_vars_constr_expr l c
-
-(* Returns the ranges of locs of the notation that are not occupied by args *)
-(* and which are then occupied by proper symbols of the notation (or spaces) *)
-
-let locs_of_notation ?loc locs ntn =
- let unloc loc = Option.cata Loc.unloc (0,0) loc in
- let (bl, el) = unloc loc in
- let locs = List.map unloc locs in
- let rec aux pos = function
- | [] -> if Int.equal pos el then [] else [(pos,el)]
- | (ba,ea)::l -> if Int.equal pos ba then aux ea l else (pos,ba)::aux ea l
- in aux bl (List.sort (fun l1 l2 -> fst l1 - fst l2) locs)
-
-let ntn_loc ?loc (args,argslist,binderslist) =
- locs_of_notation ?loc
- (List.map constr_loc (args@List.flatten argslist)@
- List.map local_binders_loc binderslist)
-
-let patntn_loc ?loc (args,argslist) =
- locs_of_notation ?loc
- (List.map cases_pattern_expr_loc (args@List.flatten argslist))
+let asymmetric_patterns = asymmetric_patterns
+let error_invalid_pattern_notation = error_invalid_pattern_notation
+let split_at_annot = split_at_annot
+let ntn_loc = ntn_loc
+let patntn_loc = patntn_loc
+let map_constr_expr_with_binders = map_constr_expr_with_binders
+let fold_constr_expr_with_binders = fold_constr_expr_with_binders
+let ids_of_cases_indtype = ids_of_cases_indtype
+let occur_var_constr_expr = occur_var_constr_expr
+let free_vars_of_constr_expr = free_vars_of_constr_expr
+let replace_vars_constr_expr = replace_vars_constr_expr
diff --git a/interp/topconstr.mli b/interp/topconstr.mli
index 922f87955..9fc02461e 100644
--- a/interp/topconstr.mli
+++ b/interp/topconstr.mli
@@ -10,40 +10,43 @@ open Loc
open Names
open Constrexpr
-(** Topconstr *)
-
+(** Topconstr: This whole module is deprecated in favor of Constrexpr_ops *)
val asymmetric_patterns : bool ref
+[@@ocaml.deprecated "use Constrexpr_ops.asymmetric_patterns"]
(** Utilities on constr_expr *)
+val split_at_annot : local_binder_expr list -> Id.t located option -> local_binder_expr list * local_binder_expr list
+[@@ocaml.deprecated "use Constrexpr_ops.split_at_annot"]
+
+val ntn_loc : ?loc:Loc.t -> constr_notation_substitution -> string -> (int * int) list
+[@@ocaml.deprecated "use Constrexpr_ops.ntn_loc"]
+val patntn_loc : ?loc:Loc.t -> cases_pattern_notation_substitution -> string -> (int * int) list
+[@@ocaml.deprecated "use Constrexpr_ops.patntn_loc"]
+
+(** For cases pattern parsing errors *)
+val error_invalid_pattern_notation : ?loc:Loc.t -> unit -> 'a
+[@@ocaml.deprecated "use Constrexpr_ops.error_invalid_pattern_notation"]
-val replace_vars_constr_expr :
- Id.t Id.Map.t -> constr_expr -> constr_expr
+(*************************************************************************)
+val replace_vars_constr_expr : Id.t Id.Map.t -> constr_expr -> constr_expr
+[@@ocaml.deprecated "use Constrexpr_ops.free_vars_of_constr_expr"]
val free_vars_of_constr_expr : constr_expr -> Id.Set.t
+[@@ocaml.deprecated "use Constrexpr_ops.free_vars_of_constr_expr"]
+
val occur_var_constr_expr : Id.t -> constr_expr -> bool
+[@@ocaml.deprecated "use Constrexpr_ops.occur_var_constr_expr"]
(** Specific function for interning "in indtype" syntax of "match" *)
val ids_of_cases_indtype : cases_pattern_expr -> Id.Set.t
-
-val split_at_annot : local_binder_expr list -> Id.t located option -> local_binder_expr list * local_binder_expr list
+[@@ocaml.deprecated "use Constrexpr_ops.ids_of_cases_indtype"]
(** Used in typeclasses *)
-
val fold_constr_expr_with_binders : (Id.t -> 'a -> 'a) ->
('a -> 'b -> constr_expr -> 'b) -> 'a -> 'b -> constr_expr -> 'b
-
-(** Used in correctness and interface; absence of var capture not guaranteed
- in pattern-matching clauses and in binders of the form [x,y:T(x)] *)
+[@@ocaml.deprecated "use Constrexpr_ops.fold_constr_expr_with_binders"]
val map_constr_expr_with_binders :
(Id.t -> 'a -> 'a) -> ('a -> constr_expr -> constr_expr) ->
'a -> constr_expr -> constr_expr
-
-val ntn_loc :
- ?loc:Loc.t -> constr_notation_substitution -> string -> (int * int) list
-val patntn_loc :
- ?loc:Loc.t -> cases_pattern_notation_substitution -> string -> (int * int) list
-
-(** For cases pattern parsing errors *)
-
-val error_invalid_pattern_notation : ?loc:Loc.t -> unit -> 'a
+[@@ocaml.deprecated "use Constrexpr_ops.map_constr_expr_with_binders"]
diff --git a/parsing/egramcoq.ml b/parsing/egramcoq.ml
index d51b8b54e..7f50fd22a 100644
--- a/parsing/egramcoq.ml
+++ b/parsing/egramcoq.ml
@@ -440,7 +440,7 @@ let make_act : type r. r target -> _ -> r gen_eval = function
CAst.make ~loc @@ CNotation (notation , env)
| ForPattern -> fun notation loc env ->
let invalid = List.exists (fun (_, b) -> not b) env.binders in
- let () = if invalid then Topconstr.error_invalid_pattern_notation ~loc () in
+ let () = if invalid then Constrexpr_ops.error_invalid_pattern_notation ~loc () in
let env = (env.constrs, env.constrlists) in
CAst.make ~loc @@ CPatNotation (notation, env, [])
diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4
index a5b58b855..82306bb9f 100644
--- a/parsing/g_vernac.ml4
+++ b/parsing/g_vernac.ml4
@@ -129,12 +129,6 @@ let test_plural_form_types loc kwd = function
warn_plural_command ~loc:!@loc kwd
| _ -> ()
-let fresh_var env c =
- Namegen.next_ident_away (Id.of_string "pat")
- (List.fold_left (fun accu id -> Id.Set.add id accu) (Topconstr.free_vars_of_constr_expr c) env)
-
-let _ = Hook.set Constrexpr_ops.fresh_var_hook fresh_var
-
(* Gallina declarations *)
GEXTEND Gram
GLOBAL: gallina gallina_ext thm_token def_body of_type_with_opt_coercion
diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml
index 0f6452de6..bce5710d6 100644
--- a/printing/ppconstr.ml
+++ b/printing/ppconstr.ml
@@ -396,7 +396,7 @@ let tag_var = tag Tag.variable
extract_prod_binders c
| { loc; v = CProdN ([[_,Name id],bk,t],
{ v = CCases (LetPatternStyle,None, [{ v = CRef (Ident (_,id'),None)},None,None],[(_,([_,[p]],b))])} ) }
- when Id.equal id id' && not (Id.Set.mem id (Topconstr.free_vars_of_constr_expr b)) ->
+ when Id.equal id id' && not (Id.Set.mem id (free_vars_of_constr_expr b)) ->
let bl,c = extract_prod_binders b in
CLocalPattern (loc, (p,None)) :: bl, c
| { loc; v = CProdN ((nal,bk,t)::bl,c) } ->
@@ -412,7 +412,7 @@ let tag_var = tag Tag.variable
extract_lam_binders c
| CLambdaN ([[_,Name id],bk,t],
{ v = CCases (LetPatternStyle,None, [{ v = CRef (Ident (_,id'),None)},None,None],[(_,([_,[p]],b))])} )
- when Id.equal id id' && not (Id.Set.mem id (Topconstr.free_vars_of_constr_expr b)) ->
+ when Id.equal id id' && not (Id.Set.mem id (free_vars_of_constr_expr b)) ->
let bl,c = extract_lam_binders b in
CLocalPattern (ce.loc,(p,None)) :: bl, c
| CLambdaN ((nal,bk,t)::bl,c) ->
@@ -430,7 +430,7 @@ let tag_var = tag Tag.variable
let rename na na' t c =
match (na,na') with
| (_,Name id), (_,Name id') ->
- (na',t,Topconstr.replace_vars_constr_expr (Id.Map.singleton id id') c)
+ (na',t,replace_vars_constr_expr (Id.Map.singleton id id') c)
| (_,Name id), (_,Anonymous) -> (na,t,c)
| _ -> (na',t,c)
diff --git a/vernac/command.ml b/vernac/command.ml
index be54f97b7..0232d7376 100644
--- a/vernac/command.ml
+++ b/vernac/command.ml
@@ -22,7 +22,6 @@ open Globnames
open Nameops
open Constrexpr
open Constrexpr_ops
-open Topconstr
open Constrintern
open Nametab
open Impargs