aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--dev/base_include2
-rw-r--r--dev/db2
-rw-r--r--dev/include2
-rw-r--r--dev/ocamldoc/docintro2
-rw-r--r--dev/top_printers.ml2
-rw-r--r--interp/constrextern.ml120
-rw-r--r--interp/constrextern.mli4
-rw-r--r--interp/constrintern.ml158
-rw-r--r--interp/constrintern.mli18
-rw-r--r--interp/doc.tex2
-rw-r--r--interp/genarg.ml6
-rw-r--r--interp/genarg.mli32
-rw-r--r--interp/implicit_quantifiers.ml26
-rw-r--r--interp/implicit_quantifiers.mli6
-rw-r--r--interp/notation.ml56
-rw-r--r--interp/notation.mli18
-rw-r--r--interp/reserve.ml46
-rw-r--r--interp/reserve.mli2
-rw-r--r--interp/topconstr.ml244
-rw-r--r--interp/topconstr.mli32
-rw-r--r--parsing/g_xml.ml432
-rw-r--r--parsing/pptactic.ml12
-rw-r--r--parsing/pptactic.mli8
-rw-r--r--parsing/prettyp.ml4
-rw-r--r--parsing/printer.ml18
-rw-r--r--parsing/printer.mli8
-rw-r--r--parsing/q_constr.ml426
-rw-r--r--plugins/decl_mode/decl_expr.mli4
-rw-r--r--plugins/decl_mode/decl_interp.ml58
-rw-r--r--plugins/decl_mode/decl_proof_instr.ml2
-rw-r--r--plugins/firstorder/instances.ml4
-rw-r--r--plugins/funind/indfun.ml28
-rw-r--r--plugins/funind/indfun_common.ml6
-rw-r--r--plugins/funind/indfun_common.mli8
-rw-r--r--plugins/funind/merge.ml102
-rw-r--r--plugins/funind/rawterm_to_relation.ml198
-rw-r--r--plugins/funind/rawterm_to_relation.mli4
-rw-r--r--plugins/funind/rawtermops.ml416
-rw-r--r--plugins/funind/rawtermops.mli90
-rw-r--r--plugins/funind/recdef.ml10
-rw-r--r--plugins/setoid_ring/newring.ml42
-rw-r--r--plugins/subtac/subtac.ml2
-rw-r--r--plugins/subtac/subtac_cases.ml24
-rw-r--r--plugins/subtac/subtac_command.ml10
-rw-r--r--plugins/subtac/subtac_pretyping.ml6
-rw-r--r--plugins/subtac/subtac_pretyping.mli2
-rw-r--r--plugins/subtac/subtac_pretyping_F.ml48
-rw-r--r--plugins/subtac/subtac_utils.ml2
-rw-r--r--plugins/subtac/subtac_utils.mli2
-rw-r--r--plugins/syntax/ascii_syntax.ml12
-rw-r--r--plugins/syntax/nat_syntax.ml12
-rw-r--r--plugins/syntax/numbers_syntax.ml64
-rw-r--r--plugins/syntax/r_syntax.ml38
-rw-r--r--plugins/syntax/string_syntax.ml12
-rw-r--r--plugins/syntax/z_syntax.ml50
-rw-r--r--pretyping/cases.ml8
-rw-r--r--pretyping/cases.mli4
-rw-r--r--pretyping/detyping.ml164
-rw-r--r--pretyping/detyping.mli22
-rw-r--r--pretyping/pattern.ml44
-rw-r--r--pretyping/pattern.mli4
-rw-r--r--pretyping/pretyping.ml82
-rw-r--r--pretyping/pretyping.mli36
-rw-r--r--pretyping/rawterm.ml266
-rw-r--r--pretyping/rawterm.mli76
-rw-r--r--proofs/evar_refiner.ml2
-rw-r--r--proofs/evar_refiner.mli2
-rw-r--r--proofs/goal.ml2
-rw-r--r--proofs/goal.mli4
-rw-r--r--proofs/proof_type.ml2
-rw-r--r--proofs/proof_type.mli2
-rw-r--r--proofs/tacexpr.ml14
-rw-r--r--proofs/tactic_debug.mli4
-rw-r--r--tactics/dhyp.ml2
-rw-r--r--tactics/evar_tactics.mli2
-rw-r--r--tactics/extraargs.ml48
-rw-r--r--tactics/extraargs.mli4
-rw-r--r--tactics/extratactics.ml412
-rw-r--r--tactics/tacinterp.ml108
-rw-r--r--tactics/tacinterp.mli10
-rw-r--r--toplevel/himsg.ml2
-rw-r--r--toplevel/vernacentries.ml10
-rw-r--r--toplevel/whelp.ml430
83 files changed, 1515 insertions, 1515 deletions
diff --git a/dev/base_include b/dev/base_include
index 23cc38fb4..19c0f4b4e 100644
--- a/dev/base_include
+++ b/dev/base_include
@@ -171,7 +171,7 @@ let parse_constr = Pcoq.parse_string Pcoq.Constr.constr;;
let parse_tac = Pcoq.parse_string Pcoq.Tactic.tactic;;
let parse_vernac = Pcoq.parse_string Pcoq.Vernac_.vernac;;
-(* build a term of type rawconstr without type-checking or resolution of
+(* build a term of type glob_constr without type-checking or resolution of
implicit syntax *)
let e s =
diff --git a/dev/db b/dev/db
index 9eed32e04..e7225e8f2 100644
--- a/dev/db
+++ b/dev/db
@@ -17,7 +17,7 @@ install_printer Top_printers.ppclindex
install_printer Top_printers.ppbigint
install_printer Top_printers.pppattern
-install_printer Top_printers.pprawconstr
+install_printer Top_printers.ppglob_constr
install_printer Top_printers.ppconstr
install_printer Top_printers.ppuni
diff --git a/dev/include b/dev/include
index 251a969b9..b72e68ac0 100644
--- a/dev/include
+++ b/dev/include
@@ -14,7 +14,7 @@
#install_printer (* pp_stdcmds *) pppp;;
#install_printer (* pattern *) pppattern;;
-#install_printer (* rawconstr *) pprawconstr;;
+#install_printer (* glob_constr *) ppglob_constr;;
#install_printer (* constr *) ppconstr;;
#install_printer (* constr_substituted *) ppsconstr;;
diff --git a/dev/ocamldoc/docintro b/dev/ocamldoc/docintro
index 20c3de5ef..33d20fc81 100644
--- a/dev/ocamldoc/docintro
+++ b/dev/ocamldoc/docintro
@@ -30,7 +30,7 @@ describes the Coq library, which is made of two parts:
describes the translation from Coq context-dependent
front abstract syntax of terms {v constr_expr v} to and from the
-context-free, untyped, globalized form of constructions {v rawconstr v}.}
+context-free, untyped, globalized form of constructions {v glob_constr v}.}
{- Parsers and printers : parsing
describes the implementation of the Coq parsers and printers.}
diff --git a/dev/top_printers.ml b/dev/top_printers.ml
index 1f05e90bd..89a6eb5e3 100644
--- a/dev/top_printers.ml
+++ b/dev/top_printers.ml
@@ -56,7 +56,7 @@ let ppconstrdb x = pp(Flags.with_option rawdebug Termops.print_constr x)
let ppterm = ppconstr
let ppsconstr x = ppconstr (Declarations.force x)
let ppconstr_univ x = Constrextern.with_universes ppconstr x
-let pprawconstr = (fun x -> pp(pr_lrawconstr x))
+let ppglob_constr = (fun x -> pp(pr_lglob_constr x))
let pppattern = (fun x -> pp(pr_constr_pattern x))
let pptype = (fun x -> try pp(pr_ltype x) with e -> pp (str (Printexc.to_string e)))
diff --git a/interp/constrextern.ml b/interp/constrextern.ml
index eb779200c..4029f6150 100644
--- a/interp/constrextern.ml
+++ b/interp/constrextern.ml
@@ -29,7 +29,7 @@ open Reserve
open Detyping
(*i*)
-(* Translation from rawconstr to front constr *)
+(* Translation from glob_constr to front constr *)
(**********************************************************************)
(* Parametrization *)
@@ -272,7 +272,7 @@ let make_pat_notation loc ntn (terms,termlists as subst) =
(fun (loc,p) -> CPatPrim (loc,p))
destPatPrim terms
- (* Better to use extern_rawconstr composed with injection/retraction ?? *)
+ (* Better to use extern_glob_constr composed with injection/retraction ?? *)
let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat =
try
if !Flags.raw_print or !print_no_symbol then raise No_match;
@@ -458,7 +458,7 @@ let rec extern_args extern scopes env args subscopes =
extern argscopes env a :: extern_args extern scopes env args subscopes
let rec remove_coercions inctx = function
- | RApp (loc,RRef (_,r),args) as c
+ | GApp (loc,GRef (_,r),args) as c
when not (!Flags.raw_print or !print_coercions)
->
let nargs = List.length args in
@@ -477,13 +477,13 @@ let rec remove_coercions inctx = function
been confused with ordinary application or would have need
a surrounding context and the coercion to funclass would
have been made explicit to match *)
- if l = [] then a' else RApp (loc,a',l)
+ if l = [] then a' else GApp (loc,a',l)
| _ -> c
with Not_found -> c)
| c -> c
let rec flatten_application = function
- | RApp (loc,RApp(_,a,l'),l) -> flatten_application (RApp (loc,a,l'@l))
+ | GApp (loc,GApp(_,a,l'),l) -> flatten_application (GApp (loc,a,l'@l))
| a -> a
(**********************************************************************)
@@ -495,7 +495,7 @@ let extern_possible_prim_token scopes r =
let (sc,n) = uninterp_prim_token r in
match availability_of_prim_token n sc scopes with
| None -> None
- | Some key -> Some (insert_delimiters (CPrim (loc_of_rawconstr r,n)) key)
+ | Some key -> Some (insert_delimiters (CPrim (loc_of_glob_constr r,n)) key)
with No_match ->
None
@@ -525,23 +525,23 @@ let rec extern inctx scopes vars r =
if !Flags.raw_print or !print_no_symbol then raise No_match;
extern_symbol scopes vars r'' (uninterp_notations r'')
with No_match -> match r' with
- | RRef (loc,ref) ->
+ | GRef (loc,ref) ->
extern_global loc (select_stronger_impargs (implicits_of_global ref))
(extern_reference loc vars ref)
- | RVar (loc,id) -> CRef (Ident (loc,id))
+ | GVar (loc,id) -> CRef (Ident (loc,id))
- | REvar (loc,n,None) when !print_meta_as_hole -> CHole (loc, None)
+ | GEvar (loc,n,None) when !print_meta_as_hole -> CHole (loc, None)
- | REvar (loc,n,l) ->
+ | GEvar (loc,n,l) ->
extern_evar loc n (Option.map (List.map (extern false scopes vars)) l)
- | RPatVar (loc,n) ->
+ | GPatVar (loc,n) ->
if !print_meta_as_hole then CHole (loc, None) else CPatVar (loc,n)
- | RApp (loc,f,args) ->
+ | GApp (loc,f,args) ->
(match f with
- | RRef (rloc,ref) ->
+ | GRef (rloc,ref) ->
let subscopes = find_arguments_scope ref in
let args =
extern_args (extern true) (snd scopes) vars args subscopes in
@@ -587,63 +587,63 @@ let rec extern inctx scopes vars r =
explicitize loc inctx [] (None,sub_extern false scopes vars f)
(List.map (sub_extern true scopes vars) args))
- | RProd (loc,Anonymous,_,t,c) ->
+ | GProd (loc,Anonymous,_,t,c) ->
(* Anonymous product are never factorized *)
CArrow (loc,extern_typ scopes vars t, extern_typ scopes vars c)
- | RLetIn (loc,na,t,c) ->
+ | GLetIn (loc,na,t,c) ->
CLetIn (loc,(loc,na),sub_extern false scopes vars t,
extern inctx scopes (add_vname vars na) c)
- | RProd (loc,na,bk,t,c) ->
+ | GProd (loc,na,bk,t,c) ->
let t = extern_typ scopes vars (anonymize_if_reserved na t) in
let (idl,c) = factorize_prod scopes (add_vname vars na) t c in
CProdN (loc,[(dummy_loc,na)::idl,Default bk,t],c)
- | RLambda (loc,na,bk,t,c) ->
+ | GLambda (loc,na,bk,t,c) ->
let t = extern_typ scopes vars (anonymize_if_reserved na t) in
let (idl,c) = factorize_lambda inctx scopes (add_vname vars na) t c in
CLambdaN (loc,[(dummy_loc,na)::idl,Default bk,t],c)
- | RCases (loc,sty,rtntypopt,tml,eqns) ->
+ | GCases (loc,sty,rtntypopt,tml,eqns) ->
let vars' =
List.fold_right (name_fold Idset.add)
(cases_predicate_names tml) vars in
let rtntypopt' = Option.map (extern_typ scopes vars') rtntypopt in
let tml = List.map (fun (tm,(na,x)) ->
let na' = match na,tm with
- Anonymous, RVar (_,id) when
- rtntypopt<>None & occur_rawconstr id (Option.get rtntypopt)
+ Anonymous, GVar (_,id) when
+ rtntypopt<>None & occur_glob_constr id (Option.get rtntypopt)
-> Some (dummy_loc,Anonymous)
| Anonymous, _ -> None
- | Name id, RVar (_,id') when id=id' -> None
+ | Name id, GVar (_,id') when id=id' -> None
| Name _, _ -> Some (dummy_loc,na) in
(sub_extern false scopes vars tm,
(na',Option.map (fun (loc,ind,n,nal) ->
let params = list_tabulate
- (fun _ -> RHole (dummy_loc,Evd.InternalHole)) n in
+ (fun _ -> GHole (dummy_loc,Evd.InternalHole)) n in
let args = List.map (function
- | Anonymous -> RHole (dummy_loc,Evd.InternalHole)
- | Name id -> RVar (dummy_loc,id)) nal in
- let t = RApp (dummy_loc,RRef (dummy_loc,IndRef ind),params@args) in
+ | Anonymous -> GHole (dummy_loc,Evd.InternalHole)
+ | Name id -> GVar (dummy_loc,id)) nal in
+ let t = GApp (dummy_loc,GRef (dummy_loc,IndRef ind),params@args) in
(extern_typ scopes vars t)) x))) tml in
let eqns = List.map (extern_eqn inctx scopes vars) eqns in
CCases (loc,sty,rtntypopt',tml,eqns)
- | RLetTuple (loc,nal,(na,typopt),tm,b) ->
+ | GLetTuple (loc,nal,(na,typopt),tm,b) ->
CLetTuple (loc,List.map (fun na -> (dummy_loc,na)) nal,
(Option.map (fun _ -> (dummy_loc,na)) typopt,
Option.map (extern_typ scopes (add_vname vars na)) typopt),
sub_extern false scopes vars tm,
extern inctx scopes (List.fold_left add_vname vars nal) b)
- | RIf (loc,c,(na,typopt),b1,b2) ->
+ | GIf (loc,c,(na,typopt),b1,b2) ->
CIf (loc,sub_extern false scopes vars c,
(Option.map (fun _ -> (dummy_loc,na)) typopt,
Option.map (extern_typ scopes (add_vname vars na)) typopt),
sub_extern inctx scopes vars b1, sub_extern inctx scopes vars b2)
- | RRec (loc,fk,idv,blv,tyv,bv) ->
+ | GRec (loc,fk,idv,blv,tyv,bv) ->
let vars' = Array.fold_right Idset.add idv vars in
(match fk with
| RFix (nv,n) ->
@@ -674,16 +674,16 @@ let rec extern inctx scopes vars r =
in
CCoFix (loc,(loc,idv.(n)),Array.to_list listdecl))
- | RSort (loc,s) -> CSort (loc,extern_rawsort s)
+ | GSort (loc,s) -> CSort (loc,extern_rawsort s)
- | RHole (loc,e) -> CHole (loc, Some e)
+ | GHole (loc,e) -> CHole (loc, Some e)
- | RCast (loc,c, CastConv (k,t)) ->
+ | GCast (loc,c, CastConv (k,t)) ->
CCast (loc,sub_extern true scopes vars c, CastConv (k,extern_typ scopes vars t))
- | RCast (loc,c, CastCoerce) ->
+ | GCast (loc,c, CastCoerce) ->
CCast (loc,sub_extern true scopes vars c, CastCoerce)
- | RDynamic (loc,d) -> CDynamic (loc,d)
+ | GDynamic (loc,d) -> CDynamic (loc,d)
and extern_typ (_,scopes) =
extern true (Some Notation.type_scope,scopes)
@@ -695,7 +695,7 @@ and factorize_prod scopes vars aty c =
if !Flags.raw_print or !print_no_symbol then raise No_match;
([],extern_symbol scopes vars c (uninterp_notations c))
with No_match -> match c with
- | RProd (loc,(Name id as na),bk,ty,c)
+ | GProd (loc,(Name id as na),bk,ty,c)
when same aty (extern_typ scopes vars (anonymize_if_reserved na ty))
& not (occur_var_constr_expr id aty) (* avoid na in ty escapes scope *)
-> let (nal,c) = factorize_prod scopes (Idset.add id vars) aty c in
@@ -707,7 +707,7 @@ and factorize_lambda inctx scopes vars aty c =
if !Flags.raw_print or !print_no_symbol then raise No_match;
([],extern_symbol scopes vars c (uninterp_notations c))
with No_match -> match c with
- | RLambda (loc,na,bk,ty,c)
+ | GLambda (loc,na,bk,ty,c)
when same aty (extern_typ scopes vars (anonymize_if_reserved na ty))
& not (occur_name na aty) (* To avoid na in ty' escapes scope *)
-> let (nal,c) =
@@ -743,11 +743,11 @@ and extern_eqn inctx scopes vars (loc,ids,pl,c) =
and extern_symbol (tmp_scope,scopes as allscopes) vars t = function
| [] -> raise No_match
| (keyrule,pat,n as _rule)::rules ->
- let loc = Rawterm.loc_of_rawconstr t in
+ let loc = Rawterm.loc_of_glob_constr t in
try
(* Adjusts to the number of arguments expected by the notation *)
let (t,args,argsscopes,argsimpls) = match t,n with
- | RApp (_,(RRef (_,ref) as f),args), Some n
+ | GApp (_,(GRef (_,ref) as f),args), Some n
when List.length args >= n ->
let args1, args2 = list_chop n args in
let subscopes =
@@ -757,15 +757,15 @@ and extern_symbol (tmp_scope,scopes as allscopes) vars t = function
select_impargs_size
(List.length args) (implicits_of_global ref) in
try list_skipn n impls with _ -> [] in
- (if n = 0 then f else RApp (dummy_loc,f,args1)),
+ (if n = 0 then f else GApp (dummy_loc,f,args1)),
args2, subscopes, impls
- | RApp (_,(RRef (_,ref) as f),args), None ->
+ | GApp (_,(GRef (_,ref) as f),args), None ->
let subscopes = find_arguments_scope ref in
let impls =
select_impargs_size
(List.length args) (implicits_of_global ref) in
f, args, subscopes, impls
- | RRef _, Some 0 -> RApp (dummy_loc,t,[]), [], [], []
+ | GRef _, Some 0 -> GApp (dummy_loc,t,[]), [], [], []
| _, None -> t, [], [], []
| _ -> raise No_match in
(* Try matching ... *)
@@ -815,10 +815,10 @@ and extern_recursion_order scopes vars = function
Option.map (extern true scopes vars) r)
-let extern_rawconstr vars c =
+let extern_glob_constr vars c =
extern false (None,[]) vars c
-let extern_rawtype vars c =
+let extern_glob_type vars c =
extern_typ (None,[]) vars c
(******************************************************************)
@@ -841,7 +841,7 @@ let extern_constr at_top env t =
let extern_type at_top env t =
let avoid = if at_top then ids_of_context env else [] in
let r = Detyping.detype at_top avoid (names_of_rel_context env) t in
- extern_rawtype (vars_of_env env) r
+ extern_glob_type (vars_of_env env) r
let extern_sort s = extern_rawsort (detype_sort s)
@@ -849,37 +849,37 @@ let extern_sort s = extern_rawsort (detype_sort s)
(* Main translation function from pattern -> constr_expr *)
let rec raw_of_pat env = function
- | PRef ref -> RRef (loc,ref)
- | PVar id -> RVar (loc,id)
- | PEvar (n,l) -> REvar (loc,n,Some (array_map_to_list (raw_of_pat env) l))
+ | PRef ref -> GRef (loc,ref)
+ | PVar id -> GVar (loc,id)
+ | PEvar (n,l) -> GEvar (loc,n,Some (array_map_to_list (raw_of_pat env) l))
| PRel n ->
let id = try match lookup_name_of_rel n env with
| Name id -> id
| Anonymous ->
- anomaly "rawconstr_of_pattern: index to an anonymous variable"
+ anomaly "glob_constr_of_pattern: index to an anonymous variable"
with Not_found -> id_of_string ("_UNBOUND_REL_"^(string_of_int n)) in
- RVar (loc,id)
- | PMeta None -> RHole (loc,Evd.InternalHole)
- | PMeta (Some n) -> RPatVar (loc,(false,n))
+ GVar (loc,id)
+ | PMeta None -> GHole (loc,Evd.InternalHole)
+ | PMeta (Some n) -> GPatVar (loc,(false,n))
| PApp (f,args) ->
- RApp (loc,raw_of_pat env f,array_map_to_list (raw_of_pat env) args)
+ GApp (loc,raw_of_pat env f,array_map_to_list (raw_of_pat env) args)
| PSoApp (n,args) ->
- RApp (loc,RPatVar (loc,(true,n)),
+ GApp (loc,GPatVar (loc,(true,n)),
List.map (raw_of_pat env) args)
| PProd (na,t,c) ->
- RProd (loc,na,Explicit,raw_of_pat env t,raw_of_pat (na::env) c)
+ GProd (loc,na,Explicit,raw_of_pat env t,raw_of_pat (na::env) c)
| PLetIn (na,t,c) ->
- RLetIn (loc,na,raw_of_pat env t, raw_of_pat (na::env) c)
+ GLetIn (loc,na,raw_of_pat env t, raw_of_pat (na::env) c)
| PLambda (na,t,c) ->
- RLambda (loc,na,Explicit,raw_of_pat env t, raw_of_pat (na::env) c)
+ GLambda (loc,na,Explicit,raw_of_pat env t, raw_of_pat (na::env) c)
| PIf (c,b1,b2) ->
- RIf (loc, raw_of_pat env c, (Anonymous,None),
+ GIf (loc, raw_of_pat env c, (Anonymous,None),
raw_of_pat env b1, raw_of_pat env b2)
| PCase ((LetStyle,[|n|],ind,None),PMeta None,tm,[|b|]) ->
let nal,b = it_destRLambda_or_LetIn_names n (raw_of_pat env b) in
- RLetTuple (loc,nal,(Anonymous,None),raw_of_pat env tm,b)
+ GLetTuple (loc,nal,(Anonymous,None),raw_of_pat env tm,b)
| PCase (_,PMeta None,tm,[||]) ->
- RCases (loc,RegularStyle,None,[raw_of_pat env tm,(Anonymous,None)],[])
+ GCases (loc,RegularStyle,None,[raw_of_pat env tm,(Anonymous,None)],[])
| PCase ((_,cstr_nargs,indo,ind_nargs),p,tm,bv) ->
let brs = Array.to_list (Array.map (raw_of_pat env) bv) in
let brns = Array.to_list cstr_nargs in
@@ -891,10 +891,10 @@ let rec raw_of_pat env = function
else
let nparams,n = Option.get ind_nargs in
return_type_of_predicate ind nparams n (raw_of_pat env p) in
- RCases (loc,RegularStyle,rtn,[raw_of_pat env tm,indnames],mat)
+ GCases (loc,RegularStyle,rtn,[raw_of_pat env tm,indnames],mat)
| PFix f -> Detyping.detype false [] env (mkFix f)
| PCoFix c -> Detyping.detype false [] env (mkCoFix c)
- | PSort s -> RSort (loc,s)
+ | PSort s -> GSort (loc,s)
let extern_constr_pattern env pat =
extern true (None,[]) Idset.empty (raw_of_pat env pat)
diff --git a/interp/constrextern.mli b/interp/constrextern.mli
index 08f089d4e..979c974ac 100644
--- a/interp/constrextern.mli
+++ b/interp/constrextern.mli
@@ -25,8 +25,8 @@ val check_same_type : constr_expr -> constr_expr -> unit
trees for printing *)
val extern_cases_pattern : Idset.t -> cases_pattern -> cases_pattern_expr
-val extern_rawconstr : Idset.t -> rawconstr -> constr_expr
-val extern_rawtype : Idset.t -> rawconstr -> constr_expr
+val extern_glob_constr : Idset.t -> glob_constr -> constr_expr
+val extern_glob_type : Idset.t -> glob_constr -> constr_expr
val extern_constr_pattern : names_context -> constr_pattern -> constr_expr
(** If [b=true] in [extern_constr b env c] then the variables in the first
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index fad3c4910..c097ce43d 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -46,7 +46,7 @@ type var_internalization_data =
type internalization_env =
(identifier * var_internalization_data) list
-type raw_binder = (name * binding_kind * rawconstr option * rawconstr)
+type glob_binder = (name * binding_kind * glob_constr option * glob_constr)
let interning_grammar = ref false
@@ -295,12 +295,12 @@ let reset_tmp_scope (ids,unb,tmp_scope,scopes) =
let rec it_mkRProd env body =
match env with
- (na, bk, _, t) :: tl -> it_mkRProd tl (RProd (dummy_loc, na, bk, t, body))
+ (na, bk, _, t) :: tl -> it_mkRProd tl (GProd (dummy_loc, na, bk, t, body))
| [] -> body
let rec it_mkRLambda env body =
match env with
- (na, bk, _, t) :: tl -> it_mkRLambda tl (RLambda (dummy_loc, na, bk, t, body))
+ (na, bk, _, t) :: tl -> it_mkRLambda tl (GLambda (dummy_loc, na, bk, t, body))
| [] -> body
(**********************************************************************)
@@ -313,11 +313,11 @@ let check_capture loc ty = function
()
let locate_if_isevar loc na = function
- | RHole _ ->
+ | GHole _ ->
(try match na with
- | Name id -> rawconstr_of_aconstr loc (Reserve.find_reserved_type id)
+ | Name id -> glob_constr_of_aconstr loc (Reserve.find_reserved_type id)
| Anonymous -> raise Not_found
- with Not_found -> RHole (loc, Evd.BinderType na))
+ with Not_found -> GHole (loc, Evd.BinderType na))
| x -> x
let check_hidden_implicit_parameters id (_,_,_,impls) =
@@ -350,9 +350,9 @@ let intern_generalized_binder ?(global_level=false) intern_type lvar
Implicit_quantifiers.combine_params_freevar ty
in
let ty' = intern_type (ids,true,tmpsc,sc) ty in
- let fvs = Implicit_quantifiers.generalizable_vars_of_rawconstr ~bound:ids ~allowed:ids' ty' in
+ let fvs = Implicit_quantifiers.generalizable_vars_of_glob_constr ~bound:ids ~allowed:ids' ty' in
let env' = List.fold_left (fun env (x, l) -> push_name_env ~global_level lvar env (l, Name x)) env fvs in
- let bl = List.map (fun (id, loc) -> (Name id, b, None, RHole (loc, Evd.BinderType (Name id)))) fvs in
+ let bl = List.map (fun (id, loc) -> (Name id, b, None, GHole (loc, Evd.BinderType (Name id)))) fvs in
let na = match na with
| Anonymous ->
if global_level then na
@@ -383,11 +383,11 @@ let intern_local_binder_aux ?(global_level=false) intern intern_type lvar (env,b
env, b @ bl)
| LocalRawDef((loc,na as locna),def) ->
(push_name_env lvar env locna,
- (na,Explicit,Some(intern env def),RHole(loc,Evd.BinderType na))::bl)
+ (na,Explicit,Some(intern env def),GHole(loc,Evd.BinderType na))::bl)
let intern_generalization intern (ids,unb,tmp_scope,scopes as env) lvar loc bk ak c =
let c = intern (ids,true,tmp_scope,scopes) c in
- let fvs = Implicit_quantifiers.generalizable_vars_of_rawconstr ~bound:ids c in
+ let fvs = Implicit_quantifiers.generalizable_vars_of_glob_constr ~bound:ids c in
let env', c' =
let abs =
let pi =
@@ -399,10 +399,10 @@ let intern_generalization intern (ids,unb,tmp_scope,scopes as env) lvar loc bk a
in
if pi then
(fun (id, loc') acc ->
- RProd (join_loc loc' loc, Name id, bk, RHole (loc', Evd.BinderType (Name id)), acc))
+ GProd (join_loc loc' loc, Name id, bk, GHole (loc', Evd.BinderType (Name id)), acc))
else
(fun (id, loc') acc ->
- RLambda (join_loc loc' loc, Name id, bk, RHole (loc', Evd.BinderType (Name id)), acc))
+ GLambda (join_loc loc' loc, Name id, bk, GHole (loc', Evd.BinderType (Name id)), acc))
in
List.fold_right (fun (id, loc as lid) (env, acc) ->
let env' = push_name_env lvar env (loc, Name id) in
@@ -426,7 +426,7 @@ let iterate_binder intern lvar (env,bl) = function
env, b @ bl)
| LocalRawDef((loc,na as locna),def) ->
(push_name_env lvar env locna,
- (na,Explicit,Some(intern env def),RHole(loc,Evd.BinderType na))::bl)
+ (na,Explicit,Some(intern env def),GHole(loc,Evd.BinderType na))::bl)
(**********************************************************************)
(* Syntax extensions *)
@@ -460,10 +460,10 @@ let traverse_binder (terms,_,_ as subst)
(renaming',env), Name id'
let rec subst_iterator y t = function
- | RVar (_,id) as x -> if id = y then t else x
- | x -> map_rawconstr (subst_iterator y t) x
+ | GVar (_,id) as x -> if id = y then t else x
+ | x -> map_glob_constr (subst_iterator y t) x
-let subst_aconstr_in_rawconstr loc intern lvar subst infos c =
+let subst_aconstr_in_glob_constr loc intern lvar subst infos c =
let (terms,termlists,binders) = subst in
let rec aux (terms,binderopt as subst') (renaming,(ids,unb,_,scopes as env)) c =
let subinfos = renaming,(ids,unb,None,scopes) in
@@ -477,10 +477,10 @@ let subst_aconstr_in_rawconstr loc intern lvar subst infos c =
intern (ids,unb,scopt,subscopes@scopes) a
with Not_found ->
try
- RVar (loc,List.assoc id renaming)
+ GVar (loc,List.assoc id renaming)
with Not_found ->
(* Happens for local notation joint with inductive/fixpoint defs *)
- RVar (loc,id)
+ GVar (loc,id)
end
| AList (x,_,iter,terminator,lassoc) ->
(try
@@ -497,7 +497,7 @@ let subst_aconstr_in_rawconstr loc intern lvar subst infos c =
let na =
try snd (coerce_to_name (fst (List.assoc id terms)))
with Not_found -> na in
- RHole (loc,Evd.BinderType na)
+ GHole (loc,Evd.BinderType na)
| ABinderList (x,_,iter,terminator) ->
(try
(* All elements of the list are in scopes (scopt,subscopes) *)
@@ -512,12 +512,12 @@ let subst_aconstr_in_rawconstr loc intern lvar subst infos c =
anomaly "Inconsistent substitution of recursive notation")
| AProd (Name id, AHole _, c') when option_mem_assoc id binderopt ->
let (na,bk,_,t) = snd (Option.get binderopt) in
- RProd (loc,na,bk,t,aux subst' infos c')
+ GProd (loc,na,bk,t,aux subst' infos c')
| ALambda (Name id,AHole _,c') when option_mem_assoc id binderopt ->
let (na,bk,_,t) = snd (Option.get binderopt) in
- RLambda (loc,na,bk,t,aux subst' infos c')
+ GLambda (loc,na,bk,t,aux subst' infos c')
| t ->
- rawconstr_of_aconstr_with_binders loc (traverse_binder subst)
+ glob_constr_of_aconstr_with_binders loc (traverse_binder subst)
(aux subst') subinfos t
in aux (terms,None) infos c
@@ -538,7 +538,7 @@ let intern_notation intern (_,_,tmp_scope,scopes as env) lvar loc ntn fullargs =
let terms = make_subst ids args in
let termlists = make_subst idsl argslist in
let binders = make_subst idsbl bll in
- subst_aconstr_in_rawconstr loc intern lvar
+ subst_aconstr_in_glob_constr loc intern lvar
(terms,termlists,binders) ([],env) c
(**********************************************************************)
@@ -558,20 +558,20 @@ let intern_var (ids,_,_,_ as genv) (ltacvars,namedctxvars,ntnvars,impls) loc id
(fun id -> CRef (Ident (loc,id)), Some (loc,ExplByName id)) expl_impls in
let tys = string_of_ty ty in
Dumpglob.dump_reference loc "<>" (string_of_id id) tys;
- RVar (loc,id), make_implicits_list impls, argsc, expl_impls
+ GVar (loc,id), make_implicits_list impls, argsc, expl_impls
with Not_found ->
(* Is [id] bound in current term or is an ltac var bound to constr *)
if Idset.mem id ids or List.mem id ltacvars
then
- RVar (loc,id), [], [], []
+ GVar (loc,id), [], [], []
(* Is [id] a notation variable *)
else if List.mem_assoc id ntnvars
then
- (set_var_scope loc id true genv ntnvars; RVar (loc,id), [], [], [])
+ (set_var_scope loc id true genv ntnvars; GVar (loc,id), [], [], [])
(* Is [id] the special variable for recursive notations *)
else if ntnvars <> [] && id = ldots_var
then
- RVar (loc,id), [], [], []
+ GVar (loc,id), [], [], []
else
(* Is [id] bound to a free name in ltac (this is an ltac error message) *)
try
@@ -589,14 +589,14 @@ let intern_var (ids,_,_,_ as genv) (ltacvars,namedctxvars,ntnvars,impls) loc id
let impls = implicits_of_global ref in
let scopes = find_arguments_scope ref in
Dumpglob.dump_reference loc "<>" (string_of_qualid (Decls.variable_secpath id)) "var";
- RRef (loc, ref), impls, scopes, []
+ GRef (loc, ref), impls, scopes, []
with _ ->
(* [id] a goal variable *)
- RVar (loc,id), [], [], []
+ GVar (loc,id), [], [], []
let find_appl_head_data = function
- | RRef (_,ref) as x -> x,implicits_of_global ref,find_arguments_scope ref,[]
- | RApp (_,RRef (_,ref),l) as x
+ | GRef (_,ref) as x -> x,implicits_of_global ref,find_arguments_scope ref,[]
+ | GApp (_,GRef (_,ref),l) as x
when l <> [] & Flags.version_strictly_greater Flags.V8_2 ->
let n = List.length l in
x,List.map (drop_first_implicits n) (implicits_of_global ref),
@@ -629,7 +629,7 @@ let intern_reference ref =
let intern_qualid loc qid intern env lvar args =
match intern_extended_global_of_qualid (loc,qid) with
| TrueGlobal ref ->
- RRef (loc, ref), args
+ GRef (loc, ref), args
| SynDef sp ->
let (ids,c) = Syntax_def.search_syntactic_definition sp in
let nids = List.length ids in
@@ -637,12 +637,12 @@ let intern_qualid loc qid intern env lvar args =
let args1,args2 = list_chop nids args in
check_no_explicitation args1;
let subst = make_subst ids (List.map fst args1) in
- subst_aconstr_in_rawconstr loc intern lvar (subst,[],[]) ([],env) c, args2
+ subst_aconstr_in_glob_constr loc intern lvar (subst,[],[]) ([],env) c, args2
(* Rule out section vars since these should have been found by intern_var *)
let intern_non_secvar_qualid loc qid intern env lvar args =
match intern_qualid loc qid intern env lvar args with
- | RRef (loc, VarRef id),_ -> error_global_not_found_loc loc qid
+ | GRef (loc, VarRef id),_ -> error_global_not_found_loc loc qid
| r -> r
let intern_applied_reference intern (_, unb, _, _ as env) lvar args = function
@@ -659,7 +659,7 @@ let intern_applied_reference intern (_, unb, _, _ as env) lvar args = function
with e ->
(* Extra allowance for non globalizing functions *)
if !interning_grammar || unb then
- (RVar (loc,id), [], [], []),args
+ (GVar (loc,id), [], [], []),args
else raise e
let interp_reference vars r =
@@ -1046,7 +1046,7 @@ let merge_impargs l args =
let check_projection isproj nargs r =
match (r,isproj) with
- | RRef (loc, ref), Some _ ->
+ | GRef (loc, ref), Some _ ->
(try
let n = Recordops.find_projection_nparams ref + 1 in
if nargs <> n then
@@ -1054,15 +1054,15 @@ let check_projection isproj nargs r =
with Not_found ->
user_err_loc
(loc,"",pr_global_env Idset.empty ref ++ str " is not a registered projection."))
- | _, Some _ -> user_err_loc (loc_of_rawconstr r, "", str "Not a projection.")
+ | _, Some _ -> user_err_loc (loc_of_glob_constr r, "", str "Not a projection.")
| _, None -> ()
let get_implicit_name n imps =
Some (Impargs.name_of_implicit (List.nth imps (n-1)))
let set_hole_implicit i b = function
- | RRef (loc,r) | RApp (_,RRef (loc,r),_) -> (loc,Evd.ImplicitArg (r,i,b))
- | RVar (loc,id) -> (loc,Evd.ImplicitArg (VarRef id,i,b))
+ | GRef (loc,r) | GApp (_,GRef (loc,r),_) -> (loc,Evd.ImplicitArg (r,i,b))
+ | GVar (loc,id) -> (loc,Evd.ImplicitArg (VarRef id,i,b))
| _ -> anomaly "Only refs have implicits"
let exists_implicit_name id =
@@ -1112,7 +1112,7 @@ let internalize sigma globalenv env allow_patvar lvar c =
intern_applied_reference intern env lvar [] ref in
(match intern_impargs c env imp subscopes l with
| [] -> c
- | l -> RApp (constr_loc x, c, l))
+ | l -> GApp (constr_loc x, c, l))
| CFix (loc, (locid,iddef), dl) ->
let lf = List.map (fun ((_, id),_,_,_,_) -> id) dl in
let dl = Array.of_list dl in
@@ -1144,7 +1144,7 @@ let internalize sigma globalenv env allow_patvar lvar c =
((n, ro), List.rev rbl,
intern_type (ids',unb,tmp_scope,scopes) ty,
intern (ids'',unb,None,scopes) bd)) dl in
- RRec (loc,RFix
+ GRec (loc,RFix
(Array.map (fun (ro,_,_,_) -> ro) idl,n),
Array.of_list lf,
Array.map (fun (_,bl,_,_) -> bl) idl,
@@ -1166,13 +1166,13 @@ let internalize sigma globalenv env allow_patvar lvar c =
(List.rev rbl,
intern_type (ids',unb,tmp_scope,scopes) ty,
intern (ids'',unb,None,scopes) bd)) dl in
- RRec (loc,RCoFix n,
+ GRec (loc,RCoFix n,
Array.of_list lf,
Array.map (fun (bl,_,_) -> bl) idl,
Array.map (fun (_,ty,_) -> ty) idl,
Array.map (fun (_,_,bd) -> bd) idl)
| CArrow (loc,c1,c2) ->
- RProd (loc, Anonymous, Explicit, intern_type env c1, intern_type env c2)
+ GProd (loc, Anonymous, Explicit, intern_type env c1, intern_type env c2)
| CProdN (loc,[],c2) ->
intern_type env c2
| CProdN (loc,(nal,bk,ty)::bll,c2) ->
@@ -1182,7 +1182,7 @@ let internalize sigma globalenv env allow_patvar lvar c =
| CLambdaN (loc,(nal,bk,ty)::bll,c2) ->
iterate_lam loc (reset_tmp_scope env) bk ty (CLambdaN (loc, bll, c2)) nal
| CLetIn (loc,na,c1,c2) ->
- RLetIn (loc, snd na, intern (reset_tmp_scope env) c1,
+ GLetIn (loc, snd na, intern (reset_tmp_scope env) c1,
intern (push_name_env lvar env na) c2)
| CNotation (loc,"- _",([CPrim (_,Numeral p)],[],[]))
when Bigint.is_strictly_pos p ->
@@ -1201,8 +1201,8 @@ let internalize sigma globalenv env allow_patvar lvar c =
let args = List.map (fun a -> (a,None)) args in
intern_applied_reference intern env lvar args ref in
check_projection isproj (List.length args) f;
- (* Rem: RApp(_,f,[]) stands for @f *)
- RApp (loc, f, intern_args env args_scopes (List.map fst args))
+ (* Rem: GApp(_,f,[]) stands for @f *)
+ GApp (loc, f, intern_args env args_scopes (List.map fst args))
| CApp (loc, (isproj,f), args) ->
let isproj,f,args = match f with
(* Compact notations like "t.(f args') args" *)
@@ -1221,8 +1221,8 @@ let internalize sigma globalenv env allow_patvar lvar c =
check_projection isproj (List.length args) c;
(match c with
(* Now compact "(f args') args" *)
- | RApp (loc', f', args') -> RApp (join_loc loc' loc, f',args'@args)
- | _ -> RApp (loc, c, args))
+ | GApp (loc', f', args') -> GApp (join_loc loc' loc, f',args'@args)
+ | _ -> GApp (loc, c, args))
| CRecord (loc, _, fs) ->
let cargs =
sort_fields true loc fs
@@ -1244,14 +1244,14 @@ let internalize sigma globalenv env allow_patvar lvar c =
tms ([],env) in
let rtnpo = Option.map (intern_type env') rtnpo in
let eqns' = List.map (intern_eqn (List.length tms) env) eqns in
- RCases (loc, sty, rtnpo, tms, List.flatten eqns')
+ GCases (loc, sty, rtnpo, tms, List.flatten eqns')
| CLetTuple (loc, nal, (na,po), b, c) ->
let env' = reset_tmp_scope env in
let ((b',(na',_)),ids) = intern_case_item env' (b,(na,None)) in
let p' = Option.map (fun p ->
let env'' = List.fold_left (push_name_env lvar) env ids in
intern_type env'' p) po in
- RLetTuple (loc, List.map snd nal, (na', p'), b',
+ GLetTuple (loc, List.map snd nal, (na', p'), b',
intern (List.fold_left (push_name_env lvar) env nal) c)
| CIf (loc, c, (na,po), b1, b2) ->
let env' = reset_tmp_scope env in
@@ -1259,23 +1259,23 @@ let internalize sigma globalenv env allow_patvar lvar c =
let p' = Option.map (fun p ->
let env'' = List.fold_left (push_name_env lvar) env ids in
intern_type env'' p) po in
- RIf (loc, c', (na', p'), intern env b1, intern env b2)
+ GIf (loc, c', (na', p'), intern env b1, intern env b2)
| CHole (loc, k) ->
- RHole (loc, match k with Some k -> k | None -> Evd.QuestionMark (Evd.Define true))
+ GHole (loc, match k with Some k -> k | None -> Evd.QuestionMark (Evd.Define true))
| CPatVar (loc, n) when allow_patvar ->
- RPatVar (loc, n)
+ GPatVar (loc, n)
| CPatVar (loc, _) ->
raise (InternalizationError (loc,IllegalMetavariable))
| CEvar (loc, n, l) ->
- REvar (loc, n, Option.map (List.map (intern env)) l)
+ GEvar (loc, n, Option.map (List.map (intern env)) l)
| CSort (loc, s) ->
- RSort(loc,s)
+ GSort(loc,s)
| CCast (loc, c1, CastConv (k, c2)) ->
- RCast (loc,intern env c1, CastConv (k, intern_type env c2))
+ GCast (loc,intern env c1, CastConv (k, intern_type env c2))
| CCast (loc, c1, CastCoerce) ->
- RCast (loc,intern env c1, CastCoerce)
+ GCast (loc,intern env c1, CastCoerce)
- | CDynamic (loc,d) -> RDynamic (loc,d)
+ | CDynamic (loc,d) -> GDynamic (loc,d)
and intern_type env = intern (set_type_scope env)
@@ -1318,17 +1318,17 @@ let internalize sigma globalenv env allow_patvar lvar c =
let tids = List.fold_right Idset.add tids Idset.empty in
let t = intern_type (tids,unb,None,scopes) t in
let loc,ind,l = match t with
- | RRef (loc,IndRef ind) -> (loc,ind,[])
- | RApp (loc,RRef (_,IndRef ind),l) -> (loc,ind,l)
- | _ -> error_bad_inductive_type (loc_of_rawconstr t) in
+ | GRef (loc,IndRef ind) -> (loc,ind,[])
+ | GApp (loc,GRef (_,IndRef ind),l) -> (loc,ind,l)
+ | _ -> error_bad_inductive_type (loc_of_glob_constr t) in
let nparams, nrealargs = inductive_nargs globalenv ind in
let nindargs = nparams + nrealargs in
if List.length l <> nindargs then
error_wrong_numarg_inductive_loc loc globalenv ind nindargs;
let nal = List.map (function
- | RHole (loc,_) -> loc,Anonymous
- | RVar (loc,id) -> loc,Name id
- | c -> user_err_loc (loc_of_rawconstr c,"",str "Not a name.")) l in
+ | GHole (loc,_) -> loc,Anonymous
+ | GVar (loc,id) -> loc,Name id
+ | c -> user_err_loc (loc_of_glob_constr c,"",str "Not a name.")) l in
let parnal,realnal = list_chop nparams nal in
if List.exists (fun (_,na) -> na <> Anonymous) parnal then
error_inductive_parameter_not_implicit loc;
@@ -1336,8 +1336,8 @@ let internalize sigma globalenv env allow_patvar lvar c =
| None ->
[], None in
let na = match tm', na with
- | RVar (loc,id), None when Idset.mem id vars -> loc,Name id
- | RRef (loc, VarRef id), None -> loc,Name id
+ | GVar (loc,id), None when Idset.mem id vars -> loc,Name id
+ | GRef (loc, VarRef id), None -> loc,Name id
| _, None -> dummy_loc,Anonymous
| _, Some (loc,na) -> loc,na in
(tm',(snd na,typ)), na::ids
@@ -1348,7 +1348,7 @@ let internalize sigma globalenv env allow_patvar lvar c =
if nal <> [] then check_capture loc1 ty na;
let body = default (push_name_env lvar env locna) bk nal in
let ty = locate_if_isevar loc1 na (intern_type env ty) in
- RProd (join_loc loc1 loc2, na, bk, ty, body)
+ GProd (join_loc loc1 loc2, na, bk, ty, body)
| [] -> intern_type env body
in
match bk with
@@ -1364,7 +1364,7 @@ let internalize sigma globalenv env allow_patvar lvar c =
if nal <> [] then check_capture loc1 ty na;
let body = default (push_name_env lvar env locna) bk nal in
let ty = locate_if_isevar loc1 na (intern_type env ty) in
- RLambda (join_loc loc1 loc2, na, bk, ty, body)
+ GLambda (join_loc loc1 loc2, na, bk, ty, body)
| [] -> intern env body
in match bk with
| Default b -> default env b nal
@@ -1391,7 +1391,7 @@ let internalize sigma globalenv env allow_patvar lvar c =
(* with implicit arguments if maximal insertion is set *)
[]
else
- RHole (set_hole_implicit (n,get_implicit_name n l) (force_inference_of imp) c) ::
+ GHole (set_hole_implicit (n,get_implicit_name n l) (force_inference_of imp) c) ::
aux (n+1) impl' subscopes' eargs rargs
end
| (imp::impl', a::rargs') ->
@@ -1423,7 +1423,7 @@ let internalize sigma globalenv env allow_patvar lvar c =
explain_internalization_error e)
(**************************************************************************)
-(* Functions to translate constr_expr into rawconstr *)
+(* Functions to translate constr_expr into glob_constr *)
(**************************************************************************)
let extract_ids env =
@@ -1477,18 +1477,18 @@ let interp_open_constr sigma env c =
let interp_open_constr_patvar sigma env c =
let raw = intern_gen false sigma env c ~allow_patvar:true in
let sigma = ref (Evd.create_evar_defs sigma) in
- let evars = ref (Gmap.empty : (identifier,rawconstr) Gmap.t) in
+ let evars = ref (Gmap.empty : (identifier,glob_constr) Gmap.t) in
let rec patvar_to_evar r = match r with
- | RPatVar (loc,(_,id)) ->
+ | GPatVar (loc,(_,id)) ->
( try Gmap.find id !evars
with Not_found ->
let ev = Evarutil.e_new_evar sigma env (Termops.new_Type()) in
let ev = Evarutil.e_new_evar sigma env ev in
- let rev = REvar (loc,(fst (Term.destEvar ev)),None) (*TODO*) in
+ let rev = GEvar (loc,(fst (Term.destEvar ev)),None) (*TODO*) in
evars := Gmap.add id rev !evars;
rev
)
- | _ -> map_rawconstr patvar_to_evar r in
+ | _ -> map_glob_constr patvar_to_evar r in
let raw = patvar_to_evar raw in
Default.understand_tcc !sigma env raw
@@ -1531,7 +1531,7 @@ type ltac_sign = identifier list * unbound_ltac_var_map
let intern_constr_pattern sigma env ?(as_type=false) ?(ltacvars=([],[])) c =
let c = intern_gen as_type ~allow_patvar:true ~ltacvars sigma env c in
- pattern_of_rawconstr c
+ pattern_of_glob_constr c
let interp_aconstr ?(impls=[]) vars recvars a =
let env = Global.env () in
@@ -1540,7 +1540,7 @@ let interp_aconstr ?(impls=[]) vars recvars a =
let c = internalize Evd.empty (Global.env()) (extract_ids env, false, None, [])
false (([],[]),Environ.named_context env,vl,impls) a in
(* Translate and check that [c] has all its free variables bound in [vars] *)
- let a = aconstr_of_rawconstr vars recvars c in
+ let a = aconstr_of_glob_constr vars recvars c in
(* Splits variables into those that are binding, bound, or both *)
(* binding and bound *)
let out_scope = function None -> None,[] | Some (a,l) -> a,l in
@@ -1552,12 +1552,12 @@ let interp_aconstr ?(impls=[]) vars recvars a =
let interp_binder sigma env na t =
let t = intern_gen true sigma env t in
- let t' = locate_if_isevar (loc_of_rawconstr t) na t in
+ let t' = locate_if_isevar (loc_of_glob_constr t) na t in
Default.understand_type sigma env t'
let interp_binder_evars evdref env na t =
let t = intern_gen true !evdref env t in
- let t' = locate_if_isevar (loc_of_rawconstr t) na t in
+ let t' = locate_if_isevar (loc_of_glob_constr t) na t in
Default.understand_tcc_evars evdref env IsType t'
open Environ
@@ -1580,7 +1580,7 @@ let interp_rawcontext_gen understand_type understand_judgment env bl =
(fun (env,params,n,impls) (na, k, b, t) ->
match b with
None ->
- let t' = locate_if_isevar (loc_of_rawconstr t) na t in
+ let t' = locate_if_isevar (loc_of_glob_constr t) na t in
let t = understand_type env t' in
let d = (na,None,t) in
let impls =
diff --git a/interp/constrintern.mli b/interp/constrintern.mli
index 6e977056c..cf9e899a6 100644
--- a/interp/constrintern.mli
+++ b/interp/constrintern.mli
@@ -18,7 +18,7 @@ open Topconstr
open Termops
open Pretyping
-(** Translation from front abstract syntax of term to untyped terms (rawconstr) *)
+(** Translation from front abstract syntax of term to untyped terms (glob_constr) *)
(** The translation performs:
@@ -68,23 +68,23 @@ type manual_implicits = (explicitation * (bool * bool * bool)) list
type ltac_sign = identifier list * unbound_ltac_var_map
-type raw_binder = (name * binding_kind * rawconstr option * rawconstr)
+type glob_binder = (name * binding_kind * glob_constr option * glob_constr)
(** {6 Internalization performs interpretation of global names and notations } *)
-val intern_constr : evar_map -> env -> constr_expr -> rawconstr
+val intern_constr : evar_map -> env -> constr_expr -> glob_constr
-val intern_type : evar_map -> env -> constr_expr -> rawconstr
+val intern_type : evar_map -> env -> constr_expr -> glob_constr
val intern_gen : bool -> evar_map -> env ->
?impls:internalization_env -> ?allow_patvar:bool -> ?ltacvars:ltac_sign ->
- constr_expr -> rawconstr
+ constr_expr -> glob_constr
val intern_pattern : env -> cases_pattern_expr ->
Names.identifier list *
((Names.identifier * Names.identifier) list * Rawterm.cases_pattern) list
-val intern_context : bool -> evar_map -> env -> local_binder list -> raw_binder list
+val intern_context : bool -> evar_map -> env -> local_binder list -> glob_binder list
(** {6 Composing internalization with pretyping } *)
@@ -142,7 +142,7 @@ val intern_constr_pattern :
val intern_reference : reference -> global_reference
(** Expands abbreviations (syndef); raise an error if not existing *)
-val interp_reference : ltac_sign -> reference -> rawconstr
+val interp_reference : ltac_sign -> reference -> glob_constr
(** Interpret binders *)
@@ -152,8 +152,8 @@ val interp_binder_evars : evar_map ref -> env -> name -> constr_expr -> types
(** Interpret contexts: returns extended env and context *)
-val interp_context_gen : (env -> rawconstr -> types) ->
- (env -> rawconstr -> unsafe_judgment) ->
+val interp_context_gen : (env -> glob_constr -> types) ->
+ (env -> glob_constr -> unsafe_judgment) ->
?global_level:bool ->
evar_map -> env -> local_binder list -> (env * rel_context) * manual_implicits
diff --git a/interp/doc.tex b/interp/doc.tex
index ddf40d6c8..4ce5811da 100644
--- a/interp/doc.tex
+++ b/interp/doc.tex
@@ -5,7 +5,7 @@
\ocwsection \label{interp}
This chapter describes the translation from \Coq\ context-dependent
front abstract syntax of terms (\verb=front=) to and from the
-context-free, untyped, globalized form of constructions (\verb=rawconstr=).
+context-free, untyped, globalized form of constructions (\verb=glob_constr=).
The modules translating back and forth the front abstract syntax are
organized as follows.
diff --git a/interp/genarg.ml b/interp/genarg.ml
index 5b221d4c0..23e17a2d4 100644
--- a/interp/genarg.ml
+++ b/interp/genarg.ml
@@ -51,11 +51,11 @@ let loc_of_or_by_notation f = function
| AN c -> f c
| ByNotation (loc,s,_) -> loc
-type rawconstr_and_expr = rawconstr * constr_expr option
+type glob_constr_and_expr = glob_constr * constr_expr option
type open_constr_expr = unit * constr_expr
-type open_rawconstr = unit * rawconstr_and_expr
+type open_glob_constr = unit * glob_constr_and_expr
-type rawconstr_pattern_and_expr = rawconstr_and_expr * Pattern.constr_pattern
+type glob_constr_pattern_and_expr = glob_constr_and_expr * Pattern.constr_pattern
type 'a with_ebindings = 'a * open_constr bindings
diff --git a/interp/genarg.mli b/interp/genarg.mli
index 963c2742e..231126d44 100644
--- a/interp/genarg.mli
+++ b/interp/genarg.mli
@@ -27,12 +27,12 @@ val loc_of_or_by_notation : ('a -> loc) -> 'a or_by_notation -> loc
(** In globalize tactics, we need to keep the initial [constr_expr] to recompute
in the environment by the effective calls to Intro, Inversion, etc
The [constr_expr] field is [None] in TacDef though *)
-type rawconstr_and_expr = rawconstr * constr_expr option
+type glob_constr_and_expr = glob_constr * constr_expr option
type open_constr_expr = unit * constr_expr
-type open_rawconstr = unit * rawconstr_and_expr
+type open_glob_constr = unit * glob_constr_and_expr
-type rawconstr_pattern_and_expr = rawconstr_and_expr * constr_pattern
+type glob_constr_pattern_and_expr = glob_constr_and_expr * constr_pattern
type 'a with_ebindings = 'a * open_constr bindings
@@ -53,11 +53,11 @@ val pr_or_and_intro_pattern : or_and_intro_pattern_expr -> Pp.std_ppcmds
{% \begin{%}verbatim{% }%}
parsing in_raw out_raw
- char stream ----> rawtype ----> constr_expr generic_argument --------|
+ char stream ----> glob_type ----> constr_expr generic_argument --------|
encapsulation decaps |
|
V
- rawtype
+ glob_type
|
globalization |
V
@@ -66,10 +66,10 @@ val pr_or_and_intro_pattern : or_and_intro_pattern_expr -> Pp.std_ppcmds
encaps |
in_glob |
V
- rawconstr generic_argument
+ glob_constr generic_argument
|
out in out_glob |
- type <--- constr generic_argument <---- type <------ rawtype <--------|
+ type <--- constr generic_argument <---- type <------ glob_type <--------|
| decaps encaps interp decaps
|
V
@@ -78,7 +78,7 @@ effective use
To distinguish between the uninterpreted (raw), globalized and
interpreted worlds, we annotate the type [generic_argument] by a
-phantom argument which is either [constr_expr], [rawconstr] or
+phantom argument which is either [constr_expr], [glob_constr] or
[constr].
Transformation for each type :
@@ -175,35 +175,35 @@ val globwit_sort : (rawsort,glevel) abstract_argument_type
val wit_sort : (sorts,tlevel) abstract_argument_type
val rawwit_constr : (constr_expr,rlevel) abstract_argument_type
-val globwit_constr : (rawconstr_and_expr,glevel) abstract_argument_type
+val globwit_constr : (glob_constr_and_expr,glevel) abstract_argument_type
val wit_constr : (constr,tlevel) abstract_argument_type
val rawwit_constr_may_eval : ((constr_expr,reference or_by_notation,constr_expr) may_eval,rlevel) abstract_argument_type
-val globwit_constr_may_eval : ((rawconstr_and_expr,evaluable_global_reference and_short_name or_var,rawconstr_pattern_and_expr) may_eval,glevel) abstract_argument_type
+val globwit_constr_may_eval : ((glob_constr_and_expr,evaluable_global_reference and_short_name or_var,glob_constr_pattern_and_expr) may_eval,glevel) abstract_argument_type
val wit_constr_may_eval : (constr,tlevel) abstract_argument_type
val rawwit_open_constr_gen : bool -> (open_constr_expr,rlevel) abstract_argument_type
-val globwit_open_constr_gen : bool -> (open_rawconstr,glevel) abstract_argument_type
+val globwit_open_constr_gen : bool -> (open_glob_constr,glevel) abstract_argument_type
val wit_open_constr_gen : bool -> (open_constr,tlevel) abstract_argument_type
val rawwit_open_constr : (open_constr_expr,rlevel) abstract_argument_type
-val globwit_open_constr : (open_rawconstr,glevel) abstract_argument_type
+val globwit_open_constr : (open_glob_constr,glevel) abstract_argument_type
val wit_open_constr : (open_constr,tlevel) abstract_argument_type
val rawwit_casted_open_constr : (open_constr_expr,rlevel) abstract_argument_type
-val globwit_casted_open_constr : (open_rawconstr,glevel) abstract_argument_type
+val globwit_casted_open_constr : (open_glob_constr,glevel) abstract_argument_type
val wit_casted_open_constr : (open_constr,tlevel) abstract_argument_type
val rawwit_constr_with_bindings : (constr_expr with_bindings,rlevel) abstract_argument_type
-val globwit_constr_with_bindings : (rawconstr_and_expr with_bindings,glevel) abstract_argument_type
+val globwit_constr_with_bindings : (glob_constr_and_expr with_bindings,glevel) abstract_argument_type
val wit_constr_with_bindings : (constr with_bindings sigma,tlevel) abstract_argument_type
val rawwit_bindings : (constr_expr bindings,rlevel) abstract_argument_type
-val globwit_bindings : (rawconstr_and_expr bindings,glevel) abstract_argument_type
+val globwit_bindings : (glob_constr_and_expr bindings,glevel) abstract_argument_type
val wit_bindings : (constr bindings sigma,tlevel) abstract_argument_type
val rawwit_red_expr : ((constr_expr,reference or_by_notation,constr_expr) red_expr_gen,rlevel) abstract_argument_type
-val globwit_red_expr : ((rawconstr_and_expr,evaluable_global_reference and_short_name or_var,rawconstr_pattern_and_expr) red_expr_gen,glevel) abstract_argument_type
+val globwit_red_expr : ((glob_constr_and_expr,evaluable_global_reference and_short_name or_var,glob_constr_pattern_and_expr) red_expr_gen,glevel) abstract_argument_type
val wit_red_expr : ((constr,evaluable_global_reference,constr_pattern) red_expr_gen,tlevel) abstract_argument_type
val wit_list0 :
diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml
index 1e97c5178..864e521bf 100644
--- a/interp/implicit_quantifiers.ml
+++ b/interp/implicit_quantifiers.ml
@@ -136,33 +136,33 @@ let add_name_to_ids set na =
| Anonymous -> set
| Name id -> Idset.add id set
-let generalizable_vars_of_rawconstr ?(bound=Idset.empty) ?(allowed=Idset.empty) =
+let generalizable_vars_of_glob_constr ?(bound=Idset.empty) ?(allowed=Idset.empty) =
let rec vars bound vs = function
- | RVar (loc,id) ->
+ | GVar (loc,id) ->
if is_freevar bound (Global.env ()) id then
if List.mem_assoc id vs then vs
else (id, loc) :: vs
else vs
- | RApp (loc,f,args) -> List.fold_left (vars bound) vs (f::args)
- | RLambda (loc,na,_,ty,c) | RProd (loc,na,_,ty,c) | RLetIn (loc,na,ty,c) ->
+ | GApp (loc,f,args) -> List.fold_left (vars bound) vs (f::args)
+ | GLambda (loc,na,_,ty,c) | GProd (loc,na,_,ty,c) | GLetIn (loc,na,ty,c) ->
let vs' = vars bound vs ty in
let bound' = add_name_to_ids bound na in
vars bound' vs' c
- | RCases (loc,sty,rtntypopt,tml,pl) ->
+ | GCases (loc,sty,rtntypopt,tml,pl) ->
let vs1 = vars_option bound vs rtntypopt in
let vs2 = List.fold_left (fun vs (tm,_) -> vars bound vs tm) vs1 tml in
List.fold_left (vars_pattern bound) vs2 pl
- | RLetTuple (loc,nal,rtntyp,b,c) ->
+ | GLetTuple (loc,nal,rtntyp,b,c) ->
let vs1 = vars_return_type bound vs rtntyp in
let vs2 = vars bound vs1 b in
let bound' = List.fold_left add_name_to_ids bound nal in
vars bound' vs2 c
- | RIf (loc,c,rtntyp,b1,b2) ->
+ | GIf (loc,c,rtntyp,b1,b2) ->
let vs1 = vars_return_type bound vs rtntyp in
let vs2 = vars bound vs1 c in
let vs3 = vars bound vs2 b1 in
vars bound vs3 b2
- | RRec (loc,fk,idl,bl,tyl,bv) ->
+ | GRec (loc,fk,idl,bl,tyl,bv) ->
let bound' = Array.fold_right Idset.add idl bound in
let vars_fix i vs fid =
let vs1,bound1 =
@@ -180,9 +180,9 @@ let generalizable_vars_of_rawconstr ?(bound=Idset.empty) ?(allowed=Idset.empty)
vars bound1 vs2 bv.(i)
in
array_fold_left_i vars_fix vs idl
- | RCast (loc,c,k) -> let v = vars bound vs c in
+ | GCast (loc,c,k) -> let v = vars bound vs c in
(match k with CastConv (_,t) -> vars bound v t | _ -> v)
- | (RSort _ | RHole _ | RRef _ | REvar _ | RPatVar _ | RDynamic _) -> vs
+ | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GDynamic _) -> vs
and vars_pattern bound vs (loc,idl,p,c) =
let bound' = List.fold_right Idset.add idl bound in
@@ -307,14 +307,14 @@ let implicits_of_rawterm ?(with_products=true) l =
else rest
in
match c with
- | RProd (loc, na, bk, t, b) ->
+ | GProd (loc, na, bk, t, b) ->
if with_products then abs loc na bk t b
else
(if bk = Implicit then
msg_warning (str "Ignoring implicit status of product binder " ++
pr_name na ++ str " and following binders");
[])
- | RLambda (loc, na, bk, t, b) -> abs loc na bk t b
- | RLetIn (loc, na, t, b) -> aux i b
+ | GLambda (loc, na, bk, t, b) -> abs loc na bk t b
+ | GLetIn (loc, na, t, b) -> aux i b
| _ -> []
in aux 1 l
diff --git a/interp/implicit_quantifiers.mli b/interp/implicit_quantifiers.mli
index fee7babe9..4c73edbf7 100644
--- a/interp/implicit_quantifiers.mli
+++ b/interp/implicit_quantifiers.mli
@@ -37,12 +37,12 @@ val free_vars_of_binders :
(** Returns the generalizable free ids in left-to-right
order with the location of their first occurence *)
-val generalizable_vars_of_rawconstr : ?bound:Idset.t -> ?allowed:Idset.t ->
- rawconstr -> (Names.identifier * loc) list
+val generalizable_vars_of_glob_constr : ?bound:Idset.t -> ?allowed:Idset.t ->
+ glob_constr -> (Names.identifier * loc) list
val make_fresh : Names.Idset.t -> Environ.env -> identifier -> identifier
-val implicits_of_rawterm : ?with_products:bool -> Rawterm.rawconstr -> (Topconstr.explicitation * (bool * bool * bool)) list
+val implicits_of_rawterm : ?with_products:bool -> Rawterm.glob_constr -> (Topconstr.explicitation * (bool * bool * bool)) list
val combine_params_freevar :
Names.Idset.t -> (global_reference * bool) option * (Names.name * Term.constr option * Term.types) ->
diff --git a/interp/notation.ml b/interp/notation.ml
index 09edd7b30..eea8afeef 100644
--- a/interp/notation.ml
+++ b/interp/notation.ml
@@ -197,9 +197,9 @@ let make_gr = function
ConstructRef((mind_of_kn(canonical_mind kn),i),j)
| VarRef id -> VarRef id
-let rawconstr_key = function
- | RApp (_,RRef (_,ref),_) -> RefKey (make_gr ref)
- | RRef (_,ref) -> RefKey (make_gr ref)
+let glob_constr_key = function
+ | GApp (_,GRef (_,ref),_) -> RefKey (make_gr ref)
+ | GRef (_,ref) -> RefKey (make_gr ref)
| _ -> Oth
let cases_pattern_key = function
@@ -219,15 +219,15 @@ let aconstr_key = function (* Rem: AApp(ARef ref,[]) stands for @ref *)
type required_module = full_path * string list
type 'a prim_token_interpreter =
- loc -> 'a -> rawconstr
+ loc -> 'a -> glob_constr
type cases_pattern_status = bool (* true = use prim token in patterns *)
type 'a prim_token_uninterpreter =
- rawconstr list * (rawconstr -> 'a option) * cases_pattern_status
+ glob_constr list * (glob_constr -> 'a option) * cases_pattern_status
type internal_prim_token_interpreter =
- loc -> prim_token -> required_module * (unit -> rawconstr)
+ loc -> prim_token -> required_module * (unit -> glob_constr)
let prim_token_interpreter_tab =
(Hashtbl.create 7 : (scope_name, internal_prim_token_interpreter) Hashtbl.t)
@@ -244,7 +244,7 @@ let declare_prim_token_interpreter sc interp (patl,uninterp,b) =
declare_scope sc;
add_prim_token_interpreter sc interp;
List.iter (fun pat ->
- Hashtbl.add prim_token_key_table (rawconstr_key pat) (sc,uninterp,b))
+ Hashtbl.add prim_token_key_table (glob_constr_key pat) (sc,uninterp,b))
patl
let mkNumeral n = Numeral n
@@ -350,7 +350,7 @@ let find_prim_token g loc p sc =
(* Try for a user-defined numerical notation *)
try
let (_,c),df = find_notation (notation_of_prim_token p) sc in
- g (rawconstr_of_aconstr loc c),df
+ g (glob_constr_of_aconstr loc c),df
with Not_found ->
(* Try for a primitive numerical notation *)
let (spdir,interp) = Hashtbl.find prim_token_interpreter_tab sc loc p in
@@ -370,7 +370,7 @@ let interp_prim_token =
interp_prim_token_gen (fun x -> x)
let interp_prim_token_cases_pattern loc p name =
- interp_prim_token_gen (cases_pattern_of_rawconstr name) loc p
+ interp_prim_token_gen (cases_pattern_of_glob_constr name) loc p
let rec interp_notation loc ntn local_scopes =
let scopes = make_current_scopes local_scopes in
@@ -380,7 +380,7 @@ let rec interp_notation loc ntn local_scopes =
(loc,"",str ("Unknown interpretation for notation \""^ntn^"\"."))
let uninterp_notations c =
- Gmapl.find (rawconstr_key c) !notations_key_table
+ Gmapl.find (glob_constr_key c) !notations_key_table
let uninterp_cases_pattern_notations c =
Gmapl.find (cases_pattern_key c) !notations_key_table
@@ -392,7 +392,7 @@ let availability_of_notation (ntn_scope,ntn) scopes =
let uninterp_prim_token c =
try
- let (sc,numpr,_) = Hashtbl.find prim_token_key_table (rawconstr_key c) in
+ let (sc,numpr,_) = Hashtbl.find prim_token_key_table (glob_constr_key c) in
match numpr c with
| None -> raise No_match
| Some n -> (sc,n)
@@ -403,7 +403,7 @@ let uninterp_prim_token_cases_pattern c =
let k = cases_pattern_key c in
let (sc,numpr,b) = Hashtbl.find prim_token_key_table k in
if not b then raise No_match;
- let na,c = rawconstr_of_closed_cases_pattern c in
+ let na,c = glob_constr_of_closed_cases_pattern c in
match numpr c with
| None -> raise No_match
| Some n -> (na,sc,n)
@@ -581,11 +581,11 @@ let pr_scope_classes sc =
hov 0 (str ("Bound to class"^(if List.tl l=[] then "" else "es")) ++
spc() ++ prlist_with_sep spc pr_class l) ++ fnl()
-let pr_notation_info prraw ntn c =
+let pr_notation_info prglob ntn c =
str "\"" ++ str ntn ++ str "\" := " ++
- prraw (rawconstr_of_aconstr dummy_loc c)
+ prglob (glob_constr_of_aconstr dummy_loc c)
-let pr_named_scope prraw scope sc =
+let pr_named_scope prglob scope sc =
(if scope = default_scope then
match Gmap.fold (fun _ _ x -> x+1) sc.notations 0 with
| 0 -> str "No lonely notation"
@@ -596,14 +596,14 @@ let pr_named_scope prraw scope sc =
++ pr_scope_classes scope
++ Gmap.fold
(fun ntn ((_,r),(_,df)) strm ->
- pr_notation_info prraw df r ++ fnl () ++ strm)
+ pr_notation_info prglob df r ++ fnl () ++ strm)
sc.notations (mt ())
-let pr_scope prraw scope = pr_named_scope prraw scope (find_scope scope)
+let pr_scope prglob scope = pr_named_scope prglob scope (find_scope scope)
-let pr_scopes prraw =
+let pr_scopes prglob =
Gmap.fold
- (fun scope sc strm -> pr_named_scope prraw scope sc ++ fnl () ++ strm)
+ (fun scope sc strm -> pr_named_scope prglob scope sc ++ fnl () ++ strm)
!scope_map (mt ())
let rec find_default ntn = function
@@ -670,7 +670,7 @@ let interp_notation_as_global_reference loc test ntn sc =
| [] -> error_notation_not_reference loc ntn
| _ -> error_ambiguous_notation loc ntn
-let locate_notation prraw ntn scope =
+let locate_notation prglob ntn scope =
let ntns = factorize_entries (browse_notation false ntn !scope_map) in
let scopes = Option.fold_right push_scope scope !scope_stack in
if ntns = [] then
@@ -683,7 +683,7 @@ let locate_notation prraw ntn scope =
prlist
(fun (sc,r,(_,df)) ->
hov 0 (
- pr_notation_info prraw df r ++ tbrk (1,2) ++
+ pr_notation_info prglob df r ++ tbrk (1,2) ++
(if sc = default_scope then mt () else (str ": " ++ str sc)) ++
tbrk (1,2) ++
(if Some sc = scope then str "(default interpretation)" else mt ())
@@ -719,10 +719,10 @@ let collect_notations stack =
(all',ntn::knownntn))
([],[]) stack)
-let pr_visible_in_scope prraw (scope,ntns) =
+let pr_visible_in_scope prglob (scope,ntns) =
let strm =
List.fold_right
- (fun (df,r) strm -> pr_notation_info prraw df r ++ fnl () ++ strm)
+ (fun (df,r) strm -> pr_notation_info prglob df r ++ fnl () ++ strm)
ntns (mt ()) in
(if scope = default_scope then
str "Lonely notation" ++ (if List.length ntns <> 1 then str "s" else mt())
@@ -730,14 +730,14 @@ let pr_visible_in_scope prraw (scope,ntns) =
str "Visible in scope " ++ str scope)
++ fnl () ++ strm
-let pr_scope_stack prraw stack =
+let pr_scope_stack prglob stack =
List.fold_left
- (fun strm scntns -> strm ++ pr_visible_in_scope prraw scntns ++ fnl ())
+ (fun strm scntns -> strm ++ pr_visible_in_scope prglob scntns ++ fnl ())
(mt ()) (collect_notations stack)
-let pr_visibility prraw = function
- | Some scope -> pr_scope_stack prraw (push_scope scope !scope_stack)
- | None -> pr_scope_stack prraw !scope_stack
+let pr_visibility prglob = function
+ | Some scope -> pr_scope_stack prglob (push_scope scope !scope_stack)
+ | None -> pr_scope_stack prglob !scope_stack
(**********************************************************************)
(* Mapping notations to concrete syntax *)
diff --git a/interp/notation.mli b/interp/notation.mli
index 84f92f874..290d5f3df 100644
--- a/interp/notation.mli
+++ b/interp/notation.mli
@@ -65,10 +65,10 @@ type required_module = full_path * string list
type cases_pattern_status = bool (** true = use prim token in patterns *)
type 'a prim_token_interpreter =
- loc -> 'a -> rawconstr
+ loc -> 'a -> glob_constr
type 'a prim_token_uninterpreter =
- rawconstr list * (rawconstr -> 'a option) * cases_pattern_status
+ glob_constr list * (glob_constr -> 'a option) * cases_pattern_status
val declare_numeral_interpreter : scope_name -> required_module ->
bigint prim_token_interpreter -> bigint prim_token_uninterpreter -> unit
@@ -80,7 +80,7 @@ val declare_string_interpreter : scope_name -> required_module ->
given scope context*)
val interp_prim_token : loc -> prim_token -> local_scopes ->
- rawconstr * (notation_location * scope_name option)
+ glob_constr * (notation_location * scope_name option)
val interp_prim_token_cases_pattern : loc -> prim_token -> name ->
local_scopes -> cases_pattern * (notation_location * scope_name option)
@@ -88,7 +88,7 @@ val interp_prim_token_cases_pattern : loc -> prim_token -> name ->
raise [No_match] if no such token *)
val uninterp_prim_token :
- rawconstr -> scope_name * prim_token
+ glob_constr -> scope_name * prim_token
val uninterp_prim_token_cases_pattern :
cases_pattern -> name * scope_name * prim_token
@@ -112,7 +112,7 @@ val interp_notation : loc -> notation -> local_scopes ->
interpretation * (notation_location * scope_name option)
(** Return the possible notations for a given term *)
-val uninterp_notations : rawconstr ->
+val uninterp_notations : glob_constr ->
(interp_rule * interpretation * int option) list
val uninterp_cases_pattern_notations : cases_pattern ->
(interp_rule * interpretation * int option) list
@@ -160,12 +160,12 @@ val make_notation_key : symbol list -> notation
val decompose_notation_key : notation -> symbol list
(** Prints scopes (expects a pure aconstr printer) *)
-val pr_scope : (rawconstr -> std_ppcmds) -> scope_name -> std_ppcmds
-val pr_scopes : (rawconstr -> std_ppcmds) -> std_ppcmds
-val locate_notation : (rawconstr -> std_ppcmds) -> notation ->
+val pr_scope : (glob_constr -> std_ppcmds) -> scope_name -> std_ppcmds
+val pr_scopes : (glob_constr -> std_ppcmds) -> std_ppcmds
+val locate_notation : (glob_constr -> std_ppcmds) -> notation ->
scope_name option -> std_ppcmds
-val pr_visibility: (rawconstr -> std_ppcmds) -> scope_name option -> std_ppcmds
+val pr_visibility: (glob_constr -> std_ppcmds) -> scope_name option -> std_ppcmds
(** {6 Printing rules for notations} *)
diff --git a/interp/reserve.ml b/interp/reserve.ml
index 2d36f2409..9d20236b8 100644
--- a/interp/reserve.ml
+++ b/interp/reserve.ml
@@ -49,42 +49,42 @@ let find_reserved_type id = Idmap.find (root_of_id id) !reserve_table
open Rawterm
let rec unloc = function
- | RVar (_,id) -> RVar (dummy_loc,id)
- | RApp (_,g,args) -> RApp (dummy_loc,unloc g, List.map unloc args)
- | RLambda (_,na,bk,ty,c) -> RLambda (dummy_loc,na,bk,unloc ty,unloc c)
- | RProd (_,na,bk,ty,c) -> RProd (dummy_loc,na,bk,unloc ty,unloc c)
- | RLetIn (_,na,b,c) -> RLetIn (dummy_loc,na,unloc b,unloc c)
- | RCases (_,sty,rtntypopt,tml,pl) ->
- RCases (dummy_loc,sty,
+ | GVar (_,id) -> GVar (dummy_loc,id)
+ | GApp (_,g,args) -> GApp (dummy_loc,unloc g, List.map unloc args)
+ | GLambda (_,na,bk,ty,c) -> GLambda (dummy_loc,na,bk,unloc ty,unloc c)
+ | GProd (_,na,bk,ty,c) -> GProd (dummy_loc,na,bk,unloc ty,unloc c)
+ | GLetIn (_,na,b,c) -> GLetIn (dummy_loc,na,unloc b,unloc c)
+ | GCases (_,sty,rtntypopt,tml,pl) ->
+ GCases (dummy_loc,sty,
(Option.map unloc rtntypopt),
List.map (fun (tm,x) -> (unloc tm,x)) tml,
List.map (fun (_,idl,p,c) -> (dummy_loc,idl,p,unloc c)) pl)
- | RLetTuple (_,nal,(na,po),b,c) ->
- RLetTuple (dummy_loc,nal,(na,Option.map unloc po),unloc b,unloc c)
- | RIf (_,c,(na,po),b1,b2) ->
- RIf (dummy_loc,unloc c,(na,Option.map unloc po),unloc b1,unloc b2)
- | RRec (_,fk,idl,bl,tyl,bv) ->
- RRec (dummy_loc,fk,idl,
+ | GLetTuple (_,nal,(na,po),b,c) ->
+ GLetTuple (dummy_loc,nal,(na,Option.map unloc po),unloc b,unloc c)
+ | GIf (_,c,(na,po),b1,b2) ->
+ GIf (dummy_loc,unloc c,(na,Option.map unloc po),unloc b1,unloc b2)
+ | GRec (_,fk,idl,bl,tyl,bv) ->
+ GRec (dummy_loc,fk,idl,
Array.map (List.map
(fun (na,k,obd,ty) -> (na,k,Option.map unloc obd, unloc ty)))
bl,
Array.map unloc tyl,
Array.map unloc bv)
- | RCast (_,c, CastConv (k,t)) -> RCast (dummy_loc,unloc c, CastConv (k,unloc t))
- | RCast (_,c, CastCoerce) -> RCast (dummy_loc,unloc c, CastCoerce)
- | RSort (_,x) -> RSort (dummy_loc,x)
- | RHole (_,x) -> RHole (dummy_loc,x)
- | RRef (_,x) -> RRef (dummy_loc,x)
- | REvar (_,x,l) -> REvar (dummy_loc,x,l)
- | RPatVar (_,x) -> RPatVar (dummy_loc,x)
- | RDynamic (_,x) -> RDynamic (dummy_loc,x)
+ | GCast (_,c, CastConv (k,t)) -> GCast (dummy_loc,unloc c, CastConv (k,unloc t))
+ | GCast (_,c, CastCoerce) -> GCast (dummy_loc,unloc c, CastCoerce)
+ | GSort (_,x) -> GSort (dummy_loc,x)
+ | GHole (_,x) -> GHole (dummy_loc,x)
+ | GRef (_,x) -> GRef (dummy_loc,x)
+ | GEvar (_,x,l) -> GEvar (dummy_loc,x,l)
+ | GPatVar (_,x) -> GPatVar (dummy_loc,x)
+ | GDynamic (_,x) -> GDynamic (dummy_loc,x)
let anonymize_if_reserved na t = match na with
| Name id as na ->
(try
if not !Flags.raw_print &
- aconstr_of_rawconstr [] [] t = find_reserved_type id
- then RHole (dummy_loc,Evd.BinderType na)
+ aconstr_of_glob_constr [] [] t = find_reserved_type id
+ then GHole (dummy_loc,Evd.BinderType na)
else t
with Not_found -> t)
| Anonymous -> t
diff --git a/interp/reserve.mli b/interp/reserve.mli
index 3bcba719c..1766f77b9 100644
--- a/interp/reserve.mli
+++ b/interp/reserve.mli
@@ -13,4 +13,4 @@ open Topconstr
val declare_reserved_type : identifier located -> aconstr -> unit
val find_reserved_type : identifier -> aconstr
-val anonymize_if_reserved : name -> rawconstr -> rawconstr
+val anonymize_if_reserved : name -> glob_constr -> glob_constr
diff --git a/interp/topconstr.ml b/interp/topconstr.ml
index e27bf6721..61549cb1f 100644
--- a/interp/topconstr.ml
+++ b/interp/topconstr.ml
@@ -18,7 +18,7 @@ open Mod_subst
(*i*)
(**********************************************************************)
-(* This is the subtype of rawconstr allowed in syntactic extensions *)
+(* This is the subtype of glob_constr allowed in syntactic extensions *)
(* For AList: first constr is iterator, second is terminator;
first id is where each argument of the list has to be substituted
@@ -26,12 +26,12 @@ open Mod_subst
boolean is associativity *)
type aconstr =
- (* Part common to rawconstr and cases_pattern *)
+ (* Part common to glob_constr and cases_pattern *)
| ARef of global_reference
| AVar of identifier
| AApp of aconstr * aconstr list
| AList of identifier * identifier * aconstr * aconstr * bool
- (* Part only in rawconstr *)
+ (* Part only in glob_constr *)
| ALambda of name * aconstr * aconstr
| AProd of name * aconstr * aconstr
| ABinderList of identifier * identifier * aconstr * aconstr
@@ -65,7 +65,7 @@ type interpretation =
(identifier * (subscopes * notation_var_instance_type)) list * aconstr
(**********************************************************************)
-(* Re-interpret a notation as a rawconstr, taking care of binders *)
+(* Re-interpret a notation as a glob_constr, taking care of binders *)
let name_to_ident = function
| Anonymous -> error "This expression should be a simple identifier."
@@ -81,43 +81,43 @@ let rec cases_pattern_fold_map loc g e = function
let e',patl' = list_fold_map (cases_pattern_fold_map loc g) e patl in
e', PatCstr (loc,cstr,patl',na')
-let rec subst_rawvars l = function
- | RVar (_,id) as r -> (try List.assoc id l with Not_found -> r)
- | RProd (loc,Name id,bk,t,c) ->
+let rec subst_glob_vars l = function
+ | GVar (_,id) as r -> (try List.assoc id l with Not_found -> r)
+ | GProd (loc,Name id,bk,t,c) ->
let id =
- try match List.assoc id l with RVar(_,id') -> id' | _ -> id
+ try match List.assoc id l with GVar(_,id') -> id' | _ -> id
with Not_found -> id in
- RProd (loc,Name id,bk,subst_rawvars l t,subst_rawvars l c)
- | RLambda (loc,Name id,bk,t,c) ->
+ GProd (loc,Name id,bk,subst_glob_vars l t,subst_glob_vars l c)
+ | GLambda (loc,Name id,bk,t,c) ->
let id =
- try match List.assoc id l with RVar(_,id') -> id' | _ -> id
+ try match List.assoc id l with GVar(_,id') -> id' | _ -> id
with Not_found -> id in
- RLambda (loc,Name id,bk,subst_rawvars l t,subst_rawvars l c)
- | r -> map_rawconstr (subst_rawvars l) r (* assume: id is not binding *)
+ GLambda (loc,Name id,bk,subst_glob_vars l t,subst_glob_vars l c)
+ | r -> map_glob_constr (subst_glob_vars l) r (* assume: id is not binding *)
let ldots_var = id_of_string ".."
-let rawconstr_of_aconstr_with_binders loc g f e = function
- | AVar id -> RVar (loc,id)
- | AApp (a,args) -> RApp (loc,f e a, List.map (f e) args)
+let glob_constr_of_aconstr_with_binders loc g f e = function
+ | AVar id -> GVar (loc,id)
+ | AApp (a,args) -> GApp (loc,f e a, List.map (f e) args)
| AList (x,y,iter,tail,swap) ->
let t = f e tail in let it = f e iter in
- let innerl = (ldots_var,t)::(if swap then [] else [x,RVar(loc,y)]) in
- let inner = RApp (loc,RVar (loc,ldots_var),[subst_rawvars innerl it]) in
- let outerl = (ldots_var,inner)::(if swap then [x,RVar(loc,y)] else []) in
- subst_rawvars outerl it
+ let innerl = (ldots_var,t)::(if swap then [] else [x,GVar(loc,y)]) in
+ let inner = GApp (loc,GVar (loc,ldots_var),[subst_glob_vars innerl it]) in
+ let outerl = (ldots_var,inner)::(if swap then [x,GVar(loc,y)] else []) in
+ subst_glob_vars outerl it
| ABinderList (x,y,iter,tail) ->
let t = f e tail in let it = f e iter in
- let innerl = [(ldots_var,t);(x,RVar(loc,y))] in
- let inner = RApp (loc,RVar (loc,ldots_var),[subst_rawvars innerl it]) in
+ let innerl = [(ldots_var,t);(x,GVar(loc,y))] in
+ let inner = GApp (loc,GVar (loc,ldots_var),[subst_glob_vars innerl it]) in
let outerl = [(ldots_var,inner)] in
- subst_rawvars outerl it
+ subst_glob_vars outerl it
| ALambda (na,ty,c) ->
- let e',na = g e na in RLambda (loc,na,Explicit,f e ty,f e' c)
+ let e',na = g e na in GLambda (loc,na,Explicit,f e ty,f e' c)
| AProd (na,ty,c) ->
- let e',na = g e na in RProd (loc,na,Explicit,f e ty,f e' c)
+ let e',na = g e na in GProd (loc,na,Explicit,f e ty,f e' c)
| ALetIn (na,b,c) ->
- let e',na = g e na in RLetIn (loc,na,f e b,f e' c)
+ let e',na = g e na in GLetIn (loc,na,f e b,f e' c)
| ACases (sty,rtntypopt,tml,eqnl) ->
let e',tml' = List.fold_right (fun (tm,(na,t)) (e',tml') ->
let e',t' = match t with
@@ -133,36 +133,36 @@ let rawconstr_of_aconstr_with_binders loc g f e = function
let ((idl,e),patl) =
list_fold_map (cases_pattern_fold_map loc fold) ([],e) patl in
(loc,idl,patl,f e rhs)) eqnl in
- RCases (loc,sty,Option.map (f e') rtntypopt,tml',eqnl')
+ GCases (loc,sty,Option.map (f e') rtntypopt,tml',eqnl')
| ALetTuple (nal,(na,po),b,c) ->
let e',nal = list_fold_map g e nal in
let e'',na = g e na in
- RLetTuple (loc,nal,(na,Option.map (f e'') po),f e b,f e' c)
+ GLetTuple (loc,nal,(na,Option.map (f e'') po),f e b,f e' c)
| AIf (c,(na,po),b1,b2) ->
let e',na = g e na in
- RIf (loc,f e c,(na,Option.map (f e') po),f e b1,f e b2)
+ GIf (loc,f e c,(na,Option.map (f e') po),f e b1,f e b2)
| ARec (fk,idl,dll,tl,bl) ->
let e,dll = array_fold_map (list_fold_map (fun e (na,oc,b) ->
let e,na = g e na in
(e,(na,Explicit,Option.map (f e) oc,f e b)))) e dll in
let e',idl = array_fold_map (to_id g) e idl in
- RRec (loc,fk,idl,dll,Array.map (f e) tl,Array.map (f e') bl)
- | ACast (c,k) -> RCast (loc,f e c,
+ GRec (loc,fk,idl,dll,Array.map (f e) tl,Array.map (f e') bl)
+ | ACast (c,k) -> GCast (loc,f e c,
match k with
| CastConv (k,t) -> CastConv (k,f e t)
| CastCoerce -> CastCoerce)
- | ASort x -> RSort (loc,x)
- | AHole x -> RHole (loc,x)
- | APatVar n -> RPatVar (loc,(false,n))
- | ARef x -> RRef (loc,x)
+ | ASort x -> GSort (loc,x)
+ | AHole x -> GHole (loc,x)
+ | APatVar n -> GPatVar (loc,(false,n))
+ | ARef x -> GRef (loc,x)
-let rec rawconstr_of_aconstr loc x =
+let rec glob_constr_of_aconstr loc x =
let rec aux () x =
- rawconstr_of_aconstr_with_binders loc (fun () id -> ((),id)) aux () x
+ glob_constr_of_aconstr_with_binders loc (fun () id -> ((),id)) aux () x
in aux () x
(****************************************************************************)
-(* Translating a rawconstr into a notation, interpreting recursive patterns *)
+(* Translating a glob_constr into a notation, interpreting recursive patterns *)
let add_id r id = r := (id :: pi1 !r, pi2 !r, pi3 !r)
let add_name r = function Anonymous -> () | Name id -> add_id r id
@@ -170,51 +170,51 @@ let add_name r = function Anonymous -> () | Name id -> add_id r id
let split_at_recursive_part c =
let sub = ref None in
let rec aux = function
- | RApp (loc0,RVar(loc,v),c::l) when v = ldots_var ->
+ | GApp (loc0,GVar(loc,v),c::l) when v = ldots_var ->
if !sub <> None then
(* Not narrowed enough to find only one recursive part *)
raise Not_found
else
(sub := Some c;
- if l = [] then RVar (loc,ldots_var)
- else RApp (loc0,RVar (loc,ldots_var),l))
- | c -> map_rawconstr aux c in
+ if l = [] then GVar (loc,ldots_var)
+ else GApp (loc0,GVar (loc,ldots_var),l))
+ | c -> map_glob_constr aux c in
let outer_iterator = aux c in
match !sub with
| None -> (* No recursive pattern found *) raise Not_found
| Some c ->
match outer_iterator with
- | RVar (_,v) when v = ldots_var -> (* Not enough context *) raise Not_found
+ | GVar (_,v) when v = ldots_var -> (* Not enough context *) raise Not_found
| _ -> outer_iterator, c
let on_true_do b f c = if b then (f c; b) else b
-let compare_rawconstr f add t1 t2 = match t1,t2 with
- | RRef (_,r1), RRef (_,r2) -> eq_gr r1 r2
- | RVar (_,v1), RVar (_,v2) -> on_true_do (v1 = v2) add (Name v1)
- | RApp (_,f1,l1), RApp (_,f2,l2) -> f f1 f2 & list_for_all2eq f l1 l2
- | RLambda (_,na1,bk1,ty1,c1), RLambda (_,na2,bk2,ty2,c2) when na1 = na2 && bk1 = bk2 -> on_true_do (f ty1 ty2 & f c1 c2) add na1
- | RProd (_,na1,bk1,ty1,c1), RProd (_,na2,bk2,ty2,c2) when na1 = na2 && bk1 = bk2 ->
+let compare_glob_constr f add t1 t2 = match t1,t2 with
+ | GRef (_,r1), GRef (_,r2) -> eq_gr r1 r2
+ | GVar (_,v1), GVar (_,v2) -> on_true_do (v1 = v2) add (Name v1)
+ | GApp (_,f1,l1), GApp (_,f2,l2) -> f f1 f2 & list_for_all2eq f l1 l2
+ | GLambda (_,na1,bk1,ty1,c1), GLambda (_,na2,bk2,ty2,c2) when na1 = na2 && bk1 = bk2 -> on_true_do (f ty1 ty2 & f c1 c2) add na1
+ | GProd (_,na1,bk1,ty1,c1), GProd (_,na2,bk2,ty2,c2) when na1 = na2 && bk1 = bk2 ->
on_true_do (f ty1 ty2 & f c1 c2) add na1
- | RHole _, RHole _ -> true
- | RSort (_,s1), RSort (_,s2) -> s1 = s2
- | RLetIn (_,na1,b1,c1), RLetIn (_,na2,b2,c2) when na1 = na2 ->
+ | GHole _, GHole _ -> true
+ | GSort (_,s1), GSort (_,s2) -> s1 = s2
+ | GLetIn (_,na1,b1,c1), GLetIn (_,na2,b2,c2) when na1 = na2 ->
on_true_do (f b1 b2 & f c1 c2) add na1
- | (RCases _ | RRec _ | RDynamic _
- | RPatVar _ | REvar _ | RLetTuple _ | RIf _ | RCast _),_
- | _,(RCases _ | RRec _ | RDynamic _
- | RPatVar _ | REvar _ | RLetTuple _ | RIf _ | RCast _)
+ | (GCases _ | GRec _ | GDynamic _
+ | GPatVar _ | GEvar _ | GLetTuple _ | GIf _ | GCast _),_
+ | _,(GCases _ | GRec _ | GDynamic _
+ | GPatVar _ | GEvar _ | GLetTuple _ | GIf _ | GCast _)
-> error "Unsupported construction in recursive notations."
- | (RRef _ | RVar _ | RApp _ | RLambda _ | RProd _
- | RHole _ | RSort _ | RLetIn _), _
+ | (GRef _ | GVar _ | GApp _ | GLambda _ | GProd _
+ | GHole _ | GSort _ | GLetIn _), _
-> false
-let rec eq_rawconstr t1 t2 = compare_rawconstr eq_rawconstr (fun _ -> ()) t1 t2
+let rec eq_glob_constr t1 t2 = compare_glob_constr eq_glob_constr (fun _ -> ()) t1 t2
let subtract_loc loc1 loc2 = make_loc (fst (unloc loc1),fst (unloc loc2)-1)
-let check_is_hole id = function RHole _ -> () | t ->
- user_err_loc (loc_of_rawconstr t,"",
+let check_is_hole id = function GHole _ -> () | t ->
+ user_err_loc (loc_of_glob_constr t,"",
strbrk "In recursive notation with binders, " ++ pr_id id ++
strbrk " is expected to come without type.")
@@ -222,40 +222,40 @@ let compare_recursive_parts found f (iterator,subc) =
let diff = ref None in
let terminator = ref None in
let rec aux c1 c2 = match c1,c2 with
- | RVar(_,v), term when v = ldots_var ->
+ | GVar(_,v), term when v = ldots_var ->
(* We found the pattern *)
assert (!terminator = None); terminator := Some term;
true
- | RApp (_,RVar(_,v),l1), RApp (_,term,l2) when v = ldots_var ->
+ | GApp (_,GVar(_,v),l1), GApp (_,term,l2) when v = ldots_var ->
(* We found the pattern, but there are extra arguments *)
(* (this allows e.g. alternative (recursive) notation of application) *)
assert (!terminator = None); terminator := Some term;
list_for_all2eq aux l1 l2
- | RVar (_,x), RVar (_,y) when x<>y ->
+ | GVar (_,x), GVar (_,y) when x<>y ->
(* We found the position where it differs *)
let lassoc = (!terminator <> None) in
let x,y = if lassoc then y,x else x,y in
!diff = None && (diff := Some (x,y,Some lassoc); true)
- | RLambda (_,Name x,_,t_x,c), RLambda (_,Name y,_,t_y,term)
- | RProd (_,Name x,_,t_x,c), RProd (_,Name y,_,t_y,term) ->
+ | GLambda (_,Name x,_,t_x,c), GLambda (_,Name y,_,t_y,term)
+ | GProd (_,Name x,_,t_x,c), GProd (_,Name y,_,t_y,term) ->
(* We found a binding position where it differs *)
check_is_hole y t_x;
check_is_hole y t_y;
!diff = None && (diff := Some (x,y,None); aux c term)
| _ ->
- compare_rawconstr aux (add_name found) c1 c2 in
+ compare_glob_constr aux (add_name found) c1 c2 in
if aux iterator subc then
match !diff with
| None ->
- let loc1 = loc_of_rawconstr iterator in
- let loc2 = loc_of_rawconstr (Option.get !terminator) in
+ let loc1 = loc_of_glob_constr iterator in
+ let loc2 = loc_of_glob_constr (Option.get !terminator) in
(* Here, we would need a loc made of several parts ... *)
user_err_loc (subtract_loc loc1 loc2,"",
str "Both ends of the recursive pattern are the same.")
| Some (x,y,Some lassoc) ->
let newfound = (pi1 !found, (x,y) :: pi2 !found, pi3 !found) in
let iterator =
- f (if lassoc then subst_rawvars [y,RVar(dummy_loc,x)] iterator
+ f (if lassoc then subst_glob_vars [y,GVar(dummy_loc,x)] iterator
else iterator) in
(* found have been collected by compare_constr *)
found := newfound;
@@ -269,7 +269,7 @@ let compare_recursive_parts found f (iterator,subc) =
else
raise Not_found
-let aconstr_and_vars_of_rawconstr a =
+let aconstr_and_vars_of_glob_constr a =
let found = ref ([],[],[]) in
let rec aux c =
let keepfound = !found in
@@ -278,7 +278,7 @@ let aconstr_and_vars_of_rawconstr a =
with Not_found ->
found := keepfound;
match c with
- | RApp (_,RVar (loc,f),[c]) when f = ldots_var ->
+ | GApp (_,GVar (loc,f),[c]) when f = ldots_var ->
(* Fall on the second part of the recursive pattern w/o having
found the first part *)
user_err_loc (loc,"",
@@ -286,12 +286,12 @@ let aconstr_and_vars_of_rawconstr a =
| c ->
aux' c
and aux' = function
- | RVar (_,id) -> add_id found id; AVar id
- | RApp (_,g,args) -> AApp (aux g, List.map aux args)
- | RLambda (_,na,bk,ty,c) -> add_name found na; ALambda (na,aux ty,aux c)
- | RProd (_,na,bk,ty,c) -> add_name found na; AProd (na,aux ty,aux c)
- | RLetIn (_,na,b,c) -> add_name found na; ALetIn (na,aux b,aux c)
- | RCases (_,sty,rtntypopt,tml,eqnl) ->
+ | GVar (_,id) -> add_id found id; AVar id
+ | GApp (_,g,args) -> AApp (aux g, List.map aux args)
+ | GLambda (_,na,bk,ty,c) -> add_name found na; ALambda (na,aux ty,aux c)
+ | GProd (_,na,bk,ty,c) -> add_name found na; AProd (na,aux ty,aux c)
+ | GLetIn (_,na,b,c) -> add_name found na; ALetIn (na,aux b,aux c)
+ | GCases (_,sty,rtntypopt,tml,eqnl) ->
let f (_,idl,pat,rhs) = List.iter (add_id found) idl; (pat,aux rhs) in
ACases (sty,Option.map aux rtntypopt,
List.map (fun (tm,(na,x)) ->
@@ -300,28 +300,28 @@ let aconstr_and_vars_of_rawconstr a =
(fun (_,_,_,nl) -> List.iter (add_name found) nl) x;
(aux tm,(na,Option.map (fun (_,ind,n,nal) -> (ind,n,nal)) x))) tml,
List.map f eqnl)
- | RLetTuple (loc,nal,(na,po),b,c) ->
+ | GLetTuple (loc,nal,(na,po),b,c) ->
add_name found na;
List.iter (add_name found) nal;
ALetTuple (nal,(na,Option.map aux po),aux b,aux c)
- | RIf (loc,c,(na,po),b1,b2) ->
+ | GIf (loc,c,(na,po),b1,b2) ->
add_name found na;
AIf (aux c,(na,Option.map aux po),aux b1,aux b2)
- | RRec (_,fk,idl,dll,tl,bl) ->
+ | GRec (_,fk,idl,dll,tl,bl) ->
Array.iter (add_id found) idl;
let dll = Array.map (List.map (fun (na,bk,oc,b) ->
if bk <> Explicit then
error "Binders marked as implicit not allowed in notations.";
add_name found na; (na,Option.map aux oc,aux b))) dll in
ARec (fk,idl,dll,Array.map aux tl,Array.map aux bl)
- | RCast (_,c,k) -> ACast (aux c,
+ | GCast (_,c,k) -> ACast (aux c,
match k with CastConv (k,t) -> CastConv (k,aux t)
| CastCoerce -> CastCoerce)
- | RSort (_,s) -> ASort s
- | RHole (_,w) -> AHole w
- | RRef (_,r) -> ARef r
- | RPatVar (_,(_,n)) -> APatVar n
- | RDynamic _ | REvar _ ->
+ | GSort (_,s) -> ASort s
+ | GHole (_,w) -> AHole w
+ | GRef (_,r) -> ARef r
+ | GPatVar (_,(_,n)) -> APatVar n
+ | GDynamic _ | GEvar _ ->
error "Existential variables not allowed in notations."
in
@@ -370,15 +370,15 @@ let check_variables vars recvars (found,foundrec,foundrecbinding) =
| NtnInternTypeIdent -> check_bound x in
List.iter check_type vars
-let aconstr_of_rawconstr vars recvars a =
- let a,found = aconstr_and_vars_of_rawconstr a in
+let aconstr_of_glob_constr vars recvars a =
+ let a,found = aconstr_and_vars_of_glob_constr a in
check_variables vars recvars found;
a
(* Substitution of kernel names, avoiding a list of bound identifiers *)
let aconstr_of_constr avoiding t =
- aconstr_of_rawconstr [] [] (Detyping.detype false avoiding [] t)
+ aconstr_of_glob_constr [] [] (Detyping.detype false avoiding [] t)
let rec subst_pat subst pat =
match pat with
@@ -508,7 +508,7 @@ let subst_interpretation subst (metas,pat) =
let bound = List.map fst metas in
(metas,subst_aconstr subst bound pat)
-(* Pattern-matching rawconstr and aconstr *)
+(* Pattern-matching glob_constr and aconstr *)
let abstract_return_type_context pi mklam tml rtno =
Option.map (fun rtn ->
@@ -518,9 +518,9 @@ let abstract_return_type_context pi mklam tml rtno =
List.fold_right mklam nal rtn)
rtno
-let abstract_return_type_context_rawconstr =
+let abstract_return_type_context_glob_constr =
abstract_return_type_context (fun (_,_,_,nal) -> nal)
- (fun na c -> RLambda(dummy_loc,na,Explicit,RHole(dummy_loc,Evd.InternalHole),c))
+ (fun na c -> GLambda(dummy_loc,na,Explicit,GHole(dummy_loc,Evd.InternalHole),c))
let abstract_return_type_context_aconstr =
abstract_return_type_context pi3
@@ -543,7 +543,7 @@ let bind_env alp (sigma,sigmalist,sigmabinders as fullsigma) var v =
else raise No_match
with Not_found ->
(* Check that no capture of binding variables occur *)
- if List.exists (fun (id,_) ->occur_rawconstr id v) alp then raise No_match;
+ if List.exists (fun (id,_) ->occur_glob_constr id v) alp then raise No_match;
(* TODO: handle the case of multiple occs in different scopes *)
((var,v)::sigma,sigmalist,sigmabinders)
@@ -565,7 +565,7 @@ let match_opt f sigma t1 t2 = match (t1,t2) with
let match_names metas (alp,sigma) na1 na2 = match (na1,na2) with
| (Name id1,Name id2) when List.mem id2 (fst metas) ->
- alp, bind_env alp sigma id2 (RVar (dummy_loc,id1))
+ alp, bind_env alp sigma id2 (GVar (dummy_loc,id1))
| (Name id1,Name id2) -> (id1,id2)::alp,sigma
| (Anonymous,Anonymous) -> alp,sigma
| _ -> raise No_match
@@ -582,13 +582,13 @@ let rec match_cases_pattern_binders metas acc pat1 pat2 =
let glue_letin_with_decls = true
let rec match_iterated_binders islambda decls = function
- | RLambda (_,na,bk,t,b) when islambda ->
+ | GLambda (_,na,bk,t,b) when islambda ->
match_iterated_binders islambda ((na,bk,None,t)::decls) b
- | RProd (_,(Name _ as na),bk,t,b) when not islambda ->
+ | GProd (_,(Name _ as na),bk,t,b) when not islambda ->
match_iterated_binders islambda ((na,bk,None,t)::decls) b
- | RLetIn (loc,na,c,b) when glue_letin_with_decls ->
+ | GLetIn (loc,na,c,b) when glue_letin_with_decls ->
match_iterated_binders islambda
- ((na,Explicit (*?*), Some c,RHole(loc,Evd.BinderType na))::decls) b
+ ((na,Explicit (*?*), Some c,GHole(loc,Evd.BinderType na))::decls) b
| b -> (decls,b)
let remove_sigma x (sigmavar,sigmalist,sigmabinders) =
@@ -630,11 +630,11 @@ let rec match_ alp (tmetas,blmetas as metas) sigma a1 a2 = match (a1,a2) with
match_alist (match_ alp) metas sigma r1 x iter termin lassoc
(* Matching recursive notations for binders: ad hoc cases supporting let-in *)
- | RLambda (_,na1,bk,t1,b1), ABinderList (x,_,ALambda (Name id2,_,b2),termin)->
+ | GLambda (_,na1,bk,t1,b1), ABinderList (x,_,ALambda (Name id2,_,b2),termin)->
let (decls,b) = match_iterated_binders true [(na1,bk,None,t1)] b1 in
(* TODO: address the possibility that termin is a Lambda itself *)
match_ alp metas (bind_binder sigma x decls) b termin
- | RProd (_,na1,bk,t1,b1), ABinderList (x,_,AProd (Name id2,_,b2),termin)
+ | GProd (_,na1,bk,t1,b1), ABinderList (x,_,AProd (Name id2,_,b2),termin)
when na1 <> Anonymous ->
let (decls,b) = match_iterated_binders false [(na1,bk,None,t1)] b1 in
(* TODO: address the possibility that termin is a Prod itself *)
@@ -644,36 +644,36 @@ let rec match_ alp (tmetas,blmetas as metas) sigma a1 a2 = match (a1,a2) with
match_abinderlist_with_app (match_ alp) metas sigma r x iter termin
(* Matching individual binders as part of a recursive pattern *)
- | RLambda (_,na,bk,t,b1), ALambda (Name id,_,b2) when List.mem id blmetas ->
+ | GLambda (_,na,bk,t,b1), ALambda (Name id,_,b2) when List.mem id blmetas ->
match_ alp metas (bind_binder sigma id [(na,bk,None,t)]) b1 b2
- | RProd (_,na,bk,t,b1), AProd (Name id,_,b2)
+ | GProd (_,na,bk,t,b1), AProd (Name id,_,b2)
when List.mem id blmetas & na <> Anonymous ->
match_ alp metas (bind_binder sigma id [(na,bk,None,t)]) b1 b2
(* Matching compositionally *)
- | RVar (_,id1), AVar id2 when alpha_var id1 id2 alp -> sigma
- | RRef (_,r1), ARef r2 when (eq_gr r1 r2) -> sigma
- | RPatVar (_,(_,n1)), APatVar n2 when n1=n2 -> sigma
- | RApp (loc,f1,l1), AApp (f2,l2) ->
+ | GVar (_,id1), AVar id2 when alpha_var id1 id2 alp -> sigma
+ | GRef (_,r1), ARef r2 when (eq_gr r1 r2) -> sigma
+ | GPatVar (_,(_,n1)), APatVar n2 when n1=n2 -> sigma
+ | GApp (loc,f1,l1), AApp (f2,l2) ->
let n1 = List.length l1 and n2 = List.length l2 in
let f1,l1,f2,l2 =
if n1 < n2 then
let l21,l22 = list_chop (n2-n1) l2 in f1,l1, AApp (f2,l21), l22
else if n1 > n2 then
- let l11,l12 = list_chop (n1-n2) l1 in RApp (loc,f1,l11),l12, f2,l2
+ let l11,l12 = list_chop (n1-n2) l1 in GApp (loc,f1,l11),l12, f2,l2
else f1,l1, f2, l2 in
List.fold_left2 (match_ alp metas) (match_ alp metas sigma f1 f2) l1 l2
- | RLambda (_,na1,_,t1,b1), ALambda (na2,t2,b2) ->
+ | GLambda (_,na1,_,t1,b1), ALambda (na2,t2,b2) ->
match_binders alp metas na1 na2 (match_ alp metas sigma t1 t2) b1 b2
- | RProd (_,na1,_,t1,b1), AProd (na2,t2,b2) ->
+ | GProd (_,na1,_,t1,b1), AProd (na2,t2,b2) ->
match_binders alp metas na1 na2 (match_ alp metas sigma t1 t2) b1 b2
- | RLetIn (_,na1,t1,b1), ALetIn (na2,t2,b2) ->
+ | GLetIn (_,na1,t1,b1), ALetIn (na2,t2,b2) ->
match_binders alp metas na1 na2 (match_ alp metas sigma t1 t2) b1 b2
- | RCases (_,sty1,rtno1,tml1,eqnl1), ACases (sty2,rtno2,tml2,eqnl2)
+ | GCases (_,sty1,rtno1,tml1,eqnl1), ACases (sty2,rtno2,tml2,eqnl2)
when sty1 = sty2
& List.length tml1 = List.length tml2
& List.length eqnl1 = List.length eqnl2 ->
- let rtno1' = abstract_return_type_context_rawconstr tml1 rtno1 in
+ let rtno1' = abstract_return_type_context_glob_constr tml1 rtno1 in
let rtno2' = abstract_return_type_context_aconstr tml2 rtno2 in
let sigma =
try Option.fold_left2 (match_ alp metas) sigma rtno1' rtno2'
@@ -682,17 +682,17 @@ let rec match_ alp (tmetas,blmetas as metas) sigma a1 a2 = match (a1,a2) with
let sigma = List.fold_left2
(fun s (tm1,_) (tm2,_) -> match_ alp metas s tm1 tm2) sigma tml1 tml2 in
List.fold_left2 (match_equations alp metas) sigma eqnl1 eqnl2
- | RLetTuple (_,nal1,(na1,to1),b1,c1), ALetTuple (nal2,(na2,to2),b2,c2)
+ | GLetTuple (_,nal1,(na1,to1),b1,c1), ALetTuple (nal2,(na2,to2),b2,c2)
when List.length nal1 = List.length nal2 ->
let sigma = match_opt (match_binders alp metas na1 na2) sigma to1 to2 in
let sigma = match_ alp metas sigma b1 b2 in
let (alp,sigma) =
List.fold_left2 (match_names metas) (alp,sigma) nal1 nal2 in
match_ alp metas sigma c1 c2
- | RIf (_,a1,(na1,to1),b1,c1), AIf (a2,(na2,to2),b2,c2) ->
+ | GIf (_,a1,(na1,to1),b1,c1), AIf (a2,(na2,to2),b2,c2) ->
let sigma = match_opt (match_binders alp metas na1 na2) sigma to1 to2 in
List.fold_left2 (match_ alp metas) sigma [a1;b1;c1] [a2;b2;c2]
- | RRec (_,fk1,idl1,dll1,tl1,bl1), ARec (fk2,idl2,dll2,tl2,bl2)
+ | GRec (_,fk1,idl1,dll1,tl1,bl1), ARec (fk2,idl2,dll2,tl2,bl2)
when match_fix_kind fk1 fk2 & Array.length idl1 = Array.length idl2 &
array_for_all2 (fun l1 l2 -> List.length l1 = List.length l2) dll1 dll2
->
@@ -705,14 +705,14 @@ let rec match_ alp (tmetas,blmetas as metas) sigma a1 a2 = match (a1,a2) with
let alp,sigma = array_fold_right2 (fun id1 id2 alsig ->
match_names metas alsig (Name id1) (Name id2)) idl1 idl2 (alp,sigma) in
array_fold_left2 (match_ alp metas) sigma bl1 bl2
- | RCast(_,c1, CastConv(_,t1)), ACast(c2, CastConv (_,t2)) ->
+ | GCast(_,c1, CastConv(_,t1)), ACast(c2, CastConv (_,t2)) ->
match_ alp metas (match_ alp metas sigma c1 c2) t1 t2
- | RCast(_,c1, CastCoerce), ACast(c2, CastCoerce) ->
+ | GCast(_,c1, CastCoerce), ACast(c2, CastCoerce) ->
match_ alp metas sigma c1 c2
- | RSort (_,s1), ASort s2 when s1 = s2 -> sigma
- | RPatVar _, AHole _ -> (*Don't hide Metas, they bind in ltac*) raise No_match
+ | GSort (_,s1), ASort s2 when s1 = s2 -> sigma
+ | GPatVar _, AHole _ -> (*Don't hide Metas, they bind in ltac*) raise No_match
| a, AHole _ -> sigma
- | (RDynamic _ | RRec _ | REvar _), _
+ | (GDynamic _ | GRec _ | GEvar _), _
| _,_ -> raise No_match
and match_binders alp metas na1 na2 sigma b1 b2 =
@@ -737,7 +737,7 @@ let match_aconstr c (metas,pat) =
with Not_found ->
(* Happens for binders bound to Anonymous *)
(* Find a better way to propagate Anonymous... *)
- RVar (dummy_loc,x) in
+ GVar (dummy_loc,x) in
List.fold_right (fun (x,(scl,typ)) (terms',termlists',binders') ->
match typ with
| NtnTypeConstr ->
diff --git a/interp/topconstr.mli b/interp/topconstr.mli
index cb4ac5e84..6e8769b85 100644
--- a/interp/topconstr.mli
+++ b/interp/topconstr.mli
@@ -17,18 +17,18 @@ open Mod_subst
(** Topconstr: definitions of [aconstr] et [constr_expr] *)
(** {6 aconstr } *)
-(** This is the subtype of rawconstr allowed in syntactic extensions
+(** This is the subtype of glob_constr allowed in syntactic extensions
No location since intended to be substituted at any place of a text
Complex expressions such as fixpoints and cofixpoints are excluded,
non global expressions such as existential variables also *)
type aconstr =
- (** Part common to [rawconstr] and [cases_pattern] *)
+ (** Part common to [glob_constr] and [cases_pattern] *)
| ARef of global_reference
| AVar of identifier
| AApp of aconstr * aconstr list
| AList of identifier * identifier * aconstr * aconstr * bool
- (** Part only in [rawconstr] *)
+ (** Part only in [glob_constr] *)
| ALambda of name * aconstr * aconstr
| AProd of name * aconstr * aconstr
| ABinderList of identifier * identifier * aconstr * aconstr
@@ -67,35 +67,35 @@ type notation_var_internalization_type =
type interpretation =
(identifier * (subscopes * notation_var_instance_type)) list * aconstr
-(** Translate a rawconstr into a notation given the list of variables
+(** Translate a glob_constr into a notation given the list of variables
bound by the notation; also interpret recursive patterns *)
-val aconstr_of_rawconstr :
+val aconstr_of_glob_constr :
(identifier * notation_var_internalization_type) list ->
- (identifier * identifier) list -> rawconstr -> aconstr
+ (identifier * identifier) list -> glob_constr -> aconstr
(** Name of the special identifier used to encode recursive notations *)
val ldots_var : identifier
-(** Equality of rawconstr (warning: only partially implemented) *)
-val eq_rawconstr : rawconstr -> rawconstr -> bool
+(** Equality of glob_constr (warning: only partially implemented) *)
+val eq_glob_constr : glob_constr -> glob_constr -> bool
-(** Re-interpret a notation as a rawconstr, taking care of binders *)
+(** Re-interpret a notation as a glob_constr, taking care of binders *)
-val rawconstr_of_aconstr_with_binders : loc ->
+val glob_constr_of_aconstr_with_binders : loc ->
('a -> name -> 'a * name) ->
- ('a -> aconstr -> rawconstr) -> 'a -> aconstr -> rawconstr
+ ('a -> aconstr -> glob_constr) -> 'a -> aconstr -> glob_constr
-val rawconstr_of_aconstr : loc -> aconstr -> rawconstr
+val glob_constr_of_aconstr : loc -> aconstr -> glob_constr
-(** [match_aconstr] matches a rawconstr against a notation interpretation;
+(** [match_aconstr] matches a glob_constr against a notation interpretation;
raise [No_match] if the matching fails *)
exception No_match
-val match_aconstr : rawconstr -> interpretation ->
- (rawconstr * subscopes) list * (rawconstr list * subscopes) list *
- (rawdecl list * subscopes) list
+val match_aconstr : glob_constr -> interpretation ->
+ (glob_constr * subscopes) list * (glob_constr list * subscopes) list *
+ (glob_decl list * subscopes) list
val match_aconstr_cases_pattern : cases_pattern -> interpretation ->
(cases_pattern * subscopes) list * (cases_pattern list * subscopes) list
diff --git a/parsing/g_xml.ml4 b/parsing/g_xml.ml4
index a35eb419d..de3fbb683 100644
--- a/parsing/g_xml.ml4
+++ b/parsing/g_xml.ml4
@@ -135,37 +135,37 @@ let compute_branches_lengths ind =
let compute_inductive_nargs ind =
Inductiveops.inductive_nargs (Global.env()) ind
-(* Interpreting constr as a rawconstr *)
+(* Interpreting constr as a glob_constr *)
let rec interp_xml_constr = function
| XmlTag (loc,"REL",al,[]) ->
- RVar (loc, get_xml_ident al)
+ GVar (loc, get_xml_ident al)
| XmlTag (loc,"VAR",al,[]) ->
error "XML parser: unable to interp free variables"
| XmlTag (loc,"LAMBDA",al,(_::_ as xl)) ->
let body,decls = list_sep_last xl in
let ctx = List.map interp_xml_decl decls in
- List.fold_right (fun (na,t) b -> RLambda (loc, na, Explicit, t, b))
+ List.fold_right (fun (na,t) b -> GLambda (loc, na, Explicit, t, b))
ctx (interp_xml_target body)
| XmlTag (loc,"PROD",al,(_::_ as xl)) ->
let body,decls = list_sep_last xl in
let ctx = List.map interp_xml_decl decls in
- List.fold_right (fun (na,t) b -> RProd (loc, na, Explicit, t, b))
+ List.fold_right (fun (na,t) b -> GProd (loc, na, Explicit, t, b))
ctx (interp_xml_target body)
| XmlTag (loc,"LETIN",al,(_::_ as xl)) ->
let body,defs = list_sep_last xl in
let ctx = List.map interp_xml_def defs in
- List.fold_right (fun (na,t) b -> RLetIn (loc, na, t, b))
+ List.fold_right (fun (na,t) b -> GLetIn (loc, na, t, b))
ctx (interp_xml_target body)
| XmlTag (loc,"APPLY",_,x::xl) ->
- RApp (loc, interp_xml_constr x, List.map interp_xml_constr xl)
+ GApp (loc, interp_xml_constr x, List.map interp_xml_constr xl)
| XmlTag (loc,"instantiate",_,
(XmlTag (_,("CONST"|"MUTIND"|"MUTCONSTRUCT"),_,_) as x)::xl) ->
- RApp (loc, interp_xml_constr x, List.map interp_xml_arg xl)
+ GApp (loc, interp_xml_constr x, List.map interp_xml_arg xl)
| XmlTag (loc,"META",al,xl) ->
- REvar (loc, get_xml_no al, Some (List.map interp_xml_substitution xl))
+ GEvar (loc, get_xml_no al, Some (List.map interp_xml_substitution xl))
| XmlTag (loc,"CONST",al,[]) ->
- RRef (loc, ConstRef (get_xml_constant al))
+ GRef (loc, ConstRef (get_xml_constant al))
| XmlTag (loc,"MUTCASE",al,x::y::yl) ->
let ind = get_xml_inductive al in
let p = interp_xml_patternsType x in
@@ -175,23 +175,23 @@ let rec interp_xml_constr = function
let mat = simple_cases_matrix_of_branches ind brns brs in
let nparams,n = compute_inductive_nargs ind in
let nal,rtn = return_type_of_predicate ind nparams n p in
- RCases (loc,RegularStyle,rtn,[tm,nal],mat)
+ GCases (loc,RegularStyle,rtn,[tm,nal],mat)
| XmlTag (loc,"MUTIND",al,[]) ->
- RRef (loc, IndRef (get_xml_inductive al))
+ GRef (loc, IndRef (get_xml_inductive al))
| XmlTag (loc,"MUTCONSTRUCT",al,[]) ->
- RRef (loc, ConstructRef (get_xml_constructor al))
+ GRef (loc, ConstructRef (get_xml_constructor al))
| XmlTag (loc,"FIX",al,xl) ->
let li,lnct = List.split (List.map interp_xml_FixFunction xl) in
let ln,lc,lt = list_split3 lnct in
let lctx = List.map (fun _ -> []) ln in
- RRec (loc, RFix (Array.of_list li, get_xml_noFun al), Array.of_list ln, Array.of_list lctx, Array.of_list lc, Array.of_list lt)
+ GRec (loc, RFix (Array.of_list li, get_xml_noFun al), Array.of_list ln, Array.of_list lctx, Array.of_list lc, Array.of_list lt)
| XmlTag (loc,"COFIX",al,xl) ->
let ln,lc,lt = list_split3 (List.map interp_xml_CoFixFunction xl) in
- RRec (loc, RCoFix (get_xml_noFun al), Array.of_list ln, [||], Array.of_list lc, Array.of_list lt)
+ GRec (loc, RCoFix (get_xml_noFun al), Array.of_list ln, [||], Array.of_list lc, Array.of_list lt)
| XmlTag (loc,"CAST",al,[x1;x2]) ->
- RCast (loc, interp_xml_term x1, CastConv (DEFAULTcast, interp_xml_type x2))
+ GCast (loc, interp_xml_term x1, CastConv (DEFAULTcast, interp_xml_type x2))
| XmlTag (loc,"SORT",al,[]) ->
- RSort (loc, get_xml_sort al)
+ GSort (loc, get_xml_sort al)
| XmlTag (loc,s,_,_) ->
user_err_loc (loc,"", str "Unexpected tag " ++ str s ++ str ".")
diff --git a/parsing/pptactic.ml b/parsing/pptactic.ml
index 6f0896cc8..0dea0d2ac 100644
--- a/parsing/pptactic.ml
+++ b/parsing/pptactic.ml
@@ -39,8 +39,8 @@ type 'a raw_extra_genarg_printer =
'a -> std_ppcmds
type 'a glob_extra_genarg_printer =
- (rawconstr_and_expr -> std_ppcmds) ->
- (rawconstr_and_expr -> std_ppcmds) ->
+ (glob_constr_and_expr -> std_ppcmds) ->
+ (glob_constr_and_expr -> std_ppcmds) ->
(tolerability -> glob_tactic_expr -> std_ppcmds) ->
'a -> std_ppcmds
@@ -958,7 +958,7 @@ let strip_prod_binders_rawterm n (ty,_) =
let rec strip_ty acc n ty =
if n=0 then (List.rev acc, (ty,None)) else
match ty with
- Rawterm.RProd(loc,na,Explicit,a,b) ->
+ Rawterm.GProd(loc,na,Explicit,a,b) ->
strip_ty (([dummy_loc,na],(a,None))::acc) (n-1) b
| _ -> error "Cannot translate fix tactic: not enough products" in
strip_ty [] n ty
@@ -995,13 +995,13 @@ and pr_raw_tactic_level env n (t:raw_tactic_expr) =
let pr_and_constr_expr pr (c,_) = pr c
let pr_pat_and_constr_expr b (c,_) =
- pr_and_constr_expr ((if b then pr_lrawconstr_env else pr_rawconstr_env)
+ pr_and_constr_expr ((if b then pr_lglob_constr_env else pr_glob_constr_env)
(Global.env())) c
let rec glob_printers =
(pr_glob_tactic_level,
- (fun env -> pr_and_constr_expr (pr_rawconstr_env env)),
- (fun env -> pr_and_constr_expr (pr_lrawconstr_env env)),
+ (fun env -> pr_and_constr_expr (pr_glob_constr_env env)),
+ (fun env -> pr_and_constr_expr (pr_lglob_constr_env env)),
pr_pat_and_constr_expr,
(fun env -> pr_or_var (pr_and_short_name (pr_evaluable_reference_env env))),
(fun env -> pr_or_var (pr_inductive env)),
diff --git a/parsing/pptactic.mli b/parsing/pptactic.mli
index 71fe03e70..21de95ba0 100644
--- a/parsing/pptactic.mli
+++ b/parsing/pptactic.mli
@@ -30,8 +30,8 @@ type 'a raw_extra_genarg_printer =
'a -> std_ppcmds
type 'a glob_extra_genarg_printer =
- (rawconstr_and_expr -> std_ppcmds) ->
- (rawconstr_and_expr -> std_ppcmds) ->
+ (glob_constr_and_expr -> std_ppcmds) ->
+ (glob_constr_and_expr -> std_ppcmds) ->
(tolerability -> glob_tactic_expr -> std_ppcmds) ->
'a -> std_ppcmds
@@ -70,9 +70,9 @@ val pr_raw_extend:
string -> raw_generic_argument list -> std_ppcmds
val pr_glob_extend:
- (rawconstr_and_expr -> std_ppcmds) -> (rawconstr_and_expr -> std_ppcmds) ->
+ (glob_constr_and_expr -> std_ppcmds) -> (glob_constr_and_expr -> std_ppcmds) ->
(tolerability -> glob_tactic_expr -> std_ppcmds) ->
- (rawconstr_pattern_and_expr -> std_ppcmds) -> int ->
+ (glob_constr_pattern_and_expr -> std_ppcmds) -> int ->
string -> glob_generic_argument list -> std_ppcmds
val pr_extend :
diff --git a/parsing/prettyp.ml b/parsing/prettyp.ml
index 4ea3b1591..3b979f519 100644
--- a/parsing/prettyp.ml
+++ b/parsing/prettyp.ml
@@ -460,10 +460,10 @@ let gallina_print_syntactic_def kn =
let sep = " := "
and qid = Nametab.shortest_qualid_of_syndef Idset.empty kn
and (vars,a) = Syntax_def.search_syntactic_definition kn in
- let c = Topconstr.rawconstr_of_aconstr dummy_loc a in
+ let c = Topconstr.glob_constr_of_aconstr dummy_loc a in
str "Notation " ++ pr_qualid qid ++
prlist_with_sep spc pr_id (List.map fst vars) ++ str sep ++
- Constrextern.without_symbols pr_lrawconstr c ++ fnl ()
+ Constrextern.without_symbols pr_lglob_constr c ++ fnl ()
let gallina_print_leaf_entry with_values ((sp,kn as oname),lobj) =
let sep = if with_values then " = " else " : "
diff --git a/parsing/printer.ml b/parsing/printer.ml
index c575fde52..a7fa08f50 100644
--- a/parsing/printer.ml
+++ b/parsing/printer.ml
@@ -86,15 +86,15 @@ let pr_ljudge_env env j =
let pr_ljudge j = pr_ljudge_env (Global.env()) j
-let pr_lrawconstr_env env c =
- pr_lconstr_expr (extern_rawconstr (Termops.vars_of_env env) c)
-let pr_rawconstr_env env c =
- pr_constr_expr (extern_rawconstr (Termops.vars_of_env env) c)
+let pr_lglob_constr_env env c =
+ pr_lconstr_expr (extern_glob_constr (Termops.vars_of_env env) c)
+let pr_glob_constr_env env c =
+ pr_constr_expr (extern_glob_constr (Termops.vars_of_env env) c)
-let pr_lrawconstr c =
- pr_lconstr_expr (extern_rawconstr Idset.empty c)
-let pr_rawconstr c =
- pr_constr_expr (extern_rawconstr Idset.empty c)
+let pr_lglob_constr c =
+ pr_lconstr_expr (extern_glob_constr Idset.empty c)
+let pr_glob_constr c =
+ pr_constr_expr (extern_glob_constr Idset.empty c)
let pr_cases_pattern t =
pr_cases_pattern_expr (extern_cases_pattern Idset.empty t)
@@ -128,7 +128,7 @@ let pr_evaluable_reference ref =
pr_global (Tacred.global_of_evaluable_reference ref)
(*let pr_rawterm t =
- pr_lconstr (Constrextern.extern_rawconstr Idset.empty t)*)
+ pr_lconstr (Constrextern.extern_glob_constr Idset.empty t)*)
(*open Pattern
diff --git a/parsing/printer.mli b/parsing/printer.mli
index 11bebbd49..2da367816 100644
--- a/parsing/printer.mli
+++ b/parsing/printer.mli
@@ -54,11 +54,11 @@ val pr_type : types -> std_ppcmds
val pr_ljudge_env : env -> unsafe_judgment -> std_ppcmds * std_ppcmds
val pr_ljudge : unsafe_judgment -> std_ppcmds * std_ppcmds
-val pr_lrawconstr_env : env -> rawconstr -> std_ppcmds
-val pr_lrawconstr : rawconstr -> std_ppcmds
+val pr_lglob_constr_env : env -> glob_constr -> std_ppcmds
+val pr_lglob_constr : glob_constr -> std_ppcmds
-val pr_rawconstr_env : env -> rawconstr -> std_ppcmds
-val pr_rawconstr : rawconstr -> std_ppcmds
+val pr_glob_constr_env : env -> glob_constr -> std_ppcmds
+val pr_glob_constr : glob_constr -> std_ppcmds
val pr_lconstr_pattern_env : env -> constr_pattern -> std_ppcmds
val pr_lconstr_pattern : constr_pattern -> std_ppcmds
diff --git a/parsing/q_constr.ml4 b/parsing/q_constr.ml4
index 348905cad..7b0576a32 100644
--- a/parsing/q_constr.ml4
+++ b/parsing/q_constr.ml4
@@ -23,14 +23,14 @@ let dloc = <:expr< Util.dummy_loc >>
let apply_ref f l =
<:expr<
- Rawterm.RApp ($dloc$, Rawterm.RRef ($dloc$, Lazy.force $f$), $mlexpr_of_list (fun x -> x) l$)
+ Rawterm.GApp ($dloc$, Rawterm.GRef ($dloc$, Lazy.force $f$), $mlexpr_of_list (fun x -> x) l$)
>>
EXTEND
GLOBAL: expr;
expr:
[ [ "PATTERN"; "["; c = constr; "]" ->
- <:expr< snd (Pattern.pattern_of_rawconstr $c$) >> ] ]
+ <:expr< snd (Pattern.pattern_of_glob_constr $c$) >> ] ]
;
sort:
[ [ "Set" -> RProp Pos
@@ -49,19 +49,19 @@ EXTEND
constr:
[ "200" RIGHTA
[ LIDENT "forall"; id = ident; ":"; c1 = constr; ","; c2 = constr ->
- <:expr< Rawterm.RProd ($dloc$,Name $id$,Rawterm.Explicit,$c1$,$c2$) >>
+ <:expr< Rawterm.GProd ($dloc$,Name $id$,Rawterm.Explicit,$c1$,$c2$) >>
| "fun"; id = ident; ":"; c1 = constr; "=>"; c2 = constr ->
- <:expr< Rawterm.RLambda ($dloc$,Name $id$,Rawterm.Explicit,$c1$,$c2$) >>
+ <:expr< Rawterm.GLambda ($dloc$,Name $id$,Rawterm.Explicit,$c1$,$c2$) >>
| "let"; id = ident; ":="; c1 = constr; "in"; c2 = constr ->
<:expr< Rawterm.RLetin ($dloc$,Name $id$,$c1$,$c2$) >>
(* fix todo *)
]
| "100" RIGHTA
[ c1 = constr; ":"; c2 = SELF ->
- <:expr< Rawterm.RCast($dloc$,$c1$,DEFAULTcast,$c2$) >> ]
+ <:expr< Rawterm.GCast($dloc$,$c1$,DEFAULTcast,$c2$) >> ]
| "90" RIGHTA
[ c1 = constr; "->"; c2 = SELF ->
- <:expr< Rawterm.RProd ($dloc$,Anonymous,Rawterm.Explicit,$c1$,$c2$) >> ]
+ <:expr< Rawterm.GProd ($dloc$,Anonymous,Rawterm.Explicit,$c1$,$c2$) >> ]
| "75" RIGHTA
[ "~"; c = constr ->
apply_ref <:expr< coq_not_ref >> [c] ]
@@ -71,15 +71,15 @@ EXTEND
| "10" LEFTA
[ f = constr; args = LIST1 NEXT ->
let args = mlexpr_of_list (fun x -> x) args in
- <:expr< Rawterm.RApp ($dloc$,$f$,$args$) >> ]
+ <:expr< Rawterm.GApp ($dloc$,$f$,$args$) >> ]
| "0"
- [ s = sort -> <:expr< Rawterm.RSort ($dloc$,s) >>
- | id = ident -> <:expr< Rawterm.RVar ($dloc$,$id$) >>
- | "_" -> <:expr< Rawterm.RHole ($dloc$, QuestionMark (Define False)) >>
- | "?"; id = ident -> <:expr< Rawterm.RPatVar($dloc$,(False,$id$)) >>
+ [ s = sort -> <:expr< Rawterm.GSort ($dloc$,s) >>
+ | id = ident -> <:expr< Rawterm.GVar ($dloc$,$id$) >>
+ | "_" -> <:expr< Rawterm.GHole ($dloc$, QuestionMark (Define False)) >>
+ | "?"; id = ident -> <:expr< Rawterm.GPatVar($dloc$,(False,$id$)) >>
| "{"; c1 = constr; "}"; "+"; "{"; c2 = constr; "}" ->
apply_ref <:expr< coq_sumbool_ref >> [c1;c2]
- | "%"; e = string -> <:expr< Rawterm.RRef ($dloc$,Lazy.force $lid:e$) >>
+ | "%"; e = string -> <:expr< Rawterm.GRef ($dloc$,Lazy.force $lid:e$) >>
| c = match_constr -> c
| "("; c = constr LEVEL "200"; ")" -> c ] ]
;
@@ -87,7 +87,7 @@ EXTEND
[ [ "match"; c = constr LEVEL "100"; (ty,nal) = match_type;
"with"; OPT"|"; br = LIST0 eqn SEP "|"; "end" ->
let br = mlexpr_of_list (fun x -> x) br in
- <:expr< Rawterm.RCases ($dloc$,$ty$,[($c$,$nal$)],$br$) >>
+ <:expr< Rawterm.GCases ($dloc$,$ty$,[($c$,$nal$)],$br$) >>
] ]
;
match_type:
diff --git a/plugins/decl_mode/decl_expr.mli b/plugins/decl_mode/decl_expr.mli
index 3b98a1dca..6c6dbf0f6 100644
--- a/plugins/decl_mode/decl_expr.mli
+++ b/plugins/decl_mode/decl_expr.mli
@@ -83,8 +83,8 @@ type raw_proof_instr =
raw_tactic_expr) gen_proof_instr
type glob_proof_instr =
- ((identifier*(Genarg.rawconstr_and_expr option)) located,
- Genarg.rawconstr_and_expr,
+ ((identifier*(Genarg.glob_constr_and_expr option)) located,
+ Genarg.glob_constr_and_expr,
Topconstr.cases_pattern_expr,
Tacexpr.glob_tactic_expr) gen_proof_instr
diff --git a/plugins/decl_mode/decl_interp.ml b/plugins/decl_mode/decl_interp.ml
index f5bc47109..d16a26550 100644
--- a/plugins/decl_mode/decl_interp.ml
+++ b/plugins/decl_mode/decl_interp.ml
@@ -21,7 +21,7 @@ open Compat
(* INTERN *)
-let raw_app (loc,hd,args) = if args =[] then hd else RApp(loc,hd,args)
+let raw_app (loc,hd,args) = if args =[] then hd else GApp(loc,hd,args)
let intern_justification_items globs =
Option.map (List.map (intern_constr globs))
@@ -184,16 +184,16 @@ let interp_constr_or_thesis check_sort sigma env = function
let abstract_one_hyp inject h raw =
match h with
Hvar (loc,(id,None)) ->
- RProd (dummy_loc,Name id, Explicit, RHole (loc,Evd.BinderType (Name id)), raw)
+ GProd (dummy_loc,Name id, Explicit, GHole (loc,Evd.BinderType (Name id)), raw)
| Hvar (loc,(id,Some typ)) ->
- RProd (dummy_loc,Name id, Explicit, fst typ, raw)
+ GProd (dummy_loc,Name id, Explicit, fst typ, raw)
| Hprop st ->
- RProd (dummy_loc,st.st_label, Explicit, inject st.st_it, raw)
+ GProd (dummy_loc,st.st_label, Explicit, inject st.st_it, raw)
-let rawconstr_of_hyps inject hyps head =
+let glob_constr_of_hyps inject hyps head =
List.fold_right (abstract_one_hyp inject) hyps head
-let raw_prop = RSort (dummy_loc,RProp Null)
+let raw_prop = GSort (dummy_loc,RProp Null)
let rec match_hyps blend names constr = function
[] -> [],substl names constr
@@ -211,7 +211,7 @@ let rec match_hyps blend names constr = function
qhyp::rhyps,head
let interp_hyps_gen inject blend sigma env hyps head =
- let constr=understand sigma env (rawconstr_of_hyps inject hyps head) in
+ let constr=understand sigma env (glob_constr_of_hyps inject hyps head) in
match_hyps blend [] constr hyps
let interp_hyps sigma env hyps = fst (interp_hyps_gen fst (fun x _ -> x) sigma env hyps raw_prop)
@@ -236,32 +236,32 @@ let rec raw_of_pat =
function
PatVar (loc,Anonymous) -> anomaly "Anonymous pattern variable"
| PatVar (loc,Name id) ->
- RVar (loc,id)
+ GVar (loc,id)
| PatCstr(loc,((ind,_) as cstr),lpat,_) ->
let mind= fst (Global.lookup_inductive ind) in
let rec add_params n q =
if n<=0 then q else
- add_params (pred n) (RHole(dummy_loc,
+ add_params (pred n) (GHole(dummy_loc,
Evd.TomatchTypeParameter(ind,n))::q) in
let args = List.map raw_of_pat lpat in
- raw_app(loc,RRef(dummy_loc,Libnames.ConstructRef cstr),
+ raw_app(loc,GRef(dummy_loc,Libnames.ConstructRef cstr),
add_params mind.Declarations.mind_nparams args)
let prod_one_hyp = function
(loc,(id,None)) ->
(fun raw ->
- RProd (dummy_loc,Name id, Explicit,
- RHole (loc,Evd.BinderType (Name id)), raw))
+ GProd (dummy_loc,Name id, Explicit,
+ GHole (loc,Evd.BinderType (Name id)), raw))
| (loc,(id,Some typ)) ->
(fun raw ->
- RProd (dummy_loc,Name id, Explicit, fst typ, raw))
+ GProd (dummy_loc,Name id, Explicit, fst typ, raw))
let prod_one_id (loc,id) raw =
- RProd (dummy_loc,Name id, Explicit,
- RHole (loc,Evd.BinderType (Name id)), raw)
+ GProd (dummy_loc,Name id, Explicit,
+ GHole (loc,Evd.BinderType (Name id)), raw)
let let_in_one_alias (id,pat) raw =
- RLetIn (dummy_loc,Name id, raw_of_pat pat, raw)
+ GLetIn (dummy_loc,Name id, raw_of_pat pat, raw)
let rec bind_primary_aliases map pat =
match pat with
@@ -331,34 +331,34 @@ let interp_cases info sigma env params (pat:cases_pattern_expr) hyps =
(if expected = 0 then str "none" else int expected) ++ spc () ++
str "expected.") in
let app_ind =
- let rind = RRef (dummy_loc,Libnames.IndRef pinfo.per_ind) in
+ let rind = GRef (dummy_loc,Libnames.IndRef pinfo.per_ind) in
let rparams = List.map detype_ground pinfo.per_params in
let rparams_rec =
List.map
(fun (loc,(id,_)) ->
- RVar (loc,id)) params in
+ GVar (loc,id)) params in
let dum_args=
- list_tabulate (fun _ -> RHole (dummy_loc,Evd.QuestionMark (Evd.Define false)))
+ list_tabulate (fun _ -> GHole (dummy_loc,Evd.QuestionMark (Evd.Define false)))
oib.Declarations.mind_nrealargs in
raw_app(dummy_loc,rind,rparams@rparams_rec@dum_args) in
let pat_vars,aliases,patt = interp_pattern env pat in
let inject = function
- Thesis (Plain) -> Rawterm.RSort(dummy_loc,RProp Null)
+ Thesis (Plain) -> Rawterm.GSort(dummy_loc,RProp Null)
| Thesis (For rec_occ) ->
if not (List.mem rec_occ pat_vars) then
errorlabstrm "suppose it is"
(str "Variable " ++ Nameops.pr_id rec_occ ++
str " does not occur in pattern.");
- Rawterm.RSort(dummy_loc,RProp Null)
+ Rawterm.GSort(dummy_loc,RProp Null)
| This (c,_) -> c in
- let term1 = rawconstr_of_hyps inject hyps raw_prop in
+ let term1 = glob_constr_of_hyps inject hyps raw_prop in
let loc_ids,npatt =
let rids=ref ([],pat_vars) in
let npatt= deanonymize rids patt in
List.rev (fst !rids),npatt in
let term2 =
- RLetIn(dummy_loc,Anonymous,
- RCast(dummy_loc,raw_of_pat npatt,
+ GLetIn(dummy_loc,Anonymous,
+ GCast(dummy_loc,raw_of_pat npatt,
CastConv (DEFAULTcast,app_ind)),term1) in
let term3=List.fold_right let_in_one_alias aliases term2 in
let term4=List.fold_right prod_one_id loc_ids term3 in
@@ -413,17 +413,17 @@ let interp_casee sigma env = function
let abstract_one_arg = function
(loc,(id,None)) ->
(fun raw ->
- RLambda (dummy_loc,Name id, Explicit,
- RHole (loc,Evd.BinderType (Name id)), raw))
+ GLambda (dummy_loc,Name id, Explicit,
+ GHole (loc,Evd.BinderType (Name id)), raw))
| (loc,(id,Some typ)) ->
(fun raw ->
- RLambda (dummy_loc,Name id, Explicit, fst typ, raw))
+ GLambda (dummy_loc,Name id, Explicit, fst typ, raw))
-let rawconstr_of_fun args body =
+let glob_constr_of_fun args body =
List.fold_right abstract_one_arg args (fst body)
let interp_fun sigma env args body =
- let constr=understand sigma env (rawconstr_of_fun args body) in
+ let constr=understand sigma env (glob_constr_of_fun args body) in
match_args destLambda [] constr args
let rec interp_bare_proof_instr info (sigma:Evd.evar_map) (env:Environ.env) = function
diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml
index 0236c3095..97277ad58 100644
--- a/plugins/decl_mode/decl_proof_instr.ml
+++ b/plugins/decl_mode/decl_proof_instr.ml
@@ -1291,7 +1291,7 @@ let understand_my_constr c gls =
let env = pf_env gls in
let nc = names_of_rel_context env in
let rawc = Detyping.detype false [] nc c in
- let rec frob = function REvar _ -> RHole (dummy_loc,QuestionMark Expand) | rc -> map_rawconstr frob rc in
+ let rec frob = function GEvar _ -> GHole (dummy_loc,QuestionMark Expand) | rc -> map_glob_constr frob rc in
Pretyping.Default.understand_tcc (sig_sig gls) env ~expected_type:(pf_concl gls) (frob rawc)
let set_refine,my_refine =
diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml
index c79a8b818..24ec23484 100644
--- a/plugins/firstorder/instances.ml
+++ b/plugins/firstorder/instances.ml
@@ -123,9 +123,9 @@ let mk_open_instance id gl m t=
let rec raux n t=
if n=0 then t else
match t with
- RLambda(loc,name,k,_,t0)->
+ GLambda(loc,name,k,_,t0)->
let t1=raux (n-1) t0 in
- RLambda(loc,name,k,RHole (dummy_loc,Evd.BinderType name),t1)
+ GLambda(loc,name,k,GHole (dummy_loc,Evd.BinderType name),t1)
| _-> anomaly "can't happen" in
let ntt=try
Pretyping.Default.understand evmap env (raux m rawt)
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index 9c3982cb5..b2b4145d7 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -125,12 +125,12 @@ let functional_induction with_clean c princl pat =
Dumpglob.continue ();
res
-let rec abstract_rawconstr c = function
+let rec abstract_glob_constr c = function
| [] -> c
- | Topconstr.LocalRawDef (x,b)::bl -> Topconstr.mkLetInC(x,b,abstract_rawconstr c bl)
+ | Topconstr.LocalRawDef (x,b)::bl -> Topconstr.mkLetInC(x,b,abstract_glob_constr c bl)
| Topconstr.LocalRawAssum (idl,k,t)::bl ->
List.fold_right (fun x b -> Topconstr.mkLambdaC([x],k,t,b)) idl
- (abstract_rawconstr c bl)
+ (abstract_glob_constr c bl)
let interp_casted_constr_with_implicits sigma env impls c =
Constrintern.intern_gen false sigma env ~impls
@@ -161,7 +161,7 @@ let build_newrecursive
try
List.map
(fun (_,bl,_,def) ->
- let def = abstract_rawconstr def bl in
+ let def = abstract_glob_constr def bl in
interp_casted_constr_with_implicits
sigma rec_sign rec_impls def
)
@@ -188,15 +188,15 @@ let rec is_rec names =
let names = List.fold_right Idset.add names Idset.empty in
let check_id id names = Idset.mem id names in
let rec lookup names = function
- | RVar(_,id) -> check_id id names
- | RRef _ | REvar _ | RPatVar _ | RSort _ | RHole _ | RDynamic _ -> false
- | RCast(_,b,_) -> lookup names b
- | RRec _ -> error "RRec not handled"
- | RIf(_,b,_,lhs,rhs) ->
+ | GVar(_,id) -> check_id id names
+ | GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ | GDynamic _ -> false
+ | GCast(_,b,_) -> lookup names b
+ | GRec _ -> error "GRec not handled"
+ | GIf(_,b,_,lhs,rhs) ->
(lookup names b) || (lookup names lhs) || (lookup names rhs)
- | RLetIn(_,na,t,b) | RLambda(_,na,_,t,b) | RProd(_,na,_,t,b) ->
+ | GLetIn(_,na,t,b) | GLambda(_,na,_,t,b) | GProd(_,na,_,t,b) ->
lookup names t || lookup (Nameops.name_fold Idset.remove na names) b
- | RLetTuple(_,nal,_,t,b) -> lookup names t ||
+ | GLetTuple(_,nal,_,t,b) -> lookup names t ||
lookup
(List.fold_left
(fun acc na -> Nameops.name_fold Idset.remove na acc)
@@ -204,8 +204,8 @@ let rec is_rec names =
nal
)
b
- | RApp(_,f,args) -> List.exists (lookup names) (f::args)
- | RCases(_,_,_,el,brl) ->
+ | GApp(_,f,args) -> List.exists (lookup names) (f::args)
+ | GCases(_,_,_,el,brl) ->
List.exists (fun (e,_) -> lookup names e) el ||
List.exists (lookup_br names) brl
and lookup_br names (_,idl,_,rt) =
@@ -222,7 +222,7 @@ let rec local_binders_length = function
let prepare_body ((name,_,args,types,_),_) rt =
let n = local_binders_length args in
-(* Pp.msgnl (str "nb lambda to chop : " ++ str (string_of_int n) ++ fnl () ++Printer.pr_rawconstr rt); *)
+(* Pp.msgnl (str "nb lambda to chop : " ++ str (string_of_int n) ++ fnl () ++Printer.pr_glob_constr rt); *)
let fun_args,rt' = chop_rlambda_n n rt in
(fun_args,rt')
diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml
index 002fb7098..9db361cf5 100644
--- a/plugins/funind/indfun_common.ml
+++ b/plugins/funind/indfun_common.ml
@@ -76,8 +76,8 @@ let chop_rlambda_n =
then List.rev acc,rt
else
match rt with
- | Rawterm.RLambda(_,name,k,t,b) -> chop_lambda_n ((name,t,false)::acc) (n-1) b
- | Rawterm.RLetIn(_,name,v,b) -> chop_lambda_n ((name,v,true)::acc) (n-1) b
+ | Rawterm.GLambda(_,name,k,t,b) -> chop_lambda_n ((name,t,false)::acc) (n-1) b
+ | Rawterm.GLetIn(_,name,v,b) -> chop_lambda_n ((name,v,true)::acc) (n-1) b
| _ ->
raise (Util.UserError("chop_rlambda_n",
str "chop_rlambda_n: Not enough Lambdas"))
@@ -90,7 +90,7 @@ let chop_rprod_n =
then List.rev acc,rt
else
match rt with
- | Rawterm.RProd(_,name,k,t,b) -> chop_prod_n ((name,t)::acc) (n-1) b
+ | Rawterm.GProd(_,name,k,t,b) -> chop_prod_n ((name,t)::acc) (n-1) b
| _ -> raise (Util.UserError("chop_rprod_n",str "chop_rprod_n: Not enough products"))
in
chop_prod_n []
diff --git a/plugins/funind/indfun_common.mli b/plugins/funind/indfun_common.mli
index c48dff0c6..d802ecf2b 100644
--- a/plugins/funind/indfun_common.mli
+++ b/plugins/funind/indfun_common.mli
@@ -35,11 +35,11 @@ val list_union_eq :
val list_add_set_eq :
('a -> 'a -> bool) -> 'a -> 'a list -> 'a list
-val chop_rlambda_n : int -> Rawterm.rawconstr ->
- (name*Rawterm.rawconstr*bool) list * Rawterm.rawconstr
+val chop_rlambda_n : int -> Rawterm.glob_constr ->
+ (name*Rawterm.glob_constr*bool) list * Rawterm.glob_constr
-val chop_rprod_n : int -> Rawterm.rawconstr ->
- (name*Rawterm.rawconstr) list * Rawterm.rawconstr
+val chop_rprod_n : int -> Rawterm.glob_constr ->
+ (name*Rawterm.glob_constr) list * Rawterm.glob_constr
val def_of_const : Term.constr -> Term.constr
val eq : Term.constr Lazy.t
diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml
index 7c460f7d3..ed8cb9cb6 100644
--- a/plugins/funind/merge.ml
+++ b/plugins/funind/merge.ml
@@ -25,7 +25,7 @@ open Rawtermops
(** {1 Utilities} *)
-(** {2 Useful operations on constr and rawconstr} *)
+(** {2 Useful operations on constr and glob_constr} *)
let rec popn i c = if i<=0 then c else pop (popn (i-1) c)
@@ -60,7 +60,7 @@ let string_of_name nme = string_of_id (id_of_name nme)
(** [isVarf f x] returns [true] if term [x] is of the form [(Var f)]. *)
let isVarf f x =
match x with
- | RVar (_,x) -> Pervasives.compare x f = 0
+ | GVar (_,x) -> Pervasives.compare x f = 0
| _ -> false
(** [ident_global_exist id] returns true if identifier [id] is linked
@@ -97,7 +97,7 @@ let prNamedConstr s c =
let prNamedRConstr s c =
begin
msg(str "");
- msg(str(s^" {§ ") ++ Printer.pr_rawconstr c ++ str " §} ");
+ msg(str(s^" {§ ") ++ Printer.pr_glob_constr c ++ str " §} ");
msg(str "");
end
let prNamedLConstr_aux lc = List.iter (prNamedConstr "\n") lc
@@ -377,11 +377,11 @@ let verify_inds mib1 mib2 =
let build_raw_params prms_decl avoid =
let dummy_constr = compose_prod (List.map (fun (x,_,z) -> x,z) prms_decl) (mkRel 1) in
let _ = prNamedConstr "DUMMY" dummy_constr in
- let dummy_rawconstr = Detyping.detype false avoid [] dummy_constr in
- let _ = prNamedRConstr "RAWDUMMY" dummy_rawconstr in
- let res,_ = raw_decompose_prod dummy_rawconstr in
+ let dummy_glob_constr = Detyping.detype false avoid [] dummy_constr in
+ let _ = prNamedRConstr "RAWDUMMY" dummy_glob_constr in
+ let res,_ = glob_decompose_prod dummy_glob_constr in
let comblist = List.combine prms_decl res in
- comblist, res , (avoid @ (Idset.elements (ids_of_rawterm dummy_rawconstr)))
+ comblist, res , (avoid @ (Idset.elements (ids_of_rawterm dummy_glob_constr)))
*)
let ids_of_rawlist avoid rawl =
@@ -511,37 +511,37 @@ exception NoMerge
let rec merge_app c1 c2 id1 id2 shift filter_shift_stable =
let lnk = Array.append shift.lnk1 shift.lnk2 in
match c1 , c2 with
- | RApp(_,f1, arr1), RApp(_,f2,arr2) when isVarf id1 f1 && isVarf id2 f2 ->
+ | GApp(_,f1, arr1), GApp(_,f2,arr2) when isVarf id1 f1 && isVarf id2 f2 ->
let _ = prstr "\nICI1!\n";Pp.flush_all() in
let args = filter_shift_stable lnk (arr1 @ arr2) in
- RApp (dummy_loc,RVar (dummy_loc,shift.ident) , args)
- | RApp(_,f1, arr1), RApp(_,f2,arr2) -> raise NoMerge
- | RLetIn(_,nme,bdy,trm) , _ ->
+ GApp (dummy_loc,GVar (dummy_loc,shift.ident) , args)
+ | GApp(_,f1, arr1), GApp(_,f2,arr2) -> raise NoMerge
+ | GLetIn(_,nme,bdy,trm) , _ ->
let _ = prstr "\nICI2!\n";Pp.flush_all() in
let newtrm = merge_app trm c2 id1 id2 shift filter_shift_stable in
- RLetIn(dummy_loc,nme,bdy,newtrm)
- | _, RLetIn(_,nme,bdy,trm) ->
+ GLetIn(dummy_loc,nme,bdy,newtrm)
+ | _, GLetIn(_,nme,bdy,trm) ->
let _ = prstr "\nICI3!\n";Pp.flush_all() in
let newtrm = merge_app c1 trm id1 id2 shift filter_shift_stable in
- RLetIn(dummy_loc,nme,bdy,newtrm)
+ GLetIn(dummy_loc,nme,bdy,newtrm)
| _ -> let _ = prstr "\nICI4!\n";Pp.flush_all() in
raise NoMerge
let rec merge_app_unsafe c1 c2 shift filter_shift_stable =
let lnk = Array.append shift.lnk1 shift.lnk2 in
match c1 , c2 with
- | RApp(_,f1, arr1), RApp(_,f2,arr2) ->
+ | GApp(_,f1, arr1), GApp(_,f2,arr2) ->
let args = filter_shift_stable lnk (arr1 @ arr2) in
- RApp (dummy_loc,RVar(dummy_loc,shift.ident) , args)
+ GApp (dummy_loc,GVar(dummy_loc,shift.ident) , args)
(* FIXME: what if the function appears in the body of the let? *)
- | RLetIn(_,nme,bdy,trm) , _ ->
+ | GLetIn(_,nme,bdy,trm) , _ ->
let _ = prstr "\nICI2 '!\n";Pp.flush_all() in
let newtrm = merge_app_unsafe trm c2 shift filter_shift_stable in
- RLetIn(dummy_loc,nme,bdy,newtrm)
- | _, RLetIn(_,nme,bdy,trm) ->
+ GLetIn(dummy_loc,nme,bdy,newtrm)
+ | _, GLetIn(_,nme,bdy,trm) ->
let _ = prstr "\nICI3 '!\n";Pp.flush_all() in
let newtrm = merge_app_unsafe c1 trm shift filter_shift_stable in
- RLetIn(dummy_loc,nme,bdy,newtrm)
+ GLetIn(dummy_loc,nme,bdy,newtrm)
| _ -> let _ = prstr "\nICI4 '!\n";Pp.flush_all() in raise NoMerge
@@ -550,24 +550,24 @@ let rec merge_app_unsafe c1 c2 shift filter_shift_stable =
calls of branch 1 with all rec calls of branch 2. *)
(* TODO: reecrire cette heuristique (jusqu'a merge_types) *)
let rec merge_rec_hyps shift accrec
- (ltyp:(Names.name * rawconstr option * rawconstr option) list)
- filter_shift_stable : (Names.name * rawconstr option * rawconstr option) list =
+ (ltyp:(Names.name * glob_constr option * glob_constr option) list)
+ filter_shift_stable : (Names.name * glob_constr option * glob_constr option) list =
let mergeonehyp t reldecl =
match reldecl with
- | (nme,x,Some (RApp(_,i,args) as ind))
+ | (nme,x,Some (GApp(_,i,args) as ind))
-> nme,x, Some (merge_app_unsafe ind t shift filter_shift_stable)
| (nme,Some _,None) -> error "letins with recursive calls not treated yet"
| (nme,None,Some _) -> assert false
| (nme,None,None) | (nme,Some _,Some _) -> assert false in
match ltyp with
| [] -> []
- | (nme,None,Some (RApp(_,f, largs) as t)) :: lt when isVarf ind2name f ->
+ | (nme,None,Some (GApp(_,f, largs) as t)) :: lt when isVarf ind2name f ->
let rechyps = List.map (mergeonehyp t) accrec in
rechyps @ merge_rec_hyps shift accrec lt filter_shift_stable
| e::lt -> e :: merge_rec_hyps shift accrec lt filter_shift_stable
-let rec build_suppl_reccall (accrec:(name * rawconstr) list) concl2 shift =
+let rec build_suppl_reccall (accrec:(name * glob_constr) list) concl2 shift =
List.map (fun (nm,tp) -> (nm,merge_app_unsafe tp concl2 shift)) accrec
@@ -577,7 +577,7 @@ let find_app (nme:identifier) ltyp =
(List.map
(fun x ->
match x with
- | _,None,Some (RApp(_,f,_)) when isVarf nme f -> raise (Found 0)
+ | _,None,Some (GApp(_,f,_)) when isVarf nme f -> raise (Found 0)
| _ -> ())
ltyp);
false
@@ -591,9 +591,9 @@ let prnt_prod_or_letin nm letbdy typ =
let rec merge_types shift accrec1
- (ltyp1:(name * rawconstr option * rawconstr option) list)
- (concl1:rawconstr) (ltyp2:(name * rawconstr option * rawconstr option) list) concl2
- : (name * rawconstr option * rawconstr option) list * rawconstr =
+ (ltyp1:(name * glob_constr option * glob_constr option) list)
+ (concl1:glob_constr) (ltyp2:(name * glob_constr option * glob_constr option) list) concl2
+ : (name * glob_constr option * glob_constr option) list * glob_constr =
let _ = prstr "MERGE_TYPES\n" in
let _ = prstr "ltyp 1 : " in
let _ = List.iter (fun (nm,lbdy,tp) -> prnt_prod_or_letin nm lbdy tp) ltyp1 in
@@ -637,7 +637,7 @@ let rec merge_types shift accrec1
rechyps , concl
| (nme,None, Some t1)as e ::lt1 ->
(match t1 with
- | RApp(_,f,carr) when isVarf ind1name f ->
+ | GApp(_,f,carr) when isVarf ind1name f ->
merge_types shift (e::accrec1) lt1 concl1 ltyp2 concl2
| _ ->
let recres, recconcl2 =
@@ -704,8 +704,8 @@ let build_link_map allargs1 allargs2 lnk =
Precond: vars sets of [typcstr1] and [typcstr2] must be disjoint.
TODO: return nothing if equalities (after linking) are contradictory. *)
-let merge_one_constructor (shift:merge_infos) (typcstr1:rawconstr)
- (typcstr2:rawconstr) : rawconstr =
+let merge_one_constructor (shift:merge_infos) (typcstr1:glob_constr)
+ (typcstr2:glob_constr) : glob_constr =
(* FIXME: les noms des parametres corerspondent en principe au
parametres du niveau mib, mais il faudrait s'en assurer *)
(* shift.nfunresprmsx last args are functional result *)
@@ -713,17 +713,17 @@ let merge_one_constructor (shift:merge_infos) (typcstr1:rawconstr)
shift.mib1.mind_nparams + shift.oib1.mind_nrealargs - shift.nfunresprms1 in
let nargs2 =
shift.mib2.mind_nparams + shift.oib2.mind_nrealargs - shift.nfunresprms2 in
- let allargs1,rest1 = raw_decompose_prod_or_letin_n nargs1 typcstr1 in
- let allargs2,rest2 = raw_decompose_prod_or_letin_n nargs2 typcstr2 in
+ let allargs1,rest1 = glob_decompose_prod_or_letin_n nargs1 typcstr1 in
+ let allargs2,rest2 = glob_decompose_prod_or_letin_n nargs2 typcstr2 in
(* Build map of linked args of [typcstr2], and apply it to [typcstr2]. *)
let linked_map = build_link_map allargs1 allargs2 shift.lnk2 in
let rest2 = change_vars linked_map rest2 in
- let hyps1,concl1 = raw_decompose_prod_or_letin rest1 in
- let hyps2,concl2' = raw_decompose_prod_or_letin rest2 in
+ let hyps1,concl1 = glob_decompose_prod_or_letin rest1 in
+ let hyps2,concl2' = glob_decompose_prod_or_letin rest2 in
let ltyp,concl2 =
merge_types shift [] (List.rev hyps1) concl1 (List.rev hyps2) concl2' in
let _ = prNamedRLDecl "ltyp result:" ltyp in
- let typ = raw_compose_prod_or_letin concl2 (List.rev ltyp) in
+ let typ = glob_compose_prod_or_letin concl2 (List.rev ltyp) in
let revargs1 =
list_filteri (fun i _ -> isArg_stable shift.lnk1.(i)) (List.rev allargs1) in
let _ = prNamedRLDecl "ltyp allargs1" allargs1 in
@@ -733,7 +733,7 @@ let merge_one_constructor (shift:merge_infos) (typcstr1:rawconstr)
let _ = prNamedRLDecl "ltyp allargs2" allargs2 in
let _ = prNamedRLDecl "ltyp revargs2" revargs2 in
let typwithprms =
- raw_compose_prod_or_letin typ (List.rev revargs2 @ List.rev revargs1) in
+ glob_compose_prod_or_letin typ (List.rev revargs2 @ List.rev revargs1) in
typwithprms
@@ -759,8 +759,8 @@ let merge_constructor_id id1 id2 shift:identifier =
constructor [(name*type)]. These are translated to rawterms
first, each of them having distinct var names. *)
let rec merge_constructors (shift:merge_infos) (avoid:Idset.t)
- (typcstr1:(identifier * rawconstr) list)
- (typcstr2:(identifier * rawconstr) list) : (identifier * rawconstr) list =
+ (typcstr1:(identifier * glob_constr) list)
+ (typcstr2:(identifier * glob_constr) list) : (identifier * glob_constr) list =
List.flatten
(List.map
(fun (id1,rawtyp1) ->
@@ -778,12 +778,12 @@ let rec merge_constructors (shift:merge_infos) (avoid:Idset.t)
info in [shift], avoiding identifiers in [avoid]. *)
let rec merge_inductive_body (shift:merge_infos) avoid (oib1:one_inductive_body)
(oib2:one_inductive_body) =
- (* building rawconstr type of constructors *)
+ (* building glob_constr type of constructors *)
let mkrawcor nme avoid typ =
(* first replace rel 1 by a varname *)
let substindtyp = substitterm 0 (mkRel 1) (mkVar nme) typ in
Detyping.detype false (Idset.elements avoid) [] substindtyp in
- let lcstr1: rawconstr list =
+ let lcstr1: glob_constr list =
Array.to_list (Array.map (mkrawcor ind1name avoid) oib1.mind_user_lc) in
(* add to avoid all indentifiers of lcstr1 *)
let avoid2 = Idset.union avoid (ids_of_rawlist avoid lcstr1) in
@@ -792,10 +792,10 @@ let rec merge_inductive_body (shift:merge_infos) avoid (oib1:one_inductive_body)
let avoid3 = Idset.union avoid (ids_of_rawlist avoid lcstr2) in
let params1 =
- try fst (raw_decompose_prod_n shift.nrecprms1 (List.hd lcstr1))
+ try fst (glob_decompose_prod_n shift.nrecprms1 (List.hd lcstr1))
with _ -> [] in
let params2 =
- try fst (raw_decompose_prod_n shift.nrecprms2 (List.hd lcstr2))
+ try fst (glob_decompose_prod_n shift.nrecprms2 (List.hd lcstr2))
with _ -> [] in
let lcstr1 = List.combine (Array.to_list oib1.mind_consnames) lcstr1 in
@@ -816,8 +816,8 @@ let rec merge_mutual_inductive_body
merge_inductive_body shift Idset.empty mib1.mind_packets.(0) mib2.mind_packets.(0)
-let rawterm_to_constr_expr x = (* build a constr_expr from a rawconstr *)
- Flags.with_option Flags.raw_print (Constrextern.extern_rawtype Idset.empty) x
+let rawterm_to_constr_expr x = (* build a constr_expr from a glob_constr *)
+ Flags.with_option Flags.raw_print (Constrextern.extern_glob_type Idset.empty) x
let merge_rec_params_and_arity prms1 prms2 shift (concl:constr) =
let params = prms2 @ prms1 in
@@ -849,7 +849,7 @@ let merge_rec_params_and_arity prms1 prms2 shift (concl:constr) =
[rawlist], named ident.
FIXME: params et cstr_expr (arity) *)
let rawterm_list_to_inductive_expr prms1 prms2 mib1 mib2 shift
- (rawlist:(identifier * rawconstr) list) =
+ (rawlist:(identifier * glob_constr) list) =
let lident = dummy_loc, shift.ident in
let bindlist , cstr_expr = (* params , arities *)
merge_rec_params_and_arity prms1 prms2 shift mkSet in
@@ -861,21 +861,21 @@ let rawterm_list_to_inductive_expr prms1 prms2 mib1 mib2 shift
-let mkProd_reldecl (rdecl:rel_declaration) (t2:rawconstr) =
+let mkProd_reldecl (rdecl:rel_declaration) (t2:glob_constr) =
match rdecl with
| (nme,None,t) ->
let traw = Detyping.detype false [] [] t in
- RProd (dummy_loc,nme,Explicit,traw,t2)
+ GProd (dummy_loc,nme,Explicit,traw,t2)
| (_,Some _,_) -> assert false
-let mkProd_reldecl (rdecl:rel_declaration) (t2:rawconstr) =
+let mkProd_reldecl (rdecl:rel_declaration) (t2:glob_constr) =
match rdecl with
| (nme,None,t) ->
let traw = Detyping.detype false [] [] t in
- RProd (dummy_loc,nme,Explicit,traw,t2)
+ GProd (dummy_loc,nme,Explicit,traw,t2)
| (_,Some _,_) -> assert false
diff --git a/plugins/funind/rawterm_to_relation.ml b/plugins/funind/rawterm_to_relation.ml
index f8da96bdf..7b67e20f3 100644
--- a/plugins/funind/rawterm_to_relation.ml
+++ b/plugins/funind/rawterm_to_relation.ml
@@ -23,14 +23,14 @@ type binder_type =
| Prod of name
| LetIn of name
-type raw_context = (binder_type*rawconstr) list
+type glob_context = (binder_type*glob_constr) list
(*
- compose_raw_context [(bt_1,n_1,t_1);......] rt returns
+ compose_glob_context [(bt_1,n_1,t_1);......] rt returns
b_1(n_1,t_1,.....,bn(n_k,t_k,rt)) where the b_i's are the
binders corresponding to the bt_i's
*)
-let compose_raw_context =
+let compose_glob_context =
let compose_binder (bt,t) acc =
match bt with
| Lambda n -> mkRLambda(n,t,acc)
@@ -47,7 +47,7 @@ let compose_raw_context =
type 'a build_entry_pre_return =
{
- context : raw_context; (* the binding context of the result *)
+ context : glob_context; (* the binding context of the result *)
value : 'a; (* The value *)
}
@@ -159,7 +159,7 @@ let apply_args ctxt body args =
| _,[] -> (* No more args *)
(ctxt,body)
| [],_ -> (* no more fun *)
- let f,args' = raw_decompose_app body in
+ let f,args' = glob_decompose_app body in
(ctxt,mkRApp(f,args'@args))
| (Lambda Anonymous,t)::ctxt',arg::args' ->
do_apply avoid ctxt' body args'
@@ -215,8 +215,8 @@ let combine_app f args =
let combine_lam n t b =
{
context = [];
- value = mkRLambda(n, compose_raw_context t.context t.value,
- compose_raw_context b.context b.value )
+ value = mkRLambda(n, compose_glob_context t.context t.value,
+ compose_glob_context b.context b.value )
}
@@ -319,15 +319,15 @@ let build_constructors_of_type ind' argl =
let pat_as_term =
mkRApp(mkRRef (ConstructRef(ind',i+1)),argl)
in
- cases_pattern_of_rawconstr Anonymous pat_as_term
+ cases_pattern_of_glob_constr Anonymous pat_as_term
)
ind.Declarations.mind_consnames
(* [find_type_of] very naive attempts to discover the type of an if or a letin *)
let rec find_type_of nb b =
- let f,_ = raw_decompose_app b in
+ let f,_ = glob_decompose_app b in
match f with
- | RRef(_,ref) ->
+ | GRef(_,ref) ->
begin
let ind_type =
match ref with
@@ -350,8 +350,8 @@ let rec find_type_of nb b =
then raise (Invalid_argument "find_type_of : not a valid inductive");
ind_type
end
- | RCast(_,b,_) -> find_type_of nb b
- | RApp _ -> assert false (* we have decomposed any application via raw_decompose_app *)
+ | GCast(_,b,_) -> find_type_of nb b
+ | GApp _ -> assert false (* we have decomposed any application via glob_decompose_app *)
| _ -> raise (Invalid_argument "not a ref")
@@ -472,7 +472,7 @@ let rec pattern_to_term_and_type env typ = function
and concatenate them (informally, each branch of a match produces a new constructor)
\end{itemize}
- WARNING: The terms constructed here are only USING the rawconstr syntax but are highly bad formed.
+ WARNING: The terms constructed here are only USING the glob_constr syntax but are highly bad formed.
We must wait to have complete all the current calculi to set the recursive calls.
At this point, each term [f t1 ... tn] (where f appears in [funnames]) is replaced by
a pseudo term [forall res, res t1 ... tn, res]. A reconstruction phase is done later.
@@ -481,15 +481,15 @@ let rec pattern_to_term_and_type env typ = function
*)
-let rec build_entry_lc env funnames avoid rt : rawconstr build_entry_return =
- observe (str " Entering : " ++ Printer.pr_rawconstr rt);
+let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
+ observe (str " Entering : " ++ Printer.pr_glob_constr rt);
match rt with
- | RRef _ | RVar _ | REvar _ | RPatVar _ | RSort _ | RHole _ ->
+ | GRef _ | GVar _ | GEvar _ | GPatVar _ | GSort _ | GHole _ ->
(* do nothing (except changing type of course) *)
mk_result [] rt avoid
- | RApp(_,_,_) ->
- let f,args = raw_decompose_app rt in
- let args_res : (rawconstr list) build_entry_return =
+ | GApp(_,_,_) ->
+ let f,args = glob_decompose_app rt in
+ let args_res : (glob_constr list) build_entry_return =
List.fold_right (* create the arguments lists of constructors and combine them *)
(fun arg ctxt_argsl ->
let arg_res = build_entry_lc env funnames ctxt_argsl.to_avoid arg in
@@ -500,19 +500,19 @@ let rec build_entry_lc env funnames avoid rt : rawconstr build_entry_return =
in
begin
match f with
- | RLambda _ ->
+ | GLambda _ ->
let rec aux t l =
match l with
| [] -> t
| u::l ->
match t with
- | RLambda(loc,na,_,nat,b) ->
- RLetIn(dummy_loc,na,u,aux b l)
+ | GLambda(loc,na,_,nat,b) ->
+ GLetIn(dummy_loc,na,u,aux b l)
| _ ->
- RApp(dummy_loc,t,l)
+ GApp(dummy_loc,t,l)
in
build_entry_lc env funnames avoid (aux f args)
- | RVar(_,id) when Idset.mem id funnames ->
+ | GVar(_,id) when Idset.mem id funnames ->
(* if we have [f t1 ... tn] with [f]$\in$[fnames]
then we create a fresh variable [res],
add [res] and its "value" (i.e. [res v1 ... vn]) to each
@@ -538,7 +538,7 @@ let rec build_entry_lc env funnames avoid rt : rawconstr build_entry_return =
args_res.result
in
{ result = new_result; to_avoid = new_avoid }
- | RVar _ | REvar _ | RPatVar _ | RHole _ | RSort _ | RRef _ ->
+ | GVar _ | GEvar _ | GPatVar _ | GHole _ | GSort _ | GRef _ ->
(* if have [g t1 ... tn] with [g] not appearing in [funnames]
then
foreach [ctxt,v1 ... vn] in [args_res] we return
@@ -552,8 +552,8 @@ let rec build_entry_lc env funnames avoid rt : rawconstr build_entry_return =
{args_res with value = mkRApp(f,args_res.value)})
args_res.result
}
- | RApp _ -> assert false (* we have collected all the app in [raw_decompose_app] *)
- | RLetIn(_,n,t,b) ->
+ | GApp _ -> assert false (* we have collected all the app in [glob_decompose_app] *)
+ | GLetIn(_,n,t,b) ->
(* if we have [(let x := v in b) t1 ... tn] ,
we discard our work and compute the list of constructor for
[let x = v in (b t1 ... tn)] up to alpha conversion
@@ -567,7 +567,7 @@ let rec build_entry_lc env funnames avoid rt : rawconstr build_entry_return =
let new_b =
replace_var_by_term
id
- (RVar(dummy_loc,id))
+ (GVar(dummy_loc,id))
b
in
(Name new_id,new_b,new_avoid)
@@ -578,26 +578,26 @@ let rec build_entry_lc env funnames avoid rt : rawconstr build_entry_return =
funnames
avoid
(mkRLetIn(new_n,t,mkRApp(new_b,args)))
- | RCases _ | RIf _ | RLetTuple _ ->
+ | GCases _ | GIf _ | GLetTuple _ ->
(* we have [(match e1, ...., en with ..... end) t1 tn]
we first compute the result from the case and
then combine each of them with each of args one
*)
let f_res = build_entry_lc env funnames args_res.to_avoid f in
combine_results combine_app f_res args_res
- | RDynamic _ ->error "Not handled RDynamic"
- | RCast(_,b,_) ->
+ | GDynamic _ ->error "Not handled GDynamic"
+ | GCast(_,b,_) ->
(* for an applied cast we just trash the cast part
and restart the work.
WARNING: We need to restart since [b] itself should be an application term
*)
build_entry_lc env funnames avoid (mkRApp(b,args))
- | RRec _ -> error "Not handled RRec"
- | RProd _ -> error "Cannot apply a type"
+ | GRec _ -> error "Not handled GRec"
+ | GProd _ -> error "Cannot apply a type"
end (* end of the application treatement *)
- | RLambda(_,n,_,t,b) ->
+ | GLambda(_,n,_,t,b) ->
(* we first compute the list of constructor
corresponding to the body of the function,
then the one corresponding to the type
@@ -612,7 +612,7 @@ let rec build_entry_lc env funnames avoid rt : rawconstr build_entry_return =
let new_env = raw_push_named (new_n,None,t) env in
let b_res = build_entry_lc new_env funnames avoid b in
combine_results (combine_lam new_n) t_res b_res
- | RProd(_,n,_,t,b) ->
+ | GProd(_,n,_,t,b) ->
(* we first compute the list of constructor
corresponding to the body of the function,
then the one corresponding to the type
@@ -622,7 +622,7 @@ let rec build_entry_lc env funnames avoid rt : rawconstr build_entry_return =
let new_env = raw_push_named (n,None,t) env in
let b_res = build_entry_lc new_env funnames avoid b in
combine_results (combine_prod n) t_res b_res
- | RLetIn(_,n,v,b) ->
+ | GLetIn(_,n,v,b) ->
(* we first compute the list of constructor
corresponding to the body of the function,
then the one corresponding to the value [t]
@@ -638,21 +638,21 @@ let rec build_entry_lc env funnames avoid rt : rawconstr build_entry_return =
in
let b_res = build_entry_lc new_env funnames avoid b in
combine_results (combine_letin n) v_res b_res
- | RCases(_,_,_,el,brl) ->
+ | GCases(_,_,_,el,brl) ->
(* we create the discrimination function
and treat the case itself
*)
let make_discr = make_discr_match brl in
build_entry_lc_from_case env funnames make_discr el brl avoid
- | RIf(_,b,(na,e_option),lhs,rhs) ->
+ | GIf(_,b,(na,e_option),lhs,rhs) ->
let b_as_constr = Pretyping.Default.understand Evd.empty env b in
let b_typ = Typing.type_of env Evd.empty b_as_constr in
let (ind,_) =
try Inductiveops.find_inductive env Evd.empty b_typ
with Not_found ->
errorlabstrm "" (str "Cannot find the inductive associated to " ++
- Printer.pr_rawconstr b ++ str " in " ++
- Printer.pr_rawconstr rt ++ str ". try again with a cast")
+ Printer.pr_glob_constr b ++ str " in " ++
+ Printer.pr_glob_constr rt ++ str ". try again with a cast")
in
let case_pats = build_constructors_of_type ind [] in
assert (Array.length case_pats = 2);
@@ -665,11 +665,11 @@ let rec build_entry_lc env funnames avoid rt : rawconstr build_entry_return =
let match_expr =
mkRCases(None,[(b,(Anonymous,None))],brl)
in
- (* Pp.msgnl (str "new case := " ++ Printer.pr_rawconstr match_expr); *)
+ (* Pp.msgnl (str "new case := " ++ Printer.pr_glob_constr match_expr); *)
build_entry_lc env funnames avoid match_expr
- | RLetTuple(_,nal,_,b,e) ->
+ | GLetTuple(_,nal,_,b,e) ->
begin
- let nal_as_rawconstr =
+ let nal_as_glob_constr =
List.map
(function
Name id -> mkRVar id
@@ -683,10 +683,10 @@ let rec build_entry_lc env funnames avoid rt : rawconstr build_entry_return =
try Inductiveops.find_inductive env Evd.empty b_typ
with Not_found ->
errorlabstrm "" (str "Cannot find the inductive associated to " ++
- Printer.pr_rawconstr b ++ str " in " ++
- Printer.pr_rawconstr rt ++ str ". try again with a cast")
+ Printer.pr_glob_constr b ++ str " in " ++
+ Printer.pr_glob_constr rt ++ str ". try again with a cast")
in
- let case_pats = build_constructors_of_type ind nal_as_rawconstr in
+ let case_pats = build_constructors_of_type ind nal_as_glob_constr in
assert (Array.length case_pats = 1);
let br =
(dummy_loc,[],[case_pats.(0)],e)
@@ -695,14 +695,14 @@ let rec build_entry_lc env funnames avoid rt : rawconstr build_entry_return =
build_entry_lc env funnames avoid match_expr
end
- | RRec _ -> error "Not handled RRec"
- | RCast(_,b,_) ->
+ | GRec _ -> error "Not handled GRec"
+ | GCast(_,b,_) ->
build_entry_lc env funnames avoid b
- | RDynamic _ -> error "Not handled RDynamic"
+ | GDynamic _ -> error "Not handled GDynamic"
and build_entry_lc_from_case env funname make_discr
(el:tomatch_tuples)
(brl:Rawterm.cases_clauses) avoid :
- rawconstr build_entry_return =
+ glob_constr build_entry_return =
match el with
| [] -> assert false (* this case correspond to match <nothing> with .... !*)
| el ->
@@ -762,7 +762,7 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve
(will be used in the following recursive calls)
*)
let new_env = List.fold_right2 add_pat_variables patl types env in
- let not_those_patterns : (identifier list -> rawconstr -> rawconstr) list =
+ let not_those_patterns : (identifier list -> glob_constr -> glob_constr) list =
List.map2
(fun pat typ ->
fun avoid pat'_as_term ->
@@ -780,7 +780,7 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve
in
mkRProd (Name id,raw_typ_of_id,acc))
pat_ids
- (raw_make_neq pat'_as_term (pattern_to_term renamed_pat))
+ (glob_make_neq pat'_as_term (pattern_to_term renamed_pat))
)
patl
types
@@ -835,7 +835,7 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve
else acc
)
idl
- [(Prod Anonymous,raw_make_eq ~typ pat_as_term e)]
+ [(Prod Anonymous,glob_make_eq ~typ pat_as_term e)]
)
patl
matched_expr.value
@@ -883,18 +883,18 @@ exception Continue
eliminates some meaningless equalities, applies some rewrites......
*)
let rec rebuild_cons env nb_args relname args crossed_types depth rt =
- observe (str "rebuilding : " ++ pr_rawconstr rt);
+ observe (str "rebuilding : " ++ pr_glob_constr rt);
match rt with
- | RProd(_,n,k,t,b) ->
+ | GProd(_,n,k,t,b) ->
let not_free_in_t id = not (is_free_in id t) in
let new_crossed_types = t::crossed_types in
begin
match t with
- | RApp(_,(RVar(_,res_id) as res_rt),args') when is_res res_id ->
+ | GApp(_,(GVar(_,res_id) as res_rt),args') when is_res res_id ->
begin
match args' with
- | (RVar(_,this_relname))::args' ->
+ | (GVar(_,this_relname))::args' ->
(*i The next call to mk_rel_id is
valid since we are constructing the graph
Ensures by: obvious
@@ -916,12 +916,12 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
| _ -> (* the first args is the name of the function! *)
assert false
end
- | RApp(loc1,RRef(loc2,eq_as_ref),[ty;RVar(loc3,id);rt])
+ | GApp(loc1,GRef(loc2,eq_as_ref),[ty;GVar(loc3,id);rt])
when eq_as_ref = Lazy.force Coqlib.coq_eq_ref && n = Anonymous
->
begin
try
- observe (str "computing new type for eq : " ++ pr_rawconstr rt);
+ observe (str "computing new type for eq : " ++ pr_glob_constr rt);
let t' =
try Pretyping.Default.understand Evd.empty env t with _ -> raise Continue
in
@@ -953,8 +953,8 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
((Util.list_chop nparam args'))
in
let rt_typ =
- RApp(Util.dummy_loc,
- RRef (Util.dummy_loc,Libnames.IndRef ind),
+ GApp(Util.dummy_loc,
+ GRef (Util.dummy_loc,Libnames.IndRef ind),
(List.map
(fun p -> Detyping.detype false []
(Termops.names_of_rel_context env)
@@ -964,9 +964,9 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
(mkRHole ()))))
in
let eq' =
- RApp(loc1,RRef(loc2,jmeq),[ty;RVar(loc3,id);rt_typ;rt])
+ GApp(loc1,GRef(loc2,jmeq),[ty;GVar(loc3,id);rt_typ;rt])
in
- observe (str "computing new type for jmeq : " ++ pr_rawconstr eq');
+ observe (str "computing new type for jmeq : " ++ pr_glob_constr eq');
let eq'_as_constr = Pretyping.Default.understand Evd.empty env eq' in
observe (str " computing new type for jmeq : done") ;
let new_args =
@@ -1033,7 +1033,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
else new_b, Idset.add id id_to_exclude
*)
| _ ->
- observe (str "computing new type for prod : " ++ pr_rawconstr rt);
+ observe (str "computing new type for prod : " ++ pr_glob_constr rt);
let t' = Pretyping.Default.understand Evd.empty env t in
let new_env = Environ.push_rel (n,None,t') env in
let new_b,id_to_exclude =
@@ -1048,11 +1048,11 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
(Idset.filter not_free_in_t id_to_exclude)
| _ -> mkRProd(n,t,new_b),Idset.filter not_free_in_t id_to_exclude
end
- | RLambda(_,n,k,t,b) ->
+ | GLambda(_,n,k,t,b) ->
begin
let not_free_in_t id = not (is_free_in id t) in
let new_crossed_types = t :: crossed_types in
- observe (str "computing new type for lambda : " ++ pr_rawconstr rt);
+ observe (str "computing new type for lambda : " ++ pr_glob_constr rt);
let t' = Pretyping.Default.understand Evd.empty env t in
match n with
| Name id ->
@@ -1067,12 +1067,12 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
then
new_b, Idset.remove id (Idset.filter not_free_in_t id_to_exclude)
else
- RProd(dummy_loc,n,k,t,new_b),Idset.filter not_free_in_t id_to_exclude
+ GProd(dummy_loc,n,k,t,new_b),Idset.filter not_free_in_t id_to_exclude
| _ -> anomaly "Should not have an anonymous function here"
(* We have renamed all the anonymous functions during alpha_renaming phase *)
end
- | RLetIn(_,n,t,b) ->
+ | GLetIn(_,n,t,b) ->
begin
let not_free_in_t id = not (is_free_in id t) in
let t' = Pretyping.Default.understand Evd.empty env t in
@@ -1086,10 +1086,10 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
match n with
| Name id when Idset.mem id id_to_exclude && depth >= nb_args ->
new_b,Idset.remove id (Idset.filter not_free_in_t id_to_exclude)
- | _ -> RLetIn(dummy_loc,n,t,new_b),
+ | _ -> GLetIn(dummy_loc,n,t,new_b),
Idset.filter not_free_in_t id_to_exclude
end
- | RLetTuple(_,nal,(na,rto),t,b) ->
+ | GLetTuple(_,nal,(na,rto),t,b) ->
assert (rto=None);
begin
let not_free_in_t id = not (is_free_in id t) in
@@ -1112,7 +1112,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
(* | Name id when Idset.mem id id_to_exclude -> *)
(* new_b,Idset.remove id (Idset.filter not_free_in_t id_to_exclude) *)
(* | _ -> *)
- RLetTuple(dummy_loc,nal,(na,None),t,new_b),
+ GLetTuple(dummy_loc,nal,(na,None),t,new_b),
Idset.filter not_free_in_t (Idset.union id_to_exclude id_to_exclude')
end
@@ -1122,12 +1122,12 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
(* debuging wrapper *)
let rebuild_cons env nb_args relname args crossed_types rt =
-(* observennl (str "rebuild_cons : rt := "++ pr_rawconstr rt ++ *)
+(* observennl (str "rebuild_cons : rt := "++ pr_glob_constr rt ++ *)
(* str "nb_args := " ++ str (string_of_int nb_args)); *)
let res =
rebuild_cons env nb_args relname args crossed_types 0 rt
in
-(* observe (str " leads to "++ pr_rawconstr (fst res)); *)
+(* observe (str " leads to "++ pr_glob_constr (fst res)); *)
res
@@ -1139,30 +1139,30 @@ let rebuild_cons env nb_args relname args crossed_types rt =
TODO: Find a valid way to deal with implicit arguments here!
*)
let rec compute_cst_params relnames params = function
- | RRef _ | RVar _ | REvar _ | RPatVar _ -> params
- | RApp(_,RVar(_,relname'),rtl) when Idset.mem relname' relnames ->
+ | GRef _ | GVar _ | GEvar _ | GPatVar _ -> params
+ | GApp(_,GVar(_,relname'),rtl) when Idset.mem relname' relnames ->
compute_cst_params_from_app [] (params,rtl)
- | RApp(_,f,args) ->
+ | GApp(_,f,args) ->
List.fold_left (compute_cst_params relnames) params (f::args)
- | RLambda(_,_,_,t,b) | RProd(_,_,_,t,b) | RLetIn(_,_,t,b) | RLetTuple(_,_,_,t,b) ->
+ | GLambda(_,_,_,t,b) | GProd(_,_,_,t,b) | GLetIn(_,_,t,b) | GLetTuple(_,_,_,t,b) ->
let t_params = compute_cst_params relnames params t in
compute_cst_params relnames t_params b
- | RCases _ ->
+ | GCases _ ->
params (* If there is still cases at this point they can only be
discriminitation ones *)
- | RSort _ -> params
- | RHole _ -> params
- | RIf _ | RRec _ | RCast _ | RDynamic _ ->
+ | GSort _ -> params
+ | GHole _ -> params
+ | GIf _ | GRec _ | GCast _ | GDynamic _ ->
raise (UserError("compute_cst_params", str "Not handled case"))
and compute_cst_params_from_app acc (params,rtl) =
match params,rtl with
| _::_,[] -> assert false (* the rel has at least nargs + 1 arguments ! *)
- | ((Name id,_,is_defined) as param)::params',(RVar(_,id'))::rtl'
+ | ((Name id,_,is_defined) as param)::params',(GVar(_,id'))::rtl'
when id_ord id id' == 0 && not is_defined ->
compute_cst_params_from_app (param::acc) (params',rtl')
| _ -> List.rev acc
-let compute_params_name relnames (args : (Names.name * Rawterm.rawconstr * bool) list array) csts =
+let compute_params_name relnames (args : (Names.name * Rawterm.glob_constr * bool) list array) csts =
let rels_params =
Array.mapi
(fun i args ->
@@ -1181,7 +1181,7 @@ let compute_params_name relnames (args : (Names.name * Rawterm.rawconstr * bool)
if array_for_all
(fun l ->
let (n',nt',is_defined') = List.nth l i in
- n = n' && Topconstr.eq_rawconstr nt nt' && is_defined = is_defined')
+ n = n' && Topconstr.eq_glob_constr nt nt' && is_defined = is_defined')
rels_params
then
l := param::!l
@@ -1204,11 +1204,11 @@ let rec rebuild_return_type rt =
let do_build_inductive
- funnames (funsargs: (Names.name * rawconstr * bool) list list)
+ funnames (funsargs: (Names.name * glob_constr * bool) list list)
returned_types
- (rtl:rawconstr list) =
+ (rtl:glob_constr list) =
let _time1 = System.get_time () in
-(* Pp.msgnl (prlist_with_sep fnl Printer.pr_rawconstr rtl); *)
+(* Pp.msgnl (prlist_with_sep fnl Printer.pr_glob_constr rtl); *)
let funnames_as_set = List.fold_right Idset.add funnames Idset.empty in
let funnames = Array.of_list funnames in
let funsargs = Array.of_list funsargs in
@@ -1233,19 +1233,19 @@ let do_build_inductive
let resa = Array.map (build_entry_lc env funnames_as_set []) rta in
let env_with_graphs =
let rel_arity i funargs = (* Reduilding arities (with parameters) *)
- let rel_first_args :(Names.name * Rawterm.rawconstr * bool ) list =
+ let rel_first_args :(Names.name * Rawterm.glob_constr * bool ) list =
funargs
in
List.fold_right
(fun (n,t,is_defined) acc ->
if is_defined
then
- Topconstr.CLetIn(dummy_loc,(dummy_loc, n),Constrextern.extern_rawconstr Idset.empty t,
+ Topconstr.CLetIn(dummy_loc,(dummy_loc, n),Constrextern.extern_glob_constr Idset.empty t,
acc)
else
Topconstr.CProdN
(dummy_loc,
- [[(dummy_loc,n)],Topconstr.default_binder_kind,Constrextern.extern_rawconstr Idset.empty t],
+ [[(dummy_loc,n)],Topconstr.default_binder_kind,Constrextern.extern_glob_constr Idset.empty t],
acc
)
)
@@ -1264,9 +1264,9 @@ let do_build_inductive
let constr i res =
List.map
(function result (* (args',concl') *) ->
- let rt = compose_raw_context result.context result.value in
+ let rt = compose_glob_context result.context result.value in
let nb_args = List.length funsargs.(i) in
- (* with_full_print (fun rt -> Pp.msgnl (str "glob constr " ++ pr_rawconstr rt)) rt; *)
+ (* with_full_print (fun rt -> Pp.msgnl (str "glob constr " ++ pr_glob_constr rt)) rt; *)
fst (
rebuild_cons env_with_graphs nb_args relnames.(i)
[]
@@ -1285,7 +1285,7 @@ let do_build_inductive
i*)
id_of_string ((string_of_id (mk_rel_id funnames.(i)))^"_"^(string_of_int !next_constructor_id))
in
- let rel_constructors i rt : (identifier*rawconstr) list =
+ let rel_constructors i rt : (identifier*glob_constr) list =
next_constructor_id := (-1);
List.map (fun constr -> (mk_constructor_id i),constr) (constr i rt)
in
@@ -1299,19 +1299,19 @@ let do_build_inductive
rel_constructors
in
let rel_arity i funargs = (* Reduilding arities (with parameters) *)
- let rel_first_args :(Names.name * Rawterm.rawconstr * bool ) list =
+ let rel_first_args :(Names.name * Rawterm.glob_constr * bool ) list =
(snd (list_chop nrel_params funargs))
in
List.fold_right
(fun (n,t,is_defined) acc ->
if is_defined
then
- Topconstr.CLetIn(dummy_loc,(dummy_loc, n),Constrextern.extern_rawconstr Idset.empty t,
+ Topconstr.CLetIn(dummy_loc,(dummy_loc, n),Constrextern.extern_glob_constr Idset.empty t,
acc)
else
Topconstr.CProdN
(dummy_loc,
- [[(dummy_loc,n)],Topconstr.default_binder_kind,Constrextern.extern_rawconstr Idset.empty t],
+ [[(dummy_loc,n)],Topconstr.default_binder_kind,Constrextern.extern_glob_constr Idset.empty t],
acc
)
)
@@ -1328,10 +1328,10 @@ let do_build_inductive
(fun (n,t,is_defined) ->
if is_defined
then
- Topconstr.LocalRawDef((dummy_loc,n), Constrextern.extern_rawconstr Idset.empty t)
+ Topconstr.LocalRawDef((dummy_loc,n), Constrextern.extern_glob_constr Idset.empty t)
else
Topconstr.LocalRawAssum
- ([(dummy_loc,n)], Topconstr.default_binder_kind, Constrextern.extern_rawconstr Idset.empty t)
+ ([(dummy_loc,n)], Topconstr.default_binder_kind, Constrextern.extern_glob_constr Idset.empty t)
)
rels_params
in
@@ -1341,7 +1341,7 @@ let do_build_inductive
false,((dummy_loc,id),
Flags.with_option
Flags.raw_print
- (Constrextern.extern_rawtype Idset.empty) ((* zeta_normalize *) t)
+ (Constrextern.extern_glob_type Idset.empty) ((* zeta_normalize *) t)
)
))
(rel_constructors)
diff --git a/plugins/funind/rawterm_to_relation.mli b/plugins/funind/rawterm_to_relation.mli
index a314050f7..772e422f8 100644
--- a/plugins/funind/rawterm_to_relation.mli
+++ b/plugins/funind/rawterm_to_relation.mli
@@ -9,8 +9,8 @@
val build_inductive :
Names.identifier list -> (* The list of function name *)
- (Names.name*Rawterm.rawconstr*bool) list list -> (* The list of function args *)
+ (Names.name*Rawterm.glob_constr*bool) list list -> (* The list of function args *)
Topconstr.constr_expr list -> (* The list of function returned type *)
- Rawterm.rawconstr list -> (* the list of body *)
+ Rawterm.glob_constr list -> (* the list of body *)
unit
diff --git a/plugins/funind/rawtermops.ml b/plugins/funind/rawtermops.ml
index e31f1452d..f372fb017 100644
--- a/plugins/funind/rawtermops.ml
+++ b/plugins/funind/rawtermops.ml
@@ -6,46 +6,46 @@ open Names
let idmap_is_empty m = m = Idmap.empty
(*
- Some basic functions to rebuild rawconstr
+ Some basic functions to rebuild glob_constr
In each of them the location is Util.dummy_loc
*)
-let mkRRef ref = RRef(dummy_loc,ref)
-let mkRVar id = RVar(dummy_loc,id)
-let mkRApp(rt,rtl) = RApp(dummy_loc,rt,rtl)
-let mkRLambda(n,t,b) = RLambda(dummy_loc,n,Explicit,t,b)
-let mkRProd(n,t,b) = RProd(dummy_loc,n,Explicit,t,b)
-let mkRLetIn(n,t,b) = RLetIn(dummy_loc,n,t,b)
-let mkRCases(rto,l,brl) = RCases(dummy_loc,Term.RegularStyle,rto,l,brl)
-let mkRSort s = RSort(dummy_loc,s)
-let mkRHole () = RHole(dummy_loc,Evd.BinderType Anonymous)
-let mkRCast(b,t) = RCast(dummy_loc,b,CastConv (Term.DEFAULTcast,t))
+let mkRRef ref = GRef(dummy_loc,ref)
+let mkRVar id = GVar(dummy_loc,id)
+let mkRApp(rt,rtl) = GApp(dummy_loc,rt,rtl)
+let mkRLambda(n,t,b) = GLambda(dummy_loc,n,Explicit,t,b)
+let mkRProd(n,t,b) = GProd(dummy_loc,n,Explicit,t,b)
+let mkRLetIn(n,t,b) = GLetIn(dummy_loc,n,t,b)
+let mkRCases(rto,l,brl) = GCases(dummy_loc,Term.RegularStyle,rto,l,brl)
+let mkRSort s = GSort(dummy_loc,s)
+let mkRHole () = GHole(dummy_loc,Evd.BinderType Anonymous)
+let mkRCast(b,t) = GCast(dummy_loc,b,CastConv (Term.DEFAULTcast,t))
(*
- Some basic functions to decompose rawconstrs
+ Some basic functions to decompose glob_constrs
These are analogous to the ones constrs
*)
-let raw_decompose_prod =
- let rec raw_decompose_prod args = function
- | RProd(_,n,k,t,b) ->
- raw_decompose_prod ((n,t)::args) b
+let glob_decompose_prod =
+ let rec glob_decompose_prod args = function
+ | GProd(_,n,k,t,b) ->
+ glob_decompose_prod ((n,t)::args) b
| rt -> args,rt
in
- raw_decompose_prod []
-
-let raw_decompose_prod_or_letin =
- let rec raw_decompose_prod args = function
- | RProd(_,n,k,t,b) ->
- raw_decompose_prod ((n,None,Some t)::args) b
- | RLetIn(_,n,t,b) ->
- raw_decompose_prod ((n,Some t,None)::args) b
+ glob_decompose_prod []
+
+let glob_decompose_prod_or_letin =
+ let rec glob_decompose_prod args = function
+ | GProd(_,n,k,t,b) ->
+ glob_decompose_prod ((n,None,Some t)::args) b
+ | GLetIn(_,n,t,b) ->
+ glob_decompose_prod ((n,Some t,None)::args) b
| rt -> args,rt
in
- raw_decompose_prod []
+ glob_decompose_prod []
-let raw_compose_prod =
+let glob_compose_prod =
List.fold_left (fun b (n,t) -> mkRProd(n,t,b))
-let raw_compose_prod_or_letin =
+let glob_compose_prod_or_letin =
List.fold_left (
fun concl decl ->
match decl with
@@ -53,37 +53,37 @@ let raw_compose_prod_or_letin =
| (n,Some bdy,None) -> mkRLetIn(n,bdy,concl)
| _ -> assert false)
-let raw_decompose_prod_n n =
- let rec raw_decompose_prod i args c =
+let glob_decompose_prod_n n =
+ let rec glob_decompose_prod i args c =
if i<=0 then args,c
else
match c with
- | RProd(_,n,_,t,b) ->
- raw_decompose_prod (i-1) ((n,t)::args) b
+ | GProd(_,n,_,t,b) ->
+ glob_decompose_prod (i-1) ((n,t)::args) b
| rt -> args,rt
in
- raw_decompose_prod n []
+ glob_decompose_prod n []
-let raw_decompose_prod_or_letin_n n =
- let rec raw_decompose_prod i args c =
+let glob_decompose_prod_or_letin_n n =
+ let rec glob_decompose_prod i args c =
if i<=0 then args,c
else
match c with
- | RProd(_,n,_,t,b) ->
- raw_decompose_prod (i-1) ((n,None,Some t)::args) b
- | RLetIn(_,n,t,b) ->
- raw_decompose_prod (i-1) ((n,Some t,None)::args) b
+ | GProd(_,n,_,t,b) ->
+ glob_decompose_prod (i-1) ((n,None,Some t)::args) b
+ | GLetIn(_,n,t,b) ->
+ glob_decompose_prod (i-1) ((n,Some t,None)::args) b
| rt -> args,rt
in
- raw_decompose_prod n []
+ glob_decompose_prod n []
-let raw_decompose_app =
+let glob_decompose_app =
let rec decompose_rapp acc rt =
-(* msgnl (str "raw_decompose_app on : "++ Printer.pr_rawconstr rt); *)
+(* msgnl (str "glob_decompose_app on : "++ Printer.pr_glob_constr rt); *)
match rt with
- | RApp(_,rt,rtl) ->
+ | GApp(_,rt,rtl) ->
decompose_rapp (List.fold_left (fun y x -> x::y) acc rtl) rt
| rt -> rt,List.rev acc
in
@@ -92,24 +92,24 @@ let raw_decompose_app =
-(* [raw_make_eq t1 t2] build the rawconstr corresponding to [t2 = t1] *)
-let raw_make_eq ?(typ= mkRHole ()) t1 t2 =
+(* [glob_make_eq t1 t2] build the glob_constr corresponding to [t2 = t1] *)
+let glob_make_eq ?(typ= mkRHole ()) t1 t2 =
mkRApp(mkRRef (Lazy.force Coqlib.coq_eq_ref),[typ;t2;t1])
-(* [raw_make_neq t1 t2] build the rawconstr corresponding to [t1 <> t2] *)
-let raw_make_neq t1 t2 =
- mkRApp(mkRRef (Lazy.force Coqlib.coq_not_ref),[raw_make_eq t1 t2])
+(* [glob_make_neq t1 t2] build the glob_constr corresponding to [t1 <> t2] *)
+let glob_make_neq t1 t2 =
+ mkRApp(mkRRef (Lazy.force Coqlib.coq_not_ref),[glob_make_eq t1 t2])
-(* [raw_make_or P1 P2] build the rawconstr corresponding to [P1 \/ P2] *)
-let raw_make_or t1 t2 = mkRApp (mkRRef(Lazy.force Coqlib.coq_or_ref),[t1;t2])
+(* [glob_make_or P1 P2] build the glob_constr corresponding to [P1 \/ P2] *)
+let glob_make_or t1 t2 = mkRApp (mkRRef(Lazy.force Coqlib.coq_or_ref),[t1;t2])
-(* [raw_make_or_list [P1;...;Pn]] build the rawconstr corresponding
+(* [glob_make_or_list [P1;...;Pn]] build the glob_constr corresponding
to [P1 \/ ( .... \/ Pn)]
*)
-let rec raw_make_or_list = function
+let rec glob_make_or_list = function
| [] -> raise (Invalid_argument "mk_or")
| [e] -> e
- | e::l -> raw_make_or e (raw_make_or_list l)
+ | e::l -> glob_make_or e (glob_make_or_list l)
let remove_name_from_mapping mapping na =
@@ -120,70 +120,70 @@ let remove_name_from_mapping mapping na =
let change_vars =
let rec change_vars mapping rt =
match rt with
- | RRef _ -> rt
- | RVar(loc,id) ->
+ | GRef _ -> rt
+ | GVar(loc,id) ->
let new_id =
try
Idmap.find id mapping
with Not_found -> id
in
- RVar(loc,new_id)
- | REvar _ -> rt
- | RPatVar _ -> rt
- | RApp(loc,rt',rtl) ->
- RApp(loc,
+ GVar(loc,new_id)
+ | GEvar _ -> rt
+ | GPatVar _ -> rt
+ | GApp(loc,rt',rtl) ->
+ GApp(loc,
change_vars mapping rt',
List.map (change_vars mapping) rtl
)
- | RLambda(loc,name,k,t,b) ->
- RLambda(loc,
+ | GLambda(loc,name,k,t,b) ->
+ GLambda(loc,
name,
k,
change_vars mapping t,
change_vars (remove_name_from_mapping mapping name) b
)
- | RProd(loc,name,k,t,b) ->
- RProd(loc,
+ | GProd(loc,name,k,t,b) ->
+ GProd(loc,
name,
k,
change_vars mapping t,
change_vars (remove_name_from_mapping mapping name) b
)
- | RLetIn(loc,name,def,b) ->
- RLetIn(loc,
+ | GLetIn(loc,name,def,b) ->
+ GLetIn(loc,
name,
change_vars mapping def,
change_vars (remove_name_from_mapping mapping name) b
)
- | RLetTuple(loc,nal,(na,rto),b,e) ->
+ | GLetTuple(loc,nal,(na,rto),b,e) ->
let new_mapping = List.fold_left remove_name_from_mapping mapping nal in
- RLetTuple(loc,
+ GLetTuple(loc,
nal,
(na, Option.map (change_vars mapping) rto),
change_vars mapping b,
change_vars new_mapping e
)
- | RCases(loc,sty,infos,el,brl) ->
- RCases(loc,sty,
+ | GCases(loc,sty,infos,el,brl) ->
+ GCases(loc,sty,
infos,
List.map (fun (e,x) -> (change_vars mapping e,x)) el,
List.map (change_vars_br mapping) brl
)
- | RIf(loc,b,(na,e_option),lhs,rhs) ->
- RIf(loc,
+ | GIf(loc,b,(na,e_option),lhs,rhs) ->
+ GIf(loc,
change_vars mapping b,
(na,Option.map (change_vars mapping) e_option),
change_vars mapping lhs,
change_vars mapping rhs
)
- | RRec _ -> error "Local (co)fixes are not supported"
- | RSort _ -> rt
- | RHole _ -> rt
- | RCast(loc,b,CastConv (k,t)) ->
- RCast(loc,change_vars mapping b, CastConv (k,change_vars mapping t))
- | RCast(loc,b,CastCoerce) ->
- RCast(loc,change_vars mapping b,CastCoerce)
- | RDynamic _ -> error "Not handled RDynamic"
+ | GRec _ -> error "Local (co)fixes are not supported"
+ | GSort _ -> rt
+ | GHole _ -> rt
+ | GCast(loc,b,CastConv (k,t)) ->
+ GCast(loc,change_vars mapping b, CastConv (k,change_vars mapping t))
+ | GCast(loc,b,CastCoerce) ->
+ GCast(loc,change_vars mapping b,CastCoerce)
+ | GDynamic _ -> error "Not handled GDynamic"
and change_vars_br mapping ((loc,idl,patl,res) as br) =
let new_mapping = List.fold_right Idmap.remove idl mapping in
if idmap_is_empty new_mapping
@@ -262,22 +262,22 @@ let get_pattern_id pat = raw_get_pattern_id pat []
let rec alpha_rt excluded rt =
let new_rt =
match rt with
- | RRef _ | RVar _ | REvar _ | RPatVar _ -> rt
- | RLambda(loc,Anonymous,k,t,b) ->
+ | GRef _ | GVar _ | GEvar _ | GPatVar _ -> rt
+ | GLambda(loc,Anonymous,k,t,b) ->
let new_id = Namegen.next_ident_away (id_of_string "_x") excluded in
let new_excluded = new_id :: excluded in
let new_t = alpha_rt new_excluded t in
let new_b = alpha_rt new_excluded b in
- RLambda(loc,Name new_id,k,new_t,new_b)
- | RProd(loc,Anonymous,k,t,b) ->
+ GLambda(loc,Name new_id,k,new_t,new_b)
+ | GProd(loc,Anonymous,k,t,b) ->
let new_t = alpha_rt excluded t in
let new_b = alpha_rt excluded b in
- RProd(loc,Anonymous,k,new_t,new_b)
- | RLetIn(loc,Anonymous,t,b) ->
+ GProd(loc,Anonymous,k,new_t,new_b)
+ | GLetIn(loc,Anonymous,t,b) ->
let new_t = alpha_rt excluded t in
let new_b = alpha_rt excluded b in
- RLetIn(loc,Anonymous,new_t,new_b)
- | RLambda(loc,Name id,k,t,b) ->
+ GLetIn(loc,Anonymous,new_t,new_b)
+ | GLambda(loc,Name id,k,t,b) ->
let new_id = Namegen.next_ident_away id excluded in
let t,b =
if new_id = id
@@ -289,8 +289,8 @@ let rec alpha_rt excluded rt =
let new_excluded = new_id::excluded in
let new_t = alpha_rt new_excluded t in
let new_b = alpha_rt new_excluded b in
- RLambda(loc,Name new_id,k,new_t,new_b)
- | RProd(loc,Name id,k,t,b) ->
+ GLambda(loc,Name new_id,k,new_t,new_b)
+ | GProd(loc,Name id,k,t,b) ->
let new_id = Namegen.next_ident_away id excluded in
let new_excluded = new_id::excluded in
let t,b =
@@ -302,8 +302,8 @@ let rec alpha_rt excluded rt =
in
let new_t = alpha_rt new_excluded t in
let new_b = alpha_rt new_excluded b in
- RProd(loc,Name new_id,k,new_t,new_b)
- | RLetIn(loc,Name id,t,b) ->
+ GProd(loc,Name new_id,k,new_t,new_b)
+ | GLetIn(loc,Name id,t,b) ->
let new_id = Namegen.next_ident_away id excluded in
let t,b =
if new_id = id
@@ -315,10 +315,10 @@ let rec alpha_rt excluded rt =
let new_excluded = new_id::excluded in
let new_t = alpha_rt new_excluded t in
let new_b = alpha_rt new_excluded b in
- RLetIn(loc,Name new_id,new_t,new_b)
+ GLetIn(loc,Name new_id,new_t,new_b)
- | RLetTuple(loc,nal,(na,rto),t,b) ->
+ | GLetTuple(loc,nal,(na,rto),t,b) ->
let rev_new_nal,new_excluded,mapping =
List.fold_left
(fun (nal,excluded,mapping) na ->
@@ -345,28 +345,28 @@ let rec alpha_rt excluded rt =
let new_t = alpha_rt new_excluded new_t in
let new_b = alpha_rt new_excluded new_b in
let new_rto = Option.map (alpha_rt new_excluded) new_rto in
- RLetTuple(loc,new_nal,(na,new_rto),new_t,new_b)
- | RCases(loc,sty,infos,el,brl) ->
+ GLetTuple(loc,new_nal,(na,new_rto),new_t,new_b)
+ | GCases(loc,sty,infos,el,brl) ->
let new_el =
List.map (function (rt,i) -> alpha_rt excluded rt, i) el
in
- RCases(loc,sty,infos,new_el,List.map (alpha_br excluded) brl)
- | RIf(loc,b,(na,e_o),lhs,rhs) ->
- RIf(loc,alpha_rt excluded b,
+ GCases(loc,sty,infos,new_el,List.map (alpha_br excluded) brl)
+ | GIf(loc,b,(na,e_o),lhs,rhs) ->
+ GIf(loc,alpha_rt excluded b,
(na,Option.map (alpha_rt excluded) e_o),
alpha_rt excluded lhs,
alpha_rt excluded rhs
)
- | RRec _ -> error "Not handled RRec"
- | RSort _ -> rt
- | RHole _ -> rt
- | RCast (loc,b,CastConv (k,t)) ->
- RCast(loc,alpha_rt excluded b,CastConv(k,alpha_rt excluded t))
- | RCast (loc,b,CastCoerce) ->
- RCast(loc,alpha_rt excluded b,CastCoerce)
- | RDynamic _ -> error "Not handled RDynamic"
- | RApp(loc,f,args) ->
- RApp(loc,
+ | GRec _ -> error "Not handled GRec"
+ | GSort _ -> rt
+ | GHole _ -> rt
+ | GCast (loc,b,CastConv (k,t)) ->
+ GCast(loc,alpha_rt excluded b,CastConv(k,alpha_rt excluded t))
+ | GCast (loc,b,CastCoerce) ->
+ GCast(loc,alpha_rt excluded b,CastCoerce)
+ | GDynamic _ -> error "Not handled GDynamic"
+ | GApp(loc,f,args) ->
+ GApp(loc,
alpha_rt excluded f,
List.map (alpha_rt excluded) args
)
@@ -386,35 +386,35 @@ and alpha_br excluded (loc,ids,patl,res) =
*)
let is_free_in id =
let rec is_free_in = function
- | RRef _ -> false
- | RVar(_,id') -> id_ord id' id == 0
- | REvar _ -> false
- | RPatVar _ -> false
- | RApp(_,rt,rtl) -> List.exists is_free_in (rt::rtl)
- | RLambda(_,n,_,t,b) | RProd(_,n,_,t,b) | RLetIn(_,n,t,b) ->
+ | GRef _ -> false
+ | GVar(_,id') -> id_ord id' id == 0
+ | GEvar _ -> false
+ | GPatVar _ -> false
+ | GApp(_,rt,rtl) -> List.exists is_free_in (rt::rtl)
+ | GLambda(_,n,_,t,b) | GProd(_,n,_,t,b) | GLetIn(_,n,t,b) ->
let check_in_b =
match n with
| Name id' -> id_ord id' id <> 0
| _ -> true
in
is_free_in t || (check_in_b && is_free_in b)
- | RCases(_,_,_,el,brl) ->
+ | GCases(_,_,_,el,brl) ->
(List.exists (fun (e,_) -> is_free_in e) el) ||
List.exists is_free_in_br brl
- | RLetTuple(_,nal,_,b,t) ->
+ | GLetTuple(_,nal,_,b,t) ->
let check_in_nal =
not (List.exists (function Name id' -> id'= id | _ -> false) nal)
in
is_free_in t || (check_in_nal && is_free_in b)
- | RIf(_,cond,_,br1,br2) ->
+ | GIf(_,cond,_,br1,br2) ->
is_free_in cond || is_free_in br1 || is_free_in br2
- | RRec _ -> raise (UserError("",str "Not handled RRec"))
- | RSort _ -> false
- | RHole _ -> false
- | RCast (_,b,CastConv (_,t)) -> is_free_in b || is_free_in t
- | RCast (_,b,CastCoerce) -> is_free_in b
- | RDynamic _ -> raise (UserError("",str "Not handled RDynamic"))
+ | GRec _ -> raise (UserError("",str "Not handled GRec"))
+ | GSort _ -> false
+ | GHole _ -> false
+ | GCast (_,b,CastConv (_,t)) -> is_free_in b || is_free_in t
+ | GCast (_,b,CastCoerce) -> is_free_in b
+ | GDynamic _ -> raise (UserError("",str "Not handled GDynamic"))
and is_free_in_br (_,ids,_,rt) =
(not (List.mem id ids)) && is_free_in rt
in
@@ -451,69 +451,69 @@ let rec pattern_to_term = function
let replace_var_by_term x_id term =
let rec replace_var_by_pattern rt =
match rt with
- | RRef _ -> rt
- | RVar(_,id) when id_ord id x_id == 0 -> term
- | RVar _ -> rt
- | REvar _ -> rt
- | RPatVar _ -> rt
- | RApp(loc,rt',rtl) ->
- RApp(loc,
+ | GRef _ -> rt
+ | GVar(_,id) when id_ord id x_id == 0 -> term
+ | GVar _ -> rt
+ | GEvar _ -> rt
+ | GPatVar _ -> rt
+ | GApp(loc,rt',rtl) ->
+ GApp(loc,
replace_var_by_pattern rt',
List.map replace_var_by_pattern rtl
)
- | RLambda(_,Name id,_,_,_) when id_ord id x_id == 0 -> rt
- | RLambda(loc,name,k,t,b) ->
- RLambda(loc,
+ | GLambda(_,Name id,_,_,_) when id_ord id x_id == 0 -> rt
+ | GLambda(loc,name,k,t,b) ->
+ GLambda(loc,
name,
k,
replace_var_by_pattern t,
replace_var_by_pattern b
)
- | RProd(_,Name id,_,_,_) when id_ord id x_id == 0 -> rt
- | RProd(loc,name,k,t,b) ->
- RProd(loc,
+ | GProd(_,Name id,_,_,_) when id_ord id x_id == 0 -> rt
+ | GProd(loc,name,k,t,b) ->
+ GProd(loc,
name,
k,
replace_var_by_pattern t,
replace_var_by_pattern b
)
- | RLetIn(_,Name id,_,_) when id_ord id x_id == 0 -> rt
- | RLetIn(loc,name,def,b) ->
- RLetIn(loc,
+ | GLetIn(_,Name id,_,_) when id_ord id x_id == 0 -> rt
+ | GLetIn(loc,name,def,b) ->
+ GLetIn(loc,
name,
replace_var_by_pattern def,
replace_var_by_pattern b
)
- | RLetTuple(_,nal,_,_,_)
+ | GLetTuple(_,nal,_,_,_)
when List.exists (function Name id -> id = x_id | _ -> false) nal ->
rt
- | RLetTuple(loc,nal,(na,rto),def,b) ->
- RLetTuple(loc,
+ | GLetTuple(loc,nal,(na,rto),def,b) ->
+ GLetTuple(loc,
nal,
(na,Option.map replace_var_by_pattern rto),
replace_var_by_pattern def,
replace_var_by_pattern b
)
- | RCases(loc,sty,infos,el,brl) ->
- RCases(loc,sty,
+ | GCases(loc,sty,infos,el,brl) ->
+ GCases(loc,sty,
infos,
List.map (fun (e,x) -> (replace_var_by_pattern e,x)) el,
List.map replace_var_by_pattern_br brl
)
- | RIf(loc,b,(na,e_option),lhs,rhs) ->
- RIf(loc, replace_var_by_pattern b,
+ | GIf(loc,b,(na,e_option),lhs,rhs) ->
+ GIf(loc, replace_var_by_pattern b,
(na,Option.map replace_var_by_pattern e_option),
replace_var_by_pattern lhs,
replace_var_by_pattern rhs
)
- | RRec _ -> raise (UserError("",str "Not handled RRec"))
- | RSort _ -> rt
- | RHole _ -> rt
- | RCast(loc,b,CastConv(k,t)) ->
- RCast(loc,replace_var_by_pattern b,CastConv(k,replace_var_by_pattern t))
- | RCast(loc,b,CastCoerce) ->
- RCast(loc,replace_var_by_pattern b,CastCoerce)
- | RDynamic _ -> raise (UserError("",str "Not handled RDynamic"))
+ | GRec _ -> raise (UserError("",str "Not handled GRec"))
+ | GSort _ -> rt
+ | GHole _ -> rt
+ | GCast(loc,b,CastConv(k,t)) ->
+ GCast(loc,replace_var_by_pattern b,CastConv(k,replace_var_by_pattern t))
+ | GCast(loc,b,CastCoerce) ->
+ GCast(loc,replace_var_by_pattern b,CastCoerce)
+ | GDynamic _ -> raise (UserError("",str "Not handled GDynamic"))
and replace_var_by_pattern_br ((loc,idl,patl,res) as br) =
if List.exists (fun id -> id_ord id x_id == 0) idl
then br
@@ -590,21 +590,21 @@ let ids_of_rawterm c =
let rec ids_of_rawterm acc c =
let idof = id_of_name in
match c with
- | RVar (_,id) -> id::acc
- | RApp (loc,g,args) ->
+ | GVar (_,id) -> id::acc
+ | GApp (loc,g,args) ->
ids_of_rawterm [] g @ List.flatten (List.map (ids_of_rawterm []) args) @ acc
- | RLambda (loc,na,k,ty,c) -> idof na :: ids_of_rawterm [] ty @ ids_of_rawterm [] c @ acc
- | RProd (loc,na,k,ty,c) -> idof na :: ids_of_rawterm [] ty @ ids_of_rawterm [] c @ acc
- | RLetIn (loc,na,b,c) -> idof na :: ids_of_rawterm [] b @ ids_of_rawterm [] c @ acc
- | RCast (loc,c,CastConv(k,t)) -> ids_of_rawterm [] c @ ids_of_rawterm [] t @ acc
- | RCast (loc,c,CastCoerce) -> ids_of_rawterm [] c @ acc
- | RIf (loc,c,(na,po),b1,b2) -> ids_of_rawterm [] c @ ids_of_rawterm [] b1 @ ids_of_rawterm [] b2 @ acc
- | RLetTuple (_,nal,(na,po),b,c) ->
+ | GLambda (loc,na,k,ty,c) -> idof na :: ids_of_rawterm [] ty @ ids_of_rawterm [] c @ acc
+ | GProd (loc,na,k,ty,c) -> idof na :: ids_of_rawterm [] ty @ ids_of_rawterm [] c @ acc
+ | GLetIn (loc,na,b,c) -> idof na :: ids_of_rawterm [] b @ ids_of_rawterm [] c @ acc
+ | GCast (loc,c,CastConv(k,t)) -> ids_of_rawterm [] c @ ids_of_rawterm [] t @ acc
+ | GCast (loc,c,CastCoerce) -> ids_of_rawterm [] c @ acc
+ | GIf (loc,c,(na,po),b1,b2) -> ids_of_rawterm [] c @ ids_of_rawterm [] b1 @ ids_of_rawterm [] b2 @ acc
+ | GLetTuple (_,nal,(na,po),b,c) ->
List.map idof nal @ ids_of_rawterm [] b @ ids_of_rawterm [] c @ acc
- | RCases (loc,sty,rtntypopt,tml,brchl) ->
+ | GCases (loc,sty,rtntypopt,tml,brchl) ->
List.flatten (List.map (fun (_,idl,patl,c) -> idl @ ids_of_rawterm [] c) brchl)
- | RRec _ -> failwith "Fix inside a constructor branch"
- | (RSort _ | RHole _ | RRef _ | REvar _ | RPatVar _ | RDynamic _) -> []
+ | GRec _ -> failwith "Fix inside a constructor branch"
+ | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GDynamic _) -> []
in
(* build the set *)
List.fold_left (fun acc x -> Idset.add x acc) Idset.empty (ids_of_rawterm [] c)
@@ -616,59 +616,59 @@ let ids_of_rawterm c =
let zeta_normalize =
let rec zeta_normalize_term rt =
match rt with
- | RRef _ -> rt
- | RVar _ -> rt
- | REvar _ -> rt
- | RPatVar _ -> rt
- | RApp(loc,rt',rtl) ->
- RApp(loc,
+ | GRef _ -> rt
+ | GVar _ -> rt
+ | GEvar _ -> rt
+ | GPatVar _ -> rt
+ | GApp(loc,rt',rtl) ->
+ GApp(loc,
zeta_normalize_term rt',
List.map zeta_normalize_term rtl
)
- | RLambda(loc,name,k,t,b) ->
- RLambda(loc,
+ | GLambda(loc,name,k,t,b) ->
+ GLambda(loc,
name,
k,
zeta_normalize_term t,
zeta_normalize_term b
)
- | RProd(loc,name,k,t,b) ->
- RProd(loc,
+ | GProd(loc,name,k,t,b) ->
+ GProd(loc,
name,
k,
zeta_normalize_term t,
zeta_normalize_term b
)
- | RLetIn(_,Name id,def,b) ->
+ | GLetIn(_,Name id,def,b) ->
zeta_normalize_term (replace_var_by_term id def b)
- | RLetIn(loc,Anonymous,def,b) -> zeta_normalize_term b
- | RLetTuple(loc,nal,(na,rto),def,b) ->
- RLetTuple(loc,
+ | GLetIn(loc,Anonymous,def,b) -> zeta_normalize_term b
+ | GLetTuple(loc,nal,(na,rto),def,b) ->
+ GLetTuple(loc,
nal,
(na,Option.map zeta_normalize_term rto),
zeta_normalize_term def,
zeta_normalize_term b
)
- | RCases(loc,sty,infos,el,brl) ->
- RCases(loc,sty,
+ | GCases(loc,sty,infos,el,brl) ->
+ GCases(loc,sty,
infos,
List.map (fun (e,x) -> (zeta_normalize_term e,x)) el,
List.map zeta_normalize_br brl
)
- | RIf(loc,b,(na,e_option),lhs,rhs) ->
- RIf(loc, zeta_normalize_term b,
+ | GIf(loc,b,(na,e_option),lhs,rhs) ->
+ GIf(loc, zeta_normalize_term b,
(na,Option.map zeta_normalize_term e_option),
zeta_normalize_term lhs,
zeta_normalize_term rhs
)
- | RRec _ -> raise (UserError("",str "Not handled RRec"))
- | RSort _ -> rt
- | RHole _ -> rt
- | RCast(loc,b,CastConv(k,t)) ->
- RCast(loc,zeta_normalize_term b,CastConv(k,zeta_normalize_term t))
- | RCast(loc,b,CastCoerce) ->
- RCast(loc,zeta_normalize_term b,CastCoerce)
- | RDynamic _ -> raise (UserError("",str "Not handled RDynamic"))
+ | GRec _ -> raise (UserError("",str "Not handled GRec"))
+ | GSort _ -> rt
+ | GHole _ -> rt
+ | GCast(loc,b,CastConv(k,t)) ->
+ GCast(loc,zeta_normalize_term b,CastConv(k,zeta_normalize_term t))
+ | GCast(loc,b,CastCoerce) ->
+ GCast(loc,zeta_normalize_term b,CastCoerce)
+ | GDynamic _ -> raise (UserError("",str "Not handled GDynamic"))
and zeta_normalize_br (loc,idl,patl,res) =
(loc,idl,patl,zeta_normalize_term res)
in
@@ -688,29 +688,29 @@ let expand_as =
in
let rec expand_as map rt =
match rt with
- | RRef _ | REvar _ | RPatVar _ | RSort _ | RHole _ -> rt
- | RVar(_,id) ->
+ | GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ -> rt
+ | GVar(_,id) ->
begin
try
Idmap.find id map
with Not_found -> rt
end
- | RApp(loc,f,args) -> RApp(loc,expand_as map f,List.map (expand_as map) args)
- | RLambda(loc,na,k,t,b) -> RLambda(loc,na,k,expand_as map t, expand_as map b)
- | RProd(loc,na,k,t,b) -> RProd(loc,na,k,expand_as map t, expand_as map b)
- | RLetIn(loc,na,v,b) -> RLetIn(loc,na, expand_as map v,expand_as map b)
- | RLetTuple(loc,nal,(na,po),v,b) ->
- RLetTuple(loc,nal,(na,Option.map (expand_as map) po),
+ | GApp(loc,f,args) -> GApp(loc,expand_as map f,List.map (expand_as map) args)
+ | GLambda(loc,na,k,t,b) -> GLambda(loc,na,k,expand_as map t, expand_as map b)
+ | GProd(loc,na,k,t,b) -> GProd(loc,na,k,expand_as map t, expand_as map b)
+ | GLetIn(loc,na,v,b) -> GLetIn(loc,na, expand_as map v,expand_as map b)
+ | GLetTuple(loc,nal,(na,po),v,b) ->
+ GLetTuple(loc,nal,(na,Option.map (expand_as map) po),
expand_as map v, expand_as map b)
- | RIf(loc,e,(na,po),br1,br2) ->
- RIf(loc,expand_as map e,(na,Option.map (expand_as map) po),
+ | GIf(loc,e,(na,po),br1,br2) ->
+ GIf(loc,expand_as map e,(na,Option.map (expand_as map) po),
expand_as map br1, expand_as map br2)
- | RRec _ -> error "Not handled RRec"
- | RDynamic _ -> error "Not handled RDynamic"
- | RCast(loc,b,CastConv(kind,t)) -> RCast(loc,expand_as map b,CastConv(kind,expand_as map t))
- | RCast(loc,b,CastCoerce) -> RCast(loc,expand_as map b,CastCoerce)
- | RCases(loc,sty,po,el,brl) ->
- RCases(loc, sty, Option.map (expand_as map) po, List.map (fun (rt,t) -> expand_as map rt,t) el,
+ | GRec _ -> error "Not handled GRec"
+ | GDynamic _ -> error "Not handled GDynamic"
+ | GCast(loc,b,CastConv(kind,t)) -> GCast(loc,expand_as map b,CastConv(kind,expand_as map t))
+ | GCast(loc,b,CastCoerce) -> GCast(loc,expand_as map b,CastCoerce)
+ | GCases(loc,sty,po,el,brl) ->
+ GCases(loc, sty, Option.map (expand_as map) po, List.map (fun (rt,t) -> expand_as map rt,t) el,
List.map (expand_as_br map) brl)
and expand_as_br map (loc,idl,cpl,rt) =
(loc,idl,cpl, expand_as (List.fold_left add_as map cpl) rt)
diff --git a/plugins/funind/rawtermops.mli b/plugins/funind/rawtermops.mli
index 465703386..9e872c236 100644
--- a/plugins/funind/rawtermops.mli
+++ b/plugins/funind/rawtermops.mli
@@ -7,60 +7,60 @@ val idmap_is_empty : 'a Names.Idmap.t -> bool
(* [get_pattern_id pat] returns a list of all the variable appearing in [pat] *)
val get_pattern_id : cases_pattern -> Names.identifier list
-(* [pattern_to_term pat] returns a rawconstr corresponding to [pat].
+(* [pattern_to_term pat] returns a glob_constr corresponding to [pat].
[pat] must not contain occurences of anonymous pattern
*)
-val pattern_to_term : cases_pattern -> rawconstr
+val pattern_to_term : cases_pattern -> glob_constr
(*
- Some basic functions to rebuild rawconstr
+ Some basic functions to rebuild glob_constr
In each of them the location is Util.dummy_loc
*)
-val mkRRef : Libnames.global_reference -> rawconstr
-val mkRVar : Names.identifier -> rawconstr
-val mkRApp : rawconstr*(rawconstr list) -> rawconstr
-val mkRLambda : Names.name * rawconstr * rawconstr -> rawconstr
-val mkRProd : Names.name * rawconstr * rawconstr -> rawconstr
-val mkRLetIn : Names.name * rawconstr * rawconstr -> rawconstr
-val mkRCases : rawconstr option * tomatch_tuples * cases_clauses -> rawconstr
-val mkRSort : rawsort -> rawconstr
-val mkRHole : unit -> rawconstr (* we only build Evd.BinderType Anonymous holes *)
-val mkRCast : rawconstr* rawconstr -> rawconstr
+val mkRRef : Libnames.global_reference -> glob_constr
+val mkRVar : Names.identifier -> glob_constr
+val mkRApp : glob_constr*(glob_constr list) -> glob_constr
+val mkRLambda : Names.name * glob_constr * glob_constr -> glob_constr
+val mkRProd : Names.name * glob_constr * glob_constr -> glob_constr
+val mkRLetIn : Names.name * glob_constr * glob_constr -> glob_constr
+val mkRCases : glob_constr option * tomatch_tuples * cases_clauses -> glob_constr
+val mkRSort : rawsort -> glob_constr
+val mkRHole : unit -> glob_constr (* we only build Evd.BinderType Anonymous holes *)
+val mkRCast : glob_constr* glob_constr -> glob_constr
(*
- Some basic functions to decompose rawconstrs
+ Some basic functions to decompose glob_constrs
These are analogous to the ones constrs
*)
-val raw_decompose_prod : rawconstr -> (Names.name*rawconstr) list * rawconstr
-val raw_decompose_prod_or_letin :
- rawconstr -> (Names.name*rawconstr option*rawconstr option) list * rawconstr
-val raw_decompose_prod_n : int -> rawconstr -> (Names.name*rawconstr) list * rawconstr
-val raw_decompose_prod_or_letin_n : int -> rawconstr ->
- (Names.name*rawconstr option*rawconstr option) list * rawconstr
-val raw_compose_prod : rawconstr -> (Names.name*rawconstr) list -> rawconstr
-val raw_compose_prod_or_letin: rawconstr ->
- (Names.name*rawconstr option*rawconstr option) list -> rawconstr
-val raw_decompose_app : rawconstr -> rawconstr*(rawconstr list)
-
-
-(* [raw_make_eq t1 t2] build the rawconstr corresponding to [t2 = t1] *)
-val raw_make_eq : ?typ:rawconstr -> rawconstr -> rawconstr -> rawconstr
-(* [raw_make_neq t1 t2] build the rawconstr corresponding to [t1 <> t2] *)
-val raw_make_neq : rawconstr -> rawconstr -> rawconstr
-(* [raw_make_or P1 P2] build the rawconstr corresponding to [P1 \/ P2] *)
-val raw_make_or : rawconstr -> rawconstr -> rawconstr
-
-(* [raw_make_or_list [P1;...;Pn]] build the rawconstr corresponding
+val glob_decompose_prod : glob_constr -> (Names.name*glob_constr) list * glob_constr
+val glob_decompose_prod_or_letin :
+ glob_constr -> (Names.name*glob_constr option*glob_constr option) list * glob_constr
+val glob_decompose_prod_n : int -> glob_constr -> (Names.name*glob_constr) list * glob_constr
+val glob_decompose_prod_or_letin_n : int -> glob_constr ->
+ (Names.name*glob_constr option*glob_constr option) list * glob_constr
+val glob_compose_prod : glob_constr -> (Names.name*glob_constr) list -> glob_constr
+val glob_compose_prod_or_letin: glob_constr ->
+ (Names.name*glob_constr option*glob_constr option) list -> glob_constr
+val glob_decompose_app : glob_constr -> glob_constr*(glob_constr list)
+
+
+(* [glob_make_eq t1 t2] build the glob_constr corresponding to [t2 = t1] *)
+val glob_make_eq : ?typ:glob_constr -> glob_constr -> glob_constr -> glob_constr
+(* [glob_make_neq t1 t2] build the glob_constr corresponding to [t1 <> t2] *)
+val glob_make_neq : glob_constr -> glob_constr -> glob_constr
+(* [glob_make_or P1 P2] build the glob_constr corresponding to [P1 \/ P2] *)
+val glob_make_or : glob_constr -> glob_constr -> glob_constr
+
+(* [glob_make_or_list [P1;...;Pn]] build the glob_constr corresponding
to [P1 \/ ( .... \/ Pn)]
*)
-val raw_make_or_list : rawconstr list -> rawconstr
+val glob_make_or_list : glob_constr list -> glob_constr
(* alpha_conversion functions *)
-(* Replace the var mapped in the rawconstr/context *)
-val change_vars : Names.identifier Names.Idmap.t -> rawconstr -> rawconstr
+(* Replace the var mapped in the glob_constr/context *)
+val change_vars : Names.identifier Names.Idmap.t -> glob_constr -> glob_constr
@@ -80,27 +80,27 @@ val change_vars : Names.identifier Names.Idmap.t -> rawconstr -> rawconstr
(* [alpha_rt avoid rt] alpha convert [rt] s.t. the result repects barendregt
conventions and does not share bound variables with avoid
*)
-val alpha_rt : Names.identifier list -> rawconstr -> rawconstr
+val alpha_rt : Names.identifier list -> glob_constr -> glob_constr
(* same as alpha_rt but for case branches *)
val alpha_br : Names.identifier list ->
Util.loc * Names.identifier list * Rawterm.cases_pattern list *
- Rawterm.rawconstr ->
+ Rawterm.glob_constr ->
Util.loc * Names.identifier list * Rawterm.cases_pattern list *
- Rawterm.rawconstr
+ Rawterm.glob_constr
(* Reduction function *)
val replace_var_by_term :
Names.identifier ->
- Rawterm.rawconstr -> Rawterm.rawconstr -> Rawterm.rawconstr
+ Rawterm.glob_constr -> Rawterm.glob_constr -> Rawterm.glob_constr
(*
[is_free_in id rt] checks if [id] is a free variable in [rt]
*)
-val is_free_in : Names.identifier -> rawconstr -> bool
+val is_free_in : Names.identifier -> glob_constr -> bool
val are_unifiable : cases_pattern -> cases_pattern -> bool
@@ -115,12 +115,12 @@ val eq_cases_pattern : cases_pattern -> cases_pattern -> bool
val ids_of_pat : cases_pattern -> Names.Idset.t
(* TODO: finish this function (Fix not treated) *)
-val ids_of_rawterm: rawconstr -> Names.Idset.t
+val ids_of_rawterm: glob_constr -> Names.Idset.t
(*
removing let_in construction in a rawterm
*)
-val zeta_normalize : Rawterm.rawconstr -> Rawterm.rawconstr
+val zeta_normalize : Rawterm.glob_constr -> Rawterm.glob_constr
-val expand_as : rawconstr -> rawconstr
+val expand_as : glob_constr -> glob_constr
diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml
index f066e39d0..1b43d4045 100644
--- a/plugins/funind/recdef.ml
+++ b/plugins/funind/recdef.ml
@@ -1107,22 +1107,22 @@ let (value_f:constr list -> global_reference -> constr) =
)
in
let fun_body =
- RCases
+ GCases
(d0,RegularStyle,None,
- [RApp(d0, RRef(d0,fterm), List.rev_map (fun x_id -> RVar(d0, x_id)) rev_x_id_l),
+ [GApp(d0, GRef(d0,fterm), List.rev_map (fun x_id -> GVar(d0, x_id)) rev_x_id_l),
(Anonymous,None)],
[d0, [v_id], [PatCstr(d0,(ind_of_ref
(delayed_force coq_sig_ref),1),
[PatVar(d0, Name v_id);
PatVar(d0, Anonymous)],
Anonymous)],
- RVar(d0,v_id)])
+ GVar(d0,v_id)])
in
let value =
List.fold_left2
(fun acc x_id a ->
- RLambda
- (d0, Name x_id, Explicit, RDynamic(d0, constr_in a),
+ GLambda
+ (d0, Name x_id, Explicit, GDynamic(d0, constr_in a),
acc
)
)
diff --git a/plugins/setoid_ring/newring.ml4 b/plugins/setoid_ring/newring.ml4
index 64cfa0d01..112a13e53 100644
--- a/plugins/setoid_ring/newring.ml4
+++ b/plugins/setoid_ring/newring.ml4
@@ -141,7 +141,7 @@ let closed_term_ast l =
let l = List.map (fun gr -> ArgArg(dummy_loc,gr)) l in
TacFun([Some(id_of_string"t")],
TacAtom(dummy_loc,TacExtend(dummy_loc,"closed_term",
- [Genarg.in_gen Genarg.globwit_constr (RVar(dummy_loc,id_of_string"t"),None);
+ [Genarg.in_gen Genarg.globwit_constr (GVar(dummy_loc,id_of_string"t"),None);
Genarg.in_gen (Genarg.wit_list1 Genarg.globwit_ref) l])))
(*
let _ = add_tacdef false ((dummy_loc,id_of_string"ring_closed_term"
diff --git a/plugins/subtac/subtac.ml b/plugins/subtac/subtac.ml
index 62fa0b2c9..0ea2290db 100644
--- a/plugins/subtac/subtac.ml
+++ b/plugins/subtac/subtac.ml
@@ -88,7 +88,7 @@ let start_proof_and_print env isevars idopt k t hook =
start_proof_com env isevars idopt k t hook;
print_subgoals ()
-let _ = Detyping.set_detype_anonymous (fun loc n -> RVar (loc, id_of_string ("Anonymous_REL_" ^ string_of_int n)))
+let _ = Detyping.set_detype_anonymous (fun loc n -> GVar (loc, id_of_string ("Anonymous_REL_" ^ string_of_int n)))
let assumption_message id =
Flags.if_verbose message ((string_of_id id) ^ " is assumed")
diff --git a/plugins/subtac/subtac_cases.ml b/plugins/subtac/subtac_cases.ml
index b6a42b6b4..5ef6f0f88 100644
--- a/plugins/subtac/subtac_cases.ml
+++ b/plugins/subtac/subtac_cases.ml
@@ -86,7 +86,7 @@ let mkSpecialLetInJudge j (na,(deppat,nondeppat,d,t)) =
type rhs =
{ rhs_env : env;
avoid_ids : identifier list;
- it : rawconstr;
+ it : glob_constr;
}
type equation =
@@ -234,7 +234,7 @@ type pattern_matching_problem =
mat : matrix;
caseloc : loc;
casestyle: case_style;
- typing_function: type_constraint -> env -> rawconstr -> unsafe_judgment }
+ typing_function: type_constraint -> env -> glob_constr -> unsafe_judgment }
(*--------------------------------------------------------------------------*
* A few functions to infer the inductive type from the patterns instead of *
@@ -366,10 +366,10 @@ let find_tomatch_tycon isevars env loc = function
| None -> empty_tycon
let coerce_row typing_fun isevars env pats (tomatch,(_,indopt)) =
- let loc = Some (loc_of_rawconstr tomatch) in
+ let loc = Some (loc_of_glob_constr tomatch) in
let tycon = find_tomatch_tycon isevars env loc indopt in
let j = typing_fun tycon env tomatch in
- let evd, j = Coercion.inh_coerce_to_base (loc_of_rawconstr tomatch) env !isevars j in
+ let evd, j = Coercion.inh_coerce_to_base (loc_of_glob_constr tomatch) env !isevars j in
isevars := evd;
let typ = nf_evar ( !isevars) j.uj_type in
let t =
@@ -527,7 +527,7 @@ let extract_rhs pb =
let occur_in_rhs na rhs =
match na with
| Anonymous -> false
- | Name id -> occur_rawconstr id rhs.it
+ | Name id -> occur_glob_constr id rhs.it
let is_dep_patt eqn = function
| PatVar (_,name) -> occur_in_rhs name eqn.rhs
@@ -1515,7 +1515,7 @@ let mk_JMeq typ x typ' y =
mkApp (delayed_force Subtac_utils.jmeq_ind, [| typ; x ; typ'; y |])
let mk_JMeq_refl typ x = mkApp (delayed_force Subtac_utils.jmeq_refl, [| typ; x |])
-let hole = RHole (dummy_loc, Evd.QuestionMark (Evd.Define true))
+let hole = GHole (dummy_loc, Evd.QuestionMark (Evd.Define true))
let constr_of_pat env isevars arsign pat avoid =
let rec typ env (ty, realargs) pat avoid =
@@ -1604,12 +1604,12 @@ let vars_of_ctx ctx =
match b with
| Some t' when kind_of_term t' = Rel 0 ->
prev,
- (RApp (dummy_loc,
- (RRef (dummy_loc, delayed_force refl_ref)), [hole; RVar (dummy_loc, prev)])) :: vars
+ (GApp (dummy_loc,
+ (GRef (dummy_loc, delayed_force refl_ref)), [hole; GVar (dummy_loc, prev)])) :: vars
| _ ->
match na with
Anonymous -> raise (Invalid_argument "vars_of_ctx")
- | Name n -> n, RVar (dummy_loc, n) :: vars)
+ | Name n -> n, GVar (dummy_loc, n) :: vars)
ctx (id_of_string "vars_of_ctx_error", [])
in List.rev y
@@ -1741,13 +1741,13 @@ let constrs_of_pats typing_fun env isevars eqns tomatchs sign neqs arity =
let branch_name = id_of_string ("branch_" ^ (string_of_int !i)) in
let branch_decl = (Name branch_name, Some (lift !i bbody), (lift !i btype)) in
let branch =
- let bref = RVar (dummy_loc, branch_name) in
+ let bref = GVar (dummy_loc, branch_name) in
match vars_of_ctx rhs_rels with
[] -> bref
- | l -> RApp (dummy_loc, bref, l)
+ | l -> GApp (dummy_loc, bref, l)
in
let branch = match ineqs with
- Some _ -> RApp (dummy_loc, branch, [ hole ])
+ Some _ -> GApp (dummy_loc, branch, [ hole ])
| None -> branch
in
incr i;
diff --git a/plugins/subtac/subtac_command.ml b/plugins/subtac/subtac_command.ml
index 852776630..24fdd679b 100644
--- a/plugins/subtac/subtac_command.ml
+++ b/plugins/subtac/subtac_command.ml
@@ -84,16 +84,16 @@ let interp_constr_judgment isevars env c =
{ uj_val = evar_nf isevars j.uj_val; uj_type = evar_nf isevars j.uj_type }
let locate_if_isevar loc na = function
- | RHole _ ->
+ | GHole _ ->
(try match na with
- | Name id -> rawconstr_of_aconstr loc (Reserve.find_reserved_type id)
+ | Name id -> glob_constr_of_aconstr loc (Reserve.find_reserved_type id)
| Anonymous -> raise Not_found
- with Not_found -> RHole (loc, Evd.BinderType na))
+ with Not_found -> GHole (loc, Evd.BinderType na))
| x -> x
let interp_binder sigma env na t =
let t = Constrintern.intern_gen true ( !sigma) env t in
- SPretyping.understand_tcc_evars sigma env IsType (locate_if_isevar (loc_of_rawconstr t) na t)
+ SPretyping.understand_tcc_evars sigma env IsType (locate_if_isevar (loc_of_glob_constr t) na t)
let interp_context_evars evdref env params =
let bl = Constrintern.intern_context false !evdref env params in
@@ -102,7 +102,7 @@ let interp_context_evars evdref env params =
(fun (env,params,n,impls) (na, k, b, t) ->
match b with
None ->
- let t' = locate_if_isevar (loc_of_rawconstr t) na t in
+ let t' = locate_if_isevar (loc_of_glob_constr t) na t in
let t = SPretyping.understand_tcc_evars evdref env IsType t' in
let d = (na,None,t) in
let impls =
diff --git a/plugins/subtac/subtac_pretyping.ml b/plugins/subtac/subtac_pretyping.ml
index 2eb7408aa..09cc17328 100644
--- a/plugins/subtac/subtac_pretyping.ml
+++ b/plugins/subtac/subtac_pretyping.ml
@@ -58,7 +58,7 @@ let my_print_rec_info env t =
str "Wf proof: " ++ my_print_constr env t.wf_proof ++ spc () ++
str "Abbreviated Type: " ++ my_print_constr env t.f_type ++ spc () ++
str "Full type: " ++ my_print_constr env t.f_fulltype
-(* trace (str "pretype for " ++ (my_print_rawconstr env c) ++ *)
+(* trace (str "pretype for " ++ (my_print_glob_constr env c) ++ *)
(* str " and tycon "++ my_print_tycon env tycon ++ *)
(* str " in environment: " ++ my_print_env env); *)
@@ -81,9 +81,9 @@ let find_with_index x l =
open Vernacexpr
-let coqintern_constr evd env : Topconstr.constr_expr -> Rawterm.rawconstr =
+let coqintern_constr evd env : Topconstr.constr_expr -> Rawterm.glob_constr =
Constrintern.intern_constr evd env
-let coqintern_type evd env : Topconstr.constr_expr -> Rawterm.rawconstr =
+let coqintern_type evd env : Topconstr.constr_expr -> Rawterm.glob_constr =
Constrintern.intern_type evd env
let env_with_binders env isevars l =
diff --git a/plugins/subtac/subtac_pretyping.mli b/plugins/subtac/subtac_pretyping.mli
index 48906b23c..2b0c8fda2 100644
--- a/plugins/subtac/subtac_pretyping.mli
+++ b/plugins/subtac/subtac_pretyping.mli
@@ -13,7 +13,7 @@ module Pretyping : Pretyping.S
val interp :
Environ.env ->
Evd.evar_map ref ->
- Rawterm.rawconstr ->
+ Rawterm.glob_constr ->
Evarutil.type_constraint -> Term.constr * Term.constr
val subtac_process : ?is_type:bool -> env -> evar_map ref -> identifier -> local_binder list ->
diff --git a/plugins/subtac/subtac_pretyping_F.ml b/plugins/subtac/subtac_pretyping_F.ml
index 3d79508c3..debd4f053 100644
--- a/plugins/subtac/subtac_pretyping_F.ml
+++ b/plugins/subtac/subtac_pretyping_F.ml
@@ -189,22 +189,22 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
(* in environment [env], with existential variables [( evdref)] and *)
(* the type constraint tycon *)
let rec pretype (tycon : type_constraint) env evdref lvar c =
-(* let _ = try Subtac_utils.trace (str "pretype " ++ Subtac_utils.my_print_rawconstr env c ++ *)
+(* let _ = try Subtac_utils.trace (str "pretype " ++ Subtac_utils.my_print_glob_constr env c ++ *)
(* str " with tycon " ++ Evarutil.pr_tycon env tycon) *)
(* with _ -> () *)
(* in *)
match c with
- | RRef (loc,ref) ->
+ | GRef (loc,ref) ->
inh_conv_coerce_to_tycon loc env evdref
(pretype_ref evdref env ref)
tycon
- | RVar (loc, id) ->
+ | GVar (loc, id) ->
inh_conv_coerce_to_tycon loc env evdref
(pretype_id loc env !evdref lvar id)
tycon
- | REvar (loc, ev, instopt) ->
+ | GEvar (loc, ev, instopt) ->
(* Ne faudrait-il pas s'assurer que hyps est bien un
sous-contexte du contexte courant, et qu'il n'y a pas de Rel "caché" *)
let hyps = evar_context (Evd.find !evdref ev) in
@@ -215,10 +215,10 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
let j = (Retyping.get_judgment_of env !evdref c) in
inh_conv_coerce_to_tycon loc env evdref j tycon
- | RPatVar (loc,(someta,n)) ->
+ | GPatVar (loc,(someta,n)) ->
anomaly "Found a pattern variable in a rawterm to type"
- | RHole (loc,k) ->
+ | GHole (loc,k) ->
let ty =
match tycon with
| Some (None, ty) -> ty
@@ -226,7 +226,7 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
e_new_evar evdref env ~src:(loc, InternalHole) (Termops.new_Type ()) in
{ uj_val = e_new_evar evdref env ~src:(loc,k) ty; uj_type = ty }
- | RRec (loc,fixkind,names,bl,lar,vdef) ->
+ | GRec (loc,fixkind,names,bl,lar,vdef) ->
let rec type_bl env ctxt = function
[] -> ctxt
| (na,k,None,ty)::bl ->
@@ -306,10 +306,10 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
make_judge (mkCoFix cofix) ftys.(i) in
inh_conv_coerce_to_tycon loc env evdref fixj tycon
- | RSort (loc,s) ->
+ | GSort (loc,s) ->
inh_conv_coerce_to_tycon loc env evdref (pretype_sort s) tycon
- | RApp (loc,f,args) ->
+ | GApp (loc,f,args) ->
let length = List.length args in
let ftycon =
let ty =
@@ -326,11 +326,11 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
| _ -> None
in
let fj = pretype ftycon env evdref lvar f in
- let floc = loc_of_rawconstr f in
+ let floc = loc_of_glob_constr f in
let rec apply_rec env n resj tycon = function
| [] -> resj
| c::rest ->
- let argloc = loc_of_rawconstr c in
+ let argloc = loc_of_glob_constr c in
let resj = evd_comb1 (Coercion.inh_app_fun env) evdref resj in
let resty = whd_betadeltaiota env !evdref resj.uj_type in
match kind_of_term resty with
@@ -364,7 +364,7 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
| _ -> resj in
inh_conv_coerce_to_tycon loc env evdref resj tycon
- | RLambda(loc,name,k,c1,c2) ->
+ | GLambda(loc,name,k,c1,c2) ->
let tycon' = evd_comb1
(fun evd tycon ->
match tycon with
@@ -382,7 +382,7 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
let resj = judge_of_abstraction env name j j' in
inh_conv_coerce_to_tycon loc env evdref resj tycon
- | RProd(loc,name,k,c1,c2) ->
+ | GProd(loc,name,k,c1,c2) ->
let j = pretype_type empty_valcon env evdref lvar c1 in
let var = (name,j.utj_val) in
let env' = Termops.push_rel_assum var env in
@@ -392,7 +392,7 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
with TypeError _ as e -> Loc.raise loc e in
inh_conv_coerce_to_tycon loc env evdref resj tycon
- | RLetIn(loc,name,c1,c2) ->
+ | GLetIn(loc,name,c1,c2) ->
let j = pretype empty_tycon env evdref lvar c1 in
let t = Termops.refresh_universes j.uj_type in
let var = (name,Some j.uj_val,t) in
@@ -401,12 +401,12 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
{ uj_val = mkLetIn (name, j.uj_val, t, j'.uj_val) ;
uj_type = subst1 j.uj_val j'.uj_type }
- | RLetTuple (loc,nal,(na,po),c,d) ->
+ | GLetTuple (loc,nal,(na,po),c,d) ->
let cj = pretype empty_tycon env evdref lvar c in
let (IndType (indf,realargs)) =
try find_rectype env !evdref cj.uj_type
with Not_found ->
- let cloc = loc_of_rawconstr c in
+ let cloc = loc_of_glob_constr c in
error_case_not_inductive_loc cloc env !evdref cj
in
let cstrs = get_constructors env indf in
@@ -466,12 +466,12 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
in
{ uj_val = v; uj_type = ccl })
- | RIf (loc,c,(na,po),b1,b2) ->
+ | GIf (loc,c,(na,po),b1,b2) ->
let cj = pretype empty_tycon env evdref lvar c in
let (IndType (indf,realargs)) =
try find_rectype env !evdref cj.uj_type
with Not_found ->
- let cloc = loc_of_rawconstr c in
+ let cloc = loc_of_glob_constr c in
error_case_not_inductive_loc cloc env !evdref cj in
let cstrs = get_constructors env indf in
if Array.length cstrs <> 2 then
@@ -536,12 +536,12 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
in
{ uj_val = v; uj_type = p }
- | RCases (loc,sty,po,tml,eqns) ->
+ | GCases (loc,sty,po,tml,eqns) ->
Cases.compile_cases loc sty
((fun vtyc env evdref -> pretype vtyc env evdref lvar),evdref)
tycon env (* loc *) (po,tml,eqns)
- | RCast (loc,c,k) ->
+ | GCast (loc,c,k) ->
let cj =
match k with
CastCoerce ->
@@ -557,7 +557,7 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
in
inh_conv_coerce_to_tycon loc env evdref cj tycon
- | RDynamic (loc,d) ->
+ | GDynamic (loc,d) ->
if (Dyn.tag d) = "constr" then
let c = constr_out d in
let j = (Retyping.get_judgment_of env !evdref c) in
@@ -568,7 +568,7 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
(* [pretype_type valcon env evdref lvar c] coerces [c] into a type *)
and pretype_type valcon env evdref lvar = function
- | RHole loc ->
+ | GHole loc ->
(match valcon with
| Some v ->
let s =
@@ -588,7 +588,7 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
utj_type = s})
| c ->
let j = pretype empty_tycon env evdref lvar c in
- let loc = loc_of_rawconstr c in
+ let loc = loc_of_glob_constr c in
let tj = evd_comb1 (Coercion.inh_coerce_to_sort loc env) evdref j in
match valcon with
| None -> tj
@@ -596,7 +596,7 @@ module SubtacPretyping_F (Coercion : Coercion.S) = struct
if e_cumul env evdref v tj.utj_val then tj
else
error_unexpected_type_loc
- (loc_of_rawconstr c) env !evdref tj.utj_val v
+ (loc_of_glob_constr c) env !evdref tj.utj_val v
let pretype_gen expand_evar fail_evar resolve_classes evdref env lvar kind c =
let c' = match kind with
diff --git a/plugins/subtac/subtac_utils.ml b/plugins/subtac/subtac_utils.ml
index 106ac4d09..43b55ca95 100644
--- a/plugins/subtac/subtac_utils.ml
+++ b/plugins/subtac/subtac_utils.ml
@@ -116,7 +116,7 @@ let my_print_rel_context env ctx = Printer.pr_rel_context env ctx
let my_print_context = Termops.print_rel_context
let my_print_named_context = Termops.print_named_context
let my_print_env = Termops.print_env
-let my_print_rawconstr = Printer.pr_rawconstr_env
+let my_print_glob_constr = Printer.pr_glob_constr_env
let my_print_evardefs = Evd.pr_evar_map
let my_print_tycon_type = Evarutil.pr_tycon_type
diff --git a/plugins/subtac/subtac_utils.mli b/plugins/subtac/subtac_utils.mli
index f56c2932e..659a67781 100644
--- a/plugins/subtac/subtac_utils.mli
+++ b/plugins/subtac/subtac_utils.mli
@@ -74,7 +74,7 @@ val my_print_context : env -> std_ppcmds
val my_print_rel_context : env -> rel_context -> std_ppcmds
val my_print_named_context : env -> std_ppcmds
val my_print_env : env -> std_ppcmds
-val my_print_rawconstr : env -> rawconstr -> std_ppcmds
+val my_print_glob_constr : env -> glob_constr -> std_ppcmds
val my_print_tycon_type : env -> type_constraint_type -> std_ppcmds
diff --git a/plugins/syntax/ascii_syntax.ml b/plugins/syntax/ascii_syntax.ml
index 79e1a6eb2..50498f0f4 100644
--- a/plugins/syntax/ascii_syntax.ml
+++ b/plugins/syntax/ascii_syntax.ml
@@ -39,9 +39,9 @@ let interp_ascii dloc p =
let rec aux n p =
if n = 0 then [] else
let mp = p mod 2 in
- RRef (dloc,if mp = 0 then glob_false else glob_true)
+ GRef (dloc,if mp = 0 then glob_false else glob_true)
:: (aux (n-1) (p/2)) in
- RApp (dloc,RRef(dloc,force glob_Ascii), aux 8 p)
+ GApp (dloc,GRef(dloc,force glob_Ascii), aux 8 p)
let interp_ascii_string dloc s =
let p =
@@ -57,12 +57,12 @@ let interp_ascii_string dloc s =
let uninterp_ascii r =
let rec uninterp_bool_list n = function
| [] when n = 0 -> 0
- | RRef (_,k)::l when k = glob_true -> 1+2*(uninterp_bool_list (n-1) l)
- | RRef (_,k)::l when k = glob_false -> 2*(uninterp_bool_list (n-1) l)
+ | GRef (_,k)::l when k = glob_true -> 1+2*(uninterp_bool_list (n-1) l)
+ | GRef (_,k)::l when k = glob_false -> 2*(uninterp_bool_list (n-1) l)
| _ -> raise Non_closed_ascii in
try
let rec aux = function
- | RApp (_,RRef (_,k),l) when k = force glob_Ascii -> uninterp_bool_list 8 l
+ | GApp (_,GRef (_,k),l) when k = force glob_Ascii -> uninterp_bool_list 8 l
| _ -> raise Non_closed_ascii in
Some (aux r)
with
@@ -78,4 +78,4 @@ let _ =
Notation.declare_string_interpreter "char_scope"
(ascii_path,ascii_module)
interp_ascii_string
- ([RRef (dummy_loc,static_glob_Ascii)], uninterp_ascii_string, true)
+ ([GRef (dummy_loc,static_glob_Ascii)], uninterp_ascii_string, true)
diff --git a/plugins/syntax/nat_syntax.ml b/plugins/syntax/nat_syntax.ml
index b673fdb9d..3d8860be4 100644
--- a/plugins/syntax/nat_syntax.ml
+++ b/plugins/syntax/nat_syntax.ml
@@ -36,11 +36,11 @@ let nat_of_int dloc n =
strbrk "working with large numbers in nat (observed threshold " ++
strbrk "may vary from 5000 to 70000 depending on your system " ++
strbrk "limits and on the command executed).");
- let ref_O = RRef (dloc, glob_O) in
- let ref_S = RRef (dloc, glob_S) in
+ let ref_O = GRef (dloc, glob_O) in
+ let ref_S = GRef (dloc, glob_S) in
let rec mk_nat acc n =
if n <> zero then
- mk_nat (RApp (dloc,ref_S, [acc])) (sub_1 n)
+ mk_nat (GApp (dloc,ref_S, [acc])) (sub_1 n)
else
acc
in
@@ -56,8 +56,8 @@ let nat_of_int dloc n =
exception Non_closed_number
let rec int_of_nat = function
- | RApp (_,RRef (_,s),[a]) when s = glob_S -> add_1 (int_of_nat a)
- | RRef (_,z) when z = glob_O -> zero
+ | GApp (_,GRef (_,s),[a]) when s = glob_S -> add_1 (int_of_nat a)
+ | GRef (_,z) when z = glob_O -> zero
| _ -> raise Non_closed_number
let uninterp_nat p =
@@ -73,4 +73,4 @@ let _ =
Notation.declare_numeral_interpreter "nat_scope"
(nat_path,["Coq";"Init";"Datatypes"])
nat_of_int
- ([RRef (dummy_loc,glob_S); RRef (dummy_loc,glob_O)], uninterp_nat, true)
+ ([GRef (dummy_loc,glob_S); GRef (dummy_loc,glob_O)], uninterp_nat, true)
diff --git a/plugins/syntax/numbers_syntax.ml b/plugins/syntax/numbers_syntax.ml
index 2a0425dca..892a5595a 100644
--- a/plugins/syntax/numbers_syntax.ml
+++ b/plugins/syntax/numbers_syntax.ml
@@ -98,9 +98,9 @@ exception Non_closed
(* parses a *non-negative* integer (from bigint.ml) into an int31
wraps modulo 2^31 *)
let int31_of_pos_bigint dloc n =
- let ref_construct = RRef (dloc, int31_construct) in
- let ref_0 = RRef (dloc, int31_0) in
- let ref_1 = RRef (dloc, int31_1) in
+ let ref_construct = GRef (dloc, int31_construct) in
+ let ref_0 = GRef (dloc, int31_0) in
+ let ref_1 = GRef (dloc, int31_1) in
let rec args counter n =
if counter <= 0 then
[]
@@ -108,7 +108,7 @@ let int31_of_pos_bigint dloc n =
let (q,r) = div2_with_rest n in
(if r then ref_1 else ref_0)::(args (counter-1) q)
in
- RApp (dloc, ref_construct, List.rev (args 31 n))
+ GApp (dloc, ref_construct, List.rev (args 31 n))
let error_negative dloc =
Util.user_err_loc (dloc, "interp_int31", Pp.str "int31 are only non-negative numbers.")
@@ -125,12 +125,12 @@ let bigint_of_int31 =
let rec args_parsing args cur =
match args with
| [] -> cur
- | (RRef (_,b))::l when b = int31_0 -> args_parsing l (mult_2 cur)
- | (RRef (_,b))::l when b = int31_1 -> args_parsing l (add_1 (mult_2 cur))
+ | (GRef (_,b))::l when b = int31_0 -> args_parsing l (mult_2 cur)
+ | (GRef (_,b))::l when b = int31_1 -> args_parsing l (add_1 (mult_2 cur))
| _ -> raise Non_closed
in
function
- | RApp (_, RRef (_, c), args) when c=int31_construct -> args_parsing args zero
+ | GApp (_, GRef (_, c), args) when c=int31_construct -> args_parsing args zero
| _ -> raise Non_closed
let uninterp_int31 i =
@@ -143,7 +143,7 @@ let uninterp_int31 i =
let _ = Notation.declare_numeral_interpreter int31_scope
(int31_path, int31_module)
interp_int31
- ([RRef (Util.dummy_loc, int31_construct)],
+ ([GRef (Util.dummy_loc, int31_construct)],
uninterp_int31,
true)
@@ -174,24 +174,24 @@ let height bi =
(* n must be a non-negative integer (from bigint.ml) *)
let word_of_pos_bigint dloc hght n =
- let ref_W0 = RRef (dloc, zn2z_W0) in
- let ref_WW = RRef (dloc, zn2z_WW) in
+ let ref_W0 = GRef (dloc, zn2z_W0) in
+ let ref_WW = GRef (dloc, zn2z_WW) in
let rec decomp hgt n =
if is_neg_or_zero hgt then
int31_of_pos_bigint dloc n
else if equal n zero then
- RApp (dloc, ref_W0, [RHole (dloc, Evd.InternalHole)])
+ GApp (dloc, ref_W0, [GHole (dloc, Evd.InternalHole)])
else
let (h,l) = split_at hgt n in
- RApp (dloc, ref_WW, [RHole (dloc, Evd.InternalHole);
+ GApp (dloc, ref_WW, [GHole (dloc, Evd.InternalHole);
decomp (sub_1 hgt) h;
decomp (sub_1 hgt) l])
in
decomp hght n
let bigN_of_pos_bigint dloc n =
- let ref_constructor i = RRef (dloc, bigN_constructor i) in
- let result h word = RApp (dloc, ref_constructor h, if less_than h n_inlined then
+ let ref_constructor i = GRef (dloc, bigN_constructor i) in
+ let result h word = GApp (dloc, ref_constructor h, if less_than h n_inlined then
[word]
else
[Nat_syntax.nat_of_int dloc (sub h n_inlined);
@@ -215,7 +215,7 @@ let interp_bigN dloc n =
let bigint_of_word =
let rec get_height rc =
match rc with
- | RApp (_,RRef(_,c), [_;lft;rght]) when c = zn2z_WW ->
+ | GApp (_,GRef(_,c), [_;lft;rght]) when c = zn2z_WW ->
let hleft = get_height lft in
let hright = get_height rght in
add_1
@@ -227,8 +227,8 @@ let bigint_of_word =
in
let rec transform hght rc =
match rc with
- | RApp (_,RRef(_,c),_) when c = zn2z_W0-> zero
- | RApp (_,RRef(_,c), [_;lft;rght]) when c=zn2z_WW-> let new_hght = sub_1 hght in
+ | GApp (_,GRef(_,c),_) when c = zn2z_W0-> zero
+ | GApp (_,GRef(_,c), [_;lft;rght]) when c=zn2z_WW-> let new_hght = sub_1 hght in
add (mult (rank new_hght)
(transform (new_hght) lft))
(transform (new_hght) rght)
@@ -240,8 +240,8 @@ let bigint_of_word =
let bigint_of_bigN rc =
match rc with
- | RApp (_,_,[one_arg]) -> bigint_of_word one_arg
- | RApp (_,_,[_;second_arg]) -> bigint_of_word second_arg
+ | GApp (_,_,[one_arg]) -> bigint_of_word one_arg
+ | GApp (_,_,[_;second_arg]) -> bigint_of_word second_arg
| _ -> raise Non_closed
let uninterp_bigN rc =
@@ -257,7 +257,7 @@ let uninterp_bigN rc =
let bigN_list_of_constructors =
let rec build i =
if less_than i (add_1 n_inlined) then
- RRef (Util.dummy_loc, bigN_constructor i)::(build (add_1 i))
+ GRef (Util.dummy_loc, bigN_constructor i)::(build (add_1 i))
else
[]
in
@@ -274,17 +274,17 @@ let _ = Notation.declare_numeral_interpreter bigN_scope
(*** Parsing for bigZ in digital notation ***)
let interp_bigZ dloc n =
- let ref_pos = RRef (dloc, bigZ_pos) in
- let ref_neg = RRef (dloc, bigZ_neg) in
+ let ref_pos = GRef (dloc, bigZ_pos) in
+ let ref_neg = GRef (dloc, bigZ_neg) in
if is_pos_or_zero n then
- RApp (dloc, ref_pos, [bigN_of_pos_bigint dloc n])
+ GApp (dloc, ref_pos, [bigN_of_pos_bigint dloc n])
else
- RApp (dloc, ref_neg, [bigN_of_pos_bigint dloc (neg n)])
+ GApp (dloc, ref_neg, [bigN_of_pos_bigint dloc (neg n)])
(* pretty printing functions for bigZ *)
let bigint_of_bigZ = function
- | RApp (_, RRef(_,c), [one_arg]) when c = bigZ_pos -> bigint_of_bigN one_arg
- | RApp (_, RRef(_,c), [one_arg]) when c = bigZ_neg ->
+ | GApp (_, GRef(_,c), [one_arg]) when c = bigZ_pos -> bigint_of_bigN one_arg
+ | GApp (_, GRef(_,c), [one_arg]) when c = bigZ_neg ->
let opp_val = bigint_of_bigN one_arg in
if equal opp_val zero then
raise Non_closed
@@ -303,19 +303,19 @@ let uninterp_bigZ rc =
let _ = Notation.declare_numeral_interpreter bigZ_scope
(bigZ_path, bigZ_module)
interp_bigZ
- ([RRef (Util.dummy_loc, bigZ_pos);
- RRef (Util.dummy_loc, bigZ_neg)],
+ ([GRef (Util.dummy_loc, bigZ_pos);
+ GRef (Util.dummy_loc, bigZ_neg)],
uninterp_bigZ,
true)
(*** Parsing for bigQ in digital notation ***)
let interp_bigQ dloc n =
- let ref_z = RRef (dloc, bigQ_z) in
- RApp (dloc, ref_z, [interp_bigZ dloc n])
+ let ref_z = GRef (dloc, bigQ_z) in
+ GApp (dloc, ref_z, [interp_bigZ dloc n])
let uninterp_bigQ rc =
try match rc with
- | RApp (_, RRef(_,c), [one_arg]) when c = bigQ_z ->
+ | GApp (_, GRef(_,c), [one_arg]) when c = bigQ_z ->
Some (bigint_of_bigZ one_arg)
| _ -> None (* we don't pretty-print yet fractions *)
with Non_closed -> None
@@ -324,5 +324,5 @@ let uninterp_bigQ rc =
let _ = Notation.declare_numeral_interpreter bigQ_scope
(bigQ_path, bigQ_module)
interp_bigQ
- ([RRef (Util.dummy_loc, bigQ_z)], uninterp_bigQ,
+ ([GRef (Util.dummy_loc, bigQ_z)], uninterp_bigQ,
true)
diff --git a/plugins/syntax/r_syntax.ml b/plugins/syntax/r_syntax.ml
index 7f093ae95..fc953e5e5 100644
--- a/plugins/syntax/r_syntax.ml
+++ b/plugins/syntax/r_syntax.ml
@@ -46,24 +46,24 @@ let four = mult_2 two
(* Unary representation of strictly positive numbers *)
let rec small_r dloc n =
- if equal one n then RRef (dloc, glob_R1)
- else RApp(dloc,RRef (dloc,glob_Rplus),
- [RRef (dloc, glob_R1);small_r dloc (sub_1 n)])
+ if equal one n then GRef (dloc, glob_R1)
+ else GApp(dloc,GRef (dloc,glob_Rplus),
+ [GRef (dloc, glob_R1);small_r dloc (sub_1 n)])
let r_of_posint dloc n =
- let r1 = RRef (dloc, glob_R1) in
+ let r1 = GRef (dloc, glob_R1) in
let r2 = small_r dloc two in
let rec r_of_pos n =
if less_than n four then small_r dloc n
else
let (q,r) = div2_with_rest n in
- let b = RApp(dloc,RRef(dloc,glob_Rmult),[r2;r_of_pos q]) in
- if r then RApp(dloc,RRef(dloc,glob_Rplus),[r1;b]) else b in
- if n <> zero then r_of_pos n else RRef(dloc,glob_R0)
+ let b = GApp(dloc,GRef(dloc,glob_Rmult),[r2;r_of_pos q]) in
+ if r then GApp(dloc,GRef(dloc,glob_Rplus),[r1;b]) else b in
+ if n <> zero then r_of_pos n else GRef(dloc,glob_R0)
let r_of_int dloc z =
if is_strictly_neg z then
- RApp (dloc, RRef(dloc,glob_Ropp), [r_of_posint dloc (neg z)])
+ GApp (dloc, GRef(dloc,glob_Ropp), [r_of_posint dloc (neg z)])
else
r_of_posint dloc z
@@ -75,33 +75,33 @@ let bignat_of_r =
(* for numbers > 1 *)
let rec bignat_of_pos = function
(* 1+1 *)
- | RApp (_,RRef (_,p), [RRef (_,o1); RRef (_,o2)])
+ | GApp (_,GRef (_,p), [GRef (_,o1); GRef (_,o2)])
when p = glob_Rplus & o1 = glob_R1 & o2 = glob_R1 -> two
(* 1+(1+1) *)
- | RApp (_,RRef (_,p1), [RRef (_,o1);
- RApp(_,RRef (_,p2),[RRef(_,o2);RRef(_,o3)])])
+ | GApp (_,GRef (_,p1), [GRef (_,o1);
+ GApp(_,GRef (_,p2),[GRef(_,o2);GRef(_,o3)])])
when p1 = glob_Rplus & p2 = glob_Rplus &
o1 = glob_R1 & o2 = glob_R1 & o3 = glob_R1 -> three
(* (1+1)*b *)
- | RApp (_,RRef (_,p), [a; b]) when p = glob_Rmult ->
+ | GApp (_,GRef (_,p), [a; b]) when p = glob_Rmult ->
if bignat_of_pos a <> two then raise Non_closed_number;
mult_2 (bignat_of_pos b)
(* 1+(1+1)*b *)
- | RApp (_,RRef (_,p1), [RRef (_,o); RApp (_,RRef (_,p2),[a;b])])
+ | GApp (_,GRef (_,p1), [GRef (_,o); GApp (_,GRef (_,p2),[a;b])])
when p1 = glob_Rplus & p2 = glob_Rmult & o = glob_R1 ->
if bignat_of_pos a <> two then raise Non_closed_number;
add_1 (mult_2 (bignat_of_pos b))
| _ -> raise Non_closed_number
in
let bignat_of_r = function
- | RRef (_,a) when a = glob_R0 -> zero
- | RRef (_,a) when a = glob_R1 -> one
+ | GRef (_,a) when a = glob_R0 -> zero
+ | GRef (_,a) when a = glob_R1 -> one
| r -> bignat_of_pos r
in
bignat_of_r
let bigint_of_r = function
- | RApp (_,RRef (_,o), [a]) when o = glob_Ropp ->
+ | GApp (_,GRef (_,o), [a]) when o = glob_Ropp ->
let n = bignat_of_r a in
if n = zero then raise Non_closed_number;
neg n
@@ -116,8 +116,8 @@ let uninterp_r p =
let _ = Notation.declare_numeral_interpreter "R_scope"
(r_path,["Coq";"Reals";"Rdefinitions"])
r_of_int
- ([RRef(dummy_loc,glob_Ropp);RRef(dummy_loc,glob_R0);
- RRef(dummy_loc,glob_Rplus);RRef(dummy_loc,glob_Rmult);
- RRef(dummy_loc,glob_R1)],
+ ([GRef(dummy_loc,glob_Ropp);GRef(dummy_loc,glob_R0);
+ GRef(dummy_loc,glob_Rplus);GRef(dummy_loc,glob_Rmult);
+ GRef(dummy_loc,glob_R1)],
uninterp_r,
false)
diff --git a/plugins/syntax/string_syntax.ml b/plugins/syntax/string_syntax.ml
index cb2b16858..61643881e 100644
--- a/plugins/syntax/string_syntax.ml
+++ b/plugins/syntax/string_syntax.ml
@@ -37,8 +37,8 @@ open Lazy
let interp_string dloc s =
let le = String.length s in
let rec aux n =
- if n = le then RRef (dloc, force glob_EmptyString) else
- RApp (dloc,RRef (dloc, force glob_String),
+ if n = le then GRef (dloc, force glob_EmptyString) else
+ GApp (dloc,GRef (dloc, force glob_String),
[interp_ascii dloc (int_of_char s.[n]); aux (n+1)])
in aux 0
@@ -46,11 +46,11 @@ let uninterp_string r =
try
let b = Buffer.create 16 in
let rec aux = function
- | RApp (_,RRef (_,k),[a;s]) when k = force glob_String ->
+ | GApp (_,GRef (_,k),[a;s]) when k = force glob_String ->
(match uninterp_ascii a with
| Some c -> Buffer.add_char b (Char.chr c); aux s
| _ -> raise Non_closed_string)
- | RRef (_,z) when z = force glob_EmptyString ->
+ | GRef (_,z) when z = force glob_EmptyString ->
Some (Buffer.contents b)
| _ ->
raise Non_closed_string
@@ -62,6 +62,6 @@ let _ =
Notation.declare_string_interpreter "string_scope"
(string_path,["Coq";"Strings";"String"])
interp_string
- ([RRef (dummy_loc,static_glob_String);
- RRef (dummy_loc,static_glob_EmptyString)],
+ ([GRef (dummy_loc,static_glob_String);
+ GRef (dummy_loc,static_glob_EmptyString)],
uninterp_string, true)
diff --git a/plugins/syntax/z_syntax.ml b/plugins/syntax/z_syntax.ml
index fb8de1f92..81588bced 100644
--- a/plugins/syntax/z_syntax.ml
+++ b/plugins/syntax/z_syntax.ml
@@ -42,13 +42,13 @@ let glob_xO = ConstructRef path_of_xO
let glob_xH = ConstructRef path_of_xH
let pos_of_bignat dloc x =
- let ref_xI = RRef (dloc, glob_xI) in
- let ref_xH = RRef (dloc, glob_xH) in
- let ref_xO = RRef (dloc, glob_xO) in
+ let ref_xI = GRef (dloc, glob_xI) in
+ let ref_xH = GRef (dloc, glob_xH) in
+ let ref_xO = GRef (dloc, glob_xO) in
let rec pos_of x =
match div2_with_rest x with
- | (q,false) -> RApp (dloc, ref_xO,[pos_of q])
- | (q,true) when q <> zero -> RApp (dloc,ref_xI,[pos_of q])
+ | (q,false) -> GApp (dloc, ref_xO,[pos_of q])
+ | (q,true) when q <> zero -> GApp (dloc,ref_xI,[pos_of q])
| (q,true) -> ref_xH
in
pos_of x
@@ -66,9 +66,9 @@ let interp_positive dloc n =
(**********************************************************************)
let rec bignat_of_pos = function
- | RApp (_, RRef (_,b),[a]) when b = glob_xO -> mult_2(bignat_of_pos a)
- | RApp (_, RRef (_,b),[a]) when b = glob_xI -> add_1(mult_2(bignat_of_pos a))
- | RRef (_, a) when a = glob_xH -> Bigint.one
+ | GApp (_, GRef (_,b),[a]) when b = glob_xO -> mult_2(bignat_of_pos a)
+ | GApp (_, GRef (_,b),[a]) when b = glob_xI -> add_1(mult_2(bignat_of_pos a))
+ | GRef (_, a) when a = glob_xH -> Bigint.one
| _ -> raise Non_closed_number
let uninterp_positive p =
@@ -84,9 +84,9 @@ let uninterp_positive p =
let _ = Notation.declare_numeral_interpreter "positive_scope"
(positive_path,positive_module)
interp_positive
- ([RRef (dummy_loc, glob_xI);
- RRef (dummy_loc, glob_xO);
- RRef (dummy_loc, glob_xH)],
+ ([GRef (dummy_loc, glob_xI);
+ GRef (dummy_loc, glob_xO);
+ GRef (dummy_loc, glob_xH)],
uninterp_positive,
true)
@@ -106,9 +106,9 @@ let n_path = make_path binnat_module "N"
let n_of_binnat dloc pos_or_neg n =
if n <> zero then
- RApp(dloc, RRef (dloc,glob_Npos), [pos_of_bignat dloc n])
+ GApp(dloc, GRef (dloc,glob_Npos), [pos_of_bignat dloc n])
else
- RRef (dloc, glob_N0)
+ GRef (dloc, glob_N0)
let error_negative dloc =
user_err_loc (dloc, "interp_N", str "No negative numbers in type \"N\".")
@@ -122,8 +122,8 @@ let n_of_int dloc n =
(**********************************************************************)
let bignat_of_n = function
- | RApp (_, RRef (_,b),[a]) when b = glob_Npos -> bignat_of_pos a
- | RRef (_, a) when a = glob_N0 -> Bigint.zero
+ | GApp (_, GRef (_,b),[a]) when b = glob_Npos -> bignat_of_pos a
+ | GRef (_, a) when a = glob_N0 -> Bigint.zero
| _ -> raise Non_closed_number
let uninterp_n p =
@@ -136,8 +136,8 @@ let uninterp_n p =
let _ = Notation.declare_numeral_interpreter "N_scope"
(n_path,binnat_module)
n_of_int
- ([RRef (dummy_loc, glob_N0);
- RRef (dummy_loc, glob_Npos)],
+ ([GRef (dummy_loc, glob_N0);
+ GRef (dummy_loc, glob_Npos)],
uninterp_n,
true)
@@ -160,18 +160,18 @@ let z_of_int dloc n =
if n <> zero then
let sgn, n =
if is_pos_or_zero n then glob_POS, n else glob_NEG, Bigint.neg n in
- RApp(dloc, RRef (dloc,sgn), [pos_of_bignat dloc n])
+ GApp(dloc, GRef (dloc,sgn), [pos_of_bignat dloc n])
else
- RRef (dloc, glob_ZERO)
+ GRef (dloc, glob_ZERO)
(**********************************************************************)
(* Printing Z via scopes *)
(**********************************************************************)
let bigint_of_z = function
- | RApp (_, RRef (_,b),[a]) when b = glob_POS -> bignat_of_pos a
- | RApp (_, RRef (_,b),[a]) when b = glob_NEG -> Bigint.neg (bignat_of_pos a)
- | RRef (_, a) when a = glob_ZERO -> Bigint.zero
+ | GApp (_, GRef (_,b),[a]) when b = glob_POS -> bignat_of_pos a
+ | GApp (_, GRef (_,b),[a]) when b = glob_NEG -> Bigint.neg (bignat_of_pos a)
+ | GRef (_, a) when a = glob_ZERO -> Bigint.zero
| _ -> raise Non_closed_number
let uninterp_z p =
@@ -185,8 +185,8 @@ let uninterp_z p =
let _ = Notation.declare_numeral_interpreter "Z_scope"
(z_path,binint_module)
z_of_int
- ([RRef (dummy_loc, glob_ZERO);
- RRef (dummy_loc, glob_POS);
- RRef (dummy_loc, glob_NEG)],
+ ([GRef (dummy_loc, glob_ZERO);
+ GRef (dummy_loc, glob_POS);
+ GRef (dummy_loc, glob_NEG)],
uninterp_z,
true)
diff --git a/pretyping/cases.ml b/pretyping/cases.ml
index c61edbc55..5115b3e3b 100644
--- a/pretyping/cases.ml
+++ b/pretyping/cases.ml
@@ -66,9 +66,9 @@ let error_needs_inversion env x t =
module type S = sig
val compile_cases :
loc -> case_style ->
- (type_constraint -> env -> evar_map ref -> rawconstr -> unsafe_judgment) * evar_map ref ->
+ (type_constraint -> env -> evar_map ref -> glob_constr -> unsafe_judgment) * evar_map ref ->
type_constraint ->
- env -> rawconstr option * tomatch_tuples * cases_clauses ->
+ env -> glob_constr option * tomatch_tuples * cases_clauses ->
unsafe_judgment
end
@@ -348,7 +348,7 @@ let find_tomatch_tycon evdref env loc = function
empty_tycon,None
let coerce_row typing_fun evdref env pats (tomatch,(_,indopt)) =
- let loc = Some (loc_of_rawconstr tomatch) in
+ let loc = Some (loc_of_glob_constr tomatch) in
let tycon,realnames = find_tomatch_tycon evdref env loc indopt in
let j = typing_fun tycon env evdref tomatch in
let typ = nf_evar !evdref j.uj_type in
@@ -1228,7 +1228,7 @@ let matx_of_eqns env tomatchl eqns =
let initial_rhs = rhs in
let rhs =
{ rhs_env = env;
- rhs_vars = free_rawvars initial_rhs;
+ rhs_vars = free_glob_vars initial_rhs;
avoid_ids = ids@(ids_of_named_context (named_context env));
it = Some initial_rhs } in
{ patterns = initial_lpat;
diff --git a/pretyping/cases.mli b/pretyping/cases.mli
index 2facf8696..015b386a5 100644
--- a/pretyping/cases.mli
+++ b/pretyping/cases.mli
@@ -61,9 +61,9 @@ type tomatch_status =
module type S = sig
val compile_cases :
loc -> case_style ->
- (type_constraint -> env -> evar_map ref -> rawconstr -> unsafe_judgment) * evar_map ref ->
+ (type_constraint -> env -> evar_map ref -> glob_constr -> unsafe_judgment) * evar_map ref ->
type_constraint ->
- env -> rawconstr option * tomatch_tuples * cases_clauses ->
+ env -> glob_constr option * tomatch_tuples * cases_clauses ->
unsafe_judgment
end
diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml
index 04dc13290..7469111bf 100644
--- a/pretyping/detyping.ml
+++ b/pretyping/detyping.ml
@@ -275,30 +275,30 @@ let is_nondep_branch c n =
let extract_nondep_branches test c b n =
let rec strip n r = if n=0 then r else
match r with
- | RLambda (_,_,_,_,t) -> strip (n-1) t
- | RLetIn (_,_,_,t) -> strip (n-1) t
+ | GLambda (_,_,_,_,t) -> strip (n-1) t
+ | GLetIn (_,_,_,t) -> strip (n-1) t
| _ -> assert false in
if test c n then Some (strip n b) else None
let it_destRLambda_or_LetIn_names n c =
let rec aux n nal c =
if n=0 then (List.rev nal,c) else match c with
- | RLambda (_,na,_,_,c) -> aux (n-1) (na::nal) c
- | RLetIn (_,na,_,c) -> aux (n-1) (na::nal) c
+ | GLambda (_,na,_,_,c) -> aux (n-1) (na::nal) c
+ | GLetIn (_,na,_,c) -> aux (n-1) (na::nal) c
| _ ->
(* eta-expansion *)
let rec next l =
let x = next_ident_away (id_of_string "x") l in
- (* Not efficient but unusual and no function to get free rawvars *)
-(* if occur_rawconstr x c then next (x::l) else x in *)
+ (* Not efficient but unusual and no function to get free glob_vars *)
+(* if occur_glob_constr x c then next (x::l) else x in *)
x
in
- let x = next (free_rawvars c) in
- let a = RVar (dl,x) in
+ let x = next (free_glob_vars c) in
+ let a = GVar (dl,x) in
aux (n-1) (Name x :: nal)
(match c with
- | RApp (loc,p,l) -> RApp (loc,c,l@[a])
- | _ -> (RApp (dl,c,[a])))
+ | GApp (loc,p,l) -> GApp (loc,c,l@[a])
+ | _ -> (GApp (dl,c,[a])))
in aux n [] c
let detype_case computable detype detype_eqns testdep avoid data p c bl =
@@ -315,7 +315,7 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl =
| Some p ->
let nl,typ = it_destRLambda_or_LetIn_names k p in
let n,typ = match typ with
- | RLambda (_,x,_,t,c) -> x, c
+ | GLambda (_,x,_,t,c) -> x, c
| _ -> Anonymous, typ in
let aliastyp =
if List.for_all ((=) Anonymous) nl then None
@@ -341,20 +341,20 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl =
| LetStyle when aliastyp = None ->
let bl' = Array.map detype bl in
let (nal,d) = it_destRLambda_or_LetIn_names consnargsl.(0) bl'.(0) in
- RLetTuple (dl,nal,(alias,pred),tomatch,d)
+ GLetTuple (dl,nal,(alias,pred),tomatch,d)
| IfStyle when aliastyp = None ->
let bl' = Array.map detype bl in
let nondepbrs =
array_map3 (extract_nondep_branches testdep) bl bl' consnargsl in
if array_for_all ((<>) None) nondepbrs then
- RIf (dl,tomatch,(alias,pred),
+ GIf (dl,tomatch,(alias,pred),
Option.get nondepbrs.(0),Option.get nondepbrs.(1))
else
let eqnl = detype_eqns constructs consnargsl bl in
- RCases (dl,tag,pred,[tomatch,(alias,aliastyp)],eqnl)
+ GCases (dl,tag,pred,[tomatch,(alias,aliastyp)],eqnl)
| _ ->
let eqnl = detype_eqns constructs consnargsl bl in
- RCases (dl,tag,pred,[tomatch,(alias,aliastyp)],eqnl)
+ GCases (dl,tag,pred,[tomatch,(alias,aliastyp)],eqnl)
let detype_sort = function
| Prop c -> RProp c
@@ -372,36 +372,36 @@ let rec detype (isgoal:bool) avoid env t =
match kind_of_term (collapse_appl t) with
| Rel n ->
(try match lookup_name_of_rel n env with
- | Name id -> RVar (dl, id)
+ | Name id -> GVar (dl, id)
| Anonymous -> !detype_anonymous dl n
with Not_found ->
let s = "_UNBOUND_REL_"^(string_of_int n)
- in RVar (dl, id_of_string s))
+ in GVar (dl, id_of_string s))
| Meta n ->
(* Meta in constr are not user-parsable and are mapped to Evar *)
- REvar (dl, n, None)
+ GEvar (dl, n, None)
| Var id ->
(try
- let _ = Global.lookup_named id in RRef (dl, VarRef id)
+ let _ = Global.lookup_named id in GRef (dl, VarRef id)
with _ ->
- RVar (dl, id))
- | Sort s -> RSort (dl,detype_sort s)
+ GVar (dl, id))
+ | Sort s -> GSort (dl,detype_sort s)
| Cast (c1,k,c2) ->
- RCast(dl,detype isgoal avoid env c1, CastConv (k, detype isgoal avoid env c2))
+ GCast(dl,detype isgoal avoid env c1, CastConv (k, detype isgoal avoid env c2))
| Prod (na,ty,c) -> detype_binder isgoal BProd avoid env na ty c
| Lambda (na,ty,c) -> detype_binder isgoal BLambda avoid env na ty c
| LetIn (na,b,_,c) -> detype_binder isgoal BLetIn avoid env na b c
| App (f,args) ->
- RApp (dl,detype isgoal avoid env f,
+ GApp (dl,detype isgoal avoid env f,
array_map_to_list (detype isgoal avoid env) args)
- | Const sp -> RRef (dl, ConstRef sp)
+ | Const sp -> GRef (dl, ConstRef sp)
| Evar (ev,cl) ->
- REvar (dl, ev,
+ GEvar (dl, ev,
Some (List.map (detype isgoal avoid env) (Array.to_list cl)))
| Ind ind_sp ->
- RRef (dl, IndRef ind_sp)
+ GRef (dl, IndRef ind_sp)
| Construct cstr_sp ->
- RRef (dl, ConstructRef cstr_sp)
+ GRef (dl, ConstructRef cstr_sp)
| Case (ci,p,c,bl) ->
let comp = computable p (ci.ci_pp_info.ind_nargs) in
detype_case comp (detype isgoal avoid env)
@@ -424,7 +424,7 @@ and detype_fix isgoal avoid env (vn,_ as nvn) (names,tys,bodies) =
let v = array_map3
(fun c t i -> share_names isgoal (i+1) [] def_avoid def_env c (lift n t))
bodies tys vn in
- RRec(dl,RFix (Array.map (fun i -> Some i, RStructRec) (fst nvn), snd nvn),Array.of_list (List.rev lfi),
+ GRec(dl,RFix (Array.map (fun i -> Some i, RStructRec) (fst nvn), snd nvn),Array.of_list (List.rev lfi),
Array.map (fun (bl,_,_) -> bl) v,
Array.map (fun (_,_,ty) -> ty) v,
Array.map (fun (_,bd,_) -> bd) v)
@@ -440,7 +440,7 @@ and detype_cofix isgoal avoid env n (names,tys,bodies) =
let v = array_map2
(fun c t -> share_names isgoal 0 [] def_avoid def_env c (lift ntys t))
bodies tys in
- RRec(dl,RCoFix n,Array.of_list (List.rev lfi),
+ GRec(dl,RCoFix n,Array.of_list (List.rev lfi),
Array.map (fun (bl,_,_) -> bl) v,
Array.map (fun (_,_,ty) -> ty) v,
Array.map (fun (_,bd,_) -> bd) v)
@@ -535,9 +535,9 @@ and detype_binder isgoal bk avoid env na ty c =
else compute_displayed_name_in flag avoid na c in
let r = detype isgoal avoid' (add_name na' env) c in
match bk with
- | BProd -> RProd (dl, na',Explicit,detype false avoid env ty, r)
- | BLambda -> RLambda (dl, na',Explicit,detype false avoid env ty, r)
- | BLetIn -> RLetIn (dl, na',detype false avoid env ty, r)
+ | BProd -> GProd (dl, na',Explicit,detype false avoid env ty, r)
+ | BLambda -> GLambda (dl, na',Explicit,detype false avoid env ty, r)
+ | BLetIn -> GLetIn (dl, na',detype false avoid env ty, r)
let rec detype_rel_context where avoid env sign =
let where = Option.map (fun c -> it_mkLambda_or_LetIn c sign) where in
@@ -569,42 +569,42 @@ let rec subst_cases_pattern subst pat =
if kn' == kn && cpl' == cpl then pat else
PatCstr (loc,((kn',i),j),cpl',n)
-let rec subst_rawconstr subst raw =
+let rec subst_glob_constr subst raw =
match raw with
- | RRef (loc,ref) ->
+ | GRef (loc,ref) ->
let ref',t = subst_global subst ref in
if ref' == ref then raw else
detype false [] [] t
- | RVar _ -> raw
- | REvar _ -> raw
- | RPatVar _ -> raw
+ | GVar _ -> raw
+ | GEvar _ -> raw
+ | GPatVar _ -> raw
- | RApp (loc,r,rl) ->
- let r' = subst_rawconstr subst r
- and rl' = list_smartmap (subst_rawconstr subst) rl in
+ | GApp (loc,r,rl) ->
+ let r' = subst_glob_constr subst r
+ and rl' = list_smartmap (subst_glob_constr subst) rl in
if r' == r && rl' == rl then raw else
- RApp(loc,r',rl')
+ GApp(loc,r',rl')
- | RLambda (loc,n,bk,r1,r2) ->
- let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
+ | GLambda (loc,n,bk,r1,r2) ->
+ let r1' = subst_glob_constr subst r1 and r2' = subst_glob_constr subst r2 in
if r1' == r1 && r2' == r2 then raw else
- RLambda (loc,n,bk,r1',r2')
+ GLambda (loc,n,bk,r1',r2')
- | RProd (loc,n,bk,r1,r2) ->
- let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
+ | GProd (loc,n,bk,r1,r2) ->
+ let r1' = subst_glob_constr subst r1 and r2' = subst_glob_constr subst r2 in
if r1' == r1 && r2' == r2 then raw else
- RProd (loc,n,bk,r1',r2')
+ GProd (loc,n,bk,r1',r2')
- | RLetIn (loc,n,r1,r2) ->
- let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
+ | GLetIn (loc,n,r1,r2) ->
+ let r1' = subst_glob_constr subst r1 and r2' = subst_glob_constr subst r2 in
if r1' == r1 && r2' == r2 then raw else
- RLetIn (loc,n,r1',r2')
+ GLetIn (loc,n,r1',r2')
- | RCases (loc,sty,rtno,rl,branches) ->
- let rtno' = Option.smartmap (subst_rawconstr subst) rtno
+ | GCases (loc,sty,rtno,rl,branches) ->
+ let rtno' = Option.smartmap (subst_glob_constr subst) rtno
and rl' = list_smartmap (fun (a,x as y) ->
- let a' = subst_rawconstr subst a in
+ let a' = subst_glob_constr subst a in
let (n,topt) = x in
let topt' = Option.smartmap
(fun (loc,(sp,i),x,y as t) ->
@@ -615,61 +615,61 @@ let rec subst_rawconstr subst raw =
(fun (loc,idl,cpl,r as branch) ->
let cpl' =
list_smartmap (subst_cases_pattern subst) cpl
- and r' = subst_rawconstr subst r in
+ and r' = subst_glob_constr subst r in
if cpl' == cpl && r' == r then branch else
(loc,idl,cpl',r'))
branches
in
if rtno' == rtno && rl' == rl && branches' == branches then raw else
- RCases (loc,sty,rtno',rl',branches')
+ GCases (loc,sty,rtno',rl',branches')
- | RLetTuple (loc,nal,(na,po),b,c) ->
- let po' = Option.smartmap (subst_rawconstr subst) po
- and b' = subst_rawconstr subst b
- and c' = subst_rawconstr subst c in
+ | GLetTuple (loc,nal,(na,po),b,c) ->
+ let po' = Option.smartmap (subst_glob_constr subst) po
+ and b' = subst_glob_constr subst b
+ and c' = subst_glob_constr subst c in
if po' == po && b' == b && c' == c then raw else
- RLetTuple (loc,nal,(na,po'),b',c')
+ GLetTuple (loc,nal,(na,po'),b',c')
- | RIf (loc,c,(na,po),b1,b2) ->
- let po' = Option.smartmap (subst_rawconstr subst) po
- and b1' = subst_rawconstr subst b1
- and b2' = subst_rawconstr subst b2
- and c' = subst_rawconstr subst c in
+ | GIf (loc,c,(na,po),b1,b2) ->
+ let po' = Option.smartmap (subst_glob_constr subst) po
+ and b1' = subst_glob_constr subst b1
+ and b2' = subst_glob_constr subst b2
+ and c' = subst_glob_constr subst c in
if c' == c & po' == po && b1' == b1 && b2' == b2 then raw else
- RIf (loc,c',(na,po'),b1',b2')
+ GIf (loc,c',(na,po'),b1',b2')
- | RRec (loc,fix,ida,bl,ra1,ra2) ->
- let ra1' = array_smartmap (subst_rawconstr subst) ra1
- and ra2' = array_smartmap (subst_rawconstr subst) ra2 in
+ | GRec (loc,fix,ida,bl,ra1,ra2) ->
+ let ra1' = array_smartmap (subst_glob_constr subst) ra1
+ and ra2' = array_smartmap (subst_glob_constr subst) ra2 in
let bl' = array_smartmap
(list_smartmap (fun (na,k,obd,ty as dcl) ->
- let ty' = subst_rawconstr subst ty in
- let obd' = Option.smartmap (subst_rawconstr subst) obd in
+ let ty' = subst_glob_constr subst ty in
+ let obd' = Option.smartmap (subst_glob_constr subst) obd in
if ty'==ty & obd'==obd then dcl else (na,k,obd',ty')))
bl in
if ra1' == ra1 && ra2' == ra2 && bl'==bl then raw else
- RRec (loc,fix,ida,bl',ra1',ra2')
+ GRec (loc,fix,ida,bl',ra1',ra2')
- | RSort _ -> raw
+ | GSort _ -> raw
- | RHole (loc,ImplicitArg (ref,i,b)) ->
+ | GHole (loc,ImplicitArg (ref,i,b)) ->
let ref',_ = subst_global subst ref in
if ref' == ref then raw else
- RHole (loc,InternalHole)
- | RHole (loc, (BinderType _ | QuestionMark _ | CasesType | InternalHole |
+ GHole (loc,InternalHole)
+ | GHole (loc, (BinderType _ | QuestionMark _ | CasesType | InternalHole |
TomatchTypeParameter _ | GoalEvar | ImpossibleCase | MatchingVar _)) ->
raw
- | RCast (loc,r1,k) ->
+ | GCast (loc,r1,k) ->
(match k with
CastConv (k,r2) ->
- let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
+ let r1' = subst_glob_constr subst r1 and r2' = subst_glob_constr subst r2 in
if r1' == r1 && r2' == r2 then raw else
- RCast (loc,r1', CastConv (k,r2'))
+ GCast (loc,r1', CastConv (k,r2'))
| CastCoerce ->
- let r1' = subst_rawconstr subst r1 in
- if r1' == r1 then raw else RCast (loc,r1',k))
- | RDynamic _ -> raw
+ let r1' = subst_glob_constr subst r1 in
+ if r1' == r1 then raw else GCast (loc,r1',k))
+ | GDynamic _ -> raw
(* Utilities to transform kernel cases to simple pattern-matching problem *)
diff --git a/pretyping/detyping.mli b/pretyping/detyping.mli
index e178a4ca3..e2644592c 100644
--- a/pretyping/detyping.mli
+++ b/pretyping/detyping.mli
@@ -17,40 +17,40 @@ open Mod_subst
val subst_cases_pattern : substitution -> cases_pattern -> cases_pattern
-val subst_rawconstr : substitution -> rawconstr -> rawconstr
+val subst_glob_constr : substitution -> glob_constr -> glob_constr
-(** [detype isgoal avoid ctx c] turns a closed [c], into a rawconstr
+(** [detype isgoal avoid ctx c] turns a closed [c], into a glob_constr
de Bruijn indexes are turned to bound names, avoiding names in [avoid]
[isgoal] tells if naming must avoid global-level synonyms as intro does
[ctx] gives the names of the free variables *)
-val detype : bool -> identifier list -> names_context -> constr -> rawconstr
+val detype : bool -> identifier list -> names_context -> constr -> glob_constr
val detype_case :
- bool -> ('a -> rawconstr) ->
+ bool -> ('a -> glob_constr) ->
(constructor array -> int array -> 'a array ->
- (loc * identifier list * cases_pattern list * rawconstr) list) ->
+ (loc * identifier list * cases_pattern list * glob_constr) list) ->
('a -> int -> bool) ->
identifier list -> inductive * case_style * int * int array * int ->
- 'a option -> 'a -> 'a array -> rawconstr
+ 'a option -> 'a -> 'a array -> glob_constr
val detype_sort : sorts -> rawsort
val detype_rel_context : constr option -> identifier list -> names_context ->
- rel_context -> rawdecl list
+ rel_context -> glob_decl list
(** look for the index of a named var or a nondep var as it is renamed *)
val lookup_name_as_displayed : env -> constr -> identifier -> int option
val lookup_index_as_renamed : env -> constr -> int -> int option
-val set_detype_anonymous : (loc -> int -> rawconstr) -> unit
+val set_detype_anonymous : (loc -> int -> glob_constr) -> unit
val force_wildcard : unit -> bool
val synthetize_type : unit -> bool
(** Utilities to transform kernel cases to simple pattern-matching problem *)
-val it_destRLambda_or_LetIn_names : int -> rawconstr -> name list * rawconstr
+val it_destRLambda_or_LetIn_names : int -> glob_constr -> name list * glob_constr
val simple_cases_matrix_of_branches :
- inductive -> int list -> rawconstr list -> cases_clauses
+ inductive -> int list -> glob_constr list -> cases_clauses
val return_type_of_predicate :
- inductive -> int -> int -> rawconstr -> predicate_pattern * rawconstr option
+ inductive -> int -> int -> glob_constr -> predicate_pattern * glob_constr option
diff --git a/pretyping/pattern.ml b/pretyping/pattern.ml
index 15bc06e02..2217074fe 100644
--- a/pretyping/pattern.ml
+++ b/pretyping/pattern.ml
@@ -248,48 +248,48 @@ let mkPLambda na b = PLambda(na,PMeta None,b)
let rev_it_mkPLambda = List.fold_right mkPLambda
let rec pat_of_raw metas vars = function
- | RVar (_,id) ->
+ | GVar (_,id) ->
(try PRel (list_index (Name id) vars)
with Not_found -> PVar id)
- | RPatVar (_,(false,n)) ->
+ | GPatVar (_,(false,n)) ->
metas := n::!metas; PMeta (Some n)
- | RRef (_,gr) ->
+ | GRef (_,gr) ->
PRef (canonical_gr gr)
(* Hack pour ne pas réécrire une interprétation complète des patterns*)
- | RApp (_, RPatVar (_,(true,n)), cl) ->
+ | GApp (_, GPatVar (_,(true,n)), cl) ->
metas := n::!metas; PSoApp (n, List.map (pat_of_raw metas vars) cl)
- | RApp (_,c,cl) ->
+ | GApp (_,c,cl) ->
PApp (pat_of_raw metas vars c,
Array.of_list (List.map (pat_of_raw metas vars) cl))
- | RLambda (_,na,bk,c1,c2) ->
+ | GLambda (_,na,bk,c1,c2) ->
name_iter (fun n -> metas := n::!metas) na;
PLambda (na, pat_of_raw metas vars c1,
pat_of_raw metas (na::vars) c2)
- | RProd (_,na,bk,c1,c2) ->
+ | GProd (_,na,bk,c1,c2) ->
name_iter (fun n -> metas := n::!metas) na;
PProd (na, pat_of_raw metas vars c1,
pat_of_raw metas (na::vars) c2)
- | RLetIn (_,na,c1,c2) ->
+ | GLetIn (_,na,c1,c2) ->
name_iter (fun n -> metas := n::!metas) na;
PLetIn (na, pat_of_raw metas vars c1,
pat_of_raw metas (na::vars) c2)
- | RSort (_,s) ->
+ | GSort (_,s) ->
PSort s
- | RHole _ ->
+ | GHole _ ->
PMeta None
- | RCast (_,c,_) ->
+ | GCast (_,c,_) ->
Flags.if_verbose
Pp.warning "Cast not taken into account in constr pattern";
pat_of_raw metas vars c
- | RIf (_,c,(_,None),b1,b2) ->
+ | GIf (_,c,(_,None),b1,b2) ->
PIf (pat_of_raw metas vars c,
pat_of_raw metas vars b1,pat_of_raw metas vars b2)
- | RLetTuple (loc,nal,(_,None),b,c) ->
- let mkRLambda c na = RLambda (loc,na,Explicit,RHole (loc,Evd.InternalHole),c) in
+ | GLetTuple (loc,nal,(_,None),b,c) ->
+ let mkRLambda c na = GLambda (loc,na,Explicit,GHole (loc,Evd.InternalHole),c) in
let c = List.fold_left mkRLambda c nal in
PCase ((LetStyle,[|1|],None,None),PMeta None,pat_of_raw metas vars b,
[|pat_of_raw metas vars c|])
- | RCases (loc,sty,p,[c,(na,indnames)],brs) ->
+ | GCases (loc,sty,p,[c,(na,indnames)],brs) ->
let pred,ind_nargs, ind = match p,indnames with
| Some p, Some (_,ind,n,nal) ->
rev_it_mkPLambda nal (mkPLambda na (pat_of_raw metas vars p)),
@@ -307,34 +307,34 @@ let rec pat_of_raw metas vars = function
pat_of_raw metas vars c, brs)
| r ->
- let loc = loc_of_rawconstr r in
- user_err_loc (loc,"pattern_of_rawconstr", Pp.str"Non supported pattern.")
+ let loc = loc_of_glob_constr r in
+ user_err_loc (loc,"pattern_of_glob_constr", Pp.str"Non supported pattern.")
and pat_of_raw_branch loc metas vars ind brs i =
let bri = List.filter
(function
(_,_,[PatCstr(_,c,lv,Anonymous)],_) -> snd c = i+1
| (loc,_,_,_) ->
- user_err_loc (loc,"pattern_of_rawconstr",
+ user_err_loc (loc,"pattern_of_glob_constr",
Pp.str "Non supported pattern.")) brs in
match bri with
| [(_,_,[PatCstr(_,(indsp,_),lv,_)],br)] ->
if ind <> None & ind <> Some indsp then
- user_err_loc (loc,"pattern_of_rawconstr",
+ user_err_loc (loc,"pattern_of_glob_constr",
Pp.str "All constructors must be in the same inductive type.");
let lna =
List.map
(function PatVar(_,na) -> na
| PatCstr(loc,_,_,_) ->
- user_err_loc (loc,"pattern_of_rawconstr",
+ user_err_loc (loc,"pattern_of_glob_constr",
Pp.str "Non supported pattern.")) lv in
let vars' = List.rev lna @ vars in
List.length lv, rev_it_mkPLambda lna (pat_of_raw metas vars' br)
- | _ -> user_err_loc (loc,"pattern_of_rawconstr",
+ | _ -> user_err_loc (loc,"pattern_of_glob_constr",
str "No unique branch for " ++ int (i+1) ++
str"-th constructor.")
-let pattern_of_rawconstr c =
+let pattern_of_glob_constr c =
let metas = ref [] in
let p = pat_of_raw metas [] c in
(!metas,p)
diff --git a/pretyping/pattern.mli b/pretyping/pattern.mli
index 7dadb78f3..a739a2888 100644
--- a/pretyping/pattern.mli
+++ b/pretyping/pattern.mli
@@ -103,11 +103,11 @@ val head_of_constr_reference : Term.constr -> global_reference
val pattern_of_constr : Evd.evar_map -> constr -> named_context * constr_pattern
-(** [pattern_of_rawconstr l c] translates a term [c] with metavariables into
+(** [pattern_of_glob_constr l c] translates a term [c] with metavariables into
a pattern; variables bound in [l] are replaced by the pattern to which they
are bound *)
-val pattern_of_rawconstr : rawconstr ->
+val pattern_of_glob_constr : glob_constr ->
patvar list * constr_pattern
val instantiate_pattern :
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index 7017edaf2..bc80296d5 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -48,7 +48,7 @@ type typing_constraint = OfType of types option | IsType
type var_map = (identifier * constr_under_binders) list
type unbound_ltac_var_map = (identifier * identifier option) list
type ltac_var_map = var_map * unbound_ltac_var_map
-type rawconstr_ltac_closure = ltac_var_map * rawconstr
+type glob_constr_ltac_closure = ltac_var_map * glob_constr
(************************************************************************)
(* This concerns Cases *)
@@ -86,7 +86,7 @@ let search_guard loc env possible_indexes fixdefs =
user_err_loc (loc,"search_guard", Pp.str errmsg)
with Found indexes -> indexes)
-(* To embed constr in rawconstr *)
+(* To embed constr in glob_constr *)
let ((constr_in : constr -> Dyn.t),
(constr_out : Dyn.t -> constr)) = Dyn.create "constr"
@@ -109,19 +109,19 @@ sig
(* Allow references to syntaxically inexistent variables (i.e., if applied on an inductive) *)
val allow_anonymous_refs : bool ref
- (* Generic call to the interpreter from rawconstr to open_constr, leaving
+ (* Generic call to the interpreter from glob_constr to open_constr, leaving
unresolved holes as evars and returning the typing contexts of
these evars. Work as [understand_gen] for the rest. *)
val understand_tcc : ?resolve_classes:bool ->
- evar_map -> env -> ?expected_type:types -> rawconstr -> open_constr
+ evar_map -> env -> ?expected_type:types -> glob_constr -> open_constr
val understand_tcc_evars : ?fail_evar:bool -> ?resolve_classes:bool ->
- evar_map ref -> env -> typing_constraint -> rawconstr -> constr
+ evar_map ref -> env -> typing_constraint -> glob_constr -> constr
(* More general entry point with evars from ltac *)
- (* Generic call to the interpreter from rawconstr to constr, failing
+ (* Generic call to the interpreter from glob_constr to constr, failing
unresolved holes in the rawterm cannot be instantiated.
In [understand_ltac expand_evars sigma env ltac_env constraint c],
@@ -134,29 +134,29 @@ sig
val understand_ltac :
bool -> evar_map -> env -> ltac_var_map ->
- typing_constraint -> rawconstr -> evar_map * constr
+ typing_constraint -> glob_constr -> evar_map * constr
- (* Standard call to get a constr from a rawconstr, resolving implicit args *)
+ (* Standard call to get a constr from a glob_constr, resolving implicit args *)
val understand : evar_map -> env -> ?expected_type:Term.types ->
- rawconstr -> constr
+ glob_constr -> constr
- (* Idem but the rawconstr is intended to be a type *)
+ (* Idem but the glob_constr is intended to be a type *)
- val understand_type : evar_map -> env -> rawconstr -> constr
+ val understand_type : evar_map -> env -> glob_constr -> constr
(* A generalization of the two previous case *)
val understand_gen : typing_constraint -> evar_map -> env ->
- rawconstr -> constr
+ glob_constr -> constr
(* Idem but returns the judgment of the understood term *)
- val understand_judgment : evar_map -> env -> rawconstr -> unsafe_judgment
+ val understand_judgment : evar_map -> env -> glob_constr -> unsafe_judgment
(* Idem but do not fail on unresolved evars *)
- val understand_judgment_tcc : evar_map ref -> env -> rawconstr -> unsafe_judgment
+ val understand_judgment_tcc : evar_map ref -> env -> glob_constr -> unsafe_judgment
(*i*)
(* Internal of Pretyping...
@@ -164,15 +164,15 @@ sig
*)
val pretype :
type_constraint -> env -> evar_map ref ->
- ltac_var_map -> rawconstr -> unsafe_judgment
+ ltac_var_map -> glob_constr -> unsafe_judgment
val pretype_type :
val_constraint -> env -> evar_map ref ->
- ltac_var_map -> rawconstr -> unsafe_type_judgment
+ ltac_var_map -> glob_constr -> unsafe_type_judgment
val pretype_gen :
bool -> bool -> bool -> evar_map ref -> env ->
- ltac_var_map -> typing_constraint -> rawconstr -> constr
+ ltac_var_map -> typing_constraint -> glob_constr -> constr
(*i*)
end
@@ -302,17 +302,17 @@ module Pretyping_F (Coercion : Coercion.S) = struct
(* in environment [env], with existential variables [evdref] and *)
(* the type constraint tycon *)
let rec pretype (tycon : type_constraint) env evdref lvar = function
- | RRef (loc,ref) ->
+ | GRef (loc,ref) ->
inh_conv_coerce_to_tycon loc env evdref
(pretype_ref evdref env ref)
tycon
- | RVar (loc, id) ->
+ | GVar (loc, id) ->
inh_conv_coerce_to_tycon loc env evdref
(pretype_id loc env !evdref lvar id)
tycon
- | REvar (loc, evk, instopt) ->
+ | GEvar (loc, evk, instopt) ->
(* Ne faudrait-il pas s'assurer que hyps est bien un
sous-contexte du contexte courant, et qu'il n'y a pas de Rel "caché" *)
let hyps = evar_filtered_context (Evd.find !evdref evk) in
@@ -323,7 +323,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
let j = (Retyping.get_judgment_of env !evdref c) in
inh_conv_coerce_to_tycon loc env evdref j tycon
- | RPatVar (loc,(someta,n)) ->
+ | GPatVar (loc,(someta,n)) ->
let ty =
match tycon with
| Some (None, ty) -> ty
@@ -332,7 +332,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
let k = MatchingVar (someta,n) in
{ uj_val = e_new_evar evdref env ~src:(loc,k) ty; uj_type = ty }
- | RHole (loc,k) ->
+ | GHole (loc,k) ->
let ty =
match tycon with
| Some (None, ty) -> ty
@@ -340,7 +340,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
e_new_evar evdref env ~src:(loc,InternalHole) (new_Type ()) in
{ uj_val = e_new_evar evdref env ~src:(loc,k) ty; uj_type = ty }
- | RRec (loc,fixkind,names,bl,lar,vdef) ->
+ | GRec (loc,fixkind,names,bl,lar,vdef) ->
let rec type_bl env ctxt = function
[] -> ctxt
| (na,bk,None,ty)::bl ->
@@ -403,16 +403,16 @@ module Pretyping_F (Coercion : Coercion.S) = struct
make_judge (mkCoFix cofix) ftys.(i) in
inh_conv_coerce_to_tycon loc env evdref fixj tycon
- | RSort (loc,s) ->
+ | GSort (loc,s) ->
inh_conv_coerce_to_tycon loc env evdref (pretype_sort s) tycon
- | RApp (loc,f,args) ->
+ | GApp (loc,f,args) ->
let fj = pretype empty_tycon env evdref lvar f in
- let floc = loc_of_rawconstr f in
+ let floc = loc_of_glob_constr f in
let rec apply_rec env n resj = function
| [] -> resj
| c::rest ->
- let argloc = loc_of_rawconstr c in
+ let argloc = loc_of_glob_constr c in
let resj = evd_comb1 (Coercion.inh_app_fun env) evdref resj in
let resty = whd_betadeltaiota env !evdref resj.uj_type in
match kind_of_term resty with
@@ -446,7 +446,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
| _ -> resj in
inh_conv_coerce_to_tycon loc env evdref resj tycon
- | RLambda(loc,name,bk,c1,c2) ->
+ | GLambda(loc,name,bk,c1,c2) ->
let (name',dom,rng) = evd_comb1 (split_tycon loc env) evdref tycon in
let dom_valcon = valcon_of_tycon dom in
let j = pretype_type dom_valcon env evdref lvar c1 in
@@ -454,7 +454,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
let j' = pretype rng (push_rel var env) evdref lvar c2 in
judge_of_abstraction env (orelse_name name name') j j'
- | RProd(loc,name,bk,c1,c2) ->
+ | GProd(loc,name,bk,c1,c2) ->
let j = pretype_type empty_valcon env evdref lvar c1 in
let j' =
if name = Anonymous then
@@ -470,10 +470,10 @@ module Pretyping_F (Coercion : Coercion.S) = struct
with TypeError _ as e -> Loc.raise loc e in
inh_conv_coerce_to_tycon loc env evdref resj tycon
- | RLetIn(loc,name,c1,c2) ->
+ | GLetIn(loc,name,c1,c2) ->
let j =
match c1 with
- | RCast (loc, c, CastConv (DEFAULTcast, t)) ->
+ | GCast (loc, c, CastConv (DEFAULTcast, t)) ->
let tj = pretype_type empty_valcon env evdref lvar t in
pretype (mk_tycon tj.utj_val) env evdref lvar c
| _ -> pretype empty_tycon env evdref lvar c1
@@ -485,12 +485,12 @@ module Pretyping_F (Coercion : Coercion.S) = struct
{ uj_val = mkLetIn (name, j.uj_val, t, j'.uj_val) ;
uj_type = subst1 j.uj_val j'.uj_type }
- | RLetTuple (loc,nal,(na,po),c,d) ->
+ | GLetTuple (loc,nal,(na,po),c,d) ->
let cj = pretype empty_tycon env evdref lvar c in
let (IndType (indf,realargs)) =
try find_rectype env !evdref cj.uj_type
with Not_found ->
- let cloc = loc_of_rawconstr c in
+ let cloc = loc_of_glob_constr c in
error_case_not_inductive_loc cloc env !evdref cj
in
let cstrs = get_constructors env indf in
@@ -551,12 +551,12 @@ module Pretyping_F (Coercion : Coercion.S) = struct
in
{ uj_val = v; uj_type = ccl })
- | RIf (loc,c,(na,po),b1,b2) ->
+ | GIf (loc,c,(na,po),b1,b2) ->
let cj = pretype empty_tycon env evdref lvar c in
let (IndType (indf,realargs)) =
try find_rectype env !evdref cj.uj_type
with Not_found ->
- let cloc = loc_of_rawconstr c in
+ let cloc = loc_of_glob_constr c in
error_case_not_inductive_loc cloc env !evdref cj in
let cstrs = get_constructors env indf in
if Array.length cstrs <> 2 then
@@ -619,12 +619,12 @@ module Pretyping_F (Coercion : Coercion.S) = struct
in
{ uj_val = v; uj_type = p }
- | RCases (loc,sty,po,tml,eqns) ->
+ | GCases (loc,sty,po,tml,eqns) ->
Cases.compile_cases loc sty
((fun vtyc env evdref -> pretype vtyc env evdref lvar),evdref)
tycon env (* loc *) (po,tml,eqns)
- | RCast (loc,c,k) ->
+ | GCast (loc,c,k) ->
let cj =
match k with
CastCoerce ->
@@ -643,7 +643,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
{ uj_val = v; uj_type = tval }
in inh_conv_coerce_to_tycon loc env evdref cj tycon
- | RDynamic (loc,d) ->
+ | GDynamic (loc,d) ->
if (Dyn.tag d) = "constr" then
let c = constr_out d in
let j = (Retyping.get_judgment_of env !evdref c) in
@@ -654,7 +654,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
(* [pretype_type valcon env evdref lvar c] coerces [c] into a type *)
and pretype_type valcon env evdref lvar = function
- | RHole loc ->
+ | GHole loc ->
(match valcon with
| Some v ->
let s =
@@ -674,7 +674,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
utj_type = s})
| c ->
let j = pretype empty_tycon env evdref lvar c in
- let loc = loc_of_rawconstr c in
+ let loc = loc_of_glob_constr c in
let tj = evd_comb1 (Coercion.inh_coerce_to_sort loc env) evdref j in
match valcon with
| None -> tj
@@ -682,7 +682,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
if e_cumul env evdref v tj.utj_val then tj
else
error_unexpected_type_loc
- (loc_of_rawconstr c) env !evdref tj.utj_val v
+ (loc_of_glob_constr c) env !evdref tj.utj_val v
let pretype_gen expand_evar fail_evar resolve_classes evdref env lvar kind c =
let c' = match kind with
diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli
index 7f707de3d..8c0270743 100644
--- a/pretyping/pretyping.mli
+++ b/pretyping/pretyping.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(** This file implements type inference. It maps [rawconstr]
+(** This file implements type inference. It maps [glob_constr]
(i.e. untyped terms whose names are located) to [constr]. In
particular, it drives complex pattern-matching problems ("match")
into elementary ones, insertion of coercions and resolution of
@@ -30,7 +30,7 @@ type typing_constraint = OfType of types option | IsType
type var_map = (identifier * Pattern.constr_under_binders) list
type unbound_ltac_var_map = (identifier * identifier option) list
type ltac_var_map = var_map * unbound_ltac_var_map
-type rawconstr_ltac_closure = ltac_var_map * rawconstr
+type glob_constr_ltac_closure = ltac_var_map * glob_constr
module type S =
sig
@@ -40,19 +40,19 @@ sig
(** Allow references to syntaxically inexistent variables (i.e., if applied on an inductive) *)
val allow_anonymous_refs : bool ref
- (** Generic call to the interpreter from rawconstr to open_constr, leaving
+ (** Generic call to the interpreter from glob_constr to open_constr, leaving
unresolved holes as evars and returning the typing contexts of
these evars. Work as [understand_gen] for the rest. *)
val understand_tcc : ?resolve_classes:bool ->
- evar_map -> env -> ?expected_type:types -> rawconstr -> open_constr
+ evar_map -> env -> ?expected_type:types -> glob_constr -> open_constr
val understand_tcc_evars : ?fail_evar:bool -> ?resolve_classes:bool ->
- evar_map ref -> env -> typing_constraint -> rawconstr -> constr
+ evar_map ref -> env -> typing_constraint -> glob_constr -> constr
(** More general entry point with evars from ltac *)
- (** Generic call to the interpreter from rawconstr to constr, failing
+ (** Generic call to the interpreter from glob_constr to constr, failing
unresolved holes in the rawterm cannot be instantiated.
In [understand_ltac expand_evars sigma env ltac_env constraint c],
@@ -65,42 +65,42 @@ sig
val understand_ltac :
bool -> evar_map -> env -> ltac_var_map ->
- typing_constraint -> rawconstr -> evar_map * constr
+ typing_constraint -> glob_constr -> evar_map * constr
- (** Standard call to get a constr from a rawconstr, resolving implicit args *)
+ (** Standard call to get a constr from a glob_constr, resolving implicit args *)
val understand : evar_map -> env -> ?expected_type:Term.types ->
- rawconstr -> constr
+ glob_constr -> constr
- (** Idem but the rawconstr is intended to be a type *)
+ (** Idem but the glob_constr is intended to be a type *)
- val understand_type : evar_map -> env -> rawconstr -> constr
+ val understand_type : evar_map -> env -> glob_constr -> constr
(** A generalization of the two previous case *)
val understand_gen : typing_constraint -> evar_map -> env ->
- rawconstr -> constr
+ glob_constr -> constr
(** Idem but returns the judgment of the understood term *)
- val understand_judgment : evar_map -> env -> rawconstr -> unsafe_judgment
+ val understand_judgment : evar_map -> env -> glob_constr -> unsafe_judgment
(** Idem but do not fail on unresolved evars *)
- val understand_judgment_tcc : evar_map ref -> env -> rawconstr -> unsafe_judgment
+ val understand_judgment_tcc : evar_map ref -> env -> glob_constr -> unsafe_judgment
(**/**)
(** Internal of Pretyping... *)
val pretype :
type_constraint -> env -> evar_map ref ->
- ltac_var_map -> rawconstr -> unsafe_judgment
+ ltac_var_map -> glob_constr -> unsafe_judgment
val pretype_type :
val_constraint -> env -> evar_map ref ->
- ltac_var_map -> rawconstr -> unsafe_type_judgment
+ ltac_var_map -> glob_constr -> unsafe_type_judgment
val pretype_gen :
bool -> bool -> bool -> evar_map ref -> env ->
- ltac_var_map -> typing_constraint -> rawconstr -> constr
+ ltac_var_map -> typing_constraint -> glob_constr -> constr
(**/**)
@@ -109,7 +109,7 @@ end
module Pretyping_F (C : Coercion.S) : S
module Default : S
-(** To embed constr in rawconstr *)
+(** To embed constr in glob_constr *)
val constr_in : constr -> Dyn.t
val constr_out : Dyn.t -> constr
diff --git a/pretyping/rawterm.ml b/pretyping/rawterm.ml
index 88dc5db42..deba9a257 100644
--- a/pretyping/rawterm.ml
+++ b/pretyping/rawterm.ml
@@ -49,29 +49,29 @@ type 'a cast_type =
| CastConv of cast_kind * 'a
| CastCoerce (* Cast to a base type (eg, an underlying inductive type) *)
-type rawconstr =
- | RRef of (loc * global_reference)
- | RVar of (loc * identifier)
- | REvar of loc * existential_key * rawconstr list option
- | RPatVar of loc * (bool * patvar) (* Used for patterns only *)
- | RApp of loc * rawconstr * rawconstr list
- | RLambda of loc * name * binding_kind * rawconstr * rawconstr
- | RProd of loc * name * binding_kind * rawconstr * rawconstr
- | RLetIn of loc * name * rawconstr * rawconstr
- | RCases of loc * case_style * rawconstr option * tomatch_tuples * cases_clauses
- | RLetTuple of loc * name list * (name * rawconstr option) *
- rawconstr * rawconstr
- | RIf of loc * rawconstr * (name * rawconstr option) * rawconstr * rawconstr
- | RRec of loc * fix_kind * identifier array * rawdecl list array *
- rawconstr array * rawconstr array
- | RSort of loc * rawsort
- | RHole of (loc * hole_kind)
- | RCast of loc * rawconstr * rawconstr cast_type
- | RDynamic of loc * Dyn.t
-
-and rawdecl = name * binding_kind * rawconstr option * rawconstr
-
-and fix_recursion_order = RStructRec | RWfRec of rawconstr | RMeasureRec of rawconstr * rawconstr option
+type glob_constr =
+ | GRef of (loc * global_reference)
+ | GVar of (loc * identifier)
+ | GEvar of loc * existential_key * glob_constr list option
+ | GPatVar of loc * (bool * patvar) (* Used for patterns only *)
+ | GApp of loc * glob_constr * glob_constr list
+ | GLambda of loc * name * binding_kind * glob_constr * glob_constr
+ | GProd of loc * name * binding_kind * glob_constr * glob_constr
+ | GLetIn of loc * name * glob_constr * glob_constr
+ | GCases of loc * case_style * glob_constr option * tomatch_tuples * cases_clauses
+ | GLetTuple of loc * name list * (name * glob_constr option) *
+ glob_constr * glob_constr
+ | GIf of loc * glob_constr * (name * glob_constr option) * glob_constr * glob_constr
+ | GRec of loc * fix_kind * identifier array * glob_decl list array *
+ glob_constr array * glob_constr array
+ | GSort of loc * rawsort
+ | GHole of (loc * hole_kind)
+ | GCast of loc * glob_constr * glob_constr cast_type
+ | GDynamic of loc * Dyn.t
+
+and glob_decl = name * binding_kind * glob_constr option * glob_constr
+
+and fix_recursion_order = RStructRec | RWfRec of glob_constr | RMeasureRec of glob_constr * glob_constr option
and fix_kind =
| RFix of ((int option * fix_recursion_order) array * int)
@@ -80,11 +80,11 @@ and fix_kind =
and predicate_pattern =
name * (loc * inductive * int * name list) option
-and tomatch_tuple = (rawconstr * predicate_pattern)
+and tomatch_tuple = (glob_constr * predicate_pattern)
and tomatch_tuples = tomatch_tuple list
-and cases_clause = (loc * identifier list * cases_pattern list * rawconstr)
+and cases_clause = (loc * identifier list * cases_pattern list * glob_constr)
and cases_clauses = cases_clause list
@@ -93,55 +93,55 @@ let cases_predicate_names tml =
| (tm,(na,None)) -> [na]
| (tm,(na,Some (_,_,_,nal))) -> na::nal) tml)
-let map_rawdecl_left_to_right f (na,k,obd,ty) =
+let map_glob_decl_left_to_right f (na,k,obd,ty) =
let comp1 = Option.map f obd in
let comp2 = f ty in
(na,k,comp1,comp2)
-let map_rawconstr_left_to_right f = function
- | RApp (loc,g,args) ->
+let map_glob_constr_left_to_right f = function
+ | GApp (loc,g,args) ->
let comp1 = f g in
let comp2 = Util.list_map_left f args in
- RApp (loc,comp1,comp2)
- | RLambda (loc,na,bk,ty,c) ->
+ GApp (loc,comp1,comp2)
+ | GLambda (loc,na,bk,ty,c) ->
let comp1 = f ty in
let comp2 = f c in
- RLambda (loc,na,bk,comp1,comp2)
- | RProd (loc,na,bk,ty,c) ->
+ GLambda (loc,na,bk,comp1,comp2)
+ | GProd (loc,na,bk,ty,c) ->
let comp1 = f ty in
let comp2 = f c in
- RProd (loc,na,bk,comp1,comp2)
- | RLetIn (loc,na,b,c) ->
+ GProd (loc,na,bk,comp1,comp2)
+ | GLetIn (loc,na,b,c) ->
let comp1 = f b in
let comp2 = f c in
- RLetIn (loc,na,comp1,comp2)
- | RCases (loc,sty,rtntypopt,tml,pl) ->
+ GLetIn (loc,na,comp1,comp2)
+ | GCases (loc,sty,rtntypopt,tml,pl) ->
let comp1 = Option.map f rtntypopt in
let comp2 = Util.list_map_left (fun (tm,x) -> (f tm,x)) tml in
let comp3 = Util.list_map_left (fun (loc,idl,p,c) -> (loc,idl,p,f c)) pl in
- RCases (loc,sty,comp1,comp2,comp3)
- | RLetTuple (loc,nal,(na,po),b,c) ->
+ GCases (loc,sty,comp1,comp2,comp3)
+ | GLetTuple (loc,nal,(na,po),b,c) ->
let comp1 = Option.map f po in
let comp2 = f b in
let comp3 = f c in
- RLetTuple (loc,nal,(na,comp1),comp2,comp3)
- | RIf (loc,c,(na,po),b1,b2) ->
+ GLetTuple (loc,nal,(na,comp1),comp2,comp3)
+ | GIf (loc,c,(na,po),b1,b2) ->
let comp1 = Option.map f po in
let comp2 = f b1 in
let comp3 = f b2 in
- RIf (loc,f c,(na,comp1),comp2,comp3)
- | RRec (loc,fk,idl,bl,tyl,bv) ->
- let comp1 = Array.map (Util.list_map_left (map_rawdecl_left_to_right f)) bl in
+ GIf (loc,f c,(na,comp1),comp2,comp3)
+ | GRec (loc,fk,idl,bl,tyl,bv) ->
+ let comp1 = Array.map (Util.list_map_left (map_glob_decl_left_to_right f)) bl in
let comp2 = Array.map f tyl in
let comp3 = Array.map f bv in
- RRec (loc,fk,idl,comp1,comp2,comp3)
- | RCast (loc,c,k) ->
+ GRec (loc,fk,idl,comp1,comp2,comp3)
+ | GCast (loc,c,k) ->
let comp1 = f c in
let comp2 = match k with CastConv (k,t) -> CastConv (k, f t) | x -> x in
- RCast (loc,comp1,comp2)
- | (RVar _ | RSort _ | RHole _ | RRef _ | REvar _ | RPatVar _ | RDynamic _) as x -> x
+ GCast (loc,comp1,comp2)
+ | (GVar _ | GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GDynamic _) as x -> x
-let map_rawconstr = map_rawconstr_left_to_right
+let map_glob_constr = map_glob_constr_left_to_right
(*
let name_app f e = function
@@ -154,54 +154,54 @@ let fold_ident g idl e =
(fun id (idl,e) -> let id,e = g id e in (id::idl,e)) idl ([],e)
in (Array.of_list idl,e)
-let map_rawconstr_with_binders_loc loc g f e = function
- | RVar (_,id) -> RVar (loc,id)
- | RApp (_,a,args) -> RApp (loc,f e a, List.map (f e) args)
- | RLambda (_,na,ty,c) ->
- let na,e = name_app g e na in RLambda (loc,na,f e ty,f e c)
- | RProd (_,na,ty,c) ->
- let na,e = name_app g e na in RProd (loc,na,f e ty,f e c)
- | RLetIn (_,na,b,c) ->
- let na,e = name_app g e na in RLetIn (loc,na,f e b,f e c)
- | RCases (_,tyopt,tml,pl) ->
+let map_glob_constr_with_binders_loc loc g f e = function
+ | GVar (_,id) -> GVar (loc,id)
+ | GApp (_,a,args) -> GApp (loc,f e a, List.map (f e) args)
+ | GLambda (_,na,ty,c) ->
+ let na,e = name_app g e na in GLambda (loc,na,f e ty,f e c)
+ | GProd (_,na,ty,c) ->
+ let na,e = name_app g e na in GProd (loc,na,f e ty,f e c)
+ | GLetIn (_,na,b,c) ->
+ let na,e = name_app g e na in GLetIn (loc,na,f e b,f e c)
+ | GCases (_,tyopt,tml,pl) ->
(* We don't modify pattern variable since we don't traverse patterns *)
let g' id e = snd (g id e) in
let h (_,idl,p,c) = (loc,idl,p,f (List.fold_right g' idl e) c) in
- RCases
+ GCases
(loc,Option.map (f e) tyopt,List.map (f e) tml, List.map h pl)
- | RRec (_,fk,idl,tyl,bv) ->
+ | GRec (_,fk,idl,tyl,bv) ->
let idl',e' = fold_ident g idl e in
- RRec (loc,fk,idl',Array.map (f e) tyl,Array.map (f e') bv)
- | RCast (_,c,t) -> RCast (loc,f e c,f e t)
- | RSort (_,x) -> RSort (loc,x)
- | RHole (_,x) -> RHole (loc,x)
- | RRef (_,x) -> RRef (loc,x)
- | REvar (_,x,l) -> REvar (loc,x,l)
- | RPatVar (_,x) -> RPatVar (loc,x)
- | RDynamic (_,x) -> RDynamic (loc,x)
+ GRec (loc,fk,idl',Array.map (f e) tyl,Array.map (f e') bv)
+ | GCast (_,c,t) -> GCast (loc,f e c,f e t)
+ | GSort (_,x) -> GSort (loc,x)
+ | GHole (_,x) -> GHole (loc,x)
+ | GRef (_,x) -> GRef (loc,x)
+ | GEvar (_,x,l) -> GEvar (loc,x,l)
+ | GPatVar (_,x) -> GPatVar (loc,x)
+ | GDynamic (_,x) -> GDynamic (loc,x)
*)
-let fold_rawconstr f acc =
+let fold_glob_constr f acc =
let rec fold acc = function
- | RVar _ -> acc
- | RApp (_,c,args) -> List.fold_left fold (fold acc c) args
- | RLambda (_,_,_,b,c) | RProd (_,_,_,b,c) | RLetIn (_,_,b,c) ->
+ | GVar _ -> acc
+ | GApp (_,c,args) -> List.fold_left fold (fold acc c) args
+ | GLambda (_,_,_,b,c) | GProd (_,_,_,b,c) | GLetIn (_,_,b,c) ->
fold (fold acc b) c
- | RCases (_,_,rtntypopt,tml,pl) ->
+ | GCases (_,_,rtntypopt,tml,pl) ->
List.fold_left fold_pattern
(List.fold_left fold (Option.fold_left fold acc rtntypopt) (List.map fst tml))
pl
- | RLetTuple (_,_,rtntyp,b,c) ->
+ | GLetTuple (_,_,rtntyp,b,c) ->
fold (fold (fold_return_type acc rtntyp) b) c
- | RIf (_,c,rtntyp,b1,b2) ->
+ | GIf (_,c,rtntyp,b1,b2) ->
fold (fold (fold (fold_return_type acc rtntyp) c) b1) b2
- | RRec (_,_,_,bl,tyl,bv) ->
+ | GRec (_,_,_,bl,tyl,bv) ->
let acc = Array.fold_left
(List.fold_left (fun acc (na,k,bbd,bty) ->
fold (Option.fold_left fold acc bbd) bty)) acc bl in
Array.fold_left fold (Array.fold_left fold acc tyl) bv
- | RCast (_,c,k) -> fold (match k with CastConv (_, t) -> fold acc t | CastCoerce -> acc) c
- | (RSort _ | RHole _ | RRef _ | REvar _ | RPatVar _ | RDynamic _) -> acc
+ | GCast (_,c,k) -> fold (match k with CastConv (_, t) -> fold acc t | CastCoerce -> acc) c
+ | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GDynamic _) -> acc
and fold_pattern acc (_,idl,p,c) = fold acc c
@@ -209,25 +209,25 @@ let fold_rawconstr f acc =
in fold acc
-let iter_rawconstr f = fold_rawconstr (fun () -> f) ()
+let iter_glob_constr f = fold_glob_constr (fun () -> f) ()
-let occur_rawconstr id =
+let occur_glob_constr id =
let rec occur = function
- | RVar (loc,id') -> id = id'
- | RApp (loc,f,args) -> (occur f) or (List.exists occur args)
- | RLambda (loc,na,bk,ty,c) -> (occur ty) or ((na <> Name id) & (occur c))
- | RProd (loc,na,bk,ty,c) -> (occur ty) or ((na <> Name id) & (occur c))
- | RLetIn (loc,na,b,c) -> (occur b) or ((na <> Name id) & (occur c))
- | RCases (loc,sty,rtntypopt,tml,pl) ->
+ | GVar (loc,id') -> id = id'
+ | GApp (loc,f,args) -> (occur f) or (List.exists occur args)
+ | GLambda (loc,na,bk,ty,c) -> (occur ty) or ((na <> Name id) & (occur c))
+ | GProd (loc,na,bk,ty,c) -> (occur ty) or ((na <> Name id) & (occur c))
+ | GLetIn (loc,na,b,c) -> (occur b) or ((na <> Name id) & (occur c))
+ | GCases (loc,sty,rtntypopt,tml,pl) ->
(occur_option rtntypopt)
or (List.exists (fun (tm,_) -> occur tm) tml)
or (List.exists occur_pattern pl)
- | RLetTuple (loc,nal,rtntyp,b,c) ->
+ | GLetTuple (loc,nal,rtntyp,b,c) ->
occur_return_type rtntyp id
or (occur b) or (not (List.mem (Name id) nal) & (occur c))
- | RIf (loc,c,rtntyp,b1,b2) ->
+ | GIf (loc,c,rtntyp,b1,b2) ->
occur_return_type rtntyp id or (occur c) or (occur b1) or (occur b2)
- | RRec (loc,fk,idl,bl,tyl,bv) ->
+ | GRec (loc,fk,idl,bl,tyl,bv) ->
not (array_for_all4 (fun fid bl ty bd ->
let rec occur_fix = function
[] -> not (occur ty) && (fid=id or not(occur bd))
@@ -239,8 +239,8 @@ let occur_rawconstr id =
(na=Name id or not(occur_fix bl)) in
occur_fix bl)
idl bl tyl bv)
- | RCast (loc,c,k) -> (occur c) or (match k with CastConv (_, t) -> occur t | CastCoerce -> false)
- | (RSort _ | RHole _ | RRef _ | REvar _ | RPatVar _ | RDynamic _) -> false
+ | GCast (loc,c,k) -> (occur c) or (match k with CastConv (_, t) -> occur t | CastCoerce -> false)
+ | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GDynamic _) -> false
and occur_pattern (loc,idl,p,c) = not (List.mem id idl) & (occur c)
@@ -256,29 +256,29 @@ let add_name_to_ids set na =
| Anonymous -> set
| Name id -> Idset.add id set
-let free_rawvars =
+let free_glob_vars =
let rec vars bounded vs = function
- | RVar (loc,id') -> if Idset.mem id' bounded then vs else Idset.add id' vs
- | RApp (loc,f,args) -> List.fold_left (vars bounded) vs (f::args)
- | RLambda (loc,na,_,ty,c) | RProd (loc,na,_,ty,c) | RLetIn (loc,na,ty,c) ->
+ | GVar (loc,id') -> if Idset.mem id' bounded then vs else Idset.add id' vs
+ | GApp (loc,f,args) -> List.fold_left (vars bounded) vs (f::args)
+ | GLambda (loc,na,_,ty,c) | GProd (loc,na,_,ty,c) | GLetIn (loc,na,ty,c) ->
let vs' = vars bounded vs ty in
let bounded' = add_name_to_ids bounded na in
vars bounded' vs' c
- | RCases (loc,sty,rtntypopt,tml,pl) ->
+ | GCases (loc,sty,rtntypopt,tml,pl) ->
let vs1 = vars_option bounded vs rtntypopt in
let vs2 = List.fold_left (fun vs (tm,_) -> vars bounded vs tm) vs1 tml in
List.fold_left (vars_pattern bounded) vs2 pl
- | RLetTuple (loc,nal,rtntyp,b,c) ->
+ | GLetTuple (loc,nal,rtntyp,b,c) ->
let vs1 = vars_return_type bounded vs rtntyp in
let vs2 = vars bounded vs1 b in
let bounded' = List.fold_left add_name_to_ids bounded nal in
vars bounded' vs2 c
- | RIf (loc,c,rtntyp,b1,b2) ->
+ | GIf (loc,c,rtntyp,b1,b2) ->
let vs1 = vars_return_type bounded vs rtntyp in
let vs2 = vars bounded vs1 c in
let vs3 = vars bounded vs2 b1 in
vars bounded vs3 b2
- | RRec (loc,fk,idl,bl,tyl,bv) ->
+ | GRec (loc,fk,idl,bl,tyl,bv) ->
let bounded' = Array.fold_right Idset.add idl bounded in
let vars_fix i vs fid =
let vs1,bounded1 =
@@ -296,9 +296,9 @@ let free_rawvars =
vars bounded1 vs2 bv.(i)
in
array_fold_left_i vars_fix vs idl
- | RCast (loc,c,k) -> let v = vars bounded vs c in
+ | GCast (loc,c,k) -> let v = vars bounded vs c in
(match k with CastConv (_,t) -> vars bounded v t | _ -> v)
- | (RSort _ | RHole _ | RRef _ | REvar _ | RPatVar _ | RDynamic _) -> vs
+ | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GDynamic _) -> vs
and vars_pattern bounded vs (loc,idl,p,c) =
let bounded' = List.fold_right Idset.add idl bounded in
@@ -315,51 +315,51 @@ let free_rawvars =
Idset.elements vs
-let loc_of_rawconstr = function
- | RRef (loc,_) -> loc
- | RVar (loc,_) -> loc
- | REvar (loc,_,_) -> loc
- | RPatVar (loc,_) -> loc
- | RApp (loc,_,_) -> loc
- | RLambda (loc,_,_,_,_) -> loc
- | RProd (loc,_,_,_,_) -> loc
- | RLetIn (loc,_,_,_) -> loc
- | RCases (loc,_,_,_,_) -> loc
- | RLetTuple (loc,_,_,_,_) -> loc
- | RIf (loc,_,_,_,_) -> loc
- | RRec (loc,_,_,_,_,_) -> loc
- | RSort (loc,_) -> loc
- | RHole (loc,_) -> loc
- | RCast (loc,_,_) -> loc
- | RDynamic (loc,_) -> loc
+let loc_of_glob_constr = function
+ | GRef (loc,_) -> loc
+ | GVar (loc,_) -> loc
+ | GEvar (loc,_,_) -> loc
+ | GPatVar (loc,_) -> loc
+ | GApp (loc,_,_) -> loc
+ | GLambda (loc,_,_,_,_) -> loc
+ | GProd (loc,_,_,_,_) -> loc
+ | GLetIn (loc,_,_,_) -> loc
+ | GCases (loc,_,_,_,_) -> loc
+ | GLetTuple (loc,_,_,_,_) -> loc
+ | GIf (loc,_,_,_,_) -> loc
+ | GRec (loc,_,_,_,_,_) -> loc
+ | GSort (loc,_) -> loc
+ | GHole (loc,_) -> loc
+ | GCast (loc,_,_) -> loc
+ | GDynamic (loc,_) -> loc
(**********************************************************************)
-(* Conversion from rawconstr to cases pattern, if possible *)
+(* Conversion from glob_constr to cases pattern, if possible *)
-let rec cases_pattern_of_rawconstr na = function
- | RVar (loc,id) when na<>Anonymous ->
+let rec cases_pattern_of_glob_constr na = function
+ | GVar (loc,id) when na<>Anonymous ->
(* Unable to manage the presence of both an alias and a variable *)
raise Not_found
- | RVar (loc,id) -> PatVar (loc,Name id)
- | RHole (loc,_) -> PatVar (loc,na)
- | RRef (loc,ConstructRef cstr) ->
+ | GVar (loc,id) -> PatVar (loc,Name id)
+ | GHole (loc,_) -> PatVar (loc,na)
+ | GRef (loc,ConstructRef cstr) ->
PatCstr (loc,cstr,[],na)
- | RApp (loc,RRef (_,ConstructRef cstr),l) ->
- PatCstr (loc,cstr,List.map (cases_pattern_of_rawconstr Anonymous) l,na)
+ | GApp (loc,GRef (_,ConstructRef cstr),l) ->
+ PatCstr (loc,cstr,List.map (cases_pattern_of_glob_constr Anonymous) l,na)
| _ -> raise Not_found
-(* Turn a closed cases pattern into a rawconstr *)
-let rec rawconstr_of_closed_cases_pattern_aux = function
+(* Turn a closed cases pattern into a glob_constr *)
+let rec glob_constr_of_closed_cases_pattern_aux = function
| PatCstr (loc,cstr,[],Anonymous) ->
- RRef (loc,ConstructRef cstr)
+ GRef (loc,ConstructRef cstr)
| PatCstr (loc,cstr,l,Anonymous) ->
- let ref = RRef (loc,ConstructRef cstr) in
- RApp (loc,ref, List.map rawconstr_of_closed_cases_pattern_aux l)
+ let ref = GRef (loc,ConstructRef cstr) in
+ GApp (loc,ref, List.map glob_constr_of_closed_cases_pattern_aux l)
| _ -> raise Not_found
-let rawconstr_of_closed_cases_pattern = function
+let glob_constr_of_closed_cases_pattern = function
| PatCstr (loc,cstr,l,na) ->
- na,rawconstr_of_closed_cases_pattern_aux (PatCstr (loc,cstr,l,Anonymous))
+ na,glob_constr_of_closed_cases_pattern_aux (PatCstr (loc,cstr,l,Anonymous))
| _ ->
raise Not_found
diff --git a/pretyping/rawterm.mli b/pretyping/rawterm.mli
index 1ab5ee7a5..95305d58c 100644
--- a/pretyping/rawterm.mli
+++ b/pretyping/rawterm.mli
@@ -50,29 +50,29 @@ type 'a cast_type =
| CastConv of cast_kind * 'a
| CastCoerce (** Cast to a base type (eg, an underlying inductive type) *)
-type rawconstr =
- | RRef of (loc * global_reference)
- | RVar of (loc * identifier)
- | REvar of loc * existential_key * rawconstr list option
- | RPatVar of loc * (bool * patvar) (** Used for patterns only *)
- | RApp of loc * rawconstr * rawconstr list
- | RLambda of loc * name * binding_kind * rawconstr * rawconstr
- | RProd of loc * name * binding_kind * rawconstr * rawconstr
- | RLetIn of loc * name * rawconstr * rawconstr
- | RCases of loc * case_style * rawconstr option * tomatch_tuples * cases_clauses
- | RLetTuple of loc * name list * (name * rawconstr option) *
- rawconstr * rawconstr
- | RIf of loc * rawconstr * (name * rawconstr option) * rawconstr * rawconstr
- | RRec of loc * fix_kind * identifier array * rawdecl list array *
- rawconstr array * rawconstr array
- | RSort of loc * rawsort
- | RHole of (loc * Evd.hole_kind)
- | RCast of loc * rawconstr * rawconstr cast_type
- | RDynamic of loc * Dyn.t
-
-and rawdecl = name * binding_kind * rawconstr option * rawconstr
-
-and fix_recursion_order = RStructRec | RWfRec of rawconstr | RMeasureRec of rawconstr * rawconstr option
+type glob_constr =
+ | GRef of (loc * global_reference)
+ | GVar of (loc * identifier)
+ | GEvar of loc * existential_key * glob_constr list option
+ | GPatVar of loc * (bool * patvar) (** Used for patterns only *)
+ | GApp of loc * glob_constr * glob_constr list
+ | GLambda of loc * name * binding_kind * glob_constr * glob_constr
+ | GProd of loc * name * binding_kind * glob_constr * glob_constr
+ | GLetIn of loc * name * glob_constr * glob_constr
+ | GCases of loc * case_style * glob_constr option * tomatch_tuples * cases_clauses
+ | GLetTuple of loc * name list * (name * glob_constr option) *
+ glob_constr * glob_constr
+ | GIf of loc * glob_constr * (name * glob_constr option) * glob_constr * glob_constr
+ | GRec of loc * fix_kind * identifier array * glob_decl list array *
+ glob_constr array * glob_constr array
+ | GSort of loc * rawsort
+ | GHole of (loc * Evd.hole_kind)
+ | GCast of loc * glob_constr * glob_constr cast_type
+ | GDynamic of loc * Dyn.t
+
+and glob_decl = name * binding_kind * glob_constr option * glob_constr
+
+and fix_recursion_order = RStructRec | RWfRec of glob_constr | RMeasureRec of glob_constr * glob_constr option
and fix_kind =
| RFix of ((int option * fix_recursion_order) array * int)
@@ -81,42 +81,42 @@ and fix_kind =
and predicate_pattern =
name * (loc * inductive * int * name list) option
-and tomatch_tuple = (rawconstr * predicate_pattern)
+and tomatch_tuple = (glob_constr * predicate_pattern)
and tomatch_tuples = tomatch_tuple list
-and cases_clause = (loc * identifier list * cases_pattern list * rawconstr)
+and cases_clause = (loc * identifier list * cases_pattern list * glob_constr)
and cases_clauses = cases_clause list
val cases_predicate_names : tomatch_tuples -> name list
-val map_rawconstr : (rawconstr -> rawconstr) -> rawconstr -> rawconstr
+val map_glob_constr : (glob_constr -> glob_constr) -> glob_constr -> glob_constr
(* Ensure traversal from left to right *)
-val map_rawconstr_left_to_right :
- (rawconstr -> rawconstr) -> rawconstr -> rawconstr
+val map_glob_constr_left_to_right :
+ (glob_constr -> glob_constr) -> glob_constr -> glob_constr
(*
-val map_rawconstr_with_binders_loc : loc ->
+val map_glob_constr_with_binders_loc : loc ->
(identifier -> 'a -> identifier * 'a) ->
- ('a -> rawconstr -> rawconstr) -> 'a -> rawconstr -> rawconstr
+ ('a -> glob_constr -> glob_constr) -> 'a -> glob_constr -> glob_constr
*)
-val fold_rawconstr : ('a -> rawconstr -> 'a) -> 'a -> rawconstr -> 'a
-val iter_rawconstr : (rawconstr -> unit) -> rawconstr -> unit
-val occur_rawconstr : identifier -> rawconstr -> bool
-val free_rawvars : rawconstr -> identifier list
-val loc_of_rawconstr : rawconstr -> loc
+val fold_glob_constr : ('a -> glob_constr -> 'a) -> 'a -> glob_constr -> 'a
+val iter_glob_constr : (glob_constr -> unit) -> glob_constr -> unit
+val occur_glob_constr : identifier -> glob_constr -> bool
+val free_glob_vars : glob_constr -> identifier list
+val loc_of_glob_constr : glob_constr -> loc
-(** Conversion from rawconstr to cases pattern, if possible
+(** Conversion from glob_constr to cases pattern, if possible
Take the current alias as parameter,
@raise Not_found if translation is impossible *)
-val cases_pattern_of_rawconstr : name -> rawconstr -> cases_pattern
+val cases_pattern_of_glob_constr : name -> glob_constr -> cases_pattern
-val rawconstr_of_closed_cases_pattern : cases_pattern -> name * rawconstr
+val glob_constr_of_closed_cases_pattern : cases_pattern -> name * glob_constr
(** {6 Reduction expressions} *)
diff --git a/proofs/evar_refiner.ml b/proofs/evar_refiner.ml
index e1d41b960..43c7e6e5a 100644
--- a/proofs/evar_refiner.ml
+++ b/proofs/evar_refiner.ml
@@ -43,7 +43,7 @@ let w_refine (evk,evi) (ltac_var,rawc) sigma =
try Pretyping.Default.understand_ltac true sigma env ltac_var
(Pretyping.OfType (Some evi.evar_concl)) rawc
with _ ->
- let loc = Rawterm.loc_of_rawconstr rawc in
+ let loc = Rawterm.loc_of_glob_constr rawc in
user_err_loc
(loc,"",Pp.str ("Instance is not well-typed in the environment of " ^
string_of_existential evk))
diff --git a/proofs/evar_refiner.mli b/proofs/evar_refiner.mli
index 7ce8a54c4..b800d0d66 100644
--- a/proofs/evar_refiner.mli
+++ b/proofs/evar_refiner.mli
@@ -17,7 +17,7 @@ open Rawterm
(** Refinement of existential variables. *)
val w_refine : evar * evar_info ->
- rawconstr_ltac_closure -> evar_map -> evar_map
+ glob_constr_ltac_closure -> evar_map -> evar_map
val instantiate_pf_com :
Evd.evar -> Topconstr.constr_expr -> Evd.evar_map -> Evd.evar_map
diff --git a/proofs/goal.ml b/proofs/goal.ml
index 9f0d48bb1..a48bc2945 100644
--- a/proofs/goal.ml
+++ b/proofs/goal.ml
@@ -175,7 +175,7 @@ module Refinable = struct
asks whether the term should have the same type as the conclusion.
[resolve_classes] is a flag on pretyping functions which, if set to true,
calls the typeclass resolver.
- The principal argument is a [rawconstr] which is then pretyped in the
+ The principal argument is a [glob_constr] which is then pretyped in the
context of a term, the remaining evars are registered to the handle.
It is the main component of the toplevel refine tactic.*)
(* spiwack: it is not entirely satisfactory to have this function here. Plus it is
diff --git a/proofs/goal.mli b/proofs/goal.mli
index b291e1a77..3d9fcc5a2 100644
--- a/proofs/goal.mli
+++ b/proofs/goal.mli
@@ -69,11 +69,11 @@ module Refinable : sig
The [check_type] argument asks whether the term should have the same
type as the conclusion. [resolve_classes] is a flag on pretyping functions
which, if set to true, calls the typeclass resolver.
- The principal argument is a [rawconstr] which is then pretyped in the
+ The principal argument is a [glob_constr] which is then pretyped in the
context of a term, the remaining evars are registered to the handle.
It is the main component of the toplevel refine tactic.*)
val constr_of_raw :
- handle -> bool -> bool -> Rawterm.rawconstr -> Term.constr sensitive
+ handle -> bool -> bool -> Rawterm.glob_constr -> Term.constr sensitive
end
diff --git a/proofs/proof_type.ml b/proofs/proof_type.ml
index ff28ba877..ebb6db213 100644
--- a/proofs/proof_type.ml
+++ b/proofs/proof_type.ml
@@ -94,7 +94,7 @@ type ltac_call_kind =
| LtacNameCall of ltac_constant
| LtacAtomCall of glob_atomic_tactic_expr * atomic_tactic_expr option ref
| LtacVarCall of identifier * glob_tactic_expr
- | LtacConstrInterp of rawconstr *
+ | LtacConstrInterp of glob_constr *
(extended_patvar_map * (identifier * identifier option) list)
type ltac_trace = (int * loc * ltac_call_kind) list
diff --git a/proofs/proof_type.mli b/proofs/proof_type.mli
index 886c0db42..cf73e0dca 100644
--- a/proofs/proof_type.mli
+++ b/proofs/proof_type.mli
@@ -124,7 +124,7 @@ type ltac_call_kind =
| LtacNameCall of ltac_constant
| LtacAtomCall of glob_atomic_tactic_expr * atomic_tactic_expr option ref
| LtacVarCall of identifier * glob_tactic_expr
- | LtacConstrInterp of rawconstr *
+ | LtacConstrInterp of glob_constr *
(extended_patvar_map * (identifier * identifier option) list)
type ltac_trace = (int * loc * ltac_call_kind) list
diff --git a/proofs/tacexpr.ml b/proofs/tacexpr.ml
index 855d2cea7..b9e22ca05 100644
--- a/proofs/tacexpr.ml
+++ b/proofs/tacexpr.ml
@@ -262,8 +262,8 @@ and ('constr,'pat,'cst,'ind,'ref,'id,'tac,'lev) gen_tactic_arg =
(* Globalized tactics *)
and glob_tactic_expr =
- (rawconstr_and_expr,
- rawconstr_and_expr * constr_pattern,
+ (glob_constr_and_expr,
+ glob_constr_and_expr * constr_pattern,
evaluable_global_reference and_short_name or_var,
inductive or_var,
ltac_constant located or_var,
@@ -307,8 +307,8 @@ type raw_red_expr =
(constr_expr, reference or_by_notation, constr_expr) red_expr_gen
type glob_atomic_tactic_expr =
- (rawconstr_and_expr,
- rawconstr_and_expr * constr_pattern,
+ (glob_constr_and_expr,
+ glob_constr_and_expr * constr_pattern,
evaluable_global_reference and_short_name or_var,
inductive or_var,
ltac_constant located or_var,
@@ -317,8 +317,8 @@ type glob_atomic_tactic_expr =
glevel) gen_atomic_tactic_expr
type glob_tactic_arg =
- (rawconstr_and_expr,
- rawconstr_and_expr * constr_pattern,
+ (glob_constr_and_expr,
+ glob_constr_and_expr * constr_pattern,
evaluable_global_reference and_short_name or_var,
inductive or_var,
ltac_constant located or_var,
@@ -329,7 +329,7 @@ type glob_tactic_arg =
type glob_generic_argument = glevel generic_argument
type glob_red_expr =
- (rawconstr_and_expr, evaluable_global_reference or_var, constr_pattern)
+ (glob_constr_and_expr, evaluable_global_reference or_var, constr_pattern)
red_expr_gen
type typed_generic_argument = tlevel generic_argument
diff --git a/proofs/tactic_debug.mli b/proofs/tactic_debug.mli
index 42d81e863..d96f4c746 100644
--- a/proofs/tactic_debug.mli
+++ b/proofs/tactic_debug.mli
@@ -22,7 +22,7 @@ val set_tactic_printer : (glob_tactic_expr ->Pp.std_ppcmds) -> unit
val set_match_pattern_printer :
(env -> constr_pattern match_pattern -> Pp.std_ppcmds) -> unit
val set_match_rule_printer :
- ((Genarg.rawconstr_and_expr * constr_pattern,glob_tactic_expr) match_rule -> Pp.std_ppcmds) ->
+ ((Genarg.glob_constr_and_expr * constr_pattern,glob_tactic_expr) match_rule -> Pp.std_ppcmds) ->
unit
(** Debug information *)
@@ -39,7 +39,7 @@ val db_constr : debug_info -> env -> constr -> unit
(** Prints the pattern rule *)
val db_pattern_rule :
- debug_info -> int -> (Genarg.rawconstr_and_expr * constr_pattern,glob_tactic_expr) match_rule -> unit
+ debug_info -> int -> (Genarg.glob_constr_and_expr * constr_pattern,glob_tactic_expr) match_rule -> unit
(** Prints a matched hypothesis *)
val db_matched_hyp :
diff --git a/tactics/dhyp.ml b/tactics/dhyp.ml
index d9addd1f0..e1e04c8ef 100644
--- a/tactics/dhyp.ml
+++ b/tactics/dhyp.ml
@@ -290,7 +290,7 @@ let applyDestructor cls discard dd gls =
match cl, dd.d_code with
| Some id, (Some x, tac) ->
let arg =
- ConstrMayEval(ConstrTerm (RRef(dummy_loc,VarRef id),None)) in
+ ConstrMayEval(ConstrTerm (GRef(dummy_loc,VarRef id),None)) in
TacLetIn (false, [(dummy_loc, x), arg], tac)
| None, (None, tac) -> tac
| _, (Some _,_) -> error "Destructor expects an hypothesis."
diff --git a/tactics/evar_tactics.mli b/tactics/evar_tactics.mli
index dc8168ca8..40b628315 100644
--- a/tactics/evar_tactics.mli
+++ b/tactics/evar_tactics.mli
@@ -11,7 +11,7 @@ open Names
open Tacexpr
open Termops
-val instantiate : int -> Tacinterp.interp_sign * Rawterm.rawconstr ->
+val instantiate : int -> Tacinterp.interp_sign * Rawterm.glob_constr ->
(identifier * hyp_location_flag, unit) location -> tactic
(*i
diff --git a/tactics/extraargs.ml4 b/tactics/extraargs.ml4
index 25a64c3dd..e31428f7c 100644
--- a/tactics/extraargs.ml4
+++ b/tactics/extraargs.ml4
@@ -103,17 +103,17 @@ let pr_occurrences = pr_occurrences () () ()
let pr_gen prc _prlc _prtac c = prc c
-let pr_rawc _prc _prlc _prtac (_,raw) = Printer.pr_rawconstr raw
+let pr_rawc _prc _prlc _prtac (_,raw) = Printer.pr_glob_constr raw
let pr_raw = pr_rawc () () ()
let interp_raw ist gl (t,_) = (ist,t)
let glob_raw = Tacinterp.intern_constr
-let subst_raw = Tacinterp.subst_rawconstr_and_expr
+let subst_raw = Tacinterp.subst_glob_constr_and_expr
ARGUMENT EXTEND raw
- TYPED AS rawconstr
+ TYPED AS glob_constr
PRINTED BY pr_rawc
INTERPRETED BY interp_raw
@@ -123,7 +123,7 @@ ARGUMENT EXTEND raw
RAW_TYPED AS constr_expr
RAW_PRINTED BY pr_gen
- GLOB_TYPED AS rawconstr_and_expr
+ GLOB_TYPED AS glob_constr_and_expr
GLOB_PRINTED BY pr_gen
[ lconstr(c) ] -> [ c ]
END
diff --git a/tactics/extraargs.mli b/tactics/extraargs.mli
index f27642678..66c251971 100644
--- a/tactics/extraargs.mli
+++ b/tactics/extraargs.mli
@@ -25,9 +25,9 @@ val wit_occurrences : (int list) typed_abstract_argument_type
val pr_occurrences : int list Rawterm.or_var -> Pp.std_ppcmds
val rawwit_raw : constr_expr raw_abstract_argument_type
-val wit_raw : (Tacinterp.interp_sign * rawconstr) typed_abstract_argument_type
+val wit_raw : (Tacinterp.interp_sign * glob_constr) typed_abstract_argument_type
val raw : constr_expr Pcoq.Gram.entry
-val pr_raw : (Tacinterp.interp_sign * Rawterm.rawconstr) -> Pp.std_ppcmds
+val pr_raw : (Tacinterp.interp_sign * Rawterm.glob_constr) -> Pp.std_ppcmds
type 'id gen_place= ('id * hyp_location_flag,unit) location
diff --git a/tactics/extratactics.ml4 b/tactics/extratactics.ml4
index 78a1f51b7..9a9ef164e 100644
--- a/tactics/extratactics.ml4
+++ b/tactics/extratactics.ml4
@@ -540,12 +540,12 @@ let subst_var_with_hole occ tid t =
let occref = if occ > 0 then ref occ else Termops.error_invalid_occurrence [occ] in
let locref = ref 0 in
let rec substrec = function
- | RVar (_,id) as x ->
+ | GVar (_,id) as x ->
if id = tid
then (decr occref; if !occref = 0 then x
- else (incr locref; RHole (make_loc (!locref,0),Evd.QuestionMark(Evd.Define true))))
+ else (incr locref; GHole (make_loc (!locref,0),Evd.QuestionMark(Evd.Define true))))
else x
- | c -> map_rawconstr_left_to_right substrec c in
+ | c -> map_glob_constr_left_to_right substrec c in
let t' = substrec t
in
if !occref > 0 then Termops.error_invalid_occurrence [occ] else t'
@@ -554,10 +554,10 @@ let subst_hole_with_term occ tc t =
let locref = ref 0 in
let occref = ref occ in
let rec substrec = function
- | RHole (_,Evd.QuestionMark(Evd.Define true)) ->
+ | GHole (_,Evd.QuestionMark(Evd.Define true)) ->
decr occref; if !occref = 0 then tc
- else (incr locref; RHole (make_loc (!locref,0),Evd.QuestionMark(Evd.Define true)))
- | c -> map_rawconstr_left_to_right substrec c
+ else (incr locref; GHole (make_loc (!locref,0),Evd.QuestionMark(Evd.Define true)))
+ | c -> map_glob_constr_left_to_right substrec c
in
substrec t
diff --git a/tactics/tacinterp.ml b/tactics/tacinterp.ml
index d4ac859ad..f193c537a 100644
--- a/tactics/tacinterp.ml
+++ b/tactics/tacinterp.ml
@@ -371,12 +371,12 @@ let intern_ltac_variable ist = function
let intern_constr_reference strict ist = function
| Ident (_,id) as r when not strict & find_hyp id ist ->
- RVar (dloc,id), Some (CRef r)
+ GVar (dloc,id), Some (CRef r)
| Ident (_,id) as r when find_ctxvar id ist ->
- RVar (dloc,id), if strict then None else Some (CRef r)
+ GVar (dloc,id), if strict then None else Some (CRef r)
| r ->
let loc,_ as lqid = qualid_of_reference r in
- RRef (loc,locate_global_with_alias lqid), if strict then None else Some (CRef r)
+ GRef (loc,locate_global_with_alias lqid), if strict then None else Some (CRef r)
let intern_move_location ist = function
| MoveAfter id -> MoveAfter (intern_hyp_or_metaid ist id)
@@ -502,7 +502,7 @@ let intern_induction_arg ist = function
if !strict_check then
(* If in a defined tactic, no intros-until *)
match intern_constr ist (CRef (Ident (dloc,id))) with
- | RVar (loc,id),_ -> ElimOnIdent (loc,id)
+ | GVar (loc,id),_ -> ElimOnIdent (loc,id)
| c -> ElimOnConstr (c,NoBindings)
else
ElimOnIdent (loc,id)
@@ -555,7 +555,7 @@ let intern_typed_pattern ist p =
let dummy_pat = PRel 0 in
(* we cannot ensure in non strict mode that the pattern is closed *)
(* keeping a constr_expr copy is too complicated and we want anyway to *)
- (* type it, so we remember the pattern as a rawconstr only *)
+ (* type it, so we remember the pattern as a glob_constr only *)
(intern_constr_gen true false ist p,dummy_pat)
let intern_typed_pattern_with_occurrences ist (l,p) =
@@ -870,7 +870,7 @@ and intern_tacarg strict ist = function
let id = id_of_string s in
if find_ltacvar id ist then
if istac then Reference (ArgVar (adjust_loc loc,id))
- else ConstrMayEval (ConstrTerm (RVar (adjust_loc loc,id), None))
+ else ConstrMayEval (ConstrTerm (GVar (adjust_loc loc,id), None))
else error_syntactic_metavariables_not_allowed loc
| TacCall (loc,f,[]) -> intern_isolated_tactic_reference strict ist f
| TacCall (loc,f,l) ->
@@ -1332,7 +1332,7 @@ let constr_list_of_VList env = function
let interp_constr_in_compound_list inj_fun dest_fun interp_fun ist env sigma l =
let try_expand_ltac_var sigma x =
try match dest_fun x with
- | RVar (_,id), _ ->
+ | GVar (_,id), _ ->
sigma,
List.map inj_fun (constr_list_of_VList env (List.assoc id ist.lfun))
| _ ->
@@ -1417,7 +1417,7 @@ let interp_may_eval f ist gl = function
f ist gl c
with e ->
debugging_exception_step ist false e (fun () ->
- str"interpretation of term " ++ pr_rawconstr_env (pf_env gl) (fst c));
+ str"interpretation of term " ++ pr_glob_constr_env (pf_env gl) (fst c));
raise e
(* Interprets a constr expression possibly to first evaluate *)
@@ -1553,11 +1553,11 @@ let interp_open_constr_with_bindings ist env sigma (c,bl) =
let loc_of_bindings = function
| NoBindings -> dummy_loc
-| ImplicitBindings l -> loc_of_rawconstr (fst (list_last l))
+| ImplicitBindings l -> loc_of_glob_constr (fst (list_last l))
| ExplicitBindings l -> pi1 (list_last l)
let interp_open_constr_with_bindings_loc ist env sigma ((c,_),bl as cb) =
- let loc1 = loc_of_rawconstr c in
+ let loc1 = loc_of_glob_constr c in
let loc2 = loc_of_bindings bl in
let loc = if loc2 = dummy_loc then loc1 else join_loc loc1 loc2 in
let sigma, cb = interp_open_constr_with_bindings ist env sigma cb in
@@ -1594,7 +1594,7 @@ let interp_induction_arg ist gl sigma arg =
if Tactics.is_quantified_hypothesis id gl then
sigma, ElimOnIdent (loc,id)
else
- let c = (RVar (loc,id),Some (CRef (Ident (loc,id)))) in
+ let c = (GVar (loc,id),Some (CRef (Ident (loc,id)))) in
let c = interp_constr ist env sigma c in
sigma, ElimOnConstr (c,NoBindings)
@@ -2068,7 +2068,7 @@ and interp_genarg ist gl x =
in_gen wit_sort
(destSort
(pf_interp_constr ist gl
- (RSort (dloc,out_gen globwit_sort x), None)))
+ (GSort (dloc,out_gen globwit_sort x), None)))
| ConstrArgType ->
in_gen wit_constr (pf_interp_constr ist gl (out_gen globwit_constr x))
| ConstrMayEvalArgType ->
@@ -2539,22 +2539,22 @@ let subst_quantified_hypothesis _ x = x
let subst_declared_or_quantified_hypothesis _ x = x
-let subst_rawconstr_and_expr subst (c,e) =
+let subst_glob_constr_and_expr subst (c,e) =
assert (e=None); (* e<>None only for toplevel tactics *)
- (Detyping.subst_rawconstr subst c,None)
+ (Detyping.subst_glob_constr subst c,None)
-let subst_rawconstr = subst_rawconstr_and_expr (* shortening *)
+let subst_glob_constr = subst_glob_constr_and_expr (* shortening *)
let subst_binding subst (loc,b,c) =
- (loc,subst_quantified_hypothesis subst b,subst_rawconstr subst c)
+ (loc,subst_quantified_hypothesis subst b,subst_glob_constr subst c)
let subst_bindings subst = function
| NoBindings -> NoBindings
- | ImplicitBindings l -> ImplicitBindings (List.map (subst_rawconstr subst) l)
+ | ImplicitBindings l -> ImplicitBindings (List.map (subst_glob_constr subst) l)
| ExplicitBindings l -> ExplicitBindings (List.map (subst_binding subst) l)
let subst_raw_with_bindings subst (c,bl) =
- (subst_rawconstr subst c, subst_bindings subst bl)
+ (subst_glob_constr subst c, subst_bindings subst bl)
let subst_induction_arg subst = function
| ElimOnConstr c -> ElimOnConstr (subst_raw_with_bindings subst c)
@@ -2598,17 +2598,17 @@ let subst_unfold subst (l,e) =
let subst_flag subst red =
{ red with rConst = List.map (subst_evaluable subst) red.rConst }
-let subst_constr_with_occurrences subst (l,c) = (l,subst_rawconstr subst c)
+let subst_constr_with_occurrences subst (l,c) = (l,subst_glob_constr subst c)
-let subst_rawconstr_or_pattern subst (c,p) =
- (subst_rawconstr subst c,subst_pattern subst p)
+let subst_glob_constr_or_pattern subst (c,p) =
+ (subst_glob_constr subst c,subst_pattern subst p)
let subst_pattern_with_occurrences subst (l,p) =
- (l,subst_rawconstr_or_pattern subst p)
+ (l,subst_glob_constr_or_pattern subst p)
let subst_redexp subst = function
| Unfold l -> Unfold (List.map (subst_unfold subst) l)
- | Fold l -> Fold (List.map (subst_rawconstr subst) l)
+ | Fold l -> Fold (List.map (subst_glob_constr subst) l)
| Cbv f -> Cbv (subst_flag subst f)
| Lazy f -> Lazy (subst_flag subst f)
| Pattern l -> Pattern (List.map (subst_constr_with_occurrences subst) l)
@@ -2616,14 +2616,14 @@ let subst_redexp subst = function
| (Red _ | Hnf | ExtraRedExpr _ | CbvVm as r) -> r
let subst_raw_may_eval subst = function
- | ConstrEval (r,c) -> ConstrEval (subst_redexp subst r,subst_rawconstr subst c)
- | ConstrContext (locid,c) -> ConstrContext (locid,subst_rawconstr subst c)
- | ConstrTypeOf c -> ConstrTypeOf (subst_rawconstr subst c)
- | ConstrTerm c -> ConstrTerm (subst_rawconstr subst c)
+ | ConstrEval (r,c) -> ConstrEval (subst_redexp subst r,subst_glob_constr subst c)
+ | ConstrContext (locid,c) -> ConstrContext (locid,subst_glob_constr subst c)
+ | ConstrTypeOf c -> ConstrTypeOf (subst_glob_constr subst c)
+ | ConstrTerm c -> ConstrTerm (subst_glob_constr subst c)
let subst_match_pattern subst = function
- | Subterm (b,ido,pc) -> Subterm (b,ido,(subst_rawconstr_or_pattern subst pc))
- | Term pc -> Term (subst_rawconstr_or_pattern subst pc)
+ | Subterm (b,ido,pc) -> Subterm (b,ido,(subst_glob_constr_or_pattern subst pc))
+ | Term pc -> Term (subst_glob_constr_or_pattern subst pc)
let rec subst_match_goal_hyps subst = function
| Hyp (locs,mp) :: tl ->
@@ -2638,39 +2638,39 @@ let rec subst_atomic subst (t:glob_atomic_tactic_expr) = match t with
(* Basic tactics *)
| TacIntroPattern _ | TacIntrosUntil _ | TacIntroMove _ as x -> x
| TacAssumption as x -> x
- | TacExact c -> TacExact (subst_rawconstr subst c)
- | TacExactNoCheck c -> TacExactNoCheck (subst_rawconstr subst c)
- | TacVmCastNoCheck c -> TacVmCastNoCheck (subst_rawconstr subst c)
+ | TacExact c -> TacExact (subst_glob_constr subst c)
+ | TacExactNoCheck c -> TacExactNoCheck (subst_glob_constr subst c)
+ | TacVmCastNoCheck c -> TacVmCastNoCheck (subst_glob_constr subst c)
| TacApply (a,ev,cb,cl) ->
TacApply (a,ev,List.map (subst_raw_with_bindings subst) cb,cl)
| TacElim (ev,cb,cbo) ->
TacElim (ev,subst_raw_with_bindings subst cb,
Option.map (subst_raw_with_bindings subst) cbo)
- | TacElimType c -> TacElimType (subst_rawconstr subst c)
+ | TacElimType c -> TacElimType (subst_glob_constr subst c)
| TacCase (ev,cb) -> TacCase (ev,subst_raw_with_bindings subst cb)
- | TacCaseType c -> TacCaseType (subst_rawconstr subst c)
+ | TacCaseType c -> TacCaseType (subst_glob_constr subst c)
| TacFix (idopt,n) as x -> x
| TacMutualFix (b,id,n,l) ->
- TacMutualFix(b,id,n,List.map (fun (id,n,c) -> (id,n,subst_rawconstr subst c)) l)
+ TacMutualFix(b,id,n,List.map (fun (id,n,c) -> (id,n,subst_glob_constr subst c)) l)
| TacCofix idopt as x -> x
| TacMutualCofix (b,id,l) ->
- TacMutualCofix (b,id, List.map (fun (id,c) -> (id,subst_rawconstr subst c)) l)
- | TacCut c -> TacCut (subst_rawconstr subst c)
+ TacMutualCofix (b,id, List.map (fun (id,c) -> (id,subst_glob_constr subst c)) l)
+ | TacCut c -> TacCut (subst_glob_constr subst c)
| TacAssert (b,na,c) ->
- TacAssert (Option.map (subst_tactic subst) b,na,subst_rawconstr subst c)
+ TacAssert (Option.map (subst_tactic subst) b,na,subst_glob_constr subst c)
| TacGeneralize cl ->
TacGeneralize (List.map (on_fst (subst_constr_with_occurrences subst))cl)
- | TacGeneralizeDep c -> TacGeneralizeDep (subst_rawconstr subst c)
- | TacLetTac (id,c,clp,b) -> TacLetTac (id,subst_rawconstr subst c,clp,b)
+ | TacGeneralizeDep c -> TacGeneralizeDep (subst_glob_constr subst c)
+ | TacLetTac (id,c,clp,b) -> TacLetTac (id,subst_glob_constr subst c,clp,b)
(* Automation tactics *)
- | TacTrivial (lems,l) -> TacTrivial (List.map (subst_rawconstr subst) lems,l)
- | TacAuto (n,lems,l) -> TacAuto (n,List.map (subst_rawconstr subst) lems,l)
+ | TacTrivial (lems,l) -> TacTrivial (List.map (subst_glob_constr subst) lems,l)
+ | TacAuto (n,lems,l) -> TacAuto (n,List.map (subst_glob_constr subst) lems,l)
| TacAutoTDB n -> TacAutoTDB n
| TacDestructHyp (b,id) -> TacDestructHyp(b,id)
| TacDestructConcl -> TacDestructConcl
| TacSuperAuto (n,l,b1,b2) -> TacSuperAuto (n,l,b1,b2)
- | TacDAuto (n,p,lems) -> TacDAuto (n,p,List.map (subst_rawconstr subst) lems)
+ | TacDAuto (n,p,lems) -> TacDAuto (n,p,List.map (subst_glob_constr subst) lems)
(* Derived basic tactics *)
| TacSimpleInductionDestruct (isrec,h) as x -> x
@@ -2679,13 +2679,13 @@ let rec subst_atomic subst (t:glob_atomic_tactic_expr) = match t with
List.map (subst_induction_arg subst) lc,
Option.map (subst_raw_with_bindings subst) cbo, ids) l, cls))
| TacDoubleInduction (h1,h2) as x -> x
- | TacDecomposeAnd c -> TacDecomposeAnd (subst_rawconstr subst c)
- | TacDecomposeOr c -> TacDecomposeOr (subst_rawconstr subst c)
+ | TacDecomposeAnd c -> TacDecomposeAnd (subst_glob_constr subst c)
+ | TacDecomposeOr c -> TacDecomposeOr (subst_glob_constr subst c)
| TacDecompose (l,c) ->
let l = List.map (subst_or_var (subst_inductive subst)) l in
- TacDecompose (l,subst_rawconstr subst c)
+ TacDecompose (l,subst_glob_constr subst c)
| TacSpecialize (n,l) -> TacSpecialize (n,subst_raw_with_bindings subst l)
- | TacLApply c -> TacLApply (subst_rawconstr subst c)
+ | TacLApply c -> TacLApply (subst_glob_constr subst c)
(* Context management *)
| TacClear _ as x -> x
@@ -2704,12 +2704,12 @@ let rec subst_atomic subst (t:glob_atomic_tactic_expr) = match t with
(* Conversion *)
| TacReduce (r,cl) -> TacReduce (subst_redexp subst r, cl)
| TacChange (op,c,cl) ->
- TacChange (Option.map (subst_rawconstr_or_pattern subst) op,
- subst_rawconstr subst c, cl)
+ TacChange (Option.map (subst_glob_constr_or_pattern subst) op,
+ subst_glob_constr subst c, cl)
(* Equivalence relations *)
| TacReflexivity | TacSymmetry _ as x -> x
- | TacTransitivity c -> TacTransitivity (Option.map (subst_rawconstr subst) c)
+ | TacTransitivity c -> TacTransitivity (Option.map (subst_glob_constr subst) c)
(* Equality and inversion *)
| TacRewrite (ev,l,cl,by) ->
@@ -2718,10 +2718,10 @@ let rec subst_atomic subst (t:glob_atomic_tactic_expr) = match t with
b,m,subst_raw_with_bindings subst c) l,
cl,Option.map (subst_tactic subst) by)
| TacInversion (DepInversion (k,c,l),hyp) ->
- TacInversion (DepInversion (k,Option.map (subst_rawconstr subst) c,l),hyp)
+ TacInversion (DepInversion (k,Option.map (subst_glob_constr subst) c,l),hyp)
| TacInversion (NonDepInversion _,_) as x -> x
| TacInversion (InversionUsing (c,cl),hyp) ->
- TacInversion (InversionUsing (subst_rawconstr subst c,cl),hyp)
+ TacInversion (InversionUsing (subst_glob_constr subst c,cl),hyp)
(* For extensions *)
| TacExtend (_loc,opn,l) ->
@@ -2808,7 +2808,7 @@ and subst_genarg subst (x:glob_generic_argument) =
| SortArgType ->
in_gen globwit_sort (out_gen globwit_sort x)
| ConstrArgType ->
- in_gen globwit_constr (subst_rawconstr subst (out_gen globwit_constr x))
+ in_gen globwit_constr (subst_glob_constr subst (out_gen globwit_constr x))
| ConstrMayEvalArgType ->
in_gen globwit_constr_may_eval (subst_raw_may_eval subst (out_gen globwit_constr_may_eval x))
| QuantHypArgType ->
@@ -2819,7 +2819,7 @@ and subst_genarg subst (x:glob_generic_argument) =
in_gen globwit_red_expr (subst_redexp subst (out_gen globwit_red_expr x))
| OpenConstrArgType b ->
in_gen (globwit_open_constr_gen b)
- ((),subst_rawconstr subst (snd (out_gen (globwit_open_constr_gen b) x)))
+ ((),subst_glob_constr subst (snd (out_gen (globwit_open_constr_gen b) x)))
| ConstrWithBindingsArgType ->
in_gen globwit_constr_with_bindings
(subst_raw_with_bindings subst (out_gen globwit_constr_with_bindings x))
diff --git a/tactics/tacinterp.mli b/tactics/tacinterp.mli
index ee3401a08..ca5acad31 100644
--- a/tactics/tacinterp.mli
+++ b/tactics/tacinterp.mli
@@ -96,11 +96,11 @@ val intern_tactic :
glob_sign -> raw_tactic_expr -> glob_tactic_expr
val intern_constr :
- glob_sign -> constr_expr -> rawconstr_and_expr
+ glob_sign -> constr_expr -> glob_constr_and_expr
val intern_constr_with_bindings :
glob_sign -> constr_expr * constr_expr Rawterm.bindings ->
- rawconstr_and_expr * rawconstr_and_expr Rawterm.bindings
+ glob_constr_and_expr * glob_constr_and_expr Rawterm.bindings
val intern_hyp :
glob_sign -> identifier Util.located -> identifier Util.located
@@ -108,8 +108,8 @@ val intern_hyp :
val subst_genarg :
substitution -> glob_generic_argument -> glob_generic_argument
-val subst_rawconstr_and_expr :
- substitution -> rawconstr_and_expr -> rawconstr_and_expr
+val subst_glob_constr_and_expr :
+ substitution -> glob_constr_and_expr -> glob_constr_and_expr
(** Interprets any expression *)
val val_interp : interp_sign -> goal sigma -> glob_tactic_expr -> value
@@ -127,7 +127,7 @@ val interp_tac_gen : (identifier * value) list -> identifier list ->
val interp_hyp : interp_sign -> goal sigma -> identifier located -> identifier
-val interp_bindings : interp_sign -> Environ.env -> Evd.evar_map -> rawconstr_and_expr Rawterm.bindings -> Evd.evar_map * constr Rawterm.bindings
+val interp_bindings : interp_sign -> Environ.env -> Evd.evar_map -> glob_constr_and_expr Rawterm.bindings -> Evd.evar_map * constr Rawterm.bindings
(** Initial call for interpretation *)
val glob_tactic : raw_tactic_expr -> glob_tactic_expr
diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml
index d14b8cccc..04e0b352c 100644
--- a/toplevel/himsg.ml
+++ b/toplevel/himsg.ml
@@ -855,7 +855,7 @@ let explain_ltac_call_trace (nrep,last,trace,loc) =
let filter =
function (id,None) -> None | (id,Some id') -> Some(id,([],mkVar id')) in
let unboundvars = list_map_filter filter unboundvars in
- quote (pr_rawconstr_env (Global.env()) c) ++
+ quote (pr_glob_constr_env (Global.env()) c) ++
(if unboundvars <> [] or vars <> [] then
strbrk " (with " ++
prlist_with_sep pr_comma
diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml
index 78235f458..1265fe02b 100644
--- a/toplevel/vernacentries.ml
+++ b/toplevel/vernacentries.ml
@@ -798,7 +798,7 @@ let vernac_reserve bl =
let sb_decl = (fun (idl,c) ->
let t = Constrintern.interp_type Evd.empty (Global.env()) c in
let t = Detyping.detype false [] [] t in
- let t = aconstr_of_rawconstr [] [] t in
+ let t = aconstr_of_glob_constr [] [] t in
List.iter (fun id -> Reserve.declare_reserved_type id t) idl)
in List.iter sb_decl bl
@@ -1121,11 +1121,11 @@ let vernac_print = function
| PrintRewriteHintDbName s -> Autorewrite.print_rewrite_hintdb s
| PrintHintDb -> Auto.print_searchtable ()
| PrintScopes ->
- pp (Notation.pr_scopes (Constrextern.without_symbols pr_lrawconstr))
+ pp (Notation.pr_scopes (Constrextern.without_symbols pr_lglob_constr))
| PrintScope s ->
- pp (Notation.pr_scope (Constrextern.without_symbols pr_lrawconstr) s)
+ pp (Notation.pr_scope (Constrextern.without_symbols pr_lglob_constr) s)
| PrintVisibility s ->
- pp (Notation.pr_visibility (Constrextern.without_symbols pr_lrawconstr) s)
+ pp (Notation.pr_visibility (Constrextern.without_symbols pr_lglob_constr) s)
| PrintAbout qid -> msg (print_about qid)
| PrintImplicit qid -> msg (print_impargs qid)
(*spiwack: prints all the axioms and section variables used by a term *)
@@ -1187,7 +1187,7 @@ let vernac_locate = function
| LocateTerm (Genarg.ByNotation (_,ntn,sc)) ->
ppnl
(Notation.locate_notation
- (Constrextern.without_symbols pr_lrawconstr) ntn sc)
+ (Constrextern.without_symbols pr_lglob_constr) ntn sc)
| LocateLibrary qid -> print_located_library qid
| LocateModule qid -> print_located_module qid
| LocateTactic qid -> print_located_tactic qid
diff --git a/toplevel/whelp.ml4 b/toplevel/whelp.ml4
index 906629997..dd947fcda 100644
--- a/toplevel/whelp.ml4
+++ b/toplevel/whelp.ml4
@@ -127,9 +127,9 @@ let uri_params f = function
let get_discharged_hyp_names sp = List.map basename (get_discharged_hyps sp)
let section_parameters = function
- | RRef (_,(ConstructRef ((induri,_),_) | IndRef (induri,_))) ->
+ | GRef (_,(ConstructRef ((induri,_),_) | IndRef (induri,_))) ->
get_discharged_hyp_names (path_of_global (IndRef(induri,0)))
- | RRef (_,(ConstRef cst as ref)) ->
+ | GRef (_,(ConstRef cst as ref)) ->
get_discharged_hyp_names (path_of_global ref)
| _ -> []
@@ -141,33 +141,33 @@ let merge vl al =
let rec uri_of_constr c =
match c with
- | RVar (_,id) -> url_id id
- | RRef (_,ref) -> uri_of_global ref
- | RHole _ | REvar _ -> url_string "?"
- | RSort (_,s) -> url_string (whelp_of_rawsort s)
+ | GVar (_,id) -> url_id id
+ | GRef (_,ref) -> uri_of_global ref
+ | GHole _ | GEvar _ -> url_string "?"
+ | GSort (_,s) -> url_string (whelp_of_rawsort s)
| _ -> url_paren (fun () -> match c with
- | RApp (_,f,args) ->
+ | GApp (_,f,args) ->
let inst,rest = merge (section_parameters f) args in
uri_of_constr f; url_char ' '; uri_params uri_of_constr inst;
url_list_with_sep " " uri_of_constr rest
- | RLambda (_,na,k,ty,c) ->
+ | GLambda (_,na,k,ty,c) ->
url_string "\\lambda "; url_of_name na; url_string ":";
uri_of_constr ty; url_string "."; uri_of_constr c
- | RProd (_,Anonymous,k,ty,c) ->
+ | GProd (_,Anonymous,k,ty,c) ->
uri_of_constr ty; url_string "\\to "; uri_of_constr c
- | RProd (_,Name id,k,ty,c) ->
+ | GProd (_,Name id,k,ty,c) ->
url_string "\\forall "; url_id id; url_string ":";
uri_of_constr ty; url_string "."; uri_of_constr c
- | RLetIn (_,na,b,c) ->
+ | GLetIn (_,na,b,c) ->
url_string "let "; url_of_name na; url_string "\\def ";
uri_of_constr b; url_string " in "; uri_of_constr c
- | RCast (_,c, CastConv (_,t)) ->
+ | GCast (_,c, CastConv (_,t)) ->
uri_of_constr c; url_string ":"; uri_of_constr t
- | RRec _ | RIf _ | RLetTuple _ | RCases _ ->
+ | GRec _ | GIf _ | GLetTuple _ | GCases _ ->
error "Whelp does not support pattern-matching and (co-)fixpoint."
- | RVar _ | RRef _ | RHole _ | REvar _ | RSort _ | RCast (_,_, CastCoerce) ->
+ | GVar _ | GRef _ | GHole _ | GEvar _ | GSort _ | GCast (_,_, CastCoerce) ->
anomaly "Written w/o parenthesis"
- | RPatVar _ | RDynamic _ ->
+ | GPatVar _ | GDynamic _ ->
anomaly "Found constructors not supported in constr") ()
let make_string f x = Buffer.reset b; f x; Buffer.contents b