src/Pure/ML/ml_pp.ML
author wenzelm
Tue, 21 Apr 2020 22:19:59 +0200
changeset 71777 3875815f5967
parent 70398 725438ceae7c
child 78649 d46006355819
permissions -rw-r--r--
clarified signature: avoid clash with Isabelle/Scala Term.OFCLASS on case-insensible file-system;
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 _ =
62819
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    11
  ML_system_pp (fn _ => fn _ => Pretty.to_polyml o Proof_Display.pp_context);
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    12
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    13
val _ =
62819
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    14
  ML_system_pp (fn depth => fn _ =>
67380
8bef51521f21 clarified implicit Pure.thy;
wenzelm
parents: 62819
diff changeset
    15
    ML_Pretty.to_polyml o Pretty.to_ML depth o Proof_Display.pp_thm Theory.get_pure);
62356
wenzelm
parents: 52426
diff changeset
    16
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    17
val _ =
62819
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    18
  ML_system_pp (fn depth => fn _ =>
67380
8bef51521f21 clarified implicit Pure.thy;
wenzelm
parents: 62819
diff changeset
    19
    ML_Pretty.to_polyml o Pretty.to_ML depth o Proof_Display.pp_cterm Theory.get_pure);
62356
wenzelm
parents: 52426
diff changeset
    20
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    21
val _ =
62819
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    22
  ML_system_pp (fn depth => fn _ =>
67380
8bef51521f21 clarified implicit Pure.thy;
wenzelm
parents: 62819
diff changeset
    23
    ML_Pretty.to_polyml o Pretty.to_ML depth o Proof_Display.pp_ctyp Theory.get_pure);
33767
f962c761a38f toplevel pretty printer for Synchronized.var;
wenzelm
parents: 33545
diff changeset
    24
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    25
val _ =
62819
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    26
  ML_system_pp (fn depth => fn _ =>
67380
8bef51521f21 clarified implicit Pure.thy;
wenzelm
parents: 62819
diff changeset
    27
    ML_Pretty.to_polyml o Pretty.to_ML depth o Proof_Display.pp_typ Theory.get_pure);
30633
cc18ae3c1c7f extra toplevel pretty-printing for Poly/ML; experimental version for Poly/ML 5.3;
wenzelm
parents:
diff changeset
    28
52426
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    29
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    30
local
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    31
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    32
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
    33
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
    34
62387
ad3eb2889f9a support for polyml-git ec49a49972c5 (branch FixedPrecisionInt);
wenzelm
parents: 62356
diff changeset
    35
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
    36
  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
    37
  else
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    38
    (case t of
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    39
      _ $ _ =>
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    40
        op :: (strip_comb t)
62387
ad3eb2889f9a support for polyml-git ec49a49972c5 (branch FixedPrecisionInt);
wenzelm
parents: 62356
diff changeset
    41
        |> 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
    42
        |> Pretty.separate " $"
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    43
        |> (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
    44
    | 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
    45
        prt_apps "Abs"
62819
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    46
         [Pretty.from_polyml (ML_system_pretty (a, dp - 1)),
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    47
          Pretty.from_polyml (ML_system_pretty (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
    48
          prt_term false (dp - 3) b]
62819
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    49
    | Const a => prt_app "Const" (Pretty.from_polyml (ML_system_pretty (a, dp - 1)))
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    50
    | Free a => prt_app "Free" (Pretty.from_polyml (ML_system_pretty (a, dp - 1)))
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    51
    | Var a => prt_app "Var" (Pretty.from_polyml (ML_system_pretty (a, dp - 1)))
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    52
    | Bound a => prt_app "Bound" (Pretty.from_polyml (ML_system_pretty (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
    53
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    54
in
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    55
62819
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    56
val _ = ML_system_pp (fn depth => fn _ => Pretty.to_polyml o prt_term false depth);
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    57
52426
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    58
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    59
local
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    60
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    61
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
    62
  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
    63
  else
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    64
    (case prf of
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    65
      _ % _ => 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
    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
    | 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
    68
        prt_apps "Abst"
62819
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    69
         [Pretty.from_polyml (ML_system_pretty (a, dp - 1)),
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    70
          Pretty.from_polyml (ML_system_pretty (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
    71
          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
    72
    | 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
    73
        prt_apps "AbsP"
62819
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    74
         [Pretty.from_polyml (ML_system_pretty (a, dp - 1)),
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    75
          Pretty.from_polyml (ML_system_pretty (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
    76
          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
    77
    | 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
    78
    | MinProof => Pretty.str "MinProof"
62819
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    79
    | PBound a => prt_app "PBound" (Pretty.from_polyml (ML_system_pretty (a, dp - 1)))
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    80
    | PAxm a => prt_app "PAxm" (Pretty.from_polyml (ML_system_pretty (a, dp - 1)))
71777
3875815f5967 clarified signature: avoid clash with Isabelle/Scala Term.OFCLASS on case-insensible file-system;
wenzelm
parents: 70398
diff changeset
    81
    | PClass a => prt_app "PClass" (Pretty.from_polyml (ML_system_pretty (a, dp - 1)))
62819
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    82
    | Oracle a => prt_app "Oracle" (Pretty.from_polyml (ML_system_pretty (a, dp - 1)))
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    83
    | PThm a => prt_app "PThm" (Pretty.from_polyml (ML_system_pretty (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
    84
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    85
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
    86
  let
52426
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    87
    val (head, args) = strip_proof prf [];
62387
ad3eb2889f9a support for polyml-git ec49a49972c5 (branch FixedPrecisionInt);
wenzelm
parents: 62356
diff changeset
    88
    val prts =
ad3eb2889f9a support for polyml-git ec49a49972c5 (branch FixedPrecisionInt);
wenzelm
parents: 62356
diff changeset
    89
      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
    90
  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
    91
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
    92
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
    93
      strip_proof p
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    94
        ((fn d =>
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
    95
          [Pretty.str " %", Pretty.brk 1,
62819
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
    96
            Pretty.from_polyml (ML_system_pretty (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
    97
  | 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
    98
      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
    99
  | 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
   100
52426
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
   101
in
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
   102
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
   103
val _ =
62819
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62673
diff changeset
   104
  ML_system_pp (fn depth => fn _ => fn prf =>
62667
254582abf067 clarified Pretty.T toplevel pp;
wenzelm
parents: 62665
diff changeset
   105
    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
   106
81e27230a8b7 ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
wenzelm
parents: 52425
diff changeset
   107
end;
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;
62665
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
   110
a78ce0c6e191 clarified modules;
wenzelm
parents: 62663
diff changeset
   111
end;