| author | wenzelm | 
| Sat, 10 Oct 2020 21:12:20 +0200 | |
| changeset 72425 | d0937d55eb90 | 
| parent 70839 | 2136e4670ad2 | 
| child 73549 | a2c589d5e1e4 | 
| permissions | -rw-r--r-- | 
| 5831 | 1 | (* Title: Pure/Isar/isar_cmd.ML | 
| 2 | Author: Markus Wenzel, TU Muenchen | |
| 3 | ||
| 30805 | 4 | Miscellaneous Isar commands. | 
| 5831 | 5 | *) | 
| 6 | ||
| 7 | signature ISAR_CMD = | |
| 8 | sig | |
| 59930 | 9 | val setup: Input.source -> theory -> theory | 
| 59064 | 10 | val local_setup: Input.source -> Proof.context -> Proof.context | 
| 11 | val parse_ast_translation: Input.source -> theory -> theory | |
| 12 | val parse_translation: Input.source -> theory -> theory | |
| 13 | val print_translation: Input.source -> theory -> theory | |
| 14 | val typed_print_translation: Input.source -> theory -> theory | |
| 15 | val print_ast_translation: Input.source -> theory -> theory | |
| 42204 | 16 | val translations: (xstring * string) Syntax.trrule list -> theory -> theory | 
| 17 | val no_translations: (xstring * string) Syntax.trrule list -> theory -> theory | |
| 59064 | 18 | val oracle: bstring * Position.range -> Input.source -> theory -> theory | 
| 19 |   val declaration: {syntax: bool, pervasive: bool} -> Input.source -> local_theory -> local_theory
 | |
| 20 | val simproc_setup: string * Position.T -> string list -> Input.source -> | |
| 62913 | 21 | local_theory -> local_theory | 
| 49889 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49866diff
changeset | 22 | val qed: Method.text_range option -> Toplevel.transition -> Toplevel.transition | 
| 
00ea087e83d8
more method position information, notably finished_pos after end of previous text;
 wenzelm parents: 
49866diff
changeset | 23 | val terminal_proof: Method.text_range * Method.text_range option -> | 
| 21350 | 24 | Toplevel.transition -> Toplevel.transition | 
| 25 | val default_proof: Toplevel.transition -> Toplevel.transition | |
| 26 | val immediate_proof: Toplevel.transition -> Toplevel.transition | |
| 27 | val done_proof: Toplevel.transition -> Toplevel.transition | |
| 28 | val skip_proof: Toplevel.transition -> Toplevel.transition | |
| 59064 | 29 | val ml_diag: bool -> Input.source -> Toplevel.transition -> Toplevel.transition | 
| 48776 
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
 wenzelm parents: 
47815diff
changeset | 30 | val diag_state: Proof.context -> Toplevel.state | 
| 
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
 wenzelm parents: 
47815diff
changeset | 31 |   val diag_goal: Proof.context -> {context: Proof.context, facts: thm list, goal: thm}
 | 
| 57605 | 32 | val pretty_theorems: bool -> Toplevel.state -> Pretty.T list | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57934diff
changeset | 33 | val print_stmts: string list * (Facts.ref * Token.src list) list | 
| 19268 | 34 | -> Toplevel.transition -> Toplevel.transition | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57934diff
changeset | 35 | val print_thms: string list * (Facts.ref * Token.src list) list | 
| 10581 | 36 | -> Toplevel.transition -> Toplevel.transition | 
| 58011 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 wenzelm parents: 
57934diff
changeset | 37 | val print_prfs: bool -> string list * (Facts.ref * Token.src list) list option | 
| 11524 
197f2e14a714
Added functions for printing primitive proof terms.
 berghofe parents: 
11017diff
changeset | 38 | -> Toplevel.transition -> Toplevel.transition | 
| 12876 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
12758diff
changeset | 39 | val print_prop: (string list * string) -> Toplevel.transition -> Toplevel.transition | 
| 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
12758diff
changeset | 40 | val print_term: (string list * string) -> Toplevel.transition -> Toplevel.transition | 
| 48792 | 41 | val print_type: (string list * (string * string option)) -> | 
| 42 | Toplevel.transition -> Toplevel.transition | |
| 5831 | 43 | end; | 
| 44 | ||
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
37198diff
changeset | 45 | structure Isar_Cmd: ISAR_CMD = | 
| 5831 | 46 | struct | 
| 47 | ||
| 48 | ||
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 49 | (** theory declarations **) | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 50 | |
| 30461 | 51 | (* generic setup *) | 
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 52 | |
| 59930 | 53 | fun setup source = | 
| 69216 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 54 | ML_Context.expression (Input.pos_of source) | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 55 |     (ML_Lex.read "Theory.setup (" @ ML_Lex.read_source source @ ML_Lex.read ")")
 | 
| 26435 | 56 | |> Context.theory_map; | 
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 57 | |
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
53171diff
changeset | 58 | fun local_setup source = | 
| 69216 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 59 | ML_Context.expression (Input.pos_of source) | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 60 |     (ML_Lex.read "Theory.local_setup (" @ ML_Lex.read_source source @ ML_Lex.read ")")
 | 
| 30461 | 61 | |> Context.proof_map; | 
| 62 | ||
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 63 | |
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 64 | (* translation functions *) | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 65 | |
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
53171diff
changeset | 66 | fun parse_ast_translation source = | 
| 69216 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 67 | ML_Context.expression (Input.pos_of source) | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 68 |     (ML_Lex.read "Theory.setup (Sign.parse_ast_translation (" @
 | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 69 | ML_Lex.read_source source @ ML_Lex.read "))") | 
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 70 | |> Context.theory_map; | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 71 | |
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
53171diff
changeset | 72 | fun parse_translation source = | 
| 69216 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 73 | ML_Context.expression (Input.pos_of source) | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 74 |     (ML_Lex.read "Theory.setup (Sign.parse_translation (" @
 | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 75 | ML_Lex.read_source source @ ML_Lex.read "))") | 
| 52143 | 76 | |> Context.theory_map; | 
| 77 | ||
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
53171diff
changeset | 78 | fun print_translation source = | 
| 69216 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 79 | ML_Context.expression (Input.pos_of source) | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 80 |     (ML_Lex.read "Theory.setup (Sign.print_translation (" @
 | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 81 | ML_Lex.read_source source @ ML_Lex.read "))") | 
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 82 | |> Context.theory_map; | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 83 | |
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
53171diff
changeset | 84 | fun typed_print_translation source = | 
| 69216 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 85 | ML_Context.expression (Input.pos_of source) | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 86 |     (ML_Lex.read "Theory.setup (Sign.typed_print_translation (" @
 | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 87 | ML_Lex.read_source source @ ML_Lex.read "))") | 
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 88 | |> Context.theory_map; | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 89 | |
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
53171diff
changeset | 90 | fun print_ast_translation source = | 
| 69216 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 91 | ML_Context.expression (Input.pos_of source) | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 92 |     (ML_Lex.read "Theory.setup (Sign.print_ast_translation (" @
 | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 93 | ML_Lex.read_source source @ ML_Lex.read "))") | 
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 94 | |> Context.theory_map; | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 95 | |
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 96 | |
| 42204 | 97 | (* translation rules *) | 
| 98 | ||
| 99 | fun read_trrules thy raw_rules = | |
| 100 | let | |
| 42360 | 101 | val ctxt = Proof_Context.init_global thy; | 
| 56006 | 102 | val read_root = | 
| 103 |       #1 o dest_Type o Proof_Context.read_type_name {proper = true, strict = false} ctxt;
 | |
| 42204 | 104 | in | 
| 56006 | 105 | raw_rules | 
| 106 | |> map (Syntax.map_trrule (fn (r, s) => Syntax_Phases.parse_ast_pattern ctxt (read_root r, s))) | |
| 42204 | 107 | end; | 
| 108 | ||
| 109 | fun translations args thy = Sign.add_trrules (read_trrules thy args) thy; | |
| 110 | fun no_translations args thy = Sign.del_trrules (read_trrules thy args) thy; | |
| 111 | ||
| 112 | ||
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 113 | (* oracles *) | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 114 | |
| 59029 
c907cbe36713
more careful ML source positions, for improved PIDE markup;
 wenzelm parents: 
58991diff
changeset | 115 | fun oracle (name, range) source = | 
| 69263 | 116 | ML_Context.expression (Input.pos_of source) | 
| 117 | (ML_Lex.read "val " @ | |
| 118 | ML_Lex.read_set_range range name @ | |
| 119 | ML_Lex.read | |
| 120 |       (" = snd (Context.>>> (Context.map_theory_result (Thm.add_oracle (" ^
 | |
| 121 | ML_Syntax.make_binding (name, #1 range) ^ ", ") @ | |
| 122 | ML_Lex.read_source source @ | |
| 123 | ML_Lex.read "))))") | |
| 124 | |> Context.theory_map; | |
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 125 | |
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 126 | |
| 22087 | 127 | (* declarations *) | 
| 128 | ||
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
53171diff
changeset | 129 | fun declaration {syntax, pervasive} source =
 | 
| 69216 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 130 | ML_Context.expression (Input.pos_of source) | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 131 | (ML_Lex.read | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 132 |       ("Theory.local_setup (Local_Theory.declaration {syntax = " ^
 | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 133 |         Bool.toString syntax ^ ", pervasive = " ^ Bool.toString pervasive ^ "} (") @
 | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 134 | ML_Lex.read_source source @ ML_Lex.read "))") | 
| 26385 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 wenzelm parents: 
26336diff
changeset | 135 | |> Context.proof_map; | 
| 22087 | 136 | |
| 137 | ||
| 22202 | 138 | (* simprocs *) | 
| 139 | ||
| 62913 | 140 | fun simproc_setup name lhss source = | 
| 69216 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 141 | ML_Context.expression (Input.pos_of source) | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 142 | (ML_Lex.read | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 143 |       ("Theory.local_setup (Simplifier.define_simproc_cmd (" ^
 | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 144 |         ML_Syntax.make_binding name ^ ") {lhss = " ^ ML_Syntax.print_strings lhss ^
 | 
| 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69187diff
changeset | 145 |         ", proc = (") @ ML_Lex.read_source source @ ML_Lex.read ")})")
 | 
| 22202 | 146 | |> Context.proof_map; | 
| 147 | ||
| 148 | ||
| 21350 | 149 | (* local endings *) | 
| 150 | ||
| 29383 | 151 | fun local_qed m = Toplevel.proof (Proof.local_qed (m, true)); | 
| 60403 | 152 | val local_terminal_proof = Toplevel.proof o Proof.local_future_terminal_proof; | 
| 29383 | 153 | val local_default_proof = Toplevel.proof Proof.local_default_proof; | 
| 154 | val local_immediate_proof = Toplevel.proof Proof.local_immediate_proof; | |
| 155 | val local_done_proof = Toplevel.proof Proof.local_done_proof; | |
| 156 | val local_skip_proof = Toplevel.proof' Proof.local_skip_proof; | |
| 21350 | 157 | |
| 158 | ||
| 159 | (* global endings *) | |
| 160 | ||
| 161 | fun global_qed m = Toplevel.end_proof (K (Proof.global_qed (m, true))); | |
| 60403 | 162 | val global_terminal_proof = Toplevel.end_proof o K o Proof.global_future_terminal_proof; | 
| 21350 | 163 | val global_default_proof = Toplevel.end_proof (K Proof.global_default_proof); | 
| 164 | val global_immediate_proof = Toplevel.end_proof (K Proof.global_immediate_proof); | |
| 165 | val global_skip_proof = Toplevel.end_proof Proof.global_skip_proof; | |
| 166 | val global_done_proof = Toplevel.end_proof (K Proof.global_done_proof); | |
| 167 | ||
| 168 | ||
| 169 | (* common endings *) | |
| 170 | ||
| 60693 
044f8bb3dd30
more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
 wenzelm parents: 
60405diff
changeset | 171 | fun qed m = local_qed m o global_qed m; | 
| 21350 | 172 | fun terminal_proof m = local_terminal_proof m o global_terminal_proof m; | 
| 173 | val default_proof = local_default_proof o global_default_proof; | |
| 174 | val immediate_proof = local_immediate_proof o global_immediate_proof; | |
| 175 | val done_proof = local_done_proof o global_done_proof; | |
| 176 | val skip_proof = local_skip_proof o global_skip_proof; | |
| 177 | ||
| 178 | ||
| 26489 | 179 | (* diagnostic ML evaluation *) | 
| 5831 | 180 | |
| 37305 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 181 | structure Diag_State = Proof_Data | 
| 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 182 | ( | 
| 69883 | 183 | type T = Toplevel.state option; | 
| 184 | fun init _ = NONE; | |
| 37305 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 185 | ); | 
| 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 186 | |
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
53171diff
changeset | 187 | fun ml_diag verbose source = Toplevel.keep (fn state => | 
| 56304 
40274e4f5ebf
redirect ML_Compiler reports more directly: only the (big) parse tree report is deferred via Execution.print (NB: this does not work for asynchronous "diag" commands);
 wenzelm parents: 
56278diff
changeset | 188 | let | 
| 
40274e4f5ebf
redirect ML_Compiler reports more directly: only the (big) parse tree report is deferred via Execution.print (NB: this does not work for asynchronous "diag" commands);
 wenzelm parents: 
56278diff
changeset | 189 | val opt_ctxt = | 
| 
40274e4f5ebf
redirect ML_Compiler reports more directly: only the (big) parse tree report is deferred via Execution.print (NB: this does not work for asynchronous "diag" commands);
 wenzelm parents: 
56278diff
changeset | 190 | try Toplevel.generic_theory_of state | 
| 69883 | 191 | |> Option.map (Context.proof_of #> Diag_State.put (SOME state)); | 
| 56304 
40274e4f5ebf
redirect ML_Compiler reports more directly: only the (big) parse tree report is deferred via Execution.print (NB: this does not work for asynchronous "diag" commands);
 wenzelm parents: 
56278diff
changeset | 192 | val flags = ML_Compiler.verbose verbose ML_Compiler.flags; | 
| 
40274e4f5ebf
redirect ML_Compiler reports more directly: only the (big) parse tree report is deferred via Execution.print (NB: this does not work for asynchronous "diag" commands);
 wenzelm parents: 
56278diff
changeset | 193 | in ML_Context.eval_source_in opt_ctxt flags source end); | 
| 37305 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 194 | |
| 69883 | 195 | fun diag_state ctxt = | 
| 196 | (case Diag_State.get ctxt of | |
| 197 | SOME st => st | |
| 69886 | 198 | | NONE => Toplevel.init_toplevel ()); | 
| 69883 | 199 | |
| 60094 | 200 | val diag_goal = Proof.goal o Toplevel.proof_of o diag_state; | 
| 37305 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 201 | |
| 53171 | 202 | val _ = Theory.setup | 
| 67147 | 203 | (ML_Antiquotation.value (Binding.qualify true "Isar" \<^binding>\<open>state\<close>) | 
| 53171 | 204 | (Scan.succeed "Isar_Cmd.diag_state ML_context") #> | 
| 67147 | 205 | ML_Antiquotation.value (Binding.qualify true "Isar" \<^binding>\<open>goal\<close>) | 
| 53171 | 206 | (Scan.succeed "Isar_Cmd.diag_goal ML_context")); | 
| 5831 | 207 | |
| 208 | ||
| 56868 
b5fb264d53ba
clarified print operations for "terms" and "theorems";
 wenzelm parents: 
56334diff
changeset | 209 | (* theorems of theory or proof context *) | 
| 17066 | 210 | |
| 56868 
b5fb264d53ba
clarified print operations for "terms" and "theorems";
 wenzelm parents: 
56334diff
changeset | 211 | fun pretty_theorems verbose st = | 
| 
b5fb264d53ba
clarified print operations for "terms" and "theorems";
 wenzelm parents: 
56334diff
changeset | 212 | if Toplevel.is_proof st then | 
| 59917 
9830c944670f
more uniform "verbose" option to print name space;
 wenzelm parents: 
59210diff
changeset | 213 | Proof_Context.pretty_local_facts verbose (Toplevel.context_of st) | 
| 56868 
b5fb264d53ba
clarified print operations for "terms" and "theorems";
 wenzelm parents: 
56334diff
changeset | 214 | else | 
| 
b5fb264d53ba
clarified print operations for "terms" and "theorems";
 wenzelm parents: 
56334diff
changeset | 215 | let | 
| 61266 | 216 | val ctxt = Toplevel.context_of st; | 
| 56868 
b5fb264d53ba
clarified print operations for "terms" and "theorems";
 wenzelm parents: 
56334diff
changeset | 217 | val prev_thys = | 
| 67391 | 218 | (case Toplevel.previous_theory_of st of | 
| 219 | SOME thy => [thy] | |
| 61266 | 220 | | NONE => Theory.parents_of (Proof_Context.theory_of ctxt)); | 
| 221 | in Proof_Display.pretty_theorems_diff verbose prev_thys ctxt end; | |
| 9513 | 222 | |
| 12060 | 223 | |
| 19268 | 224 | (* print theorems, terms, types etc. *) | 
| 225 | ||
| 19385 
c0f2f8224ea8
print_term etc.: actually observe print mode in final output;
 wenzelm parents: 
19268diff
changeset | 226 | local | 
| 
c0f2f8224ea8
print_term etc.: actually observe print mode in final output;
 wenzelm parents: 
19268diff
changeset | 227 | |
| 38331 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 228 | fun string_of_stmts ctxt args = | 
| 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 229 | Attrib.eval_thms ctxt args | 
| 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 230 | |> map (Element.pretty_statement ctxt Thm.theoremK) | 
| 19385 
c0f2f8224ea8
print_term etc.: actually observe print mode in final output;
 wenzelm parents: 
19268diff
changeset | 231 | |> Pretty.chunks2 |> Pretty.string_of; | 
| 5880 | 232 | |
| 38331 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 233 | fun string_of_thms ctxt args = | 
| 51583 | 234 |   Pretty.string_of (Proof_Context.pretty_fact ctxt ("", Attrib.eval_thms ctxt args));
 | 
| 5895 | 235 | |
| 19385 
c0f2f8224ea8
print_term etc.: actually observe print mode in final output;
 wenzelm parents: 
19268diff
changeset | 236 | fun string_of_prfs full state arg = | 
| 32859 
204f749f35a9
replaced Proof.get_goal state by Proof.flat_goal state, which provides the standard view on goals for (semi)automated tools;
 wenzelm parents: 
32804diff
changeset | 237 | Pretty.string_of | 
| 
204f749f35a9
replaced Proof.get_goal state by Proof.flat_goal state, which provides the standard view on goals for (semi)automated tools;
 wenzelm parents: 
32804diff
changeset | 238 | (case arg of | 
| 15531 | 239 | NONE => | 
| 12125 
316d11f760f7
Commands prf and full_prf can now also be used to display proof term
 berghofe parents: 
12069diff
changeset | 240 | let | 
| 38331 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 241 |           val {context = ctxt, goal = thm} = Proof.simple_goal (Toplevel.proof_of state);
 | 
| 70449 | 242 | val thy = Proof_Context.theory_of ctxt; | 
| 28814 | 243 | val prf = Thm.proof_of thm; | 
| 17066 | 244 | val prop = Thm.full_prop_of thm; | 
| 32859 
204f749f35a9
replaced Proof.get_goal state by Proof.flat_goal state, which provides the standard view on goals for (semi)automated tools;
 wenzelm parents: 
32804diff
changeset | 245 | val prf' = Proofterm.rewrite_proof_notypes ([], []) prf; | 
| 12125 
316d11f760f7
Commands prf and full_prf can now also be used to display proof term
 berghofe parents: 
12069diff
changeset | 246 | in | 
| 33388 | 247 | Proof_Syntax.pretty_proof ctxt | 
| 70449 | 248 | (if full then Proofterm.reconstruct_proof thy prop prf' else prf') | 
| 12125 
316d11f760f7
Commands prf and full_prf can now also be used to display proof term
 berghofe parents: 
12069diff
changeset | 249 | end | 
| 38331 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 250 | | SOME srcs => | 
| 63624 
994d1a1105ef
more informative 'prf' and 'full_prf', based on HOL/Proofs/ex/XML_Data.thy;
 wenzelm parents: 
62922diff
changeset | 251 | let | 
| 
994d1a1105ef
more informative 'prf' and 'full_prf', based on HOL/Proofs/ex/XML_Data.thy;
 wenzelm parents: 
62922diff
changeset | 252 | val ctxt = Toplevel.context_of state; | 
| 70839 
2136e4670ad2
clarified standard_proof_of: prefer expand_proof over somewhat adhoc strip_thm_proof;
 wenzelm parents: 
70449diff
changeset | 253 | val pretty_proof = Proof_Syntax.pretty_standard_proof_of ctxt full; | 
| 63624 
994d1a1105ef
more informative 'prf' and 'full_prf', based on HOL/Proofs/ex/XML_Data.thy;
 wenzelm parents: 
62922diff
changeset | 254 | in Pretty.chunks (map pretty_proof (Attrib.eval_thms ctxt srcs)) end); | 
| 11524 
197f2e14a714
Added functions for printing primitive proof terms.
 berghofe parents: 
11017diff
changeset | 255 | |
| 38331 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 256 | fun string_of_prop ctxt s = | 
| 5831 | 257 | let | 
| 24508 
c8b82fec6447
replaced ProofContext.read_term/prop by general Syntax.read_term/prop;
 wenzelm parents: 
24314diff
changeset | 258 | val prop = Syntax.read_prop ctxt s; | 
| 70308 | 259 | val ctxt' = Proof_Context.augment prop ctxt; | 
| 26704 
51ee753cc2e3
token translations: context dependent, result Pretty.T;
 wenzelm parents: 
26694diff
changeset | 260 | in Pretty.string_of (Pretty.quote (Syntax.pretty_term ctxt' prop)) end; | 
| 5831 | 261 | |
| 38331 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 262 | fun string_of_term ctxt s = | 
| 5831 | 263 | let | 
| 24508 
c8b82fec6447
replaced ProofContext.read_term/prop by general Syntax.read_term/prop;
 wenzelm parents: 
24314diff
changeset | 264 | val t = Syntax.read_term ctxt s; | 
| 5831 | 265 | val T = Term.type_of t; | 
| 70308 | 266 | val ctxt' = Proof_Context.augment t ctxt; | 
| 5831 | 267 | in | 
| 19385 
c0f2f8224ea8
print_term etc.: actually observe print mode in final output;
 wenzelm parents: 
19268diff
changeset | 268 | Pretty.string_of | 
| 26704 
51ee753cc2e3
token translations: context dependent, result Pretty.T;
 wenzelm parents: 
26694diff
changeset | 269 | (Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t), Pretty.fbrk, | 
| 
51ee753cc2e3
token translations: context dependent, result Pretty.T;
 wenzelm parents: 
26694diff
changeset | 270 | Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' T)]) | 
| 9128 | 271 | end; | 
| 5831 | 272 | |
| 48792 | 273 | fun string_of_type ctxt (s, NONE) = | 
| 274 | let val T = Syntax.read_typ ctxt s | |
| 275 | in Pretty.string_of (Pretty.quote (Syntax.pretty_typ ctxt T)) end | |
| 276 | | string_of_type ctxt (s1, SOME s2) = | |
| 277 | let | |
| 278 | val ctxt' = Config.put show_sorts true ctxt; | |
| 279 | val raw_T = Syntax.parse_typ ctxt' s1; | |
| 280 | val S = Syntax.read_sort ctxt' s2; | |
| 281 | val T = | |
| 282 | Syntax.check_term ctxt' | |
| 283 | (Logic.mk_type raw_T |> Type.constraint (Term.itselfT (Type_Infer.anyT S))) | |
| 284 | |> Logic.dest_type; | |
| 285 | in Pretty.string_of (Pretty.quote (Syntax.pretty_typ ctxt' T)) end; | |
| 9128 | 286 | |
| 23935 | 287 | fun print_item string_of (modes, arg) = Toplevel.keep (fn state => | 
| 38331 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 288 | Print_Mode.with_modes modes (fn () => writeln (string_of state arg)) ()); | 
| 19385 
c0f2f8224ea8
print_term etc.: actually observe print mode in final output;
 wenzelm parents: 
19268diff
changeset | 289 | |
| 
c0f2f8224ea8
print_term etc.: actually observe print mode in final output;
 wenzelm parents: 
19268diff
changeset | 290 | in | 
| 9128 | 291 | |
| 38331 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 292 | val print_stmts = print_item (string_of_stmts o Toplevel.context_of); | 
| 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 293 | val print_thms = print_item (string_of_thms o Toplevel.context_of); | 
| 19385 
c0f2f8224ea8
print_term etc.: actually observe print mode in final output;
 wenzelm parents: 
19268diff
changeset | 294 | val print_prfs = print_item o string_of_prfs; | 
| 38331 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 295 | val print_prop = print_item (string_of_prop o Toplevel.context_of); | 
| 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 296 | val print_term = print_item (string_of_term o Toplevel.context_of); | 
| 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 297 | val print_type = print_item (string_of_type o Toplevel.context_of); | 
| 5831 | 298 | |
| 19385 
c0f2f8224ea8
print_term etc.: actually observe print mode in final output;
 wenzelm parents: 
19268diff
changeset | 299 | end; | 
| 
c0f2f8224ea8
print_term etc.: actually observe print mode in final output;
 wenzelm parents: 
19268diff
changeset | 300 | |
| 5831 | 301 | end; |