src/HOL/ex/Cartouche_Examples.thy
author wenzelm
Mon, 03 Nov 2014 14:50:27 +0100
changeset 58893 9e0ecb66d6a7
parent 58889 5b7a9633cfa8
child 58928 23d0ffd48006
permissions -rw-r--r--
eliminated unused int_only flag (see also c12484a27367); just proper commands;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 56500
diff changeset
     1
section {* Some examples with text cartouches *}
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
     2
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
     3
theory Cartouche_Examples
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
     4
imports Main
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
     5
keywords
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
     6
  "cartouche" "term_cartouche" :: diag and
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
     7
  "text_cartouche" :: thy_decl
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
     8
begin
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
     9
56499
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    10
subsection {* Regular outer syntax *}
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    11
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    12
text \<open>Text cartouches may be used in the outer syntax category "text",
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    13
  as alternative to the traditional "verbatim" tokens.  An example is
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    14
  this text block.\<close>  -- \<open>The same works for small side-comments.\<close>
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    15
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    16
notepad
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    17
begin
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    18
  txt \<open>Moreover, cartouches work as additional syntax in the
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    19
    "altstring" category, for literal fact references.  For example:\<close>
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    20
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    21
  fix x y :: 'a
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    22
  assume "x = y"
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    23
  note \<open>x = y\<close>
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    24
  have "x = y" by (rule \<open>x = y\<close>)
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    25
  from \<open>x = y\<close> have "x = y" .
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    26
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    27
  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
    28
    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
    29
    [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
    30
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
    31
  have "x = y"
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
    32
    by (tactic \<open>rtac @{thm \<open>x = y\<close>} 1\<close>)  -- \<open>more cartouches involving ML\<close>
56499
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    33
end
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    34
7e0178c84994 allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents: 56304
diff changeset
    35
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    36
subsection {* Outer syntax: cartouche within command syntax *}
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    37
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    38
ML {*
58893
9e0ecb66d6a7 eliminated unused int_only flag (see also c12484a27367);
wenzelm
parents: 58889
diff changeset
    39
  Outer_Syntax.command @{command_spec "cartouche"} ""
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    40
    (Parse.cartouche >> (fn s =>
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    41
      Toplevel.imperative (fn () => writeln s)))
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    42
*}
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    43
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    44
cartouche \<open>abc\<close>
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    45
cartouche \<open>abc \<open>\<alpha>\<beta>\<gamma>\<close> xzy\<close>
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    46
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    47
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    48
subsection {* Inner syntax: string literals via cartouche *}
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    49
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    50
ML {*
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    51
  local
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    52
    val mk_nib =
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    53
      Syntax.const o Lexicon.mark_const o
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    54
        fst o Term.dest_Const o HOLogic.mk_nibble;
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    55
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);
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    62
      in Syntax.const @{const_syntax Char} $ mk_nib (c div 16) $ mk_nib (c mod 16) end;
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    63
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    64
    fun mk_string [] = Const (@{const_syntax Nil}, @{typ string})
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    65
      | mk_string (s :: ss) =
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    66
          Syntax.const @{const_syntax Cons} $ mk_char s $ mk_string ss;
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
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    72
          [(c as Const (@{syntax_const "_constrain"}, _)) $ Free (s, _) $ p] =>
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    73
            (case Term_Position.decode_position p of
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
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;
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    79
*}
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    80
55133
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
    81
syntax "_cartouche_string" :: "cartouche_position \<Rightarrow> string"  ("_")
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    82
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    83
parse_translation {*
55133
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
    84
  [(@{syntax_const "_cartouche_string"},
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)))]
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    86
*}
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
    87
55133
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
    88
term "\<open>\<close>"
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
    89
term "\<open>abc\<close>"
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
    90
term "\<open>abc\<close> @ \<open>xyz\<close>"
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
    91
term "\<open>\<newline>\<close>"
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
    92
term "\<open>\001\010\100\<close>"
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    93
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    94
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    95
subsection {* Alternate outer and inner syntax: string literals *}
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    96
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    97
subsubsection {* Nested quotes *}
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
    98
55133
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
    99
syntax "_string_string" :: "string_position \<Rightarrow> string"  ("_")
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   100
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   101
parse_translation {*
55133
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   102
  [(@{syntax_const "_string_string"}, K (string_tr Lexicon.explode_string))]
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   103
*}
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   104
55133
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   105
term "\"\""
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   106
term "\"abc\""
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   107
term "\"abc\" @ \"xyz\""
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   108
term "\"\<newline>\""
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   109
term "\"\001\010\100\""
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   110
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   111
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   112
subsubsection {* Term cartouche and regular quotes *}
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   113
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   114
ML {*
58893
9e0ecb66d6a7 eliminated unused int_only flag (see also c12484a27367);
wenzelm
parents: 58889
diff changeset
   115
  Outer_Syntax.command @{command_spec "term_cartouche"} ""
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   116
    (Parse.inner_syntax Parse.cartouche >> (fn s =>
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   117
      Toplevel.keep (fn state =>
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   118
        let
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   119
          val ctxt = Toplevel.context_of state;
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   120
          val t = Syntax.read_term ctxt s;
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   121
        in writeln (Syntax.string_of_term ctxt t) end)))
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   122
*}
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   123
55133
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   124
term_cartouche \<open>""\<close>
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   125
term_cartouche \<open>"abc"\<close>
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   126
term_cartouche \<open>"abc" @ "xyz"\<close>
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   127
term_cartouche \<open>"\<newline>"\<close>
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   128
term_cartouche \<open>"\001\010\100"\<close>
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
   129
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   130
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   131
subsubsection {* Further nesting: antiquotations *}
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
setup -- "ML antiquotation"
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   134
{*
56072
31e427387ab5 tuned signature -- clarified module name;
wenzelm
parents: 55828
diff changeset
   135
  ML_Antiquotation.inline @{binding term_cartouche}
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   136
    (Args.context -- Scan.lift (Parse.inner_syntax Parse.cartouche) >>
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   137
      (fn (ctxt, s) => ML_Syntax.atomic (ML_Syntax.print_term (Syntax.read_term ctxt s))))
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   138
*}
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   139
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   140
setup -- "document antiquotation"
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   141
{*
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   142
  Thy_Output.antiquotation @{binding ML_cartouche}
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 55133
diff changeset
   143
    (Scan.lift Args.cartouche_source_position) (fn {context, ...} => fn source =>
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   144
      let
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   145
        val toks =
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   146
          ML_Lex.read Position.none "fn _ => (" @
56278
2576d3a40ed6 separate tokenization and language context for SML: no symbols, no antiquotes;
wenzelm
parents: 56275
diff changeset
   147
          ML_Lex.read_source false source @
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   148
          ML_Lex.read Position.none ");";
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
   149
        val _ = ML_Context.eval_in (SOME context) ML_Compiler.flags (#pos source) toks;
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   150
      in "" end);
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   151
*}
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   152
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   153
ML {*
55133
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   154
  @{term_cartouche \<open>""\<close>};
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   155
  @{term_cartouche \<open>"abc"\<close>};
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   156
  @{term_cartouche \<open>"abc" @ "xyz"\<close>};
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   157
  @{term_cartouche \<open>"\<newline>"\<close>};
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   158
  @{term_cartouche \<open>"\001\010\100"\<close>};
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   159
*}
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   160
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   161
text {*
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   162
  @{ML_cartouche
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   163
    \<open>
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   164
      (
55133
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   165
        @{term_cartouche \<open>""\<close>};
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   166
        @{term_cartouche \<open>"abc"\<close>};
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   167
        @{term_cartouche \<open>"abc" @ "xyz"\<close>};
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   168
        @{term_cartouche \<open>"\<newline>"\<close>};
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   169
        @{term_cartouche \<open>"\001\010\100"\<close>}
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   170
      )
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   171
    \<close>
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   172
  }
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   173
*}
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   174
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   175
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   176
subsubsection {* Uniform nesting of sub-languages: document source, ML, term, string literals *}
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   177
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   178
ML {*
55110
917e799f19da avoid breakdown of document preparation, which does not understand cartouche tokens yet;
wenzelm
parents: 55109
diff changeset
   179
  Outer_Syntax.command
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   180
    @{command_spec "text_cartouche"} ""
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   181
    (Parse.opt_target -- Parse.source_position Parse.cartouche >> Isar_Cmd.local_theory_markup)
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   182
*}
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   183
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   184
text_cartouche
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   185
\<open>
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   186
  @{ML_cartouche
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   187
    \<open>
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   188
      (
55133
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   189
        @{term_cartouche \<open>""\<close>};
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   190
        @{term_cartouche \<open>"abc"\<close>};
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   191
        @{term_cartouche \<open>"abc" @ "xyz"\<close>};
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   192
        @{term_cartouche \<open>"\<newline>"\<close>};
687656233ad8 simplified inner syntax;
wenzelm
parents: 55110
diff changeset
   193
        @{term_cartouche \<open>"\001\010\100"\<close>}
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   194
      )
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   195
    \<close>
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   196
  }
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   197
\<close>
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   198
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   199
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   200
subsection {* Proof method syntax: ML tactic expression *}
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   201
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   202
ML {*
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   203
structure ML_Tactic:
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   204
sig
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   205
  val set: (Proof.context -> tactic) -> Proof.context -> Proof.context
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 55133
diff changeset
   206
  val ml_tactic: Symbol_Pos.source -> Proof.context -> tactic
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   207
end =
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   208
struct
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   209
  structure Data = Proof_Data(type T = Proof.context -> tactic fun init _ = K no_tac);
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   210
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   211
  val set = Data.put;
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   212
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 55133
diff changeset
   213
  fun ml_tactic source ctxt =
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   214
    let
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   215
      val ctxt' = ctxt |> Context.proof_map
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 55133
diff changeset
   216
        (ML_Context.expression (#pos source)
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   217
          "fun tactic (ctxt : Proof.context) : tactic"
56278
2576d3a40ed6 separate tokenization and language context for SML: no symbols, no antiquotes;
wenzelm
parents: 56275
diff changeset
   218
          "Context.map_proof (ML_Tactic.set tactic)" (ML_Lex.read_source false source));
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   219
    in Data.get ctxt' ctxt end;
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   220
end;
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   221
*}
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   222
55048
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   223
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   224
subsubsection {* Explicit version: method with cartouche argument *}
55048
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   225
55047
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   226
method_setup ml_tactic = {*
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   227
  Scan.lift Args.cartouche_source_position
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   228
    >> (fn arg => fn ctxt => SIMPLE_METHOD (ML_Tactic.ml_tactic arg ctxt))
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   229
*}
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   230
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   231
lemma "A \<and> B \<longrightarrow> B \<and> A"
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   232
  apply (ml_tactic \<open>rtac @{thm impI} 1\<close>)
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   233
  apply (ml_tactic \<open>etac @{thm conjE} 1\<close>)
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   234
  apply (ml_tactic \<open>rtac @{thm conjI} 1\<close>)
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   235
  apply (ml_tactic \<open>ALLGOALS atac\<close>)
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   236
  done
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   237
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   238
lemma "A \<and> B \<longrightarrow> B \<and> A" by (ml_tactic \<open>blast_tac ctxt 1\<close>)
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   239
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   240
ML {* @{lemma "A \<and> B \<longrightarrow> B \<and> A" by (ml_tactic \<open>blast_tac ctxt 1\<close>)} *}
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   241
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   242
text {* @{ML "@{lemma \"A \<and> B \<longrightarrow> B \<and> A\" by (ml_tactic \<open>blast_tac ctxt 1\<close>)}"} *}
660398f1d806 more examples;
wenzelm
parents: 55036
diff changeset
   243
55048
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   244
55109
ecff9e26360c more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents: 55048
diff changeset
   245
subsubsection {* Implicit version: method with special name "cartouche" (dynamic!) *}
55048
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   246
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   247
method_setup "cartouche" = {*
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   248
  Scan.lift Args.cartouche_source_position
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   249
    >> (fn arg => fn ctxt => SIMPLE_METHOD (ML_Tactic.ml_tactic arg ctxt))
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   250
*}
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   251
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   252
lemma "A \<and> B \<longrightarrow> B \<and> A"
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   253
  apply \<open>rtac @{thm impI} 1\<close>
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   254
  apply \<open>etac @{thm conjE} 1\<close>
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   255
  apply \<open>rtac @{thm conjI} 1\<close>
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   256
  apply \<open>ALLGOALS atac\<close>
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   257
  done
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   258
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   259
lemma "A \<and> B \<longrightarrow> B \<and> A"
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   260
  by (\<open>rtac @{thm impI} 1\<close>, \<open>etac @{thm conjE} 1\<close>, \<open>rtac @{thm conjI} 1\<close>, \<open>atac 1\<close>+)
ce34a2934386 implicit "cartouche" method (experimental, undocumented);
wenzelm
parents: 55047
diff changeset
   261
55033
8e8243975860 support for nested text cartouches;
wenzelm
parents:
diff changeset
   262
end