src/Pure/Isar/isar_cmd.ML
author wenzelm
Sun, 02 Nov 2014 15:27:37 +0100
changeset 58868 c5e1cce7ace3
parent 58201 5bf56c758e02
child 58875 ab1c65b015c3
permissions -rw-r--r--
uniform heading commands work in any context, even in theory header; discontinued obsolete 'sect', 'subsect', 'subsubsect'; marked obsolete 'header' as legacy;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5831
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/isar_cmd.ML
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
     2
    Author:     Markus Wenzel, TU Muenchen
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
     3
30805
wenzelm
parents: 30801
diff changeset
     4
Miscellaneous Isar commands.
5831
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
     5
*)
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
     6
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
     7
signature ISAR_CMD =
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
     8
sig
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
     9
  val global_setup: Symbol_Pos.source -> theory -> theory
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    10
  val local_setup: Symbol_Pos.source -> Proof.context -> Proof.context
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    11
  val parse_ast_translation: Symbol_Pos.source -> theory -> theory
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    12
  val parse_translation: Symbol_Pos.source -> theory -> theory
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    13
  val print_translation: Symbol_Pos.source -> theory -> theory
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    14
  val typed_print_translation: Symbol_Pos.source -> theory -> theory
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    15
  val print_ast_translation: Symbol_Pos.source -> theory -> theory
42204
b3277168c1e7 added Position.reports convenience;
wenzelm
parents: 41435
diff changeset
    16
  val translations: (xstring * string) Syntax.trrule list -> theory -> theory
b3277168c1e7 added Position.reports convenience;
wenzelm
parents: 41435
diff changeset
    17
  val no_translations: (xstring * string) Syntax.trrule list -> theory -> theory
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    18
  val oracle: bstring * Position.T -> Symbol_Pos.source -> theory -> theory
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 57934
diff changeset
    19
  val add_defs: (bool * bool) * ((binding * string) * Token.src list) list -> theory -> theory
40784
177e8cea3e09 added 'syntax_declaration' command;
wenzelm
parents: 40743
diff changeset
    20
  val declaration: {syntax: bool, pervasive: bool} ->
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    21
    Symbol_Pos.source -> local_theory -> local_theory
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    22
  val simproc_setup: string * Position.T -> string list -> Symbol_Pos.source ->
42464
ae16b8abf1a8 proper binding/report of defined simprocs;
wenzelm
parents: 42425
diff changeset
    23
    string list -> local_theory -> local_theory
28084
a05ca48ef263 type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents: 28083
diff changeset
    24
  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: 28083
diff changeset
    25
  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: 28083
diff changeset
    26
  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: 28083
diff changeset
    27
  val thus: (Attrib.binding * (string * string list) list) list -> bool -> Proof.state -> Proof.state
49889
00ea087e83d8 more method position information, notably finished_pos after end of previous text;
wenzelm
parents: 49866
diff changeset
    28
  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: 49866
diff changeset
    29
  val terminal_proof: Method.text_range * Method.text_range option ->
21350
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
    30
    Toplevel.transition -> Toplevel.transition
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
    31
  val default_proof: Toplevel.transition -> Toplevel.transition
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
    32
  val immediate_proof: Toplevel.transition -> Toplevel.transition
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
    33
  val done_proof: Toplevel.transition -> Toplevel.transition
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
    34
  val skip_proof: Toplevel.transition -> Toplevel.transition
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    35
  val ml_diag: bool -> Symbol_Pos.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: 47815
diff changeset
    36
  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: 47815
diff changeset
    37
  val diag_goal: Proof.context -> {context: Proof.context, facts: thm list, goal: thm}
57605
8e0a7eaffe47 tuned messages;
wenzelm
parents: 56868
diff changeset
    38
  val pretty_theorems: bool -> Toplevel.state -> Pretty.T list
49569
7b6aaf446496 tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
wenzelm
parents: 49561
diff changeset
    39
  val thy_deps: Toplevel.transition -> Toplevel.transition
7b6aaf446496 tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
wenzelm
parents: 49561
diff changeset
    40
  val locale_deps: Toplevel.transition -> Toplevel.transition
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 57934
diff changeset
    41
  val print_stmts: string list * (Facts.ref * Token.src list) list
19268
5a575522fd26 added print_stmts;
wenzelm
parents: 19057
diff changeset
    42
    -> Toplevel.transition -> Toplevel.transition
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 57934
diff changeset
    43
  val print_thms: string list * (Facts.ref * Token.src list) list
10581
74e542a299f0 dignostic commands: comment;
wenzelm
parents: 9731
diff changeset
    44
    -> Toplevel.transition -> Toplevel.transition
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 57934
diff changeset
    45
  val print_prfs: bool -> string list * (Facts.ref * Token.src list) list option
11524
197f2e14a714 Added functions for printing primitive proof terms.
berghofe
parents: 11017
diff changeset
    46
    -> Toplevel.transition -> Toplevel.transition
12876
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12758
diff changeset
    47
  val print_prop: (string list * string) -> Toplevel.transition -> Toplevel.transition
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12758
diff changeset
    48
  val print_term: (string list * string) -> Toplevel.transition -> Toplevel.transition
48792
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48776
diff changeset
    49
  val print_type: (string list * (string * string option)) ->
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48776
diff changeset
    50
    Toplevel.transition -> Toplevel.transition
58868
c5e1cce7ace3 uniform heading commands work in any context, even in theory header;
wenzelm
parents: 58201
diff changeset
    51
  val local_theory_markup: (xstring * Position.T) option * Symbol_Pos.source ->
17262
63cf42df2723 add_chapter/section/subsection/subsubsection/text: optional locale specification;
wenzelm
parents: 17228
diff changeset
    52
    Toplevel.transition -> Toplevel.transition
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    53
  val proof_markup: Symbol_Pos.source -> Toplevel.transition -> Toplevel.transition
58868
c5e1cce7ace3 uniform heading commands work in any context, even in theory header;
wenzelm
parents: 58201
diff changeset
    54
  val heading_markup: (xstring * Position.T) option * Symbol_Pos.source ->
c5e1cce7ace3 uniform heading commands work in any context, even in theory header;
wenzelm
parents: 58201
diff changeset
    55
    Toplevel.transition -> Toplevel.transition
5831
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
    56
end;
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
    57
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37198
diff changeset
    58
structure Isar_Cmd: ISAR_CMD =
5831
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
    59
struct
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
    60
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
    61
22116
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
    62
(** theory declarations **)
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
    63
30461
00323c45ea83 added 'local_setup' command;
wenzelm
parents: 30367
diff changeset
    64
(* generic setup *)
22116
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
    65
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    66
fun global_setup source =
56278
2576d3a40ed6 separate tokenization and language context for SML: no symbols, no antiquotes;
wenzelm
parents: 56275
diff changeset
    67
  ML_Lex.read_source false source
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    68
  |> ML_Context.expression (#pos source) "val setup: theory -> theory" "Context.map_theory setup"
26435
bdce320cd426 eliminated delayed theory setup
wenzelm
parents: 26425
diff changeset
    69
  |> Context.theory_map;
22116
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
    70
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    71
fun local_setup source =
56278
2576d3a40ed6 separate tokenization and language context for SML: no symbols, no antiquotes;
wenzelm
parents: 56275
diff changeset
    72
  ML_Lex.read_source false source
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    73
  |> ML_Context.expression (#pos source) "val setup: local_theory -> local_theory" "Context.map_proof setup"
30461
00323c45ea83 added 'local_setup' command;
wenzelm
parents: 30367
diff changeset
    74
  |> Context.proof_map;
00323c45ea83 added 'local_setup' command;
wenzelm
parents: 30367
diff changeset
    75
22116
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
    76
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
    77
(* translation functions *)
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
    78
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    79
fun parse_ast_translation source =
56278
2576d3a40ed6 separate tokenization and language context for SML: no symbols, no antiquotes;
wenzelm
parents: 56275
diff changeset
    80
  ML_Lex.read_source false source
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    81
  |> ML_Context.expression (#pos source)
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51717
diff changeset
    82
    "val parse_ast_translation: (string * (Proof.context -> Ast.ast list -> Ast.ast)) list"
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51717
diff changeset
    83
    "Context.map_theory (Sign.parse_ast_translation parse_ast_translation)"
22116
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
    84
  |> Context.theory_map;
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
    85
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    86
fun parse_translation source =
56278
2576d3a40ed6 separate tokenization and language context for SML: no symbols, no antiquotes;
wenzelm
parents: 56275
diff changeset
    87
  ML_Lex.read_source false source
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    88
  |> ML_Context.expression (#pos source)
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51717
diff changeset
    89
    "val parse_translation: (string * (Proof.context -> term list -> term)) list"
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51717
diff changeset
    90
    "Context.map_theory (Sign.parse_translation parse_translation)"
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51717
diff changeset
    91
  |> Context.theory_map;
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51717
diff changeset
    92
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    93
fun print_translation source =
56278
2576d3a40ed6 separate tokenization and language context for SML: no symbols, no antiquotes;
wenzelm
parents: 56275
diff changeset
    94
  ML_Lex.read_source false source
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
    95
  |> ML_Context.expression (#pos source)
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51717
diff changeset
    96
    "val print_translation: (string * (Proof.context -> term list -> term)) list"
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51717
diff changeset
    97
    "Context.map_theory (Sign.print_translation print_translation)"
22116
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
    98
  |> Context.theory_map;
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
    99
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
   100
fun typed_print_translation source =
56278
2576d3a40ed6 separate tokenization and language context for SML: no symbols, no antiquotes;
wenzelm
parents: 56275
diff changeset
   101
  ML_Lex.read_source false source
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
   102
  |> ML_Context.expression (#pos source)
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51717
diff changeset
   103
    "val typed_print_translation: (string * (Proof.context -> typ -> term list -> term)) list"
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51717
diff changeset
   104
    "Context.map_theory (Sign.typed_print_translation typed_print_translation)"
22116
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
   105
  |> Context.theory_map;
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
   106
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
   107
fun print_ast_translation source =
56278
2576d3a40ed6 separate tokenization and language context for SML: no symbols, no antiquotes;
wenzelm
parents: 56275
diff changeset
   108
  ML_Lex.read_source false source
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
   109
  |> ML_Context.expression (#pos source)
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51717
diff changeset
   110
    "val print_ast_translation: (string * (Proof.context -> Ast.ast list -> Ast.ast)) list"
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51717
diff changeset
   111
    "Context.map_theory (Sign.print_ast_translation print_ast_translation)"
22116
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
   112
  |> Context.theory_map;
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
   113
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
   114
42204
b3277168c1e7 added Position.reports convenience;
wenzelm
parents: 41435
diff changeset
   115
(* translation rules *)
b3277168c1e7 added Position.reports convenience;
wenzelm
parents: 41435
diff changeset
   116
b3277168c1e7 added Position.reports convenience;
wenzelm
parents: 41435
diff changeset
   117
fun read_trrules thy raw_rules =
b3277168c1e7 added Position.reports convenience;
wenzelm
parents: 41435
diff changeset
   118
  let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   119
    val ctxt = Proof_Context.init_global thy;
56006
6a4dcaf53664 more formal read_root;
wenzelm
parents: 56005
diff changeset
   120
    val read_root =
6a4dcaf53664 more formal read_root;
wenzelm
parents: 56005
diff changeset
   121
      #1 o dest_Type o Proof_Context.read_type_name {proper = true, strict = false} ctxt;
42204
b3277168c1e7 added Position.reports convenience;
wenzelm
parents: 41435
diff changeset
   122
  in
56006
6a4dcaf53664 more formal read_root;
wenzelm
parents: 56005
diff changeset
   123
    raw_rules
6a4dcaf53664 more formal read_root;
wenzelm
parents: 56005
diff changeset
   124
    |> map (Syntax.map_trrule (fn (r, s) => Syntax_Phases.parse_ast_pattern ctxt (read_root r, s)))
42204
b3277168c1e7 added Position.reports convenience;
wenzelm
parents: 41435
diff changeset
   125
  end;
b3277168c1e7 added Position.reports convenience;
wenzelm
parents: 41435
diff changeset
   126
b3277168c1e7 added Position.reports convenience;
wenzelm
parents: 41435
diff changeset
   127
fun translations args thy = Sign.add_trrules (read_trrules thy args) thy;
b3277168c1e7 added Position.reports convenience;
wenzelm
parents: 41435
diff changeset
   128
fun no_translations args thy = Sign.del_trrules (read_trrules thy args) thy;
b3277168c1e7 added Position.reports convenience;
wenzelm
parents: 41435
diff changeset
   129
b3277168c1e7 added Position.reports convenience;
wenzelm
parents: 41435
diff changeset
   130
22116
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
   131
(* oracles *)
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
   132
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
   133
fun oracle (name, pos) source =
27871
4ef76f8788ad oracle, header/local_theory/proof_markup: pass SymbolPos.text;
wenzelm
parents: 27853
diff changeset
   134
  let
56278
2576d3a40ed6 separate tokenization and language context for SML: no symbols, no antiquotes;
wenzelm
parents: 56275
diff changeset
   135
    val body = ML_Lex.read_source false source;
37198
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 37146
diff changeset
   136
    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: 37146
diff changeset
   137
      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: 37146
diff changeset
   138
       ("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: 37146
diff changeset
   139
        \  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: 37146
diff changeset
   140
        \  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: 37146
diff changeset
   141
        \in\n\
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
   142
        \  val " ^ name ^
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
   143
        " = snd (Context.>>> (Context.map_theory_result (Thm.add_oracle (binding, body))));\n\
37198
3af985b10550 replaced ML_Lex.read_antiq by more concise ML_Lex.read, which includes full read/report with explicit position information;
wenzelm
parents: 37146
diff changeset
   144
        \end;\n");
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: 56278
diff changeset
   145
  in
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: 56278
diff changeset
   146
    Context.theory_map
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: 56278
diff changeset
   147
      (ML_Context.exec (fn () => ML_Context.eval ML_Compiler.flags (#pos source) ants))
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: 56278
diff changeset
   148
  end;
22116
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
   149
6917be2e647d added various ML setup functions (from sign.ML, pure_thy.ML);
wenzelm
parents: 22087
diff changeset
   150
50737
f310d1735d93 tuned -- less indirection;
wenzelm
parents: 49889
diff changeset
   151
(* old-style defs *)
21350
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   152
35852
4e3fe0b8687b minor renovation of old-style 'axioms' -- make it an alias of iterated 'axiomatization';
wenzelm
parents: 35141
diff changeset
   153
fun add_defs ((unchecked, overloaded), args) thy =
57683
cc0aa6528890 old 'defs' is legacy --- slightly odd side-entry that bypasses regular Local_Theory.define interface;
wenzelm
parents: 57605
diff changeset
   154
 (legacy_feature "Old 'defs' command -- use 'definition' (with 'overloading') instead";
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39507
diff changeset
   155
  thy |>
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39507
diff changeset
   156
    (if unchecked then Global_Theory.add_defs_unchecked_cmd else Global_Theory.add_defs_cmd)
47815
43f677b3ae91 clarified signature;
wenzelm
parents: 46922
diff changeset
   157
      overloaded
43f677b3ae91 clarified signature;
wenzelm
parents: 46922
diff changeset
   158
      (map (fn ((b, ax), srcs) => ((b, ax), map (Attrib.attribute_cmd_global thy) srcs)) args)
57683
cc0aa6528890 old 'defs' is legacy --- slightly odd side-entry that bypasses regular Local_Theory.define interface;
wenzelm
parents: 57605
diff changeset
   159
  |> snd);
21350
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   160
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   161
22087
a13299166175 added 'declaration' command;
wenzelm
parents: 21955
diff changeset
   162
(* declarations *)
a13299166175 added 'declaration' command;
wenzelm
parents: 21955
diff changeset
   163
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
   164
fun declaration {syntax, pervasive} source =
56278
2576d3a40ed6 separate tokenization and language context for SML: no symbols, no antiquotes;
wenzelm
parents: 56275
diff changeset
   165
  ML_Lex.read_source false source
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
   166
  |> ML_Context.expression (#pos source)
26455
1757a6e049f4 reorganized signature of ML_Context;
wenzelm
parents: 26435
diff changeset
   167
    "val declaration: Morphism.declaration"
45291
57cd50f98fdc uniform Local_Theory.declaration with explicit params;
wenzelm
parents: 44338
diff changeset
   168
    ("Context.map_proof (Local_Theory.declaration {syntax = " ^ Bool.toString syntax ^ ", \
57cd50f98fdc uniform Local_Theory.declaration with explicit params;
wenzelm
parents: 44338
diff changeset
   169
      \pervasive = " ^ Bool.toString pervasive ^ "} declaration)")
26385
ae7564661e76 ML runtime compilation: pass position, tuned signature;
wenzelm
parents: 26336
diff changeset
   170
  |> Context.proof_map;
22087
a13299166175 added 'declaration' command;
wenzelm
parents: 21955
diff changeset
   171
a13299166175 added 'declaration' command;
wenzelm
parents: 21955
diff changeset
   172
22202
0544af1a5117 added simproc_setup;
wenzelm
parents: 22116
diff changeset
   173
(* simprocs *)
0544af1a5117 added simproc_setup;
wenzelm
parents: 22116
diff changeset
   174
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
   175
fun simproc_setup name lhss source identifier =
56278
2576d3a40ed6 separate tokenization and language context for SML: no symbols, no antiquotes;
wenzelm
parents: 56275
diff changeset
   176
  ML_Lex.read_source false source
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
   177
  |> ML_Context.expression (#pos source)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51583
diff changeset
   178
    "val proc: Morphism.morphism -> Proof.context -> cterm -> thm option"
42464
ae16b8abf1a8 proper binding/report of defined simprocs;
wenzelm
parents: 42425
diff changeset
   179
    ("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: 37146
diff changeset
   180
      \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: 37146
diff changeset
   181
      \identifier = Library.maps ML_Context.thms " ^ ML_Syntax.print_strings identifier ^ "})")
22202
0544af1a5117 added simproc_setup;
wenzelm
parents: 22116
diff changeset
   182
  |> Context.proof_map;
0544af1a5117 added simproc_setup;
wenzelm
parents: 22116
diff changeset
   183
0544af1a5117 added simproc_setup;
wenzelm
parents: 22116
diff changeset
   184
21350
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   185
(* goals *)
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   186
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   187
fun goal opt_chain goal stmt int =
29383
223f18cfbb32 qed/after_qed: singleton result;
wenzelm
parents: 29360
diff changeset
   188
  opt_chain #> goal NONE (K I) stmt int;
21350
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   189
36323
655e2d74de3a modernized naming conventions of main Isar proof elements;
wenzelm
parents: 36178
diff changeset
   190
val have = goal I Proof.have_cmd;
655e2d74de3a modernized naming conventions of main Isar proof elements;
wenzelm
parents: 36178
diff changeset
   191
val hence = goal Proof.chain Proof.have_cmd;
655e2d74de3a modernized naming conventions of main Isar proof elements;
wenzelm
parents: 36178
diff changeset
   192
val show = goal I Proof.show_cmd;
655e2d74de3a modernized naming conventions of main Isar proof elements;
wenzelm
parents: 36178
diff changeset
   193
val thus = goal Proof.chain Proof.show_cmd;
21350
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   194
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   195
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   196
(* local endings *)
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   197
29383
223f18cfbb32 qed/after_qed: singleton result;
wenzelm
parents: 29360
diff changeset
   198
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: 31819
diff changeset
   199
val local_terminal_proof = Toplevel.proof' o Proof.local_future_terminal_proof;
29383
223f18cfbb32 qed/after_qed: singleton result;
wenzelm
parents: 29360
diff changeset
   200
val local_default_proof = Toplevel.proof Proof.local_default_proof;
223f18cfbb32 qed/after_qed: singleton result;
wenzelm
parents: 29360
diff changeset
   201
val local_immediate_proof = Toplevel.proof Proof.local_immediate_proof;
223f18cfbb32 qed/after_qed: singleton result;
wenzelm
parents: 29360
diff changeset
   202
val local_done_proof = Toplevel.proof Proof.local_done_proof;
223f18cfbb32 qed/after_qed: singleton result;
wenzelm
parents: 29360
diff changeset
   203
val local_skip_proof = Toplevel.proof' Proof.local_skip_proof;
21350
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   204
27562
bcb01eb565ee removed obsolete history commands;
wenzelm
parents: 27493
diff changeset
   205
val skip_local_qed = Toplevel.skip_proof (fn i => if i > 1 then i - 1 else raise Toplevel.UNDEF);
21350
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   206
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   207
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   208
(* global endings *)
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   209
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   210
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: 48918
diff changeset
   211
val global_terminal_proof = Toplevel.end_proof o Proof.global_future_terminal_proof;
21350
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   212
val global_default_proof = Toplevel.end_proof (K Proof.global_default_proof);
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   213
val global_immediate_proof = Toplevel.end_proof (K Proof.global_immediate_proof);
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   214
val global_skip_proof = Toplevel.end_proof Proof.global_skip_proof;
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   215
val global_done_proof = Toplevel.end_proof (K Proof.global_done_proof);
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   216
28375
c879d88d038a eliminated polymorphic equality;
wenzelm
parents: 28290
diff changeset
   217
val skip_global_qed = Toplevel.skip_proof_to_theory (fn n => n = 1);
21350
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   218
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   219
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   220
(* common endings *)
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   221
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   222
fun qed m = local_qed m o global_qed m o skip_local_qed o skip_global_qed;
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   223
fun terminal_proof m = local_terminal_proof m o global_terminal_proof m;
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   224
val default_proof = local_default_proof o global_default_proof;
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   225
val immediate_proof = local_immediate_proof o global_immediate_proof;
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   226
val done_proof = local_done_proof o global_done_proof;
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   227
val skip_proof = local_skip_proof o global_skip_proof;
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   228
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21003
diff changeset
   229
26489
e83dc4bb9ab4 removed obsolete use_XXX;
wenzelm
parents: 26463
diff changeset
   230
(* diagnostic ML evaluation *)
5831
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
   231
37305
9763792e4ac7 diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
wenzelm
parents: 37216
diff changeset
   232
structure Diag_State = Proof_Data
9763792e4ac7 diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
wenzelm
parents: 37216
diff changeset
   233
(
9763792e4ac7 diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
wenzelm
parents: 37216
diff changeset
   234
  type T = Toplevel.state;
9763792e4ac7 diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
wenzelm
parents: 37216
diff changeset
   235
  fun init _ = Toplevel.toplevel;
9763792e4ac7 diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
wenzelm
parents: 37216
diff changeset
   236
);
9763792e4ac7 diagnostic commands 'ML_val' and 'ML_command' may refer to antiquotations @{Isar.state} and @{Isar.goal};
wenzelm
parents: 37216
diff changeset
   237
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 53171
diff changeset
   238
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: 56278
diff changeset
   239
  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: 56278
diff changeset
   240
    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: 56278
diff changeset
   241
      try Toplevel.generic_theory_of state
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: 56278
diff changeset
   242
      |> Option.map (Context.proof_of #> Diag_State.put state);
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: 56278
diff changeset
   243
    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: 56278
diff changeset
   244
  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: 37216
diff changeset
   245
48776
37cd53e69840 faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents: 47815
diff changeset
   246
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: 37216
diff changeset
   247
48776
37cd53e69840 faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents: 47815
diff changeset
   248
fun diag_goal ctxt =
37cd53e69840 faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
wenzelm
parents: 47815
diff changeset
   249
  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: 37216
diff changeset
   250
    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: 37216
diff changeset
   251
53171
a5e54d4d9081 added Theory.setup convenience;
wenzelm
parents: 52549
diff changeset
   252
val _ = Theory.setup
56204
f70e69208a8c more antiquotations;
wenzelm
parents: 56158
diff changeset
   253
  (ML_Antiquotation.value (Binding.qualify true "Isar" @{binding state})
53171
a5e54d4d9081 added Theory.setup convenience;
wenzelm
parents: 52549
diff changeset
   254
    (Scan.succeed "Isar_Cmd.diag_state ML_context") #>
56204
f70e69208a8c more antiquotations;
wenzelm
parents: 56158
diff changeset
   255
   ML_Antiquotation.value (Binding.qualify true "Isar" @{binding goal})
53171
a5e54d4d9081 added Theory.setup convenience;
wenzelm
parents: 52549
diff changeset
   256
    (Scan.succeed "Isar_Cmd.diag_goal ML_context"));
5831
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
   257
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
   258
56868
b5fb264d53ba clarified print operations for "terms" and "theorems";
wenzelm
parents: 56334
diff changeset
   259
(* theorems of theory or proof context *)
17066
b53f050bc37d back: removed ill-defined '!' option;
wenzelm
parents: 16812
diff changeset
   260
56868
b5fb264d53ba clarified print operations for "terms" and "theorems";
wenzelm
parents: 56334
diff changeset
   261
fun pretty_theorems verbose st =
b5fb264d53ba clarified print operations for "terms" and "theorems";
wenzelm
parents: 56334
diff changeset
   262
  if Toplevel.is_proof st then
57605
8e0a7eaffe47 tuned messages;
wenzelm
parents: 56868
diff changeset
   263
    Proof_Context.pretty_local_facts (Toplevel.context_of st) verbose
56868
b5fb264d53ba clarified print operations for "terms" and "theorems";
wenzelm
parents: 56334
diff changeset
   264
  else
b5fb264d53ba clarified print operations for "terms" and "theorems";
wenzelm
parents: 56334
diff changeset
   265
    let
b5fb264d53ba clarified print operations for "terms" and "theorems";
wenzelm
parents: 56334
diff changeset
   266
      val thy = Toplevel.theory_of st;
b5fb264d53ba clarified print operations for "terms" and "theorems";
wenzelm
parents: 56334
diff changeset
   267
      val prev_thys =
b5fb264d53ba clarified print operations for "terms" and "theorems";
wenzelm
parents: 56334
diff changeset
   268
        (case Toplevel.previous_context_of st of
b5fb264d53ba clarified print operations for "terms" and "theorems";
wenzelm
parents: 56334
diff changeset
   269
          SOME prev => [Proof_Context.theory_of prev]
b5fb264d53ba clarified print operations for "terms" and "theorems";
wenzelm
parents: 56334
diff changeset
   270
        | NONE => Theory.parents_of thy);
b5fb264d53ba clarified print operations for "terms" and "theorems";
wenzelm
parents: 56334
diff changeset
   271
    in Proof_Display.pretty_theorems_diff verbose prev_thys thy end;
9513
8531c18d9181 unknown_theory/proof/context;
wenzelm
parents: 9454
diff changeset
   272
12060
f85eddf6a4fb added print_locales, print_locale;
wenzelm
parents: 12055
diff changeset
   273
50737
f310d1735d93 tuned -- less indirection;
wenzelm
parents: 49889
diff changeset
   274
(* display dependencies *)
5831
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
   275
22485
3a7d623485fa added theory dependency graph
haftmann
parents: 22340
diff changeset
   276
val thy_deps = Toplevel.unknown_theory o Toplevel.keep (fn state =>
3a7d623485fa added theory dependency graph
haftmann
parents: 22340
diff changeset
   277
  let
3a7d623485fa added theory dependency graph
haftmann
parents: 22340
diff changeset
   278
    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: 37305
diff changeset
   279
    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: 37305
diff changeset
   280
42425
2aa907d5ee4f added Theory.nodes_of convenience;
wenzelm
parents: 42360
diff changeset
   281
    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: 22573
diff changeset
   282
      let
a165d9ed08b8 simplified thy_deps using Theory.ancestors_of (in order of creation);
wenzelm
parents: 22573
diff changeset
   283
        val name = Context.theory_name node;
a165d9ed08b8 simplified thy_deps using Theory.ancestors_of (in order of creation);
wenzelm
parents: 22573
diff changeset
   284
        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: 37305
diff changeset
   285
        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: 37305
diff changeset
   286
        val unfold = (session = thy_session);
49561
26fc70e983c2 separate module Graph_Display;
wenzelm
parents: 49012
diff changeset
   287
      in
26fc70e983c2 separate module Graph_Display;
wenzelm
parents: 49012
diff changeset
   288
       {name = name, ID = name, parents = parents, dir = session,
26fc70e983c2 separate module Graph_Display;
wenzelm
parents: 49012
diff changeset
   289
        unfold = unfold, path = "", content = []}
26fc70e983c2 separate module Graph_Display;
wenzelm
parents: 49012
diff changeset
   290
      end);
26fc70e983c2 separate module Graph_Display;
wenzelm
parents: 49012
diff changeset
   291
  in Graph_Display.display_graph gr end);
22485
3a7d623485fa added theory dependency graph
haftmann
parents: 22340
diff changeset
   292
49569
7b6aaf446496 tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
wenzelm
parents: 49561
diff changeset
   293
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: 49561
diff changeset
   294
  let
7b6aaf446496 tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
wenzelm
parents: 49561
diff changeset
   295
    val thy = Toplevel.theory_of state;
7b6aaf446496 tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
wenzelm
parents: 49561
diff changeset
   296
    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: 49561
diff changeset
   297
     {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: 49561
diff changeset
   298
      dir = "", unfold = true, path = "", content = [body]});
7b6aaf446496 tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
wenzelm
parents: 49561
diff changeset
   299
  in Graph_Display.display_graph gr end);
7b6aaf446496 tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
wenzelm
parents: 49561
diff changeset
   300
26184
64ee6a2ca6d6 Added unused_thms command.
berghofe
parents: 26070
diff changeset
   301
19268
5a575522fd26 added print_stmts;
wenzelm
parents: 19057
diff changeset
   302
(* print theorems, terms, types etc. *)
5a575522fd26 added print_stmts;
wenzelm
parents: 19057
diff changeset
   303
19385
c0f2f8224ea8 print_term etc.: actually observe print mode in final output;
wenzelm
parents: 19268
diff changeset
   304
local
c0f2f8224ea8 print_term etc.: actually observe print mode in final output;
wenzelm
parents: 19268
diff changeset
   305
38331
6e72f31999ac prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
wenzelm
parents: 38139
diff changeset
   306
fun string_of_stmts ctxt args =
6e72f31999ac prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
wenzelm
parents: 38139
diff changeset
   307
  Attrib.eval_thms ctxt args
6e72f31999ac prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
wenzelm
parents: 38139
diff changeset
   308
  |> map (Element.pretty_statement ctxt Thm.theoremK)
19385
c0f2f8224ea8 print_term etc.: actually observe print mode in final output;
wenzelm
parents: 19268
diff changeset
   309
  |> Pretty.chunks2 |> Pretty.string_of;
5880
feec44106a8e add print_theorems;
wenzelm
parents: 5831
diff changeset
   310
38331
6e72f31999ac prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
wenzelm
parents: 38139
diff changeset
   311
fun string_of_thms ctxt args =
51583
9100c8e66b69 item markup for Proof_Context.pretty_fact;
wenzelm
parents: 50737
diff changeset
   312
  Pretty.string_of (Proof_Context.pretty_fact ctxt ("", Attrib.eval_thms ctxt args));
5895
457b42674b57 added print_thm;
wenzelm
parents: 5880
diff changeset
   313
19385
c0f2f8224ea8 print_term etc.: actually observe print mode in final output;
wenzelm
parents: 19268
diff changeset
   314
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: 32804
diff changeset
   315
  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: 32804
diff changeset
   316
    (case arg of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
   317
      NONE =>
12125
316d11f760f7 Commands prf and full_prf can now also be used to display proof term
berghofe
parents: 12069
diff changeset
   318
        let
38331
6e72f31999ac prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
wenzelm
parents: 38139
diff changeset
   319
          val {context = ctxt, goal = thm} = Proof.simple_goal (Toplevel.proof_of state);
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   320
          val thy = Proof_Context.theory_of ctxt;
28814
463c9e9111ae clarified Thm.proof_body_of vs. Thm.proof_of;
wenzelm
parents: 28799
diff changeset
   321
          val prf = Thm.proof_of thm;
17066
b53f050bc37d back: removed ill-defined '!' option;
wenzelm
parents: 16812
diff changeset
   322
          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: 32804
diff changeset
   323
          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: 12069
diff changeset
   324
        in
33388
d64545e6cba5 modernized structure Proof_Syntax;
wenzelm
parents: 33369
diff changeset
   325
          Proof_Syntax.pretty_proof ctxt
17066
b53f050bc37d back: removed ill-defined '!' option;
wenzelm
parents: 16812
diff changeset
   326
            (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: 12069
diff changeset
   327
        end
38331
6e72f31999ac prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
wenzelm
parents: 38139
diff changeset
   328
    | SOME srcs =>
6e72f31999ac prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
wenzelm
parents: 38139
diff changeset
   329
        let val ctxt = Toplevel.context_of state
6e72f31999ac prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
wenzelm
parents: 38139
diff changeset
   330
        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: 38139
diff changeset
   331
        |> Pretty.chunks);
11524
197f2e14a714 Added functions for printing primitive proof terms.
berghofe
parents: 11017
diff changeset
   332
38331
6e72f31999ac prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
wenzelm
parents: 38139
diff changeset
   333
fun string_of_prop ctxt s =
5831
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
   334
  let
24508
c8b82fec6447 replaced ProofContext.read_term/prop by general Syntax.read_term/prop;
wenzelm
parents: 24314
diff changeset
   335
    val prop = Syntax.read_prop ctxt s;
26704
51ee753cc2e3 token translations: context dependent, result Pretty.T;
wenzelm
parents: 26694
diff changeset
   336
    val ctxt' = Variable.auto_fixes prop ctxt;
51ee753cc2e3 token translations: context dependent, result Pretty.T;
wenzelm
parents: 26694
diff changeset
   337
  in Pretty.string_of (Pretty.quote (Syntax.pretty_term ctxt' prop)) end;
5831
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
   338
38331
6e72f31999ac prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
wenzelm
parents: 38139
diff changeset
   339
fun string_of_term ctxt s =
5831
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
   340
  let
24508
c8b82fec6447 replaced ProofContext.read_term/prop by general Syntax.read_term/prop;
wenzelm
parents: 24314
diff changeset
   341
    val t = Syntax.read_term ctxt s;
5831
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
   342
    val T = Term.type_of t;
26704
51ee753cc2e3 token translations: context dependent, result Pretty.T;
wenzelm
parents: 26694
diff changeset
   343
    val ctxt' = Variable.auto_fixes t ctxt;
5831
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
   344
  in
19385
c0f2f8224ea8 print_term etc.: actually observe print mode in final output;
wenzelm
parents: 19268
diff changeset
   345
    Pretty.string_of
26704
51ee753cc2e3 token translations: context dependent, result Pretty.T;
wenzelm
parents: 26694
diff changeset
   346
      (Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t), Pretty.fbrk,
51ee753cc2e3 token translations: context dependent, result Pretty.T;
wenzelm
parents: 26694
diff changeset
   347
        Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' T)])
9128
35abf6308ab0 rearranged print commands;
wenzelm
parents: 9009
diff changeset
   348
  end;
5831
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
   349
48792
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48776
diff changeset
   350
fun string_of_type ctxt (s, NONE) =
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48776
diff changeset
   351
      let val T = Syntax.read_typ ctxt s
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48776
diff changeset
   352
      in Pretty.string_of (Pretty.quote (Syntax.pretty_typ ctxt T)) end
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48776
diff changeset
   353
  | string_of_type ctxt (s1, SOME s2) =
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48776
diff changeset
   354
      let
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48776
diff changeset
   355
        val ctxt' = Config.put show_sorts true ctxt;
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48776
diff changeset
   356
        val raw_T = Syntax.parse_typ ctxt' s1;
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48776
diff changeset
   357
        val S = Syntax.read_sort ctxt' s2;
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48776
diff changeset
   358
        val T =
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48776
diff changeset
   359
          Syntax.check_term ctxt'
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48776
diff changeset
   360
            (Logic.mk_type raw_T |> Type.constraint (Term.itselfT (Type_Infer.anyT S)))
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48776
diff changeset
   361
          |> Logic.dest_type;
4aa5b965f70e support for 'typ' with explicit sort constraint;
wenzelm
parents: 48776
diff changeset
   362
      in Pretty.string_of (Pretty.quote (Syntax.pretty_typ ctxt' T)) end;
9128
35abf6308ab0 rearranged print commands;
wenzelm
parents: 9009
diff changeset
   363
23935
2a4e42ec9a54 PrintMode.with_modes;
wenzelm
parents: 23897
diff changeset
   364
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: 38139
diff changeset
   365
  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: 19268
diff changeset
   366
c0f2f8224ea8 print_term etc.: actually observe print mode in final output;
wenzelm
parents: 19268
diff changeset
   367
in
9128
35abf6308ab0 rearranged print commands;
wenzelm
parents: 9009
diff changeset
   368
38331
6e72f31999ac prefer plain Attrib.eval_thms with plain Proof.context instead of Proof.state;
wenzelm
parents: 38139
diff changeset
   369
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: 38139
diff changeset
   370
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: 19268
diff changeset
   371
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: 38139
diff changeset
   372
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: 38139
diff changeset
   373
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: 38139
diff changeset
   374
val print_type = print_item (string_of_type o Toplevel.context_of);
5831
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
   375
19385
c0f2f8224ea8 print_term etc.: actually observe print mode in final output;
wenzelm
parents: 19268
diff changeset
   376
end;
c0f2f8224ea8 print_term etc.: actually observe print mode in final output;
wenzelm
parents: 19268
diff changeset
   377
12938
a646d0467d81 markup commands (from isar_thy.ML) with proper check of antiquotations;
wenzelm
parents: 12876
diff changeset
   378
a646d0467d81 markup commands (from isar_thy.ML) with proper check of antiquotations;
wenzelm
parents: 12876
diff changeset
   379
(* markup commands *)
a646d0467d81 markup commands (from isar_thy.ML) with proper check of antiquotations;
wenzelm
parents: 12876
diff changeset
   380
48918
6e5fd4585512 check side-comments of command spans (normally filtered out in Outer_Syntax.toplevel_source);
wenzelm
parents: 48881
diff changeset
   381
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: 48881
diff changeset
   382
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: 12876
diff changeset
   383
58868
c5e1cce7ace3 uniform heading commands work in any context, even in theory header;
wenzelm
parents: 58201
diff changeset
   384
fun reject_target NONE = ()
c5e1cce7ace3 uniform heading commands work in any context, even in theory header;
wenzelm
parents: 58201
diff changeset
   385
  | reject_target (SOME (_, pos)) =
c5e1cce7ace3 uniform heading commands work in any context, even in theory header;
wenzelm
parents: 58201
diff changeset
   386
      error ("Illegal target specification -- not a theory context" ^ Position.here pos);
c5e1cce7ace3 uniform heading commands work in any context, even in theory header;
wenzelm
parents: 58201
diff changeset
   387
c5e1cce7ace3 uniform heading commands work in any context, even in theory header;
wenzelm
parents: 58201
diff changeset
   388
fun heading_markup (loc, txt) =
c5e1cce7ace3 uniform heading commands work in any context, even in theory header;
wenzelm
parents: 58201
diff changeset
   389
  Toplevel.keep (fn state =>
c5e1cce7ace3 uniform heading commands work in any context, even in theory header;
wenzelm
parents: 58201
diff changeset
   390
    if Toplevel.is_toplevel state then
c5e1cce7ace3 uniform heading commands work in any context, even in theory header;
wenzelm
parents: 58201
diff changeset
   391
     (legacy_feature "Obsolete 'header' command -- use 'chapter', 'section' etc. instead";
c5e1cce7ace3 uniform heading commands work in any context, even in theory header;
wenzelm
parents: 58201
diff changeset
   392
      reject_target loc;
c5e1cce7ace3 uniform heading commands work in any context, even in theory header;
wenzelm
parents: 58201
diff changeset
   393
      Thy_Output.check_text txt state)
c5e1cce7ace3 uniform heading commands work in any context, even in theory header;
wenzelm
parents: 58201
diff changeset
   394
    else raise Toplevel.UNDEF) o
c5e1cce7ace3 uniform heading commands work in any context, even in theory header;
wenzelm
parents: 58201
diff changeset
   395
  local_theory_markup (loc, txt) o
c5e1cce7ace3 uniform heading commands work in any context, even in theory header;
wenzelm
parents: 58201
diff changeset
   396
  Toplevel.present_proof (fn state => (reject_target loc; Thy_Output.check_text txt state));
c5e1cce7ace3 uniform heading commands work in any context, even in theory header;
wenzelm
parents: 58201
diff changeset
   397
5831
996361157cfb Non-logical toplevel commands.
wenzelm
parents:
diff changeset
   398
end;