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