src/Pure/ML/ml_pp.ML
author wenzelm
Fri, 18 Mar 2016 17:51:57 +0100
changeset 62667 254582abf067
parent 62665 a78ce0c6e191
child 62673 b5c57430b9dd
permissions -rw-r--r--
clarified Pretty.T toplevel pp;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
     1
(*  Title:      Pure/ML/ml_pp.ML
38327
d6afb77b0f6d more precise and more maintainable dependencies;
wenzelm
parents: 33767
diff changeset
     2
    Author:     Makarius
30633
cc18ae3c1c7f extra toplevel pretty-printing for Poly/ML; experimental version for Poly/ML 5.3;
wenzelm
parents:
diff changeset
     3
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
     4
ML toplevel pretty-printing for logical entities.
30633
cc18ae3c1c7f extra toplevel pretty-printing for Poly/ML; experimental version for Poly/ML 5.3;
wenzelm
parents:
diff changeset
     5
*)
cc18ae3c1c7f extra toplevel pretty-printing for Poly/ML; experimental version for Poly/ML 5.3;
wenzelm
parents:
diff changeset
     6
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
     7
structure ML_PP: sig end =
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
     8
struct
62356
wenzelm
parents: 52426
diff changeset
     9
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    10
val _ =
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    11
  PolyML.addPrettyPrinter
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    12
    (fn _ => fn _ => Pretty.to_polyml o Proof_Display.pp_context);
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    13
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    14
val _ =
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    15
  PolyML.addPrettyPrinter
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    16
    (fn _ => fn _ => Pretty.to_polyml o Proof_Display.pp_thm Thy_Info.pure_theory);
62356
wenzelm
parents: 52426
diff changeset
    17
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    18
val _ =
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    19
  PolyML.addPrettyPrinter
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    20
    (fn _ => fn _ => Pretty.to_polyml o Proof_Display.pp_cterm Thy_Info.pure_theory);
62356
wenzelm
parents: 52426
diff changeset
    21
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    22
val _ =
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    23
  PolyML.addPrettyPrinter
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    24
    (fn _ => fn _ => Pretty.to_polyml o Proof_Display.pp_ctyp Thy_Info.pure_theory);
33767
f962c761a38f toplevel pretty printer for Synchronized.var;
wenzelm
parents: 33545
diff changeset
    25
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    26
val _ =
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    27
  PolyML.addPrettyPrinter
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    28
    (fn _ => fn _ => Pretty.to_polyml o Proof_Display.pp_typ Thy_Info.pure_theory);
30633
cc18ae3c1c7f extra toplevel pretty-printing for Poly/ML; experimental version for Poly/ML 5.3;
wenzelm
parents:
diff changeset
    29
52426
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    30
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    31
local
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    32
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    33
fun prt_app name prt = Pretty.block [Pretty.str (name ^ " "), prt];
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    34
fun prt_apps name = Pretty.enum "," (name ^ " (") ")";
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    35
62387
ad3eb2889f9a support for polyml-git ec49a49972c5 (branch FixedPrecisionInt);
wenzelm
parents: 62356
diff changeset
    36
fun prt_term parens (dp: FixedInt.int) t =
52426
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    37
  if dp <= 0 then Pretty.str "..."
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    38
  else
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    39
    (case t of
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    40
      _ $ _ =>
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    41
        op :: (strip_comb t)
62387
ad3eb2889f9a support for polyml-git ec49a49972c5 (branch FixedPrecisionInt);
wenzelm
parents: 62356
diff changeset
    42
        |> map_index (fn (i, u) => prt_term true (dp - FixedInt.fromInt i - 1) u)
52426
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    43
        |> Pretty.separate " $"
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    44
        |> (if parens then Pretty.enclose "(" ")" else Pretty.block)
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    45
    | Abs (a, T, b) =>
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    46
        prt_apps "Abs"
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    47
         [Pretty.from_polyml (PolyML.prettyRepresentation (a, dp - 1)),
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    48
          Pretty.from_polyml (PolyML.prettyRepresentation (T, dp - 2)),
52426
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    49
          prt_term false (dp - 3) b]
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    50
    | Const a => prt_app "Const" (Pretty.from_polyml (PolyML.prettyRepresentation (a, dp - 1)))
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    51
    | Free a => prt_app "Free" (Pretty.from_polyml (PolyML.prettyRepresentation (a, dp - 1)))
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    52
    | Var a => prt_app "Var" (Pretty.from_polyml (PolyML.prettyRepresentation (a, dp - 1)))
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    53
    | Bound a => prt_app "Bound" (Pretty.from_polyml (PolyML.prettyRepresentation (a, dp - 1))));
52426
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    54
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    55
in
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    56
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    57
val _ = PolyML.addPrettyPrinter (fn depth => fn _ => Pretty.to_polyml o prt_term false depth);
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    58
52426
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    59
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    60
local
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    61
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    62
fun prt_proof parens dp prf =
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    63
  if dp <= 0 then Pretty.str "..."
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    64
  else
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    65
    (case prf of
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    66
      _ % _ => prt_proofs parens dp prf
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    67
    | _ %% _ => prt_proofs parens dp prf
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    68
    | Abst (a, T, b) =>
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    69
        prt_apps "Abst"
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    70
         [Pretty.from_polyml (PolyML.prettyRepresentation (a, dp - 1)),
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    71
          Pretty.from_polyml (PolyML.prettyRepresentation (T, dp - 2)),
52426
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    72
          prt_proof false (dp - 3) b]
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    73
    | AbsP (a, t, b) =>
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    74
        prt_apps "AbsP"
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    75
         [Pretty.from_polyml (PolyML.prettyRepresentation (a, dp - 1)),
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    76
          Pretty.from_polyml (PolyML.prettyRepresentation (t, dp - 2)),
52426
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    77
          prt_proof false (dp - 3) b]
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    78
    | Hyp t => prt_app "Hyp" (prt_term true (dp - 1) t)
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    79
    | MinProof => Pretty.str "MinProof"
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    80
    | PBound a => prt_app "PBound" (Pretty.from_polyml (PolyML.prettyRepresentation (a, dp - 1)))
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    81
    | PAxm a => prt_app "PAxm" (Pretty.from_polyml (PolyML.prettyRepresentation (a, dp - 1)))
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    82
    | OfClass a => prt_app "OfClass" (Pretty.from_polyml (PolyML.prettyRepresentation (a, dp - 1)))
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    83
    | Oracle a => prt_app "Oracle" (Pretty.from_polyml (PolyML.prettyRepresentation (a, dp - 1)))
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    84
    | Promise a => prt_app "Promise" (Pretty.from_polyml (PolyML.prettyRepresentation (a, dp - 1)))
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    85
    | PThm a => prt_app "PThm" (Pretty.from_polyml (PolyML.prettyRepresentation (a, dp - 1))))
52426
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    86
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    87
and prt_proofs parens dp prf =
33545
d8903f0002e5 home-grown pretty printer for term -- Poly/ML 5.3.0 does not observe infix status of constructors (notably $);
wenzelm
parents: 33538
diff changeset
    88
  let
52426
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    89
    val (head, args) = strip_proof prf [];
62387
ad3eb2889f9a support for polyml-git ec49a49972c5 (branch FixedPrecisionInt);
wenzelm
parents: 62356
diff changeset
    90
    val prts =
ad3eb2889f9a support for polyml-git ec49a49972c5 (branch FixedPrecisionInt);
wenzelm
parents: 62356
diff changeset
    91
      head (dp - 1) :: flat (map_index (fn (i, prt) => prt (dp - FixedInt.fromInt i - 2)) args);
52426
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    92
  in if parens then Pretty.enclose "(" ")" prts else Pretty.block prts end
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    93
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    94
and strip_proof (p % t) res =
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    95
      strip_proof p
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    96
        ((fn d =>
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    97
          [Pretty.str " %", Pretty.brk 1,
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    98
            Pretty.from_polyml (PolyML.prettyRepresentation (t, d))]) :: res)
52426
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    99
  | strip_proof (p %% q) res =
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
   100
      strip_proof p ((fn d => [Pretty.str " %%", Pretty.brk 1, prt_proof true d q]) :: res)
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
   101
  | strip_proof p res = (fn d => prt_proof true d p, res);
33545
d8903f0002e5 home-grown pretty printer for term -- Poly/ML 5.3.0 does not observe infix status of constructors (notably $);
wenzelm
parents: 33538
diff changeset
   102
52426
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
   103
in
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
   104
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
   105
val _ =
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
   106
  PolyML.addPrettyPrinter (fn depth => fn _ => fn prf =>
62667
254582abf067 clarified Pretty.T toplevel pp;
wenzelm
parents: 62665
diff changeset
   107
    ML_Pretty.to_polyml (Pretty.to_ML ~1 (prt_proof false depth prf)));
52426
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
   108
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
   109
end;
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
   110
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
   111
end;
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
   112
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
   113
end;