src/Pure/Thy/term_style.ML
author wenzelm
Wed, 06 Dec 2017 18:59:33 +0100
changeset 67147 dea94b1aabc3
parent 62969 9f394a16c557
child 74561 8e6c973003c8
permissions -rw-r--r--
prefer control symbol antiquotations;
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
56030
ef2ffd622264 modernized data managed via Name_Space.table;
wenzelm
parents: 56029
diff changeset
     9
  val setup: binding -> (Proof.context -> term -> term) parser -> theory -> theory
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
    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
56030
ef2ffd622264 modernized data managed via Name_Space.table;
wenzelm
parents: 56029
diff changeset
    16
(* theory data *)
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    17
56030
ef2ffd622264 modernized data managed via Name_Space.table;
wenzelm
parents: 56029
diff changeset
    18
structure Data = Theory_Data
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22107
diff changeset
    19
(
56030
ef2ffd622264 modernized data managed via Name_Space.table;
wenzelm
parents: 56029
diff changeset
    20
  type T = (Proof.context -> term -> term) parser Name_Space.table;
ef2ffd622264 modernized data managed via Name_Space.table;
wenzelm
parents: 56029
diff changeset
    21
  val empty : T = Name_Space.empty_table "antiquotation_style";
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    22
  val extend = I;
56030
ef2ffd622264 modernized data managed via Name_Space.table;
wenzelm
parents: 56029
diff changeset
    23
  fun merge data : T = Name_Space.merge_tables data;
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22107
diff changeset
    24
);
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    25
56030
ef2ffd622264 modernized data managed via Name_Space.table;
wenzelm
parents: 56029
diff changeset
    26
val get_data = Data.get o Proof_Context.theory_of;
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    27
56030
ef2ffd622264 modernized data managed via Name_Space.table;
wenzelm
parents: 56029
diff changeset
    28
fun setup binding style thy =
ef2ffd622264 modernized data managed via Name_Space.table;
wenzelm
parents: 56029
diff changeset
    29
  Data.map (#2 o Name_Space.define (Context.Theory thy) true (binding, style)) thy;
22107
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
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
    32
(* 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
    33
56030
ef2ffd622264 modernized data managed via Name_Space.table;
wenzelm
parents: 56029
diff changeset
    34
fun parse_single ctxt =
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 61814
diff changeset
    35
  Parse.token Parse.name ::: Parse.args >> (fn src0 =>
56030
ef2ffd622264 modernized data managed via Name_Space.table;
wenzelm
parents: 56029
diff changeset
    36
    let
61814
1ca1142e1711 clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents: 61470
diff changeset
    37
      val (src, parse) = Token.check_src ctxt get_data src0;
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 56204
diff changeset
    38
      val (f, _) = Token.syntax (Scan.lift parse) src ctxt;
56030
ef2ffd622264 modernized data managed via Name_Space.table;
wenzelm
parents: 56029
diff changeset
    39
    in f ctxt end);
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
    40
50637
81d6fe75ea5b tuned whitespace;
wenzelm
parents: 50592
diff changeset
    41
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
    42
  (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
    43
      >> 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
    44
  || 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
    45
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
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    47
(* predefined styles *)
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    48
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    49
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
    50
  let
59970
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 58011
diff changeset
    51
    val concl = Object_Logic.drop_judgment ctxt (Logic.strip_imp_concl t);
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    52
  in
50637
81d6fe75ea5b tuned whitespace;
wenzelm
parents: 50592
diff changeset
    53
    (case concl of
56030
ef2ffd622264 modernized data managed via Name_Space.table;
wenzelm
parents: 56029
diff changeset
    54
      _ $ l $ r => proj (l, r)
50637
81d6fe75ea5b tuned whitespace;
wenzelm
parents: 50592
diff changeset
    55
    | _ => 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
    56
  end);
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    57
50638
wenzelm
parents: 50637
diff changeset
    58
val style_prem = Parse.nat >> (fn i => fn ctxt => fn t =>
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    59
  let
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    60
    val prems = Logic.strip_imp_prems t;
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    61
  in
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    62
    if i <= length prems then nth prems (i - 1)
50637
81d6fe75ea5b tuned whitespace;
wenzelm
parents: 50592
diff changeset
    63
    else
81d6fe75ea5b tuned whitespace;
wenzelm
parents: 50592
diff changeset
    64
      error ("Not enough premises for prem " ^ string_of_int i ^
81d6fe75ea5b tuned whitespace;
wenzelm
parents: 50592
diff changeset
    65
        " in propositon: " ^ Syntax.string_of_term ctxt t)
32891
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    66
  end);
d403b99287ff new generalized concept for term styles
haftmann
parents: 32890
diff changeset
    67
53021
d0fa3f446b9d discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents: 50638
diff changeset
    68
fun sub_symbols (d :: s :: ss) =
61470
c42960228a81 clarified Symbol.is_control;
wenzelm
parents: 59970
diff changeset
    69
      if Symbol.is_ascii_digit d andalso not (Symbol.is_control s)
56203
76c72f4d0667 clarified bootstrap process: switch to ML with context and antiquotations earlier;
wenzelm
parents: 56201
diff changeset
    70
      then d :: "\<^sub>" :: sub_symbols (s :: ss)
41685
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    71
      else d :: s :: ss
53021
d0fa3f446b9d discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents: 50638
diff changeset
    72
  | sub_symbols cs = cs;
41685
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    73
53021
d0fa3f446b9d discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents: 50638
diff changeset
    74
val sub_name = implode o rev o sub_symbols o rev o Symbol.explode;
41685
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    75
53021
d0fa3f446b9d discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents: 50638
diff changeset
    76
fun sub_term (Free (n, T)) = Free (sub_name n, T)
d0fa3f446b9d discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents: 50638
diff changeset
    77
  | sub_term (Var ((n, idx), T)) =
d0fa3f446b9d discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents: 50638
diff changeset
    78
      if idx <> 0 then Var ((sub_name (n ^ string_of_int idx), 0), T)
d0fa3f446b9d discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents: 50638
diff changeset
    79
      else Var ((sub_name n, 0), T)
d0fa3f446b9d discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents: 50638
diff changeset
    80
  | sub_term (t $ u) = sub_term t $ sub_term u
d0fa3f446b9d discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents: 50638
diff changeset
    81
  | sub_term (Abs (n, T, b)) = Abs (sub_name n, T, sub_term b)
d0fa3f446b9d discontinued special treatment of \<^isub> and \<^isup> in rendering or editor front-end;
wenzelm
parents: 50638
diff changeset
    82
  | sub_term t = t;
41685
e29ea98a76ce term style 'isub': ad-hoc subscripting of variables that end with digits (x1, x23, ...)
krauss
parents: 36950
diff changeset
    83
53171
a5e54d4d9081 added Theory.setup convenience;
wenzelm
parents: 53021
diff changeset
    84
val _ = Theory.setup
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 62969
diff changeset
    85
 (setup \<^binding>\<open>lhs\<close> (style_lhs_rhs fst) #>
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 62969
diff changeset
    86
  setup \<^binding>\<open>rhs\<close> (style_lhs_rhs snd) #>
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 62969
diff changeset
    87
  setup \<^binding>\<open>prem\<close> style_prem #>
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 62969
diff changeset
    88
  setup \<^binding>\<open>concl\<close> (Scan.succeed (K Logic.strip_imp_concl)) #>
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 62969
diff changeset
    89
  setup \<^binding>\<open>sub\<close> (Scan.succeed (K sub_term)));
22107
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    90
926afa3361e1 renamed Isar/term_style.ML to Thy/term_style.ML;
wenzelm
parents:
diff changeset
    91
end;