| author | wenzelm | 
| Thu, 11 Oct 2012 15:06:27 +0200 | |
| changeset 49817 | 85b37aece3b3 | 
| parent 49569 | 7b6aaf446496 | 
| child 49866 | 619acbd72664 | 
| 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 | |
| 30575 | 9 | val global_setup: Symbol_Pos.text * Position.T -> theory -> theory | 
| 10 | val local_setup: Symbol_Pos.text * Position.T -> Proof.context -> Proof.context | |
| 11 | val parse_ast_translation: bool * (Symbol_Pos.text * Position.T) -> theory -> theory | |
| 12 | val parse_translation: bool * (Symbol_Pos.text * Position.T) -> theory -> theory | |
| 13 | val print_translation: bool * (Symbol_Pos.text * Position.T) -> theory -> theory | |
| 14 | val typed_print_translation: bool * (Symbol_Pos.text * Position.T) -> theory -> theory | |
| 15 | val print_ast_translation: bool * (Symbol_Pos.text * Position.T) -> theory -> theory | |
| 42204 | 16 | val translations: (xstring * string) Syntax.trrule list -> theory -> theory | 
| 17 | val no_translations: (xstring * string) Syntax.trrule list -> theory -> theory | |
| 30573 | 18 | val oracle: bstring * Position.T -> Symbol_Pos.text * Position.T -> theory -> theory | 
| 35852 
4e3fe0b8687b
minor renovation of old-style 'axioms' -- make it an alias of iterated 'axiomatization';
 wenzelm parents: 
35141diff
changeset | 19 | val add_axioms: (Attrib.binding * string) list -> theory -> theory | 
| 29579 | 20 | val add_defs: (bool * bool) * ((binding * string) * Attrib.src list) list -> theory -> theory | 
| 40784 | 21 |   val declaration: {syntax: bool, pervasive: bool} ->
 | 
| 22 | Symbol_Pos.text * Position.T -> local_theory -> local_theory | |
| 42464 | 23 | val simproc_setup: string * Position.T -> string list -> Symbol_Pos.text * Position.T -> | 
| 24 | string list -> local_theory -> local_theory | |
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36174diff
changeset | 25 | val hide_class: bool -> xstring list -> theory -> theory | 
| 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36174diff
changeset | 26 | val hide_type: bool -> xstring list -> theory -> theory | 
| 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36174diff
changeset | 27 | val hide_const: bool -> xstring list -> theory -> theory | 
| 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36174diff
changeset | 28 | val hide_fact: bool -> xstring list -> theory -> theory | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 29 | val have: (Attrib.binding * (string * string list) list) list -> bool -> Proof.state -> Proof.state | 
| 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 30 | val hence: (Attrib.binding * (string * string list) list) list -> bool -> Proof.state -> Proof.state | 
| 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 31 | val show: (Attrib.binding * (string * string list) list) list -> bool -> Proof.state -> Proof.state | 
| 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 32 | val thus: (Attrib.binding * (string * string list) list) list -> bool -> Proof.state -> Proof.state | 
| 21350 | 33 | val qed: Method.text option -> Toplevel.transition -> Toplevel.transition | 
| 34 | val terminal_proof: Method.text * Method.text option -> | |
| 35 | Toplevel.transition -> Toplevel.transition | |
| 36 | val default_proof: Toplevel.transition -> Toplevel.transition | |
| 37 | val immediate_proof: Toplevel.transition -> Toplevel.transition | |
| 38 | val done_proof: Toplevel.transition -> Toplevel.transition | |
| 39 | val skip_proof: Toplevel.transition -> Toplevel.transition | |
| 30575 | 40 | val ml_diag: bool -> Symbol_Pos.text * Position.T -> 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 | 41 | 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 | 42 |   val diag_goal: Proof.context -> {context: Proof.context, facts: thm list, goal: thm}
 | 
| 48881 
46e053eda5dd
clarified Parse.path vs. Parse.explode -- prefer errors in proper transaction context;
 wenzelm parents: 
48792diff
changeset | 43 | val display_drafts: string list -> Toplevel.transition -> Toplevel.transition | 
| 
46e053eda5dd
clarified Parse.path vs. Parse.explode -- prefer errors in proper transaction context;
 wenzelm parents: 
48792diff
changeset | 44 | val print_drafts: string list -> Toplevel.transition -> Toplevel.transition | 
| 7308 | 45 | val print_context: Toplevel.transition -> Toplevel.transition | 
| 20621 | 46 | val print_theory: bool -> Toplevel.transition -> Toplevel.transition | 
| 5831 | 47 | val print_syntax: Toplevel.transition -> Toplevel.transition | 
| 21725 | 48 | val print_abbrevs: Toplevel.transition -> Toplevel.transition | 
| 21003 
37492b0062c6
renamed print_lthms to print_facts, do not insist on proof state;
 wenzelm parents: 
20978diff
changeset | 49 | val print_facts: Toplevel.transition -> Toplevel.transition | 
| 24115 | 50 | val print_configs: Toplevel.transition -> Toplevel.transition | 
| 33515 
d066e8369a33
print_theorems: suppress concealed (global) facts, unless "!" option is given;
 wenzelm parents: 
33456diff
changeset | 51 | val print_theorems: bool -> Toplevel.transition -> Toplevel.transition | 
| 12060 | 52 | val print_locales: Toplevel.transition -> Toplevel.transition | 
| 46922 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
45666diff
changeset | 53 | val print_locale: bool * (xstring * Position.T) -> Toplevel.transition -> Toplevel.transition | 
| 
3717f3878714
source positions for locale and class expressions;
 wenzelm parents: 
45666diff
changeset | 54 | val print_registrations: xstring * Position.T -> Toplevel.transition -> Toplevel.transition | 
| 41435 | 55 | val print_dependencies: bool * Expression.expression -> Toplevel.transition | 
| 56 | -> Toplevel.transition | |
| 5831 | 57 | val print_attributes: Toplevel.transition -> Toplevel.transition | 
| 16026 | 58 | val print_simpset: Toplevel.transition -> Toplevel.transition | 
| 12382 | 59 | val print_rules: Toplevel.transition -> Toplevel.transition | 
| 9219 | 60 | val print_trans_rules: Toplevel.transition -> Toplevel.transition | 
| 5831 | 61 | val print_methods: Toplevel.transition -> Toplevel.transition | 
| 9219 | 62 | val print_antiquotations: Toplevel.transition -> Toplevel.transition | 
| 49569 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 wenzelm parents: 
49561diff
changeset | 63 | val thy_deps: Toplevel.transition -> Toplevel.transition | 
| 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 wenzelm parents: 
49561diff
changeset | 64 | val locale_deps: Toplevel.transition -> Toplevel.transition | 
| 20574 | 65 | val class_deps: Toplevel.transition -> Toplevel.transition | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26186diff
changeset | 66 | val thm_deps: (Facts.ref * Attrib.src list) list -> Toplevel.transition -> Toplevel.transition | 
| 26184 | 67 | val unused_thms: (string list * string list option) option -> | 
| 68 | Toplevel.transition -> Toplevel.transition | |
| 5831 | 69 | val print_binds: Toplevel.transition -> Toplevel.transition | 
| 8369 | 70 | val print_cases: Toplevel.transition -> Toplevel.transition | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26186diff
changeset | 71 | val print_stmts: string list * (Facts.ref * Attrib.src list) list | 
| 19268 | 72 | -> Toplevel.transition -> Toplevel.transition | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26186diff
changeset | 73 | val print_thms: string list * (Facts.ref * Attrib.src list) list | 
| 10581 | 74 | -> Toplevel.transition -> Toplevel.transition | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26186diff
changeset | 75 | val print_prfs: bool -> string list * (Facts.ref * Attrib.src list) list option | 
| 11524 
197f2e14a714
Added functions for printing primitive proof terms.
 berghofe parents: 
11017diff
changeset | 76 | -> Toplevel.transition -> Toplevel.transition | 
| 12876 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 wenzelm parents: 
12758diff
changeset | 77 | 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 | 78 | val print_term: (string list * string) -> Toplevel.transition -> Toplevel.transition | 
| 48792 | 79 | val print_type: (string list * (string * string option)) -> | 
| 80 | Toplevel.transition -> Toplevel.transition | |
| 30573 | 81 | val header_markup: Symbol_Pos.text * Position.T -> Toplevel.transition -> Toplevel.transition | 
| 45488 
6d71d9e52369
pass positions for named targets, for formal links in the document model;
 wenzelm parents: 
45291diff
changeset | 82 | val local_theory_markup: (xstring * Position.T) option * (Symbol_Pos.text * Position.T) -> | 
| 17262 
63cf42df2723
add_chapter/section/subsection/subsubsection/text: optional locale specification;
 wenzelm parents: 
17228diff
changeset | 83 | Toplevel.transition -> Toplevel.transition | 
| 30573 | 84 | val proof_markup: Symbol_Pos.text * Position.T -> Toplevel.transition -> Toplevel.transition | 
| 5831 | 85 | end; | 
| 86 | ||
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
37198diff
changeset | 87 | structure Isar_Cmd: ISAR_CMD = | 
| 5831 | 88 | struct | 
| 89 | ||
| 90 | ||
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 91 | (** theory declarations **) | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 92 | |
| 30461 | 93 | (* generic setup *) | 
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 94 | |
| 30461 | 95 | fun global_setup (txt, pos) = | 
| 37198 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 96 | ML_Lex.read pos txt | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 97 | |> ML_Context.expression pos "val setup: theory -> theory" "Context.map_theory setup" | 
| 26435 | 98 | |> Context.theory_map; | 
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 99 | |
| 30461 | 100 | fun local_setup (txt, pos) = | 
| 37198 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 101 | ML_Lex.read pos txt | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 102 | |> ML_Context.expression pos "val setup: local_theory -> local_theory" "Context.map_proof setup" | 
| 30461 | 103 | |> Context.proof_map; | 
| 104 | ||
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 105 | |
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 106 | (* translation functions *) | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 107 | |
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 108 | local | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 109 | |
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 110 | fun advancedT false = "" | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 111 | | advancedT true = "Proof.context -> "; | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 112 | |
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 113 | fun advancedN false = "" | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 114 | | advancedN true = "advanced_"; | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 115 | |
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 116 | in | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 117 | |
| 26385 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 wenzelm parents: 
26336diff
changeset | 118 | fun parse_ast_translation (a, (txt, pos)) = | 
| 37198 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 119 | ML_Lex.read pos txt | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 120 | |> ML_Context.expression pos | 
| 26455 | 121 |     ("val parse_ast_translation: (string * (" ^ advancedT a ^
 | 
| 42224 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 122 | "Ast.ast list -> Ast.ast)) list") | 
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 123 |     ("Context.map_theory (Sign.add_" ^ advancedN a ^ "trfuns (parse_ast_translation, [], [], []))")
 | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 124 | |> Context.theory_map; | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 125 | |
| 26385 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 wenzelm parents: 
26336diff
changeset | 126 | fun parse_translation (a, (txt, pos)) = | 
| 37198 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 127 | ML_Lex.read pos txt | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 128 | |> ML_Context.expression pos | 
| 26455 | 129 |     ("val parse_translation: (string * (" ^ advancedT a ^
 | 
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 130 | "term list -> term)) list") | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 131 |     ("Context.map_theory (Sign.add_" ^ advancedN a ^ "trfuns ([], parse_translation, [], []))")
 | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 132 | |> Context.theory_map; | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 133 | |
| 26385 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 wenzelm parents: 
26336diff
changeset | 134 | fun print_translation (a, (txt, pos)) = | 
| 37198 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 135 | ML_Lex.read pos txt | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 136 | |> ML_Context.expression pos | 
| 26455 | 137 |     ("val print_translation: (string * (" ^ advancedT a ^
 | 
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 138 | "term list -> term)) list") | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 139 |     ("Context.map_theory (Sign.add_" ^ advancedN a ^ "trfuns ([], [], print_translation, []))")
 | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 140 | |> Context.theory_map; | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 141 | |
| 26385 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 wenzelm parents: 
26336diff
changeset | 142 | fun print_ast_translation (a, (txt, pos)) = | 
| 37198 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 143 | ML_Lex.read pos txt | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 144 | |> ML_Context.expression pos | 
| 26455 | 145 |     ("val print_ast_translation: (string * (" ^ advancedT a ^
 | 
| 42224 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42204diff
changeset | 146 | "Ast.ast list -> Ast.ast)) list") | 
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 147 |     ("Context.map_theory (Sign.add_" ^ advancedN a ^ "trfuns ([], [], [], print_ast_translation))")
 | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 148 | |> Context.theory_map; | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 149 | |
| 26385 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 wenzelm parents: 
26336diff
changeset | 150 | fun typed_print_translation (a, (txt, pos)) = | 
| 37198 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 151 | ML_Lex.read pos txt | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 152 | |> ML_Context.expression pos | 
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42243diff
changeset | 153 |     ("val typed_print_translation: (string * (" ^ advancedT a ^ "typ -> term list -> term)) list")
 | 
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 154 |     ("Context.map_theory (Sign.add_" ^ advancedN a ^ "trfunsT typed_print_translation)")
 | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 155 | |> Context.theory_map; | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 156 | |
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 157 | end; | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 158 | |
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 159 | |
| 42204 | 160 | (* translation rules *) | 
| 161 | ||
| 162 | fun read_trrules thy raw_rules = | |
| 163 | let | |
| 42360 | 164 | val ctxt = Proof_Context.init_global thy; | 
| 42204 | 165 | in | 
| 166 | raw_rules |> map (Syntax.map_trrule (fn (r, s) => | |
| 42360 | 167 | Syntax_Phases.parse_ast_pattern ctxt (Proof_Context.intern_type ctxt r, s))) | 
| 42204 | 168 | end; | 
| 169 | ||
| 170 | fun translations args thy = Sign.add_trrules (read_trrules thy args) thy; | |
| 171 | fun no_translations args thy = Sign.del_trrules (read_trrules thy args) thy; | |
| 172 | ||
| 173 | ||
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 174 | (* oracles *) | 
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 175 | |
| 30334 | 176 | fun oracle (name, pos) (body_txt, body_pos) = | 
| 27871 
4ef76f8788ad
oracle, header/local_theory/proof_markup: pass SymbolPos.text;
 wenzelm parents: 
27853diff
changeset | 177 | let | 
| 37198 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 178 | val body = ML_Lex.read body_pos body_txt; | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 179 | val ants = | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 180 | ML_Lex.read Position.none | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 181 |        ("local\n\
 | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 182 | \ val binding = " ^ ML_Syntax.make_binding (name, pos) ^ ";\n\ | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 183 |         \  val body = ") @ body @ ML_Lex.read Position.none (";\n\
 | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 184 | \in\n\ | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 185 | \ val " ^ name ^ " = snd (Context.>>> (Context.map_theory_result (Thm.add_oracle (binding, body))));\n\ | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 186 | \end;\n"); | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 187 | in Context.theory_map (ML_Context.exec (fn () => ML_Context.eval false body_pos ants)) end; | 
| 22116 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 188 | |
| 
6917be2e647d
added various ML setup functions (from sign.ML, pure_thy.ML);
 wenzelm parents: 
22087diff
changeset | 189 | |
| 35852 
4e3fe0b8687b
minor renovation of old-style 'axioms' -- make it an alias of iterated 'axiomatization';
 wenzelm parents: 
35141diff
changeset | 190 | (* old-style axioms *) | 
| 21350 | 191 | |
| 35894 | 192 | val add_axioms = fold (snd oo Specification.axiom_cmd); | 
| 21350 | 193 | |
| 35852 
4e3fe0b8687b
minor renovation of old-style 'axioms' -- make it an alias of iterated 'axiomatization';
 wenzelm parents: 
35141diff
changeset | 194 | fun add_defs ((unchecked, overloaded), args) thy = | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39507diff
changeset | 195 | thy |> | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39507diff
changeset | 196 | (if unchecked then Global_Theory.add_defs_unchecked_cmd else Global_Theory.add_defs_cmd) | 
| 47815 | 197 | overloaded | 
| 198 | (map (fn ((b, ax), srcs) => ((b, ax), map (Attrib.attribute_cmd_global thy) srcs)) args) | |
| 35852 
4e3fe0b8687b
minor renovation of old-style 'axioms' -- make it an alias of iterated 'axiomatization';
 wenzelm parents: 
35141diff
changeset | 199 | |> snd; | 
| 21350 | 200 | |
| 201 | ||
| 22087 | 202 | (* declarations *) | 
| 203 | ||
| 40784 | 204 | fun declaration {syntax, pervasive} (txt, pos) =
 | 
| 37198 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 205 | ML_Lex.read pos txt | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 206 | |> ML_Context.expression pos | 
| 26455 | 207 | "val declaration: Morphism.declaration" | 
| 45291 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 wenzelm parents: 
44338diff
changeset | 208 |     ("Context.map_proof (Local_Theory.declaration {syntax = " ^ Bool.toString syntax ^ ", \
 | 
| 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 wenzelm parents: 
44338diff
changeset | 209 | \pervasive = " ^ Bool.toString pervasive ^ "} declaration)") | 
| 26385 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 wenzelm parents: 
26336diff
changeset | 210 | |> Context.proof_map; | 
| 22087 | 211 | |
| 212 | ||
| 22202 | 213 | (* simprocs *) | 
| 214 | ||
| 37198 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 215 | fun simproc_setup name lhss (txt, pos) identifier = | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 216 | ML_Lex.read pos txt | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 217 | |> ML_Context.expression pos | 
| 22239 | 218 | "val proc: Morphism.morphism -> Simplifier.simpset -> cterm -> thm option" | 
| 42464 | 219 |     ("Context.map_proof (Simplifier.def_simproc_cmd {name = " ^ ML_Syntax.make_binding name ^ ", \
 | 
| 37198 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 220 | \lhss = " ^ ML_Syntax.print_strings lhss ^ ", proc = proc, \ | 
| 
3af985b10550
replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
 wenzelm parents: 
37146diff
changeset | 221 | \identifier = Library.maps ML_Context.thms " ^ ML_Syntax.print_strings identifier ^ "})") | 
| 22202 | 222 | |> Context.proof_map; | 
| 223 | ||
| 224 | ||
| 26671 
c95590e01df5
added hide_names command (formerly Sign.hide_names), support fact name space;
 wenzelm parents: 
26626diff
changeset | 225 | (* hide names *) | 
| 
c95590e01df5
added hide_names command (formerly Sign.hide_names), support fact name space;
 wenzelm parents: 
26626diff
changeset | 226 | |
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36174diff
changeset | 227 | fun hide_names intern check hide fully xnames thy = | 
| 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36174diff
changeset | 228 | let | 
| 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36174diff
changeset | 229 | val names = map (intern thy) xnames; | 
| 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36174diff
changeset | 230 | val bads = filter_out (check thy) names; | 
| 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36174diff
changeset | 231 | in | 
| 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36174diff
changeset | 232 | if null bads then fold (hide fully) names thy | 
| 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36174diff
changeset | 233 |     else error ("Attempt to hide undeclared item(s): " ^ commas_quote bads)
 | 
| 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36174diff
changeset | 234 | end; | 
| 26671 
c95590e01df5
added hide_names command (formerly Sign.hide_names), support fact name space;
 wenzelm parents: 
26626diff
changeset | 235 | |
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36174diff
changeset | 236 | val hide_class = hide_names Sign.intern_class (can o Sign.certify_class) Sign.hide_class; | 
| 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36174diff
changeset | 237 | val hide_type = hide_names Sign.intern_type Sign.declared_tyname Sign.hide_type; | 
| 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36174diff
changeset | 238 | val hide_const = hide_names Sign.intern_const Sign.declared_const Sign.hide_const; | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39507diff
changeset | 239 | val hide_fact = | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39507diff
changeset | 240 | hide_names Global_Theory.intern_fact Global_Theory.defined_fact Global_Theory.hide_fact; | 
| 26671 
c95590e01df5
added hide_names command (formerly Sign.hide_names), support fact name space;
 wenzelm parents: 
26626diff
changeset | 241 | |
| 
c95590e01df5
added hide_names command (formerly Sign.hide_names), support fact name space;
 wenzelm parents: 
26626diff
changeset | 242 | |
| 21350 | 243 | (* goals *) | 
| 244 | ||
| 245 | fun goal opt_chain goal stmt int = | |
| 29383 | 246 | opt_chain #> goal NONE (K I) stmt int; | 
| 21350 | 247 | |
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36178diff
changeset | 248 | val have = goal I Proof.have_cmd; | 
| 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36178diff
changeset | 249 | val hence = goal Proof.chain Proof.have_cmd; | 
| 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36178diff
changeset | 250 | val show = goal I Proof.show_cmd; | 
| 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36178diff
changeset | 251 | val thus = goal Proof.chain Proof.show_cmd; | 
| 21350 | 252 | |
| 253 | ||
| 254 | (* local endings *) | |
| 255 | ||
| 29383 | 256 | fun local_qed m = Toplevel.proof (Proof.local_qed (m, true)); | 
| 32061 
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
 wenzelm parents: 
31819diff
changeset | 257 | val local_terminal_proof = Toplevel.proof' o Proof.local_future_terminal_proof; | 
| 29383 | 258 | val local_default_proof = Toplevel.proof Proof.local_default_proof; | 
| 259 | val local_immediate_proof = Toplevel.proof Proof.local_immediate_proof; | |
| 260 | val local_done_proof = Toplevel.proof Proof.local_done_proof; | |
| 261 | val local_skip_proof = Toplevel.proof' Proof.local_skip_proof; | |
| 21350 | 262 | |
| 27562 | 263 | val skip_local_qed = Toplevel.skip_proof (fn i => if i > 1 then i - 1 else raise Toplevel.UNDEF); | 
| 21350 | 264 | |
| 265 | ||
| 266 | (* global endings *) | |
| 267 | ||
| 268 | fun global_qed m = Toplevel.end_proof (K (Proof.global_qed (m, true))); | |
| 49012 
8686c36fa27d
refined treatment of forked proofs at transaction boundaries, including proof commands (see also 7ee000ce5390);
 wenzelm parents: 
48918diff
changeset | 269 | val global_terminal_proof = Toplevel.end_proof o Proof.global_future_terminal_proof; | 
| 21350 | 270 | val global_default_proof = Toplevel.end_proof (K Proof.global_default_proof); | 
| 271 | val global_immediate_proof = Toplevel.end_proof (K Proof.global_immediate_proof); | |
| 272 | val global_skip_proof = Toplevel.end_proof Proof.global_skip_proof; | |
| 273 | val global_done_proof = Toplevel.end_proof (K Proof.global_done_proof); | |
| 274 | ||
| 28375 | 275 | val skip_global_qed = Toplevel.skip_proof_to_theory (fn n => n = 1); | 
| 21350 | 276 | |
| 277 | ||
| 278 | (* common endings *) | |
| 279 | ||
| 280 | fun qed m = local_qed m o global_qed m o skip_local_qed o skip_global_qed; | |
| 281 | fun terminal_proof m = local_terminal_proof m o global_terminal_proof m; | |
| 282 | val default_proof = local_default_proof o global_default_proof; | |
| 283 | val immediate_proof = local_immediate_proof o global_immediate_proof; | |
| 284 | val done_proof = local_done_proof o global_done_proof; | |
| 285 | val skip_proof = local_skip_proof o global_skip_proof; | |
| 286 | ||
| 287 | ||
| 26489 | 288 | (* diagnostic ML evaluation *) | 
| 5831 | 289 | |
| 37305 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 290 | 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 | 291 | ( | 
| 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 292 | type T = Toplevel.state; | 
| 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 293 | fun init _ = Toplevel.toplevel; | 
| 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 294 | ); | 
| 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 295 | |
| 26489 | 296 | fun ml_diag verbose (txt, pos) = Toplevel.keep (fn state => | 
| 37305 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 297 | let val opt_ctxt = | 
| 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 298 | try Toplevel.generic_theory_of state | 
| 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 299 | |> Option.map (Context.proof_of #> Diag_State.put state) | 
| 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 300 | in ML_Context.eval_text_in opt_ctxt verbose pos txt end); | 
| 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 301 | |
| 48776 
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
 wenzelm parents: 
47815diff
changeset | 302 | val diag_state = Diag_State.get; | 
| 37305 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 303 | |
| 48776 
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
 wenzelm parents: 
47815diff
changeset | 304 | fun diag_goal ctxt = | 
| 
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
 wenzelm parents: 
47815diff
changeset | 305 | Proof.goal (Toplevel.proof_of (diag_state ctxt)) | 
| 37305 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 306 | handle Toplevel.UNDEF => error "No goal present"; | 
| 
9763792e4ac7
diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
 wenzelm parents: 
37216diff
changeset | 307 | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42464diff
changeset | 308 | val _ = | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42464diff
changeset | 309 | Context.>> (Context.map_theory | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42464diff
changeset | 310 | (ML_Antiquote.value (Binding.qualify true "Isar" (Binding.name "state")) | 
| 48776 
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
 wenzelm parents: 
47815diff
changeset | 311 | (Scan.succeed "Isar_Cmd.diag_state ML_context") #> | 
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42464diff
changeset | 312 | ML_Antiquote.value (Binding.qualify true "Isar" (Binding.name "goal")) | 
| 48776 
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
 wenzelm parents: 
47815diff
changeset | 313 | (Scan.succeed "Isar_Cmd.diag_goal ML_context"))); | 
| 5831 | 314 | |
| 315 | ||
| 14934 | 316 | (* present draft files *) | 
| 317 | ||
| 48881 
46e053eda5dd
clarified Parse.path vs. Parse.explode -- prefer errors in proper transaction context;
 wenzelm parents: 
48792diff
changeset | 318 | fun display_drafts names = Toplevel.imperative (fn () => | 
| 
46e053eda5dd
clarified Parse.path vs. Parse.explode -- prefer errors in proper transaction context;
 wenzelm parents: 
48792diff
changeset | 319 | let | 
| 
46e053eda5dd
clarified Parse.path vs. Parse.explode -- prefer errors in proper transaction context;
 wenzelm parents: 
48792diff
changeset | 320 | val paths = map Path.explode names; | 
| 
46e053eda5dd
clarified Parse.path vs. Parse.explode -- prefer errors in proper transaction context;
 wenzelm parents: 
48792diff
changeset | 321 | val outfile = File.shell_path (Present.drafts (getenv "ISABELLE_DOC_FORMAT") paths); | 
| 40743 | 322 |   in Isabelle_System.isabelle_tool "display" ("-c " ^ outfile ^ " &"); () end);
 | 
| 14934 | 323 | |
| 48881 
46e053eda5dd
clarified Parse.path vs. Parse.explode -- prefer errors in proper transaction context;
 wenzelm parents: 
48792diff
changeset | 324 | fun print_drafts names = Toplevel.imperative (fn () => | 
| 
46e053eda5dd
clarified Parse.path vs. Parse.explode -- prefer errors in proper transaction context;
 wenzelm parents: 
48792diff
changeset | 325 | let | 
| 
46e053eda5dd
clarified Parse.path vs. Parse.explode -- prefer errors in proper transaction context;
 wenzelm parents: 
48792diff
changeset | 326 | val paths = map Path.explode names; | 
| 
46e053eda5dd
clarified Parse.path vs. Parse.explode -- prefer errors in proper transaction context;
 wenzelm parents: 
48792diff
changeset | 327 | val outfile = File.shell_path (Present.drafts "ps" paths); | 
| 40743 | 328 |   in Isabelle_System.isabelle_tool "print" ("-c " ^ outfile); () end);
 | 
| 14934 | 329 | |
| 330 | ||
| 9513 | 331 | (* print parts of theory and proof context *) | 
| 5831 | 332 | |
| 7308 | 333 | val print_context = Toplevel.keep Toplevel.print_state_context; | 
| 9513 | 334 | |
| 20621 | 335 | fun print_theory verbose = Toplevel.unknown_theory o | 
| 33389 | 336 | Toplevel.keep (Pretty.writeln o Proof_Display.pretty_full_theory verbose o Toplevel.theory_of); | 
| 9513 | 337 | |
| 21663 | 338 | val print_syntax = Toplevel.unknown_context o | 
| 42360 | 339 | Toplevel.keep (Proof_Context.print_syntax o Toplevel.context_of); | 
| 9513 | 340 | |
| 21725 | 341 | val print_abbrevs = Toplevel.unknown_context o | 
| 42360 | 342 | Toplevel.keep (Proof_Context.print_abbrevs o Toplevel.context_of); | 
| 21725 | 343 | |
| 35141 | 344 | val print_facts = Toplevel.unknown_context o | 
| 42360 | 345 | Toplevel.keep (Proof_Context.print_lthms o Toplevel.context_of); | 
| 21003 
37492b0062c6
renamed print_lthms to print_facts, do not insist on proof state;
 wenzelm parents: 
20978diff
changeset | 346 | |
| 35141 | 347 | val print_configs = Toplevel.unknown_context o | 
| 348 | Toplevel.keep (Attrib.print_configs o Toplevel.context_of); | |
| 23989 | 349 | |
| 35141 | 350 | val print_theorems_proof = | 
| 42360 | 351 | Toplevel.keep (Proof_Context.print_lthms o Proof.context_of o Toplevel.proof_of); | 
| 17066 | 352 | |
| 33515 
d066e8369a33
print_theorems: suppress concealed (global) facts, unless "!" option is given;
 wenzelm parents: 
33456diff
changeset | 353 | fun print_theorems_theory verbose = Toplevel.keep (fn state => | 
| 18588 | 354 | Toplevel.theory_of state |> | 
| 30801 | 355 | (case Toplevel.previous_context_of state of | 
| 42360 | 356 | SOME prev => Proof_Display.print_theorems_diff verbose (Proof_Context.theory_of prev) | 
| 33515 
d066e8369a33
print_theorems: suppress concealed (global) facts, unless "!" option is given;
 wenzelm parents: 
33456diff
changeset | 357 | | NONE => Proof_Display.print_theorems verbose)); | 
| 18588 | 358 | |
| 33515 
d066e8369a33
print_theorems: suppress concealed (global) facts, unless "!" option is given;
 wenzelm parents: 
33456diff
changeset | 359 | fun print_theorems verbose = | 
| 
d066e8369a33
print_theorems: suppress concealed (global) facts, unless "!" option is given;
 wenzelm parents: 
33456diff
changeset | 360 | Toplevel.unknown_context o print_theorems_theory verbose o print_theorems_proof; | 
| 9513 | 361 | |
| 12060 | 362 | val print_locales = Toplevel.unknown_theory o | 
| 29360 | 363 | Toplevel.keep (Locale.print_locales o Toplevel.theory_of); | 
| 12060 | 364 | |
| 33515 
d066e8369a33
print_theorems: suppress concealed (global) facts, unless "!" option is given;
 wenzelm parents: 
33456diff
changeset | 365 | fun print_locale (verbose, name) = Toplevel.unknown_theory o | 
| 
d066e8369a33
print_theorems: suppress concealed (global) facts, unless "!" option is given;
 wenzelm parents: 
33456diff
changeset | 366 | Toplevel.keep (fn state => Locale.print_locale (Toplevel.theory_of state) verbose name); | 
| 15596 | 367 | |
| 32804 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarin parents: 
32149diff
changeset | 368 | fun print_registrations name = Toplevel.unknown_context o | 
| 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarin parents: 
32149diff
changeset | 369 | Toplevel.keep (fn state => | 
| 38109 | 370 | Locale.print_registrations (Toplevel.context_of state) name); | 
| 32804 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 ballarin parents: 
32149diff
changeset | 371 | |
| 41435 | 372 | fun print_dependencies (clean, expression) = Toplevel.unknown_context o | 
| 373 | Toplevel.keep (fn state => | |
| 374 | Expression.print_dependencies (Toplevel.context_of state) clean expression); | |
| 375 | ||
| 9513 | 376 | val print_attributes = Toplevel.unknown_theory o | 
| 377 | Toplevel.keep (Attrib.print_attributes o Toplevel.theory_of); | |
| 378 | ||
| 16026 | 379 | val print_simpset = Toplevel.unknown_context o | 
| 30357 | 380 | Toplevel.keep (fn state => | 
| 381 | let val ctxt = Toplevel.context_of state | |
| 32149 
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 wenzelm parents: 
32091diff
changeset | 382 | in Pretty.writeln (Simplifier.pretty_ss ctxt (simpset_of ctxt)) end); | 
| 16026 | 383 | |
| 12382 | 384 | val print_rules = Toplevel.unknown_context o | 
| 33369 | 385 | Toplevel.keep (Context_Rules.print_rules o Toplevel.context_of); | 
| 12382 | 386 | |
| 9513 | 387 | val print_trans_rules = Toplevel.unknown_context o | 
| 18639 | 388 | Toplevel.keep (Calculation.print_rules o Toplevel.context_of); | 
| 9513 | 389 | |
| 390 | val print_methods = Toplevel.unknown_theory o | |
| 391 | Toplevel.keep (Method.print_methods o Toplevel.theory_of); | |
| 392 | ||
| 43564 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43560diff
changeset | 393 | val print_antiquotations = | 
| 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43560diff
changeset | 394 | Toplevel.keep (Thy_Output.print_antiquotations o Toplevel.context_of); | 
| 5831 | 395 | |
| 22485 | 396 | val thy_deps = Toplevel.unknown_theory o Toplevel.keep (fn state => | 
| 397 | let | |
| 398 | val thy = Toplevel.theory_of state; | |
| 37866 
cd1d1bc7684c
thy_deps: more direct comparison of sessions, which is presumably what "unfold" is meant to indicate here -- also avoid referring to accidental theory loader state;
 wenzelm parents: 
37305diff
changeset | 399 | val thy_session = Present.session_name thy; | 
| 
cd1d1bc7684c
thy_deps: more direct comparison of sessions, which is presumably what "unfold" is meant to indicate here -- also avoid referring to accidental theory loader state;
 wenzelm parents: 
37305diff
changeset | 400 | |
| 42425 | 401 | val gr = rev (Theory.nodes_of thy) |> map (fn node => | 
| 22602 
a165d9ed08b8
simplified thy_deps using Theory.ancestors_of (in order of creation);
 wenzelm parents: 
22573diff
changeset | 402 | let | 
| 
a165d9ed08b8
simplified thy_deps using Theory.ancestors_of (in order of creation);
 wenzelm parents: 
22573diff
changeset | 403 | val name = Context.theory_name node; | 
| 
a165d9ed08b8
simplified thy_deps using Theory.ancestors_of (in order of creation);
 wenzelm parents: 
22573diff
changeset | 404 | val parents = map Context.theory_name (Theory.parents_of node); | 
| 37866 
cd1d1bc7684c
thy_deps: more direct comparison of sessions, which is presumably what "unfold" is meant to indicate here -- also avoid referring to accidental theory loader state;
 wenzelm parents: 
37305diff
changeset | 405 | val session = Present.session_name node; | 
| 
cd1d1bc7684c
thy_deps: more direct comparison of sessions, which is presumably what "unfold" is meant to indicate here -- also avoid referring to accidental theory loader state;
 wenzelm parents: 
37305diff
changeset | 406 | val unfold = (session = thy_session); | 
| 49561 | 407 | in | 
| 408 |        {name = name, ID = name, parents = parents, dir = session,
 | |
| 409 | unfold = unfold, path = "", content = []} | |
| 410 | end); | |
| 411 | in Graph_Display.display_graph gr end); | |
| 22485 | 412 | |
| 49569 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 wenzelm parents: 
49561diff
changeset | 413 | val locale_deps = Toplevel.unknown_theory o Toplevel.keep (fn state => | 
| 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 wenzelm parents: 
49561diff
changeset | 414 | let | 
| 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 wenzelm parents: 
49561diff
changeset | 415 | val thy = Toplevel.theory_of state; | 
| 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 wenzelm parents: 
49561diff
changeset | 416 |     val gr = Locale.pretty_locale_deps thy |> map (fn {name, parents, body} =>
 | 
| 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 wenzelm parents: 
49561diff
changeset | 417 |      {name = Locale.extern thy name, ID = name, parents = parents,
 | 
| 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 wenzelm parents: 
49561diff
changeset | 418 | dir = "", unfold = true, path = "", content = [body]}); | 
| 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 wenzelm parents: 
49561diff
changeset | 419 | in Graph_Display.display_graph gr end); | 
| 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 wenzelm parents: 
49561diff
changeset | 420 | |
| 20574 | 421 | val class_deps = Toplevel.unknown_theory o Toplevel.keep (fn state => | 
| 422 | let | |
| 42358 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 wenzelm parents: 
42247diff
changeset | 423 | val ctxt = Toplevel.context_of state; | 
| 42360 | 424 |     val {classes = (space, algebra), ...} = Type.rep_tsig (Proof_Context.tsig_of ctxt);
 | 
| 36328 
4d9deabf6474
replaced Sorts.rep_algebra by slightly more abstract selectors classes_of/arities_of;
 wenzelm parents: 
36323diff
changeset | 425 | val classes = Sorts.classes_of algebra; | 
| 20574 | 426 | fun entry (c, (i, (_, cs))) = | 
| 44338 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 wenzelm parents: 
43564diff
changeset | 427 |       (i, {name = Name_Space.extern ctxt space c, ID = c, parents = Graph.Keys.dest cs,
 | 
| 49561 | 428 | dir = "", unfold = true, path = "", content = []}); | 
| 20574 | 429 | val gr = | 
| 430 | Graph.fold (cons o entry) classes [] | |
| 431 | |> sort (int_ord o pairself #1) |> map #2; | |
| 49561 | 432 | in Graph_Display.display_graph gr end); | 
| 20574 | 433 | |
| 9513 | 434 | fun thm_deps args = Toplevel.unknown_theory o Toplevel.keep (fn state => | 
| 37870 
dd9cfc512b7f
thm_deps/unused_thms: Context.get_theory based on proper theory ancestry, not accidental theory loader state;
 wenzelm parents: 
37866diff
changeset | 435 | Thm_Deps.thm_deps (Toplevel.theory_of state) | 
| 38331 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 436 | (Attrib.eval_thms (Toplevel.context_of state) args)); | 
| 9454 | 437 | |
| 5831 | 438 | |
| 26184 | 439 | (* find unused theorems *) | 
| 440 | ||
| 26186 
9af968b694d9
unused_thms: print via official context (ProofContext.pretty_fact),
 wenzelm parents: 
26184diff
changeset | 441 | fun unused_thms opt_range = Toplevel.keep (fn state => | 
| 
9af968b694d9
unused_thms: print via official context (ProofContext.pretty_fact),
 wenzelm parents: 
26184diff
changeset | 442 | let | 
| 
9af968b694d9
unused_thms: print via official context (ProofContext.pretty_fact),
 wenzelm parents: 
26184diff
changeset | 443 | val thy = Toplevel.theory_of state; | 
| 
9af968b694d9
unused_thms: print via official context (ProofContext.pretty_fact),
 wenzelm parents: 
26184diff
changeset | 444 | val ctxt = Toplevel.context_of state; | 
| 42360 | 445 | fun pretty_thm (a, th) = Proof_Context.pretty_fact ctxt (a, [th]); | 
| 37870 
dd9cfc512b7f
thm_deps/unused_thms: Context.get_theory based on proper theory ancestry, not accidental theory loader state;
 wenzelm parents: 
37866diff
changeset | 446 | val get_theory = Context.get_theory thy; | 
| 26186 
9af968b694d9
unused_thms: print via official context (ProofContext.pretty_fact),
 wenzelm parents: 
26184diff
changeset | 447 | in | 
| 33391 | 448 | Thm_Deps.unused_thms | 
| 26186 
9af968b694d9
unused_thms: print via official context (ProofContext.pretty_fact),
 wenzelm parents: 
26184diff
changeset | 449 | (case opt_range of | 
| 26694 | 450 | NONE => (Theory.parents_of thy, [thy]) | 
| 37870 
dd9cfc512b7f
thm_deps/unused_thms: Context.get_theory based on proper theory ancestry, not accidental theory loader state;
 wenzelm parents: 
37866diff
changeset | 451 | | SOME (xs, NONE) => (map get_theory xs, [thy]) | 
| 
dd9cfc512b7f
thm_deps/unused_thms: Context.get_theory based on proper theory ancestry, not accidental theory loader state;
 wenzelm parents: 
37866diff
changeset | 452 | | SOME (xs, SOME ys) => (map get_theory xs, map get_theory ys)) | 
| 26186 
9af968b694d9
unused_thms: print via official context (ProofContext.pretty_fact),
 wenzelm parents: 
26184diff
changeset | 453 | |> map pretty_thm |> Pretty.chunks |> Pretty.writeln | 
| 
9af968b694d9
unused_thms: print via official context (ProofContext.pretty_fact),
 wenzelm parents: 
26184diff
changeset | 454 | end); | 
| 26184 | 455 | |
| 456 | ||
| 5831 | 457 | (* print proof context contents *) | 
| 458 | ||
| 35141 | 459 | val print_binds = Toplevel.unknown_context o | 
| 42360 | 460 | Toplevel.keep (Proof_Context.print_binds o Toplevel.context_of); | 
| 9513 | 461 | |
| 35141 | 462 | val print_cases = Toplevel.unknown_context o | 
| 42360 | 463 | Toplevel.keep (Proof_Context.print_cases o Toplevel.context_of); | 
| 5831 | 464 | |
| 465 | ||
| 19268 | 466 | (* print theorems, terms, types etc. *) | 
| 467 | ||
| 19385 
c0f2f8224ea8
print_term etc.: actually observe print mode in final output;
 wenzelm parents: 
19268diff
changeset | 468 | local | 
| 
c0f2f8224ea8
print_term etc.: actually observe print mode in final output;
 wenzelm parents: 
19268diff
changeset | 469 | |
| 38331 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 470 | fun string_of_stmts ctxt args = | 
| 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 471 | Attrib.eval_thms ctxt args | 
| 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 472 | |> map (Element.pretty_statement ctxt Thm.theoremK) | 
| 19385 
c0f2f8224ea8
print_term etc.: actually observe print mode in final output;
 wenzelm parents: 
19268diff
changeset | 473 | |> Pretty.chunks2 |> Pretty.string_of; | 
| 5880 | 474 | |
| 38331 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 475 | fun string_of_thms ctxt args = | 
| 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 476 | Pretty.string_of (Display.pretty_thms ctxt (Attrib.eval_thms ctxt args)); | 
| 5895 | 477 | |
| 19385 
c0f2f8224ea8
print_term etc.: actually observe print mode in final output;
 wenzelm parents: 
19268diff
changeset | 478 | 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 | 479 | 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 | 480 | (case arg of | 
| 15531 | 481 | NONE => | 
| 12125 
316d11f760f7
Commands prf and full_prf can now also be used to display proof term
 berghofe parents: 
12069diff
changeset | 482 | let | 
| 38331 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 483 |           val {context = ctxt, goal = thm} = Proof.simple_goal (Toplevel.proof_of state);
 | 
| 42360 | 484 | val thy = Proof_Context.theory_of ctxt; | 
| 28814 | 485 | val prf = Thm.proof_of thm; | 
| 17066 | 486 | 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 | 487 | 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 | 488 | in | 
| 33388 | 489 | Proof_Syntax.pretty_proof ctxt | 
| 17066 | 490 | (if full then Reconstruct.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 | 491 | end | 
| 38331 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 492 | | SOME srcs => | 
| 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 493 | let val ctxt = Toplevel.context_of state | 
| 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 494 | in map (Proof_Syntax.pretty_proof_of ctxt full) (Attrib.eval_thms ctxt srcs) end | 
| 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 495 | |> Pretty.chunks); | 
| 11524 
197f2e14a714
Added functions for printing primitive proof terms.
 berghofe parents: 
11017diff
changeset | 496 | |
| 38331 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 497 | fun string_of_prop ctxt s = | 
| 5831 | 498 | let | 
| 24508 
c8b82fec6447
replaced ProofContext.read_term/prop by general Syntax.read_term/prop;
 wenzelm parents: 
24314diff
changeset | 499 | val prop = Syntax.read_prop ctxt s; | 
| 26704 
51ee753cc2e3
token translations: context dependent, result Pretty.T;
 wenzelm parents: 
26694diff
changeset | 500 | val ctxt' = Variable.auto_fixes prop ctxt; | 
| 
51ee753cc2e3
token translations: context dependent, result Pretty.T;
 wenzelm parents: 
26694diff
changeset | 501 | in Pretty.string_of (Pretty.quote (Syntax.pretty_term ctxt' prop)) end; | 
| 5831 | 502 | |
| 38331 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 503 | fun string_of_term ctxt s = | 
| 5831 | 504 | let | 
| 24508 
c8b82fec6447
replaced ProofContext.read_term/prop by general Syntax.read_term/prop;
 wenzelm parents: 
24314diff
changeset | 505 | val t = Syntax.read_term ctxt s; | 
| 5831 | 506 | val T = Term.type_of t; | 
| 26704 
51ee753cc2e3
token translations: context dependent, result Pretty.T;
 wenzelm parents: 
26694diff
changeset | 507 | val ctxt' = Variable.auto_fixes t ctxt; | 
| 5831 | 508 | in | 
| 19385 
c0f2f8224ea8
print_term etc.: actually observe print mode in final output;
 wenzelm parents: 
19268diff
changeset | 509 | Pretty.string_of | 
| 26704 
51ee753cc2e3
token translations: context dependent, result Pretty.T;
 wenzelm parents: 
26694diff
changeset | 510 | (Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t), Pretty.fbrk, | 
| 
51ee753cc2e3
token translations: context dependent, result Pretty.T;
 wenzelm parents: 
26694diff
changeset | 511 | Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' T)]) | 
| 9128 | 512 | end; | 
| 5831 | 513 | |
| 48792 | 514 | fun string_of_type ctxt (s, NONE) = | 
| 515 | let val T = Syntax.read_typ ctxt s | |
| 516 | in Pretty.string_of (Pretty.quote (Syntax.pretty_typ ctxt T)) end | |
| 517 | | string_of_type ctxt (s1, SOME s2) = | |
| 518 | let | |
| 519 | val ctxt' = Config.put show_sorts true ctxt; | |
| 520 | val raw_T = Syntax.parse_typ ctxt' s1; | |
| 521 | val S = Syntax.read_sort ctxt' s2; | |
| 522 | val T = | |
| 523 | Syntax.check_term ctxt' | |
| 524 | (Logic.mk_type raw_T |> Type.constraint (Term.itselfT (Type_Infer.anyT S))) | |
| 525 | |> Logic.dest_type; | |
| 526 | in Pretty.string_of (Pretty.quote (Syntax.pretty_typ ctxt' T)) end; | |
| 9128 | 527 | |
| 23935 | 528 | 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 | 529 | 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 | 530 | |
| 
c0f2f8224ea8
print_term etc.: actually observe print mode in final output;
 wenzelm parents: 
19268diff
changeset | 531 | in | 
| 9128 | 532 | |
| 38331 
6e72f31999ac
prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
 wenzelm parents: 
38139diff
changeset | 533 | 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 | 534 | 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 | 535 | 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 | 536 | 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 | 537 | 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 | 538 | val print_type = print_item (string_of_type o Toplevel.context_of); | 
| 5831 | 539 | |
| 19385 
c0f2f8224ea8
print_term etc.: actually observe print mode in final output;
 wenzelm parents: 
19268diff
changeset | 540 | end; | 
| 
c0f2f8224ea8
print_term etc.: actually observe print mode in final output;
 wenzelm parents: 
19268diff
changeset | 541 | |
| 12938 
a646d0467d81
markup commands (from isar_thy.ML) with proper check of antiquotations;
 wenzelm parents: 
12876diff
changeset | 542 | |
| 
a646d0467d81
markup commands (from isar_thy.ML) with proper check of antiquotations;
 wenzelm parents: 
12876diff
changeset | 543 | (* markup commands *) | 
| 
a646d0467d81
markup commands (from isar_thy.ML) with proper check of antiquotations;
 wenzelm parents: 
12876diff
changeset | 544 | |
| 27853 
916038f77be6
simplified markup commands -- removed obsolete Present.results, always check text;
 wenzelm parents: 
27730diff
changeset | 545 | fun header_markup txt = Toplevel.keep (fn state => | 
| 48918 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 wenzelm parents: 
48881diff
changeset | 546 | if Toplevel.is_toplevel state then Thy_Output.check_text txt state | 
| 27853 
916038f77be6
simplified markup commands -- removed obsolete Present.results, always check text;
 wenzelm parents: 
27730diff
changeset | 547 | else raise Toplevel.UNDEF); | 
| 12953 
7d5bd53555d8
markup commands: proper theory/proof transactions!
 wenzelm parents: 
12938diff
changeset | 548 | |
| 48918 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 wenzelm parents: 
48881diff
changeset | 549 | fun local_theory_markup (loc, txt) = Toplevel.present_local_theory loc (Thy_Output.check_text txt); | 
| 
6e5fd4585512
check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
 wenzelm parents: 
48881diff
changeset | 550 | val proof_markup = Toplevel.present_proof o Thy_Output.check_text; | 
| 12938 
a646d0467d81
markup commands (from isar_thy.ML) with proper check of antiquotations;
 wenzelm parents: 
12876diff
changeset | 551 | |
| 5831 | 552 | end; |