aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--CHANGES2
-rw-r--r--doc/refman/RefMan-ext.tex6
-rw-r--r--parsing/g_vernac.ml423
-rw-r--r--parsing/ppvernac.ml10
-rw-r--r--toplevel/vernacentries.ml12
-rw-r--r--toplevel/vernacexpr.ml2
6 files changed, 42 insertions, 13 deletions
diff --git a/CHANGES b/CHANGES
index 63900a2b76..eb906205e9 100644
--- a/CHANGES
+++ b/CHANGES
@@ -163,6 +163,8 @@ Vernacular commands
variables before the colon at the start of interactive proofs.
- New command "Declare Reduction <id> := <conv_expr>", allowing to write
later "Eval <id> in ...". This command accepts a Local variant.
+- Syntax of Implicit Type now supports more than one block of variables of
+ a given type.
Library
diff --git a/doc/refman/RefMan-ext.tex b/doc/refman/RefMan-ext.tex
index 42374e368f..01df199ab9 100644
--- a/doc/refman/RefMan-ext.tex
+++ b/doc/refman/RefMan-ext.tex
@@ -1585,6 +1585,12 @@ Lemma cons_inj_bool : forall (m n:bool) l, n :: l = m :: l -> n = m.
\begin{Variants}
\item {\tt Implicit Type {\ident} : {\type}}\\
This is useful for declaring the implicit type of a single variable.
+\item
+ {\tt Implicit Types\,%
+(\,{\ident$_{1,1}$}\ldots{\ident$_{1,k_1}$}\,{\tt :}\,{\term$_1$} {\tt )}\,%
+\ldots\,{\tt (}\,{\ident$_{n,1}$}\ldots{\ident$_{n,k_n}$}\,{\tt :}\,%
+{\term$_n$} {\tt )}.}\\
+ Adds $n$ blocks of implicit types with different specifications.
\end{Variants}
diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4
index a0f9dcaefa..008a6ac510 100644
--- a/parsing/g_vernac.ml4
+++ b/parsing/g_vernac.ml4
@@ -120,6 +120,12 @@ let test_plurial_form = function
"Keywords Variables/Hypotheses/Parameters expect more than one assumption"
| _ -> ()
+let test_plurial_form_types = function
+ | [([_],_)] ->
+ Flags.if_verbose warning
+ "Keywords Implicit Types expect more than one type"
+ | _ -> ()
+
(* Gallina declarations *)
GEXTEND Gram
GLOBAL: gallina gallina_ext thm_token def_body of_type_with_opt_coercion
@@ -546,8 +552,12 @@ GEXTEND Gram
List.map (fun (id,b,f) -> (ExplByName id,b,f)) l ] ->
VernacDeclareImplicits (use_section_locality (),qid,pos)
- | IDENT "Implicit"; ["Type" | IDENT "Types"];
- idl = LIST1 identref; ":"; c = lconstr -> VernacReserve (idl,c)
+ | IDENT "Implicit"; "Type"; bl = reserv_list ->
+ VernacReserve bl
+
+ | IDENT "Implicit"; IDENT "Types"; bl = reserv_list ->
+ test_plurial_form_types bl;
+ VernacReserve bl
| IDENT "Generalizable";
gen = [IDENT "All"; IDENT "Variables" -> Some []
@@ -575,6 +585,15 @@ GEXTEND Gram
(Option.default [] sup)
| -> (loc, Anonymous), [] ] ]
;
+ reserv_list:
+ [ [ bl = LIST1 reserv_tuple -> bl | b = simple_reserv -> [b] ] ]
+ ;
+ reserv_tuple:
+ [ [ "("; a = simple_reserv; ")" -> a ] ]
+ ;
+ simple_reserv:
+ [ [ idl = LIST1 identref; ":"; c = lconstr -> (idl,c) ] ]
+ ;
END
diff --git a/parsing/ppvernac.ml b/parsing/ppvernac.ml
index ff480a5d44..9ebf77adba 100644
--- a/parsing/ppvernac.ml
+++ b/parsing/ppvernac.ml
@@ -851,11 +851,11 @@ let rec pr_vernac = function
hov 1 (pr_section_locality local ++ str"Implicit Arguments " ++
spc() ++ pr_smart_global q ++ spc() ++
str"[" ++ prlist_with_sep sep pr_explanation imps ++ str"]")
- | VernacReserve (idl,c) ->
- hov 1 (str"Implicit Type" ++
- str (if List.length idl > 1 then "s " else " ") ++
- prlist_with_sep spc pr_lident idl ++ str " :" ++ spc () ++
- pr_lconstr c)
+ | VernacReserve bl ->
+ let n = List.length (List.flatten (List.map fst bl)) in
+ hov 2 (str"Implicit Type" ++
+ str (if n > 1 then "s " else " ") ++
+ pr_ne_params_list pr_lconstr_expr (List.map (fun sb -> false,sb) bl))
| VernacGeneralizable (local, g) ->
hov 1 (pr_locality local ++ str"Generalizable Variable" ++
match g with
diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml
index d77257a0b7..886f4033d8 100644
--- a/toplevel/vernacentries.ml
+++ b/toplevel/vernacentries.ml
@@ -771,10 +771,12 @@ let vernac_declare_implicits local r = function
| None ->
Impargs.declare_implicits local (smart_global r)
-let vernac_reserve idl c =
- let t = Constrintern.interp_type Evd.empty (Global.env()) c in
- let t = Detyping.detype false [] [] t in
- List.iter (fun id -> Reserve.declare_reserved_type id t) idl
+let vernac_reserve bl =
+ let sb_decl = (fun (idl,c) ->
+ let t = Constrintern.interp_type Evd.empty (Global.env()) c in
+ let t = Detyping.detype false [] [] t in
+ List.iter (fun id -> Reserve.declare_reserved_type id t) idl)
+ in List.iter sb_decl bl
let vernac_generalizable = Implicit_quantifiers.declare_generalizable
@@ -1386,7 +1388,7 @@ let interp c = match c with
| VernacHints (local,dbnames,hints) -> vernac_hints local dbnames hints
| VernacSyntacticDefinition (id,c,l,b) ->vernac_syntactic_definition id c l b
| VernacDeclareImplicits (local,qid,l) ->vernac_declare_implicits local qid l
- | VernacReserve (idl,c) -> vernac_reserve idl c
+ | VernacReserve bl -> vernac_reserve bl
| VernacGeneralizable (local,gen) -> vernac_generalizable local gen
| VernacSetOpacity (local,qidl) -> vernac_set_opacity local qidl
| VernacSetOption (locality,key,v) -> vernac_set_option locality key v
diff --git a/toplevel/vernacexpr.ml b/toplevel/vernacexpr.ml
index 4f7beeb0b2..0db000a485 100644
--- a/toplevel/vernacexpr.ml
+++ b/toplevel/vernacexpr.ml
@@ -312,7 +312,7 @@ type vernac_expr =
locality_flag * onlyparsing_flag
| VernacDeclareImplicits of locality_flag * reference or_by_notation *
(explicitation * bool * bool) list option
- | VernacReserve of lident list * constr_expr
+ | VernacReserve of simple_binder list
| VernacGeneralizable of locality_flag * (lident list) option
| VernacSetOpacity of
locality_flag * (Conv_oracle.level * reference or_by_notation list) list