author | nipkow |
Sat, 29 Sep 2018 21:05:32 +0200 | |
changeset 69082 | 0405d06f08f3 |
parent 67505 | ceb324e34c14 |
child 73761 | ef1a18e20ace |
permissions | -rw-r--r-- |
47269 | 1 |
(* Title: HOL/Library/LaTeXsugar.thy |
55116 | 2 |
Author: Gerwin Klein, Tobias Nipkow, Norbert Schirmer |
47269 | 3 |
Copyright 2005 NICTA and TUM |
4 |
*) |
|
5 |
||
6 |
(*<*) |
|
7 |
theory LaTeXsugar |
|
8 |
imports Main |
|
9 |
begin |
|
10 |
||
11 |
(* DUMMY *) |
|
63935
aa1fe1103ab8
raw control symbols are superseded by Latex.embed_raw;
wenzelm
parents:
63120
diff
changeset
|
12 |
consts DUMMY :: 'a ("\<^latex>\<open>\\_\<close>") |
47269 | 13 |
|
14 |
(* THEOREMS *) |
|
15 |
notation (Rule output) |
|
63935
aa1fe1103ab8
raw control symbols are superseded by Latex.embed_raw;
wenzelm
parents:
63120
diff
changeset
|
16 |
Pure.imp ("\<^latex>\<open>\\mbox{}\\inferrule{\\mbox{\<close>_\<^latex>\<open>}}\<close>\<^latex>\<open>{\\mbox{\<close>_\<^latex>\<open>}}\<close>") |
47269 | 17 |
|
18 |
syntax (Rule output) |
|
19 |
"_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop" |
|
63935
aa1fe1103ab8
raw control symbols are superseded by Latex.embed_raw;
wenzelm
parents:
63120
diff
changeset
|
20 |
("\<^latex>\<open>\\mbox{}\\inferrule{\<close>_\<^latex>\<open>}\<close>\<^latex>\<open>{\\mbox{\<close>_\<^latex>\<open>}}\<close>") |
47269 | 21 |
|
22 |
"_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" |
|
63935
aa1fe1103ab8
raw control symbols are superseded by Latex.embed_raw;
wenzelm
parents:
63120
diff
changeset
|
23 |
("\<^latex>\<open>\\mbox{\<close>_\<^latex>\<open>}\\\\\<close>/ _") |
47269 | 24 |
|
63935
aa1fe1103ab8
raw control symbols are superseded by Latex.embed_raw;
wenzelm
parents:
63120
diff
changeset
|
25 |
"_asm" :: "prop \<Rightarrow> asms" ("\<^latex>\<open>\\mbox{\<close>_\<^latex>\<open>}\<close>") |
47269 | 26 |
|
27 |
notation (Axiom output) |
|
63935
aa1fe1103ab8
raw control symbols are superseded by Latex.embed_raw;
wenzelm
parents:
63120
diff
changeset
|
28 |
"Trueprop" ("\<^latex>\<open>\\mbox{}\\inferrule{\\mbox{}}{\\mbox{\<close>_\<^latex>\<open>}}\<close>") |
47269 | 29 |
|
30 |
notation (IfThen output) |
|
63935
aa1fe1103ab8
raw control symbols are superseded by Latex.embed_raw;
wenzelm
parents:
63120
diff
changeset
|
31 |
Pure.imp ("\<^latex>\<open>{\\normalsize{}\<close>If\<^latex>\<open>\\,}\<close> _/ \<^latex>\<open>{\\normalsize \\,\<close>then\<^latex>\<open>\\,}\<close>/ _.") |
47269 | 32 |
syntax (IfThen output) |
33 |
"_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop" |
|
63935
aa1fe1103ab8
raw control symbols are superseded by Latex.embed_raw;
wenzelm
parents:
63120
diff
changeset
|
34 |
("\<^latex>\<open>{\\normalsize{}\<close>If\<^latex>\<open>\\,}\<close> _ /\<^latex>\<open>{\\normalsize \\,\<close>then\<^latex>\<open>\\,}\<close>/ _.") |
aa1fe1103ab8
raw control symbols are superseded by Latex.embed_raw;
wenzelm
parents:
63120
diff
changeset
|
35 |
"_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("\<^latex>\<open>\\mbox{\<close>_\<^latex>\<open>}\<close> /\<^latex>\<open>{\\normalsize \\,\<close>and\<^latex>\<open>\\,}\<close>/ _") |
aa1fe1103ab8
raw control symbols are superseded by Latex.embed_raw;
wenzelm
parents:
63120
diff
changeset
|
36 |
"_asm" :: "prop \<Rightarrow> asms" ("\<^latex>\<open>\\mbox{\<close>_\<^latex>\<open>}\<close>") |
47269 | 37 |
|
38 |
notation (IfThenNoBox output) |
|
63935
aa1fe1103ab8
raw control symbols are superseded by Latex.embed_raw;
wenzelm
parents:
63120
diff
changeset
|
39 |
Pure.imp ("\<^latex>\<open>{\\normalsize{}\<close>If\<^latex>\<open>\\,}\<close> _/ \<^latex>\<open>{\\normalsize \\,\<close>then\<^latex>\<open>\\,}\<close>/ _.") |
47269 | 40 |
syntax (IfThenNoBox output) |
41 |
"_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop" |
|
63935
aa1fe1103ab8
raw control symbols are superseded by Latex.embed_raw;
wenzelm
parents:
63120
diff
changeset
|
42 |
("\<^latex>\<open>{\\normalsize{}\<close>If\<^latex>\<open>\\,}\<close> _ /\<^latex>\<open>{\\normalsize \\,\<close>then\<^latex>\<open>\\,}\<close>/ _.") |
aa1fe1103ab8
raw control symbols are superseded by Latex.embed_raw;
wenzelm
parents:
63120
diff
changeset
|
43 |
"_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("_ /\<^latex>\<open>{\\normalsize \\,\<close>and\<^latex>\<open>\\,}\<close>/ _") |
47269 | 44 |
"_asm" :: "prop \<Rightarrow> asms" ("_") |
45 |
||
67463 | 46 |
setup \<open> |
47 |
Thy_Output.antiquotation_pretty_source \<^binding>\<open>const_typ\<close> (Scan.lift Args.embedded_inner_syntax) |
|
48 |
(fn ctxt => fn c => |
|
69082 | 49 |
let val tc = Proof_Context.read_const {proper = false, strict = false} ctxt c in |
67505
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67463
diff
changeset
|
50 |
Pretty.block [Thy_Output.pretty_term ctxt tc, Pretty.str " ::", |
ceb324e34c14
clarified signature: items with \isasep are special;
wenzelm
parents:
67463
diff
changeset
|
51 |
Pretty.brk 1, Syntax.pretty_typ ctxt (fastype_of tc)] |
67463 | 52 |
end) |
67406 | 53 |
\<close> |
49629 | 54 |
|
47269 | 55 |
end |
67399 | 56 |
(*>*) |