src/HOL/ex/Cartouche_Examples.thy
author haftmann
Thu, 19 Jun 2025 17:15:40 +0200
changeset 82734 89347c0cc6a3
parent 81558 b57996a0688c
permissions -rw-r--r--
treat map_filter similar to list_all, list_ex, list_ex1
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
59089
da2fef2faa83 tuned header;
wenzelm
parents: 59068
diff changeset
     1
(*  Title:      HOL/ex/Cartouche_Examples.thy
da2fef2faa83 tuned header;
wenzelm
parents: 59068
diff changeset
     2
    Author:     Makarius
da2fef2faa83 tuned header;
wenzelm
parents: 59068
diff changeset
     3
*)
da2fef2faa83 tuned header;
wenzelm
parents: 59068
diff changeset
     4
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
     5
section \<open>Some examples with text cartouches\<close>
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
     6
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
     7
theory Cartouche_Examples
74837
f0e0fc82b0b9 removed pointless 'text_cartouche' command: regular 'text' already supports cartouches;
wenzelm
parents: 74834
diff changeset
     8
  imports Main
f0e0fc82b0b9 removed pointless 'text_cartouche' command: regular 'text' already supports cartouches;
wenzelm
parents: 74834
diff changeset
     9
  keywords "cartouche" :: diag
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    10
begin
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    11
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
    12
subsection \<open>Regular outer syntax\<close>
56499
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    13
63054
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 62910
diff changeset
    14
text \<open>Text cartouches may be used in the outer syntax category \<open>text\<close>,
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 62910
diff changeset
    15
  as alternative to the traditional \<open>verbatim\<close> tokens.  An example is
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61457
diff changeset
    16
  this text block.\<close>  \<comment> \<open>The same works for small side-comments.\<close>
56499
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    17
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    18
notepad
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    19
begin
63120
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
    20
  txt \<open>Cartouches work as additional syntax for embedded language tokens
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
    21
    (types, terms, props) and as a replacement for the \<open>altstring\<close> category
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
    22
    (for literal fact references). For example:\<close>
56499
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    23
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    24
  fix x y :: 'a
63120
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
    25
  assume \<open>x = y\<close>
56499
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    26
  note \<open>x = y\<close>
63120
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
    27
  have \<open>x = y\<close> by (rule \<open>x = y\<close>)
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
    28
  from \<open>x = y\<close> have \<open>x = y\<close> .
56499
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    29
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    30
  txt \<open>Of course, this can be nested inside formal comments and
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    31
    antiquotations, e.g. like this @{thm \<open>x = y\<close>} or this @{thm sym
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    32
    [OF \<open>x = y\<close>]}.\<close>
56500
90f17a04567d proper Args.name vs. Args.text as documented (in contrast to adhoc union in 75aaee32893d, which had to cope with more limited Args.T);
wenzelm
parents: 56499
diff changeset
    33
63120
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
    34
  have \<open>x = y\<close>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
    35
    by (tactic \<open>resolve_tac \<^context> @{thms \<open>x = y\<close>} 1\<close>)
63120
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
    36
      \<comment> \<open>more cartouches involving ML\<close>
56499
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    37
end
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    38
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    39
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
    40
subsection \<open>Outer syntax: cartouche within command syntax\<close>
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    41
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
    42
ML \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
    43
  Outer_Syntax.command \<^command_keyword>\<open>cartouche\<close> ""
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    44
    (Parse.cartouche >> (fn s =>
60189
0d3a62127057 tuned signature;
wenzelm
parents: 59936
diff changeset
    45
      Toplevel.keep (fn _ => writeln s)))
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
    46
\<close>
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    47
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    48
cartouche \<open>abc\<close>
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    49
cartouche \<open>abc \<open>\<alpha>\<beta>\<gamma>\<close> xzy\<close>
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    50
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    51
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
    52
subsection \<open>Inner syntax: string literals via cartouche\<close>
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    53
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
    54
ML \<open>
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    55
  local
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    56
    fun mk_char (s, pos) =
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    57
      let
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    58
        val c =
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    59
          if Symbol.is_ascii s then ord s
55036
87797f8f3152 proper \<newline>;
wenzelm
parents: 55033
diff changeset
    60
          else if s = "\<newline>" then 10
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    61
          else error ("String literal contains illegal symbol: " ^ quote s ^ Position.here pos);
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
    62
      in list_comb (Syntax.const \<^const_syntax>\<open>Char\<close>, String_Syntax.mk_bits_syntax 8 c) end;
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    63
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
    64
    fun mk_string [] = Const (\<^const_syntax>\<open>Nil\<close>, \<^typ>\<open>string\<close>)
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    65
      | mk_string (s :: ss) =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
    66
          Syntax.const \<^const_syntax>\<open>Cons\<close> $ mk_char s $ mk_string ss;
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    67
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    68
  in
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    69
    fun string_tr content args =
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    70
      let fun err () = raise TERM ("string_tr", args) in
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    71
        (case args of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
    72
          [(c as Const (\<^syntax_const>\<open>_constrain\<close>, _)) $ Free (s, _) $ p] =>
81218
94bace5078ba support multiple positions (non-empty list);
wenzelm
parents: 80914
diff changeset
    73
            (case Term_Position.decode_position1 p of
81558
b57996a0688c clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents: 81218
diff changeset
    74
              SOME {pos, ...} => c $ mk_string (content (s, pos)) $ p
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    75
            | NONE => err ())
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    76
        | _ => err ())
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    77
      end;
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    78
  end;
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
    79
\<close>
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    80
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 74837
diff changeset
    81
syntax "_cartouche_string" :: \<open>cartouche_position \<Rightarrow> string\<close>  (\<open>_\<close>)
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    82
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
    83
parse_translation \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
    84
  [(\<^syntax_const>\<open>_cartouche_string\<close>,
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    85
    K (string_tr (Symbol_Pos.cartouche_content o Symbol_Pos.explode)))]
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
    86
\<close>
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    87
63120
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
    88
term \<open>\<open>\<close>\<close>
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
    89
term \<open>\<open>abc\<close>\<close>
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
    90
term \<open>\<open>abc\<close> @ \<open>xyz\<close>\<close>
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
    91
term \<open>\<open>\<newline>\<close>\<close>
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    92
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    93
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
    94
subsection \<open>Alternate outer and inner syntax: string literals\<close>
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    95
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
    96
subsubsection \<open>Nested quotes\<close>
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    97
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 74837
diff changeset
    98
syntax "_string_string" :: \<open>string_position \<Rightarrow> string\<close>  (\<open>_\<close>)
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    99
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
   100
parse_translation \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   101
  [(\<^syntax_const>\<open>_string_string\<close>, K (string_tr Lexicon.explode_string))]
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
   102
\<close>
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   103
63120
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
   104
term \<open>""\<close>
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
   105
term \<open>"abc"\<close>
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
   106
term \<open>"abc" @ "xyz"\<close>
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
   107
term \<open>"\<newline>"\<close>
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
   108
term \<open>"\001\010\100"\<close>
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
   109
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   110
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
   111
subsubsection \<open>Further nesting: antiquotations\<close>
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   112
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
   113
ML \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   114
  \<^term>\<open>""\<close>;
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   115
  \<^term>\<open>"abc"\<close>;
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   116
  \<^term>\<open>"abc" @ "xyz"\<close>;
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   117
  \<^term>\<open>"\<newline>"\<close>;
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   118
  \<^term>\<open>"\001\010\100"\<close>;
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
   119
\<close>
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   120
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
   121
text \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   122
  \<^ML>\<open>
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   123
      (
69604
d80b2df54d31 isabelle update -u control_cartouches;
wenzelm
parents: 69597
diff changeset
   124
        \<^term>\<open>""\<close>;
d80b2df54d31 isabelle update -u control_cartouches;
wenzelm
parents: 69597
diff changeset
   125
        \<^term>\<open>"abc"\<close>;
d80b2df54d31 isabelle update -u control_cartouches;
wenzelm
parents: 69597
diff changeset
   126
        \<^term>\<open>"abc" @ "xyz"\<close>;
d80b2df54d31 isabelle update -u control_cartouches;
wenzelm
parents: 69597
diff changeset
   127
        \<^term>\<open>"\<newline>"\<close>;
d80b2df54d31 isabelle update -u control_cartouches;
wenzelm
parents: 69597
diff changeset
   128
        \<^term>\<open>"\001\010\100"\<close>
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   129
      )
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   130
    \<close>
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
   131
\<close>
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   132
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   133
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
   134
subsubsection \<open>Uniform nesting of sub-languages: document source, ML, term, string literals\<close>
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   135
74837
f0e0fc82b0b9 removed pointless 'text_cartouche' command: regular 'text' already supports cartouches;
wenzelm
parents: 74834
diff changeset
   136
text
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   137
\<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   138
  \<^ML>\<open>
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   139
      (
69604
d80b2df54d31 isabelle update -u control_cartouches;
wenzelm
parents: 69597
diff changeset
   140
        \<^term>\<open>""\<close>;
d80b2df54d31 isabelle update -u control_cartouches;
wenzelm
parents: 69597
diff changeset
   141
        \<^term>\<open>"abc"\<close>;
d80b2df54d31 isabelle update -u control_cartouches;
wenzelm
parents: 69597
diff changeset
   142
        \<^term>\<open>"abc" @ "xyz"\<close>;
d80b2df54d31 isabelle update -u control_cartouches;
wenzelm
parents: 69597
diff changeset
   143
        \<^term>\<open>"\<newline>"\<close>;
d80b2df54d31 isabelle update -u control_cartouches;
wenzelm
parents: 69597
diff changeset
   144
        \<^term>\<open>"\001\010\100"\<close>
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   145
      )
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   146
    \<close>
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   147
\<close>
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   148
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   149
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
   150
subsection \<open>Proof method syntax: ML tactic expression\<close>
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   151
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
   152
ML \<open>
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   153
structure ML_Tactic:
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   154
sig
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   155
  val set: (Proof.context -> tactic) -> Proof.context -> Proof.context
59064
a8bcb5a446c8 more abstract type Input.source;
wenzelm
parents: 58999
diff changeset
   156
  val ml_tactic: Input.source -> Proof.context -> tactic
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   157
end =
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   158
struct
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   159
  structure Data = Proof_Data(type T = Proof.context -> tactic fun init _ = K no_tac);
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   160
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   161
  val set = Data.put;
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   162
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 55133
diff changeset
   163
  fun ml_tactic source ctxt =
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   164
    let
69216
1a52baa70aed clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
wenzelm
parents: 69187
diff changeset
   165
      val ctxt' = ctxt
1a52baa70aed clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
wenzelm
parents: 69187
diff changeset
   166
        |> Context.proof_map (ML_Context.expression (Input.pos_of source)
1a52baa70aed clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
wenzelm
parents: 69187
diff changeset
   167
          (ML_Lex.read "Theory.local_setup (ML_Tactic.set (fn ctxt: Proof.context => (" @
1a52baa70aed clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
wenzelm
parents: 69187
diff changeset
   168
           ML_Lex.read_source source @ ML_Lex.read ")))"));
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   169
    in Data.get ctxt' ctxt end;
69216
1a52baa70aed clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
wenzelm
parents: 69187
diff changeset
   170
end
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
   171
\<close>
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   172
55048
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   173
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
   174
subsubsection \<open>Explicit version: method with cartouche argument\<close>
55048
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   175
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
   176
method_setup ml_tactic = \<open>
59809
87641097d0f3 tuned signature;
wenzelm
parents: 59184
diff changeset
   177
  Scan.lift Args.cartouche_input
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   178
    >> (fn arg => fn ctxt => SIMPLE_METHOD (ML_Tactic.ml_tactic arg ctxt))
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
   179
\<close>
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   180
63120
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
   181
lemma \<open>A \<and> B \<longrightarrow> B \<and> A\<close>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   182
  apply (ml_tactic \<open>resolve_tac \<^context> @{thms impI} 1\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   183
  apply (ml_tactic \<open>eresolve_tac \<^context> @{thms conjE} 1\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   184
  apply (ml_tactic \<open>resolve_tac \<^context> @{thms conjI} 1\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   185
  apply (ml_tactic \<open>ALLGOALS (assume_tac \<^context>)\<close>)
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   186
  done
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   187
63120
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
   188
lemma \<open>A \<and> B \<longrightarrow> B \<and> A\<close> by (ml_tactic \<open>blast_tac ctxt 1\<close>)
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   189
63120
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
   190
ML \<open>@{lemma \<open>A \<and> B \<longrightarrow> B \<and> A\<close> by (ml_tactic \<open>blast_tac ctxt 1\<close>)}\<close>
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   191
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   192
text \<open>\<^ML>\<open>@{lemma \<open>A \<and> B \<longrightarrow> B \<and> A\<close> by (ml_tactic \<open>blast_tac ctxt 1\<close>)}\<close>\<close>
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   193
55048
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   194
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
   195
subsubsection \<open>Implicit version: method with special name "cartouche" (dynamic!)\<close>
55048
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   196
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
   197
method_setup "cartouche" = \<open>
59809
87641097d0f3 tuned signature;
wenzelm
parents: 59184
diff changeset
   198
  Scan.lift Args.cartouche_input
55048
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   199
    >> (fn arg => fn ctxt => SIMPLE_METHOD (ML_Tactic.ml_tactic arg ctxt))
59068
8606f2ee11b1 update_cartouches;
wenzelm
parents: 59067
diff changeset
   200
\<close>
55048
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   201
63120
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
   202
lemma \<open>A \<and> B \<longrightarrow> B \<and> A\<close>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   203
  apply \<open>resolve_tac \<^context> @{thms impI} 1\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   204
  apply \<open>eresolve_tac \<^context> @{thms conjE} 1\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   205
  apply \<open>resolve_tac \<^context> @{thms conjI} 1\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   206
  apply \<open>ALLGOALS (assume_tac \<^context>)\<close>
55048
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   207
  done
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   208
63120
629a4c5e953e embedded content may be delimited via cartouches;
wenzelm
parents: 63054
diff changeset
   209
lemma \<open>A \<and> B \<longrightarrow> B \<and> A\<close>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   210
  by (\<open>resolve_tac \<^context> @{thms impI} 1\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   211
    \<open>eresolve_tac \<^context> @{thms conjE} 1\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   212
    \<open>resolve_tac \<^context> @{thms conjI} 1\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69216
diff changeset
   213
    \<open>assume_tac \<^context> 1\<close>+)
55048
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   214
59112
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59089
diff changeset
   215
59128
7b1931111e37 more examples;
wenzelm
parents: 59112
diff changeset
   216
subsection \<open>ML syntax\<close>
59112
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59089
diff changeset
   217
59128
7b1931111e37 more examples;
wenzelm
parents: 59112
diff changeset
   218
text \<open>Input source with position information:\<close>
59112
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59089
diff changeset
   219
ML \<open>
59135
wenzelm
parents: 59128
diff changeset
   220
  val s: Input.source = \<open>abc123def456\<close>;
59184
830bb7ddb3ab explicit message channels for "state", "information";
wenzelm
parents: 59135
diff changeset
   221
  Output.information ("Look here!" ^ Position.here (Input.pos_of s));
59112
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59089
diff changeset
   222
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59089
diff changeset
   223
  \<open>abc123def456\<close> |> Input.source_explode |> List.app (fn (s, pos) =>
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59089
diff changeset
   224
    if Symbol.is_digit s then Position.report pos Markup.ML_numeral else ());
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59089
diff changeset
   225
\<close>
e670969f34df expand ML cartouches to Input.source;
wenzelm
parents: 59089
diff changeset
   226
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
   227
end