| author | wenzelm | 
| Tue, 08 Oct 2024 12:10:35 +0200 | |
| changeset 81125 | ec121999a9cb | 
| parent 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 15469 | 1 | (* Title: HOL/Library/LaTeXsugar.thy | 
| 55077 | 2 | Author: Gerwin Klein, Tobias Nipkow, Norbert Schirmer | 
| 15469 | 3 | Copyright 2005 NICTA and TUM | 
| 4 | *) | |
| 5 | ||
| 6 | (*<*) | |
| 7 | theory LaTeXsugar | |
| 30663 
0b6aff7451b2
Main is (Complex_Main) base entry point in library theories
 haftmann parents: 
30304diff
changeset | 8 | imports Main | 
| 15469 | 9 | begin | 
| 10 | ||
| 78471 | 11 | (* Boxing *) | 
| 12 | ||
| 13 | definition mbox :: "'a \<Rightarrow> 'a" where | |
| 14 | "mbox x = x" | |
| 15 | ||
| 16 | definition mbox0 :: "'a \<Rightarrow> 'a" where | |
| 17 | "mbox0 x = x" | |
| 18 | ||
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 19 | notation (latex) mbox (\<open>\<^latex>\<open>\mbox{\<close>_\<^latex>\<open>}\<close>\<close> [999] 999)
 | 
| 78471 | 20 | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 21 | notation (latex) mbox0 (\<open>\<^latex>\<open>\mbox{\<close>_\<^latex>\<open>}\<close>\<close> [0] 999)
 | 
| 78471 | 22 | |
| 15469 | 23 | (* LOGIC *) | 
| 21210 | 24 | notation (latex output) | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 25 |   If  (\<open>(\<^latex>\<open>\textsf{\<close>if\<^latex>\<open>}\<close> (_)/ \<^latex>\<open>\textsf{\<close>then\<^latex>\<open>}\<close> (_)/ \<^latex>\<open>\textsf{\<close>else\<^latex>\<open>}\<close> (_))\<close> 10)
 | 
| 19674 | 26 | |
| 15469 | 27 | syntax (latex output) | 
| 28 | ||
| 29 | "_Let" :: "[letbinds, 'a] => 'a" | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 30 |   (\<open>(\<^latex>\<open>\textsf{\<close>let\<^latex>\<open>}\<close> (_)/ \<^latex>\<open>\textsf{\<close>in\<^latex>\<open>}\<close> (_))\<close> 10)
 | 
| 15469 | 31 | |
| 32 | "_case_syntax":: "['a, cases_syn] => 'b" | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 33 |   (\<open>(\<^latex>\<open>\textsf{\<close>case\<^latex>\<open>}\<close> _ \<^latex>\<open>\textsf{\<close>of\<^latex>\<open>}\<close>/ _)\<close> 10)
 | 
| 15469 | 34 | |
| 16110 | 35 | |
| 15469 | 36 | (* SETS *) | 
| 37 | ||
| 38 | (* empty set *) | |
| 27688 | 39 | notation (latex) | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 40 | "Set.empty" (\<open>\<emptyset>\<close>) | 
| 15469 | 41 | |
| 42 | (* insert *) | |
| 43 | translations | |
| 31462 | 44 |   "{x} \<union> A" <= "CONST insert x A"
 | 
| 15690 | 45 |   "{x,y}" <= "{x} \<union> {y}"
 | 
| 15469 | 46 |   "{x,y} \<union> A" <= "{x} \<union> ({y} \<union> A)"
 | 
| 27688 | 47 |   "{x}" <= "{x} \<union> \<emptyset>"
 | 
| 15469 | 48 | |
| 49 | (* set comprehension *) | |
| 50 | syntax (latex output) | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 51 |   "_Collect" :: "pttrn => bool => 'a set"              (\<open>(1{_ | _})\<close>)
 | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 52 |   "_CollectIn" :: "pttrn => 'a set => bool => 'a set"   (\<open>(1{_ \<in> _ | _})\<close>)
 | 
| 15469 | 53 | translations | 
| 54 |   "_Collect p P"      <= "{p. P}"
 | |
| 15690 | 55 |   "_Collect p P"      <= "{p|xs. P}"
 | 
| 41757 | 56 |   "_CollectIn p A P"  <= "{p : A. P}"
 | 
| 15469 | 57 | |
| 56977 | 58 | (* card *) | 
| 59 | notation (latex output) | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 60 | card (\<open>|_|\<close>) | 
| 56977 | 61 | |
| 15469 | 62 | (* LISTS *) | 
| 63 | ||
| 64 | (* Cons *) | |
| 21210 | 65 | notation (latex) | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 66 | Cons (\<open>_ \<cdot>/ _\<close> [66,65] 65) | 
| 15469 | 67 | |
| 68 | (* length *) | |
| 21210 | 69 | notation (latex output) | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 70 | length (\<open>|_|\<close>) | 
| 15469 | 71 | |
| 72 | (* nth *) | |
| 21210 | 73 | notation (latex output) | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 74 |   nth  (\<open>_\<^latex>\<open>\ensuremath{_{[\mathit{\<close>_\<^latex>\<open>}]}}\<close>\<close> [1000,0] 1000)
 | 
| 15469 | 75 | |
| 76 | (* DUMMY *) | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 77 | consts DUMMY :: 'a (\<open>\<^latex>\<open>\_\<close>\<close>) | 
| 15469 | 78 | |
| 79 | (* THEOREMS *) | |
| 35251 
e244adbbc28f
modernized notation -- to make it work for authentic syntax;
 wenzelm parents: 
31462diff
changeset | 80 | notation (Rule output) | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 81 |   Pure.imp  (\<open>\<^latex>\<open>\mbox{}\inferrule{\mbox{\<close>_\<^latex>\<open>}}\<close>\<^latex>\<open>{\mbox{\<close>_\<^latex>\<open>}}\<close>\<close>)
 | 
| 35251 
e244adbbc28f
modernized notation -- to make it work for authentic syntax;
 wenzelm parents: 
31462diff
changeset | 82 | |
| 15469 | 83 | syntax (Rule output) | 
| 84 | "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop" | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 85 |   (\<open>\<^latex>\<open>\mbox{}\inferrule{\<close>_\<^latex>\<open>}\<close>\<^latex>\<open>{\mbox{\<close>_\<^latex>\<open>}}\<close>\<close>)
 | 
| 15469 | 86 | |
| 87 | "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 88 |   (\<open>\<^latex>\<open>\mbox{\<close>_\<^latex>\<open>}\\\<close>/ _\<close>)
 | 
| 15469 | 89 | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 90 |   "_asm" :: "prop \<Rightarrow> asms" (\<open>\<^latex>\<open>\mbox{\<close>_\<^latex>\<open>}\<close>\<close>)
 | 
| 15469 | 91 | |
| 35251 
e244adbbc28f
modernized notation -- to make it work for authentic syntax;
 wenzelm parents: 
31462diff
changeset | 92 | notation (Axiom output) | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 93 |   "Trueprop"  (\<open>\<^latex>\<open>\mbox{}\inferrule{\mbox{}}{\mbox{\<close>_\<^latex>\<open>}}\<close>\<close>)
 | 
| 22328 
cc403d881873
added print-mode Axiom to print theorems without premises with a rule on top.
 schirmer parents: 
21210diff
changeset | 94 | |
| 35251 
e244adbbc28f
modernized notation -- to make it work for authentic syntax;
 wenzelm parents: 
31462diff
changeset | 95 | notation (IfThen output) | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 96 |   Pure.imp  (\<open>\<^latex>\<open>{\normalsize{}\<close>If\<^latex>\<open>\,}\<close> _/ \<^latex>\<open>{\normalsize \,\<close>then\<^latex>\<open>\,}\<close>/ _.\<close>)
 | 
| 15469 | 97 | syntax (IfThen output) | 
| 98 | "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop" | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 99 |   (\<open>\<^latex>\<open>{\normalsize{}\<close>If\<^latex>\<open>\,}\<close> _ /\<^latex>\<open>{\normalsize \,\<close>then\<^latex>\<open>\,}\<close>/ _.\<close>)
 | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 100 |   "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" (\<open>\<^latex>\<open>\mbox{\<close>_\<^latex>\<open>}\<close> /\<^latex>\<open>{\normalsize \,\<close>and\<^latex>\<open>\,}\<close>/ _\<close>)
 | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 101 |   "_asm" :: "prop \<Rightarrow> asms" (\<open>\<^latex>\<open>\mbox{\<close>_\<^latex>\<open>}\<close>\<close>)
 | 
| 15469 | 102 | |
| 35251 
e244adbbc28f
modernized notation -- to make it work for authentic syntax;
 wenzelm parents: 
31462diff
changeset | 103 | notation (IfThenNoBox output) | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 104 |   Pure.imp  (\<open>\<^latex>\<open>{\normalsize{}\<close>If\<^latex>\<open>\,}\<close> _/ \<^latex>\<open>{\normalsize \,\<close>then\<^latex>\<open>\,}\<close>/ _.\<close>)
 | 
| 15469 | 105 | syntax (IfThenNoBox output) | 
| 106 | "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop" | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 107 |   (\<open>\<^latex>\<open>{\normalsize{}\<close>If\<^latex>\<open>\,}\<close> _ /\<^latex>\<open>{\normalsize \,\<close>then\<^latex>\<open>\,}\<close>/ _.\<close>)
 | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 108 |   "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" (\<open>_ /\<^latex>\<open>{\normalsize \,\<close>and\<^latex>\<open>\,}\<close>/ _\<close>)
 | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
78471diff
changeset | 109 | "_asm" :: "prop \<Rightarrow> asms" (\<open>_\<close>) | 
| 15469 | 110 | |
| 67463 | 111 | setup \<open> | 
| 73761 | 112 | Document_Output.antiquotation_pretty_source_embedded \<^binding>\<open>const_typ\<close> | 
| 74563 | 113 | (Scan.lift Parse.embedded_inner_syntax) | 
| 67463 | 114 | (fn ctxt => fn c => | 
| 69081 
0b403ce1e8f8
const_typ also works for fixed variables - useful primarily for locales
 nipkow parents: 
67505diff
changeset | 115 |       let val tc = Proof_Context.read_const {proper = false, strict = false} ctxt c in
 | 
| 73761 | 116 | Pretty.block [Document_Output.pretty_term ctxt tc, Pretty.str " ::", | 
| 67505 
ceb324e34c14
clarified signature: items with \isasep are special;
 wenzelm parents: 
67463diff
changeset | 117 | Pretty.brk 1, Syntax.pretty_typ ctxt (fastype_of tc)] | 
| 67463 | 118 | end) | 
| 60500 | 119 | \<close> | 
| 49628 | 120 | |
| 63414 | 121 | setup\<open> | 
| 122 | let | |
| 123 | fun dummy_pats (wrap $ (eq $ lhs $ rhs)) = | |
| 124 | let | |
| 125 | val rhs_vars = Term.add_vars rhs []; | |
| 126 | fun dummy (v as Var (ixn as (_, T))) = | |
| 69593 | 127 | if member ((=) ) rhs_vars ixn then v else Const (\<^const_name>\<open>DUMMY\<close>, T) | 
| 63414 | 128 | | dummy (t $ u) = dummy t $ dummy u | 
| 129 | | dummy (Abs (n, T, b)) = Abs (n, T, dummy b) | |
| 130 | | dummy t = t; | |
| 131 | in wrap $ (eq $ dummy lhs $ rhs) end | |
| 132 | in | |
| 69593 | 133 | Term_Style.setup \<^binding>\<open>dummy_pats\<close> (Scan.succeed (K dummy_pats)) | 
| 63414 | 134 | end | 
| 135 | \<close> | |
| 136 | ||
| 66527 | 137 | setup \<open> | 
| 138 | let | |
| 139 | ||
| 140 | fun eta_expand Ts t xs = case t of | |
| 141 | Abs(x,T,t) => | |
| 142 | let val (t', xs') = eta_expand (T::Ts) t xs | |
| 143 | in (Abs (x, T, t'), xs') end | |
| 144 | | _ => | |
| 145 | let | |
| 146 | val (a,ts) = strip_comb t (* assume a atomic *) | |
| 147 | val (ts',xs') = fold_map (eta_expand Ts) ts xs | |
| 148 | val t' = list_comb (a, ts'); | |
| 149 | val Bs = binder_types (fastype_of1 (Ts,t)); | |
| 150 | val n = Int.min (length Bs, length xs'); | |
| 151 | val bs = map Bound ((n - 1) downto 0); | |
| 152 | val xBs = ListPair.zip (xs',Bs); | |
| 153 | val xs'' = drop n xs'; | |
| 154 | val t'' = fold_rev Term.abs xBs (list_comb(t', bs)) | |
| 155 | in (t'', xs'') end | |
| 156 | ||
| 157 | val style_eta_expand = | |
| 158 | (Scan.repeat Args.name) >> (fn xs => fn ctxt => fn t => fst (eta_expand [] t xs)) | |
| 159 | ||
| 69593 | 160 | in Term_Style.setup \<^binding>\<open>eta_expand\<close> style_eta_expand end | 
| 66527 | 161 | \<close> | 
| 162 | ||
| 15469 | 163 | end | 
| 55077 | 164 | (*>*) |