| author | paulson | 
| Tue, 23 May 2000 12:30:29 +0200 | |
| changeset 8927 | 1cf815412d78 | 
| parent 8897 | fb1436ca3b2e | 
| child 8966 | 916966f68907 | 
| permissions | -rw-r--r-- | 
| 5830 | 1  | 
(* Title: Pure/Isar/isar_thy.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Markus Wenzel, TU Muenchen  | 
|
| 8804 | 4  | 
License: GPL (GNU GENERAL PUBLIC LICENSE)  | 
| 5830 | 5  | 
|
| 6371 | 6  | 
Pure/Isar derived theory operations.  | 
| 5830 | 7  | 
*)  | 
8  | 
||
9  | 
signature ISAR_THY =  | 
|
10  | 
sig  | 
|
| 7734 | 11  | 
val add_header: Comment.text -> Toplevel.transition -> Toplevel.transition  | 
| 6552 | 12  | 
val add_chapter: Comment.text -> theory -> theory  | 
13  | 
val add_section: Comment.text -> theory -> theory  | 
|
14  | 
val add_subsection: Comment.text -> theory -> theory  | 
|
15  | 
val add_subsubsection: Comment.text -> theory -> theory  | 
|
| 7172 | 16  | 
val add_text: Comment.text -> theory -> theory  | 
| 7862 | 17  | 
val add_text_raw: Comment.text -> theory -> theory  | 
| 7172 | 18  | 
val add_sect: Comment.text -> ProofHistory.T -> ProofHistory.T  | 
19  | 
val add_subsect: Comment.text -> ProofHistory.T -> ProofHistory.T  | 
|
20  | 
val add_subsubsect: Comment.text -> ProofHistory.T -> ProofHistory.T  | 
|
21  | 
val add_txt: Comment.text -> ProofHistory.T -> ProofHistory.T  | 
|
| 7862 | 22  | 
val add_txt_raw: Comment.text -> ProofHistory.T -> ProofHistory.T  | 
| 8724 | 23  | 
val global_path: Comment.text -> theory -> theory  | 
24  | 
val local_path: Comment.text -> theory -> theory  | 
|
25  | 
val hide_space: (string * xstring list) * Comment.text -> theory -> theory  | 
|
26  | 
val hide_space_i: (string * string list) * Comment.text -> theory -> theory  | 
|
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
27  | 
val add_classes: (bclass * xclass list) list * Comment.text -> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
28  | 
val add_classes_i: (bclass * class list) list * Comment.text -> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
29  | 
val add_classrel: (xclass * xclass) * Comment.text -> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
30  | 
val add_classrel_i: (class * class) * Comment.text -> theory -> theory  | 
| 8897 | 31  | 
val add_defsort: string * Comment.text -> theory -> theory  | 
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
32  | 
val add_defsort_i: sort * Comment.text -> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
33  | 
val add_nonterminals: (bstring * Comment.text) list -> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
34  | 
val add_tyabbrs: ((bstring * string list * string * mixfix) * Comment.text) list  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
35  | 
-> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
36  | 
val add_tyabbrs_i: ((bstring * string list * typ * mixfix) * Comment.text) list  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
37  | 
-> theory -> theory  | 
| 8897 | 38  | 
val add_arities: ((xstring * string list * string) * Comment.text) list -> theory -> theory  | 
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
39  | 
val add_arities_i: ((string * sort list * sort) * Comment.text) list -> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
40  | 
val add_typedecl: (bstring * string list * mixfix) * Comment.text -> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
41  | 
val add_consts: ((bstring * string * mixfix) * Comment.text) list -> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
42  | 
val add_consts_i: ((bstring * typ * mixfix) * Comment.text) list -> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
43  | 
val add_modesyntax: string * bool -> ((bstring * string * mixfix) * Comment.text) list  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
44  | 
-> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
45  | 
val add_modesyntax_i: string * bool -> ((bstring * typ * mixfix) * Comment.text) list  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
46  | 
-> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
47  | 
val add_trrules: ((xstring * string) Syntax.trrule * Comment.text) list -> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
48  | 
val add_trrules_i: (Syntax.ast Syntax.trrule * Comment.text) list -> theory -> theory  | 
| 8227 | 49  | 
val add_judgment: (bstring * string * mixfix) * Comment.text -> theory -> theory  | 
50  | 
val add_judgment_i: (bstring * typ * mixfix) * Comment.text -> theory -> theory  | 
|
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
51  | 
val add_axioms: (((bstring * string) * Args.src list) * Comment.text) list -> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
52  | 
val add_axioms_i: (((bstring * term) * theory attribute list) * Comment.text) list  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
53  | 
-> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
54  | 
val add_defs: (((bstring * string) * Args.src list) * Comment.text) list -> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
55  | 
val add_defs_i: (((bstring * term) * theory attribute list) * Comment.text) list  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
56  | 
-> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
57  | 
val add_constdefs: (((bstring * string * mixfix) * Comment.text) * (string * Comment.text)) list  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
58  | 
-> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
59  | 
val add_constdefs_i: (((bstring * typ * mixfix) * Comment.text) * (term * Comment.text)) list  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
60  | 
-> theory -> theory  | 
| 6371 | 61  | 
val apply_theorems: (xstring * Args.src list) list -> theory -> theory * thm list  | 
62  | 
val apply_theorems_i: (thm * theory attribute list) list -> theory -> theory * thm list  | 
|
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
63  | 
val have_theorems: ((bstring * Args.src list) * (xstring * Args.src list) list) * Comment.text  | 
| 6371 | 64  | 
-> theory -> theory  | 
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
65  | 
val have_theorems_i: ((bstring * theory attribute list) * (thm * theory attribute list) list)  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
66  | 
* Comment.text -> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
67  | 
val have_lemmas: ((bstring * Args.src list) * (xstring * Args.src list) list) * Comment.text  | 
| 5915 | 68  | 
-> theory -> theory  | 
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
69  | 
val have_lemmas_i: ((bstring * theory attribute list) * (thm * theory attribute list) list)  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
70  | 
* Comment.text -> theory -> theory  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
71  | 
val have_facts: ((string * Args.src list) * (string * Args.src list) list) * Comment.text  | 
| 5915 | 72  | 
-> ProofHistory.T -> ProofHistory.T  | 
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
73  | 
val have_facts_i: ((string * Proof.context attribute list) *  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
74  | 
(thm * Proof.context attribute list) list) * Comment.text -> ProofHistory.T -> ProofHistory.T  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
75  | 
val from_facts: (string * Args.src list) list * Comment.text -> ProofHistory.T -> ProofHistory.T  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
76  | 
val from_facts_i: (thm * Proof.context attribute list) list * Comment.text  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
77  | 
-> ProofHistory.T -> ProofHistory.T  | 
| 6879 | 78  | 
val with_facts: (string * Args.src list) list * Comment.text -> ProofHistory.T -> ProofHistory.T  | 
79  | 
val with_facts_i: (thm * Proof.context attribute list) list * Comment.text  | 
|
80  | 
-> ProofHistory.T -> ProofHistory.T  | 
|
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
81  | 
val chain: Comment.text -> ProofHistory.T -> ProofHistory.T  | 
| 7417 | 82  | 
val fix: ((string list * string option) * Comment.text) list -> ProofHistory.T -> ProofHistory.T  | 
| 7666 | 83  | 
val fix_i: ((string list * typ option) * Comment.text) list -> ProofHistory.T -> ProofHistory.T  | 
| 8615 | 84  | 
val let_bind: ((string list * string) * Comment.text) list -> ProofHistory.T -> ProofHistory.T  | 
85  | 
val let_bind_i: ((term list * term) * Comment.text) list -> ProofHistory.T -> ProofHistory.T  | 
|
| 8450 | 86  | 
val invoke_case: (string * Args.src list) * Comment.text -> ProofHistory.T -> ProofHistory.T  | 
87  | 
val invoke_case_i: (string * Proof.context attribute list) * Comment.text  | 
|
88  | 
-> ProofHistory.T -> ProofHistory.T  | 
|
| 6937 | 89  | 
val theorem: (bstring * Args.src list * (string * (string list * string list)))  | 
90  | 
* Comment.text -> bool -> theory -> ProofHistory.T  | 
|
91  | 
val theorem_i: (bstring * theory attribute list * (term * (term list * term list)))  | 
|
92  | 
* Comment.text -> bool -> theory -> ProofHistory.T  | 
|
93  | 
val lemma: (bstring * Args.src list * (string * (string list * string list)))  | 
|
94  | 
* Comment.text -> bool -> theory -> ProofHistory.T  | 
|
95  | 
val lemma_i: (bstring * theory attribute list * (term * (term list * term list)))  | 
|
96  | 
* Comment.text -> bool -> theory -> ProofHistory.T  | 
|
| 7271 | 97  | 
val assume: ((string * Args.src list * (string * (string list * string list)) list)  | 
98  | 
* Comment.text) list -> ProofHistory.T -> ProofHistory.T  | 
|
99  | 
val assume_i: ((string * Proof.context attribute list * (term * (term list * term list)) list)  | 
|
100  | 
* Comment.text) list -> ProofHistory.T -> ProofHistory.T  | 
|
101  | 
val presume: ((string * Args.src list * (string * (string list * string list)) list)  | 
|
102  | 
* Comment.text) list -> ProofHistory.T -> ProofHistory.T  | 
|
103  | 
val presume_i: ((string * Proof.context attribute list * (term * (term list * term list)) list)  | 
|
104  | 
* Comment.text) list -> ProofHistory.T -> ProofHistory.T  | 
|
| 6952 | 105  | 
val local_def: (string * Args.src list * ((string * string option) * (string * string list)))  | 
106  | 
* Comment.text -> ProofHistory.T -> ProofHistory.T  | 
|
| 7666 | 107  | 
val local_def_i: (string * Args.src list * ((string * typ option) * (term * term list)))  | 
| 6952 | 108  | 
* Comment.text -> ProofHistory.T -> ProofHistory.T  | 
| 6937 | 109  | 
val show: (string * Args.src list * (string * (string list * string list)))  | 
110  | 
* Comment.text -> ProofHistory.T -> ProofHistory.T  | 
|
111  | 
val show_i: (string * Proof.context attribute list * (term * (term list * term list)))  | 
|
112  | 
* Comment.text -> ProofHistory.T -> ProofHistory.T  | 
|
113  | 
val have: (string * Args.src list * (string * (string list * string list)))  | 
|
114  | 
* Comment.text -> ProofHistory.T -> ProofHistory.T  | 
|
115  | 
val have_i: (string * Proof.context attribute list * (term * (term list * term list)))  | 
|
116  | 
* Comment.text -> ProofHistory.T -> ProofHistory.T  | 
|
117  | 
val thus: (string * Args.src list * (string * (string list * string list)))  | 
|
118  | 
* Comment.text -> ProofHistory.T -> ProofHistory.T  | 
|
119  | 
val thus_i: (string * Proof.context attribute list * (term * (term list * term list)))  | 
|
120  | 
* Comment.text -> ProofHistory.T -> ProofHistory.T  | 
|
121  | 
val hence: (string * Args.src list * (string * (string list * string list)))  | 
|
122  | 
* Comment.text -> ProofHistory.T -> ProofHistory.T  | 
|
123  | 
val hence_i: (string * Proof.context attribute list * (term * (term list * term list)))  | 
|
124  | 
* Comment.text -> ProofHistory.T -> ProofHistory.T  | 
|
| 5830 | 125  | 
val begin_block: ProofHistory.T -> ProofHistory.T  | 
126  | 
val next_block: ProofHistory.T -> ProofHistory.T  | 
|
127  | 
val end_block: ProofHistory.T -> ProofHistory.T  | 
|
| 8681 | 128  | 
val defer: int option * Comment.text -> ProofHistory.T -> ProofHistory.T  | 
129  | 
val prefer: int * Comment.text -> ProofHistory.T -> ProofHistory.T  | 
|
130  | 
val apply: Method.text * Comment.text -> ProofHistory.T -> ProofHistory.T  | 
|
131  | 
val apply_end: Method.text * Comment.text -> ProofHistory.T -> ProofHistory.T  | 
|
| 7002 | 132  | 
val proof: (Comment.interest * (Method.text * Comment.interest) option) * Comment.text  | 
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
133  | 
-> ProofHistory.T -> ProofHistory.T  | 
| 7002 | 134  | 
val qed: (Method.text * Comment.interest) option * Comment.text  | 
| 6937 | 135  | 
-> Toplevel.transition -> Toplevel.transition  | 
| 7002 | 136  | 
val terminal_proof: ((Method.text * Comment.interest) * (Method.text * Comment.interest) option)  | 
137  | 
* Comment.text -> Toplevel.transition -> Toplevel.transition  | 
|
138  | 
val immediate_proof: Comment.text -> Toplevel.transition -> Toplevel.transition  | 
|
139  | 
val default_proof: Comment.text -> Toplevel.transition -> Toplevel.transition  | 
|
140  | 
val skip_proof: Comment.text -> Toplevel.transition -> Toplevel.transition  | 
|
| 8210 | 141  | 
val forget_proof: Comment.text -> Toplevel.transition -> Toplevel.transition  | 
| 7012 | 142  | 
val get_thmss: (string * Args.src list) list -> Proof.state -> thm list  | 
| 6879 | 143  | 
val also: ((string * Args.src list) list * Comment.interest) option * Comment.text  | 
144  | 
-> Toplevel.transition -> Toplevel.transition  | 
|
145  | 
val also_i: (thm list * Comment.interest) option * Comment.text  | 
|
146  | 
-> Toplevel.transition -> Toplevel.transition  | 
|
147  | 
val finally: ((string * Args.src list) list * Comment.interest) option * Comment.text  | 
|
148  | 
-> Toplevel.transition -> Toplevel.transition  | 
|
149  | 
val finally_i: (thm list * Comment.interest) option * Comment.text  | 
|
150  | 
-> Toplevel.transition -> Toplevel.transition  | 
|
| 8562 | 151  | 
val moreover: Comment.text -> Toplevel.transition -> Toplevel.transition  | 
| 8588 | 152  | 
val ultimately: Comment.text -> Toplevel.transition -> Toplevel.transition  | 
| 5830 | 153  | 
val parse_ast_translation: string -> theory -> theory  | 
154  | 
val parse_translation: string -> theory -> theory  | 
|
155  | 
val print_translation: string -> theory -> theory  | 
|
156  | 
val typed_print_translation: string -> theory -> theory  | 
|
157  | 
val print_ast_translation: string -> theory -> theory  | 
|
158  | 
val token_translation: string -> theory -> theory  | 
|
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
159  | 
val add_oracle: (bstring * string) * Comment.text -> theory -> theory  | 
| 6331 | 160  | 
val begin_theory: string -> string list -> (string * bool) list -> theory  | 
161  | 
val end_theory: theory -> theory  | 
|
| 7932 | 162  | 
val kill_theory: string -> unit  | 
| 6246 | 163  | 
val theory: string * string list * (string * bool) list  | 
164  | 
-> Toplevel.transition -> Toplevel.transition  | 
|
| 7960 | 165  | 
  val init_context: ('a -> unit) -> 'a -> Toplevel.transition -> Toplevel.transition
 | 
| 6246 | 166  | 
val context: string -> Toplevel.transition -> Toplevel.transition  | 
| 5830 | 167  | 
end;  | 
168  | 
||
169  | 
structure IsarThy: ISAR_THY =  | 
|
170  | 
struct  | 
|
171  | 
||
172  | 
||
173  | 
(** derived theory and proof operations **)  | 
|
174  | 
||
| 8090 | 175  | 
(* markup commands *)  | 
| 5959 | 176  | 
|
| 7734 | 177  | 
fun add_header comment =  | 
178  | 
Toplevel.keep (fn state => if Toplevel.is_toplevel state then () else raise Toplevel.UNDEF);  | 
|
179  | 
||
| 8090 | 180  | 
fun add_text _ x = x;  | 
| 8079 | 181  | 
fun add_text_raw _ x = x;  | 
| 7633 | 182  | 
|
| 8090 | 183  | 
fun add_heading add present txt thy =  | 
| 7633 | 184  | 
(Context.setmp (Some thy) present (Comment.string_of txt); add txt thy);  | 
185  | 
||
| 8090 | 186  | 
val add_chapter = add_heading add_text Present.section;  | 
187  | 
val add_section = add_heading add_text Present.section;  | 
|
188  | 
val add_subsection = add_heading add_text Present.subsection;  | 
|
189  | 
val add_subsubsection = add_heading add_text Present.subsubsection;  | 
|
| 5959 | 190  | 
|
| 8090 | 191  | 
fun add_txt (_: Comment.text) = ProofHistory.apply I;  | 
192  | 
val add_txt_raw = add_txt;  | 
|
| 7332 | 193  | 
val add_sect = add_txt;  | 
194  | 
val add_subsect = add_txt;  | 
|
195  | 
val add_subsubsect = add_txt;  | 
|
| 7172 | 196  | 
|
| 5959 | 197  | 
|
| 8724 | 198  | 
(* name spaces *)  | 
199  | 
||
200  | 
fun global_path comment_ignore = PureThy.global_path;  | 
|
201  | 
fun local_path comment_ignore = PureThy.local_path;  | 
|
202  | 
||
| 8885 | 203  | 
local  | 
204  | 
||
205  | 
val kinds =  | 
|
206  | 
[(Sign.classK, fn sg => fn c => c mem Sign.classes sg),  | 
|
207  | 
(Sign.typeK, fn sg => fn c =>  | 
|
208  | 
    let val {tycons, abbrs, ...} = Type.rep_tsig (Sign.tsig_of sg)
 | 
|
209  | 
in is_some (Symtab.lookup (tycons, c)) orelse is_some (Symtab.lookup (abbrs, c)) end),  | 
|
210  | 
(Sign.constK, is_some oo Sign.const_type)];  | 
|
| 8724 | 211  | 
|
| 8885 | 212  | 
fun gen_hide intern ((kind, xnames), comment_ignore) thy =  | 
213  | 
(case assoc (kinds, kind) of  | 
|
214  | 
Some check =>  | 
|
215  | 
let  | 
|
216  | 
val sg = Theory.sign_of thy;  | 
|
217  | 
val names = map (intern sg kind) xnames;  | 
|
218  | 
val bads = filter_out (check sg) names;  | 
|
219  | 
in  | 
|
220  | 
if null bads then Theory.hide_space_i (kind, names) thy  | 
|
221  | 
        else error ("Attempt to hide undeclared item(s): " ^ commas_quote bads)
 | 
|
222  | 
end  | 
|
223  | 
  | None => error ("Bad name space specification: " ^ quote kind));
 | 
|
224  | 
||
225  | 
in  | 
|
226  | 
||
227  | 
fun hide_space x = gen_hide Sign.intern x;  | 
|
228  | 
fun hide_space_i x = gen_hide (K (K I)) x;  | 
|
229  | 
||
230  | 
end;  | 
|
| 8724 | 231  | 
|
232  | 
||
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
233  | 
(* signature and syntax *)  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
234  | 
|
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
235  | 
val add_classes = Theory.add_classes o Comment.ignore;  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
236  | 
val add_classes_i = Theory.add_classes_i o Comment.ignore;  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
237  | 
val add_classrel = Theory.add_classrel o single o Comment.ignore;  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
238  | 
val add_classrel_i = Theory.add_classrel_i o single o Comment.ignore;  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
239  | 
val add_defsort = Theory.add_defsort o Comment.ignore;  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
240  | 
val add_defsort_i = Theory.add_defsort_i o Comment.ignore;  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
241  | 
val add_nonterminals = Theory.add_nonterminals o map Comment.ignore;  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
242  | 
val add_tyabbrs = Theory.add_tyabbrs o map Comment.ignore;  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
243  | 
val add_tyabbrs_i = Theory.add_tyabbrs_i o map Comment.ignore;  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
244  | 
val add_arities = Theory.add_arities o map Comment.ignore;  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
245  | 
val add_arities_i = Theory.add_arities_i o map Comment.ignore;  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
246  | 
val add_typedecl = PureThy.add_typedecls o single o Comment.ignore;  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
247  | 
val add_consts = Theory.add_consts o map Comment.ignore;  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
248  | 
val add_consts_i = Theory.add_consts_i o map Comment.ignore;  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
249  | 
fun add_modesyntax mode = Theory.add_modesyntax mode o map Comment.ignore;  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
250  | 
fun add_modesyntax_i mode = Theory.add_modesyntax_i mode o map Comment.ignore;  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
251  | 
val add_trrules = Theory.add_trrules o map Comment.ignore;  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
252  | 
val add_trrules_i = Theory.add_trrules_i o map Comment.ignore;  | 
| 8227 | 253  | 
val add_judgment = AutoBind.add_judgment o Comment.ignore;  | 
254  | 
val add_judgment_i = AutoBind.add_judgment_i o Comment.ignore;  | 
|
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
255  | 
|
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
256  | 
|
| 5830 | 257  | 
(* axioms and defs *)  | 
258  | 
||
| 5915 | 259  | 
fun add_axms f args thy =  | 
260  | 
f (map (fn (x, srcs) => (x, map (Attrib.global_attribute thy) srcs)) args) thy;  | 
|
261  | 
||
| 8428 | 262  | 
val add_axioms = add_axms (#1 oo PureThy.add_axioms) o map Comment.ignore;  | 
263  | 
val add_axioms_i = (#1 oo PureThy.add_axioms_i) o map Comment.ignore;  | 
|
264  | 
val add_defs = add_axms (#1 oo PureThy.add_defs) o map Comment.ignore;  | 
|
265  | 
val add_defs_i = (#1 oo PureThy.add_defs_i) o map Comment.ignore;  | 
|
| 6371 | 266  | 
|
267  | 
||
268  | 
(* constdefs *)  | 
|
269  | 
||
270  | 
fun gen_add_constdefs consts defs args thy =  | 
|
271  | 
thy  | 
|
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
272  | 
|> consts (map (Comment.ignore o fst) args)  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
273  | 
|> defs (map (fn (((c, _, mx), _), (s, _)) =>  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
274  | 
(((Thm.def_name (Syntax.const_name c mx), s), []), Comment.none)) args);  | 
| 6371 | 275  | 
|
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
276  | 
fun add_constdefs args = gen_add_constdefs Theory.add_consts add_defs args;  | 
| 
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
277  | 
fun add_constdefs_i args = gen_add_constdefs Theory.add_consts_i add_defs_i args;  | 
| 5915 | 278  | 
|
279  | 
||
280  | 
(* theorems *)  | 
|
281  | 
||
282  | 
fun gen_have_thmss get attrib f ((name, more_srcs), th_srcs) x =  | 
|
283  | 
f name (map (attrib x) more_srcs)  | 
|
284  | 
(map (fn (s, srcs) => (get x s, map (attrib x) srcs)) th_srcs) x;  | 
|
285  | 
||
| 7572 | 286  | 
fun global_have_thmss kind f (x as ((name, _), _)) thy =  | 
287  | 
let val (thy', thms) = gen_have_thmss PureThy.get_thms Attrib.global_attribute f x thy in  | 
|
288  | 
if kind <> "" then Context.setmp (Some thy') (Present.results kind name) thms else ();  | 
|
289  | 
(thy', thms)  | 
|
290  | 
end;  | 
|
| 5915 | 291  | 
|
| 6371 | 292  | 
fun local_have_thmss x =  | 
293  | 
gen_have_thmss (ProofContext.get_thms o Proof.context_of)  | 
|
294  | 
(Attrib.local_attribute o Proof.theory_of) x;  | 
|
295  | 
||
| 7012 | 296  | 
fun gen_have_thmss_i f ((name, more_atts), th_atts) =  | 
| 6371 | 297  | 
f name more_atts (map (apfst single) th_atts);  | 
298  | 
||
299  | 
fun have_lemss name atts = PureThy.have_thmss name (atts @ [Drule.tag_lemma]);  | 
|
| 5915 | 300  | 
|
301  | 
||
| 7572 | 302  | 
fun apply_theorems th_srcs = global_have_thmss "" PureThy.have_thmss (("", []), th_srcs);
 | 
| 7476 | 303  | 
fun apply_theorems_i th_srcs = gen_have_thmss_i PureThy.have_thmss (("", []), th_srcs);
 | 
| 7572 | 304  | 
val have_theorems = #1 oo (global_have_thmss "theorems" PureThy.have_thmss o Comment.ignore);  | 
| 7476 | 305  | 
val have_theorems_i = #1 oo (gen_have_thmss_i PureThy.have_thmss o Comment.ignore);  | 
| 7572 | 306  | 
val have_lemmas = #1 oo (global_have_thmss "lemmas" have_lemss o Comment.ignore);  | 
| 7476 | 307  | 
val have_lemmas_i = #1 oo (gen_have_thmss_i have_lemss o Comment.ignore);  | 
| 6879 | 308  | 
val have_facts = ProofHistory.apply o local_have_thmss (Proof.have_thmss []) o Comment.ignore;  | 
| 7012 | 309  | 
val have_facts_i = ProofHistory.apply o gen_have_thmss_i (Proof.have_thmss []) o Comment.ignore;  | 
| 6371 | 310  | 
|
| 5915 | 311  | 
|
| 6371 | 312  | 
(* forward chaining *)  | 
313  | 
||
| 6879 | 314  | 
fun gen_from_facts f = ProofHistory.apply o (Proof.chain oo curry f ("", []));
 | 
315  | 
||
316  | 
fun add_thmss name atts ths_atts state =  | 
|
317  | 
Proof.have_thmss (Proof.the_facts state) name atts ths_atts state;  | 
|
| 6371 | 318  | 
|
| 6879 | 319  | 
val from_facts = gen_from_facts (local_have_thmss (Proof.have_thmss [])) o Comment.ignore;  | 
| 7012 | 320  | 
val from_facts_i = gen_from_facts (gen_have_thmss_i (Proof.have_thmss [])) o Comment.ignore;  | 
| 6879 | 321  | 
val with_facts = gen_from_facts (local_have_thmss add_thmss) o Comment.ignore;  | 
| 7012 | 322  | 
val with_facts_i = gen_from_facts (gen_have_thmss_i add_thmss) o Comment.ignore;  | 
| 6879 | 323  | 
|
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
324  | 
fun chain comment_ignore = ProofHistory.apply Proof.chain;  | 
| 5830 | 325  | 
|
326  | 
||
327  | 
(* context *)  | 
|
328  | 
||
| 7417 | 329  | 
val fix = ProofHistory.apply o Proof.fix o map Comment.ignore;  | 
330  | 
val fix_i = ProofHistory.apply o Proof.fix_i o map Comment.ignore;  | 
|
| 8615 | 331  | 
val let_bind = ProofHistory.apply o Proof.let_bind o map Comment.ignore;  | 
332  | 
val let_bind_i = ProofHistory.apply o Proof.let_bind_i o map Comment.ignore;  | 
|
| 8450 | 333  | 
|
334  | 
fun invoke_case ((name, src), comment_ignore) = ProofHistory.apply (fn state =>  | 
|
335  | 
Proof.invoke_case (name, map (Attrib.local_attribute (Proof.theory_of state)) src) state);  | 
|
336  | 
val invoke_case_i = ProofHistory.apply o Proof.invoke_case o Comment.ignore;  | 
|
| 5830 | 337  | 
|
338  | 
||
339  | 
(* statements *)  | 
|
340  | 
||
| 7271 | 341  | 
local  | 
342  | 
||
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
343  | 
fun global_statement f (name, src, s) int thy =  | 
| 6688 | 344  | 
ProofHistory.init (Toplevel.undo_limit int)  | 
345  | 
(f name (map (Attrib.global_attribute thy) src) s thy);  | 
|
| 5830 | 346  | 
|
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
347  | 
fun global_statement_i f (name, atts, t) int thy =  | 
| 6688 | 348  | 
ProofHistory.init (Toplevel.undo_limit int) (f name atts t thy);  | 
| 6501 | 349  | 
|
| 
6904
 
4125d6b6d8f9
removed proof history nesting commands (not useful);
 
wenzelm 
parents: 
6890 
diff
changeset
 | 
350  | 
fun local_statement f g (name, src, s) = ProofHistory.apply (fn state =>  | 
| 6501 | 351  | 
f name (map (Attrib.local_attribute (Proof.theory_of state)) src) s (g state));  | 
| 6371 | 352  | 
|
| 
6904
 
4125d6b6d8f9
removed proof history nesting commands (not useful);
 
wenzelm 
parents: 
6890 
diff
changeset
 | 
353  | 
fun local_statement_i f g (name, atts, t) = ProofHistory.apply (f name atts t o g);  | 
| 5830 | 354  | 
|
| 8090 | 355  | 
fun multi_local_attribute state (name, src, s) =  | 
356  | 
(name, map (Attrib.local_attribute (Proof.theory_of state)) src, s);  | 
|
357  | 
||
| 7271 | 358  | 
fun multi_statement f args = ProofHistory.apply (fn state =>  | 
| 7678 | 359  | 
f (map (multi_local_attribute state) args) state);  | 
| 7271 | 360  | 
|
361  | 
fun multi_statement_i f args = ProofHistory.apply (f args);  | 
|
362  | 
||
363  | 
in  | 
|
364  | 
||
| 6952 | 365  | 
val theorem = global_statement Proof.theorem o Comment.ignore;  | 
366  | 
val theorem_i = global_statement_i Proof.theorem_i o Comment.ignore;  | 
|
367  | 
val lemma = global_statement Proof.lemma o Comment.ignore;  | 
|
368  | 
val lemma_i = global_statement_i Proof.lemma_i o Comment.ignore;  | 
|
| 7271 | 369  | 
val assume = multi_statement Proof.assume o map Comment.ignore;  | 
370  | 
val assume_i = multi_statement_i Proof.assume_i o map Comment.ignore;  | 
|
371  | 
val presume = multi_statement Proof.presume o map Comment.ignore;  | 
|
372  | 
val presume_i = multi_statement_i Proof.presume_i o map Comment.ignore;  | 
|
| 6952 | 373  | 
val local_def = local_statement LocalDefs.def I o Comment.ignore;  | 
374  | 
val local_def_i = local_statement LocalDefs.def_i I o Comment.ignore;  | 
|
| 7176 | 375  | 
val show = local_statement (Proof.show Seq.single) I o Comment.ignore;  | 
376  | 
val show_i = local_statement_i (Proof.show_i Seq.single) I o Comment.ignore;  | 
|
377  | 
val have = local_statement (Proof.have Seq.single) I o Comment.ignore;  | 
|
378  | 
val have_i = local_statement_i (Proof.have_i Seq.single) I o Comment.ignore;  | 
|
379  | 
val thus = local_statement (Proof.show Seq.single) Proof.chain o Comment.ignore;  | 
|
380  | 
val thus_i = local_statement_i (Proof.show_i Seq.single) Proof.chain o Comment.ignore;  | 
|
381  | 
val hence = local_statement (Proof.have Seq.single) Proof.chain o Comment.ignore;  | 
|
382  | 
val hence_i = local_statement_i (Proof.have_i Seq.single) Proof.chain o Comment.ignore;  | 
|
| 5830 | 383  | 
|
| 7271 | 384  | 
end;  | 
385  | 
||
| 5830 | 386  | 
|
387  | 
(* blocks *)  | 
|
388  | 
||
| 
6904
 
4125d6b6d8f9
removed proof history nesting commands (not useful);
 
wenzelm 
parents: 
6890 
diff
changeset
 | 
389  | 
val begin_block = ProofHistory.apply Proof.begin_block;  | 
| 5830 | 390  | 
val next_block = ProofHistory.apply Proof.next_block;  | 
| 6937 | 391  | 
val end_block = ProofHistory.applys Proof.end_block;  | 
| 5830 | 392  | 
|
393  | 
||
| 8165 | 394  | 
(* shuffle state *)  | 
395  | 
||
| 8236 | 396  | 
fun shuffle meth i = Method.refine (Method.Basic (K (meth i))) o Proof.assert_no_chain;  | 
| 8204 | 397  | 
|
| 8681 | 398  | 
fun defer (i, comment_ignore) = ProofHistory.applys (shuffle Method.defer i);  | 
399  | 
fun prefer (i, comment_ignore) = ProofHistory.applys (shuffle Method.prefer i);  | 
|
| 8165 | 400  | 
|
401  | 
||
| 5830 | 402  | 
(* backward steps *)  | 
403  | 
||
| 8881 | 404  | 
fun apply (m, comment_ignore) = ProofHistory.applys  | 
405  | 
(Seq.map (Proof.goal_facts (K [])) o Method.refine m o Proof.assert_backward);  | 
|
406  | 
||
407  | 
fun apply_end (m, comment_ignore) = ProofHistory.applys  | 
|
408  | 
(Method.refine_end m o Proof.assert_forward);  | 
|
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
409  | 
|
| 7002 | 410  | 
val proof = ProofHistory.applys o Method.proof o  | 
411  | 
apsome Comment.ignore_interest o Comment.ignore_interest' o Comment.ignore;  | 
|
| 6404 | 412  | 
|
413  | 
||
| 6732 | 414  | 
(* print result *)  | 
415  | 
||
| 7572 | 416  | 
local  | 
417  | 
||
| 6732 | 418  | 
fun pretty_result {kind, name, thm} =
 | 
| 7590 | 419  | 
Pretty.block [Pretty.str (kind ^ (if name = "" then "" else " " ^ name) ^ ":"),  | 
| 7572 | 420  | 
Pretty.fbrk, Proof.pretty_thm thm];  | 
| 6986 | 421  | 
|
422  | 
fun pretty_rule thm =  | 
|
| 7353 | 423  | 
Pretty.block [Pretty.str "trying to solve goal by rule:", Pretty.fbrk, Proof.pretty_thm thm];  | 
| 6732 | 424  | 
|
| 7572 | 425  | 
in  | 
426  | 
||
| 6732 | 427  | 
val print_result = Pretty.writeln o pretty_result;  | 
428  | 
||
| 6986 | 429  | 
fun cond_print_result_rule int =  | 
430  | 
if int then (print_result, Pretty.writeln o pretty_rule)  | 
|
431  | 
else (K (), K ());  | 
|
432  | 
||
433  | 
fun proof'' f = Toplevel.proof' (f o cond_print_result_rule);  | 
|
| 6732 | 434  | 
|
| 7572 | 435  | 
end;  | 
436  | 
||
| 6732 | 437  | 
|
| 6404 | 438  | 
(* local endings *)  | 
439  | 
||
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
440  | 
val local_qed =  | 
| 
6904
 
4125d6b6d8f9
removed proof history nesting commands (not useful);
 
wenzelm 
parents: 
6890 
diff
changeset
 | 
441  | 
proof'' o (ProofHistory.applys oo Method.local_qed) o apsome Comment.ignore_interest;  | 
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
442  | 
|
| 6937 | 443  | 
fun ignore_interests (x, y) = (Comment.ignore_interest x, apsome Comment.ignore_interest y);  | 
444  | 
||
| 6688 | 445  | 
val local_terminal_proof =  | 
| 6937 | 446  | 
proof'' o (ProofHistory.applys oo Method.local_terminal_proof) o ignore_interests;  | 
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
447  | 
|
| 
6904
 
4125d6b6d8f9
removed proof history nesting commands (not useful);
 
wenzelm 
parents: 
6890 
diff
changeset
 | 
448  | 
val local_immediate_proof = proof'' (ProofHistory.applys o Method.local_immediate_proof);  | 
| 
 
4125d6b6d8f9
removed proof history nesting commands (not useful);
 
wenzelm 
parents: 
6890 
diff
changeset
 | 
449  | 
val local_default_proof = proof'' (ProofHistory.applys o Method.local_default_proof);  | 
| 
 
4125d6b6d8f9
removed proof history nesting commands (not useful);
 
wenzelm 
parents: 
6890 
diff
changeset
 | 
450  | 
val local_skip_proof = proof'' (ProofHistory.applys o SkipProof.local_skip_proof);  | 
| 6404 | 451  | 
|
452  | 
||
453  | 
(* global endings *)  | 
|
454  | 
||
455  | 
fun global_result finish = Toplevel.proof_to_theory (fn prf =>  | 
|
456  | 
let  | 
|
457  | 
val state = ProofHistory.current prf;  | 
|
458  | 
val _ = if Proof.at_bottom state then () else raise Toplevel.UNDEF;  | 
|
| 7572 | 459  | 
    val (thy, res as {kind, name, thm}) = finish state;
 | 
460  | 
in print_result res; Context.setmp (Some thy) (Present.result kind name) thm; thy end);  | 
|
| 6404 | 461  | 
|
| 6952 | 462  | 
val global_qed = global_result o Method.global_qed o apsome Comment.ignore_interest;  | 
| 6937 | 463  | 
val global_terminal_proof = global_result o Method.global_terminal_proof o ignore_interests;  | 
| 6404 | 464  | 
val global_immediate_proof = global_result Method.global_immediate_proof;  | 
465  | 
val global_default_proof = global_result Method.global_default_proof;  | 
|
| 
6888
 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 
wenzelm 
parents: 
6885 
diff
changeset
 | 
466  | 
val global_skip_proof = global_result SkipProof.global_skip_proof;  | 
| 6404 | 467  | 
|
468  | 
||
469  | 
(* common endings *)  | 
|
470  | 
||
| 7002 | 471  | 
fun qed (meth, comment) = local_qed meth o global_qed meth;  | 
472  | 
fun terminal_proof (meths, comment) = local_terminal_proof meths o global_terminal_proof meths;  | 
|
473  | 
fun immediate_proof comment = local_immediate_proof o global_immediate_proof;  | 
|
474  | 
fun default_proof comment = local_default_proof o global_default_proof;  | 
|
475  | 
fun skip_proof comment = local_skip_proof o global_skip_proof;  | 
|
| 5830 | 476  | 
|
| 8210 | 477  | 
fun forget_proof comment = Toplevel.proof_to_theory (Proof.theory_of o ProofHistory.current);  | 
478  | 
||
| 5830 | 479  | 
|
| 6774 | 480  | 
(* calculational proof commands *)  | 
481  | 
||
| 6879 | 482  | 
local  | 
483  | 
||
| 7417 | 484  | 
fun cond_print_calc int thms =  | 
485  | 
if int then Pretty.writeln (Pretty.big_list "calculation:" (map Proof.pretty_thm thms))  | 
|
| 6774 | 486  | 
else ();  | 
487  | 
||
488  | 
fun proof''' f = Toplevel.proof' (f o cond_print_calc);  | 
|
489  | 
||
| 6879 | 490  | 
fun gen_calc get f (args, _) prt state =  | 
491  | 
f (apsome (fn (srcs, _) => get srcs state) args) prt state;  | 
|
492  | 
||
493  | 
in  | 
|
494  | 
||
| 7012 | 495  | 
fun get_thmss srcs = Proof.the_facts o local_have_thmss (Proof.have_thmss []) (("", []), srcs); 
 | 
496  | 
fun get_thmss_i thms _ = thms;  | 
|
497  | 
||
| 6879 | 498  | 
fun also x = proof''' (ProofHistory.applys o gen_calc get_thmss Calculation.also x);  | 
499  | 
fun also_i x = proof''' (ProofHistory.applys o gen_calc get_thmss_i Calculation.also x);  | 
|
500  | 
fun finally x = proof''' (ProofHistory.applys o gen_calc get_thmss Calculation.finally x);  | 
|
501  | 
fun finally_i x = proof''' (ProofHistory.applys o gen_calc get_thmss_i Calculation.finally x);  | 
|
| 8562 | 502  | 
fun moreover comment = proof''' (ProofHistory.apply o Calculation.moreover);  | 
| 8588 | 503  | 
fun ultimately comment = proof''' (ProofHistory.apply o Calculation.ultimately);  | 
| 6879 | 504  | 
|
505  | 
end;  | 
|
| 6774 | 506  | 
|
507  | 
||
| 5830 | 508  | 
(* translation functions *)  | 
509  | 
||
510  | 
val parse_ast_translation =  | 
|
| 
8349
 
611342539639
moved use_mltext, use_mltext_theory, use_let, use_setup to context.ML;
 
wenzelm 
parents: 
8236 
diff
changeset
 | 
511  | 
Context.use_let "parse_ast_translation: (string * (Syntax.ast list -> Syntax.ast)) list"  | 
| 5830 | 512  | 
"Theory.add_trfuns (parse_ast_translation, [], [], [])";  | 
513  | 
||
514  | 
val parse_translation =  | 
|
| 
8349
 
611342539639
moved use_mltext, use_mltext_theory, use_let, use_setup to context.ML;
 
wenzelm 
parents: 
8236 
diff
changeset
 | 
515  | 
Context.use_let "parse_translation: (string * (term list -> term)) list"  | 
| 5830 | 516  | 
"Theory.add_trfuns ([], parse_translation, [], [])";  | 
517  | 
||
518  | 
val print_translation =  | 
|
| 
8349
 
611342539639
moved use_mltext, use_mltext_theory, use_let, use_setup to context.ML;
 
wenzelm 
parents: 
8236 
diff
changeset
 | 
519  | 
Context.use_let "print_translation: (string * (term list -> term)) list"  | 
| 5830 | 520  | 
"Theory.add_trfuns ([], [], print_translation, [])";  | 
521  | 
||
522  | 
val print_ast_translation =  | 
|
| 
8349
 
611342539639
moved use_mltext, use_mltext_theory, use_let, use_setup to context.ML;
 
wenzelm 
parents: 
8236 
diff
changeset
 | 
523  | 
Context.use_let "print_ast_translation: (string * (Syntax.ast list -> Syntax.ast)) list"  | 
| 5830 | 524  | 
"Theory.add_trfuns ([], [], [], print_ast_translation)";  | 
525  | 
||
526  | 
val typed_print_translation =  | 
|
| 
8349
 
611342539639
moved use_mltext, use_mltext_theory, use_let, use_setup to context.ML;
 
wenzelm 
parents: 
8236 
diff
changeset
 | 
527  | 
Context.use_let "typed_print_translation: (string * (bool -> typ -> term list -> term)) list"  | 
| 5830 | 528  | 
"Theory.add_trfunsT typed_print_translation";  | 
529  | 
||
530  | 
val token_translation =  | 
|
| 
8349
 
611342539639
moved use_mltext, use_mltext_theory, use_let, use_setup to context.ML;
 
wenzelm 
parents: 
8236 
diff
changeset
 | 
531  | 
Context.use_let "token_translation: (string * string * (string -> string * int)) list"  | 
| 5830 | 532  | 
"Theory.add_tokentrfuns token_translation";  | 
533  | 
||
534  | 
||
535  | 
(* add_oracle *)  | 
|
536  | 
||
| 
6728
 
b51b25db7bc6
added formal comment arguments almost everywhere (still ignored);
 
wenzelm 
parents: 
6688 
diff
changeset
 | 
537  | 
fun add_oracle ((name, txt), comment_ignore) =  | 
| 
8349
 
611342539639
moved use_mltext, use_mltext_theory, use_let, use_setup to context.ML;
 
wenzelm 
parents: 
8236 
diff
changeset
 | 
538  | 
Context.use_let  | 
| 5830 | 539  | 
"oracle: bstring * (Sign.sg * Object.T -> term)"  | 
540  | 
"Theory.add_oracle oracle"  | 
|
541  | 
    ("(" ^ quote name ^ ", " ^ txt ^ ")");
 | 
|
542  | 
||
543  | 
||
| 6688 | 544  | 
(* theory init and exit *)  | 
| 5830 | 545  | 
|
| 8804 | 546  | 
fun gen_begin_theory upd name parents files =  | 
547  | 
ThyInfo.begin_theory Present.begin_theory upd name parents (map (apfst Path.unpack) files);  | 
|
| 5830 | 548  | 
|
| 8804 | 549  | 
val begin_theory = gen_begin_theory false;  | 
| 7909 | 550  | 
|
551  | 
fun end_theory thy =  | 
|
| 7932 | 552  | 
if ThyInfo.check_known_thy (PureThy.get_name thy) then ThyInfo.end_theory thy  | 
| 7909 | 553  | 
else thy;  | 
554  | 
||
| 7932 | 555  | 
val kill_theory = ThyInfo.if_known_thy ThyInfo.remove_thy;  | 
| 6688 | 556  | 
|
| 7103 | 557  | 
|
| 8804 | 558  | 
fun bg_theory (name, parents, files) int = gen_begin_theory int name parents files;  | 
| 6331 | 559  | 
fun en_theory thy = (end_theory thy; ());  | 
560  | 
||
| 7932 | 561  | 
fun theory spec = Toplevel.init_theory (bg_theory spec) en_theory (kill_theory o PureThy.get_name);  | 
| 6246 | 562  | 
|
563  | 
||
564  | 
(* context switch *)  | 
|
565  | 
||
| 7960 | 566  | 
fun fetch_context f x =  | 
567  | 
(case Context.pass None f x of  | 
|
568  | 
((), None) => raise Toplevel.UNDEF  | 
|
569  | 
| ((), Some thy) => thy);  | 
|
570  | 
||
571  | 
fun init_context f x = Toplevel.init_theory (fn _ => fetch_context f x) (K ()) (K ());  | 
|
| 7953 | 572  | 
|
573  | 
val context = init_context (ThyInfo.quiet_update_thy true);  | 
|
| 6246 | 574  | 
|
| 5830 | 575  | 
|
576  | 
end;  |