summaryrefslogtreecommitdiff
path: root/parsing/g_vernac.ml4
diff options
context:
space:
mode:
Diffstat (limited to 'parsing/g_vernac.ml4')
-rw-r--r--parsing/g_vernac.ml4488
1 files changed, 247 insertions, 241 deletions
diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4
index e61be53a..61b1de82 100644
--- a/parsing/g_vernac.ml4
+++ b/parsing/g_vernac.ml4
@@ -1,23 +1,25 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
open Pp
-open Compat
open CErrors
open Util
open Names
+open Vernacexpr
open Constrexpr
open Constrexpr_ops
open Extend
-open Vernacexpr
open Decl_kinds
+open Declarations
open Misctypes
-open Tok (* necessary for camlp4 *)
+open Tok (* necessary for camlp5 *)
open Pcoq
open Pcoq.Prim
@@ -41,7 +43,6 @@ let def_body = Gram.entry_create "vernac:def_body"
let decl_notation = Gram.entry_create "vernac:decl_notation"
let record_field = Gram.entry_create "vernac:record_field"
let of_type_with_opt_coercion = Gram.entry_create "vernac:of_type_with_opt_coercion"
-let subgoal_command = Gram.entry_create "proof_mode:subgoal_command"
let instance_name = Gram.entry_create "vernac:instance_name"
let section_subset_expr = Gram.entry_create "vernac:section_subset_expr"
@@ -53,53 +54,57 @@ let make_bullet s =
| '*' -> Star n
| _ -> assert false
+let parse_compat_version ?(allow_old = true) = let open Flags in function
+ | "8.8" -> Current
+ | "8.7" -> V8_7
+ | "8.6" -> V8_6
+ | ("8.5" | "8.4" | "8.3" | "8.2" | "8.1" | "8.0") as s ->
+ CErrors.user_err ~hdr:"get_compat_version"
+ Pp.(str "Compatibility with version " ++ str s ++ str " not supported.")
+ | s ->
+ CErrors.user_err ~hdr:"get_compat_version"
+ Pp.(str "Unknown compatibility version \"" ++ str s ++ str "\".")
+
GEXTEND Gram
- GLOBAL: vernac gallina_ext noedit_mode subprf subgoal_command;
- vernac: FIRST
- [ [ IDENT "Time"; c = located_vernac -> VernacTime c
+ GLOBAL: vernac_control gallina_ext noedit_mode subprf;
+ vernac_control: FIRST
+ [ [ IDENT "Time"; c = located_vernac -> VernacTime (false,c)
| IDENT "Redirect"; s = ne_string; c = located_vernac -> VernacRedirect (s, c)
- | IDENT "Timeout"; n = natural; v = vernac -> VernacTimeout(n,v)
- | IDENT "Fail"; v = vernac -> VernacFail v
-
- | IDENT "Local"; v = vernac_poly -> VernacLocal (true, v)
- | IDENT "Global"; v = vernac_poly -> VernacLocal (false, v)
-
- (* Stm backdoor *)
- | IDENT "Stm"; IDENT "JoinDocument"; "." -> VernacStm JoinDocument
- | IDENT "Stm"; IDENT "Finish"; "." -> VernacStm Finish
- | IDENT "Stm"; IDENT "Wait"; "." -> VernacStm Wait
- | IDENT "Stm"; IDENT "PrintDag"; "." -> VernacStm PrintDag
- | IDENT "Stm"; IDENT "Observe"; id = INT; "." ->
- VernacStm (Observe (Stateid.of_int (int_of_string id)))
- | IDENT "Stm"; IDENT "Command"; v = vernac_aux -> VernacStm (Command v)
- | IDENT "Stm"; IDENT "PGLast"; v = vernac_aux -> VernacStm (PGLast v)
+ | IDENT "Timeout"; n = natural; v = vernac_control -> VernacTimeout(n,v)
+ | IDENT "Fail"; v = vernac_control -> VernacFail v
+ | (f, v) = vernac -> VernacExpr(f, v) ]
+ ]
+ ;
+ vernac:
+ [ [ IDENT "Local"; (f, v) = vernac_poly -> (VernacLocal true :: f, v)
+ | IDENT "Global"; (f, v) = vernac_poly -> (VernacLocal false :: f, v)
| v = vernac_poly -> v ]
]
;
vernac_poly:
- [ [ IDENT "Polymorphic"; v = vernac_aux -> VernacPolymorphic (true, v)
- | IDENT "Monomorphic"; v = vernac_aux -> VernacPolymorphic (false, v)
+ [ [ IDENT "Polymorphic"; (f, v) = vernac_aux -> (VernacPolymorphic true :: f, v)
+ | IDENT "Monomorphic"; (f, v) = vernac_aux -> (VernacPolymorphic false :: f, v)
| v = vernac_aux -> v ]
]
;
vernac_aux:
(* Better to parse "." here: in case of failure (e.g. in coerce_to_var), *)
(* "." is still in the stream and discard_to_dot works correctly *)
- [ [ IDENT "Program"; g = gallina; "." -> VernacProgram g
- | IDENT "Program"; g = gallina_ext; "." -> VernacProgram g
- | g = gallina; "." -> g
- | g = gallina_ext; "." -> g
- | c = command; "." -> c
- | c = syntax; "." -> c
- | c = subprf -> c
+ [ [ IDENT "Program"; g = gallina; "." -> ([VernacProgram], g)
+ | IDENT "Program"; g = gallina_ext; "." -> ([VernacProgram], g)
+ | g = gallina; "." -> ([], g)
+ | g = gallina_ext; "." -> ([], g)
+ | c = command; "." -> ([], c)
+ | c = syntax; "." -> ([], c)
+ | c = subprf -> ([], c)
] ]
;
vernac_aux: LAST
- [ [ prfcom = command_entry -> prfcom ] ]
+ [ [ prfcom = command_entry -> ([], prfcom) ] ]
;
noedit_mode:
- [ [ c = subgoal_command -> c None] ]
+ [ [ c = query_command -> c None] ]
;
subprf:
@@ -109,17 +114,8 @@ GEXTEND Gram
] ]
;
- subgoal_command:
- [ [ c = query_command; "." ->
- begin function
- | Some (SelectNth g) -> c (Some g)
- | None -> c None
- | _ ->
- VernacError (UserError ("",str"Typing and evaluation commands, cannot be used with the \"all:\" selector."))
- end ] ]
- ;
located_vernac:
- [ [ v = vernac -> !@loc, v ] ]
+ [ [ v = vernac_control -> CAst.make ~loc:!@loc v ] ]
;
END
@@ -137,47 +133,57 @@ let test_plural_form_types loc kwd = function
warn_plural_command ~loc:!@loc kwd
| _ -> ()
-let fresh_var env c =
- Namegen.next_ident_away (Id.of_string "pat")
- (env @ Id.Set.elements (Topconstr.free_vars_of_constr_expr c))
+let lname_of_lident : lident -> lname =
+ CAst.map (fun s -> Name s)
-let _ = Hook.set Constrexpr_ops.fresh_var_hook fresh_var
+let name_of_ident_decl : ident_decl -> name_decl =
+ on_fst lname_of_lident
(* Gallina declarations *)
GEXTEND Gram
GLOBAL: gallina gallina_ext thm_token def_body of_type_with_opt_coercion
- record_field decl_notation rec_definition pidentref;
+ record_field decl_notation rec_definition ident_decl univ_decl;
gallina:
(* Definition, Theorem, Variable, Axiom, ... *)
- [ [ thm = thm_token; id = pidentref; bl = binders; ":"; c = lconstr;
+ [ [ thm = thm_token; id = ident_decl; bl = binders; ":"; c = lconstr;
l = LIST0
- [ "with"; id = pidentref; bl = binders; ":"; c = lconstr ->
- (Some id,(bl,c,None)) ] ->
- VernacStartTheoremProof (thm, (Some id,(bl,c,None))::l, false)
+ [ "with"; id = ident_decl; bl = binders; ":"; c = lconstr ->
+ (id,(bl,c)) ] ->
+ VernacStartTheoremProof (thm, (id,(bl,c))::l)
| stre = assumption_token; nl = inline; bl = assum_list ->
VernacAssumption (stre, nl, bl)
| (kwd,stre) = assumptions_token; nl = inline; bl = assum_list ->
test_plural_form loc kwd bl;
VernacAssumption (stre, nl, bl)
- | d = def_token; id = pidentref; b = def_body ->
- VernacDefinition (d, id, b)
+ | d = def_token; id = ident_decl; b = def_body ->
+ VernacDefinition (d, name_of_ident_decl id, b)
| IDENT "Let"; id = identref; b = def_body ->
- VernacDefinition ((Some Discharge, Definition), (id, None), b)
+ VernacDefinition ((DoDischarge, Let), (lname_of_lident id, None), b)
(* Gallina inductive declarations *)
- | priv = private_token; f = finite_token;
+ | cum = cumulativity_token; priv = private_token; f = finite_token;
indl = LIST1 inductive_definition SEP "with" ->
let (k,f) = f in
- let indl=List.map (fun ((a,b,c,d),e) -> ((a,b,c,k,d),e)) indl in
- VernacInductive (priv,f,indl)
+ let indl=List.map (fun ((a,b,c,d),e) -> ((a,b,c,k,d),e)) indl in
+ let cum =
+ match cum with
+ Some true -> LocalCumulativity
+ | Some false -> LocalNonCumulativity
+ | None ->
+ if Flags.is_polymorphic_inductive_cumulativity () then
+ GlobalCumulativity
+ else
+ GlobalNonCumulativity
+ in
+ VernacInductive (cum, priv,f,indl)
| "Fixpoint"; recs = LIST1 rec_definition SEP "with" ->
- VernacFixpoint (None, recs)
+ VernacFixpoint (NoDischarge, recs)
| IDENT "Let"; "Fixpoint"; recs = LIST1 rec_definition SEP "with" ->
- VernacFixpoint (Some Discharge, recs)
+ VernacFixpoint (DoDischarge, recs)
| "CoFixpoint"; corecs = LIST1 corec_definition SEP "with" ->
- VernacCoFixpoint (None, corecs)
+ VernacCoFixpoint (NoDischarge, corecs)
| IDENT "Let"; "CoFixpoint"; corecs = LIST1 corec_definition SEP "with" ->
- VernacCoFixpoint (Some Discharge, corecs)
+ VernacCoFixpoint (DoDischarge, corecs)
| IDENT "Scheme"; l = LIST1 scheme SEP "with" -> VernacScheme l
| IDENT "Combined"; IDENT "Scheme"; id = identref; IDENT "from";
l = LIST1 identref SEP "," -> VernacCombinedScheme (id, l)
@@ -199,60 +205,82 @@ GEXTEND Gram
| IDENT "Property" -> Property ] ]
;
def_token:
- [ [ "Definition" -> (None, Definition)
- | IDENT "Example" -> (None, Example)
- | IDENT "SubClass" -> (None, SubClass) ] ]
+ [ [ "Definition" -> (NoDischarge,Definition)
+ | IDENT "Example" -> (NoDischarge,Example)
+ | IDENT "SubClass" -> (NoDischarge,SubClass) ] ]
;
assumption_token:
- [ [ "Hypothesis" -> (Some Discharge, Logical)
- | "Variable" -> (Some Discharge, Definitional)
- | "Axiom" -> (None, Logical)
- | "Parameter" -> (None, Definitional)
- | IDENT "Conjecture" -> (None, Conjectural) ] ]
+ [ [ "Hypothesis" -> (DoDischarge, Logical)
+ | "Variable" -> (DoDischarge, Definitional)
+ | "Axiom" -> (NoDischarge, Logical)
+ | "Parameter" -> (NoDischarge, Definitional)
+ | IDENT "Conjecture" -> (NoDischarge, Conjectural) ] ]
;
assumptions_token:
- [ [ IDENT "Hypotheses" -> ("Hypotheses", (Some Discharge, Logical))
- | IDENT "Variables" -> ("Variables", (Some Discharge, Definitional))
- | IDENT "Axioms" -> ("Axioms", (None, Logical))
- | IDENT "Parameters" -> ("Parameters", (None, Definitional))
- | IDENT "Conjectures" -> ("Conjectures", (None, Conjectural)) ] ]
+ [ [ IDENT "Hypotheses" -> ("Hypotheses", (DoDischarge, Logical))
+ | IDENT "Variables" -> ("Variables", (DoDischarge, Definitional))
+ | IDENT "Axioms" -> ("Axioms", (NoDischarge, Logical))
+ | IDENT "Parameters" -> ("Parameters", (NoDischarge, Definitional))
+ | IDENT "Conjectures" -> ("Conjectures", (NoDischarge, Conjectural)) ] ]
;
inline:
[ [ IDENT "Inline"; "("; i = INT; ")" -> InlineAt (int_of_string i)
| IDENT "Inline" -> DefaultInline
| -> NoInline] ]
;
- pidentref:
- [ [ i = identref; l = OPT [ "@{" ; l = LIST0 identref; "}" -> l ] -> (i,l) ] ]
- ;
univ_constraint:
[ [ l = universe_level; ord = [ "<" -> Univ.Lt | "=" -> Univ.Eq | "<=" -> Univ.Le ];
r = universe_level -> (l, ord, r) ] ]
;
+ univ_decl :
+ [ [ "@{" ; l = LIST0 identref; ext = [ "+" -> true | -> false ];
+ cs = [ "|"; l' = LIST0 univ_constraint SEP ",";
+ ext = [ "+" -> true | -> false ]; "}" -> (l',ext)
+ | ext = [ "}" -> true | "|}" -> false ] -> ([], ext) ]
+ ->
+ { univdecl_instance = l;
+ univdecl_extensible_instance = ext;
+ univdecl_constraints = fst cs;
+ univdecl_extensible_constraints = snd cs }
+ ] ]
+ ;
+ ident_decl:
+ [ [ i = identref; l = OPT univ_decl -> (i, l)
+ ] ]
+ ;
finite_token:
- [ [ "Inductive" -> (Inductive_kw,Finite)
- | "CoInductive" -> (CoInductive,CoFinite)
- | "Variant" -> (Variant,BiFinite)
+ [ [ IDENT "Inductive" -> (Inductive_kw,Finite)
+ | IDENT "CoInductive" -> (CoInductive,CoFinite)
+ | IDENT "Variant" -> (Variant,BiFinite)
| IDENT "Record" -> (Record,BiFinite)
| IDENT "Structure" -> (Structure,BiFinite)
| IDENT "Class" -> (Class true,BiFinite) ] ]
;
+ cumulativity_token:
+ [ [ IDENT "Cumulative" -> Some true | IDENT "NonCumulative" -> Some false | -> None ] ]
+ ;
private_token:
[ [ IDENT "Private" -> true | -> false ] ]
;
(* Simple definitions *)
def_body:
[ [ bl = binders; ":="; red = reduce; c = lconstr ->
- let (bl, c) = expand_pattern_binders mkCLambdaN bl c in
- (match c with
- CCast(_,c, CastConv t) -> DefineBody (bl, red, c, Some t)
+ if List.exists (function CLocalPattern _ -> true | _ -> false) bl
+ then
+ (* FIXME: "red" will be applied to types in bl and Cast with remain *)
+ let c = mkCLambdaN ~loc:!@loc bl c in
+ DefineBody ([], red, c, None)
+ else
+ (match c with
+ | { CAst.v = CCast(c, CastConv t) } -> DefineBody (bl, red, c, Some t)
| _ -> DefineBody (bl, red, c, None))
| bl = binders; ":"; t = lconstr; ":="; red = reduce; c = lconstr ->
let ((bl, c), tyo) =
- if List.exists (function LocalPattern _ -> true | _ -> false) bl
+ if List.exists (function CLocalPattern _ -> true | _ -> false) bl
then
- let c = CCast (!@loc, c, CastConv t) in
- (expand_pattern_binders mkCLambdaN bl c, None)
+ (* FIXME: "red" will be applied to types in bl and Cast with remain *)
+ let c = CAst.make ~loc:!@loc @@ CCast (c, CastConv t) in
+ (([],mkCLambdaN ~loc:!@loc bl c), None)
else ((bl, c), Some t)
in
DefineBody (bl, red, c, tyo)
@@ -260,7 +288,7 @@ GEXTEND Gram
ProveBody (bl, t) ] ]
;
reduce:
- [ [ IDENT "Eval"; r = Tactic.red_expr; "in" -> Some r
+ [ [ IDENT "Eval"; r = red_expr; "in" -> Some r
| -> None ] ]
;
one_decl_notation:
@@ -277,7 +305,7 @@ GEXTEND Gram
| -> RecordDecl (None, []) ] ]
;
inductive_definition:
- [ [ oc = opt_coercion; id = pidentref; indpar = binders;
+ [ [ oc = opt_coercion; id = ident_decl; indpar = binders;
c = OPT [ ":"; c = lconstr -> c ];
lc=opt_constructors_or_fields; ntn = decl_notation ->
(((oc,id),indpar,c,lc),ntn) ] ]
@@ -303,20 +331,20 @@ GEXTEND Gram
;
(* (co)-fixpoints *)
rec_definition:
- [ [ id = pidentref;
+ [ [ id = ident_decl;
bl = binders_fixannot;
ty = type_cstr;
def = OPT [":="; def = lconstr -> def]; ntn = decl_notation ->
let bl, annot = bl in ((id,annot,bl,ty,def),ntn) ] ]
;
corec_definition:
- [ [ id = pidentref; bl = binders; ty = type_cstr;
+ [ [ id = ident_decl; bl = binders; ty = type_cstr;
def = OPT [":="; def = lconstr -> def]; ntn = decl_notation ->
((id,bl,ty,def),ntn) ] ]
;
type_cstr:
[ [ ":"; c=lconstr -> c
- | -> CHole (!@loc, None, Misctypes.IntroAnonymous, None) ] ]
+ | -> CAst.make ~loc:!@loc @@ CHole (None, Misctypes.IntroAnonymous, None) ] ]
;
(* Inductive schemes *)
scheme:
@@ -325,13 +353,13 @@ GEXTEND Gram
;
scheme_kind:
[ [ IDENT "Induction"; "for"; ind = smart_global;
- IDENT "Sort"; s = sort-> InductionScheme(true,ind,s)
+ IDENT "Sort"; s = sort_family-> InductionScheme(true,ind,s)
| IDENT "Minimality"; "for"; ind = smart_global;
- IDENT "Sort"; s = sort-> InductionScheme(false,ind,s)
+ IDENT "Sort"; s = sort_family-> InductionScheme(false,ind,s)
| IDENT "Elimination"; "for"; ind = smart_global;
- IDENT "Sort"; s = sort-> CaseScheme(true,ind,s)
+ IDENT "Sort"; s = sort_family-> CaseScheme(true,ind,s)
| IDENT "Case"; "for"; ind = smart_global;
- IDENT "Sort"; s = sort-> CaseScheme(false,ind,s)
+ IDENT "Sort"; s = sort_family-> CaseScheme(false,ind,s)
| IDENT "Equality"; "for" ; ind = smart_global -> EqualityScheme(ind) ] ]
;
(* Various Binders *)
@@ -340,8 +368,8 @@ GEXTEND Gram
binder_nodef:
[ [ b = binder_let ->
(match b with
- LocalRawAssum(l,ty) -> (l,ty)
- | LocalRawDef _ ->
+ CLocalAssum(l,ty) -> (l,ty)
+ | CLocalDef _ ->
Util.user_err_loc
(loc,"fix_param",Pp.str"defined binder not allowed here.")) ] ]
;
@@ -360,19 +388,19 @@ GEXTEND Gram
;
record_binder_body:
[ [ l = binders; oc = of_type_with_opt_coercion;
- t = lconstr -> fun id -> (oc,AssumExpr (id,mkCProdN (!@loc) l t))
+ t = lconstr -> fun id -> (oc,AssumExpr (id,mkCProdN ~loc:!@loc l t))
| l = binders; oc = of_type_with_opt_coercion;
t = lconstr; ":="; b = lconstr -> fun id ->
- (oc,DefExpr (id,mkCLambdaN (!@loc) l b,Some (mkCProdN (!@loc) l t)))
+ (oc,DefExpr (id,mkCLambdaN ~loc:!@loc l b,Some (mkCProdN ~loc:!@loc l t)))
| l = binders; ":="; b = lconstr -> fun id ->
- match b with
- | CCast(_,b, (CastConv t|CastVM t|CastNative t)) ->
- (None,DefExpr(id,mkCLambdaN (!@loc) l b,Some (mkCProdN (!@loc) l t)))
+ match b.CAst.v with
+ | CCast(b', (CastConv t|CastVM t|CastNative t)) ->
+ (None,DefExpr(id,mkCLambdaN ~loc:!@loc l b',Some (mkCProdN ~loc:!@loc l t)))
| _ ->
- (None,DefExpr(id,mkCLambdaN (!@loc) l b,None)) ] ]
+ (None,DefExpr(id,mkCLambdaN ~loc:!@loc l b,None)) ] ]
;
record_binder:
- [ [ id = name -> (None,AssumExpr(id,CHole (!@loc, None, Misctypes.IntroAnonymous, None)))
+ [ [ id = name -> (None,AssumExpr(id, CAst.make ~loc:!@loc @@ CHole (None, Misctypes.IntroAnonymous, None)))
| id = name; f = record_binder_body -> f id ] ]
;
assum_list:
@@ -382,16 +410,16 @@ GEXTEND Gram
[ [ "("; a = simple_assum_coe; ")" -> a ] ]
;
simple_assum_coe:
- [ [ idl = LIST1 pidentref; oc = of_type_with_opt_coercion; c = lconstr ->
+ [ [ idl = LIST1 ident_decl; oc = of_type_with_opt_coercion; c = lconstr ->
(not (Option.is_empty oc),(idl,c)) ] ]
;
constructor_type:
[[ l = binders;
t= [ coe = of_type_with_opt_coercion; c = lconstr ->
- fun l id -> (not (Option.is_empty coe),(id,mkCProdN (!@loc) l c))
+ fun l id -> (not (Option.is_empty coe),(id,mkCProdN ~loc:!@loc l c))
| ->
- fun l id -> (false,(id,mkCProdN (!@loc) l (CHole (!@loc, None, Misctypes.IntroAnonymous, None)))) ]
+ fun l id -> (false,(id,mkCProdN ~loc:!@loc l (CAst.make ~loc:!@loc @@ CHole (None, Misctypes.IntroAnonymous, None)))) ]
-> t l
]]
;
@@ -413,7 +441,7 @@ let only_starredidentrefs =
Gram.Entry.of_parser "test_only_starredidentrefs"
(fun strm ->
let rec aux n =
- match get_tok (Util.stream_nth n strm) with
+ match Util.stream_nth n strm with
| KEYWORD "." -> ()
| KEYWORD ")" -> ()
| (IDENT _ | KEYWORD "Type" | KEYWORD "*") -> aux (n+1)
@@ -522,25 +550,26 @@ GEXTEND Gram
(* Module expressions *)
module_expr:
[ [ me = module_expr_atom -> me
- | me1 = module_expr; me2 = module_expr_atom -> CMapply (!@loc,me1,me2)
+ | me1 = module_expr; me2 = module_expr_atom -> CAst.make ~loc:!@loc @@ CMapply (me1,me2)
] ]
;
module_expr_atom:
- [ [ qid = qualid -> CMident qid | "("; me = module_expr; ")" -> me ] ]
+ [ [ qid = qualid -> CAst.make ~loc:!@loc @@ CMident (qid.CAst.v) | "("; me = module_expr; ")" -> me ] ]
;
with_declaration:
- [ [ "Definition"; fqid = fullyqualid; ":="; c = Constr.lconstr ->
- CWith_Definition (fqid,c)
+ [ [ "Definition"; fqid = fullyqualid; udecl = OPT univ_decl; ":="; c = Constr.lconstr ->
+ CWith_Definition (fqid,udecl,c)
| IDENT "Module"; fqid = fullyqualid; ":="; qid = qualid ->
CWith_Module (fqid,qid)
] ]
;
module_type:
- [ [ qid = qualid -> CMident qid
+ [ [ qid = qualid -> CAst.make ~loc:!@loc @@ CMident (qid.CAst.v)
| "("; mt = module_type; ")" -> mt
- | mty = module_type; me = module_expr_atom -> CMapply (!@loc,mty,me)
+ | mty = module_type; me = module_expr_atom ->
+ CAst.make ~loc:!@loc @@ CMapply (mty,me)
| mty = module_type; "with"; decl = with_declaration ->
- CMwith (!@loc,mty,decl)
+ CAst.make ~loc:!@loc @@ CMwith (mty,decl)
] ]
;
(* Proof using *)
@@ -552,8 +581,8 @@ GEXTEND Gram
starredidentref:
[ [ i = identref -> SsSingl i
| i = identref; "*" -> SsFwdClose(SsSingl i)
- | "Type" -> SsSingl (!@loc, Id.of_string "Type")
- | "Type"; "*" -> SsFwdClose (SsSingl (!@loc, Id.of_string "Type")) ]]
+ | "Type" -> SsType
+ | "Type"; "*" -> SsFwdClose SsType ]]
;
ssexpr:
[ "35"
@@ -580,12 +609,6 @@ let warn_deprecated_implicit_arguments =
CWarnings.create ~name:"deprecated-implicit-arguments" ~category:"deprecated"
(fun () -> strbrk "Implicit Arguments is deprecated; use Arguments instead")
-let warn_deprecated_arguments_syntax =
- CWarnings.create ~name:"deprecated-arguments-syntax" ~category:"deprecated"
- (fun () -> strbrk "The \"/\" and \"!\" modifiers have an effect only "
- ++ strbrk "in the first arguments list. The syntax allowing"
- ++ strbrk " them to appear in other lists is deprecated.")
-
(* Extensions: implicits, coercions, etc. *)
GEXTEND Gram
GLOBAL: gallina_ext instance_name hint_info;
@@ -601,43 +624,29 @@ GEXTEND Gram
VernacSetStrategy l
(* Canonical structure *)
| IDENT "Canonical"; IDENT "Structure"; qid = global ->
- VernacCanonical (AN qid)
+ VernacCanonical CAst.(make ~loc:!@loc @@ AN qid)
| IDENT "Canonical"; IDENT "Structure"; ntn = by_notation ->
- VernacCanonical (ByNotation ntn)
- | IDENT "Canonical"; IDENT "Structure"; qid = global;
- d = def_body ->
+ VernacCanonical CAst.(make ~loc:!@loc @@ ByNotation ntn)
+ | IDENT "Canonical"; IDENT "Structure"; qid = global; d = def_body ->
let s = coerce_reference_to_id qid in
- VernacDefinition
- ((Some Global,CanonicalStructure),((Loc.ghost,s),None),d)
+ VernacDefinition ((NoDischarge,CanonicalStructure),((CAst.make (Name s)),None),d)
(* Coercions *)
| IDENT "Coercion"; qid = global; d = def_body ->
let s = coerce_reference_to_id qid in
- VernacDefinition ((None,Coercion),((Loc.ghost,s),None),d)
- | IDENT "Coercion"; IDENT "Local"; qid = global; d = def_body ->
- let s = coerce_reference_to_id qid in
- VernacDefinition ((Some Decl_kinds.Local,Coercion),((Loc.ghost,s),None),d)
- | IDENT "Identity"; IDENT "Coercion"; IDENT "Local"; f = identref;
- ":"; s = class_rawexpr; ">->"; t = class_rawexpr ->
- VernacIdentityCoercion (true, f, s, t)
+ VernacDefinition ((NoDischarge,Coercion),((CAst.make (Name s)),None),d)
| IDENT "Identity"; IDENT "Coercion"; f = identref; ":";
s = class_rawexpr; ">->"; t = class_rawexpr ->
- VernacIdentityCoercion (false, f, s, t)
- | IDENT "Coercion"; IDENT "Local"; qid = global; ":";
- s = class_rawexpr; ">->"; t = class_rawexpr ->
- VernacCoercion (true, AN qid, s, t)
- | IDENT "Coercion"; IDENT "Local"; ntn = by_notation; ":";
- s = class_rawexpr; ">->"; t = class_rawexpr ->
- VernacCoercion (true, ByNotation ntn, s, t)
+ VernacIdentityCoercion (f, s, t)
| IDENT "Coercion"; qid = global; ":"; s = class_rawexpr; ">->";
t = class_rawexpr ->
- VernacCoercion (false, AN qid, s, t)
+ VernacCoercion (CAst.make ~loc:!@loc @@ AN qid, s, t)
| IDENT "Coercion"; ntn = by_notation; ":"; s = class_rawexpr; ">->";
t = class_rawexpr ->
- VernacCoercion (false, ByNotation ntn, s, t)
+ VernacCoercion (CAst.make ~loc:!@loc @@ ByNotation ntn, s, t)
- | IDENT "Context"; c = binders ->
- VernacContext c
+ | IDENT "Context"; c = LIST1 binder ->
+ VernacContext (List.flatten c)
| IDENT "Instance"; namesup = instance_name; ":";
expl = [ "!" -> Decl_kinds.Implicit | -> Decl_kinds.Explicit ] ; t = operconstr LEVEL "200";
@@ -663,10 +672,7 @@ GEXTEND Gram
args = LIST0 argument_spec_block;
more_implicits = OPT
[ ","; impl = LIST1
- [ impl = LIST0 more_implicits_block ->
- let warn_deprecated = List.exists fst impl in
- if warn_deprecated then warn_deprecated_arguments_syntax ~loc:!@loc ();
- List.flatten (List.map snd impl)]
+ [ impl = LIST0 more_implicits_block -> List.flatten impl]
SEP "," -> impl
];
mods = OPT [ ":"; l = LIST1 arguments_modifier SEP "," -> l ] ->
@@ -679,7 +685,7 @@ GEXTEND Gram
if Option.is_empty !slash_position then
(slash_position := Some i; parse_args i args)
else
- error "The \"/\" modifier can occur only once"
+ user_err Pp.(str "The \"/\" modifier can occur only once")
in
let args = parse_args 0 (List.flatten args) in
let more_implicits = Option.default [] more_implicits in
@@ -739,7 +745,7 @@ GEXTEND Gram
;
argument_spec: [
[ b = OPT "!"; id = name ; s = OPT scope ->
- snd id, not (Option.is_empty b), Option.map (fun x -> !@loc, x) s
+ id.CAst.v, not (Option.is_empty b), Option.map (fun x -> CAst.make ~loc:!@loc x) s
]
];
(* List of arguments implicit status, scope, modifiers *)
@@ -752,43 +758,37 @@ GEXTEND Gram
| "/" -> [`Slash]
| "("; items = LIST1 argument_spec; ")"; sc = OPT scope ->
let f x = match sc, x with
- | None, x -> x | x, None -> Option.map (fun y -> !@loc, y) x
- | Some _, Some _ -> error "scope declared twice" in
+ | None, x -> x | x, None -> Option.map (fun y -> CAst.make ~loc:!@loc y) x
+ | Some _, Some _ -> user_err Pp.(str "scope declared twice") in
List.map (fun (name,recarg_like,notation_scope) ->
`Id { name=name; recarg_like=recarg_like;
notation_scope=f notation_scope;
implicit_status = NotImplicit}) items
| "["; items = LIST1 argument_spec; "]"; sc = OPT scope ->
let f x = match sc, x with
- | None, x -> x | x, None -> Option.map (fun y -> !@loc, y) x
- | Some _, Some _ -> error "scope declared twice" in
+ | None, x -> x | x, None -> Option.map (fun y -> CAst.make ~loc:!@loc y) x
+ | Some _, Some _ -> user_err Pp.(str "scope declared twice") in
List.map (fun (name,recarg_like,notation_scope) ->
`Id { name=name; recarg_like=recarg_like;
notation_scope=f notation_scope;
implicit_status = Implicit}) items
| "{"; items = LIST1 argument_spec; "}"; sc = OPT scope ->
let f x = match sc, x with
- | None, x -> x | x, None -> Option.map (fun y -> !@loc, y) x
- | Some _, Some _ -> error "scope declared twice" in
+ | None, x -> x | x, None -> Option.map (fun y -> CAst.make ~loc:!@loc y) x
+ | Some _, Some _ -> user_err Pp.(str "scope declared twice") in
List.map (fun (name,recarg_like,notation_scope) ->
`Id { name=name; recarg_like=recarg_like;
notation_scope=f notation_scope;
implicit_status = MaximallyImplicit}) items
]
];
- name_or_bang: [
- [ b = OPT "!"; id = name ->
- not (Option.is_empty b), id
- ]
- ];
(* Same as [argument_spec_block], but with only implicit status and names *)
more_implicits_block: [
- [ (bang,name) = name_or_bang -> (bang, [(snd name, Vernacexpr.NotImplicit)])
- | "/" -> (true (* Should warn about deprecated syntax *), [])
- | "["; items = LIST1 name_or_bang; "]" ->
- (List.exists fst items, List.map (fun (_,(_,name)) -> (name, Vernacexpr.Implicit)) items)
- | "{"; items = LIST1 name_or_bang; "}" ->
- (List.exists fst items, List.map (fun (_,(_,name)) -> (name, Vernacexpr.MaximallyImplicit)) items)
+ [ name = name -> [(name.CAst.v, Vernacexpr.NotImplicit)]
+ | "["; items = LIST1 name; "]" ->
+ List.map (fun name -> (name.CAst.v, Vernacexpr.Implicit)) items
+ | "{"; items = LIST1 name; "}" ->
+ List.map (fun name -> (name.CAst.v, Vernacexpr.MaximallyImplicit)) items
]
];
strategy_level:
@@ -799,10 +799,10 @@ GEXTEND Gram
| IDENT "transparent" -> Conv_oracle.transparent ] ]
;
instance_name:
- [ [ name = pidentref; sup = OPT binders ->
- (let ((loc,id),l) = name in ((loc, Name id),l)),
+ [ [ name = ident_decl; sup = OPT binders ->
+ (CAst.map (fun id -> Name id) (fst name), snd name),
(Option.default [] sup)
- | -> ((!@loc, Anonymous), None), [] ] ]
+ | -> ((CAst.make ~loc:!@loc Anonymous), None), [] ] ]
;
hint_info:
[ [ "|"; i = OPT natural; pat = OPT constr_pattern ->
@@ -822,12 +822,19 @@ GEXTEND Gram
END
GEXTEND Gram
- GLOBAL: command query_command class_rawexpr;
+ GLOBAL: command query_command class_rawexpr gallina_ext;
+
+ gallina_ext:
+ [ [ IDENT "Export"; "Set"; table = option_table; v = option_value ->
+ VernacSetOption (true, table, v)
+ | IDENT "Export"; IDENT "Unset"; table = option_table ->
+ VernacUnsetOption (true, table)
+ ] ];
command:
[ [ IDENT "Comments"; l = LIST0 comment -> VernacComments l
- (* Hack! Should be in grammar_ext, but camlp4 factorize badly *)
+ (* Hack! Should be in grammar_ext, but camlp5 factorizes badly *)
| IDENT "Declare"; IDENT "Instance"; namesup = instance_name; ":";
expl = [ "!" -> Decl_kinds.Implicit | -> Decl_kinds.Explicit ] ; t = operconstr LEVEL "200";
info = hint_info ->
@@ -871,7 +878,7 @@ GEXTEND Gram
(* Printing (careful factorization of entries) *)
| IDENT "Print"; p = printable -> VernacPrint p
- | IDENT "Print"; qid = smart_global -> VernacPrint (PrintName qid)
+ | IDENT "Print"; qid = smart_global; l = OPT univ_name_list -> VernacPrint (PrintName (qid,l))
| IDENT "Print"; IDENT "Module"; "Type"; qid = global ->
VernacPrint (PrintModuleType qid)
| IDENT "Print"; IDENT "Module"; qid = global ->
@@ -887,24 +894,9 @@ GEXTEND Gram
(* For acting on parameter tables *)
| "Set"; table = option_table; v = option_value ->
- begin match v with
- | StringValue s ->
- (* We make a special case for warnings because appending is their
- natural semantics *)
- if CString.List.equal table ["Warnings"] then
- VernacSetAppendOption (table, s)
- else
- let (last, prefix) = List.sep_last table in
- if String.equal last "Append" && not (List.is_empty prefix) then
- VernacSetAppendOption (prefix, s)
- else
- VernacSetOption (table, v)
- | _ -> VernacSetOption (table, v)
- end
- | "Set"; table = option_table ->
- VernacSetOption (table,BoolValue true)
+ VernacSetOption (false, table, v)
| IDENT "Unset"; table = option_table ->
- VernacUnsetOption table
+ VernacUnsetOption (false, table)
| IDENT "Print"; IDENT "Table"; table = option_table ->
VernacPrintOption table
@@ -929,33 +921,34 @@ GEXTEND Gram
VernacRemoveOption ([table], v) ]]
;
query_command: (* TODO: rapprocher Eval et Check *)
- [ [ IDENT "Eval"; r = Tactic.red_expr; "in"; c = lconstr ->
+ [ [ IDENT "Eval"; r = red_expr; "in"; c = lconstr; "." ->
fun g -> VernacCheckMayEval (Some r, g, c)
- | IDENT "Compute"; c = lconstr ->
+ | IDENT "Compute"; c = lconstr; "." ->
fun g -> VernacCheckMayEval (Some (Genredexpr.CbvVm None), g, c)
- | IDENT "Check"; c = lconstr ->
+ | IDENT "Check"; c = lconstr; "." ->
fun g -> VernacCheckMayEval (None, g, c)
(* Searching the environment *)
- | IDENT "About"; qid = smart_global ->
- fun g -> VernacPrint (PrintAbout (qid,g))
- | IDENT "SearchHead"; c = constr_pattern; l = in_or_out_modules ->
+ | IDENT "About"; qid = smart_global; l = OPT univ_name_list; "." ->
+ fun g -> VernacPrint (PrintAbout (qid,l,g))
+ | IDENT "SearchHead"; c = constr_pattern; l = in_or_out_modules; "." ->
fun g -> VernacSearch (SearchHead c,g, l)
- | IDENT "SearchPattern"; c = constr_pattern; l = in_or_out_modules ->
+ | IDENT "SearchPattern"; c = constr_pattern; l = in_or_out_modules; "." ->
fun g -> VernacSearch (SearchPattern c,g, l)
- | IDENT "SearchRewrite"; c = constr_pattern; l = in_or_out_modules ->
+ | IDENT "SearchRewrite"; c = constr_pattern; l = in_or_out_modules; "." ->
fun g -> VernacSearch (SearchRewrite c,g, l)
- | IDENT "Search"; s = searchabout_query; l = searchabout_queries ->
+ | IDENT "Search"; s = searchabout_query; l = searchabout_queries; "." ->
let (sl,m) = l in fun g -> VernacSearch (SearchAbout (s::sl),g, m)
(* compatibility: SearchAbout *)
- | IDENT "SearchAbout"; s = searchabout_query; l = searchabout_queries ->
+ | IDENT "SearchAbout"; s = searchabout_query; l = searchabout_queries; "." ->
fun g -> let (sl,m) = l in VernacSearch (SearchAbout (s::sl),g, m)
(* compatibility: SearchAbout with "[ ... ]" *)
| IDENT "SearchAbout"; "["; sl = LIST1 searchabout_query; "]";
- l = in_or_out_modules -> fun g -> VernacSearch (SearchAbout sl,g, l)
+ l = in_or_out_modules; "." ->
+ fun g -> VernacSearch (SearchAbout sl,g, l)
] ]
;
printable:
- [ [ IDENT "Term"; qid = smart_global -> PrintName qid
+ [ [ IDENT "Term"; qid = smart_global; l = OPT univ_name_list -> PrintName (qid,l)
| IDENT "All" -> PrintFullContext
| IDENT "Section"; s = global -> PrintSectionContext s
| IDENT "Grammar"; ent = IDENT ->
@@ -963,7 +956,7 @@ GEXTEND Gram
PrintGrammar ent
| IDENT "LoadPath"; dir = OPT dirpath -> PrintLoadPath dir
| IDENT "Modules" ->
- error "Print Modules is obsolete; use Print Libraries instead"
+ user_err Pp.(str "Print Modules is obsolete; use Print Libraries instead")
| IDENT "Libraries" -> PrintModules
| IDENT "ML"; IDENT "Path" -> PrintMLLoadPath
@@ -1006,11 +999,11 @@ GEXTEND Gram
| IDENT "Term"; qid = smart_global -> LocateTerm qid
| IDENT "File"; f = ne_string -> LocateFile f
| IDENT "Library"; qid = global -> LocateLibrary qid
- | IDENT "Module"; qid = global -> LocateModule qid
- | IDENT "Ltac"; qid = global -> LocateTactic qid ] ]
+ | IDENT "Module"; qid = global -> LocateModule qid ] ]
;
option_value:
- [ [ n = integer -> IntValue (Some n)
+ [ [ -> BoolValue true
+ | n = integer -> IntValue (Some n)
| s = STRING -> StringValue s ] ]
;
option_ref_value:
@@ -1056,6 +1049,9 @@ GEXTEND Gram
| -> ([],SearchOutside [])
] ]
;
+ univ_name_list:
+ [ [ "@{" ; l = LIST0 name; "}" -> l ] ]
+ ;
END;
GEXTEND Gram
@@ -1078,15 +1074,15 @@ GEXTEND Gram
(* Tactic Debugger *)
| IDENT "Debug"; IDENT "On" ->
- VernacSetOption (["Ltac";"Debug"], BoolValue true)
+ VernacSetOption (false, ["Ltac";"Debug"], BoolValue true)
| IDENT "Debug"; IDENT "Off" ->
- VernacSetOption (["Ltac";"Debug"], BoolValue false)
+ VernacSetOption (false, ["Ltac";"Debug"], BoolValue false)
(* registration of a custom reduction *)
| IDENT "Declare"; IDENT "Reduction"; s = IDENT; ":=";
- r = Tactic.red_expr ->
+ r = red_expr ->
VernacDeclareReduction (s,r)
] ];
@@ -1099,11 +1095,11 @@ GEXTEND Gram
GLOBAL: syntax;
syntax:
- [ [ IDENT "Open"; local = obsolete_locality; IDENT "Scope"; sc = IDENT ->
- VernacOpenCloseScope (local,(true,sc))
+ [ [ IDENT "Open"; IDENT "Scope"; sc = IDENT ->
+ VernacOpenCloseScope (true,sc)
- | IDENT "Close"; local = obsolete_locality; IDENT "Scope"; sc = IDENT ->
- VernacOpenCloseScope (local,(false,sc))
+ | IDENT "Close"; IDENT "Scope"; sc = IDENT ->
+ VernacOpenCloseScope (false,sc)
| IDENT "Delimit"; IDENT "Scope"; sc = IDENT; "with"; key = IDENT ->
VernacDelimiters (sc, Some key)
@@ -1113,33 +1109,31 @@ GEXTEND Gram
| IDENT "Bind"; IDENT "Scope"; sc = IDENT; "with";
refl = LIST1 class_rawexpr -> VernacBindScope (sc,refl)
- | IDENT "Infix"; local = obsolete_locality;
- op = ne_lstring; ":="; p = constr;
+ | IDENT "Infix"; op = ne_lstring; ":="; p = constr;
modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ];
sc = OPT [ ":"; sc = IDENT -> sc ] ->
- VernacInfix (local,(op,modl),p,sc)
- | IDENT "Notation"; local = obsolete_locality; id = identref;
+ VernacInfix ((op,modl),p,sc)
+ | IDENT "Notation"; id = identref;
idl = LIST0 ident; ":="; c = constr; b = only_parsing ->
VernacSyntacticDefinition
- (id,(idl,c),local,b)
- | IDENT "Notation"; local = obsolete_locality; s = ne_lstring; ":=";
+ (id,(idl,c),b)
+ | IDENT "Notation"; s = lstring; ":=";
c = constr;
modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ];
sc = OPT [ ":"; sc = IDENT -> sc ] ->
- VernacNotation (local,c,(s,modl),sc)
+ VernacNotation (c,(s,modl),sc)
| IDENT "Format"; IDENT "Notation"; n = STRING; s = STRING; fmt = STRING ->
VernacNotationAddFormat (n,s,fmt)
| IDENT "Reserved"; IDENT "Infix"; s = ne_lstring;
l = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ] ->
- Metasyntax.check_infix_modifiers l;
- let (loc,s) = s in
- VernacSyntaxExtension (false,((loc,"x '"^s^"' y"),l))
+ let s = CAst.map (fun s -> "x '"^s^"' y") s in
+ VernacSyntaxExtension (true,(s,l))
- | IDENT "Reserved"; IDENT "Notation"; local = obsolete_locality;
+ | IDENT "Reserved"; IDENT "Notation";
s = ne_lstring;
l = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ]
- -> VernacSyntaxExtension (local,(s,l))
+ -> VernacSyntaxExtension (false, (s,l))
(* "Print" "Grammar" should be here but is in "command" entry in order
to factorize with other "Print"-based vernac entries *)
@@ -1149,12 +1143,9 @@ GEXTEND Gram
[ [ "("; IDENT "only"; IDENT "parsing"; ")" ->
Some Flags.Current
| "("; IDENT "compat"; s = STRING; ")" ->
- Some (Coqinit.get_compat_version s)
+ Some (parse_compat_version s)
| -> None ] ]
;
- obsolete_locality:
- [ [ IDENT "Local" -> true | -> false ] ]
- ;
level:
[ [ IDENT "level"; n = natural -> NumLevel n
| IDENT "next"; IDENT "level" -> NextLevel ] ]
@@ -1167,15 +1158,17 @@ GEXTEND Gram
| IDENT "only"; IDENT "printing" -> SetOnlyPrinting
| IDENT "only"; IDENT "parsing" -> SetOnlyParsing
| IDENT "compat"; s = STRING ->
- SetCompatVersion (Coqinit.get_compat_version s)
- | IDENT "format"; s1 = [s = STRING -> (!@loc,s)];
- s2 = OPT [s = STRING -> (!@loc,s)] ->
+ SetCompatVersion (parse_compat_version s)
+ | IDENT "format"; s1 = [s = STRING -> CAst.make ~loc:!@loc s];
+ s2 = OPT [s = STRING -> CAst.make ~loc:!@loc s] ->
begin match s1, s2 with
- | (_,k), Some s -> SetFormat(k,s)
+ | { CAst.v = k }, Some s -> SetFormat(k,s)
| s, None -> SetFormat ("text",s) end
| x = IDENT; ","; l = LIST1 [id = IDENT -> id ] SEP ","; "at";
lev = level -> SetItemLevel (x::l,lev)
| x = IDENT; "at"; lev = level -> SetItemLevel ([x],lev)
+ | x = IDENT; "at"; lev = level; b = constr_as_binder_kind -> SetItemLevelAsBinder ([x],b,Some lev)
+ | x = IDENT; b = constr_as_binder_kind -> SetItemLevelAsBinder ([x],b,None)
| x = IDENT; typ = syntax_extension_type -> SetEntryType (x,typ)
] ]
;
@@ -1183,7 +1176,20 @@ GEXTEND Gram
[ [ IDENT "ident" -> ETName | IDENT "global" -> ETReference
| IDENT "bigint" -> ETBigint
| IDENT "binder" -> ETBinder true
+ | IDENT "constr"; n = OPT at_level; b = constr_as_binder_kind -> ETConstrAsBinder (b,n)
+ | IDENT "pattern" -> ETPattern (false,None)
+ | IDENT "pattern"; "at"; IDENT "level"; n = natural -> ETPattern (false,Some n)
+ | IDENT "strict"; IDENT "pattern" -> ETPattern (true,None)
+ | IDENT "strict"; IDENT "pattern"; "at"; IDENT "level"; n = natural -> ETPattern (true,Some n)
| IDENT "closed"; IDENT "binder" -> ETBinder false
] ]
;
+ at_level:
+ [ [ "at"; n = level -> n ] ]
+ ;
+ constr_as_binder_kind:
+ [ [ "as"; IDENT "ident" -> AsIdent
+ | "as"; IDENT "pattern" -> AsIdentOrPattern
+ | "as"; IDENT "strict"; IDENT "pattern" -> AsStrictPattern ] ]
+ ;
END