| author | blanchet | 
| Mon, 11 Oct 2010 18:03:18 +0700 | |
| changeset 39979 | b13515940b53 | 
| parent 39616 | 8052101883c3 | 
| child 40292 | ba13793594f0 | 
| permissions | -rw-r--r-- | 
| 21637 | 1  | 
(* Title: Pure/ProofGeneral/preferences.ML  | 
2  | 
Author: David Aspinall and Markus Wenzel  | 
|
3  | 
||
4  | 
User preferences for Isabelle which are maintained by the interface.  | 
|
5  | 
*)  | 
|
6  | 
||
| 24329 | 7  | 
signature PREFERENCES =  | 
| 21637 | 8  | 
sig  | 
| 30980 | 9  | 
val category_display: string  | 
10  | 
val category_advanced_display: string  | 
|
11  | 
val category_tracing: string  | 
|
12  | 
val category_proof: string  | 
|
| 28587 | 13  | 
type preference =  | 
14  | 
   {name: string,
 | 
|
15  | 
descr: string,  | 
|
16  | 
default: string,  | 
|
17  | 
pgiptype: PgipTypes.pgiptype,  | 
|
18  | 
get: unit -> string,  | 
|
19  | 
set: string -> unit}  | 
|
20  | 
  val generic_pref: ('a -> string) -> (string -> 'a) -> PgipTypes.pgiptype ->
 | 
|
| 32738 | 21  | 
'a Unsynchronized.ref -> string -> string -> preference  | 
22  | 
val string_pref: string Unsynchronized.ref -> string -> string -> preference  | 
|
23  | 
val int_pref: int Unsynchronized.ref -> string -> string -> preference  | 
|
24  | 
val nat_pref: int Unsynchronized.ref -> string -> string -> preference  | 
|
25  | 
val bool_pref: bool Unsynchronized.ref -> string -> string -> preference  | 
|
| 28587 | 26  | 
type T = (string * preference list) list  | 
27  | 
val pure_preferences: T  | 
|
28  | 
val remove: string -> T -> T  | 
|
29  | 
val add: string -> preference -> T -> T  | 
|
30  | 
val set_default: string * string -> T -> T  | 
|
| 21637 | 31  | 
end  | 
32  | 
||
| 28587 | 33  | 
structure Preferences: PREFERENCES =  | 
| 21637 | 34  | 
struct  | 
35  | 
||
| 30980 | 36  | 
(* categories *)  | 
37  | 
||
38  | 
val category_display = "Display";  | 
|
39  | 
val category_advanced_display = "Advanced Display";  | 
|
40  | 
val category_tracing = "Tracing";  | 
|
41  | 
val category_proof = "Proof"  | 
|
42  | 
||
43  | 
||
| 28587 | 44  | 
(* preferences and preference tables *)  | 
| 21637 | 45  | 
|
| 28587 | 46  | 
type preference =  | 
47  | 
 {name: string,
 | 
|
48  | 
descr: string,  | 
|
49  | 
default: string,  | 
|
50  | 
pgiptype: PgipTypes.pgiptype,  | 
|
51  | 
get: unit -> string,  | 
|
52  | 
set: string -> unit};  | 
|
53  | 
||
54  | 
fun mkpref raw_get raw_set typ name descr : preference =  | 
|
55  | 
let  | 
|
56  | 
fun get () = CRITICAL raw_get;  | 
|
57  | 
fun set x = CRITICAL (fn () => raw_set x);  | 
|
58  | 
  in {name = name, descr = descr, pgiptype = typ, get = get, set = set, default = get ()} end;
 | 
|
| 21637 | 59  | 
|
60  | 
||
| 28587 | 61  | 
(* generic preferences *)  | 
| 21637 | 62  | 
|
| 28587 | 63  | 
fun generic_pref read write typ r =  | 
64  | 
mkpref (fn () => read (! r)) (fn x => r := write x) typ;  | 
|
65  | 
||
66  | 
val string_pref = generic_pref I I PgipTypes.Pgipstring;  | 
|
| 21637 | 67  | 
|
| 24329 | 68  | 
val int_pref =  | 
| 28587 | 69  | 
generic_pref PgipTypes.int_to_pgstring (PgipTypes.read_pgipint (NONE, NONE))  | 
70  | 
(PgipTypes.Pgipint (NONE, NONE));  | 
|
71  | 
||
| 21637 | 72  | 
val nat_pref =  | 
| 28587 | 73  | 
generic_pref PgipTypes.int_to_pgstring PgipTypes.read_pgipnat PgipTypes.Pgipnat;  | 
| 21637 | 74  | 
|
75  | 
val bool_pref =  | 
|
| 28587 | 76  | 
generic_pref PgipTypes.bool_to_pgstring PgipTypes.read_pgipbool PgipTypes.Pgipbool;  | 
77  | 
||
78  | 
||
79  | 
(* preferences of Pure *)  | 
|
| 21637 | 80  | 
|
| 
39616
 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 
wenzelm 
parents: 
39165 
diff
changeset
 | 
81  | 
val proof_pref = Unsynchronized.setmp Proofterm.proofs 1 (fn () =>  | 
| 28587 | 82  | 
let  | 
83  | 
fun get () = PgipTypes.bool_to_pgstring (! Proofterm.proofs >= 2);  | 
|
84  | 
fun set s = Proofterm.proofs := (if PgipTypes.read_pgipbool s then 2 else 1);  | 
|
| 30980 | 85  | 
in mkpref get set PgipTypes.Pgipbool "full-proofs" "Record full proof objects internally" end) ();  | 
| 21637 | 86  | 
|
| 
32061
 
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
 
wenzelm 
parents: 
30985 
diff
changeset
 | 
87  | 
val parallel_proof_pref =  | 
| 
 
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
 
wenzelm 
parents: 
30985 
diff
changeset
 | 
88  | 
let  | 
| 
 
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
 
wenzelm 
parents: 
30985 
diff
changeset
 | 
89  | 
fun get () = PgipTypes.bool_to_pgstring (! Goal.parallel_proofs >= 1);  | 
| 
 
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
 
wenzelm 
parents: 
30985 
diff
changeset
 | 
90  | 
fun set s = Goal.parallel_proofs := (if PgipTypes.read_pgipbool s then 1 else 0);  | 
| 
 
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
 
wenzelm 
parents: 
30985 
diff
changeset
 | 
91  | 
in mkpref get set PgipTypes.Pgipbool "parallel-proofs" "Check proofs in parallel" end;  | 
| 
 
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
 
wenzelm 
parents: 
30985 
diff
changeset
 | 
92  | 
|
| 28587 | 93  | 
val thm_depsN = "thm_deps";  | 
| 24329 | 94  | 
val thm_deps_pref =  | 
| 28587 | 95  | 
let  | 
96  | 
fun get () = PgipTypes.bool_to_pgstring (print_mode_active thm_depsN);  | 
|
97  | 
fun set s =  | 
|
| 32738 | 98  | 
if PgipTypes.read_pgipbool s  | 
99  | 
then Unsynchronized.change print_mode (insert (op =) thm_depsN)  | 
|
100  | 
else Unsynchronized.change print_mode (remove (op =) thm_depsN);  | 
|
| 28587 | 101  | 
in  | 
102  | 
mkpref get set PgipTypes.Pgipbool "theorem-dependencies"  | 
|
103  | 
"Track theorem dependencies within Proof General"  | 
|
104  | 
end;  | 
|
| 21637 | 105  | 
|
106  | 
val print_depth_pref =  | 
|
| 28587 | 107  | 
let  | 
108  | 
fun get () = PgipTypes.int_to_pgstring (get_print_depth ());  | 
|
109  | 
val set = print_depth o PgipTypes.read_pgipnat;  | 
|
110  | 
in mkpref get set PgipTypes.Pgipnat "print-depth" "Setting for the ML print depth" end;  | 
|
| 21637 | 111  | 
|
112  | 
||
| 24329 | 113  | 
val display_preferences =  | 
| 
39134
 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 
wenzelm 
parents: 
39128 
diff
changeset
 | 
114  | 
[bool_pref show_types_default  | 
| 28587 | 115  | 
"show-types"  | 
116  | 
"Include types in display of Isabelle terms",  | 
|
| 
39134
 
917b4b6ba3d2
turned show_sorts/show_types into proper configuration options;
 
wenzelm 
parents: 
39128 
diff
changeset
 | 
117  | 
bool_pref show_sorts_default  | 
| 28587 | 118  | 
"show-sorts"  | 
119  | 
"Include sorts in display of Isabelle terms",  | 
|
| 
39050
 
600de0485859
turned show_consts into proper configuration option;
 
wenzelm 
parents: 
38980 
diff
changeset
 | 
120  | 
bool_pref show_consts_default  | 
| 28587 | 121  | 
"show-consts"  | 
122  | 
"Show types of consts in Isabelle goal display",  | 
|
123  | 
bool_pref long_names  | 
|
124  | 
"long-names"  | 
|
125  | 
"Show fully qualified names in Isabelle terms",  | 
|
| 
39137
 
ccb53edd59f0
turned show_brackets into proper configuration option;
 
wenzelm 
parents: 
39134 
diff
changeset
 | 
126  | 
bool_pref show_brackets_default  | 
| 28587 | 127  | 
"show-brackets"  | 
128  | 
"Show full bracketing in Isabelle terms",  | 
|
| 
39125
 
f45d332a90e3
pretty_goals: turned some global references and function arguments into configuration options (goals_limit = 10, goals_total = true, show_main_goal = false) depending on the context;
 
wenzelm 
parents: 
39050 
diff
changeset
 | 
129  | 
bool_pref Goal_Display.show_main_goal_default  | 
| 28587 | 130  | 
"show-main-goal"  | 
131  | 
"Show main goal in proof state display",  | 
|
| 
39128
 
93a7365fb4ee
turned eta_contract into proper configuration option;
 
wenzelm 
parents: 
39125 
diff
changeset
 | 
132  | 
bool_pref Syntax.eta_contract_default  | 
| 28587 | 133  | 
"eta-contract"  | 
134  | 
"Print terms eta-contracted"];  | 
|
| 21637 | 135  | 
|
136  | 
val advanced_display_preferences =  | 
|
| 
39125
 
f45d332a90e3
pretty_goals: turned some global references and function arguments into configuration options (goals_limit = 10, goals_total = true, show_main_goal = false) depending on the context;
 
wenzelm 
parents: 
39050 
diff
changeset
 | 
137  | 
[nat_pref Goal_Display.goals_limit_default  | 
| 28587 | 138  | 
"goals-limit"  | 
139  | 
"Setting for maximum number of goals printed",  | 
|
140  | 
print_depth_pref,  | 
|
| 
38980
 
af73cf0dc31f
turned show_question_marks into proper configuration option;
 
wenzelm 
parents: 
36006 
diff
changeset
 | 
141  | 
bool_pref show_question_marks_default  | 
| 28587 | 142  | 
"show-question-marks"  | 
143  | 
"Show leading question mark of variable name"];  | 
|
| 21637 | 144  | 
|
| 24329 | 145  | 
val tracing_preferences =  | 
| 
36006
 
7ddc33baf959
the configuration option 'trace_simp' now uses the reference of the ProofGeneral settings menu as (dynamic) default value
 
boehmes 
parents: 
35995 
diff
changeset
 | 
146  | 
[bool_pref trace_simp_default  | 
| 
35995
 
26e820d27e0a
re-introduce reference to control simplifier tracing (needed for ProofGeneral settings menu) (cf. 12bb31230550)
 
boehmes 
parents: 
35979 
diff
changeset
 | 
147  | 
"trace-simplifier"  | 
| 
 
26e820d27e0a
re-introduce reference to control simplifier tracing (needed for ProofGeneral settings menu) (cf. 12bb31230550)
 
boehmes 
parents: 
35979 
diff
changeset
 | 
148  | 
"Trace simplification rules.",  | 
| 
 
26e820d27e0a
re-introduce reference to control simplifier tracing (needed for ProofGeneral settings menu) (cf. 12bb31230550)
 
boehmes 
parents: 
35979 
diff
changeset
 | 
149  | 
nat_pref trace_simp_depth_limit  | 
| 28587 | 150  | 
"trace-simplifier-depth"  | 
151  | 
"Trace simplifier depth limit.",  | 
|
152  | 
bool_pref trace_rules  | 
|
153  | 
"trace-rules"  | 
|
154  | 
"Trace application of the standard rules",  | 
|
155  | 
bool_pref Pattern.trace_unify_fail  | 
|
156  | 
"trace-unification"  | 
|
157  | 
"Output error diagnostics during unification",  | 
|
158  | 
bool_pref Output.timing  | 
|
159  | 
"global-timing"  | 
|
160  | 
"Whether to enable timing in Isabelle.",  | 
|
161  | 
bool_pref Toplevel.debug  | 
|
162  | 
"debugging"  | 
|
163  | 
"Whether to enable debugging.",  | 
|
164  | 
thm_deps_pref];  | 
|
| 21637 | 165  | 
|
| 24329 | 166  | 
val proof_preferences =  | 
| 
39616
 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 
wenzelm 
parents: 
39165 
diff
changeset
 | 
167  | 
[Unsynchronized.setmp quick_and_dirty true (fn () =>  | 
| 30980 | 168  | 
bool_pref quick_and_dirty  | 
169  | 
"quick-and-dirty"  | 
|
170  | 
"Take a few short cuts") (),  | 
|
| 28587 | 171  | 
bool_pref Toplevel.skip_proofs  | 
172  | 
"skip-proofs"  | 
|
173  | 
"Skip over proofs (interactive-only)",  | 
|
174  | 
proof_pref,  | 
|
175  | 
nat_pref Multithreading.max_threads  | 
|
176  | 
"max-threads"  | 
|
| 
29435
 
a5f84ac14609
added parallel_proofs flag (default true, cf. usedir option -Q), which can be disabled in low-memory situations;
 
wenzelm 
parents: 
28591 
diff
changeset
 | 
177  | 
"Maximum number of threads",  | 
| 
32061
 
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
 
wenzelm 
parents: 
30985 
diff
changeset
 | 
178  | 
parallel_proof_pref];  | 
| 21637 | 179  | 
|
| 28587 | 180  | 
val pure_preferences =  | 
| 30980 | 181  | 
[(category_display, display_preferences),  | 
182  | 
(category_advanced_display, advanced_display_preferences),  | 
|
183  | 
(category_tracing, tracing_preferences),  | 
|
184  | 
(category_proof, proof_preferences)];  | 
|
| 21637 | 185  | 
|
| 28587 | 186  | 
|
187  | 
(* table of categories and preferences; names must be unique *)  | 
|
188  | 
||
189  | 
type T = (string * preference list) list;  | 
|
| 
22214
 
6e9ab159512f
Add operations on preference tables (remove, set_default).
 
aspinall 
parents: 
22130 
diff
changeset
 | 
190  | 
|
| 28587 | 191  | 
fun remove name (tab: T) = tab |> map  | 
192  | 
(fn (cat, prefs) => (cat, filter_out (curry op = name o #name) prefs));  | 
|
| 
22214
 
6e9ab159512f
Add operations on preference tables (remove, set_default).
 
aspinall 
parents: 
22130 
diff
changeset
 | 
193  | 
|
| 28587 | 194  | 
fun set_default (setname, newdefault) (tab: T) = tab |> map  | 
195  | 
(fn (cat, prefs) =>  | 
|
196  | 
    (cat, prefs |> map (fn (pref as {name, descr, default, pgiptype, get, set}) =>
 | 
|
197  | 
if name = setname then  | 
|
198  | 
(set newdefault;  | 
|
199  | 
          {name =name , descr = descr, default = newdefault,
 | 
|
200  | 
pgiptype = pgiptype, get = get, set = set})  | 
|
201  | 
else pref)));  | 
|
| 
22214
 
6e9ab159512f
Add operations on preference tables (remove, set_default).
 
aspinall 
parents: 
22130 
diff
changeset
 | 
202  | 
|
| 28587 | 203  | 
fun add cname (pref: preference) (tab: T) = tab |> map  | 
204  | 
(fn (cat, prefs) =>  | 
|
205  | 
if cat <> cname then (cat, prefs)  | 
|
206  | 
else  | 
|
207  | 
      if exists (fn {name, ...} => name = #name pref) prefs
 | 
|
| 28591 | 208  | 
      then (warning ("Preference already exists: " ^ quote (#name pref)); (cat, prefs))
 | 
| 30985 | 209  | 
else (cat, prefs @ [pref]));  | 
| 
28066
 
611e504c1191
extended interface to preferences to allow adding new ones
 
nipkow 
parents: 
25440 
diff
changeset
 | 
210  | 
|
| 28587 | 211  | 
end;  |