| author | wenzelm | 
| Tue, 09 May 2017 21:25:32 +0200 | |
| changeset 65794 | a880f41a8d0f | 
| parent 62819 | d3ff367a16a0 | 
| child 67380 | 8bef51521f21 | 
| permissions | -rw-r--r-- | 
| 62665 | 1  | 
(* Title: Pure/ML/ml_pp.ML  | 
| 38327 | 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 | 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 | 7  | 
structure ML_PP: sig end =  | 
8  | 
struct  | 
|
| 62356 | 9  | 
|
| 62665 | 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 | 12  | 
|
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 _ =>  | 
| 62673 | 15  | 
ML_Pretty.to_polyml o Pretty.to_ML depth o Proof_Display.pp_thm Thy_Info.pure_theory);  | 
| 62356 | 16  | 
|
| 62665 | 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 _ =>  | 
| 62673 | 19  | 
ML_Pretty.to_polyml o Pretty.to_ML depth o Proof_Display.pp_cterm Thy_Info.pure_theory);  | 
| 62356 | 20  | 
|
| 62665 | 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 _ =>  | 
| 62673 | 23  | 
ML_Pretty.to_polyml o Pretty.to_ML depth o Proof_Display.pp_ctyp Thy_Info.pure_theory);  | 
| 33767 | 24  | 
|
| 62665 | 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 _ =>  | 
| 62673 | 27  | 
ML_Pretty.to_polyml o Pretty.to_ML depth 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
 | 
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 | 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)))  | 
| 
 
d3ff367a16a0
careful export of type-dependent functions, without losing their special status;
 
wenzelm 
parents: 
62673 
diff
changeset
 | 
81  | 
| OfClass a => prt_app "OfClass" (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
 | 
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  | 
| Promise a => prt_app "Promise" (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
 | 
84  | 
| 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
 | 
85  | 
|
| 
 
81e27230a8b7
ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
 
wenzelm 
parents: 
52425 
diff
changeset
 | 
86  | 
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
 | 
87  | 
let  | 
| 
52426
 
81e27230a8b7
ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
 
wenzelm 
parents: 
52425 
diff
changeset
 | 
88  | 
val (head, args) = strip_proof prf [];  | 
| 
62387
 
ad3eb2889f9a
support for polyml-git ec49a49972c5 (branch FixedPrecisionInt);
 
wenzelm 
parents: 
62356 
diff
changeset
 | 
89  | 
val prts =  | 
| 
 
ad3eb2889f9a
support for polyml-git ec49a49972c5 (branch FixedPrecisionInt);
 
wenzelm 
parents: 
62356 
diff
changeset
 | 
90  | 
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
 | 
91  | 
  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
 | 
92  | 
|
| 
 
81e27230a8b7
ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
 
wenzelm 
parents: 
52425 
diff
changeset
 | 
93  | 
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
 | 
94  | 
strip_proof p  | 
| 62665 | 95  | 
((fn d =>  | 
96  | 
[Pretty.str " %", Pretty.brk 1,  | 
|
| 
62819
 
d3ff367a16a0
careful export of type-dependent functions, without losing their special status;
 
wenzelm 
parents: 
62673 
diff
changeset
 | 
97  | 
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
 | 
98  | 
| 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
 | 
99  | 
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
 | 
100  | 
| 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
 | 
101  | 
|
| 
52426
 
81e27230a8b7
ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
 
wenzelm 
parents: 
52425 
diff
changeset
 | 
102  | 
in  | 
| 
 
81e27230a8b7
ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
 
wenzelm 
parents: 
52425 
diff
changeset
 | 
103  | 
|
| 
 
81e27230a8b7
ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
 
wenzelm 
parents: 
52425 
diff
changeset
 | 
104  | 
val _ =  | 
| 
62819
 
d3ff367a16a0
careful export of type-dependent functions, without losing their special status;
 
wenzelm 
parents: 
62673 
diff
changeset
 | 
105  | 
ML_system_pp (fn depth => fn _ => fn prf =>  | 
| 62667 | 106  | 
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
 | 
107  | 
|
| 
 
81e27230a8b7
ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
 
wenzelm 
parents: 
52425 
diff
changeset
 | 
108  | 
end;  | 
| 
 
81e27230a8b7
ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
 
wenzelm 
parents: 
52425 
diff
changeset
 | 
109  | 
|
| 
 
81e27230a8b7
ML pretty printer for proof terms, with specifica treatment of infix application (similar to terms);
 
wenzelm 
parents: 
52425 
diff
changeset
 | 
110  | 
end;  | 
| 62665 | 111  | 
|
112  | 
end;  |