author | wenzelm |
Fri, 01 Sep 2017 14:58:19 +0200 | |
changeset 66590 | 8e1aac4eed11 |
parent 63120 | 629a4c5e953e |
child 67382 | 39e4668ded4f |
permissions | -rw-r--r-- |
59089 | 1 |
(* Title: HOL/ex/Cartouche_Examples.thy |
2 |
Author: Makarius |
|
3 |
*) |
|
4 |
||
59068 | 5 |
section \<open>Some examples with text cartouches\<close> |
55033 | 6 |
|
7 |
theory Cartouche_Examples |
|
8 |
imports Main |
|
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
9 |
keywords |
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
10 |
"cartouche" :: diag and |
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
11 |
"text_cartouche" :: thy_decl |
55033 | 12 |
begin |
13 |
||
59068 | 14 |
subsection \<open>Regular outer syntax\<close> |
56499
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents:
56304
diff
changeset
|
15 |
|
63054 | 16 |
text \<open>Text cartouches may be used in the outer syntax category \<open>text\<close>, |
17 |
as alternative to the traditional \<open>verbatim\<close> tokens. An example is |
|
61933 | 18 |
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
|
19 |
|
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents:
56304
diff
changeset
|
20 |
notepad |
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents:
56304
diff
changeset
|
21 |
begin |
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
22 |
txt \<open>Cartouches work as additional syntax for embedded language tokens |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
23 |
(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
|
24 |
(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
|
25 |
|
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents:
56304
diff
changeset
|
26 |
fix x y :: 'a |
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
27 |
assume \<open>x = y\<close> |
56499
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents:
56304
diff
changeset
|
28 |
note \<open>x = y\<close> |
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
29 |
have \<open>x = y\<close> by (rule \<open>x = y\<close>) |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
30 |
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
|
31 |
|
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents:
56304
diff
changeset
|
32 |
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
|
33 |
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
|
34 |
[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
|
35 |
|
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
36 |
have \<open>x = y\<close> |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
37 |
by (tactic \<open>resolve_tac @{context} @{thms \<open>x = y\<close>} 1\<close>) |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
38 |
\<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
|
39 |
end |
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents:
56304
diff
changeset
|
40 |
|
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
wenzelm
parents:
56304
diff
changeset
|
41 |
|
59068 | 42 |
subsection \<open>Outer syntax: cartouche within command syntax\<close> |
55033 | 43 |
|
59068 | 44 |
ML \<open> |
59936
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
wenzelm
parents:
59809
diff
changeset
|
45 |
Outer_Syntax.command @{command_keyword cartouche} "" |
55033 | 46 |
(Parse.cartouche >> (fn s => |
60189 | 47 |
Toplevel.keep (fn _ => writeln s))) |
59068 | 48 |
\<close> |
55033 | 49 |
|
50 |
cartouche \<open>abc\<close> |
|
51 |
cartouche \<open>abc \<open>\<alpha>\<beta>\<gamma>\<close> xzy\<close> |
|
52 |
||
53 |
||
59068 | 54 |
subsection \<open>Inner syntax: string literals via cartouche\<close> |
55033 | 55 |
|
59068 | 56 |
ML \<open> |
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
57 |
local |
55033 | 58 |
fun mk_char (s, pos) = |
59 |
let |
|
60 |
val c = |
|
61 |
if Symbol.is_ascii s then ord s |
|
55036 | 62 |
else if s = "\<newline>" then 10 |
55033 | 63 |
else error ("String literal contains illegal symbol: " ^ quote s ^ Position.here pos); |
62597 | 64 |
in Syntax.const @{const_syntax Char} $ HOLogic.mk_numeral c end; |
55033 | 65 |
|
66 |
fun mk_string [] = Const (@{const_syntax Nil}, @{typ string}) |
|
67 |
| mk_string (s :: ss) = |
|
68 |
Syntax.const @{const_syntax Cons} $ mk_char s $ mk_string ss; |
|
69 |
||
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
70 |
in |
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
71 |
fun string_tr content args = |
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
72 |
let fun err () = raise TERM ("string_tr", args) in |
55033 | 73 |
(case args of |
74 |
[(c as Const (@{syntax_const "_constrain"}, _)) $ Free (s, _) $ p] => |
|
75 |
(case Term_Position.decode_position p of |
|
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
76 |
SOME (pos, _) => c $ mk_string (content (s, pos)) $ p |
55033 | 77 |
| NONE => err ()) |
78 |
| _ => err ()) |
|
79 |
end; |
|
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
80 |
end; |
59068 | 81 |
\<close> |
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
82 |
|
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
83 |
syntax "_cartouche_string" :: \<open>cartouche_position \<Rightarrow> string\<close> ("_") |
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
84 |
|
59068 | 85 |
parse_translation \<open> |
55133 | 86 |
[(@{syntax_const "_cartouche_string"}, |
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
87 |
K (string_tr (Symbol_Pos.cartouche_content o Symbol_Pos.explode)))] |
59068 | 88 |
\<close> |
55033 | 89 |
|
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
90 |
term \<open>\<open>\<close>\<close> |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
91 |
term \<open>\<open>abc\<close>\<close> |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
92 |
term \<open>\<open>abc\<close> @ \<open>xyz\<close>\<close> |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
93 |
term \<open>\<open>\<newline>\<close>\<close> |
55109
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 |
|
59068 | 96 |
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
|
97 |
|
59068 | 98 |
subsubsection \<open>Nested quotes\<close> |
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
99 |
|
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
100 |
syntax "_string_string" :: \<open>string_position \<Rightarrow> string\<close> ("_") |
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
101 |
|
59068 | 102 |
parse_translation \<open> |
55133 | 103 |
[(@{syntax_const "_string_string"}, K (string_tr Lexicon.explode_string))] |
59068 | 104 |
\<close> |
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
105 |
|
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
106 |
term \<open>""\<close> |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
107 |
term \<open>"abc"\<close> |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
108 |
term \<open>"abc" @ "xyz"\<close> |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
109 |
term \<open>"\<newline>"\<close> |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
110 |
term \<open>"\001\010\100"\<close> |
55033 | 111 |
|
55047 | 112 |
|
59068 | 113 |
subsubsection \<open>Further nesting: antiquotations\<close> |
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
114 |
|
59068 | 115 |
ML \<open> |
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
116 |
@{term \<open>""\<close>}; |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
117 |
@{term \<open>"abc"\<close>}; |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
118 |
@{term \<open>"abc" @ "xyz"\<close>}; |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
119 |
@{term \<open>"\<newline>"\<close>}; |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
120 |
@{term \<open>"\001\010\100"\<close>}; |
59068 | 121 |
\<close> |
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
122 |
|
59068 | 123 |
text \<open> |
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
124 |
@{ML |
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
125 |
\<open> |
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
126 |
( |
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
127 |
@{term \<open>""\<close>}; |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
128 |
@{term \<open>"abc"\<close>}; |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
129 |
@{term \<open>"abc" @ "xyz"\<close>}; |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
130 |
@{term \<open>"\<newline>"\<close>}; |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
131 |
@{term \<open>"\001\010\100"\<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 |
\<close> |
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
134 |
} |
59068 | 135 |
\<close> |
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
136 |
|
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
137 |
|
59068 | 138 |
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
|
139 |
|
59068 | 140 |
ML \<open> |
55110
917e799f19da
avoid breakdown of document preparation, which does not understand cartouche tokens yet;
wenzelm
parents:
55109
diff
changeset
|
141 |
Outer_Syntax.command |
59936
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
wenzelm
parents:
59809
diff
changeset
|
142 |
@{command_keyword text_cartouche} "" |
61457 | 143 |
(Parse.opt_target -- Parse.input Parse.cartouche |
144 |
>> Thy_Output.document_command {markdown = true}) |
|
59068 | 145 |
\<close> |
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
146 |
|
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
147 |
text_cartouche |
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
148 |
\<open> |
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
149 |
@{ML |
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
150 |
\<open> |
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
151 |
( |
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
152 |
@{term \<open>""\<close>}; |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
153 |
@{term \<open>"abc"\<close>}; |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
154 |
@{term \<open>"abc" @ "xyz"\<close>}; |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
155 |
@{term \<open>"\<newline>"\<close>}; |
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
156 |
@{term \<open>"\001\010\100"\<close>} |
55109
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
157 |
) |
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
158 |
\<close> |
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 |
\<close> |
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
161 |
|
ecff9e26360c
more cartouche examples, including uniform nesting of sub-languages;
wenzelm
parents:
55048
diff
changeset
|
162 |
|
59068 | 163 |
subsection \<open>Proof method syntax: ML tactic expression\<close> |
55047 | 164 |
|
59068 | 165 |
ML \<open> |
55047 | 166 |
structure ML_Tactic: |
167 |
sig |
|
168 |
val set: (Proof.context -> tactic) -> Proof.context -> Proof.context |
|
59064 | 169 |
val ml_tactic: Input.source -> Proof.context -> tactic |
55047 | 170 |
end = |
171 |
struct |
|
172 |
structure Data = Proof_Data(type T = Proof.context -> tactic fun init _ = K no_tac); |
|
173 |
||
174 |
val set = Data.put; |
|
175 |
||
55828
42ac3cfb89f6
clarified language markup: added "delimited" property;
wenzelm
parents:
55133
diff
changeset
|
176 |
fun ml_tactic source ctxt = |
55047 | 177 |
let |
178 |
val ctxt' = ctxt |> Context.proof_map |
|
59064 | 179 |
(ML_Context.expression (Input.range_of source) "tactic" "Proof.context -> tactic" |
58991
92b6f4e68c5a
more careful ML source positions, for improved PIDE markup;
wenzelm
parents:
58978
diff
changeset
|
180 |
"Context.map_proof (ML_Tactic.set tactic)" |
59067 | 181 |
(ML_Lex.read "fn ctxt: Proof.context =>" @ ML_Lex.read_source false source)); |
55047 | 182 |
in Data.get ctxt' ctxt end; |
183 |
end; |
|
59068 | 184 |
\<close> |
55047 | 185 |
|
55048
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
wenzelm
parents:
55047
diff
changeset
|
186 |
|
59068 | 187 |
subsubsection \<open>Explicit version: method with cartouche argument\<close> |
55048
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
wenzelm
parents:
55047
diff
changeset
|
188 |
|
59068 | 189 |
method_setup ml_tactic = \<open> |
59809 | 190 |
Scan.lift Args.cartouche_input |
55047 | 191 |
>> (fn arg => fn ctxt => SIMPLE_METHOD (ML_Tactic.ml_tactic arg ctxt)) |
59068 | 192 |
\<close> |
55047 | 193 |
|
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
194 |
lemma \<open>A \<and> B \<longrightarrow> B \<and> A\<close> |
60754 | 195 |
apply (ml_tactic \<open>resolve_tac @{context} @{thms impI} 1\<close>) |
196 |
apply (ml_tactic \<open>eresolve_tac @{context} @{thms conjE} 1\<close>) |
|
197 |
apply (ml_tactic \<open>resolve_tac @{context} @{thms conjI} 1\<close>) |
|
198 |
apply (ml_tactic \<open>ALLGOALS (assume_tac @{context})\<close>) |
|
55047 | 199 |
done |
200 |
||
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
201 |
lemma \<open>A \<and> B \<longrightarrow> B \<and> A\<close> by (ml_tactic \<open>blast_tac ctxt 1\<close>) |
55047 | 202 |
|
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
203 |
ML \<open>@{lemma \<open>A \<and> B \<longrightarrow> B \<and> A\<close> by (ml_tactic \<open>blast_tac ctxt 1\<close>)}\<close> |
55047 | 204 |
|
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
205 |
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 | 206 |
|
55048
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
wenzelm
parents:
55047
diff
changeset
|
207 |
|
59068 | 208 |
subsubsection \<open>Implicit version: method with special name "cartouche" (dynamic!)\<close> |
55048
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
wenzelm
parents:
55047
diff
changeset
|
209 |
|
59068 | 210 |
method_setup "cartouche" = \<open> |
59809 | 211 |
Scan.lift Args.cartouche_input |
55048
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
wenzelm
parents:
55047
diff
changeset
|
212 |
>> (fn arg => fn ctxt => SIMPLE_METHOD (ML_Tactic.ml_tactic arg ctxt)) |
59068 | 213 |
\<close> |
55048
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
wenzelm
parents:
55047
diff
changeset
|
214 |
|
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
215 |
lemma \<open>A \<and> B \<longrightarrow> B \<and> A\<close> |
60754 | 216 |
apply \<open>resolve_tac @{context} @{thms impI} 1\<close> |
217 |
apply \<open>eresolve_tac @{context} @{thms conjE} 1\<close> |
|
218 |
apply \<open>resolve_tac @{context} @{thms conjI} 1\<close> |
|
219 |
apply \<open>ALLGOALS (assume_tac @{context})\<close> |
|
55048
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
wenzelm
parents:
55047
diff
changeset
|
220 |
done |
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
wenzelm
parents:
55047
diff
changeset
|
221 |
|
63120
629a4c5e953e
embedded content may be delimited via cartouches;
wenzelm
parents:
63054
diff
changeset
|
222 |
lemma \<open>A \<and> B \<longrightarrow> B \<and> A\<close> |
60754 | 223 |
by (\<open>resolve_tac @{context} @{thms impI} 1\<close>, |
224 |
\<open>eresolve_tac @{context} @{thms conjE} 1\<close>, |
|
225 |
\<open>resolve_tac @{context} @{thms conjI} 1\<close>, |
|
226 |
\<open>assume_tac @{context} 1\<close>+) |
|
55048
ce34a2934386
implicit "cartouche" method (experimental, undocumented);
wenzelm
parents:
55047
diff
changeset
|
227 |
|
59112 | 228 |
|
59128 | 229 |
subsection \<open>ML syntax\<close> |
59112 | 230 |
|
59128 | 231 |
text \<open>Input source with position information:\<close> |
59112 | 232 |
ML \<open> |
59135 | 233 |
val s: Input.source = \<open>abc123def456\<close>; |
59184
830bb7ddb3ab
explicit message channels for "state", "information";
wenzelm
parents:
59135
diff
changeset
|
234 |
Output.information ("Look here!" ^ Position.here (Input.pos_of s)); |
59112 | 235 |
|
236 |
\<open>abc123def456\<close> |> Input.source_explode |> List.app (fn (s, pos) => |
|
237 |
if Symbol.is_digit s then Position.report pos Markup.ML_numeral else ()); |
|
238 |
\<close> |
|
239 |
||
55033 | 240 |
end |