aboutsummaryrefslogtreecommitdiffhomepage
path: root/printing
diff options
context:
space:
mode:
authorGravatar Regis-Gianas <yrg@pps.univ-paris-diderot.fr>2014-10-30 22:31:27 +0100
committerGravatar Regis-Gianas <yrg@pps.univ-paris-diderot.fr>2014-11-04 22:51:35 +0100
commit2656c62e9be7a856f12c0da0740869d193273c4d (patch)
tree96792a9ccb608e795857f784bc6373eebfd4635c /printing
parent0aa24d51d2606549da87ed42085f612f2dbb1428 (diff)
Ppannotation: New.
Define the annotations stored in semi-structured pretty-prints. Ppconstrsig: New. Contains the signature of a pretty-printer for ppconstr. Ppconstr: Export a new rich pretty-printer for constr_expr and co.
Diffstat (limited to 'printing')
-rw-r--r--printing/ppannotation.ml14
-rw-r--r--printing/ppannotation.mli17
-rw-r--r--printing/ppconstr.ml17
-rw-r--r--printing/ppconstr.mli101
-rw-r--r--printing/ppconstrsig.mli105
-rw-r--r--printing/printing.mllib2
6 files changed, 168 insertions, 88 deletions
diff --git a/printing/ppannotation.ml b/printing/ppannotation.ml
new file mode 100644
index 000000000..549a0dab8
--- /dev/null
+++ b/printing/ppannotation.ml
@@ -0,0 +1,14 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Constrexpr
+open Ppextend
+
+type t =
+ | AUnparsing of unparsing
+ | AConstrExpr of constr_expr
diff --git a/printing/ppannotation.mli b/printing/ppannotation.mli
new file mode 100644
index 000000000..54f70aeda
--- /dev/null
+++ b/printing/ppannotation.mli
@@ -0,0 +1,17 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** This module defines the annotations that are attached to
+ semi-structured pretty-printing of Coq syntactic objects. *)
+
+open Constrexpr
+open Ppextend
+
+type t =
+ | AUnparsing of unparsing
+ | AConstrExpr of constr_expr
diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml
index fe8b582b1..99ca023a8 100644
--- a/printing/ppconstr.ml
+++ b/printing/ppconstr.ml
@@ -765,3 +765,20 @@ include Make (struct
let tag_unparsing = do_not_tag
let tag_constr_expr = do_not_tag
end)
+
+module RichPp (Indexer : sig
+ val index : Ppannotation.t -> string
+end) = struct
+
+ include Make (struct
+ open Ppannotation
+
+ let tag_unparsing unp t =
+ Pp.open_tag (Indexer.index (AUnparsing unp)) ++ t ++ Pp.close_tag ()
+
+ let tag_constr_expr e t =
+ Pp.open_tag (Indexer.index (AConstrExpr e)) ++ t ++ Pp.close_tag ()
+ end)
+
+end
+
diff --git a/printing/ppconstr.mli b/printing/ppconstr.mli
index 79399094d..a4c262b92 100644
--- a/printing/ppconstr.mli
+++ b/printing/ppconstr.mli
@@ -6,95 +6,20 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Loc
-open Pp
-open Libnames
-open Constrexpr
-open Names
-open Misctypes
-open Locus
-open Genredexpr
+(** This module implements pretty-printers for constr_expr syntactic
+ objects and their subcomponents. *)
-val extract_lam_binders :
- constr_expr -> local_binder list * constr_expr
-val extract_prod_binders :
- constr_expr -> local_binder list * constr_expr
-val split_fix :
- int -> constr_expr -> constr_expr ->
- local_binder list * constr_expr * constr_expr
+(** The default pretty-printers produce {!Pp.std_ppcmds} that are
+ interpreted as raw strings. *)
+include Ppconstrsig.Pp
-val prec_less : int -> int * Ppextend.parenRelation -> bool
+(** The rich pretty-printers produce {!Pp.std_ppcmds} that are
+ interpreted as annotated strings. The annotations can be
+ retrieved using {!RichPp.rich_pp}. Their definitions is
+ located in {!Ppannotation.t}.
-val pr_tight_coma : unit -> std_ppcmds
+ Please refer to {!RichPp} to know what are the requirements over
+ [Indexer.index] behavior. *)
+module RichPp (Indexer : sig val index : Ppannotation.t -> string end)
+ : Ppconstrsig.Pp
-val pr_or_var : ('a -> std_ppcmds) -> 'a or_var -> std_ppcmds
-val pr_metaid : Id.t -> std_ppcmds
-
-val pr_lident : Id.t located -> std_ppcmds
-val pr_lname : Name.t located -> std_ppcmds
-
-val pr_with_comments : Loc.t -> std_ppcmds -> std_ppcmds
-val pr_com_at : int -> std_ppcmds
-val pr_sep_com :
- (unit -> std_ppcmds) ->
- (constr_expr -> std_ppcmds) ->
- constr_expr -> std_ppcmds
-
-val pr_id : Id.t -> std_ppcmds
-val pr_name : Name.t -> 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_red_expr :
- ('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) ->
- ('c -> std_ppcmds) -> ('a,'b,'c) may_eval -> std_ppcmds
-
-val pr_glob_sort : glob_sort -> std_ppcmds
-val pr_guard_annot : (constr_expr -> std_ppcmds) ->
- local_binder list ->
- ('a * Names.Id.t) option * recursion_order_expr ->
- std_ppcmds
-
-val pr_binders : local_binder list -> std_ppcmds
-val pr_constr_pattern_expr : constr_pattern_expr -> std_ppcmds
-val pr_lconstr_pattern_expr : constr_pattern_expr -> std_ppcmds
-val pr_constr_expr : constr_expr -> std_ppcmds
-val pr_lconstr_expr : constr_expr -> std_ppcmds
-val pr_cases_pattern_expr : cases_pattern_expr -> std_ppcmds
-
-type term_pr = {
- pr_constr_expr : constr_expr -> std_ppcmds;
- pr_lconstr_expr : constr_expr -> std_ppcmds;
- pr_constr_pattern_expr : constr_pattern_expr -> std_ppcmds;
- pr_lconstr_pattern_expr : constr_pattern_expr -> std_ppcmds
-}
-
-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 [lsimpleconstr] 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 (_,GProp 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 lsimpleconstr : 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/printing/ppconstrsig.mli b/printing/ppconstrsig.mli
new file mode 100644
index 000000000..7ac650c98
--- /dev/null
+++ b/printing/ppconstrsig.mli
@@ -0,0 +1,105 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Loc
+open Pp
+open Libnames
+open Constrexpr
+open Names
+open Misctypes
+open Locus
+open Genredexpr
+
+module type Pp = sig
+
+ val extract_lam_binders :
+ constr_expr -> local_binder list * constr_expr
+ val extract_prod_binders :
+ constr_expr -> local_binder list * constr_expr
+ val split_fix :
+ 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
+ val pr_metaid : Id.t -> std_ppcmds
+
+ val pr_lident : Id.t located -> std_ppcmds
+ val pr_lname : Name.t located -> std_ppcmds
+
+ val pr_with_comments : Loc.t -> std_ppcmds -> std_ppcmds
+ val pr_com_at : int -> std_ppcmds
+ val pr_sep_com :
+ (unit -> std_ppcmds) ->
+ (constr_expr -> std_ppcmds) ->
+ constr_expr -> std_ppcmds
+
+ val pr_id : Id.t -> std_ppcmds
+ val pr_name : Name.t -> 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_red_expr :
+ ('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) ->
+ ('c -> std_ppcmds) -> ('a,'b,'c) may_eval -> std_ppcmds
+
+ val pr_glob_sort : glob_sort -> std_ppcmds
+ val pr_guard_annot : (constr_expr -> std_ppcmds) ->
+ local_binder list ->
+ ('a * Names.Id.t) option * recursion_order_expr ->
+ std_ppcmds
+
+ val pr_binders : local_binder list -> std_ppcmds
+ val pr_constr_pattern_expr : constr_pattern_expr -> std_ppcmds
+ val pr_lconstr_pattern_expr : constr_pattern_expr -> std_ppcmds
+ val pr_constr_expr : constr_expr -> std_ppcmds
+ val pr_lconstr_expr : constr_expr -> std_ppcmds
+ val pr_cases_pattern_expr : cases_pattern_expr -> std_ppcmds
+
+ type term_pr = {
+ pr_constr_expr : constr_expr -> std_ppcmds;
+ pr_lconstr_expr : constr_expr -> std_ppcmds;
+ pr_constr_pattern_expr : constr_pattern_expr -> std_ppcmds;
+ pr_lconstr_pattern_expr : constr_pattern_expr -> std_ppcmds
+ }
+
+ 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 [lsimpleconstr] 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 (_,GProp 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 lsimpleconstr : precedence
+ val ltop : precedence
+ val modular_constr_pr :
+ ((unit->std_ppcmds) -> precedence -> constr_expr -> std_ppcmds) ->
+ (unit->std_ppcmds) -> precedence -> constr_expr -> std_ppcmds
+
+end
+
diff --git a/printing/printing.mllib b/printing/printing.mllib
index 2a8f1030f..28c57edba 100644
--- a/printing/printing.mllib
+++ b/printing/printing.mllib
@@ -1,6 +1,8 @@
Genprint
Pputils
+Ppannotation
Ppconstr
+Ppconstrsig
Printer
Pptactic
Printmod