summaryrefslogtreecommitdiff
path: root/parsing
diff options
context:
space:
mode:
authorGravatar Stephane Glondu <steph@glondu.net>2010-07-21 09:46:51 +0200
committerGravatar Stephane Glondu <steph@glondu.net>2010-07-21 09:46:51 +0200
commit5b7eafd0f00a16d78f99a27f5c7d5a0de77dc7e6 (patch)
tree631ad791a7685edafeb1fb2e8faeedc8379318ae /parsing
parentda178a880e3ace820b41d38b191d3785b82991f5 (diff)
Imported Upstream snapshot 8.3~beta0+13298
Diffstat (limited to 'parsing')
-rw-r--r--parsing/argextend.ml457
-rw-r--r--parsing/egrammar.ml261
-rw-r--r--parsing/egrammar.mli57
-rw-r--r--parsing/extend.ml62
-rw-r--r--parsing/extend.mli55
-rw-r--r--parsing/extrawit.ml62
-rw-r--r--parsing/extrawit.mli51
-rw-r--r--parsing/g_ascii_syntax.ml83
-rw-r--r--parsing/g_constr.ml4109
-rw-r--r--parsing/g_decl_mode.ml498
-rw-r--r--parsing/g_intsyntax.ml344
-rw-r--r--parsing/g_ltac.ml426
-rw-r--r--parsing/g_minicoq.mli31
-rw-r--r--parsing/g_natsyntax.ml78
-rw-r--r--parsing/g_natsyntax.mli2
-rw-r--r--parsing/g_prim.ml424
-rw-r--r--parsing/g_proofs.ml428
-rw-r--r--parsing/g_rsyntax.ml125
-rw-r--r--parsing/g_string_syntax.ml69
-rw-r--r--parsing/g_tactic.ml4182
-rw-r--r--parsing/g_vernac.ml4504
-rw-r--r--parsing/g_xml.ml429
-rw-r--r--parsing/g_zsyntax.ml194
-rw-r--r--parsing/g_zsyntax.mli2
-rw-r--r--parsing/grammar.mllib84
-rw-r--r--parsing/highparsing.mllib7
-rw-r--r--parsing/lexer.ml4157
-rw-r--r--parsing/lexer.mli12
-rw-r--r--parsing/parsing.mllib12
-rw-r--r--parsing/pcoq.ml4369
-rw-r--r--parsing/pcoq.mli245
-rw-r--r--parsing/ppconstr.ml141
-rw-r--r--parsing/ppconstr.mli40
-rw-r--r--parsing/ppdecl_proof.ml122
-rw-r--r--parsing/ppdecl_proof.mli2
-rw-r--r--parsing/pptactic.ml317
-rw-r--r--parsing/pptactic.mli33
-rw-r--r--parsing/ppvernac.ml508
-rw-r--r--parsing/ppvernac.mli4
-rw-r--r--parsing/prettyp.ml281
-rw-r--r--parsing/prettyp.mli9
-rw-r--r--parsing/printer.ml156
-rw-r--r--parsing/printer.mli12
-rw-r--r--parsing/printmod.ml95
-rw-r--r--parsing/q_constr.ml436
-rw-r--r--parsing/q_coqast.ml463
-rw-r--r--parsing/q_util.ml480
-rw-r--r--parsing/q_util.mli5
-rw-r--r--parsing/search.ml214
-rw-r--r--parsing/search.mli50
-rw-r--r--parsing/tacextend.ml481
-rw-r--r--parsing/tactic_printer.ml79
-rw-r--r--parsing/tactic_printer.mli2
-rw-r--r--parsing/vernacextend.ml462
54 files changed, 2514 insertions, 3297 deletions
diff --git a/parsing/argextend.ml4 b/parsing/argextend.ml4
index bd6be424..89edbb12 100644
--- a/parsing/argextend.ml4
+++ b/parsing/argextend.ml4
@@ -8,11 +8,12 @@
(*i camlp4use: "pa_extend.cmo q_MLast.cmo" i*)
-(* $Id: argextend.ml4 11622 2008-11-23 08:45:56Z herbelin $ *)
+(* $Id$ *)
open Genarg
open Q_util
open Q_coqast
+open Egrammar
let join_loc = Util.join_loc
let loc = Util.dummy_loc
@@ -39,7 +40,7 @@ let rec make_rawwit loc = function
| List0ArgType t -> <:expr< Genarg.wit_list0 $make_rawwit loc t$ >>
| List1ArgType t -> <:expr< Genarg.wit_list1 $make_rawwit loc t$ >>
| OptArgType t -> <:expr< Genarg.wit_opt $make_rawwit loc t$ >>
- | PairArgType (t1,t2) ->
+ | PairArgType (t1,t2) ->
<:expr< Genarg.wit_pair $make_rawwit loc t1$ $make_rawwit loc t2$ >>
| ExtraArgType s -> <:expr< $lid:"rawwit_"^s$ >>
@@ -64,7 +65,7 @@ let rec make_globwit loc = function
| List0ArgType t -> <:expr< Genarg.wit_list0 $make_globwit loc t$ >>
| List1ArgType t -> <:expr< Genarg.wit_list1 $make_globwit loc t$ >>
| OptArgType t -> <:expr< Genarg.wit_opt $make_globwit loc t$ >>
- | PairArgType (t1,t2) ->
+ | PairArgType (t1,t2) ->
<:expr< Genarg.wit_pair $make_globwit loc t1$ $make_globwit loc t2$ >>
| ExtraArgType s -> <:expr< $lid:"globwit_"^s$ >>
@@ -89,25 +90,31 @@ let rec make_wit loc = function
| List0ArgType t -> <:expr< Genarg.wit_list0 $make_wit loc t$ >>
| List1ArgType t -> <:expr< Genarg.wit_list1 $make_wit loc t$ >>
| OptArgType t -> <:expr< Genarg.wit_opt $make_wit loc t$ >>
- | PairArgType (t1,t2) ->
+ | PairArgType (t1,t2) ->
<:expr< Genarg.wit_pair $make_wit loc t1$ $make_wit loc t2$ >>
| ExtraArgType s -> <:expr< $lid:"wit_"^s$ >>
let make_act loc act pil =
let rec make = function
| [] -> <:expr< Gramext.action (fun loc -> ($act$ : 'a)) >>
- | None :: tl -> <:expr< Gramext.action (fun _ -> $make tl$) >>
- | Some (p, t) :: tl ->
+ | GramNonTerminal (_,t,_,Some p) :: tl ->
+ let p = Names.string_of_id p in
<:expr<
- Gramext.action
+ Gramext.action
(fun $lid:p$ ->
let _ = Genarg.in_gen $make_rawwit loc t$ $lid:p$ in $make tl$)
- >> in
+ >>
+ | (GramTerminal _ | GramNonTerminal (_,_,_,None)) :: tl ->
+ <:expr< Gramext.action (fun _ -> $make tl$) >> in
make (List.rev pil)
+let make_prod_item = function
+ | GramTerminal s -> <:expr< (Gramext.Stoken (Lexer.terminal $str:s$)) >>
+ | GramNonTerminal (_,_,g,_) ->
+ <:expr< Pcoq.symbol_of_prod_entry_key $mlexpr_of_prod_entry_key g$ >>
+
let make_rule loc (prods,act) =
- let (symbs,pil) = List.split prods in
- <:expr< ($mlexpr_of_list (fun x -> x) symbs$,$make_act loc act pil$) >>
+ <:expr< ($mlexpr_of_list make_prod_item prods$,$make_act loc act prods$) >>
let declare_tactic_argument loc s typ pr f g h rawtyppr globtyppr cl =
let rawtyp, rawpr = match rawtyppr with
@@ -124,14 +131,14 @@ let declare_tactic_argument loc s typ pr f g h rawtyppr globtyppr cl =
(Genarg.in_gen $make_rawwit loc rawtyp$ x)) >>
| Some f -> <:expr< $lid:f$>> in
let interp = match f with
- | None ->
+ | None ->
<:expr< fun ist gl x ->
out_gen $make_wit loc typ$
(Tacinterp.interp_genarg ist gl
(Genarg.in_gen $make_globwit loc globtyp$ x)) >>
| Some f -> <:expr< $lid:f$>> in
let substitute = match h with
- | None ->
+ | None ->
<:expr< fun s x ->
out_gen $make_globwit loc globtyp$
(Tacinterp.subst_genarg s
@@ -144,6 +151,8 @@ let declare_tactic_argument loc s typ pr f g h rawtyppr globtyppr cl =
let rules = mlexpr_of_list (make_rule loc) (List.rev cl) in
<:str_item<
declare
+ open Pcoq;
+ open Extrawit;
value ($lid:"wit_"^s$, $lid:"globwit_"^s$, $lid:"rawwit_"^s$) =
Genarg.create_arg $se$;
value $lid:s$ = Pcoq.create_generic_entry $se$ $rawwit$;
@@ -154,7 +163,7 @@ let declare_tactic_argument loc s typ pr f g h rawtyppr globtyppr cl =
(Genarg.in_gen $wit$ ($interp$ ist gl (out_gen $globwit$ x)))),
(fun subst x ->
(Genarg.in_gen $globwit$ ($substitute$ subst (out_gen $globwit$ x)))));
- Pcoq.Gram.extend ($lid:s$ : Pcoq.Gram.Entry.e 'a) None
+ Pcoq.Gram.extend ($lid:s$ : Pcoq.Gram.Entry.e 'a) None
[(None, None, $rules$)];
Pptactic.declare_extra_genarg_pprule
($rawwit$, $lid:rawpr$)
@@ -174,11 +183,13 @@ let declare_vernac_argument loc s pr cl =
| Some pr -> <:expr< fun _ _ _ -> $lid:pr$ >> in
<:str_item<
declare
+ open Pcoq;
+ open Extrawit;
value (($lid:"wit_"^s$:Genarg.abstract_argument_type unit Genarg.tlevel),
($lid:"globwit_"^s$:Genarg.abstract_argument_type unit Genarg.glevel),
$lid:"rawwit_"^s$) = Genarg.create_arg $se$;
value $lid:s$ = Pcoq.create_generic_entry $se$ $rawwit$;
- Pcoq.Gram.extend ($lid:s$ : Pcoq.Gram.Entry.e 'a) None
+ Pcoq.Gram.extend ($lid:s$ : Pcoq.Gram.Entry.e 'a) None
[(None, None, $rules$)];
Pptactic.declare_extra_genarg_pprule
($rawwit$, $pr_rules$)
@@ -202,10 +213,10 @@ EXTEND
h = OPT [ "SUBSTITUTED"; "BY"; f = LIDENT -> f ];
rawtyppr =
(* Necessary if the globalized type is different from the final type *)
- OPT [ "RAW_TYPED"; "AS"; t = argtype;
+ OPT [ "RAW_TYPED"; "AS"; t = argtype;
"RAW_PRINTED"; "BY"; pr = LIDENT -> (t,pr) ];
globtyppr =
- OPT [ "GLOB_TYPED"; "AS"; t = argtype;
+ OPT [ "GLOB_TYPED"; "AS"; t = argtype;
"GLOB_PRINTED"; "BY"; pr = LIDENT -> (t,pr) ];
OPT "|"; l = LIST1 argrule SEP "|";
"END" ->
@@ -221,13 +232,13 @@ EXTEND
declare_vernac_argument loc s pr l ] ]
;
argtype:
- [ "2"
+ [ "2"
[ e1 = argtype; "*"; e2 = argtype -> PairArgType (e1, e2) ]
| "1"
[ e = argtype; LIDENT "list" -> List0ArgType e
| e = argtype; LIDENT "option" -> OptArgType e ]
| "0"
- [ e = LIDENT -> fst (interp_entry_name loc e "")
+ [ e = LIDENT -> fst (interp_entry_name false None e "")
| "("; e = argtype; ")" -> e ] ]
;
argrule:
@@ -235,13 +246,15 @@ EXTEND
;
genarg:
[ [ e = LIDENT; "("; s = LIDENT; ")" ->
- let t, g = interp_entry_name loc e "" in (g, Some (s,t))
+ let t, g = interp_entry_name false None e "" in
+ GramNonTerminal (loc, t, g, Some (Names.id_of_string s))
| e = LIDENT; "("; s = LIDENT; ","; sep = STRING; ")" ->
- let t, g = interp_entry_name loc e sep in (g, Some (s,t))
+ let t, g = interp_entry_name false None e sep in
+ GramNonTerminal (loc, t, g, Some (Names.id_of_string s))
| s = STRING ->
if String.length s > 0 && Util.is_letter s.[0] then
- Compat.using Pcoq.lexer ("", s);
- (<:expr< (Gramext.Stoken (Lexer.terminal $str:s$)) >>, None)
+ Lexer.add_token ("", s);
+ GramTerminal s
] ]
;
END
diff --git a/parsing/egrammar.ml b/parsing/egrammar.ml
index 43836dbb..023ec0f3 100644
--- a/parsing/egrammar.ml
+++ b/parsing/egrammar.ml
@@ -6,12 +6,13 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: egrammar.ml 11512 2008-10-27 12:28:36Z herbelin $ *)
+(* $Id$ *)
open Pp
open Util
open Pcoq
open Extend
+open Ppextend
open Topconstr
open Genarg
open Libnames
@@ -21,9 +22,9 @@ open Names
open Vernacexpr
(**************************************************************************)
-(*
+(*
* --- Note on the mapping of grammar productions to camlp4 actions ---
- *
+ *
* Translation of environments: a production
* [ nt1(x1) ... nti(xi) ] -> act(x1..xi)
* is written (with camlp4 conventions):
@@ -33,9 +34,9 @@ open Vernacexpr
* the make_*_action family build the following closure:
*
* ((fun env ->
- * (fun vi ->
+ * (fun vi ->
* (fun env -> ...
- *
+ *
* (fun v1 ->
* (fun env -> gram_action .. env act)
* ((x1,v1)::env))
@@ -47,69 +48,106 @@ open Vernacexpr
(**********************************************************************)
(** Declare Notations grammar rules *)
-type prod_item =
- | Term of Token.pattern
- | NonTerm of constr_production_entry *
- (Names.identifier * constr_production_entry) option
+let constr_expr_of_name (loc,na) = match na with
+ | Anonymous -> CHole (loc,None)
+ | Name id -> CRef (Ident (loc,id))
+
+let cases_pattern_expr_of_name (loc,na) = match na with
+ | Anonymous -> CPatAtom (loc,None)
+ | Name id -> CPatAtom (loc,Some (Ident (loc,id)))
+
+type grammar_constr_prod_item =
+ | GramConstrTerminal of Token.pattern
+ | GramConstrNonTerminal of constr_prod_entry_key * identifier option
+ | GramConstrListMark of int * bool
+ (* tells action rule to make a list of the n previous parsed items;
+ concat with last parsed list if true *)
type 'a action_env = 'a list * 'a list list
let make_constr_action
(f : loc -> constr_expr action_env -> constr_expr) pil =
let rec make (env,envlist as fullenv : constr_expr action_env) = function
- | [] ->
- Gramext.action (fun loc -> f loc fullenv)
- | None :: tl -> (* parse a non-binding item *)
- Gramext.action (fun _ -> make fullenv tl)
- | Some (p, (ETConstr _| ETOther _)) :: tl -> (* constr non-terminal *)
+ | [] ->
+ Gramext.action (fun loc -> f loc fullenv)
+ | (GramConstrTerminal _ | GramConstrNonTerminal (_,None)) :: tl ->
+ (* parse a non-binding item *)
+ Gramext.action (fun _ -> make fullenv tl)
+ | GramConstrNonTerminal (typ, Some _) :: tl ->
+ (* parse a binding non-terminal *)
+ (match typ with
+ | (ETConstr _| ETOther _) ->
Gramext.action (fun (v:constr_expr) -> make (v :: env, envlist) tl)
- | Some (p, ETReference) :: tl -> (* non-terminal *)
+ | ETReference ->
Gramext.action (fun (v:reference) -> make (CRef v :: env, envlist) tl)
- | Some (p, ETIdent) :: tl -> (* non-terminal *)
- Gramext.action (fun (v:identifier) ->
- make (CRef (Ident (dummy_loc,v)) :: env, envlist) tl)
- | Some (p, ETBigint) :: tl -> (* non-terminal *)
+ | ETName ->
+ Gramext.action (fun (na:name located) ->
+ make (constr_expr_of_name na :: env, envlist) tl)
+ | ETBigint ->
Gramext.action (fun (v:Bigint.bigint) ->
make (CPrim (dummy_loc,Numeral v) :: env, envlist) tl)
- | Some (p, ETConstrList _) :: tl ->
- Gramext.action (fun (v:constr_expr list) -> make (env, v::envlist) tl)
- | Some (p, ETPattern) :: tl ->
- failwith "Unexpected entry of type cases pattern" in
+ | ETConstrList (_,n) ->
+ Gramext.action (fun (v:constr_expr list) -> make (env, v::envlist) tl)
+ | ETPattern ->
+ failwith "Unexpected entry of type cases pattern")
+ | GramConstrListMark (n,b) :: tl ->
+ (* Rebuild expansions of ConstrList *)
+ let heads,env = list_chop n env in
+ if b then make (env,(heads@List.hd envlist)::List.tl envlist) tl
+ else make (env,heads::envlist) tl
+ in
make ([],[]) (List.rev pil)
let make_cases_pattern_action
(f : loc -> cases_pattern_expr action_env -> cases_pattern_expr) pil =
let rec make (env,envlist as fullenv : cases_pattern_expr action_env) = function
- | [] ->
- Gramext.action (fun loc -> f loc fullenv)
- | None :: tl -> (* parse a non-binding item *)
- Gramext.action (fun _ -> make fullenv tl)
- | Some (p, ETConstr _) :: tl -> (* pattern non-terminal *)
+ | [] ->
+ Gramext.action (fun loc -> f loc fullenv)
+ | (GramConstrTerminal _ | GramConstrNonTerminal (_,None)) :: tl ->
+ (* parse a non-binding item *)
+ Gramext.action (fun _ -> make fullenv tl)
+ | GramConstrNonTerminal (typ, Some _) :: tl ->
+ (* parse a binding non-terminal *)
+ (match typ with
+ | ETConstr _ -> (* pattern non-terminal *)
Gramext.action (fun (v:cases_pattern_expr) -> make (v::env,envlist) tl)
- | Some (p, ETReference) :: tl -> (* non-terminal *)
+ | ETReference ->
Gramext.action (fun (v:reference) ->
make (CPatAtom (dummy_loc,Some v) :: env, envlist) tl)
- | Some (p, ETIdent) :: tl -> (* non-terminal *)
- Gramext.action (fun (v:identifier) ->
- make
- (CPatAtom (dummy_loc,Some (Ident (dummy_loc,v)))::env, envlist) tl)
- | Some (p, ETBigint) :: tl -> (* non-terminal *)
+ | ETName ->
+ Gramext.action (fun (na:name located) ->
+ make (cases_pattern_expr_of_name na :: env, envlist) tl)
+ | ETBigint ->
Gramext.action (fun (v:Bigint.bigint) ->
make (CPatPrim (dummy_loc,Numeral v) :: env, envlist) tl)
- | Some (p, ETConstrList _) :: tl ->
- Gramext.action (fun (v:cases_pattern_expr list) ->
- make (env, v :: envlist) tl)
- | Some (p, (ETPattern | ETOther _)) :: tl ->
- failwith "Unexpected entry of type cases pattern or other" in
+ | ETConstrList (_,_) ->
+ Gramext.action (fun (vl:cases_pattern_expr list) ->
+ make (env, vl :: envlist) tl)
+ | (ETPattern | ETOther _) ->
+ failwith "Unexpected entry of type cases pattern or other")
+ | GramConstrListMark (n,b) :: tl ->
+ (* Rebuild expansions of ConstrList *)
+ let heads,env = list_chop n env in
+ if b then make (env,(heads@List.hd envlist)::List.tl envlist) tl
+ else make (env,heads::envlist) tl
+ in
make ([],[]) (List.rev pil)
-let make_constr_prod_item univ assoc from forpat = function
- | Term tok -> (Gramext.Stoken tok, None)
- | NonTerm (nt, ovar) ->
- let eobj = symbol_of_production assoc from forpat nt in
- (eobj, ovar)
-
-let prepare_empty_levels entry (pos,p4assoc,name,reinit) =
+let rec make_constr_prod_item assoc from forpat = function
+ | GramConstrTerminal tok :: l ->
+ Gramext.Stoken tok :: make_constr_prod_item assoc from forpat l
+ | GramConstrNonTerminal (nt, ovar) :: l ->
+ symbol_of_constr_prod_entry_key assoc from forpat nt
+ :: make_constr_prod_item assoc from forpat l
+ | GramConstrListMark _ :: l ->
+ make_constr_prod_item assoc from forpat l
+ | [] ->
+ []
+
+let prepare_empty_levels forpat (pos,p4assoc,name,reinit) =
+ let entry =
+ if forpat then weaken_entry Constr.pattern
+ else weaken_entry Constr.operconstr in
grammar_extend entry pos reinit [(name, p4assoc, [])]
let pure_sublevels level symbs =
@@ -119,26 +157,25 @@ let pure_sublevels level symbs =
| _ ->
failwith "") symbs
-let extend_constr (entry,level) (n,assoc) mkact forpat pt =
- let univ = get_univ "constr" in
- let pil = List.map (make_constr_prod_item univ assoc n forpat) pt in
- let (symbs,ntl) = List.split pil in
+let extend_constr (entry,level) (n,assoc) mkact forpat rules =
+ List.iter (fun pt ->
+ let symbs = make_constr_prod_item assoc n forpat pt in
let pure_sublevels = pure_sublevels level symbs in
let needed_levels = register_empty_levels forpat pure_sublevels in
let pos,p4assoc,name,reinit = find_position forpat assoc level in
- List.iter (prepare_empty_levels entry) needed_levels;
- grammar_extend entry pos reinit [(name, p4assoc, [symbs, mkact ntl])]
+ List.iter (prepare_empty_levels forpat) needed_levels;
+ grammar_extend entry pos reinit [(name, p4assoc, [symbs, mkact pt])]) rules
-let extend_constr_notation (n,assoc,ntn,rule) =
+let extend_constr_notation (n,assoc,ntn,rules) =
(* Add the notation in constr *)
let mkact loc env = CNotation (loc,ntn,env) in
- let e = get_constr_entry false (ETConstr (n,())) in
- extend_constr e (ETConstr(n,()),assoc) (make_constr_action mkact) false rule;
+ let e = interp_constr_entry_key false (ETConstr (n,())) in
+ extend_constr e (ETConstr(n,()),assoc) (make_constr_action mkact) false rules;
(* Add the notation in cases_pattern *)
let mkact loc env = CPatNotation (loc,ntn,env) in
- let e = get_constr_entry true (ETConstr (n,())) in
+ let e = interp_constr_entry_key true (ETConstr (n,())) in
extend_constr e (ETConstr (n,()),assoc) (make_cases_pattern_action mkact)
- true rule
+ true rules
(**********************************************************************)
(** Making generic actions in type generic_argument *)
@@ -146,7 +183,7 @@ let extend_constr_notation (n,assoc,ntn,rule) =
let make_generic_action
(f:loc -> ('b * raw_generic_argument) list -> 'a) 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)
@@ -160,24 +197,21 @@ let make_rule univ f g pt =
(symbs, act)
(**********************************************************************)
-(** Grammar extensions declared at ML level *)
+(** Grammar extensions declared at ML level *)
-type grammar_tactic_production =
- | TacTerm of string
- | TacNonTerm of
- loc * (Gram.te Gramext.g_symbol * argument_type) * string option
+type grammar_prod_item =
+ | GramTerminal of string
+ | GramNonTerminal of
+ loc * argument_type * Gram.te prod_entry_key * identifier option
let make_prod_item = function
- | TacTerm s -> (Gramext.Stoken (Lexer.terminal s), None)
- | TacNonTerm (_,(nont,t), po) -> (nont, Option.map (fun p -> (p,t)) po)
+ | GramTerminal s -> (Gramext.Stoken (Lexer.terminal s), None)
+ | GramNonTerminal (_,t,e,po) ->
+ (symbol_of_prod_entry_key e, Option.map (fun p -> (p,t)) po)
(* Tactic grammar extensions *)
-let tac_exts = ref []
-let get_extend_tactic_grammars () = !tac_exts
-
let extend_tactic_grammar s gl =
- tac_exts := (s,gl) :: !tac_exts;
let univ = get_univ "tactic" in
let mkact loc l = Tacexpr.TacExtend (loc,s,List.map snd l) in
let rules = List.map (make_rule univ mkact make_prod_item) gl in
@@ -196,53 +230,7 @@ let extend_vernac_command_grammar s gl =
Gram.extend Vernac_.command None [(None, None, List.rev rules)]
(**********************************************************************)
-(** Grammar declaration for Tactic Notation (Coq level) *)
-
-(* Interpretation of the grammar entry names *)
-
-let find_index s t =
- let t,n = repr_ident (id_of_string t) in
- if s <> t or n = None then raise Not_found;
- Option.get n
-
-let rec interp_entry_name up_level s =
- let l = String.length s in
- if l > 8 & String.sub s 0 3 = "ne_" & String.sub s (l-5) 5 = "_list" then
- let t, g = interp_entry_name up_level (String.sub s 3 (l-8)) in
- List1ArgType t, Gramext.Slist1 g
- else if l > 5 & String.sub s (l-5) 5 = "_list" then
- let t, g = interp_entry_name up_level (String.sub s 0 (l-5)) in
- List0ArgType t, Gramext.Slist0 g
- else if l > 4 & String.sub s (l-4) 4 = "_opt" then
- let t, g = interp_entry_name up_level (String.sub s 0 (l-4)) in
- OptArgType t, Gramext.Sopt g
- else
- let s = if s = "hyp" then "var" else s in
- try
- let i = find_index "tactic" s in
- ExtraArgType s,
- if up_level<>5 && i=up_level then Gramext.Sself else
- if up_level<>5 && i=up_level-1 then Gramext.Snext else
- Gramext.Snterml(Pcoq.Gram.Entry.obj Tactic.tactic_expr,string_of_int i)
- with Not_found ->
- let e =
- (* Qualified entries are no longer in use *)
- try get_entry (get_univ "tactic") s
- with _ ->
- try get_entry (get_univ "prim") s
- with _ ->
- try get_entry (get_univ "constr") s
- with _ -> error ("Unknown entry "^s^".")
- in
- let o = object_of_typed_entry e in
- let t = type_of_typed_entry e in
- t,Gramext.Snterm (Pcoq.Gram.Entry.obj o)
-
-let make_vprod_item n = function
- | VTerm s -> (Gramext.Stoken (Lexer.terminal s), None)
- | VNonTerm (loc, nt, po) ->
- let (etyp, e) = interp_entry_name n nt in
- e, Option.map (fun p -> (p,etyp)) po
+(** Grammar declaration for Tactic Notation (Coq level) *)
let get_tactic_entry n =
if n = 0 then
@@ -251,43 +239,42 @@ let get_tactic_entry n =
weaken_entry Tactic.binder_tactic, None
else if 1<=n && n<5 then
weaken_entry Tactic.tactic_expr, Some (Gramext.Level (string_of_int n))
- else
+ else
error ("Invalid Tactic Notation level: "^(string_of_int n)^".")
(* Declaration of the tactic grammar rule *)
-let head_is_ident = function VTerm _::_ -> true | _ -> false
+let head_is_ident = function GramTerminal _::_ -> true | _ -> false
let add_tactic_entry (key,lev,prods,tac) =
let univ = get_univ "tactic" in
let entry, pos = get_tactic_entry lev in
- let mkprod = make_vprod_item lev in
- let rules =
+ let rules =
if lev = 0 then begin
if not (head_is_ident prods) then
error "Notation for simple tactic must start with an identifier.";
let mkact s tac loc l =
(TacAlias(loc,s,l,tac):raw_atomic_tactic_expr) in
- make_rule univ (mkact key tac) mkprod prods
+ make_rule univ (mkact key tac) make_prod_item prods
end
else
- let mkact s tac loc l =
+ let mkact s tac loc l =
(TacAtom(loc,TacAlias(loc,s,l,tac)):raw_tactic_expr) in
- make_rule univ (mkact key tac) mkprod prods in
+ make_rule univ (mkact key tac) make_prod_item prods in
synchronize_level_positions ();
grammar_extend entry pos None [(None, None, List.rev [rules])]
(**********************************************************************)
(** State of the grammar extensions *)
-type notation_grammar =
- int * Gramext.g_assoc option * notation * prod_item list
+type notation_grammar =
+ int * Gramext.g_assoc option * notation * grammar_constr_prod_item list list
type all_grammar_command =
- | Notation of Notation.level * notation_grammar
+ | Notation of (precedence * tolerability list) * notation_grammar
| TacticGrammar of
- (string * int * grammar_production list *
- (Names.dir_path * Tacexpr.glob_tactic_expr))
+ (string * int * grammar_prod_item list *
+ (dir_path * Tacexpr.glob_tactic_expr))
let (grammar_state : all_grammar_command list ref) = ref []
@@ -297,14 +284,6 @@ let extend_grammar gram =
| TacticGrammar g -> add_tactic_entry g);
grammar_state := gram :: !grammar_state
-let reset_extend_grammars_v8 () =
- let te = List.rev !tac_exts in
- let tv = List.rev !vernac_exts in
- tac_exts := [];
- vernac_exts := [];
- List.iter (fun (s,gl) -> print_string ("Resinstalling "^s); flush stdout; extend_tactic_grammar s gl) te;
- List.iter (fun (s,gl) -> extend_vernac_command_grammar s gl) tv
-
let recover_notation_grammar ntn prec =
let l = map_succeed (function
| Notation (prec',(_,_,ntn',_ as x)) when prec = prec' & ntn = ntn' -> x
@@ -319,7 +298,7 @@ type frozen_t = all_grammar_command list * Lexer.frozen_t
let freeze () = (!grammar_state, Lexer.freeze ())
-(* We compare the current state of the grammar and the state to unfreeze,
+(* We compare the current state of the grammar and the state to unfreeze,
by computing the longest common suffixes *)
let factorize_grams l1 l2 =
if l1 == l2 then ([], [], l1) else list_share_tails l1 l2
@@ -339,7 +318,7 @@ let unfreeze (grams, lex) =
grammar_state := common;
Lexer.unfreeze lex;
List.iter extend_grammar (List.rev redo)
-
+
let init_grammar () =
remove_grammars (number_of_entries !grammar_state);
grammar_state := []
@@ -349,10 +328,8 @@ let init () =
open Summary
-let _ =
+let _ =
declare_summary "GRAMMAR_LEXER"
{ freeze_function = freeze;
unfreeze_function = unfreeze;
- init_function = init;
- survive_module = false;
- survive_section = false }
+ init_function = init }
diff --git a/parsing/egrammar.mli b/parsing/egrammar.mli
index d2d912d1..1228b40c 100644
--- a/parsing/egrammar.mli
+++ b/parsing/egrammar.mli
@@ -6,62 +6,63 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: egrammar.mli 10122 2007-09-15 10:35:59Z letouzey $ i*)
+(*i $Id$ i*)
(*i*)
open Util
+open Names
open Topconstr
open Pcoq
open Extend
open Vernacexpr
open Ppextend
open Rawterm
+open Genarg
open Mod_subst
(*i*)
(** Mapping of grammar productions to camlp4 actions
- Used for Coq-level Notation and Tactic Notation,
+ Used for Coq-level Notation and Tactic Notation,
and for ML-level tactic and vernac extensions
*)
-type prod_item =
- | Term of Token.pattern
- | NonTerm of constr_production_entry *
- (Names.identifier * constr_production_entry) option
+(* For constr notations *)
-type notation_grammar =
- int * Gramext.g_assoc option * notation * prod_item list
+type grammar_constr_prod_item =
+ | GramConstrTerminal of Token.pattern
+ | GramConstrNonTerminal of constr_prod_entry_key * identifier option
+ | GramConstrListMark of int * bool
+ (* tells action rule to make a list of the n previous parsed items;
+ concat with last parsed list if true *)
+
+type notation_grammar =
+ int * Gramext.g_assoc option * notation * grammar_constr_prod_item list list
+
+(* For tactic and vernac notations *)
+
+type grammar_prod_item =
+ | GramTerminal of string
+ | GramNonTerminal of loc * argument_type *
+ Gram.te prod_entry_key * identifier option
+
+(* Adding notations *)
type all_grammar_command =
| Notation of (precedence * tolerability list) * notation_grammar
| TacticGrammar of
- (string * int * grammar_production list *
- (Names.dir_path * Tacexpr.glob_tactic_expr))
+ (string * int * grammar_prod_item list *
+ (dir_path * Tacexpr.glob_tactic_expr))
val extend_grammar : all_grammar_command -> unit
-(* Add grammar rules for tactics *)
-
-type grammar_tactic_production =
- | TacTerm of string
- | TacNonTerm of
- loc * (Compat.token Gramext.g_symbol * Genarg.argument_type) * string option
-
val extend_tactic_grammar :
- string -> grammar_tactic_production list list -> unit
+ string -> grammar_prod_item list list -> unit
val extend_vernac_command_grammar :
- string -> grammar_tactic_production list list -> unit
-(*
-val get_extend_tactic_grammars :
- unit -> (string * grammar_tactic_production list list) list
-*)
+ string -> grammar_prod_item list list -> unit
+
val get_extend_vernac_grammars :
- unit -> (string * grammar_tactic_production list list) list
-(*
-val reset_extend_grammars_v8 : unit -> unit
-*)
-val interp_entry_name : int -> string -> entry_type * Compat.token Gramext.g_symbol
+ unit -> (string * grammar_prod_item list list) list
val recover_notation_grammar :
notation -> (precedence * tolerability list) -> notation_grammar
diff --git a/parsing/extend.ml b/parsing/extend.ml
index f4c98291..7643120f 100644
--- a/parsing/extend.ml
+++ b/parsing/extend.ml
@@ -6,52 +6,58 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-
-(*i $Id: extend.ml 7761 2005-12-30 10:52:19Z herbelin $ i*)
+(*i $Id$ i*)
open Util
-open Pp
-open Gramext
-open Names
-open Ppextend
-open Topconstr
-open Genarg
(**********************************************************************)
-(* constr entry keys *)
+(* General entry keys *)
+
+(* This intermediate abstract representation of entries can *)
+(* both be reified into mlexpr for the ML extensions and *)
+(* dynamically interpreted as entries for the Coq level extensions *)
+
+type 'a prod_entry_key =
+ | Alist1 of 'a prod_entry_key
+ | Alist1sep of 'a prod_entry_key * string
+ | Alist0 of 'a prod_entry_key
+ | Alist0sep of 'a prod_entry_key * string
+ | Aopt of 'a prod_entry_key
+ | Amodifiers of 'a prod_entry_key
+ | Aself
+ | Anext
+ | Atactic of int
+ | Agram of 'a Gramext.g_entry
+ | Aentry of string * string
+
+(**********************************************************************)
+(* Entry keys for constr notations *)
type side = Left | Right
type production_position =
- | BorderProd of side * Gramext.g_assoc option (* true=left; false=right *)
+ | BorderProd of side * Gramext.g_assoc option
| InternalProd
type production_level =
| NextLevel
| NumLevel of int
-type ('lev,'pos) constr_entry_key =
- | ETIdent | ETReference | ETBigint
+type ('lev,'pos) constr_entry_key_gen =
+ | ETName | ETReference | ETBigint
| ETConstr of ('lev * 'pos)
| ETPattern
| ETOther of string * string
| ETConstrList of ('lev * 'pos) * Token.pattern list
-type constr_production_entry =
- (production_level,production_position) constr_entry_key
-type constr_entry =
- (int,unit) constr_entry_key
-type simple_constr_production_entry =
- (production_level,unit) constr_entry_key
-
-(**********************************************************************)
-(* syntax modifiers *)
+(* Entries level (left-hand-side of grammar rules) *)
+type constr_entry_key =
+ (int,unit) constr_entry_key_gen
-type syntax_modifier =
- | SetItemLevel of string list * production_level
- | SetLevel of int
- | SetAssoc of Gramext.g_assoc
- | SetEntryType of string * simple_constr_production_entry
- | SetOnlyParsing
- | SetFormat of string located
+(* Entries used in productions (in right-hand-side of grammar rules) *)
+type constr_prod_entry_key =
+ (production_level,production_position) constr_entry_key_gen
+(* Entries used in productions, vernac side (e.g. "x bigint" or "x ident") *)
+type simple_constr_prod_entry_key =
+ (production_level,unit) constr_entry_key_gen
diff --git a/parsing/extend.mli b/parsing/extend.mli
index 80de7108..7643120f 100644
--- a/parsing/extend.mli
+++ b/parsing/extend.mli
@@ -6,45 +6,58 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: extend.mli 7761 2005-12-30 10:52:19Z herbelin $ i*)
+(*i $Id$ i*)
open Util
(**********************************************************************)
-(* constr entry keys *)
+(* General entry keys *)
+
+(* This intermediate abstract representation of entries can *)
+(* both be reified into mlexpr for the ML extensions and *)
+(* dynamically interpreted as entries for the Coq level extensions *)
+
+type 'a prod_entry_key =
+ | Alist1 of 'a prod_entry_key
+ | Alist1sep of 'a prod_entry_key * string
+ | Alist0 of 'a prod_entry_key
+ | Alist0sep of 'a prod_entry_key * string
+ | Aopt of 'a prod_entry_key
+ | Amodifiers of 'a prod_entry_key
+ | Aself
+ | Anext
+ | Atactic of int
+ | Agram of 'a Gramext.g_entry
+ | Aentry of string * string
+
+(**********************************************************************)
+(* Entry keys for constr notations *)
type side = Left | Right
type production_position =
- | BorderProd of side * Gramext.g_assoc option (* true=left; false=right *)
+ | BorderProd of side * Gramext.g_assoc option
| InternalProd
type production_level =
| NextLevel
| NumLevel of int
-type ('lev,'pos) constr_entry_key =
- | ETIdent | ETReference | ETBigint
+type ('lev,'pos) constr_entry_key_gen =
+ | ETName | ETReference | ETBigint
| ETConstr of ('lev * 'pos)
| ETPattern
| ETOther of string * string
| ETConstrList of ('lev * 'pos) * Token.pattern list
-type constr_production_entry =
- (production_level,production_position) constr_entry_key
-type constr_entry =
- (int,unit) constr_entry_key
-type simple_constr_production_entry =
- (production_level,unit) constr_entry_key
-
-(**********************************************************************)
-(* syntax modifiers *)
+(* Entries level (left-hand-side of grammar rules) *)
+type constr_entry_key =
+ (int,unit) constr_entry_key_gen
-type syntax_modifier =
- | SetItemLevel of string list * production_level
- | SetLevel of int
- | SetAssoc of Gramext.g_assoc
- | SetEntryType of string * simple_constr_production_entry
- | SetOnlyParsing
- | SetFormat of string located
+(* Entries used in productions (in right-hand-side of grammar rules) *)
+type constr_prod_entry_key =
+ (production_level,production_position) constr_entry_key_gen
+(* Entries used in productions, vernac side (e.g. "x bigint" or "x ident") *)
+type simple_constr_prod_entry_key =
+ (production_level,unit) constr_entry_key_gen
diff --git a/parsing/extrawit.ml b/parsing/extrawit.ml
new file mode 100644
index 00000000..122730f7
--- /dev/null
+++ b/parsing/extrawit.ml
@@ -0,0 +1,62 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id$ i*)
+
+open Util
+open Genarg
+
+(* This file defines extra argument types *)
+
+(* Tactics as arguments *)
+
+let tactic_main_level = 5
+
+let (wit_tactic0,globwit_tactic0,rawwit_tactic0) = create_arg "tactic0"
+let (wit_tactic1,globwit_tactic1,rawwit_tactic1) = create_arg "tactic1"
+let (wit_tactic2,globwit_tactic2,rawwit_tactic2) = create_arg "tactic2"
+let (wit_tactic3,globwit_tactic3,rawwit_tactic3) = create_arg "tactic3"
+let (wit_tactic4,globwit_tactic4,rawwit_tactic4) = create_arg "tactic4"
+let (wit_tactic5,globwit_tactic5,rawwit_tactic5) = create_arg "tactic5"
+
+let wit_tactic = function
+ | 0 -> wit_tactic0
+ | 1 -> wit_tactic1
+ | 2 -> wit_tactic2
+ | 3 -> wit_tactic3
+ | 4 -> wit_tactic4
+ | 5 -> wit_tactic5
+ | n -> anomaly ("Unavailable tactic level: "^string_of_int n)
+
+let globwit_tactic = function
+ | 0 -> globwit_tactic0
+ | 1 -> globwit_tactic1
+ | 2 -> globwit_tactic2
+ | 3 -> globwit_tactic3
+ | 4 -> globwit_tactic4
+ | 5 -> globwit_tactic5
+ | n -> anomaly ("Unavailable tactic level: "^string_of_int n)
+
+let rawwit_tactic = function
+ | 0 -> rawwit_tactic0
+ | 1 -> rawwit_tactic1
+ | 2 -> rawwit_tactic2
+ | 3 -> rawwit_tactic3
+ | 4 -> rawwit_tactic4
+ | 5 -> rawwit_tactic5
+ | n -> anomaly ("Unavailable tactic level: "^string_of_int n)
+
+let tactic_genarg_level s =
+ if String.length s = 7 && String.sub s 0 6 = "tactic" then
+ let c = s.[6] in if '5' >= c && c >= '0' then Some (Char.code c - 48)
+ else None
+ else None
+
+let is_tactic_genarg = function
+| ExtraArgType s -> tactic_genarg_level s <> None
+| _ -> false
diff --git a/parsing/extrawit.mli b/parsing/extrawit.mli
new file mode 100644
index 00000000..9eff5dc1
--- /dev/null
+++ b/parsing/extrawit.mli
@@ -0,0 +1,51 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id$ i*)
+
+open Util
+open Genarg
+open Tacexpr
+
+(* This file defines extra argument types *)
+
+(* Tactics as arguments *)
+
+val tactic_main_level : int
+
+val rawwit_tactic : int -> (raw_tactic_expr,rlevel) abstract_argument_type
+val globwit_tactic : int -> (glob_tactic_expr,glevel) abstract_argument_type
+val wit_tactic : int -> (glob_tactic_expr,tlevel) abstract_argument_type
+
+val rawwit_tactic0 : (raw_tactic_expr,rlevel) abstract_argument_type
+val globwit_tactic0 : (glob_tactic_expr,glevel) abstract_argument_type
+val wit_tactic0 : (glob_tactic_expr,tlevel) abstract_argument_type
+
+val rawwit_tactic1 : (raw_tactic_expr,rlevel) abstract_argument_type
+val globwit_tactic1 : (glob_tactic_expr,glevel) abstract_argument_type
+val wit_tactic1 : (glob_tactic_expr,tlevel) abstract_argument_type
+
+val rawwit_tactic2 : (raw_tactic_expr,rlevel) abstract_argument_type
+val globwit_tactic2 : (glob_tactic_expr,glevel) abstract_argument_type
+val wit_tactic2 : (glob_tactic_expr,tlevel) abstract_argument_type
+
+val rawwit_tactic3 : (raw_tactic_expr,rlevel) abstract_argument_type
+val globwit_tactic3 : (glob_tactic_expr,glevel) abstract_argument_type
+val wit_tactic3 : (glob_tactic_expr,tlevel) abstract_argument_type
+
+val rawwit_tactic4 : (raw_tactic_expr,rlevel) abstract_argument_type
+val globwit_tactic4 : (glob_tactic_expr,glevel) abstract_argument_type
+val wit_tactic4 : (glob_tactic_expr,tlevel) abstract_argument_type
+
+val rawwit_tactic5 : (raw_tactic_expr,rlevel) abstract_argument_type
+val globwit_tactic5 : (glob_tactic_expr,glevel) abstract_argument_type
+val wit_tactic5 : (glob_tactic_expr,tlevel) abstract_argument_type
+
+val is_tactic_genarg : argument_type -> bool
+
+val tactic_genarg_level : string -> int option
diff --git a/parsing/g_ascii_syntax.ml b/parsing/g_ascii_syntax.ml
deleted file mode 100644
index 944e2338..00000000
--- a/parsing/g_ascii_syntax.ml
+++ /dev/null
@@ -1,83 +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: g_ascii_syntax.ml 11309 2008-08-06 10:30:35Z herbelin $ i*)
-
-open Pp
-open Util
-open Names
-open Pcoq
-open Rawterm
-open Topconstr
-open Libnames
-open Coqlib
-open Bigint
-
-exception Non_closed_ascii
-
-let make_dir l = make_dirpath (List.map id_of_string (List.rev l))
-let make_kn dir id = Libnames.encode_kn (make_dir dir) (id_of_string id)
-let make_path dir id = Libnames.make_path (make_dir dir) (id_of_string id)
-
-let ascii_module = ["Coq";"Strings";"Ascii"]
-
-let ascii_path = make_path ascii_module "ascii"
-
-let ascii_kn = make_kn ascii_module "ascii"
-let path_of_Ascii = ((ascii_kn,0),1)
-let static_glob_Ascii = ConstructRef path_of_Ascii
-
-let make_reference id = find_reference "Ascii interpretation" ascii_module id
-let glob_Ascii = lazy (make_reference "Ascii")
-
-open Lazy
-
-let interp_ascii dloc p =
- let rec aux n p =
- if n = 0 then [] else
- let mp = p mod 2 in
- RRef (dloc,if mp = 0 then glob_false else glob_true)
- :: (aux (n-1) (p/2)) in
- RApp (dloc,RRef(dloc,force glob_Ascii), aux 8 p)
-
-let interp_ascii_string dloc s =
- let p =
- if String.length s = 1 then int_of_char s.[0]
- else
- if String.length s = 3 & is_digit s.[0] & is_digit s.[1] & is_digit s.[2]
- then int_of_string s
- else
- user_err_loc (dloc,"interp_ascii_string",
- str "Expects a single character or a three-digits ascii code.") in
- interp_ascii dloc p
-
-let uninterp_ascii r =
- let rec uninterp_bool_list n = function
- | [] when n = 0 -> 0
- | RRef (_,k)::l when k = glob_true -> 1+2*(uninterp_bool_list (n-1) l)
- | RRef (_,k)::l when k = glob_false -> 2*(uninterp_bool_list (n-1) l)
- | _ -> raise Non_closed_ascii in
- try
- let rec aux = function
- | RApp (_,RRef (_,k),l) when k = force glob_Ascii -> uninterp_bool_list 8 l
- | _ -> raise Non_closed_ascii in
- Some (aux r)
- with
- Non_closed_ascii -> None
-
-let make_ascii_string n =
- if n>=32 && n<=126 then String.make 1 (char_of_int n)
- else Printf.sprintf "%03d" n
-
-let uninterp_ascii_string r = Option.map make_ascii_string (uninterp_ascii r)
-
-let _ =
- Notation.declare_string_interpreter "char_scope"
- (ascii_path,ascii_module)
- interp_ascii_string
- ([RRef (dummy_loc,static_glob_Ascii)], uninterp_ascii_string, true)
diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4
index cdce13e6..393125e2 100644
--- a/parsing/g_constr.ml4
+++ b/parsing/g_constr.ml4
@@ -8,8 +8,9 @@
(*i camlp4use: "pa_extend.cmo" i*)
-(* $Id: g_constr.ml4 11709 2008-12-20 11:42:15Z msozeau $ *)
+(* $Id$ *)
+open Pp
open Pcoq
open Constr
open Prim
@@ -22,7 +23,7 @@ open Topconstr
open Util
let constr_kw =
- [ "forall"; "fun"; "match"; "fix"; "cofix"; "with"; "in"; "for";
+ [ "forall"; "fun"; "match"; "fix"; "cofix"; "with"; "in"; "for";
"end"; "as"; "let"; "if"; "then"; "else"; "return";
"Prop"; "Set"; "Type"; ".("; "_"; "..";
"`{"; "`("; "{|"; "|}" ]
@@ -33,35 +34,21 @@ let mk_cast = function
(c,(_,None)) -> c
| (c,(_,Some ty)) -> CCast(join_loc (constr_loc c) (constr_loc ty), c, CastConv (DEFAULTcast, ty))
-let mk_lam = function
- ([],c) -> c
- | (bl,c) -> CLambdaN(constr_loc c, bl,c)
-
let loc_of_binder_let = function
| LocalRawAssum ((loc,_)::_,_,_)::_ -> loc
| LocalRawDef ((loc,_),_)::_ -> loc
| _ -> dummy_loc
let binders_of_lidents l =
- List.map (fun (loc, id) ->
- LocalRawAssum ([loc, Name id], Default Rawterm.Explicit,
+ List.map (fun (loc, id) ->
+ LocalRawAssum ([loc, Name id], Default Rawterm.Explicit,
CHole (loc, Some (Evd.BinderType (Name id))))) l
-
-let rec index_and_rec_order_of_annot loc bl ann =
- match names_of_local_assums bl,ann with
- | [loc,Name id], (None, r) -> Some (loc, id), r
- | lids, (Some (loc, n), ro) ->
- if List.exists (fun (_, x) -> x = Name n) lids then
- Some (loc, n), ro
- else user_err_loc(loc,"index_of_annot", Pp.str"No such fix variable.")
- | _, (None, r) -> None, r
let mk_fixb (id,bl,ann,body,(loc,tyc)) =
- let n,ro = index_and_rec_order_of_annot (fst id) bl ann in
let ty = match tyc with
Some ty -> ty
| None -> CHole (loc, None) in
- (id,(n,ro),bl,ty,body)
+ (id,ann,bl,ty,body)
let mk_cofixb (id,bl,ann,body,(loc,tyc)) =
let _ = Option.map (fun (aloc,_) ->
@@ -74,7 +61,7 @@ let mk_cofixb (id,bl,ann,body,(loc,tyc)) =
(id,bl,ty,body)
let mk_fix(loc,kw,id,dcls) =
- if kw then
+ if kw then
let fb = List.map mk_fixb dcls in
CFix(loc,id,fb)
else
@@ -84,9 +71,6 @@ let mk_fix(loc,kw,id,dcls) =
let mk_single_fix (loc,kw,dcl) =
let (id,_,_,_,_) = dcl in mk_fix(loc,kw,id,[dcl])
-let binder_constr =
- create_constr_entry (get_univ "constr") "binder_constr"
-
(* Hack to parse "(x:=t)" as an explicit argument without conflicts with the *)
(* admissible notation "(x t)" *)
let lpar_id_coloneq =
@@ -108,16 +92,16 @@ let impl_ident =
Gram.Entry.of_parser "impl_ident"
(fun strm ->
match Stream.npeek 1 strm with
- | [(_,"{")] ->
+ | [(_,"{")] ->
(match Stream.npeek 2 strm with
| [_;("IDENT",("wf"|"struct"|"measure"))] ->
raise Stream.Failure
- | [_;("IDENT",s)] ->
+ | [_;("IDENT",s)] ->
Stream.junk strm; Stream.junk strm;
Names.id_of_string s
| _ -> raise Stream.Failure)
| _ -> raise Stream.Failure)
-
+
let ident_colon =
Gram.Entry.of_parser "ident_colon"
(fun strm ->
@@ -141,7 +125,7 @@ let ident_with =
Names.id_of_string s
| _ -> raise Stream.Failure)
| _ -> raise Stream.Failure)
-
+
let aliasvar = function CPatAlias (_, _, id) -> Some (Name id) | _ -> None
GEXTEND Gram
@@ -176,21 +160,21 @@ GEXTEND Gram
[ [ c = operconstr LEVEL "200" -> c ] ]
;
constr:
- [ [ c = operconstr LEVEL "8" -> c
+ [ [ c = operconstr LEVEL "8" -> c
| "@"; f=global -> CAppExpl(loc,(None,f),[]) ] ]
;
operconstr:
[ "200" RIGHTA
[ c = binder_constr -> c ]
| "100" RIGHTA
- [ c1 = operconstr; "<:"; c2 = binder_constr ->
+ [ c1 = operconstr; "<:"; c2 = binder_constr ->
CCast(loc,c1, CastConv (VMcast,c2))
- | c1 = operconstr; "<:"; c2 = SELF ->
+ | c1 = operconstr; "<:"; c2 = SELF ->
CCast(loc,c1, CastConv (VMcast,c2))
- | c1 = operconstr; ":";c2 = binder_constr ->
+ | c1 = operconstr; ":";c2 = binder_constr ->
+ CCast(loc,c1, CastConv (DEFAULTcast,c2))
+ | c1 = operconstr; ":"; c2 = SELF ->
CCast(loc,c1, CastConv (DEFAULTcast,c2))
- | c1 = operconstr; ":"; c2 = SELF ->
- CCast(loc,c1, CastConv (DEFAULTcast,c2))
| c1 = operconstr; ":>" ->
CCast(loc,c1, CastCoerce) ]
| "99" RIGHTA [ ]
@@ -212,7 +196,7 @@ GEXTEND Gram
CApp(loc,(Some (List.length args+1),CRef f),args@[c,None])
| c=operconstr; ".("; "@"; f=global;
args=LIST0 (operconstr LEVEL "9"); ")" ->
- CAppExpl(loc,(Some (List.length args+1),f),args@[c])
+ CAppExpl(loc,(Some (List.length args+1),f),args@[c])
| c=operconstr; "%"; key=IDENT -> CDelimiters (loc,key,c) ]
| "0"
[ c=atomic_constr -> c
@@ -229,13 +213,13 @@ GEXTEND Gram
CGeneralization (loc, Explicit, None, c)
] ]
;
- forall:
- [ [ "forall" -> ()
+ forall:
+ [ [ "forall" -> ()
| IDENT "Π" -> ()
] ]
;
- lambda:
- [ [ "fun" -> ()
+ lambda:
+ [ [ "fun" -> ()
| IDENT "λ" -> ()
] ]
;
@@ -246,7 +230,7 @@ GEXTEND Gram
] ]
;
record_field_declaration:
- [ [ id = identref; params = LIST0 identref; ":="; c = lconstr ->
+ [ [ id = global; params = LIST0 identref; ":="; c = lconstr ->
(id, Topconstr.abstract_constr_expr c (binders_of_lidents params)) ] ]
;
binder_constr:
@@ -273,10 +257,10 @@ GEXTEND Gram
| "let"; "'"; p=pattern; ":="; c1 = operconstr LEVEL "200";
"in"; c2 = operconstr LEVEL "200" ->
CCases (loc, LetPatternStyle, None, [(c1,(None,None))], [(loc, [(loc,[p])], c2)])
- | "let"; "'"; p=pattern; ":="; c1 = operconstr LEVEL "200";
+ | "let"; "'"; p=pattern; ":="; c1 = operconstr LEVEL "200";
rt = case_type; "in"; c2 = operconstr LEVEL "200" ->
CCases (loc, LetPatternStyle, Some rt, [(c1, (aliasvar p, None))], [(loc, [(loc, [p])], c2)])
- | "let"; "'"; p=pattern; "in"; t = operconstr LEVEL "200";
+ | "let"; "'"; p=pattern; "in"; t = operconstr LEVEL "200";
":="; c1 = operconstr LEVEL "200"; rt = case_type;
"in"; c2 = operconstr LEVEL "200" ->
CCases (loc, LetPatternStyle, Some rt, [(c1, (aliasvar p, Some t))], [(loc, [(loc, [p])], c2)])
@@ -315,7 +299,8 @@ GEXTEND Gram
;
fix_decl:
[ [ id=identref; bl=binders_let_fixannot; ty=type_cstr; ":=";
- c=operconstr LEVEL "200" -> (id,fst bl,snd bl,c,ty) ] ]
+ c=operconstr LEVEL "200" ->
+ (id,fst bl,snd bl,c,ty) ] ]
;
match_constr:
[ [ "match"; ci=LIST1 case_item SEP ","; ty=OPT case_type; "with";
@@ -333,8 +318,8 @@ GEXTEND Gram
;
return_type:
[ [ a = OPT [ na = OPT["as"; id=name -> snd id];
- ty = case_type -> (na,ty) ] ->
- match a with
+ ty = case_type -> (na,ty) ] ->
+ match a with
| None -> None, None
| Some (na,t) -> (na, Some t)
] ]
@@ -349,6 +334,9 @@ GEXTEND Gram
[ [ pll = LIST1 mult_pattern SEP "|";
"=>"; rhs = lconstr -> (loc,pll,rhs) ] ]
;
+ recordpattern:
+ [ [ id = global; ":="; pat = pattern -> (id, pat) ] ]
+ ;
pattern:
[ "200" RIGHTA [ ]
| "100" RIGHTA
@@ -358,7 +346,7 @@ GEXTEND Gram
[ p = pattern; lp = LIST1 NEXT ->
(match p with
| CPatAtom (_, Some r) -> CPatCstr (loc, r, lp)
- | _ -> Util.user_err_loc
+ | _ -> Util.user_err_loc
(cases_pattern_expr_loc p, "compound_pattern",
Pp.str "Constructor expected."))
| p = pattern; "as"; id = ident ->
@@ -367,6 +355,7 @@ GEXTEND Gram
[ c = pattern; "%"; key=IDENT -> CPatDelimiters (loc,key,c) ]
| "0"
[ r = Prim.reference -> CPatAtom (loc,Some r)
+ | "{|"; pat = LIST0 recordpattern SEP ";" ; "|}" -> CPatRecord (loc, pat)
| "_" -> CPatAtom (loc,None)
| "("; p = pattern LEVEL "200"; ")" ->
(match p with
@@ -377,9 +366,9 @@ GEXTEND Gram
| s = string -> CPatPrim (loc, String s) ] ]
;
binder_list:
- [ [ idl=LIST1 name; bl=binders_let ->
+ [ [ idl=LIST1 name; bl=binders_let ->
LocalRawAssum (idl,Default Explicit,CHole (loc, Some (Evd.BinderType (snd (List.hd idl)))))::bl
- | idl=LIST1 name; ":"; c=lconstr ->
+ | idl=LIST1 name; ":"; c=lconstr ->
[LocalRawAssum (idl,Default Explicit,c)]
| cl = binders_let -> cl
] ]
@@ -397,15 +386,15 @@ GEXTEND Gram
fixannot:
[ [ "{"; IDENT "struct"; id=identref; "}" -> (Some id, CStructRec)
| "{"; IDENT "wf"; rel=constr; id=OPT identref; "}" -> (id, CWfRec rel)
- | "{"; IDENT "measure"; rel=constr; id=OPT identref; "}" -> (id, CMeasureRec rel)
+ | "{"; IDENT "measure"; m=constr; id=OPT identref;
+ rel=OPT constr; "}" -> (id, CMeasureRec (m,rel))
] ]
;
binders_let_fixannot:
- [ [ id=impl_ident; assum=binder_assum; bl = binders_let_fixannot ->
+ [ [ id=impl_ident; assum=binder_assum; bl = binders_let_fixannot ->
(assum (loc, Name id) :: fst bl), snd bl
| f = fixannot -> [], f
- | b = binder_let; bl = binders_let_fixannot ->
- b @ fst bl, snd bl
+ | b = binder_let; bl = binders_let_fixannot -> b @ fst bl, snd bl
| -> [], (None, CStructRec)
] ]
;
@@ -416,21 +405,21 @@ GEXTEND Gram
binder_let:
[ [ id=name ->
[LocalRawAssum ([id],Default Explicit,CHole (loc, None))]
- | "("; id=name; idl=LIST1 name; ":"; c=lconstr; ")" ->
+ | "("; id=name; idl=LIST1 name; ":"; c=lconstr; ")" ->
[LocalRawAssum (id::idl,Default Explicit,c)]
- | "("; id=name; ":"; c=lconstr; ")" ->
+ | "("; id=name; ":"; c=lconstr; ")" ->
[LocalRawAssum ([id],Default Explicit,c)]
| "("; id=name; ":="; c=lconstr; ")" ->
[LocalRawDef (id,c)]
- | "("; id=name; ":"; t=lconstr; ":="; c=lconstr; ")" ->
+ | "("; id=name; ":"; t=lconstr; ":="; c=lconstr; ")" ->
[LocalRawDef (id,CCast (join_loc (constr_loc t) loc,c, CastConv (DEFAULTcast,t)))]
| "{"; id=name; "}" ->
[LocalRawAssum ([id],Default Implicit,CHole (loc, None))]
- | "{"; id=name; idl=LIST1 name; ":"; c=lconstr; "}" ->
+ | "{"; id=name; idl=LIST1 name; ":"; c=lconstr; "}" ->
[LocalRawAssum (id::idl,Default Implicit,c)]
- | "{"; id=name; ":"; c=lconstr; "}" ->
+ | "{"; id=name; ":"; c=lconstr; "}" ->
[LocalRawAssum ([id],Default Implicit,c)]
- | "{"; id=name; idl=LIST1 name; "}" ->
+ | "{"; id=name; idl=LIST1 name; "}" ->
List.map (fun id -> LocalRawAssum ([id],Default Implicit,CHole (loc, None))) (id::idl)
| "`("; tc = LIST1 typeclass_constraint SEP "," ; ")" ->
List.map (fun (n, b, t) -> LocalRawAssum ([n], Generalized (Implicit, Explicit, b), t)) tc
@@ -440,8 +429,8 @@ GEXTEND Gram
;
binder:
[ [ id=name -> ([id],Default Explicit,CHole (loc, None))
- | "("; idl=LIST1 name; ":"; c=lconstr; ")" -> (idl,Default Explicit,c)
- | "{"; idl=LIST1 name; ":"; c=lconstr; "}" -> (idl,Default Implicit,c)
+ | "("; idl=LIST1 name; ":"; c=lconstr; ")" -> (idl,Default Explicit,c)
+ | "{"; idl=LIST1 name; ":"; c=lconstr; "}" -> (idl,Default Implicit,c)
] ]
;
typeclass_constraint:
@@ -454,7 +443,7 @@ GEXTEND Gram
(loc, Anonymous), false, c
] ]
;
-
+
type_cstr:
[ [ c=OPT [":"; c=lconstr -> c] -> (loc,c) ] ]
;
diff --git a/parsing/g_decl_mode.ml4 b/parsing/g_decl_mode.ml4
index 35fc064b..e73e54e7 100644
--- a/parsing/g_decl_mode.ml4
+++ b/parsing/g_decl_mode.ml4
@@ -9,7 +9,7 @@
(*i camlp4deps: "parsing/grammar.cma" i*)
(*i camlp4use: "pa_extend.cmo q_MLast.cmo" i*)
-(* $Id: g_decl_mode.ml4 12414 2009-10-25 18:50:41Z corbinea $ *)
+(* $Id$ *)
open Decl_expr
@@ -29,7 +29,7 @@ let none_is_empty = function
GEXTEND Gram
GLOBAL: proof_instr;
thesis :
- [[ "thesis" -> Plain
+ [[ "thesis" -> Plain
| "thesis"; "for"; i=ident -> (For i)
]];
statement :
@@ -42,9 +42,9 @@ GLOBAL: proof_instr;
[[ t=thesis -> Thesis t ] |
[ c=constr -> This c
]];
- statement_or_thesis :
+ statement_or_thesis :
[
- [ t=thesis -> {st_label=Anonymous;st_it=Thesis t} ]
+ [ t=thesis -> {st_label=Anonymous;st_it=Thesis t} ]
|
[ i=ident ; ":" ; cot=constr_or_thesis -> {st_label=Name i;st_it=cot}
| i=ident -> {st_label=Anonymous;
@@ -52,25 +52,25 @@ GLOBAL: proof_instr;
| c=constr -> {st_label=Anonymous;st_it=This c}
]
];
- justification_items :
- [[ -> Some []
- | IDENT "by"; l=LIST1 constr SEP "," -> Some l
- | IDENT "by"; "*" -> None ]]
+ justification_items :
+ [[ -> Some []
+ | "by"; l=LIST1 constr SEP "," -> Some l
+ | "by"; "*" -> None ]]
;
- justification_method :
- [[ -> None
+ justification_method :
+ [[ -> None
| "using"; tac = tactic -> Some tac ]]
;
simple_cut_or_thesis :
[[ ls = statement_or_thesis;
j = justification_items;
- taco = justification_method
+ taco = justification_method
-> {cut_stat=ls;cut_by=j;cut_using=taco} ]]
;
simple_cut :
[[ ls = statement;
j = justification_items;
- taco = justification_method
+ taco = justification_method
-> {cut_stat=ls;cut_by=j;cut_using=taco} ]]
;
elim_type:
@@ -82,40 +82,40 @@ GLOBAL: proof_instr;
| IDENT "focus" -> B_focus
| IDENT "proof" -> B_proof
| et=elim_type -> B_elim et ]]
- ;
+ ;
elim_obj:
[[ IDENT "on"; c=constr -> Real c
| IDENT "of"; c=simple_cut -> Virtual c ]]
- ;
+ ;
elim_step:
- [[ IDENT "consider" ;
+ [[ IDENT "consider" ;
h=consider_vars ; IDENT "from" ; c=constr -> Pconsider (c,h)
| IDENT "per"; et=elim_type; obj=elim_obj -> Pper (et,obj)
| IDENT "suffices"; ls=suff_clause;
j = justification_items;
- taco = justification_method
- -> Psuffices {cut_stat=ls;cut_by=j;cut_using=taco} ]]
+ taco = justification_method
+ -> Psuffices {cut_stat=ls;cut_by=j;cut_using=taco} ]]
;
rew_step :
- [[ "~=" ; c=simple_cut -> (Rhs,c)
+ [[ "~=" ; c=simple_cut -> (Rhs,c)
| "=~" ; c=simple_cut -> (Lhs,c)]]
;
cut_step:
[[ "then"; tt=elim_step -> Pthen tt
| "then"; c=simple_cut_or_thesis -> Pthen (Pcut c)
- | IDENT "thus"; tt=rew_step -> Pthus (let s,c=tt in Prew (s,c))
+ | IDENT "thus"; tt=rew_step -> Pthus (let s,c=tt in Prew (s,c))
| IDENT "thus"; c=simple_cut_or_thesis -> Pthus (Pcut c)
| IDENT "hence"; c=simple_cut_or_thesis -> Phence (Pcut c)
| tt=elim_step -> tt
- | tt=rew_step -> let s,c=tt in Prew (s,c);
+ | tt=rew_step -> let s,c=tt in Prew (s,c);
| IDENT "have"; c=simple_cut_or_thesis -> Pcut c;
| IDENT "claim"; c=statement -> Pclaim c;
- | IDENT "focus"; IDENT "on"; c=statement -> Pfocus c;
+ | IDENT "focus"; IDENT "on"; c=statement -> Pfocus c;
| "end"; bt = block_type -> Pend bt;
| IDENT "escape" -> Pescape ]]
;
(* examiner s'il est possible de faire R _ et _ R pour R une relation qcq*)
- loc_id:
+ loc_id:
[[ id=ident -> fun x -> (loc,(id,x)) ]];
hyp:
[[ id=loc_id -> id None ;
@@ -124,27 +124,27 @@ GLOBAL: proof_instr;
consider_vars:
[[ name=hyp -> [Hvar name]
| name=hyp; ","; v=consider_vars -> (Hvar name) :: v
- | name=hyp;
+ | name=hyp;
IDENT "such"; IDENT "that"; h=consider_hyps -> (Hvar name)::h
]]
;
- consider_hyps:
+ consider_hyps:
[[ st=statement; IDENT "and"; h=consider_hyps -> Hprop st::h
- | st=statement; IDENT "and";
+ | st=statement; IDENT "and";
IDENT "consider" ; v=consider_vars -> Hprop st::v
| st=statement -> [Hprop st]
]]
- ;
+ ;
assume_vars:
[[ name=hyp -> [Hvar name]
| name=hyp; ","; v=assume_vars -> (Hvar name) :: v
- | name=hyp;
+ | name=hyp;
IDENT "such"; IDENT "that"; h=assume_hyps -> (Hvar name)::h
]]
;
- assume_hyps:
+ assume_hyps:
[[ st=statement; IDENT "and"; h=assume_hyps -> Hprop st::h
- | st=statement; IDENT "and";
+ | st=statement; IDENT "and";
IDENT "we"; IDENT "have" ; v=assume_vars -> Hprop st::v
| st=statement -> [Hprop st]
]]
@@ -152,38 +152,38 @@ GLOBAL: proof_instr;
assume_clause:
[[ IDENT "we" ; IDENT "have" ; v=assume_vars -> v
| h=assume_hyps -> h ]]
- ;
+ ;
suff_vars:
[[ name=hyp; IDENT "to"; IDENT "show" ; c = constr_or_thesis ->
[Hvar name],c
- | name=hyp; ","; v=suff_vars ->
+ | name=hyp; ","; v=suff_vars ->
let (q,c) = v in ((Hvar name) :: q),c
- | name=hyp;
- IDENT "such"; IDENT "that"; h=suff_hyps ->
+ | name=hyp;
+ IDENT "such"; IDENT "that"; h=suff_hyps ->
let (q,c) = h in ((Hvar name) :: q),c
]];
- suff_hyps:
- [[ st=statement; IDENT "and"; h=suff_hyps ->
+ suff_hyps:
+ [[ st=statement; IDENT "and"; h=suff_hyps ->
let (q,c) = h in (Hprop st::q),c
- | st=statement; IDENT "and";
- IDENT "to" ; IDENT "have" ; v=suff_vars ->
+ | st=statement; IDENT "and";
+ IDENT "to" ; IDENT "have" ; v=suff_vars ->
let (q,c) = v in (Hprop st::q),c
- | st=statement; IDENT "to"; IDENT "show" ; c = constr_or_thesis ->
+ | st=statement; IDENT "to"; IDENT "show" ; c = constr_or_thesis ->
[Hprop st],c
]]
;
suff_clause:
[[ IDENT "to" ; IDENT "have" ; v=suff_vars -> v
| h=suff_hyps -> h ]]
- ;
+ ;
let_vars:
[[ name=hyp -> [Hvar name]
| name=hyp; ","; v=let_vars -> (Hvar name) :: v
- | name=hyp; IDENT "be";
+ | name=hyp; IDENT "be";
IDENT "such"; IDENT "that"; h=let_hyps -> (Hvar name)::h
]]
;
- let_hyps:
+ let_hyps:
[[ st=statement; IDENT "and"; h=let_hyps -> Hprop st::h
| st=statement; IDENT "and"; "let"; v=let_vars -> Hprop st::v
| st=statement -> [Hprop st]
@@ -194,19 +194,19 @@ GLOBAL: proof_instr;
| name=hyp; IDENT "such"; IDENT "that"; h=given_hyps -> (Hvar name)::h
]]
;
- given_hyps:
+ given_hyps:
[[ st=statement; IDENT "and"; h=given_hyps -> Hprop st::h
| st=statement; IDENT "and"; IDENT "given"; v=given_vars -> Hprop st::v
| st=statement -> [Hprop st]
]];
suppose_vars:
- [[name=hyp -> [Hvar name]
+ [[name=hyp -> [Hvar name]
|name=hyp; ","; v=suppose_vars -> (Hvar name) :: v
- |name=hyp; OPT[IDENT "be"];
+ |name=hyp; OPT[IDENT "be"];
IDENT "such"; IDENT "that"; h=suppose_hyps -> (Hvar name)::h
]]
;
- suppose_hyps:
+ suppose_hyps:
[[ st=statement_or_thesis; IDENT "and"; h=suppose_hyps -> Hprop st::h
| st=statement_or_thesis; IDENT "and"; IDENT "we"; IDENT "have";
v=suppose_vars -> Hprop st::v
@@ -223,17 +223,17 @@ GLOBAL: proof_instr;
po=OPT[ "with"; p=LIST1 hyp SEP ","-> p ] ;
ho=OPT[ IDENT "and" ; h=suppose_clause -> h ] ->
Pcase (none_is_empty po,c,none_is_empty ho)
- | "let" ; v=let_vars -> Plet v
+ | "let" ; v=let_vars -> Plet v
| IDENT "take"; witnesses = LIST1 constr SEP "," -> Ptake witnesses
| IDENT "assume"; h=assume_clause -> Passume h
| IDENT "given"; h=given_vars -> Pgiven h
- | IDENT "define"; id=ident; args=LIST0 hyp;
+ | IDENT "define"; id=ident; args=LIST0 hyp;
"as"; body=constr -> Pdefine(id,args,body)
| IDENT "reconsider"; id=ident; "as" ; typ=constr -> Pcast (This id,typ)
| IDENT "reconsider"; t=thesis; "as" ; typ=constr -> Pcast (Thesis t ,typ)
]]
;
- emphasis :
+ emphasis :
[[ -> 0
| "*" -> 1
| "**" -> 2
@@ -249,4 +249,4 @@ GLOBAL: proof_instr;
;
END;;
-
+
diff --git a/parsing/g_intsyntax.ml b/parsing/g_intsyntax.ml
deleted file mode 100644
index 64fa0b49..00000000
--- a/parsing/g_intsyntax.ml
+++ /dev/null
@@ -1,344 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(*i $Id: g_intsyntax.ml 12509 2009-11-12 15:52:50Z letouzey $ i*)
-
-(* digit-based syntax for int31, bigN bigZ and bigQ *)
-
-open Bigint
-open Libnames
-open Rawterm
-
-(*** Constants for locating the int31 and bigN ***)
-
-
-
-let make_dir l = Names.make_dirpath (List.map Names.id_of_string (List.rev l))
-let make_path dir id = Libnames.make_path (make_dir dir) (Names.id_of_string id)
-
-(* copied on g_zsyntax.ml, where it is said to be a temporary hack*)
-(* takes a path an identifier in the form of a string list and a string,
- returns a kernel_name *)
-let make_kn dir id = Libnames.encode_kn (make_dir dir) (Names.id_of_string id)
-
-
-(* int31 stuff *)
-let int31_module = ["Coq"; "Numbers"; "Cyclic"; "Int31"; "Int31"]
-let int31_path = make_path int31_module "int31"
-let int31_id = make_kn int31_module
-let int31_scope = "int31_scope"
-
-let int31_construct = ConstructRef ((int31_id "int31",0),1)
-
-let int31_0 = ConstructRef ((int31_id "digits",0),1)
-let int31_1 = ConstructRef ((int31_id "digits",0),2)
-
-
-(* bigN stuff*)
-let zn2z_module = ["Coq"; "Numbers"; "Cyclic"; "DoubleCyclic"; "DoubleType"]
-let zn2z_path = make_path zn2z_module "zn2z"
-let zn2z_id = make_kn zn2z_module
-
-let zn2z_W0 = ConstructRef ((zn2z_id "zn2z",0),1)
-let zn2z_WW = ConstructRef ((zn2z_id "zn2z",0),2)
-
-let bigN_module = ["Coq"; "Numbers"; "Natural"; "BigN"; "BigN" ]
-let bigN_path = make_path (bigN_module@["BigN"]) "t"
-(* big ugly hack *)
-let bigN_id id = (Obj.magic ((Names.MPdot ((Names.MPfile (make_dir bigN_module)),
- Names.mk_label "BigN")),
- [], Names.id_of_string id) : Names.kernel_name)
-let bigN_scope = "bigN_scope"
-
-(* number of inlined level of bigN (actually the level 0 to n_inlined-1 are inlined) *)
-let n_inlined = of_string "7"
-let bigN_constructor =
- (* converts a bigint into an int the ugly way *)
- let rec to_int i =
- if equal i zero then
- 0
- else
- let (quo,rem) = div2_with_rest i in
- if rem then
- 2*(to_int quo)+1
- else
- 2*(to_int quo)
- in
- fun i ->
- ConstructRef ((bigN_id "t_",0),
- if less_than i n_inlined then
- (to_int i)+1
- else
- (to_int n_inlined)+1
- )
-
-(*bigZ stuff*)
-let bigZ_module = ["Coq"; "Numbers"; "Integer"; "BigZ"; "BigZ" ]
-let bigZ_path = make_path (bigZ_module@["BigZ"]) "t"
-(* big ugly hack bis *)
-let bigZ_id id = (Obj.magic ((Names.MPdot ((Names.MPfile (make_dir bigZ_module)),
- Names.mk_label "BigZ")),
- [], Names.id_of_string id) : Names.kernel_name)
-let bigZ_scope = "bigZ_scope"
-
-let bigZ_pos = ConstructRef ((bigZ_id "t_",0),1)
-let bigZ_neg = ConstructRef ((bigZ_id "t_",0),2)
-
-
-(*bigQ stuff*)
-let bigQ_module = ["Coq"; "Numbers"; "Rational"; "BigQ"; "BigQ"]
-let qmake_module = ["Coq"; "Numbers"; "Rational"; "BigQ"; "QMake"]
-let bigQ_path = make_path (bigQ_module@["BigQ"]) "t"
-(* big ugly hack bis *)
-let bigQ_id id = (Obj.magic ((Names.MPdot ((Names.MPfile (make_dir bigQ_module)),
- Names.mk_label "BigQ")),
- [], Names.id_of_string id) : Names.kernel_name)
-let bigQ_scope = "bigQ_scope"
-
-let bigQ_z = ConstructRef ((bigQ_id "t_",0),1)
-
-
-(*** Definition of the Non_closed exception, used in the pretty printing ***)
-exception Non_closed
-
-(*** Parsing for int31 in digital notation ***)
-
-(* parses a *non-negative* integer (from bigint.ml) into an int31
- wraps modulo 2^31 *)
-let int31_of_pos_bigint dloc n =
- let ref_construct = RRef (dloc, int31_construct) in
- let ref_0 = RRef (dloc, int31_0) in
- let ref_1 = RRef (dloc, int31_1) in
- let rec args counter n =
- if counter <= 0 then
- []
- else
- let (q,r) = div2_with_rest n in
- (if r then ref_1 else ref_0)::(args (counter-1) q)
- in
- RApp (dloc, ref_construct, List.rev (args 31 n))
-
-let error_negative dloc =
- Util.user_err_loc (dloc, "interp_int31", Pp.str "int31 are only non-negative numbers.")
-
-let interp_int31 dloc n =
- if is_pos_or_zero n then
- int31_of_pos_bigint dloc n
- else
- error_negative dloc
-
-(* Pretty prints an int31 *)
-
-let bigint_of_int31 =
- let rec args_parsing args cur =
- match args with
- | [] -> cur
- | (RRef (_,b))::l when b = int31_0 -> args_parsing l (mult_2 cur)
- | (RRef (_,b))::l when b = int31_1 -> args_parsing l (add_1 (mult_2 cur))
- | _ -> raise Non_closed
- in
- function
- | RApp (_, RRef (_, c), args) when c=int31_construct -> args_parsing args zero
- | _ -> raise Non_closed
-
-let uninterp_int31 i =
- try
- Some (bigint_of_int31 i)
- with Non_closed ->
- None
-
-(* Actually declares the interpreter for int31 *)
-let _ = Notation.declare_numeral_interpreter int31_scope
- (int31_path, int31_module)
- interp_int31
- ([RRef (Util.dummy_loc, int31_construct)],
- uninterp_int31,
- true)
-
-
-(*** Parsing for bigN in digital notation ***)
-(* the base for bigN (in Coq) that is 2^31 in our case *)
-let base = pow two (of_string "31")
-
-(* base of the bigN of height N : *)
-let rank n = pow base (pow two n)
-
-(* splits a number bi at height n, that is the rest needs 2^n int31 to be stored
- it is expected to be used only when the quotient would also need 2^n int31 to be
- stored *)
-let split_at n bi =
- euclid bi (rank (sub_1 n))
-
-(* search the height of the Coq bigint needed to represent the integer bi *)
-let height bi =
- let rec height_aux n =
- if less_than bi (rank n) then
- n
- else
- height_aux (add_1 n)
- in
- height_aux zero
-
-
-(* n must be a non-negative integer (from bigint.ml) *)
-let word_of_pos_bigint dloc hght n =
- let ref_W0 = RRef (dloc, zn2z_W0) in
- let ref_WW = RRef (dloc, zn2z_WW) in
- let rec decomp hgt n =
- if is_neg_or_zero hgt then
- int31_of_pos_bigint dloc n
- else if equal n zero then
- RApp (dloc, ref_W0, [RHole (dloc, Evd.InternalHole)])
- else
- let (h,l) = split_at hgt n in
- RApp (dloc, ref_WW, [RHole (dloc, Evd.InternalHole);
- decomp (sub_1 hgt) h;
- decomp (sub_1 hgt) l])
- in
- decomp hght n
-
-let bigN_of_pos_bigint dloc n =
- let ref_constructor i = RRef (dloc, bigN_constructor i) in
- let result h word = RApp (dloc, ref_constructor h, if less_than h n_inlined then
- [word]
- else
- [G_natsyntax.nat_of_int dloc (sub h n_inlined);
- word])
- in
- let hght = height n in
- result hght (word_of_pos_bigint dloc hght n)
-
-let bigN_error_negative dloc =
- Util.user_err_loc (dloc, "interp_bigN", Pp.str "bigN are only non-negative numbers.")
-
-let interp_bigN dloc n =
- if is_pos_or_zero n then
- bigN_of_pos_bigint dloc n
- else
- bigN_error_negative dloc
-
-
-(* Pretty prints a bigN *)
-
-let bigint_of_word =
- let rec get_height rc =
- match rc with
- | RApp (_,RRef(_,c), [_;lft;rght]) when c = zn2z_WW ->
- let hleft = get_height lft in
- let hright = get_height rght in
- add_1
- (if less_than hleft hright then
- hright
- else
- hleft)
- | _ -> zero
- in
- let rec transform hght rc =
- match rc with
- | RApp (_,RRef(_,c),_) when c = zn2z_W0-> zero
- | RApp (_,RRef(_,c), [_;lft;rght]) when c=zn2z_WW-> let new_hght = sub_1 hght in
- add (mult (rank new_hght)
- (transform (new_hght) lft))
- (transform (new_hght) rght)
- | _ -> bigint_of_int31 rc
- in
- fun rc ->
- let hght = get_height rc in
- transform hght rc
-
-let bigint_of_bigN rc =
- match rc with
- | RApp (_,_,[one_arg]) -> bigint_of_word one_arg
- | RApp (_,_,[_;second_arg]) -> bigint_of_word second_arg
- | _ -> raise Non_closed
-
-let uninterp_bigN rc =
- try
- Some (bigint_of_bigN rc)
- with Non_closed ->
- None
-
-
-(* declare the list of constructors of bigN used in the declaration of the
- numeral interpreter *)
-
-let bigN_list_of_constructors =
- let rec build i =
- if less_than i (add_1 n_inlined) then
- RRef (Util.dummy_loc, bigN_constructor i)::(build (add_1 i))
- else
- []
- in
- build zero
-
-(* Actually declares the interpreter for bigN *)
-let _ = Notation.declare_numeral_interpreter bigN_scope
- (bigN_path, bigN_module)
- interp_bigN
- (bigN_list_of_constructors,
- uninterp_bigN,
- true)
-
-
-(*** Parsing for bigZ in digital notation ***)
-let interp_bigZ dloc n =
- let ref_pos = RRef (dloc, bigZ_pos) in
- let ref_neg = RRef (dloc, bigZ_neg) in
- if is_pos_or_zero n then
- RApp (dloc, ref_pos, [bigN_of_pos_bigint dloc n])
- else
- RApp (dloc, ref_neg, [bigN_of_pos_bigint dloc (neg n)])
-
-(* pretty printing functions for bigZ *)
-let bigint_of_bigZ = function
- | RApp (_, RRef(_,c), [one_arg]) when c = bigZ_pos -> bigint_of_bigN one_arg
- | RApp (_, RRef(_,c), [one_arg]) when c = bigZ_neg ->
- let opp_val = bigint_of_bigN one_arg in
- if equal opp_val zero then
- raise Non_closed
- else
- neg opp_val
- | _ -> raise Non_closed
-
-
-let uninterp_bigZ rc =
- try
- Some (bigint_of_bigZ rc)
- with Non_closed ->
- None
-
-(* Actually declares the interpreter for bigN *)
-let _ = Notation.declare_numeral_interpreter bigZ_scope
- (bigZ_path, bigZ_module)
- interp_bigZ
- ([RRef (Util.dummy_loc, bigZ_pos);
- RRef (Util.dummy_loc, bigZ_neg)],
- uninterp_bigZ,
- true)
-
-(*** Parsing for bigQ in digital notation ***)
-let interp_bigQ dloc n =
- let ref_z = RRef (dloc, bigQ_z) in
- let ref_pos = RRef (dloc, bigZ_pos) in
- let ref_neg = RRef (dloc, bigZ_neg) in
- if is_pos_or_zero n then
- RApp (dloc, ref_z,
- [RApp (dloc, ref_pos, [bigN_of_pos_bigint dloc n])])
- else
- RApp (dloc, ref_z,
- [RApp (dloc, ref_neg, [bigN_of_pos_bigint dloc (neg n)])])
-
-let uninterp_bigQ rc = None
-
-
-(* Actually declares the interpreter for bigQ *)
-let _ = Notation.declare_numeral_interpreter bigQ_scope
- (bigQ_path, bigQ_module)
- interp_bigQ
- ([], uninterp_bigQ,
- true)
diff --git a/parsing/g_ltac.ml4 b/parsing/g_ltac.ml4
index 316bf8e1..0e97b2a7 100644
--- a/parsing/g_ltac.ml4
+++ b/parsing/g_ltac.ml4
@@ -8,7 +8,7 @@
(*i camlp4use: "pa_extend.cmo" i*)
-(* $Id: g_ltac.ml4 11576 2008-11-10 19:13:15Z msozeau $ *)
+(* $Id$ *)
open Pp
open Util
@@ -35,7 +35,7 @@ GEXTEND Gram
tactic_then_last:
[ [ "|"; lta = LIST0 OPT tactic_expr SEP "|" ->
Array.map (function None -> TacId [] | Some t -> t) (Array.of_list lta)
- | -> [||]
+ | -> [||]
] ]
;
tactic_then_gen:
@@ -54,7 +54,7 @@ GEXTEND Gram
[ ta0 = tactic_expr; ";"; ta1 = binder_tactic -> TacThen (ta0, [||], ta1, [||])
| ta0 = tactic_expr; ";"; ta1 = tactic_expr -> TacThen (ta0, [||], ta1, [||])
| ta0 = tactic_expr; ";"; "["; (first,tail) = tactic_then_gen; "]" ->
- match tail with
+ match tail with
| Some (t,last) -> TacThen (ta0, Array.of_list first, t, last)
| None -> TacThens (ta0,first) ]
| "3" RIGHTA
@@ -94,7 +94,7 @@ GEXTEND Gram
TacArg(MetaIdArg (loc,false,id))
| IDENT "constr"; ":"; c = Constr.constr ->
TacArg(ConstrMayEval(ConstrTerm c))
- | IDENT "ipattern"; ":"; ipat = simple_intropattern ->
+ | IDENT "ipattern"; ":"; ipat = simple_intropattern ->
TacArg(IntroPattern ipat)
| r = reference; la = LIST0 tactic_arg ->
TacArg(TacCall (loc,r,la)) ]
@@ -107,7 +107,7 @@ GEXTEND Gram
[ RIGHTA
[ "fun"; it = LIST1 input_fun ; "=>"; body = tactic_expr LEVEL "5" ->
TacFun (it,body)
- | "let"; isrec = [IDENT "rec" -> true | -> false];
+ | "let"; isrec = [IDENT "rec" -> true | -> false];
llc = LIST1 let_clause SEP "with"; "in";
body = tactic_expr LEVEL "5" -> TacLetIn (isrec,llc,body)
| IDENT "info"; tc = tactic_expr LEVEL "5" -> TacInfo tc ] ]
@@ -153,7 +153,7 @@ GEXTEND Gram
[ [ "match" -> false | "lazymatch" -> true ] ]
;
input_fun:
- [ [ "_" -> None
+ [ [ "_" -> None
| l = ident -> Some l ] ]
;
let_clause:
@@ -172,11 +172,11 @@ GEXTEND Gram
| pc = Constr.lconstr_pattern -> Term pc ] ]
;
match_hyps:
- [ [ na = name; ":"; mp = match_pattern -> Hyp (na, mp)
- | na = name; ":="; "["; mpv = match_pattern; "]"; ":"; mpt = match_pattern -> Def (na, mpv, mpt)
- | na = name; ":="; mpv = match_pattern ->
+ [ [ na = name; ":"; mp = match_pattern -> Hyp (na, mp)
+ | na = name; ":="; "["; mpv = match_pattern; "]"; ":"; mpt = match_pattern -> Def (na, mpv, mpt)
+ | na = name; ":="; mpv = match_pattern ->
let t, ty =
- match mpv with
+ match mpv with
| Term t -> (match t with
| CCast (loc, t, CastConv (_, ty)) -> Term t, Some (Term ty)
| _ -> mpv, None)
@@ -213,7 +213,7 @@ GEXTEND Gram
[ [ ":=" -> false
| "::=" -> true ] ]
;
-
+
(* Definitions for tactics *)
tacdef_body:
[ [ name = Constr.global; it=LIST1 input_fun; redef = ltac_def_kind; body = tactic_expr ->
@@ -224,9 +224,9 @@ GEXTEND Gram
tactic:
[ [ tac = tactic_expr -> tac ] ]
;
- Vernac_.command:
+ Vernac_.command:
[ [ IDENT "Ltac";
l = LIST1 tacdef_body SEP "with" ->
- VernacDeclareTacticDefinition (true, l) ] ]
+ VernacDeclareTacticDefinition (use_module_locality (), true, l) ] ]
;
END
diff --git a/parsing/g_minicoq.mli b/parsing/g_minicoq.mli
deleted file mode 100644
index 345d9575..00000000
--- a/parsing/g_minicoq.mli
+++ /dev/null
@@ -1,31 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(*i $Id: g_minicoq.mli 5920 2004-07-16 20:01:26Z herbelin $ i*)
-
-(*i*)
-open Pp
-open Names
-open Term
-open Environ
-(*i*)
-
-val term : constr Grammar.Entry.e
-
-type command =
- | Definition of identifier * constr option * constr
- | Parameter of identifier * constr
- | Variable of identifier * constr
- | Inductive of
- (identifier * constr) list *
- (identifier * constr * (identifier * constr) list) list
- | Check of constr
-
-val command : command Grammar.Entry.e
-
-val pr_term : path_kind -> env -> constr -> std_ppcmds
diff --git a/parsing/g_natsyntax.ml b/parsing/g_natsyntax.ml
deleted file mode 100644
index 8804d81a..00000000
--- a/parsing/g_natsyntax.ml
+++ /dev/null
@@ -1,78 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(* $Id: g_natsyntax.ml 10348 2007-12-06 17:36:14Z aspiwack $ *)
-
-(* This file defines the printer for natural numbers in [nat] *)
-
-(*i*)
-open Pcoq
-open Pp
-open Util
-open Names
-open Coqlib
-open Rawterm
-open Libnames
-open Bigint
-open Coqlib
-open Notation
-open Pp
-open Util
-open Names
-(*i*)
-
-(**********************************************************************)
-(* Parsing via scopes *)
-(* For example, (nat_of_string "3") is <<(S (S (S O)))>> *)
-
-let nat_of_int dloc n =
- if is_pos_or_zero n then begin
- if less_than (of_string "5000") n then
- Flags.if_warn msg_warning
- (strbrk "Stack overflow or segmentation fault happens when " ++
- strbrk "working with large numbers in nat (observed threshold " ++
- strbrk "may vary from 5000 to 70000 depending on your system " ++
- strbrk "limits and on the command executed).");
- let ref_O = RRef (dloc, glob_O) in
- let ref_S = RRef (dloc, glob_S) in
- let rec mk_nat acc n =
- if n <> zero then
- mk_nat (RApp (dloc,ref_S, [acc])) (sub_1 n)
- else
- acc
- in
- mk_nat ref_O n
- end
- else
- user_err_loc (dloc, "nat_of_int",
- str "Cannot interpret a negative number as a number of type nat")
-
-(************************************************************************)
-(* Printing via scopes *)
-
-exception Non_closed_number
-
-let rec int_of_nat = function
- | RApp (_,RRef (_,s),[a]) when s = glob_S -> add_1 (int_of_nat a)
- | RRef (_,z) when z = glob_O -> zero
- | _ -> raise Non_closed_number
-
-let uninterp_nat p =
- try
- Some (int_of_nat p)
- with
- Non_closed_number -> None
-
-(************************************************************************)
-(* Declare the primitive parsers and printers *)
-
-let _ =
- Notation.declare_numeral_interpreter "nat_scope"
- (nat_path,["Coq";"Init";"Datatypes"])
- nat_of_int
- ([RRef (dummy_loc,glob_S); RRef (dummy_loc,glob_O)], uninterp_nat, true)
diff --git a/parsing/g_natsyntax.mli b/parsing/g_natsyntax.mli
index 09095959..bf3921bf 100644
--- a/parsing/g_natsyntax.mli
+++ b/parsing/g_natsyntax.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: g_natsyntax.mli 11087 2008-06-10 13:29:52Z letouzey $ i*)
+(*i $Id$ i*)
(* Nice syntax for naturals. *)
diff --git a/parsing/g_prim.ml4 b/parsing/g_prim.ml4
index 76225d77..6e7acd3f 100644
--- a/parsing/g_prim.ml4
+++ b/parsing/g_prim.ml4
@@ -8,7 +8,7 @@
(*i camlp4use: "pa_extend.cmo" i*)
-(*i $Id: g_prim.ml4 11525 2008-10-30 22:18:54Z amahboub $ i*)
+(*i $Id$ i*)
open Pcoq
open Names
@@ -34,10 +34,10 @@ let my_int_of_string loc s =
Util.user_err_loc (loc,"",Pp.str "Cannot support a so large number.")
GEXTEND Gram
- GLOBAL:
+ GLOBAL:
bigint natural integer identref name ident var preident
- fullyqualid qualid reference dirpath
- ne_string string pattern_ident pattern_identref;
+ fullyqualid qualid reference dirpath ne_lstring
+ ne_string string pattern_ident pattern_identref by_notation smart_global;
preident:
[ [ s = IDENT -> s ] ]
;
@@ -45,7 +45,7 @@ GEXTEND Gram
[ [ s = IDENT -> id_of_string s ] ]
;
pattern_ident:
- [ [ LEFTQMARK; id = ident -> id ] ]
+ [ [ s = LEFTQMARK; id = ident -> id ] ]
;
pattern_identref:
[ [ id = pattern_ident -> (loc, id) ] ]
@@ -71,7 +71,7 @@ GEXTEND Gram
;
basequalid:
[ [ id = ident; (l,id')=fields -> local_make_qualid (l@[id]) id'
- | id = ident -> make_short_qualid id
+ | id = ident -> qualid_of_ident id
] ]
;
name:
@@ -84,14 +84,24 @@ GEXTEND Gram
| id = ident -> Ident (loc,id)
] ]
;
+ by_notation:
+ [ [ s = ne_string; sc = OPT ["%"; key = IDENT -> key ] -> (loc,s,sc) ] ]
+ ;
+ smart_global:
+ [ [ c = reference -> Genarg.AN c
+ | ntn = by_notation -> Genarg.ByNotation ntn ] ]
+ ;
qualid:
[ [ qid = basequalid -> loc, qid ] ]
;
ne_string:
- [ [ s = STRING ->
+ [ [ s = STRING ->
if s="" then Util.user_err_loc(loc,"",Pp.str"Empty string."); s
] ]
;
+ ne_lstring:
+ [ [ s = ne_string -> (loc,s) ] ]
+ ;
dirpath:
[ [ id = ident; l = LIST0 field ->
make_dirpath (l@[id]) ] ]
diff --git a/parsing/g_proofs.ml4 b/parsing/g_proofs.ml4
index 655bb267..39e577b8 100644
--- a/parsing/g_proofs.ml4
+++ b/parsing/g_proofs.ml4
@@ -8,7 +8,7 @@
(*i camlp4use: "pa_extend.cmo" i*)
-(* $Id: g_proofs.ml4 11784 2009-01-14 11:36:32Z herbelin $ *)
+(* $Id$ *)
open Pcoq
@@ -53,7 +53,7 @@ GEXTEND Gram
| IDENT "Save"; id = identref ->
VernacEndProof (Proved (true,Some (id,None)))
| IDENT "Defined" -> VernacEndProof (Proved (false,None))
- | IDENT "Defined"; id=identref ->
+ | IDENT "Defined"; id=identref ->
VernacEndProof (Proved (false,Some (id,None)))
| IDENT "Suspend" -> VernacSuspend
| IDENT "Resume" -> VernacResume None
@@ -82,7 +82,7 @@ GEXTEND Gram
| IDENT "Show"; IDENT "Thesis" -> VernacShow ShowThesis
| IDENT "Explain"; IDENT "Proof"; l = LIST0 integer ->
VernacShow (ExplainProof l)
- | IDENT "Explain"; IDENT "Proof"; IDENT "Tree"; l = LIST0 integer ->
+ | IDENT "Explain"; IDENT "Proof"; IDENT "Tree"; l = LIST0 integer ->
VernacShow (ExplainTree l)
| IDENT "Go"; n = natural -> VernacGo (GoTo n)
| IDENT "Go"; IDENT "top" -> VernacGo GoTop
@@ -90,16 +90,22 @@ GEXTEND Gram
| IDENT "Go"; IDENT "next" -> VernacGo GoNext
| IDENT "Guarded" -> VernacCheckGuard
(* Hints for Auto and EAuto *)
- | IDENT "Create"; IDENT "HintDb" ;
+ | IDENT "Create"; IDENT "HintDb" ;
id = IDENT ; b = [ "discriminated" -> true | -> false ] ->
- VernacCreateHintDb (use_locality (), id, b)
+ VernacCreateHintDb (use_module_locality (), id, b)
| IDENT "Hint"; local = obsolete_locality; h = hint;
- dbnames = opt_hintbases ->
- VernacHints (enforce_locality_of local,dbnames, h)
-
+ dbnames = opt_hintbases ->
+ VernacHints (enforce_module_locality local,dbnames, h)
+
+(* Declare "Resolve" directly so as to be able to later extend with
+ "Resolve ->" and "Resolve <-" *)
+ | IDENT "Hint"; IDENT "Resolve"; lc = LIST1 constr; n = OPT natural;
+ dbnames = opt_hintbases ->
+ VernacHints (use_module_locality (),dbnames,
+ HintsResolve (List.map (fun x -> (n, true, x)) lc))
(*This entry is not commented, only for debug*)
- | IDENT "PrintConstr"; c = constr ->
+ | IDENT "PrintConstr"; c = constr ->
VernacExtend ("PrintConstr",
[Genarg.in_gen Genarg.rawwit_constr c])
] ];
@@ -108,7 +114,7 @@ GEXTEND Gram
[ [ IDENT "Local" -> true | -> false ] ]
;
hint:
- [ [ IDENT "Resolve"; lc = LIST1 constr; n = OPT [ n = natural -> n ] ->
+ [ [ IDENT "Resolve"; lc = LIST1 constr; n = OPT natural ->
HintsResolve (List.map (fun x -> (n, true, x)) lc)
| IDENT "Immediate"; lc = LIST1 constr -> HintsImmediate lc
| IDENT "Transparent"; lc = LIST1 global -> HintsTransparency (lc, true)
@@ -118,7 +124,7 @@ GEXTEND Gram
| IDENT "Extern"; n = natural; c = OPT constr_pattern ; "=>";
tac = tactic ->
HintsExtern (n,c,tac)
- | IDENT "Destruct";
+ | IDENT "Destruct";
id = ident; ":=";
pri = natural;
dloc = destruct_location;
diff --git a/parsing/g_rsyntax.ml b/parsing/g_rsyntax.ml
deleted file mode 100644
index b3425899..00000000
--- a/parsing/g_rsyntax.ml
+++ /dev/null
@@ -1,125 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(*i $Id: g_rsyntax.ml 10739 2008-04-01 14:45:20Z herbelin $ i*)
-
-open Pp
-open Util
-open Names
-open Pcoq
-open Topconstr
-open Libnames
-
-exception Non_closed_number
-
-(**********************************************************************)
-(* Parsing R via scopes *)
-(**********************************************************************)
-
-open Libnames
-open Rawterm
-open Bigint
-
-let make_dir l = make_dirpath (List.map id_of_string (List.rev l))
-let rdefinitions = make_dir ["Coq";"Reals";"Rdefinitions"]
-let make_path dir id = Libnames.make_path dir (id_of_string id)
-
-let r_path = make_path rdefinitions "R"
-
-(* TODO: temporary hack *)
-let make_path dir id = Libnames.encode_con dir (id_of_string id)
-
-let r_kn = make_path rdefinitions "R"
-let glob_R = ConstRef r_kn
-let glob_R1 = ConstRef (make_path rdefinitions "R1")
-let glob_R0 = ConstRef (make_path rdefinitions "R0")
-let glob_Ropp = ConstRef (make_path rdefinitions "Ropp")
-let glob_Rplus = ConstRef (make_path rdefinitions "Rplus")
-let glob_Rmult = ConstRef (make_path rdefinitions "Rmult")
-
-let two = mult_2 one
-let three = add_1 two
-let four = mult_2 two
-
-(* Unary representation of strictly positive numbers *)
-let rec small_r dloc n =
- if equal one n then RRef (dloc, glob_R1)
- else RApp(dloc,RRef (dloc,glob_Rplus),
- [RRef (dloc, glob_R1);small_r dloc (sub_1 n)])
-
-let r_of_posint dloc n =
- let r1 = RRef (dloc, glob_R1) in
- let r2 = small_r dloc two in
- let rec r_of_pos n =
- if less_than n four then small_r dloc n
- else
- let (q,r) = div2_with_rest n in
- let b = RApp(dloc,RRef(dloc,glob_Rmult),[r2;r_of_pos q]) in
- if r then RApp(dloc,RRef(dloc,glob_Rplus),[r1;b]) else b in
- if n <> zero then r_of_pos n else RRef(dloc,glob_R0)
-
-let r_of_int dloc z =
- if is_strictly_neg z then
- RApp (dloc, RRef(dloc,glob_Ropp), [r_of_posint dloc (neg z)])
- else
- r_of_posint dloc z
-
-(**********************************************************************)
-(* Printing R via scopes *)
-(**********************************************************************)
-
-let bignat_of_r =
-(* for numbers > 1 *)
-let rec bignat_of_pos = function
- (* 1+1 *)
- | RApp (_,RRef (_,p), [RRef (_,o1); RRef (_,o2)])
- when p = glob_Rplus & o1 = glob_R1 & o2 = glob_R1 -> two
- (* 1+(1+1) *)
- | RApp (_,RRef (_,p1), [RRef (_,o1);
- RApp(_,RRef (_,p2),[RRef(_,o2);RRef(_,o3)])])
- when p1 = glob_Rplus & p2 = glob_Rplus &
- o1 = glob_R1 & o2 = glob_R1 & o3 = glob_R1 -> three
- (* (1+1)*b *)
- | RApp (_,RRef (_,p), [a; b]) when p = glob_Rmult ->
- if bignat_of_pos a <> two then raise Non_closed_number;
- mult_2 (bignat_of_pos b)
- (* 1+(1+1)*b *)
- | RApp (_,RRef (_,p1), [RRef (_,o); RApp (_,RRef (_,p2),[a;b])])
- when p1 = glob_Rplus & p2 = glob_Rmult & o = glob_R1 ->
- if bignat_of_pos a <> two then raise Non_closed_number;
- add_1 (mult_2 (bignat_of_pos b))
- | _ -> raise Non_closed_number
-in
-let bignat_of_r = function
- | RRef (_,a) when a = glob_R0 -> zero
- | RRef (_,a) when a = glob_R1 -> one
- | r -> bignat_of_pos r
-in
-bignat_of_r
-
-let bigint_of_r = function
- | RApp (_,RRef (_,o), [a]) when o = glob_Ropp ->
- let n = bignat_of_r a in
- if n = zero then raise Non_closed_number;
- neg n
- | a -> bignat_of_r a
-
-let uninterp_r p =
- try
- Some (bigint_of_r p)
- with Non_closed_number ->
- None
-
-let _ = Notation.declare_numeral_interpreter "R_scope"
- (r_path,["Coq";"Reals";"Rdefinitions"])
- r_of_int
- ([RRef(dummy_loc,glob_Ropp);RRef(dummy_loc,glob_R0);
- RRef(dummy_loc,glob_Rplus);RRef(dummy_loc,glob_Rmult);
- RRef(dummy_loc,glob_R1)],
- uninterp_r,
- false)
diff --git a/parsing/g_string_syntax.ml b/parsing/g_string_syntax.ml
deleted file mode 100644
index 6a650987..00000000
--- a/parsing/g_string_syntax.ml
+++ /dev/null
@@ -1,69 +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: g_string_syntax.ml 10739 2008-04-01 14:45:20Z herbelin $ i*)
-
-open Pp
-open Util
-open Names
-open Pcoq
-open Libnames
-open Topconstr
-open G_ascii_syntax
-open Rawterm
-open Coqlib
-
-exception Non_closed_string
-
-(* make a string term from the string s *)
-
-let string_module = ["Coq";"Strings";"String"]
-
-let string_path = make_path string_module "string"
-
-let string_kn = make_kn string_module "string"
-let static_glob_EmptyString = ConstructRef ((string_kn,0),1)
-let static_glob_String = ConstructRef ((string_kn,0),2)
-
-let make_reference id = find_reference "String interpretation" string_module id
-let glob_String = lazy (make_reference "String")
-let glob_EmptyString = lazy (make_reference "EmptyString")
-
-open Lazy
-
-let interp_string dloc s =
- let le = String.length s in
- let rec aux n =
- if n = le then RRef (dloc, force glob_EmptyString) else
- RApp (dloc,RRef (dloc, force glob_String),
- [interp_ascii dloc (int_of_char s.[n]); aux (n+1)])
- in aux 0
-
-let uninterp_string r =
- try
- let b = Buffer.create 16 in
- let rec aux = function
- | RApp (_,RRef (_,k),[a;s]) when k = force glob_String ->
- (match uninterp_ascii a with
- | Some c -> Buffer.add_char b (Char.chr c); aux s
- | _ -> raise Non_closed_string)
- | RRef (_,z) when z = force glob_EmptyString ->
- Some (Buffer.contents b)
- | _ ->
- raise Non_closed_string
- in aux r
- with
- Non_closed_string -> None
-
-let _ =
- Notation.declare_string_interpreter "string_scope"
- (string_path,["Coq";"Strings";"String"])
- interp_string
- ([RRef (dummy_loc,static_glob_String);
- RRef (dummy_loc,static_glob_EmptyString)],
- uninterp_string, true)
diff --git a/parsing/g_tactic.ml4 b/parsing/g_tactic.ml4
index ad093507..c845daf2 100644
--- a/parsing/g_tactic.ml4
+++ b/parsing/g_tactic.ml4
@@ -8,7 +8,7 @@
(*i camlp4use: "pa_extend.cmo" i*)
-(* $Id: g_tactic.ml4 12009 2009-03-23 22:55:37Z herbelin $ *)
+(* $Id$ *)
open Pp
open Pcoq
@@ -22,7 +22,7 @@ open Termops
let all_with delta = make_red_flag [FBeta;FIota;FZeta;delta]
-let tactic_kw = [ "->"; "<-" ]
+let tactic_kw = [ "->"; "<-" ; "by" ]
let _ = List.iter (fun s -> Lexer.add_token("",s)) tactic_kw
(* Hack to parse "(x:=t)" as an explicit argument without conflicts with the *)
@@ -147,9 +147,29 @@ let induction_arg_of_constr (c,lbind as clbind) =
with _ -> ElimOnConstr clbind
else ElimOnConstr clbind
+let mkTacCase with_evar = function
+ | [([ElimOnConstr cl],None,(None,None))],None ->
+ TacCase (with_evar,cl)
+ (* Reinterpret numbers as a notation for terms *)
+ | [([(ElimOnAnonHyp n)],None,(None,None))],None ->
+ TacCase (with_evar,
+ (CPrim (dummy_loc, Numeral (Bigint.of_string (string_of_int n))),
+ NoBindings))
+ (* Reinterpret ident as notations for variables in the context *)
+ (* because we don't know if they are quantified or not *)
+ | [([ElimOnIdent id],None,(None,None))],None ->
+ TacCase (with_evar,(CRef (Ident id),NoBindings))
+ | ic ->
+ if List.exists (fun (cl,a,b) ->
+ List.exists (function ElimOnAnonHyp _ -> true | _ -> false) cl)
+ (fst ic)
+ then
+ error "Use of numbers as direct arguments of 'case' is not supported.";
+ TacInductionDestruct (false,with_evar,ic)
+
let rec mkCLambdaN_simple_loc loc bll c =
match bll with
- | ((loc1,_)::_ as idl,bk,t) :: bll ->
+ | ((loc1,_)::_ as idl,bk,t) :: bll ->
CLambdaN (loc,[idl,bk,t],mkCLambdaN_simple_loc (join_loc loc1 loc) bll c)
| ([],_,_) :: bll -> mkCLambdaN_simple_loc loc bll c
| [] -> c
@@ -166,11 +186,39 @@ let map_int_or_var f = function
| Rawterm.ArgArg x -> Rawterm.ArgArg (f x)
| Rawterm.ArgVar _ as y -> y
+let all_concl_occs_clause = { onhyps=Some[]; concl_occs=all_occurrences_expr }
+
+let has_no_specified_occs cl =
+ (cl.onhyps = None ||
+ List.for_all (fun ((occs,_),_) -> occs = all_occurrences_expr)
+ (Option.get cl.onhyps))
+ && (cl.concl_occs = all_occurrences_expr
+ || cl.concl_occs = no_occurrences_expr)
+
+let merge_occurrences loc cl = function
+ | None ->
+ if has_no_specified_occs cl then (None, cl)
+ else
+ user_err_loc (loc,"",str "Found an \"at\" clause without \"with\" clause.")
+ | Some (occs,p) ->
+ (Some p,
+ if occs = all_occurrences_expr then cl
+ else if cl = all_concl_occs_clause then { onhyps=Some[]; concl_occs=occs }
+ else match cl.onhyps with
+ | Some [(occs',id),l] when
+ occs' = all_occurrences_expr && cl.concl_occs = no_occurrences_expr ->
+ { cl with onhyps=Some[(occs,id),l] }
+ | _ ->
+ if has_no_specified_occs cl then
+ user_err_loc (loc,"",str "Unable to interpret the \"at\" clause; move it in the \"in\" clause.")
+ else
+ user_err_loc (loc,"",str "Cannot use clause \"at\" twice."))
+
(* Auxiliary grammar rules *)
GEXTEND Gram
GLOBAL: simple_tactic constr_with_bindings quantified_hypothesis
- bindings red_expr int_or_var open_constr casted_open_constr
+ bindings red_expr int_or_var open_constr casted_open_constr
simple_intropattern;
int_or_var:
@@ -183,7 +231,7 @@ GEXTEND Gram
;
(* An identifier or a quotation meta-variable *)
id_or_meta:
- [ [ id = identref -> AI id
+ [ [ id = identref -> AI id
(* This is used in quotations *)
| id = METAIDENT -> MetaId (loc,id) ] ]
@@ -215,19 +263,14 @@ GEXTEND Gram
| c1 = constr; "at"; occs = occs_nums; "with"; c2 = constr ->
(Some (occs,c1), c2) ] ]
;
- smart_global:
- [ [ c = global -> AN c
- | s = ne_string; sc = OPT ["%"; key = IDENT -> key ] ->
- ByNotation (loc,s,sc) ] ]
- ;
occs_nums:
[ [ nl = LIST1 nat_or_var -> no_occurrences_expr_but nl
| "-"; n = nat_or_var; nl = LIST0 int_or_var ->
(* have used int_or_var instead of nat_or_var for compatibility *)
all_occurrences_expr_but (List.map (map_int_or_var abs) (n::nl)) ] ]
- ;
+ ;
occs:
- [ [ "at"; occs = occs_nums -> occs | -> all_occurrences_expr_but [] ] ]
+ [ [ "at"; occs = occs_nums -> occs | -> all_occurrences_expr ] ]
;
pattern_occ:
[ [ c = constr; nl = occs -> (nl,c) ] ]
@@ -242,13 +285,13 @@ GEXTEND Gram
[ [ "["; tc = LIST1 intropatterns SEP "|"; "]" -> loc,IntroOrAndPattern tc
| "()" -> loc,IntroOrAndPattern [[]]
| "("; si = simple_intropattern; ")" -> loc,IntroOrAndPattern [[si]]
- | "("; si = simple_intropattern; ",";
- tc = LIST1 simple_intropattern SEP "," ; ")" ->
+ | "("; si = simple_intropattern; ",";
+ tc = LIST1 simple_intropattern SEP "," ; ")" ->
loc,IntroOrAndPattern [si::tc]
- | "("; si = simple_intropattern; "&";
- tc = LIST1 simple_intropattern SEP "&" ; ")" ->
+ | "("; si = simple_intropattern; "&";
+ tc = LIST1 simple_intropattern SEP "&" ; ")" ->
(* (A & B & C) is translated into (A,(B,C)) *)
- let rec pairify = function
+ let rec pairify = function
| ([]|[_]|[_;_]) as l -> IntroOrAndPattern [l]
| t::q -> IntroOrAndPattern [[t;(loc_of_ne_list q,pairify q)]]
in loc,pairify (si::tc) ] ]
@@ -256,10 +299,12 @@ GEXTEND Gram
naming_intropattern:
[ [ prefix = pattern_ident -> loc, IntroFresh prefix
| "?" -> loc, IntroAnonymous
- | id = ident -> loc, IntroIdentifier id ] ]
+ | id = ident -> loc, IntroIdentifier id
+ | "*" -> loc, IntroForthcoming true
+ | "**" -> loc, IntroForthcoming false ] ]
;
intropattern_modifier:
- [ [ IDENT "_eqn";
+ [ [ IDENT "_eqn";
id = [ ":"; id = naming_intropattern -> id | -> loc, IntroAnonymous ]
-> id ] ]
;
@@ -357,7 +402,7 @@ GEXTEND Gram
clause_dft_concl:
[ [ "in"; cl = in_clause -> cl
| occs=occs -> {onhyps=Some[]; concl_occs=occs}
- | -> {onhyps=Some[]; concl_occs=all_occurrences_expr} ] ]
+ | -> all_concl_occs_clause ] ]
;
clause_dft_all:
[ [ "in"; cl = in_clause -> cl
@@ -378,14 +423,14 @@ GEXTEND Gram
[ [ "in"; id = id_or_meta; ipat = as_ipat -> Some (id,ipat)
| -> None ] ]
;
- orient:
- [ [ "->" -> true
+ orient:
+ [ [ "->" -> true
| "<-" -> false
| -> true ]]
;
simple_binder:
[ [ na=name -> ([na],Default Explicit,CHole (loc, None))
- | "("; nal=LIST1 name; ":"; c=lconstr; ")" -> (nal,Default Explicit,c)
+ | "("; nal=LIST1 name; ":"; c=lconstr; ")" -> (nal,Default Explicit,c)
] ]
;
fixdecl:
@@ -401,7 +446,7 @@ GEXTEND Gram
(loc,id,bl,None,ty) ] ]
;
bindings_with_parameters:
- [ [ check_for_coloneq; "("; id = ident; bl = LIST0 simple_binder;
+ [ [ check_for_coloneq; "("; id = ident; bl = LIST0 simple_binder;
":="; c = lconstr; ")" -> (id, mkCLambdaN_simple bl c) ] ]
;
hintbases:
@@ -433,17 +478,17 @@ GEXTEND Gram
[ [ "as"; id = ident -> Names.Name id | -> Names.Anonymous ] ]
;
by_tactic:
- [ [ IDENT "by"; tac = tactic_expr LEVEL "3" -> TacComplete tac
+ [ [ "by"; tac = tactic_expr LEVEL "3" -> TacComplete tac
| -> TacId [] ] ]
;
opt_by_tactic:
- [ [ IDENT "by"; tac = tactic_expr LEVEL "3" -> Some tac
+ [ [ "by"; tac = tactic_expr LEVEL "3" -> Some tac
| -> None ] ]
;
- rename :
+ rename :
[ [ id1 = id_or_meta; IDENT "into"; id2 = id_or_meta -> (id1,id2) ] ]
;
- rewriter :
+ rewriter :
[ [ "!"; c = constr_with_bindings -> (RepeatPlus,c)
| ["?"| LEFTQMARK]; c = constr_with_bindings -> (RepeatStar,c)
| n = natural; "!"; c = constr_with_bindings -> (Precisely n,c)
@@ -452,12 +497,18 @@ GEXTEND Gram
| c = constr_with_bindings -> (Precisely 1, c)
] ]
;
- oriented_rewriter :
+ oriented_rewriter :
[ [ b = orient; p = rewriter -> let (m,c) = p in (b,m,c) ] ]
- ;
+ ;
induction_clause:
- [ [ lc = LIST1 induction_arg; ipats = with_induction_names;
- el = OPT eliminator; cl = opt_clause -> (lc,el,ipats,cl) ] ]
+ [ [ lc = LIST1 induction_arg; ipats = with_induction_names;
+ el = OPT eliminator -> (lc,el,ipats) ] ]
+ ;
+ one_induction_clause:
+ [ [ ic = induction_clause; cl = opt_clause -> ([ic],cl) ] ]
+ ;
+ induction_clause_list:
+ [ [ ic = LIST1 induction_clause SEP ","; cl = opt_clause -> (ic,cl) ] ]
;
move_location:
[ [ IDENT "after"; id = id_or_meta -> MoveAfter id
@@ -466,9 +517,9 @@ GEXTEND Gram
| "at"; IDENT "top" -> MoveToEnd false ] ]
;
simple_tactic:
- [ [
+ [ [
(* Basic tactics *)
- IDENT "intros"; IDENT "until"; id = quantified_hypothesis ->
+ IDENT "intros"; IDENT "until"; id = quantified_hypothesis ->
TacIntrosUntil id
| IDENT "intros"; pl = intropatterns -> TacIntroPattern pl
| IDENT "intro"; id = ident; hto = move_location ->
@@ -482,7 +533,7 @@ GEXTEND Gram
| IDENT "exact_no_check"; c = constr -> TacExactNoCheck c
| IDENT "vm_cast_no_check"; c = constr -> TacVmCastNoCheck c
- | IDENT "apply"; cl = LIST1 constr_with_bindings SEP ",";
+ | IDENT "apply"; cl = LIST1 constr_with_bindings SEP ",";
inhyp = in_hyp_as -> TacApply (true,false,cl,inhyp)
| IDENT "eapply"; cl = LIST1 constr_with_bindings SEP ",";
inhyp = in_hyp_as -> TacApply (true,true,cl,inhyp)
@@ -495,8 +546,8 @@ GEXTEND Gram
| IDENT "eelim"; cl = constr_with_bindings; el = OPT eliminator ->
TacElim (true,cl,el)
| IDENT "elimtype"; c = constr -> TacElimType c
- | IDENT "case"; cl = constr_with_bindings -> TacCase (false,cl)
- | IDENT "ecase"; cl = constr_with_bindings -> TacCase (true,cl)
+ | IDENT "case"; icl = induction_clause_list -> mkTacCase false icl
+ | IDENT "ecase"; icl = induction_clause_list -> mkTacCase true icl
| IDENT "casetype"; c = constr -> TacCaseType c
| "fix"; n = natural -> TacFix (None,n)
| "fix"; id = ident; n = natural -> TacFix (Some id,n)
@@ -519,11 +570,11 @@ GEXTEND Gram
TacLetTac (na,c,p,false)
(* Begin compatibility *)
- | IDENT "assert"; test_lpar_id_coloneq; "("; (loc,id) = identref; ":=";
- c = lconstr; ")" ->
+ | IDENT "assert"; test_lpar_id_coloneq; "("; (loc,id) = identref; ":=";
+ c = lconstr; ")" ->
TacAssert (None,Some (loc,IntroIdentifier id),c)
- | IDENT "assert"; test_lpar_id_colon; "("; (loc,id) = identref; ":";
- c = lconstr; ")"; tac=by_tactic ->
+ | IDENT "assert"; test_lpar_id_colon; "("; (loc,id) = identref; ":";
+ c = lconstr; ")"; tac=by_tactic ->
TacAssert (Some tac,Some (loc,IntroIdentifier id),c)
(* End compatibility *)
@@ -538,8 +589,8 @@ GEXTEND Gram
| IDENT "generalize"; c = constr; l = LIST1 constr ->
let gen_everywhere c = ((all_occurrences_expr,c),Names.Anonymous) in
TacGeneralize (List.map gen_everywhere (c::l))
- | IDENT "generalize"; c = constr; lookup_at_as_coma; nl = occs;
- na = as_name;
+ | IDENT "generalize"; c = constr; lookup_at_as_coma; nl = occs;
+ na = as_name;
l = LIST0 [","; c = pattern_occ; na = as_name -> (c,na)] ->
TacGeneralize (((nl,c),na)::l)
| IDENT "generalize"; IDENT "dependent"; c = constr -> TacGeneralizeDep c
@@ -551,18 +602,18 @@ GEXTEND Gram
(* Derived basic tactics *)
| IDENT "simple"; IDENT"induction"; h = quantified_hypothesis ->
TacSimpleInductionDestruct (true,h)
- | IDENT "induction"; ic = induction_clause ->
- TacInductionDestruct (true,false,[ic])
- | IDENT "einduction"; ic = induction_clause ->
- TacInductionDestruct(true,true,[ic])
+ | IDENT "induction"; ic = one_induction_clause ->
+ TacInductionDestruct (true,false,ic)
+ | IDENT "einduction"; ic = one_induction_clause ->
+ TacInductionDestruct(true,true,ic)
| IDENT "double"; IDENT "induction"; h1 = quantified_hypothesis;
h2 = quantified_hypothesis -> TacDoubleInduction (h1,h2)
| IDENT "simple"; IDENT "destruct"; h = quantified_hypothesis ->
TacSimpleInductionDestruct (false,h)
- | IDENT "destruct"; ic = induction_clause ->
- TacInductionDestruct(false,false,[ic])
- | IDENT "edestruct"; ic = induction_clause ->
- TacInductionDestruct(false,true,[ic])
+ | IDENT "destruct"; icl = induction_clause_list ->
+ TacInductionDestruct(false,false,icl)
+ | IDENT "edestruct"; icl = induction_clause_list ->
+ TacInductionDestruct(false,true,icl)
| IDENT "decompose"; IDENT "record" ; c = constr -> TacDecomposeAnd c
| IDENT "decompose"; IDENT "sum"; c = constr -> TacDecomposeOr c
| IDENT "decompose"; "["; l = LIST1 smart_global; "]"; c = constr
@@ -600,10 +651,11 @@ GEXTEND Gram
| IDENT "eleft"; bl = with_bindings -> TacLeft (true,bl)
| IDENT "right"; bl = with_bindings -> TacRight (false,bl)
| IDENT "eright"; bl = with_bindings -> TacRight (true,bl)
- | IDENT "split"; bl = with_bindings -> TacSplit (false,false,bl)
- | IDENT "esplit"; bl = with_bindings -> TacSplit (true,false,bl)
- | "exists"; bl = opt_bindings -> TacSplit (false,true,bl)
- | IDENT "eexists"; bl = opt_bindings -> TacSplit (true,true,bl)
+ | IDENT "split"; bl = with_bindings -> TacSplit (false,false,[bl])
+ | IDENT "esplit"; bl = with_bindings -> TacSplit (true,false,[bl])
+ | "exists"; bll = LIST1 opt_bindings SEP "," -> TacSplit (false,true,bll)
+ | IDENT "eexists"; bll = LIST1 opt_bindings SEP "," ->
+ TacSplit (true,true,bll)
| IDENT "constructor"; n = num_or_meta; l = with_bindings ->
TacConstructor (false,n,l)
| IDENT "econstructor"; n = num_or_meta; l = with_bindings ->
@@ -614,33 +666,34 @@ GEXTEND Gram
(* Equivalence relations *)
| IDENT "reflexivity" -> TacReflexivity
| IDENT "symmetry"; cl = clause_dft_concl -> TacSymmetry cl
- | IDENT "transitivity"; c = constr -> TacTransitivity c
+ | IDENT "transitivity"; c = constr -> TacTransitivity (Some c)
+ | IDENT "etransitivity" -> TacTransitivity None
(* Equality and inversion *)
- | IDENT "rewrite"; l = LIST1 oriented_rewriter SEP ",";
+ | IDENT "rewrite"; l = LIST1 oriented_rewriter SEP ",";
cl = clause_dft_concl; t=opt_by_tactic -> TacRewrite (false,l,cl,t)
- | IDENT "erewrite"; l = LIST1 oriented_rewriter SEP ",";
+ | IDENT "erewrite"; l = LIST1 oriented_rewriter SEP ",";
cl = clause_dft_concl; t=opt_by_tactic -> TacRewrite (true,l,cl,t)
| IDENT "dependent"; k =
[ IDENT "simple"; IDENT "inversion" -> SimpleInversion
| IDENT "inversion" -> FullInversion
| IDENT "inversion_clear" -> FullInversionClear ];
- hyp = quantified_hypothesis;
+ hyp = quantified_hypothesis;
ids = with_inversion_names; co = OPT ["with"; c = constr -> c] ->
TacInversion (DepInversion (k,co,ids),hyp)
| IDENT "simple"; IDENT "inversion";
hyp = quantified_hypothesis; ids = with_inversion_names;
cl = in_hyp_list ->
TacInversion (NonDepInversion (SimpleInversion, cl, ids), hyp)
- | IDENT "inversion";
+ | IDENT "inversion";
hyp = quantified_hypothesis; ids = with_inversion_names;
cl = in_hyp_list ->
TacInversion (NonDepInversion (FullInversion, cl, ids), hyp)
- | IDENT "inversion_clear";
- hyp = quantified_hypothesis; ids = with_inversion_names;
+ | IDENT "inversion_clear";
+ hyp = quantified_hypothesis; ids = with_inversion_names;
cl = in_hyp_list ->
TacInversion (NonDepInversion (FullInversionClear, cl, ids), hyp)
- | IDENT "inversion"; hyp = quantified_hypothesis;
+ | IDENT "inversion"; hyp = quantified_hypothesis;
"using"; c = constr; cl = in_hyp_list ->
TacInversion (InversionUsing (c,cl), hyp)
@@ -648,7 +701,8 @@ GEXTEND Gram
| r = red_tactic; cl = clause_dft_concl -> TacReduce (r, cl)
(* Change ne doit pas s'appliquer dans un Definition t := Eval ... *)
| IDENT "change"; (oc,c) = conversion; cl = clause_dft_concl ->
- TacChange (oc,c,cl)
+ let p,cl = merge_occurrences loc cl oc in
+ TacChange (p,c,cl)
] ]
;
END;;
diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4
index f727dfea..36dd0de1 100644
--- a/parsing/g_vernac.ml4
+++ b/parsing/g_vernac.ml4
@@ -9,20 +9,20 @@
(*i camlp4deps: "parsing/grammar.cma" i*)
(*i camlp4use: "pa_extend.cmo" i*)
-(* $Id: g_vernac.ml4 13197 2010-06-25 22:36:17Z letouzey $ *)
+(* $Id$ *)
open Pp
open Util
open Names
open Topconstr
+open Extend
open Vernacexpr
open Pcoq
open Decl_mode
open Tactic
open Decl_kinds
open Genarg
-open Extend
open Ppextend
open Goptions
@@ -50,6 +50,7 @@ let decl_notation = Gram.Entry.create "vernac:decl_notation"
let typeclass_context = Gram.Entry.create "vernac:typeclass_context"
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 instance_name = Gram.Entry.create "vernac:instance_name"
let get_command_entry () =
match Decl_mode.get_current_mode () with
@@ -58,34 +59,34 @@ let get_command_entry () =
| Mode_none -> noedit_mode
let default_command_entry =
- Gram.Entry.of_parser "command_entry"
+ Gram.Entry.of_parser "command_entry"
(fun strm -> Gram.Entry.parse_token (get_command_entry ()) strm)
let no_hook _ _ = ()
GEXTEND Gram
GLOBAL: vernac gallina_ext tactic_mode proof_mode noedit_mode;
vernac: FIRST
- [ [ IDENT "Time"; locality; v = vernac_aux ->
- check_locality (); VernacTime v
- | locality; v = vernac_aux ->
- check_locality (); v ] ]
+ [ [ IDENT "Time"; v = vernac -> VernacTime v
+ | IDENT "Timeout"; n = natural; v = vernac -> VernacTimeout(n,v)
+ | IDENT "Fail"; v = vernac -> VernacFail v
+ | locality; 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 *)
- [ [ g = gallina; "." -> g
+ [ [ g = gallina; "." -> g
| g = gallina_ext; "." -> g
- | c = command; "." -> c
+ | c = command; "." -> c
| c = syntax; "." -> c
| "["; l = LIST1 located_vernac; "]"; "." -> VernacList l
] ]
;
- vernac_aux: LAST
+ vernac_aux: LAST
[ [ prfcom = default_command_entry -> prfcom ] ]
;
locality:
- [ [ IDENT "Local" -> locality_flag := Some true
- | IDENT "Global" -> locality_flag := Some false
+ [ [ IDENT "Local" -> locality_flag := Some (loc,true)
+ | IDENT "Global" -> locality_flag := Some (loc,false)
| -> locality_flag := None ] ]
;
noedit_mode:
@@ -104,11 +105,11 @@ GEXTEND Gram
VernacSolve(g,tac,use_dft_tac)) ] ]
;
proof_mode:
- [ [ instr = proof_instr; "." -> VernacProofInstr instr ] ]
+ [ [ instr = proof_instr; "." -> VernacProofInstr instr ] ]
;
proof_mode: LAST
[ [ c=subgoal_command -> c (Some 1) ] ]
- ;
+ ;
located_vernac:
[ [ v = vernac -> loc, v ] ]
;
@@ -120,10 +121,11 @@ let test_plurial_form = function
"Keywords Variables/Hypotheses/Parameters expect more than one assumption"
| _ -> ()
-let no_coercion loc (c,x) =
- if c then Util.user_err_loc
- (loc,"no_coercion",str"No coercion allowed here.");
- x
+let test_plurial_form_types = function
+ | [([_],_)] ->
+ Flags.if_verbose warning
+ "Keywords Implicit Types expect more than one type"
+ | _ -> ()
(* Gallina declarations *)
GEXTEND Gram
@@ -133,27 +135,27 @@ GEXTEND Gram
gallina:
(* Definition, Theorem, Variable, Axiom, ... *)
[ [ thm = thm_token; id = identref; bl = binders_let; ":"; c = lconstr;
- l = LIST0
+ l = LIST0
[ "with"; id = identref; bl = binders_let; ":"; c = lconstr ->
- (Some id,(bl,c)) ] ->
- VernacStartTheoremProof (thm,(Some id,(bl,c))::l, false, no_hook)
- | stre = assumption_token; nl = inline; bl = assum_list ->
+ (Some id,(bl,c,None)) ] ->
+ VernacStartTheoremProof (thm,(Some id,(bl,c,None))::l, false, no_hook)
+ | stre = assumption_token; nl = inline; bl = assum_list ->
VernacAssumption (stre, nl, bl)
| stre = assumptions_token; nl = inline; bl = assum_list ->
test_plurial_form bl;
VernacAssumption (stre, nl, bl)
- | IDENT "Boxed";"Definition";id = identref; b = def_body ->
+ | IDENT "Boxed";"Definition";id = identref; b = def_body ->
VernacDefinition ((Global,true,Definition), id, b, no_hook)
- | IDENT "Unboxed";"Definition";id = identref; b = def_body ->
+ | IDENT "Unboxed";"Definition";id = identref; b = def_body ->
VernacDefinition ((Global,false,Definition), id, b, no_hook)
- | (f,d) = def_token; id = identref; b = def_body ->
+ | (f,d) = def_token; id = identref; b = def_body ->
VernacDefinition (d, id, b, f)
(* Gallina inductive declarations *)
| 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 (f,indl)
+ let (k,f) = f in
+ let indl=List.map (fun ((a,b,c,d),e) -> ((a,b,c,k,d),e)) indl in
+ VernacInductive (f,false,indl)
| IDENT "Boxed";"Fixpoint"; recs = LIST1 rec_definition SEP "with" ->
VernacFixpoint (recs,true)
| IDENT "Unboxed";"Fixpoint"; recs = LIST1 rec_definition SEP "with" ->
@@ -163,21 +165,21 @@ GEXTEND Gram
| "CoFixpoint"; corecs = LIST1 corec_definition SEP "with" ->
VernacCoFixpoint (corecs,false)
| IDENT "Scheme"; l = LIST1 scheme SEP "with" -> VernacScheme l
- | IDENT "Combined"; IDENT "Scheme"; id = identref; IDENT "from";
+ | IDENT "Combined"; IDENT "Scheme"; id = identref; IDENT "from";
l = LIST1 identref SEP "," -> VernacCombinedScheme (id, l) ] ]
;
gallina_ext:
- [ [ b = record_token; oc = opt_coercion; name = identref;
- ps = binders_let;
+ [ [ b = record_token; infer = infer_token; oc = opt_coercion; name = identref;
+ ps = binders_let;
s = OPT [ ":"; s = lconstr -> s ];
cfs = [ ":="; l = constructor_list_or_record_decl -> l
| -> RecordDecl (None, []) ] ->
let (recf,indf) = b in
- VernacInductive (indf,[((oc,name),ps,s,recf,cfs),None])
+ VernacInductive (indf,infer,[((oc,name),ps,s,recf,cfs),[]])
] ]
;
typeclass_context:
- [ [ "["; l=LIST1 typeclass_constraint SEP ","; "]" -> l
+ [ [ "["; l=LIST1 typeclass_constraint SEP ","; "]" -> l
| -> [] ] ]
;
thm_token:
@@ -190,14 +192,14 @@ GEXTEND Gram
| IDENT "Property" -> Property ] ]
;
def_token:
- [ [ "Definition" ->
+ [ [ "Definition" ->
no_hook, (Global, Flags.boxed_definitions(), Definition)
- | IDENT "Let" ->
+ | IDENT "Let" ->
no_hook, (Local, Flags.boxed_definitions(), Definition)
- | IDENT "Example" ->
+ | IDENT "Example" ->
no_hook, (Global, Flags.boxed_definitions(), Example)
| IDENT "SubClass" ->
- Class.add_subclass_hook, (use_locality_exp (), false, SubClass) ] ]
+ Class.add_subclass_hook, (use_locality_exp (), false, SubClass) ] ]
;
assumption_token:
[ [ "Hypothesis" -> (Local, Logical)
@@ -219,9 +221,12 @@ GEXTEND Gram
[ [ "Inductive" -> (Inductive_kw,Finite)
| "CoInductive" -> (CoInductive,CoFinite) ] ]
;
+ infer_token:
+ [ [ IDENT "Infer" -> true | -> false ] ]
+ ;
record_token:
[ [ IDENT "Record" -> (Record,BiFinite)
- | IDENT "Structure" -> (Structure,BiFinite)
+ | IDENT "Structure" -> (Structure,BiFinite)
| IDENT "Class" -> (Class true,BiFinite) ] ]
;
(* Simple definitions *)
@@ -239,25 +244,29 @@ GEXTEND Gram
[ [ IDENT "Eval"; r = Tactic.red_expr; "in" -> Some r
| -> None ] ]
;
+ one_decl_notation:
+ [ [ ntn = ne_lstring; ":="; c = constr;
+ scopt = OPT [ ":"; sc = IDENT -> sc] -> (ntn,c,scopt) ] ]
+ ;
decl_notation:
- [ [ OPT [ "where"; ntn = ne_string; ":="; c = constr;
- scopt = OPT [ ":"; sc = IDENT -> sc] -> (ntn,c,scopt) ] ] ]
- ;
+ [ [ "where"; l = LIST1 one_decl_notation SEP IDENT "and" -> l
+ | -> [] ] ]
+ ;
(* Inductives and records *)
inductive_definition:
- [ [ id = identref; oc = opt_coercion; indpar = binders_let;
+ [ [ id = identref; oc = opt_coercion; indpar = binders_let;
c = OPT [ ":"; c = lconstr -> c ];
":="; lc = constructor_list_or_record_decl; ntn = decl_notation ->
(((oc,id),indpar,c,lc),ntn) ] ]
;
constructor_list_or_record_decl:
[ [ "|"; l = LIST1 constructor SEP "|" -> Constructors l
- | id = identref ; c = constructor_type; "|"; l = LIST0 constructor SEP "|" ->
+ | id = identref ; c = constructor_type; "|"; l = LIST0 constructor SEP "|" ->
Constructors ((c id)::l)
| id = identref ; c = constructor_type -> Constructors [ c id ]
- | cstr = identref; "{"; fs = LIST0 record_field SEP ";"; "}" ->
- RecordDecl (Some cstr,fs)
- | "{";fs = LIST0 record_field SEP ";"; "}" -> RecordDecl (None,fs)
+ | cstr = identref; "{"; fs = LIST0 record_field SEP ";"; "}" ->
+ RecordDecl (Some cstr,fs)
+ | "{";fs = LIST0 record_field SEP ";"; "}" -> RecordDecl (None,fs)
| -> Constructors [] ] ]
;
(*
@@ -271,36 +280,19 @@ GEXTEND Gram
;
(* (co)-fixpoints *)
rec_definition:
- [ [ id = identref;
+ [ [ id = identref;
bl = binders_let_fixannot;
- ty = type_cstr;
- ":="; def = lconstr; ntn = decl_notation ->
- let bl, annot = bl in
- let names = names_of_local_assums bl in
- let ni =
- match fst annot with
- Some (loc, id) ->
- (if List.exists (fun (_, id') -> Name id = id') names then
- Some (loc, id)
- else Util.user_err_loc
- (loc,"Fixpoint",
- str "No argument named " ++ Nameops.pr_id id ++ str"."))
- | None ->
- (* If there is only one argument, it is the recursive one,
- otherwise, we search the recursive index later *)
- match names with
- | [(loc, Name na)] -> Some (loc, na)
- | _ -> None
- in
- ((id,(ni,snd annot),bl,ty,def),ntn) ] ]
+ 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 = identref; bl = binders_let; ty = type_cstr; ":=";
- def = lconstr; ntn = decl_notation ->
+ [ [ id = identref; bl = binders_let; ty = type_cstr;
+ def = OPT [":="; def = lconstr -> def]; ntn = decl_notation ->
((id,bl,ty,def),ntn) ] ]
;
type_cstr:
- [ [ ":"; c=lconstr -> c
+ [ [ ":"; c=lconstr -> c
| -> CHole (loc, None) ] ]
;
(* Inductive schemes *)
@@ -309,11 +301,11 @@ GEXTEND Gram
| id = identref; ":="; kind = scheme_kind -> (Some id,kind) ] ]
;
scheme_kind:
- [ [ IDENT "Induction"; "for"; ind = global;
+ [ [ IDENT "Induction"; "for"; ind = smart_global;
IDENT "Sort"; s = sort-> InductionScheme(true,ind,s)
- | IDENT "Minimality"; "for"; ind = global;
+ | IDENT "Minimality"; "for"; ind = smart_global;
IDENT "Sort"; s = sort-> InductionScheme(false,ind,s)
- | IDENT "Equality"; "for" ; ind = global -> EqualityScheme(ind) ] ]
+ | IDENT "Equality"; "for" ; ind = smart_global -> EqualityScheme(ind) ] ]
;
(* Various Binders *)
(*
@@ -331,16 +323,22 @@ GEXTEND Gram
record_field:
[ [ bd = record_binder; ntn = decl_notation -> bd,ntn ] ]
;
+ record_binder_body:
+ [ [ l = binders_let; oc = of_type_with_opt_coercion;
+ t = lconstr -> fun id -> (oc,AssumExpr (id,mkCProdN loc l t))
+ | l = binders_let; oc = of_type_with_opt_coercion;
+ t = lconstr; ":="; b = lconstr -> fun id ->
+ (oc,DefExpr (id,mkCLambdaN loc l b,Some (mkCProdN loc l t)))
+ | l = binders_let; ":="; b = lconstr -> fun id ->
+ match b with
+ | CCast(_,b, Rawterm.CastConv (_, t)) ->
+ (false,DefExpr(id,mkCLambdaN loc l b,Some (mkCProdN loc l t)))
+ | _ ->
+ (false,DefExpr(id,mkCLambdaN loc l b,None)) ] ]
+ ;
record_binder:
[ [ id = name -> (false,AssumExpr(id,CHole (loc, None)))
- | id = name; oc = of_type_with_opt_coercion; t = lconstr ->
- (oc,AssumExpr (id,t))
- | id = name; oc = of_type_with_opt_coercion;
- t = lconstr; ":="; b = lconstr -> (oc,DefExpr (id,b,Some t))
- | id = name; ":="; b = lconstr ->
- match b with
- CCast(_,b, Rawterm.CastConv (_, t)) -> (false,DefExpr(id,b,Some t))
- | _ -> (false,DefExpr(id,b,None)) ] ]
+ | id = name; f = record_binder_body -> f id ] ]
;
assum_list:
[ [ bl = LIST1 assum_coe -> bl | b = simple_assum_coe -> [b] ] ]
@@ -349,12 +347,12 @@ GEXTEND Gram
[ [ "("; a = simple_assum_coe; ")" -> a ] ]
;
simple_assum_coe:
- [ [ idl = LIST1 identref; oc = of_type_with_opt_coercion; c = lconstr ->
+ [ [ idl = LIST1 identref; oc = of_type_with_opt_coercion; c = lconstr ->
(oc,(idl,c)) ] ]
;
constructor_type:
- [[ l = binders_let;
+ [[ l = binders_let;
t= [ coe = of_type_with_opt_coercion; c = lconstr ->
fun l id -> (coe,(id,mkCProdN loc l c))
| ->
@@ -380,18 +378,17 @@ GEXTEND Gram
gallina_ext:
[ [ (* Interactive module declaration *)
- IDENT "Module"; export = export_token; id = identref;
- bl = LIST0 module_binder; mty_o = OPT of_module_type;
- mexpr_o = OPT is_module_expr ->
- VernacDefineModule (export, id, bl, mty_o, mexpr_o)
-
- | IDENT "Module"; "Type"; id = identref;
- bl = LIST0 module_binder; mty_o = OPT is_module_type ->
- VernacDeclareModuleType (id, bl, mty_o)
-
- | IDENT "Declare"; IDENT "Module"; export = export_token; id = identref;
- bl = LIST0 module_binder; ":"; mty = module_type ->
- VernacDeclareModule (export, id, bl, (mty,true))
+ IDENT "Module"; export = export_token; id = identref;
+ bl = LIST0 module_binder; sign = of_module_type;
+ body = is_module_expr ->
+ VernacDefineModule (export, id, bl, sign, body)
+ | IDENT "Module"; "Type"; id = identref;
+ bl = LIST0 module_binder; sign = check_module_types;
+ body = is_module_type ->
+ VernacDeclareModuleType (id, bl, sign, body)
+ | IDENT "Declare"; IDENT "Module"; export = export_token; id = identref;
+ bl = LIST0 module_binder; ":"; mty = module_type_inl ->
+ VernacDeclareModule (export, id, bl, mty)
(* Section beginning *)
| IDENT "Section"; id = identref -> VernacBeginSection id
| IDENT "Chapter"; id = identref -> VernacBeginSection id
@@ -402,43 +399,66 @@ GEXTEND Gram
(* Requiring an already compiled module *)
| IDENT "Require"; export = export_token; qidl = LIST1 global ->
VernacRequire (export, None, qidl)
- | IDENT "Require"; export = export_token; filename = ne_string ->
+ | IDENT "Require"; export = export_token; filename = ne_string ->
VernacRequireFrom (export, None, filename)
| IDENT "Import"; qidl = LIST1 global -> VernacImport (false,qidl)
- | IDENT "Export"; qidl = LIST1 global -> VernacImport (true,qidl)
- | IDENT "Include"; expr = module_expr -> VernacInclude(CIME(expr))
- | IDENT "Include"; "Type"; expr = module_type -> VernacInclude(CIMTE(expr)) ] ]
+ | IDENT "Export"; qidl = LIST1 global -> VernacImport (true,qidl)
+ | IDENT "Include"; e = module_expr_inl; l = LIST0 ext_module_expr ->
+ VernacInclude(e::l)
+ | IDENT "Include"; "Type"; e = module_type_inl; l = LIST0 ext_module_type ->
+ warning "Include Type is deprecated; use Include instead";
+ VernacInclude(e::l) ] ]
;
export_token:
[ [ IDENT "Import" -> Some false
| IDENT "Export" -> Some true
| -> None ] ]
;
+ ext_module_type:
+ [ [ "<+"; mty = module_type_inl -> mty ] ]
+ ;
+ ext_module_expr:
+ [ [ "<+"; mexpr = module_expr_inl -> mexpr ] ]
+ ;
+ check_module_type:
+ [ [ "<:"; mty = module_type_inl -> mty ] ]
+ ;
+ check_module_types:
+ [ [ mtys = LIST0 check_module_type -> mtys ] ]
+ ;
of_module_type:
- [ [ ":"; mty = module_type -> (mty, true)
- | "<:"; mty = module_type -> (mty, false) ] ]
+ [ [ ":"; mty = module_type_inl -> Enforce mty
+ | mtys = check_module_types -> Check mtys ] ]
;
is_module_type:
- [ [ ":="; mty = module_type -> mty ] ]
+ [ [ ":="; mty = module_type_inl ; l = LIST0 ext_module_type -> (mty::l)
+ | -> [] ] ]
;
is_module_expr:
- [ [ ":="; mexpr = module_expr -> mexpr ] ]
+ [ [ ":="; mexpr = module_expr_inl; l = LIST0 ext_module_expr -> (mexpr::l)
+ | -> [] ] ]
+ ;
+ module_expr_inl:
+ [ [ "!"; me = module_expr -> (me,false)
+ | me = module_expr -> (me,true) ] ]
+ ;
+ module_type_inl:
+ [ [ "!"; me = module_type -> (me,false)
+ | me = module_type -> (me,true) ] ]
;
-
(* Module binder *)
module_binder:
[ [ "("; export = export_token; idl = LIST1 identref; ":";
- mty = module_type; ")" -> (export,idl,mty) ] ]
+ mty = module_type_inl; ")" -> (export,idl,mty) ] ]
;
-
(* Module expressions *)
module_expr:
[ [ me = module_expr_atom -> me
- | me1 = module_expr; me2 = module_expr_atom -> CMEapply (me1,me2)
+ | me1 = module_expr; me2 = module_expr_atom -> CMapply (me1,me2)
] ]
;
module_expr_atom:
- [ [ qid = qualid -> CMEident qid | "("; me = module_expr; ")" -> me ] ]
+ [ [ qid = qualid -> CMident qid | "("; me = module_expr; ")" -> me ] ]
;
with_declaration:
[ [ "Definition"; fqid = fullyqualid; ":="; c = Constr.lconstr ->
@@ -447,94 +467,106 @@ GEXTEND Gram
CWith_Module (fqid,qid)
] ]
;
- module_type_atom:
- [ [ qid = qualid -> CMTEident qid
- | mty = module_type_atom; me = module_expr_atom -> CMTEapply (mty,me)
- ] ]
- ;
module_type:
- [ [ mty = module_type_atom -> mty
- | mty = module_type; "with"; decl = with_declaration -> CMTEwith (mty,decl)
+ [ [ qid = qualid -> CMident qid
+ | "("; mt = module_type; ")" -> mt
+ | mty = module_type; me = module_expr_atom -> CMapply (mty,me)
+ | mty = module_type; "with"; decl = with_declaration -> CMwith (mty,decl)
] ]
;
END
-(* Extensions: implicits, coercions, etc. *)
+(* Extensions: implicits, coercions, etc. *)
GEXTEND Gram
- GLOBAL: gallina_ext;
+ GLOBAL: gallina_ext instance_name;
gallina_ext:
[ [ (* Transparent and Opaque *)
- IDENT "Transparent"; l = LIST1 global ->
+ IDENT "Transparent"; l = LIST1 smart_global ->
VernacSetOpacity (use_non_locality (),[Conv_oracle.transparent,l])
- | IDENT "Opaque"; l = LIST1 global ->
+ | IDENT "Opaque"; l = LIST1 smart_global ->
VernacSetOpacity (use_non_locality (),[Conv_oracle.Opaque, l])
| IDENT "Strategy"; l =
- LIST1 [ lev=strategy_level; "["; q=LIST1 global; "]" -> (lev,q)] ->
+ LIST1 [ lev=strategy_level; "["; q=LIST1 smart_global; "]" -> (lev,q)] ->
VernacSetOpacity (use_locality (),l)
(* Canonical structure *)
| IDENT "Canonical"; IDENT "Structure"; qid = global ->
- VernacCanonical qid
- | IDENT "Canonical"; IDENT "Structure"; qid = global; d = def_body ->
- let s = coerce_global_to_id qid in
- VernacDefinition
+ VernacCanonical (AN qid)
+ | IDENT "Canonical"; IDENT "Structure"; ntn = by_notation ->
+ VernacCanonical (ByNotation ntn)
+ | IDENT "Canonical"; IDENT "Structure"; qid = global;
+ d = def_body ->
+ let s = coerce_reference_to_id qid in
+ VernacDefinition
((Global,false,CanonicalStructure),(dummy_loc,s),d,
(fun _ -> Recordops.declare_canonical_structure))
(* Coercions *)
| IDENT "Coercion"; qid = global; d = def_body ->
- let s = coerce_global_to_id qid in
+ let s = coerce_reference_to_id qid in
VernacDefinition ((use_locality_exp (),false,Coercion),(dummy_loc,s),d,Class.add_coercion_hook)
| IDENT "Coercion"; IDENT "Local"; qid = global; d = def_body ->
- let s = coerce_global_to_id qid in
- VernacDefinition ((enforce_locality_exp (),false,Coercion),(dummy_loc,s),d,Class.add_coercion_hook)
+ let s = coerce_reference_to_id qid in
+ VernacDefinition ((enforce_locality_exp true,false,Coercion),(dummy_loc,s),d,Class.add_coercion_hook)
| IDENT "Identity"; IDENT "Coercion"; IDENT "Local"; f = identref;
- ":"; s = class_rawexpr; ">->"; t = class_rawexpr ->
- VernacIdentityCoercion (enforce_locality_exp (), f, s, t)
+ ":"; s = class_rawexpr; ">->"; t = class_rawexpr ->
+ VernacIdentityCoercion (enforce_locality_exp true, f, s, t)
| IDENT "Identity"; IDENT "Coercion"; f = identref; ":";
- s = class_rawexpr; ">->"; t = class_rawexpr ->
+ s = class_rawexpr; ">->"; t = class_rawexpr ->
VernacIdentityCoercion (use_locality_exp (), f, s, t)
| IDENT "Coercion"; IDENT "Local"; qid = global; ":";
- s = class_rawexpr; ">->"; t = class_rawexpr ->
- VernacCoercion (enforce_locality_exp (), qid, s, t)
+ s = class_rawexpr; ">->"; t = class_rawexpr ->
+ VernacCoercion (enforce_locality_exp true, AN qid, s, t)
+ | IDENT "Coercion"; IDENT "Local"; ntn = by_notation; ":";
+ s = class_rawexpr; ">->"; t = class_rawexpr ->
+ VernacCoercion (enforce_locality_exp true, ByNotation ntn, s, t)
| IDENT "Coercion"; qid = global; ":"; s = class_rawexpr; ">->";
t = class_rawexpr ->
- VernacCoercion (use_locality_exp (), qid, s, t)
-
- | IDENT "Context"; c = binders_let ->
+ VernacCoercion (use_locality_exp (), AN qid, s, t)
+ | IDENT "Coercion"; ntn = by_notation; ":"; s = class_rawexpr; ">->";
+ t = class_rawexpr ->
+ VernacCoercion (use_locality_exp (), ByNotation ntn, s, t)
+
+ | IDENT "Context"; c = binders_let ->
VernacContext c
-
- | IDENT "Instance"; name = identref; sup = OPT binders_let; ":";
+
+ | IDENT "Instance"; namesup = instance_name; ":";
expl = [ "!" -> Rawterm.Implicit | -> Rawterm.Explicit ] ; t = operconstr LEVEL "200";
- pri = OPT [ "|"; i = natural -> i ] ;
- props = [ ":="; "{"; r = record_declaration; "}" -> r |
+ pri = OPT [ "|"; i = natural -> i ] ;
+ props = [ ":="; "{"; r = record_declaration; "}" -> r |
":="; c = lconstr -> c | -> CRecord (loc, None, []) ] ->
- let sup =
- match sup with
- None -> []
- | Some l -> l
- in
- let n =
- let (loc, id) = name in
- (loc, Name id)
- in
- VernacInstance (not (use_non_locality ()), sup, (n, expl, t), props, pri)
-
- | IDENT "Existing"; IDENT "Instance"; is = identref -> VernacDeclareInstance is
+ VernacInstance (false, not (use_non_locality ()),
+ snd namesup, (fst namesup, expl, t), props, pri)
+
+ | IDENT "Existing"; IDENT "Instance"; is = global ->
+ VernacDeclareInstance (not (use_section_locality ()), is)
+
+ | IDENT "Existing"; IDENT "Class"; is = global -> VernacDeclareClass is
(* Implicit *)
- | IDENT "Implicit"; IDENT "Arguments"; qid = global;
- pos = OPT [ "["; l = LIST0 implicit_name; "]" ->
+ | IDENT "Implicit"; IDENT "Arguments"; qid = smart_global;
+ pos = OPT [ "["; l = LIST0 implicit_name; "]" ->
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 []
+ | IDENT "No"; IDENT "Variables" -> None
+ | ["Variable" | IDENT "Variables"];
+ idl = LIST1 identref -> Some idl ] ->
+ VernacGeneralizable (use_non_locality (), gen) ] ]
;
implicit_name:
[ [ "!"; id = ident -> (id, false, true)
| id = ident -> (id,false,false)
- | "["; "!"; id = ident; "]" -> (id,true,true)
+ | "["; "!"; id = ident; "]" -> (id,true,true)
| "["; id = ident; "]" -> (id,true, false) ] ]
;
strategy_level:
@@ -544,6 +576,22 @@ GEXTEND Gram
| "-"; n=INT -> Conv_oracle.Level (- int_of_string n)
| IDENT "transparent" -> Conv_oracle.transparent ] ]
;
+ instance_name:
+ [ [ name = identref; sup = OPT binders_let ->
+ (let (loc,id) = name in (loc, Name id)),
+ (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
GEXTEND Gram
@@ -552,6 +600,14 @@ GEXTEND Gram
command:
[ [ IDENT "Comments"; l = LIST0 comment -> VernacComments l
+ (* Hack! Should be in grammar_ext, but camlp4 factorize badly *)
+ | IDENT "Declare"; IDENT "Instance"; namesup = instance_name; ":";
+ expl = [ "!" -> Rawterm.Implicit | -> Rawterm.Explicit ] ; t = operconstr LEVEL "200";
+ pri = OPT [ "|"; i = natural -> i ] ->
+ VernacInstance (true, not (use_non_locality ()),
+ snd namesup, (fst namesup, expl, t),
+ CRecord (loc, None, []), pri)
+
(* System directory *)
| IDENT "Pwd" -> VernacChdir None
| IDENT "Cd" -> VernacChdir None
@@ -559,14 +615,13 @@ GEXTEND Gram
(* Toplevel control *)
| IDENT "Drop" -> VernacToplevelControl Drop
- | IDENT "ProtectedLoop" -> VernacToplevelControl ProtectedLoop
| IDENT "Quit" -> VernacToplevelControl Quit
| IDENT "Load"; verbosely = [ IDENT "Verbose" -> true | -> false ];
s = [ s = ne_string -> s | s = IDENT -> s ] ->
VernacLoad (verbosely, s)
| IDENT "Declare"; IDENT "ML"; IDENT "Module"; l = LIST1 ne_string ->
- VernacDeclareMLModule l
+ VernacDeclareMLModule (use_locality (), l)
| IDENT "Dump"; IDENT "Universes"; fopt = OPT ne_string ->
error "This command is deprecated, use Print Universes"
@@ -576,7 +631,7 @@ GEXTEND Gram
(* Managing load paths *)
| IDENT "Add"; IDENT "LoadPath"; dir = ne_string; alias = as_dirpath ->
VernacAddLoadPath (false, dir, alias)
- | IDENT "Add"; IDENT "Rec"; IDENT "LoadPath"; dir = ne_string;
+ | IDENT "Add"; IDENT "Rec"; IDENT "LoadPath"; dir = ne_string;
alias = as_dirpath -> VernacAddLoadPath (true, dir, alias)
| IDENT "Remove"; IDENT "LoadPath"; dir = ne_string ->
VernacRemoveLoadPath dir
@@ -594,24 +649,24 @@ GEXTEND Gram
(* Printing (careful factorization of entries) *)
| IDENT "Print"; p = printable -> VernacPrint p
- | IDENT "Print"; qid = global -> VernacPrint (PrintName qid)
- | IDENT "Print"; IDENT "Module"; "Type"; qid = global ->
+ | IDENT "Print"; qid = smart_global -> VernacPrint (PrintName qid)
+ | IDENT "Print"; IDENT "Module"; "Type"; qid = global ->
VernacPrint (PrintModuleType qid)
- | IDENT "Print"; IDENT "Module"; qid = global ->
+ | IDENT "Print"; IDENT "Module"; qid = global ->
VernacPrint (PrintModule qid)
| IDENT "Inspect"; n = natural -> VernacPrint (PrintInspect n)
- | IDENT "About"; qid = global -> VernacPrint (PrintAbout qid)
+ | IDENT "About"; qid = smart_global -> VernacPrint (PrintAbout qid)
(* Searching the environment *)
- | IDENT "Search"; qid = global; l = in_or_out_modules ->
- VernacSearch (SearchHead qid, l)
+ | IDENT "Search"; c = constr_pattern; l = in_or_out_modules ->
+ VernacSearch (SearchHead c, l)
| IDENT "SearchPattern"; c = constr_pattern; l = in_or_out_modules ->
VernacSearch (SearchPattern c, l)
| IDENT "SearchRewrite"; c = constr_pattern; l = in_or_out_modules ->
VernacSearch (SearchRewrite c, l)
- | IDENT "SearchAbout";
+ | IDENT "SearchAbout";
sl = [ "[";
- l = LIST1 [
+ l = LIST1 [
b = positive_search_mark; s = ne_string; sc = OPT scope
-> b, SearchString (s,sc)
| b = positive_search_mark; p = constr_pattern
@@ -619,7 +674,7 @@ GEXTEND Gram
]; "]" -> l
| p = constr_pattern -> [true,SearchSubPattern p]
| s = ne_string; sc = OPT scope -> [true,SearchString (s,sc)] ];
- l = in_or_out_modules ->
+ l = in_or_out_modules ->
VernacSearch (SearchAbout sl, l)
| IDENT "Add"; IDENT "ML"; IDENT "Path"; dir = ne_string ->
@@ -629,23 +684,23 @@ GEXTEND Gram
(* Pour intervenir sur les tables de paramètres *)
| "Set"; table = option_table; v = option_value ->
- VernacSetOption (table,v)
+ VernacSetOption (use_locality_full(),table,v)
| "Set"; table = option_table ->
- VernacSetOption (table,BoolValue true)
+ VernacSetOption (use_locality_full(),table,BoolValue true)
| IDENT "Unset"; table = option_table ->
- VernacUnsetOption table
+ VernacUnsetOption (use_locality_full(),table)
| IDENT "Print"; IDENT "Table"; table = option_table ->
VernacPrintOption table
| IDENT "Add"; table = IDENT; field = IDENT; v = LIST1 option_ref_value
- -> VernacAddOption (SecondaryTable (table,field), v)
+ -> VernacAddOption ([table;field], v)
(* Un value global ci-dessous va être caché par un field au dessus! *)
(* En fait, on donne priorité aux tables secondaires *)
(* Pas de syntaxe pour les tables tertiaires pour cause de conflit *)
(* (mais de toutes façons, pas utilisées) *)
| IDENT "Add"; table = IDENT; v = LIST1 option_ref_value ->
- VernacAddOption (PrimaryTable table, v)
+ VernacAddOption ([table], v)
| IDENT "Test"; table = option_table; "for"; v = LIST1 option_ref_value
-> VernacMemOption (table, v)
@@ -653,9 +708,9 @@ GEXTEND Gram
VernacPrintOption table
| IDENT "Remove"; table = IDENT; field = IDENT; v= LIST1 option_ref_value
- -> VernacRemoveOption (SecondaryTable (table,field), v)
+ -> VernacRemoveOption ([table;field], v)
| IDENT "Remove"; table = IDENT; v = LIST1 option_ref_value ->
- VernacRemoveOption (PrimaryTable table, v)
+ VernacRemoveOption ([table], v)
| IDENT "proof" -> VernacDeclProof
| "return" -> VernacReturn ]]
@@ -669,14 +724,14 @@ GEXTEND Gram
fun g -> VernacCheckMayEval (None, g, c) ] ]
;
printable:
- [ [ IDENT "Term"; qid = global -> PrintName qid
+ [ [ IDENT "Term"; qid = smart_global -> PrintName qid
| IDENT "All" -> PrintFullContext
| IDENT "Section"; s = global -> PrintSectionContext s
| IDENT "Grammar"; ent = IDENT ->
(* This should be in "syntax" section but is here for factorization*)
PrintGrammar ent
| IDENT "LoadPath"; dir = OPT dirpath -> PrintLoadPath dir
- | IDENT "Modules" ->
+ | IDENT "Modules" ->
error "Print Modules is obsolete; use Print Libraries instead"
| IDENT "Libraries" -> PrintModules
@@ -685,40 +740,37 @@ GEXTEND Gram
| IDENT "Graph" -> PrintGraph
| IDENT "Classes" -> PrintClasses
| IDENT "TypeClasses" -> PrintTypeClasses
- | IDENT "Instances"; qid = global -> PrintInstances qid
+ | IDENT "Instances"; qid = smart_global -> PrintInstances qid
| IDENT "Ltac"; qid = global -> PrintLtac qid
| IDENT "Coercions" -> PrintCoercions
| IDENT "Coercion"; IDENT "Paths"; s = class_rawexpr; t = class_rawexpr
-> PrintCoercionPaths (s,t)
| IDENT "Canonical"; IDENT "Projections" -> PrintCanonicalConversions
| IDENT "Tables" -> PrintTables
-(* Obsolete: was used for cooking V6.3 recipes ??
- | IDENT "Proof"; qid = global -> PrintOpaqueName qid
-*)
| IDENT "Hint" -> PrintHintGoal
- | IDENT "Hint"; qid = global -> PrintHint qid
+ | IDENT "Hint"; qid = smart_global -> PrintHint qid
| IDENT "Hint"; "*" -> PrintHintDb
| IDENT "HintDb"; s = IDENT -> PrintHintDbName s
| "Rewrite"; IDENT "HintDb"; s = IDENT -> PrintRewriteHintDbName s
| IDENT "Scopes" -> PrintScopes
| IDENT "Scope"; s = IDENT -> PrintScope s
| IDENT "Visibility"; s = OPT IDENT -> PrintVisibility s
- | IDENT "Implicit"; qid = global -> PrintImplicit qid
+ | IDENT "Implicit"; qid = smart_global -> PrintImplicit qid
| IDENT "Universes"; fopt = OPT ne_string -> PrintUniverses fopt
- | IDENT "Assumptions"; qid = global -> PrintAssumptions (false, qid)
- | IDENT "Opaque"; IDENT "Dependencies"; qid = global -> PrintAssumptions (true, qid) ] ]
+ | IDENT "Assumptions"; qid = smart_global -> PrintAssumptions (false, qid)
+ | IDENT "Opaque"; IDENT "Dependencies"; qid = smart_global -> PrintAssumptions (true, qid) ] ]
;
class_rawexpr:
[ [ IDENT "Funclass" -> FunClass
| IDENT "Sortclass" -> SortClass
- | qid = global -> RefClass qid ] ]
+ | qid = smart_global -> RefClass qid ] ]
;
locatable:
- [ [ qid = global -> LocateTerm qid
+ [ [ qid = smart_global -> LocateTerm qid
| IDENT "File"; f = ne_string -> LocateFile f
| IDENT "Library"; qid = global -> LocateLibrary qid
| IDENT "Module"; qid = global -> LocateModule qid
- | s = ne_string -> LocateNotation s ] ]
+ | IDENT "Ltac"; qid = global -> LocateTactic qid ] ]
;
option_value:
[ [ n = integer -> IntValue n
@@ -729,9 +781,7 @@ GEXTEND Gram
| s = STRING -> StringRefValue s ] ]
;
option_table:
- [ [ f1 = IDENT; f2 = IDENT; f3 = IDENT -> TertiaryTable (f1,f2,f3)
- | f1 = IDENT; f2 = IDENT -> SecondaryTable (f1,f2)
- | f1 = IDENT -> PrimaryTable f1 ] ]
+ [ [ fl = LIST1 IDENT -> fl ]]
;
as_dirpath:
[ [ d = OPT [ "as"; d = dirpath -> d ] -> d ] ]
@@ -756,7 +806,7 @@ END;
GEXTEND Gram
command:
- [ [
+ [ [
(* State management *)
IDENT "Write"; IDENT "State"; s = IDENT -> VernacWriteState s
| IDENT "Write"; IDENT "State"; s = ne_string -> VernacWriteState s
@@ -770,15 +820,21 @@ GEXTEND Gram
| IDENT "Back" -> VernacBack 1
| IDENT "Back"; n = natural -> VernacBack n
| IDENT "BackTo"; n = natural -> VernacBackTo n
- | IDENT "Backtrack"; n = natural ; m = natural ; p = natural ->
+ | IDENT "Backtrack"; n = natural ; m = natural ; p = natural ->
VernacBacktrack (n,m,p)
(* Tactic Debugger *)
- | IDENT "Debug"; IDENT "On" ->
- VernacSetOption (SecondaryTable ("Ltac","Debug"), BoolValue true)
+ | IDENT "Debug"; IDENT "On" ->
+ VernacSetOption (None,["Ltac";"Debug"], BoolValue true)
| IDENT "Debug"; IDENT "Off" ->
- VernacSetOption (SecondaryTable ("Ltac","Debug"), BoolValue false)
+ VernacSetOption (None,["Ltac";"Debug"], BoolValue false)
+
+(* registration of a custom reduction *)
+
+ | IDENT "Declare"; IDENT "Reduction"; s = IDENT; ":=";
+ r = Tactic.red_expr ->
+ VernacDeclareReduction (use_locality(),s,r)
] ];
END
@@ -790,47 +846,54 @@ GEXTEND Gram
GLOBAL: syntax;
syntax:
- [ [ IDENT "Open"; local = obsolete_locality; IDENT "Scope"; sc = IDENT ->
- VernacOpenCloseScope (enforce_locality_of local,true,sc)
+ [ [ IDENT "Open"; local = obsolete_locality; IDENT "Scope"; sc = IDENT ->
+ VernacOpenCloseScope (enforce_section_locality local,true,sc)
- | IDENT "Close"; local = obsolete_locality; IDENT "Scope"; sc = IDENT ->
- VernacOpenCloseScope (enforce_locality_of local,false,sc)
+ | IDENT "Close"; local = obsolete_locality; IDENT "Scope"; sc = IDENT ->
+ VernacOpenCloseScope (enforce_section_locality local,false,sc)
| IDENT "Delimit"; IDENT "Scope"; sc = IDENT; "with"; key = IDENT ->
VernacDelimiters (sc,key)
- | IDENT "Bind"; IDENT "Scope"; sc = IDENT; "with";
+ | IDENT "Bind"; IDENT "Scope"; sc = IDENT; "with";
refl = LIST1 class_rawexpr -> VernacBindScope (sc,refl)
- | IDENT "Arguments"; IDENT "Scope"; qid = global;
- "["; scl = LIST0 opt_scope; "]" ->
- VernacArgumentsScope (use_non_locality (),qid,scl)
+ | IDENT "Arguments"; IDENT "Scope"; qid = smart_global;
+ "["; scl = LIST0 opt_scope; "]" ->
+ VernacArgumentsScope (use_section_locality (),qid,scl)
| IDENT "Infix"; local = obsolete_locality;
- op = ne_string; ":="; p = global;
+ op = ne_lstring; ":="; p = constr;
modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ];
sc = OPT [ ":"; sc = IDENT -> sc ] ->
- VernacInfix (enforce_locality_of local,(op,modl),p,sc)
- | IDENT "Notation"; local = obsolete_locality; id = identref;
+ VernacInfix (enforce_module_locality local,(op,modl),p,sc)
+ | IDENT "Notation"; local = obsolete_locality; id = identref;
idl = LIST0 ident; ":="; c = constr;
b = [ "("; IDENT "only"; IDENT "parsing"; ")" -> true | -> false ] ->
- VernacSyntacticDefinition (id,(idl,c),enforce_locality_of local,b)
- | IDENT "Notation"; local = obsolete_locality; s = ne_string; ":=";
+ VernacSyntacticDefinition
+ (id,(idl,c),enforce_module_locality local,b)
+ | IDENT "Notation"; local = obsolete_locality; s = ne_lstring; ":=";
c = constr;
modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ];
sc = OPT [ ":"; sc = IDENT -> sc ] ->
- VernacNotation (enforce_locality_of local,c,(s,modl),sc)
+ VernacNotation (enforce_module_locality local,c,(s,modl),sc)
- | IDENT "Tactic"; IDENT "Notation"; n = tactic_level;
+ | IDENT "Tactic"; IDENT "Notation"; n = tactic_level;
pil = LIST1 production_item; ":="; t = Tactic.tactic
-> VernacTacticNotation (n,pil,t)
- | IDENT "Reserved"; IDENT "Notation"; local = obsolete_locality;
- s = ne_string;
+ | 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 (use_module_locality(),((loc,"x '"^s^"' y"),l))
+
+ | IDENT "Reserved"; IDENT "Notation"; local = obsolete_locality;
+ s = ne_lstring;
l = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ]
- -> VernacSyntaxExtension (enforce_locality_of local,(s,l))
+ -> VernacSyntaxExtension (enforce_module_locality local,(s,l))
- (* "Print" "Grammar" should be here but is in "command" entry in order
+ (* "Print" "Grammar" should be here but is in "command" entry in order
to factorize with other "Print"-based vernac entries *)
] ]
;
@@ -846,7 +909,7 @@ GEXTEND Gram
;
syntax_modifier:
[ [ x = IDENT; "at"; lev = level -> SetItemLevel ([x],lev)
- | x = IDENT; ","; l = LIST1 IDENT SEP ","; "at";
+ | x = IDENT; ","; l = LIST1 IDENT SEP ","; "at";
lev = level -> SetItemLevel (x::l,lev)
| "at"; IDENT "level"; n = natural -> SetLevel n
| IDENT "left"; IDENT "associativity" -> SetAssoc Gramext.LeftA
@@ -857,7 +920,7 @@ GEXTEND Gram
| IDENT "format"; s = [s = STRING -> (loc,s)] -> SetFormat s ] ]
;
syntax_extension_type:
- [ [ IDENT "ident" -> ETIdent | IDENT "global" -> ETReference
+ [ [ IDENT "ident" -> ETName | IDENT "global" -> ETReference
| IDENT "bigint" -> ETBigint
] ]
;
@@ -865,8 +928,9 @@ GEXTEND Gram
[ [ "_" -> None | sc = IDENT -> Some sc ] ]
;
production_item:
- [ [ s = ne_string -> VTerm s
- | nt = IDENT; po = OPT [ "("; p = ident; ")" -> p ] ->
- VNonTerm (loc,nt,po) ] ]
+ [ [ s = ne_string -> TacTerm s
+ | nt = IDENT;
+ po = OPT [ "("; p = ident; sep = [ -> "" | ","; sep = STRING -> sep ];
+ ")" -> (p,sep) ] -> TacNonTerm (loc,nt,po) ] ]
;
END
diff --git a/parsing/g_xml.ml4 b/parsing/g_xml.ml4
index 72d5d275..0f702904 100644
--- a/parsing/g_xml.ml4
+++ b/parsing/g_xml.ml4
@@ -8,7 +8,7 @@
(*i camlp4use: "pa_extend.cmo" i*)
-(* $Id: g_xml.ml4 11309 2008-08-06 10:30:35Z herbelin $ *)
+(* $Id$ *)
open Pp
open Util
@@ -30,7 +30,7 @@ type xml = XmlTag of loc * string * attribute list * xml list
let check_tags loc otag ctag =
if otag <> ctag then
- user_err_loc (loc,"",str "closing xml tag " ++ str ctag ++
+ user_err_loc (loc,"",str "closing xml tag " ++ str ctag ++
str "does not match open xml tag " ++ str otag ++ str ".")
let xml_eoi = (Gram.Entry.create "xml_eoi" : xml Gram.Entry.e)
@@ -57,6 +57,9 @@ END
(* Errors *)
+let error_expect_two_arguments loc =
+ user_err_loc (loc,"",str "wrong number of arguments (expect two).")
+
let error_expect_one_argument loc =
user_err_loc (loc,"",str "wrong number of arguments (expect one).")
@@ -68,12 +71,8 @@ let error_expect_no_argument loc =
let nmtoken (loc,a) =
try int_of_string a
with Failure _ -> user_err_loc (loc,"",str "nmtoken expected.")
-
-let interp_xml_attr_qualid = function
- | "uri", s -> qualid_of_string s
- | _ -> error "Ill-formed xml attribute"
-let get_xml_attr s al =
+let get_xml_attr s al =
try List.assoc s al
with Not_found -> error ("No attribute "^s)
@@ -144,7 +143,7 @@ let compute_inductive_nargs ind =
let rec interp_xml_constr = function
| XmlTag (loc,"REL",al,[]) ->
RVar (loc, get_xml_ident al)
- | XmlTag (loc,"VAR",al,[]) ->
+ | XmlTag (loc,"VAR",al,[]) ->
error "XML parser: unable to interp free variables"
| XmlTag (loc,"LAMBDA",al,(_::_ as xl)) ->
let body,decls = list_sep_last xl in
@@ -201,7 +200,7 @@ let rec interp_xml_constr = function
and interp_xml_tag s = function
| XmlTag (loc,tag,al,xl) when tag=s -> (loc,al,xl)
- | XmlTag (loc,tag,_,_) -> user_err_loc (loc, "",
+ | XmlTag (loc,tag,_,_) -> user_err_loc (loc, "",
str "Expect tag " ++ str s ++ str " but find " ++ str s ++ str ".")
and interp_xml_constr_alias s x =
@@ -232,17 +231,17 @@ and interp_xml_recursionOrder x =
let (loc, al, l) = interp_xml_tag "RecursionOrder" x in
let (locs, s) = get_xml_attr "type" al in
match s with
- "Structural" ->
+ "Structural" ->
(match l with [] -> RStructRec
| _ -> error_expect_no_argument loc)
- | "WellFounded" ->
+ | "WellFounded" ->
(match l with
[c] -> RWfRec (interp_xml_type c)
| _ -> error_expect_one_argument loc)
- | "Measure" ->
+ | "Measure" ->
(match l with
- [c] -> RMeasureRec (interp_xml_type c)
- | _ -> error_expect_one_argument loc)
+ [m;r] -> RMeasureRec (interp_xml_type m, Some (interp_xml_type r))
+ | _ -> error_expect_two_arguments loc)
| _ ->
user_err_loc (locs,"",str "Invalid recursion order.")
@@ -262,7 +261,7 @@ and interp_xml_CoFixFunction x =
match interp_xml_tag "CoFixFunction" x with
| (loc,al,[x1;x2]) ->
(get_xml_name al, interp_xml_type x1, interp_xml_body x2)
- | (loc,_,_) ->
+ | (loc,_,_) ->
error_expect_one_argument loc
(* Interpreting tactic argument *)
diff --git a/parsing/g_zsyntax.ml b/parsing/g_zsyntax.ml
deleted file mode 100644
index 8b9c0d22..00000000
--- a/parsing/g_zsyntax.ml
+++ /dev/null
@@ -1,194 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(* $Id: g_zsyntax.ml 11309 2008-08-06 10:30:35Z herbelin $ *)
-
-open Pcoq
-open Pp
-open Util
-open Names
-open Topconstr
-open Libnames
-open Bigint
-
-exception Non_closed_number
-
-(**********************************************************************)
-(* Parsing positive via scopes *)
-(**********************************************************************)
-
-open Libnames
-open Rawterm
-let make_dir l = make_dirpath (List.map id_of_string (List.rev l))
-let positive_module = ["Coq";"NArith";"BinPos"]
-let make_path dir id = Libnames.make_path (make_dir dir) (id_of_string id)
-
-let positive_path = make_path positive_module "positive"
-
-(* TODO: temporary hack *)
-let make_kn dir id = Libnames.encode_kn dir id
-
-let positive_kn =
- make_kn (make_dir positive_module) (id_of_string "positive")
-let glob_positive = IndRef (positive_kn,0)
-let path_of_xI = ((positive_kn,0),1)
-let path_of_xO = ((positive_kn,0),2)
-let path_of_xH = ((positive_kn,0),3)
-let glob_xI = ConstructRef path_of_xI
-let glob_xO = ConstructRef path_of_xO
-let glob_xH = ConstructRef path_of_xH
-
-let pos_of_bignat dloc x =
- let ref_xI = RRef (dloc, glob_xI) in
- let ref_xH = RRef (dloc, glob_xH) in
- let ref_xO = RRef (dloc, glob_xO) in
- let rec pos_of x =
- match div2_with_rest x with
- | (q,false) -> RApp (dloc, ref_xO,[pos_of q])
- | (q,true) when q <> zero -> RApp (dloc,ref_xI,[pos_of q])
- | (q,true) -> ref_xH
- in
- pos_of x
-
-let error_non_positive dloc =
- user_err_loc (dloc, "interp_positive",
- str "Only strictly positive numbers in type \"positive\".")
-
-let interp_positive dloc n =
- if is_strictly_pos n then pos_of_bignat dloc n
- else error_non_positive dloc
-
-(**********************************************************************)
-(* Printing positive via scopes *)
-(**********************************************************************)
-
-let rec bignat_of_pos = function
- | RApp (_, RRef (_,b),[a]) when b = glob_xO -> mult_2(bignat_of_pos a)
- | RApp (_, RRef (_,b),[a]) when b = glob_xI -> add_1(mult_2(bignat_of_pos a))
- | RRef (_, a) when a = glob_xH -> Bigint.one
- | _ -> raise Non_closed_number
-
-let uninterp_positive p =
- try
- Some (bignat_of_pos p)
- with Non_closed_number ->
- None
-
-(************************************************************************)
-(* Declaring interpreters and uninterpreters for positive *)
-(************************************************************************)
-
-let _ = Notation.declare_numeral_interpreter "positive_scope"
- (positive_path,positive_module)
- interp_positive
- ([RRef (dummy_loc, glob_xI);
- RRef (dummy_loc, glob_xO);
- RRef (dummy_loc, glob_xH)],
- uninterp_positive,
- true)
-
-(**********************************************************************)
-(* Parsing N via scopes *)
-(**********************************************************************)
-
-let binnat_module = ["Coq";"NArith";"BinNat"]
-let n_kn = make_kn (make_dir binnat_module) (id_of_string "N")
-let glob_n = IndRef (n_kn,0)
-let path_of_N0 = ((n_kn,0),1)
-let path_of_Npos = ((n_kn,0),2)
-let glob_N0 = ConstructRef path_of_N0
-let glob_Npos = ConstructRef path_of_Npos
-
-let n_path = make_path binnat_module "N"
-
-let n_of_binnat dloc pos_or_neg n =
- if n <> zero then
- RApp(dloc, RRef (dloc,glob_Npos), [pos_of_bignat dloc n])
- else
- RRef (dloc, glob_N0)
-
-let error_negative dloc =
- user_err_loc (dloc, "interp_N", str "No negative numbers in type \"N\".")
-
-let n_of_int dloc n =
- if is_pos_or_zero n then n_of_binnat dloc true n
- else error_negative dloc
-
-(**********************************************************************)
-(* Printing N via scopes *)
-(**********************************************************************)
-
-let bignat_of_n = function
- | RApp (_, RRef (_,b),[a]) when b = glob_Npos -> bignat_of_pos a
- | RRef (_, a) when a = glob_N0 -> Bigint.zero
- | _ -> raise Non_closed_number
-
-let uninterp_n p =
- try Some (bignat_of_n p)
- with Non_closed_number -> None
-
-(************************************************************************)
-(* Declaring interpreters and uninterpreters for N *)
-
-let _ = Notation.declare_numeral_interpreter "N_scope"
- (n_path,binnat_module)
- n_of_int
- ([RRef (dummy_loc, glob_N0);
- RRef (dummy_loc, glob_Npos)],
- uninterp_n,
- true)
-
-(**********************************************************************)
-(* Parsing Z via scopes *)
-(**********************************************************************)
-
-let binint_module = ["Coq";"ZArith";"BinInt"]
-let z_path = make_path binint_module "Z"
-let z_kn = make_kn (make_dir binint_module) (id_of_string "Z")
-let glob_z = IndRef (z_kn,0)
-let path_of_ZERO = ((z_kn,0),1)
-let path_of_POS = ((z_kn,0),2)
-let path_of_NEG = ((z_kn,0),3)
-let glob_ZERO = ConstructRef path_of_ZERO
-let glob_POS = ConstructRef path_of_POS
-let glob_NEG = ConstructRef path_of_NEG
-
-let z_of_int dloc n =
- if n <> zero then
- let sgn, n =
- if is_pos_or_zero n then glob_POS, n else glob_NEG, Bigint.neg n in
- RApp(dloc, RRef (dloc,sgn), [pos_of_bignat dloc n])
- else
- RRef (dloc, glob_ZERO)
-
-(**********************************************************************)
-(* Printing Z via scopes *)
-(**********************************************************************)
-
-let bigint_of_z = function
- | RApp (_, RRef (_,b),[a]) when b = glob_POS -> bignat_of_pos a
- | RApp (_, RRef (_,b),[a]) when b = glob_NEG -> Bigint.neg (bignat_of_pos a)
- | RRef (_, a) when a = glob_ZERO -> Bigint.zero
- | _ -> raise Non_closed_number
-
-let uninterp_z p =
- try
- Some (bigint_of_z p)
- with Non_closed_number -> None
-
-(************************************************************************)
-(* Declaring interpreters and uninterpreters for Z *)
-
-let _ = Notation.declare_numeral_interpreter "Z_scope"
- (z_path,binint_module)
- z_of_int
- ([RRef (dummy_loc, glob_ZERO);
- RRef (dummy_loc, glob_POS);
- RRef (dummy_loc, glob_NEG)],
- uninterp_z,
- true)
diff --git a/parsing/g_zsyntax.mli b/parsing/g_zsyntax.mli
index 11e0b6ac..b7e6e994 100644
--- a/parsing/g_zsyntax.mli
+++ b/parsing/g_zsyntax.mli
@@ -6,6 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: g_zsyntax.mli 5920 2004-07-16 20:01:26Z herbelin $ i*)
+(*i $Id$ i*)
(* Nice syntax for integers. *)
diff --git a/parsing/grammar.mllib b/parsing/grammar.mllib
new file mode 100644
index 00000000..248a8ad9
--- /dev/null
+++ b/parsing/grammar.mllib
@@ -0,0 +1,84 @@
+Coq_config
+
+Profile
+Pp_control
+Pp
+Compat
+Flags
+Segmenttree
+Unicodetable
+Util
+Bigint
+Dyn
+Hashcons
+Predicate
+Rtree
+Option
+
+Names
+Univ
+Esubst
+Term
+Mod_subst
+Sign
+Cbytecodes
+Copcodes
+Cemitcodes
+Declarations
+Retroknowledge
+Pre_env
+Cbytegen
+Environ
+Conv_oracle
+Closure
+Reduction
+Type_errors
+Entries
+Modops
+Inductive
+Typeops
+Indtypes
+Cooking
+Term_typing
+Subtyping
+Mod_typing
+Safe_typing
+
+Nameops
+Libnames
+Summary
+Nametab
+Libobject
+Lib
+Goptions
+Decl_kinds
+Global
+Termops
+Namegen
+Evd
+Reductionops
+Inductiveops
+Rawterm
+Detyping
+Pattern
+Topconstr
+Genarg
+Ppextend
+Tacexpr
+Lexer
+Extend
+Vernacexpr
+Extrawit
+Pcoq
+Q_util
+Q_coqast
+
+Egrammar
+Argextend
+Tacextend
+Vernacextend
+
+G_prim
+G_tactic
+G_ltac
+G_constr
diff --git a/parsing/highparsing.mllib b/parsing/highparsing.mllib
new file mode 100644
index 00000000..3eb27abb
--- /dev/null
+++ b/parsing/highparsing.mllib
@@ -0,0 +1,7 @@
+G_constr
+G_vernac
+G_prim
+G_proofs
+G_tactic
+G_ltac
+G_decl_mode
diff --git a/parsing/lexer.ml4 b/parsing/lexer.ml4
index 52b5ede7..4ec61a23 100644
--- a/parsing/lexer.ml4
+++ b/parsing/lexer.ml4
@@ -6,13 +6,12 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: lexer.ml4 12891 2010-03-30 11:40:02Z herbelin $ i*)
-
-
-(*i camlp4use: "pr_o.cmo" i*)
+(*i camlp4use: "pr_o.cmo pa_macro.cmo" i*)
(* Add pr_o.cmo to circumvent a useless-warning bug when preprocessed with
* ast-based camlp4 *)
+(*i $Id$ i*)
+
open Pp
open Util
open Token
@@ -22,7 +21,7 @@ open Token
module CharMap = Map.Make (struct type t = char let compare = compare end)
-type ttree = {
+type ttree = {
node : string option;
branch : ttree CharMap.t }
@@ -30,7 +29,7 @@ let empty_ttree = { node = None; branch = CharMap.empty }
let ttree_add ttree str =
let rec insert tt i =
- if i == String.length str then
+ if i == String.length str then
{node = Some str; branch = tt.branch}
else
let c = str.[i] in
@@ -43,7 +42,7 @@ let ttree_add ttree str =
CharMap.add c (insert tt' (i + 1)) tt.branch
in
{ node = tt.node; branch = br }
- in
+ in
insert ttree 0
(* Search a string in a dictionary: raises [Not_found]
@@ -51,15 +50,30 @@ let ttree_add ttree str =
let ttree_find ttree str =
let rec proc_rec tt i =
- if i == String.length str then
- match tt.node with
- | Some s -> s
- | None -> raise Not_found
- else
- proc_rec (CharMap.find str.[i] tt.branch) (i+1)
- in
+ if i == String.length str then tt
+ else proc_rec (CharMap.find str.[i] tt.branch) (i+1)
+ in
proc_rec ttree 0
+(* Removes a string from a dictionary: returns an equal dictionary
+ if the word not present. *)
+let ttree_remove ttree str =
+ let rec remove tt i =
+ if i == String.length str then
+ {node = None; branch = tt.branch}
+ else
+ let c = str.[i] in
+ let br =
+ match try Some (CharMap.find c tt.branch) with Not_found -> None with
+ | Some tt' ->
+ CharMap.add c (remove tt' (i + 1)) (CharMap.remove c tt.branch)
+ | None -> tt.branch
+ in
+ { node = tt.node; branch = br }
+ in
+ remove ttree 0
+
+
(* Errors occuring while lexing (explained as "Lexer error: ...") *)
type error =
@@ -108,7 +122,7 @@ let check_utf8_trailing_byte cs c =
(* but don't certify full utf8 compliance (e.g. no emptyness check) *)
let lookup_utf8_tail c cs =
let c1 = Char.code c in
- if c1 land 0x40 = 0 or c1 land 0x38 = 0x38 then error_utf8 cs
+ if c1 land 0x40 = 0 or c1 land 0x38 = 0x38 then error_utf8 cs
else
let n, unicode =
if c1 land 0x20 = 0 then
@@ -121,20 +135,20 @@ let lookup_utf8_tail c cs =
match Stream.npeek 3 cs with
| [_;c2;c3] ->
check_utf8_trailing_byte cs c2; check_utf8_trailing_byte cs c3;
- 3, (c1 land 0x0F) lsl 12 + (Char.code c2 land 0x3F) lsl 6 +
+ 3, (c1 land 0x0F) lsl 12 + (Char.code c2 land 0x3F) lsl 6 +
(Char.code c3 land 0x3F)
| _ -> error_utf8 cs
else match Stream.npeek 4 cs with
| [_;c2;c3;c4] ->
check_utf8_trailing_byte cs c2; check_utf8_trailing_byte cs c3;
- check_utf8_trailing_byte cs c4;
+ check_utf8_trailing_byte cs c4;
4, (c1 land 0x07) lsl 18 + (Char.code c2 land 0x3F) lsl 12 +
(Char.code c3 land 0x3F) lsl 6 + (Char.code c4 land 0x3F)
| _ -> error_utf8 cs
in
try classify_unicode unicode, n
with UnsupportedUtf8 -> error_unsupported_unicode_character n cs
-
+
let lookup_utf8 cs =
match Stream.peek cs with
| Some ('\x00'..'\x7F') -> AsciiChar
@@ -171,15 +185,17 @@ let check_keyword str =
(* Keyword and symbol dictionary *)
let token_tree = ref empty_ttree
-let find_keyword s = ttree_find !token_tree s
-
-let is_keyword s =
- try let _ = ttree_find !token_tree s in true with Not_found -> false
+let is_keyword s =
+ try match (ttree_find !token_tree s).node with None -> false | Some _ -> true
+ with Not_found -> false
let add_keyword str =
check_keyword str;
token_tree := ttree_add !token_tree str
+let remove_keyword str =
+ token_tree := ttree_remove !token_tree str
+
(* Adding a new token (keyword or special token). *)
let add_token (con, str) = match con with
| "" -> add_keyword str
@@ -235,10 +251,18 @@ let rec number len = parser
let escape len c = store len c
-let rec string bp len = parser
+let rec string in_comments bp len = parser
| [< ''"'; esc=(parser [<''"' >] -> true | [< >] -> false); s >] ->
- if esc then string bp (store len '"') s else len
- | [< 'c; s >] -> string bp (store len c) s
+ if esc then string in_comments bp (store len '"') s else len
+ | [< ''*'; s >] ->
+ (parser
+ | [< '')'; s >] ->
+ if in_comments then
+ msg_warning (str "Not interpreting \"*)\" as the end of current non-terminated comment because it occurs in a non-terminated string of the comment.");
+ string in_comments bp (store (store len '*') ')') s
+ | [< >] ->
+ string in_comments bp (store len '*') s) s
+ | [< 'c; s >] -> string in_comments bp (store len c) s
| [< _ = Stream.empty >] ep -> err (bp, ep) Unterminated_string
(* Hook for exporting comment into xml theory files *)
@@ -254,8 +278,8 @@ let between_com = ref true
type com_state = int option * string * bool
let restore_com_state (o,s,b) =
- comment_begin := o;
- Buffer.clear current; Buffer.add_string current s;
+ comment_begin := o;
+ Buffer.clear current; Buffer.add_string current s;
between_com := b
let dflt_com = (None,"",true)
let com_state () =
@@ -310,13 +334,13 @@ let rec comm_string bp = parser
| [< >] -> real_push_char '\\'); s >]
-> comm_string bp s
| [< _ = Stream.empty >] ep -> err (bp, ep) Unterminated_string
- | [< 'c; s >] -> real_push_char c; comm_string bp s
+ | [< 'c; s >] -> real_push_char c; comm_string bp s
let rec comment bp = parser bp2
| [< ''(';
_ = (parser
| [< ''*'; s >] -> push_string "(*"; comment bp s
- | [< >] -> push_string "(" );
+ | [< >] -> push_string "(" );
s >] -> comment bp s
| [< ''*';
_ = parser
@@ -324,7 +348,7 @@ let rec comment bp = parser bp2
| [< s >] -> real_push_char '*'; comment bp s >] -> ()
| [< ''"'; s >] ->
if Flags.do_beautify() then (push_string"\"";comm_string bp2 s)
- else ignore (string bp2 0 s);
+ else ignore (string true bp2 0 s);
comment bp s
| [< _ = Stream.empty >] ep -> err (bp, ep) Unterminated_comment
| [< 'z; s >] -> real_push_char z; comment bp s
@@ -340,12 +364,12 @@ let rec progress_further last nj tt cs =
and update_longest_valid_token last nj tt cs =
match tt.node with
| Some _ as last' ->
- for i=1 to nj do Stream.junk cs done;
+ for i=1 to nj do Stream.junk cs done;
progress_further last' 0 tt cs
| None ->
progress_further last nj tt cs
-(* nr is the number of char peeked since last valid token *)
+(* nj is the number of char peeked since last valid token *)
(* n the number of char in utf8 block *)
and progress_utf8 last nj n c tt cs =
try
@@ -358,7 +382,7 @@ and progress_utf8 last nj n c tt cs =
List.iter (check_utf8_trailing_byte cs) l;
let tt = List.fold_left (fun tt c -> CharMap.find c tt.branch) tt l in
update_longest_valid_token last (nj+n) tt cs
- | _ ->
+ | _ ->
error_utf8 cs
with Not_found ->
last
@@ -366,6 +390,12 @@ and progress_utf8 last nj n c tt cs =
and progress_from_byte last nj tt cs c =
progress_utf8 last nj (utf8_char_size cs c) c tt cs
+let find_keyword id s =
+ let tt = ttree_find !token_tree id in
+ match progress_further tt.node 0 tt s with
+ | None -> raise Not_found
+ | Some c -> c
+
(* Must be a special token *)
let process_chars bp c cs =
let t = progress_from_byte None (-1) !token_tree cs c in
@@ -379,7 +409,7 @@ let process_chars bp c cs =
let parse_after_dollar bp =
parser
- | [< ' ('a'..'z' | 'A'..'Z' | '_' as c); len = ident_tail (store 0 c) >] ->
+ | [< ' ('a'..'z' | 'A'..'Z' | '_' as c); len = ident_tail (store 0 c) >] ->
("METAIDENT", get_buff len)
| [< s >] ->
match lookup_utf8 s with
@@ -394,9 +424,9 @@ let parse_after_dot bp c =
("FIELD", get_buff len)
| [< s >] ->
match lookup_utf8 s with
- | Utf8Token (UnicodeLetter, n) ->
+ | Utf8Token (UnicodeLetter, n) ->
("FIELD", get_buff (ident_tail (nstore n 0 s) s))
- | AsciiChar | Utf8Token _ | EmptyStream ->
+ | AsciiChar | Utf8Token _ | EmptyStream ->
fst (process_chars bp c s)
(* Parse what follows a question mark *)
@@ -409,6 +439,7 @@ let parse_after_qmark bp s =
| Utf8Token (UnicodeLetter, _) -> ("LEFTQMARK", "")
| AsciiChar | Utf8Token _ | EmptyStream -> fst (process_chars bp '?' s)
+
(* Parse a token in a char stream *)
let rec next_token = parser bp
| [< '' ' | '\t' | '\n' |'\r' as c; s >] ->
@@ -422,14 +453,14 @@ let rec next_token = parser bp
| [< ''?'; s >] ep ->
let t = parse_after_qmark bp s in comment_stop bp; (t, (ep, bp))
| [< ' ('a'..'z' | 'A'..'Z' | '_' as c);
- len = ident_tail (store 0 c) >] ep ->
- let id = get_buff len in
+ len = ident_tail (store 0 c); s >] ep ->
+ let id = get_buff len in
comment_stop bp;
- (try ("", find_keyword id) with Not_found -> ("IDENT", id)), (bp, ep)
+ (try ("", find_keyword id s) with Not_found -> ("IDENT", id)), (bp, ep)
| [< ' ('0'..'9' as c); len = number (store 0 c) >] ep ->
comment_stop bp;
(("INT", get_buff len), (bp, ep))
- | [< ''\"'; len = string bp 0 >] ep ->
+ | [< ''\"'; len = string false bp 0 >] ep ->
comment_stop bp;
(("STRING", get_buff len), (bp, ep))
| [< ' ('(' as c);
@@ -448,8 +479,8 @@ let rec next_token = parser bp
let id = get_buff len in
let ep = Stream.count s in
comment_stop bp;
- (try ("",find_keyword id) with Not_found -> ("IDENT",id)), (bp, ep)
- | AsciiChar | Utf8Token ((UnicodeSymbol | UnicodeIdentPart), _) ->
+ (try ("",find_keyword id s) with Not_found -> ("IDENT",id)), (bp, ep)
+ | AsciiChar | Utf8Token ((UnicodeSymbol | UnicodeIdentPart), _) ->
let t = process_chars bp (Stream.next s) s in
comment_stop bp; t
| EmptyStream ->
@@ -514,16 +545,38 @@ let token_text = function
| ("STRING", "") -> "string"
| ("EOI", "") -> "end of input"
| (con, "") -> con
- | (con, prm) -> con ^ " \"" ^ prm ^ "\""
+ | (con, prm) -> con ^ " \"" ^ prm ^ "\""
-let tparse (p_con, p_prm) =
- None
- (*i was
- if p_prm = "" then
- (parser [< '(con, prm) when con = p_con >] -> prm)
- else
- (parser [< '(con, prm) when con = p_con && prm = p_prm >] -> prm)
- i*)
+(* The lexer of Coq *)
+
+(* Note: removing a token.
+ We do nothing because [remove_token] is called only when removing a grammar
+ rule with [Grammar.delete_rule]. The latter command is called only when
+ unfreezing the state of the grammar entries (see GRAMMAR summary, file
+ env/metasyntax.ml). Therefore, instead of removing tokens one by one,
+ we unfreeze the state of the lexer. This restores the behaviour of the
+ lexer. B.B. *)
+
+IFDEF CAMLP5 THEN
+
+let lexer = {
+ Token.tok_func = func;
+ Token.tok_using = add_token;
+ Token.tok_removing = (fun _ -> ());
+ Token.tok_match = default_match;
+ Token.tok_comm = None;
+ Token.tok_text = token_text }
+
+ELSE
+
+let lexer = {
+ Token.func = func;
+ Token.using = add_token;
+ Token.removing = (fun _ -> ());
+ Token.tparse = (fun _ -> None);
+ Token.text = token_text }
+
+END
(* Terminal symbols interpretation *)
@@ -534,7 +587,7 @@ let is_ident_not_keyword s =
let is_number s =
let rec aux i =
- String.length s = i or
+ String.length s = i or
match s.[i] with '0'..'9' -> aux (i+1) | _ -> false
in aux 0
diff --git a/parsing/lexer.mli b/parsing/lexer.mli
index f1ab6446..1b40d7f1 100644
--- a/parsing/lexer.mli
+++ b/parsing/lexer.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: lexer.mli 7732 2005-12-26 13:51:24Z herbelin $ i*)
+(*i $Id$ i*)
open Pp
open Util
@@ -21,9 +21,9 @@ type error =
exception Error of error
val add_token : string * string -> unit
+val remove_keyword : string -> unit
val is_keyword : string -> bool
-val func : char Stream.t -> (string * string) Stream.t * (int -> loc)
val location_function : int -> loc
(* for coqdoc *)
@@ -34,10 +34,6 @@ val restore_location_table : location_table -> unit
val check_ident : string -> unit
val check_keyword : string -> unit
-val tparse : string * string -> ((string * string) Stream.t -> string) option
-
-val token_text : string * string -> string
-
type frozen_t
val freeze : unit -> frozen_t
val unfreeze : frozen_t -> unit
@@ -50,3 +46,7 @@ val restore_com_state: com_state -> unit
val set_xml_output_comment : (string -> unit) -> unit
val terminal : string -> string * string
+
+(* The lexer of Coq *)
+
+val lexer : Compat.lexer
diff --git a/parsing/parsing.mllib b/parsing/parsing.mllib
new file mode 100644
index 00000000..c0c1817d
--- /dev/null
+++ b/parsing/parsing.mllib
@@ -0,0 +1,12 @@
+Extend
+Extrawit
+Pcoq
+Egrammar
+G_xml
+Ppconstr
+Printer
+Pptactic
+Ppdecl_proof
+Tactic_printer
+Printmod
+Prettyp
diff --git a/parsing/pcoq.ml4 b/parsing/pcoq.ml4
index d2d81cd1..7120f72d 100644
--- a/parsing/pcoq.ml4
+++ b/parsing/pcoq.ml4
@@ -8,7 +8,7 @@
(*i camlp4use: "pa_extend.cmo pa_macro.cmo" i*)
-(*i $Id: pcoq.ml4 12055 2009-04-07 18:19:05Z herbelin $ i*)
+(*i $Id$ i*)
open Pp
open Util
@@ -19,55 +19,28 @@ open Rawterm
open Topconstr
open Genarg
open Tacexpr
+open Extrawit
open Ppextend
-(* The lexer of Coq *)
-
-(* Note: removing a token.
- We do nothing because [remove_token] is called only when removing a grammar
- rule with [Grammar.delete_rule]. The latter command is called only when
- unfreezing the state of the grammar entries (see GRAMMAR summary, file
- env/metasyntax.ml). Therefore, instead of removing tokens one by one,
- we unfreeze the state of the lexer. This restores the behaviour of the
- lexer. B.B. *)
-
-IFDEF CAMLP5 THEN
+(* The parser of Coq *)
-let lexer = {
- Token.tok_func = Lexer.func;
- Token.tok_using = Lexer.add_token;
- Token.tok_removing = (fun _ -> ());
- Token.tok_match = Token.default_match;
- (* Token.parse = Lexer.tparse; *)
- Token.tok_comm = None;
- Token.tok_text = Lexer.token_text }
+IFDEF CAMLP5 THEN
module L =
struct
type te = string * string
- let lexer = lexer
+ let lexer = Lexer.lexer
end
-(* The parser of Coq *)
-
module G = Grammar.GMake(L)
-ELSE
-
-let lexer = {
- Token.func = Lexer.func;
- Token.using = Lexer.add_token;
- Token.removing = (fun _ -> ());
- Token.tparse = Lexer.tparse;
- Token.text = Lexer.token_text }
+ELSE
module L =
struct
- let lexer = lexer
+ let lexer = Lexer.lexer
end
-(* The parser of Coq *)
-
module G = Grammar.Make(L)
END
@@ -82,7 +55,7 @@ let grammar_delete e pos reinit rls =
99 and 200. We didn't find a good solution to this problem
(e.g. using G.extend to know if the level exists results in a
printed error message as side effect). As a consequence an
- extension at 99 or 8 (and for pattern 200 too) inside a section
+ extension at 99 or 8 (and for pattern 200 too) inside a section
corrupts the parser. *)
List.iter (fun (pil,_) -> G.delete_rule e pil) (List.rev lev))
@@ -90,7 +63,7 @@ let grammar_delete e pos reinit rls =
if reinit <> None then
let lev = match pos with Some (Gramext.Level n) -> n | _ -> assert false in
let pos =
- if lev = "200" then Gramext.First
+ if lev = "200" then Gramext.First
else Gramext.After (string_of_int (int_of_string lev + 1)) in
G.extend e (Some pos) [Some lev,reinit,[]];
@@ -134,12 +107,17 @@ end
open Gramtypes
+type camlp4_rule =
+ Compat.token Gramext.g_symbol list * Gramext.g_action
+
+type camlp4_entry_rules =
+ (* first two parameters are name and assoc iff a level is created *)
+ string option * Gramext.g_assoc option * camlp4_rule list
+
type ext_kind =
| ByGrammar of
grammar_object G.Entry.e * Gramext.position option *
- (string option * Gramext.g_assoc option *
- (Compat.token Gramext.g_symbol list * Gramext.g_action) list) list *
- Gramext.g_assoc option
+ camlp4_entry_rules list * Gramext.g_assoc option
| ByGEXTEND of (unit -> unit) * (unit -> unit)
let camlp4_state = ref []
@@ -156,7 +134,15 @@ module Gram =
:: !camlp4_state;
G.extend e pos rls
let delete_rule e pil =
- errorlabstrm "Pcoq.delete_rule" (str "GDELETE_RULE forbidden.")
+ (* spiwack: if you use load an ML module which contains GDELETE_RULE
+ in a section, God kills a kitty. As it would corrupt remove_grammars.
+ There does not seem to be a good way to undo a delete rule. As deleting
+ takes fewer arguments than extending. The production rule isn't returned
+ by delete_rule. If we could retrieve the necessary information, then
+ ByGEXTEND provides just the framework we need to allow this in section.
+ I'm not entirely sure it makes sense, but at least it would be more correct.
+ *)
+ G.delete_rule e pil
end
@@ -212,14 +198,14 @@ let map_entry f en =
let parse_string f x =
let strm = Stream.of_string x in Gram.Entry.parse f (Gram.parsable strm)
-type gram_universe = (string, typed_entry) Hashtbl.t
+type gram_universe = string * (string, typed_entry) Hashtbl.t
let trace = ref false
-(* The univ_tab is not part of the state. It contains all the grammar that
+(* The univ_tab is not part of the state. It contains all the grammars that
exist or have existed before in the session. *)
-let univ_tab = (Hashtbl.create 7 : (string, string * gram_universe) Hashtbl.t)
+let univ_tab = (Hashtbl.create 7 : (string, gram_universe) Hashtbl.t)
let create_univ s =
let u = s, Hashtbl.create 29 in Hashtbl.add univ_tab s u; u
@@ -229,49 +215,27 @@ let uconstr = create_univ "constr"
let utactic = create_univ "tactic"
let uvernac = create_univ "vernac"
-let create_univ_if_new s =
- (* compatibilite *)
- let s = if s = "command" then (warning "'command' grammar universe is obsolete; use name 'constr' instead"; "constr") else s in
+let get_univ s =
try
Hashtbl.find univ_tab s
- with Not_found ->
- if !trace then begin
- Printf.eprintf "[Creating univ %s]\n" s; flush stderr; ()
- end;
- let u = s, Hashtbl.create 29 in Hashtbl.add univ_tab s u; u
+ with Not_found ->
+ anomaly ("Unknown grammar universe: "^s)
-let get_univ = create_univ_if_new
+let get_entry (u, utab) s = Hashtbl.find utab s
-let get_entry (u, utab) s =
+let get_entry_type (u, utab) s =
try
- Hashtbl.find utab s
- with Not_found ->
+ type_of_typed_entry (get_entry (u, utab) s)
+ with Not_found ->
errorlabstrm "Pcoq.get_entry"
(str "Unknown grammar entry " ++ str u ++ str ":" ++ str s ++ str ".")
let new_entry etyp (u, utab) s =
+ if !trace then (Printf.eprintf "[Creating entry %s:%s]\n" u s; flush stderr);
let ename = u ^ ":" ^ s in
let e = in_typed_entry etyp (Gram.Entry.create ename) in
Hashtbl.add utab s e; e
-let entry_type (u, utab) s =
- try
- let e = Hashtbl.find utab s in
- Some (type_of_typed_entry e)
- with Not_found -> None
-
-let get_entry_type (u,n) = type_of_typed_entry (get_entry (get_univ u) n)
-
-let create_entry_if_new (u, utab) s etyp =
- try
- if type_of_typed_entry (Hashtbl.find utab s) <> etyp then
- failwith ("Entry " ^ u ^ ":" ^ s ^ " already exists with another type")
- with Not_found ->
- if !trace then begin
- Printf.eprintf "[Creating entry %s:%s]\n" u s; flush stderr; ()
- end;
- let _ = new_entry etyp (u, utab) s in ()
-
let create_entry (u, utab) s etyp =
try
let e = Hashtbl.find utab s in
@@ -279,105 +243,13 @@ let create_entry (u, utab) s etyp =
failwith ("Entry " ^ u ^ ":" ^ s ^ " already exists with another type");
e
with Not_found ->
- if !trace then begin
- Printf.eprintf "[Creating entry %s:%s]\n" u s; flush stderr; ()
- end;
new_entry etyp (u, utab) s
-let create_constr_entry u s =
- 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 <> etyp then
- failwith ("Entry " ^ u ^ ":" ^ s ^ " already exists with another type");
- 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 (inGramObj wit e); e
-
-let get_generic_entry s =
- let (u,utab) = utactic in
- try
- object_of_typed_entry (Hashtbl.find utab s)
- with Not_found ->
- error ("Unknown grammar entry "^u^":"^s^".")
-
-let get_generic_entry_type (u,utab) s =
- try type_of_typed_entry (Hashtbl.find utab s)
- with Not_found ->
- error ("Unknown grammar entry "^u^":"^s^".")
-
-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 = extyp then
- entry
- else begin
- prerr_endline
- ("Grammar entry " ^ u ^ ":" ^ s ^
- " redefined with another type;\n older entry hidden.");
- Hashtbl.remove utab s;
- new_entry etyp (u, utab) s
- end
- with Not_found ->
- new_entry etyp (u, utab) s
-
-(* Tactics as arguments *)
-
-let tactic_main_level = 5
-
-let (wit_tactic0,globwit_tactic0,rawwit_tactic0) = create_arg "tactic0"
-let (wit_tactic1,globwit_tactic1,rawwit_tactic1) = create_arg "tactic1"
-let (wit_tactic2,globwit_tactic2,rawwit_tactic2) = create_arg "tactic2"
-let (wit_tactic3,globwit_tactic3,rawwit_tactic3) = create_arg "tactic3"
-let (wit_tactic4,globwit_tactic4,rawwit_tactic4) = create_arg "tactic4"
-let (wit_tactic5,globwit_tactic5,rawwit_tactic5) = create_arg "tactic5"
-
-let wit_tactic = function
- | 0 -> wit_tactic0
- | 1 -> wit_tactic1
- | 2 -> wit_tactic2
- | 3 -> wit_tactic3
- | 4 -> wit_tactic4
- | 5 -> wit_tactic5
- | n -> anomaly ("Unavailable tactic level: "^string_of_int n)
-
-let globwit_tactic = function
- | 0 -> globwit_tactic0
- | 1 -> globwit_tactic1
- | 2 -> globwit_tactic2
- | 3 -> globwit_tactic3
- | 4 -> globwit_tactic4
- | 5 -> globwit_tactic5
- | n -> anomaly ("Unavailable tactic level: "^string_of_int n)
-
-let rawwit_tactic = function
- | 0 -> rawwit_tactic0
- | 1 -> rawwit_tactic1
- | 2 -> rawwit_tactic2
- | 3 -> rawwit_tactic3
- | 4 -> rawwit_tactic4
- | 5 -> rawwit_tactic5
- | n -> anomaly ("Unavailable tactic level: "^string_of_int n)
-
-let tactic_genarg_level s =
- if String.length s = 7 && String.sub s 0 6 = "tactic" then
- let c = s.[6] in if '5' >= c && c >= '0' then Some (Char.code c - 48)
- else None
- else None
-
-let is_tactic_genarg = function
-| ExtraArgType s -> tactic_genarg_level s <> None
-| _ -> false
+let create_constr_entry s =
+ outGramObj rawwit_constr (create_entry uconstr s ConstrArgType)
+let create_generic_entry s wit =
+ outGramObj wit (create_entry utactic s (unquote wit))
(* [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 *)
@@ -386,12 +258,12 @@ let make_gen_entry (u,univ) rawwit s =
let e = Gram.Entry.create (u ^ ":" ^ s) in
Hashtbl.add univ s (inGramObj rawwit e); e
-(* Grammar entries *)
+(* Initial grammar entries *)
module Prim =
struct
let gec_gen x = make_gen_entry uprim x
-
+
(* 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"
@@ -401,6 +273,8 @@ module Prim =
let bigint = Gram.Entry.create "Prim.bigint"
let string = gec_gen rawwit_string "string"
let reference = make_gen_entry uprim rawwit_ref "reference"
+ let by_notation = Gram.Entry.create "by_notation"
+ let smart_global = Gram.Entry.create "smart_global"
(* parsed like ident but interpreted as a term *)
let var = gec_gen rawwit_var "var"
@@ -418,10 +292,10 @@ module Prim =
let dirpath = Gram.Entry.create "Prim.dirpath"
let ne_string = Gram.Entry.create "Prim.ne_string"
+ let ne_lstring = Gram.Entry.create "Prim.ne_lstring"
end
-
module Constr =
struct
let gec_constr = make_gen_entry uconstr rawwit_constr
@@ -432,12 +306,11 @@ module Constr =
let operconstr = gec_constr "operconstr"
let constr_eoi = eoi_entry constr
let lconstr = gec_constr "lconstr"
- let binder_constr = create_constr_entry uconstr "binder_constr"
+ let binder_constr = create_constr_entry "binder_constr"
let ident = make_gen_entry uconstr rawwit_ident "ident"
let global = make_gen_entry uconstr rawwit_ref "global"
let sort = make_gen_entry uconstr rawwit_sort "sort"
let pattern = Gram.Entry.create "constr:pattern"
- let annot = Gram.Entry.create "constr:annot"
let constr_pattern = gec_constr "constr_pattern"
let lconstr_pattern = gec_constr "lconstr_pattern"
let binder = Gram.Entry.create "constr:binder"
@@ -462,7 +335,7 @@ module Tactic =
(* Entries that can be refered via the string -> Gram.Entry.e table *)
(* Typically for tactic user extensions *)
- let open_constr =
+ let open_constr =
make_gen_entry utactic (rawwit_open_constr_gen false) "open_constr"
let casted_open_constr =
make_gen_entry utactic (rawwit_open_constr_gen true) "casted_open_constr"
@@ -475,7 +348,7 @@ 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_intropattern =
+ let simple_intropattern =
make_gen_entry utactic rawwit_intro_pattern "simple_intropattern"
(* Main entries for ltac *)
@@ -490,8 +363,6 @@ module Tactic =
end
-
-
module Vernac_ =
struct
let gec_vernac s = Gram.Entry.create ("vernac:" ^ s)
@@ -501,24 +372,22 @@ module Vernac_ =
let gallina_ext = gec_vernac "gallina_ext"
let command = gec_vernac "command"
let syntax = gec_vernac "syntax_command"
- let vernac = gec_vernac "Vernac_.vernac"
-
- (* MMode *)
-
+ let vernac = gec_vernac "Vernac.vernac"
let proof_instr = Gram.Entry.create "proofmode:instr"
- (* /MMode *)
-
let vernac_eoi = eoi_entry vernac
- end
-let main_entry = Gram.Entry.create "vernac"
+ (* Main vernac entry *)
+ let main_entry = Gram.Entry.create "vernac"
+ GEXTEND Gram
+ main_entry:
+ [ [ a = vernac -> Some (loc,a) | EOI -> None ] ]
+ ;
+ END
-GEXTEND Gram
- main_entry:
- [ [ a = Vernac_.vernac -> Some (loc,a) | EOI -> None ] ]
- ;
-END
+ end
+
+let main_entry = Vernac_.main_entry
(**********************************************************************)
(* This determines (depending on the associativity of the current
@@ -527,7 +396,7 @@ END
left border and into "constr LEVEL n" elsewhere), to the level below
(to be translated into "NEXT") or to an below wrt associativity (to be
translated in camlp4 into "constr" without level) or to another level
- (to be translated into "constr LEVEL n")
+ (to be translated into "constr LEVEL n")
The boolean is true if the entry was existing _and_ empty; this to
circumvent a weakness of camlp4/camlp5 whose undo mechanism is not the
@@ -554,7 +423,7 @@ let default_pattern_levels =
1,Gramext.LeftA,false;
0,Gramext.RightA,false]
-let level_stack =
+let level_stack =
ref [(default_levels, default_pattern_levels)]
(* At a same level, LeftA takes precedence over RightA and NoneA *)
@@ -574,7 +443,7 @@ let create_assoc = function
let error_level_assoc p current expected =
let pr_assoc = function
| Gramext.LeftA -> str "left"
- | Gramext.RightA -> str "right"
+ | Gramext.RightA -> str "right"
| Gramext.NonA -> str "non" in
errorlabstrm ""
(str "Level " ++ int p ++ str " is already declared " ++
@@ -640,13 +509,16 @@ let register_empty_levels forpat levels =
let find_position forpat assoc level =
find_position_gen forpat false assoc level
-(* Synchronise the stack of level updates *)
+(* Synchronise the stack of level updates *)
let synchronize_level_positions () =
let _ = find_position true None None in ()
+(**********************************************************************)
+(* Binding constr entry keys to entries *)
+
(* Camlp4 levels do not treat NonA: use RightA with a NEXT on the left *)
let camlp4_assoc = function
- | Some Gramext.NonA | Some Gramext.RightA -> Gramext.RightA
+ | Some Gramext.NonA | Some Gramext.RightA -> Gramext.RightA
| None | Some Gramext.LeftA -> Gramext.LeftA
(* [adjust_level assoc from prod] where [assoc] and [from] are the name
@@ -690,22 +562,20 @@ let compute_entry allow_create adjust forpat = function
(if forpat then weaken_entry Constr.pattern
else weaken_entry Constr.operconstr),
adjust (n,q), false
- | ETIdent -> weaken_entry Constr.ident, None, false
+ | ETName -> weaken_entry Prim.name, None, false
| ETBigint -> weaken_entry Prim.bigint, None, false
| ETReference -> weaken_entry Constr.global, None, false
| ETPattern -> weaken_entry Constr.pattern, None, false
- | ETOther ("constr","annot") ->
- weaken_entry Constr.annot, None, false
| ETConstrList _ -> error "List of entries cannot be registered."
| ETOther (u,n) ->
let u = get_univ u in
let e =
try get_entry u n
- with e when allow_create -> create_entry u n ConstrArgType in
+ with Not_found when allow_create -> create_entry u n ConstrArgType in
object_of_typed_entry e, None, true
(* This computes the name of the level where to add a new rule *)
-let get_constr_entry forpat = function
+let interp_constr_entry_key forpat = function
| ETConstr(200,()) when not forpat ->
weaken_entry Constr.binder_constr, None
| e ->
@@ -714,15 +584,18 @@ let get_constr_entry forpat = function
(* This computes the name to give to a production knowing the name and
associativity of the level where it must be added *)
-let get_constr_production_entry ass from forpat en =
+let interp_constr_prod_entry_key ass from forpat en =
compute_entry false (adjust_level ass from) forpat en
+(**********************************************************************)
+(* Binding constr entry keys to symbols *)
+
let is_self from e =
match from, e with
ETConstr(n,()), ETConstr(NumLevel n',
BorderProd(Right, _ (* Some(Gramext.NonA|Gramext.LeftA) *))) -> false
| ETConstr(n,()), ETConstr(NumLevel n',BorderProd(Left,_)) -> n=n'
- | (ETIdent,ETIdent | ETReference, ETReference | ETBigint,ETBigint
+ | (ETName,ETName | ETReference, ETReference | ETBigint,ETBigint
| ETPattern, ETPattern) -> true
| ETOther(s1,s2), ETOther(s1',s2') -> s1=s1' & s2=s2'
| _ -> false
@@ -733,7 +606,7 @@ let is_binder_level from e =
ETConstr(NumLevel 200,(BorderProd(Right,_)|InternalProd)) -> true
| _ -> false
-let rec symbol_of_production assoc from forpat typ =
+let rec symbol_of_constr_prod_entry_key assoc from forpat typ =
if is_binder_level from typ then
if forpat then
Gramext.Snterml (Gram.Entry.obj Constr.pattern,"200")
@@ -744,28 +617,90 @@ let rec symbol_of_production assoc from forpat typ =
else
match typ with
| ETConstrList (typ',[]) ->
- Gramext.Slist1 (symbol_of_production assoc from forpat (ETConstr typ'))
+ Gramext.Slist1 (symbol_of_constr_prod_entry_key assoc from forpat (ETConstr typ'))
| ETConstrList (typ',tkl) ->
Gramext.Slist1sep
- (symbol_of_production assoc from forpat (ETConstr typ'),
+ (symbol_of_constr_prod_entry_key assoc from forpat (ETConstr typ'),
Gramext.srules
[List.map (fun x -> Gramext.Stoken x) tkl,
List.fold_right (fun _ v -> Gramext.action (fun _ -> v)) tkl
(Gramext.action (fun loc -> ()))])
| _ ->
- match get_constr_production_entry assoc from forpat typ with
+ match interp_constr_prod_entry_key assoc from forpat typ with
| (eobj,None,_) -> Gramext.Snterm (Gram.Entry.obj eobj)
| (eobj,Some None,_) -> Gramext.Snext
- | (eobj,Some (Some (lev,cur)),_) ->
+ | (eobj,Some (Some (lev,cur)),_) ->
Gramext.Snterml (Gram.Entry.obj eobj,constr_level lev)
-(*****************************)
-(* Coercions between entries *)
-
-let coerce_reference_to_id = function
- | Ident (_,id) -> id
- | Qualid (loc,_) ->
- user_err_loc (loc, "coerce_reference_to_id",
- str "This expression should be a simple identifier.")
+(**********************************************************************)
+(* Binding general entry keys to symbol *)
+
+let rec symbol_of_prod_entry_key = function
+ | Alist1 s -> Gramext.Slist1 (symbol_of_prod_entry_key s)
+ | Alist1sep (s,sep) ->
+ Gramext.Slist1sep (symbol_of_prod_entry_key s, Gramext.Stoken ("",sep))
+ | Alist0 s -> Gramext.Slist0 (symbol_of_prod_entry_key s)
+ | Alist0sep (s,sep) ->
+ Gramext.Slist0sep (symbol_of_prod_entry_key s, Gramext.Stoken ("",sep))
+ | Aopt s -> Gramext.Sopt (symbol_of_prod_entry_key s)
+ | Amodifiers s ->
+ Gramext.srules
+ [([], Gramext.action(fun _loc -> []));
+ ([Gramext.Stoken ("", "(");
+ Gramext.Slist1sep ((symbol_of_prod_entry_key s), Gramext.Stoken ("", ","));
+ Gramext.Stoken ("", ")")],
+ Gramext.action (fun _ l _ _loc -> l))]
+ | Aself -> Gramext.Sself
+ | Anext -> Gramext.Snext
+ | Atactic 5 -> Gramext.Snterm (Gram.Entry.obj Tactic.binder_tactic)
+ | Atactic n ->
+ Gramext.Snterml (Gram.Entry.obj Tactic.tactic_expr, string_of_int n)
+ | Agram s -> Gramext.Snterm s
+ | Aentry (u,s) ->
+ Gramext.Snterm (Gram.Entry.obj
+ (object_of_typed_entry (get_entry (get_univ u) s)))
-let coerce_global_to_id = coerce_reference_to_id
+(**********************************************************************)
+(* Interpret entry names of the form "ne_constr_list" as entry keys *)
+
+let rec interp_entry_name static up_level s sep =
+ let l = String.length s in
+ if l > 8 & String.sub s 0 3 = "ne_" & String.sub s (l-5) 5 = "_list" then
+ let t, g = interp_entry_name static up_level (String.sub s 3 (l-8)) "" in
+ List1ArgType t, Alist1 g
+ else if l > 12 & String.sub s 0 3 = "ne_" &
+ String.sub s (l-9) 9 = "_list_sep" then
+ let t, g = interp_entry_name static up_level (String.sub s 3 (l-12)) "" in
+ List1ArgType t, Alist1sep (g,sep)
+ else if l > 5 & String.sub s (l-5) 5 = "_list" then
+ let t, g = interp_entry_name static up_level (String.sub s 0 (l-5)) "" in
+ List0ArgType t, Alist0 g
+ else if l > 9 & String.sub s (l-9) 9 = "_list_sep" then
+ let t, g = interp_entry_name static up_level (String.sub s 0 (l-9)) "" in
+ List0ArgType t, Alist0sep (g,sep)
+ else if l > 4 & String.sub s (l-4) 4 = "_opt" then
+ let t, g = interp_entry_name static up_level (String.sub s 0 (l-4)) "" in
+ OptArgType t, Aopt g
+ else if l > 5 & String.sub s (l-5) 5 = "_mods" then
+ let t, g = interp_entry_name static up_level (String.sub s 0 (l-1)) "" in
+ List0ArgType t, Amodifiers g
+ else
+ let s = if s = "hyp" then "var" else s in
+ let t, se =
+ match Extrawit.tactic_genarg_level s with
+ | Some n when Some n = up_level & up_level <> Some 5 -> None, Aself
+ | Some n when Some (n+1) = up_level & up_level <> Some 5 -> None, Anext
+ | Some n -> None, Atactic n
+ | None ->
+ try Some (get_entry uprim s), Aentry ("prim",s) with Not_found ->
+ try Some (get_entry uconstr s), Aentry ("constr",s) with Not_found ->
+ try Some (get_entry utactic s), Aentry ("tactic",s) with Not_found ->
+ if static then
+ error ("Unknown entry "^s^".")
+ else
+ None, Aentry ("",s) in
+ let t =
+ match t with
+ | Some t -> type_of_typed_entry t
+ | None -> ExtraArgType s in
+ t, se
diff --git a/parsing/pcoq.mli b/parsing/pcoq.mli
index 0a4b349f..ed370a99 100644
--- a/parsing/pcoq.mli
+++ b/parsing/pcoq.mli
@@ -6,47 +6,131 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: pcoq.mli 11784 2009-01-14 11:36:32Z herbelin $ i*)
+(*i $Id$ i*)
open Util
open Names
open Rawterm
open Extend
+open Vernacexpr
open Genarg
open Topconstr
open Tacexpr
-open Vernacexpr
open Libnames
-(* The lexer and parser of Coq. *)
-
-val lexer : Compat.lexer
+(**********************************************************************)
+(* The parser of Coq *)
module Gram : Grammar.S with type te = Compat.token
+(**********************************************************************)
+(* The parser of Coq is built from three kinds of rule declarations:
+ - dynamic rules declared at the evaluation of Coq files (using
+ e.g. Notation, Infix, or Tactic Notation)
+ - static rules explicitly defined in files g_*.ml4
+ - static rules macro-generated by ARGUMENT EXTEND, TACTIC EXTEND and
+ VERNAC EXTEND (see e.g. file extratactics.ml4)
+*)
+
+(* Dynamic extension of rules
+
+ For constr notations, dynamic addition of new rules is done in
+ several steps:
+
+ - "x + y" (user gives a notation string of type Topconstr.notation)
+ | (together with a constr entry level, e.g. 50, and indications of)
+ | (subentries, e.g. x in constr next level and y constr same level)
+ |
+ | spliting into tokens by Metasyntax.split_notation_string
+ V
+ [String "x"; String "+"; String "y"] : symbol_token list
+ |
+ | interpreted as a mixed parsing/printing production
+ | by Metasyntax.analyse_notation_tokens
+ V
+ [NonTerminal "x"; Terminal "+"; NonTerminal "y"] : symbol list
+ |
+ | translated to a parsing production by Metasyntax.make_production
+ V
+ [GramConstrNonTerminal (ETConstr (NextLevel,(BorderProd Left,LeftA)),
+ Some "x");
+ GramConstrTerminal ("","+");
+ GramConstrNonTerminal (ETConstr (NextLevel,(BorderProd Right,LeftA)),
+ Some "y")]
+ : grammar_constr_prod_item list
+ |
+ | Egrammar.make_constr_prod_item
+ V
+ Gramext.g_symbol list which is sent to camlp4
+
+ For user level tactic notations, dynamic addition of new rules is
+ also done in several steps:
+
+ - "f" constr(x) (user gives a Tactic Notation command)
+ |
+ | parsing
+ V
+ [TacTerm "f"; TacNonTerm ("constr", Some "x")]
+ : grammar_tactic_prod_item_expr list
+ |
+ | Metasyntax.interp_prod_item
+ V
+ [GramTerminal "f";
+ GramNonTerminal (ConstrArgType, Aentry ("constr","constr"), Some "x")]
+ : grammar_prod_item list
+ |
+ | Egrammar.make_prod_item
+ V
+ Gramext.g_symbol list
+
+ For TACTIC/VERNAC/ARGUMENT EXTEND, addition of new rules is done as follows:
+
+ - "f" constr(x) (developer gives an EXTEND rule)
+ |
+ | macro-generation in tacextend.ml4/vernacextend.ml4/argextend.ml4
+ V
+ [GramTerminal "f";
+ GramNonTerminal (ConstrArgType, Aentry ("constr","constr"), Some "x")]
+ |
+ | Egrammar.make_prod_item
+ V
+ Gramext.g_symbol list
+
+*)
+
(* The superclass of all grammar entries *)
type grammar_object
+type camlp4_rule =
+ Compat.token Gramext.g_symbol list * Gramext.g_action
+
+type camlp4_entry_rules =
+ (* first two parameters are name and assoc iff a level is created *)
+ string option * Gramext.g_assoc option * camlp4_rule list
+
+(* Add one extension at some camlp4 position of some camlp4 entry *)
+val grammar_extend :
+ grammar_object Gram.Entry.e -> Gramext.position option ->
+ (* for reinitialization if ever needed: *) Gramext.g_assoc option ->
+ camlp4_entry_rules list -> unit
+
+(* Remove the last n extensions *)
+val remove_grammars : int -> unit
+
+
+
+
(* The type of typed grammar objects *)
type typed_entry
+(* The possible types for extensible grammars *)
type entry_type = argument_type
val type_of_typed_entry : typed_entry -> entry_type
val object_of_typed_entry : typed_entry -> grammar_object Gram.Entry.e
val weaken_entry : 'a Gram.Entry.e -> grammar_object Gram.Entry.e
-val get_constr_entry :
- bool -> constr_entry -> grammar_object Gram.Entry.e * int option
-
-val grammar_extend :
- grammar_object Gram.Entry.e -> Gramext.position option ->
- (* for reinitialization if ever: *) Gramext.g_assoc option ->
- (string option * Gramext.g_assoc option *
- (Compat.token Gramext.g_symbol list * Gramext.g_action) list) list
- -> unit
-
-val remove_grammars : int -> unit
+(* Temporary activate camlp4 verbosity *)
val camlp4_verbosity : bool -> ('a -> unit) -> 'a -> unit
@@ -56,73 +140,28 @@ 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
-(* Entry types *)
-
-(* Table of Coq's grammar entries *)
+(**********************************************************************)
+(* Table of Coq statically defined grammar entries *)
type gram_universe
-val create_univ_if_new : string -> string * gram_universe
-val get_univ : string -> string * gram_universe
-val get_entry : string * gram_universe -> string -> typed_entry
-
-val entry_type : string * gram_universe -> string -> entry_type option
-
-val get_entry_type : string * string -> entry_type
-val create_entry_if_new :
- string * gram_universe -> string -> entry_type -> unit
-val create_entry :
- string * gram_universe -> string -> entry_type -> typed_entry
-val force_entry_type :
- string * gram_universe -> string -> entry_type -> typed_entry
-
-val create_constr_entry :
- string * gram_universe -> string -> constr_expr Gram.Entry.e
-val create_generic_entry : string -> ('a, rlevel) 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
-
-(* Tactics as arguments *)
-
-val tactic_main_level : int
+(* There are four predefined universes: "prim", "constr", "tactic", "vernac" *)
-val rawwit_tactic : int -> (raw_tactic_expr,rlevel) abstract_argument_type
-val globwit_tactic : int -> (glob_tactic_expr,glevel) abstract_argument_type
-val wit_tactic : int -> (glob_tactic_expr,tlevel) abstract_argument_type
+val get_univ : string -> gram_universe
-val rawwit_tactic0 : (raw_tactic_expr,rlevel) abstract_argument_type
-val globwit_tactic0 : (glob_tactic_expr,glevel) abstract_argument_type
-val wit_tactic0 : (glob_tactic_expr,tlevel) abstract_argument_type
+val uprim : gram_universe
+val uconstr : gram_universe
+val utactic : gram_universe
+val uvernac : gram_universe
-val rawwit_tactic1 : (raw_tactic_expr,rlevel) abstract_argument_type
-val globwit_tactic1 : (glob_tactic_expr,glevel) abstract_argument_type
-val wit_tactic1 : (glob_tactic_expr,tlevel) abstract_argument_type
+(*
+val get_entry : gram_universe -> string -> typed_entry
+val get_entry_type : gram_universe -> string -> entry_type
+*)
-val rawwit_tactic2 : (raw_tactic_expr,rlevel) abstract_argument_type
-val globwit_tactic2 : (glob_tactic_expr,glevel) abstract_argument_type
-val wit_tactic2 : (glob_tactic_expr,tlevel) abstract_argument_type
-
-val rawwit_tactic3 : (raw_tactic_expr,rlevel) abstract_argument_type
-val globwit_tactic3 : (glob_tactic_expr,glevel) abstract_argument_type
-val wit_tactic3 : (glob_tactic_expr,tlevel) abstract_argument_type
-
-val rawwit_tactic4 : (raw_tactic_expr,rlevel) abstract_argument_type
-val globwit_tactic4 : (glob_tactic_expr,glevel) abstract_argument_type
-val wit_tactic4 : (glob_tactic_expr,tlevel) abstract_argument_type
-
-val rawwit_tactic5 : (raw_tactic_expr,rlevel) abstract_argument_type
-val globwit_tactic5 : (glob_tactic_expr,glevel) abstract_argument_type
-val wit_tactic5 : (glob_tactic_expr,tlevel) abstract_argument_type
-
-val is_tactic_genarg : argument_type -> bool
-
-val tactic_genarg_level : string -> int option
-
-(* The main entry: reads an optional vernac command *)
-
-val main_entry : (loc * vernac_expr) option Gram.Entry.e
-
-(* Initial state of the grammar *)
+val create_entry : gram_universe -> string -> entry_type -> typed_entry
+val create_generic_entry : string -> ('a, rlevel) abstract_argument_type ->
+ 'a Gram.Entry.e
module Prim :
sig
@@ -143,8 +182,11 @@ module Prim :
val qualid : qualid located Gram.Entry.e
val fullyqualid : identifier list located Gram.Entry.e
val reference : reference Gram.Entry.e
+ val by_notation : (loc * string * string option) Gram.Entry.e
+ val smart_global : reference or_by_notation Gram.Entry.e
val dirpath : dir_path Gram.Entry.e
val ne_string : string Gram.Entry.e
+ val ne_lstring : string located Gram.Entry.e
val var : identifier located Gram.Entry.e
end
@@ -159,7 +201,6 @@ module Constr :
val global : reference Gram.Entry.e
val sort : rawsort Gram.Entry.e
val pattern : cases_pattern_expr Gram.Entry.e
- val annot : constr_expr Gram.Entry.e
val constr_pattern : constr_expr Gram.Entry.e
val lconstr_pattern : constr_expr Gram.Entry.e
val binder : (name located list * binder_kind * constr_expr) Gram.Entry.e
@@ -171,10 +212,10 @@ module Constr :
val appl_arg : (constr_expr * explicitation located option) Gram.Entry.e
end
-module Module :
+module Module :
sig
val module_expr : module_ast Gram.Entry.e
- val module_type : module_type_ast Gram.Entry.e
+ val module_type : module_ast Gram.Entry.e
end
module Tactic :
@@ -184,7 +225,7 @@ module Tactic :
val casted_open_constr : open_constr_expr Gram.Entry.e
val constr_with_bindings : constr_expr with_bindings Gram.Entry.e
val bindings : constr_expr bindings Gram.Entry.e
- val constr_may_eval : (constr_expr,reference or_by_notation) may_eval Gram.Entry.e
+ val constr_may_eval : (constr_expr,reference or_by_notation,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
@@ -205,27 +246,43 @@ module Vernac_ :
val command : vernac_expr Gram.Entry.e
val syntax : vernac_expr Gram.Entry.e
val vernac : vernac_expr Gram.Entry.e
-
- (* MMode *)
-
+ val vernac_eoi : vernac_expr Gram.Entry.e
val proof_instr : Decl_expr.raw_proof_instr Gram.Entry.e
+ end
+
+(* The main entry: reads an optional vernac command *)
+val main_entry : (loc * vernac_expr) option Gram.Entry.e
- (*/ MMode *)
+(**********************************************************************)
+(* Mapping formal entries into concrete ones *)
- val vernac_eoi : vernac_expr Gram.Entry.e
- end
+(* Binding constr entry keys to entries and symbols *)
-(* Binding entry names to campl4 entries *)
+val interp_constr_entry_key : bool (* true for cases_pattern *) ->
+ constr_entry_key -> grammar_object Gram.Entry.e * int option
-val symbol_of_production : Gramext.g_assoc option -> constr_entry ->
- bool -> constr_production_entry -> Compat.token Gramext.g_symbol
+val symbol_of_constr_prod_entry_key : Gramext.g_assoc option ->
+ constr_entry_key -> bool -> constr_prod_entry_key ->
+ Compat.token Gramext.g_symbol
-(* Registering/resetting the level of an entry *)
+(* Binding general entry keys to symbols *)
-val find_position :
+val symbol_of_prod_entry_key :
+ Gram.te prod_entry_key -> Gram.te Gramext.g_symbol
+
+(**********************************************************************)
+(* Interpret entry names of the form "ne_constr_list" as entry keys *)
+
+val interp_entry_name : bool (* true to fail on unknown entry *) ->
+ int option -> string -> string -> entry_type * Gram.te prod_entry_key
+
+(**********************************************************************)
+(* Registering/resetting the level of a constr entry *)
+
+val find_position :
bool (* true if for creation in pattern entry; false if in constr entry *) ->
Gramext.g_assoc option -> int option ->
- Gramext.position option * Gramext.g_assoc option * string option *
+ Gramext.position option * Gramext.g_assoc option * string option *
(* for reinitialization: *) Gramext.g_assoc option
val synchronize_level_positions : unit -> unit
@@ -234,6 +291,4 @@ val register_empty_levels : bool -> int list ->
(Gramext.position option * Gramext.g_assoc option *
string option * Gramext.g_assoc option) list
-val remove_levels : int -> unit
-
-val coerce_global_to_id : reference -> identifier
+val remove_levels : int -> unit
diff --git a/parsing/ppconstr.ml b/parsing/ppconstr.ml
index d5357d86..06f3f381 100644
--- a/parsing/ppconstr.ml
+++ b/parsing/ppconstr.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: ppconstr.ml 11739 2009-01-02 19:33:19Z herbelin $ *)
+(* $Id$ *)
(*i*)
open Util
@@ -94,14 +94,14 @@ let pr_delimiters key strm =
let pr_generalization bk ak c =
let hd, tl =
- match bk with
+ match bk with
| Implicit -> "{", "}"
| Explicit -> "(", ")"
- in (* TODO: syntax Abstraction Kind *)
+ in (* TODO: syntax Abstraction Kind *)
str "`" ++ str hd ++ c ++ str tl
let pr_com_at n =
- if Flags.do_beautify() && n <> 0 then comment n
+ if Flags.do_beautify() && n <> 0 then comment n
else mt()
let pr_with_comments loc pp = pr_located (fun x -> x) (loc,pp)
@@ -114,7 +114,7 @@ let pr_optc pr = function
let pr_in_comment pr x = str "(* " ++ pr x ++ str " *)"
-let pr_universe = Univ.pr_uni
+let pr_universe = Univ.pr_uni
let pr_rawsort = function
| RProp Term.Null -> str "Prop"
@@ -124,13 +124,14 @@ let pr_rawsort = function
let pr_id = pr_id
let pr_name = pr_name
let pr_qualid = pr_qualid
+let pr_patvar = pr_id
let pr_expl_args pr (a,expl) =
match expl with
| None -> pr (lapp,L) a
| Some (_,ExplByPos (n,_id)) ->
anomaly("Explicitation by position not implemented")
- | Some (_,ExplByName id) ->
+ | Some (_,ExplByName id) ->
str "(" ++ pr_id id ++ str ":=" ++ pr ltop a ++ str ")"
let pr_opt_type pr = function
@@ -164,16 +165,21 @@ let pr_evar pr n l =
(match l with
| Some l ->
spc () ++ pr_in_comment
- (fun l ->
- str"[" ++ hov 0 (prlist_with_sep pr_coma (pr ltop) l) ++ str"]")
+ (fun l ->
+ str"[" ++ hov 0 (prlist_with_sep pr_comma (pr ltop) l) ++ str"]")
(List.rev l)
| None -> mt()))
let las = lapp
let lpator = 100
+let lpatrec = 0
let rec pr_patt sep inh p =
let (strm,prec) = match p with
+ | CPatRecord (_, l) ->
+ let pp (c, p) =
+ pr_reference c ++ spc() ++ str ":=" ++ pr_patt spc (lpatrec, Any) p in
+ str "{| " ++ prlist_with_sep pr_semicolon pp l ++ str " |}", lpatrec
| CPatAlias (_,p,id) ->
pr_patt mt (las,E) p ++ str " as " ++ pr_id id, las
| CPatCstr (_,c,[]) -> pr_reference c, latom
@@ -200,7 +206,7 @@ let pr_eqn pr (loc,pl,rhs) =
spc() ++ hov 4
(pr_with_comments loc
(str "| " ++
- hov 0 (prlist_with_sep pr_bar (prlist_with_sep sep_v (pr_patt ltop)) pl
+ hov 0 (prlist_with_sep pr_bar (prlist_with_sep sep_v (pr_patt ltop)) pl
++ str " =>") ++
pr_sep_com spc (pr ltop) rhs))
@@ -213,22 +219,22 @@ let begin_of_binders = function
| b::_ -> begin_of_binder b
| _ -> 0
-let surround_impl k p =
+let surround_impl k p =
match k with
| Explicit -> str"(" ++ p ++ str")"
| Implicit -> str"{" ++ p ++ str"}"
-let surround_binder k p =
+let surround_binder k p =
match k with
| Default b -> hov 1 (surround_impl b p)
- | Generalized (b, b', t) ->
+ | Generalized (b, b', t) ->
hov 1 (surround_impl b' (surround_impl b p))
-
+
let surround_implicit k p =
match k with
| Default Explicit -> p
| Default Implicit -> (str"{" ++ p ++ str"}")
- | Generalized (b, b', t) ->
+ | Generalized (b, b', t) ->
surround_impl b' (surround_impl b p)
let pr_binder many pr (nal,k,t) =
@@ -281,7 +287,7 @@ let rec extract_lam_binders = function
let bl,c = extract_lam_binders (CLambdaN(loc,bl,c)) in
LocalRawAssum (nal,bk,t) :: bl, c
| c -> [], c
-
+
let split_lambda = function
| CLambdaN (loc,[[na],bk,t],c) -> (na,t,c)
| CLambdaN (loc,([na],bk,t)::bl,c) -> (na,t,CLambdaN(loc,bl,c))
@@ -293,7 +299,7 @@ let rename na na' t c =
| (_,Name id), (_,Name id') -> (na',t,replace_vars_constr_expr [id,id'] c)
| (_,Name id), (_,Anonymous) -> (na,t,c)
| _ -> (na',t,c)
-
+
let split_product na' = function
| CArrow (loc,t,c) -> (na',t,c)
| CProdN (loc,[[na],bk,t],c) -> rename na na' t c
@@ -324,7 +330,7 @@ let merge_binders (na1,bk1,ty1) cofun (na2,bk2,ty2) codom =
Constrextern.check_same_type ty1 ty2;
ty2 in
(LocalRawAssum ([na],bk1,ty), codom)
-
+
let rec strip_domain bvar cofun c =
match c with
| CArrow(loc,a,b) ->
@@ -401,13 +407,14 @@ let pr_fixdecl pr prd dangling_with_for ((_,id),(n,ro),bl,t,c) =
let annot =
match ro with
CStructRec ->
- if List.length bl > 1 && n <> None then
+ if List.length bl > 1 && n <> None then
spc() ++ str "{struct " ++ pr_id (snd (Option.get n)) ++ str"}"
- else mt()
+ else mt()
| CWfRec c ->
spc () ++ str "{wf " ++ pr lsimple c ++ pr_id (snd (Option.get n)) ++ str"}"
- | CMeasureRec c ->
- spc () ++ str "{measure " ++ pr lsimple c ++ pr_id (snd (Option.get n)) ++ str"}"
+ | CMeasureRec (m,r) ->
+ spc () ++ str "{measure " ++ pr lsimple m ++ pr_id (snd (Option.get n)) ++
+ (match r with None -> mt() | Some r -> str" on " ++ pr lsimple r) ++ str"}"
in
pr_recursive_decl pr prd dangling_with_for id bl annot t c
@@ -427,11 +434,11 @@ let is_var id = function
| _ -> false
let tm_clash = function
- | (CRef (Ident (_,id)), Some (CApp (_,_,nal)))
+ | (CRef (Ident (_,id)), Some (CApp (_,_,nal)))
when List.exists (function CRef (Ident (_,id')),_ -> id=id' | _ -> false)
nal
-> Some id
- | (CRef (Ident (_,id)), Some (CAppExpl (_,_,nal)))
+ | (CRef (Ident (_,id)), Some (CAppExpl (_,_,nal)))
when List.exists (function CRef (Ident (_,id')) -> id=id' | _ -> false)
nal
-> Some id
@@ -444,7 +451,7 @@ let pr_asin pr (na,indnalopt) =
(match indnalopt with
| None -> mt ()
| Some t -> spc () ++ str "in " ++ pr lsimple t)
-
+
let pr_case_item pr (tm,asin) =
hov 0 (pr (lcast,E) tm ++ pr_asin pr asin)
@@ -473,7 +480,7 @@ let pr_appexpl pr f l =
let pr_app pr a l =
hov 2 (
- pr (lapp,L) a ++
+ pr (lapp,L) a ++
prlist (fun a -> spc () ++ pr_expl_args pr a) l)
let pr_forall () =
@@ -486,18 +493,24 @@ let pr_fun () =
let pr_fun_sep = lazy (if !Flags.unicode_syntax then str "," else str " =>")
-let rec pr sep inherited a =
+
+let pr_dangling_with_for sep pr inherited a =
+ match a with
+ | (CFix (_,_,[_])|CCoFix(_,_,[_])) -> pr sep (latom,E) a
+ | _ -> pr sep inherited a
+
+let pr pr sep inherited a =
let (strm,prec) = match a with
| CRef r -> pr_reference r, latom
| CFix (_,id,fix) ->
hov 0 (str"fix " ++
pr_recursive
- (pr_fixdecl (pr mt) (pr_dangling_with_for mt)) (snd id) fix),
+ (pr_fixdecl (pr mt) (pr_dangling_with_for mt pr)) (snd id) fix),
lfix
| CCoFix (_,id,cofix) ->
hov 0 (str "cofix " ++
pr_recursive
- (pr_cofixdecl (pr mt) (pr_dangling_with_for mt)) (snd id) cofix),
+ (pr_cofixdecl (pr mt) (pr_dangling_with_for mt pr)) (snd id) cofix),
lfix
| CArrow (_,a,b) ->
hov 0 (pr mt (larrow,L) a ++ str " ->" ++
@@ -547,29 +560,29 @@ let rec pr sep inherited a =
let c,l1 = list_sep_last l1 in
assert (snd c = None);
let p = pr_proj (pr mt) pr_app (fst c) f l1 in
- if l2<>[] then
+ if l2<>[] then
p ++ prlist (fun a -> spc () ++ pr_expl_args (pr mt) a) l2, lapp
else
p, lproj
| CApp (_,(None,a),l) -> pr_app (pr mt) a l, lapp
| CRecord (_,w,l) ->
- let beg =
+ let beg =
match w with
- | None -> spc ()
+ | None -> spc ()
| Some t -> spc () ++ pr spc ltop t ++ spc () ++ str"with" ++ spc ()
in
- hv 0 (str"{" ++ beg ++
- prlist_with_sep (fun () -> spc () ++ str";" ++ spc ())
- (fun ((_,id), c) -> pr_id id ++ spc () ++ str":=" ++ spc () ++ pr spc ltop c)
- l), latom
+ hv 0 (str"{|" ++ beg ++
+ prlist_with_sep pr_semicolon
+ (fun (id, c) -> h 1 (pr_reference id ++ spc () ++ str":=" ++ pr spc ltop c)) l
+ ++ str" |}"), latom
| CCases (_,LetPatternStyle,rtntypopt,[c,asin],[(_,[(loc,[p])],b)]) ->
hv 0 (
- str "let '" ++
- hov 0 (pr_patt ltop p ++
- pr_asin (pr_dangling_with_for mt) asin ++
- str " :=" ++ pr spc ltop c ++
- pr_case_type (pr_dangling_with_for mt) rtntypopt ++
+ str "let '" ++
+ hov 0 (pr_patt ltop p ++
+ pr_asin (pr_dangling_with_for mt pr) asin ++
+ str " :=" ++ pr spc ltop c ++
+ pr_case_type (pr_dangling_with_for mt pr) rtntypopt ++
str " in" ++ pr spc ltop b)),
lletpattern
| CCases(_,_,rtntypopt,c,eqns) ->
@@ -577,8 +590,8 @@ let rec pr sep inherited a =
(hv 0 (str "match" ++ brk (1,2) ++
hov 0 (
prlist_with_sep sep_v
- (pr_case_item (pr_dangling_with_for mt)) c
- ++ pr_case_type (pr_dangling_with_for mt) rtntypopt) ++
+ (pr_case_item (pr_dangling_with_for mt pr)) c
+ ++ pr_case_type (pr_dangling_with_for mt pr) rtntypopt) ++
spc () ++ str "with") ++
prlist (pr_eqn (pr mt)) eqns ++ spc() ++ str "end"),
latom
@@ -601,7 +614,7 @@ let rec pr sep inherited a =
hov 0 (str "then" ++ pr (fun () -> brk (1,1)) ltop b1) ++ spc () ++
hov 0 (str "else" ++ pr (fun () -> brk (1,1)) ltop b2)),
lif
-
+
| CHole _ -> str "_", latom
| CEvar (_,n,l) -> pr_evar (pr mt) n l, latom
| CPatVar (_,(_,p)) -> str "?" ++ pr_patvar p, latom
@@ -625,12 +638,6 @@ let rec pr sep inherited a =
pr_with_comments loc
(sep() ++ if prec_less prec inherited then strm else surround strm)
-and pr_dangling_with_for sep inherited a =
- match a with
- | (CFix (_,_,[_])|CCoFix(_,_,[_])) -> pr sep (latom,E) a
- | _ -> pr sep inherited a
-
-let pr = pr mt
let rec strip_context n iscast t =
if n = 0 then
@@ -644,7 +651,7 @@ let rec strip_context n iscast t =
else
let bl', c = strip_context (n-n') iscast
(if bll=[] then c else CLambdaN (loc,bll,c)) in
- LocalRawAssum (nal,bk,t) :: bl', c
+ LocalRawAssum (nal,bk,t) :: bl', c
| CProdN (loc,(nal,bk,t)::bll,c) ->
let n' = List.length nal in
if n' > n then
@@ -653,12 +660,12 @@ let rec strip_context n iscast t =
else
let bl', c = strip_context (n-n') iscast
(if bll=[] then c else CProdN (loc,bll,c)) in
- LocalRawAssum (nal,bk,t) :: bl', c
+ LocalRawAssum (nal,bk,t) :: bl', c
| CArrow (loc,t,c) ->
let bl', c = strip_context (n-1) iscast c in
LocalRawAssum ([loc,Anonymous],default_binder_kind,t) :: bl', c
| CCast (_,c,_) -> strip_context n false c
- | CLetIn (_,na,b,c) ->
+ | CLetIn (_,na,b,c) ->
let bl', c = strip_context (n-1) iscast c in
LocalRawDef (na,b) :: bl', c
| _ -> anomaly "strip_context"
@@ -670,6 +677,11 @@ type term_pr = {
pr_lconstr_pattern_expr : constr_pattern_expr -> std_ppcmds
}
+type precedence = Ppextend.precedence * Ppextend.parenRelation
+let modular_constr_pr = pr
+let rec fix rf x =rf (fix rf) x
+let pr = fix modular_constr_pr mt
+
let default_term_pr = {
pr_constr_expr = pr lsimple;
pr_lconstr_expr = pr ltop;
@@ -690,16 +702,13 @@ let pr_cases_pattern_expr = pr_patt ltop
let pr_binders = pr_undelimited_binders (pr ltop)
-let pr_with_occurrences_with_trailer pr occs trailer =
+let pr_with_occurrences pr occs =
match occs with
- ((false,[]),c) -> pr c ++ trailer
+ ((false,[]),c) -> pr c
| ((nowhere_except_in,nl),c) ->
hov 1 (pr c ++ spc() ++ str"at " ++
(if nowhere_except_in then mt() else str "- ") ++
- hov 0 (prlist_with_sep spc (pr_or_var int) nl) ++ trailer)
-
-let pr_with_occurrences pr occs =
- pr_with_occurrences_with_trailer pr occs (mt())
+ hov 0 (prlist_with_sep spc (pr_or_var int) nl))
let pr_red_flag pr r =
(if r.rBeta then pr_arg str "beta" else mt ()) ++
@@ -716,34 +725,34 @@ open Genarg
let pr_metaid id = str"?" ++ pr_id id
-let pr_red_expr (pr_constr,pr_lconstr,pr_ref) = function
+let pr_red_expr (pr_constr,pr_lconstr,pr_ref,pr_pattern) = function
| Red false -> str "red"
| Hnf -> str "hnf"
- | Simpl o -> str "simpl" ++ pr_opt (pr_with_occurrences pr_constr) o
+ | Simpl o -> str "simpl" ++ pr_opt (pr_with_occurrences pr_pattern) o
| Cbv f ->
if f = {rBeta=true;rIota=true;rZeta=true;rDelta=true;rConst=[]} then
str "compute"
else
hov 1 (str "cbv" ++ pr_red_flag pr_ref f)
- | Lazy f ->
+ | Lazy f ->
hov 1 (str "lazy" ++ pr_red_flag pr_ref f)
| Unfold l ->
hov 1 (str "unfold" ++ spc() ++
- prlist_with_sep pr_coma (pr_with_occurrences pr_ref) l)
+ prlist_with_sep pr_comma (pr_with_occurrences pr_ref) l)
| Fold l -> hov 1 (str "fold" ++ prlist (pr_arg pr_constr) l)
| Pattern l ->
hov 1 (str "pattern" ++
- pr_arg (prlist_with_sep pr_coma (pr_with_occurrences pr_constr)) l)
-
+ pr_arg (prlist_with_sep pr_comma (pr_with_occurrences pr_constr)) l)
+
| Red true -> error "Shouldn't be accessible from user."
| ExtraRedExpr s -> str s
| CbvVm -> str "vm_compute"
-let rec pr_may_eval test prc prlc pr2 = function
+let rec pr_may_eval test prc prlc pr2 pr3 = function
| ConstrEval (r,c) ->
hov 0
(str "eval" ++ brk (1,1) ++
- pr_red_expr (prc,prlc,pr2) r ++
+ pr_red_expr (prc,prlc,pr2,pr3) r ++
str " in" ++ spc() ++ prc c)
| ConstrContext ((_,id),c) ->
hov 0
diff --git a/parsing/ppconstr.mli b/parsing/ppconstr.mli
index 0d0c8f56..b33ed682 100644
--- a/parsing/ppconstr.mli
+++ b/parsing/ppconstr.mli
@@ -6,8 +6,8 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-
-(*i $Id: ppconstr.mli 11739 2009-01-02 19:33:19Z herbelin $ i*)
+
+(*i $Id$ i*)
open Pp
open Environ
@@ -28,11 +28,11 @@ val extract_def_binders :
constr_expr -> constr_expr ->
local_binder list * constr_expr * constr_expr
val split_fix :
- int -> constr_expr -> constr_expr ->
+ int -> constr_expr -> constr_expr ->
local_binder list * constr_expr * constr_expr
val prec_less : int -> int * Ppextend.parenRelation -> bool
-
+
val pr_tight_coma : unit -> std_ppcmds
val pr_or_var : ('a -> std_ppcmds) -> 'a or_var -> std_ppcmds
@@ -51,17 +51,16 @@ val pr_sep_com :
val pr_id : identifier -> std_ppcmds
val pr_name : name -> std_ppcmds
val pr_qualid : qualid -> std_ppcmds
+val pr_patvar : patvar -> std_ppcmds
val pr_with_occurrences :
('a -> std_ppcmds) -> 'a with_occurrences -> std_ppcmds
-val pr_with_occurrences_with_trailer :
- ('a -> std_ppcmds) -> 'a with_occurrences -> std_ppcmds -> std_ppcmds
val pr_red_expr :
- ('a -> std_ppcmds) * ('a -> std_ppcmds) * ('b -> std_ppcmds) ->
- ('a,'b) red_expr_gen -> std_ppcmds
+ ('a -> std_ppcmds) * ('a -> std_ppcmds) * ('b -> std_ppcmds) * ('c -> std_ppcmds) ->
+ ('a,'b,'c) red_expr_gen -> std_ppcmds
val pr_may_eval :
- ('a -> std_ppcmds) -> ('a -> std_ppcmds) -> ('b -> std_ppcmds) ->
- ('a,'b) may_eval -> std_ppcmds
+ ('a -> std_ppcmds) -> ('a -> std_ppcmds) -> ('b -> std_ppcmds) ->
+ ('c -> std_ppcmds) -> ('a,'b,'c) may_eval -> std_ppcmds
val pr_rawsort : rawsort -> std_ppcmds
@@ -81,3 +80,24 @@ type term_pr = {
val set_term_pr : term_pr -> unit
val default_term_pr : term_pr
+
+(* The modular constr printer.
+ [modular_constr_pr pr s p t] prints the head of the term [t] and calls
+ [pr] on its subterms.
+ [s] is typically {!Pp.mt} and [p] is [lsimple] for "constr" printers and [ltop]
+ for "lconstr" printers (spiwack: we might need more specification here).
+ We can make a new modular constr printer by overriding certain branches,
+ for instance if we want to build a printer which prints "Prop" as "Omega"
+ instead we can proceed as follows:
+ let my_modular_constr_pr pr s p = function
+ | CSort (_,RProp Null) -> str "Omega"
+ | t -> modular_constr_pr pr s p t
+ Which has the same type. We can turn a modular printer into a printer by
+ taking its fixpoint. *)
+
+type precedence
+val lsimple : precedence
+val ltop : precedence
+val modular_constr_pr :
+ ((unit->std_ppcmds) -> precedence -> constr_expr -> std_ppcmds) ->
+ (unit->std_ppcmds) -> precedence -> constr_expr -> std_ppcmds
diff --git a/parsing/ppdecl_proof.ml b/parsing/ppdecl_proof.ml
index 3b9a002f..abcbedfa 100644
--- a/parsing/ppdecl_proof.ml
+++ b/parsing/ppdecl_proof.ml
@@ -6,45 +6,45 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: ppdecl_proof.ml 10739 2008-04-01 14:45:20Z herbelin $ *)
+(* $Id$ *)
-open Util
+open Util
open Pp
open Decl_expr
-open Names
+open Names
open Nameops
let pr_constr = Printer.pr_constr_env
let pr_tac = Pptactic.pr_glob_tactic
-let pr_pat mpat = Ppconstr.pr_cases_pattern_expr mpat.pat_expr
+let pr_pat mpat = Ppconstr.pr_cases_pattern_expr mpat.pat_expr
let pr_label = function
Anonymous -> mt ()
- | Name id -> pr_id id ++ spc () ++ str ":" ++ spc ()
+ | Name id -> pr_id id ++ spc () ++ str ":" ++ spc ()
let pr_justification_items env = function
Some [] -> mt ()
- | Some (_::_ as l) ->
- spc () ++ str "by" ++ spc () ++
+ | Some (_::_ as l) ->
+ spc () ++ str "by" ++ spc () ++
prlist_with_sep (fun () -> str ",") (pr_constr env) l
| None -> spc () ++ str "by *"
let pr_justification_method env = function
None -> mt ()
- | Some tac ->
+ | Some tac ->
spc () ++ str "using" ++ spc () ++ pr_tac env tac
-let pr_statement pr_it env st =
+let pr_statement pr_it env st =
pr_label st.st_label ++ pr_it env st.st_it
let pr_or_thesis pr_this env = function
Thesis Plain -> str "thesis"
- | Thesis (For id) ->
- str "thesis" ++ spc() ++ str "for" ++ spc () ++ pr_id id
+ | Thesis (For id) ->
+ str "thesis" ++ spc() ++ str "for" ++ spc () ++ pr_id id
| This c -> pr_this env c
-let pr_cut pr_it env c =
- hov 1 (pr_it env c.cut_stat) ++
+let pr_cut pr_it env c =
+ hov 1 (pr_it env c.cut_stat) ++
pr_justification_items env c.cut_by ++
pr_justification_method env c.cut_using
@@ -54,45 +54,45 @@ let type_or_thesis = function
let _I x = x
-let rec print_hyps pconstr gtyp env sep _be _have hyps =
+let rec print_hyps pconstr gtyp env sep _be _have hyps =
let pr_sep = if sep then str "and" ++ spc () else mt () in
- match hyps with
- (Hvar _ ::_) as rest ->
- spc () ++ pr_sep ++ str _have ++
+ match hyps with
+ (Hvar _ ::_) as rest ->
+ spc () ++ pr_sep ++ str _have ++
print_vars pconstr gtyp env false _be _have rest
- | Hprop st :: rest ->
+ | Hprop st :: rest ->
begin
let nenv =
match st.st_label with
Anonymous -> env
| Name id -> Environ.push_named (id,None,gtyp st.st_it) env in
- spc() ++ pr_sep ++ pr_statement pconstr env st ++
+ spc() ++ pr_sep ++ pr_statement pconstr env st ++
print_hyps pconstr gtyp nenv true _be _have rest
end
| [] -> mt ()
and print_vars pconstr gtyp env sep _be _have vars =
match vars with
- Hvar st :: rest ->
+ Hvar st :: rest ->
begin
- let nenv =
+ let nenv =
match st.st_label with
Anonymous -> anomaly "anonymous variable"
| Name id -> Environ.push_named (id,None,st.st_it) env in
- let pr_sep = if sep then pr_coma () else mt () in
+ let pr_sep = if sep then pr_comma () else mt () in
spc() ++ pr_sep ++
pr_statement pr_constr env st ++
print_vars pconstr gtyp nenv true _be _have rest
end
| (Hprop _ :: _) as rest ->
- let _st = if _be then
- str "be such that"
- else
+ let _st = if _be then
+ str "be such that"
+ else
str "such that" in
spc() ++ _st ++ print_hyps pconstr gtyp env false _be _have rest
| [] -> mt ()
-let pr_suffices_clause env (hyps,c) =
+let pr_suffices_clause env (hyps,c) =
print_hyps pr_constr _I env false false "to have" hyps ++ spc () ++
str "to show" ++ spc () ++ pr_or_thesis pr_constr env c
@@ -110,68 +110,68 @@ let pr_side = function
let rec pr_bare_proof_instr _then _thus env = function
| Pescape -> str "escape"
- | Pthen i -> pr_bare_proof_instr true _thus env i
- | Pthus i -> pr_bare_proof_instr _then true env i
+ | Pthen i -> pr_bare_proof_instr true _thus env i
+ | Pthus i -> pr_bare_proof_instr _then true env i
| Phence i -> pr_bare_proof_instr true true env i
- | Pcut c ->
+ | Pcut c ->
begin
match _then,_thus with
- false,false -> str "have" ++ spc () ++
+ false,false -> str "have" ++ spc () ++
pr_cut (pr_statement (pr_or_thesis pr_constr)) env c
- | false,true -> str "thus" ++ spc () ++
+ | false,true -> str "thus" ++ spc () ++
pr_cut (pr_statement (pr_or_thesis pr_constr)) env c
| true,false -> str "then" ++ spc () ++
pr_cut (pr_statement (pr_or_thesis pr_constr)) env c
- | true,true -> str "hence" ++ spc () ++
+ | true,true -> str "hence" ++ spc () ++
pr_cut (pr_statement (pr_or_thesis pr_constr)) env c
end
| Psuffices c ->
- str "suffices" ++ pr_cut pr_suffices_clause env c
+ str "suffices" ++ pr_cut pr_suffices_clause env c
| Prew (sid,c) ->
(if _thus then str "thus" else str " ") ++ spc () ++
pr_side sid ++ spc () ++ pr_cut (pr_statement pr_constr) env c
- | Passume hyps ->
+ | Passume hyps ->
str "assume" ++ print_hyps pr_constr _I env false false "we have" hyps
- | Plet hyps ->
+ | Plet hyps ->
str "let" ++ print_vars pr_constr _I env false true "let" hyps
| Pclaim st ->
str "claim" ++ spc () ++ pr_statement pr_constr env st
| Pfocus st ->
str "focus on" ++ spc () ++ pr_statement pr_constr env st
| Pconsider (id,hyps) ->
- str "consider" ++ print_vars pr_constr _I env false false "consider" hyps
- ++ spc () ++ str "from " ++ pr_constr env id
+ str "consider" ++ print_vars pr_constr _I env false false "consider" hyps
+ ++ spc () ++ str "from " ++ pr_constr env id
| Pgiven hyps ->
str "given" ++ print_vars pr_constr _I env false false "given" hyps
- | Ptake witl ->
- str "take" ++ spc () ++
- prlist_with_sep pr_coma (pr_constr env) witl
+ | Ptake witl ->
+ str "take" ++ spc () ++
+ prlist_with_sep pr_comma (pr_constr env) witl
| Pdefine (id,args,body) ->
- str "define" ++ spc () ++ pr_id id ++ spc () ++
- prlist_with_sep spc
- (fun st -> str "(" ++
- pr_statement pr_constr env st ++ str ")") args ++ spc () ++
- str "as" ++ (pr_constr env body)
- | Pcast (id,typ) ->
- str "reconsider" ++ spc () ++
- pr_or_thesis (fun _ -> pr_id) env id ++ spc () ++
- str "as" ++ spc () ++ (pr_constr env typ)
- | Psuppose hyps ->
- str "suppose" ++
+ str "define" ++ spc () ++ pr_id id ++ spc () ++
+ prlist_with_sep spc
+ (fun st -> str "(" ++
+ pr_statement pr_constr env st ++ str ")") args ++ spc () ++
+ str "as" ++ (pr_constr env body)
+ | Pcast (id,typ) ->
+ str "reconsider" ++ spc () ++
+ pr_or_thesis (fun _ -> pr_id) env id ++ spc () ++
+ str "as" ++ spc () ++ (pr_constr env typ)
+ | Psuppose hyps ->
+ str "suppose" ++
print_hyps pr_constr _I env false false "we have" hyps
| Pcase (params,pat,hyps) ->
str "suppose it is" ++ spc () ++ pr_pat pat ++
- (if params = [] then mt () else
- (spc () ++ str "with" ++ spc () ++
- prlist_with_sep spc
- (fun st -> str "(" ++
- pr_statement pr_constr env st ++ str ")") params ++ spc ()))
+ (if params = [] then mt () else
+ (spc () ++ str "with" ++ spc () ++
+ prlist_with_sep spc
+ (fun st -> str "(" ++
+ pr_statement pr_constr env st ++ str ")") params ++ spc ()))
++
- (if hyps = [] then mt () else
- (spc () ++ str "and" ++
+ (if hyps = [] then mt () else
+ (spc () ++ str "and" ++
print_hyps (pr_or_thesis pr_constr) type_or_thesis
env false false "we have" hyps))
- | Pper (et,c) ->
+ | Pper (et,c) ->
str "per" ++ spc () ++ pr_elim_type et ++ spc () ++
pr_casee env c
| Pend (B_elim et) -> str "end" ++ spc () ++ pr_elim_type et
@@ -184,7 +184,7 @@ let pr_emph = function
| 3 -> str "*** "
| _ -> anomaly "unknown emphasis"
-let pr_proof_instr env instr =
- pr_emph instr.emph ++ spc () ++
+let pr_proof_instr env instr =
+ pr_emph instr.emph ++ spc () ++
pr_bare_proof_instr false false env instr.instr
diff --git a/parsing/ppdecl_proof.mli b/parsing/ppdecl_proof.mli
index b0f0e110..fd6fb663 100644
--- a/parsing/ppdecl_proof.mli
+++ b/parsing/ppdecl_proof.mli
@@ -1,2 +1,2 @@
-val pr_proof_instr : Environ.env -> Decl_expr.proof_instr -> Pp.std_ppcmds
+val pr_proof_instr : Environ.env -> Decl_expr.proof_instr -> Pp.std_ppcmds
diff --git a/parsing/pptactic.ml b/parsing/pptactic.ml
index f52ebc76..466c69eb 100644
--- a/parsing/pptactic.ml
+++ b/parsing/pptactic.ml
@@ -6,11 +6,11 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: pptactic.ml 12581 2009-12-13 15:02:33Z herbelin $ *)
+(* $Id$ *)
open Pp
open Names
-open Nameops
+open Namegen
open Util
open Tacexpr
open Rawterm
@@ -36,8 +36,8 @@ let declare_extra_tactic_pprule (s,tags,prods) =
let exists_extra_tactic_pprule s tags = Hashtbl.mem prtac_tab (s,tags)
type 'a raw_extra_genarg_printer =
- (constr_expr -> std_ppcmds) ->
- (constr_expr -> std_ppcmds) ->
+ (constr_expr -> std_ppcmds) ->
+ (constr_expr -> std_ppcmds) ->
(tolerability -> raw_tactic_expr -> std_ppcmds) ->
'a -> std_ppcmds
@@ -48,8 +48,8 @@ type 'a glob_extra_genarg_printer =
'a -> std_ppcmds
type 'a extra_genarg_printer =
- (Term.constr -> std_ppcmds) ->
- (Term.constr -> std_ppcmds) ->
+ (Term.constr -> std_ppcmds) ->
+ (Term.constr -> std_ppcmds) ->
(tolerability -> glob_tactic_expr -> std_ppcmds) ->
'a -> std_ppcmds
@@ -57,7 +57,7 @@ let genarg_pprule = ref Stringmap.empty
let declare_extra_genarg_pprule (rawwit, f) (globwit, g) (wit, h) =
let s = match unquote wit with
- | ExtraArgType s -> s
+ | ExtraArgType s -> s
| _ -> error
"Can declare a pretty-printing rule only for extra argument types."
in
@@ -84,13 +84,13 @@ let pr_or_by_notation f = function
let pr_located pr (loc,x) = pr x
-let pr_evaluable_reference = function
+let pr_evaluable_reference = function
| EvalVarRef id -> pr_id id
| EvalConstRef sp -> pr_global (Libnames.ConstRef sp)
let pr_quantified_hypothesis = function
| AnonHyp n -> int n
- | NamedHyp id -> pr_id id
+ | NamedHyp id -> pr_id id
let pr_quantified_hypothesis_arg h = spc () ++ pr_quantified_hypothesis h
@@ -103,7 +103,7 @@ let pr_bindings prc prlc = function
brk (1,1) ++ str "with" ++ brk (1,1) ++
prlist_with_sep spc prc l
| ExplicitBindings l ->
- brk (1,1) ++ str "with" ++ brk (1,1) ++
+ brk (1,1) ++ str "with" ++ brk (1,1) ++
prlist_with_sep spc (fun b -> str"(" ++ pr_binding prlc b ++ str")") l
| NoBindings -> mt ()
@@ -112,7 +112,7 @@ let pr_bindings_no_with prc prlc = function
brk (1,1) ++
prlist_with_sep spc prc l
| ExplicitBindings l ->
- brk (1,1) ++
+ brk (1,1) ++
prlist_with_sep spc (fun b -> str"(" ++ pr_binding prlc b ++ str")") l
| NoBindings -> mt ()
@@ -139,7 +139,7 @@ let out_bindings = function
let if_pattern_ident b pr c = (if b then str "?" else mt()) ++ pr c
-let rec pr_raw_generic prc prlc prtac prref (x:Genarg.rlevel Genarg.generic_argument) =
+let rec pr_raw_generic prc prlc prtac prpat prref (x:Genarg.rlevel Genarg.generic_argument) =
match Genarg.genarg_tag x with
| BoolArgType -> str (if out_gen rawwit_bool x then "true" else "false")
| IntArgType -> int (out_gen rawwit_int x)
@@ -153,35 +153,36 @@ let rec pr_raw_generic prc prlc prtac prref (x:Genarg.rlevel Genarg.generic_argu
| SortArgType -> pr_rawsort (out_gen rawwit_sort x)
| ConstrArgType -> prc (out_gen rawwit_constr x)
| ConstrMayEvalArgType ->
- pr_may_eval prc prlc (pr_or_by_notation prref)
+ pr_may_eval prc prlc (pr_or_by_notation prref) prpat
(out_gen rawwit_constr_may_eval x)
| QuantHypArgType -> pr_quantified_hypothesis (out_gen rawwit_quant_hyp x)
| RedExprArgType ->
- pr_red_expr (prc,prlc,pr_or_by_notation prref)
+ pr_red_expr (prc,prlc,pr_or_by_notation prref,prpat)
(out_gen rawwit_red_expr x)
| OpenConstrArgType b -> prc (snd (out_gen (rawwit_open_constr_gen b) x))
- | ConstrWithBindingsArgType ->
+ | ConstrWithBindingsArgType ->
pr_with_bindings prc prlc (out_gen rawwit_constr_with_bindings x)
- | BindingsArgType ->
+ | BindingsArgType ->
pr_bindings_no_with prc prlc (out_gen rawwit_bindings x)
- | List0ArgType _ ->
- hov 0 (pr_sequence (pr_raw_generic prc prlc prtac prref)
+ | List0ArgType _ ->
+ hov 0 (pr_sequence (pr_raw_generic prc prlc prtac prpat prref)
(fold_list0 (fun a l -> a::l) x []))
| List1ArgType _ ->
- hov 0 (pr_sequence (pr_raw_generic prc prlc prtac prref)
+ hov 0 (pr_sequence (pr_raw_generic prc prlc prtac prpat prref)
(fold_list1 (fun a l -> a::l) x []))
- | OptArgType _ -> hov 0 (fold_opt (pr_raw_generic prc prlc prtac prref) (mt()) x)
+ | OptArgType _ -> hov 0 (fold_opt (pr_raw_generic prc prlc prtac prpat prref) (mt()) x)
| PairArgType _ ->
hov 0
(fold_pair
- (fun a b -> pr_sequence (pr_raw_generic prc prlc prtac prref) [a;b])
+ (fun a b -> pr_sequence (pr_raw_generic prc prlc prtac prpat prref)
+ [a;b])
x)
- | ExtraArgType s ->
+ | ExtraArgType s ->
try pi1 (Stringmap.find s !genarg_pprule) prc prlc prtac x
with Not_found -> str "[no printer for " ++ str s ++ str "]"
-let rec pr_glob_generic prc prlc prtac x =
+let rec pr_glob_generic prc prlc prtac prpat x =
match Genarg.genarg_tag x with
| BoolArgType -> str (if out_gen globwit_bool x then "true" else "false")
| IntArgType -> int (out_gen globwit_int x)
@@ -196,38 +197,38 @@ let rec pr_glob_generic prc prlc prtac x =
| ConstrArgType -> prc (out_gen globwit_constr x)
| ConstrMayEvalArgType ->
pr_may_eval prc prlc
- (pr_or_var (pr_and_short_name pr_evaluable_reference))
+ (pr_or_var (pr_and_short_name pr_evaluable_reference)) prpat
(out_gen globwit_constr_may_eval x)
| QuantHypArgType ->
pr_quantified_hypothesis (out_gen globwit_quant_hyp x)
| RedExprArgType ->
- pr_red_expr
- (prc,prlc,pr_or_var (pr_and_short_name pr_evaluable_reference))
+ pr_red_expr
+ (prc,prlc,pr_or_var (pr_and_short_name pr_evaluable_reference),prpat)
(out_gen globwit_red_expr x)
| OpenConstrArgType b -> prc (snd (out_gen (globwit_open_constr_gen b) x))
- | ConstrWithBindingsArgType ->
+ | ConstrWithBindingsArgType ->
pr_with_bindings prc prlc (out_gen globwit_constr_with_bindings x)
- | BindingsArgType ->
+ | BindingsArgType ->
pr_bindings_no_with prc prlc (out_gen globwit_bindings x)
- | List0ArgType _ ->
- hov 0 (pr_sequence (pr_glob_generic prc prlc prtac)
+ | List0ArgType _ ->
+ hov 0 (pr_sequence (pr_glob_generic prc prlc prtac prpat)
(fold_list0 (fun a l -> a::l) x []))
| List1ArgType _ ->
- hov 0 (pr_sequence (pr_glob_generic prc prlc prtac)
+ hov 0 (pr_sequence (pr_glob_generic prc prlc prtac prpat)
(fold_list1 (fun a l -> a::l) x []))
- | OptArgType _ -> hov 0 (fold_opt (pr_glob_generic prc prlc prtac) (mt()) x)
+ | OptArgType _ -> hov 0 (fold_opt (pr_glob_generic prc prlc prtac prpat) (mt()) x)
| PairArgType _ ->
hov 0
(fold_pair
- (fun a b -> pr_sequence (pr_glob_generic prc prlc prtac) [a;b])
+ (fun a b -> pr_sequence (pr_glob_generic prc prlc prtac prpat) [a;b])
x)
- | ExtraArgType s ->
+ | ExtraArgType s ->
try pi2 (Stringmap.find s !genarg_pprule) prc prlc prtac x
with Not_found -> str "[no printer for " ++ str s ++ str "]"
open Closure
-let rec pr_generic prc prlc prtac x =
+let rec pr_generic prc prlc prtac prpat x =
match Genarg.genarg_tag x with
| BoolArgType -> str (if out_gen wit_bool x then "true" else "false")
| IntArgType -> int (out_gen wit_int x)
@@ -243,25 +244,27 @@ let rec pr_generic prc prlc prtac x =
| ConstrMayEvalArgType -> prc (out_gen wit_constr_may_eval x)
| QuantHypArgType -> pr_quantified_hypothesis (out_gen wit_quant_hyp x)
| RedExprArgType ->
- pr_red_expr (prc,prlc,pr_evaluable_reference) (out_gen wit_red_expr x)
+ pr_red_expr (prc,prlc,pr_evaluable_reference,prpat)
+ (out_gen wit_red_expr x)
| OpenConstrArgType b -> prc (snd (out_gen (wit_open_constr_gen b) x))
| ConstrWithBindingsArgType ->
- let (c,b) = out_gen wit_constr_with_bindings x in
- pr_with_bindings prc prlc (c,out_bindings b)
- | BindingsArgType ->
- pr_bindings_no_with prc prlc (out_bindings (out_gen wit_bindings x))
+ let (c,b) = (out_gen wit_constr_with_bindings x).Evd.it in
+ pr_with_bindings prc prlc (c,b)
+ | BindingsArgType ->
+ pr_bindings_no_with prc prlc (out_gen wit_bindings x).Evd.it
| List0ArgType _ ->
- hov 0 (pr_sequence (pr_generic prc prlc prtac)
+ hov 0 (pr_sequence (pr_generic prc prlc prtac prpat)
(fold_list0 (fun a l -> a::l) x []))
| List1ArgType _ ->
- hov 0 (pr_sequence (pr_generic prc prlc prtac)
+ hov 0 (pr_sequence (pr_generic prc prlc prtac prpat)
(fold_list1 (fun a l -> a::l) x []))
- | OptArgType _ -> hov 0 (fold_opt (pr_generic prc prlc prtac) (mt()) x)
+ | OptArgType _ -> hov 0 (fold_opt (pr_generic prc prlc prtac prpat) (mt()) x)
| PairArgType _ ->
hov 0
- (fold_pair (fun a b -> pr_sequence (pr_generic prc prlc prtac) [a;b])
+ (fold_pair (fun a b -> pr_sequence (pr_generic prc prlc prtac prpat)
+ [a;b])
x)
- | ExtraArgType s ->
+ | ExtraArgType s ->
try pi3 (Stringmap.find s !genarg_pprule) prc prlc prtac x
with Not_found -> str "[no printer for " ++ str s ++ str "]"
@@ -275,7 +278,7 @@ let pr_tacarg_using_rule pr_gen l=
pr_sequence (fun x -> x) (tacarg_using_rule_token pr_gen l)
let pr_extend_gen pr_gen lev s l =
- try
+ try
let tags = List.map genarg_tag l in
let (lev',pl) = Hashtbl.find prtac_tab (s,tags) in
let p = pr_tacarg_using_rule pr_gen (pl,l) in
@@ -283,12 +286,12 @@ let pr_extend_gen pr_gen lev s l =
with Not_found ->
str s ++ spc() ++ pr_sequence pr_gen l ++ str" (* Generic printer *)"
-let pr_raw_extend prc prlc prtac =
- pr_extend_gen (pr_raw_generic prc prlc prtac pr_reference)
-let pr_glob_extend prc prlc prtac =
- pr_extend_gen (pr_glob_generic prc prlc prtac)
-let pr_extend prc prlc prtac =
- pr_extend_gen (pr_generic (fun c -> prc (Evd.empty,c)) (fun c -> prlc (Evd.empty,c)) prtac)
+let pr_raw_extend prc prlc prtac prpat =
+ pr_extend_gen (pr_raw_generic prc prlc prtac prpat pr_reference)
+let pr_glob_extend prc prlc prtac prpat =
+ pr_extend_gen (pr_glob_generic prc prlc prtac prpat)
+let pr_extend prc prlc prtac prpat =
+ pr_extend_gen (pr_generic prc prlc prtac prpat)
(**********************************************************************)
(* The tactic printer *)
@@ -320,14 +323,14 @@ let pr_arg pr x = spc () ++ pr x
let pr_ltac_constant sp =
pr_qualid (Nametab.shortest_qualid_of_tactic sp)
-let pr_evaluable_reference_env env = function
+let pr_evaluable_reference_env env = function
| EvalVarRef id -> pr_id id
- | EvalConstRef sp ->
+ | EvalConstRef sp ->
Nametab.pr_global_env (Termops.vars_of_env env) (Libnames.ConstRef sp)
let pr_quantified_hypothesis = function
| AnonHyp n -> int n
- | NamedHyp id -> pr_id id
+ | NamedHyp id -> pr_id id
let pr_quantified_hypothesis_arg h = spc () ++ pr_quantified_hypothesis h
@@ -362,7 +365,7 @@ let pr_with_constr prc = function
let pr_with_induction_names = function
| None, None -> mt ()
| eqpat, ipat ->
- spc () ++ hov 1 (str "as" ++ pr_opt pr_intro_pattern eqpat ++
+ spc () ++ hov 1 (str "as" ++ pr_opt pr_intro_pattern eqpat ++
pr_opt pr_intro_pattern ipat)
let pr_as_intro_pattern ipat =
@@ -410,23 +413,27 @@ let pr_by_tactic prt = function
let pr_hyp_location pr_id = function
| occs, InHyp -> spc () ++ pr_with_occurrences pr_id occs
| occs, InHypTypeOnly ->
- spc () ++
+ spc () ++
pr_with_occurrences (fun id -> str "(type of " ++ pr_id id ++ str ")") occs
| occs, InHypValueOnly ->
- spc () ++
+ spc () ++
pr_with_occurrences (fun id -> str "(value of " ++ pr_id id ++ str ")") occs
let pr_in pp = spc () ++ hov 0 (str "in" ++ pp)
-let pr_simple_clause pr_id = function
+let pr_simple_hyp_clause pr_id = function
| [] -> mt ()
| l -> pr_in (spc () ++ prlist_with_sep spc pr_id l)
let pr_in_hyp_as pr_id = function
| None -> mt ()
- | Some (id,ipat) -> pr_simple_clause pr_id [id] ++ pr_as_ipat ipat
+ | Some (id,ipat) -> pr_simple_hyp_clause pr_id [id] ++ pr_as_ipat ipat
-let pr_clauses pr_id = function
+let pr_clauses default_is_concl pr_id = function
+ | { onhyps=Some []; concl_occs=occs }
+ when occs = all_occurrences_expr & default_is_concl = Some true -> mt ()
+ | { onhyps=None; concl_occs=occs }
+ when occs = all_occurrences_expr & default_is_concl = Some false -> mt ()
| { onhyps=None; concl_occs=occs } ->
if occs = no_occurrences_expr then pr_in (str " * |-")
else pr_in (pr_with_occurrences (fun () -> str " *") (occs,()))
@@ -441,13 +448,13 @@ let pr_clause_pattern pr_id = function
| (glopt,l) ->
str " in" ++
prlist
- (fun (id,nl) -> prlist (pr_arg int) nl
+ (fun (id,nl) -> prlist (pr_arg int) nl
++ spc () ++ pr_id id) l ++
pr_opt (fun nl -> prlist_with_sep spc int nl ++ str " Goal") glopt
let pr_orient b = if b then mt () else str " <-"
-let pr_multi = function
+let pr_multi = function
| Precisely 1 -> mt ()
| Precisely n -> pr_int n ++ str "!"
| UpTo n -> pr_int n ++ str "?"
@@ -485,15 +492,15 @@ let pr_match_rule m pr pr_pat = function
spc () ++ str "=>" ++ brk (1,4) ++ pr t
(*
| Pat (rl,mp,t) ->
- hv 0 (prlist_with_sep pr_coma (pr_match_hyps pr_pat) rl ++
- (if rl <> [] then spc () else mt ()) ++
+ hv 0 (prlist_with_sep pr_comma (pr_match_hyps pr_pat) rl ++
+ (if rl <> [] then spc () else mt ()) ++
hov 0 (str "|-" ++ spc () ++ pr_match_pattern pr_pat mp ++ spc () ++
str "=>" ++ brk (1,4) ++ pr t))
*)
| Pat (rl,mp,t) ->
hov 0 (
- hv 0 (prlist_with_sep pr_coma (pr_match_hyps pr_pat) rl) ++
- (if rl <> [] then spc () else mt ()) ++
+ hv 0 (prlist_with_sep pr_comma (pr_match_hyps pr_pat) rl) ++
+ (if rl <> [] then spc () else mt ()) ++
hov 0 (
str "|-" ++ spc () ++ pr_match_pattern pr_pat mp ++ spc () ++
str "=>" ++ brk (1,4) ++ pr t))
@@ -504,7 +511,7 @@ let pr_funvar = function
| Some id -> spc () ++ pr_id id
let pr_let_clause k pr (id,(bl,t)) =
- hov 0 (str k ++ pr_lident id ++ prlist pr_funvar bl ++
+ hov 0 (str k ++ pr_lident id ++ prlist pr_funvar bl ++
str " :=" ++ brk (1,1) ++ pr (TacArg t))
let pr_let_clauses recflag pr = function
@@ -538,8 +545,8 @@ let pr_hintbases = function
let pr_auto_using prc = function
| [] -> mt ()
- | l -> spc () ++
- hov 2 (str "using" ++ spc () ++ prlist_with_sep pr_coma prc l)
+ | l -> spc () ++
+ hov 2 (str "using" ++ spc () ++ prlist_with_sep pr_comma prc l)
let pr_autoarg_adding = function
| [] -> mt ()
@@ -555,12 +562,6 @@ let pr_autoarg_usingTDB = function
| true -> spc () ++ str "using tdb"
| false -> mt ()
-let rec pr_tacarg_using_rule pr_gen = function
- | Egrammar.TacTerm s :: l, al -> spc () ++ str s ++ pr_tacarg_using_rule pr_gen (l,al)
- | Egrammar.TacNonTerm _ :: l, a :: al -> pr_gen a ++ pr_tacarg_using_rule pr_gen (l,al)
- | [], [] -> mt ()
- | _ -> failwith "Inconsistent arguments of extended tactic"
-
let pr_then () = str ";"
let ltop = (5,E)
@@ -587,7 +588,7 @@ open Closure
used only at the glob and typed level: it is used to feed the
constr printers *)
-let make_pr_tac
+let make_pr_tac
(pr_tac_level,pr_constr,pr_lconstr,pr_pat,
pr_cst,pr_ind,pr_ref,pr_ident,
pr_extend,strip_prod_binders) env =
@@ -596,6 +597,8 @@ let make_pr_tac
constr and cst printers; hence we can make some abbreviations *)
let pr_constr = pr_constr env in
let pr_lconstr = pr_lconstr env in
+let pr_lpat = pr_pat true in
+let pr_pat = pr_pat false in
let pr_cst = pr_cst env in
let pr_ind = pr_ind env in
let pr_tac_level = pr_tac_level env in
@@ -604,8 +607,8 @@ let pr_tac_level = pr_tac_level env in
let pr_bindings = pr_bindings pr_lconstr pr_constr in
let pr_ex_bindings = pr_bindings_gen true pr_lconstr pr_constr in
let pr_with_bindings = pr_with_bindings pr_lconstr pr_constr in
-let pr_extend = pr_extend pr_constr pr_lconstr pr_tac_level in
-let pr_red_expr = pr_red_expr (pr_constr,pr_lconstr,pr_cst) in
+let pr_extend = pr_extend pr_constr pr_lconstr pr_tac_level pr_pat in
+let pr_red_expr = pr_red_expr (pr_constr,pr_lconstr,pr_cst,pr_pat) in
let pr_constrarg c = spc () ++ pr_constr c in
let pr_lconstrarg c = spc () ++ pr_lconstr c in
@@ -632,7 +635,7 @@ let pr_fix_tac (id,n,c) =
match list_chop (n-1) nal with
_, (_,Name id) :: _ -> id, (nal,ty)::bll
| bef, (loc,Anonymous) :: aft ->
- let id = next_ident_away_from (id_of_string"y") avoid in
+ let id = next_ident_away (id_of_string"y") avoid in
id, ((bef@(loc,Name id)::aft, ty)::bll)
| _ -> assert false
else
@@ -650,7 +653,7 @@ let pr_fix_tac (id,n,c) =
let annot =
if List.length names = 1 then mt()
else spc() ++ str"{struct " ++ pr_id idarg ++ str"}" in
- hov 1 (str"(" ++ pr_id id ++
+ hov 1 (str"(" ++ pr_id id ++
prlist pr_binder_fix bll ++ annot ++ str" :" ++
pr_lconstrarg ty ++ str")") in
(* spc() ++
@@ -687,7 +690,7 @@ and pr_atom1 = function
(* Basic tactics *)
| TacIntroPattern [] as t -> pr_atom0 t
- | TacIntroPattern (_::_ as p) ->
+ | TacIntroPattern (_::_ as p) ->
hov 1 (str "intros" ++ spc () ++ prlist_with_sep spc pr_intro_pattern p)
| TacIntrosUntil h ->
hv 1 (str "intros until" ++ pr_arg pr_quantified_hypothesis h)
@@ -701,11 +704,11 @@ and pr_atom1 = function
| TacVmCastNoCheck c -> hov 1 (str "vm_cast_no_check" ++ pr_constrarg c)
| TacApply (a,ev,cb,inhyp) ->
hov 1 ((if a then mt() else str "simple ") ++
- str (with_evars ev "apply") ++ spc () ++
- prlist_with_sep pr_coma pr_with_bindings cb ++
+ str (with_evars ev "apply") ++ spc () ++
+ prlist_with_sep pr_comma pr_with_bindings cb ++
pr_in_hyp_as pr_ident inhyp)
| TacElim (ev,cb,cbo) ->
- hov 1 (str (with_evars ev "elim") ++ pr_arg pr_with_bindings cb ++
+ hov 1 (str (with_evars ev "elim") ++ pr_arg pr_with_bindings cb ++
pr_opt pr_eliminator cbo)
| TacElimType c -> hov 1 (str "elimtype" ++ pr_constrarg c)
| TacCase (ev,cb) ->
@@ -722,16 +725,16 @@ and pr_atom1 = function
hov 1 (str "cofix" ++ spc () ++ pr_id id ++ spc() ++
str"with " ++ prlist_with_sep spc pr_cofix_tac l)
| TacCut c -> hov 1 (str "cut" ++ pr_constrarg c)
- | TacAssert (Some tac,ipat,c) ->
- hov 1 (str "assert" ++
- pr_assumption pr_lconstr pr_constr ipat c ++
+ | TacAssert (Some tac,ipat,c) ->
+ hov 1 (str "assert" ++
+ pr_assumption pr_lconstr pr_constr ipat c ++
pr_by_tactic (pr_tac_level ltop) tac)
- | TacAssert (None,ipat,c) ->
+ | TacAssert (None,ipat,c) ->
hov 1 (str "pose proof" ++
pr_assertion pr_lconstr pr_constr ipat c)
| TacGeneralize l ->
hov 1 (str "generalize" ++ spc () ++
- prlist_with_sep pr_coma (fun (cl,na) ->
+ prlist_with_sep pr_comma (fun (cl,na) ->
pr_with_occurrences pr_constr cl ++ pr_as_name na)
l)
| TacGeneralizeDep c ->
@@ -743,7 +746,7 @@ and pr_atom1 = function
hov 1 ((if b then str "set" else str "remember") ++
(if b then pr_pose pr_lconstr else pr_pose_as_style)
pr_constr na c ++
- pr_clauses pr_ident cl)
+ pr_clauses (Some b) pr_ident cl)
(* | TacInstantiate (n,c,ConclLocation ()) ->
hov 1 (str "instantiate" ++ spc() ++
hov 1 (str"(" ++ pr_arg int n ++ str" :=" ++
@@ -751,24 +754,24 @@ and pr_atom1 = function
| TacInstantiate (n,c,HypLocation (id,hloc)) ->
hov 1 (str "instantiate" ++ spc() ++
hov 1 (str"(" ++ pr_arg int n ++ str" :=" ++
- pr_lconstrarg c ++ str ")" )
+ pr_lconstrarg c ++ str ")" )
++ str "in" ++ pr_hyp_location pr_ident (id,[],(hloc,ref None)))
*)
(* Derived basic tactics *)
| TacSimpleInductionDestruct (isrec,h) ->
hov 1 (str "simple " ++ str (if isrec then "induction" else "destruct")
++ pr_arg pr_quantified_hypothesis h)
- | TacInductionDestruct (isrec,ev,l) ->
+ | TacInductionDestruct (isrec,ev,(l,cl)) ->
hov 1 (str (with_evars ev (if isrec then "induction" else "destruct")) ++
spc () ++
- prlist_with_sep pr_coma (fun (h,e,ids,cl) ->
+ prlist_with_sep pr_comma (fun (h,e,ids) ->
prlist_with_sep spc (pr_induction_arg pr_lconstr pr_constr) h ++
pr_with_induction_names ids ++
- pr_opt pr_eliminator e ++
- pr_opt_no_spc (pr_clauses pr_ident) cl) l)
+ pr_opt pr_eliminator e) l ++
+ pr_opt_no_spc (pr_clauses None pr_ident) cl)
| TacDoubleInduction (h1,h2) ->
hov 1
- (str "double induction" ++
+ (str "double induction" ++
pr_arg pr_quantified_hypothesis h1 ++
pr_arg pr_quantified_hypothesis h2)
| TacDecomposeAnd c ->
@@ -780,22 +783,22 @@ and pr_atom1 = function
hov 0 (str "[" ++ prlist_with_sep spc pr_ind l
++ str "]" ++ pr_constrarg c))
| TacSpecialize (n,c) ->
- hov 1 (str "specialize" ++ spc () ++ pr_opt int n ++
+ hov 1 (str "specialize" ++ spc () ++ pr_opt int n ++
pr_with_bindings c)
- | TacLApply c ->
+ | TacLApply c ->
hov 1 (str "lapply" ++ pr_constrarg c)
(* Automation tactics *)
| TacTrivial ([],Some []) as x -> pr_atom0 x
| TacTrivial (lems,db) ->
- hov 0 (str "trivial" ++
+ hov 0 (str "trivial" ++
pr_auto_using pr_constr lems ++ pr_hintbases db)
| TacAuto (None,[],Some []) as x -> pr_atom0 x
| TacAuto (n,lems,db) ->
- hov 0 (str "auto" ++ pr_opt (pr_or_var int) n ++
+ hov 0 (str "auto" ++ pr_opt (pr_or_var int) n ++
pr_auto_using pr_constr lems ++ pr_hintbases db)
| TacDAuto (n,p,lems) ->
- hov 1 (str "auto" ++ pr_opt (pr_or_var int) n ++ str "decomp" ++
+ hov 1 (str "auto" ++ pr_opt (pr_or_var int) n ++ str "decomp" ++
pr_opt int p ++ pr_auto_using pr_constr lems)
(* Context management *)
@@ -809,59 +812,58 @@ and pr_atom1 = function
(* Rem: only b = true is available for users *)
assert b;
hov 1
- (str "move" ++ brk (1,1) ++ pr_ident id1 ++
+ (str "move" ++ brk (1,1) ++ pr_ident id1 ++
pr_move_location pr_ident id2)
| TacRename l ->
hov 1
(str "rename" ++ brk (1,1) ++
- prlist_with_sep
+ prlist_with_sep
(fun () -> str "," ++ brk (1,1))
- (fun (i1,i2) ->
+ (fun (i1,i2) ->
pr_ident i1 ++ spc () ++ str "into" ++ spc () ++ pr_ident i2)
l)
- | TacRevert l ->
- hov 1 (str "revert" ++ spc () ++ prlist_with_sep spc pr_ident l)
+ | TacRevert l ->
+ hov 1 (str "revert" ++ spc () ++ prlist_with_sep spc pr_ident l)
(* Constructors *)
| TacLeft (ev,l) -> hov 1 (str (with_evars ev "left") ++ pr_bindings l)
| TacRight (ev,l) -> hov 1 (str (with_evars ev "right") ++ pr_bindings l)
- | TacSplit (ev,false,l) -> hov 1 (str (with_evars ev "split") ++ pr_bindings l)
- | TacSplit (ev,true,l) -> hov 1 (str (with_evars ev "exists") ++ pr_ex_bindings l)
+ | TacSplit (ev,false,l) -> hov 1 (str (with_evars ev "split") ++ prlist_with_sep pr_comma pr_bindings l)
+ | TacSplit (ev,true,l) -> hov 1 (str (with_evars ev "exists") ++ prlist_with_sep (fun () -> str",") pr_ex_bindings l)
| TacAnyConstructor (ev,Some t) ->
hov 1 (str (with_evars ev "constructor") ++ pr_arg (pr_tac_level (latom,E)) t)
| TacAnyConstructor (ev,None) as t -> pr_atom0 t
| TacConstructor (ev,n,l) ->
- hov 1 (str (with_evars ev "constructor") ++
+ hov 1 (str (with_evars ev "constructor") ++
pr_or_metaid pr_intarg n ++ pr_bindings l)
- (* Conversion *)
+ (* Conversion *)
| TacReduce (r,h) ->
hov 1 (pr_red_expr r ++
- pr_clauses pr_ident h)
- | TacChange (occ,c,h) ->
+ pr_clauses (Some true) pr_ident h)
+ | TacChange (op,c,h) ->
hov 1 (str "change" ++ brk (1,1) ++
- (match occ with
+ (match op with
None -> mt()
- | Some occlc ->
- pr_with_occurrences_with_trailer pr_constr occlc
- (spc () ++ str "with ")) ++
- pr_constr c ++ pr_clauses pr_ident h)
+ | Some p -> pr_pat p ++ spc () ++ str "with ") ++
+ pr_constr c ++ pr_clauses (Some true) pr_ident h)
(* Equivalence relations *)
| TacReflexivity as x -> pr_atom0 x
- | TacSymmetry cls -> str "symmetry " ++ pr_clauses pr_ident cls
- | TacTransitivity c -> str "transitivity" ++ pr_constrarg c
+ | TacSymmetry cls -> str "symmetry " ++ pr_clauses (Some true) pr_ident cls
+ | TacTransitivity (Some c) -> str "transitivity" ++ pr_constrarg c
+ | TacTransitivity None -> str "etransitivity"
(* Equality and inversion *)
- | TacRewrite (ev,l,cl,by) ->
- hov 1 (str (with_evars ev "rewrite") ++
+ | TacRewrite (ev,l,cl,by) ->
+ hov 1 (str (with_evars ev "rewrite") ++
prlist_with_sep
(fun () -> str ","++spc())
- (fun (b,m,c) ->
+ (fun (b,m,c) ->
pr_orient b ++ spc() ++ pr_multi m ++ pr_with_bindings c)
l
- ++ pr_clauses pr_ident cl
- ++ (match by with Some by -> pr_by_tactic (pr_tac_level ltop) by | None -> mt()))
+ ++ pr_clauses (Some true) pr_ident cl
+ ++ (match by with Some by -> pr_by_tactic (pr_tac_level ltop) by | None -> mt()))
| TacInversion (DepInversion (k,c,ids),hyp) ->
hov 1 (str "dependent " ++ pr_induction_kind k ++ spc () ++
pr_quantified_hypothesis hyp ++
@@ -869,11 +871,11 @@ and pr_atom1 = function
| TacInversion (NonDepInversion (k,cl,ids),hyp) ->
hov 1 (pr_induction_kind k ++ spc () ++
pr_quantified_hypothesis hyp ++
- pr_with_inversion_names ids ++ pr_simple_clause pr_ident cl)
+ pr_with_inversion_names ids ++ pr_simple_hyp_clause pr_ident cl)
| TacInversion (InversionUsing (c,cl),hyp) ->
- hov 1 (str "inversion" ++ spc() ++ pr_quantified_hypothesis hyp ++
- spc () ++ str "using" ++ spc () ++ pr_constr c ++
- pr_simple_clause pr_ident cl)
+ hov 1 (str "inversion" ++ spc() ++ pr_quantified_hypothesis hyp ++
+ spc () ++ str "using" ++ spc () ++ pr_constr c ++
+ pr_simple_hyp_clause pr_ident cl)
in
@@ -881,7 +883,7 @@ let rec pr_tac inherited tac =
let (strm,prec) = match tac with
| TacAbstract (t,None) ->
str "abstract " ++ pr_tac (labstract,L) t, labstract
- | TacAbstract (t,Some s) ->
+ | TacAbstract (t,Some s) ->
hov 0
(str "abstract (" ++ pr_tac (labstract,L) t ++ str")" ++ spc () ++
str "using " ++ pr_id s),
@@ -896,16 +898,16 @@ let rec pr_tac inherited tac =
hov 0 (pr_lazy lz ++ str "match " ++ pr_tac ltop t ++ str " with"
++ prlist
(fun r -> fnl () ++ str "| " ++
- pr_match_rule true (pr_tac ltop) pr_pat r)
+ pr_match_rule true (pr_tac ltop) pr_lpat r)
lrul
++ fnl() ++ str "end"),
lmatch
| TacMatchGoal (lz,lr,lrul) ->
- hov 0 (pr_lazy lz ++
+ hov 0 (pr_lazy lz ++
str (if lr then "match reverse goal with" else "match goal with")
++ prlist
(fun r -> fnl () ++ str "| " ++
- pr_match_rule false (pr_tac ltop) pr_pat r)
+ pr_match_rule false (pr_tac ltop) pr_lpat r)
lrul
++ fnl() ++ str "end"),
lmatch
@@ -914,7 +916,7 @@ let rec pr_tac inherited tac =
prlist pr_funvar lvar ++ str " =>" ++ spc () ++
pr_tac (lfun,E) body),
lfun
- | TacThens (t,tl) ->
+ | TacThens (t,tl) ->
hov 1 (pr_tac (lseq,E) t ++ pr_then () ++ spc () ++
pr_seq_body (pr_tac ltop) tl),
lseq
@@ -930,7 +932,7 @@ let rec pr_tac inherited tac =
hov 1 (str "try" ++ spc () ++ pr_tac (ltactical,E) t),
ltactical
| TacDo (n,t) ->
- hov 1 (str "do " ++ pr_or_var int n ++ spc () ++
+ hov 1 (str "do " ++ pr_or_var int n ++ spc () ++
pr_tac (ltactical,E) t),
ltactical
| TacRepeat t ->
@@ -946,7 +948,7 @@ let rec pr_tac inherited tac =
hov 1 (pr_tac (lorelse,L) t1 ++ str " ||" ++ brk (1,1) ++
pr_tac (lorelse,E) t2),
lorelse
- | TacFail (n,l) ->
+ | TacFail (n,l) ->
str "fail" ++ (if n=ArgArg 0 then mt () else pr_arg (pr_or_var int) n) ++
prlist (pr_arg (pr_message_token pr_ident)) l, latom
| TacFirst tl ->
@@ -967,7 +969,7 @@ let rec pr_tac inherited tac =
| TacArg(ConstrMayEval (ConstrTerm c)) ->
str "constr:" ++ pr_constr c, latom
| TacArg(ConstrMayEval c) ->
- pr_may_eval pr_constr pr_lconstr pr_cst c, leval
+ pr_may_eval pr_constr pr_lconstr pr_cst pr_pat c, leval
| TacArg(TacFreshId l) -> str "fresh" ++ pr_fresh_ids l, latom
| TacArg(Integer n) -> int n, latom
| TacArg(TacCall(loc,f,[])) -> pr_ref f, latom
@@ -989,11 +991,10 @@ and pr_tacarg = function
| IntroPattern ipat -> str "ipattern:" ++ pr_intro_pattern ipat
| TacVoid -> str "()"
| Reference r -> pr_ref r
- | ConstrMayEval c ->
- pr_may_eval pr_constr pr_lconstr pr_cst c
+ | ConstrMayEval c -> pr_may_eval pr_constr pr_lconstr pr_cst pr_pat c
| TacFreshId l -> str "fresh" ++ pr_fresh_ids l
| TacExternal (_,com,req,la) ->
- str "external" ++ spc() ++ qs com ++ spc() ++ qs req ++
+ str "external" ++ spc() ++ qs com ++ spc() ++ qs req ++
spc() ++ prlist_with_sep spc pr_tacarg la
| (TacCall _|Tacexp _|Integer _) as a ->
str "ltac:" ++ pr_tac (latom,E) (TacArg a)
@@ -1009,22 +1010,25 @@ let strip_prod_binders_rawterm n (ty,_) =
| _ -> error "Cannot translate fix tactic: not enough products" in
strip_ty [] n ty
-let strip_prod_binders_constr n (sigma,ty) =
+let strip_prod_binders_constr n ty =
let rec strip_ty acc n ty =
- if n=0 then (List.rev acc, (sigma,ty)) else
+ if n=0 then (List.rev acc, ty) else
match Term.kind_of_term ty with
Term.Prod(na,a,b) ->
- strip_ty (([dummy_loc,na],(sigma,a))::acc) (n-1) b
+ strip_ty (([dummy_loc,na],a)::acc) (n-1) b
| _ -> error "Cannot translate fix tactic: not enough products" in
strip_ty [] n ty
let drop_env f _env = f
+let pr_constr_or_lconstr_pattern_expr b =
+ if b then pr_lconstr_pattern_expr else pr_constr_pattern_expr
+
let rec raw_printers =
- (pr_raw_tactic_level,
+ (pr_raw_tactic_level,
drop_env pr_constr_expr,
drop_env pr_lconstr_expr,
- pr_lconstr_pattern_expr,
+ pr_constr_or_lconstr_pattern_expr,
drop_env (pr_or_by_notation pr_reference),
drop_env (pr_or_by_notation pr_reference),
pr_reference,
@@ -1040,11 +1044,15 @@ and pr_raw_match_rule env t =
let pr_and_constr_expr pr (c,_) = pr c
+let pr_pat_and_constr_expr b (c,_) =
+ pr_and_constr_expr ((if b then pr_lrawconstr_env else pr_rawconstr_env)
+ (Global.env())) c
+
let rec glob_printers =
- (pr_glob_tactic_level,
+ (pr_glob_tactic_level,
(fun env -> pr_and_constr_expr (pr_rawconstr_env env)),
(fun env -> pr_and_constr_expr (pr_lrawconstr_env env)),
- (fun c -> pr_lconstr_pattern_env (Global.env()) c),
+ pr_pat_and_constr_expr,
(fun env -> pr_or_var (pr_and_short_name (pr_evaluable_reference_env env))),
(fun env -> pr_or_var (pr_inductive env)),
pr_ltac_or_var (pr_located pr_ltac_constant),
@@ -1058,11 +1066,14 @@ and pr_glob_tactic_level env n (t:glob_tactic_expr) =
and pr_glob_match_rule env t =
snd (make_pr_tac glob_printers env) t
+let pr_constr_or_lconstr_pattern b =
+ if b then pr_lconstr_pattern else pr_constr_pattern
+
let typed_printers =
(pr_glob_tactic_level,
- pr_open_constr_env,
- pr_open_lconstr_env,
- pr_lconstr_pattern,
+ pr_constr_env,
+ pr_lconstr_env,
+ pr_constr_or_lconstr_pattern,
pr_evaluable_reference_env,
pr_inductive,
pr_ltac_constant,
@@ -1084,9 +1095,10 @@ let _ = Tactic_debug.set_match_pattern_printer
let _ = Tactic_debug.set_match_rule_printer
(fun rl ->
- pr_match_rule false (pr_glob_tactic (Global.env())) pr_constr_pattern rl)
+ pr_match_rule false (pr_glob_tactic (Global.env()))
+ (fun (_,p) -> pr_constr_pattern p) rl)
-open Pcoq
+open Extrawit
let pr_tac_polymorphic n _ _ prtac = prtac (n,E)
@@ -1096,3 +1108,4 @@ let _ = for i=0 to 5 do
(globwit_tactic i, pr_tac_polymorphic i)
(wit_tactic i, pr_tac_polymorphic i)
done
+
diff --git a/parsing/pptactic.mli b/parsing/pptactic.mli
index e24f666f..446dc9f9 100644
--- a/parsing/pptactic.mli
+++ b/parsing/pptactic.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: pptactic.mli 11309 2008-08-06 10:30:35Z herbelin $ i*)
+(*i $Id$ i*)
open Pp
open Genarg
@@ -15,6 +15,7 @@ open Pretyping
open Proof_type
open Topconstr
open Rawterm
+open Pattern
open Ppextend
open Environ
open Evd
@@ -25,8 +26,8 @@ val pr_and_short_name : ('a -> std_ppcmds) -> 'a and_short_name -> std_ppcmds
val pr_or_by_notation : ('a -> std_ppcmds) -> 'a or_by_notation -> std_ppcmds
type 'a raw_extra_genarg_printer =
- (constr_expr -> std_ppcmds) ->
- (constr_expr -> std_ppcmds) ->
+ (constr_expr -> std_ppcmds) ->
+ (constr_expr -> std_ppcmds) ->
(tolerability -> raw_tactic_expr -> std_ppcmds) ->
'a -> std_ppcmds
@@ -37,13 +38,13 @@ type 'a glob_extra_genarg_printer =
'a -> std_ppcmds
type 'a extra_genarg_printer =
- (Term.constr -> std_ppcmds) ->
- (Term.constr -> std_ppcmds) ->
+ (Term.constr -> std_ppcmds) ->
+ (Term.constr -> std_ppcmds) ->
(tolerability -> glob_tactic_expr -> std_ppcmds) ->
'a -> std_ppcmds
(* if the boolean is false then the extension applies only to old syntax *)
-val declare_extra_genarg_pprule :
+val declare_extra_genarg_pprule :
('c raw_abstract_argument_type * 'c raw_extra_genarg_printer) ->
('a glob_abstract_argument_type * 'a glob_extra_genarg_printer) ->
('b typed_abstract_argument_type * 'b extra_genarg_printer) -> unit
@@ -51,31 +52,35 @@ val declare_extra_genarg_pprule :
type grammar_terminals = string option list
(* if the boolean is false then the extension applies only to old syntax *)
-val declare_extra_tactic_pprule :
+val declare_extra_tactic_pprule :
string * argument_type list * (int * grammar_terminals) -> unit
val exists_extra_tactic_pprule : string -> argument_type list -> bool
-val pr_raw_generic :
+val pr_raw_generic :
(constr_expr -> std_ppcmds) ->
(constr_expr -> std_ppcmds) ->
(tolerability -> raw_tactic_expr -> std_ppcmds) ->
- (Libnames.reference -> std_ppcmds) -> constr_expr generic_argument ->
+ (constr_expr -> std_ppcmds) ->
+ (Libnames.reference -> std_ppcmds) -> rlevel generic_argument ->
std_ppcmds
val pr_raw_extend:
(constr_expr -> std_ppcmds) -> (constr_expr -> std_ppcmds) ->
- (tolerability -> raw_tactic_expr -> std_ppcmds) -> int ->
+ (tolerability -> raw_tactic_expr -> std_ppcmds) ->
+ (constr_expr -> std_ppcmds) -> int ->
string -> raw_generic_argument list -> std_ppcmds
val pr_glob_extend:
(rawconstr_and_expr -> std_ppcmds) -> (rawconstr_and_expr -> std_ppcmds) ->
- (tolerability -> glob_tactic_expr -> std_ppcmds) -> int ->
+ (tolerability -> glob_tactic_expr -> std_ppcmds) ->
+ (rawconstr_pattern_and_expr -> std_ppcmds) -> int ->
string -> glob_generic_argument list -> std_ppcmds
val pr_extend :
- (open_constr -> std_ppcmds) -> (open_constr -> std_ppcmds) ->
- (tolerability -> glob_tactic_expr -> std_ppcmds) -> int ->
+ (Term.constr -> std_ppcmds) -> (Term.constr -> std_ppcmds) ->
+ (tolerability -> glob_tactic_expr -> std_ppcmds) ->
+ (constr_pattern -> std_ppcmds) -> int ->
string -> typed_generic_argument list -> std_ppcmds
val pr_ltac_constant : Nametab.ltac_constant -> std_ppcmds
@@ -83,7 +88,7 @@ val pr_ltac_constant : Nametab.ltac_constant -> std_ppcmds
val pr_raw_tactic : env -> raw_tactic_expr -> std_ppcmds
val pr_raw_tactic_level : env -> tolerability -> raw_tactic_expr -> std_ppcmds
-
+
val pr_glob_tactic : env -> glob_tactic_expr -> std_ppcmds
val pr_tactic : env -> Proof_type.tactic_expr -> std_ppcmds
diff --git a/parsing/ppvernac.ml b/parsing/ppvernac.ml
index 5f4ea5a6..ca41c633 100644
--- a/parsing/ppvernac.ml
+++ b/parsing/ppvernac.ml
@@ -6,12 +6,12 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: ppvernac.ml 12187 2009-06-13 19:36:59Z msozeau $ *)
+(* $Id$ *)
open Pp
open Names
open Nameops
-open Nametab
+open Nametab
open Util
open Extend
open Vernacexpr
@@ -50,7 +50,9 @@ let pr_lname = function
(loc,Name id) -> pr_lident (loc,id)
| lna -> pr_located pr_name lna
-let pr_ltac_id = Libnames.pr_reference
+let pr_smart_global = pr_or_by_notation pr_reference
+
+let pr_ltac_ref = Libnames.pr_reference
let pr_module = Libnames.pr_reference
@@ -60,20 +62,20 @@ let sep_end () = str"."
(* Warning: [pr_raw_tactic] globalises and fails if globalisation fails *)
-let pr_raw_tactic_env l env t =
+let pr_raw_tactic_env l env t =
pr_glob_tactic env (Tacinterp.glob_tactic_env l env t)
let pr_gen env t =
- pr_raw_generic
+ pr_raw_generic
pr_constr_expr
pr_lconstr_expr
- (pr_raw_tactic_level env) pr_reference t
+ (pr_raw_tactic_level env) pr_constr_expr pr_reference t
let pr_raw_tactic tac = pr_raw_tactic (Global.env()) tac
let rec extract_signature = function
| [] -> []
- | Egrammar.TacNonTerm (_,(_,t),_) :: l -> t :: extract_signature l
+ | Egrammar.GramNonTerminal (_,t,_,_) :: l -> t :: extract_signature l
| _::l -> extract_signature l
let rec match_vernac_rule tys = function
@@ -105,7 +107,7 @@ let pr_prec = function
| None -> mt()
let pr_set_entry_type = function
- | ETIdent -> str"ident"
+ | ETName -> str"ident"
| ETReference -> str"global"
| ETPattern -> str"pattern"
| ETConstr _ -> str"constr"
@@ -119,9 +121,11 @@ let strip_meta id =
else id
let pr_production_item = function
- | VNonTerm (loc,nt,Some p) -> str nt ++ str"(" ++ pr_id (strip_meta p) ++ str")"
- | VNonTerm (loc,nt,None) -> str nt
- | VTerm s -> qs s
+ | TacNonTerm (loc,nt,Some (p,sep)) ->
+ let pp_sep = if sep <> "" then str "," ++ quote (str sep) else mt () in
+ str nt ++ str"(" ++ pr_id (strip_meta p) ++ pp_sep ++ str")"
+ | TacNonTerm (loc,nt,None) -> str nt
+ | TacTerm s -> qs s
let pr_comment pr_c = function
| CommentConstr c -> pr_c c
@@ -133,20 +137,28 @@ let pr_in_out_modules = function
| SearchOutside [] -> mt()
| SearchOutside l -> spc() ++ str"outside" ++ spc() ++ prlist_with_sep sep pr_module l
-let pr_search_about (b,c) =
+let pr_search_about (b,c) =
(if b then str "-" else mt()) ++
match c with
| SearchSubPattern p -> pr_constr_pattern_expr p
| SearchString (s,sc) -> qs s ++ pr_opt (fun sc -> str "%" ++ str sc) sc
let pr_search a b pr_p = match a with
- | SearchHead qid -> str"Search" ++ spc() ++ pr_reference qid ++ pr_in_out_modules b
+ | SearchHead c -> str"Search" ++ spc() ++ pr_p c ++ pr_in_out_modules b
| SearchPattern c -> str"SearchPattern" ++ spc() ++ pr_p c ++ pr_in_out_modules b
| SearchRewrite c -> str"SearchRewrite" ++ spc() ++ pr_p c ++ pr_in_out_modules b
| SearchAbout sl -> str"SearchAbout" ++ spc() ++ str "[" ++ prlist_with_sep spc pr_search_about sl ++ str "]" ++ pr_in_out_modules b
+let pr_locality_full = function
+ | None -> mt()
+ | Some true -> str"Local "
+ | Some false -> str"Global "
let pr_locality local = if local then str "Local " else str ""
let pr_non_locality local = if local then str "" else str "Global "
+let pr_section_locality local =
+ if Lib.sections_are_opened () && not local then str "Global "
+ else if not (Lib.sections_are_opened ()) && local then str "Local "
+ else mt ()
let pr_explanation (e,b,f) =
let a = match e with
@@ -158,21 +170,18 @@ let pr_explanation (e,b,f) =
let pr_class_rawexpr = function
| FunClass -> str"Funclass"
| SortClass -> str"Sortclass"
- | RefClass qid -> pr_reference qid
+ | RefClass qid -> pr_smart_global qid
let pr_option_ref_value = function
| QualidRefValue id -> pr_reference id
| StringRefValue s -> qs s
-let pr_printoption a b = match a with
- | Goptions.PrimaryTable table -> str table ++ pr_opt (prlist_with_sep sep pr_option_ref_value) b
- | Goptions.SecondaryTable (table,field) -> str table ++ spc() ++ str field ++ pr_opt (prlist_with_sep sep pr_option_ref_value) b
- | Goptions.TertiaryTable (table,field1,field2) -> str table ++ spc() ++
- str field1 ++ spc() ++ str field2 ++
- pr_opt (prlist_with_sep sep pr_option_ref_value) b
+let pr_printoption table b =
+ prlist_with_sep spc str table ++
+ pr_opt (prlist_with_sep sep pr_option_ref_value) b
-let pr_set_option a b =
- let pr_opt_value = function
+let pr_set_option a b =
+ let pr_opt_value = function
| IntValue n -> spc() ++ int n
| StringValue s -> spc() ++ str s
| BoolValue b -> mt()
@@ -188,13 +197,13 @@ let pr_opt_hintbases l = match l with
| [] -> mt()
| _ as z -> str":" ++ spc() ++ prlist_with_sep sep str z
-let pr_hints local db h pr_c pr_pat =
+let pr_hints local db h pr_c pr_pat =
let opth = pr_opt_hintbases db in
let pph =
match h with
| HintsResolve l ->
- str "Resolve " ++ prlist_with_sep sep
- (fun (pri, _, c) -> pr_c c ++
+ str "Resolve " ++ prlist_with_sep sep
+ (fun (pri, _, c) -> pr_c c ++
match pri with Some x -> spc () ++ str"(" ++ int x ++ str")" | None -> mt ())
l
| HintsImmediate l ->
@@ -202,11 +211,11 @@ let pr_hints local db h pr_c pr_pat =
| HintsUnfold l ->
str "Unfold " ++ prlist_with_sep sep pr_reference l
| HintsTransparency (l, b) ->
- str (if b then "Transparent " else "Opaque ") ++ prlist_with_sep sep
+ str (if b then "Transparent " else "Opaque ") ++ prlist_with_sep sep
pr_reference l
| HintsConstructors c ->
str"Constructors" ++ spc() ++ prlist_with_sep spc pr_reference c
- | HintsExtern (n,c,tac) ->
+ | HintsExtern (n,c,tac) ->
let pat = match c with None -> mt () | Some pat -> pr_pat pat in
str "Extern" ++ spc() ++ int n ++ spc() ++ pat ++ str" =>" ++
spc() ++ pr_raw_tactic tac
@@ -225,48 +234,45 @@ let pr_with_declaration pr_c = function
str"Module" ++ spc() ++ pr_lfqid id ++ str" := " ++
pr_located pr_qualid qid
-let rec pr_module_type pr_c = function
- | CMTEident qid -> spc () ++ pr_located pr_qualid qid
- | CMTEwith (mty,decl) ->
- let m = pr_module_type pr_c mty in
+let rec pr_module_ast pr_c = function
+ | CMident qid -> spc () ++ pr_located pr_qualid qid
+ | CMwith (mty,decl) ->
+ let m = pr_module_ast pr_c mty in
let p = pr_with_declaration pr_c decl in
m ++ spc() ++ str"with" ++ spc() ++ p
- | CMTEapply (fexpr,mexpr)->
- let f = pr_module_type pr_c fexpr in
- let m = pr_module_expr mexpr in
- f ++ spc () ++ m
-
-and pr_module_expr = function
- | CMEident qid -> pr_located pr_qualid qid
- | CMEapply (me1,(CMEident _ as me2)) ->
- pr_module_expr me1 ++ spc() ++ pr_module_expr me2
- | CMEapply (me1,me2) ->
- pr_module_expr me1 ++ spc() ++
- hov 1 (str"(" ++ pr_module_expr me2 ++ str")")
-
-let pr_of_module_type prc (mty,b) =
- str (if b then ":" else "<:") ++
- pr_module_type prc mty
+ | CMapply (me1,(CMident _ as me2)) ->
+ pr_module_ast pr_c me1 ++ spc() ++ pr_module_ast pr_c me2
+ | CMapply (me1,me2) ->
+ pr_module_ast pr_c me1 ++ spc() ++
+ hov 1 (str"(" ++ pr_module_ast pr_c me2 ++ str")")
+
+let pr_module_ast_inl pr_c (mast,b) =
+ (if b then mt () else str "!") ++ pr_module_ast pr_c mast
+
+let pr_of_module_type prc = function
+ | Enforce mty -> str ":" ++ pr_module_ast_inl prc mty
+ | Check mtys ->
+ prlist_strict (fun m -> str "<:" ++ pr_module_ast_inl prc m) mtys
let pr_require_token = function
| Some true -> str "Export "
| Some false -> str "Import "
| None -> mt()
-let pr_module_vardecls pr_c (export,idl,mty) =
- let m = pr_module_type pr_c mty in
+let pr_module_vardecls pr_c (export,idl,(mty,inl)) =
+ let m = pr_module_ast pr_c mty in
(* Update the Nametab for interpreting the body of module/modtype *)
let lib_dir = Lib.library_dp() in
List.iter (fun (_,id) ->
Declaremods.process_module_bindings [id]
[make_mbid lib_dir (string_of_id id),
- Modintern.interp_modtype (Global.env()) mty]) idl;
+ (Modintern.interp_modtype (Global.env()) mty, inl)]) idl;
(* Builds the stream *)
spc() ++
hov 1 (str"(" ++ pr_require_token export ++
prlist_with_sep spc pr_lident idl ++ str":" ++ m ++ str")")
-let pr_module_binders l pr_c =
+let pr_module_binders l pr_c =
(* Effet de bord complexe pour garantir la declaration des noms des
modules parametres dans la Nametab des l'appel de pr_module_binders
malgre l'aspect paresseux des streams *)
@@ -279,10 +285,9 @@ let pr_type_option pr_c = function
| CHole (loc, k) -> mt()
| _ as c -> brk(0,2) ++ str":" ++ pr_c c
-let pr_decl_notation prc =
- pr_opt (fun (ntn,c,scopt) -> fnl () ++
- str "where " ++ qs ntn ++ str " := " ++ prc c ++
- pr_opt (fun sc -> str ": " ++ str sc) scopt)
+let pr_decl_notation prc ((loc,ntn),c,scopt) =
+ fnl () ++ str "where " ++ qs ntn ++ str " := " ++ prc c ++
+ pr_opt (fun sc -> str ": " ++ str sc) scopt
let pr_vbinders l =
hv 0 (pr_binders l)
@@ -293,23 +298,45 @@ let pr_binders_arg =
let pr_and_type_binders_arg bl =
pr_binders_arg bl
+let names_of_binder = function
+ | LocalRawAssum (nal,_,_) -> nal
+ | LocalRawDef (_,_) -> []
+
+let pr_guard_annot bl (n,ro) =
+ match n with
+ | None -> mt ()
+ | Some (loc, id) ->
+ match (ro : Topconstr.recursion_order_expr) with
+ | CStructRec ->
+ let ids = List.flatten (List.map names_of_binder bl) in
+ if List.length ids > 1 then
+ spc() ++ str "{struct " ++ pr_id id ++ str"}"
+ else mt()
+ | CWfRec c ->
+ spc() ++ str "{wf " ++ pr_lconstr_expr c ++ spc() ++
+ pr_id id ++ str"}"
+ | CMeasureRec (m,r) ->
+ spc() ++ str "{measure " ++ pr_lconstr_expr m ++ spc() ++
+ pr_id id ++ (match r with None -> mt() | Some r -> str" on " ++
+ pr_lconstr_expr r) ++ str"}"
+
let pr_onescheme (idop,schem) =
- match schem with
+ match schem with
| InductionScheme (dep,ind,s) ->
(match idop with
| Some id -> hov 0 (pr_lident id ++ str" :=") ++ spc()
| None -> spc ()
) ++
hov 0 ((if dep then str"Induction for" else str"Minimality for")
- ++ spc() ++ pr_reference ind) ++ spc() ++
+ ++ spc() ++ pr_smart_global ind) ++ spc() ++
hov 0 (str"Sort" ++ spc() ++ pr_rawsort s)
- | EqualityScheme ind ->
+ | EqualityScheme ind ->
(match idop with
| Some id -> hov 0 (pr_lident id ++ str" :=") ++ spc()
| None -> spc()
) ++
hov 0 (str"Equality for")
- ++ spc() ++ pr_reference ind
+ ++ spc() ++ pr_smart_global ind
let begin_of_inductive = function
[] -> 0
@@ -318,7 +345,7 @@ let begin_of_inductive = function
let pr_class_rawexpr = function
| FunClass -> str"Funclass"
| SortClass -> str"Sortclass"
- | RefClass qid -> pr_reference qid
+ | RefClass qid -> pr_smart_global qid
let pr_assumption_token many = function
| (Local,Logical) ->
@@ -327,10 +354,10 @@ let pr_assumption_token many = function
str (if many then "Variables" else "Variable")
| (Global,Logical) ->
str (if many then "Axioms" else "Axiom")
- | (Global,Definitional) ->
+ | (Global,Definitional) ->
str (if many then "Parameters" else "Parameter")
| (Global,Conjectural) -> str"Conjecture"
- | (Local,Conjectural) ->
+ | (Local,Conjectural) ->
anomaly "Don't know how to beautify a local conjecture"
let pr_params pr_c (xl,(c,t)) =
@@ -374,14 +401,14 @@ let pr_syntax_modifier = function
let pr_syntax_modifiers = function
| [] -> mt()
- | l -> spc() ++
+ | l -> spc() ++
hov 1 (str"(" ++ prlist_with_sep sep_v2 pr_syntax_modifier l ++ str")")
let print_level n =
if n <> 0 then str " (at level " ++ int n ++ str ")" else mt ()
let pr_grammar_tactic_rule n (_,pil,t) =
- hov 2 (str "Tactic Notation" ++ print_level n ++ spc() ++
+ hov 2 (str "Tactic Notation" ++ print_level n ++ spc() ++
hov 0 (prlist_with_sep sep pr_production_item pil ++
spc() ++ str":=" ++ spc() ++ pr_raw_tactic t))
@@ -392,7 +419,7 @@ let pr_box b = let pr_boxkind = function
| PpHOVB n -> str"hov" ++ spc() ++ int n
| PpTB -> str"t"
in str"<" ++ pr_boxkind b ++ str">"
-
+
let pr_paren_reln_or_extern = function
| None,L -> str"L"
| None,E -> str"E"
@@ -400,6 +427,14 @@ let pr_paren_reln_or_extern = function
| Some pprim,Prec p -> qs pprim ++ spc() ++ str":" ++ spc() ++ int p
| _ -> mt()
+let pr_statement head (id,(bl,c,guard)) =
+ assert (id<>None);
+ hov 0
+ (head ++ pr_lident (Option.get id) ++ spc() ++
+ (match bl with [] -> mt() | _ -> pr_binders bl ++ spc()) ++
+ pr_opt (pr_guard_annot bl) guard ++
+ str":" ++ pr_spc_lconstr c)
+
(**************************************)
(* Pretty printer for vernac commands *)
(**************************************)
@@ -409,7 +444,7 @@ let pr_constrarg c = spc () ++ pr_constr c in
let pr_lconstrarg c = spc () ++ pr_lconstr c in
let pr_intarg n = spc () ++ int n in
(* let pr_lident_constr sep (i,c) = pr_lident i ++ sep ++ pr_constrarg c in *)
-let pr_record_field (x, ntn) =
+let pr_record_field (x, ntn) =
let prx = match x with
| (oc,AssumExpr (id,t)) ->
hov 1 (pr_lname id ++
@@ -423,15 +458,15 @@ let pr_record_field (x, ntn) =
| None ->
hov 1 (pr_lname id ++ str" :=" ++ spc() ++
pr_lconstr b)) in
- prx ++ pr_decl_notation pr_constr ntn
+ prx ++ prlist (pr_decl_notation pr_constr) ntn
in
-let pr_record_decl b c fs =
+let pr_record_decl b c fs =
pr_opt pr_lident c ++ str"{" ++
hv 0 (prlist_with_sep pr_semicolon pr_record_field fs ++ str"}")
in
let rec pr_vernac = function
-
+
(* Proof management *)
| VernacAbortAll -> str "Abort All"
| VernacRestart -> str"Restart"
@@ -442,17 +477,17 @@ let rec pr_vernac = function
| VernacResume id -> str"Resume" ++ pr_opt pr_lident id
| VernacUndo i -> if i=1 then str"Undo" else str"Undo" ++ pr_intarg i
| VernacUndoTo i -> str"Undo" ++ spc() ++ str"To" ++ pr_intarg i
- | VernacBacktrack (i,j,k) ->
+ | VernacBacktrack (i,j,k) ->
str "Backtrack" ++ spc() ++ prlist_with_sep sep int [i;j;k]
| VernacFocus i -> str"Focus" ++ pr_opt int i
- | VernacGo g ->
+ | VernacGo g ->
let pr_goable = function
| GoTo i -> int i
| GoTop -> str"top"
| GoNext -> str"next"
- | GoPrev -> str"prev"
+ | GoPrev -> str"prev"
in str"Go" ++ spc() ++ pr_goable g
- | VernacShow s ->
+ | VernacShow s ->
let pr_showable = function
| ShowGoal n -> str"Show" ++ pr_opt int n
| ShowGoalImplicitly n -> str"Show Implicit Arguments" ++ pr_opt int n
@@ -466,7 +501,7 @@ let rec pr_vernac = function
| ShowMatch id -> str"Show Match " ++ pr_lident id
| ShowThesis -> str "Show Thesis"
| ExplainProof l -> str"Explain Proof" ++ spc() ++ prlist_with_sep sep int l
- | ExplainTree l -> str"Explain Proof Tree" ++ spc() ++ prlist_with_sep sep int l
+ | ExplainTree l -> str"Explain Proof Tree" ++ spc() ++ prlist_with_sep sep int l
in pr_showable s
| VernacCheckGuard -> str"Guarded"
@@ -485,15 +520,18 @@ let rec pr_vernac = function
| VernacList l ->
hov 2 (str"[" ++ spc() ++
prlist (fun v -> pr_located pr_vernac v ++ sep_end () ++ fnl()) l
- ++ spc() ++ str"]")
+ ++ spc() ++ str"]")
| VernacLoad (f,s) -> str"Load" ++ if f then (spc() ++ str"Verbose"
++ spc()) else spc() ++ qs s
| VernacTime v -> str"Time" ++ spc() ++ pr_vernac v
-
- (* Syntax *)
+ | VernacTimeout(n,v) -> str"Timeout " ++ int n ++ spc() ++ pr_vernac v
+ | VernacFail v -> str"Fail" ++ spc() ++ pr_vernac v
+
+ (* Syntax *)
| VernacTacticNotation (n,r,e) -> pr_grammar_tactic_rule n ("",r,e)
| VernacOpenCloseScope (local,opening,sc) ->
- str (if opening then "Open " else "Close ") ++ pr_locality local ++
+ pr_section_locality local ++
+ str (if opening then "Open " else "Close ") ++
str "Scope" ++ spc() ++ str sc
| VernacDelimiters (sc,key) ->
str"Delimit Scope" ++ spc () ++ str sc ++
@@ -501,33 +539,34 @@ let rec pr_vernac = function
| VernacBindScope (sc,cll) ->
str"Bind Scope" ++ spc () ++ str sc ++
spc() ++ str "with " ++ prlist_with_sep spc pr_class_rawexpr cll
- | VernacArgumentsScope (local,q,scl) -> let pr_opt_scope = function
+ | VernacArgumentsScope (local,q,scl) -> let pr_opt_scope = function
| None -> str"_"
- | Some sc -> str sc in
- str"Arguments Scope" ++ spc() ++ pr_non_locality local ++ pr_reference q
+ | Some sc -> str sc in
+ pr_section_locality local ++ str"Arguments Scope" ++ spc() ++
+ pr_smart_global q
++ spc() ++ str"[" ++ prlist_with_sep sep pr_opt_scope scl ++ str"]"
- | VernacInfix (local,(s,mv),q,sn) -> (* A Verifier *)
- hov 0 (hov 0 (str"Infix " ++ pr_locality local
- ++ qs s ++ str " :=" ++ spc() ++ pr_reference q) ++
+ | VernacInfix (local,((_,s),mv),q,sn) -> (* A Verifier *)
+ hov 0 (hov 0 (pr_locality local ++ str"Infix "
+ ++ qs s ++ str " :=" ++ pr_constrarg q) ++
pr_syntax_modifiers mv ++
(match sn with
| None -> mt()
| Some sc -> spc() ++ str":" ++ spc() ++ str sc))
- | VernacNotation (local,c,(s,l),opt) ->
+ | VernacNotation (local,c,((_,s),l),opt) ->
let ps =
let n = String.length s in
- if n > 2 & s.[0] = '\'' & s.[n-1] = '\''
+ if n > 2 & s.[0] = '\'' & s.[n-1] = '\''
then
let s' = String.sub s 1 (n-2) in
if String.contains s' '\'' then qs s else str s'
else qs s in
- hov 2( str"Notation" ++ spc() ++ pr_locality local ++ ps ++
+ hov 2 (pr_locality local ++ str"Notation" ++ spc() ++ ps ++
str " :=" ++ pr_constrarg c ++ pr_syntax_modifiers l ++
(match opt with
| None -> mt()
| Some sc -> str" :" ++ spc() ++ str sc))
| VernacSyntaxExtension (local,(s,l)) ->
- str"Reserved Notation" ++ spc() ++ pr_locality local ++ qs s ++
+ pr_locality local ++ str"Reserved Notation" ++ spc() ++ pr_located qs s ++
pr_syntax_modifiers l
(* Gallina *)
@@ -537,7 +576,7 @@ let rec pr_vernac = function
| None -> mt()
| Some r ->
str"Eval" ++ spc() ++
- pr_red_expr (pr_constr, pr_lconstr, pr_or_by_notation pr_reference) r ++
+ pr_red_expr (pr_constr, pr_lconstr, pr_smart_global, pr_constr) r ++
str" in" ++ spc() in
let pr_def_body = function
| DefineBody (bl,red,body,d) ->
@@ -555,11 +594,6 @@ let rec pr_vernac = function
| Some cc -> str" :=" ++ spc() ++ cc))
| VernacStartTheoremProof (ki,l,_,_) ->
- let pr_statement head (id,(bl,c)) =
- hov 0
- (head ++ pr_opt pr_lident id ++ spc() ++
- (match bl with [] -> mt() | _ -> pr_binders bl ++ spc()) ++
- str":" ++ pr_spc_lconstr c) in
hov 1 (pr_statement (pr_thm_token ki) (List.hd l) ++
prlist (pr_statement (spc () ++ str "with")) (List.tl l))
@@ -568,15 +602,15 @@ let rec pr_vernac = function
| None -> if opac then str"Qed" else str"Defined"
| Some (id,th) -> (match th with
| None -> (if opac then str"Save" else str"Defined") ++ spc() ++ pr_lident id
- | Some tok -> str"Save" ++ spc() ++ pr_thm_token tok ++ spc() ++ pr_lident id))
+ | Some tok -> str"Save" ++ spc() ++ pr_thm_token tok ++ spc() ++ pr_lident id))
| VernacExactProof c ->
hov 2 (str"Proof" ++ pr_lconstrarg c)
| VernacAssumption (stre,_,l) ->
let n = List.length (List.flatten (List.map fst (List.map snd l))) in
hov 2
- (pr_assumption_token (n > 1) stre ++ spc() ++
+ (pr_assumption_token (n > 1) stre ++ spc() ++
pr_ne_params_list pr_lconstr_expr l)
- | VernacInductive (f,l) ->
+ | VernacInductive (f,i,l) ->
let pr_constructor (coe,(id,c)) =
hov 2 (pr_lident id ++ str" " ++
@@ -588,79 +622,52 @@ let rec pr_vernac = function
pr_com_at (begin_of_inductive l) ++
fnl() ++
str (if List.length l = 1 then " " else " | ") ++
- prlist_with_sep (fun _ -> fnl() ++ str" | ") pr_constructor l
- | RecordDecl (c,fs) ->
+ prlist_with_sep (fun _ -> fnl() ++ str" | ") pr_constructor l
+ | RecordDecl (c,fs) ->
spc() ++
pr_record_decl b c fs in
let pr_oneind key (((coe,id),indpar,s,k,lc),ntn) =
- let kw =
- str (match k with Record -> "Record" | Structure -> "Structure"
- | Inductive_kw -> "Inductive" | CoInductive -> "CoInductive"
- | Class b -> if b then "Definitional Class" else "Class")
- in
- hov 0 (
- kw ++ spc() ++
- (if coe then str" > " else str" ") ++ pr_lident id ++
- pr_and_type_binders_arg indpar ++ spc() ++
- Option.cata (fun s -> str":" ++ spc() ++ pr_lconstr_expr s) (mt()) s ++
- str" :=") ++ pr_constructor_list k lc ++
- pr_decl_notation pr_constr ntn
+ hov 0 (
+ str key ++ spc() ++
+ (if i then str"Infer " else str"") ++
+ (if coe then str"> " else str"") ++ pr_lident id ++
+ pr_and_type_binders_arg indpar ++ spc() ++
+ Option.cata (fun s -> str":" ++ spc() ++ pr_lconstr_expr s) (mt()) s ++
+ str" :=") ++ pr_constructor_list k lc ++
+ prlist (pr_decl_notation pr_constr) ntn
in
- hov 1 (pr_oneind (if (Decl_kinds.recursivity_flag_of_kind f) then "Inductive" else "CoInductive") (List.hd l))
- ++
+ let key =
+ let (_,_,_,k,_),_ = List.hd l in
+ match k with Record -> "Record" | Structure -> "Structure"
+ | Inductive_kw -> "Inductive" | CoInductive -> "CoInductive"
+ | Class b -> if b then "Definitional Class" else "Class" in
+ hov 1 (pr_oneind key (List.hd l)) ++
(prlist (fun ind -> fnl() ++ hov 1 (pr_oneind "with" ind)) (List.tl l))
| VernacFixpoint (recs,b) ->
- let name_of_binder = function
- | LocalRawAssum (nal,_,_) -> nal
- | LocalRawDef (_,_) -> [] in
let pr_onerec = function
- | ((loc,id),(n,ro),bl,type_,def),ntn ->
- let (bl',def,type_) =
- if Flags.do_beautify() then extract_def_binders def type_
- else ([],def,type_) in
- let bl = bl @ bl' in
- let ids = List.flatten (List.map name_of_binder bl) in
- let annot =
- match n with
- | None -> mt ()
- | Some (loc, id) ->
- match (ro : Topconstr.recursion_order_expr) with
- CStructRec ->
- if List.length ids > 1 then
- spc() ++ str "{struct " ++ pr_id id ++ str"}"
- else mt()
- | CWfRec c ->
- spc() ++ str "{wf " ++ pr_lconstr_expr c ++ spc() ++
- pr_id id ++ str"}"
- | CMeasureRec c ->
- spc() ++ str "{measure " ++ pr_lconstr_expr c ++ spc() ++
- pr_id id ++ str"}"
- in
+ | ((loc,id),ro,bl,type_,def),ntn ->
+ let annot = pr_guard_annot bl ro in
pr_id id ++ pr_binders_arg bl ++ annot ++ spc()
++ pr_type_option (fun c -> spc() ++ pr_lconstr_expr c) type_
- ++ str" :=" ++ brk(1,1) ++ pr_lconstr def ++
- pr_decl_notation pr_constr ntn
+ ++ pr_opt (fun def -> str" :=" ++ brk(1,2) ++ pr_lconstr def) def ++
+ prlist (pr_decl_notation pr_constr) ntn
in
let start = if b then "Boxed Fixpoint" else "Fixpoint" in
- hov 1 (str start ++ spc() ++
- prlist_with_sep (fun _ -> fnl() ++ fnl() ++ str"with ") pr_onerec recs)
+ hov 0 (str start ++ spc() ++
+ prlist_with_sep (fun _ -> fnl() ++ str"with ") pr_onerec recs)
| VernacCoFixpoint (corecs,b) ->
let pr_onecorec (((loc,id),bl,c,def),ntn) =
- let (bl',def,c) =
- if Flags.do_beautify() then extract_def_binders def c
- else ([],def,c) in
- let bl = bl @ bl' in
pr_id id ++ spc() ++ pr_binders bl ++ spc() ++ str":" ++
spc() ++ pr_lconstr_expr c ++
- str" :=" ++ brk(1,1) ++ pr_lconstr def ++
- pr_decl_notation pr_constr ntn
+ pr_opt (fun def -> str" :=" ++ brk(1,2) ++ pr_lconstr def) def ++
+ prlist (pr_decl_notation pr_constr) ntn
in
let start = if b then "Boxed CoFixpoint" else "CoFixpoint" in
- hov 1 (str start ++ spc() ++
- prlist_with_sep (fun _ -> fnl() ++ str"with ") pr_onecorec corecs)
+ hov 0 (str start ++ spc() ++
+ prlist_with_sep (fun _ -> fnl() ++ str"with ") pr_onecorec corecs)
| VernacScheme l ->
hov 2 (str"Scheme" ++ spc() ++
prlist_with_sep (fun _ -> fnl() ++ str"with ") pr_onescheme l)
@@ -668,7 +675,7 @@ let rec pr_vernac = function
hov 2 (str"Combined Scheme" ++ spc() ++
pr_lident id ++ spc() ++ str"from" ++ spc() ++
prlist_with_sep (fun _ -> fnl() ++ str", ") pr_lident l)
-
+
(* Gallina extensions *)
| VernacBeginSection id -> hov 2 (str"Section" ++ spc () ++ pr_lident id)
@@ -684,26 +691,38 @@ let rec pr_vernac = function
| VernacImport (f,l) ->
(if f then str"Export" else str"Import") ++ spc() ++
prlist_with_sep sep pr_import_module l
- | VernacCanonical q -> str"Canonical Structure" ++ spc() ++ pr_reference q
+ | VernacCanonical q -> str"Canonical Structure" ++ spc() ++ pr_smart_global q
| VernacCoercion (s,id,c1,c2) ->
hov 1 (
str"Coercion" ++ (match s with | Local -> spc() ++
str"Local" ++ spc() | Global -> spc()) ++
- pr_reference id ++ spc() ++ str":" ++ spc() ++ pr_class_rawexpr c1 ++
+ pr_smart_global id ++ spc() ++ str":" ++ spc() ++ pr_class_rawexpr c1 ++
spc() ++ str">->" ++ spc() ++ pr_class_rawexpr c2)
| VernacIdentityCoercion (s,id,c1,c2) ->
hov 1 (
str"Identity Coercion" ++ (match s with | Local -> spc() ++
- str"Local" ++ spc() | Global -> spc()) ++ pr_lident id ++
+ str"Local" ++ spc() | Global -> spc()) ++ pr_lident id ++
spc() ++ str":" ++ spc() ++ pr_class_rawexpr c1 ++ spc() ++ str">->" ++
spc() ++ pr_class_rawexpr c2)
-
- | VernacInstance (glob, sup, (instid, bk, cl), props, pri) ->
+
+
+(* | VernacClass (id, par, ar, sup, props) -> *)
+(* hov 1 ( *)
+(* str"Class" ++ spc () ++ pr_lident id ++ *)
+(* (\* prlist_with_sep (spc) (pr_lident_constr (spc() ++ str ":" ++ spc())) par ++ *\) *)
+(* pr_and_type_binders_arg par ++ *)
+(* (match ar with Some ar -> spc () ++ str":" ++ spc() ++ pr_rawsort (snd ar) | None -> mt()) ++ *)
+(* spc () ++ str":=" ++ spc () ++ *)
+(* prlist_with_sep (fun () -> str";" ++ spc()) *)
+(* (fun (lid,oc,c) -> pr_lident_constr ((if oc then str" :>" else str" :") ++ spc()) (lid,c)) props ) *)
+
+ | VernacInstance (abst,glob, sup, (instid, bk, cl), props, pri) ->
hov 1 (
pr_non_locality (not glob) ++
- str"Instance" ++ spc () ++
+ (if abst then str"Declare " else mt ()) ++
+ str"Instance" ++ spc () ++
pr_and_type_binders_arg sup ++
- str"=>" ++ spc () ++
+ str"=>" ++ spc () ++
(match snd instid with Name id -> pr_lident (fst instid, id) ++ spc () ++ str":" ++ spc () | Anonymous -> mt ()) ++
pr_constr_expr cl ++ spc () ++
spc () ++ str":=" ++ spc () ++
@@ -713,37 +732,40 @@ let rec pr_vernac = function
hov 1 (
str"Context" ++ spc () ++ str"[" ++ spc () ++
pr_and_type_binders_arg l ++ spc () ++ str "]")
-
- | VernacDeclareInstance id ->
- hov 1 (str"Instance" ++ spc () ++ pr_lident id)
-
+
+ | VernacDeclareInstance (glob, id) ->
+ hov 1 (pr_non_locality (not glob) ++
+ str"Existing" ++ spc () ++ str"Instance" ++ spc () ++ pr_reference id)
+
+ | VernacDeclareClass id ->
+ hov 1 (str"Existing" ++ spc () ++ str"Class" ++ spc () ++ pr_reference id)
+
(* Modules and Module Types *)
- | VernacDefineModule (export,m,bl,ty,bd) ->
- let b = pr_module_binders_list bl pr_lconstr in
+ | VernacDefineModule (export,m,bl,tys,bd) ->
+ let b = pr_module_binders_list bl pr_lconstr in
hov 2 (str"Module" ++ spc() ++ pr_require_token export ++
pr_lident m ++ b ++
- pr_opt (pr_of_module_type pr_lconstr) ty ++
- pr_opt (fun me -> str ":= " ++ pr_module_expr me) bd)
+ pr_of_module_type pr_lconstr tys ++
+ (if bd = [] then mt () else str ":= ") ++
+ prlist_with_sep (fun () -> str " <+ ")
+ (pr_module_ast_inl pr_lconstr) bd)
| VernacDeclareModule (export,id,bl,m1) ->
- let b = pr_module_binders_list bl pr_lconstr in
+ let b = pr_module_binders_list bl pr_lconstr in
hov 2 (str"Declare Module" ++ spc() ++ pr_require_token export ++
pr_lident id ++ b ++
- pr_of_module_type pr_lconstr m1)
- | VernacDeclareModuleType (id,bl,m) ->
- let b = pr_module_binders_list bl pr_lconstr in
+ pr_module_ast_inl pr_lconstr m1)
+ | VernacDeclareModuleType (id,bl,tyl,m) ->
+ let b = pr_module_binders_list bl pr_lconstr in
+ let pr_mt = pr_module_ast_inl pr_lconstr in
hov 2 (str"Module Type " ++ pr_lident id ++ b ++
- pr_opt (fun mt -> str ":= " ++ pr_module_type pr_lconstr mt) m)
- | VernacInclude (in_ast) ->
- begin
- match in_ast with
- | CIMTE mty ->
- hov 2 (str"Include" ++
- (fun mt -> str " " ++ pr_module_type pr_lconstr mt) mty)
- | CIME mexpr ->
- hov 2 (str"Include" ++
- (fun me -> str " " ++ pr_module_expr me) mexpr)
- end
+ prlist_strict (fun m -> str " <: " ++ pr_mt m) tyl ++
+ (if m = [] then mt () else str ":= ") ++
+ prlist_with_sep (fun () -> str " <+ ") pr_mt m)
+ | VernacInclude (mexprs) ->
+ let pr_m = pr_module_ast_inl pr_lconstr in
+ hov 2 (str"Include " ++
+ prlist_with_sep (fun () -> str " <+ ") pr_m mexprs)
(* Solving *)
| VernacSolve (i,tac,deftac) ->
(if i = 1 then mt() else int i ++ str ": ") ++
@@ -755,12 +777,12 @@ let rec pr_vernac = function
str"Existential " ++ int i ++ pr_lconstrarg c
(* MMode *)
-
+
| VernacProofInstr instr -> anomaly "Not implemented"
- | VernacDeclProof -> str "proof"
+ | VernacDeclProof -> str "proof"
| VernacReturn -> str "return"
- (* /MMode *)
+ (* /MMode *)
(* Auxiliary file and library management *)
| VernacRequireFrom (exp,spe,f) -> hov 2
@@ -774,62 +796,73 @@ let rec pr_vernac = function
(str"Add" ++
(if fl then str" Rec " else spc()) ++
str"LoadPath" ++ spc() ++ qs s ++
- (match d with
+ (match d with
| None -> mt()
- | Some dir -> spc() ++ str"as" ++ spc() ++ pr_dirpath dir))
+ | Some dir -> spc() ++ str"as" ++ spc() ++ pr_dirpath dir))
| VernacRemoveLoadPath s -> str"Remove LoadPath" ++ qs s
| VernacAddMLPath (fl,s) ->
str"Add" ++ (if fl then str" Rec " else spc()) ++ str"ML Path" ++ qs s
- | VernacDeclareMLModule l ->
+ | VernacDeclareMLModule (local, l) ->
+ pr_locality local ++
hov 2 (str"Declare ML Module" ++ spc() ++ prlist_with_sep sep qs l)
| VernacChdir s -> str"Cd" ++ pr_opt qs s
(* Commands *)
- | VernacDeclareTacticDefinition (rc,l) ->
+ | VernacDeclareTacticDefinition (local,rc,l) ->
let pr_tac_body (id, redef, body) =
let idl, body =
match body with
| Tacexpr.TacFun (idl,b) -> idl,b
| _ -> [], body in
- pr_ltac_id id ++
+ pr_ltac_ref id ++
prlist (function None -> str " _"
| Some id -> spc () ++ pr_id id) idl
++ (if redef then str" ::=" else str" :=") ++ brk(1,1) ++
let idl = List.map Option.get (List.filter (fun x -> not (x=None)) idl)in
- pr_raw_tactic_env
- (idl @ List.map coerce_global_to_id
+ pr_raw_tactic_env
+ (idl @ List.map coerce_reference_to_id
(List.map (fun (x, _, _) -> x) (List.filter (fun (_, redef, _) -> not redef) l)))
(Global.env())
body in
hov 1
- ((str "Ltac ") ++
+ (pr_locality local ++ str "Ltac " ++
prlist_with_sep (fun () -> fnl() ++ str"with ") pr_tac_body l)
| VernacCreateHintDb (local,dbname,b) ->
- hov 1 (str "Create " ++ pr_locality local ++ str "HintDb " ++ str dbname ++ (if b then str" discriminated" else mt ()))
+ hov 1 (pr_locality local ++ str "Create " ++ str "HintDb " ++
+ str dbname ++ (if b then str" discriminated" else mt ()))
| VernacHints (local,dbnames,h) ->
pr_hints local dbnames h pr_constr pr_constr_pattern_expr
| VernacSyntacticDefinition (id,(ids,c),local,onlyparsing) ->
hov 2
- (str"Notation " ++ pr_locality local ++ pr_lident id ++
+ (pr_locality local ++ str"Notation " ++ pr_lident id ++
prlist_with_sep spc pr_id ids ++ str" :=" ++ pr_constrarg c ++
pr_syntax_modifiers (if onlyparsing then [SetOnlyParsing] else []))
| VernacDeclareImplicits (local,q,None) ->
- hov 2 (str"Implicit Arguments" ++ spc() ++ pr_reference q)
+ hov 2 (pr_section_locality local ++ str"Implicit Arguments" ++ spc() ++
+ pr_smart_global q)
| VernacDeclareImplicits (local,q,Some imps) ->
- hov 1 (str"Implicit Arguments " ++ pr_non_locality local ++
- spc() ++ pr_reference q ++ spc() ++
+ 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
+ | None -> str "s none"
+ | Some [] -> str "s all"
+ | Some idl ->
+ str (if List.length idl > 1 then "s " else " ") ++
+ prlist_with_sep spc pr_lident idl)
| VernacSetOpacity(b,[k,l]) when k=Conv_oracle.transparent ->
hov 1 (str"Transparent" ++ pr_non_locality b ++
- spc() ++ prlist_with_sep sep pr_reference l)
+ spc() ++ prlist_with_sep sep pr_smart_global l)
| VernacSetOpacity(b,[Conv_oracle.Opaque,l]) ->
hov 1 (str"Opaque" ++ pr_non_locality b ++
- spc() ++ prlist_with_sep sep pr_reference l)
+ spc() ++ prlist_with_sep sep pr_smart_global l)
| VernacSetOpacity (local,l) ->
let pr_lev = function
Conv_oracle.Opaque -> str"opaque"
@@ -838,28 +871,32 @@ let rec pr_vernac = function
| Conv_oracle.Level n -> int n in
let pr_line (l,q) =
hov 2 (pr_lev l ++ spc() ++
- str"[" ++ prlist_with_sep sep pr_reference q ++ str"]") in
- hov 1 (pr_locality local ++ str"Strategy" ++ spc() ++
+ str"[" ++ prlist_with_sep sep pr_smart_global q ++ str"]") in
+ hov 1 (pr_non_locality local ++ str"Strategy" ++ spc() ++
hv 0 (prlist_with_sep sep pr_line l))
- | VernacUnsetOption na ->
- hov 1 (str"Unset" ++ spc() ++ pr_printoption na None)
- | VernacSetOption (na,v) -> hov 2 (str"Set" ++ spc() ++ pr_set_option na v)
+ | VernacUnsetOption (l,na) ->
+ hov 1 (pr_locality_full l ++ str"Unset" ++ spc() ++ pr_printoption na None)
+ | VernacSetOption (l,na,v) ->
+ hov 2 (pr_locality_full l ++ str"Set" ++ spc() ++ pr_set_option na v)
| VernacAddOption (na,l) -> hov 2 (str"Add" ++ spc() ++ pr_printoption na (Some l))
| VernacRemoveOption (na,l) -> hov 2 (str"Remove" ++ spc() ++ pr_printoption na (Some l))
| VernacMemOption (na,l) -> hov 2 (str"Test" ++ spc() ++ pr_printoption na (Some l))
| VernacPrintOption na -> hov 2 (str"Test" ++ spc() ++ pr_printoption na None)
- | VernacCheckMayEval (r,io,c) ->
- let pr_mayeval r c = match r with
+ | VernacCheckMayEval (r,io,c) ->
+ let pr_mayeval r c = match r with
| Some r0 ->
hov 2 (str"Eval" ++ spc() ++
- pr_red_expr (pr_constr,pr_lconstr,pr_or_by_notation pr_reference) r0 ++
+ pr_red_expr (pr_constr,pr_lconstr,pr_smart_global, pr_constr) r0 ++
spc() ++ str"in" ++ spc () ++ pr_constr c)
- | None -> hov 2 (str"Check" ++ spc() ++ pr_constr c)
- in
- (if io = None then mt() else int (Option.get io) ++ str ": ") ++
+ | None -> hov 2 (str"Check" ++ spc() ++ pr_constr c)
+ in
+ (if io = None then mt() else int (Option.get io) ++ str ": ") ++
pr_mayeval r c
| VernacGlobalCheck c -> hov 2 (str"Type" ++ pr_constrarg c)
- | VernacPrint p ->
+ | VernacDeclareReduction (b,s,r) ->
+ pr_locality b ++ str "Declare Reduction " ++ str s ++ str " := " ++
+ pr_red_expr (pr_constr,pr_lconstr,pr_smart_global, pr_constr) r
+ | VernacPrint p ->
let pr_printable = function
| PrintFullContext -> str"Print All"
| PrintSectionContext s ->
@@ -873,54 +910,53 @@ let rec pr_vernac = function
| PrintGraph -> str"Print Graph"
| PrintClasses -> str"Print Classes"
| PrintTypeClasses -> str"Print TypeClasses"
- | PrintInstances qid -> str"Print Instances" ++ spc () ++ pr_reference qid
- | PrintLtac qid -> str"Print Ltac" ++ spc() ++ pr_reference qid
+ | PrintInstances qid -> str"Print Instances" ++ spc () ++ pr_smart_global qid
+ | PrintLtac qid -> str"Print Ltac" ++ spc() ++ pr_ltac_ref qid
| PrintCoercions -> str"Print Coercions"
| PrintCoercionPaths (s,t) -> str"Print Coercion Paths" ++ spc() ++ pr_class_rawexpr s ++ spc() ++ pr_class_rawexpr t
| PrintCanonicalConversions -> str"Print Canonical Structures"
| PrintTables -> str"Print Tables"
- | PrintOpaqueName qid -> str"Print Term" ++ spc() ++ pr_reference qid
| PrintHintGoal -> str"Print Hint"
- | PrintHint qid -> str"Print Hint" ++ spc() ++ pr_reference qid
+ | PrintHint qid -> str"Print Hint" ++ spc() ++ pr_smart_global qid
| PrintHintDb -> str"Print Hint *"
| PrintHintDbName s -> str"Print HintDb" ++ spc() ++ str s
| PrintRewriteHintDbName s -> str"Print Rewrite HintDb" ++ spc() ++ str s
| PrintUniverses fopt -> str"Dump Universes" ++ pr_opt str fopt
- | PrintName qid -> str"Print" ++ spc() ++ pr_reference qid
+ | PrintName qid -> str"Print" ++ spc() ++ pr_smart_global qid
| PrintModuleType qid -> str"Print Module Type" ++ spc() ++ pr_reference qid
| PrintModule qid -> str"Print Module" ++ spc() ++ pr_reference qid
| PrintInspect n -> str"Inspect" ++ spc() ++ int n
| PrintScopes -> str"Print Scopes"
- | PrintScope s -> str"Print Scope" ++ spc() ++ str s
- | PrintVisibility s -> str"Print Visibility" ++ pr_opt str s
- | PrintAbout qid -> str"About" ++ spc() ++ pr_reference qid
- | PrintImplicit qid -> str"Print Implicit" ++ spc() ++ pr_reference qid
-(* spiwack: command printing all the axioms and section variables used in a
+ | PrintScope s -> str"Print Scope" ++ spc() ++ str s
+ | PrintVisibility s -> str"Print Visibility" ++ pr_opt str s
+ | PrintAbout qid -> str"About" ++ spc() ++ pr_smart_global qid
+ | PrintImplicit qid -> str"Print Implicit" ++ spc() ++ pr_smart_global qid
+(* spiwack: command printing all the axioms and section variables used in a
term *)
| PrintAssumptions (b,qid) -> (if b then str"Print Assumptions" else str"Print Opaque Dependencies")
- ++spc()++pr_reference qid
+ ++ spc() ++ pr_smart_global qid
in pr_printable p
| VernacSearch (sea,sea_r) -> pr_search sea sea_r pr_constr_pattern_expr
- | VernacLocate loc ->
+ | VernacLocate loc ->
let pr_locate =function
- | LocateTerm qid -> pr_reference qid
+ | LocateTerm qid -> pr_smart_global qid
| LocateFile f -> str"File" ++ spc() ++ qs f
| LocateLibrary qid -> str"Library" ++ spc () ++ pr_module qid
| LocateModule qid -> str"Module" ++ spc () ++ pr_module qid
- | LocateNotation s -> qs s
+ | LocateTactic qid -> str"Ltac" ++ spc () ++ pr_ltac_ref qid
in str"Locate" ++ spc() ++ pr_locate loc
| VernacComments l ->
hov 2
(str"Comments" ++ spc() ++ prlist_with_sep sep (pr_comment pr_constr) l)
| VernacNop -> mt()
-
+
(* Toplevel control *)
| VernacToplevelControl exn -> pr_topcmd exn
(* For extension *)
| VernacExtend (s,c) -> pr_extend s c
| VernacProof (Tacexpr.TacId _) -> str "Proof"
- | VernacProof te -> str "Proof with" ++ spc() ++ pr_raw_tactic te
+ | VernacProof te -> str "Proof with" ++ spc() ++ pr_raw_tactic te
and pr_extend s cl =
let pr_arg a =
@@ -931,15 +967,15 @@ and pr_extend s cl =
let rl = match_vernac_rule (List.map Genarg.genarg_tag cl) rls in
let start,rl,cl =
match rl with
- | Egrammar.TacTerm s :: rl -> str s, rl, cl
- | Egrammar.TacNonTerm _ :: rl -> pr_arg (List.hd cl), rl, List.tl cl
+ | Egrammar.GramTerminal s :: rl -> str s, rl, cl
+ | Egrammar.GramNonTerminal _ :: rl -> pr_arg (List.hd cl), rl, List.tl cl
| [] -> anomaly "Empty entry" in
let (pp,_) =
List.fold_left
(fun (strm,args) pi ->
let pp,args = match pi with
- | Egrammar.TacNonTerm _ -> (pr_arg (List.hd args), List.tl args)
- | Egrammar.TacTerm s -> (str s, args) in
+ | Egrammar.GramNonTerminal _ -> (pr_arg (List.hd args), List.tl args)
+ | Egrammar.GramTerminal s -> (str s, args) in
(strm ++ spc() ++ pp), args)
(start,cl) rl in
hov 1 pp
diff --git a/parsing/ppvernac.mli b/parsing/ppvernac.mli
index 21d983f5..c24744f3 100644
--- a/parsing/ppvernac.mli
+++ b/parsing/ppvernac.mli
@@ -5,8 +5,8 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-
-(*i $Id: ppvernac.mli 7744 2005-12-27 09:16:06Z herbelin $ i*)
+
+(*i $Id$ i*)
open Pp
open Genarg
diff --git a/parsing/prettyp.ml b/parsing/prettyp.ml
index 1e50bc51..21baeb58 100644
--- a/parsing/prettyp.ml
+++ b/parsing/prettyp.ml
@@ -10,7 +10,7 @@
* on May-June 2006 for implementation of abstraction of pretty-printing of objects.
*)
-(* $Id: prettyp.ml 12187 2009-06-13 19:36:59Z msozeau $ *)
+(* $Id$ *)
open Pp
open Util
@@ -62,20 +62,20 @@ let with_line_skip p = if ismt p then mt() else (fnl () ++ p)
(********************************)
(** Printing implicit arguments *)
-
+
let conjugate_to_be = function [_] -> "is" | _ -> "are"
let pr_implicit imp = pr_id (name_of_implicit imp)
let print_impl_args_by_name max = function
| [] -> mt ()
- | impls ->
- hov 0 (str (plural (List.length impls) "Argument") ++ spc() ++
- prlist_with_sep pr_coma pr_implicit impls ++ spc() ++
+ | impls ->
+ hov 0 (str (plural (List.length impls) "Argument") ++ spc() ++
+ prlist_with_sep pr_comma pr_implicit impls ++ spc() ++
str (conjugate_to_be impls) ++ str" implicit" ++
(if max then strbrk " and maximally inserted" else mt())) ++ fnl()
-let print_impl_args l =
+let print_impl_args l =
let imps = List.filter is_status_implicit l in
let maximps = List.filter Impargs.maximal_insertion_of imps in
let nonmaximps = list_subtract imps maximps in
@@ -87,25 +87,25 @@ let print_impl_args l =
let print_ref reduce ref =
let typ = Global.type_of_global ref in
- let typ =
+ let typ =
if reduce then
let ctx,ccl = Reductionops.splay_prod_assum (Global.env()) Evd.empty typ
- in it_mkProd_or_LetIn ccl ctx
+ in it_mkProd_or_LetIn ccl ctx
else typ in
hov 0 (pr_global ref ++ str " :" ++ spc () ++ pr_ltype typ) ++ fnl ()
let print_argument_scopes = function
| [Some sc] -> str"Argument scope is [" ++ str sc ++ str"]" ++ fnl()
| l when not (List.for_all ((=) None) l) ->
- hov 2 (str"Argument scopes are" ++ spc() ++
- str "[" ++
+ hov 2 (str"Argument scopes are" ++ spc() ++
+ str "[" ++
prlist_with_sep spc (function Some sc -> str sc | None -> str "_") l ++
str "]") ++ fnl()
| _ -> mt()
-let need_expansion impl ref =
+let need_expansion impl ref =
let typ = Global.type_of_global ref in
- let ctx = fst (decompose_prod_assum typ) in
+ let ctx = (prod_assum typ) in
let nprods = List.length (List.filter (fun (_,b,_) -> b=None) ctx) in
impl <> [] & List.length impl >= nprods &
let _,lastimpl = list_chop nprods impl in
@@ -116,7 +116,7 @@ type opacity =
| TransparentMaybeOpacified of Conv_oracle.level
let opacity env = function
- | VarRef v when pi2 (Environ.lookup_named v env) <> None ->
+ | VarRef v when pi2 (Environ.lookup_named v env) <> None ->
Some(TransparentMaybeOpacified (Conv_oracle.get_strategy(VarKey v)))
| ConstRef cst ->
let cb = Environ.lookup_constant cst env in
@@ -129,7 +129,7 @@ let opacity env = function
let print_opacity ref =
match opacity (Global.env()) ref with
| None -> mt ()
- | Some s -> pr_global ref ++ str " is " ++
+ | Some s -> pr_global ref ++ str " is " ++
str (match s with
| FullyOpaque -> "opaque"
| TransparentMaybeOpacified Conv_oracle.Opaque ->
@@ -140,14 +140,14 @@ let print_opacity ref =
"transparent (with expansion weight "^string_of_int n^")"
| TransparentMaybeOpacified Conv_oracle.Expand ->
"transparent (with minimal expansion weight)") ++ fnl()
-
+
let print_name_infos ref =
let impl = implicits_of_global ref in
let scopes = Notation.find_arguments_scope ref in
- let type_for_implicit =
+ let type_for_implicit =
if need_expansion impl ref then
(* Need to reduce since implicits are computed with products flattened *)
- str "Expanded type for implicit arguments" ++ fnl () ++
+ str "Expanded type for implicit arguments" ++ fnl () ++
print_ref true ref ++ fnl()
else mt() in
type_for_implicit ++ print_impl_args impl ++ print_argument_scopes scopes
@@ -155,14 +155,14 @@ let print_name_infos ref =
let print_id_args_data test pr id l =
if List.exists test l then
str"For " ++ pr_id id ++ str": " ++ pr l
- else
+ else
mt()
let print_args_data_of_inductive_ids get test pr sp mipv =
prvecti
- (fun i mip ->
+ (fun i mip ->
print_id_args_data test pr mip.mind_typename (get (IndRef (sp,i))) ++
- prvecti
+ prvecti
(fun j idc ->
print_id_args_data test pr idc (get (ConstructRef ((sp,i),j+1))))
mip.mind_consnames)
@@ -173,7 +173,7 @@ let print_inductive_implicit_args =
implicits_of_global is_status_implicit print_impl_args
let print_inductive_argument_scopes =
- print_args_data_of_inductive_ids
+ print_args_data_of_inductive_ids
Notation.find_arguments_scope ((<>) None) print_argument_scopes
(*********************)
@@ -190,8 +190,8 @@ let locate_any_name ref =
let module N = Nametab in
let (loc,qid) = qualid_of_reference ref in
try Term (N.locate qid)
- with Not_found ->
- try Syntactic (N.locate_syntactic_definition qid)
+ with Not_found ->
+ try Syntactic (N.locate_syndef qid)
with Not_found ->
try Dir (N.locate_dir qid)
with Not_found ->
@@ -205,9 +205,9 @@ let pr_located_qualid = function
| IndRef _ -> "Inductive"
| ConstructRef _ -> "Constructor"
| VarRef _ -> "Variable" in
- str ref_str ++ spc () ++ pr_sp (Nametab.sp_of_global ref)
+ str ref_str ++ spc () ++ pr_path (Nametab.path_of_global ref)
| Syntactic kn ->
- str "Notation" ++ spc () ++ pr_sp (Nametab.sp_of_syntactic_definition kn)
+ str "Notation" ++ spc () ++ pr_path (Nametab.path_of_syndef kn)
| Dir dir ->
let s,dir = match dir with
| DirOpenModule (dir,_) -> "Open Module", dir
@@ -218,8 +218,8 @@ let pr_located_qualid = function
in
str s ++ spc () ++ pr_dirpath dir
| ModuleType (qid,_) ->
- str "Module Type" ++ spc () ++ pr_sp (Nametab.full_name_modtype qid)
- | Undefined qid ->
+ str "Module Type" ++ spc () ++ pr_path (Nametab.full_name_modtype qid)
+ | Undefined qid ->
pr_qualid qid ++ spc () ++ str "not a defined object."
let print_located_qualid ref =
@@ -228,10 +228,10 @@ let print_located_qualid ref =
let expand = function
| TrueGlobal ref ->
Term ref, N.shortest_qualid_of_global Idset.empty ref
- | SyntacticDef kn ->
+ | SynDef kn ->
Syntactic kn, N.shortest_qualid_of_syndef Idset.empty kn in
- match List.map expand (N.extended_locate_all qid) with
- | [] ->
+ match List.map expand (N.locate_extended_all qid) with
+ | [] ->
let (dir,id) = repr_qualid qid in
if dir = empty_dirpath then
str "No object of basename " ++ pr_id id
@@ -291,7 +291,7 @@ let print_constructors envpar names types =
prlist_with_sep (fun () -> brk(1,0) ++ str "| ")
(fun (id,c) -> pr_id id ++ str " : " ++ pr_lconstr_env envpar c)
(Array.to_list (array_map2 (fun n t -> (n,t)) names types))
- in
+ in
hv 0 (str " " ++ pc)
let build_inductive sp tyi =
@@ -300,7 +300,7 @@ let build_inductive sp tyi =
let args = extended_rel_list 0 params in
let env = Global.env() in
let fullarity = match mip.mind_arity with
- | Monomorphic ar -> ar.mind_user_arity
+ | Monomorphic ar -> ar.mind_user_arity
| Polymorphic ar ->
it_mkProd_or_LetIn (mkSort (Type ar.poly_level)) mip.mind_arity_ctxt in
let arity = hnf_prod_applist env fullarity args in
@@ -335,7 +335,7 @@ let get_fields =
let id = match na with Name id -> id | Anonymous -> id_of_string "_" in
prodec_rec ((id,false,substl subst b)::l) (mkVar id::subst) c
| _ -> List.rev l
- in
+ in
prodec_rec [] []
let pr_record (sp,tyi) =
@@ -345,15 +345,15 @@ let pr_record (sp,tyi) =
let fields = get_fields cstrtypes.(0) in
hov 0 (
hov 0 (
- str "Record " ++ pr_global (IndRef (sp,tyi)) ++ brk(1,4) ++
+ str "Record " ++ pr_global (IndRef (sp,tyi)) ++ brk(1,4) ++
print_params env params ++
- str ": " ++ pr_lconstr_env envpar arity ++ brk(1,2) ++
- str ":= " ++ pr_id cstrnames.(0)) ++
+ str ": " ++ pr_lconstr_env envpar arity ++ brk(1,2) ++
+ str ":= " ++ pr_id cstrnames.(0)) ++
brk(1,2) ++
- hv 2 (str "{" ++
- prlist_with_sep (fun () -> str ";" ++ brk(1,0))
- (fun (id,b,c) ->
- str " " ++ pr_id id ++ str (if b then " : " else " := ") ++
+ hv 2 (str "{ " ++
+ prlist_with_sep (fun () -> str ";" ++ brk(2,0))
+ (fun (id,b,c) ->
+ pr_id id ++ str (if b then " : " else " := ") ++
pr_lconstr_env envpar c) fields) ++ str" }")
let gallina_print_inductive sp =
@@ -364,11 +364,11 @@ let gallina_print_inductive sp =
pr_record (List.hd names)
else
pr_mutual_inductive mib.mind_finite names) ++ fnl () ++
- with_line_skip
+ with_line_skip
(print_inductive_implicit_args sp mipv ++
print_inductive_argument_scopes sp mipv)
-let print_named_decl id =
+let print_named_decl id =
gallina_print_named_decl (Global.lookup_named id) ++ fnl ()
let gallina_print_section_variable id =
@@ -391,26 +391,26 @@ let print_constant with_values sep sp =
let val_0 = cb.const_body in
let typ = ungeneralized_type_of_constant_type cb.const_type in
hov 0 (
- match val_0 with
- | None ->
- str"*** [ " ++
- print_basename sp ++ str " : " ++ cut () ++ pr_ltype typ ++
+ match val_0 with
+ | None ->
+ str"*** [ " ++
+ print_basename sp ++ str " : " ++ cut () ++ pr_ltype typ ++
str" ]"
- | _ ->
+ | _ ->
print_basename sp ++ str sep ++ cut () ++
(if with_values then print_typed_body (val_0,typ) else pr_ltype typ))
++ fnl ()
let gallina_print_constant_with_infos sp =
- print_constant true " = " sp ++
+ print_constant true " = " sp ++
with_line_skip (print_name_infos (ConstRef sp))
let gallina_print_syntactic_def kn =
let sep = " := "
and qid = Nametab.shortest_qualid_of_syndef Idset.empty kn
- and (vars,a) = Syntax_def.search_syntactic_definition dummy_loc kn in
+ and (vars,a) = Syntax_def.search_syntactic_definition kn in
let c = Topconstr.rawconstr_of_aconstr dummy_loc a in
- str "Notation " ++ pr_qualid qid ++
+ str "Notation " ++ pr_qualid qid ++
prlist_with_sep spc pr_id (List.map fst vars) ++ str sep ++
Constrextern.without_symbols pr_lrawconstr c ++ fnl ()
@@ -419,42 +419,42 @@ let gallina_print_leaf_entry with_values ((sp,kn as oname),lobj) =
and tag = object_tag lobj in
match (oname,tag) with
| (_,"VARIABLE") ->
- (* Outside sections, VARIABLES still exist but only with universes
+ (* Outside sections, VARIABLES still exist but only with universes
constraints *)
(try Some(print_named_decl (basename sp)) with Not_found -> None)
| (_,"CONSTANT") ->
Some (print_constant with_values sep (constant_of_kn kn))
| (_,"INDUCTIVE") ->
- Some (gallina_print_inductive kn)
+ Some (gallina_print_inductive (mind_of_kn kn))
| (_,"MODULE") ->
- let (mp,_,l) = repr_kn kn in
+ let (mp,_,l) = repr_kn kn in
Some (print_module with_values (MPdot (mp,l)))
| (_,"MODULE TYPE") ->
- let (mp,_,l) = repr_kn kn in
+ let (mp,_,l) = repr_kn kn in
Some (print_modtype (MPdot (mp,l)))
| (_,("AUTOHINT"|"GRAMMAR"|"SYNTAXCONSTANT"|"PPSYNTAX"|"TOKEN"|"CLASS"|
"COERCION"|"REQUIRE"|"END-SECTION"|"STRUCTURE")) -> None
(* To deal with forgotten cases... *)
| (_,s) -> None
-let gallina_print_library_entry with_values ent =
+let gallina_print_library_entry with_values ent =
let pr_name (sp,_) = pr_id (basename sp) in
match ent with
- | (oname,Lib.Leaf lobj) ->
+ | (oname,Lib.Leaf lobj) ->
gallina_print_leaf_entry with_values (oname,lobj)
- | (oname,Lib.OpenedSection (dir,_)) ->
+ | (oname,Lib.OpenedSection (dir,_)) ->
Some (str " >>>>>>> Section " ++ pr_name oname)
- | (oname,Lib.ClosedSection _) ->
+ | (oname,Lib.ClosedSection _) ->
Some (str " >>>>>>> Closed Section " ++ pr_name oname)
| (_,Lib.CompilingLibrary (dir,_)) ->
Some (str " >>>>>>> Library " ++ pr_dirpath dir)
| (oname,Lib.OpenedModule _) ->
Some (str " >>>>>>> Module " ++ pr_name oname)
- | (oname,Lib.ClosedModule _) ->
+ | (oname,Lib.ClosedModule _) ->
Some (str " >>>>>>> Closed Module " ++ pr_name oname)
| (oname,Lib.OpenedModtype _) ->
Some (str " >>>>>>> Module Type " ++ pr_name oname)
- | (oname,Lib.ClosedModtype _) ->
+ | (oname,Lib.ClosedModtype _) ->
Some (str " >>>>>>> Closed Module Type " ++ pr_name oname)
| (_,Lib.FrozenState _) ->
None
@@ -464,14 +464,14 @@ let gallina_print_leaf_entry with_values c =
| None -> mt ()
| Some pp -> pp ++ fnl()
-let gallina_print_context with_values =
+let gallina_print_context with_values =
let rec prec n = function
- | h::rest when n = None or Option.get n > 0 ->
+ | h::rest when n = None or Option.get n > 0 ->
(match gallina_print_library_entry with_values h with
| None -> prec n rest
| Some pp -> prec (Option.map ((+) (-1)) n) rest ++ pp ++ fnl ())
| _ -> mt ()
- in
+ in
prec
let gallina_print_eval red_fun env evmap _ {uj_val=trm;uj_type=typ} =
@@ -520,16 +520,16 @@ let print_typed_value x = print_typed_value_in_env (Global.env ()) x
let print_judgment env {uj_val=trm;uj_type=typ} =
print_typed_value_in_env env (trm, typ)
-
+
let print_safe_judgment env j =
let trm = Safe_typing.j_val j in
let typ = Safe_typing.j_type j in
print_typed_value_in_env env (trm, typ)
-
+
(*********************)
(* *)
-let print_full_context () =
+let print_full_context () =
print_context true None (Lib.contents_after None)
let print_full_context_typ () =
@@ -540,33 +540,34 @@ let print_full_pure_context () =
| ((_,kn),Lib.Leaf lobj)::rest ->
let pp = match object_tag lobj with
| "CONSTANT" ->
- let con = constant_of_kn kn in
+ let con = Global.constant_of_delta (constant_of_kn kn) in
let cb = Global.lookup_constant con in
let val_0 = cb.const_body in
let typ = ungeneralized_type_of_constant_type cb.const_type in
hov 0 (
- match val_0 with
+ match val_0 with
| None ->
str (if cb.const_opaque then "Axiom " else "Parameter ") ++
print_basename con ++ str " : " ++ cut () ++ pr_ltype typ
| Some v ->
if cb.const_opaque then
- str "Theorem " ++ print_basename con ++ cut () ++
+ str "Theorem " ++ print_basename con ++ cut () ++
str " : " ++ pr_ltype typ ++ str "." ++ fnl () ++
str "Proof " ++ print_body val_0
else
- str "Definition " ++ print_basename con ++ cut () ++
+ str "Definition " ++ print_basename con ++ cut () ++
str " : " ++ pr_ltype typ ++ cut () ++ str " := " ++
print_body val_0) ++ str "." ++ fnl () ++ fnl ()
| "INDUCTIVE" ->
- let (mib,mip) = Global.lookup_inductive (kn,0) in
+ let mind = Global.mind_of_delta (mind_of_kn kn) in
+ let (mib,mip) = Global.lookup_inductive (mind,0) in
let mipv = mib.mind_packets in
- let names = list_tabulate (fun x -> (kn,x)) (Array.length mipv) in
- pr_mutual_inductive mib.mind_finite names ++ str "." ++
+ let names = list_tabulate (fun x -> (mind,x)) (Array.length mipv) in
+ pr_mutual_inductive mib.mind_finite names ++ str "." ++
fnl () ++ fnl ()
| "MODULE" ->
(* TODO: make it reparsable *)
- let (mp,_,l) = repr_kn kn in
+ let (mp,_,l) = repr_kn kn in
print_module true (MPdot (mp,l)) ++ str "." ++ fnl () ++ fnl ()
| "MODULE TYPE" ->
(* TODO: make it reparsable *)
@@ -576,7 +577,7 @@ let print_full_pure_context () =
| _ -> mt () in
prec rest ++ pp
| _::rest -> prec rest
- | _ -> mt () in
+ | _ -> mt () in
prec (Lib.contents_after None)
(* For printing an inductive definition with
@@ -584,14 +585,14 @@ let print_full_pure_context () =
assume that the declaration of constructors and eliminations
follows the definition of the inductive type *)
-let list_filter_vec f vec =
- let rec frec n lf =
- if n < 0 then lf
- else if f vec.(n) then
+let list_filter_vec f vec =
+ let rec frec n lf =
+ if n < 0 then lf
+ else if f vec.(n) then
frec (n-1) (vec.(n)::lf)
- else
+ else
frec (n-1) lf
- in
+ in
frec (Array.length vec -1) []
(* This is designed to print the contents of an opened section *)
@@ -608,19 +609,18 @@ let read_sec_context r =
error "Cannot print the contents of a closed section."
(* LEM: Actually, we could if we wanted to. *)
| [] -> []
- | hd::rest -> get_cxt (hd::in_cxt) rest
+ | hd::rest -> get_cxt (hd::in_cxt) rest
in
let cxt = (Lib.contents_after None) in
List.rev (get_cxt [] cxt)
-let print_sec_context sec =
+let print_sec_context sec =
print_context true None (read_sec_context sec)
let print_sec_context_typ sec =
print_context false None (read_sec_context sec)
-let print_name ref =
- match locate_any_name ref with
+let print_any_name = function
| Term (ConstRef sp) -> print_constant_with_infos sp
| Term (IndRef (sp,_)) -> print_inductive sp
| Term (ConstructRef ((sp,_),_)) -> print_inductive sp
@@ -631,49 +631,44 @@ let print_name ref =
| ModuleType (_,kn) -> print_modtype kn
| Undefined qid ->
try (* Var locale de but, pas var de section... donc pas d'implicits *)
- let dir,str = repr_qualid qid in
+ let dir,str = repr_qualid qid in
if (repr_dirpath dir) <> [] then raise Not_found;
- let (_,c,typ) = Global.lookup_named str in
+ let (_,c,typ) = Global.lookup_named str in
(print_named_decl (str,c,typ))
with Not_found ->
- try
- let sp = Nametab.locate_obj qid in
- let (oname,lobj) =
- let (oname,entry) =
- List.find (fun en -> (fst (fst en)) = sp) (Lib.contents_after None)
- in
- match entry with
- | Lib.Leaf obj -> (oname,obj)
- | _ -> raise Not_found
- in
- print_leaf_entry true (oname,lobj)
- with Not_found ->
errorlabstrm
"print_name" (pr_qualid qid ++ spc () ++ str "not a defined object.")
-let print_opaque_name qid =
+let print_name = function
+ | Genarg.ByNotation (loc,ntn,sc) ->
+ print_any_name
+ (Term (Notation.interp_notation_as_global_reference loc (fun _ -> true)
+ ntn sc))
+ | Genarg.AN ref ->
+ print_any_name (locate_any_name ref)
+
+let print_opaque_name qid =
let env = Global.env () in
match global qid with
| ConstRef cst ->
let cb = Global.lookup_constant cst in
if cb.const_body <> None then
print_constant_with_infos cst
- else
+ else
error "Not a defined constant."
| IndRef (sp,_) ->
print_inductive sp
- | ConstructRef cstr ->
+ | ConstructRef cstr ->
let ty = Inductiveops.type_of_constructor env cstr in
print_typed_value (mkConstruct cstr, ty)
| VarRef id ->
- let (_,c,ty) = lookup_named id env in
+ let (_,c,ty) = lookup_named id env in
print_named_decl (id,c,ty)
-let print_about ref =
- let k = locate_any_name ref in
+let print_about_any k =
begin match k with
| Term ref ->
- print_ref false ref ++ fnl () ++ print_name_infos ref ++
+ print_ref false ref ++ fnl () ++ print_name_infos ref ++
print_opacity ref
| Syntactic kn ->
print_syntactic_def kn
@@ -682,26 +677,34 @@ let print_about ref =
++
hov 0 (str "Expands to: " ++ pr_located_qualid k)
+let print_about = function
+ | Genarg.ByNotation (loc,ntn,sc) ->
+ print_about_any
+ (Term (Notation.interp_notation_as_global_reference loc (fun _ -> true)
+ ntn sc))
+ | Genarg.AN ref ->
+ print_about_any (locate_any_name ref)
+
let print_impargs ref =
- let ref = Nametab.global ref in
+ let ref = Smartlocate.smart_global ref in
let impl = implicits_of_global ref in
let has_impl = List.filter is_status_implicit impl <> [] in
(* Need to reduce since implicits are computed with products flattened *)
print_ref (need_expansion impl ref) ref ++ fnl() ++
- (if has_impl then print_impl_args impl
+ (if has_impl then print_impl_args impl
else (str "No implicit arguments" ++ fnl ()))
-let unfold_head_fconst =
+let unfold_head_fconst =
let rec unfrec k = match kind_of_term k with
- | Const cst -> constant_value (Global.env ()) cst
+ | Const cst -> constant_value (Global.env ()) cst
| Lambda (na,t,b) -> mkLambda (na,t,unfrec b)
| App (f,v) -> appvect (unfrec f,v)
| _ -> k
- in
+ in
unfrec
(* for debug *)
-let inspect depth =
+let inspect depth =
print_context false (Some depth) (Lib.contents_after None)
@@ -715,8 +718,8 @@ let print_coercion_value v = pr_lconstr (get_coercion_value v)
let print_class i =
let cl,_ = class_info_from_index i in
pr_class cl
-
-let print_path ((i,j),p) =
+
+let print_path ((i,j),p) =
hov 2 (
str"[" ++ hov 0 (prlist_with_sep pr_semicolon print_coercion_value p) ++
str"] : ") ++
@@ -724,45 +727,39 @@ let print_path ((i,j),p) =
let _ = Classops.install_path_printer print_path
-let print_graph () =
+let print_graph () =
prlist_with_sep pr_fnl print_path (inheritance_graph())
-let print_classes () =
+let print_classes () =
prlist_with_sep pr_spc pr_class (classes())
-let print_coercions () =
+let print_coercions () =
prlist_with_sep pr_spc print_coercion_value (coercions())
-
+
let index_of_class cl =
- try
+ try
fst (class_info cl)
- with _ ->
+ with _ ->
errorlabstrm "index_of_class"
(pr_class cl ++ spc() ++ str "not a defined class.")
-let print_path_between cls clt =
+let print_path_between cls clt =
let i = index_of_class cls in
let j = index_of_class clt in
- let p =
- try
- lookup_path_between_class (i,j)
- with _ ->
+ let p =
+ try
+ lookup_path_between_class (i,j)
+ with _ ->
errorlabstrm "index_cl_of_id"
(str"No path between " ++ pr_class cls ++ str" and " ++ pr_class clt
++ str ".")
in
print_path ((i,j),p)
-let pr_cs_pattern = function
- Const_cs c -> pr_global c
- | Prod_cs -> str "_ -> _"
- | Default_cs -> str "_"
- | Sort_cs s -> pr_sort_family s
-
let print_canonical_projections () =
- prlist_with_sep pr_fnl
- (fun ((r1,r2),o) -> pr_cs_pattern r2 ++
- str " <- " ++
+ prlist_with_sep pr_fnl
+ (fun ((r1,r2),o) -> pr_cs_pattern r2 ++
+ str " <- " ++
pr_global r1 ++ str " ( " ++ pr_lconstr o.o_DEF ++ str " )")
(canonical_projections ())
@@ -773,25 +770,25 @@ let print_canonical_projections () =
open Typeclasses
-let pr_typeclass env t =
+let pr_typeclass env t =
print_ref false t.cl_impl
let print_typeclasses () =
let env = Global.env () in
prlist_with_sep fnl (pr_typeclass env) (typeclasses ())
-
-let pr_instance env i =
+
+let pr_instance env i =
(* gallina_print_constant_with_infos i.is_impl *)
(* lighter *)
- print_ref false (ConstRef (instance_impl i))
-
+ print_ref false (instance_impl i)
+
let print_all_instances () =
let env = Global.env () in
- let inst = all_instances () in
+ let inst = all_instances () in
prlist_with_sep fnl (pr_instance env) inst
let print_instances r =
let env = Global.env () in
- let inst = instances r in
+ let inst = instances r in
prlist_with_sep fnl (pr_instance env) inst
-
+
diff --git a/parsing/prettyp.mli b/parsing/prettyp.mli
index ec2228c7..ba1977e8 100644
--- a/parsing/prettyp.mli
+++ b/parsing/prettyp.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: prettyp.mli 11576 2008-11-10 19:13:15Z msozeau $ i*)
+(*i $Id$ i*)
(*i*)
open Pp
@@ -18,6 +18,7 @@ open Environ
open Reductionops
open Libnames
open Nametab
+open Genarg
(*i*)
(* A Pretty-Printer for the Calculus of Inductive Constructions. *)
@@ -40,10 +41,10 @@ val print_eval :
(* This function is exported for the graphical user-interface pcoq *)
val build_inductive : mutual_inductive -> int ->
global_reference * rel_context * types * identifier array * types array
-val print_name : reference -> std_ppcmds
+val print_name : reference or_by_notation -> std_ppcmds
val print_opaque_name : reference -> std_ppcmds
-val print_about : reference -> std_ppcmds
-val print_impargs : reference -> std_ppcmds
+val print_about : reference or_by_notation -> std_ppcmds
+val print_impargs : reference or_by_notation -> std_ppcmds
(*i
val print_extracted_name : identifier -> std_ppcmds
diff --git a/parsing/printer.ml b/parsing/printer.ml
index 0c673fbd..7f5087a8 100644
--- a/parsing/printer.ml
+++ b/parsing/printer.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: printer.ml 12187 2009-06-13 19:36:59Z msozeau $ *)
+(* $Id$ *)
open Pp
open Util
@@ -20,7 +20,6 @@ open Global
open Declare
open Libnames
open Nametab
-open Ppconstr
open Evd
open Proof_type
open Decl_mode
@@ -30,11 +29,11 @@ open Ppconstr
open Constrextern
open Tacexpr
-let emacs_str s alts =
+let emacs_str s alts =
match !Flags.print_emacs, !Flags.print_emacs_safechar with
| true, true -> alts
| true , false -> s
- | false,_ -> ""
+ | false,_ -> ""
(**********************************************************************)
(** Terms *)
@@ -59,6 +58,19 @@ let pr_constr t = pr_constr_env (Global.env()) t
let pr_open_lconstr (_,c) = pr_lconstr c
let pr_open_constr (_,c) = pr_constr c
+let pr_constr_under_binders_env_gen pr env (ids,c) =
+ (* Warning: clashes can occur with variables of same name in env but *)
+ (* we also need to preserve the actual names of the patterns *)
+ (* So what to do? *)
+ let assums = List.map (fun id -> (Name id,(* dummy *) mkProp)) ids in
+ pr (push_rels_assum assums env) c
+
+let pr_constr_under_binders_env = pr_constr_under_binders_env_gen pr_constr_env
+let pr_lconstr_under_binders_env = pr_constr_under_binders_env_gen pr_lconstr_env
+
+let pr_constr_under_binders c = pr_constr_under_binders_env (Global.env()) c
+let pr_lconstr_under_binders c = pr_lconstr_under_binders_env (Global.env()) c
+
let pr_type_core at_top env t =
pr_constr_expr (extern_type at_top env t)
let pr_ltype_core at_top env t =
@@ -78,7 +90,7 @@ let pr_ljudge j = pr_ljudge_env (Global.env()) j
let pr_lrawconstr_env env c =
pr_lconstr_expr (extern_rawconstr (vars_of_env env) c)
-let pr_rawconstr_env env c =
+let pr_rawconstr_env env c =
pr_constr_expr (extern_rawconstr (vars_of_env env) c)
let pr_lrawconstr c =
@@ -115,10 +127,7 @@ let pr_inductive env ind = pr_lconstr_env env (mkInd ind)
let pr_constructor env cstr = pr_lconstr_env env (mkConstruct cstr)
let pr_evaluable_reference ref =
- let ref' = match ref with
- | EvalConstRef const -> ConstRef const
- | EvalVarRef sp -> VarRef sp in
- pr_global ref'
+ pr_global (Tacred.global_of_evaluable_reference ref)
(*let pr_rawterm t =
pr_lconstr (Constrextern.extern_rawconstr Idset.empty t)*)
@@ -134,7 +143,7 @@ let pr_var_decl env (id,c,typ) =
let pbody = match c with
| None -> (mt ())
| Some c ->
- (* Force evaluation *)
+ (* Force evaluation *)
let pb = pr_lconstr_env env c in
let pb = if isCast c then surround pb else pb in
(str" := " ++ pb ++ cut () ) in
@@ -146,7 +155,7 @@ let pr_rel_decl env (na,c,typ) =
let pbody = match c with
| None -> mt ()
| Some c ->
- (* Force evaluation *)
+ (* Force evaluation *)
let pb = pr_lconstr_env env c in
let pb = if isCast c then surround pb else pb in
(str":=" ++ spc () ++ pb ++ spc ()) in
@@ -166,7 +175,7 @@ let pr_named_context_of env =
let psl = List.rev (fold_named_context make_decl_list env ~init:[]) in
hv 0 (prlist_with_sep (fun _ -> ws 2) (fun x -> x) psl)
-let pr_named_context env ne_context =
+let pr_named_context env ne_context =
hv 0 (Sign.fold_named_context
(fun d pps -> pps ++ ws 2 ++ pr_var_decl env d)
ne_context ~init:(mt ()))
@@ -183,14 +192,14 @@ let pr_context_unlimited env =
fold_named_context
(fun env d pps ->
let pidt = pr_var_decl env d in (pps ++ fnl () ++ pidt))
- env ~init:(mt ())
+ env ~init:(mt ())
in
let db_env =
fold_rel_context
(fun env d pps ->
let pnat = pr_rel_decl env d in (pps ++ fnl () ++ pnat))
env ~init:(mt ())
- in
+ in
(sign_env ++ db_env)
let pr_ne_context_of header env =
@@ -201,21 +210,21 @@ let pr_ne_context_of header env =
let pr_context_limit n env =
let named_context = Environ.named_context env in
let lgsign = List.length named_context in
- if n >= lgsign then
+ if n >= lgsign then
pr_context_unlimited env
else
let k = lgsign-n in
let _,sign_env =
fold_named_context
(fun env d (i,pps) ->
- if i < k then
+ if i < k then
(i+1, (pps ++str "."))
else
let pidt = pr_var_decl env d in
(i+1, (pps ++ fnl () ++
str (emacs_str (String.make 1 (Char.chr 253)) "") ++
pidt)))
- env ~init:(0,(mt ()))
+ env ~init:(0,(mt ()))
in
let db_env =
fold_rel_context
@@ -225,10 +234,10 @@ let pr_context_limit n env =
str (emacs_str (String.make 1 (Char.chr 253)) "") ++
pnat))
env ~init:(mt ())
- in
+ in
(sign_env ++ db_env)
-let pr_context_of env = match Flags.print_hyps_limit () with
+let pr_context_of env = match Flags.print_hyps_limit () with
| None -> hv 0 (pr_context_unlimited env)
| Some n -> hv 0 (pr_context_limit n env)
@@ -238,33 +247,33 @@ let pr_restricted_named_context among env =
hv 0 (fold_named_context
(fun env ((id,_,_) as d) pps ->
if true || Idset.mem id among then
- pps ++
+ pps ++
fnl () ++ str (emacs_str (String.make 1 (Char.chr 253)) "") ++
pr_var_decl env d
- else
+ else
pps)
env ~init:(mt ()))
-let pr_predicate pr_elt (b, elts) =
+let pr_predicate pr_elt (b, elts) =
let pr_elts = prlist_with_sep spc pr_elt elts in
if b then
- str"all" ++
+ str"all" ++
(if elts = [] then mt () else str" except: " ++ pr_elts)
else
if elts = [] then str"none" else pr_elts
-
+
let pr_cpred p = pr_predicate pr_con (Cpred.elements p)
let pr_idpred p = pr_predicate Nameops.pr_id (Idpred.elements p)
-let pr_transparent_state (ids, csts) =
+let pr_transparent_state (ids, csts) =
hv 0 (str"VARIABLES: " ++ pr_idpred ids ++ fnl () ++
str"CONSTANTS: " ++ pr_cpred csts ++ fnl ())
let pr_subgoal_metas metas env=
- let pr_one (meta,typ) =
- str "?" ++ int meta ++ str " : " ++
- hov 0 (pr_ltype_env_at_top env typ) ++ fnl () ++
+ let pr_one (meta,typ) =
+ str "?" ++ int meta ++ str " : " ++
+ hov 0 (pr_ltype_env_at_top env typ) ++ fnl () ++
str (emacs_str (String.make 1 (Char.chr 253)) "") in
hv 0 (prlist_with_sep mt pr_one metas)
@@ -276,9 +285,9 @@ let default_pr_goal g =
mt (), mt (),
pr_context_of env,
pr_ltype_env_at_top env g.evar_concl
- else
+ else
(str " *** Declarative Mode ***" ++ fnl ()++fnl ()),
- (str"thesis := " ++ fnl ()),
+ (str "thesis := " ++ fnl ()),
pr_context_of env,
pr_ltype_env_at_top env g.evar_concl
in
@@ -287,7 +296,7 @@ let default_pr_goal g =
str (emacs_str (String.make 1 (Char.chr 253)) "") ++
str "============================" ++ fnl () ++
thesis ++ str " " ++ pc) ++ fnl ()
-
+
(* display the conclusion of a goal *)
let pr_concl n g =
let env = evar_env g in
@@ -296,13 +305,13 @@ let pr_concl n g =
str "subgoal " ++ int n ++ str " is:" ++ cut () ++ str" " ++ pc
(* display evar type: a context and a type *)
-let pr_evgl_sign gl =
+let pr_evgl_sign gl =
let ps = pr_named_context_of (evar_unfiltered_env gl) in
let _,l = list_filter2 (fun b c -> not b) (evar_filter gl,evar_context gl) in
let ids = List.rev (List.map pi1 l) in
let warn =
if ids = [] then mt () else
- (str "(" ++ prlist_with_sep pr_coma pr_id ids ++ str " cannot be used)")
+ (str "(" ++ prlist_with_sep pr_comma pr_id ids ++ str " cannot be used)")
in
let pc = pr_lconstr gl.evar_concl in
hov 0 (str"[" ++ ps ++ spc () ++ str"|- " ++ pc ++ str"]" ++ spc () ++ warn)
@@ -311,10 +320,10 @@ let pr_evgl_sign gl =
let rec pr_evars_int i = function
| [] -> (mt ())
| (ev,evd)::rest ->
- let pegl = pr_evgl_sign evd in
+ let pegl = pr_evgl_sign evd in
let pei = pr_evars_int (i+1) rest in
(hov 0 (str "Existential " ++ int i ++ str " =" ++ spc () ++
- str (string_of_existential ev) ++ str " : " ++ pegl)) ++
+ str (string_of_existential ev) ++ str " : " ++ pegl)) ++
fnl () ++ pei
let default_pr_subgoal n =
@@ -324,27 +333,27 @@ let default_pr_subgoal n =
if p = 1 then
let pg = default_pr_goal g in
v 0 (str "subgoal " ++ int n ++ str " is:" ++ cut () ++ pg)
- else
+ else
prrec (p-1) rest
- in
+ in
prrec n
(* Print open subgoals. Checks for uninstantiated existential variables *)
-let default_pr_subgoals close_cmd sigma = function
- | [] ->
+let default_pr_subgoals close_cmd sigma = function
+ | [] ->
begin
match close_cmd with
Some cmd ->
- (str "Subproof completed, now type " ++ str cmd ++
+ (str "Subproof completed, now type " ++ str cmd ++
str "." ++ fnl ())
| None ->
- let exl = Evarutil.non_instantiated sigma in
- if exl = [] then
- (str"Proof completed." ++ fnl ())
+ let exl = Evarutil.non_instantiated sigma in
+ if exl = [] then
+ (str"Proof completed." ++ fnl ())
else
let pei = pr_evars_int 1 exl in
(str "No more subgoals but non-instantiated existential " ++
- str "variables :" ++fnl () ++ (hov 0 pei))
+ str "variables:" ++ fnl () ++ (hov 0 pei))
end
| [g] ->
let pg = default_pr_goal g in
@@ -355,11 +364,11 @@ let default_pr_subgoals close_cmd sigma = function
| g::rest ->
let pc = pr_concl n g in
let prest = pr_rec (n+1) rest in
- (cut () ++ pc ++ prest)
+ (cut () ++ pc ++ prest)
in
let pg1 = default_pr_goal g1 in
let prest = pr_rec 2 rest in
- v 0 (int(List.length rest+1) ++ str" subgoals" ++ cut ()
+ v 0 (int(List.length rest+1) ++ str" subgoals" ++ cut ()
++ pg1 ++ prest ++ fnl ())
@@ -390,24 +399,19 @@ let pr_goal x = !printer_pr.pr_goal x
(* End abstraction layer *)
(**********************************************************************)
-let pr_subgoals_of_pfts pfts =
- let close_cmd = Decl_mode.get_end_command pfts in
- let gls = fst (Refiner.frontier (proof_of_pftreestate pfts)) in
- let sigma = (top_goal_of_pftreestate pfts).sigma in
- pr_subgoals close_cmd sigma gls
-
let pr_open_subgoals () =
let pfts = get_pftreestate () in
+ let gls = fst (frontier (proof_of_pftreestate pfts)) in
match focus() with
- | 0 ->
- pr_subgoals_of_pfts pfts
- | n ->
- let pf = proof_of_pftreestate pfts in
- let gls = fst (frontier pf) in
+ | 0 ->
+ let sigma = (top_goal_of_pftreestate pfts).sigma in
+ let close_cmd = Decl_mode.get_end_command pfts in
+ pr_subgoals close_cmd sigma gls
+ | n ->
assert (n > List.length gls);
- if List.length gls < 2 then
+ if List.length gls < 2 then
pr_subgoal n gls
- else
+ else
(* LEM TODO: this way of saying how many subgoals has to be abstracted out*)
v 0 (int(List.length gls) ++ str" subgoals" ++ cut () ++
pr_subgoal n gls)
@@ -419,25 +423,25 @@ let pr_nth_open_subgoal n =
(* Elementary tactics *)
let pr_prim_rule = function
- | Intro id ->
+ | Intro id ->
str"intro " ++ pr_id id
-
+
| Cut (b,replace,id,t) ->
if b then
(* TODO: express "replace" *)
(str"assert " ++ str"(" ++ pr_id id ++ str":" ++ pr_lconstr t ++ str")")
else
let cl = if replace then str"clear " ++ pr_id id ++ str"; " else mt() in
- (str"cut " ++ pr_constr t ++
+ (str"cut " ++ pr_constr t ++
str ";[" ++ cl ++ str"intro " ++ pr_id id ++ str"|idtac]")
-
+
| FixRule (f,n,[],_) ->
(str"fix " ++ pr_id f ++ str"/" ++ int n)
-
- | FixRule (f,n,others,j) ->
+
+ | FixRule (f,n,others,j) ->
if j<>0 then warning "Unsupported printing of \"fix\"";
let rec print_mut = function
- | (f,n,ar)::oth ->
+ | (f,n,ar)::oth ->
pr_id f ++ str"/" ++ int n ++ str" : " ++ pr_lconstr ar ++ print_mut oth
| [] -> mt () in
(str"fix " ++ pr_id f ++ str"/" ++ int n ++
@@ -453,26 +457,26 @@ let pr_prim_rule = function
(pr_id f ++ str" : " ++ pr_lconstr ar ++ print_mut oth)
| [] -> mt () in
(str"cofix " ++ pr_id f ++ str" with " ++ print_mut others)
- | Refine c ->
+ | Refine c ->
str(if occur_meta c then "refine " else "exact ") ++
Constrextern.with_meta_as_hole pr_constr c
-
+
| Convert_concl (c,_) ->
(str"change " ++ pr_constr c)
-
+
| Convert_hyp (id,None,t) ->
(str"change " ++ pr_constr t ++ spc () ++ str"in " ++ pr_id id)
| Convert_hyp (id,Some c,t) ->
(str"change " ++ pr_constr c ++ spc () ++ str"in "
++ pr_id id ++ str" (type of " ++ pr_id id ++ str ")")
-
+
| Thin ids ->
(str"clear " ++ prlist_with_sep pr_spc pr_id ids)
-
+
| ThinBody ids ->
(str"clearbody " ++ prlist_with_sep pr_spc pr_id ids)
-
+
| Move (withdep,id1,id2) ->
(str (if withdep then "dependent " else "") ++
str"move " ++ pr_id id1 ++ pr_move_location pr_id id2)
@@ -497,7 +501,7 @@ let prterm = pr_lconstr
(* spiwack: printer function for sets of Environ.assumption.
It is used primarily by the Print Assumption command. *)
-let pr_assumptionset env s =
+let pr_assumptionset env s =
if (Environ.ContextObjectMap.is_empty s) then
str "Closed under the global context"
else
@@ -506,7 +510,7 @@ let pr_assumptionset env s =
let (v,a,o) = r in
match t with
| Variable id -> ( Some (
- Option.default (fnl ()) v
+ Option.default (fnl ()) v
++ str (string_of_id id)
++ str " : "
++ pr_ltype typ
@@ -536,7 +540,7 @@ let pr_assumptionset env s =
)
s (None,None,None)
in
- let (vars,axioms,opaque) =
+ let (vars,axioms,opaque) =
( Option.map (fun p -> str "Section Variables:" ++ p) vars ,
Option.map (fun p -> str "Axioms:" ++ p) axioms ,
Option.map (fun p -> str "Opaque constants:" ++ p) opaque
@@ -550,7 +554,7 @@ let cmap_to_list m = Cmap.fold (fun k v acc -> v :: acc) m []
open Typeclasses
let pr_instance i =
- pr_global (ConstRef (instance_impl i))
+ pr_global (instance_impl i)
let pr_instance_gmap insts =
prlist_with_sep fnl (fun (gr, insts) ->
diff --git a/parsing/printer.mli b/parsing/printer.mli
index 40bb9122..bfb7dfe1 100644
--- a/parsing/printer.mli
+++ b/parsing/printer.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: printer.mli 11309 2008-08-06 10:30:35Z herbelin $ i*)
+(*i $Id$ i*)
(*i*)
open Pp
@@ -42,6 +42,12 @@ val pr_open_constr : open_constr -> std_ppcmds
val pr_open_lconstr_env : env -> open_constr -> std_ppcmds
val pr_open_lconstr : open_constr -> std_ppcmds
+val pr_constr_under_binders_env : env -> constr_under_binders -> std_ppcmds
+val pr_constr_under_binders : constr_under_binders -> std_ppcmds
+
+val pr_lconstr_under_binders_env : env -> constr_under_binders -> std_ppcmds
+val pr_lconstr_under_binders : constr_under_binders -> std_ppcmds
+
val pr_ltype_env_at_top : env -> types -> std_ppcmds
val pr_ltype_env : env -> types -> std_ppcmds
val pr_ltype : types -> std_ppcmds
@@ -112,8 +118,8 @@ val pr_evars_int : int -> (evar * evar_info) list -> std_ppcmds
val pr_prim_rule : prim_rule -> std_ppcmds
(* Emacs/proof general support *)
-(* (emacs_str s alts) outputs
- - s if emacs mode & unicode allowed,
+(* (emacs_str s alts) outputs
+ - s if emacs mode & unicode allowed,
- alts if emacs mode and & unicode not allowed
- nothing otherwise *)
val emacs_str : string -> string -> string
diff --git a/parsing/printmod.ml b/parsing/printmod.ml
index 596ce6b2..eb6e88c9 100644
--- a/parsing/printmod.ml
+++ b/parsing/printmod.ml
@@ -13,13 +13,13 @@ open Declarations
open Nameops
open Libnames
-let get_new_id locals id =
- let rec get_id l id =
+let get_new_id locals id =
+ let rec get_id l id =
let dir = make_dirpath [id] in
if not (Nametab.exists_module dir) then
id
else
- get_id (id::l) (Nameops.next_ident_away id l)
+ get_id (id::l) (Namegen.next_ident_away id l)
in
get_id (List.map snd locals) id
@@ -27,21 +27,21 @@ let rec print_local_modpath locals = function
| MPbound mbid -> pr_id (List.assoc mbid locals)
| MPdot(mp,l) ->
print_local_modpath locals mp ++ str "." ++ pr_lab l
- | MPself _ | MPfile _ -> raise Not_found
+ | MPfile _ -> raise Not_found
-let print_modpath locals mp =
+let print_modpath locals mp =
try (* must be with let because streams are lazy! *)
- let qid = Nametab.shortest_qualid_of_module mp in
+ let qid = Nametab.shortest_qualid_of_module mp in
pr_qualid qid
with
| Not_found -> print_local_modpath locals mp
-let print_kn locals kn =
+let print_kn locals kn =
try
- let qid = Nametab.shortest_qualid_of_modtype kn in
+ let qid = Nametab.shortest_qualid_of_modtype kn in
pr_qualid qid
with
- Not_found ->
+ Not_found ->
try
print_local_modpath locals kn
with
@@ -52,108 +52,107 @@ let rec flatten_app mexpr l = match mexpr with
| mexpr -> mexpr::l
let rec print_module name locals with_body mb =
- let body = match with_body, mb.mod_expr with
- | false, _
+ let body = match with_body, mb.mod_expr with
+ | false, _
| true, None -> mt()
- | true, Some mexpr ->
+ | true, Some mexpr ->
spc () ++ str ":= " ++ print_modexpr locals mexpr
in
- let modtype = match mb.mod_type with
- None -> str ""
- | Some t -> spc () ++ str": " ++
+
+ let modtype =
+ match mb.mod_type with
+ | t -> spc () ++ str": " ++
print_modtype locals t
in
hv 2 (str "Module " ++ name ++ modtype ++ body)
-and print_modtype locals mty =
+and print_modtype locals mty =
match mty with
| SEBident kn -> print_kn locals kn
| SEBfunctor (mbid,mtb1,mtb2) ->
- (* let env' = Modops.add_module (MPbid mbid)
- (Modops.body_of_type mtb) env
- in *)
+ (* let env' = Modops.add_module (MPbid mbid)
+ (Modops.body_of_type mtb) env
+ in *)
let locals' = (mbid, get_new_id locals (id_of_mbid mbid))
::locals in
- hov 2 (str "Funsig" ++ spc () ++ str "(" ++
- pr_id (id_of_mbid mbid) ++ str " : " ++
- print_modtype locals mtb1.typ_expr ++
+ hov 2 (str "Funsig" ++ spc () ++ str "(" ++
+ pr_id (id_of_mbid mbid) ++ str " : " ++
+ print_modtype locals mtb1.typ_expr ++
str ")" ++ spc() ++ print_modtype locals' mtb2)
- | SEBstruct (msid,sign) ->
- hv 2 (str "Sig" ++ spc () ++ print_sig locals msid sign ++ brk (1,-2) ++ str "End")
- | SEBapply (mexpr,marg,_) ->
+ | SEBstruct (sign) ->
+ hv 2 (str "Sig" ++ spc () ++ print_sig locals sign ++ brk (1,-2) ++ str "End")
+ | SEBapply (mexpr,marg,_) ->
let lapp = flatten_app mexpr [marg] in
let fapp = List.hd lapp in
let mapp = List.tl lapp in
- hov 3 (str"(" ++ (print_modtype locals fapp) ++ spc () ++
+ hov 3 (str"(" ++ (print_modtype locals fapp) ++ spc () ++
prlist_with_sep spc (print_modexpr locals) mapp ++ str")")
| SEBwith(seb,With_definition_body(idl,cb))->
let s = (String.concat "." (List.map string_of_id idl)) in
- hov 2 (print_modtype locals seb ++ spc() ++ str "with" ++ spc() ++
+ hov 2 (print_modtype locals seb ++ spc() ++ str "with" ++ spc() ++
str "Definition"++ spc() ++ str s ++ spc() ++ str ":="++ spc())
- | SEBwith(seb,With_module_body(idl,mp,_,_))->
+ | SEBwith(seb,With_module_body(idl,mp))->
let s =(String.concat "." (List.map string_of_id idl)) in
- hov 2 (print_modtype locals seb ++ spc() ++ str "with" ++ spc() ++
+ hov 2 (print_modtype locals seb ++ spc() ++ str "with" ++ spc() ++
str "Module"++ spc() ++ str s ++ spc() ++ str ":="++ spc())
-and print_sig locals msid sign =
+and print_sig locals sign =
let print_spec (l,spec) = (match spec with
| SFBconst {const_body=Some _; const_opaque=false} -> str "Definition "
| SFBconst {const_body=None}
| SFBconst {const_opaque=true} -> str "Parameter "
| SFBmind _ -> str "Inductive "
| SFBmodule _ -> str "Module "
- | SFBalias (mp,_,_) -> str "Module "
| SFBmodtype _ -> str "Module Type ") ++ str (string_of_label l)
in
prlist_with_sep spc print_spec sign
-and print_struct locals msid struc =
+and print_struct locals struc =
let print_body (l,body) = (match body with
| SFBconst {const_body=Some _; const_opaque=false} -> str "Definition "
| SFBconst {const_body=Some _; const_opaque=true} -> str "Theorem "
| SFBconst {const_body=None} -> str "Parameter "
| SFBmind _ -> str "Inductive "
| SFBmodule _ -> str "Module "
- | SFBalias (mp,_,_) -> str "Module "
| SFBmodtype _ -> str "Module Type ") ++ str (string_of_label l)
in
prlist_with_sep spc print_body struc
-and print_modexpr locals mexpr = match mexpr with
+and print_modexpr locals mexpr = match mexpr with
| SEBident mp -> print_modpath locals mp
| SEBfunctor (mbid,mty,mexpr) ->
-(* let env' = Modops.add_module (MPbid mbid) (Modops.body_of_type mtb) env
+(* let env' = Modops.add_module (MPbid mbid) (Modops.body_of_type mtb) env
in *)
let locals' = (mbid, get_new_id locals (id_of_mbid mbid))::locals in
- hov 2 (str "Functor" ++ spc() ++ str"(" ++ pr_id(id_of_mbid mbid) ++
- str ":" ++ print_modtype locals mty.typ_expr ++
+ hov 2 (str "Functor" ++ spc() ++ str"(" ++ pr_id(id_of_mbid mbid) ++
+ str ":" ++ print_modtype locals mty.typ_expr ++
str ")" ++ spc () ++ print_modexpr locals' mexpr)
- | SEBstruct (msid, struc) ->
- hv 2 (str "Struct" ++ spc () ++ print_struct locals msid struc ++ brk (1,-2) ++ str "End")
- | SEBapply (mexpr,marg,_) ->
+ | SEBstruct ( struc) ->
+ hv 2 (str "Struct" ++ spc () ++ print_struct locals struc ++ brk (1,-2) ++ str "End")
+ | SEBapply (mexpr,marg,_) ->
let lapp = flatten_app mexpr [marg] in
hov 3 (str"(" ++ prlist_with_sep spc (print_modexpr locals) lapp ++ str")")
| SEBwith (_,_)-> anomaly "Not avaible yet"
-let rec printable_body dir =
- let dir = dirpath_prefix dir in
- dir = empty_dirpath ||
- try
+let rec printable_body dir =
+ let dir = pop_dirpath dir in
+ dir = empty_dirpath ||
+ try
match Nametab.locate_dir (qualid_of_dirpath dir) with
DirOpenModtype _ -> false
| DirModule _ | DirOpenModule _ -> printable_body dir
| _ -> true
- with
+ with
Not_found -> true
-let print_module with_body mp =
+let print_module with_body mp =
let name = print_modpath [] mp in
print_module name [] with_body (Global.lookup_module mp) ++ fnl ()
-let print_modtype kn =
+let print_modtype kn =
let mtb = Global.lookup_modtype kn in
let name = print_kn [] kn in
- str "Module Type " ++ name ++ str " = " ++
+ str "Module Type " ++ name ++ str " = " ++
print_modtype [] mtb.typ_expr ++ fnl ()
diff --git a/parsing/q_constr.ml4 b/parsing/q_constr.ml4
index 37817389..093910b4 100644
--- a/parsing/q_constr.ml4
+++ b/parsing/q_constr.ml4
@@ -8,7 +8,7 @@
(*i camlp4use: "pa_extend.cmo q_MLast.cmo" i*)
-(* $Id: q_constr.ml4 11576 2008-11-10 19:13:15Z msozeau $ *)
+(* $Id$ *)
open Rawterm
open Term
@@ -21,8 +21,8 @@ open Pcaml
let loc = dummy_loc
let dloc = <:expr< Util.dummy_loc >>
-let apply_ref f l =
- <:expr<
+let apply_ref f l =
+ <:expr<
Rawterm.RApp ($dloc$, Rawterm.RRef ($dloc$, Lazy.force $f$), $mlexpr_of_list (fun x -> x) l$)
>>
@@ -57,13 +57,13 @@ EXTEND
(* fix todo *)
]
| "100" RIGHTA
- [ c1 = constr; ":"; c2 = SELF ->
+ [ c1 = constr; ":"; c2 = SELF ->
<:expr< Rawterm.RCast($dloc$,$c1$,DEFAULTcast,$c2$) >> ]
| "90" RIGHTA
- [ c1 = constr; "->"; c2 = SELF ->
+ [ c1 = constr; "->"; c2 = SELF ->
<:expr< Rawterm.RProd ($dloc$,Anonymous,Rawterm.Explicit,$c1$,$c2$) >> ]
| "75" RIGHTA
- [ "~"; c = constr ->
+ [ "~"; c = constr ->
apply_ref <:expr< coq_not_ref >> [c] ]
| "70" RIGHTA
[ c1 = constr; "="; c2 = NEXT; ":>"; t = NEXT ->
@@ -85,26 +85,26 @@ EXTEND
;
match_constr:
[ [ "match"; c = constr LEVEL "100"; (ty,nal) = match_type;
- "with"; OPT"|"; br = LIST0 eqn SEP "|"; "end" ->
+ "with"; OPT"|"; br = LIST0 eqn SEP "|"; "end" ->
let br = mlexpr_of_list (fun x -> x) br in
- <:expr< Rawterm.RCases ($dloc$,$ty$,[($c$,$nal$)],$br$) >>
+ <:expr< Rawterm.RCases ($dloc$,$ty$,[($c$,$nal$)],$br$) >>
] ]
;
match_type:
- [ [ "as"; id = ident; "in"; ind = LIDENT; nal = LIST0 name;
- "return"; ty = constr LEVEL "100" ->
+ [ [ "as"; id = ident; "in"; ind = LIDENT; nal = LIST0 name;
+ "return"; ty = constr LEVEL "100" ->
let nal = mlexpr_of_list (fun x -> x) nal in
- <:expr< Some $ty$ >>,
- <:expr< (Name $id$, Some ($dloc$,$lid:ind$,$nal$)) >>
+ <:expr< Some $ty$ >>,
+ <:expr< (Name $id$, Some ($dloc$,$lid:ind$,$nal$)) >>
| -> <:expr< None >>, <:expr< (Anonymous, None) >> ] ]
;
eqn:
- [ [ (lid,pl) = pattern; "=>"; rhs = constr ->
+ [ [ (lid,pl) = pattern; "=>"; rhs = constr ->
let lid = mlexpr_of_list (fun x -> x) lid in
- <:expr< ($dloc$,$lid$,[$pl$],$rhs$) >>
+ <:expr< ($dloc$,$lid$,[$pl$],$rhs$) >>
] ]
;
- pattern:
+ pattern:
[ [ "%"; e = string; lip = LIST0 patvar ->
let lp = mlexpr_of_list (fun (_,x) -> x) lip in
let lid = List.flatten (List.map fst lip) in
@@ -113,13 +113,13 @@ EXTEND
| "("; p = pattern; ")" -> p ] ]
;
patvar:
- [ [ "_" -> [], <:expr< Rawterm.PatVar ($dloc$,Anonymous) >>
- | id = ident -> [id], <:expr< Rawterm.PatVar ($dloc$,Name $id$) >>
+ [ [ "_" -> [], <:expr< Rawterm.PatVar ($dloc$,Anonymous) >>
+ | id = ident -> [id], <:expr< Rawterm.PatVar ($dloc$,Name $id$) >>
] ]
;
END;;
-(* Example
+(* Example
open Coqlib
let a = PATTERN [ match ?X with %path_of_S n => n | %path_of_O => ?X end ]
*)
diff --git a/parsing/q_coqast.ml4 b/parsing/q_coqast.ml4
index aeee632c..1bd5af53 100644
--- a/parsing/q_coqast.ml4
+++ b/parsing/q_coqast.ml4
@@ -8,7 +8,7 @@
(*i camlp4use: "q_MLast.cmo pa_macro.cmo" i*)
-(* $Id: q_coqast.ml4 11735 2009-01-02 17:22:31Z herbelin $ *)
+(* $Id$ *)
open Util
open Names
@@ -28,11 +28,11 @@ IFDEF CAMLP5 THEN DEFINE NOP END
let anti loc x =
let e =
let loc =
- IFDEF NOP THEN
+ IFDEF NOP THEN
loc
- ELSE
+ ELSE
(1, snd loc - fst loc)
- END
+ END
in <:expr< $lid:purge_str x$ >>
in
<:expr< $anti:e$ >>
@@ -47,7 +47,7 @@ let mlexpr_of_ident id =
let mlexpr_of_name = function
| Names.Anonymous -> <:expr< Names.Anonymous >>
- | Names.Name id ->
+ | Names.Name id ->
<:expr< Names.Name (Names.id_of_string $str:Names.string_of_id id$) >>
let mlexpr_of_dirpath dir =
@@ -68,13 +68,14 @@ let mlexpr_of_loc loc = <:expr< $dloc$ >>
let mlexpr_of_by_notation f = function
| Genarg.AN x -> <:expr< Genarg.AN $f x$ >>
- | Genarg.ByNotation (loc,s,sco) ->
+ | Genarg.ByNotation (loc,s,sco) ->
<:expr< Genarg.ByNotation $dloc$ $str:s$ $mlexpr_of_option mlexpr_of_string sco$ >>
let mlexpr_of_intro_pattern = function
| Genarg.IntroWildcard -> <:expr< Genarg.IntroWildcard >>
| Genarg.IntroAnonymous -> <:expr< Genarg.IntroAnonymous >>
| Genarg.IntroFresh id -> <:expr< Genarg.IntroFresh (mlexpr_of_ident $dloc$ id) >>
+ | Genarg.IntroForthcoming b -> <:expr< Genarg.IntroForthcoming (mlexpr_of_bool $dloc$ b) >>
| Genarg.IntroIdentifier id ->
<:expr< Genarg.IntroIdentifier (mlexpr_of_ident $dloc$ id) >>
| Genarg.IntroOrAndPattern _ | Genarg.IntroRewrite _ ->
@@ -133,14 +134,14 @@ let mlexpr_of_red_flags {
let mlexpr_of_explicitation = function
| Topconstr.ExplByName id -> <:expr< Topconstr.ExplByName $mlexpr_of_ident id$ >>
| Topconstr.ExplByPos (n,_id) -> <:expr< Topconstr.ExplByPos $mlexpr_of_int n$ >>
-
+
let mlexpr_of_binding_kind = function
| Rawterm.Implicit -> <:expr< Rawterm.Implicit >>
| Rawterm.Explicit -> <:expr< Rawterm.Explicit >>
let mlexpr_of_binder_kind = function
| Topconstr.Default b -> <:expr< Topconstr.Default $mlexpr_of_binding_kind b$ >>
- | Topconstr.Generalized (b,b',b'') ->
+ | Topconstr.Generalized (b,b',b'') ->
<:expr< Topconstr.TypeClass $mlexpr_of_binding_kind b$
$mlexpr_of_binding_kind b'$ $mlexpr_of_bool b''$ >>
@@ -152,7 +153,7 @@ let rec mlexpr_of_constr = function
| 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
+ | Topconstr.CProdN (loc,l,a) -> <:expr< Topconstr.CProdN $dloc$ $mlexpr_of_list
(mlexpr_of_triple (mlexpr_of_list (mlexpr_of_pair (fun _ -> dloc) mlexpr_of_name)) mlexpr_of_binder_kind mlexpr_of_constr) l$ $mlexpr_of_constr a$ >>
| Topconstr.CLambdaN (loc,l,a) -> <:expr< Topconstr.CLambdaN $dloc$ $mlexpr_of_list (mlexpr_of_triple (mlexpr_of_list (mlexpr_of_pair (fun _ -> dloc) mlexpr_of_name)) mlexpr_of_binder_kind mlexpr_of_constr) l$ $mlexpr_of_constr a$ >>
| Topconstr.CLetIn (loc,_,_,_) -> failwith "mlexpr_of_constr: TODO"
@@ -163,10 +164,10 @@ let rec mlexpr_of_constr = function
| Topconstr.CHole (loc, Some _) -> failwith "mlexpr_of_constr: TODO CHole (Some _)"
| Topconstr.CNotation(_,ntn,subst) ->
<:expr< Topconstr.CNotation $dloc$ $mlexpr_of_string ntn$
- $mlexpr_of_pair
+ $mlexpr_of_pair
(mlexpr_of_list mlexpr_of_constr)
(mlexpr_of_list (mlexpr_of_list mlexpr_of_constr)) subst$ >>
- | Topconstr.CPatVar (loc,n) ->
+ | Topconstr.CPatVar (loc,n) ->
<:expr< Topconstr.CPatVar $dloc$ $mlexpr_of_pair mlexpr_of_bool mlexpr_of_ident n$ >>
| _ -> failwith "mlexpr_of_constr: TODO"
@@ -215,7 +216,7 @@ let rec mlexpr_of_argtype loc = function
| Genarg.List0ArgType t -> <:expr< Genarg.List0ArgType $mlexpr_of_argtype loc t$ >>
| Genarg.List1ArgType t -> <:expr< Genarg.List1ArgType $mlexpr_of_argtype loc t$ >>
| Genarg.OptArgType t -> <:expr< Genarg.OptArgType $mlexpr_of_argtype loc t$ >>
- | Genarg.PairArgType (t1,t2) ->
+ | Genarg.PairArgType (t1,t2) ->
let t1 = mlexpr_of_argtype loc t1 in
let t2 = mlexpr_of_argtype loc t2 in
<:expr< Genarg.PairArgType $t1$ $t2$ >>
@@ -236,10 +237,10 @@ let mlexpr_of_binding_kind = function
| Rawterm.ExplicitBindings l ->
let l = mlexpr_of_list (mlexpr_of_triple mlexpr_of_loc mlexpr_of_quantified_hypothesis mlexpr_of_constr) l in
<:expr< Rawterm.ExplicitBindings $l$ >>
- | Rawterm.ImplicitBindings l ->
+ | Rawterm.ImplicitBindings l ->
let l = mlexpr_of_list mlexpr_of_constr l in
<:expr< Rawterm.ImplicitBindings $l$ >>
- | Rawterm.NoBindings ->
+ | Rawterm.NoBindings ->
<:expr< Rawterm.NoBindings >>
let mlexpr_of_binding = mlexpr_of_pair mlexpr_of_binding_kind mlexpr_of_constr
@@ -255,7 +256,7 @@ let mlexpr_of_move_location f = function
let mlexpr_of_induction_arg = function
| Tacexpr.ElimOnConstr c ->
<:expr< Tacexpr.ElimOnConstr $mlexpr_of_constr_with_binding c$ >>
- | Tacexpr.ElimOnIdent (_,id) ->
+ | Tacexpr.ElimOnIdent (_,id) ->
<:expr< Tacexpr.ElimOnIdent $dloc$ $mlexpr_of_ident id$ >>
| Tacexpr.ElimOnAnonHyp n ->
<:expr< Tacexpr.ElimOnAnonHyp $mlexpr_of_int n$ >>
@@ -346,11 +347,11 @@ let rec mlexpr_of_atomic_tactic = function
<:expr< Tacexpr.TacCut $mlexpr_of_constr c$ >>
| Tacexpr.TacAssert (t,ipat,c) ->
let ipat = mlexpr_of_option (mlexpr_of_located mlexpr_of_intro_pattern) ipat in
- <:expr< Tacexpr.TacAssert $mlexpr_of_option mlexpr_of_tactic t$ $ipat$
+ <:expr< Tacexpr.TacAssert $mlexpr_of_option mlexpr_of_tactic t$ $ipat$
$mlexpr_of_constr c$ >>
| Tacexpr.TacGeneralize cl ->
<:expr< Tacexpr.TacGeneralize
- $mlexpr_of_list
+ $mlexpr_of_list
(mlexpr_of_pair mlexpr_of_occ_constr mlexpr_of_name) cl$ >>
| Tacexpr.TacGeneralizeDep c ->
<:expr< Tacexpr.TacGeneralizeDep $mlexpr_of_constr c$ >>
@@ -365,14 +366,14 @@ let rec mlexpr_of_atomic_tactic = function
<:expr< Tacexpr.TacSimpleInductionDestruct $mlexpr_of_bool isrec$
$mlexpr_of_quantified_hypothesis h$ >>
| Tacexpr.TacInductionDestruct (isrec,ev,l) ->
- <:expr< Tacexpr.TacInductionDestruct $mlexpr_of_bool isrec$ $mlexpr_of_bool ev$
- $mlexpr_of_list (mlexpr_of_quadruple
+ <:expr< Tacexpr.TacInductionDestruct $mlexpr_of_bool isrec$ $mlexpr_of_bool ev$
+ $mlexpr_of_pair (mlexpr_of_list (mlexpr_of_triple
(mlexpr_of_list mlexpr_of_induction_arg)
(mlexpr_of_option mlexpr_of_constr_with_binding)
(mlexpr_of_pair
(mlexpr_of_option (mlexpr_of_located mlexpr_of_intro_pattern))
- (mlexpr_of_option (mlexpr_of_located mlexpr_of_intro_pattern)))
- (mlexpr_of_option mlexpr_of_clause)) l$ >>
+ (mlexpr_of_option (mlexpr_of_located mlexpr_of_intro_pattern)))))
+ (mlexpr_of_option mlexpr_of_clause) l$ >>
(* Context management *)
| Tacexpr.TacClear (b,l) ->
@@ -393,7 +394,7 @@ let rec mlexpr_of_atomic_tactic = function
<:expr< Tacexpr.TacRight $mlexpr_of_bool ev$ $mlexpr_of_binding_kind l$>>
| Tacexpr.TacSplit (ev,b,l) ->
<:expr< Tacexpr.TacSplit
- ($mlexpr_of_bool ev$,$mlexpr_of_bool b$,$mlexpr_of_binding_kind l$)>>
+ ($mlexpr_of_bool ev$,$mlexpr_of_bool b$,$mlexpr_of_list mlexpr_of_binding_kind l$)>>
| Tacexpr.TacAnyConstructor (ev,t) ->
<:expr< Tacexpr.TacAnyConstructor $mlexpr_of_bool ev$ $mlexpr_of_option mlexpr_of_tactic t$>>
| Tacexpr.TacConstructor (ev,n,l) ->
@@ -404,15 +405,15 @@ let rec mlexpr_of_atomic_tactic = function
| Tacexpr.TacReduce (r,cl) ->
let l = mlexpr_of_clause cl in
<:expr< Tacexpr.TacReduce $mlexpr_of_red_expr r$ $l$ >>
- | Tacexpr.TacChange (occl,c,cl) ->
+ | Tacexpr.TacChange (p,c,cl) ->
let l = mlexpr_of_clause cl in
- let g = mlexpr_of_option mlexpr_of_occ_constr in
- <:expr< Tacexpr.TacChange $g occl$ $mlexpr_of_constr c$ $l$ >>
+ let g = mlexpr_of_option mlexpr_of_constr in
+ <:expr< Tacexpr.TacChange $g p$ $mlexpr_of_constr c$ $l$ >>
(* Equivalence relations *)
| Tacexpr.TacReflexivity -> <:expr< Tacexpr.TacReflexivity >>
| Tacexpr.TacSymmetry ido -> <:expr< Tacexpr.TacSymmetry $mlexpr_of_clause ido$ >>
- | Tacexpr.TacTransitivity c -> <:expr< Tacexpr.TacTransitivity $mlexpr_of_constr c$ >>
+ | Tacexpr.TacTransitivity c -> <:expr< Tacexpr.TacTransitivity $mlexpr_of_option mlexpr_of_constr c$ >>
(* Automation tactics *)
| Tacexpr.TacAuto (n,lems,l) ->
@@ -436,7 +437,7 @@ let rec mlexpr_of_atomic_tactic = function
and mlexpr_of_tactic : (Tacexpr.raw_tactic_expr -> MLast.expr) = function
| Tacexpr.TacAtom (loc,t) ->
<:expr< Tacexpr.TacAtom $dloc$ $mlexpr_of_atomic_tactic t$ >>
- | Tacexpr.TacThen (t1,[||],t2,[||]) ->
+ | Tacexpr.TacThen (t1,[||],t2,[||]) ->
<:expr< Tacexpr.TacThen $mlexpr_of_tactic t1$ [||] $mlexpr_of_tactic t2$ [||]>>
| Tacexpr.TacThens (t,tl) ->
<:expr< Tacexpr.TacThens $mlexpr_of_tactic t$ $mlexpr_of_list mlexpr_of_tactic tl$>>
@@ -454,7 +455,7 @@ and mlexpr_of_tactic : (Tacexpr.raw_tactic_expr -> MLast.expr) = function
<:expr< Tacexpr.TacRepeat $mlexpr_of_tactic t$ >>
| Tacexpr.TacProgress t ->
<:expr< Tacexpr.TacProgress $mlexpr_of_tactic t$ >>
- | Tacexpr.TacId l ->
+ | Tacexpr.TacId l ->
<:expr< Tacexpr.TacId $mlexpr_of_list mlexpr_of_message_token l$ >>
| Tacexpr.TacFail (n,l) ->
<:expr< Tacexpr.TacFail $mlexpr_of_or_var mlexpr_of_int n$ $mlexpr_of_list mlexpr_of_message_token l$ >>
@@ -476,7 +477,7 @@ and mlexpr_of_tactic : (Tacexpr.raw_tactic_expr -> MLast.expr) = function
$mlexpr_of_tactic t$
$mlexpr_of_list (mlexpr_of_match_rule mlexpr_of_tactic) l$>>
| Tacexpr.TacMatchGoal (lz,lr,l) ->
- <:expr< Tacexpr.TacMatchGoal
+ <:expr< Tacexpr.TacMatchGoal
$mlexpr_of_bool lz$
$mlexpr_of_bool lr$
$mlexpr_of_list (mlexpr_of_match_rule mlexpr_of_tactic) l$>>
@@ -494,7 +495,7 @@ and mlexpr_of_tactic : (Tacexpr.raw_tactic_expr -> MLast.expr) = function
and mlexpr_of_tactic_arg = function
| Tacexpr.MetaIdArg (loc,true,id) -> anti loc id
- | Tacexpr.MetaIdArg (loc,false,id) ->
+ | Tacexpr.MetaIdArg (loc,false,id) ->
<:expr< Tacexpr.ConstrMayEval (Rawterm.ConstrTerm $anti loc id$) >>
| Tacexpr.TacCall (loc,t,tl) ->
<:expr< Tacexpr.TacCall $dloc$ $mlexpr_of_reference t$ $mlexpr_of_list mlexpr_of_tactic_arg tl$>>
@@ -522,7 +523,7 @@ let ftac e =
let ep s = patt_of_expr (ee s) in
Quotation.ExAst (ee, ep)
-let _ =
+let _ =
Quotation.add "constr" (fconstr Pcoq.Constr.constr_eoi);
Quotation.add "tactic" (ftac Pcoq.Tactic.tactic_eoi);
Quotation.default := "constr"
diff --git a/parsing/q_util.ml4 b/parsing/q_util.ml4
index da4329bb..7b9037d9 100644
--- a/parsing/q_util.ml4
+++ b/parsing/q_util.ml4
@@ -8,17 +8,19 @@
(*i camlp4use: "q_MLast.cmo" i*)
-(* $Id: q_util.ml4 11309 2008-08-06 10:30:35Z herbelin $ *)
+(* $Id$ *)
(* This file defines standard combinators to build ml expressions *)
open Util
+open Extrawit
+open Pcoq
let not_impl name x =
let desc =
if Obj.is_block (Obj.repr x) then
"tag = " ^ string_of_int (Obj.tag (Obj.repr x))
- else
+ else
"int_val = " ^ string_of_int (Obj.magic x)
in
failwith ("<Q_util." ^ name ^ ", not impl: " ^ desc)
@@ -78,64 +80,16 @@ open Vernacexpr
open Pcoq
open Genarg
-let modifiers e =
-<:expr< Gramext.srules
- [([], Gramext.action(fun _loc -> []));
- ([Gramext.Stoken ("", "(");
- Gramext.Slist1sep ($e$, Gramext.Stoken ("", ","));
- Gramext.Stoken ("", ")")],
- Gramext.action (fun _ l _ _loc -> l))]
- >>
-
-let rec interp_entry_name loc s sep =
- let l = String.length s in
- if l > 8 & String.sub s 0 3 = "ne_" & String.sub s (l-5) 5 = "_list" then
- let t, g = interp_entry_name loc (String.sub s 3 (l-8)) "" in
- List1ArgType t, <:expr< Gramext.Slist1 $g$ >>
- else if l > 12 & String.sub s 0 3 = "ne_" &
- String.sub s (l-9) 9 = "_list_sep" then
- let t, g = interp_entry_name loc (String.sub s 3 (l-12)) "" in
- let sep = <:expr< Gramext.Stoken("",$str:sep$) >> in
- List1ArgType t, <:expr< Gramext.Slist1sep $g$ $sep$ >>
- else if l > 5 & String.sub s (l-5) 5 = "_list" then
- let t, g = interp_entry_name loc (String.sub s 0 (l-5)) "" in
- List0ArgType t, <:expr< Gramext.Slist0 $g$ >>
- else if l > 9 & String.sub s (l-9) 9 = "_list_sep" then
- let t, g = interp_entry_name loc (String.sub s 0 (l-9)) "" in
- let sep = <:expr< Gramext.Stoken("",$str:sep$) >> in
- List0ArgType t, <:expr< Gramext.Slist0sep $g$ $sep$ >>
- else if l > 4 & String.sub s (l-4) 4 = "_opt" then
- let t, g = interp_entry_name loc (String.sub s 0 (l-4)) "" in
- OptArgType t, <:expr< Gramext.Sopt $g$ >>
- else if l > 5 & String.sub s (l-5) 5 = "_mods" then
- let t, g = interp_entry_name loc (String.sub s 0 (l-1)) "" in
- List0ArgType t, modifiers g
- else
- let s = if s = "hyp" then "var" else s in
- let t, se, lev =
- match tactic_genarg_level s with
- | Some 5 ->
- Some (ExtraArgType s), <:expr< Tactic. binder_tactic >>, None
- | Some n ->
- Some (ExtraArgType s), <:expr< Tactic. tactic_expr >>, Some n
- | None ->
- match Pcoq.entry_type (Pcoq.get_univ "prim") s with
- | Some _ as x -> x, <:expr< Prim. $lid:s$ >>, None
- | None ->
- match Pcoq.entry_type (Pcoq.get_univ "constr") s with
- | Some _ as x -> x, <:expr< Constr. $lid:s$ >>, None
- | None ->
- match Pcoq.entry_type (Pcoq.get_univ "tactic") s with
- | Some _ as x -> x, <:expr< Tactic. $lid:s$ >>, None
- | None -> None, <:expr< $lid:s$ >>, None in
- let t =
- match t with
- | Some t -> t
- | None -> ExtraArgType s in
- let entry = match lev with
- | Some n ->
- let s = string_of_int n in
- <:expr< Gramext.Snterml (Pcoq.Gram.Entry.obj $se$, $str:s$) >>
- | None ->
- <:expr< Gramext.Snterm (Pcoq.Gram.Entry.obj $se$) >>
- in t, entry
+let rec mlexpr_of_prod_entry_key = function
+ | Extend.Alist1 s -> <:expr< Extend.Alist1 $mlexpr_of_prod_entry_key s$ >>
+ | Extend.Alist1sep (s,sep) -> <:expr< Extend.Alist1sep $mlexpr_of_prod_entry_key s$ $str:sep$ >>
+ | Extend.Alist0 s -> <:expr< Extend.Alist0 $mlexpr_of_prod_entry_key s$ >>
+ | Extend.Alist0sep (s,sep) -> <:expr< Extend.Alist0sep $mlexpr_of_prod_entry_key s$ $str:sep$ >>
+ | Extend.Aopt s -> <:expr< Extend.Aopt $mlexpr_of_prod_entry_key s$ >>
+ | Extend.Amodifiers s -> <:expr< Extend.Amodifiers $mlexpr_of_prod_entry_key s$ >>
+ | Extend.Aself -> <:expr< Extend.Aself >>
+ | Extend.Anext -> <:expr< Extend.Anext >>
+ | Extend.Atactic n -> <:expr< Extend.Atactic $mlexpr_of_int n$ >>
+ | Extend.Agram s -> anomaly "Agram not supported"
+ | Extend.Aentry ("",s) -> <:expr< Extend.Agram (Gram.Entry.obj $lid:s$) >>
+ | Extend.Aentry (u,s) -> <:expr< Extend.Aentry $str:u$ $str:s$ >>
diff --git a/parsing/q_util.mli b/parsing/q_util.mli
index b950e68f..7c0fec9a 100644
--- a/parsing/q_util.mli
+++ b/parsing/q_util.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: q_util.mli 11309 2008-08-06 10:30:35Z herbelin $ i*)
+(*i $Id$ i*)
val patt_of_expr : MLast.expr -> MLast.patt
@@ -32,5 +32,4 @@ val mlexpr_of_string : string -> MLast.expr
val mlexpr_of_option : ('a -> MLast.expr) -> 'a option -> MLast.expr
-val interp_entry_name : Util.loc -> string -> string ->
- Pcoq.entry_type * MLast.expr
+val mlexpr_of_prod_entry_key : Pcoq.Gram.te Extend.prod_entry_key -> MLast.expr
diff --git a/parsing/search.ml b/parsing/search.ml
deleted file mode 100644
index 8b1551b6..00000000
--- a/parsing/search.ml
+++ /dev/null
@@ -1,214 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(* $Id: search.ml 11739 2009-01-02 19:33:19Z herbelin $ *)
-
-open Pp
-open Util
-open Names
-open Nameops
-open Term
-open Rawterm
-open Declarations
-open Libobject
-open Declare
-open Environ
-open Pattern
-open Matching
-open Printer
-open Libnames
-open Nametab
-
-(* The functions print_constructors and crible implement the behavior needed
- for the Coq searching commands.
- These functions take as first argument the procedure
- that will be called to treat each entry. This procedure receives the name
- of the object, the assumptions that will make it possible to print its type,
- and the constr term that represent its type. *)
-
-let print_constructors indsp fn env nconstr =
- for i = 1 to nconstr do
- fn (ConstructRef (indsp,i)) env (Inductiveops.type_of_constructor env (indsp,i))
- done
-
-let rec head_const c = match kind_of_term c with
- | Prod (_,_,d) -> head_const d
- | LetIn (_,_,_,d) -> head_const d
- | App (f,_) -> head_const f
- | Cast (d,_,_) -> head_const d
- | _ -> c
-
-(* General search, restricted to head constant if [only_head] *)
-
-let gen_crible refopt (fn : global_reference -> env -> constr -> unit) =
- let env = Global.env () in
- let crible_rec (sp,_) lobj = match object_tag lobj with
- | "VARIABLE" ->
- (try
- let (id,_,typ) = Global.lookup_named (basename sp) in
- if refopt = None
- || head_const typ = constr_of_global (Option.get refopt)
- then
- fn (VarRef id) env typ
- with Not_found -> (* we are in a section *) ())
- | "CONSTANT" ->
- let cst = locate_constant (qualid_of_sp sp) in
- let typ = Typeops.type_of_constant env cst in
- if refopt = None
- || head_const typ = constr_of_global (Option.get refopt)
- then
- fn (ConstRef cst) env typ
- | "INDUCTIVE" ->
- let kn = locate_mind (qualid_of_sp sp) in
- let mib = Global.lookup_mind kn in
- (match refopt with
- | Some (IndRef ((kn',tyi) as ind)) when kn=kn' ->
- print_constructors ind fn env
- (Array.length (mib.mind_packets.(tyi).mind_user_lc))
- | Some _ -> ()
- | _ ->
- Array.iteri
- (fun i mip -> print_constructors (kn,i) fn env
- (Array.length mip.mind_user_lc)) mib.mind_packets)
- | _ -> ()
- in
- try
- Declaremods.iter_all_segments crible_rec
- with Not_found ->
- ()
-
-let crible ref = gen_crible (Some ref)
-
-(* Fine Search. By Yves Bertot. *)
-
-exception No_section_path
-
-let rec head c =
- let c = strip_outer_cast c in
- match kind_of_term c with
- | Prod (_,_,c) -> head c
- | LetIn (_,_,_,c) -> head c
- | _ -> c
-
-let constr_to_section_path c = match kind_of_term c with
- | Const sp -> sp
- | _ -> raise No_section_path
-
-let xor a b = (a or b) & (not (a & b))
-
-let plain_display ref a c =
- let pc = pr_lconstr_env a c in
- let pr = pr_global ref in
- msg (hov 2 (pr ++ str":" ++ spc () ++ pc) ++ fnl ())
-
-let filter_by_module (module_list:dir_path list) (accept:bool)
- (ref:global_reference) _ _ =
- try
- let sp = sp_of_global ref in
- let sl = dirpath sp in
- let rec filter_aux = function
- | m :: tl -> (not (is_dirpath_prefix_of m sl)) && (filter_aux tl)
- | [] -> true
- in
- xor accept (filter_aux module_list)
- with No_section_path ->
- false
-
-let gref_eq =
- IndRef (Libnames.encode_kn Coqlib.logic_module (id_of_string "eq"), 0)
-let gref_eqT =
- IndRef (Libnames.encode_kn Coqlib.logic_type_module (id_of_string "eqT"), 0)
-
-let mk_rewrite_pattern1 eq pattern =
- PApp (PRef eq, [| PMeta None; pattern; PMeta None |])
-
-let mk_rewrite_pattern2 eq pattern =
- PApp (PRef eq, [| PMeta None; PMeta None; pattern |])
-
-let pattern_filter pat _ a c =
- try
- try
- is_matching pat (head c)
- with _ ->
- is_matching
- pat (head (Typing.type_of (Global.env()) Evd.empty c))
- with UserError _ ->
- false
-
-let filtered_search filter_function display_function ref =
- crible ref
- (fun s a c -> if filter_function s a c then display_function s a c)
-
-let rec id_from_pattern = function
- | PRef gr -> gr
-(* should be appear as a PRef (VarRef sp) !!
- | PVar id -> Nametab.locate (make_qualid [] (string_of_id id))
- *)
- | PApp (p,_) -> id_from_pattern p
- | _ -> error "The pattern is not simple enough."
-
-let raw_pattern_search extra_filter display_function pat =
- let name = id_from_pattern pat in
- filtered_search
- (fun s a c -> (pattern_filter pat s a c) & extra_filter s a c)
- display_function name
-
-let raw_search_rewrite extra_filter display_function pattern =
- filtered_search
- (fun s a c ->
- ((pattern_filter (mk_rewrite_pattern1 gref_eq pattern) s a c) ||
- (pattern_filter (mk_rewrite_pattern2 gref_eq pattern) s a c))
- && extra_filter s a c)
- display_function gref_eq
-
-let text_pattern_search extra_filter =
- raw_pattern_search extra_filter plain_display
-
-let text_search_rewrite extra_filter =
- raw_search_rewrite extra_filter plain_display
-
-let filter_by_module_from_list = function
- | [], _ -> (fun _ _ _ -> true)
- | l, outside -> filter_by_module l (not outside)
-
-let search_by_head ref inout =
- filtered_search (filter_by_module_from_list inout) plain_display ref
-
-let search_rewrite pat inout =
- text_search_rewrite (filter_by_module_from_list inout) pat
-
-let search_pattern pat inout =
- text_pattern_search (filter_by_module_from_list inout) pat
-
-
-let gen_filtered_search filter_function display_function =
- gen_crible None
- (fun s a c -> if filter_function s a c then display_function s a c)
-
-(** SearchAbout *)
-
-let name_of_reference ref = string_of_id (id_of_global ref)
-
-type glob_search_about_item =
- | GlobSearchSubPattern of constr_pattern
- | GlobSearchString of string
-
-let search_about_item (itemref,typ) = function
- | GlobSearchSubPattern pat -> is_matching_appsubterm ~closed:false pat typ
- | GlobSearchString s -> string_string_contains (name_of_reference itemref) s
-
-let raw_search_about filter_modules display_function l =
- let filter ref' env typ =
- filter_modules ref' env typ &&
- List.for_all (fun (b,i) -> b = search_about_item (ref',typ) i) l &&
- not (string_string_contains (name_of_reference ref') "_subproof")
- in
- gen_filtered_search filter display_function
-
-let search_about ref inout =
- raw_search_about (filter_by_module_from_list inout) plain_display ref
diff --git a/parsing/search.mli b/parsing/search.mli
deleted file mode 100644
index 7d12d26f..00000000
--- a/parsing/search.mli
+++ /dev/null
@@ -1,50 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(*i $Id: search.mli 11739 2009-01-02 19:33:19Z herbelin $ i*)
-
-open Pp
-open Names
-open Term
-open Environ
-open Pattern
-open Libnames
-open Nametab
-
-(*s Search facilities. *)
-
-type glob_search_about_item =
- | GlobSearchSubPattern of constr_pattern
- | GlobSearchString of string
-
-val search_by_head : global_reference -> dir_path list * bool -> unit
-val search_rewrite : constr_pattern -> dir_path list * bool -> unit
-val search_pattern : constr_pattern -> dir_path list * bool -> unit
-val search_about :
- (bool * glob_search_about_item) list -> dir_path list * bool -> unit
-
-(* The filtering function that is by standard search facilities.
- It can be passed as argument to the raw search functions.
- It is used in pcoq. *)
-
-val filter_by_module_from_list :
- dir_path list * bool -> global_reference -> env -> 'a -> bool
-
-(* raw search functions can be used for various extensions.
- They are also used for pcoq. *)
-val gen_filtered_search : (global_reference -> env -> constr -> bool) ->
- (global_reference -> env -> constr -> unit) -> unit
-val filtered_search : (global_reference -> env -> constr -> bool) ->
- (global_reference -> env -> constr -> unit) -> global_reference -> unit
-val raw_pattern_search : (global_reference -> env -> constr -> bool) ->
- (global_reference -> env -> constr -> unit) -> constr_pattern -> unit
-val raw_search_rewrite : (global_reference -> env -> constr -> bool) ->
- (global_reference -> env -> constr -> unit) -> constr_pattern -> unit
-val raw_search_about : (global_reference -> env -> constr -> bool) ->
- (global_reference -> env -> constr -> unit) ->
- (bool * glob_search_about_item) list -> unit
diff --git a/parsing/tacextend.ml4 b/parsing/tacextend.ml4
index 695677a3..517e34aa 100644
--- a/parsing/tacextend.ml4
+++ b/parsing/tacextend.ml4
@@ -8,30 +8,28 @@
(*i camlp4use: "pa_extend.cmo q_MLast.cmo" i*)
-(* $Id: tacextend.ml4 11309 2008-08-06 10:30:35Z herbelin $ *)
+(* $Id$ *)
open Util
open Genarg
open Q_util
open Q_coqast
open Argextend
-
-let loc = Util.dummy_loc
-let default_loc = <:expr< Util.dummy_loc >>
-
-type grammar_tactic_production_expr =
- | TacTerm of string
- | TacNonTerm of Util.loc * Genarg.argument_type * MLast.expr * string option
+open Pcoq
+open Extrawit
+open Egrammar
let rec make_patt = function
| [] -> <:patt< [] >>
- | TacNonTerm(loc',_,_,Some p)::l ->
+ | GramNonTerminal(loc',_,_,Some p)::l ->
+ let p = Names.string_of_id p in
<:patt< [ $lid:p$ :: $make_patt l$ ] >>
| _::l -> make_patt l
let rec make_when loc = function
| [] -> <:expr< True >>
- | TacNonTerm(loc',t,_,Some p)::l ->
+ | GramNonTerminal(loc',t,_,Some p)::l ->
+ let p = Names.string_of_id p in
let l = make_when loc l in
let loc = join_loc loc' loc in
let t = mlexpr_of_argtype loc' t in
@@ -40,14 +38,15 @@ let rec make_when loc = function
let rec make_let e = function
| [] -> e
- | TacNonTerm(loc,t,_,Some p)::l ->
+ | GramNonTerminal(loc,t,_,Some p)::l ->
+ let p = Names.string_of_id p in
let loc = join_loc loc (MLast.loc_of_expr e) in
let e = make_let e l in
let v = <:expr< Genarg.out_gen $make_wit loc t$ $lid:p$ >> in
- let v =
+ let v =
(* Special case for tactics which must be stored in algebraic
form to avoid marshalling closures and to be reprinted *)
- if Pcoq.is_tactic_genarg t then
+ if is_tactic_genarg t then
<:expr< ($v$, Tacinterp.eval_tactic $v$) >>
else v in
<:expr< let $lid:p$ = $v$ in $e$ >>
@@ -60,7 +59,7 @@ let add_clause s (pt,e) l =
let rec extract_signature = function
| [] -> []
- | TacNonTerm (_,t,_,_) :: l -> t :: extract_signature l
+ | GramNonTerminal (_,t,_,_) :: l -> t :: extract_signature l
| _::l -> extract_signature l
let check_unicity s l =
@@ -78,13 +77,15 @@ let make_clauses s l =
let rec make_args = function
| [] -> <:expr< [] >>
- | TacNonTerm(loc,t,_,Some p)::l ->
+ | GramNonTerminal(loc,t,_,Some p)::l ->
+ let p = Names.string_of_id p in
<:expr< [ Genarg.in_gen $make_wit loc t$ $lid:p$ :: $make_args l$ ] >>
| _::l -> make_args l
let rec make_eval_tactic e = function
| [] -> e
- | TacNonTerm(loc,tag,_,Some p)::l when Pcoq.is_tactic_genarg tag ->
+ | GramNonTerminal(loc,tag,_,Some p)::l when is_tactic_genarg tag ->
+ let p = Names.string_of_id p in
let loc = join_loc loc (MLast.loc_of_expr e) in
let e = make_eval_tactic e l in
(* Special case for tactics which must be stored in algebraic
@@ -94,26 +95,27 @@ let rec make_eval_tactic e = function
let rec make_fun e = function
| [] -> e
- | TacNonTerm(loc,_,_,Some p)::l ->
+ | GramNonTerminal(loc,_,_,Some p)::l ->
+ let p = Names.string_of_id p in
<:expr< fun $lid:p$ -> $make_fun e l$ >>
| _::l -> make_fun e l
-let mlexpr_of_grammar_production = function
- | TacTerm s ->
- <:expr< Egrammar.TacTerm $mlexpr_of_string s$ >>
- | TacNonTerm (loc,nt,g,sopt) ->
- <:expr< Egrammar.TacNonTerm $default_loc$ ($g$,$mlexpr_of_argtype loc nt$) $mlexpr_of_option mlexpr_of_string sopt$ >>
+let mlexpr_terminals_of_grammar_tactic_prod_item_expr = function
+ | GramTerminal s -> <:expr< Some $mlexpr_of_string s$ >>
+ | GramNonTerminal (loc,nt,_,sopt) -> <:expr< None >>
-let mlexpr_terminals_of_grammar_production = function
- | TacTerm s -> <:expr< Some $mlexpr_of_string s$ >>
- | TacNonTerm (loc,nt,g,sopt) -> <:expr< None >>
+let make_prod_item = function
+ | GramTerminal s -> <:expr< Egrammar.GramTerminal $str:s$ >>
+ | GramNonTerminal (loc,nt,g,sopt) ->
+ <:expr< Egrammar.GramNonTerminal $default_loc$ $mlexpr_of_argtype loc nt$
+ $mlexpr_of_prod_entry_key g$ $mlexpr_of_option mlexpr_of_ident sopt$ >>
let mlexpr_of_clause =
- mlexpr_of_list (fun (a,b) -> mlexpr_of_list mlexpr_of_grammar_production a)
+ mlexpr_of_list (fun (a,b) -> mlexpr_of_list make_prod_item a)
let rec make_tags loc = function
| [] -> <:expr< [] >>
- | TacNonTerm(loc',t,_,Some p)::l ->
+ | GramNonTerminal(loc',t,_,Some p)::l ->
let l = make_tags loc l in
let loc = join_loc loc' loc in
let t = mlexpr_of_argtype loc' t in
@@ -123,7 +125,7 @@ let rec make_tags loc = function
let make_one_printing_rule se (pt,e) =
let level = mlexpr_of_int 0 in (* only level 0 supported here *)
let loc = MLast.loc_of_expr e in
- let prods = mlexpr_of_list mlexpr_terminals_of_grammar_production pt in
+ let prods = mlexpr_of_list mlexpr_terminals_of_grammar_tactic_prod_item_expr pt in
<:expr< ($se$, $make_tags loc pt$, ($level$, $prods$)) >>
let make_printing_rule se = mlexpr_of_list (make_one_printing_rule se)
@@ -136,10 +138,10 @@ let rec contains_epsilon = function
| ExtraArgType("hintbases") -> true
| _ -> false
let is_atomic = function
- | TacTerm s :: l when
+ | GramTerminal s :: l when
List.for_all (function
- TacTerm _ -> false
- | TacNonTerm(_,t,_,_) -> contains_epsilon t) l
+ GramTerminal _ -> false
+ | GramNonTerminal(_,t,_,_) -> contains_epsilon t) l
-> [s]
| _ -> []
@@ -150,7 +152,7 @@ let declare_tactic loc s cl =
let hide_tac (p,e) =
(* reste a definir les fonctions cachees avec des noms frais *)
let stac = "h_"^s in
- let e =
+ let e =
make_fun
<:expr<
Refiner.abstract_extended_tactic $mlexpr_of_string s$ $make_args p$ $make_eval_tactic e p$
@@ -165,6 +167,7 @@ let declare_tactic loc s cl =
<:str_item<
declare
open Pcoq;
+ open Extrawit;
declare $list:hidden$ end;
try
let _=Tacinterp.add_tactic $se$ (fun [ $list:make_clauses s cl$ ]) in
@@ -191,8 +194,8 @@ EXTEND
;
tacrule:
[ [ "["; l = LIST1 tacargs; "]"; "->"; "["; e = Pcaml.expr; "]"
- ->
- if match List.hd l with TacNonTerm _ -> true | _ -> false then
+ ->
+ if match List.hd l with GramNonTerminal _ -> true | _ -> false then
(* En attendant la syntaxe de tacticielles *)
failwith "Tactic syntax must start with an identifier";
(l,e)
@@ -200,14 +203,14 @@ EXTEND
;
tacargs:
[ [ e = LIDENT; "("; s = LIDENT; ")" ->
- let t, g = Q_util.interp_entry_name loc e "" in
- TacNonTerm (loc, t, g, Some s)
+ let t, g = interp_entry_name false None e "" in
+ GramNonTerminal (loc, t, g, Some (Names.id_of_string s))
| e = LIDENT; "("; s = LIDENT; ","; sep = STRING; ")" ->
- let t, g = Q_util.interp_entry_name loc e sep in
- TacNonTerm (loc, t, g, Some s)
+ let t, g = interp_entry_name false None e sep in
+ GramNonTerminal (loc, t, g, Some (Names.id_of_string s))
| s = STRING ->
if s = "" then Util.user_err_loc (loc,"",Pp.str "Empty terminal.");
- TacTerm s
+ GramTerminal s
] ]
;
END
diff --git a/parsing/tactic_printer.ml b/parsing/tactic_printer.ml
index e0836984..c09b3431 100644
--- a/parsing/tactic_printer.ml
+++ b/parsing/tactic_printer.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: tactic_printer.ml 11313 2008-08-07 11:15:03Z barras $ *)
+(* $Id$ *)
open Pp
open Util
@@ -23,30 +23,30 @@ let pr_tactic = function
| TacArg (Tacexp t) ->
(*top tactic from tacinterp*)
Pptactic.pr_glob_tactic (Global.env()) t
- | t ->
+ | t ->
Pptactic.pr_tactic (Global.env()) t
-let pr_proof_instr instr =
+let pr_proof_instr instr =
Ppdecl_proof.pr_proof_instr (Global.env()) instr
let pr_rule = function
| Prim r -> hov 0 (pr_prim_rule r)
| Nested(cmpd,_) ->
begin
- match cmpd with
+ match cmpd with
| Tactic (texp,_) -> hov 0 (pr_tactic texp)
| Proof_instr (_,instr) -> hov 0 (pr_proof_instr instr)
end
| Daimon -> str "<Daimon>"
- | Decl_proof _ -> str "proof"
+ | Decl_proof _ -> str "proof"
let uses_default_tac = function
| Nested(Tactic(_,dflt),_) -> dflt
| _ -> false
(* Does not print change of evars *)
-let pr_rule_dot = function
- | Prim Change_evars ->str "PC: ch_evars" ++ mt ()
+let pr_rule_dot = function
+ | Prim Change_evars ->str "PC: ch_evars" ++ mt ()
(* PC: this might be redundant *)
| r ->
pr_rule r ++ if uses_default_tac r then str "..." else str"."
@@ -66,27 +66,27 @@ exception Different
let thin_sign osign sign =
Sign.fold_named_context
(fun (id,c,ty as d) sign ->
- try
+ try
if Sign.lookup_named id osign = (id,c,ty) then sign
else raise Different
with Not_found | Different -> Environ.push_named_context_val d sign)
sign ~init:Environ.empty_named_context_val
-let rec print_proof sigma osign pf =
+let rec print_proof _sigma osign pf =
let hyps = Environ.named_context_of_val pf.goal.evar_hyps in
let hyps' = thin_sign osign hyps in
match pf.ref with
- | None ->
+ | None ->
hov 0 (pr_goal {pf.goal with evar_hyps=hyps'})
| Some(r,spfl) ->
- hov 0
+ hov 0
(hov 0 (pr_goal {pf.goal with evar_hyps=hyps'}) ++
spc () ++ str" BY " ++
hov 0 (pr_rule r) ++ fnl () ++
str" " ++
- hov 0 (prlist_with_sep pr_fnl (print_proof sigma hyps) spfl))
-
-let pr_change gl =
+ hov 0 (prlist_with_sep pr_fnl (print_proof _sigma hyps) spfl))
+
+let pr_change gl =
str"change " ++
pr_lconstr_env (Global.env_of_context gl.evar_hyps) gl.evar_concl ++ str"."
@@ -94,9 +94,9 @@ let print_decl_script tac_printer ?(nochange=true) sigma pf =
let rec print_prf pf =
match pf.ref with
| None ->
- (if nochange then
+ (if nochange then
(str"<Your Proof Text here>")
- else
+ else
pr_change pf.goal)
++ fnl ()
| Some (Daimon,[]) -> str "(* Some proof has been skipped here *)"
@@ -114,17 +114,17 @@ let print_decl_script tac_printer ?(nochange=true) sigma pf =
(if opened then mt () else str "end claim." ++ fnl ()) ++
print_prf cont
| Pfocus _,[body;cont] ->
- hov 2 (pr_rule_dot_fnl rule ++ print_prf body) ++
+ hov 2 (pr_rule_dot_fnl rule ++ print_prf body) ++
fnl () ++
(if opened then mt () else str "end focus." ++ fnl ()) ++
print_prf cont
| (Psuppose _ |Pcase (_,_,_)),[body;cont] ->
hov 2 (pr_rule_dot_fnl rule ++ print_prf body) ++ fnl () ++
- print_prf cont
+ print_prf cont
| _,[next] ->
pr_rule_dot_fnl rule ++ print_prf next
| _,[] ->
- pr_rule_dot rule
+ pr_rule_dot rule
| _,_ -> anomaly "unknown branching instruction"
end
| _ -> anomaly "Not Applicable" in
@@ -134,19 +134,19 @@ let print_script ?(nochange=true) sigma pf =
let rec print_prf pf =
match pf.ref with
| None ->
- (if nochange then
- (str"<Your Tactic Text here>")
- else
+ (if nochange then
+ (str"<Your Tactic Text here>")
+ else
pr_change pf.goal)
++ fnl ()
| Some(Decl_proof opened,script) ->
assert (List.length script = 1);
begin
- if nochange then (mt ()) else (pr_change pf.goal ++ fnl ())
+ if nochange then (mt ()) else (pr_change pf.goal ++ fnl ())
end ++
begin
- hov 0 (str "proof." ++ fnl () ++
- print_decl_script print_prf
+ hov 0 (str "proof." ++ fnl () ++
+ print_decl_script print_prf
~nochange sigma (List.hd script))
end ++ fnl () ++
begin
@@ -167,7 +167,7 @@ let print_treescript ?(nochange=true) sigma pf =
let rec print_prf pf =
match pf.ref with
| None ->
- if nochange then
+ if nochange then
if pf.goal.evar_extra=None then str"<Your Tactic Text here>"
else str"<Your Proof Text here>"
else pr_change pf.goal
@@ -176,10 +176,10 @@ let print_treescript ?(nochange=true) sigma pf =
begin
if nochange then mt () else pr_change pf.goal ++ fnl ()
end ++
- hov 0
+ hov 0
begin str "proof." ++ fnl () ++
- print_decl_script print_prf ~nochange sigma (List.hd script)
- end ++ fnl () ++
+ print_decl_script print_prf ~nochange sigma (List.hd script)
+ end ++ fnl () ++
begin
if opened then mt () else (str "end proof." ++ fnl ())
end
@@ -197,28 +197,29 @@ let rec print_info_script sigma osign pf =
match pf.ref with
| None -> (mt ())
| Some(r,spfl) ->
- (pr_rule r ++
+ (pr_rule r ++
match spfl with
| [pf1] ->
- if pf1.ref = None then
+ if pf1.ref = None then
(str "." ++ fnl ())
- else
+ else
(str";" ++ brk(1,3) ++
- print_info_script sigma
+ print_info_script sigma
(Environ.named_context_of_val sign) pf1)
| _ -> (str"." ++ fnl () ++
prlist_with_sep pr_fnl
- (print_info_script sigma
+ (print_info_script sigma
(Environ.named_context_of_val sign)) spfl))
-let format_print_info_script sigma osign pf =
+let format_print_info_script sigma osign pf =
hov 0 (print_info_script sigma osign pf)
-
-let print_subscript sigma sign pf =
- if is_tactic_proof pf then
+
+let print_subscript sigma sign pf =
+ if is_tactic_proof pf then
format_print_info_script sigma sign (subproof_of_proof pf)
- else
+ else
format_print_info_script sigma sign pf
let _ = Refiner.set_info_printer print_subscript
+let _ = Refiner.set_proof_printer print_proof
diff --git a/parsing/tactic_printer.mli b/parsing/tactic_printer.mli
index affc5ec2..d46f19c6 100644
--- a/parsing/tactic_printer.mli
+++ b/parsing/tactic_printer.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: tactic_printer.mli 11313 2008-08-07 11:15:03Z barras $ i*)
+(*i $Id$ i*)
(*i*)
open Pp
diff --git a/parsing/vernacextend.ml4 b/parsing/vernacextend.ml4
index a7b27e21..e8a3094b 100644
--- a/parsing/vernacextend.ml4
+++ b/parsing/vernacextend.ml4
@@ -8,38 +8,21 @@
(*i camlp4use: "pa_extend.cmo q_MLast.cmo" i*)
-(* $Id: vernacextend.ml4 11622 2008-11-23 08:45:56Z herbelin $ *)
+(* $Id$ *)
open Util
open Genarg
open Q_util
open Q_coqast
open Argextend
-
-let loc = Util.dummy_loc
-let default_loc = <:expr< Util.dummy_loc >>
-
-type grammar_tactic_production_expr =
- | VernacTerm of string
- | VernacNonTerm of Util.loc * Genarg.argument_type * MLast.expr * string option
-let rec make_patt = function
- | [] -> <:patt< [] >>
- | VernacNonTerm(_,_,_,Some p)::l ->
- <:patt< [ $lid:p$ :: $make_patt l$ ] >>
- | _::l -> make_patt l
-
-let rec make_when loc = function
- | [] -> <:expr< True >>
- | VernacNonTerm(loc',t,_,Some p)::l ->
- let l = make_when loc l in
- let loc = join_loc loc' loc in
- let t = mlexpr_of_argtype loc' t in
- <:expr< Genarg.genarg_tag $lid:p$ = $t$ && $l$ >>
- | _::l -> make_when loc l
+open Tacextend
+open Pcoq
+open Egrammar
let rec make_let e = function
| [] -> e
- | VernacNonTerm(loc,t,_,Some p)::l ->
+ | GramNonTerminal(loc,t,_,Some p)::l ->
+ let p = Names.string_of_id p in
let loc = join_loc loc (MLast.loc_of_expr e) in
let e = make_let e l in
<:expr< let $lid:p$ = Genarg.out_gen $make_rawwit loc t$ $lid:p$ in $e$ >>
@@ -50,11 +33,6 @@ let add_clause s (_,pt,e) l =
let w = Some (make_when (MLast.loc_of_expr e) pt) in
(p, w, make_let e pt)::l
-let rec extract_signature = function
- | [] -> []
- | VernacNonTerm (_,t,_,_) :: l -> t :: extract_signature l
- | _::l -> extract_signature l
-
let check_unicity s l =
let l' = List.map (fun (_,l,_) -> extract_signature l) l in
if not (Util.list_distinct l') then
@@ -68,22 +46,9 @@ let make_clauses s l =
(<:patt< _ >>,None,<:expr< failwith "Vernac extension: cannot occur" >>) in
List.fold_right (add_clause s) l [default]
-let rec make_fun e = function
- | [] -> e
- | VernacNonTerm(loc,_,_,Some p)::l ->
- <:expr< fun $lid:p$ -> $make_fun e l$ >>
- | _::l -> make_fun e l
-
-let mlexpr_of_grammar_production = function
- | VernacTerm s ->
- <:expr< Egrammar.TacTerm $mlexpr_of_string s$ >>
- | VernacNonTerm (loc,nt,g,sopt) ->
- <:expr< Egrammar.TacNonTerm $default_loc$ ($g$,$mlexpr_of_argtype loc nt$) $mlexpr_of_option mlexpr_of_string sopt$ >>
-
let mlexpr_of_clause =
mlexpr_of_list
- (fun (a,b,c) ->
- mlexpr_of_list mlexpr_of_grammar_production (VernacTerm a::b))
+ (fun (a,b,c) -> mlexpr_of_list make_prod_item (GramTerminal a::b))
let declare_command loc s cl =
let gl = mlexpr_of_clause cl in
@@ -91,6 +56,7 @@ let declare_command loc s cl =
<:str_item<
declare
open Pcoq;
+ open Extrawit;
try Vernacinterp.vinterp_add $mlexpr_of_string s$ (fun [ $list:icl$ ])
with e -> Pp.pp (Cerrors.explain_exn e);
Egrammar.extend_vernac_command_grammar $mlexpr_of_string s$ $gl$;
@@ -109,20 +75,20 @@ EXTEND
;
rule:
[ [ "["; s = STRING; l = LIST0 args; "]"; "->"; "["; e = Pcaml.expr; "]"
- ->
+ ->
if s = "" then Util.user_err_loc (loc,"",Pp.str"Command name is empty.");
(s,l,<:expr< fun () -> $e$ >>)
] ]
;
args:
[ [ e = LIDENT; "("; s = LIDENT; ")" ->
- let t, g = Q_util.interp_entry_name loc e "" in
- VernacNonTerm (loc, t, g, Some s)
+ let t, g = interp_entry_name false None e "" in
+ GramNonTerminal (loc, t, g, Some (Names.id_of_string s))
| e = LIDENT; "("; s = LIDENT; ","; sep = STRING; ")" ->
- let t, g = Q_util.interp_entry_name loc e sep in
- VernacNonTerm (loc, t, g, Some s)
+ let t, g = interp_entry_name false None e sep in
+ GramNonTerminal (loc, t, g, Some (Names.id_of_string s))
| s = STRING ->
- VernacTerm s
+ GramTerminal s
] ]
;
END