| author | blanchet | 
| Tue, 16 Sep 2014 19:23:37 +0200 | |
| changeset 58353 | c9f374b64d99 | 
| parent 56500 | 90f17a04567d | 
| child 58889 | 5b7a9633cfa8 | 
| permissions | -rw-r--r-- | 
| 55033 | 1 | header {* Some examples with text cartouches *}
 | 
| 2 | ||
| 3 | theory Cartouche_Examples | |
| 4 | imports Main | |
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 5 | keywords | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 6 | "cartouche" "term_cartouche" :: diag and | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 7 | "text_cartouche" :: thy_decl | 
| 55033 | 8 | begin | 
| 9 | ||
| 56499 
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
 wenzelm parents: 
56304diff
changeset | 10 | subsection {* Regular outer syntax *}
 | 
| 
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
 wenzelm parents: 
56304diff
changeset | 11 | |
| 
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
 wenzelm parents: 
56304diff
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: 
56304diff
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: 
56304diff
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: 
56304diff
changeset | 15 | |
| 
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
 wenzelm parents: 
56304diff
changeset | 16 | notepad | 
| 
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
 wenzelm parents: 
56304diff
changeset | 17 | begin | 
| 
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
 wenzelm parents: 
56304diff
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: 
56304diff
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: 
56304diff
changeset | 20 | |
| 
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
 wenzelm parents: 
56304diff
changeset | 21 | fix x y :: 'a | 
| 
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
 wenzelm parents: 
56304diff
changeset | 22 | assume "x = y" | 
| 
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
 wenzelm parents: 
56304diff
changeset | 23 | note \<open>x = y\<close> | 
| 
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
 wenzelm parents: 
56304diff
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: 
56304diff
changeset | 25 | from \<open>x = y\<close> have "x = y" . | 
| 
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
 wenzelm parents: 
56304diff
changeset | 26 | |
| 
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
 wenzelm parents: 
56304diff
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: 
56304diff
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: 
56304diff
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: 
56499diff
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: 
56499diff
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: 
56499diff
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: 
56304diff
changeset | 33 | end | 
| 
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
 wenzelm parents: 
56304diff
changeset | 34 | |
| 
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
 wenzelm parents: 
56304diff
changeset | 35 | |
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 36 | subsection {* Outer syntax: cartouche within command syntax *}
 | 
| 55033 | 37 | |
| 38 | ML {*
 | |
| 39 |   Outer_Syntax.improper_command @{command_spec "cartouche"} ""
 | |
| 40 | (Parse.cartouche >> (fn s => | |
| 41 | Toplevel.imperative (fn () => writeln s))) | |
| 42 | *} | |
| 43 | ||
| 44 | cartouche \<open>abc\<close> | |
| 45 | cartouche \<open>abc \<open>\<alpha>\<beta>\<gamma>\<close> xzy\<close> | |
| 46 | ||
| 47 | ||
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 48 | subsection {* Inner syntax: string literals via cartouche *}
 | 
| 55033 | 49 | |
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 50 | ML {*
 | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 51 | local | 
| 55033 | 52 | val mk_nib = | 
| 53 | Syntax.const o Lexicon.mark_const o | |
| 54 | fst o Term.dest_Const o HOLogic.mk_nibble; | |
| 55 | ||
| 56 | fun mk_char (s, pos) = | |
| 57 | let | |
| 58 | val c = | |
| 59 | if Symbol.is_ascii s then ord s | |
| 55036 | 60 | else if s = "\<newline>" then 10 | 
| 55033 | 61 |           else error ("String literal contains illegal symbol: " ^ quote s ^ Position.here pos);
 | 
| 62 |       in Syntax.const @{const_syntax Char} $ mk_nib (c div 16) $ mk_nib (c mod 16) end;
 | |
| 63 | ||
| 64 |     fun mk_string [] = Const (@{const_syntax Nil}, @{typ string})
 | |
| 65 | | mk_string (s :: ss) = | |
| 66 |           Syntax.const @{const_syntax Cons} $ mk_char s $ mk_string ss;
 | |
| 67 | ||
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 68 | in | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 69 | fun string_tr content args = | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 70 |       let fun err () = raise TERM ("string_tr", args) in
 | 
| 55033 | 71 | (case args of | 
| 72 |           [(c as Const (@{syntax_const "_constrain"}, _)) $ Free (s, _) $ p] =>
 | |
| 73 | (case Term_Position.decode_position p of | |
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 74 | SOME (pos, _) => c $ mk_string (content (s, pos)) $ p | 
| 55033 | 75 | | NONE => err ()) | 
| 76 | | _ => err ()) | |
| 77 | end; | |
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 78 | end; | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 79 | *} | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 80 | |
| 55133 | 81 | syntax "_cartouche_string" :: "cartouche_position \<Rightarrow> string"  ("_")
 | 
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 82 | |
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 83 | parse_translation {*
 | 
| 55133 | 84 |   [(@{syntax_const "_cartouche_string"},
 | 
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 85 | K (string_tr (Symbol_Pos.cartouche_content o Symbol_Pos.explode)))] | 
| 55033 | 86 | *} | 
| 87 | ||
| 55133 | 88 | term "\<open>\<close>" | 
| 89 | term "\<open>abc\<close>" | |
| 90 | term "\<open>abc\<close> @ \<open>xyz\<close>" | |
| 91 | term "\<open>\<newline>\<close>" | |
| 92 | term "\<open>\001\010\100\<close>" | |
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 93 | |
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 94 | |
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 95 | subsection {* Alternate outer and inner syntax: string literals *}
 | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 96 | |
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 97 | subsubsection {* Nested quotes *}
 | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 98 | |
| 55133 | 99 | syntax "_string_string" :: "string_position \<Rightarrow> string"  ("_")
 | 
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 100 | |
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 101 | parse_translation {*
 | 
| 55133 | 102 |   [(@{syntax_const "_string_string"}, K (string_tr Lexicon.explode_string))]
 | 
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 103 | *} | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 104 | |
| 55133 | 105 | term "\"\"" | 
| 106 | term "\"abc\"" | |
| 107 | term "\"abc\" @ \"xyz\"" | |
| 108 | term "\"\<newline>\"" | |
| 109 | term "\"\001\010\100\"" | |
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 110 | |
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 111 | |
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 112 | subsubsection {* Term cartouche and regular quotes *}
 | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 113 | |
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 114 | ML {*
 | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 115 |   Outer_Syntax.improper_command @{command_spec "term_cartouche"} ""
 | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 116 | (Parse.inner_syntax Parse.cartouche >> (fn s => | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 117 | Toplevel.keep (fn state => | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 118 | let | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 119 | val ctxt = Toplevel.context_of state; | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 120 | val t = Syntax.read_term ctxt s; | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 121 | in writeln (Syntax.string_of_term ctxt t) end))) | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 122 | *} | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 123 | |
| 55133 | 124 | term_cartouche \<open>""\<close> | 
| 125 | term_cartouche \<open>"abc"\<close> | |
| 126 | term_cartouche \<open>"abc" @ "xyz"\<close> | |
| 127 | term_cartouche \<open>"\<newline>"\<close> | |
| 128 | term_cartouche \<open>"\001\010\100"\<close> | |
| 55033 | 129 | |
| 55047 | 130 | |
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 131 | subsubsection {* Further nesting: antiquotations *}
 | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 132 | |
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 133 | setup -- "ML antiquotation" | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 134 | {*
 | 
| 56072 | 135 |   ML_Antiquotation.inline @{binding term_cartouche}
 | 
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 136 | (Args.context -- Scan.lift (Parse.inner_syntax Parse.cartouche) >> | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
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: 
55048diff
changeset | 138 | *} | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 139 | |
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 140 | setup -- "document antiquotation" | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 141 | {*
 | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 142 |   Thy_Output.antiquotation @{binding ML_cartouche}
 | 
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
55133diff
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: 
55048diff
changeset | 144 | let | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 145 | val toks = | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 146 |           ML_Lex.read Position.none "fn _ => (" @
 | 
| 56278 
2576d3a40ed6
separate tokenization and language context for SML: no symbols, no antiquotes;
 wenzelm parents: 
56275diff
changeset | 147 | ML_Lex.read_source false source @ | 
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
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: 
56278diff
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: 
55048diff
changeset | 150 | in "" end); | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 151 | *} | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 152 | |
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 153 | ML {*
 | 
| 55133 | 154 |   @{term_cartouche \<open>""\<close>};
 | 
| 155 |   @{term_cartouche \<open>"abc"\<close>};
 | |
| 156 |   @{term_cartouche \<open>"abc" @ "xyz"\<close>};
 | |
| 157 |   @{term_cartouche \<open>"\<newline>"\<close>};
 | |
| 158 |   @{term_cartouche \<open>"\001\010\100"\<close>};
 | |
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 159 | *} | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 160 | |
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 161 | text {*
 | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 162 |   @{ML_cartouche
 | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 163 | \<open> | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 164 | ( | 
| 55133 | 165 |         @{term_cartouche \<open>""\<close>};
 | 
| 166 |         @{term_cartouche \<open>"abc"\<close>};
 | |
| 167 |         @{term_cartouche \<open>"abc" @ "xyz"\<close>};
 | |
| 168 |         @{term_cartouche \<open>"\<newline>"\<close>};
 | |
| 169 |         @{term_cartouche \<open>"\001\010\100"\<close>}
 | |
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 170 | ) | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 171 | \<close> | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 172 | } | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 173 | *} | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 174 | |
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 175 | |
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
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: 
55048diff
changeset | 177 | |
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 178 | ML {*
 | 
| 55110 
917e799f19da
avoid breakdown of document preparation, which does not understand cartouche tokens yet;
 wenzelm parents: 
55109diff
changeset | 179 | Outer_Syntax.command | 
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 180 |     @{command_spec "text_cartouche"} ""
 | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
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: 
55048diff
changeset | 182 | *} | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 183 | |
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 184 | text_cartouche | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 185 | \<open> | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 186 |   @{ML_cartouche
 | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 187 | \<open> | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 188 | ( | 
| 55133 | 189 |         @{term_cartouche \<open>""\<close>};
 | 
| 190 |         @{term_cartouche \<open>"abc"\<close>};
 | |
| 191 |         @{term_cartouche \<open>"abc" @ "xyz"\<close>};
 | |
| 192 |         @{term_cartouche \<open>"\<newline>"\<close>};
 | |
| 193 |         @{term_cartouche \<open>"\001\010\100"\<close>}
 | |
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 194 | ) | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 195 | \<close> | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 196 | } | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 197 | \<close> | 
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 198 | |
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 199 | |
| 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 200 | subsection {* Proof method syntax: ML tactic expression *}
 | 
| 55047 | 201 | |
| 202 | ML {*
 | |
| 203 | structure ML_Tactic: | |
| 204 | sig | |
| 205 | val set: (Proof.context -> tactic) -> Proof.context -> Proof.context | |
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
55133diff
changeset | 206 | val ml_tactic: Symbol_Pos.source -> Proof.context -> tactic | 
| 55047 | 207 | end = | 
| 208 | struct | |
| 209 | structure Data = Proof_Data(type T = Proof.context -> tactic fun init _ = K no_tac); | |
| 210 | ||
| 211 | val set = Data.put; | |
| 212 | ||
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
55133diff
changeset | 213 | fun ml_tactic source ctxt = | 
| 55047 | 214 | let | 
| 215 | val ctxt' = ctxt |> Context.proof_map | |
| 55828 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 wenzelm parents: 
55133diff
changeset | 216 | (ML_Context.expression (#pos source) | 
| 55047 | 217 | "fun tactic (ctxt : Proof.context) : tactic" | 
| 56278 
2576d3a40ed6
separate tokenization and language context for SML: no symbols, no antiquotes;
 wenzelm parents: 
56275diff
changeset | 218 | "Context.map_proof (ML_Tactic.set tactic)" (ML_Lex.read_source false source)); | 
| 55047 | 219 | in Data.get ctxt' ctxt end; | 
| 220 | end; | |
| 221 | *} | |
| 222 | ||
| 55048 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
55047diff
changeset | 223 | |
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 224 | subsubsection {* Explicit version: method with cartouche argument *}
 | 
| 55048 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
55047diff
changeset | 225 | |
| 55047 | 226 | method_setup ml_tactic = {*
 | 
| 227 | Scan.lift Args.cartouche_source_position | |
| 228 | >> (fn arg => fn ctxt => SIMPLE_METHOD (ML_Tactic.ml_tactic arg ctxt)) | |
| 229 | *} | |
| 230 | ||
| 231 | lemma "A \<and> B \<longrightarrow> B \<and> A" | |
| 232 |   apply (ml_tactic \<open>rtac @{thm impI} 1\<close>)
 | |
| 233 |   apply (ml_tactic \<open>etac @{thm conjE} 1\<close>)
 | |
| 234 |   apply (ml_tactic \<open>rtac @{thm conjI} 1\<close>)
 | |
| 235 | apply (ml_tactic \<open>ALLGOALS atac\<close>) | |
| 236 | done | |
| 237 | ||
| 238 | lemma "A \<and> B \<longrightarrow> B \<and> A" by (ml_tactic \<open>blast_tac ctxt 1\<close>) | |
| 239 | ||
| 240 | ML {* @{lemma "A \<and> B \<longrightarrow> B \<and> A" by (ml_tactic \<open>blast_tac ctxt 1\<close>)} *}
 | |
| 241 | ||
| 242 | text {* @{ML "@{lemma \"A \<and> B \<longrightarrow> B \<and> A\" by (ml_tactic \<open>blast_tac ctxt 1\<close>)}"} *}
 | |
| 243 | ||
| 55048 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
55047diff
changeset | 244 | |
| 55109 
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
 wenzelm parents: 
55048diff
changeset | 245 | subsubsection {* Implicit version: method with special name "cartouche" (dynamic!) *}
 | 
| 55048 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
55047diff
changeset | 246 | |
| 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
55047diff
changeset | 247 | method_setup "cartouche" = {*
 | 
| 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
55047diff
changeset | 248 | Scan.lift Args.cartouche_source_position | 
| 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
55047diff
changeset | 249 | >> (fn arg => fn ctxt => SIMPLE_METHOD (ML_Tactic.ml_tactic arg ctxt)) | 
| 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
55047diff
changeset | 250 | *} | 
| 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
55047diff
changeset | 251 | |
| 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
55047diff
changeset | 252 | lemma "A \<and> B \<longrightarrow> B \<and> A" | 
| 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
55047diff
changeset | 253 |   apply \<open>rtac @{thm impI} 1\<close>
 | 
| 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
55047diff
changeset | 254 |   apply \<open>etac @{thm conjE} 1\<close>
 | 
| 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
55047diff
changeset | 255 |   apply \<open>rtac @{thm conjI} 1\<close>
 | 
| 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
55047diff
changeset | 256 | apply \<open>ALLGOALS atac\<close> | 
| 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
55047diff
changeset | 257 | done | 
| 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
55047diff
changeset | 258 | |
| 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
55047diff
changeset | 259 | lemma "A \<and> B \<longrightarrow> B \<and> A" | 
| 
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
 wenzelm parents: 
55047diff
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: 
55047diff
changeset | 261 | |
| 55033 | 262 | end |