aboutsummaryrefslogtreecommitdiffhomepage
path: root/parsing
diff options
context:
space:
mode:
authorGravatar herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7>2002-11-14 18:37:54 +0000
committerGravatar herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7>2002-11-14 18:37:54 +0000
commite88e0b2140bdd2d194a52bc09f8338b5667d0f92 (patch)
tree67ca22f77ddb98725456e5f9a0b5ad613ae28da5 /parsing
parente4efb857fa9053c41e4c030256bd17de7e24542f (diff)
Réforme de l'interprétation des termes :
- Le parsing se fait maintenant via "constr_expr" au lieu de "Coqast.t" - "Coqast.t" reste pour l'instant pour le pretty-printing. Un deuxième pretty-printer dans ppconstr.ml est basé sur "constr_expr". - Nouveau répertoire "interp" qui hérite de la partie interprétation qui se trouvait avant dans "parsing" (constrintern.ml remplace astterm.ml; constrextern.ml est l'équivalent de termast.ml pour le nouveau printer; topconstr.ml; contient la définition de "constr_expr"; modintern.ml remplace astmod.ml) - Libnames.reference tend à remplacer Libnames.qualid git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@3235 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'parsing')
-rw-r--r--parsing/argextend.ml410
-rwxr-xr-xparsing/ast.ml251
-rwxr-xr-xparsing/ast.mli72
-rw-r--r--parsing/astmod.ml133
-rw-r--r--parsing/astmod.mli25
-rw-r--r--parsing/astterm.ml949
-rw-r--r--parsing/astterm.mli101
-rw-r--r--parsing/coqast.ml86
-rw-r--r--parsing/coqast.mli62
-rw-r--r--parsing/coqlib.ml285
-rw-r--r--parsing/coqlib.mli133
-rw-r--r--parsing/egrammar.ml160
-rw-r--r--parsing/egrammar.mli9
-rw-r--r--parsing/esyntax.ml37
-rw-r--r--parsing/esyntax.mli6
-rw-r--r--parsing/extend.ml101
-rw-r--r--parsing/extend.mli61
-rw-r--r--parsing/g_basevernac.ml4109
-rw-r--r--parsing/g_cases.ml459
-rw-r--r--parsing/g_constr.ml4349
-rw-r--r--parsing/g_ltac.ml449
-rw-r--r--parsing/g_minicoq.ml42
-rw-r--r--parsing/g_module.ml466
-rw-r--r--parsing/g_prim.ml447
-rw-r--r--parsing/g_proofs.ml433
-rw-r--r--parsing/g_rsyntax.ml57
-rw-r--r--parsing/g_tactic.ml472
-rw-r--r--parsing/g_vernac.ml4157
-rw-r--r--parsing/g_zsyntax.ml69
-rw-r--r--parsing/g_zsyntax.mli5
-rw-r--r--parsing/genarg.ml181
-rw-r--r--parsing/genarg.mli208
-rw-r--r--parsing/pcoq.ml4244
-rw-r--r--parsing/pcoq.mli106
-rw-r--r--parsing/ppconstr.ml248
-rw-r--r--parsing/ppconstr.mli28
-rw-r--r--parsing/pptactic.ml54
-rw-r--r--parsing/pptactic.mli20
-rw-r--r--parsing/prettyp.ml6
-rw-r--r--parsing/prettyp.mli8
-rw-r--r--parsing/printer.ml2
-rw-r--r--parsing/q_coqast.ml487
-rw-r--r--parsing/search.ml1
-rw-r--r--parsing/symbols.ml320
-rw-r--r--parsing/symbols.mli60
-rw-r--r--parsing/tacextend.ml48
-rw-r--r--parsing/termast.ml19
-rw-r--r--parsing/vernacextend.ml47
48 files changed, 1293 insertions, 3869 deletions
diff --git a/parsing/argextend.ml4 b/parsing/argextend.ml4
index 54eeca754..320774836 100644
--- a/parsing/argextend.ml4
+++ b/parsing/argextend.ml4
@@ -24,7 +24,8 @@ let rec make_rawwit loc = function
| StringArgType -> <:expr< Genarg.rawwit_string >>
| PreIdentArgType -> <:expr< Genarg.rawwit_pre_ident >>
| IdentArgType -> <:expr< Genarg.rawwit_ident >>
- | QualidArgType -> <:expr< Genarg.rawwit_qualid >>
+ | RefArgType -> <:expr< Genarg.rawwit_ref >>
+ | SortArgType -> <:expr< Genarg.rawwit_sort >>
| ConstrArgType -> <:expr< Genarg.rawwit_constr >>
| ConstrMayEvalArgType -> <:expr< Genarg.rawwit_constr_may_eval >>
| QuantHypArgType -> <:expr< Genarg.rawwit_quant_hyp >>
@@ -46,8 +47,9 @@ let rec make_wit loc = function
| StringArgType -> <:expr< Genarg.wit_string >>
| PreIdentArgType -> <:expr< Genarg.wit_pre_ident >>
| IdentArgType -> <:expr< Genarg.wit_ident >>
- | QualidArgType -> <:expr< Genarg.wit_qualid >>
+ | RefArgType -> <:expr< Genarg.wit_ref >>
| QuantHypArgType -> <:expr< Genarg.wit_quant_hyp >>
+ | SortArgType -> <:expr< Genarg.wit_sort >>
| ConstrArgType -> <:expr< Genarg.wit_constr >>
| ConstrMayEvalArgType -> <:expr< Genarg.wit_constr_may_eval >>
| TacticArgType -> <:expr< Genarg.wit_tactic >>
@@ -148,9 +150,7 @@ let rec interp_entry_name loc s =
| None -> None, <:expr< $lid:s$ >> in
let t =
match t with
- | Some (GenAstType t) -> t
- | Some _ ->
- failwith "Only entries of generic type can be used in extension"
+ | Some t -> t
| None ->
(* Pp.warning_with Pp_control.err_ft
("Unknown primitive grammar entry: "^s);*)
diff --git a/parsing/ast.ml b/parsing/ast.ml
index 52a390af1..ae677979f 100755
--- a/parsing/ast.ml
+++ b/parsing/ast.ml
@@ -13,12 +13,11 @@ open Util
open Names
open Libnames
open Coqast
+open Topconstr
open Genarg
let isMeta s = String.length s <> 0 & s.[0]='$'
-let dummy_loc = (0,0)
-
let loc = function
| Node (loc,_,_) -> loc
| Nvar (loc,_) -> loc
@@ -31,17 +30,6 @@ let loc = function
| Path (loc,_) -> loc
| Dynamic (loc,_) -> loc
-type entry_type =
- | PureAstType
- | IntAstType
- | IdentAstType
- | AstListType
- | TacticAtomAstType
- | ThmTokenAstType
- | DynamicAstType
- | ReferenceAstType
- | GenAstType of Genarg.argument_type
-
(* patterns of ast *)
type astpat =
| Pquote of t
@@ -107,19 +95,28 @@ let id_of_ast = function
(* semantic actions of grammar rules *)
type act =
- | Act of pat
+ | Act of constr_expr
| ActCase of act * (pat * act) list
| ActCaseList of act * (pat * act) list
(* values associated to variables *)
+(*
+type typed_ast =
+ | AstListNode of Coqast.t list
+ | PureAstNode of Coqast.t
+*)
type typed_ast =
| AstListNode of Coqast.t list
| PureAstNode of Coqast.t
type ast_action_type = ETast | ETastl
+type dynamic_grammar =
+ | ConstrNode of constr_expr
+ | CasesPatternNode of cases_pattern_expr
+
type grammar_action =
- | SimpleAction of loc * typed_ast
+ | SimpleAction of loc * dynamic_grammar
| CaseAction of
loc * grammar_action * ast_action_type * (t list * grammar_action) list
@@ -211,56 +208,25 @@ let rec coerce_to_var = function
(loc ast,"Ast.coerce_to_var",
(str"This expression should be a simple identifier"))
-let coerce_to_id a = match coerce_to_var a with
+let coerce_to_id_ast a = match coerce_to_var a with
| Nvar (_,id) -> id
| ast -> user_err_loc
(loc ast,"Ast.coerce_to_id",
str"This expression should be a simple identifier")
-let coerce_qualid_to_id (loc,qid) = match repr_qualid qid with
- | dir, id when dir = empty_dirpath -> id
- | _ ->
- user_err_loc (loc, "Ast.coerce_qualid_to_id",
- str"This expression should be a simple identifier")
+let coerce_to_id = function
+ | CRef (Ident (_,id)) -> id
+ | a -> user_err_loc
+ (constr_loc a,"Ast.coerce_to_id",
+ str"This expression should be a simple identifier")
let coerce_reference_to_id = function
- | RIdent (_,id) -> id
- | RQualid (loc,_) ->
+ | Ident (_,id) -> id
+ | Qualid (loc,_) ->
user_err_loc (loc, "Ast.coerce_reference_to_id",
str"This expression should be a simple identifier")
-(* This is to interpret the macro $ABSTRACT used in binders *)
-(* $ABSTRACT should occur in this configuration : *)
-(* ($ABSTRACT name (s1 a1 ($LIST l1)) ... (s2 an ($LIST ln)) b) *)
-(* where li is id11::...::id1p1 and it produces the ast *)
-(* (s1' a1 [id11]...[id1p1](... (sn' an [idn1]...[idnpn]b)...)) *)
-(* where s1' is overwritten by name if s1 is $BINDER otherwise s1 *)
-
-let slam_ast (_,fin) id ast =
- match id with
- | Coqast.Nvar ((deb,_), s) ->
- let name = if s = id_of_string "_" then None else Some s in
- Coqast.Slam ((deb,fin), name, ast)
- | Coqast.Nmeta ((deb,_), s) -> Coqast.Smetalam ((deb,fin), s, ast)
- | _ -> invalid_arg "Ast.slam_ast"
-
-let abstract_binder_ast (_,fin as loc) name a b =
- match a with
- | Coqast.Node((deb,_),s,d::l) ->
- let s' = if s="BINDER" then name else s in
- Coqast.Node((deb,fin),s', [d; List.fold_right (slam_ast loc) l b])
- | _ -> invalid_arg "Bad usage of $ABSTRACT macro"
-
-let abstract_binders_ast loc name a b =
- match a with
- | Coqast.Node(_,"BINDERS",l) ->
- List.fold_right (abstract_binder_ast loc name) l b
- | _ -> invalid_arg "Bad usage of $ABSTRACT macro"
-
-let mkCastC(a,b) = ope("CAST",[a;b])
-let mkLambdaC(x,a,b) = ope("LAMBDA",[a;slam(Some x,b)])
-let mkLetInC(x,a,b) = ope("LETIN",[a;slam(Some x,b)])
-let mkProdC (x,a,b) = ope("PROD",[a;slam(Some x,b)])
+let coerce_global_to_id = coerce_reference_to_id
(* Pattern-matching on ast *)
@@ -284,8 +250,8 @@ let env_assoc sigma k (loc,v) =
let env_assoc_nvars sigma (dloc,v) =
match env_assoc_value dloc v sigma with
- | AstListNode al -> List.map coerce_to_id al
- | PureAstNode ast -> [coerce_to_id ast]
+ | AstListNode al -> List.map coerce_to_id_ast al
+ | PureAstNode ast -> [coerce_to_id_ast ast]
let build_lams dloc idl ast =
List.fold_right (fun id lam -> Slam(dloc,Some id,lam)) idl ast
@@ -488,7 +454,9 @@ let rec pat_of_ast env ast =
| Node(_,op,args) ->
let (pargs, env') = patl_of_astl env args in
(Pnode(op,pargs), env')
- | (Path _|Num _|Id _|Str _|Nvar _) -> (Pquote (set_loc dummy_loc ast), env)
+(* Compatibility with new parsing mode *)
+ | Nvar(loc,id) when (string_of_id id).[0] = '$' -> make_astvar env loc (string_of_id id) Tany
+ | (Path _|Num _|Id _|Str _ |Nvar _) -> (Pquote (set_loc dummy_loc ast), env)
| Dynamic(loc,_) ->
invalid_arg_loc(loc,"pat_of_ast: dynamic")
@@ -505,27 +473,15 @@ and patl_of_astl env astl =
type entry_env = (string * ast_action_type) list
-(*
-let to_pat env = function
- | AstListNode al -> let p,e = patl_of_astl env al in AstListPat p, e
- | PureAstNode a -> let p,e = pat_of_ast env a in PureAstPat p, e
-*)
-
let to_pat = pat_of_ast
-(*
- match ast with
- | Node(_,"ASTPAT",[p]) -> pat_of_ast env p
- | _ -> invalid_arg_loc (loc ast,"Ast.to_pat")
-*)
-
-
(* Substitution *)
(* Locations in quoted ast are wrong (they refer to the right hand
side of a grammar rule). A default location dloc is used whenever
we create an ast constructor. Locations in the binding list are trusted. *)
+(* For old ast printer *)
let rec pat_sub dloc sigma pat =
match pat with
| Pmeta(pv,c) -> env_assoc sigma c (dloc,pv)
@@ -549,6 +505,7 @@ and patl_sub dloc sigma pl =
(* Converting and checking free meta-variables *)
+(* For old ast printer *)
let type_of_meta env loc pv =
try
List.assoc pv env
@@ -556,6 +513,7 @@ let type_of_meta env loc pv =
user_err_loc (loc,"Ast.type_of_meta",
(str"variable " ++ str pv ++ str" is unbound"))
+(* For old ast printer *)
let check_ast_meta env loc pv =
match type_of_meta env loc pv with
| ETast -> ()
@@ -563,6 +521,7 @@ let check_ast_meta env loc pv =
user_err_loc (loc,"Ast.check_ast_meta",
(str"variable " ++ str pv ++ str" is not of ast type"))
+(* For old ast printer *)
let rec val_of_ast env = function
| Nmeta(loc,pv) ->
check_ast_meta env loc pv;
@@ -593,48 +552,8 @@ and vall_of_astl env = function
str"variable " ++ str pv ++ str" is not a List")
| ast::asttl -> Pcons (val_of_ast env ast, vall_of_astl env asttl)
| [] -> Pnil
-(*
-let rec val_of_ast_constr env = function
-(*
- | ConstrEval (r,c) -> ConstrEvalPat (r,val_of_ast_constr env c)
- | ConstrContext (x,c) -> ConstrContextPat (x,val_of_ast_constr env c)
-*)
- | ConstrTerm c -> ConstrTermPat (val_of_ast env c)
-*)
-(*
-let rec check_pat_meta env = function
- | Pquote _ -> ()
- | Pmeta(s,Tany) -> check_ast_meta env loc s
- | Pmeta(s,_) -> anomaly "not well-formed pattern"
- | Pmeta_slam(s,b) ->
- let _ = type_of_meta env loc s in (* ids are coerced to id lists *)
- check_pat_meta env b
- | Pslam(_,b) -> check_pat_meta env b
- | Pnode(op,Plmeta (locv,pv)) ->
- if type_of_meta env locv pv <> ETastl then
- user_err_loc (locv,"Ast.vall_of_astl",
- [< 'sTR"variable "; 'sTR pv; 'sTR" is not a List" >])
- | Pnode(op,l) -> check_patlist_meta env l
-
-and check_patlist_meta env = function
- | Plmeta (locv,pv) ->
- if type_of_meta env locv pv <> ETastl then
- user_err_loc (locv,"Ast.vall_of_astl",
- [< 'sTR"variable "; 'sTR pv; 'sTR" is not a List" >])
- | Pcons(Pmeta(pv,Tlist),l) ->
- if l = Pnil then anomaly "not well-formed pattern list";
- if type_of_meta env locv pv <> ETastl then
- user_err_loc (locv,"Ast.vall_of_astl",
- [< 'sTR"variable "; 'sTR pv; 'sTR" is not a List" >])
- else check_patlist_meta env l
- | Pcons(p,l) -> check_pat_meta env p; check_patlist_meta env l
- | Pnil -> ()
-
-let check_typed_pat_meta env = function
- | AstListPat cl -> check_patlist_meta env cl
- | PureAstPat c -> check_pat_meta env c
-*)
+(* For old ast printer *)
let rec occur_var_ast s = function
| Node(_,"QUALID",_::_::_) -> false
| Node(_,"QUALID",[Nvar(_,s2)]) -> s = s2
@@ -645,104 +564,9 @@ let rec occur_var_ast s = function
| Id _ | Str _ | Num _ | Path _ -> false
| Dynamic _ -> (* Hum... what to do here *) false
-let rec replace_vars_ast l = function
- | Node(loc,op,args) -> Node (loc,op, List.map (replace_vars_ast l) args)
- | Nvar(loc,s) as a -> (try Nvar (loc, List.assoc s l) with Not_found -> a)
- | Smetalam _ | Nmeta _ -> anomaly "replace_var: metas should not occur here"
- | Slam(loc,None,body) -> Slam(loc,None,replace_vars_ast l body)
- | Slam(loc,Some s,body) as a ->
- if List.mem_assoc s l then a else
- Slam(loc,Some s,replace_vars_ast l body)
- | Id _ | Str _ | Num _ | Path _ as a -> a
- | Dynamic _ as a -> (* Hum... what to do here *) a
-
-(* Ast with cases and metavariables *)
-
-let print_sig = function
- | [] ->
- mt ()
- | sigma ->
- str"with constraints :" ++ brk(1,1) ++
- v 0 (prlist_with_sep pr_spc
- (fun (x,v) -> str x ++ str" = " ++ hov 0 (print_val v))
- sigma)
-
-let case_failed loc sigma e pats =
- user_err_loc
- (loc,"Ast.eval_act",
- str"Grammar case failure. The ast" ++ spc () ++ print_ast e ++
- spc () ++ str"does not match any of the patterns :" ++
- brk(1,1) ++ v 0 (prlist_with_sep pr_spc print_astpat pats) ++ fnl () ++
- print_sig sigma)
-
-let caselist_failed loc sigma el pats =
- user_err_loc
- (loc,"Ast.eval_act",
- str"Grammar case failure. The ast list" ++ brk(1,1) ++ print_astl el ++
- spc () ++ str"does not match any of the patterns :" ++
- brk(1,1) ++ v 0 (prlist_with_sep pr_spc print_astlpat pats) ++ fnl () ++
- print_sig sigma)
-
-let myfst = function
- | PureAstPat p, a -> p
- | _ -> error "Expects a pure ast"
-
-let myfstl = function
- | AstListPat p, a -> p
- | _ -> error "Expects an ast list"
-
-let rec eval_act dloc sigma = function
- | Act (AstListPat patl) -> AstListNode (patl_sub dloc sigma patl)
- | Act (PureAstPat pat) -> PureAstNode (pat_sub dloc sigma pat)
- | ActCase(e,ml) ->
- (match eval_act dloc sigma e with
- | (PureAstNode esub) ->
- (match first_match myfst sigma esub ml with
- | Some((_,a),sigma_pat) -> eval_act dloc sigma_pat a
- | _ -> case_failed dloc sigma esub (List.map myfst ml))
- | _ -> grammar_type_error (dloc,"Ast.eval_act"))
- | ActCaseList(e,ml) ->
- (match eval_act dloc sigma e with
- | AstListNode elsub ->
- (match first_matchl myfstl sigma elsub ml with
- | Some((_,a),sigma_pat) -> eval_act dloc sigma_pat a
- | _ -> caselist_failed dloc sigma elsub (List.map myfstl ml))
- | _ -> grammar_type_error (dloc,"Ast.eval_act"))
-
-let val_of_typed_ast loc env = function
- | ETast, PureAstNode c -> PureAstPat (val_of_ast env c)
- | ETastl, AstListNode cl -> AstListPat (vall_of_astl env cl)
- | (ETast|ETastl), _ ->
- invalid_arg_loc (loc,"Ast.act_of_ast: ill-typed")
-
-(* TODO: case sur des variables uniquement -> pas de pb de conflit Ast/List *)
-let rec act_of_ast vars etyp = function
- | CaseAction (loc,a,atyp,cl) ->
- let pa = act_of_ast vars atyp a in
- (match atyp with
- | ETastl ->
- let acl = List.map (caselist vars etyp) cl in
- ActCaseList (pa,acl)
- | _ ->
- let acl = List.map (case loc vars etyp) cl in
- ActCase (pa,acl))
- | SimpleAction (loc,a) -> Act (val_of_typed_ast loc vars (etyp,a))
-
-and case loc vars etyp = function
- | [p],a ->
- let (apl,penv) = pat_of_ast vars p in
- let aa = act_of_ast penv etyp a in
- (PureAstPat apl,aa)
- | _ ->
- user_err_loc
- (loc, "Ast.case", str"case pattern for an ast should be a single ast")
-
-and caselist vars etyp (pl,a) =
- let (apl,penv) = patl_of_astl vars pl in
- let aa = act_of_ast penv etyp a in
- (AstListPat apl,aa)
-let to_act_check_vars = act_of_ast
+(**********************************************************************)
+(* Object substitution in modules *)
let rec subst_astpat subst = function
| Pquote a -> Pquote (subst_ast subst a)
@@ -758,12 +582,3 @@ and subst_astpatlist subst = function
let subst_pat subst = function
| AstListPat pl -> AstListPat (subst_astpatlist subst pl)
| PureAstPat p -> PureAstPat (subst_astpat subst p)
-
-let rec subst_act subst = function
- | Act p -> Act (subst_pat subst p)
- | ActCase (a,l) ->
- ActCase (subst_act subst a,
- List.map (fun (p,a) -> subst_pat subst p, subst_act subst a) l)
- | ActCaseList (a,l) ->
- ActCaseList (subst_act subst a,
- List.map (fun (p,a) -> subst_pat subst p, subst_act subst a) l)
diff --git a/parsing/ast.mli b/parsing/ast.mli
index 9fd8e9cc9..1faaf78a7 100755
--- a/parsing/ast.mli
+++ b/parsing/ast.mli
@@ -10,19 +10,17 @@
(*i*)
open Pp
+open Util
open Names
open Libnames
open Coqast
+open Topconstr
open Genarg
(*i*)
(* Abstract syntax trees. *)
-val dummy_loc : Coqast.loc
-val loc : Coqast.t -> Coqast.loc
-(*
-val vernac_loc : Coqast.vernac_ast -> Coqast.loc
-*)
+val loc : Coqast.t -> loc
(* ast constructors with dummy location *)
val ope : string * Coqast.t list -> Coqast.t
@@ -34,9 +32,9 @@ val string : string -> Coqast.t
val path : kernel_name -> Coqast.t
val dynamic : Dyn.t -> Coqast.t
-val set_loc : Coqast.loc -> Coqast.t -> Coqast.t
+val set_loc : loc -> Coqast.t -> Coqast.t
-val path_section : Coqast.loc -> kernel_name -> Coqast.t
+val path_section : loc -> kernel_name -> Coqast.t
val section_path : kernel_name -> kernel_name
(* ast destructors *)
@@ -45,19 +43,6 @@ val id_of_ast : Coqast.t -> string
val nvar_of_ast : Coqast.t -> identifier
val meta_of_ast : Coqast.t -> string
-(* ast processing datatypes *)
-
-type entry_type =
- | PureAstType
- | IntAstType
- | IdentAstType
- | AstListType
- | TacticAtomAstType
- | ThmTokenAstType
- | DynamicAstType
- | ReferenceAstType
- | GenAstType of Genarg.argument_type
-
(* patterns of ast *)
type astpat =
| Pquote of t
@@ -79,7 +64,7 @@ type pat =
(* semantic actions of grammar rules *)
type act =
- | Act of pat
+ | Act of constr_expr
| ActCase of act * (pat * act) list
| ActCaseList of act * (pat * act) list
@@ -90,28 +75,21 @@ type typed_ast =
type ast_action_type = ETast | ETastl
+type dynamic_grammar =
+ | ConstrNode of constr_expr
+ | CasesPatternNode of cases_pattern_expr
+
type grammar_action =
- | SimpleAction of loc * typed_ast
+ | SimpleAction of loc * dynamic_grammar
| CaseAction of
loc * grammar_action * ast_action_type * (t list * grammar_action) list
type env = (string * typed_ast) list
-val coerce_to_var : Coqast.t -> Coqast.t
-
-val coerce_to_id : Coqast.t -> identifier
-
-val coerce_qualid_to_id : qualid Util.located -> identifier
-
-val coerce_reference_to_id : reference_expr -> identifier
-
-val abstract_binders_ast :
- Coqast.loc -> string -> Coqast.t -> Coqast.t -> Coqast.t
+val coerce_to_id : constr_expr -> identifier
-val mkCastC : Coqast.t * Coqast.t -> Coqast.t
-val mkLambdaC : identifier * Coqast.t * Coqast.t -> Coqast.t
-val mkLetInC : identifier * Coqast.t * Coqast.t -> Coqast.t
-val mkProdC : identifier * Coqast.t * Coqast.t -> Coqast.t
+val coerce_global_to_id : reference -> identifier
+val coerce_reference_to_id : reference -> identifier
exception No_match of string
@@ -126,32 +104,20 @@ val print_astlpat : patlist -> std_ppcmds
type entry_env = (string * ast_action_type) list
-val grammar_type_error : Coqast.loc * string -> 'a
+val grammar_type_error : loc * string -> 'a
(* Converting and checking free meta-variables *)
-val pat_sub : Coqast.loc -> env -> astpat -> Coqast.t
-val val_of_ast : entry_env -> Coqast.t -> astpat
-val vall_of_astl : entry_env -> Coqast.t list -> patlist
-
-val pat_of_ast : entry_env -> Coqast.t -> astpat * entry_env
+(* For old ast printer *)
+val pat_sub : loc -> env -> astpat -> Coqast.t
+val val_of_ast : entry_env -> Coqast.t -> astpat
val alpha_eq : Coqast.t * Coqast.t -> bool
val alpha_eq_val : typed_ast * typed_ast -> bool
-
val occur_var_ast : identifier -> Coqast.t -> bool
-val replace_vars_ast : (identifier * identifier) list -> Coqast.t -> Coqast.t
-
-val bind_env : env -> string -> typed_ast -> env
-val ast_match : env -> astpat -> Coqast.t -> env
-val astl_match : env -> patlist -> Coqast.t list -> env
val find_all_matches : ('a -> astpat) -> env -> t -> 'a list -> ('a * env) list
val first_matchl : ('a -> patlist) -> env -> Coqast.t list -> 'a list ->
('a * env) option
-
val to_pat : entry_env -> Coqast.t -> (astpat * entry_env)
-val eval_act : Coqast.loc -> env -> act -> typed_ast
-val to_act_check_vars : entry_env -> ast_action_type -> grammar_action -> act
-
+(* Object substitution in modules *)
val subst_astpat : Names.substitution -> astpat -> astpat
-val subst_act : Names.substitution -> act -> act
diff --git a/parsing/astmod.ml b/parsing/astmod.ml
deleted file mode 100644
index cbb19fa0b..000000000
--- a/parsing/astmod.ml
+++ /dev/null
@@ -1,133 +0,0 @@
-(***********************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
-(* \VV/ *************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(***********************************************************************)
-
-(* $Id$ *)
-
-open Pp
-open Util
-open Names
-open Entries
-open Libnames
-open Coqast
-open Astterm
-
-let rec make_mp mp = function
- [] -> mp
- | h::tl -> make_mp (MPdot(mp, label_of_id h)) tl
-
-(*
-(* Since module components are not put in the nametab we try to locate
-the module prefix *)
-exception BadRef
-
-let lookup_qualid (modtype:bool) qid =
- let rec make_mp mp = function
- [] -> mp
- | h::tl -> make_mp (MPdot(mp, label_of_id h)) tl
- in
- let rec find_module_prefix dir n =
- if n<0 then raise Not_found;
- let dir',dir'' = list_chop n dir in
- let id',dir''' =
- match dir'' with
- | hd::tl -> hd,tl
- | _ -> anomaly "This list should not be empty!"
- in
- let qid' = make_qualid dir' id' in
- try
- match Nametab.locate qid' with
- | ModRef mp -> mp,dir'''
- | _ -> raise BadRef
- with
- Not_found -> find_module_prefix dir (pred n)
- in
- try Nametab.locate qid
- with Not_found ->
- let (dir,id) = repr_qualid qid in
- let pref_mp,dir' = find_module_prefix dir (List.length dir - 1) in
- let mp =
- List.fold_left (fun mp id -> MPdot (mp, label_of_id id)) pref_mp dir'
- in
- if modtype then
- ModTypeRef (make_ln mp (label_of_id id))
- else
- ModRef (MPdot (mp,label_of_id id))
-
-*)
-
-(* Search for the head of [qid] in [binders].
- If found, returns the module_path/kernel_name created from the dirpath
- and the basename. Searches Nametab otherwise.
-*)
-
-let lookup_module qid =
- Nametab.locate_module qid
-
-let lookup_modtype qid =
- Nametab.locate_modtype qid
-
-let transl_with_decl env = function
- | Node(loc,"WITHMODULE",[id_ast;qid_ast]) ->
- let id = match id_ast with
- Nvar(_,id) -> id
- | _ -> anomaly "Identifier AST expected"
- in
- let qid = match qid_ast with
- | Node (loc, "QUALID", astl) ->
- interp_qualid astl
- | _ -> anomaly "QUALID expected"
- in
- With_Module (id,lookup_module qid)
- | Node(loc,"WITHDEFINITION",[id_ast;cast]) ->
- let id = match id_ast with
- Nvar(_,id) -> id
- | _ -> anomaly "Identifier AST expected"
- in
- let c = interp_constr Evd.empty env cast in
- With_Definition (id,c)
- | _ -> anomaly "Unexpected AST"
-
-let rec interp_modtype env = function
- | Node(loc,"MODTYPEQID",qid_ast) -> begin match qid_ast with
- | [Node (loc, "QUALID", astl)] ->
- let qid = interp_qualid astl in begin
- try
- MTEident (lookup_modtype qid)
- with
- | Not_found ->
- Modops.error_not_a_modtype (*loc*) (string_of_qualid qid)
- end
- | _ -> anomaly "QUALID expected"
- end
- | Node(loc,"MODTYPEWITH",[mty_ast;decl_ast]) ->
- let mty = interp_modtype env mty_ast in
- let decl = transl_with_decl env decl_ast in
- MTEwith(mty,decl)
- | _ -> anomaly "TODO: transl_modtype: I can handle qualid module types only"
-
-
-let rec interp_modexpr env = function
- | Node(loc,"MODEXPRQID",qid_ast) -> begin match qid_ast with
- | [Node (loc, "QUALID", astl)] ->
- let qid = interp_qualid astl in begin
- try
- MEident (lookup_module qid)
- with
- | Not_found ->
- Modops.error_not_a_module (*loc*) (string_of_qualid qid)
- end
- | _ -> anomaly "QUALID expected"
- end
- | Node(_,"MODEXPRAPP",[ast1;ast2]) ->
- let me1 = interp_modexpr env ast1 in
- let me2 = interp_modexpr env ast2 in
- MEapply(me1,me2)
- | Node(_,"MODEXPRAPP",_) ->
- anomaly "transl_modexpr: MODEXPRAPP must have two arguments"
- | _ -> anomaly "transl_modexpr: I can handle MODEXPRQID or MODEXPRAPP only..."
-
diff --git a/parsing/astmod.mli b/parsing/astmod.mli
deleted file mode 100644
index 49e061a0b..000000000
--- a/parsing/astmod.mli
+++ /dev/null
@@ -1,25 +0,0 @@
-(***********************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
-(* \VV/ *************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(***********************************************************************)
-
-(*i $Id$ i*)
-
-(*i*)
-open Names
-open Declarations
-open Environ
-open Entries
-open Evd
-(*i*)
-
-(* Module expressions and module types are interpreted relatively to
- eventual functor or funsig arguments. *)
-
-val interp_modtype : env -> Coqast.t -> module_type_entry
-
-val interp_modexpr : env -> Coqast.t -> module_expr
-
diff --git a/parsing/astterm.ml b/parsing/astterm.ml
deleted file mode 100644
index bbd9b49e6..000000000
--- a/parsing/astterm.ml
+++ /dev/null
@@ -1,949 +0,0 @@
-(***********************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
-(* \VV/ *************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(***********************************************************************)
-
-(* $Id$ *)
-
-open Pp
-open Util
-open Options
-open Names
-open Nameops
-open Sign
-open Term
-open Termops
-open Environ
-open Evd
-open Reductionops
-open Libnames
-open Impargs
-open Declare
-open Rawterm
-open Pattern
-open Typing
-open Pretyping
-open Evarutil
-open Ast
-open Coqast
-open Nametab
-open Symbols
-
-(*Takes a list of variables which must not be globalized*)
-let from_list l = List.fold_right Idset.add l Idset.empty
-
-(* when an head ident is not a constructor in pattern *)
-let mssg_hd_is_not_constructor s =
- (str "The symbol " ++ pr_id s ++ str " should be a constructor")
-
-(* checking linearity of a list of ids in patterns *)
-let non_linearl_mssg id =
- (str "The variable " ++ str(string_of_id id) ++
- str " is bound several times in pattern")
-
-let error_capture_loc loc s =
- user_err_loc
- (loc,"ast_to_rawconstr",
- (str "The variable " ++ pr_id s ++ str " occurs in its type"))
-
-let error_expl_impl_loc loc =
- user_err_loc
- (loc,"ast_to_rawconstr",
- (str "Found an explicitely given implicit argument but was expecting" ++
- fnl () ++ str "a regular one"))
-
-let error_metavar_loc loc =
- user_err_loc
- (loc,"ast_to_rawconstr",
- (str "Metavariable numbers must be positive"))
-
-let rec has_duplicate = function
- | [] -> None
- | x::l -> if List.mem x l then (Some x) else has_duplicate l
-
-let loc_of_lhs lhs = join_loc (loc (List.hd lhs)) (loc (list_last lhs))
-
-let check_linearity lhs ids =
- match has_duplicate ids with
- | Some id ->
- user_err_loc (loc_of_lhs lhs,"ast_to_eqn",non_linearl_mssg id)
- | None -> ()
-
-let mal_formed_mssg () =
- (str "malformed macro of multiple case")
-
-(* determines if some pattern variable starts with uppercase *)
-let warning_uppercase loc uplid = (* Comment afficher loc ?? *)
- let vars =
- prlist_with_sep
- (fun () -> (str ", ")) (* We avoid spc (), else it breaks the line *)
- (fun v -> (str (string_of_id v))) uplid in
- let (s1,s2) = if List.length uplid = 1 then (" ","s ") else ("s "," ") in
- warn (str ("the variable"^s1) ++ vars ++
- str (" start"^s2^"with an upper case letter in pattern"))
-
-let is_uppercase_var v =
- match (string_of_id v).[0] with
- 'A'..'Z' -> true
- | _ -> false
-
-let check_uppercase loc ids =
-(* A quoi ça sert ? Pour l'extraction vers ML ? Maintenant elle est externe
- let uplid = List.filter is_uppercase_var ids in
- if uplid <> [] then warning_uppercase loc uplid
-*)
- ()
-
-(* check that the number of pattern matches the number of matched args *)
-let mssg_number_of_patterns n pl =
- str"Expecting " ++ int n ++ str" pattern(s) but found " ++
- int (List.length pl) ++ str" in "
-
-let check_number_of_pattern loc n l =
- if n<>(List.length l) then
- user_err_loc (loc,"check_number_of_pattern",mssg_number_of_patterns n l)
-
-(****************************************************************)
-(* Arguments normally implicit in the "Implicit Arguments mode" *)
-(* but explicitely given *)
-
-(* Dump of globalization (to be used by coqdoc) *)
-
-let add_glob loc ref =
-(*i
- let sp = Nametab.sp_of_global (Global.env ()) ref in
- let dir,_ = repr_path sp in
- let rec find_module d =
- try
- let qid = let dir,id = split_dirpath d in make_qualid dir id in
- let _ = Nametab.locate_loaded_library qid in d
- with Not_found -> find_module (dirpath_prefix d)
- in
- let s = string_of_dirpath (find_module dir) in
- i*)
- let sp = Nametab.sp_of_global None ref in
- let id = let _,id = repr_path sp in string_of_id id in
- let dp = string_of_dirpath (Declare.library_part ref) in
- dump_string (Printf.sprintf "R%d %s.%s\n" (fst loc) dp id)
-
-(* Translation of references *)
-
-let ast_to_sp = function
- | Path(loc,sp) ->
- (try
- section_path sp
- with Invalid_argument _ | Failure _ ->
- anomaly_loc(loc,"Astterm.ast_to_sp",
- (str"ill-formed section-path")))
- | ast -> anomaly_loc(Ast.loc ast,"Astterm.ast_to_sp",
- (str"not a section-path"))
-
-let is_underscore id = (id = wildcard)
-
-let name_of_nvar s =
- if is_underscore s then Anonymous else Name s
-
-let ident_of_nvar loc s =
- if is_underscore s then
- user_err_loc (loc,"ident_of_nvar", (str "Unexpected wildcard"))
- else s
-
-let interp_qualid p =
- let outnvar = function
- | Nvar (loc,s) -> s
- | _ -> anomaly "interp_qualid: ill-formed qualified identifier" in
- match p with
- | [] -> anomaly "interp_qualid: empty qualified identifier"
- | l ->
- let p, r = list_chop (List.length l -1) (List.map outnvar l) in
- make_qualid (make_dirpath (List.rev p)) (List.hd r)
-
-let maybe_variable = function
- | [Nvar (_,s)] -> Some s
- | _ -> None
-
-let ids_of_ctxt ctxt =
- Array.to_list
- (Array.map
- (function c -> match kind_of_term c with
- | Var id -> id
- | _ ->
- error
- "Astterm: arbitrary substitution of references not yet implemented")
- ctxt)
-
-type pattern_qualid_kind =
- | ConstrPat of loc * constructor
- | VarPat of loc * identifier
-
-let may_allow_variable loc allow_var l =
- match maybe_variable l with
- | Some s when allow_var ->
- (* Why a warning since there is no warning when writing [globname:T]...
- warning ("Defined reference "^(string_of_qualid qid)
- ^" is here considered as a matching variable");
- *)
- VarPat (loc,s)
- | _ ->
- user_err_loc (loc,"maybe_constructor",
- str "This reference does not denote a constructor: " ++
- str (string_of_qualid (interp_qualid l)))
-
-let maybe_constructor allow_var = function
- | Node(loc,"QUALID",l) ->
- let qid = interp_qualid l in
- (try match extended_locate qid with
- | SyntacticDef sp ->
- (match Syntax_def.search_syntactic_definition loc sp with
- | RRef (_,(ConstructRef c as x)) ->
- if !dump then add_glob loc x;
- ConstrPat (loc,c)
- | _ ->
- user_err_loc (loc,"maybe_constructor",
- str "This syntactic definition should be aliased to a constructor"))
- | TrueGlobal r ->
- let rec unf = function
- | ConstRef cst ->
- (try
- unf
- (reference_of_constr (constant_value (Global.env()) cst))
- with
- NotEvaluableConst _ | Not_found ->
- may_allow_variable loc allow_var l)
- | ConstructRef c ->
- if !dump then add_glob loc r;
- ConstrPat (loc,c)
- | _ -> may_allow_variable loc allow_var l
- in unf r
- with Not_found ->
- match maybe_variable l with
- | Some s when allow_var -> VarPat (loc,s)
- | _ -> error ("Unknown qualified constructor: "
- ^(string_of_qualid qid)))
-
- (* This may happen in quotations *)
- | Node(loc,"MUTCONSTRUCT",[sp;Num(_,ti);Num(_,n)]) ->
- (* Buggy: needs to compute the context *)
- let c = (ast_to_sp sp,ti),n in
- if !dump then add_glob loc (ConstructRef c);
- ConstrPat (loc,c)
-
- | Path(loc,kn) ->
- (let dir,id = decode_kn kn in
- let sp = make_path dir id in
- match absolute_reference sp with
- | ConstructRef c as r ->
- if !dump then add_glob loc (ConstructRef c);
- ConstrPat (loc,c)
- | _ ->
- error ("Unknown absolute constructor name: "^(string_of_path sp)))
- | Node(loc,("CONST"|"SECVAR"|"EVAR"|"MUTIND"|"SYNCONST" as key), l) ->
- user_err_loc (loc,"ast_to_pattern",
- (str "Found a pattern involving global references which are not constructors"
-))
-
- | _ -> anomaly "ast_to_pattern: badly-formed ast for Cases pattern"
-
-let ast_to_global loc = function
- | ("SYNCONST", [sp]) ->
- Syntax_def.search_syntactic_definition loc (ast_to_sp sp), [], []
- | ("EVAR", [(Num (_,ev))]) ->
- REvar (loc, ev), [], []
- | ast ->
- let ref = match ast with
- | ("CONST", [sp]) -> ConstRef (ast_to_sp sp)
- | ("SECVAR", [Nvar (_,s)]) -> VarRef s
- | ("MUTIND", [sp;Num(_,tyi)]) -> IndRef (ast_to_sp sp, tyi)
- | ("MUTCONSTRUCT", [sp;Num(_,ti);Num(_,n)]) ->
- ConstructRef ((ast_to_sp sp,ti),n)
- | _ -> anomaly_loc (loc,"ast_to_global",
- (str "Bad ast for this global a reference"))
- in
- RRef (loc, ref), implicits_of_global ref, find_arguments_scope ref
-
-(*
-let ref_from_constr c = match kind_of_term c with
- | Const (sp,ctxt) -> RConst (sp, ast_to_constr_ctxt ctxt)
- | Evar (ev,ctxt) -> REVar (ev, ast_to_constr_ctxt ctxt)
- | Construct (csp,ctxt) -> RConstruct (csp, ast_to_constr_ctxt ctxt)
- | Ind (isp,ctxt) -> RInd (isp, ast_to_constr_ctxt ctxt)
- | Var id -> RVar id (* utilisé pour coercion_value (tmp) *)
- | _ -> anomaly "Not a reference"
-*)
-
-(* [vars1] is a set of name to avoid (used for the tactic language);
- [vars2] is the set of global variables, env is the set of variables
- abstracted until this point *)
-
-let ast_to_var (env,impls,_) (vars1,vars2) loc id =
- let imps, subscopes =
- if Idset.mem id env or List.mem id vars1
- then
- try List.assoc id impls, []
- with Not_found -> [], []
- else
- let _ = lookup_named id vars2 in
- (* Car Fixpoint met les fns définies tmporairement comme vars de sect *)
- try
- let ref = VarRef id in
- implicits_of_global ref, find_arguments_scope ref
- with _ -> [], []
- in RVar (loc, id), imps, subscopes
-
-(**********************************************************************)
-
-let rawconstr_of_var env vars loc id =
- try
- let (r,_,_) = ast_to_var env vars loc id in r
- with Not_found ->
- Pretype_errors.error_var_not_found_loc loc id
-
-let rawconstr_of_qualid_gen env vars loc qid =
- (* Is it a bound variable? *)
- try
- match repr_qualid qid with
- | d,s when repr_dirpath d = [] -> ast_to_var env vars loc s
- | _ -> raise Not_found
- with Not_found ->
- (* Is it a global reference or a syntactic definition? *)
- try match Nametab.extended_locate qid with
- | TrueGlobal ref ->
- if !dump then add_glob loc ref;
- RRef (loc, ref), implicits_of_global ref, find_arguments_scope ref
- | SyntacticDef sp ->
- Syntax_def.search_syntactic_definition loc sp, [], []
- with Not_found ->
- error_global_not_found_loc loc qid
-
-let rawconstr_of_qualid env vars loc qid =
- let (r,_,_) = rawconstr_of_qualid_gen env vars loc qid in r
-
-let mkLambdaC (x,a,b) = ope("LAMBDA",[a;slam(Some x,b)])
-let mkLambdaCit = List.fold_right (fun (x,a) b -> mkLambdaC(x,a,b))
-let mkProdC (x,a,b) = ope("PROD",[a;slam(Some x,b)])
-let mkProdCit = List.fold_right (fun (x,a) b -> mkProdC(x,a,b))
-
-let destruct_binder = function
- | Node(_,"BINDER",c::idl) -> List.map (fun id -> (nvar_of_ast id,c)) idl
- | _ -> anomaly "BINDER is expected"
-
-let apply_scope_env (ids,impls,scopes as env) = function
- | [] -> env, []
- | (Some sc)::scl -> (ids,impls,sc::scopes), scl
- | None::scl -> env, scl
-
-(* [merge_aliases] returns the sets of all aliases encountered at this
- point and a substitution mapping extra aliases to the first one *)
-let merge_aliases (ids,subst as aliases) = function
- | Anonymous -> aliases
- | Name id ->
- ids@[id],
- if ids=[] then subst
- else (id, List.hd ids)::subst
-
-let alias_of = function
- | ([],_) -> Anonymous
- | (id::_,_) -> Name id
-
-let message_redundant_alias (s1,s2) =
- warning ("Alias variable "^(string_of_id s1)
- ^" is merged with "^(string_of_id s2))
-
-let rec ast_to_pattern scopes aliases = function
- | Node(_,"PATTAS",[Nvar (loc,s); p]) ->
- let aliases' = merge_aliases aliases (name_of_nvar s) in
- ast_to_pattern scopes aliases' p
-
- | Node(_,"PATTCONSTRUCT", head::((_::_) as pl)) ->
- (match maybe_constructor false head with
- | ConstrPat (loc,c) ->
- let (idsl,pl') =
- List.split (List.map (ast_to_pattern scopes ([],[])) pl) in
- (aliases::(List.flatten idsl),
- PatCstr (loc,c,pl',alias_of aliases))
- | VarPat (loc,s) ->
-(*
- user_err_loc (loc,"ast_to_pattern",mssg_hd_is_not_constructor s)
-*)
- assert false)
- | Node(_,"PATTNUMERAL", [Str(loc,n)]) ->
- ([aliases],
- Symbols.interp_numeral_as_pattern loc (Bignat.POS (Bignat.of_string n))
- (alias_of aliases) scopes)
-
- | Node(_,"PATTNEGNUMERAL", [Str(loc,n)]) ->
- ([aliases],
- Symbols.interp_numeral_as_pattern loc (Bignat.NEG (Bignat.of_string n))
- (alias_of aliases) scopes)
-
- | Node(_,"PATTDELIMITERS", [Str(_,sc);e]) ->
- ast_to_pattern (sc::scopes) aliases e
-
- | ast ->
- (match maybe_constructor true ast with
- | ConstrPat (loc,c) -> ([aliases], PatCstr (loc,c,[],alias_of aliases))
- | VarPat (loc,s) ->
- let aliases = merge_aliases aliases (name_of_nvar s) in
- ([aliases], PatVar (loc,alias_of aliases)))
-
-let rec ast_to_fix = function
- | [] -> ([],[],[],[])
- | Node(_,"NUMFDECL", [Nvar(_,fi); Num(_,ni); astA; astT])::rest ->
- let (lf,ln,lA,lt) = ast_to_fix rest in
- (fi::lf, (ni-1)::ln, astA::lA, astT::lt)
- | Node(_,"FDECL", [Nvar(_,fi); Node(_,"BINDERS",bl); astA; astT])::rest->
- let binders = List.flatten (List.map destruct_binder bl) in
- let ni = List.length binders - 1 in
- let (lf,ln,lA,lt) = ast_to_fix rest in
- (fi::lf, ni::ln, (mkProdCit binders astA)::lA,
- (mkLambdaCit binders astT)::lt)
- | _ -> anomaly "FDECL or NUMFDECL is expected"
-
-let rec ast_to_cofix = function
- | [] -> ([],[],[])
- | Node(_,"CFDECL", [Nvar(_,fi); astA; astT])::rest ->
- let (lf,lA,lt) = ast_to_cofix rest in
- (fi::lf, astA::lA, astT::lt)
- | _ -> anomaly "CFDECL is expected"
-
-let error_fixname_unbound s is_cofix loc name =
- user_err_loc
- (loc,"ast_to (COFIX)",
- str "The name" ++ spc () ++ pr_id name ++
- spc () ++ str "is not bound in the corresponding" ++ spc () ++
- str ((if is_cofix then "co" else "")^"fixpoint definition"))
-(*
-let rec collapse_env n env = if n=0 then env else
- add_rel_decl (Anonymous,()) (collapse_env (n-1) (snd (uncons_rel_env env)))
-*)
-
-let check_capture loc s ty = function
- | Slam _ when occur_var_ast s ty -> error_capture_loc loc s
- | _ -> ()
-
-let locate_if_isevar loc id = function
- | RHole _ -> RHole (loc, AbstractionType id)
- | x -> x
-
-let set_hole_implicit i = function
- | RRef (loc,r) -> (loc,ImplicitArg (r,i))
- | RVar (loc,id) -> (loc,ImplicitArg (VarRef id,i))
- | _ -> anomaly "Only refs have implicits"
-
-(*
-let check_only_implicits t imp =
- let rec aux env n t =
- match kind_of_term (whd_betadeltaiota env t) with
- | Prod (x,a,b) -> (aux (push_rel (x,None,a) env) (n+1) b)
- | _ -> n
- in
- let env = Global.env () in
- imp = interval 1 (aux env 0 (get_type_of env Evd.empty t))
-*)
-
-let build_expression loc1 loc2 (ref,impls) args =
- let rec add_args n = function
- | imp::impls,args when is_status_implicit imp ->
- (RHole (set_hole_implicit n (RRef (loc2,ref))))
- ::add_args (n+1) (impls,args)
- | _::impls,a::args -> a::add_args (n+1) (impls,args)
- | [], args -> args
- | _ -> anomalylabstrm "astterm"
- (str "Incorrect signature " ++ pr_global_env None ref ++ str " as an infix") in
- RApp (loc1,RRef (loc2,ref),add_args 1 (impls,args))
-
-let ast_to_rawconstr sigma env allow_soapp lvar =
- let rec dbrec (ids,impls,scopes as env) = function
- | Nvar(loc,s) ->
- rawconstr_of_var env lvar loc s
-
- | Node(loc,"QUALID", l) ->
- let (c,imp,subscopes) =
- rawconstr_of_qualid_gen env lvar loc (interp_qualid l)
- in
- (match ast_to_impargs c env imp subscopes [] with
- [] -> c
- | l -> RApp (loc, c, l))
-
- | Node(loc,"FIX", (Nvar (locid,iddef))::ldecl) ->
- let (lf,ln,lA,lt) = ast_to_fix ldecl in
- let n =
- try
- (list_index (ident_of_nvar locid iddef) lf) -1
- with Not_found ->
- error_fixname_unbound "ast_to_rawconstr (FIX)" false locid iddef in
- let ext_ids = List.fold_right Idset.add lf ids in
- let defl = Array.of_list (List.map (dbrec (ext_ids,impls,scopes)) lt) in
- let arityl = Array.of_list (List.map (dbrec env) lA) in
- RRec (loc,RFix (Array.of_list ln,n), Array.of_list lf, arityl, defl)
-
- | Node(loc,"COFIX", (Nvar(locid,iddef))::ldecl) ->
- let (lf,lA,lt) = ast_to_cofix ldecl in
- let n =
- try
- (list_index (ident_of_nvar locid iddef) lf) -1
- with Not_found ->
- error_fixname_unbound "ast_to_rawconstr (COFIX)" true locid iddef
- in
- let ext_ids = List.fold_right Idset.add lf ids in
- let defl = Array.of_list (List.map (dbrec (ext_ids,impls,scopes)) lt) in
- let arityl = Array.of_list (List.map (dbrec env) lA) in
- RRec (loc,RCoFix n, Array.of_list lf, arityl, defl)
-
- | Node(loc,("PROD"|"LAMBDA"|"LETIN" as k), [c1;Slam(locna,ona,c2)]) ->
- let na,ids' = match ona with
- | Some id -> Name id, Idset.add id ids
- | _ -> Anonymous, ids in
- let c1' = dbrec env c1 and c2' = dbrec (ids',impls,scopes) c2 in
- (match k with
- | "PROD" -> RProd (loc, na, c1', c2')
- | "LAMBDA" -> RLambda (loc, na, locate_if_isevar locna na c1', c2')
- | "LETIN" -> RLetIn (loc, na, c1', c2')
- | _ -> assert false)
-
- | Node(_,("PRODLIST"|"LAMBDALIST" as s), [c1;(Slam _ as c2)]) ->
- iterated_binder s 0 c1 env c2
-
- | Node(loc1,"NOTATION", Str(loc2,ntn)::args) ->
- Symbols.interp_notation ntn scopes (List.map (dbrec env) args)
-
- | Node(_,"NUMERAL", [Str(loc,n)]) ->
- Symbols.interp_numeral loc (Bignat.POS (Bignat.of_string n))
- scopes
-
- | Node(_,"NEGNUMERAL", [Str(loc,n)]) ->
- Symbols.interp_numeral loc (Bignat.NEG (Bignat.of_string n))
- scopes
-
- | Node(_,"DELIMITERS", [Str(_,sc);e]) ->
- dbrec (ids,impls,sc::scopes) e
-
- | Node(loc,"APPLISTEXPL", f::args) ->
- let (f,_,subscopes) = match f with
- | Node(locs,"QUALID",p) ->
- rawconstr_of_qualid_gen env lvar locs (interp_qualid p)
- | _ ->
- (dbrec env f, [], []) in
- RApp (loc,f,ast_to_args env subscopes args)
-
- | Node(loc,"APPLIST", f::args) ->
- let (c, impargs, subscopes) =
- match f with
- | Node(locs,"QUALID",p) ->
- rawconstr_of_qualid_gen env lvar locs (interp_qualid p)
- (* For globalized references (e.g. in Infix) *)
- | Node(loc,
- ("CONST"|"SECVAR"|"EVAR"|"MUTIND"|"MUTCONSTRUCT"|"SYNCONST" as key),
- l) ->
- ast_to_global loc (key,l)
- | _ -> (dbrec env f, [], [])
- in
- RApp (loc, c, ast_to_impargs c env impargs subscopes args)
-
- | Node(loc,"CASES", p:: Node(_,"TOMATCH",tms):: eqns) ->
- let po = match p with
- | Str(_,"SYNTH") -> None
- | _ -> Some(dbrec env p) in
- RCases (loc,PrintCases,po,
- List.map (dbrec env) tms,
- List.map (ast_to_eqn (List.length tms) env) eqns)
-
- | Node(loc,(("CASE"|"IF"|"LET"|"MATCH")as tag), p::c::cl) ->
- let po = match p with
- | Str(_,"SYNTH") -> None
- | _ -> Some(dbrec env p) in
- let isrec = match tag with
- | "MATCH" -> true | ("LET"|"CASE"|"IF") -> false
- | _ -> anomaly "ast_to: wrong tag in old case expression" in
- ROldCase (loc,isrec,po,dbrec env c,
- Array.of_list (List.map (dbrec env) cl))
-
- | Node(loc,"ISEVAR",[]) -> RHole (loc, QuestionMark)
- | Node(loc,"META",[Num(_,n)]) ->
- if n<0 then error_metavar_loc loc else RMeta (loc, n)
-
- | Node(loc,"PROP", []) -> RSort(loc,RProp Null)
- | Node(loc,"SET", []) -> RSort(loc,RProp Pos)
- | Node(loc,"TYPE", _) -> RSort(loc,RType None)
-
- (* This case mainly parses things build in a quotation *)
- | Node(loc,("CONST"|"SECVAR"|"EVAR"|"MUTIND"|"MUTCONSTRUCT"|"SYNCONST" as key),l) ->
- let (r,_,_) = ast_to_global loc (key,l) in r
-
- | Node(loc,"CAST", [c1;c2]) ->
- RCast (loc,dbrec env c1,dbrec env c2)
-
- | Node(loc,"SOAPP", args) when allow_soapp ->
- (match List.map (dbrec env) args with
- (* Hack special pour l'interprétation des constr_pattern *)
- | RMeta (locn,n) :: args -> RApp (loc,RMeta (locn,- n), args)
- | RHole _ :: _ -> anomaly "Metavariable for 2nd-order pattern-matching cannot be anonymous"
- | _ -> anomaly "Bad arguments for second-order pattern-matching")
-
- | Node(loc,"SQUASH",_) ->
- user_err_loc(loc,"ast_to_rawconstr",
- (str "Ill-formed specification"))
-
- | Node(loc,opn,tl) ->
- anomaly ("ast_to_rawconstr found operator "^opn^" with "^
- (string_of_int (List.length tl))^" arguments")
-
- | Dynamic (loc,d) -> RDynamic (loc,d)
-
- | _ -> anomaly "ast_to_rawconstr: unexpected ast"
-
- and ast_to_eqn n (ids,impls,scopes as env) = function
- | Node(loc,"EQN",rhs::lhs) ->
- let (idsl_substl_list,pl) =
- List.split (List.map (ast_to_pattern scopes ([],[])) lhs) in
- let idsl, substl = List.split (List.flatten idsl_substl_list) in
- let eqn_ids = List.flatten idsl in
- let subst = List.flatten substl in
- (* Linearity implies the order in ids is irrelevant *)
- check_linearity lhs eqn_ids;
- check_uppercase loc eqn_ids;
- check_number_of_pattern loc n pl;
- let rhs = replace_vars_ast subst rhs in
- List.iter message_redundant_alias subst;
- let env_ids = List.fold_right Idset.add eqn_ids ids in
- (loc, eqn_ids,pl,dbrec (env_ids,impls,scopes) rhs)
- | _ -> anomaly "ast_to_rawconstr: ill-formed ast for Cases equation"
-
- and iterated_binder oper n ty (ids,impls,scopes as env) = function
- | Slam(loc,ona,body) ->
- let na,ids' = match ona with
- | Some id ->
- if n>0 then check_capture loc id ty body;
- Name id, Idset.add id ids
- | _ -> Anonymous, ids
- in
- let r = iterated_binder oper (n+1) ty (ids',impls,scopes) body in
- (match oper with
- | "PRODLIST" -> RProd(loc, na, dbrec env ty, r)
- | "LAMBDALIST" ->
- RLambda(loc, na, locate_if_isevar loc na (dbrec env ty), r)
- | _ -> assert false)
- | body -> dbrec env body
-
- and ast_to_impargs c env l subscopes args =
- let rec aux n l subscopes args =
- let (enva,subscopes') = apply_scope_env env subscopes in
- match (l,args) with
- | (imp::l',Node(loc, "EXPL", [Num(_,j);a])::args') ->
- if is_status_implicit imp & j>=n then
- if j=n then
- (dbrec enva a)::(aux (n+1) l' subscopes' args')
- else
- (RHole (set_hole_implicit n c))::(aux (n+1) l' subscopes' args)
- else
- if not (is_status_implicit imp) then
- error ("Bad explicitation number: found "^
- (string_of_int j)^" but was expecting a regular argument")
- else
- error ("Bad explicitation number: found "^
- (string_of_int j)^" but was expecting "^(string_of_int n))
- | (imp::l',a::args') ->
- if is_status_implicit imp then
- (RHole (set_hole_implicit n c))::(aux (n+1) l' subscopes' args)
- else
- (dbrec enva a)::(aux (n+1) l' subscopes' args')
- | ([],args) -> ast_to_args env subscopes args
- | (_::l',[]) ->
- if List.for_all is_status_implicit l then
- (RHole (set_hole_implicit n c))::(aux (n+1) l' subscopes args)
- else []
- in
- aux 1 l subscopes args
-
- and ast_to_args env subscopes = function
- | Node(loc, "EXPL", _)::args' ->
- (* To deal with errors *)
- error_expl_impl_loc loc
- | a::args ->
- let enva, subscopes = apply_scope_env env subscopes in
- (dbrec enva a) :: (ast_to_args env subscopes args)
- | [] -> []
-
- and interp_binding env = function
- | Node(_,"BINDING", [Num(_,n);Node(loc,"CONSTR",[c])]) ->
- (AnonHyp n,dbrec env c)
- | Node(_,"BINDING", [Nvar(loc0,s); Node(loc1,"CONSTR",[c])]) ->
- (NamedHyp (ident_of_nvar loc0 s), dbrec env c)
- | x ->
- errorlabstrm "bind_interp"
- (str "Not the expected form in binding" ++ print_ast x)
-
- in
- dbrec env
-
-(**************************************************************************)
-(* Globalization of AST quotations (mainly used to get statically *)
-(* bound idents in grammar or pretty-printing rules) *)
-(**************************************************************************)
-
-let ast_of_ref_loc loc ref = set_loc loc (Termast.ast_of_ref ref)
-
-let ast_of_syndef loc sp = Node (loc, "SYNCONST", [path_section loc sp])
-
-let ast_of_extended_ref_loc loc = function
- | TrueGlobal ref -> ast_of_ref_loc loc ref
- | SyntacticDef kn -> ast_of_syndef loc kn
-
-let ast_of_extended_ref = ast_of_extended_ref_loc dummy_loc
-
-let ast_of_var env ast id =
- if isMeta (string_of_id id) or Idset.mem id env then ast
- else raise Not_found
-
-let ast_hole = Node (dummy_loc, "ISEVAR", [])
-
-let implicits_of_extended_reference = function
- | TrueGlobal ref -> implicits_of_global ref
- | SyntacticDef _ -> []
-
-let warning_globalize qid =
- warning ("Could not globalize " ^ (string_of_qualid qid))
-
-let globalize_qualid (loc,qid) =
- try
- let ref = Nametab.extended_locate qid in
- ast_of_extended_ref_loc loc ref
- with Not_found ->
- if_verbose warning_globalize qid;
- Termast.ast_of_qualid qid
-
-let adjust_qualid env loc ast qid =
- (* Is it a bound variable? *)
- try
- match repr_qualid qid with
- | d,id when repr_dirpath d = [] -> ast_of_var env ast id
- | _ -> raise Not_found
- with Not_found ->
- (* Is it a global reference or a syntactic definition? *)
- try
- let ref = Nametab.extended_locate qid in
- ast_of_extended_ref_loc loc ref
- with Not_found ->
- if_verbose warning_globalize qid;
- ast
-
-let ast_adjust_consts sigma =
- let rec dbrec env = function
- | Node(loc, ("APPLIST" as key), (Node(locs,"QUALID",p) as ast)::args) ->
- let f = adjust_qualid env loc ast (interp_qualid p) in
- Node(loc, key, f :: List.map (dbrec env) args)
- | Nmeta (loc, s) as ast -> ast
- | Nvar (loc, id) as ast ->
- if Idset.mem id env then ast
- else adjust_qualid env loc ast (make_short_qualid id)
- | Node (loc, "QUALID", p) as ast ->
- adjust_qualid env loc ast (interp_qualid p)
- | Slam (loc, None, t) -> Slam (loc, None, dbrec env t)
- | Slam (loc, Some na, t) ->
- let env' = Idset.add na env in
- Slam (loc, Some na, dbrec env' t)
- | Node (loc, opn, tl) -> Node (loc, opn, List.map (dbrec env) tl)
- | x -> x
-
- in
- dbrec
-
-let globalize_constr ast =
- let sign = Global.named_context () in
- ast_adjust_consts Evd.empty (from_list (ids_of_named_context sign)) ast
-
-(* Globalizes ast expressing constructions in tactics or vernac *)
-(* Actually, it is incomplete, see vernacinterp.ml and tacinterp.ml *)
-(* Used mainly to parse Grammar and Syntax expressions *)
-let rec glob_ast sigma env =
- function
- Node (loc, "CONSTR", [c]) ->
- Node (loc, "CONSTR", [ast_adjust_consts sigma env c])
- | Node (loc, "CONSTRLIST", l) ->
- Node (loc, "CONSTRLIST", List.map (ast_adjust_consts sigma env) l)
- | Slam (loc, None, t) -> Slam (loc, None, glob_ast sigma env t)
- | Slam (loc, Some na, t) ->
- let env' = Idset.add na env in
- Slam (loc, Some na, glob_ast sigma env' t)
- | Node (loc, opn, tl) -> Node (loc, opn, List.map (glob_ast sigma env) tl)
- | x -> x
-
-let globalize_ast ast =
- let sign = Global.named_context () in
- glob_ast Evd.empty (from_list (ids_of_named_context sign)) ast
-
-(**************************************************************************)
-(* Functions to translate ast into rawconstr *)
-(**************************************************************************)
-
-let interp_rawconstr_gen sigma env impls allow_soapp lvar com =
- ast_to_rawconstr sigma
- (from_list (ids_of_rel_context (rel_context env)), impls, Symbols.current_scopes ())
- allow_soapp (lvar,env) com
-
-let interp_rawconstr sigma env com =
- interp_rawconstr_gen sigma env [] false [] com
-
-let interp_rawconstr_with_implicits sigma env impls com =
- interp_rawconstr_gen sigma env impls false [] com
-
-(*The same as interp_rawconstr but with a list of variables which must not be
- globalized*)
-
-let interp_rawconstr_wo_glob sigma env lvar com =
- interp_rawconstr_gen sigma env [] false lvar com
-
-(*********************************************************************)
-(* V6 compat: Functions before in ex-trad *)
-
-(* Functions to parse and interpret constructions *)
-
-(* To embed constr in Coqast.t *)
-let constrIn t = Dynamic (dummy_loc,constr_in t)
-let constrOut = function
- | Dynamic (_,d) ->
- if (Dyn.tag d) = "constr" then
- constr_out d
- else
- anomalylabstrm "constrOut" (str "Dynamic tag should be constr")
- | ast ->
- anomalylabstrm "constrOut"
- (str "Not a Dynamic ast: " ++ print_ast ast)
-
-let interp_global_constr env (loc,qid) =
- let c =
- rawconstr_of_qualid (Idset.empty,[],current_scopes()) ([],env) loc qid
- in
- understand Evd.empty env c
-
-let interp_constr sigma env c =
- understand sigma env (interp_rawconstr sigma env c)
-
-let interp_openconstr sigma env c =
- understand_gen_tcc sigma env [] [] None (interp_rawconstr sigma env c)
-
-let interp_casted_openconstr sigma env c typ =
- understand_gen_tcc sigma env [] [] (Some typ) (interp_rawconstr sigma env c)
-
-let interp_type sigma env c =
- understand_type sigma env (interp_rawconstr sigma env c)
-
-let interp_type_with_implicits sigma env impls c =
- understand_type sigma env (interp_rawconstr_with_implicits sigma env impls c)
-
-let interp_sort = function
- | Node(loc,"PROP", []) -> Prop Null
- | Node(loc,"SET", []) -> Prop Pos
- | Node(loc,"TYPE", _) -> new_Type_sort ()
- | a -> user_err_loc (Ast.loc a,"interp_sort", (str "Not a sort"))
-
-let interp_elimination_sort = function
- | Node(loc,"PROP", []) -> InProp
- | Node(loc,"SET", []) -> InSet
- | Node(loc,"TYPE", _) -> InType
- | a -> user_err_loc (Ast.loc a,"interp_sort", (str "Not a sort"))
-
-let judgment_of_rawconstr sigma env c =
- understand_judgment sigma env (interp_rawconstr sigma env c)
-
-let type_judgment_of_rawconstr sigma env c =
- understand_type_judgment sigma env (interp_rawconstr sigma env c)
-
-(*To retype a list of key*constr with undefined key*)
-let retype_list sigma env lst =
- List.fold_right (fun (x,csr) a ->
- try (x,Retyping.get_judgment_of env sigma csr)::a with
- | Anomaly _ -> a) lst []
-
-(* List.map (fun (x,csr) -> (x,Retyping.get_judgment_of env sigma csr)) lst*)
-
-(* Interprets a constr according to two lists *)
-(* of instantiations (variables and metas) *)
-(* Note: typ is retyped *)
-let interp_constr_gen sigma env lvar lmeta com exptyp =
- let c = interp_rawconstr_gen sigma env [] false (List.map fst lvar) com
- and rtype lst = retype_list sigma env lst in
- understand_gen sigma env (rtype lvar) (rtype lmeta) exptyp c;;
-
-(*Interprets a casted constr according to two lists of instantiations
- (variables and metas)*)
-let interp_openconstr_gen sigma env lvar lmeta com exptyp =
- let c = interp_rawconstr_gen sigma env [] false (List.map fst lvar) com
- and rtype lst = retype_list sigma env lst in
- understand_gen_tcc sigma env (rtype lvar) (rtype lmeta) exptyp c;;
-
-let interp_casted_constr sigma env com typ =
- understand_gen sigma env [] [] (Some typ) (interp_rawconstr sigma env com)
-
-(* To process patterns, we need a translation from AST to term
- without typing at all. *)
-
-let ctxt_of_ids ids = Array.of_list (List.map mkVar ids)
-(*
-let rec pat_of_ref metas vars = function
- | RConst (sp,ctxt) -> RConst (sp, ast_to_rawconstr_ctxt ctxt)
- | RInd (ip,ctxt) -> RInd (ip, ast_to_rawconstr_ctxt ctxt)
- | RConstruct(cp,ctxt) ->RConstruct(cp, ast_to_rawconstr_ctxt ctxt)
- | REVar (n,ctxt) -> REVar (n, ast_to_rawconstr_ctxt ctxt)
- | RVar _ -> assert false (* Capturé dans pattern_of_raw *)
-*)
-let rec pat_of_raw metas vars lvar = function
- | RVar (_,id) ->
- (try PRel (list_index (Name id) vars)
- with Not_found ->
- try List.assoc id lvar
- with Not_found -> PVar id)
- | RMeta (_,n) ->
- metas := n::!metas; PMeta (Some n)
- | RRef (_,r) ->
- PRef r
- (* Hack pour ne pas réécrire une interprétation complète des patterns*)
- | RApp (_, RMeta (_,n), cl) when n<0 ->
- PSoApp (- n, List.map (pat_of_raw metas vars lvar) cl)
- | RApp (_,c,cl) ->
- PApp (pat_of_raw metas vars lvar c,
- Array.of_list (List.map (pat_of_raw metas vars lvar) cl))
- | RLambda (_,na,c1,c2) ->
- PLambda (na, pat_of_raw metas vars lvar c1,
- pat_of_raw metas (na::vars) lvar c2)
- | RProd (_,na,c1,c2) ->
- PProd (na, pat_of_raw metas vars lvar c1,
- pat_of_raw metas (na::vars) lvar c2)
- | RLetIn (_,na,c1,c2) ->
- PLetIn (na, pat_of_raw metas vars lvar c1,
- pat_of_raw metas (na::vars) lvar c2)
- | RSort (_,s) ->
- PSort s
- | RHole _ ->
- PMeta None
- | RCast (_,c,t) ->
- warning "Cast not taken into account in constr pattern";
- pat_of_raw metas vars lvar c
- | ROldCase (_,false,po,c,br) ->
- PCase (option_app (pat_of_raw metas vars lvar) po,
- pat_of_raw metas vars lvar c,
- Array.map (pat_of_raw metas vars lvar) br)
- | _ ->
- error "pattern_of_rawconstr: not implemented"
-
-let pattern_of_rawconstr lvar c =
- let metas = ref [] in
- let p = pat_of_raw metas [] lvar c in
- (!metas,p)
-
-let interp_constrpattern_gen sigma env lvar com =
- let c =
- ast_to_rawconstr sigma
- (from_list (ids_of_rel_context (rel_context env)), [], Symbols.current_scopes ())
- true (List.map fst lvar,env) com
- and nlvar = List.map (fun (id,c) -> (id,pattern_of_constr c)) lvar in
- try
- pattern_of_rawconstr nlvar c
- with e ->
- Stdpp.raise_with_loc (Ast.loc com) e
-
-let interp_constrpattern sigma env com =
- interp_constrpattern_gen sigma env [] com
diff --git a/parsing/astterm.mli b/parsing/astterm.mli
deleted file mode 100644
index 3a871cd53..000000000
--- a/parsing/astterm.mli
+++ /dev/null
@@ -1,101 +0,0 @@
-(***********************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
-(* \VV/ *************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(***********************************************************************)
-
-(*i $Id$ i*)
-
-(*i*)
-open Names
-open Term
-open Sign
-open Evd
-open Environ
-open Libnames
-open Rawterm
-open Pattern
-(*i*)
-
-(* Translation from AST to terms. *)
-
-(* To embed constr in Coqast.t *)
-val constrIn : constr -> Coqast.t
-val constrOut : Coqast.t -> constr
-
-(* Interprets global names, including syntactic defs and section variables *)
-val interp_global_constr : env -> qualid Util.located -> constr
-
-val interp_rawconstr : evar_map -> env -> Coqast.t -> rawconstr
-val interp_rawconstr_gen :
- evar_map -> env -> (identifier * Impargs.implicits_list) list ->
- bool -> identifier list -> Coqast.t -> rawconstr
-val interp_constr : evar_map -> env -> Coqast.t -> constr
-val interp_casted_constr : evar_map -> env -> Coqast.t -> types -> constr
-val interp_type : evar_map -> env -> Coqast.t -> types
-val interp_sort : Coqast.t -> sorts
-
-val interp_elimination_sort : Coqast.t -> sorts_family
-
-val interp_openconstr :
- evar_map -> env -> Coqast.t -> evar_map * constr
-val interp_casted_openconstr :
- evar_map -> env -> Coqast.t -> constr -> evar_map * constr
-
-(* [interp_type_with_implicits] extends [interp_type] by allowing
- implicits arguments in the ``rel'' part of [env]; the extra
- argument associates a list of implicit positions to identifiers
- declared in the rel_context of [env] *)
-val interp_type_with_implicits :
- evar_map -> env ->
- (identifier * Impargs.implicits_list) list -> Coqast.t -> types
-
-val judgment_of_rawconstr : evar_map -> env -> Coqast.t -> unsafe_judgment
-val type_judgment_of_rawconstr :
- evar_map -> env -> Coqast.t -> unsafe_type_judgment
-
-(*Interprets a constr according to two lists of instantiations (variables and
- metas), possibly casting it*)
-val interp_constr_gen :
- evar_map -> env -> (identifier * constr) list ->
- (int * constr) list -> Coqast.t -> constr option -> constr
-
-(*Interprets a constr according to two lists of instantiations (variables and
- metas), possibly casting it, and turning unresolved evar into metas*)
-val interp_openconstr_gen :
- evar_map -> env -> (identifier * constr) list ->
- (int * constr) list -> Coqast.t -> constr option
- -> evar_map * constr
-
-(*Interprets constr patterns according to a list of instantiations
- (variables)*)
-val interp_constrpattern_gen :
- evar_map -> env -> (identifier * constr) list -> Coqast.t ->
- int list * constr_pattern
-
-val interp_constrpattern :
- evar_map -> env -> Coqast.t -> int list * constr_pattern
-
-(*s Globalization of AST quotations (mainly used to get statically
- bound idents in grammar or pretty-printing rules) *)
-val globalize_constr : Coqast.t -> Coqast.t
-val globalize_ast : Coqast.t -> Coqast.t
-val globalize_qualid : qualid Util.located -> Coqast.t
-
-val ast_of_extended_ref_loc : loc -> Libnames.extended_global_reference -> Coqast.t
-
-(* This transforms args of a qualid keyword into a qualified ident *)
-(* it does no relocation *)
-val interp_qualid : Coqast.t list -> qualid
-
-(*i Translation rules from V6 to V7:
-
-constr_of_com_casted -> interp_casted_constr
-constr_of_com_sort -> interp_type
-constr_of_com -> interp_constr
-rawconstr_of_com -> interp_rawconstr [+ env instead of sign]
-type_of_com -> types_of_com Evd.empty
-constr_of_com1 true -> interp_type
-i*)
diff --git a/parsing/coqast.ml b/parsing/coqast.ml
index c0ecc618b..65519b673 100644
--- a/parsing/coqast.ml
+++ b/parsing/coqast.ml
@@ -9,12 +9,11 @@
(* $Id$ *)
(*i*)
+open Util
open Names
open Libnames
(*i*)
-type loc = int * int
-
type t =
| Node of loc * string * t list
| Nmeta of loc * string
@@ -122,86 +121,3 @@ let rec subst_ast subst ast = match ast with
| Str _
| Id _
| Dynamic _ -> ast
-
-open Util
-open Rawterm
-open Term
-
-type scope_name = string
-
-type reference_expr =
- | RQualid of qualid located
- | RIdent of identifier located
-
-type explicitation = int
-
-type cases_pattern =
- | CPatAlias of loc * cases_pattern * identifier
- | CPatCstr of loc * reference_expr * cases_pattern list
- | CPatAtom of loc * reference_expr option
- | CPatNumeral of loc * Bignat.bigint
- | CPatDelimiters of loc * scope_name * cases_pattern
-
-type ordered_case_style = CIf | CLet | CMatch | CCase
-
-type constr_ast =
- | CRef of reference_expr
- | CFix of loc * identifier located * fixpoint_expr list
- | CCoFix of loc * identifier located * cofixpoint_expr list
- | CArrow of loc * constr_ast * constr_ast
- | CProdN of loc * (name located list * constr_ast) list * constr_ast
- | CLambdaN of loc * (name located list * constr_ast) list * constr_ast
- | CLetIn of loc * identifier located * constr_ast * constr_ast
- | CAppExpl of loc * reference_expr * constr_ast list
- | CApp of loc * constr_ast * (constr_ast * explicitation option) list
- | CCases of loc * case_style * constr_ast option * constr_ast list *
- (loc * cases_pattern list * constr_ast) list
- | COrderedCase of loc * ordered_case_style * constr_ast option * constr_ast * constr_ast list
- | CHole of loc
- | CMeta of loc * int
- | CSort of loc * rawsort
- | CCast of loc * constr_ast * constr_ast
- | CNotation of loc * string * constr_ast list
- | CNumeral of loc * Bignat.bigint
- | CDelimiters of loc * scope_name * constr_ast
- | CDynamic of loc * Dyn.t
-
-and local_binder = name located list * constr_ast
-
-and fixpoint_expr = identifier * local_binder list * constr_ast * constr_ast
-
-and cofixpoint_expr = identifier * constr_ast * constr_ast
-
-let constr_loc = function
- | CRef (RIdent (loc,_)) -> loc
- | CRef (RQualid (loc,_)) -> loc
- | CFix (loc,_,_) -> loc
- | CCoFix (loc,_,_) -> loc
- | CArrow (loc,_,_) -> loc
- | CProdN (loc,_,_) -> loc
- | CLambdaN (loc,_,_) -> loc
- | CLetIn (loc,_,_,_) -> loc
- | CAppExpl (loc,_,_) -> loc
- | CApp (loc,_,_) -> loc
- | CCases (loc,_,_,_,_) -> loc
- | COrderedCase (loc,_,_,_,_) -> loc
- | CHole loc -> loc
- | CMeta (loc,_) -> loc
- | CSort (loc,_) -> loc
- | CCast (loc,_,_) -> loc
- | CNotation (loc,_,_) -> loc
- | CNumeral (loc,_) -> loc
- | CDelimiters (loc,_,_) -> loc
- | CDynamic _ -> dummy_loc
-
-let cases_pattern_loc = function
- | CPatAlias (loc,_,_) -> loc
- | CPatCstr (loc,_,_) -> loc
- | CPatAtom (loc,_) -> loc
- | CPatNumeral (loc,_) -> loc
- | CPatDelimiters (loc,_,_) -> loc
-
-let replace_vars_constr_ast l t =
- if l = [] then t else failwith "replace_constr_ast: TODO"
-
-let occur_var_constr_ast id t = Pp.warning "occur_var_constr_ast:TODO"; true
diff --git a/parsing/coqast.mli b/parsing/coqast.mli
index 52b19c6bc..5b8c9d7d7 100644
--- a/parsing/coqast.mli
+++ b/parsing/coqast.mli
@@ -9,14 +9,13 @@
(*i $Id$ i*)
(*i*)
+open Util
open Names
open Libnames
(*i*)
(* Abstract syntax trees. *)
-type loc = int * int
-
type t =
| Node of loc * string * t list
| Nmeta of loc * string
@@ -50,62 +49,3 @@ val fold_tactic_expr :
('a -> t -> 'a) -> ('a -> tactic_expr -> 'a) -> 'a -> tactic_expr -> 'a
val iter_tactic_expr : (tactic_expr -> unit) -> tactic_expr -> unit
*)
-
-
-open Util
-open Rawterm
-open Term
-
-type scope_name = string
-
-type reference_expr =
- | RQualid of qualid located
- | RIdent of identifier located
-
-type explicitation = int
-
-type cases_pattern =
- | CPatAlias of loc * cases_pattern * identifier
- | CPatCstr of loc * reference_expr * cases_pattern list
- | CPatAtom of loc * reference_expr option
- | CPatNumeral of loc * Bignat.bigint
- | CPatDelimiters of loc * scope_name * cases_pattern
-
-type ordered_case_style = CIf | CLet | CMatch | CCase
-
-type constr_ast =
- | CRef of reference_expr
- | CFix of loc * identifier located * fixpoint_expr list
- | CCoFix of loc * identifier located * cofixpoint_expr list
- | CArrow of loc * constr_ast * constr_ast
- | CProdN of loc * (name located list * constr_ast) list * constr_ast
- | CLambdaN of loc * (name located list * constr_ast) list * constr_ast
- | CLetIn of loc * identifier located * constr_ast * constr_ast
- | CAppExpl of loc * reference_expr * constr_ast list
- | CApp of loc * constr_ast * (constr_ast * explicitation option) list
- | CCases of loc * case_style * constr_ast option * constr_ast list *
- (loc * cases_pattern list * constr_ast) list
- | COrderedCase of loc * ordered_case_style * constr_ast option * constr_ast * constr_ast list
- | CHole of loc
- | CMeta of loc * int
- | CSort of loc * rawsort
- | CCast of loc * constr_ast * constr_ast
- | CNotation of loc * string * constr_ast list
- | CNumeral of loc * Bignat.bigint
- | CDelimiters of loc * scope_name * constr_ast
- | CDynamic of loc * Dyn.t
-
-and local_binder = name located list * constr_ast
-
-and fixpoint_expr = identifier * local_binder list * constr_ast * constr_ast
-
-and cofixpoint_expr = identifier * constr_ast * constr_ast
-
-val constr_loc : constr_ast -> loc
-
-val cases_pattern_loc : cases_pattern -> loc
-
-val replace_vars_constr_ast :
- (identifier * identifier) list -> constr_ast -> constr_ast
-
-val occur_var_constr_ast : identifier -> constr_ast -> bool
diff --git a/parsing/coqlib.ml b/parsing/coqlib.ml
deleted file mode 100644
index 5c0fef4aa..000000000
--- a/parsing/coqlib.ml
+++ /dev/null
@@ -1,285 +0,0 @@
-(***********************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
-(* \VV/ *************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(***********************************************************************)
-
-(* $Id$ *)
-
-open Util
-open Names
-open Term
-open Libnames
-open Declare
-open Pattern
-open Nametab
-
-let make_dir l = make_dirpath (List.map id_of_string (List.rev l))
-let coq_id = id_of_string "Coq"
-let init_id = id_of_string "Init"
-let arith_id = id_of_string "Arith"
-let datatypes_id = id_of_string "Datatypes"
-
-let logic_module = make_dir ["Coq";"Init";"Logic"]
-let logic_type_module = make_dir ["Coq";"Init";"Logic_Type"]
-let datatypes_module = make_dir ["Coq";"Init";"Datatypes"]
-let arith_module = make_dir ["Coq";"Arith";"Arith"]
-
-(* TODO: temporary hack *)
-let make_path dir id = Libnames.encode_kn dir id
-
-let nat_path = make_path datatypes_module (id_of_string "nat")
-
-let glob_nat = IndRef (nat_path,0)
-
-let path_of_O = ((nat_path,0),1)
-let path_of_S = ((nat_path,0),2)
-let glob_O = ConstructRef path_of_O
-let glob_S = ConstructRef path_of_S
-
-let eq_path = make_path logic_module (id_of_string "eq")
-let eqT_path = make_path logic_type_module (id_of_string "eqT")
-
-let glob_eq = IndRef (eq_path,0)
-let glob_eqT = IndRef (eqT_path,0)
-
-let reference dir s =
- let dir = make_dir ("Coq"::"Init"::[dir]) in
- let id = id_of_string s in
- try
- Nametab.absolute_reference (Libnames.make_path dir id)
- with Not_found ->
- anomaly ("Coqlib: cannot find "^(string_of_qualid (make_qualid dir id)))
-
-let constant dir s = Declare.constr_of_reference (reference dir s)
-
-type coq_sigma_data = {
- proj1 : constr;
- proj2 : constr;
- elim : constr;
- intro : constr;
- typ : constr }
-
-type 'a delayed = unit -> 'a
-
-let build_sigma_set () =
- { proj1 = constant "Specif" "projS1";
- proj2 = constant "Specif" "projS2";
- elim = constant "Specif" "sigS_rec";
- intro = constant "Specif" "existS";
- typ = constant "Specif" "sigS" }
-
-let build_sigma_type () =
- { proj1 = constant "Specif" "projT1";
- proj2 = constant "Specif" "projT2";
- elim = constant "Specif" "sigT_rec";
- intro = constant "Specif" "existT";
- typ = constant "Specif" "sigT" }
-
-(* Equalities *)
-type coq_leibniz_eq_data = {
- eq : constr delayed;
- ind : constr delayed;
- rrec : constr delayed option;
- rect : constr delayed option;
- congr: constr delayed;
- sym : constr delayed }
-
-let constant dir id = lazy (constant dir id)
-
-(* Equality on Set *)
-let coq_eq_eq = constant "Logic" "eq"
-let coq_eq_ind = constant "Logic" "eq_ind"
-let coq_eq_rec = constant "Logic" "eq_rec"
-let coq_eq_rect = constant "Logic" "eq_rect"
-let coq_eq_congr = constant "Logic" "f_equal"
-let coq_eq_sym = constant "Logic" "sym_eq"
-let coq_f_equal2 = constant "Logic" "f_equal2"
-
-let build_coq_eq_data = {
- eq = (fun () -> Lazy.force coq_eq_eq);
- ind = (fun () -> Lazy.force coq_eq_ind);
- rrec = Some (fun () -> Lazy.force coq_eq_rec);
- rect = Some (fun () -> Lazy.force coq_eq_rect);
- congr = (fun () -> Lazy.force coq_eq_congr);
- sym = (fun () -> Lazy.force coq_eq_sym) }
-
-let build_coq_eq = build_coq_eq_data.eq
-let build_coq_f_equal2 () = Lazy.force coq_f_equal2
-
-(* Specif *)
-let coq_sumbool = constant "Specif" "sumbool"
-
-let build_coq_sumbool () = Lazy.force coq_sumbool
-
-(* Equality on Type *)
-let coq_eqT_eq = constant "Logic_Type" "eqT"
-let coq_eqT_ind = constant "Logic_Type" "eqT_ind"
-let coq_eqT_congr =constant "Logic_Type" "congr_eqT"
-let coq_eqT_sym = constant "Logic_Type" "sym_eqT"
-
-let build_coq_eqT_data = {
- eq = (fun () -> Lazy.force coq_eqT_eq);
- ind = (fun () -> Lazy.force coq_eqT_ind);
- rrec = None;
- rect = None;
- congr = (fun () -> Lazy.force coq_eqT_congr);
- sym = (fun () -> Lazy.force coq_eqT_sym) }
-
-let build_coq_eqT = build_coq_eqT_data.eq
-let build_coq_sym_eqT = build_coq_eqT_data.sym
-
-(* Equality on Type as a Type *)
-let coq_idT_eq = constant "Logic_Type" "identityT"
-let coq_idT_ind = constant "Logic_Type" "identityT_ind"
-let coq_idT_rec = constant "Logic_Type" "identityT_rec"
-let coq_idT_rect = constant "Logic_Type" "identityT_rect"
-let coq_idT_congr = constant "Logic_Type" "congr_idT"
-let coq_idT_sym = constant "Logic_Type" "sym_idT"
-
-let build_coq_idT_data = {
- eq = (fun () -> Lazy.force coq_idT_eq);
- ind = (fun () -> Lazy.force coq_idT_ind);
- rrec = Some (fun () -> Lazy.force coq_idT_rec);
- rect = Some (fun () -> Lazy.force coq_idT_rect);
- congr = (fun () -> Lazy.force coq_idT_congr);
- sym = (fun () -> Lazy.force coq_idT_sym) }
-
-(* Empty Type *)
-let coq_EmptyT = constant "Logic_Type" "EmptyT"
-
-(* Unit Type and its unique inhabitant *)
-let coq_UnitT = constant "Logic_Type" "UnitT"
-let coq_IT = constant "Logic_Type" "IT"
-
-(* The False proposition *)
-let coq_False = constant "Logic" "False"
-
-(* The True proposition and its unique proof *)
-let coq_True = constant "Logic" "True"
-let coq_I = constant "Logic" "I"
-
-(* Connectives *)
-let coq_not = constant "Logic" "not"
-let coq_and = constant "Logic" "and"
-let coq_or = constant "Logic" "or"
-let coq_ex = constant "Logic" "ex"
-
-(* Runtime part *)
-let build_coq_EmptyT () = Lazy.force coq_EmptyT
-let build_coq_UnitT () = Lazy.force coq_UnitT
-let build_coq_IT () = Lazy.force coq_IT
-
-let build_coq_True () = Lazy.force coq_True
-let build_coq_I () = Lazy.force coq_I
-
-let build_coq_False () = Lazy.force coq_False
-let build_coq_not () = Lazy.force coq_not
-let build_coq_and () = Lazy.force coq_and
-let build_coq_or () = Lazy.force coq_or
-let build_coq_ex () = Lazy.force coq_ex
-
-(****************************************************************************)
-(* Patterns *)
-(* This needs to have interp_constrpattern available ...
-let parse_astconstr s =
- try
- Pcoq.parse_string Pcoq.Constr.constr_eoi s
- with Stdpp.Exc_located (_ , (Stream.Failure | Stream.Error _)) ->
- error "Syntax error : not a construction"
-
-let parse_pattern s =
- Astterm.interp_constrpattern Evd.empty (Global.env()) (parse_astconstr s)
-
-let coq_eq_pattern =
- lazy (snd (parse_pattern "(Coq.Init.Logic.eq ?1 ?2 ?3)"))
-let coq_eqT_pattern =
- lazy (snd (parse_pattern "(Coq.Init.Logic_Type.eqT ?1 ?2 ?3)"))
-let coq_idT_pattern =
- lazy (snd (parse_pattern "(Coq.Init.Logic_Type.identityT ?1 ?2 ?3)"))
-let coq_existS_pattern =
- lazy (snd (parse_pattern "(Coq.Init.Specif.existS ?1 ?2 ?3 ?4)"))
-let coq_existT_pattern =
- lazy (snd (parse_pattern "(Coq.Init.Specif.existT ?1 ?2 ?3 ?4)"))
-let coq_not_pattern =
- lazy (snd (parse_pattern "(Coq.Init.Logic.not ?)"))
-let coq_imp_False_pattern =
- lazy (snd (parse_pattern "? -> Coq.Init.Logic.False"))
-let coq_imp_False_pattern =
- lazy (snd (parse_pattern "? -> Coq.Init.Logic.False"))
-let coq_eqdec_partial_pattern
- lazy (snd (parse_pattern "(sumbool (eq ?1 ?2 ?3) ?4)"))
-let coq_eqdec_pattern
- lazy (snd (parse_pattern "(x,y:?1){<?1>x=y}+{~(<?1>x=y)}"))
-*)
-
-(* The following is less readable but does not depend on parsing *)
-let coq_eq_ref = lazy (reference "Logic" "eq")
-let coq_eqT_ref = lazy (reference "Logic_Type" "eqT")
-let coq_idT_ref = lazy (reference "Logic_Type" "identityT")
-let coq_existS_ref = lazy (reference "Specif" "existS")
-let coq_existT_ref = lazy (reference "Specif" "existT")
-let coq_not_ref = lazy (reference "Logic" "not")
-let coq_False_ref = lazy (reference "Logic" "False")
-let coq_sumbool_ref = lazy (reference "Specif" "sumbool")
-let coq_sig_ref = lazy (reference "Specif" "sig")
-
-(* Pattern "(sig ?1 ?2)" *)
-let coq_sig_pattern =
- lazy (PApp (PRef (Lazy.force coq_sig_ref),
- [| PMeta (Some 1); PMeta (Some 2) |]))
-
-(* Patterns "(eq ?1 ?2 ?3)", "(eqT ?1 ?2 ?3)" and "(idT ?1 ?2 ?3)" *)
-let coq_eq_pattern_gen eq =
- lazy (PApp(PRef (Lazy.force eq), Array.init 3 (fun i -> PMeta (Some (i+1)))))
-let coq_eq_pattern = coq_eq_pattern_gen coq_eq_ref
-let coq_eqT_pattern = coq_eq_pattern_gen coq_eqT_ref
-let coq_idT_pattern = coq_eq_pattern_gen coq_idT_ref
-
-(* Patterns "(existS ?1 ?2 ?3 ?4)" and "(existT ?1 ?2 ?3 ?4)" *)
-let coq_ex_pattern_gen ex =
- lazy (PApp(PRef (Lazy.force ex), Array.init 4 (fun i -> PMeta (Some (i+1)))))
-let coq_existS_pattern = coq_ex_pattern_gen coq_existS_ref
-let coq_existT_pattern = coq_ex_pattern_gen coq_existT_ref
-
-(* Patterns "~ ?" and "? -> False" *)
-let coq_not_pattern = lazy(PApp(PRef (Lazy.force coq_not_ref), [|PMeta None|]))
-let imp a b = PProd (Anonymous, a, b)
-let coq_imp_False_pattern =
- lazy (imp (PMeta None) (PRef (Lazy.force coq_False_ref)))
-
-(* Pattern "(sumbool (eq ?1 ?2 ?3) ?4)" *)
-let coq_eqdec_partial_pattern =
- lazy
- (PApp
- (PRef (Lazy.force coq_sumbool_ref),
- [| Lazy.force coq_eq_pattern; PMeta (Some 4) |]))
-
-(* The expected form of the goal for the tactic Decide Equality *)
-
-(* Pattern "(x,y:?1){<?1>x=y}+{~(<?1>x=y)}" *)
-(* i.e. "(x,y:?1)(sumbool (eq ?1 x y) ~(eq ?1 x y))" *)
-let x = Name (id_of_string "x")
-let y = Name (id_of_string "y")
-let coq_eqdec_pattern =
- lazy
- (PProd (x, PMeta (Some 1), PProd (y, PMeta (Some 1),
- PApp (PRef (Lazy.force coq_sumbool_ref),
- [| PApp (PRef (Lazy.force coq_eq_ref),
- [| PMeta (Some 1); PRel 2; PRel 1 |]);
- PApp (PRef (Lazy.force coq_not_ref),
- [|PApp (PRef (Lazy.force coq_eq_ref),
- [| PMeta (Some 1); PRel 2; PRel 1 |])|]) |]))))
-
-let build_coq_eq_pattern () = Lazy.force coq_eq_pattern
-let build_coq_eqT_pattern () = Lazy.force coq_eqT_pattern
-let build_coq_idT_pattern () = Lazy.force coq_idT_pattern
-let build_coq_existS_pattern () = Lazy.force coq_existS_pattern
-let build_coq_existT_pattern () = Lazy.force coq_existT_pattern
-let build_coq_not_pattern () = Lazy.force coq_not_pattern
-let build_coq_imp_False_pattern () = Lazy.force coq_imp_False_pattern
-let build_coq_eqdec_partial_pattern () = Lazy.force coq_eqdec_partial_pattern
-let build_coq_eqdec_pattern () = Lazy.force coq_eqdec_pattern
-let build_coq_sig_pattern () = Lazy.force coq_sig_pattern
diff --git a/parsing/coqlib.mli b/parsing/coqlib.mli
deleted file mode 100644
index dbe99e399..000000000
--- a/parsing/coqlib.mli
+++ /dev/null
@@ -1,133 +0,0 @@
-(***********************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
-(* \VV/ *************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(***********************************************************************)
-
-(* $Id$ *)
-
-(*i*)
-open Names
-open Libnames
-open Nametab
-open Term
-open Pattern
-(*i*)
-
-(*s This module collects the global references, constructions and
- patterns of the standard library used in ocaml files *)
-
-(*s Global references *)
-
-(* Modules *)
-val logic_module : dir_path
-val logic_type_module : dir_path
-
-(* Natural numbers *)
-val glob_nat : global_reference
-val path_of_O : constructor
-val path_of_S : constructor
-val glob_O : global_reference
-val glob_S : global_reference
-
-(* Equality *)
-val glob_eq : global_reference
-val glob_eqT : global_reference
-
-(*s Constructions and patterns related to Coq initial state are unknown
- at compile time. Therefore, we can only provide methods to build
- them at runtime. This is the purpose of the [constr delayed] and
- [constr_pattern delayed] types. Objects of this time needs to be
- applied to [()] to get the actual constr or pattern at runtime *)
-
-type 'a delayed = unit -> 'a
-
-(*s For Equality tactics *)
-type coq_sigma_data = {
- proj1 : constr;
- proj2 : constr;
- elim : constr;
- intro : constr;
- typ : constr }
-
-val build_sigma_set : unit -> coq_sigma_data
-val build_sigma_type : unit -> coq_sigma_data
-
-type coq_leibniz_eq_data = {
- eq : constr delayed;
- ind : constr delayed;
- rrec : constr delayed option;
- rect : constr delayed option;
- congr: constr delayed;
- sym : constr delayed }
-
-val build_coq_eq_data : coq_leibniz_eq_data
-val build_coq_eqT_data : coq_leibniz_eq_data
-val build_coq_idT_data : coq_leibniz_eq_data
-
-val build_coq_f_equal2 : constr delayed
-val build_coq_eqT : constr delayed
-val build_coq_sym_eqT : constr delayed
-
-(* Empty Type *)
-val build_coq_EmptyT : constr delayed
-
-(* Unit Type and its unique inhabitant *)
-val build_coq_UnitT : constr delayed
-val build_coq_IT : constr delayed
-
-(* Specif *)
-val build_coq_sumbool : constr delayed
-
-(*s Connectives *)
-(* The False proposition *)
-val build_coq_False : constr delayed
-
-(* The True proposition and its unique proof *)
-val build_coq_True : constr delayed
-val build_coq_I : constr delayed
-
-(* Negation *)
-val build_coq_not : constr delayed
-
-(* Conjunction *)
-val build_coq_and : constr delayed
-
-(* Disjunction *)
-val build_coq_or : constr delayed
-
-(* Existential quantifier *)
-val build_coq_ex : constr delayed
-
-(**************************** Patterns ************************************)
-(* ["(eq ?1 ?2 ?3)"] *)
-val build_coq_eq_pattern : constr_pattern delayed
-
-(* ["(eqT ?1 ?2 ?3)"] *)
-val build_coq_eqT_pattern : constr_pattern delayed
-
-(* ["(identityT ?1 ?2 ?3)"] *)
-val build_coq_idT_pattern : constr_pattern delayed
-
-(* ["(existS ?1 ?2 ?3 ?4)"] *)
-val build_coq_existS_pattern : constr_pattern delayed
-
-(* ["(existT ?1 ?2 ?3 ?4)"] *)
-val build_coq_existT_pattern : constr_pattern delayed
-
-(* ["(not ?)"] *)
-val build_coq_not_pattern : constr_pattern delayed
-
-(* ["? -> False"] *)
-val build_coq_imp_False_pattern : constr_pattern delayed
-
-(* ["(sumbool (eq ?1 ?2 ?3) ?4)"] *)
-val build_coq_eqdec_partial_pattern : constr_pattern delayed
-
-(* ["! (x,y:?1). (sumbool (eq ?1 x y) ~(eq ?1 x y))"] *)
-val build_coq_eqdec_pattern : constr_pattern delayed
-
-(* ["(sig ?1 ?2)"] *)
-val build_coq_sig_pattern : constr_pattern delayed
diff --git a/parsing/egrammar.ml b/parsing/egrammar.ml
index 731bb5e64..cec7e4458 100644
--- a/parsing/egrammar.ml
+++ b/parsing/egrammar.ml
@@ -12,20 +12,25 @@ open Pp
open Util
open Extend
open Pcoq
-open Coqast
+open Topconstr
open Ast
open Genarg
+open Libnames
(* State of the grammar extensions *)
type all_grammar_command =
- | AstGrammar of grammar_command
+ | Notation of (string * notation * prod_item list)
+ | Delimiters of (scope_name * prod_item list * prod_item list)
+ | Grammar of grammar_command
| TacticGrammar of
(string * (string * grammar_production list) * Tacexpr.raw_tactic_expr)
list
let subst_all_grammar_command subst = function
- | AstGrammar gc -> AstGrammar (subst_grammar_command subst gc)
+ | Notation _ -> anomaly "Notation not in GRAMMAR summary"
+ | Delimiters _ -> anomaly "Delimiters not in GRAMMAR summary"
+ | Grammar gc -> Grammar (subst_grammar_command subst gc)
| TacticGrammar g -> TacticGrammar g (* TODO ... *)
let (grammar_state : all_grammar_command list ref) = ref []
@@ -45,24 +50,8 @@ let specify_name name e =
Failure("during interpretation of grammar rule "^name^", "^s)
| e -> e
-let gram_action (name, etyp) env act dloc =
- try
- let v = Ast.eval_act dloc env act in
- match etyp, v with
- | (PureAstType, PureAstNode ast) -> Obj.repr ast
- | (AstListType, AstListNode astl) -> Obj.repr astl
- | (GenAstType ConstrArgType, PureAstNode ast) -> Obj.repr ast
- | _ -> grammar_type_error (dloc, "Egrammar.gram_action")
- with e ->
- let (loc, exn) =
- match e with
- | Stdpp.Exc_located (loce, lexn) -> (loce, lexn)
- | e -> (dloc, e)
- in
- Stdpp.raise_with_loc loc (specify_name name exn)
-
(* Translation of environments: a production
- * [ nt1($x1) ... nti($xi) ] -> act($x1..$xi)
+ * [ nt1(x1) ... nti(xi) ] -> act(x1..xi)
* is written (with camlp4 conventions):
* (fun vi -> .... (fun v1 -> act(v1 .. vi) )..)
* where v1..vi are the values generated by non-terminals nt1..nti.
@@ -75,23 +64,42 @@ let gram_action (name, etyp) env act dloc =
*
* (fun v1 ->
* (fun env -> gram_action .. env act)
- * (($x1,v1)::env))
+ * ((x1,v1)::env))
* ...)
- * (($xi,vi)::env)))
+ * ((xi,vi)::env)))
* [])
*)
-let make_act name_typ a pil =
- let act_without_arg env = Gramext.action (gram_action name_typ env a)
- and make_prod_item act_tl = function
- | None -> (* parse a non-binding item *)
- (fun env -> Gramext.action (fun _ -> act_tl env))
- | Some (p, ETast) -> (* non-terminal *)
- (fun env -> Gramext.action (fun v -> act_tl((p,PureAstNode v)::env)))
- | Some (p, ETastl) -> (* non-terminal *)
- (fun env -> Gramext.action (fun v -> act_tl((p,AstListNode v)::env)))
- in
- (List.fold_left make_prod_item act_without_arg pil) []
+open Names
+
+let make_act f pil =
+ let rec make env = function
+ | [] ->
+ Gramext.action (fun loc -> f loc env)
+ | None :: tl -> (* parse a non-binding item *)
+ Gramext.action (fun _ -> make env tl)
+ | Some (p, ETConstr) :: tl -> (* non-terminal *)
+ Gramext.action (fun (v:constr_expr) -> make ((p,v) :: env) tl)
+ | Some (p, ETReference) :: tl -> (* non-terminal *)
+ Gramext.action (fun (v:reference) -> make ((p,CRef v) :: env) tl)
+ | Some (p, ETIdent) :: tl -> (* non-terminal *)
+ Gramext.action (fun (v:identifier) ->
+ make ((p,CRef (Ident (dummy_loc,v))) :: env) tl) in
+ make [] (List.rev pil)
+
+let make_cases_pattern_act f pil =
+ let rec make env = function
+ | [] ->
+ Gramext.action (fun loc -> f loc env)
+ | None :: tl -> (* parse a non-binding item *)
+ Gramext.action (fun _ -> make env tl)
+ | Some (p, ETConstr) :: tl -> (* non-terminal *)
+ Gramext.action (fun v -> make ((p,v) :: env) tl)
+ | Some (p, ETReference) :: tl -> (* non-terminal *)
+ Gramext.action (fun v -> make ((p,CPatAtom (dummy_loc,Some v)) :: env) tl)
+ | Some (p, ETIdent) :: tl ->
+ error "ident entry not admitted in patterns cases syntax extensions" in
+ make [] (List.rev pil)
(* Grammar extension command. Rules are assumed correct.
* Type-checking of grammar rules is done during the translation of
@@ -101,7 +109,8 @@ let make_act name_typ a pil =
* Extend.of_ast) *)
let get_entry_type (u,n) =
- Gram.Entry.obj (object_of_typed_entry (get_entry (get_univ u) n))
+ if u = "constr" & n = "pattern" then Gram.Entry.obj Constr.pattern
+ else Gram.Entry.obj (object_of_typed_entry (get_entry (get_univ u) n))
let rec build_prod_item univ = function
| ProdList0 s -> Gramext.Slist0 (build_prod_item univ s)
@@ -117,26 +126,36 @@ let symbol_of_prod_item univ = function
let eobj = build_prod_item univ nt in
(eobj, ovar)
+(*
let make_rule univ etyp rule =
let pil = List.map (symbol_of_prod_item univ) rule.gr_production in
let (symbs,ntl) = List.split pil in
let act = make_act (rule.gr_name,etyp) rule.gr_action ntl in
(symbs, act)
+*)
+
+let make_rule univ etyp rule =
+ let pil = List.map (symbol_of_prod_item univ) rule.gr_production in
+ let (symbs,ntl) = List.split pil in
+ let f loc env = CGrammar (loc, rule.gr_action, env) in
+ let act = make_act f ntl in
+ (symbs, act)
+
(* Rules of a level are entered in reverse order, so that the first rules
are applied before the last ones *)
let extend_entry univ (te, etyp, ass, rls) =
let rules = List.rev (List.map (make_rule univ etyp) rls) in
- grammar_extend te None [(None, ass, rules)]
+ grammar_extend (object_of_typed_entry te) None [(None, ass, rules)]
(* Defines new entries. If the entry already exists, check its type *)
let define_entry univ {ge_name=n; ge_type=typ; gl_assoc=ass; gl_rules=rls} =
- let typ = if typ = ETast then GenAstType ConstrArgType else AstListType in
+ let typ = entry_type_of_constr_entry_type typ in
let e = force_entry_type univ n typ in
(e,typ,ass,rls)
(* Add a bunch of grammar rules. Does not check if it is well formed *)
-let extend_grammar gram =
+let extend_grammar_rules gram =
let univ = get_univ gram.gc_univ in
let tl = List.map (define_entry univ) gram.gc_entries in
List.iter (extend_entry univ) tl
@@ -154,32 +173,56 @@ let make_prod_item = function
let make_gen_act f pil =
let rec make env = function
| [] ->
- Gramext.action (fun loc -> f env)
+ Gramext.action (fun loc -> f loc env)
| None :: tl -> (* parse a non-binding item *)
Gramext.action (fun _ -> make env tl)
| Some (p, t) :: tl -> (* non-terminal *)
Gramext.action (fun v -> make ((p,in_generic t v) :: env) tl) in
make [] (List.rev pil)
-let make_rule univ f g (s',pt) =
- let hd = Gramext.Stoken ("IDENT", s') in
+let extend_constr entry make_act pt =
+ let univ = get_univ "constr" in
+ let pil = List.map (symbol_of_prod_item univ) pt in
+ let (symbs,ntl) = List.split pil in
+ let act = make_act ntl in
+ grammar_extend entry None [(None, None, [symbs, act])]
+
+let constr_entry name =
+ object_of_typed_entry (get_entry (get_univ "constr") name)
+
+let extend_constr_notation (name,ntn,rule) =
+ let mkact loc env = CNotation (loc,ntn,env) in
+ extend_constr (constr_entry name) (make_act mkact) rule
+
+let extend_constr_grammar (name,c,rule) =
+ let mkact loc env = CGrammar (loc,c,env) in
+ extend_constr (constr_entry name) (make_act mkact) rule
+
+let extend_constr_delimiters (sc,rule,pat_rule) =
+ let mkact loc env = CDelimiters (loc,sc,snd (List.hd env)) in
+ extend_constr (constr_entry "constr8") (make_act mkact) rule;
+ let mkact loc env = CPatDelimiters (loc,sc,snd (List.hd env)) in
+ extend_constr Constr.pattern (make_cases_pattern_act mkact) pat_rule
+
+(* These grammars are not a removable *)
+let make_rule univ f g (s,pt) =
+ let hd = Gramext.Stoken ("IDENT", s) in
let pil = (hd,None) :: List.map g pt in
let (symbs,ntl) = List.split pil in
let act = make_gen_act f ntl in
(symbs, act)
-(* These grammars are not a removable *)
let extend_tactic_grammar s gl =
let univ = get_univ "tactic" in
- let make_act l = Tacexpr.TacExtend (s,List.map snd l) in
- let rules = List.rev (List.map (make_rule univ make_act make_prod_item) gl)
- in Gram.extend Tactic.simple_tactic None [(None, None, rules)]
+ let make_act loc l = Tacexpr.TacExtend (loc,s,List.map snd l) in
+ let rules = List.map (make_rule univ make_act make_prod_item) gl in
+ Gram.extend Tactic.simple_tactic None [(None, None, List.rev rules)]
let extend_vernac_command_grammar s gl =
let univ = get_univ "vernac" in
- let make_act l = Vernacexpr.VernacExtend (s,List.map snd l) in
- let rules = List.rev (List.map (make_rule univ make_act make_prod_item) gl)
- in Gram.extend Vernac_.command None [(None, None, rules)]
+ let make_act loc l = Vernacexpr.VernacExtend (s,List.map snd l) in
+ let rules = List.map (make_rule univ make_act make_prod_item) gl in
+ Gram.extend Vernac_.command None [(None, None, List.rev rules)]
let rec interp_entry_name u s =
let l = String.length s in
@@ -196,9 +239,7 @@ let rec interp_entry_name u s =
let n = Extend.rename_command s in
let e = get_entry (get_univ u) n in
let o = object_of_typed_entry e in
- let t = match type_of_typed_entry e with
- | GenAstType t -> t
- | _ -> failwith "Only entries of generic type can be used in alias" in
+ let t = type_of_typed_entry e in
t, Gramext.Snterm (Pcoq.Gram.Entry.obj o)
let qualified_nterm current_univ = function
@@ -214,16 +255,17 @@ let make_vprod_item univ = function
let add_tactic_entries gl =
let univ = get_univ "tactic" in
- let make_act s tac l = Tacexpr.TacAlias (s,l,tac) in
- let rules =
- List.rev (List.map (fun (s,l,tac) -> make_rule univ (make_act s tac) (make_vprod_item "tactic") l) gl)
- in
- let tacentry = get_entry (get_univ "tactic") "simple_tactic" in
- grammar_extend tacentry None [(None, None, rules)]
+ let make_act s tac loc l = Tacexpr.TacAlias (s,l,tac) in
+ let f (s,l,tac) =
+ make_rule univ (make_act s tac) (make_vprod_item "tactic") l in
+ let rules = List.map f gl in
+ grammar_extend Tactic.simple_tactic None [(None, None, List.rev rules)]
let extend_grammar gram =
(match gram with
- | AstGrammar g -> extend_grammar g
+ | Notation a -> extend_constr_notation a
+ | Delimiters a -> extend_constr_delimiters a
+ | Grammar g -> extend_grammar_rules g
| TacticGrammar l -> add_tactic_entries l);
grammar_state := gram :: !grammar_state
@@ -243,7 +285,9 @@ let factorize_grams l1 l2 =
let number_of_entries gcl =
List.fold_left
(fun n -> function
- | AstGrammar gc -> n + (List.length gc.gc_entries)
+ | Notation _ -> n + 1
+ | Delimiters _ -> n + 2 (* One rule for constr, one for pattern *)
+ | Grammar gc -> n + (List.length gc.gc_entries)
| TacticGrammar l -> n + 1)
0 gcl
diff --git a/parsing/egrammar.mli b/parsing/egrammar.mli
index 73f9e424e..ff3f6284b 100644
--- a/parsing/egrammar.mli
+++ b/parsing/egrammar.mli
@@ -9,11 +9,12 @@
(*i $Id$ i*)
(*i*)
-open Coqast
+open Topconstr
open Ast
open Coqast
open Vernacexpr
open Extend
+open Rawterm
(*i*)
type frozen_t
@@ -23,11 +24,15 @@ val unfreeze : frozen_t -> unit
val init : unit -> unit
type all_grammar_command =
- | AstGrammar of grammar_command
+ | Notation of (string * notation * prod_item list)
+ | Delimiters of (scope_name * prod_item list * prod_item list)
+ | Grammar of grammar_command
| TacticGrammar of (string * (string * grammar_production list) * Tacexpr.raw_tactic_expr) list
val extend_grammar : all_grammar_command -> unit
+val extend_constr_grammar : string * aconstr * prod_item list -> unit
+
(* Add grammar rules for tactics *)
type grammar_tactic_production =
| TacTerm of string
diff --git a/parsing/esyntax.ml b/parsing/esyntax.ml
index 9f802563b..76f4b3f19 100644
--- a/parsing/esyntax.ml
+++ b/parsing/esyntax.ml
@@ -15,11 +15,13 @@ open Libnames
open Coqast
open Ast
open Extend
+open Ppextend
open Vernacexpr
open Names
open Nametab
+open Topconstr
open Symbols
-
+
(*** Syntax keys ***)
(* We define keys for ast and astpats. This is a kind of hash
@@ -84,30 +86,20 @@ let se_key se = spat_key se.syn_astpat
let from_name_table = ref Gmap.empty
let from_key_table = ref Gmapl.empty
-let infix_symbols_map = ref Stringmap.empty
-let infix_names_map = ref Spmap.empty
-
(* Summary operations *)
type frozen_t = (string * string, astpat syntax_entry) Gmap.t *
- (string * key, astpat syntax_entry) Gmapl.t *
- section_path Stringmap.t * string list Spmap.t
+ (string * key, astpat syntax_entry) Gmapl.t
let freeze () =
- (!from_name_table, !from_key_table, !infix_symbols_map, !infix_names_map)
+ (!from_name_table, !from_key_table)
-let unfreeze (fnm,fkm,infs,infn) =
+let unfreeze (fnm,fkm) =
from_name_table := fnm;
- from_key_table := fkm;
- infix_symbols_map := infs;
- infix_names_map := infn
+ from_key_table := fkm
let init () =
from_name_table := Gmap.empty;
from_key_table := Gmapl.empty
-(*
- infix_symbols_map := Stringmap.empty;
- infix_names_map := Spmap.empty
-*)
let find_syntax_entry whatfor gt =
let gt_keys = ast_keys gt in
@@ -140,9 +132,9 @@ let add_ppobject {sc_univ=wf;sc_entries=sel} = List.iter (add_rule wf) sel
(* Pretty-printing machinery *)
-type std_printer = Genarg.constr_ast -> std_ppcmds
+type std_printer = Coqast.t -> std_ppcmds
type unparsing_subfunction = string -> tolerability option -> std_printer
-type primitive_printer = Genarg.constr_ast -> std_ppcmds option
+type primitive_printer = Coqast.t -> std_ppcmds option
(* Module of primitive printers *)
module Ppprim =
@@ -187,9 +179,7 @@ let _ = declare_primitive_printer "print_as" default_scope print_as_printer
(* Handle infix symbols *)
let pr_parenthesis inherited se strm =
- let rule_prec = (se.syn_id, se.syn_prec) in
- let no_paren = tolerable_prec inherited rule_prec in
- if no_paren then
+ if tolerable_prec inherited se.syn_prec then
strm
else
(str"(" ++ strm ++ str")")
@@ -212,7 +202,7 @@ let print_delimiters inh se strm = function
let print_syntax_entry sub_pr scopes env se =
let rec print_hunk rule_prec scopes = function
| PH(e,externpr,reln) ->
- let node = Ast.pat_sub Ast.dummy_loc env e in
+ let node = Ast.pat_sub dummy_loc env e in
let printer =
match externpr with (* May branch to an other printer *)
| Some c ->
@@ -228,8 +218,7 @@ let print_syntax_entry sub_pr scopes env se =
| UNP_BOX (b,sub) -> ppcmd_of_box b (prlist (print_hunk rule_prec scopes) sub)
| UNP_SYMBOLIC _ -> anomaly "handled by call_primitive_parser"
in
- let rule_prec = (se.syn_id, se.syn_prec) in
- prlist (print_hunk rule_prec scopes) se.syn_hunks
+ prlist (print_hunk se.syn_prec scopes) se.syn_hunks
let call_primitive_parser rec_pr otherwise inherited scopes (se,env) =
try (
@@ -242,7 +231,7 @@ let call_primitive_parser rec_pr otherwise inherited scopes (se,env) =
| None -> otherwise ()
| Some (dlm,scopes) ->
(* We can use this printer *)
- let node = Ast.pat_sub Ast.dummy_loc env e in
+ let node = Ast.pat_sub dummy_loc env e in
match pr node with
| Some strm -> print_delimiters inherited se strm dlm
| None -> otherwise ())
diff --git a/parsing/esyntax.mli b/parsing/esyntax.mli
index cf1b0de3f..9ee6b9f0a 100644
--- a/parsing/esyntax.mli
+++ b/parsing/esyntax.mli
@@ -13,6 +13,8 @@ open Pp
open Extend
open Vernacexpr
open Symbols
+open Ppextend
+open Topconstr
(*i*)
(* Syntax entry tables. *)
@@ -33,9 +35,9 @@ val warning_verbose : bool ref
(* Pretty-printing *)
-type std_printer = Genarg.constr_ast -> std_ppcmds
+type std_printer = Coqast.t -> std_ppcmds
type unparsing_subfunction = string -> tolerability option -> std_printer
-type primitive_printer = Genarg.constr_ast -> std_ppcmds option
+type primitive_printer = Coqast.t -> std_ppcmds option
(* Module of constr primitive printers [old style - no scope] *)
module Ppprim :
diff --git a/parsing/extend.ml b/parsing/extend.ml
index a469a648f..0e1f72536 100644
--- a/parsing/extend.ml
+++ b/parsing/extend.ml
@@ -10,9 +10,16 @@
(*i $Id$ i*)
open Util
-open Gramext
open Pp
+open Gramext
+open Names
open Ast
+open Ppextend
+open Topconstr
+open Genarg
+
+type entry_type = argument_type
+type constr_entry_type = ETConstr | ETIdent | ETReference
type nonterm_prod =
| ProdList0 of nonterm_prod
@@ -22,16 +29,16 @@ type nonterm_prod =
type prod_item =
| Term of Token.pattern
- | NonTerm of nonterm_prod * (string * ast_action_type) option
+ | NonTerm of nonterm_prod * (Names.identifier * constr_entry_type) option
type grammar_rule = {
gr_name : string;
gr_production : prod_item list;
- gr_action : act }
+ gr_action : aconstr }
type grammar_entry = {
ge_name : string;
- ge_type : ast_action_type;
+ ge_type : constr_entry_type;
gl_assoc : Gramext.g_assoc option;
gl_rules : grammar_rule list }
@@ -40,18 +47,40 @@ type grammar_command = {
gc_entries : grammar_entry list }
type grammar_associativity = Gramext.g_assoc option
+
+(**********************************************************************)
+(* Globalisation and type-checking of Grammar actions *)
+
+type entry_context = (identifier * constr_entry_type) list
+
+let ast_to_rawconstr = ref (fun _ _ -> AVar (id_of_string "Z"))
+let set_ast_to_rawconstr f = ast_to_rawconstr := f
+
+let act_of_ast vars = function
+ | SimpleAction (loc,ConstrNode a) -> !ast_to_rawconstr vars a
+ | SimpleAction (loc,CasesPatternNode a) -> failwith "TODO:act_of_ast: cases_pattern"
+ | CaseAction _ -> failwith "case/let not supported"
+
+let to_act_check_vars = act_of_ast
+
+type syntax_modifier =
+ | SetItemLevel of string * int
+ | SetLevel of int
+ | SetAssoc of Gramext.g_assoc
+ | SetEntryType of string * constr_entry_type
+
type nonterm =
| NtShort of string
| NtQual of string * string
type grammar_production =
| VTerm of string
- | VNonTerm of loc * nonterm * string option
+ | VNonTerm of loc * nonterm * Names.identifier option
type raw_grammar_rule = string * grammar_production list * grammar_action
type raw_grammar_entry =
- string * ast_action_type * grammar_associativity * raw_grammar_rule list
+ string * constr_entry_type * grammar_associativity * raw_grammar_rule list
let subst_grammar_rule subst gr =
- { gr with gr_action = subst_act subst gr.gr_action }
+ { gr with gr_action = subst_aconstr subst gr.gr_action }
let subst_grammar_entry subst ge =
{ ge with gl_rules = List.map (subst_grammar_rule subst) ge.gl_rules }
@@ -116,15 +145,20 @@ let qualified_nterm current_univ = function
| NtQual (univ, en) -> (rename_command univ, rename_command en)
| NtShort en -> (current_univ, rename_command en)
+let entry_type_of_constr_entry_type = function
+ | ETConstr -> ConstrArgType
+ | ETIdent -> IdentArgType
+ | ETReference -> RefArgType
+
+let constr_entry_of_entry = function
+ | ConstrArgType -> ETConstr
+ | IdentArgType -> ETIdent
+ | RefArgType -> ETReference
+ | _ -> error "Cannot arbitrarily extend non constr/ident/ref entries"
+
let nterm loc (get_entry_type,univ) nont =
let nt = qualified_nterm univ nont in
- try
- let et = match get_entry_type nt with
- | PureAstType -> ETast
- | GenAstType Genarg.ConstrArgType -> ETast
- | AstListType -> ETastl
- | _ -> error "Cannot arbitrarily extend non ast entries"
- in (nt,et)
+ try (nt,constr_entry_of_entry (get_entry_type nt))
with Not_found ->
let (s,n) = nt in
user_err_loc(loc,"Externd.nterm",str("unknown grammar entry: "^s^":"^n))
@@ -132,7 +166,7 @@ let nterm loc (get_entry_type,univ) nont =
let prod_item univ env = function
| VTerm s -> env, Term (terminal s)
| VNonTerm (loc, nt, Some p) ->
- let (nont, etyp) = nterm loc univ nt in
+ let (nont, etyp) = nterm loc univ nt in
((p, etyp) :: env, NonTerm (ProdPrimitive nont, Some (p,etyp)))
| VNonTerm (loc, nt, None) ->
let (nont, etyp) = nterm loc univ nt in
@@ -148,7 +182,7 @@ let rec prod_item_list univ penv pil =
let gram_rule univ etyp (name,pil,act) =
let (pilc, act_env) = prod_item_list univ [] pil in
- let a = Ast.to_act_check_vars act_env etyp act in
+ let a = to_act_check_vars act_env act in
{ gr_name = name; gr_production = pilc; gr_action = a }
let gram_entry univ (nt, etyp, ass, rl) =
@@ -162,21 +196,6 @@ let interp_grammar_command univ ge entryl =
{ gc_univ = univ;
gc_entries = List.map (gram_entry (ge,univ)) entryl }
-(*s Pretty-print. *)
-
-(* Dealing with precedences *)
-
-type precedence = int * int * int
-
-type parenRelation = L | E | Any | Prec of precedence
-
-type ppbox =
- | PpHB of int
- | PpHOVB of int
- | PpHVB of int
- | PpVB of int
- | PpTB
-
(* unparsing objects *)
type 'pat unparsing_hunk =
@@ -212,29 +231,23 @@ let rec subst_hunk subst_pat subst hunk = match hunk with
highest precedence), and the child's one, follow the given
relation. *)
-type tolerability = (string * precedence) * parenRelation
-
+(*
let compare_prec (a1,b1,c1) (a2,b2,c2) =
match (a1=a2),(b1=b2),(c1=c2) with
| true,true,true -> 0
| true,true,false -> c1-c2
| true,false,_ -> b1-b2
| false,_,_ -> a1-a2
+*)
+let compare_prec a1 a2 = a1-a2
-let tolerable_prec oparent_prec_reln (_,child_prec) =
+let tolerable_prec oparent_prec_reln child_prec =
match oparent_prec_reln with
- | Some ((_,pprec), L) -> (compare_prec child_prec pprec) < 0
- | Some ((_,pprec), E) -> (compare_prec child_prec pprec) <= 0
+ | Some (pprec, L) -> (compare_prec child_prec pprec) < 0
+ | Some (pprec, E) -> (compare_prec child_prec pprec) <= 0
| Some (_, Prec level) -> (compare_prec child_prec level) <= 0
| _ -> true
-let ppcmd_of_box = function
- | PpHB n -> h n
- | PpHOVB n -> hov n
- | PpHVB n -> hv n
- | PpVB n -> v n
- | PpTB -> t
-
type 'pat syntax_entry = {
syn_id : string;
syn_prec: precedence;
@@ -265,7 +278,7 @@ let subst_syntax_command subst_pat subst scomm =
{ scomm with sc_entries = sc_entries' }
type syntax_rule = string * Coqast.t * Coqast.t unparsing_hunk list
-type syntax_entry_ast = precedence * syntax_rule list
+type raw_syntax_entry = precedence * syntax_rule list
let rec interp_unparsing env = function
| PH (ast,ext,pr) -> PH (Ast.val_of_ast env ast,ext,pr)
diff --git a/parsing/extend.mli b/parsing/extend.mli
index 7294a2bb0..13e3ee067 100644
--- a/parsing/extend.mli
+++ b/parsing/extend.mli
@@ -9,10 +9,20 @@
(*i $Id$ i*)
(*i*)
-
open Pp
+open Util
+open Names
open Ast
open Coqast
+open Ppextend
+open Topconstr
+open Genarg
+(*i*)
+
+type entry_type = argument_type
+type constr_entry_type = ETConstr | ETIdent | ETReference
+
+val entry_type_of_constr_entry_type : constr_entry_type -> entry_type
type nonterm_prod =
| ProdList0 of nonterm_prod
@@ -22,16 +32,16 @@ type nonterm_prod =
type prod_item =
| Term of Token.pattern
- | NonTerm of nonterm_prod * (string * ast_action_type) option
+ | NonTerm of nonterm_prod * (Names.identifier * constr_entry_type) option
type grammar_rule = {
gr_name : string;
gr_production : prod_item list;
- gr_action : Ast.act }
+ gr_action : aconstr }
type grammar_entry = {
ge_name : string;
- ge_type : ast_action_type;
+ ge_type : constr_entry_type;
gl_assoc : Gramext.g_assoc option;
gl_rules : grammar_rule list }
@@ -40,15 +50,27 @@ type grammar_command = {
gc_entries : grammar_entry list }
type grammar_associativity = Gramext.g_assoc option
+
+(* Globalisation and type-checking of Grammar actions *)
+type entry_context = (identifier * constr_entry_type) list
+val to_act_check_vars : entry_context -> grammar_action -> aconstr
+val set_ast_to_rawconstr : (entry_context -> constr_expr -> aconstr) -> unit
+
+type syntax_modifier =
+ | SetItemLevel of string * int
+ | SetLevel of int
+ | SetAssoc of Gramext.g_assoc
+ | SetEntryType of string * constr_entry_type
+
type nonterm =
| NtShort of string
| NtQual of string * string
type grammar_production =
| VTerm of string
- | VNonTerm of loc * nonterm * string option
+ | VNonTerm of loc * nonterm * Names.identifier option
type raw_grammar_rule = string * grammar_production list * grammar_action
type raw_grammar_entry =
- string * ast_action_type * grammar_associativity * raw_grammar_rule list
+ string * constr_entry_type * grammar_associativity * raw_grammar_rule list
val terminal : string -> string * string
@@ -57,21 +79,6 @@ val rename_command : string -> string
val subst_grammar_command :
Names.substitution -> grammar_command -> grammar_command
-(*s Pretty-print. *)
-
-(* Dealing with precedences *)
-
-type precedence = int * int * int
-
-type parenRelation = L | E | Any | Prec of precedence
-
-type ppbox =
- | PpHB of int
- | PpHOVB of int
- | PpHVB of int
- | PpVB of int
- | PpTB
-
(* unparsing objects *)
type 'pat unparsing_hunk =
@@ -97,11 +104,7 @@ type 'pat unparsing_hunk =
highest precedence), and the child's one, follow the given
relation. *)
-type tolerability = (string * precedence) * parenRelation
-
-val tolerable_prec : tolerability option -> (string * precedence) -> bool
-
-val ppcmd_of_box : ppbox -> std_ppcmds -> std_ppcmds
+val tolerable_prec : tolerability option -> precedence -> bool
type 'pat syntax_entry = {
syn_id : string;
@@ -123,11 +126,11 @@ val subst_syntax_command :
Names.substitution -> 'pat syntax_command -> 'pat syntax_command
type syntax_rule = string * Coqast.t * Coqast.t unparsing_hunk list
-type syntax_entry_ast = precedence * syntax_rule list
+type raw_syntax_entry = precedence * syntax_rule list
val interp_grammar_command :
- string -> (string * string -> entry_type) ->
+ string -> (string * string -> Genarg.argument_type) ->
raw_grammar_entry list -> grammar_command
val interp_syntax_entry :
- string -> syntax_entry_ast list -> Ast.astpat syntax_command
+ string -> raw_syntax_entry list -> Ast.astpat syntax_command
diff --git a/parsing/g_basevernac.ml4 b/parsing/g_basevernac.ml4
index 1d056cf5b..77f587894 100644
--- a/parsing/g_basevernac.ml4
+++ b/parsing/g_basevernac.ml4
@@ -6,12 +6,11 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(*i camlp4deps: "parsing/grammar.cma" i*)
-
(* $Id$ *)
open Coqast
open Extend
+open Ppextend
open Vernacexpr
open Pcoq
open Vernac_
@@ -25,7 +24,7 @@ GEXTEND Gram
class_rawexpr:
[ [ IDENT "FUNCLASS" -> FunClass
| IDENT "SORTCLASS" -> SortClass
- | qid = Prim.qualid -> RefClass qid ] ]
+ | qid = global -> RefClass qid ] ]
;
END;
@@ -54,9 +53,9 @@ GEXTEND Gram
| IDENT "Dump"; IDENT "Universes"; fopt = OPT STRING ->
VernacPrint (PrintUniverses fopt)
- | IDENT "Locate"; qid = qualid -> VernacLocate (LocateTerm qid)
+ | IDENT "Locate"; qid = global -> VernacLocate (LocateTerm qid)
| IDENT "Locate"; IDENT "File"; f = STRING -> VernacLocate (LocateFile f)
- | IDENT "Locate"; IDENT "Library"; qid = qualid ->
+ | IDENT "Locate"; IDENT "Library"; qid = global ->
VernacLocate (LocateLibrary qid)
(* Managing load paths *)
@@ -77,16 +76,16 @@ GEXTEND Gram
(* Printing (careful factorization of entries) *)
| IDENT "Print"; p = printable -> VernacPrint p
- | IDENT "Print"; qid = qualid -> VernacPrint (PrintName qid)
+ | IDENT "Print"; qid = global -> VernacPrint (PrintName qid)
| IDENT "Print" -> VernacPrint PrintLocalContext
- | IDENT "Print"; IDENT "Module"; "Type"; qid = qualid ->
+ | IDENT "Print"; IDENT "Module"; "Type"; qid = global ->
VernacPrint (PrintModuleType qid)
- | IDENT "Print"; IDENT "Module"; qid = qualid ->
+ | IDENT "Print"; IDENT "Module"; qid = global ->
VernacPrint (PrintModule qid)
| IDENT "Inspect"; n = natural -> VernacPrint (PrintInspect n)
(* Searching the environment *)
- | IDENT "Search"; qid = Prim.qualid; l = in_or_out_modules ->
+ | IDENT "Search"; qid = global; l = in_or_out_modules ->
VernacSearch (SearchHead qid, l)
| IDENT "SearchPattern"; c = constr_pattern; l = in_or_out_modules ->
VernacSearch (SearchPattern c, l)
@@ -135,7 +134,7 @@ GEXTEND Gram
| IDENT "Add"; table = IDENT; field = IDENT; v = LIST1 option_ref_value
-> VernacAddOption (SecondaryTable (table,field), v)
- (* Un value qualid ci-dessous va être caché par un field au dessus! *)
+ (* Un value global ci-dessous va être caché par un field au dessus! *)
| IDENT "Add"; table = IDENT; v = LIST1 option_ref_value ->
VernacAddOption (PrimaryTable table, v)
@@ -155,7 +154,7 @@ GEXTEND Gram
;
printable:
[ [ IDENT "All" -> PrintFullContext
- | IDENT "Section"; s = qualid -> PrintSectionContext s
+ | IDENT "Section"; s = global -> PrintSectionContext s
| "Grammar"; uni = IDENT; ent = IDENT ->
(* This should be in "syntax" section but is here for factorization*)
PrintGrammar (uni, ent)
@@ -170,9 +169,9 @@ GEXTEND Gram
| IDENT "Coercion"; IDENT "Paths"; s = class_rawexpr; t = class_rawexpr
-> PrintCoercionPaths (s,t)
| IDENT "Tables" -> PrintTables
- | "Proof"; qid = qualid -> PrintOpaqueName qid
+ | "Proof"; qid = global -> PrintOpaqueName qid
| IDENT "Hint" -> PrintHintGoal
- | IDENT "Hint"; qid = qualid -> PrintHint qid
+ | IDENT "Hint"; qid = global -> PrintHint qid
| IDENT "Hint"; "*" -> PrintHintDb
| IDENT "HintDb"; s = IDENT -> PrintHintDbName s ] ]
;
@@ -181,15 +180,15 @@ GEXTEND Gram
| s = STRING -> StringValue s ] ]
;
option_ref_value:
- [ [ id = qualid -> QualidRefValue id
+ [ [ id = global -> QualidRefValue id
| s = STRING -> StringRefValue s ] ]
;
as_dirpath:
[ [ d = OPT [ "as"; d = dirpath -> d ] -> d ] ]
;
in_or_out_modules:
- [ [ IDENT "inside"; l = LIST1 qualid -> SearchInside l
- | IDENT "outside"; l = LIST1 qualid -> SearchOutside l
+ [ [ IDENT "inside"; l = LIST1 global -> SearchInside l
+ | IDENT "outside"; l = LIST1 global -> SearchOutside l
| -> SearchOutside [] ] ]
;
END;
@@ -218,48 +217,57 @@ GEXTEND Gram
| "Syntax"; u = univ; el = LIST1 syntax_entry SEP ";" ->
VernacSyntax (u,el)
+ | "Syntax"; IDENT "Extension"; s = STRING;
+ l = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ]
+ -> VernacSyntaxExtension (s,l)
+
| IDENT "Open"; IDENT "Scope"; sc = IDENT -> VernacOpenScope sc
| IDENT "Delimiters"; left = STRING; sc = IDENT; right = STRING ->
VernacDelimiters (sc,(left,right))
- | IDENT "Arguments"; IDENT "Scope"; qid = qualid;
+ | IDENT "Arguments"; IDENT "Scope"; qid = global;
"["; scl = LIST0 opt_scope; "]" -> VernacArgumentsScope (qid,scl)
- (* Faudrait une version de qualidarg dans Prim pour respecter l'ordre *)
- | IDENT "Infix"; a = entry_prec; n = natural; op = STRING; p = qualid;
+ | IDENT "Infix"; a = entry_prec; n = natural; op = STRING; p = global;
sc = OPT [ ":"; sc = IDENT -> sc ] -> VernacInfix (a,n,op,p,sc)
- | IDENT "Distfix"; a = entry_prec; n = natural; s = STRING; p = qualid;
+ | IDENT "Distfix"; a = entry_prec; n = natural; s = STRING; p = global;
sc = OPT [ ":"; sc = IDENT -> sc ] -> VernacDistfix (a,n,s,p,sc)
| IDENT "Notation"; a = entry_prec; n = natural; s = STRING; c = constr;
- precl = [ "("; l = LIST1 var_precedence SEP ","; ")" -> l | -> [] ];
- sc = OPT [ ":"; sc = IDENT -> sc ] ->
- VernacNotation (a,n,s,c,precl,sc)
+ modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ];
+ sc = OPT [ ":"; sc = IDENT -> sc ] ->
+ let a = match a with None -> Gramext.LeftA | Some a -> a in
+ VernacNotation (s,c,(SetAssoc a)::(SetLevel n)::modl,sc)
(* "Print" "Grammar" should be here but is in "command" entry in order
to factorize with other "Print"-based vernac entries *)
] ]
;
- var_precedence:
- [ [ x = IDENT; IDENT "at"; IDENT "level"; n = natural -> (x,n) ] ]
+ syntax_modifier:
+ [ [ x = IDENT; IDENT "at"; IDENT "level"; n = natural -> SetItemLevel (x,n)
+ | IDENT "at"; IDENT "level"; n = natural -> SetLevel n
+ | IDENT "left"; IDENT "associativity" -> SetAssoc Gramext.LeftA
+ | IDENT "right"; IDENT "associativity" -> SetAssoc Gramext.RightA
+ | IDENT "no"; IDENT "associativity" -> SetAssoc Gramext.NonA
+ | x = IDENT; typ = syntax_extension_type -> SetEntryType (x,typ) ] ]
+ ;
+ syntax_extension_type:
+ [ [ IDENT "ident" -> ETIdent | IDENT "global" -> ETReference ] ]
;
opt_scope:
[ [ IDENT "_" -> None | sc = IDENT -> Some sc ] ]
;
(* Syntax entries for Grammar. Only grammar_entry is exported *)
grammar_entry:
- [[ nont = located_ident; etyp = set_entry_type; ":=";
+ [[ nont = IDENT; etyp = set_entry_type; ":=";
ep = entry_prec; OPT "|"; rl = LIST0 grammar_rule SEP "|" ->
(nont,etyp,ep,rl) ]]
;
- located_ident:
- [[ id = IDENT -> (loc,id) ]]
- ;
entry_prec:
[[ IDENT "LEFTA" -> Some Gramext.LeftA
| IDENT "RIGHTA" -> Some Gramext.RightA
| IDENT "NONA" -> Some Gramext.NonA
- | -> None ]]
+ | -> None ]]
;
grammar_tactic_rule:
[[ name = rule_name; "["; s = STRING; pil = LIST0 production_item; "]";
@@ -274,9 +282,9 @@ GEXTEND Gram
;
production_item:
[[ s = STRING -> VTerm s
- | nt = non_terminal; po = OPT [ "("; p = Prim.metaident; ")" -> p ] ->
+ | nt = non_terminal; po = OPT [ "("; p = METAIDENT; ")" -> p ] ->
match po with
- | Some p -> VNonTerm (loc,nt,Some (Ast.meta_of_ast p))
+ | Some p -> VNonTerm (loc,nt,Some (Names.id_of_string p))
| _ -> VNonTerm (loc,nt,None) ]]
;
non_terminal:
@@ -294,8 +302,9 @@ GEXTEND Gram
[ [ nm = IDENT; "["; s = astpat; "]"; "->"; u = unparsing -> (nm,s,u) ] ]
;
precedence:
- [ [ a = natural -> (a,0,0)
- | "["; a1 = natural; a2 = natural; a3 = natural; "]" -> (a1,a2,a3) ] ]
+ [ [ a = natural -> a
+(* | "["; a1 = natural; a2 = natural; a3 = natural; "]" -> (a1,a2,a3)*)
+ ] ]
;
unparsing:
[ [ "["; ll = LIST0 next_hunks; "]" -> ll ] ]
@@ -313,7 +322,8 @@ GEXTEND Gram
| e = Prim.ast; oprec = OPT [ ":"; pr = paren_reln_or_extern -> pr ] ->
match oprec with
| Some (ext,pr) -> PH (e,ext,pr)
- | None -> PH (e,None,Any) ]]
+ | None -> PH (e,None,Any)
+ ]]
;
box:
[ [ "<"; bk = box_kind; ">" -> bk ] ]
@@ -335,11 +345,11 @@ GEXTEND Gram
;
(* meta-syntax entries *)
astpat:
- [ [ "<<" ; a = Prim.ast; ">>" -> a
- | a = default_action_parser ->
- match a with
- | Ast.PureAstNode a -> a
- | _ -> failwith "Cannot deal with non pure ast expression" ] ]
+ [ [ "<<" ; a = Prim.ast; ">>" -> a
+ | a = Constr.constr ->
+ Termast.ast_of_rawconstr
+ (Constrintern.interp_rawconstr Evd.empty (Global.env()) a)
+ ] ]
;
action:
[ [ IDENT "let"; p = Prim.astlist; et = set_internal_entry_type;
@@ -356,15 +366,20 @@ GEXTEND Gram
| [ ":"; IDENT "ast" -> () | -> () ] -> Ast.ETast ]]
;
set_entry_type:
- [[ ":"; et = entry_type -> set_default_action_parser et; entry_type_of_parser et
- | -> None ]]
+ [[ ":"; et = entry_type -> set_default_action_parser et;
+ let a = match et with
+ | ConstrParser -> ETConstr
+ | CasesPatternParser ->
+ failwith "entry_type_of_parser: cases_pattern, TODO" in
+ a
+ | -> ETConstr ]]
;
entry_type:
- [[ IDENT "ast"; IDENT "list" -> AstListParser
- | IDENT "ast" -> AstParser
+ [[ IDENT "ast"; IDENT "list" -> Util.error "type ast list no longer supported"
+ | IDENT "ast" -> Util.error "type ast no longer supported"
| IDENT "constr" -> ConstrParser
- | IDENT "cases_pattern" -> CasesPatternParser
- | IDENT "tactic" -> TacticParser
- | IDENT "vernac" -> VernacParser ]]
+ | IDENT "pattern" -> CasesPatternParser
+ | IDENT "tactic" -> assert false
+ | IDENT "vernac" -> Util.error "vernac extensions no longer supported" ] ]
;
END
diff --git a/parsing/g_cases.ml4 b/parsing/g_cases.ml4
index e84a89092..67b6165da 100644
--- a/parsing/g_cases.ml4
+++ b/parsing/g_cases.ml4
@@ -6,66 +6,63 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(*i camlp4deps: "parsing/grammar.cma" i*)
-
(* $Id$ *)
open Pcoq
open Constr
+open Topconstr
+open Term
+open Libnames
+
+open Prim
+
+let pair loc =
+ Qualid (loc, Libnames.qualid_of_string "Coq.Init.Datatypes.pair")
GEXTEND Gram
GLOBAL: constr1 pattern;
pattern:
- [ [ qid = global -> qid
+ [ [ r = Prim.reference -> CPatAtom (loc,Some r)
+ | IDENT "_" -> CPatAtom (loc,None)
(* Hack to parse syntax "(n)" as a natural number *)
| "("; G_constr.test_int_rparen; n = INT; ")" ->
- let n = Coqast.Str (loc,n) in
- <:ast< (PATTDELIMITERS "nat_scope" (PATTNUMERAL $n)) >>
+ let n = CPatNumeral (loc,Bignat.POS (Bignat.of_string n)) in
+ CPatDelimiters (loc,"nat_scope",n)
| "("; p = compound_pattern; ")" -> p
- | n = INT ->
- let n = Coqast.Str (loc,n) in <:ast< (PATTNUMERAL $n) >>
- | "-"; n = INT ->
- let n = Coqast.Str (loc,n) in <:ast< (PATTNEGNUMERAL $n) >>
+ | n = INT -> CPatNumeral (loc,Bignat.POS (Bignat.of_string n))
+ | "-"; n = INT -> CPatNumeral (loc,Bignat.NEG (Bignat.of_string n))
] ]
;
compound_pattern:
- [ [ p = pattern ; lp = ne_pattern_list ->
+ [ [ p = pattern ; lp = LIST1 pattern ->
(match p with
- | Coqast.Node(_,"QUALID",_) -> ()
+ | CPatAtom (_, Some r) -> CPatCstr (loc, r, lp)
| _ -> Util.user_err_loc
- (loc, "compound_pattern", Pp.str "Constructor expected"));
- <:ast< (PATTCONSTRUCT $p ($LIST $lp)) >>
- | p = pattern; "as"; id = Prim.var ->
- <:ast< (PATTAS $id $p)>>
+ (loc, "compound_pattern", Pp.str "Constructor expected"))
+ | p = pattern; "as"; id = base_ident ->
+ CPatAlias (loc, p, id)
| p1 = pattern; ","; p2 = pattern ->
- <:ast< (PATTCONSTRUCT Coq.Init.Datatypes.pair $p1 $p2) >>
+ CPatCstr (loc, pair loc, [p1; p2])
| p = pattern -> p ] ]
;
- ne_pattern_list:
- [ [ c1 = pattern; cl = ne_pattern_list -> c1 :: cl
- | c1 = pattern -> [c1] ] ]
- ;
equation:
- [ [ lhs = ne_pattern_list; "=>"; rhs = constr9 ->
- <:ast< (EQN $rhs ($LIST $lhs)) >> ] ]
+ [ [ lhs = LIST1 pattern; "=>"; rhs = constr9 -> (loc,lhs,rhs) ] ]
;
ne_eqn_list:
- [ [ e = equation; "|"; leqn = ne_eqn_list -> e :: leqn
- | e = equation -> [e] ] ]
+ [ [ leqn = LIST1 equation SEP "|" -> leqn ] ]
;
constr1:
- [ [ "<"; l1 = lconstr; ">"; "Cases"; lc = ne_constr_list; "of";
+ [ [ "<"; p = lconstr; ">"; "Cases"; lc = LIST1 constr; "of";
OPT "|"; eqs = ne_eqn_list; "end" ->
- <:ast< (CASES $l1 (TOMATCH ($LIST $lc)) ($LIST $eqs)) >>
+ CCases (loc, Some p, lc, eqs)
| "Cases"; lc = ne_constr_list; "of";
OPT "|"; eqs = ne_eqn_list; "end" ->
- <:ast< (CASES "SYNTH" (TOMATCH ($LIST $lc)) ($LIST $eqs)) >>
- | "<"; l1 = lconstr; ">"; "Cases"; lc = ne_constr_list; "of";
- "end" ->
- <:ast< (CASES $l1 (TOMATCH ($LIST $lc))) >>
+ CCases (loc, None, lc, eqs)
+ | "<"; p = lconstr; ">"; "Cases"; lc = ne_constr_list; "of"; "end" ->
+ CCases (loc, Some p, lc, [])
| "Cases"; lc = ne_constr_list; "of"; "end" ->
- <:ast< (CASES "SYNTH" (TOMATCH ($LIST $lc))) >> ] ]
+ CCases (loc, None, lc, []) ] ]
;
END;
diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4
index 620b6a800..057494597 100644
--- a/parsing/g_constr.ml4
+++ b/parsing/g_constr.ml4
@@ -6,12 +6,60 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(*i camlp4deps: "parsing/grammar.cma" i*)
-
(* $Id$ *)
open Pcoq
open Constr
+open Rawterm
+open Term
+open Names
+open Libnames
+open Prim
+open Topconstr
+
+(* For the very old syntax of fixpoints *)
+let split_lambda = function
+ | CLambdaN (loc,[[na],t],c) -> (na,t,c)
+ | CLambdaN (loc,([na],t)::bl,c) -> (na,t,CLambdaN(loc,bl,c))
+ | CLambdaN (loc,(na::nal,t)::bl,c) -> (na,t,CLambdaN(loc,(nal,t)::bl,c))
+ | _ -> Util.error "ill-formed fixpoint body"
+
+let split_product = function
+ | CArrow (loc,t,c) -> c
+ | CProdN (loc,[[na],t],c) -> c
+ | CProdN (loc,([na],t)::bl,c) -> CProdN(loc,bl,c)
+ | CProdN (loc,(na::nal,t)::bl,c) -> CProdN(loc,(nal,t)::bl,c)
+ | _ -> Util.error "ill-formed fixpoint body"
+
+let rec split_fix n typ def =
+ if n = 0 then ([],typ,def)
+ else
+ let (na,t,def) = split_lambda def in
+ let typ = split_product typ in
+ let (bl,typ,def) = split_fix (n-1) typ def in
+ (([na],t)::bl,typ,def)
+
+let coerce_to_var = function
+ | CRef (Ident (_,id)) -> id
+ | ast -> Util.user_err_loc
+ (constr_loc ast,"Ast.coerce_to_var",
+ (Pp.str"This expression should be a simple identifier"))
+
+let coerce_to_name = function
+ | CRef (Ident (loc,id)) -> (loc, Name id)
+ | ast -> Util.user_err_loc
+ (constr_loc ast,"Ast.coerce_to_var",
+ (Pp.str"This expression should be a simple identifier"))
+
+open Util
+
+let rec abstract_constr loc c = function
+ | [] -> c
+ | LocalRawDef ((loc',x),b)::bl ->
+ CLetIn (join_loc loc' loc, (loc', x), b, abstract_constr loc c bl)
+ | LocalRawAssum (nal,t)::bl ->
+ let loc' = join_loc (fst (List.hd nal)) loc in
+ CLambdaN(loc', [nal, t], abstract_constr loc c bl)
(* Hack to parse "(n)" as nat without conflicts with the (useless) *)
(* admissible notation "(n)" *)
@@ -30,40 +78,22 @@ let test_int_rparen =
GEXTEND Gram
GLOBAL: constr0 constr1 constr2 constr3 lassoc_constr4 constr5
constr6 constr7 constr8 constr9 constr10 lconstr constr
- ne_ident_comma_list ne_constr_list sort ne_binders_list qualid
- global constr_pattern ident numarg;
- ident:
- [ [ id = Prim.var -> id
+ ne_name_comma_list ne_constr_list sort
+ global constr_pattern Constr.ident;
+ Constr.ident:
+ [ [ id = Prim.ident -> id
- (* This is used in quotations *)
- | id = Prim.metaident -> id ] ]
+ (* This is used in quotations and Syntax *)
+ | id = METAIDENT -> id_of_string id ] ]
;
global:
- [ [ l = qualid -> l
+ [ [ r = Prim.reference -> r
(* This is used in quotations *)
- | id = Prim.metaident -> id ] ]
- ;
- qualid:
- [ [ id = Prim.var; l = fields -> <:ast< (QUALID $id ($LIST $l)) >>
- | id = Prim.var -> <:ast< (QUALID $id) >>
- ] ]
- ;
- fields:
- [ [ id = FIELD; l = fields -> <:ast< ($VAR $id) >> :: l
- | id = FIELD -> [ <:ast< ($VAR $id) >> ]
- ] ]
- ;
- raw_constr:
- [ [ c = Prim.ast -> c ] ]
+ | id = METAIDENT -> Ident (loc,id_of_string id) ] ]
;
constr:
- [ [ c = constr8 -> (* Genarg.ConstrTerm *) c
-(* | IDENT "Inst"; id = Prim.rawident; "["; c = constr; "]" ->
- Genarg.ConstrContext (id, c)
- | IDENT "Eval"; rtc = Tactic.raw_red_tactic; "in"; c = constr ->
- Genarg.ConstrEval (rtc,c)
- | IDENT "Check"; c = constr8 -> <:ast<(CHECK $c)>> *)] ]
+ [ [ c = constr8 -> c ] ]
;
lconstr:
[ [ c = constr10 -> c ] ]
@@ -72,101 +102,85 @@ GEXTEND Gram
[ [ c = constr -> c ] ]
;
ne_constr_list:
- [ [ c1 = constr; cl = ne_constr_list -> c1::cl
- | c1 = constr -> [c1] ] ]
+ [ [ cl = LIST1 constr -> cl ] ]
;
sort:
- [ [ "Set" -> <:ast< (SET) >>
- | "Prop" -> <:ast< (PROP) >>
- | "Type" -> <:ast< (TYPE) >> ] ]
+ [ [ "Set" -> RProp Pos
+ | "Prop" -> RProp Null
+ | "Type" -> RType None ] ]
;
constr0:
- [ [ "?" -> <:ast< (ISEVAR) >>
- | "?"; n = Prim.natural ->
- let n = Coqast.Num (loc,n) in <:ast< (META $n) >>
- | bl = binders; c = constr -> <:ast< ($ABSTRACT "LAMBDALIST" $bl $c) >>
+ [ [ "?" -> CHole loc
+ | "?"; n = Prim.natural -> CMeta (loc, n)
+ | bll = binders; c = constr -> abstract_constr loc c bll
(* Hack to parse syntax "(n)" as a natural number *)
| "("; test_int_rparen; n = INT; ")" ->
- let n = Coqast.Str (loc,n) in
- <:ast< (DELIMITERS "nat_scope" (NUMERAL $n)) >>
- | "("; lc1 = lconstr; ":"; c = constr; body = product_tail ->
- let id = Ast.coerce_to_var lc1 in
- <:ast< (PROD $c [$id]$body) >>
+ let n = CNumeral (loc,Bignat.POS (Bignat.of_string n)) in
+ CDelimiters (loc,"nat_scope",n)
+ | "("; lc1 = lconstr; ":"; c = constr; (bl,body) = product_tail ->
+ let id = coerce_to_name lc1 in
+ CProdN (loc, ([id], c)::bl, body)
+(* TODO: Syntaxe (_:t...)t et (_,x...)t *)
| "("; lc1 = lconstr; ","; lc2 = lconstr; ":"; c = constr;
- body = product_tail ->
- let id1 = Ast.coerce_to_var lc1 in
- let id2 = Ast.coerce_to_var lc2 in
- <:ast< (PRODLIST $c [$id1][$id2]$body) >>
+ (bl,body) = product_tail ->
+ let id1 = coerce_to_name lc1 in
+ let id2 = coerce_to_name lc2 in
+ CProdN (loc, ([id1; id2], c)::bl, body)
| "("; lc1 = lconstr; ","; lc2 = lconstr; ",";
- idl = ne_ident_comma_list; ":"; c = constr; body = product_tail ->
- let id1 = Ast.coerce_to_var lc1 in
- let id2 = Ast.coerce_to_var lc2 in
-(* <:ast< (PRODLIST $c [$id1][$id2]($SLAM $idl $body)) >>*)
- <:ast< ($ABSTRACT "PRODLIST" (BINDERS (BINDER $c $id1 $id2 ($LIST $idl))) $body) >>
+ idl = ne_name_comma_list; ":"; c = constr; (bl,body) = product_tail ->
+ let id1 = coerce_to_name lc1 in
+ let id2 = coerce_to_name lc2 in
+ CProdN (loc, (id1::id2::idl, c)::bl, body)
| "("; lc1 = lconstr; ")" -> lc1
| "("; lc1 = lconstr; ")"; "@"; "["; cl = ne_constr_list; "]" ->
- <:ast< (SOAPP $lc1 ($LIST $cl)) >>
- | IDENT "Fix"; id = ident; "{"; fbinders = fixbinders; "}" ->
- <:ast< (FIX $id ($LIST $fbinders)) >>
- | IDENT "CoFix"; id = ident; "{"; fbinders = cofixbinders; "}" ->
- <:ast< (COFIX $id ($LIST $fbinders)) >>
- | s = sort -> s
- | v = global -> v
- | n = INT ->
- let n = Coqast.Str (loc,n) in <:ast< (NUMERAL $n) >>
- | "-"; n = INT ->
- let n = Coqast.Str (loc,n) in <:ast< (NEGNUMERAL $n) >>
- | "!"; f = global ->
- <:ast< (APPLISTEXPL $f) >>
- ] ]
+ (match lc1 with
+ | CMeta (loc2,n) ->
+ CApp (loc,CMeta (loc2, -n), List.map (fun c -> c, None) cl)
+ | _ ->
+ Util.error "Second-order pattern-matching expects a head metavariable")
+ | IDENT "Fix"; id = identref; "{"; fbinders = fixbinders; "}" ->
+ CFix (loc, id, fbinders)
+ | IDENT "CoFix"; id = identref; "{"; fbinders = cofixbinders; "}" ->
+ CCoFix (loc, id, fbinders)
+ | s = sort -> CSort (loc, s)
+ | v = global -> CRef v
+ | n = INT -> CNumeral (loc,Bignat.POS (Bignat.of_string n))
+ | "-"; n = INT -> CNumeral (loc,Bignat.NEG (Bignat.of_string n))
+ | "!"; f = global -> CAppExpl (loc,f,[])
+ ] ]
;
constr1:
- [ [ "<"; ":"; IDENT "ast"; ":"; "<"; c = raw_constr; ">>" -> c
- | "<"; l1 = lconstr; ">"; IDENT "Match"; c = constr; "with";
- cl = ne_constr_list; "end" ->
- <:ast< (MATCH $l1 $c ($LIST $cl)) >>
- | "<"; l1 = lconstr; ">"; IDENT "Match"; c = constr; "with"; "end"
- ->
- <:ast< (MATCH $l1 $c) >>
- | "<"; l1 = lconstr; ">"; IDENT "Case"; c = constr; "of";
- cl = ne_constr_list; "end" ->
- <:ast< (CASE $l1 $c ($LIST $cl)) >>
- | "<"; l1 = lconstr; ">"; IDENT "Case"; c = constr; "of"; "end" ->
- <:ast< (CASE $l1 $c) >>
- | IDENT "Case"; c = constr; "of"; cl = ne_constr_list; "end" ->
- <:ast< (CASE "SYNTH" $c ($LIST $cl)) >>
- | IDENT "Case"; c = constr; "of"; "end" ->
- <:ast< (CASE "SYNTH" $c) >>
- | IDENT "Match"; c = constr; "with"; cl = ne_constr_list; "end" ->
- <:ast< (MATCH "SYNTH" $c ($LIST $cl)) >>
- | IDENT "let"; "("; b = ne_ident_comma_list; ")"; "=";
+ [ [ "<"; p = lconstr; ">"; IDENT "Match"; c = constr; "with";
+ cl = LIST0 constr; "end" ->
+ COrderedCase (loc, MatchStyle, Some p, c, cl)
+ | "<"; p = lconstr; ">"; IDENT "Case"; c = constr; "of";
+ cl = LIST0 constr; "end" ->
+ COrderedCase (loc, RegularStyle, Some p, c, cl)
+ | IDENT "Case"; c = constr; "of"; cl = LIST0 constr; "end" ->
+ COrderedCase (loc, RegularStyle, None, c, cl)
+ | IDENT "Match"; c = constr; "with"; cl = LIST1 constr; "end" ->
+ COrderedCase (loc, MatchStyle, None, c, cl)
+ | IDENT "let"; "("; b = ne_name_comma_list; ")"; "=";
c = constr; "in"; c1 = constr->
- <:ast< (LET "SYNTH" $c ($ABSTRACT "LAMBDALIST"
- (BINDERS (BINDER (ISEVAR) ($LIST $b))) $c1)) >>
- | IDENT "let"; id1 = ident ; "="; c = opt_casted_constr;
- "in"; c1 = constr ->
- <:ast< (LETIN $c [$id1]$c1) >>
-(*
- | IDENT "let"; id1 = ident ; "="; c = opt_casted_constr;
- "in"; c1 = constr ->
- <:ast< (LETIN $c [$id1]$c1) >>
-*)
+ (* TODO: right loc *)
+ COrderedCase
+ (loc, LetStyle, None, c, [CLambdaN (loc, [b, CHole loc], c1)])
+ | IDENT "let"; na = name; "="; c = opt_casted_constr; "in"; c1 = constr
+ -> CLetIn (loc, na, c, c1)
| IDENT "if"; c1 = constr; IDENT "then"; c2 = constr;
IDENT "else"; c3 = constr ->
- <:ast< ( IF "SYNTH" $c1 $c2 $c3) >>
-(* EN ATTENTE DE REMPLACER CE QUI EST DANS Program.v ... *)
- | "<"; l1 = lconstr; ">";
- IDENT "let"; "("; b = ne_ident_comma_list; ")"; "=";
+ COrderedCase (loc, IfStyle, None, c1, [c2; c3])
+ | "<"; p = lconstr; ">";
+ IDENT "let"; "("; b = ne_name_comma_list; ")"; "=";
c = constr; "in"; c1 = constr ->
-(* <:ast< (CASE "NOREC" $l1 $c (LAMBDALIST (ISEVAR) ($SLAM $b $c1))) >>*)
- <:ast< (LET $l1 $c ($ABSTRACT "LAMBDALIST" (BINDERS
- (BINDER (ISEVAR) ($LIST $b))) $c1)) >>
- | "<"; l1 = lconstr; ">";
+ (* TODO: right loc *)
+ COrderedCase (loc, LetStyle, Some p, c,
+ [CLambdaN (loc, [b, CHole loc], c1)])
+ | "<"; p = lconstr; ">";
IDENT "if"; c1 = constr; IDENT "then";
c2 = constr; IDENT "else"; c3 = constr ->
- <:ast< (IF $l1 $c1 $c2 $c3) >>
- | c = constr0 -> c
- ] ]
+ COrderedCase (loc, IfStyle, Some p, c1, [c2; c3])
+ | c = constr0 -> c ] ]
;
constr2: (* ~ will be here *)
[ [ c = constr1 -> c ] ]
@@ -188,113 +202,98 @@ GEXTEND Gram
;
constr8: (* <-> will be here *)
[ [ c1 = constr7 -> c1
- | c1 = constr7; "->"; c2 = constr8 -> <:ast< (PROD $c1 [<>]$c2) >> ] ]
+ | c1 = constr7; "->"; c2 = constr8 -> CArrow (loc, c1, c2) ] ]
;
constr9:
[ [ c1 = constr8 -> c1
- | c1 = constr8; "::"; c2 = constr8 -> <:ast< (CAST $c1 $c2) >> ] ]
- ;
- numarg:
- [ [ n = Prim.natural -> Coqast.Num (loc, n) ] ]
- ;
- simple_binding:
- [ [ id = ident; ":="; c = constr -> <:ast< (BINDING $id $c) >>
- | n = numarg; ":="; c = constr -> <:ast< (BINDING $n $c) >> ] ]
- ;
- simple_binding_list:
- [ [ b = simple_binding; l = simple_binding_list -> b :: l | -> [] ] ]
- ;
- binding_list:
- [ [ c1 = constr; ":="; c2 = constr; bl = simple_binding_list ->
- Coqast.Node
- (loc, "EXPLICITBINDINGS",
- (Coqast.Node (loc, "BINDING", [Ast.coerce_to_var c1; c2]) :: bl))
- | n = numarg; ":="; c = constr; bl = simple_binding_list ->
- <:ast<(EXPLICITBINDINGS (BINDING $n $c) ($LIST $bl))>>
- | c1 = constr; bl = LIST0 constr ->
- <:ast<(IMPLICITBINDINGS $c1 ($LIST $bl))>> ] ]
+ | c1 = constr8; "::"; c2 = constr8 -> CCast (loc, c1, c2) ] ]
;
constr10:
- [ [ "!"; f = global; args = LIST0 constr9 ->
- <:ast< (APPLISTEXPL $f ($LIST $args)) >>
+ [ [ "!"; f = global; args = LIST0 constr9 -> CAppExpl (loc, f, args)
+(*
| "!"; f = global; "with"; b = binding_list ->
<:ast< (APPLISTWITH $f $b) >>
- | f = constr9; args = LIST1 constr91 ->
- <:ast< (APPLIST $f ($LIST $args)) >>
+*)
+ | f = constr9; args = LIST1 constr91 -> CApp (loc, f, args)
| f = constr9 -> f ] ]
;
- ne_ident_comma_list:
- [ [ id = ident; ","; idl = ne_ident_comma_list -> id :: idl
- | id = ident -> [id] ] ]
+ ne_name_comma_list:
+ [ [ nal = LIST1 name SEP "," -> nal ] ]
;
- ident_comma_list_tail:
- [ [ ","; idl = ne_ident_comma_list -> idl
+ name_comma_list_tail:
+
+ [ [ ","; idl = ne_name_comma_list -> idl
| -> [] ] ]
;
opt_casted_constr:
- [ [ c = constr; ":"; t = constr -> <:ast< (CAST $c $t) >>
+ [ [ c = constr; ":"; t = constr -> CCast (loc, c, t)
| c = constr -> c ] ]
;
- vardecls: (* This is interpreted by Pcoq.abstract_binder *)
- [ [ id = ident; idl = ident_comma_list_tail; c = type_option ->
- <:ast< (BINDER $c $id ($LIST $idl)) >>
- | id = ident; ":="; c = opt_casted_constr ->
- <:ast< (LETIN $c $id) >>
- | id = ident; "="; c = opt_casted_constr ->
- <:ast< (LETIN $c $id) >> ] ]
+ vardecls:
+ [ [ na = name; nal = name_comma_list_tail; c = type_option ->
+ LocalRawAssum (na::nal,c)
+ | na = name; "="; c = opt_casted_constr ->
+ LocalRawDef (na, c)
+ | na = name; ":="; c = opt_casted_constr ->
+ LocalRawDef (na, c)
+
+ (* This is used in quotations *)
+ | id = METAIDENT; c = type_option -> LocalRawAssum ([loc, Name (id_of_string id)], c)
+ ] ]
;
ne_vardecls_list:
[ [ id = vardecls; ";"; idl = ne_vardecls_list -> id :: idl
| id = vardecls -> [id] ] ]
;
binders:
- [ [ "["; bl = ne_vardecls_list; "]" -> <:ast< (BINDERS ($LIST $bl)) >> ] ]
- ;
- rawbinders:
[ [ "["; bl = ne_vardecls_list; "]" -> bl ] ]
;
- ne_binders_list:
- [ [ bl = rawbinders; bll = ne_binders_list -> bl @ bll
- | bl = rawbinders -> bl ] ]
+ simple_params:
+ [ [ idl = LIST1 name SEP ","; ":"; c = constr -> (idl, c)
+ | idl = LIST1 name SEP "," -> (idl, CHole loc)
+ ] ]
+ ;
+ simple_binders:
+ [ [ "["; bll = LIST1 simple_params SEP ";"; "]" -> bll ] ]
+ ;
+ ne_simple_binders_list:
+ [ [ bll = LIST1 simple_binders -> List.flatten bll ] ]
;
type_option:
[ [ ":"; c = constr -> c
- | -> <:ast< (ISEVAR) >> ] ]
+ | -> CHole loc ] ]
;
constr91:
- [ [ n = INT; "!"; c1 = constr9 ->
- let n = Coqast.Num (loc,int_of_string n) in <:ast< (EXPL $n $c1) >>
- | n = INT ->
- let n = Coqast.Str (loc,n) in <:ast< (NUMERAL $n) >>
- | c1 = constr9 -> c1 ] ]
+ [ [ n = natural; "!"; c = constr9 -> (c, Some n)
+ | n = natural ->
+ (CNumeral (loc, Bignat.POS (Bignat.of_string (string_of_int n))), None)
+ | c = constr9 -> (c, None) ] ]
;
fixbinder:
- [ [ id = ident; "/"; recarg = Prim.natural; ":"; type_ = constr;
- ":="; def = constr ->
- let recarg = Coqast.Num (loc,recarg) in
- <:ast< (NUMFDECL $id $recarg $type_ $def) >>
- | id = ident; bl = ne_binders_list; ":"; type_ = constr;
+ [ [ id = base_ident; "/"; recarg = natural; ":"; type_ = constr;
":="; def = constr ->
- <:ast< (FDECL $id (BINDERS ($LIST $bl)) $type_ $def) >> ] ]
+ Options.if_verbose Pp.warning
+ "Checking of the fixpoint type not done for very-old-style fixpoint";
+ let (bl, typ, def) = split_fix recarg type_ def in (id, bl, typ, def)
+ | id = base_ident; bl = ne_simple_binders_list; ":"; type_ = constr;
+ ":="; def = constr ->
+ (id, bl, type_, def) ] ]
;
fixbinders:
- [ [ fb = fixbinder; "with"; fbs = fixbinders -> fb::fbs
- | fb = fixbinder -> [fb] ] ]
+ [ [ fbs = LIST1 fixbinder SEP "with" -> fbs ] ]
;
cofixbinder:
- [ [ id = ident; ":"; type_ = constr; ":="; def = constr ->
- <:ast< (CFDECL $id $type_ $def) >> ] ]
+ [ [ id = base_ident; ":"; type_ = constr; ":="; def = constr ->
+ (id, type_, def) ] ]
;
cofixbinders:
- [ [ fb = cofixbinder; "with"; fbs = cofixbinders -> fb::fbs
- | fb = cofixbinder -> [fb] ] ]
+ [ [ fbs = LIST1 cofixbinder SEP "with" -> fbs ] ]
;
product_tail:
- [ [ ";"; idl = ne_ident_comma_list;
- ":"; c = constr; c2 = product_tail ->
- <:ast< ($ABSTRACT "PRODLIST" (BINDERS (BINDER $c ($LIST $idl))) $c2) >>
- | ";"; idl = ne_ident_comma_list; c2 = product_tail ->
- <:ast< ($ABSTRACT "PRODLIST" (BINDERS (BINDER (ISEVAR) ($LIST $idl))) $c2) >>
- | ")"; c = constr -> c ] ]
+ [ [ ";"; idl = ne_name_comma_list; ":"; c = constr;
+ (bl,c2) = product_tail -> ((idl, c)::bl, c2)
+ | ";"; idl = ne_name_comma_list; (bl,c2) = product_tail ->
+ ((idl, CHole (fst (List.hd idl)))::bl, c2)
+ | ")"; c = constr -> ([], c) ] ]
;
END;;
diff --git a/parsing/g_ltac.ml4 b/parsing/g_ltac.ml4
index a7c37160a..21206e6db 100644
--- a/parsing/g_ltac.ml4
+++ b/parsing/g_ltac.ml4
@@ -6,14 +6,12 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(*i camlp4deps: "parsing/grammar.cma" i*)
-
(* $Id$ *)
open Pp
open Util
open Ast
-open Coqast
+open Topconstr
open Rawterm
open Tacexpr
open Ast
@@ -23,15 +21,16 @@ open Qast
else
open Pcoq
+open Prim
open Tactic
ifdef Quotify then
open Q
type let_clause_kind =
- | LETTOPCLAUSE of Names.identifier * Genarg.constr_ast
+ | LETTOPCLAUSE of Names.identifier * constr_expr
| LETCLAUSE of
- (Names.identifier Util.located * Genarg.constr_ast may_eval option * raw_tactic_arg)
+ (Names.identifier Util.located * constr_expr may_eval option * raw_tactic_arg)
ifdef Quotify then
module Prelude = struct
@@ -69,20 +68,20 @@ GEXTEND Gram
GLOBAL: tactic_atom tactic_atom0 tactic_expr input_fun;
*)
input_fun:
- [ [ l = Prim.ident -> Some l
+ [ [ l = base_ident -> Some l
| "()" -> None ] ]
;
let_clause:
- [ [ id = Prim.rawident; "="; te = tactic_letarg -> LETCLAUSE (id, None, te)
- | id = Prim.ident; ":"; c = Constr.constr; ":="; "Proof" ->
+ [ [ id = identref; "="; te = tactic_letarg -> LETCLAUSE (id, None, te)
+ | id = base_ident; ":"; c = Constr.constr; ":="; "Proof" ->
LETTOPCLAUSE (id, c)
- | id = Prim.rawident; ":"; c = constrarg; ":="; te = tactic_letarg ->
+ | id = identref; ":"; c = constrarg; ":="; te = tactic_letarg ->
LETCLAUSE (id, Some c, te)
- | id = Prim.ident; ":"; c = Constr.constr ->
+ | id = base_ident; ":"; c = Constr.constr ->
LETTOPCLAUSE (id, c) ] ]
;
rec_clause:
- [ [ name = Prim.rawident; it = LIST1 input_fun; "->"; body = tactic_expr ->
+ [ [ name = identref; it = LIST1 input_fun; "->"; body = tactic_expr ->
(name,(it,body)) ] ]
;
match_pattern:
@@ -92,7 +91,7 @@ GEXTEND Gram
| pc = Constr.constr_pattern -> Term pc ] ]
;
match_hyps:
- [ [ id = Prim.rawident; ":"; mp = match_pattern -> Hyp (id, mp)
+ [ [ id = identref; ":"; mp = match_pattern -> Hyp (id, mp)
| IDENT "_"; ":"; mp = match_pattern -> NoHypId mp ] ]
;
match_context_rule:
@@ -126,7 +125,7 @@ GEXTEND Gram
;
tactic_expr3:
[ [ IDENT "Try"; ta = tactic_expr3 -> TacTry ta
- | IDENT "Do"; n = Prim.natural; ta = tactic_expr3 -> TacDo (n,ta)
+ | IDENT "Do"; n = natural; ta = tactic_expr3 -> TacDo (n,ta)
| IDENT "Repeat"; ta = tactic_expr3 -> TacRepeat ta
| IDENT "Progress"; ta = tactic_expr3 -> TacProgress ta
| IDENT "Info"; tc = tactic_expr3 -> TacInfo tc
@@ -179,7 +178,7 @@ GEXTEND Gram
TacMatch (c,mrl)
(*To do: put Abstract in Refiner*)
| IDENT "Abstract"; tc = tactic_expr -> TacAbstract (tc,None)
- | IDENT "Abstract"; tc = tactic_expr; "using"; s = Prim.ident ->
+ | IDENT "Abstract"; tc = tactic_expr; "using"; s = base_ident ->
TacAbstract (tc,Some s)
(*End of To do*)
| IDENT "First" ; "["; l = LIST0 tactic_expr SEP "|"; "]" ->
@@ -188,7 +187,7 @@ GEXTEND Gram
TacSolve l
| IDENT "Idtac" -> TacId
| IDENT "Fail" -> TacFail fail_default_value
- | IDENT "Fail"; n = Prim.natural -> TacFail n
+ | IDENT "Fail"; n = natural -> TacFail n
| st = simple_tactic -> TacAtom (loc,st)
| "("; a = tactic_expr; ")" -> a
| a = tactic_arg -> TacArg a
@@ -203,7 +202,7 @@ GEXTEND Gram
parsed as lqualid! *)
[ [ IDENT "Eval"; rtc = red_expr; "in"; c = Constr.constr ->
ConstrMayEval (ConstrEval (rtc,c))
- | IDENT "Inst"; id = Prim.rawident; "["; c = Constr.constr; "]" ->
+ | IDENT "Inst"; id = identref; "["; c = Constr.constr; "]" ->
ConstrMayEval (ConstrContext (id,c))
| IDENT "Check"; c = Constr.constr ->
ConstrMayEval (ConstrTypeOf c)
@@ -213,7 +212,7 @@ GEXTEND Gram
tactic_arg1:
[ [ IDENT "Eval"; rtc = red_expr; "in"; c = Constr.constr ->
ConstrMayEval (ConstrEval (rtc,c))
- | IDENT "Inst"; id = Prim.rawident; "["; c = Constr.constr; "]" ->
+ | IDENT "Inst"; id = identref; "["; c = Constr.constr; "]" ->
ConstrMayEval (ConstrContext (id,c))
| IDENT "Check"; c = Constr.constr ->
ConstrMayEval (ConstrTypeOf c)
@@ -225,14 +224,14 @@ GEXTEND Gram
[ [ "("; a = tactic_expr; ")" -> Tacexp a
| "()" -> TacVoid
| qid = lqualid -> Reference qid
- | n = Prim.integer -> Integer n
+ | n = integer -> Integer n
| id = METAIDENT -> MetaIdArg (loc,id)
- | "?" -> ConstrMayEval (ConstrTerm <:ast< (ISEVAR) >>)
- | "?"; n = Prim.natural -> MetaNumArg (loc,n)
+ | "?" -> ConstrMayEval (ConstrTerm (CHole loc))
+ | "?"; n = natural -> MetaNumArg (loc,n)
| "'"; c = Constr.constr -> ConstrMayEval (ConstrTerm c) ] ]
;
lqualid:
- [ [ ref = Prim.reference -> ref ] ]
+ [ [ ref = reference -> ref ] ]
;
(* Definitions for tactics *)
@@ -241,18 +240,18 @@ GEXTEND Gram
| IDENT "Tactic" ] ]
;
vrec_clause:
- [ [ name = Prim.rawident; it=LIST1 input_fun; ":="; body = tactic_expr ->
+ [ [ name = identref; it=LIST1 input_fun; ":="; body = tactic_expr ->
(name, TacFunRec (name, (it, body)))
- | name = Prim.rawident; ":="; body = tactic_expr ->
+ | name = identref; ":="; body = tactic_expr ->
(name, body) ] ]
;
tactic:
[ [ tac = tactic_expr -> tac ] ]
;
Vernac_.command:
- [ [ deftok; "Definition"; name = Prim.rawident; ":="; body = tactic ->
+ [ [ deftok; "Definition"; name = identref; ":="; body = tactic ->
Vernacexpr.VernacDeclareTacticDefinition (loc, [name, body])
- | deftok; "Definition"; name = Prim.rawident; largs=LIST1 input_fun;
+ | deftok; "Definition"; name = identref; largs=LIST1 input_fun;
":="; body=tactic_expr ->
Vernacexpr.VernacDeclareTacticDefinition
(loc, [name, TacFun (largs,body)])
diff --git a/parsing/g_minicoq.ml4 b/parsing/g_minicoq.ml4
index 8c5df17a7..5ea97ae7d 100644
--- a/parsing/g_minicoq.ml4
+++ b/parsing/g_minicoq.ml4
@@ -6,8 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(*i camlp4deps: "parsing/grammar.cma" i*)
-
(* $Id$ *)
open Pp
diff --git a/parsing/g_module.ml4 b/parsing/g_module.ml4
index 56db0cb59..a3714c43b 100644
--- a/parsing/g_module.ml4
+++ b/parsing/g_module.ml4
@@ -6,8 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(*i camlp4deps: "parsing/grammar.cma" i*)
-
(* $Id$ *)
open Pp
@@ -16,75 +14,33 @@ open Pcoq
open Prim
open Module
open Util
+open Topconstr
(* Grammar rules for module expressions and types *)
GEXTEND Gram
- GLOBAL: ne_binders_list module_expr
- module_type;
+ GLOBAL: module_expr module_type;
- ident:
- [ [ id = Prim.var -> id ] ]
- ;
-
- ident_comma_list_tail:
- [ [ ","; idl = LIST0 ident SEP "," -> idl
- | -> [] ] ]
- ;
-
- qualid:
- [ [ id = Prim.var; l = fields -> <:ast< (QUALID $id ($LIST $l)) >>
- | id = Prim.var -> <:ast< (QUALID $id) >>
- ] ]
- ;
- fields:
- [ [ id = FIELD; l = fields -> <:ast< ($VAR $id) >> :: l
- | id = FIELD -> [ <:ast< ($VAR $id) >> ]
- ] ]
- ;
-
- vardecls: (* This is interpreted by Pcoq.abstract_binder *)
- [ [ id = ident; idl = ident_comma_list_tail;
- ":"; mty = module_type ->
- <:ast< (BINDER $mty $id ($LIST $idl)) >> ] ]
- ;
-
- ne_vardecls_list:
- [ [ id = vardecls; ";"; idl = ne_vardecls_list -> id :: idl
- | id = vardecls -> [id] ] ]
- ;
-
- rawbinders:
- [ [ "["; bl = ne_vardecls_list; "]" -> bl ] ]
- ;
-
- ne_binders_list:
- [ [ bl = rawbinders; bll = ne_binders_list -> bl @ bll
- | bl = rawbinders -> bl ] ]
- ;
-
module_expr:
- [ [ qid = qualid -> <:ast< (MODEXPRQID $qid) >>
- | me1 = module_expr; me2 = module_expr ->
- <:ast< (MODEXPRAPP $me1 $me2) >>
- | "("; me = module_expr; ")" ->
- me
+ [ [ qid = qualid -> CMEident qid
+ | me1 = module_expr; me2 = module_expr -> CMEapply (me1,me2)
+ | "("; me = module_expr; ")" -> me
(* ... *)
] ]
;
with_declaration:
- [ [ "Definition"; id = ident; ":="; c = Constr.constr ->
- <:ast< (WITHDEFINITION $id $c) >>
- | IDENT "Module"; id = ident; ":="; qid = qualid ->
- <:ast< (WITHMODULE $id $qid) >>
+ [ [ "Definition"; id = base_ident; ":="; c = Constr.constr ->
+ CWith_Definition (id,c)
+ | IDENT "Module"; id = base_ident; ":="; qid = qualid ->
+ CWith_Module (id,qid)
] ]
;
module_type:
- [ [ qid = qualid -> <:ast< (MODTYPEQID $qid) >>
+ [ [ qid = qualid -> CMTEident qid
(* ... *)
| mty = module_type; "with"; decl = with_declaration ->
- <:ast< (MODTYPEWITH $mty $decl)>> ] ]
+ CMTEwith (mty,decl) ] ]
;
END
diff --git a/parsing/g_prim.ml4 b/parsing/g_prim.ml4
index 5363be633..f65ebd64d 100644
--- a/parsing/g_prim.ml4
+++ b/parsing/g_prim.ml4
@@ -8,14 +8,11 @@
(*i $Id$ i*)
-(*
-camlp4o pa_ifdef.cmo pa_extend.cmo pr_o.cmo pr_extend.cmo -quotify -DQuotify -impl parsing/g_prim.ml4
-*)
-
open Coqast
open Pcoq
open Names
open Libnames
+open Topconstr
ifdef Quotify then
open Qast
@@ -72,23 +69,30 @@ ifdef Quotify then
open Q
GEXTEND Gram
- GLOBAL: var ident natural metaident integer string preident ast astpat
- astact astlist qualid reference dirpath rawident;
+ GLOBAL: ident natural integer string preident ast
+ astlist qualid reference dirpath identref name base_ident var;
+ (* Compatibility: Prim.var is a synonym of Prim.ident *)
+ var:
+ [ [ id = ident -> id ] ]
+ ;
metaident:
[ [ s = METAIDENT -> Nmeta (loc,s) ] ]
;
- var:
- [ [ id = ident -> Nvar(loc, id) ] ]
- ;
preident:
[ [ s = IDENT -> s ] ]
;
- ident:
+ base_ident:
[ [ s = IDENT -> local_id_of_string s ] ]
;
- rawident:
- [ [ id = ident -> (loc,id) ] ]
+ name:
+ [ [ IDENT "_" -> (loc, Anonymous) | id = base_ident -> (loc, Name id) ] ]
+ ;
+ identref:
+ [ [ id = base_ident -> (loc,id) ] ]
+ ;
+ ident:
+ [ [ id = base_ident -> id ] ]
;
natural:
[ [ i = INT -> local_make_posint i ] ]
@@ -101,7 +105,8 @@ GEXTEND Gram
[ [ s = FIELD -> local_id_of_string s ] ]
;
dirpath:
- [ [ id = ident; l = LIST0 field -> local_make_dirpath (local_append l id) ] ]
+ [ [ id = base_ident; l = LIST0 field ->
+ local_make_dirpath (local_append l id) ] ]
;
fields:
[ [ id = field; (l,id') = fields -> (local_append l id,id')
@@ -109,26 +114,26 @@ GEXTEND Gram
] ]
;
basequalid:
- [ [ id = ident; (l,id')=fields -> local_make_qualid (local_append l id) id'
- | id = ident -> local_make_short_qualid id
+ [ [ id = base_ident; (l,id')=fields -> local_make_qualid (local_append l id) id'
+ | id = base_ident -> local_make_short_qualid id
] ]
;
qualid:
[ [ qid = basequalid -> loc, qid ] ]
;
reference:
- [ [ id = ident; (l,id') = fields ->
- Coqast.RQualid (loc, local_make_qualid (local_append l id) id')
- | id = ident -> Coqast.RIdent (loc,id)
+ [ [ id = base_ident; (l,id') = fields ->
+ Qualid (loc, local_make_qualid (local_append l id) id')
+ | id = base_ident -> Ident (loc,id)
] ]
;
string:
[ [ s = STRING -> s ] ]
;
astpath:
- [ [ id = ident; (l,a) = fields ->
+ [ [ id = base_ident; (l,a) = fields ->
Path(loc, local_make_path (local_append l id) a)
- | id = ident -> Nvar(loc, id)
+ | id = base_ident -> Nvar(loc, id)
] ]
;
(* ast *)
@@ -156,6 +161,6 @@ GEXTEND Gram
| "'"; a = ast -> Node(loc,"$QUOTE",[a]) ] ]
;
astlist:
- [ [ l = LIST0 Prim.ast -> l ] ]
+ [ [ l = LIST0 ast -> l ] ]
;
END
diff --git a/parsing/g_proofs.ml4 b/parsing/g_proofs.ml4
index d4a00346b..52100764d 100644
--- a/parsing/g_proofs.ml4
+++ b/parsing/g_proofs.ml4
@@ -6,8 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(*i camlp4deps: "parsing/grammar.cma" i*)
-
(* $Id$ *)
open Pcoq
@@ -15,9 +13,10 @@ open Pp
open Tactic
open Util
open Vernac_
-open Coqast
+open Topconstr
open Vernacexpr
open Prim
+open Constr
(* Proof commands *)
GEXTEND Gram
@@ -42,17 +41,17 @@ GEXTEND Gram
*)
| IDENT "Abort" -> VernacAbort None
| IDENT "Abort"; IDENT "All" -> VernacAbortAll
- | IDENT "Abort"; id = ident -> VernacAbort (Some id)
+ | IDENT "Abort"; id = identref -> VernacAbort (Some id)
| "Qed" -> VernacEndProof (true,None)
| IDENT "Save" -> VernacEndProof (true,None)
| IDENT "Defined" -> VernacEndProof (false,None)
- | IDENT "Defined"; id = ident -> VernacEndProof (false,Some (id,None))
- | IDENT "Save"; tok = thm_token; id = ident ->
+ | IDENT "Defined"; id=base_ident -> VernacEndProof (false,Some (id,None))
+ | IDENT "Save"; tok = thm_token; id = base_ident ->
VernacEndProof (true,Some (id,Some tok))
- | IDENT "Save"; id = ident -> VernacEndProof (true,Some (id,None))
+ | IDENT "Save"; id = base_ident -> VernacEndProof (true,Some (id,None))
| IDENT "Suspend" -> VernacSuspend
| IDENT "Resume" -> VernacResume None
- | IDENT "Resume"; id = ident -> VernacResume (Some id)
+ | IDENT "Resume"; id = identref -> VernacResume (Some id)
| IDENT "Restart" -> VernacRestart
| "Proof"; c = Constr.constr -> VernacExactProof c
| IDENT "Undo" -> VernacUndo 1
@@ -86,13 +85,13 @@ GEXTEND Gram
| IDENT "HintDestruct";
dloc = destruct_location;
- id = ident;
+ id = base_ident;
hyptyp = Constr.constr_pattern;
pri = natural;
"["; tac = tactic; "]" ->
VernacHintDestruct (id,dloc,hyptyp,pri,tac)
- | IDENT "Hint"; hintname = ident; dbnames = opt_hintbases; ":="; h = hint
+ | IDENT "Hint"; hintname = base_ident; dbnames = opt_hintbases; ":="; h = hint
-> VernacHints (dbnames, h hintname)
| IDENT "Hints"; (dbnames,h) = hints -> VernacHints (dbnames, h)
@@ -107,17 +106,17 @@ GEXTEND Gram
hint:
[ [ IDENT "Resolve"; c = Constr.constr -> fun name -> HintsResolve [Some name, c]
| IDENT "Immediate"; c = Constr.constr -> fun name -> HintsImmediate [Some name, c]
- | IDENT "Unfold"; qid = qualid -> fun name -> HintsUnfold [Some name,qid]
- | IDENT "Constructors"; c = qualid -> fun n -> HintsConstructors (n,c)
+ | IDENT "Unfold"; qid = global -> fun name -> HintsUnfold [Some name,qid]
+ | IDENT "Constructors"; c = global -> fun n -> HintsConstructors (n,c)
| IDENT "Extern"; n = natural; c = Constr.constr8 ; tac = tactic ->
fun name -> HintsExtern (name,n,c,tac) ] ]
;
hints:
- [ [ IDENT "Resolve"; l = LIST1 Constr.qualid; dbnames = opt_hintbases ->
- (dbnames, HintsResolve (List.map (fun qid -> (None, qid)) l))
- | IDENT "Immediate"; l = LIST1 Constr.qualid; dbnames = opt_hintbases ->
- (dbnames, HintsImmediate (List.map (fun qid -> (None, qid)) l))
- | IDENT "Unfold"; l = LIST1 qualid; dbnames = opt_hintbases ->
+ [ [ IDENT "Resolve"; l = LIST1 global; dbnames = opt_hintbases ->
+ (dbnames, HintsResolve (List.map (fun qid -> (None, CRef qid)) l))
+ | IDENT "Immediate"; l = LIST1 global; dbnames = opt_hintbases ->
+ (dbnames, HintsImmediate (List.map (fun qid -> (None, CRef qid)) l))
+ | IDENT "Unfold"; l = LIST1 global; dbnames = opt_hintbases ->
(dbnames, HintsUnfold (List.map (fun qid -> (None,qid)) l)) ] ]
;
END
diff --git a/parsing/g_rsyntax.ml b/parsing/g_rsyntax.ml
index 6c5829627..e39b8125c 100644
--- a/parsing/g_rsyntax.ml
+++ b/parsing/g_rsyntax.ml
@@ -13,24 +13,37 @@ open Util
open Names
open Pcoq
open Extend
+open Topconstr
+open Libnames
let get_r_sign loc =
- let ast_of_id id = Astterm.globalize_constr (Nvar(loc,id)) in
- ((ast_of_id (id_of_string "R0"),
- ast_of_id (id_of_string "R1"),
- ast_of_id (id_of_string "Rplus"),
- ast_of_id (id_of_string "NRplus")))
+ let mkid id =
+ mkRefC (Qualid (loc,Libnames.make_short_qualid id))
+ in
+ ((mkid (id_of_string "R0"),
+ mkid (id_of_string "R1"),
+ mkid (id_of_string "Rplus"),
+ mkid (id_of_string "NRplus")))
+
+let get_r_sign_ast loc =
+ let mkid id =
+ Termast.ast_of_ref (Nametab.locate (Libnames.make_short_qualid id))
+ in
+ ((mkid (id_of_string "R0"),
+ mkid (id_of_string "R1"),
+ mkid (id_of_string "Rplus"),
+ mkid (id_of_string "NRplus")))
(* Parsing via Grammar *)
let r_of_int n dloc =
- let (ast0,ast1,astp,_) = get_r_sign dloc in
+ let (a0,a1,plus,_) = get_r_sign dloc in
let rec mk_r n =
- if n <= 0 then
- ast0
+ if n <= 0 then
+ a0
else if n = 1 then
- ast1
+ a1
else
- Node(dloc,"APPLIST", [astp; ast1; mk_r (n-1)])
+ mkAppC (plus, [a1; mk_r (n-1)])
in
mk_r n
@@ -49,33 +62,33 @@ let _ =
exception Non_closed_number
-let rec int_of_r_rec ast1 astp p =
+let rec int_of_r_rec a1 plus p =
match p with
- | Node (_,"APPLIST", [b; a; c]) when alpha_eq(b,astp) &&
- alpha_eq(a,ast1) ->
- (int_of_r_rec ast1 astp c)+1
- | a when alpha_eq(a,ast1) -> 1
+ | Node (_,"APPLIST", [b; a; c]) when alpha_eq(b,plus) &&
+ alpha_eq(a,a1) ->
+ (int_of_r_rec a1 plus c)+1
+ | a when alpha_eq(a,a1) -> 1
| _ -> raise Non_closed_number
let int_of_r p =
- let (_,ast1,astp,_) = get_r_sign dummy_loc in
+ let (_,a1,plus,_) = get_r_sign_ast dummy_loc in
try
- Some (int_of_r_rec ast1 astp p)
+ Some (int_of_r_rec a1 plus p)
with
Non_closed_number -> None
let replace_plus p =
- let (_,ast1,_,astnr) = get_r_sign dummy_loc in
- ope ("REXPR",[ope("APPLIST", [astnr; ast1; p])])
+ let (_,a1,_,astnr) = get_r_sign_ast dummy_loc in
+ ope ("REXPR",[ope("APPLIST", [astnr; a1; p])])
let r_printer std_pr p =
- let (_,ast1,astp,_) = get_r_sign dummy_loc in
+ let (_,a1,plus,_) = get_r_sign dummy_loc in
match (int_of_r p) with
| Some i -> str (string_of_int (i+1))
| None -> std_pr (replace_plus p)
let r_printer_outside std_pr p =
- let (_,ast1,astp,_) = get_r_sign dummy_loc in
+ let (_,a1,plus,_) = get_r_sign dummy_loc in
match (int_of_r p) with
| Some i -> str "``" ++ str (string_of_int (i+1)) ++ str "``"
| None -> std_pr (replace_plus p)
@@ -144,7 +157,7 @@ let _ = Symbols.declare_numeral_interpreter "R_scope" (r_of_int,None)
exception Non_closed_number
let bignat_of_pos p =
- let (_,one,plus,_) = get_r_sign dummy_loc in
+ let (_,one,plus,_) = get_r_sign_ast dummy_loc in
let rec transl = function
| Node (_,"APPLIST",[p; o; a]) when alpha_eq(p,plus) & alpha_eq(o,one)
-> add_1(transl a)
diff --git a/parsing/g_tactic.ml4 b/parsing/g_tactic.ml4
index 286642375..341752f45 100644
--- a/parsing/g_tactic.ml4
+++ b/parsing/g_tactic.ml4
@@ -6,8 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(*i camlp4deps: "parsing/grammar.cma" i*)
-
(* $Id$ *)
open Pp
@@ -27,7 +25,7 @@ open Tactic
(* Functions overloaded by quotifier *)
let induction_arg_of_constr c =
- try ElimOnIdent (Ast.loc c,coerce_to_id c) with _ -> ElimOnConstr c
+ try ElimOnIdent (Topconstr.constr_loc c,coerce_to_id c) with _ -> ElimOnConstr c
let local_compute = [FBeta;FIota;FDeltaBut [];FZeta]
@@ -80,13 +78,13 @@ GEXTEND Gram
int_or_var:
[ [ n = integer -> Genarg.ArgArg n
- | id = ident -> Genarg.ArgVar (loc,id) ] ]
+ | id = identref -> Genarg.ArgVar id ] ]
;
autoarg_depth:
[ [ n = OPT natural -> n ] ]
;
autoarg_adding:
- [ [ IDENT "Adding" ; "["; l = LIST1 qualid; "]" -> l | -> [] ] ]
+ [ [ IDENT "Adding" ; "["; l = LIST1 global; "]" -> l | -> [] ] ]
;
autoarg_destructing:
[ [ IDENT "Destructing" -> true | -> false ] ]
@@ -100,17 +98,17 @@ GEXTEND Gram
;
(* Either an hypothesis or a ltac ref (variable or pattern metavariable) *)
id_or_ltac_ref:
- [ [ id = ident -> AN (loc,id)
+ [ [ id = base_ident -> AN id
| "?"; n = natural -> MetaNum (loc,n) ] ]
;
(* Either a global ref or a ltac ref (variable or pattern metavariable) *)
- qualid_or_ltac_ref:
- [ [ (loc,qid) = qualid -> AN (loc,qid)
+ global_or_ltac_ref:
+ [ [ qid = global -> AN qid
| "?"; n = natural -> MetaNum (loc,n) ] ]
;
(* An identifier or a quotation meta-variable *)
id_or_meta:
- [ [ id = rawident -> AI id
+ [ [ id = identref -> AI id
(* This is used in quotations *)
| id = METAIDENT -> MetaId (loc,id) ] ]
@@ -122,7 +120,7 @@ GEXTEND Gram
] ]
;
constrarg:
- [ [ IDENT "Inst"; id = rawident; "["; c = constr; "]" ->
+ [ [ IDENT "Inst"; id = identref; "["; c = constr; "]" ->
ConstrContext (id, c)
| IDENT "Eval"; rtc = Tactic.red_expr; "in"; c = constr ->
ConstrEval (rtc,c)
@@ -138,7 +136,7 @@ GEXTEND Gram
] ]
;
quantified_hypothesis:
- [ [ id = ident -> NamedHyp id
+ [ [ id = base_ident -> NamedHyp id
| n = natural -> AnonHyp n ] ]
;
pattern_occ:
@@ -161,11 +159,11 @@ GEXTEND Gram
[ [ "["; tc = LIST1 intropatterns SEP "|" ; "]" -> IntroOrAndPattern tc
| "("; tc = LIST1 simple_intropattern SEP "," ; ")" -> IntroOrAndPattern [tc]
| IDENT "_" -> IntroWildcard
- | id = ident -> IntroIdentifier id
+ | id = base_ident -> IntroIdentifier id
] ]
;
simple_binding:
- [ [ id = ident; ":="; c = constr -> (NamedHyp id, c)
+ [ [ id = base_ident; ":="; c = constr -> (NamedHyp id, c)
| n = natural; ":="; c = constr -> (AnonHyp n, c) ] ]
;
binding_list:
@@ -183,15 +181,15 @@ GEXTEND Gram
[ [ "with"; bl = binding_list -> bl | -> NoBindings ] ]
;
unfold_occ:
- [ [ nl = LIST0 integer; c = qualid_or_ltac_ref -> (nl,c) ] ]
+ [ [ nl = LIST0 integer; c = global_or_ltac_ref -> (nl,c) ] ]
;
red_flag:
[ [ IDENT "Beta" -> FBeta
| IDENT "Delta" -> FDeltaBut []
| IDENT "Iota" -> FIota
| IDENT "Zeta" -> FZeta
- | IDENT "Delta"; "["; idl = LIST1 qualid_or_ltac_ref; "]" -> FConst idl
- | IDENT "Delta"; "-"; "["; idl = LIST1 qualid_or_ltac_ref; "]" -> FDeltaBut idl
+ | IDENT "Delta"; "["; idl = LIST1 global_or_ltac_ref; "]" -> FConst idl
+ | IDENT "Delta"; "-"; "["; idl = LIST1 global_or_ltac_ref; "]" -> FDeltaBut idl
] ]
;
red_tactic:
@@ -227,10 +225,10 @@ GEXTEND Gram
| -> [] ] ]
;
fixdecl:
- [ [ id = ident; "/"; n = natural; ":"; c = constr -> (id,n,c) ] ]
+ [ [ id = base_ident; "/"; n = natural; ":"; c = constr -> (id,n,c) ] ]
;
cofixdecl:
- [ [ id = ident; ":"; c = constr -> (id,c) ] ]
+ [ [ id = base_ident; ":"; c = constr -> (id,c) ] ]
;
hintbases:
[ [ "with"; "*" -> None
@@ -241,7 +239,7 @@ GEXTEND Gram
[ [ "using"; el = constr_with_bindings -> el ] ]
;
with_names:
- [ [ "as"; "["; ids = LIST1 (LIST0 Prim.ident) SEP "|"; "]" -> ids
+ [ [ "as"; "["; ids = LIST1 (LIST0 base_ident) SEP "|"; "]" -> ids
| -> [] ] ]
;
simple_tactic:
@@ -250,11 +248,11 @@ GEXTEND Gram
IDENT "Intros"; IDENT "until"; id = quantified_hypothesis ->
TacIntrosUntil id
| IDENT "Intros"; pl = intropatterns -> TacIntroPattern pl
- | IDENT "Intro"; id = ident; IDENT "after"; id2 = rawident ->
+ | IDENT "Intro"; id = base_ident; IDENT "after"; id2 = identref ->
TacIntroMove (Some id, Some id2)
- | IDENT "Intro"; IDENT "after"; id2 = rawident ->
+ | IDENT "Intro"; IDENT "after"; id2 = identref ->
TacIntroMove (None, Some id2)
- | IDENT "Intro"; id = ident -> TacIntroMove (Some id, None)
+ | IDENT "Intro"; id = base_ident -> TacIntroMove (Some id, None)
| IDENT "Intro" -> TacIntroMove (None, None)
| IDENT "Assumption" -> TacAssumption
@@ -269,12 +267,12 @@ GEXTEND Gram
| IDENT "Case"; cl = constr_with_bindings -> TacCase cl
| IDENT "CaseType"; c = constr -> TacCaseType c
| IDENT "Fix"; n = natural -> TacFix (None,n)
- | IDENT "Fix"; id = ident; n = natural -> TacFix (Some id,n)
- | IDENT "Fix"; id = ident; n = natural; "with"; fd = LIST0 fixdecl ->
+ | IDENT "Fix"; id = base_ident; n = natural -> TacFix (Some id,n)
+ | IDENT "Fix"; id = base_ident; n = natural; "with"; fd = LIST0 fixdecl ->
TacMutualFix (id,n,fd)
| IDENT "Cofix" -> TacCofix None
- | IDENT "Cofix"; id = ident -> TacCofix (Some id)
- | IDENT "Cofix"; id = ident; "with"; fd = LIST0 cofixdecl ->
+ | IDENT "Cofix"; id = base_ident -> TacCofix (Some id)
+ | IDENT "Cofix"; id = base_ident; "with"; fd = LIST0 cofixdecl ->
TacMutualCofix (id,fd)
| IDENT "Cut"; c = constr -> TacCut c
@@ -288,7 +286,7 @@ GEXTEND Gram
| IDENT "Pose"; b = constr -> TacForward (true,Names.Anonymous,b)
| IDENT "Generalize"; lc = LIST1 constr -> TacGeneralize lc
| IDENT "Generalize"; IDENT "Dependent"; c = constr -> TacGeneralizeDep c
- | IDENT "LetTac"; id = ident; ":="; c = constr; p = clause_pattern
+ | IDENT "LetTac"; id = base_ident; ":="; c = constr; p = clause_pattern
-> TacLetTac (id,c,p)
| IDENT "Instantiate"; n = natural; c = constr -> TacInstantiate (n,c)
@@ -307,7 +305,7 @@ GEXTEND Gram
ids = with_names -> TacNewDestruct (c,el,ids)
| IDENT "Decompose"; IDENT "Record" ; c = constr -> TacDecomposeAnd c
| IDENT "Decompose"; IDENT "Sum"; c = constr -> TacDecomposeOr c
- | IDENT "Decompose"; "["; l = LIST1 qualid_or_ltac_ref; "]"; c = constr
+ | IDENT "Decompose"; "["; l = LIST1 global_or_ltac_ref; "]"; c = constr
-> TacDecompose (l,c)
(* Automation tactic *)
@@ -315,8 +313,8 @@ GEXTEND Gram
| IDENT "Auto"; n = OPT natural; db = hintbases -> TacAuto (n, db)
| IDENT "AutoTDB"; n = OPT natural -> TacAutoTDB n
- | IDENT "CDHyp"; id = rawident -> TacDestructHyp (true,id)
- | IDENT "DHyp"; id = rawident -> TacDestructHyp (false,id)
+ | IDENT "CDHyp"; id = identref -> TacDestructHyp (true,id)
+ | IDENT "DHyp"; id = identref -> TacDestructHyp (false,id)
| IDENT "DConcl" -> TacDestructConcl
| IDENT "SuperAuto"; l = autoargs -> TacSuperAuto l
| IDENT "Auto"; n = OPT natural; IDENT "Decomp"; p = OPT natural ->
@@ -325,9 +323,9 @@ GEXTEND Gram
(* Context management *)
| IDENT "Clear"; l = LIST1 id_or_ltac_ref -> TacClear l
| IDENT "ClearBody"; l = LIST1 id_or_ltac_ref -> TacClearBody l
- | IDENT "Move"; id1 = rawident; IDENT "after"; id2 = rawident ->
+ | IDENT "Move"; id1 = identref; IDENT "after"; id2 = identref ->
TacMove (true,id1,id2)
- | IDENT "Rename"; id1 = rawident; IDENT "into"; id2 = rawident ->
+ | IDENT "Rename"; id1 = identref; IDENT "into"; id2 = identref ->
TacRename (id1,id2)
(* Constructors *)
@@ -353,14 +351,6 @@ GEXTEND Gram
(* Unused ??
| IDENT "ML"; s = string -> ExtraTactic<:ast< (MLTACTIC $s) >>
*)
-
- (* | [ id = identarg; l = constrarg_list ->
- match (isMeta (nvar_of_ast id), l) with
- | (true, []) -> id
- | (false, _) -> <:ast< (CALL $id ($LIST $l)) >>
- | _ -> Util.user_err_loc
- (loc, "G_tactic.meta_tactic",
- (str"Cannot apply arguments to a meta-tactic."))
- ] *)] ]
+ ] ]
;
END;;
diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4
index 1a582b293..f347ac20e 100644
--- a/parsing/g_vernac.ml4
+++ b/parsing/g_vernac.ml4
@@ -6,11 +6,10 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(*i camlp4deps: "parsing/grammar.cma" i*)
-
(* $Id$ *)
-open Coqast
+open Names
+open Topconstr
open Vernacexpr
open Pcoq
open Pp
@@ -26,7 +25,7 @@ let join_binders (idl,c) = List.map (fun id -> (id,c)) idl
open Genarg
-let evar_constr loc = <:ast< (ISEVAR) >>
+let evar_constr loc = CHole loc
(* Rem: do not join the different GEXTEND into one, it breaks native *)
(* compilation on PowerPC and Sun architectures *)
@@ -40,10 +39,10 @@ GEXTEND Gram
| g = gallina_ext; "." -> g
| c = command; "." -> c
| c = syntax; "." -> c
- | n = Prim.natural; ":"; v = goal_vernac; "." -> v n
+ | n = natural; ":"; v = goal_vernac; "." -> v n
| "["; l = vernac_list_tail -> VernacList l
(* This is for "Grammar vernac" rules *)
- | id = Prim.metaident -> VernacVar (Ast.nvar_of_ast id) ] ]
+ | id = METAIDENT -> VernacVar (Names.id_of_string id) ] ]
;
goal_vernac:
[ [ tac = Tactic.tactic -> fun n -> VernacSolve (n,tac)
@@ -71,8 +70,8 @@ GEXTEND Gram
] ]
;
constr_body:
- [ [ ":="; c = constr; ":"; t = constr -> <:ast< (CAST $c $t) >>
- | ":"; t = constr; ":="; c = constr -> <:ast< (CAST $c $t) >>
+ [ [ ":="; c = constr; ":"; t = constr -> CCast(loc,c,t)
+ | ":"; t = constr; ":="; c = constr -> CCast(loc,c,t)
| ":="; c = constr -> c ] ]
;
vernac_list_tail:
@@ -123,31 +122,34 @@ GEXTEND Gram
| ":" -> false ] ]
;
params:
- [ [ idl = LIST1 ident SEP ","; coe = of_type_with_opt_coercion; c = constr
+ [ [ idl = LIST1 base_ident SEP ","; coe = of_type_with_opt_coercion; c = constr
-> List.map (fun c -> (coe,c)) (join_binders (idl,c))
] ]
;
ne_params_list:
[ [ ll = LIST1 params SEP ";" -> List.flatten ll ] ]
;
-ident_comma_list_tail:
- [ [ ","; idl = LIST1 ident SEP "," -> idl | -> [] ] ]
+ name_comma_list_tail:
+ [ [ ","; nal = LIST1 name SEP "," -> nal | -> [] ] ]
+ ;
+ ident_comma_list_tail:
+ [ [ ","; nal = LIST1 base_ident SEP "," -> nal | -> [] ] ]
;
type_option:
[ [ ":"; c = constr -> c
| -> evar_constr loc ] ]
;
opt_casted_constr:
- [ [ c = constr; ":"; t = constr -> <:ast< (CAST $c $t) >>
+ [ [ c = constr; ":"; t = constr -> CCast(loc,c,t)
| c = constr -> c ] ]
;
vardecls:
- [ [ id = ident; idl = ident_comma_list_tail; c = type_option ->
- LocalRawAssum (id::idl,c)
- | id = ident; "="; c = opt_casted_constr ->
- LocalRawDef (id,c)
- | id = ident; ":="; c = opt_casted_constr ->
- LocalRawDef (id,c)
+ [ [ na = name; nal = name_comma_list_tail; c = type_option
+ -> LocalRawAssum (na::nal,c)
+ | na = name; "="; c = opt_casted_constr ->
+ LocalRawDef (na,c)
+ | na = name; ":="; c = opt_casted_constr ->
+ LocalRawDef (na,c)
] ]
;
binders:
@@ -172,9 +174,9 @@ ident_comma_list_tail:
;
gallina:
(* Definition, Theorem, Variable, Axiom, ... *)
- [ [ thm = thm_token; id = ident; bl = binders_list; ":"; c = constr ->
+ [ [ thm = thm_token; id = base_ident; bl = binders_list; ":"; c = constr ->
VernacStartTheoremProof (thm, id, (bl, c), false, (fun _ _ -> ()))
- | (f,d) = def_token; id = ident; b = def_body ->
+ | (f,d) = def_token; id = base_ident; b = def_body ->
VernacDefinition (d, id, b, f)
| stre = assumption_token; bl = ne_params_list ->
VernacAssumption (stre, bl)
@@ -192,7 +194,7 @@ ident_comma_list_tail:
[ [ IDENT "Record" -> () | IDENT "Structure" -> () ] ]
;
constructor:
- [ [ id = ident; coe = of_type_with_opt_coercion; c = constr ->
+ [ [ id = base_ident; coe = of_type_with_opt_coercion; c = constr ->
(coe,(id,c)) ] ]
;
ne_constructor_list:
@@ -209,7 +211,7 @@ ident_comma_list_tail:
| ind = oneind_old_style -> [ind] ] ]
;
oneind_old_style:
- [ [ id = ident; ":"; c = constr; ":="; lc = constructor_list ->
+ [ [ id = base_ident; ":"; c = constr; ":="; lc = constructor_list ->
(id,c,lc) ] ]
;
block:
@@ -217,7 +219,7 @@ ident_comma_list_tail:
| ind = oneind -> [ind] ] ]
;
oneind:
- [ [ id = ident; indpar = indpar; ":"; c = constr; ":=";
+ [ [ id = base_ident; indpar = indpar; ":"; c = constr; ":=";
lc = constructor_list -> (id,indpar,c,lc) ] ]
;
indpar:
@@ -229,7 +231,7 @@ ident_comma_list_tail:
| -> false ] ]
;
onescheme:
- [ [ id = ident; ":="; dep = dep; ind = qualid; IDENT "Sort";
+ [ [ id = base_ident; ":="; dep = dep; ind = global; IDENT "Sort";
s = sort -> (id,dep,ind,s) ] ]
;
schemes:
@@ -240,34 +242,34 @@ ident_comma_list_tail:
| IDENT "Minimality"; IDENT "for" -> false ] ]
;
onerec:
- [ [ id = ident; idl = ne_simple_binders_list; ":"; c = constr;
+ [ [ id = base_ident; idl = ne_fix_binders; ":"; c = constr;
":="; def = constr -> (id,idl,c,def) ] ]
;
specifrec:
[ [ l = LIST1 onerec SEP "with" -> l ] ]
;
onecorec:
- [ [ id = ident; ":"; c = constr; ":="; def = constr ->
+ [ [ id = base_ident; ":"; c = constr; ":="; def = constr ->
(id,c,def) ] ]
;
specifcorec:
[ [ l = LIST1 onecorec SEP "with" -> l ] ]
;
record_field:
- [ [ id = ident; oc = of_type_with_opt_coercion; t = constr ->
+ [ [ id = base_ident; oc = of_type_with_opt_coercion; t = constr ->
(oc,AssumExpr (id,t))
- | id = ident; oc = of_type_with_opt_coercion; t = constr;
+ | id = base_ident; oc = of_type_with_opt_coercion; t = constr;
":="; b = constr ->
(oc,DefExpr (id,b,Some t))
- | id = ident; ":="; b = constr ->
+ | id = base_ident; ":="; b = constr ->
(false,DefExpr (id,b,None)) ] ]
;
fields:
[ [ fs = LIST0 record_field SEP ";" -> fs ] ]
;
simple_params:
- [ [ idl = LIST1 ident SEP ","; ":"; c = constr -> join_binders (idl, c)
- | idl = LIST1 ident SEP "," -> join_binders (idl, evar_constr dummy_loc)
+ [ [ idl = LIST1 base_ident SEP ","; ":"; c = constr -> join_binders (idl, c)
+ | idl = LIST1 base_ident SEP "," -> join_binders (idl, evar_constr dummy_loc)
] ]
;
simple_binders:
@@ -276,8 +278,19 @@ ident_comma_list_tail:
ne_simple_binders_list:
[ [ bll = LIST1 simple_binders -> List.flatten bll ] ]
;
+ fix_params:
+ [ [ idl = LIST1 name SEP ","; ":"; c = constr -> (idl, c)
+ | idl = LIST1 name SEP "," -> (idl, evar_constr dummy_loc)
+ ] ]
+ ;
+ fix_binders:
+ [ [ "["; bll = LIST1 fix_params SEP ";"; "]" -> bll ] ]
+ ;
+ ne_fix_binders:
+ [ [ bll = LIST1 fix_binders -> List.flatten bll ] ]
+ ;
rec_constructor:
- [ [ c = ident -> Some c
+ [ [ c = base_ident -> Some c
| -> None ] ]
;
gallina_ext:
@@ -285,7 +298,7 @@ ident_comma_list_tail:
indl = block_old_style ->
let indl' = List.map (fun (id,ar,c) -> (id,bl,ar,c)) indl in
VernacInductive (f,indl')
- | record_token; oc = opt_coercion; name = ident; ps = indpar; ":";
+ | record_token; oc = opt_coercion; name = base_ident; ps = indpar; ":";
s = sort; ":="; c = rec_constructor; "{"; fs = fields; "}" ->
VernacRecord ((oc,name),ps,s,c,fs)
] ]
@@ -296,25 +309,25 @@ ident_comma_list_tail:
| "Fixpoint"; recs = specifrec -> VernacFixpoint recs
| "CoFixpoint"; corecs = specifcorec -> VernacCoFixpoint corecs
| IDENT "Scheme"; l = schemes -> VernacScheme l
- | f = finite_token; s = sort; id = ident; indpar = indpar; ":=";
+ | f = finite_token; s = csort; id = base_ident; indpar = indpar; ":=";
lc = constructor_list ->
VernacInductive (f,[id,indpar,s,lc])
| f = finite_token; indl = block ->
VernacInductive (f,indl) ] ]
;
+ csort:
+ [ [ s = sort -> CSort (loc,s) ] ]
+ ;
gallina_ext:
[ [
(* Sections *)
- IDENT "Section"; id = ident -> VernacBeginSection id
- | IDENT "Chapter"; id = ident -> VernacBeginSection id ] ]
-(* | IDENT "Module"; id = ident ->
- warning "Module is currently unsupported"; VernacNop *)
+ IDENT "Section"; id = base_ident -> VernacBeginSection id
+ | IDENT "Chapter"; id = base_ident -> VernacBeginSection id ] ]
;
module_vardecls: (* This is interpreted by Pcoq.abstract_binder *)
- [ [ id = ident; idl = ident_comma_list_tail;
- ":"; mty = Module.module_type ->
- (id::idl,mty) ] ]
+ [ [ id = base_ident; idl = ident_comma_list_tail; ":"; mty = Module.module_type
+ -> (id::idl,mty) ] ]
;
module_binders:
[ [ "["; bl = LIST1 module_vardecls SEP ";"; "]" -> bl ] ]
@@ -334,64 +347,64 @@ ident_comma_list_tail:
gallina_ext:
[ [
(* Interactive module declaration *)
- IDENT "Module"; id = ident; bl = module_binders_list;
+ IDENT "Module"; id = base_ident; bl = module_binders_list;
mty_o = OPT of_module_type; mexpr_o = OPT is_module_expr ->
VernacDeclareModule (id, bl, mty_o, mexpr_o)
- | IDENT "Module"; "Type"; id = ident;
+ | IDENT "Module"; "Type"; id = base_ident;
bl = module_binders_list; mty_o = OPT is_module_type ->
VernacDeclareModuleType (id, bl, mty_o)
(* This end a Section a Module or a Module Type *)
- | IDENT "End"; id = ident -> VernacEndSegment id
+ | IDENT "End"; id = base_ident -> VernacEndSegment id
(* Transparent and Opaque *)
- | IDENT "Transparent"; l = LIST1 qualid -> VernacSetOpacity (false, l)
- | IDENT "Opaque"; l = LIST1 qualid -> VernacSetOpacity (true, l)
+ | IDENT "Transparent"; l = LIST1 global -> VernacSetOpacity (false, l)
+ | IDENT "Opaque"; l = LIST1 global -> VernacSetOpacity (true, l)
(* Canonical structure *)
- | IDENT "Canonical"; IDENT "Structure"; qid = qualid ->
+ | IDENT "Canonical"; IDENT "Structure"; qid = global ->
VernacCanonical qid
- | IDENT "Canonical"; IDENT "Structure"; qid = qualid; d = def_body ->
- let s = Ast.coerce_qualid_to_id qid in
+ | IDENT "Canonical"; IDENT "Structure"; qid = global; d = def_body ->
+ let s = Ast.coerce_global_to_id qid in
VernacDefinition (Global,s,d,Recordobj.add_object_hook)
(* Rem: LOBJECT, OBJCOERCION, LOBJCOERCION have been removed
(they were unused and undocumented) *)
(* Coercions *)
- | IDENT "Coercion"; qid = qualid; d = def_body ->
- let s = Ast.coerce_qualid_to_id qid in
+ | IDENT "Coercion"; qid = global; d = def_body ->
+ let s = Ast.coerce_global_to_id qid in
VernacDefinition (Global,s,d,Class.add_coercion_hook)
- | IDENT "Coercion"; IDENT "Local"; qid = qualid; d = def_body ->
- let s = Ast.coerce_qualid_to_id qid in
+ | IDENT "Coercion"; IDENT "Local"; qid = global; d = def_body ->
+ let s = Ast.coerce_global_to_id qid in
VernacDefinition (Local,s,d,Class.add_coercion_hook)
- | IDENT "Identity"; IDENT "Coercion"; IDENT "Local"; f = Prim.ident;
+ | IDENT "Identity"; IDENT "Coercion"; IDENT "Local"; f = base_ident;
":"; s = class_rawexpr; ">->"; t = class_rawexpr ->
VernacIdentityCoercion (Local, f, s, t)
- | IDENT "Identity"; IDENT "Coercion"; f = Prim.ident; ":";
+ | IDENT "Identity"; IDENT "Coercion"; f = base_ident; ":";
s = class_rawexpr; ">->"; t = class_rawexpr ->
VernacIdentityCoercion (Global, f, s, t)
- | IDENT "Coercion"; IDENT "Local"; qid = qualid; ":";
+ | IDENT "Coercion"; IDENT "Local"; qid = global; ":";
s = class_rawexpr; ">->"; t = class_rawexpr ->
VernacCoercion (Local, qid, s, t)
- | IDENT "Coercion"; qid = qualid; ":"; s = class_rawexpr; ">->";
+ | IDENT "Coercion"; qid = global; ":"; s = class_rawexpr; ">->";
t = class_rawexpr ->
VernacCoercion (Global, qid, s, t)
- | IDENT "Class"; IDENT "Local"; c = qualid ->
+ | IDENT "Class"; IDENT "Local"; c = global ->
Pp.warning "Class is obsolete"; VernacNop
- | IDENT "Class"; c = qualid ->
+ | IDENT "Class"; c = global ->
Pp.warning "Class is obsolete"; VernacNop
(* Implicit *)
- | IDENT "Syntactic"; "Definition"; id = ident; ":="; c = constr;
+ | IDENT "Syntactic"; "Definition"; id = base_ident; ":="; c = constr;
n = OPT [ "|"; n = natural -> n ] ->
VernacSyntacticDefinition (id,c,n)
- | IDENT "Implicits"; qid = qualid; "["; l = LIST0 natural; "]" ->
+ | IDENT "Implicits"; qid = global; "["; l = LIST0 natural; "]" ->
VernacDeclareImplicits (qid,Some l)
- | IDENT "Implicits"; qid = qualid -> VernacDeclareImplicits (qid,None)
+ | IDENT "Implicits"; qid = global -> VernacDeclareImplicits (qid,None)
(* For compatibility *)
| IDENT "Implicit"; IDENT "Arguments"; IDENT "On" ->
@@ -436,23 +449,17 @@ GEXTEND Gram
<:ast< (CompileFile ($STR $verbosely) ($STR $only_spec)
($STR $mname) ($STR $fname))>>
*)
- | IDENT "Read"; IDENT "Module"; qidl = LIST1 qualid ->
+ | IDENT "Read"; IDENT "Module"; qidl = LIST1 global ->
VernacRequire (None, None, qidl)
| IDENT "Require"; export = export_token; specif = specif_token;
- qidl = LIST1 qualid -> VernacRequire (Some export, specif, qidl)
+ qidl = LIST1 global -> VernacRequire (Some export, specif, qidl)
| IDENT "Require"; export = export_token; specif = specif_token;
- id = Prim.ident; filename = STRING ->
+ id = base_ident; filename = STRING ->
VernacRequireFrom (export, specif, id, filename)
-(*
- | IDENT "Write"; IDENT "Module"; id = identarg -> ExtraVernac
- <:ast< (WriteModule $id) >>
- | IDENT "Write"; IDENT "Module"; id = identarg; s = stringarg -> ExtraVernac
- <:ast< (WriteModule $id $s) >>
-*)
| IDENT "Declare"; IDENT "ML"; IDENT "Module"; l = LIST1 STRING ->
VernacDeclareMLModule l
- | IDENT "Import"; qidl = LIST1 qualid -> VernacImport (false,qidl)
- | IDENT "Export"; qidl = LIST1 qualid -> VernacImport (true,qidl)
+ | IDENT "Import"; qidl = LIST1 global -> VernacImport (false,qidl)
+ | IDENT "Export"; qidl = LIST1 global -> VernacImport (true,qidl)
]
]
;
@@ -471,10 +478,10 @@ GEXTEND Gram
| IDENT "Restore"; IDENT "State"; s = STRING -> VernacRestoreState s
(* Resetting *)
- | IDENT "Reset"; id = Prim.ident -> VernacResetName id
+ | IDENT "Reset"; id = identref -> VernacResetName id
| IDENT "Reset"; IDENT "Initial" -> VernacResetInitial
| IDENT "Back" -> VernacBack 1
- | IDENT "Back"; n = Prim.natural -> VernacBack n
+ | IDENT "Back"; n = natural -> VernacBack n
(* Tactic Debugger *)
| IDENT "Debug"; IDENT "On" -> VernacDebug true
diff --git a/parsing/g_zsyntax.ml b/parsing/g_zsyntax.ml
index 56ded0837..7b3c3e391 100644
--- a/parsing/g_zsyntax.ml
+++ b/parsing/g_zsyntax.ml
@@ -15,40 +15,55 @@ open Util
open Names
open Ast
open Extend
+open Topconstr
+open Libnames
let get_z_sign loc =
- let ast_of_id id = Astterm.globalize_constr (Nvar(loc,id)) in
- ((ast_of_id (id_of_string "xI"),
- ast_of_id (id_of_string "xO"),
- ast_of_id (id_of_string "xH")),
- (ast_of_id (id_of_string "ZERO"),
- ast_of_id (id_of_string "POS"),
- ast_of_id (id_of_string "NEG")))
+ let mkid id =
+ mkRefC (Qualid (loc,Libnames.make_short_qualid id))
+ in
+ ((mkid (id_of_string "xI"),
+ mkid (id_of_string "xO"),
+ mkid (id_of_string "xH")),
+ (mkid (id_of_string "ZERO"),
+ mkid (id_of_string "POS"),
+ mkid (id_of_string "NEG")))
open Bignat
-let pos_of_bignat astxI astxO astxH x =
+let pos_of_bignat xI xO xH x =
let rec pos_of x =
match div2_with_rest x with
- | (q, true) when is_nonzero q -> ope("APPLIST", [astxI; pos_of q])
- | (q, false) -> ope("APPLIST", [astxO; pos_of q])
- | (_, true) -> astxH
+ | (q, true) when is_nonzero q -> mkAppC (xI, [pos_of q])
+ | (q, false) -> mkAppC (xO, [pos_of q])
+ | (_, true) -> xH
in
pos_of x
let z_of_string pos_or_neg s dloc =
- let ((astxI,astxO,astxH),(astZERO,astPOS,astNEG)) = get_z_sign dloc in
+ let ((xI,xO,xH),(aZERO,aPOS,aNEG)) = get_z_sign dloc in
let v = Bignat.of_string s in
if is_nonzero v then
if pos_or_neg then
- ope("APPLIST",[astPOS; pos_of_bignat astxI astxO astxH v])
+ mkAppC (aPOS, [pos_of_bignat xI xO xH v])
else
- ope("APPLIST",[astNEG; pos_of_bignat astxI astxO astxH v])
+ mkAppC (aNEG, [pos_of_bignat xI xO xH v])
else
- astZERO
+ aZERO
exception Non_closed_number
+let get_z_sign_ast loc =
+ let ast_of_id id =
+ Termast.ast_of_ref (Nametab.locate (Libnames.make_short_qualid id))
+ in
+ ((ast_of_id (id_of_string "xI"),
+ ast_of_id (id_of_string "xO"),
+ ast_of_id (id_of_string "xH")),
+ (ast_of_id (id_of_string "ZERO"),
+ ast_of_id (id_of_string "POS"),
+ ast_of_id (id_of_string "NEG")))
+
let rec bignat_of_pos c1 c2 c3 p =
match p with
| Node (_,"APPLIST", [b; a]) when alpha_eq(b,c1) ->
@@ -58,9 +73,9 @@ let rec bignat_of_pos c1 c2 c3 p =
| a when alpha_eq(a,c3) -> Bignat.one
| _ -> raise Non_closed_number
-let bignat_option_of_pos astxI astxO astxH p =
+let bignat_option_of_pos xI xO xH p =
try
- Some (bignat_of_pos astxO astxI astxH p)
+ Some (bignat_of_pos xO xI xH p)
with Non_closed_number ->
None
@@ -68,8 +83,8 @@ let pr_pos a = hov 0 (str "POS" ++ brk (1,1) ++ a)
let pr_neg a = hov 0 (str "NEG" ++ brk (1,1) ++ a)
let inside_printer posneg std_pr p =
- let ((astxI,astxO,astxH),_) = get_z_sign dummy_loc in
- match (bignat_option_of_pos astxI astxO astxH p) with
+ let ((xI,xO,xH),_) = get_z_sign_ast dummy_loc in
+ match (bignat_option_of_pos xI xO xH p) with
| Some n ->
if posneg then
(str (Bignat.to_string n))
@@ -82,8 +97,8 @@ let inside_printer posneg std_pr p =
let outside_zero_printer std_pr p = str "`0`"
let outside_printer posneg std_pr p =
- let ((astxI,astxO,astxH),_) = get_z_sign dummy_loc in
- match (bignat_option_of_pos astxI astxO astxH p) with
+ let ((xI,xO,xH),_) = get_z_sign_ast dummy_loc in
+ match (bignat_option_of_pos xI xO xH p) with
| Some n ->
if posneg then
(str "`" ++ str (Bignat.to_string n) ++ str "`")
@@ -177,15 +192,19 @@ let z_of_int dloc z =
let _ = Symbols.declare_numeral_interpreter "Z_scope" (z_of_int,None)
(***********************************************************************)
+(* Printer for positive *)
+
+
+(***********************************************************************)
(* Printers *)
exception Non_closed_number
let bignat_of_pos p =
- let ((astxI,astxO,astxH),_) = get_z_sign dummy_loc in
- let c1 = astxO in
- let c2 = astxI in
- let c3 = astxH in
+ let ((xI,xO,xH),_) = get_z_sign_ast dummy_loc in
+ let c1 = xO in
+ let c2 = xI in
+ let c3 = xH in
let rec transl = function
| Node (_,"APPLIST",[b; a]) when alpha_eq(b,c1) -> mult_2(transl a)
| Node (_,"APPLIST",[b; a]) when alpha_eq(b,c2) -> add_1(mult_2(transl a))
diff --git a/parsing/g_zsyntax.mli b/parsing/g_zsyntax.mli
index afda96bd9..93b40191c 100644
--- a/parsing/g_zsyntax.mli
+++ b/parsing/g_zsyntax.mli
@@ -10,4 +10,7 @@
(* Nice syntax for integers. *)
-val z_of_string : bool -> string -> Coqast.loc -> Coqast.t
+open Util
+open Topconstr
+
+val z_of_string : bool -> string -> loc -> constr_expr
diff --git a/parsing/genarg.ml b/parsing/genarg.ml
deleted file mode 100644
index e0d3b8019..000000000
--- a/parsing/genarg.ml
+++ /dev/null
@@ -1,181 +0,0 @@
-(***********************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
-(* \VV/ *************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(***********************************************************************)
-
-(* $Id$ *)
-
-open Util
-open Names
-open Nametab
-open Rawterm
-
-type argument_type =
- (* Basic types *)
- | BoolArgType
- | IntArgType
- | IntOrVarArgType
- | StringArgType
- | PreIdentArgType
- | IdentArgType
- | QualidArgType
- (* Specific types *)
- | ConstrArgType
- | ConstrMayEvalArgType
- | QuantHypArgType
- | TacticArgType
- | CastedOpenConstrArgType
- | ConstrWithBindingsArgType
- | RedExprArgType
- | List0ArgType of argument_type
- | List1ArgType of argument_type
- | OptArgType of argument_type
- | PairArgType of argument_type * argument_type
- | ExtraArgType of string
-
-type 'a or_var = ArgArg of 'a | ArgVar of loc * identifier
-
-type constr_ast = Coqast.t
-
-(* Dynamics but tagged by a type expression *)
-
-type ('a,'b) generic_argument = argument_type * Obj.t
-
-let dyntab = ref ([] : string list)
-
-type ('a,'b,'c) abstract_argument_type = argument_type
-
-let create_arg s =
- if List.mem s !dyntab then
- anomaly ("Genarg.create: already declared generic argument " ^ s);
- dyntab := s :: !dyntab;
- let t = ExtraArgType s in
- (t,t)
-
-let exists_argtype s = List.mem s !dyntab
-
-type open_constr = Evd.evar_map * Term.constr
-type open_rawconstr = Coqast.t
-
-let rawwit_bool = BoolArgType
-let wit_bool = BoolArgType
-
-let rawwit_int = IntArgType
-let wit_int = IntArgType
-
-let rawwit_int_or_var = IntOrVarArgType
-let wit_int_or_var = IntOrVarArgType
-
-let rawwit_string = StringArgType
-let wit_string = StringArgType
-
-let rawwit_ident = IdentArgType
-let wit_ident = IdentArgType
-
-let rawwit_pre_ident = PreIdentArgType
-let wit_pre_ident = PreIdentArgType
-
-let rawwit_qualid = QualidArgType
-let wit_qualid = QualidArgType
-
-let rawwit_quant_hyp = QuantHypArgType
-let wit_quant_hyp = QuantHypArgType
-
-let rawwit_constr = ConstrArgType
-let wit_constr = ConstrArgType
-
-let rawwit_constr_may_eval = ConstrMayEvalArgType
-let wit_constr_may_eval = ConstrMayEvalArgType
-
-let rawwit_tactic = TacticArgType
-let wit_tactic = TacticArgType
-
-let rawwit_casted_open_constr = CastedOpenConstrArgType
-let wit_casted_open_constr = CastedOpenConstrArgType
-
-let rawwit_constr_with_bindings = ConstrWithBindingsArgType
-let wit_constr_with_bindings = ConstrWithBindingsArgType
-
-let rawwit_red_expr = RedExprArgType
-let wit_red_expr = RedExprArgType
-
-let wit_list0 t = List0ArgType t
-
-let wit_list1 t = List1ArgType t
-
-let wit_opt t = OptArgType t
-
-let wit_pair t1 t2 = PairArgType (t1,t2)
-
-let in_gen t o = (t,Obj.repr o)
-let out_gen t (t',o) = if t = t' then Obj.magic o else failwith "out_gen"
-let genarg_tag (s,_) = s
-
-let fold_list0 f = function
- | (List0ArgType t as u, l) ->
- List.fold_right (fun x -> f (in_gen t x)) (Obj.magic l)
- | _ -> failwith "Genarg: not a list0"
-
-let fold_list1 f = function
- | (List1ArgType t as u, l) ->
- List.fold_right (fun x -> f (in_gen t x)) (Obj.magic l)
- | _ -> failwith "Genarg: not a list1"
-
-let fold_opt f a = function
- | (OptArgType t as u, l) ->
- (match Obj.magic l with
- | None -> a
- | Some x -> f (in_gen t x))
- | _ -> failwith "Genarg: not a opt"
-
-let fold_pair f = function
- | (PairArgType (t1,t2) as u, l) ->
- let (x1,x2) = Obj.magic l in
- f (in_gen t1 x1) (in_gen t2 x2)
- | _ -> failwith "Genarg: not a pair"
-
-let app_list0 f = function
- | (List0ArgType t as u, l) ->
- let o = Obj.magic l in
- (u, Obj.repr (List.map (fun x -> out_gen t (f (in_gen t x))) o))
- | _ -> failwith "Genarg: not a list0"
-
-let app_list1 f = function
- | (List1ArgType t as u, l) ->
- let o = Obj.magic l in
- (u, Obj.repr (List.map (fun x -> out_gen t (f (in_gen t x))) o))
- | _ -> failwith "Genarg: not a list1"
-
-let app_opt f = function
- | (OptArgType t as u, l) ->
- let o = Obj.magic l in
- (u, Obj.repr (option_app (fun x -> out_gen t (f (in_gen t x))) o))
- | _ -> failwith "Genarg: not an opt"
-
-let app_pair f1 f2 = function
- | (PairArgType (t1,t2) as u, l) ->
- let (o1,o2) = Obj.magic l in
- let o1 = out_gen t1 (f1 (in_gen t1 o1)) in
- let o2 = out_gen t2 (f2 (in_gen t2 o2)) in
- (u, Obj.repr (o1,o2))
- | _ -> failwith "Genarg: not a pair"
-
-let or_var_app f = function
- | ArgArg x -> ArgArg (f x)
- | ArgVar _ as x -> x
-
-let smash_var_app t f g = function
- | ArgArg x -> f x
- | ArgVar (_,id) ->
- let (u, _ as x) = g id in
- if t <> u then failwith "Genarg: a variable bound to a wrong type";
- x
-
-let unquote x = x
-
-type an_arg_of_this_type = Obj.t
-
-let in_generic t x = (t, Obj.repr x)
diff --git a/parsing/genarg.mli b/parsing/genarg.mli
deleted file mode 100644
index 2991d237a..000000000
--- a/parsing/genarg.mli
+++ /dev/null
@@ -1,208 +0,0 @@
-(***********************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
-(* \VV/ *************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(***********************************************************************)
-
-(* $Id$ *)
-
-open Util
-open Names
-open Term
-open Libnames
-open Rawterm
-
-type 'a or_var = ArgArg of 'a | ArgVar of loc * identifier
-
-type constr_ast = Coqast.t
-
-type open_constr = Evd.evar_map * Term.constr
-type open_rawconstr = constr_ast
-
-
-(* The route of a generic argument, from parsing to evaluation
-
- parsing in_raw out_raw
- char stream ----> rawtype ----> rawconstr generic_argument ---->
- |
- | interp
- V
- type <---- constr generic_argument <----
- out in
-
-To distinguish between the uninterpreted (raw) and the interpreted
-worlds, we annotate the type generic_argument by a phantom argument
-which is either constr_ast or constr (actually we add also a second
-argument raw_tactic_expr and tactic, but this is only for technical
-reasons, because these types are undefined at the type of compilation
-of Genarg).
-
-Transformation for each type :
-tag f raw open type cooked closed type
-
-BoolArgType bool bool
-IntArgType int int
-IntOrVarArgType int or_var int
-StringArgType string (parsed w/ "") string
-IdentArgType identifier identifier
-PreIdentArgType string (parsed w/o "") string
-QualidArgType qualid located global_reference
-ConstrArgType constr_ast constr
-ConstrMayEvalArgType constr_ast may_eval constr
-QuantHypArgType quantified_hypothesis quantified_hypothesis
-TacticArgType raw_tactic_expr tactic
-CastedOpenConstrArgType constr_ast open_constr
-ConstrWithBindingsArgType constr_ast with_bindings constr with_bindings
-List0ArgType of argument_type
-List1ArgType of argument_type
-OptArgType of argument_type
-ExtraArgType of string '_a '_b
-*)
-
-type ('a,'co,'ta) abstract_argument_type
-
-val rawwit_bool : (bool,'co,'ta) abstract_argument_type
-val wit_bool : (bool,'co,'ta) abstract_argument_type
-
-val rawwit_int : (int,'co,'ta) abstract_argument_type
-val wit_int : (int,'co,'ta) abstract_argument_type
-
-val rawwit_int_or_var : (int or_var,'co,'ta) abstract_argument_type
-val wit_int_or_var : (int or_var,'co,'ta) abstract_argument_type
-
-val rawwit_string : (string,'co,'ta) abstract_argument_type
-val wit_string : (string,'co,'ta) abstract_argument_type
-
-val rawwit_ident : (identifier,'co,'ta) abstract_argument_type
-val wit_ident : (identifier,'co,'ta) abstract_argument_type
-
-val rawwit_pre_ident : (string,'co,'ta) abstract_argument_type
-val wit_pre_ident : (string,'co,'ta) abstract_argument_type
-
-val rawwit_qualid : (qualid located,constr_ast,'ta) abstract_argument_type
-val wit_qualid : (global_reference,constr,'ta) abstract_argument_type
-
-val rawwit_quant_hyp : (quantified_hypothesis,'co,'ta) abstract_argument_type
-val wit_quant_hyp : (quantified_hypothesis,'co,'ta) abstract_argument_type
-
-val rawwit_constr : (constr_ast,constr_ast,'ta) abstract_argument_type
-val wit_constr : (constr,constr,'ta) abstract_argument_type
-
-val rawwit_constr_may_eval : (constr_ast may_eval,constr_ast,'ta) abstract_argument_type
-val wit_constr_may_eval : (constr,constr,'ta) abstract_argument_type
-
-val rawwit_casted_open_constr : (open_rawconstr,constr_ast,'ta) abstract_argument_type
-val wit_casted_open_constr : (open_constr,constr,'ta) abstract_argument_type
-
-val rawwit_constr_with_bindings : (constr_ast with_bindings,constr_ast,'ta) abstract_argument_type
-val wit_constr_with_bindings : (constr with_bindings,constr,'ta) abstract_argument_type
-
-val rawwit_red_expr : ((constr_ast,qualid or_metanum) red_expr_gen,constr_ast,'ta) abstract_argument_type
-val wit_red_expr : ((constr,Closure.evaluable_global_reference) red_expr_gen,constr,'ta) abstract_argument_type
-
-(* TODO: transformer tactic en extra arg *)
-val rawwit_tactic : ('ta,constr_ast,'ta) abstract_argument_type
-val wit_tactic : ('ta,constr,'ta) abstract_argument_type
-
-val wit_list0 :
- ('a,'co,'ta) abstract_argument_type -> ('a list,'co,'ta) abstract_argument_type
-
-val wit_list1 :
- ('a,'co,'ta) abstract_argument_type -> ('a list,'co,'ta) abstract_argument_type
-
-val wit_opt :
- ('a,'co,'ta) abstract_argument_type -> ('a option,'co,'ta) abstract_argument_type
-
-val wit_pair :
- ('a,'co,'ta) abstract_argument_type ->
- ('b,'co,'ta) abstract_argument_type ->
- ('a * 'b,'co,'ta) abstract_argument_type
-
-(* 'a generic_argument = (Sigma t:type. t[constr/'a]) *)
-type ('a,'b) generic_argument
-
-val fold_list0 :
- (('a,'b) generic_argument -> 'c -> 'c) -> ('a,'b) generic_argument -> 'c -> 'c
-
-val fold_list1 :
- (('a,'b) generic_argument -> 'c -> 'c) -> ('a,'b) generic_argument -> 'c -> 'c
-
-val fold_opt :
- (('a,'b) generic_argument -> 'c) -> 'c -> ('a,'b) generic_argument -> 'c
-
-val fold_pair :
- (('a,'b) generic_argument -> ('a,'b) generic_argument -> 'c) ->
- ('a,'b) generic_argument -> 'c
-
-(* [app_list0] fails if applied to an argument not of tag [List0 t]
- for some [t]; it's the responsability of the caller to ensure it *)
-
-val app_list0 : (('a,'b) generic_argument -> ('c,'d) generic_argument) ->
-('a,'b) generic_argument -> ('c,'d) generic_argument
-
-val app_list1 : (('a,'b) generic_argument -> ('c,'d) generic_argument) ->
-('a,'b) generic_argument -> ('c,'d) generic_argument
-
-val app_opt : (('a,'b) generic_argument -> ('c,'d) generic_argument) ->
-('a,'b) generic_argument -> ('c,'d) generic_argument
-
-val app_pair :
- (('a,'b) generic_argument -> ('c,'d) generic_argument) ->
- (('a,'b) generic_argument -> ('c,'d) generic_argument)
- -> ('a,'b) generic_argument -> ('c,'d) generic_argument
-
-(* Manque l'ordre supérieur, on aimerait ('co,'ta) 'a; manque aussi le
- polymorphism, on aimerait que 'b et 'c restent polymorphes à l'appel
- de create *)
-val create_arg : string ->
- ('rawa,'rawco,'rawta) abstract_argument_type
- * ('a,'co,'ta) abstract_argument_type
-
-val exists_argtype : string -> bool
-
-type argument_type =
- | BoolArgType
- | IntArgType
- | IntOrVarArgType
- | StringArgType
- | PreIdentArgType
- | IdentArgType
- | QualidArgType
- | ConstrArgType
- | ConstrMayEvalArgType
- | QuantHypArgType
- | TacticArgType
- | CastedOpenConstrArgType
- | ConstrWithBindingsArgType
- | RedExprArgType
- | List0ArgType of argument_type
- | List1ArgType of argument_type
- | OptArgType of argument_type
- | PairArgType of argument_type * argument_type
- | ExtraArgType of string
-
-val genarg_tag : ('a,'b) generic_argument -> argument_type
-
-val unquote : ('a,'co,'ta) abstract_argument_type -> argument_type
-
-(* We'd like
-
- [in_generic: !b:type, !a:argument_type -> (f a) -> b generic_argument]
-
- with f a = b if a is Constr, f a = c if a is Tactic, otherwise f a = |a|
-
- in_generic is not typable; we replace the second argument by an absurd
- type (with no introduction rule)
-*)
-type an_arg_of_this_type
-
-val in_generic :
- argument_type -> an_arg_of_this_type -> ('a,'b) generic_argument
-
-val in_gen :
- ('a,'co,'ta) abstract_argument_type -> 'a -> ('co,'ta) generic_argument
-val out_gen :
- ('a,'co,'ta) abstract_argument_type -> ('co,'ta) generic_argument -> 'a
-
diff --git a/parsing/pcoq.ml4 b/parsing/pcoq.ml4
index 67322863a..9c206565e 100644
--- a/parsing/pcoq.ml4
+++ b/parsing/pcoq.ml4
@@ -10,9 +10,14 @@
open Pp
open Util
+open Names
+open Libnames
+open Rawterm
+open Topconstr
open Ast
open Genarg
open Tacexpr
+open Extend
(* The lexer of Coq *)
@@ -46,59 +51,39 @@ let grammar_delete e rls =
List.iter (fun (pil,_) -> G.delete_rule e pil) (List.rev lev))
(List.rev rls)
+(* grammar_object is the superclass of all grammar entry *)
module type Gramobj =
sig
type grammar_object
- type 'a entry
-
- val in_entry : 'a -> 'b G.Entry.e -> 'a entry
- val object_of_entry : 'a entry -> grammar_object G.Entry.e
- val type_of_entry : 'a entry -> 'a
+ val weaken_entry : 'a G.Entry.e -> grammar_object G.Entry.e
end
module Gramobj : Gramobj =
struct
type grammar_object = Obj.t
- type 'a entry = 'a * grammar_object G.Entry.e
-
- let in_entry t e = (t,Obj.magic e)
- let object_of_entry (t,e) = e
- let type_of_entry (t,e) = t
+ let weaken_entry e = Obj.magic e
end
type grammar_object = Gramobj.grammar_object
-let in_typed_entry = Gramobj.in_entry
-let type_of_typed_entry = Gramobj.type_of_entry
-let object_of_typed_entry = Gramobj.object_of_entry
-type typed_entry = entry_type Gramobj.entry
+type typed_entry = entry_type * grammar_object G.Entry.e
+let in_typed_entry t e = (t,Gramobj.weaken_entry e)
+let type_of_typed_entry (t,e) = t
+let object_of_typed_entry (t,e) = e
module type Gramtypes =
sig
open Decl_kinds
- val inAstListType : Coqast.t list G.Entry.e -> typed_entry
- val inTacticAtomAstType : raw_atomic_tactic_expr G.Entry.e -> typed_entry
- val inThmTokenAstType : theorem_kind G.Entry.e -> typed_entry
- val inDynamicAstType : typed_ast G.Entry.e -> typed_entry
- val inReferenceAstType : Coqast.reference_expr G.Entry.e -> typed_entry
- val inPureAstType : constr_ast G.Entry.e -> typed_entry
- val inGenAstType : 'a raw_abstract_argument_type ->
- 'a G.Entry.e -> typed_entry
- val outGenAstType : 'a raw_abstract_argument_type -> typed_entry -> 'a G.Entry.e
+ val inGramObj : 'a raw_abstract_argument_type -> 'a G.Entry.e -> typed_entry
+ val outGramObj : 'a raw_abstract_argument_type -> typed_entry -> 'a G.Entry.e
end
module Gramtypes : Gramtypes =
struct
- let inAstListType = in_typed_entry AstListType
- let inTacticAtomAstType = in_typed_entry TacticAtomAstType
- let inThmTokenAstType = in_typed_entry ThmTokenAstType
- let inDynamicAstType = in_typed_entry DynamicAstType
- let inReferenceAstType = in_typed_entry ReferenceAstType
- let inPureAstType = in_typed_entry (GenAstType ConstrArgType)
- let inGenAstType rawwit = in_typed_entry (GenAstType (unquote rawwit))
-
- let outGenAstType (a:'a raw_abstract_argument_type) o =
- if type_of_typed_entry o <> GenAstType (unquote a)
- then anomaly "outGenAstType: wrong type";
+ let inGramObj rawwit = in_typed_entry (unquote rawwit)
+ let outGramObj (a:'a raw_abstract_argument_type) o =
+ if type_of_typed_entry o <> unquote a
+ then anomaly "outGramObj: wrong type";
+ (* downcast from grammar_object *)
Obj.magic (object_of_typed_entry o)
end
@@ -106,7 +91,7 @@ open Gramtypes
type ext_kind =
| ByGrammar of
- typed_entry * Gramext.position option *
+ grammar_object G.Entry.e * Gramext.position option *
(string option * Gramext.g_assoc option *
(Token.t Gramext.g_symbol list * Gramext.g_action) list) list
| ByGEXTEND of (unit -> unit) * (unit -> unit)
@@ -138,22 +123,20 @@ module Gram =
(* This extension command is used by the Grammar constr *)
let grammar_extend te pos rls =
- camlp4_state := ByGrammar (te,pos,rls) :: !camlp4_state;
+ camlp4_state := ByGrammar (Gramobj.weaken_entry te,pos,rls) :: !camlp4_state;
let a = !Gramext.warning_verbose in
Gramext.warning_verbose := Options.is_verbose ();
- G.extend (object_of_typed_entry te) pos rls;
+ G.extend te pos rls;
Gramext.warning_verbose := a
(* n is the number of extended entries (not the number of Grammar commands!)
to remove. *)
-let remove_grammar rls te = grammar_delete (object_of_typed_entry te) rls
-
let rec remove_grammars n =
if n>0 then
(match !camlp4_state with
| [] -> anomaly "Pcoq.remove_grammars: too many rules to remove"
| ByGrammar(g,_,rls)::t ->
- remove_grammar rls g;
+ grammar_delete g rls;
camlp4_state := t;
remove_grammars (n-1)
| ByGEXTEND (undo,redo)::t ->
@@ -187,14 +170,6 @@ let parse_string f x =
let strm = Stream.of_string x in Gram.Entry.parse f (Gram.parsable strm)
(*
-let slam_ast (_,fin) id ast =
- match id with
- | Coqast.Nvar (loc, s) -> Coqast.Slam (loc, Some s, ast)
- | Coqast.Nmeta (loc, s) -> Coqast.Smetalam (loc, s, ast)
- | _ -> invalid_arg "Ast.slam_ast"
-*)
-
-(*
let entry_type ast =
match ast with
| Coqast.Id (_, "LIST") -> ETastl
@@ -216,7 +191,7 @@ let trace = ref false
(* The univ_tab is not part of the state. It contains all the grammar that
exist or have existed before in the session. *)
-let univ_tab = Hashtbl.create 7
+let univ_tab = (Hashtbl.create 7 : (string, string * gram_universe) Hashtbl.t)
let create_univ s =
let u = s, Hashtbl.create 29 in Hashtbl.add univ_tab s u; u
@@ -283,22 +258,22 @@ let create_entry (u, utab) s etyp =
new_entry etyp (u, utab) s
let create_constr_entry u s =
- outGenAstType rawwit_constr (create_entry u s (GenAstType ConstrArgType))
+ outGramObj rawwit_constr (create_entry u s ConstrArgType)
let create_generic_entry s wit =
let (u,utab) = utactic in
let etyp = unquote wit in
try
let e = Hashtbl.find utab s in
- if type_of_typed_entry e <> GenAstType etyp then
+ if type_of_typed_entry e <> etyp then
failwith ("Entry " ^ u ^ ":" ^ s ^ " already exists with another type");
- outGenAstType wit e
+ outGramObj wit e
with Not_found ->
if !trace then begin
Printf.eprintf "[Creating entry %s:%s]\n" u s; flush stderr; ()
end;
let e = Gram.Entry.create s in
- Hashtbl.add utab s (inGenAstType wit e); e
+ Hashtbl.add utab s (inGramObj wit e); e
let get_generic_entry s =
let (u,utab) = utactic in
@@ -308,10 +283,7 @@ let get_generic_entry s =
error ("unknown grammar entry "^u^":"^s)
let get_generic_entry_type (u,utab) s =
- try
- match type_of_typed_entry (Hashtbl.find utab s) with
- | GenAstType t -> t
- | _ -> error "Not a generic type"
+ try type_of_typed_entry (Hashtbl.find utab s)
with Not_found ->
error ("unknown grammar entry "^u^":"^s)
@@ -319,8 +291,6 @@ let force_entry_type (u, utab) s etyp =
try
let entry = Hashtbl.find utab s in
let extyp = type_of_typed_entry entry in
- if etyp = PureAstType && extyp = GenAstType ConstrArgType then
- entry else
if etyp = extyp then
entry
else begin
@@ -333,45 +303,55 @@ let force_entry_type (u, utab) s etyp =
with Not_found ->
new_entry etyp (u, utab) s
-(* Grammar entries *)
+(* [make_gen_entry] builds entries extensible by giving its name (a string) *)
+(* For entries extensible only via the ML name, Gram.Entry.create is enough *)
-let make_entry (u,univ) in_fun s =
+let make_gen_entry (u,univ) rawwit s =
let e = Gram.Entry.create (u ^ ":" ^ s) in
- Hashtbl.add univ s (in_fun e); e
+ Hashtbl.add univ s (inGramObj rawwit e); e
-let make_gen_entry u rawwit = make_entry u (inGenAstType rawwit)
+(* Grammar entries *)
module Prim =
struct
let gec_gen x = make_gen_entry uprim x
- let gec = make_entry uprim inPureAstType
- let gec_list = make_entry uprim inAstListType
+ (* Entries that can be refered via the string -> Gram.Entry.e table *)
+ (* Typically for tactic or vernac extensions *)
let preident = gec_gen rawwit_pre_ident "preident"
let ident = gec_gen rawwit_ident "ident"
- let rawident = Gram.Entry.create "Prim.rawident"
let natural = gec_gen rawwit_int "natural"
let integer = gec_gen rawwit_int "integer"
let string = gec_gen rawwit_string "string"
- let qualid = gec_gen rawwit_qualid "qualid"
- let reference = make_entry uprim inReferenceAstType "reference"
+ let reference = make_gen_entry uprim rawwit_ref "reference"
+
+ (* A synonym of ident, for compatibility *)
+ let var = gec_gen rawwit_ident "var"
+
+ let name = Gram.Entry.create "Prim.name"
+ let identref = Gram.Entry.create "Prim.identref"
+
+ (* A synonym of ident - maybe ident will be located one day *)
+ let base_ident = Gram.Entry.create "Prim.base_ident"
+
+ let qualid = Gram.Entry.create "Prim.qualid"
let dirpath = Gram.Entry.create "Prim.dirpath"
- let astpat = make_entry uprim inDynamicAstType "astpat"
- let ast = gec "ast"
- let astlist = gec_list "astlist"
+
+ (* For old ast printer *)
+ let astpat = Gram.Entry.create "Prim.astpat"
+ let ast = Gram.Entry.create "Prim.ast"
+ let astlist = Gram.Entry.create "Prim.astlist"
let ast_eoi = eoi_entry ast
- let astact = gec "astact"
- let metaident = gec "metaident"
- let var = gec "var"
+ let astact = Gram.Entry.create "Prim.astact"
end
module Constr =
struct
- let gec = make_entry uconstr inPureAstType
let gec_constr = make_gen_entry uconstr rawwit_constr
- let gec_list = make_entry uconstr inAstListType
+ let gec_constr_list = make_gen_entry uconstr (wit_list0 rawwit_constr)
+ (* Entries that can be refered via the string -> Gram.Entry.e table *)
let constr = gec_constr "constr"
let constr0 = gec_constr "constr0"
let constr1 = gec_constr "constr1"
@@ -387,35 +367,30 @@ module Constr =
let constr10 = gec_constr "constr10"
let constr_eoi = eoi_entry constr
let lconstr = gec_constr "lconstr"
- let ident = gec "ident"
- let qualid = gec "qualid"
- let global = gec "global"
- let ne_ident_comma_list = gec_list "ne_ident_comma_list"
- let ne_constr_list = gec_list "ne_constr_list"
- let sort = gec_constr "sort"
- let pattern = gec "pattern"
- let constr_pattern = gec "constr_pattern"
- let ne_binders_list = gec_list "ne_binders_list"
- let numarg = gec "numarg"
- end
+ let sort = make_gen_entry uconstr rawwit_sort "sort"
+ let ident = make_gen_entry uconstr rawwit_ident "ident"
+ let global = make_gen_entry uconstr rawwit_ref "global"
+
+ let ne_name_comma_list = Gram.Entry.create "constr:ne_name_comma_list"
+ let ne_constr_list = gec_constr_list "ne_constr_list"
+ let pattern = Gram.Entry.create "constr:pattern"
+ let constr_pattern = gec_constr "constr_pattern"
+ end
module Module =
struct
- let gec = make_entry umodule inPureAstType
- let gec_list = make_entry umodule inAstListType
-
- let ident = gec "ident"
- let qualid = gec_list "qualid"
- let ne_binders_list = gec_list "ne_binders_list"
- let module_expr = gec "module_expr"
- let module_type = gec "module_type"
+ let module_expr = Gram.Entry.create "module_expr"
+ let module_type = Gram.Entry.create "module_type"
end
module Tactic =
struct
- let gec = make_entry utactic inPureAstType
- let gec_list = make_entry utactic inAstListType
+ (* Main entry for extensions *)
+ let simple_tactic = Gram.Entry.create "tactic:simple_tactic"
+
+ (* Entries that can be refered via the string -> Gram.Entry.e table *)
+ (* Typically for tactic user extensions *)
let castedopenconstr =
make_gen_entry utactic rawwit_casted_open_constr "castedopenconstr"
let constr_with_bindings =
@@ -425,23 +400,31 @@ module Tactic =
make_gen_entry utactic rawwit_quant_hyp "quantified_hypothesis"
let int_or_var = make_gen_entry utactic rawwit_int_or_var "int_or_var"
let red_expr = make_gen_entry utactic rawwit_red_expr "red_expr"
- let simple_tactic = make_entry utactic inTacticAtomAstType "simple_tactic"
+
+ (* Main entries for ltac *)
let tactic_arg = Gram.Entry.create "tactic:tactic_arg"
let tactic = make_gen_entry utactic rawwit_tactic "tactic"
+
+ (* Main entry for quotations *)
let tactic_eoi = eoi_entry tactic
end
module Vernac_ =
struct
- let thm_token = make_entry uvernac inThmTokenAstType "thm_token"
- let class_rawexpr = Gram.Entry.create "vernac:class_rawexpr"
let gec_vernac s = Gram.Entry.create ("vernac:" ^ s)
+
+ (* The different kinds of vernacular commands *)
let gallina = gec_vernac "gallina"
let gallina_ext = gec_vernac "gallina_ext"
let command = gec_vernac "command"
let syntax = gec_vernac "syntax_command"
let vernac = gec_vernac "Vernac_.vernac"
+
+ (* Various vernacular entries needed to be exported *)
+ let thm_token = Gram.Entry.create "vernac:thm_token"
+ let class_rawexpr = Gram.Entry.create "vernac:class_rawexpr"
+
let vernac_eoi = eoi_entry vernac
end
@@ -462,8 +445,12 @@ open Tactic
open Vernac_
(* current file and toplevel/vernac.ml *)
+let globalizer = ref (fun x -> failwith "No globalizer")
+let set_globalizer f = globalizer := f
-let define_quotation default s e =
+let f = (ast : Coqast.t G.Entry.e)
+
+let define_ast_quotation default s (e:Coqast.t G.Entry.e) =
(if default then
GEXTEND Gram
ast: [ [ "<<"; c = e; ">>" -> c ] ];
@@ -487,31 +474,16 @@ let define_quotation default s e =
*)
END)
-let _ = define_quotation false "ast" ast in ()
-
-let gecdyn s =
- let e = Gram.Entry.create ("Dyn." ^ s) in
- Hashtbl.add (snd uconstr) s (inDynamicAstType e); e
+(*
+let _ = define_ast_quotation false "ast" ast in ()
+*)
-let dynconstr = gecdyn "dynconstr"
-let dyncasespattern = gecdyn "dyncasespattern"
-let dyntactic = gecdyn "dyntactic"
-let dynvernac = gecdyn "dynvernac"
-let dynastlist = gecdyn "dynastlist"
-let dynast = gecdyn "dynast"
-
-let globalizer = ref (fun x -> x)
-let set_globalizer f = globalizer := f
+let dynconstr = Gram.Entry.create "Constr.dynconstr"
+let dyncasespattern = Gram.Entry.create "Constr.dyncasespattern"
GEXTEND Gram
- dynconstr: [ [ a = Constr.constr -> !globalizer (PureAstNode a) ] ];
- dyncasespattern: [ [ a = Constr.pattern -> !globalizer (PureAstNode a) ] ];
-(*
- dyntactic: [ [ a = Tactic.tactic -> !globalizer (TacticAstNode a) ] ];
- dynvernac: [ [ a = Vernac_.vernac -> !globalizer(VernacAstNode a) ] ];
-*)
- dynastlist: [ [ a = Prim.astlist -> AstListNode a ] ];
- dynast: [ [ a = ast -> PureAstNode a ] ];
+ dynconstr: [ [ a = Constr.constr -> ConstrNode a ] ];
+ dyncasespattern: [ [ a = Constr.pattern -> CasesPatternNode a ] ];
END
(**********************************************************************)
@@ -519,41 +491,27 @@ END
(* and Syntax pattern, according to the universe of the rule defined *)
type parser_type =
- | AstListParser
- | AstParser
| ConstrParser
| CasesPatternParser
- | TacticParser
- | VernacParser
-let default_action_parser_ref = ref dynast
+let default_action_parser_ref = ref dynconstr
let get_default_action_parser () = !default_action_parser_ref
-let entry_type_from_name = function
- | "constr" -> GenAstType ConstrArgType
- | "tactic" -> failwith "Not supported"
- | "vernac" -> failwith "Not supported"
- | s -> GenAstType ConstrArgType
-
let entry_type_of_parser = function
- | AstListParser -> Some AstListType
- | _ -> None
+ | ConstrParser -> Some ConstrArgType
+ | CasesPatternParser -> failwith "entry_type_of_parser: cases_pattern, TODO"
let parser_type_from_name = function
| "constr" -> ConstrParser
| "cases_pattern" -> CasesPatternParser
- | "tactic" -> TacticParser
- | "vernac" -> VernacParser
- | s -> AstParser
+ | "tactic" -> assert false
+ | "vernac" -> error "No longer supported"
+ | s -> ConstrParser
let set_default_action_parser = function
- | AstListParser -> default_action_parser_ref := dynastlist
- | AstParser -> default_action_parser_ref := dynast
| ConstrParser -> default_action_parser_ref := dynconstr
| CasesPatternParser -> default_action_parser_ref := dyncasespattern
- | TacticParser -> default_action_parser_ref := dyntactic
- | VernacParser -> default_action_parser_ref := dynvernac
let default_action_parser =
Gram.Entry.of_parser "default_action_parser"
diff --git a/parsing/pcoq.mli b/parsing/pcoq.mli
index b4a5bc9a7..a0f5a55c0 100644
--- a/parsing/pcoq.mli
+++ b/parsing/pcoq.mli
@@ -8,12 +8,16 @@
(*i $Id$ i*)
+open Util
open Names
-open Tacexpr
+open Rawterm
open Ast
open Genarg
+open Topconstr
open Tacexpr
open Vernacexpr
+open Libnames
+open Extend
(* The lexer and parser of Coq. *)
@@ -24,11 +28,11 @@ module Gram : Grammar.S with type te = Token.t
type grammar_object
type typed_entry
-val type_of_typed_entry : typed_entry -> entry_type
+val type_of_typed_entry : typed_entry -> Extend.entry_type
val object_of_typed_entry : typed_entry -> grammar_object Gram.Entry.e
val grammar_extend :
- typed_entry -> Gramext.position option ->
+ 'a Gram.Entry.e -> Gramext.position option ->
(string option * Gramext.g_assoc option *
(Token.t Gramext.g_symbol list * Gramext.g_action) list) list
-> unit
@@ -41,12 +45,6 @@ val parse_string : 'a Gram.Entry.e -> string -> 'a
val eoi_entry : 'a Gram.Entry.e -> 'a Gram.Entry.e
val map_entry : ('a -> 'b) -> 'a Gram.Entry.e -> 'b Gram.Entry.e
-(*
-val slam_ast : Coqast.loc -> Coqast.t -> Coqast.t -> Coqast.t
-val abstract_binders_ast :
- Coqast.loc -> string -> Coqast.t -> Coqast.t -> Coqast.t
-*)
-
(* Entry types *)
(* Table of Coq's grammar entries *)
@@ -67,35 +65,30 @@ val force_entry_type :
string * gram_universe -> string -> entry_type -> typed_entry
val create_constr_entry :
- string * gram_universe -> string -> Coqast.t Gram.Entry.e
-val create_generic_entry : string -> ('a, constr_ast,raw_tactic_expr) abstract_argument_type -> 'a Gram.Entry.e
+ string * gram_universe -> string -> constr_expr Gram.Entry.e
+val create_generic_entry : string -> ('a, constr_expr,raw_tactic_expr) abstract_argument_type -> 'a Gram.Entry.e
val get_generic_entry : string -> grammar_object Gram.Entry.e
val get_generic_entry_type : string * gram_universe -> string -> Genarg.argument_type
type parser_type =
- | AstListParser
- | AstParser
| ConstrParser
| CasesPatternParser
- | TacticParser
- | VernacParser
-val entry_type_from_name : string -> entry_type
val entry_type_of_parser : parser_type -> entry_type option
val parser_type_from_name : string -> parser_type
-(* Quotations *)
-val define_quotation : bool -> string -> (Coqast.t Gram.Entry.e) -> unit
-val set_globalizer : (typed_ast -> typed_ast) -> unit
+(* Quotations in ast parser *)
+val define_ast_quotation : bool -> string -> (Coqast.t Gram.Entry.e) -> unit
+val set_globalizer : (constr_expr -> Coqast.t) -> unit
(* The default parser for actions in grammar rules *)
-val default_action_parser : typed_ast Gram.Entry.e
+val default_action_parser : dynamic_grammar Gram.Entry.e
val set_default_action_parser : parser_type -> unit
(* The main entry: reads an optional vernac command *)
-val main_entry : (Coqast.loc * vernac_expr) option Gram.Entry.e
+val main_entry : (loc * vernac_expr) option Gram.Entry.e
(* Initial state of the grammar *)
@@ -106,64 +99,63 @@ module Prim :
open Libnames
val preident : string Gram.Entry.e
val ident : identifier Gram.Entry.e
- val rawident : identifier located Gram.Entry.e
+ val name : name located Gram.Entry.e
+ val identref : identifier located Gram.Entry.e
+ val base_ident : identifier Gram.Entry.e
val natural : int Gram.Entry.e
val integer : int Gram.Entry.e
val string : string Gram.Entry.e
val qualid : qualid located Gram.Entry.e
- val reference : Coqast.reference_expr Gram.Entry.e
+ val reference : reference Gram.Entry.e
val dirpath : dir_path Gram.Entry.e
val astpat: typed_ast Gram.Entry.e
val ast : Coqast.t Gram.Entry.e
val astlist : Coqast.t list Gram.Entry.e
val ast_eoi : Coqast.t Gram.Entry.e
- val astact : Coqast.t Gram.Entry.e
- val metaident : Coqast.t Gram.Entry.e
- val var : Coqast.t Gram.Entry.e
+ val var : identifier Gram.Entry.e
end
module Constr :
sig
- val constr : Coqast.t Gram.Entry.e
- val constr0 : Coqast.t Gram.Entry.e
- val constr1 : Coqast.t Gram.Entry.e
- val constr2 : Coqast.t Gram.Entry.e
- val constr3 : Coqast.t Gram.Entry.e
- val lassoc_constr4 : Coqast.t Gram.Entry.e
- val constr5 : Coqast.t Gram.Entry.e
- val constr6 : Coqast.t Gram.Entry.e
- val constr7 : Coqast.t Gram.Entry.e
- val constr8 : Coqast.t Gram.Entry.e
- val constr9 : Coqast.t Gram.Entry.e
- val constr91 : Coqast.t Gram.Entry.e
- val constr10 : Coqast.t Gram.Entry.e
- val constr_eoi : constr_ast Gram.Entry.e
- val lconstr : Coqast.t Gram.Entry.e
- val ident : Coqast.t Gram.Entry.e
- val qualid : Coqast.t Gram.Entry.e
- val global : Coqast.t Gram.Entry.e
- val ne_ident_comma_list : Coqast.t list Gram.Entry.e
- val ne_constr_list : Coqast.t list Gram.Entry.e
- val sort : Coqast.t Gram.Entry.e
- val pattern : Coqast.t Gram.Entry.e
- val constr_pattern : Coqast.t Gram.Entry.e
- val ne_binders_list : Coqast.t list Gram.Entry.e
- val numarg : Coqast.t Gram.Entry.e
+ val constr : constr_expr Gram.Entry.e
+ val constr0 : constr_expr Gram.Entry.e
+ val constr1 : constr_expr Gram.Entry.e
+ val constr2 : constr_expr Gram.Entry.e
+ val constr3 : constr_expr Gram.Entry.e
+ val lassoc_constr4 : constr_expr Gram.Entry.e
+ val constr5 : constr_expr Gram.Entry.e
+ val constr6 : constr_expr Gram.Entry.e
+ val constr7 : constr_expr Gram.Entry.e
+ val constr8 : constr_expr Gram.Entry.e
+ val constr9 : constr_expr Gram.Entry.e
+ val constr91 : constr_expr Gram.Entry.e
+ val constr10 : constr_expr Gram.Entry.e
+ val constr_eoi : constr_expr Gram.Entry.e
+ val lconstr : constr_expr Gram.Entry.e
+ val ident : identifier Gram.Entry.e
+ val global : reference Gram.Entry.e
+ val ne_name_comma_list : name located list Gram.Entry.e
+ val ne_constr_list : constr_expr list Gram.Entry.e
+ val sort : rawsort Gram.Entry.e
+ val pattern : cases_pattern_expr Gram.Entry.e
+ val constr_pattern : constr_expr Gram.Entry.e
+(*
+ val ne_binders_list : local_binder list Gram.Entry.e
+*)
end
module Module :
sig
- val ne_binders_list : Coqast.t list Gram.Entry.e
- val module_expr : Coqast.t Gram.Entry.e
- val module_type : Coqast.t Gram.Entry.e
+ val module_expr : module_ast Gram.Entry.e
+ val module_type : module_type_ast Gram.Entry.e
end
module Tactic :
sig
open Rawterm
- val castedopenconstr : constr_ast Gram.Entry.e
- val constr_with_bindings : constr_ast with_bindings Gram.Entry.e
- val constrarg : constr_ast may_eval Gram.Entry.e
+ val castedopenconstr : constr_expr Gram.Entry.e
+ val constr_with_bindings : constr_expr with_bindings Gram.Entry.e
+ val constrarg : constr_expr may_eval Gram.Entry.e
val quantified_hypothesis : quantified_hypothesis Gram.Entry.e
val int_or_var : int or_var Gram.Entry.e
val red_expr : raw_red_expr Gram.Entry.e
diff --git a/parsing/ppconstr.ml b/parsing/ppconstr.ml
index 76430e1c4..6dd9211bb 100644
--- a/parsing/ppconstr.ml
+++ b/parsing/ppconstr.ml
@@ -1,25 +1,26 @@
-(****************************************************************************)
-(* *)
-(* The Coq Proof Assistant *)
-(* *)
-(* Projet Coq *)
-(* *)
-(* INRIA LRI-CNRS ENS-CNRS *)
-(* Rocquencourt Orsay Lyon *)
-(* *)
-(****************************************************************************)
-
-(* $:Id$ *)
+(***********************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
+(* \VV/ *************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(***********************************************************************)
+(* $Id$ *)
+
+(*i*)
open Ast
+open Util
open Pp
open Nametab
open Names
open Nameops
open Libnames
open Coqast
-open Extend
-open Util
+open Ppextend
+open Topconstr
+open Term
+(*i*)
let dfltpr ast = (str"#GENTERM " ++ print_ast ast);;
@@ -70,25 +71,190 @@ let wrap_exception = function
str"<PP error: non-printable term>"
| s -> raise s
-let constr_syntax_universe = "constr"
-(* This is starting precedence for printing constructions or tactics *)
-(* Level 9 means no parentheses except for applicative terms (at level 10) *)
-let constr_initial_prec = Some ((constr_syntax_universe,(9,0,0)),Extend.L)
+let latom = 0
+let lannot = 1
+let lprod = 8
+let lcast = 9
+let lapp = 10
+let ltop = (8,E)
-let gentermpr_fail gt =
- Esyntax.genprint globpr constr_syntax_universe constr_initial_prec gt
+let prec_less child (parent,assoc) =
+ (if assoc = E then (<=) else (<)) child parent
-let gentermpr gt =
- try gentermpr_fail gt
- with s -> wrap_exception s
+let env_assoc_value v env =
+ try List.assoc v env
+ with Not_found ->
+ anomaly ("Printing metavariable "^(string_of_id v)^" is unbound")
-(* [at_top] means ids of env must be avoided in bound variables *)
-let gentermpr_core at_top env t =
- gentermpr (Termast.ast_of_constr at_top env t)
+open Symbols
+
+let rec print_hunk pr env = function
+ | UnpMetaVar (e,prec) -> pr prec (env_assoc_value e env)
+ | UnpTerminal s -> str s
+ | UnpBox (b,sub) -> ppcmd_of_box b (prlist (print_hunk pr env) sub)
+ | UnpCut cut -> ppcmd_of_cut cut
+
+let pr_notation pr s env =
+ let unpl, level = find_notation_printing_rule s in
+ prlist (print_hunk pr env) unpl, level
+
+let pr_delimiters x = failwith "pr_delimiters: TODO"
+
+open Rawterm
+
+let pr_opt pr = function
+ | None -> mt ()
+ | Some x -> spc () ++ pr x
+
+let pr_universe u = str "<univ>"
+
+let pr_sort = function
+ | RProp Term.Null -> str "Prop"
+ | RProp Term.Pos -> str "Set"
+ | RType u -> str "Type" ++ pr_opt pr_universe u
+
+let pr_explicitation = function
+ | None -> mt ()
+ | Some n -> int n ++ str "!"
+
+let pr_expl_args pr (a,expl) =
+ pr_explicitation expl ++ pr (latom,E) a
+
+let pr_opt_type pr = function
+ | CHole _ -> mt ()
+ | t -> cut () ++ str ":" ++ pr ltop t
+
+let pr_tight_coma () = str "," ++ cut ()
+
+let pr_name = function
+ | Anonymous -> mt ()
+ | Name id -> pr_id id
+
+let pr_located pr (loc,x) = pr x
-let pr_constr = gentermpr
+let pr_let_binder pr x a =
+ hv 0 (
+ str "[" ++ brk(0,1) ++
+ pr_name x ++ brk(0,1) ++
+ str ":=" ++ brk(0,1) ++
+ pr ltop a ++ brk(0,1) ++
+ str "]")
-let pr_pattern = gentermpr
+let pr_binder pr (nal,t) =
+ hov 0 (
+ prlist_with_sep pr_tight_coma (pr_located pr_name) nal ++
+ pr_opt_type pr t)
+
+let pr_binders pr bl =
+ hv 0 (prlist_with_sep pr_semicolon (pr_binder pr) bl)
+
+let pr_recursive_decl pr id b t c =
+ pr_id id ++
+ brk (1,2) ++ str ": " ++ pr ltop t ++ str ":=" ++
+ brk (1,2) ++ pr ltop c
+
+let pr_fixdecl pr (id,bl,t,c) =
+ pr_recursive_decl pr id
+ (brk (1,2) ++ str "[" ++ pr_binders pr bl ++ str "]") t c
+
+let pr_cofixdecl pr (id,t,c) =
+ pr_recursive_decl pr id (mt ()) t c
+
+let pr_recursive s pr_decl id = function
+ | [] -> anomaly "(co)fixpoint with no definition"
+ | d1::dl ->
+ hov 0 (
+ str "Fix " ++ pr_id id ++ brk (0,2) ++ str "{" ++
+ (v 0 (
+ (hov 0 (pr_decl d1)) ++
+ (prlist (fun fix -> fnl () ++ hov 0 (str "with" ++ pr_decl fix))
+ dl))) ++
+ str "}")
+
+let pr_fix pr = pr_recursive "Fix" (pr_fixdecl pr)
+let pr_cofix pr = pr_recursive "Fix" (pr_cofixdecl pr)
+
+let rec pr_arrow pr = function
+ | CArrow (_,a,b) -> pr (lprod,L) a ++ cut () ++ str "->" ++ pr_arrow pr b
+ | a -> pr (lprod,E) a
+
+let pr_annotation pr t = str "<" ++ pr ltop t ++ str ">"
+
+let pr_cases _ _ _ = str "<CASES(TODO)>"
+
+let rec pr inherited a =
+ let (strm,prec) = match a with
+ | CRef r -> pr_reference r, latom
+ | CFix (_,id,fix) -> pr_fix pr (snd id) fix, latom
+ | CCoFix (_,id,cofix) -> pr_cofix pr (snd id) cofix, latom
+ | CArrow _ -> hv 0 (pr_arrow pr a), lprod
+ | CProdN (_,bl,a) ->
+ hov 0 (
+ str "(" ++ pr_binders pr bl ++ str ")" ++ brk(0,1) ++ pr ltop a), lprod
+ | CLambdaN (_,bl,a) ->
+ hov 0 (
+ str "[" ++ pr_binders pr bl ++ str "]" ++ brk(0,1) ++ pr ltop a), lprod
+ | CLetIn (_,x,a,b) ->
+ hov 0 (pr_let_binder pr (snd x) a ++ cut () ++ pr ltop b), lprod
+ | CAppExpl (_,f,l) ->
+ hov 0 (
+ str "!" ++ pr_reference f ++
+ prlist (fun a -> brk (1,1) ++ pr (latom,E) a) l), lapp
+ | CApp (_,a,l) ->
+ hov 0 (
+ pr (latom,E) a ++
+ prlist (fun a -> brk (1,1) ++ pr_expl_args pr a) l), lapp
+ | CCases (_,po,c,eqns) ->
+ pr_cases po c eqns, lannot
+ | COrderedCase (_,IfStyle,po,c,[b1;b2]) ->
+ (* On force les parenthèses autour d'un "if" sous-terme (même si le
+ parsing est lui plus tolérant) *)
+ hov 0 (
+ pr_opt (pr_annotation pr) po ++
+ hv 0 (
+ str "if" ++ pr ltop c ++ spc () ++
+ hov 0 (str "then" ++ brk (1,1) ++ pr ltop b1) ++ spc () ++
+ hov 0 (str "else" ++ brk (1,1) ++ pr ltop b2))), lapp
+ | COrderedCase (_,LetStyle,po,c,[CLambdaN(_,[_,_ as bd],b)]) ->
+ hov 0 (
+ pr_opt (pr_annotation pr) po ++
+ hv 0 (
+ str "let" ++ brk (1,1) ++
+ hov 0 (str "(" ++ pr_binder pr bd ++ str ")") ++
+ str "=" ++ brk (1,1) ++
+ pr ltop c ++ spc () ++
+ str "in " ++ pr ltop b)), lapp
+ | COrderedCase (_,(MatchStyle|RegularStyle as style),po,c,bl) ->
+ hov 0 (
+ hov 0 (
+ pr_opt (pr_annotation pr) po ++ brk (0,2) ++
+ hov 0 (
+ str (if style=RegularStyle then "Case" else "Match") ++
+ brk (1,1) ++ pr ltop c ++ spc () ++
+ str (if style=RegularStyle then "of" else "with") ++
+ brk (1,3) ++
+ hov 0 (prlist (fun b -> pr ltop b ++ fnl ()) bl)) ++
+ str "end")), lannot
+ | COrderedCase (_,_,_,_,_) ->
+ anomaly "malformed if or destructuring let"
+ | CHole _ -> str "?", latom
+ | CMeta (_,p) -> str "?" ++ int p, latom
+ | CSort (_,s) -> pr_sort s, latom
+ | CCast (_,a,b) ->
+ hv 0 (pr (lcast,L) a ++ cut () ++ str "::" ++ pr (lcast,E) b), lcast
+ | CNotation (_,s,env) -> pr_notation pr s env
+ | CGrammar _ -> failwith "CGrammar: TODO"
+ | CNumeral (_,p) -> Bignat.pr_bigint p, latom
+ | CDelimiters (_,sc,a) -> failwith "pr_delim: TODO"
+(* pr_delimiters (find_delimiters) (pr_constr_expr a)*)
+ | CDynamic _ -> str "<dynamic>", latom
+ in
+ if prec_less prec inherited then strm
+ else str"(" ++ strm ++ str")"
+
+let pr_constr = pr ltop
+
+let pr_pattern x = (* gentermpr*) failwith "pr_pattern: TODO"
let pr_qualid qid = str (string_of_qualid qid)
@@ -110,7 +276,7 @@ let pr_red_flag pr r =
open Genarg
let pr_metanum pr = function
- | AN (_,x) -> pr x
+ | AN x -> pr x
| MetaNum (_,n) -> str "?" ++ int n
let pr_red_expr (pr_constr,pr_ref) = function
@@ -139,7 +305,7 @@ let pr_red_expr (pr_constr,pr_ref) = function
let rec pr_may_eval pr = function
| ConstrEval (r,c) ->
hov 0
- (str "Eval" ++ brk (1,1) ++ pr_red_expr (pr,pr_metanum pr_qualid) r ++
+ (str "Eval" ++ brk (1,1) ++ pr_red_expr (pr,pr_metanum pr_reference) r ++
spc () ++ str "in" ++ brk (1,1) ++ pr c)
| ConstrContext ((_,id),c) ->
hov 0
@@ -147,3 +313,25 @@ let rec pr_may_eval pr = function
str "[" ++ pr c ++ str "]")
| ConstrTypeOf c -> hov 0 (str "Check " ++ brk (1,1) ++ pr c)
| ConstrTerm c -> pr c
+
+(**********************************************************************)
+let constr_syntax_universe = "constr"
+(* This is starting precedence for printing constructions or tactics *)
+(* Level 9 means no parentheses except for applicative terms (at level 10) *)
+let constr_initial_prec = Some (9,Ppextend.L)
+
+let gentermpr_fail gt =
+ Esyntax.genprint globpr constr_syntax_universe constr_initial_prec gt
+
+let gentermpr gt =
+ try gentermpr_fail gt
+ with s -> wrap_exception s
+
+(* [at_top] means ids of env must be avoided in bound variables *)
+let gentermpr_core at_top env t =
+ gentermpr (Termast.ast_of_constr at_top env t)
+(*
+let gentermpr_core at_top env t =
+ pr_constr (Constrextern.extern_constr at_top env t)
+*)
+
diff --git a/parsing/ppconstr.mli b/parsing/ppconstr.mli
index 04225ef7a..bd95637fa 100644
--- a/parsing/ppconstr.mli
+++ b/parsing/ppconstr.mli
@@ -1,13 +1,10 @@
-(****************************************************************************)
-(* *)
-(* The Coq Proof Assistant *)
-(* *)
-(* Projet Coq *)
-(* *)
-(* INRIA LRI-CNRS ENS-CNRS *)
-(* Rocquencourt Orsay Lyon *)
-(* *)
-(****************************************************************************)
+(***********************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
+(* \VV/ *************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(***********************************************************************)
(* $Id$ *)
@@ -18,16 +15,23 @@ open Libnames
open Pcoq
open Rawterm
open Extend
+open Coqast
+open Topconstr
+open Names
val pr_global : global_reference -> std_ppcmds
+
val gentermpr : Coqast.t -> std_ppcmds
val gentermpr_core : bool -> env -> constr -> std_ppcmds
+val pr_opt : ('a -> std_ppcmds) -> 'a option -> std_ppcmds
+val pr_name : name -> std_ppcmds
val pr_qualid : qualid -> std_ppcmds
val pr_red_expr :
('a -> std_ppcmds) * ('b -> std_ppcmds) ->
('a,'b) red_expr_gen -> std_ppcmds
-val pr_pattern : Tacexpr.pattern_ast -> std_ppcmds
-val pr_constr : Genarg.constr_ast -> std_ppcmds
+val pr_sort : rawsort -> std_ppcmds
+val pr_pattern : Tacexpr.pattern_expr -> std_ppcmds
+val pr_constr : constr_expr -> std_ppcmds
val pr_may_eval : ('a -> std_ppcmds) -> 'a may_eval -> std_ppcmds
diff --git a/parsing/pptactic.ml b/parsing/pptactic.ml
index 2abdc6813..6571e0af8 100644
--- a/parsing/pptactic.ml
+++ b/parsing/pptactic.ml
@@ -1,13 +1,10 @@
-(****************************************************************************)
-(* *)
-(* The Coq Proof Assistant *)
-(* *)
-(* Projet Coq *)
-(* *)
-(* INRIA LRI-CNRS ENS-CNRS *)
-(* Rocquencourt Orsay Lyon *)
-(* *)
-(****************************************************************************)
+(***********************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
+(* \VV/ *************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(***********************************************************************)
(* $Id$ *)
@@ -19,8 +16,9 @@ open Extend
open Ppconstr
open Tacexpr
open Rawterm
-open Coqast
+open Topconstr
open Genarg
+open Libnames
(* Extensions *)
let prtac_tab = Hashtbl.create 17
@@ -56,22 +54,14 @@ let pr_rawtac0 =
let pr_arg pr x = spc () ++ pr x
-let pr_name = function
- | Anonymous -> mt ()
- | Name id -> spc () ++ pr_id id
-
let pr_metanum pr = function
- | AN (_,x) -> pr x
+ | AN x -> pr x
| MetaNum (_,n) -> str "?" ++ int n
let pr_or_var pr = function
| ArgArg x -> pr x
| ArgVar (_,s) -> pr_id s
-let pr_opt pr = function
- | None -> mt ()
- | Some x -> spc () ++ pr x
-
let pr_or_meta pr = function
| AI x -> pr x
| _ -> failwith "pr_hyp_location: unexpected quotation meta-variable"
@@ -189,7 +179,7 @@ let pr_autoarg_adding = function
| [] -> mt ()
| l ->
spc () ++ str "Adding [" ++
- hv 0 (prlist_with_sep spc (fun (_,x) -> pr_qualid x) l) ++ str "]"
+ hv 0 (prlist_with_sep spc pr_reference l) ++ str "]"
let pr_autoarg_destructing = function
| true -> spc () ++ str "Destructing"
@@ -207,14 +197,15 @@ let rec pr_rawgen prtac x =
| StringArgType -> spc () ++ str "\"" ++ str (out_gen rawwit_string x) ++ str "\""
| PreIdentArgType -> pr_arg str (out_gen rawwit_pre_ident x)
| IdentArgType -> pr_arg pr_id (out_gen rawwit_ident x)
- | QualidArgType -> pr_arg pr_qualid (snd (out_gen rawwit_qualid x))
+ | RefArgType -> pr_arg pr_reference (out_gen rawwit_ref x)
+ | SortArgType -> pr_arg pr_sort (out_gen rawwit_sort x)
| ConstrArgType -> pr_arg pr_constr (out_gen rawwit_constr x)
| ConstrMayEvalArgType ->
pr_arg (pr_may_eval pr_constr) (out_gen rawwit_constr_may_eval x)
| QuantHypArgType ->
pr_arg pr_quantified_hypothesis (out_gen rawwit_quant_hyp x)
| RedExprArgType ->
- pr_arg (pr_red_expr (pr_constr,pr_metanum pr_qualid)) (out_gen rawwit_red_expr x)
+ pr_arg (pr_red_expr (pr_constr,pr_metanum pr_reference)) (out_gen rawwit_red_expr x)
| TacticArgType -> pr_arg prtac (out_gen rawwit_tactic x)
| CastedOpenConstrArgType ->
pr_arg pr_casted_open_constr (out_gen rawwit_casted_open_constr x)
@@ -264,7 +255,8 @@ let rec pr_generic prtac x =
| StringArgType -> spc () ++ str "\"" ++ str (out_gen wit_string x) ++ str "\""
| PreIdentArgType -> pr_arg str (out_gen wit_pre_ident x)
| IdentArgType -> pr_arg pr_id (out_gen wit_ident x)
- | QualidArgType -> pr_arg pr_global (out_gen wit_qualid x)
+ | RefArgType -> pr_arg pr_global (out_gen wit_ref x)
+ | SortArgType -> pr_arg Printer.prterm (Term.mkSort (out_gen wit_sort x))
| ConstrArgType -> pr_arg Printer.prterm (out_gen wit_constr x)
| ConstrMayEvalArgType ->
pr_arg Printer.prterm (out_gen wit_constr_may_eval x)
@@ -329,7 +321,7 @@ let rec pr_atom0 = function
(* Main tactic printer *)
and pr_atom1 = function
- | TacExtend (s,l) -> pr_extend !pr_rawtac s l
+ | TacExtend (_,s,l) -> pr_extend !pr_rawtac s l
| TacAlias (s,l,_) -> pr_extend !pr_rawtac s (List.map snd l)
(* Basic tactics *)
@@ -372,9 +364,9 @@ and pr_atom1 = function
| TacTrueCut (Some id,c) ->
hov 1 (str "Assert" ++ spc () ++ pr_id id ++ str ":" ++ pr_constr c)
| TacForward (false,na,c) ->
- hov 1 (str "Assert" ++ pr_name na ++ str ":=" ++ pr_constr c)
+ hov 1 (str "Assert" ++ pr_arg pr_name na ++ str ":=" ++ pr_constr c)
| TacForward (true,na,c) ->
- hov 1 (str "Pose" ++ pr_name na ++ str ":=" ++ pr_constr c)
+ hov 1 (str "Pose" ++ pr_arg pr_name na ++ str ":=" ++ pr_constr c)
| TacGeneralize l ->
hov 1 (str "Generalize" ++ spc () ++ prlist_with_sep spc pr_constr l)
| TacGeneralizeDep c ->
@@ -566,10 +558,6 @@ and pr6 = function
| TacArg c -> pr_tacarg c
-and pr_reference = function
- | RQualid (_,qid) -> pr_qualid qid
- | RIdent (_,id) -> pr_id id
-
and pr_tacarg0 = function
| TacDynamic (_,t) -> str ("<dynamic ["^(Dyn.tag t)^"]>")
| MetaNumArg (_,n) -> str ("?" ^ string_of_int n )
@@ -596,8 +584,8 @@ in (prtac,pr0,pr_match_rule)
let (pr_raw_tactic,pr_raw_tactic0,pr_match_rule) =
make_pr_tac
(Ppconstr.pr_constr,
- pr_metanum pr_qualid,
- pr_qualid,
+ pr_metanum pr_reference,
+ pr_reference,
pr_or_meta (fun (loc,id) -> pr_id id),
pr_raw_extend)
diff --git a/parsing/pptactic.mli b/parsing/pptactic.mli
index b049a6c47..a3963571c 100644
--- a/parsing/pptactic.mli
+++ b/parsing/pptactic.mli
@@ -1,13 +1,10 @@
-(****************************************************************************)
-(* *)
-(* The Coq Proof Assistant *)
-(* *)
-(* Projet Coq *)
-(* *)
-(* INRIA LRI-CNRS ENS-CNRS *)
-(* Rocquencourt Orsay Lyon *)
-(* *)
-(****************************************************************************)
+(***********************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
+(* \VV/ *************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(***********************************************************************)
(* $Id$ *)
@@ -28,7 +25,8 @@ val declare_extra_tactic_pprule :
string * Egrammar.grammar_tactic_production list)
-> unit
-val pr_match_rule : bool -> (raw_tactic_expr -> std_ppcmds) -> (pattern_ast,raw_tactic_expr) match_rule -> std_ppcmds
+val pr_match_rule : bool -> (raw_tactic_expr -> std_ppcmds) ->
+ (pattern_expr,raw_tactic_expr) match_rule -> std_ppcmds
val pr_raw_tactic : raw_tactic_expr -> std_ppcmds
diff --git a/parsing/prettyp.ml b/parsing/prettyp.ml
index 0f1157f1d..d963d8644 100644
--- a/parsing/prettyp.ml
+++ b/parsing/prettyp.ml
@@ -398,7 +398,8 @@ let list_filter_vec f vec =
frec (Array.length vec -1) []
(* This is designed to print the contents of an opened section *)
-let read_sec_context (loc,qid) =
+let read_sec_context r =
+ let loc,qid = qualid_of_reference r in
let dir =
try Nametab.locate_section qid
with Not_found ->
@@ -430,7 +431,8 @@ let print_eval red_fun env {uj_val=trm;uj_type=typ} =
let ntrm = red_fun env Evd.empty trm in
(str " = " ++ print_judgment env {uj_val = ntrm; uj_type = typ})
-let print_name (loc,qid) =
+let print_name r =
+ let loc,qid = qualid_of_reference r in
try
let sp = Nametab.locate_obj qid in
let (oname,lobj) =
diff --git a/parsing/prettyp.mli b/parsing/prettyp.mli
index 2d175f1f9..54d952ed5 100644
--- a/parsing/prettyp.mli
+++ b/parsing/prettyp.mli
@@ -30,8 +30,8 @@ val print_context : bool -> Lib.library_segment -> std_ppcmds
val print_library_entry : bool -> (object_name * Lib.node) -> std_ppcmds
val print_full_context : unit -> std_ppcmds
val print_full_context_typ : unit -> std_ppcmds
-val print_sec_context : qualid located -> std_ppcmds
-val print_sec_context_typ : qualid located -> std_ppcmds
+val print_sec_context : reference -> std_ppcmds
+val print_sec_context_typ : reference -> std_ppcmds
val print_judgment : env -> unsafe_judgment -> std_ppcmds
val print_safe_judgment : env -> Safe_typing.judgment -> std_ppcmds
val print_eval :
@@ -40,8 +40,8 @@ val print_eval :
val build_inductive : mutual_inductive -> int ->
global_reference * rel_context * types * identifier array * types array
val print_mutual : mutual_inductive -> std_ppcmds
-val print_name : qualid located -> std_ppcmds
-val print_opaque_name : qualid located -> std_ppcmds
+val print_name : reference -> std_ppcmds
+val print_opaque_name : reference -> std_ppcmds
val print_local_context : unit -> std_ppcmds
(*i
diff --git a/parsing/printer.ml b/parsing/printer.ml
index 5867d8143..6305cd650 100644
--- a/parsing/printer.ml
+++ b/parsing/printer.ml
@@ -35,7 +35,7 @@ let tactic_syntax_universe = "tactic"
(* This is starting precedence for printing constructions or tactics *)
(* Level 9 means no parentheses except for applicative terms (at level 10) *)
-let tactic_initial_prec = Some ((tactic_syntax_universe,(9,0,0)),Extend.L)
+let tactic_initial_prec = Some ((tactic_syntax_universe,(9,0,0)),Ppextend.L)
let prterm_env_at_top env = gentermpr_core true env
let prterm_env env = gentermpr_core false env
diff --git a/parsing/q_coqast.ml4 b/parsing/q_coqast.ml4
index 9b1977f0e..06ccc6bea 100644
--- a/parsing/q_coqast.ml4
+++ b/parsing/q_coqast.ml4
@@ -110,13 +110,15 @@ and expr_list_of_var_list sl =
(* We don't give location for tactic quotation! *)
let loc = dummy_loc
+let dloc = <:expr< (0,0) >>
+
let mlexpr_of_ident id =
<:expr< Names.id_of_string $str:Names.string_of_id id$ >>
let mlexpr_of_name = function
| Names.Anonymous -> <:expr< Names.Anonymous >>
| Names.Name id ->
- <:expr< Names.Names (Names.id_of_string $str:Names.string_of_id id$) >>
+ <:expr< Names.Name (Names.id_of_string $str:Names.string_of_id id$) >>
let mlexpr_of_dirpath dir =
let l = Names.repr_dirpath dir in
@@ -127,8 +129,8 @@ let mlexpr_of_qualid qid =
<:expr< make_qualid $mlexpr_of_dirpath dir$ $mlexpr_of_ident id$ >>
let mlexpr_of_reference = function
- | Coqast.RQualid (loc,qid) -> <:expr< Coqast.RQualid loc $mlexpr_of_qualid qid$ >>
- | Coqast.RIdent (loc,id) -> <:expr< Coqast.RIdent loc $mlexpr_of_ident id$ >>
+ | Libnames.Qualid (loc,qid) -> <:expr< Libnames.Qualid $dloc$ $mlexpr_of_qualid qid$ >>
+ | Libnames.Ident (loc,id) -> <:expr< Libnames.Ident $dloc$ $mlexpr_of_ident id$ >>
let mlexpr_of_bool = function
| true -> <:expr< True >>
@@ -138,14 +140,14 @@ let mlexpr_of_intro_pattern = function
| Tacexpr.IntroOrAndPattern _ -> failwith "mlexpr_of_intro_pattern: TODO"
| Tacexpr.IntroWildcard -> <:expr< Tacexpr.IntroWildcard >>
| Tacexpr.IntroIdentifier id ->
- <:expr< Tacexpr.IntroIdentifier (mlexpr_of_ident loc id) >>
+ <:expr< Tacexpr.IntroIdentifier (mlexpr_of_ident $dloc$ id) >>
let mlexpr_of_ident_option = mlexpr_of_option (mlexpr_of_ident)
let mlexpr_of_or_metanum f = function
- | Rawterm.AN (_,a) -> <:expr< Rawterm.AN loc $f a$ >>
+ | Rawterm.AN a -> <:expr< Rawterm.AN $f a$ >>
| Rawterm.MetaNum (_,n) ->
- <:expr< Rawterm.MetaNum loc $mlexpr_of_int n$ >>
+ <:expr< Rawterm.MetaNum $dloc$ $mlexpr_of_int n$ >>
let mlexpr_of_or_metaid f = function
| Tacexpr.AI a -> <:expr< Tacexpr.AI $f a$ >>
@@ -155,7 +157,7 @@ let mlexpr_of_quantified_hypothesis = function
| Rawterm.AnonHyp n -> <:expr< Rawterm.AnonHyp $mlexpr_of_int n$ >>
| Rawterm.NamedHyp id -> <:expr< Rawterm.NamedHyp $mlexpr_of_ident id$ >>
-let mlexpr_of_located f (loc,x) = <:expr< (loc, $f x$) >>
+let mlexpr_of_located f (loc,x) = <:expr< ($dloc$, $f x$) >>
let mlexpr_of_hyp_location = function
| Tacexpr.InHyp id ->
@@ -181,10 +183,25 @@ let mlexpr_of_red_flags {
Rawterm.rIota = $mlexpr_of_bool bi$;
Rawterm.rZeta = $mlexpr_of_bool bz$;
Rawterm.rDelta = $mlexpr_of_bool bd$;
- Rawterm.rConst = $mlexpr_of_list (mlexpr_of_or_metanum mlexpr_of_qualid) l$
+ Rawterm.rConst = $mlexpr_of_list (mlexpr_of_or_metanum mlexpr_of_reference) l$
} >>
-let mlexpr_of_constr = mlexpr_of_ast
+let rec mlexpr_of_constr = function
+ | Topconstr.CRef r -> <:expr< Topconstr.CRef $mlexpr_of_reference r$ >>
+ | Topconstr.CFix (loc,_,_) -> failwith "mlexpr_of_constr: TODO"
+ | Topconstr.CCoFix (loc,_,_) -> failwith "mlexpr_of_constr: TODO"
+ | Topconstr.CArrow (loc,a,b) ->
+ <:expr< Topconstr.CArrow $dloc$ $mlexpr_of_constr a$ $mlexpr_of_constr b$ >>
+ | Topconstr.CProdN (loc,l,a) -> <:expr< Topconstr.CProdN $dloc$ $mlexpr_of_list (mlexpr_of_pair (mlexpr_of_list (mlexpr_of_pair (fun _ -> dloc) mlexpr_of_name)) mlexpr_of_constr) l$ $mlexpr_of_constr a$ >>
+ | Topconstr.CLambdaN (loc,l,a) -> <:expr< Topconstr.CLambdaN $dloc$ $mlexpr_of_list (mlexpr_of_pair (mlexpr_of_list (mlexpr_of_pair (fun _ -> dloc) mlexpr_of_name)) mlexpr_of_constr) l$ $mlexpr_of_constr a$ >>
+ | Topconstr.CLetIn (loc,_,_,_) -> failwith "mlexpr_of_constr: TODO"
+ | Topconstr.CAppExpl (loc,a,l) -> <:expr< Topconstr.CAppExpl $dloc$ $mlexpr_of_reference a$ $mlexpr_of_list mlexpr_of_constr l$ >>
+ | Topconstr.CApp (loc,a,l) -> <:expr< Topconstr.CApp $dloc$ $mlexpr_of_constr a$ $mlexpr_of_list (mlexpr_of_pair mlexpr_of_constr (mlexpr_of_option mlexpr_of_int)) l$ >>
+ | Topconstr.CCases (loc,_,_,_) -> failwith "mlexpr_of_constr: TODO"
+ | Topconstr.COrderedCase (loc,_,_,_,_) -> failwith "mlexpr_of_constr: TODO"
+ | Topconstr.CHole loc -> <:expr< Topconstr.CHole $dloc$ >>
+ | Topconstr.CMeta (loc,n) -> <:expr< Topconstr.CMeta $dloc$ $mlexpr_of_int n$ >>
+ | _ -> failwith "mlexpr_of_constr: TODO"
let mlexpr_of_red_expr = function
| Rawterm.Red b -> <:expr< Rawterm.Red $mlexpr_of_bool b$ >>
@@ -196,7 +213,7 @@ let mlexpr_of_red_expr = function
<:expr< Rawterm.Lazy $mlexpr_of_red_flags f$ >>
| Rawterm.Unfold l ->
let f1 = mlexpr_of_list mlexpr_of_int in
- let f2 = mlexpr_of_or_metanum mlexpr_of_qualid in
+ let f2 = mlexpr_of_or_metanum mlexpr_of_reference in
let f = mlexpr_of_list (mlexpr_of_pair f1 f2) in
<:expr< Rawterm.Unfold $f l$ >>
| Rawterm.Fold l ->
@@ -213,7 +230,7 @@ let rec mlexpr_of_argtype loc = function
| Genarg.BoolArgType -> <:expr< Genarg.BoolArgType >>
| Genarg.IntArgType -> <:expr< Genarg.IntArgType >>
| Genarg.IntOrVarArgType -> <:expr< Genarg.IntOrVarArgType >>
- | Genarg.QualidArgType -> <:expr< Genarg.QualidArgType >>
+ | Genarg.RefArgType -> <:expr< Genarg.RefArgType >>
| Genarg.PreIdentArgType -> <:expr< Genarg.PreIdentArgType >>
| Genarg.IdentArgType -> <:expr< Genarg.IdentArgType >>
| Genarg.StringArgType -> <:expr< Genarg.StringArgType >>
@@ -222,6 +239,7 @@ let rec mlexpr_of_argtype loc = function
| Genarg.ConstrWithBindingsArgType -> <:expr< Genarg.ConstrWithBindingsArgType >>
| Genarg.RedExprArgType -> <:expr< Genarg.RedExprArgType >>
| Genarg.TacticArgType -> <:expr< Genarg.TacticArgType >>
+ | Genarg.SortArgType -> <:expr< Genarg.SortArgType >>
| Genarg.ConstrArgType -> <:expr< Genarg.ConstrArgType >>
| Genarg.ConstrMayEvalArgType -> <:expr< Genarg.ConstrMayEvalArgType >>
| Genarg.List0ArgType t -> <:expr< Genarg.List0ArgType $mlexpr_of_argtype loc t$ >>
@@ -258,7 +276,7 @@ let mlexpr_of_induction_arg = function
| Tacexpr.ElimOnConstr c ->
<:expr< Tacexpr.ElimOnConstr $mlexpr_of_constr c$ >>
| Tacexpr.ElimOnIdent (_,id) ->
- <:expr< Tacexpr.ElimOnIdent loc $mlexpr_of_ident id$ >>
+ <:expr< Tacexpr.ElimOnIdent $dloc$ $mlexpr_of_ident id$ >>
| Tacexpr.ElimOnAnonHyp n ->
<:expr< Tacexpr.ElimOnAnonHyp $mlexpr_of_int n$ >>
@@ -269,7 +287,7 @@ let mlexpr_of_constr_with_binding =
let mlexpr_of_clause_pattern _ = failwith "mlexpr_of_clause_pattern: TODO"
-let mlexpr_of_pattern_ast = mlexpr_of_ast
+let mlexpr_of_pattern_ast = mlexpr_of_constr
let mlexpr_of_entry_type = function
_ -> failwith "mlexpr_of_entry_type: TODO"
@@ -418,14 +436,14 @@ let rec mlexpr_of_atomic_tactic = function
(*
| Tacexpr.TacExtend (s,l) ->
let l = mlexpr_of_list mlexpr_of_tactic_arg l in
- let loc = MLast.loc_of_expr l in
+ let $dloc$ = MLast.loc_of_expr l in
<:expr< Tacexpr.TacExtend $mlexpr_of_string s$ $l$ >>
*)
| _ -> failwith "Quotation of atomic tactic expressions: TODO"
and mlexpr_of_tactic = function
| Tacexpr.TacAtom (loc,t) ->
- <:expr< Tacexpr.TacAtom loc $mlexpr_of_atomic_tactic t$ >>
+ <:expr< Tacexpr.TacAtom $dloc$ $mlexpr_of_atomic_tactic t$ >>
| Tacexpr.TacThen (t1,t2) ->
<:expr< Tacexpr.TacThen $mlexpr_of_tactic t1$ $mlexpr_of_tactic t2$ >>
| Tacexpr.TacThens (t,tl) ->
@@ -444,9 +462,8 @@ and mlexpr_of_tactic = function
<:expr< Tacexpr.TacRepeat $mlexpr_of_tactic t$ >>
| Tacexpr.TacProgress t ->
<:expr< Tacexpr.TacProgress $mlexpr_of_tactic t$ >>
- | Tacexpr.TacId -> let loc = dummy_loc in <:expr< Tacexpr.TacId >>
- | Tacexpr.TacFail n ->
- let loc = dummy_loc in <:expr< Tacexpr.TacFail $int:string_of_int n$ >>
+ | Tacexpr.TacId -> <:expr< Tacexpr.TacId >>
+ | Tacexpr.TacFail n -> <:expr< Tacexpr.TacFail $int:string_of_int n$ >>
(*
| Tacexpr.TacInfo t -> TacInfo (loc,f t)
@@ -456,7 +473,7 @@ and mlexpr_of_tactic = function
| Tacexpr.TacLetIn (l,t) ->
let f =
mlexpr_of_triple
- (mlexpr_of_pair (fun _ -> <:expr< loc >>) mlexpr_of_ident)
+ (mlexpr_of_pair (fun _ -> dloc) mlexpr_of_ident)
(mlexpr_of_option (mlexpr_of_may_eval mlexpr_of_constr))
mlexpr_of_tactic_arg in
<:expr< Tacexpr.TacLetIn $mlexpr_of_list f l$ $mlexpr_of_tactic t$ >>
@@ -469,11 +486,11 @@ and mlexpr_of_tactic = function
$mlexpr_of_bool lr$
$mlexpr_of_list (mlexpr_of_match_rule mlexpr_of_tactic) l$>>
(*
- | Tacexpr.TacFun of loc * tactic_fun_ast
- | Tacexpr.TacFunRec of loc * identifier * tactic_fun_ast
+ | Tacexpr.TacFun of $dloc$ * tactic_fun_ast
+ | Tacexpr.TacFunRec of $dloc$ * identifier * tactic_fun_ast
*)
(*
- | Tacexpr.TacArg (Tacexpr.AstTacArg (Coqast.Nmeta loc id)) -> anti loc id
+ | Tacexpr.TacArg (Tacexpr.AstTacArg (Coqast.Nmeta $dloc$ id)) -> anti loc id
*)
| Tacexpr.TacArg (Tacexpr.MetaIdArg (_,id)) -> anti loc id
| Tacexpr.TacArg t ->
@@ -483,35 +500,15 @@ and mlexpr_of_tactic = function
and mlexpr_of_tactic_arg = function
| Tacexpr.MetaIdArg (loc,id) -> anti loc id
| Tacexpr.MetaNumArg (loc,n) ->
- <:expr< Tacexpr.MetaNumArg loc $mlexpr_of_int n$ >>
-(*
- | Tacexpr.Identifier id ->
- <:expr< Tacexpr.Identifier $mlexpr_of_ident id$ >>
-*)
-(*
- | Tacexpr.AstTacArg t ->
- <:expr< Tacexpr.AstTacArg $mlexpr_of_ast t$ >>
-*)
+ <:expr< Tacexpr.MetaNumArg $dloc$ $mlexpr_of_int n$ >>
| Tacexpr.TacCall (loc,t,tl) ->
- <:expr< Tacexpr.TacCall loc $mlexpr_of_reference t$ $mlexpr_of_list mlexpr_of_tactic_arg tl$>>
+ <:expr< Tacexpr.TacCall $dloc$ $mlexpr_of_reference t$ $mlexpr_of_list mlexpr_of_tactic_arg tl$>>
| Tacexpr.Tacexp t ->
<:expr< Tacexpr.Tacexp $mlexpr_of_tactic t$ >>
| Tacexpr.ConstrMayEval c ->
<:expr< Tacexpr.ConstrMayEval $mlexpr_of_may_eval mlexpr_of_constr c$ >>
-(*
- | Tacexpr.Constr c ->
- <:expr< Tacexpr.Constr $mlexpr_of_constr c$ >>
-*)
-(*
- | Tacexpr.Qualid q ->
- <:expr< Tacexpr.Qualid $mlexpr_of_qualid q$ >>
-*)
| Tacexpr.Reference r ->
<:expr< Tacexpr.Reference $mlexpr_of_reference r$ >>
-(*
- | Tacexpr.TacArgGen q ->
- <:expr< Tacexpr.TacArgGen $mlexpr_of_argtype q$ >>
-*)
| _ -> failwith "mlexpr_of_tactic_arg: TODO"
let f e =
@@ -542,5 +539,5 @@ let _ =
Quotation.add "constr" (fconstr Pcoq.Constr.constr_eoi);
Quotation.add "tactic" (ftac Pcoq.Tactic.tactic_eoi);
(* Quotation.add "vernac" (f Pcoq.Vernac_.vernac_eoi);*)
- Quotation.add "ast" (f Pcoq.Prim.ast_eoi);
+(* Quotation.add "ast" (f Pcoq.Prim.ast_eoi);*)
Quotation.default := "constr"
diff --git a/parsing/search.ml b/parsing/search.ml
index e1723a1d1..c771a7737 100644
--- a/parsing/search.ml
+++ b/parsing/search.ml
@@ -18,7 +18,6 @@ open Declarations
open Libobject
open Declare
open Coqast
-open Astterm
open Environ
open Pattern
open Printer
diff --git a/parsing/symbols.ml b/parsing/symbols.ml
deleted file mode 100644
index cc76d4aa0..000000000
--- a/parsing/symbols.ml
+++ /dev/null
@@ -1,320 +0,0 @@
-open Util
-open Pp
-open Names
-open Nametab
-open Summary
-open Rawterm
-open Bignat
-
-(* A scope is a set of notations; it includes
-
- - a set of ML interpreters/parsers for positive (e.g. 0, 1, 15, ...) and
- negative numbers (e.g. -0, -2, -13, ...). These interpreters may
- fail if a number has no interpretation in the scope (e.g. there is
- no interpretation for negative numbers in [nat]); interpreters both for
- terms and patterns can be set; these interpreters are in permanent table
- [numeral_interpreter_tab]
- - a set of ML printers for expressions denoting numbers parsable in
- this scope (permanently declared in [Esyntax.primitive_printer_tab])
- - a set of interpretations for infix (more generally distfix) notations
- - an optional pair of delimiters which, when occurring in a syntactic
- expression, set this scope to be the current scope
-*)
-
-let pr_bigint = function
- | POS n -> str (Bignat.to_string n)
- | NEG n -> str "-" ++ str (Bignat.to_string n)
-
-(**********************************************************************)
-(* Scope of symbols *)
-
-type level = Extend.precedence * Extend.precedence list
-type notation = string
-type scope_name = string
-type delimiters = string * string
-type scope = {
- notations: (rawconstr * level) Stringmap.t;
- delimiters: delimiters option
-}
-type scopes = scope_name list
-
-(* Scopes table: scope_name -> symbol_interpretation *)
-let scope_map = ref Stringmap.empty
-
-let empty_scope = {
- notations = Stringmap.empty;
- delimiters = None
-}
-
-let default_scope = "core_scope"
-
-let _ = Stringmap.add default_scope empty_scope !scope_map
-
-let scope_stack = ref [default_scope]
-
-let current_scopes () = !scope_stack
-
-(* TODO: push nat_scope, z_scope, ... in scopes summary *)
-
-(**********************************************************************)
-(* Interpreting numbers (not in summary because functional objects) *)
-
-type numeral_interpreter_name = string
-type numeral_interpreter =
- (loc -> bigint -> rawconstr)
- * (loc -> bigint -> name -> cases_pattern) option
-
-let numeral_interpreter_tab =
- (Hashtbl.create 17 : (numeral_interpreter_name,numeral_interpreter)Hashtbl.t)
-
-let declare_numeral_interpreter sc t =
- Hashtbl.add numeral_interpreter_tab sc t
-
-let lookup_numeral_interpreter s =
- try
- Hashtbl.find numeral_interpreter_tab s
- with Not_found ->
- error ("No interpretation for numerals in scope "^s)
-
-(* For loading without opening *)
-let declare_scope scope =
- try let _ = Stringmap.find scope !scope_map in ()
- with Not_found ->
-(* Options.if_verbose message ("Creating scope "^scope);*)
- scope_map := Stringmap.add scope empty_scope !scope_map
-
-let find_scope scope =
- try Stringmap.find scope !scope_map
- with Not_found -> error ("Scope "^scope^" is not declared")
-
-let check_scope sc = let _ = find_scope sc in ()
-
-let declare_delimiters scope dlm =
- let sc = find_scope scope in
- if sc.delimiters <> None && Options.is_verbose () then
- warning ("Overwriting previous delimiters in "^scope);
- let sc = { sc with delimiters = Some dlm } in
- scope_map := Stringmap.add scope sc !scope_map
-
-(* The mapping between notations and production *)
-
-let declare_notation prec nt c scope =
- let sc = find_scope scope in
- if Stringmap.mem nt sc.notations && Options.is_verbose () then
- warning ("Notation "^nt^" is already used in scope "^scope);
- let sc = { sc with notations = Stringmap.add nt (c,prec) sc.notations } in
- scope_map := Stringmap.add scope sc !scope_map
-
-open Coqast
-
-let rec subst_meta_rawconstr subst = function
- | RMeta (_,n) -> List.nth subst (n-1)
- | t -> map_rawconstr (subst_meta_rawconstr subst) t
-
-let rec find_interpretation f = function
- | scope::scopes ->
- (try f (find_scope scope)
- with Not_found -> find_interpretation f scopes)
- | [] -> raise Not_found
-
-let rec interp_notation ntn scopes args =
- let f scope =
- let (c,_) = Stringmap.find ntn scope.notations in
- subst_meta_rawconstr args c in
- try find_interpretation f scopes
- with Not_found -> anomaly ("Unknown interpretation for notation "^ntn)
-
-let find_notation_with_delimiters scope =
- match (Stringmap.find scope !scope_map).delimiters with
- | Some dlm -> Some (Some dlm)
- | None -> None
-
-let rec find_notation_without_delimiters ntn_scope ntn = function
- | scope::scopes ->
- (* Is the expected printer attached to the most recently open scope? *)
- if scope = ntn_scope then
- Some None
- else
- (* If the most recently open scope has a printer for this pattern
- but not the expected one then we need delimiters *)
- if Stringmap.mem ntn (Stringmap.find scope !scope_map).notations then
- find_notation_with_delimiters ntn_scope
- else
- find_notation_without_delimiters ntn_scope ntn scopes
- | [] ->
- find_notation_with_delimiters ntn_scope
-
-let find_notation ntn_scope ntn scopes =
- match
- find_notation_without_delimiters ntn_scope ntn scopes
- with
- | None -> None
- | Some None -> Some (None,scopes)
- | Some x -> Some (x,ntn_scope::scopes)
-
-let exists_notation_in_scope scope prec ntn r =
- try Stringmap.find ntn (Stringmap.find scope !scope_map).notations = (r,prec)
- with Not_found -> false
-
-let exists_notation_prec prec nt sc =
- try snd (Stringmap.find nt sc.notations) = prec with Not_found -> false
-
-let exists_notation prec nt =
- Stringmap.fold (fun scn sc b -> b or exists_notation_prec prec nt sc)
- !scope_map false
-
-(* We have to print delimiters; look for the more recent defined one *)
-(* Do we need to print delimiters? To know it, we look for a numeral *)
-(* printer available in the current stack of scopes *)
-let find_numeral_with_delimiters scope =
- match (Stringmap.find scope !scope_map).delimiters with
- | Some dlm -> Some (Some dlm)
- | None -> None
-
-let rec find_numeral_without_delimiters printer_scope = function
- | scope :: scopes ->
- (* Is the expected printer attached to the most recently open scope? *)
- if scope = printer_scope then
- Some None
- else
- (* If the most recently open scope has a printer for numerals
- but not the expected one then we need delimiters *)
- if not (Hashtbl.mem numeral_interpreter_tab scope) then
- find_numeral_without_delimiters printer_scope scopes
- else
- find_numeral_with_delimiters printer_scope
- | [] ->
- (* Can we switch to [scope]? Yes if it has defined delimiters *)
- find_numeral_with_delimiters printer_scope
-
-let find_numeral_printer printer_scope scopes =
- match
- find_numeral_without_delimiters printer_scope scopes
- with
- | None -> None
- | Some None -> Some (None,scopes)
- | Some x -> Some (x,printer_scope::scopes)
-
-(* This is the map associating the scope a numeral printer belongs to *)
-(*
-let numeral_printer_map = ref (Stringmap.empty : scope_name Stringmap.t)
-*)
-
-let rec interp_numeral loc n = function
- | scope :: scopes ->
- (try fst (lookup_numeral_interpreter scope) loc n
- with Not_found -> interp_numeral loc n scopes)
- | [] ->
- user_err_loc (loc,"interp_numeral",
- str "No interpretation for numeral " ++ pr_bigint n)
-
-let rec interp_numeral_as_pattern loc n name = function
- | scope :: scopes ->
- (try
- match snd (lookup_numeral_interpreter scope) with
- | None -> raise Not_found
- | Some g -> g loc n name
- with Not_found -> interp_numeral_as_pattern loc n name scopes)
- | [] ->
- user_err_loc (loc,"interp_numeral_as_pattern",
- str "No interpretation for numeral " ++ pr_bigint n)
-
-(* Exportation of scopes *)
-let cache_scope (_,sc) =
- check_scope sc;
- scope_stack := sc :: !scope_stack
-
-let subst_scope (_,subst,sc) = sc
-
-open Libobject
-
-let (inScope,outScope) =
- declare_object {(default_object "SCOPE") with
- cache_function = cache_scope;
- open_function = (fun i o -> if i=1 then cache_scope o);
- subst_function = subst_scope;
- classify_function = (fun (_,o) -> Substitute o);
- export_function = (fun x -> Some x) }
-
-let open_scope sc = Lib.add_anonymous_leaf (inScope sc)
-
-
-(* Special scopes associated to arguments of a global reference *)
-
-open Libnames
-
-module RefOrdered =
- struct
- type t = global_reference
- let compare = Pervasives.compare
- end
-
-module Refmap = Map.Make(RefOrdered)
-
-let arguments_scope = ref Refmap.empty
-
-let cache_arguments_scope (_,(r,scl)) =
- List.iter (option_iter check_scope) scl;
- arguments_scope := Refmap.add r scl !arguments_scope
-
-let subst_arguments_scope (_,subst,(r,scl)) = (subst_global subst r,scl)
-
-let (inArgumentsScope,outArgumentsScope) =
- declare_object {(default_object "ARGUMENTS-SCOPE") with
- cache_function = cache_arguments_scope;
- open_function = (fun i o -> if i=1 then cache_arguments_scope o);
- subst_function = subst_arguments_scope;
- classify_function = (fun (_,o) -> Substitute o);
- export_function = (fun x -> Some x) }
-
-let declare_arguments_scope r scl =
- Lib.add_anonymous_leaf (inArgumentsScope (r,scl))
-
-let find_arguments_scope r =
- try Refmap.find r !arguments_scope
- with Not_found -> []
-
-(* Printing *)
-
-let pr_delimiters = function
- | None -> str "No delimiters"
- | Some (l,r) -> str "Delimiters are " ++ str l ++ str " and " ++ str r
-
-let pr_notation prraw ntn r =
- str ntn ++ str " stands for " ++ prraw r
-
-let pr_named_scope prraw scope sc =
- str "Scope " ++ str scope ++ fnl ()
- ++ pr_delimiters sc.delimiters ++ fnl ()
- ++ Stringmap.fold
- (fun ntn (r,_) strm -> pr_notation prraw ntn r ++ fnl () ++ strm)
- sc.notations (mt ())
-
-let pr_scope prraw scope = pr_named_scope prraw scope (find_scope scope)
-
-let pr_scopes prraw =
- Stringmap.fold
- (fun scope sc strm -> pr_named_scope prraw scope sc ++ fnl () ++ strm)
- !scope_map (mt ())
-
-(* Synchronisation with reset *)
-
-let freeze () = (!scope_map, !scope_stack, !arguments_scope)
-
-let unfreeze (scm,scs,asc) =
- scope_map := scm;
- scope_stack := scs;
- arguments_scope := asc
-
-let init () = ()
-(*
- scope_map := Strinmap.empty;
- scope_stack := Stringmap.empty
-*)
-
-let _ =
- declare_summary "symbols"
- { freeze_function = freeze;
- unfreeze_function = unfreeze;
- init_function = init;
- survive_section = false }
diff --git a/parsing/symbols.mli b/parsing/symbols.mli
deleted file mode 100644
index f5b26b877..000000000
--- a/parsing/symbols.mli
+++ /dev/null
@@ -1,60 +0,0 @@
-open Names
-open Util
-open Nametab
-open Rawterm
-open Bignat
-
-(* A numeral interpreter is the pair of an interpreter for _integer_
- numbers in terms and an optional interpreter in pattern, if
- negative numbers are not supported, the interpreter must fail with
- an appropriate error message *)
-
-type numeral_interpreter_name = string
-type numeral_interpreter =
- (loc -> bigint -> rawconstr)
- * (loc -> bigint -> name -> cases_pattern) option
-
-(* A scope is a set of interpreters for symbols + optional
- interpreter and printers for integers + optional delimiters *)
-
-type level = Extend.precedence * Extend.precedence list
-type scope_name = string
-type delimiters = string * string
-type scope
-type scopes = scope_name list
-
-val default_scope : scope_name
-val current_scopes : unit -> scopes
-val open_scope : scope_name -> unit
-val declare_scope : scope_name -> unit
-
-(* Declare delimiters for printing *)
-val declare_delimiters : scope_name -> delimiters -> unit
-
-(* Declare, interpret, and look for a printer for numeral *)
-val declare_numeral_interpreter :
- numeral_interpreter_name -> numeral_interpreter -> unit
-val interp_numeral : loc -> bigint -> scopes -> rawconstr
-val interp_numeral_as_pattern: loc -> bigint -> name -> scopes -> cases_pattern
-val find_numeral_printer : string -> scopes ->
- (delimiters option * scopes) option
-
-(* Declare, interpret, and look for a printer for symbolic notations *)
-type notation = string
-val declare_notation : level -> notation -> rawconstr -> scope_name -> unit
-val interp_notation : notation -> scopes -> rawconstr list -> rawconstr
-val find_notation : scope_name -> notation -> scopes ->
- (delimiters option * scopes) option
-val exists_notation_in_scope :
- scope_name -> level -> notation -> rawconstr -> bool
-val exists_notation : level -> notation -> bool
-
-(* Declare and look for scopes associated to arguments of a global ref *)
-open Libnames
-val declare_arguments_scope: global_reference -> scope_name option list -> unit
-val find_arguments_scope : global_reference -> scope_name option list
-
-(* Printing scopes *)
-open Pp
-val pr_scope : (rawconstr -> std_ppcmds) -> scope_name -> std_ppcmds
-val pr_scopes : (rawconstr -> std_ppcmds) -> std_ppcmds
diff --git a/parsing/tacextend.ml4 b/parsing/tacextend.ml4
index eb9577902..593fb0169 100644
--- a/parsing/tacextend.ml4
+++ b/parsing/tacextend.ml4
@@ -11,7 +11,6 @@
open Genarg
open Q_util
open Q_coqast
-open Ast
let join_loc (deb1,_) (_,fin2) = (deb1,fin2)
let loc = (0,0)
@@ -43,8 +42,9 @@ let rec make_wit loc = function
| StringArgType -> <:expr< Genarg.wit_string >>
| PreIdentArgType -> <:expr< Genarg.wit_pre_ident >>
| IdentArgType -> <:expr< Genarg.wit_ident >>
- | QualidArgType -> <:expr< Genarg.wit_qualid >>
+ | RefArgType -> <:expr< Genarg.wit_ref >>
| QuantHypArgType -> <:expr< Genarg.wit_quant_hyp >>
+ | SortArgType -> <:expr< Genarg.wit_sort >>
| ConstrArgType -> <:expr< Genarg.wit_constr >>
| ConstrMayEvalArgType -> <:expr< Genarg.wit_constr_may_eval >>
| TacticArgType -> <:expr< Genarg.wit_tactic >>
@@ -179,9 +179,7 @@ let rec interp_entry_name loc s =
| None -> None, <:expr< $lid:s$ >> in
let t =
match t with
- | Some (GenAstType t) -> t
- | Some _ ->
- failwith "Only entries of generic type can be used in extension"
+ | Some t -> t
| None ->
(* Pp.warning_with Pp_control.err_ft
("Unknown primitive grammar entry: "^s);*)
diff --git a/parsing/termast.ml b/parsing/termast.ml
index 1b9c38758..bacfa24ce 100644
--- a/parsing/termast.ml
+++ b/parsing/termast.ml
@@ -217,18 +217,21 @@ let rec ast_of_raw = function
(* Pour compatibilité des theories, il faut LAMBDALIST partout *)
ope("LAMBDALIST",[ast_of_raw t;a])
- | RCases (_,printinfo,typopt,tml,eqns) ->
+ | RCases (_,typopt,tml,eqns) ->
let pred = ast_of_rawopt typopt in
- let tag = match printinfo with
- | PrintIf -> "FORCEIF"
- | PrintLet -> "FORCELET"
- | PrintCases -> "CASES"
- in
+ let tag = "CASES" in
let asttomatch = ope("TOMATCH", List.map ast_of_raw tml) in
let asteqns = List.map ast_of_eqn eqns in
ope(tag,pred::asttomatch::asteqns)
- | ROldCase (_,isrec,typopt,tm,bv) ->
+ | ROrderedCase (_,st,typopt,tm,bv) ->
+ let tag = match st with
+ | IfStyle -> "FORCEIF"
+ | LetStyle -> "FORCELET"
+ | RegularStyle -> "CASES"
+ | MatchStyle -> "MATCH"
+ in
+
(* warning "Old Case syntax"; *)
ope("CASE",(ast_of_rawopt typopt)
::(ast_of_raw tm)
@@ -387,7 +390,7 @@ let rec ast_of_pattern tenv env = function
let tag = if n=1 then "LAMBDA" else "LAMBDALIST" in
ope(tag,[ast_of_pattern tenv env t;a])
- | PCase (typopt,tm,bv) ->
+ | PCase (st,typopt,tm,bv) ->
warning "Old Case syntax";
ope("MUTCASE",(ast_of_patopt tenv env typopt)
::(ast_of_pattern tenv env tm)
diff --git a/parsing/vernacextend.ml4 b/parsing/vernacextend.ml4
index d9d56770e..a910c1c06 100644
--- a/parsing/vernacextend.ml4
+++ b/parsing/vernacextend.ml4
@@ -42,7 +42,8 @@ let rec make_rawwit loc = function
| StringArgType -> <:expr< Genarg.rawwit_string >>
| PreIdentArgType -> <:expr< Genarg.rawwit_pre_ident >>
| IdentArgType -> <:expr< Genarg.rawwit_ident >>
- | QualidArgType -> <:expr< Genarg.rawwit_qualid >>
+ | RefArgType -> <:expr< Genarg.rawwit_ref >>
+ | SortArgType -> <:expr< Genarg.rawwit_sort >>
| ConstrArgType -> <:expr< Genarg.rawwit_constr >>
| ConstrMayEvalArgType -> <:expr< Genarg.rawwit_constr_may_eval >>
| QuantHypArgType -> <:expr< Genarg.rawwit_quant_hyp >>
@@ -147,9 +148,7 @@ let rec interp_entry_name loc s =
| None -> None, <:expr< $lid:s$ >> in
let t =
match t with
- | Some (GenAstType t) -> t
- | Some _ ->
- failwith "Only entries of generic type can be used in extension"
+ | Some t -> t
| None ->
(* Pp.warning_with Pp_control.err_ft
("Unknown primitive grammar entry: "^s);*)