aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7>2003-10-16 14:52:09 +0000
committerGravatar herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7>2003-10-16 14:52:09 +0000
commit4d61b0c12eb94e220740e39e10a7678cdc0a5a85 (patch)
tree6f070cffd72473b1c22301929b617d90eb6ce842
parent214759f4b785489932031d0651a7c8a85f49ab1d (diff)
Debranchement de l'affichage systematique des projections avec la notation pointee; soumis maintenant a l'activation de l'option 'Set Printing Projections'
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@4649 85f007b7-540e-0410-9357-904b9bb8a0f7
-rw-r--r--interp/constrextern.ml7
-rw-r--r--toplevel/vernacentries.ml8
2 files changed, 13 insertions, 2 deletions
diff --git a/interp/constrextern.ml b/interp/constrextern.ml
index ea17a9fba..1be3b805a 100644
--- a/interp/constrextern.ml
+++ b/interp/constrextern.ml
@@ -57,6 +57,9 @@ let print_universes = ref false
(* This suppresses printing of numeral and symbols *)
let print_no_symbol = ref false
+(* This governs printing of projections using the dot notation symbols *)
+let print_projections = ref false
+
let print_meta_as_hole = ref false
let with_arguments f = Options.with_option print_arguments f
@@ -465,12 +468,12 @@ let occur_name na aty =
| Anonymous -> false
let is_projection nargs = function
- | Some r ->
+ | Some r when !print_projections ->
(try
let n = Recordops.find_projection_nparams r + 1 in
if n <= nargs then Some n else None
with Not_found -> None)
- | None -> None
+ | _ -> None
let stdlib = function
| Some r ->
diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml
index b9150a30e..3f979d9b1 100644
--- a/toplevel/vernacentries.ml
+++ b/toplevel/vernacentries.ml
@@ -722,6 +722,14 @@ let _ =
let _ =
declare_bool_option
{ optsync = true;
+ optname = "projection printing using dot notation";
+ optkey = (SecondaryTable ("Printing","Projections"));
+ optread = (fun () -> !Constrextern.print_projections);
+ optwrite = (fun b -> Constrextern.print_projections := b) }
+
+let _ =
+ declare_bool_option
+ { optsync = true;
optname = "symbols printing";
optkey = (SecondaryTable ("Printing",if !Options.v7 then "Symbols" else "Notations"));
optread = (fun () -> not !Constrextern.print_no_symbol);