src/Pure/display.ML
changeset 61270 28eb608b9b59
parent 61245 b77bf45efe21
parent 61269 64a5bce1f498
child 61271 0478ba10152a
equal deleted inserted replaced
61245:b77bf45efe21 61270:28eb608b9b59
     1 (*  Title:      Pure/display.ML
       
     2     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
       
     3     Author:     Makarius
       
     4 
       
     5 Printing of theorems, results etc.
       
     6 *)
       
     7 
       
     8 signature BASIC_DISPLAY =
       
     9 sig
       
    10   val show_consts: bool Config.T
       
    11   val show_hyps_raw: Config.raw
       
    12   val show_hyps: bool Config.T
       
    13   val show_tags_raw: Config.raw
       
    14   val show_tags: bool Config.T
       
    15 end;
       
    16 
       
    17 signature DISPLAY =
       
    18 sig
       
    19   include BASIC_DISPLAY
       
    20   val pretty_thm_raw: Proof.context -> {quote: bool, show_hyps: bool} -> thm -> Pretty.T
       
    21   val pretty_thm: Proof.context -> thm -> Pretty.T
       
    22   val pretty_thm_item: Proof.context -> thm -> Pretty.T
       
    23   val pretty_thm_global: theory -> thm -> Pretty.T
       
    24   val string_of_thm: Proof.context -> thm -> string
       
    25   val string_of_thm_global: theory -> thm -> string
       
    26   val pretty_full_theory: bool -> theory -> Pretty.T list
       
    27 end;
       
    28 
       
    29 structure Display: DISPLAY =
       
    30 struct
       
    31 
       
    32 (** options **)
       
    33 
       
    34 val show_consts = Goal_Display.show_consts;
       
    35 
       
    36 val show_hyps_raw = Config.declare ("show_hyps", @{here}) (fn _ => Config.Bool false);
       
    37 val show_hyps = Config.bool show_hyps_raw;
       
    38 
       
    39 val show_tags_raw = Config.declare ("show_tags", @{here}) (fn _ => Config.Bool false);
       
    40 val show_tags = Config.bool show_tags_raw;
       
    41 
       
    42 
       
    43 
       
    44 (** print thm **)
       
    45 
       
    46 fun pretty_tag (name, arg) = Pretty.strs [name, quote arg];
       
    47 val pretty_tags = Pretty.list "[" "]" o map pretty_tag;
       
    48 
       
    49 fun pretty_thm_raw ctxt {quote, show_hyps = show_hyps'} raw_th =
       
    50   let
       
    51     val show_tags = Config.get ctxt show_tags;
       
    52     val show_hyps = Config.get ctxt show_hyps;
       
    53 
       
    54     val th = raw_th
       
    55       |> perhaps (try (Thm.transfer (Proof_Context.theory_of ctxt)))
       
    56       |> perhaps (try Thm.strip_shyps);
       
    57 
       
    58     val hyps = if show_hyps then Thm.hyps_of th else Thm.undeclared_hyps (Context.Proof ctxt) th;
       
    59     val extra_shyps = Thm.extra_shyps th;
       
    60     val tags = Thm.get_tags th;
       
    61     val tpairs = Thm.tpairs_of th;
       
    62 
       
    63     val q = if quote then Pretty.quote else I;
       
    64     val prt_term = q o Syntax.pretty_term ctxt;
       
    65 
       
    66 
       
    67     val hlen = length extra_shyps + length hyps + length tpairs;
       
    68     val hsymbs =
       
    69       if hlen = 0 then []
       
    70       else if show_hyps orelse show_hyps' then
       
    71         [Pretty.brk 2, Pretty.list "[" "]"
       
    72           (map (q o Goal_Display.pretty_flexpair ctxt) tpairs @ map prt_term hyps @
       
    73            map (Syntax.pretty_sort ctxt) extra_shyps)]
       
    74       else [Pretty.brk 2, Pretty.str ("[" ^ replicate_string hlen "." ^ "]")];
       
    75     val tsymbs =
       
    76       if null tags orelse not show_tags then []
       
    77       else [Pretty.brk 1, pretty_tags tags];
       
    78   in Pretty.block (prt_term (Thm.prop_of th) :: (hsymbs @ tsymbs)) end;
       
    79 
       
    80 fun pretty_thm ctxt = pretty_thm_raw ctxt {quote = false, show_hyps = true};
       
    81 fun pretty_thm_item ctxt th = Pretty.item [pretty_thm ctxt th];
       
    82 
       
    83 fun pretty_thm_global thy =
       
    84   pretty_thm_raw (Syntax.init_pretty_global thy) {quote = false, show_hyps = false};
       
    85 
       
    86 val string_of_thm = Pretty.string_of oo pretty_thm;
       
    87 val string_of_thm_global = Pretty.string_of oo pretty_thm_global;
       
    88 
       
    89 
       
    90 
       
    91 (** print theory **)
       
    92 
       
    93 (* pretty_full_theory *)
       
    94 
       
    95 fun pretty_full_theory verbose thy =
       
    96   let
       
    97     val ctxt = Syntax.init_pretty_global thy;
       
    98 
       
    99     fun prt_cls c = Syntax.pretty_sort ctxt [c];
       
   100     fun prt_sort S = Syntax.pretty_sort ctxt S;
       
   101     fun prt_arity t (c, Ss) = Syntax.pretty_arity ctxt (t, Ss, [c]);
       
   102     fun prt_typ ty = Pretty.quote (Syntax.pretty_typ ctxt ty);
       
   103     val prt_typ_no_tvars = prt_typ o Logic.unvarifyT_global;
       
   104     fun prt_term t = Pretty.quote (Syntax.pretty_term ctxt t);
       
   105     val prt_term_no_vars = prt_term o Logic.unvarify_global;
       
   106     fun prt_const (c, ty) = [Pretty.mark_str c, Pretty.str " ::", Pretty.brk 1, prt_typ_no_tvars ty];
       
   107     val prt_const' = Defs.pretty_const ctxt;
       
   108 
       
   109     fun pretty_classrel (c, []) = prt_cls c
       
   110       | pretty_classrel (c, cs) = Pretty.block
       
   111           (prt_cls c :: Pretty.str " <" :: Pretty.brk 1 :: Pretty.commas (map prt_cls cs));
       
   112 
       
   113     fun pretty_default S = Pretty.block
       
   114       [Pretty.str "default sort:", Pretty.brk 1, prt_sort S];
       
   115 
       
   116     val tfrees = map (fn v => TFree (v, []));
       
   117     fun pretty_type syn (t, (Type.LogicalType n)) =
       
   118           if syn then NONE
       
   119           else SOME (prt_typ (Type (t, tfrees (Name.invent Name.context Name.aT n))))
       
   120       | pretty_type syn (t, (Type.Abbreviation (vs, U, syn'))) =
       
   121           if syn <> syn' then NONE
       
   122           else SOME (Pretty.block
       
   123             [prt_typ (Type (t, tfrees vs)), Pretty.str " =", Pretty.brk 1, prt_typ U])
       
   124       | pretty_type syn (t, Type.Nonterminal) =
       
   125           if not syn then NONE
       
   126           else SOME (prt_typ (Type (t, [])));
       
   127 
       
   128     val pretty_arities = maps (fn (t, ars) => map (prt_arity t) ars);
       
   129 
       
   130     fun pretty_abbrev (c, (ty, t)) = Pretty.block
       
   131       (prt_const (c, ty) @ [Pretty.str " ==", Pretty.brk 1, prt_term_no_vars t]);
       
   132 
       
   133     fun pretty_axm (a, t) =
       
   134       Pretty.block [Pretty.mark_str a, Pretty.str ":", Pretty.brk 1, prt_term_no_vars t];
       
   135 
       
   136     fun pretty_finals reds = Pretty.block
       
   137       (Pretty.str "final:" :: Pretty.brk 1 :: Pretty.commas (map (prt_const' o fst) reds));
       
   138 
       
   139     fun pretty_reduct (lhs, rhs) = Pretty.block
       
   140       ([prt_const' lhs, Pretty.str "  ->", Pretty.brk 2] @
       
   141         Pretty.commas (map prt_const' (sort_by #1 rhs)));
       
   142 
       
   143     fun pretty_restrict (const, name) =
       
   144       Pretty.block ([prt_const' const, Pretty.brk 2, Pretty.str ("(from " ^ quote name ^ ")")]);
       
   145 
       
   146     val defs = Theory.defs_of thy;
       
   147     val {restricts, reducts} = Defs.dest defs;
       
   148     val tsig = Sign.tsig_of thy;
       
   149     val consts = Sign.consts_of thy;
       
   150     val {const_space, constants, constraints} = Consts.dest consts;
       
   151     val extern_const = Name_Space.extern ctxt const_space;
       
   152     val {classes, default, types, ...} = Type.rep_tsig tsig;
       
   153     val (class_space, class_algebra) = classes;
       
   154     val classes = Sorts.classes_of class_algebra;
       
   155     val arities = Sorts.arities_of class_algebra;
       
   156 
       
   157     val clsses =
       
   158       Name_Space.extern_entries verbose ctxt class_space
       
   159         (map (fn ((c, _), cs) => (c, Sign.minimize_sort thy cs)) (Graph.dest classes))
       
   160       |> map (apfst #1);
       
   161     val tdecls = Name_Space.extern_table verbose ctxt types |> map (apfst #1);
       
   162     val arties =
       
   163       Name_Space.extern_entries verbose ctxt (Type.type_space tsig) (Symtab.dest arities)
       
   164       |> map (apfst #1);
       
   165 
       
   166     val cnsts = Name_Space.markup_entries verbose ctxt const_space constants;
       
   167     val log_cnsts = map_filter (fn (c, (ty, NONE)) => SOME (c, ty) | _ => NONE) cnsts;
       
   168     val abbrevs = map_filter (fn (c, (ty, SOME t)) => SOME (c, (ty, t)) | _ => NONE) cnsts;
       
   169     val cnstrs = Name_Space.markup_entries verbose ctxt const_space constraints;
       
   170 
       
   171     val axms = Name_Space.markup_table verbose ctxt (Theory.axiom_table thy);
       
   172 
       
   173     fun prune_const c = not verbose andalso Consts.is_concealed consts c;
       
   174     val (reds0, (reds1, reds2)) = filter_out (prune_const o fst o fst) reducts
       
   175       |> map (fn (lhs, rhs) =>
       
   176         (apfst extern_const lhs, map (apfst extern_const) (filter_out (prune_const o fst) rhs)))
       
   177       |> sort_by (#1 o #1)
       
   178       |> List.partition (null o #2)
       
   179       ||> List.partition (Defs.plain_args o #2 o #1);
       
   180     val rests = restricts |> map (apfst (apfst extern_const)) |> sort_by (#1 o #1);
       
   181   in
       
   182     [Pretty.strs ("names:" :: Context.display_names thy)] @
       
   183     [Pretty.big_list "classes:" (map pretty_classrel clsses),
       
   184       pretty_default default,
       
   185       Pretty.big_list "syntactic types:" (map_filter (pretty_type true) tdecls),
       
   186       Pretty.big_list "logical types:" (map_filter (pretty_type false) tdecls),
       
   187       Pretty.big_list "type arities:" (pretty_arities arties),
       
   188       Pretty.big_list "logical consts:" (map (Pretty.block o prt_const) log_cnsts),
       
   189       Pretty.big_list "abbreviations:" (map pretty_abbrev abbrevs),
       
   190       Pretty.big_list "constraints:" (map (Pretty.block o prt_const) cnstrs),
       
   191       Pretty.big_list "axioms:" (map pretty_axm axms),
       
   192       Pretty.block
       
   193         (Pretty.breaks (Pretty.str "oracles:" ::
       
   194           map Pretty.mark_str (Thm.extern_oracles verbose ctxt))),
       
   195       Pretty.big_list "definitions:"
       
   196         [pretty_finals reds0,
       
   197          Pretty.big_list "non-overloaded:" (map pretty_reduct reds1),
       
   198          Pretty.big_list "overloaded:" (map pretty_reduct reds2),
       
   199          Pretty.big_list "pattern restrictions:" (map pretty_restrict rests)]]
       
   200   end;
       
   201 
       
   202 end;
       
   203 
       
   204 structure Basic_Display: BASIC_DISPLAY = Display;
       
   205 open Basic_Display;