| author | wenzelm | 
| Sat, 01 Jul 2000 19:45:23 +0200 | |
| changeset 9222 | 92ad2341179d | 
| parent 9219 | 84af672218b9 | 
| child 9273 | 798673f65f02 | 
| permissions | -rw-r--r-- | 
| 5831 | 1  | 
(* Title: Pure/Isar/isar_cmd.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Markus Wenzel, TU Muenchen  | 
|
| 8807 | 4  | 
License: GPL (GNU GENERAL PUBLIC LICENSE)  | 
| 5831 | 5  | 
|
6  | 
Non-logical toplevel commands.  | 
|
7  | 
*)  | 
|
8  | 
||
9  | 
signature ISAR_CMD =  | 
|
10  | 
sig  | 
|
| 7462 | 11  | 
val welcome: Toplevel.transition -> Toplevel.transition  | 
12  | 
val init_toplevel: Toplevel.transition -> Toplevel.transition  | 
|
| 5831 | 13  | 
val exit: Toplevel.transition -> Toplevel.transition  | 
14  | 
val quit: Toplevel.transition -> Toplevel.transition  | 
|
| 7908 | 15  | 
val touch_child_thys: string -> Toplevel.transition -> Toplevel.transition  | 
| 7101 | 16  | 
val touch_all_thys: Toplevel.transition -> Toplevel.transition  | 
17  | 
val touch_thy: string -> Toplevel.transition -> Toplevel.transition  | 
|
18  | 
val remove_thy: string -> Toplevel.transition -> Toplevel.transition  | 
|
| 7931 | 19  | 
val kill_thy: string -> Toplevel.transition -> Toplevel.transition  | 
| 8463 | 20  | 
val pr: string list * int option -> Toplevel.transition -> Toplevel.transition  | 
| 8453 | 21  | 
val disable_pr: Toplevel.transition -> Toplevel.transition  | 
22  | 
val enable_pr: Toplevel.transition -> Toplevel.transition  | 
|
| 6734 | 23  | 
val cannot_undo: string -> Toplevel.transition -> Toplevel.transition  | 
| 7729 | 24  | 
val clear_undos_theory: int -> Toplevel.transition -> Toplevel.transition  | 
| 6742 | 25  | 
val redo: Toplevel.transition -> Toplevel.transition  | 
26  | 
val undos_proof: int -> Toplevel.transition -> Toplevel.transition  | 
|
| 7936 | 27  | 
val kill_proof_notify: (unit -> unit) -> Toplevel.transition -> Toplevel.transition  | 
| 6742 | 28  | 
val kill_proof: Toplevel.transition -> Toplevel.transition  | 
29  | 
val undo_theory: Toplevel.transition -> Toplevel.transition  | 
|
| 6686 | 30  | 
val undo: Toplevel.transition -> Toplevel.transition  | 
| 8498 | 31  | 
val kill: Toplevel.transition -> Toplevel.transition  | 
| 5831 | 32  | 
val use: string -> Toplevel.transition -> Toplevel.transition  | 
33  | 
val use_mltext_theory: string -> Toplevel.transition -> Toplevel.transition  | 
|
| 7891 | 34  | 
val use_mltext: bool -> string -> Toplevel.transition -> Toplevel.transition  | 
| 5831 | 35  | 
val use_commit: Toplevel.transition -> Toplevel.transition  | 
36  | 
val cd: string -> Toplevel.transition -> Toplevel.transition  | 
|
37  | 
val pwd: Toplevel.transition -> Toplevel.transition  | 
|
38  | 
val use_thy: string -> Toplevel.transition -> Toplevel.transition  | 
|
| 6694 | 39  | 
val use_thy_only: string -> Toplevel.transition -> Toplevel.transition  | 
| 6195 | 40  | 
val update_thy: string -> Toplevel.transition -> Toplevel.transition  | 
| 7101 | 41  | 
val update_thy_only: string -> Toplevel.transition -> Toplevel.transition  | 
| 7124 | 42  | 
val pretty_setmargin: int -> Toplevel.transition -> Toplevel.transition  | 
| 7308 | 43  | 
val print_context: Toplevel.transition -> Toplevel.transition  | 
| 5831 | 44  | 
val print_theory: Toplevel.transition -> Toplevel.transition  | 
45  | 
val print_syntax: Toplevel.transition -> Toplevel.transition  | 
|
| 5880 | 46  | 
val print_theorems: Toplevel.transition -> Toplevel.transition  | 
| 5831 | 47  | 
val print_attributes: Toplevel.transition -> Toplevel.transition  | 
| 9219 | 48  | 
val print_trans_rules: Toplevel.transition -> Toplevel.transition  | 
| 5831 | 49  | 
val print_methods: Toplevel.transition -> Toplevel.transition  | 
| 9219 | 50  | 
val print_antiquotations: Toplevel.transition -> Toplevel.transition  | 
| 7615 | 51  | 
val print_thms_containing: xstring list -> Toplevel.transition -> Toplevel.transition  | 
| 5831 | 52  | 
val print_binds: Toplevel.transition -> Toplevel.transition  | 
53  | 
val print_lthms: Toplevel.transition -> Toplevel.transition  | 
|
| 8369 | 54  | 
val print_cases: Toplevel.transition -> Toplevel.transition  | 
| 8463 | 55  | 
val print_thms: string list * (string * Args.src list) list  | 
56  | 
-> Toplevel.transition -> Toplevel.transition  | 
|
57  | 
val print_prop: string list * string -> Toplevel.transition -> Toplevel.transition  | 
|
58  | 
val print_term: string list * string -> Toplevel.transition -> Toplevel.transition  | 
|
59  | 
val print_type: string list * string -> Toplevel.transition -> Toplevel.transition  | 
|
| 5831 | 60  | 
end;  | 
61  | 
||
62  | 
structure IsarCmd: ISAR_CMD =  | 
|
63  | 
struct  | 
|
64  | 
||
65  | 
||
| 7462 | 66  | 
(* variations on init / exit *)  | 
67  | 
||
68  | 
val init_toplevel = Toplevel.imperative (fn () => raise Toplevel.RESTART);  | 
|
69  | 
val welcome = Toplevel.imperative (writeln o Session.welcome);  | 
|
| 5831 | 70  | 
|
71  | 
val exit = Toplevel.keep (fn state =>  | 
|
72  | 
(Context.set_context (try Toplevel.theory_of state);  | 
|
| 6635 | 73  | 
writeln "Leaving the Isar loop. Invoke 'loop();' to restart.";  | 
| 5831 | 74  | 
raise Toplevel.TERMINATE));  | 
75  | 
||
76  | 
val quit = Toplevel.imperative quit;  | 
|
77  | 
||
| 7101 | 78  | 
|
79  | 
(* touch theories *)  | 
|
80  | 
||
| 7908 | 81  | 
fun touch_child_thys name = Toplevel.imperative (fn () => ThyInfo.touch_child_thys name);  | 
| 7101 | 82  | 
val touch_all_thys = Toplevel.imperative ThyInfo.touch_all_thys;  | 
83  | 
fun touch_thy name = Toplevel.imperative (fn () => ThyInfo.touch_thy name);  | 
|
84  | 
fun remove_thy name = Toplevel.imperative (fn () => ThyInfo.remove_thy name);  | 
|
| 7931 | 85  | 
fun kill_thy name = Toplevel.imperative (fn () => IsarThy.kill_theory name);  | 
| 7101 | 86  | 
|
| 5831 | 87  | 
|
| 8453 | 88  | 
(* print state *)  | 
89  | 
||
| 8463 | 90  | 
fun with_modes modes e =  | 
91  | 
Library.setmp print_mode (modes @ ! print_mode) e ();  | 
|
92  | 
||
93  | 
fun pr (ms, limit) = Toplevel.keep (fn state =>  | 
|
94  | 
((case limit of Some n => goals_limit := n | None => ()); Toplevel.quiet := false;  | 
|
| 8886 | 95  | 
with_modes ms (fn () => Toplevel.print_state true state)));  | 
| 8453 | 96  | 
|
97  | 
val disable_pr = Toplevel.imperative (fn () => Toplevel.quiet := true);  | 
|
98  | 
val enable_pr = Toplevel.imperative (fn () => Toplevel.quiet := false);  | 
|
99  | 
||
100  | 
||
| 6686 | 101  | 
(* history commands *)  | 
102  | 
||
| 6734 | 103  | 
fun cannot_undo txt = Toplevel.imperative (fn () => error ("Cannot undo " ^ quote txt));
 | 
| 7729 | 104  | 
val clear_undos_theory = Toplevel.history o History.clear;  | 
| 6686 | 105  | 
val redo = Toplevel.history History.redo o Toplevel.proof ProofHistory.redo;  | 
106  | 
||
| 6742 | 107  | 
fun undos_proof n = Toplevel.proof (fn prf =>  | 
108  | 
if ProofHistory.is_initial prf then raise Toplevel.UNDEF else funpow n ProofHistory.undo prf);  | 
|
| 6686 | 109  | 
|
| 7936 | 110  | 
fun kill_proof_notify (f: unit -> unit) = Toplevel.history (fn hist =>  | 
| 6742 | 111  | 
(case History.current hist of  | 
112  | 
Toplevel.Theory _ => raise Toplevel.UNDEF  | 
|
| 7936 | 113  | 
| Toplevel.Proof _ => (f (); History.undo hist)));  | 
114  | 
||
115  | 
val kill_proof = kill_proof_notify (K ());  | 
|
| 6686 | 116  | 
|
| 6742 | 117  | 
val undo_theory = Toplevel.history (fn hist =>  | 
118  | 
if History.is_initial hist then raise Toplevel.UNDEF else History.undo hist);  | 
|
| 6686 | 119  | 
|
| 7413 | 120  | 
val undo = Toplevel.kill o undo_theory o undos_proof 1;  | 
| 8498 | 121  | 
val kill = Toplevel.kill o kill_proof;  | 
| 6686 | 122  | 
|
123  | 
||
| 5831 | 124  | 
(* use ML text *)  | 
125  | 
||
| 6264 | 126  | 
fun use name = Toplevel.keep (fn state =>  | 
127  | 
Context.setmp (try Toplevel.theory_of state) ThyInfo.use name);  | 
|
| 5831 | 128  | 
|
129  | 
(*passes changes of theory context*)  | 
|
| 
8349
 
611342539639
moved use_mltext, use_mltext_theory, use_let, use_setup to context.ML;
 
wenzelm 
parents: 
7936 
diff
changeset
 | 
130  | 
val use_mltext_theory = Toplevel.theory' o Context.use_mltext_theory;  | 
| 5831 | 131  | 
|
132  | 
(*ignore result theory context*)  | 
|
| 7891 | 133  | 
fun use_mltext v txt = Toplevel.keep' (fn verb => fn state =>  | 
| 
8349
 
611342539639
moved use_mltext, use_mltext_theory, use_let, use_setup to context.ML;
 
wenzelm 
parents: 
7936 
diff
changeset
 | 
134  | 
(Context.use_mltext txt (v andalso verb) (try Toplevel.theory_of state)));  | 
| 5831 | 135  | 
|
136  | 
(*Note: commit is dynamically bound*)  | 
|
| 7891 | 137  | 
val use_commit = use_mltext false "commit();";  | 
| 5831 | 138  | 
|
139  | 
||
140  | 
(* current working directory *)  | 
|
141  | 
||
| 7790 | 142  | 
fun cd dir = Toplevel.imperative (fn () => (File.cd (Path.unpack dir)));  | 
143  | 
val pwd = Toplevel.imperative (fn () => writeln (Path.pack (File.pwd ())));  | 
|
| 5831 | 144  | 
|
145  | 
||
146  | 
(* load theory files *)  | 
|
147  | 
||
| 6243 | 148  | 
fun use_thy name = Toplevel.imperative (fn () => Context.save ThyInfo.use_thy name);  | 
| 6694 | 149  | 
fun use_thy_only name = Toplevel.imperative (fn () => Context.save ThyInfo.use_thy_only name);  | 
| 6243 | 150  | 
fun update_thy name = Toplevel.imperative (fn () => Context.save ThyInfo.update_thy name);  | 
| 7124 | 151  | 
fun update_thy_only name =  | 
152  | 
Toplevel.imperative (fn () => Context.save ThyInfo.update_thy_only name);  | 
|
153  | 
||
154  | 
||
155  | 
(* pretty_setmargin *)  | 
|
156  | 
||
157  | 
fun pretty_setmargin n = Toplevel.imperative (fn () => Pretty.setmargin n);  | 
|
| 5831 | 158  | 
|
159  | 
||
| 7308 | 160  | 
(* print theory *)  | 
| 5831 | 161  | 
|
| 7308 | 162  | 
val print_context = Toplevel.keep Toplevel.print_state_context;  | 
| 5831 | 163  | 
val print_theory = Toplevel.keep (PureThy.print_theory o Toplevel.theory_of);  | 
164  | 
val print_syntax = Toplevel.keep (Display.print_syntax o Toplevel.theory_of);  | 
|
| 5880 | 165  | 
val print_theorems = Toplevel.keep (PureThy.print_theorems o Toplevel.theory_of);  | 
| 5831 | 166  | 
val print_attributes = Toplevel.keep (Attrib.print_attributes o Toplevel.theory_of);  | 
| 9219 | 167  | 
val print_trans_rules = Toplevel.keep (Toplevel.node_case Calculation.print_global_rules  | 
168  | 
(Calculation.print_local_rules o Proof.context_of));  | 
|
| 5831 | 169  | 
val print_methods = Toplevel.keep (Method.print_methods o Toplevel.theory_of);  | 
| 9219 | 170  | 
val print_antiquotations = Toplevel.imperative IsarOutput.print_antiquotations;  | 
| 5831 | 171  | 
|
| 7615 | 172  | 
fun print_thms_containing cs = Toplevel.keep (fn state =>  | 
173  | 
ThmDatabase.print_thms_containing (Toplevel.theory_of state) cs);  | 
|
174  | 
||
| 5831 | 175  | 
|
176  | 
(* print proof context contents *)  | 
|
177  | 
||
178  | 
val print_binds = Toplevel.keep (ProofContext.print_binds o Proof.context_of o Toplevel.proof_of);  | 
|
179  | 
val print_lthms = Toplevel.keep (ProofContext.print_thms o Proof.context_of o Toplevel.proof_of);  | 
|
| 8369 | 180  | 
val print_cases = Toplevel.keep (ProofContext.print_cases o Proof.context_of o Toplevel.proof_of);  | 
| 5831 | 181  | 
|
182  | 
||
| 7012 | 183  | 
(* print theorems / types / terms / props *)  | 
| 5880 | 184  | 
|
| 9128 | 185  | 
fun string_of_thms state ms args = with_modes ms (fn () =>  | 
186  | 
Pretty.string_of (Proof.pretty_thms (IsarThy.get_thmss args state)));  | 
|
| 5895 | 187  | 
|
| 9128 | 188  | 
fun string_of_prop state ms s =  | 
| 5831 | 189  | 
let  | 
| 9128 | 190  | 
val sign = Proof.sign_of state;  | 
191  | 
val prop = ProofContext.read_prop (Proof.context_of state) s;  | 
|
192  | 
in with_modes ms (fn () => Pretty.string_of (Pretty.quote (Sign.pretty_term sign prop))) end;  | 
|
| 5831 | 193  | 
|
| 9128 | 194  | 
fun string_of_term state ms s =  | 
| 5831 | 195  | 
let  | 
| 9128 | 196  | 
val sign = Proof.sign_of state;  | 
197  | 
val t = ProofContext.read_term (Proof.context_of state) s;  | 
|
| 5831 | 198  | 
val T = Term.type_of t;  | 
199  | 
in  | 
|
| 9128 | 200  | 
with_modes ms (fn () => Pretty.string_of  | 
| 8463 | 201  | 
(Pretty.block [Pretty.quote (Sign.pretty_term sign t), Pretty.fbrk,  | 
202  | 
Pretty.str "::", Pretty.brk 1, Pretty.quote (Sign.pretty_typ sign T)]))  | 
|
| 9128 | 203  | 
end;  | 
| 5831 | 204  | 
|
| 9128 | 205  | 
fun string_of_type state ms s =  | 
| 5831 | 206  | 
let  | 
| 9128 | 207  | 
val sign = Proof.sign_of state;  | 
208  | 
val T = ProofContext.read_typ (Proof.context_of state) s;  | 
|
209  | 
in with_modes ms (fn () => Pretty.string_of (Pretty.quote (Sign.pretty_typ sign T))) end;  | 
|
210  | 
||
211  | 
fun print_item string_of (x, y) = Toplevel.keep (fn state =>  | 
|
212  | 
writeln (string_of (Toplevel.node_case Proof.init_state Proof.enter_forward state) x y));  | 
|
213  | 
||
214  | 
val print_thms = print_item string_of_thms;  | 
|
215  | 
val print_prop = print_item string_of_prop;  | 
|
216  | 
val print_term = print_item string_of_term;  | 
|
217  | 
val print_type = print_item string_of_type;  | 
|
| 5831 | 218  | 
|
219  | 
||
220  | 
end;  |