src/Pure/Thy/term_style.ML
author wenzelm
Sun, 26 May 2013 19:27:32 +0200
changeset 52156 576aceb343dc
parent 50638 eedc01eca736
child 53021 d0fa3f446b9d
permissions -rw-r--r--
tuned signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Thy/term_style.ML
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
     3
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
     4
Styles for term printing.
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
     5
*)
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
     6
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
     7
signature TERM_STYLE =
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
     8
sig
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
     9
  val setup: string -> (Proof.context -> term -> term) parser -> theory -> theory
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    10
  val parse: (term -> term) context_parser
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    11
end;
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    12
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    13
structure Term_Style: TERM_STYLE =
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    14
struct
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    15
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    16
(* style data *)
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    17
23655
d2d1138e0ddc replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents: 23577
diff changeset
    18
fun err_dup_style name =
d2d1138e0ddc replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents: 23577
diff changeset
    19
  error ("Duplicate declaration of antiquote style: " ^ quote name);
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    20
42016
wenzelm
parents: 41686
diff changeset
    21
structure Styles = Theory_Data
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22107
diff changeset
    22
(
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    23
  type T = ((Proof.context -> term -> term) parser * stamp) Symtab.table;
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    24
  val empty = Symtab.empty;
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    25
  val extend = I;
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33184
diff changeset
    26
  fun merge data : T = Symtab.merge (eq_snd (op =)) data
23655
d2d1138e0ddc replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents: 23577
diff changeset
    27
    handle Symtab.DUP dup => err_dup_style dup;
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22107
diff changeset
    28
);
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    29
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    30
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    31
(* accessors *)
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    32
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    33
fun the_style thy name =
42016
wenzelm
parents: 41686
diff changeset
    34
  (case Symtab.lookup (Styles.get thy) name of
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    35
    NONE => error ("Unknown antiquote style: " ^ quote name)
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    36
  | SOME (style, _) => style);
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    37
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    38
fun setup name style thy =
42016
wenzelm
parents: 41686
diff changeset
    39
  Styles.map (Symtab.update_new (name, (style, stamp ()))) thy
23655
d2d1138e0ddc replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents: 23577
diff changeset
    40
    handle Symtab.DUP _ => err_dup_style name;
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    41
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    42
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    43
(* style parsing *)
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    44
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 35625
diff changeset
    45
fun parse_single ctxt = Parse.position (Parse.xname -- Args.parse)
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    46
  >> (fn x as ((name, _), _) => fst (Args.context_syntax "style"
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42016
diff changeset
    47
       (Scan.lift (the_style (Proof_Context.theory_of ctxt) name))
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    48
         (Args.src x) ctxt |>> (fn f => f ctxt)));
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    49
50637
81d6fe75ea5b tuned whitespace;
wenzelm
parents: 50592
diff changeset
    50
val parse = Args.context :|-- (fn ctxt => Scan.lift
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    51
  (Args.parens (parse_single ctxt ::: Scan.repeat (Args.$$$ "," |-- parse_single ctxt))
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    52
      >> fold I
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    53
  || Scan.succeed I));
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    54
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    55
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    56
(* predefined styles *)
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    57
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    58
fun style_lhs_rhs proj = Scan.succeed (fn ctxt => fn t =>
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    59
  let
35625
9c818cab0dd0 modernized structure Object_Logic;
wenzelm
parents: 33522
diff changeset
    60
    val concl =
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42016
diff changeset
    61
      Object_Logic.drop_judgment (Proof_Context.theory_of ctxt) (Logic.strip_imp_concl t)
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    62
  in
50637
81d6fe75ea5b tuned whitespace;
wenzelm
parents: 50592
diff changeset
    63
    (case concl of
81d6fe75ea5b tuned whitespace;
wenzelm
parents: 50592
diff changeset
    64
      (_ $ l $ r) => proj (l, r)
81d6fe75ea5b tuned whitespace;
wenzelm
parents: 50592
diff changeset
    65
    | _ => error ("Binary operator expected in term: " ^ Syntax.string_of_term ctxt concl))
32890
77df12652210 generalized term styles: transformations may depend on arguments; modernized term_style module; antiquotations thm, prop and term accepting term styles
haftmann
parents: 29606
diff changeset
    66
  end);
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    67
50638
wenzelm
parents: 50637
diff changeset
    68
val style_prem = Parse.nat >> (fn i => fn ctxt => fn t =>
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    69
  let
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    70
    val prems = Logic.strip_imp_prems t;
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    71
  in
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    72
    if i <= length prems then nth prems (i - 1)
50637
81d6fe75ea5b tuned whitespace;
wenzelm
parents: 50592
diff changeset
    73
    else
81d6fe75ea5b tuned whitespace;
wenzelm
parents: 50592
diff changeset
    74
      error ("Not enough premises for prem " ^ string_of_int i ^
81d6fe75ea5b tuned whitespace;
wenzelm
parents: 50592
diff changeset
    75
        " in propositon: " ^ Syntax.string_of_term ctxt t)
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    76
  end);
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    77
41685
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    78
fun isub_symbols (d :: s :: ss) =
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    79
      if Symbol.is_ascii_digit d andalso not (String.isPrefix ("\\<^") s)
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    80
      then d :: "\\<^isub>" :: isub_symbols (s :: ss)
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    81
      else d :: s :: ss
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    82
  | isub_symbols cs = cs;
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    83
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    84
val isub_name = implode o rev o isub_symbols o rev o Symbol.explode;
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    85
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    86
fun isub_term (Free (n, T)) = Free (isub_name n, T)
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    87
  | isub_term (Var ((n, idx), T)) =
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    88
      if idx <> 0 then Var ((isub_name (n ^ string_of_int idx), 0), T)
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    89
      else Var ((isub_name n, 0), T)
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    90
  | isub_term (t $ u) = isub_term t $ isub_term u
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    91
  | isub_term (Abs (n, T, b)) = Abs (isub_name n, T, isub_term b)
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    92
  | isub_term t = t;
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    93
26463
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26435
diff changeset
    94
val _ = Context.>> (Context.map_theory
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    95
 (setup "lhs" (style_lhs_rhs fst) #>
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    96
  setup "rhs" (style_lhs_rhs snd) #>
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    97
  setup "prem" style_prem #>
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    98
  setup "concl" (Scan.succeed (K Logic.strip_imp_concl)) #>
50592
a39250169636 discontinued legacy antiquotations and styles
haftmann
parents: 43277
diff changeset
    99
  setup "isub" (Scan.succeed (K isub_term))));
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
   100
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
   101
end;