src/Pure/display.ML
author wenzelm
Sun, 18 Mar 2012 13:04:22 +0100
changeset 47876 421760a1efe7
parent 44211 84472e198515
child 50575 11430dd89e35
permissions -rw-r--r--
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
more explicit Context.generic for Name_Space.declare/define and derivatives (NB: naming changed after Proof_Context.init_global);
prefer Context.pretty in low-level operations of structure Sorts/Type (defer full Syntax.init_pretty until error output);
simplified signatures;
     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_default: bool Unsynchronized.ref
    11   val show_consts: bool Config.T
    12   val show_hyps_raw: Config.raw
    13   val show_hyps: bool Config.T
    14   val show_tags_raw: Config.raw
    15   val show_tags: bool Config.T
    16 end;
    17 
    18 signature DISPLAY =
    19 sig
    20   include BASIC_DISPLAY
    21   val pretty_thm_raw: Proof.context -> {quote: bool, show_hyps: bool, show_status: bool} ->
    22     thm -> Pretty.T
    23   val pretty_thm_aux: Proof.context -> bool -> thm -> Pretty.T
    24   val pretty_thm: Proof.context -> thm -> Pretty.T
    25   val pretty_thm_global: theory -> thm -> Pretty.T
    26   val pretty_thm_without_context: thm -> Pretty.T
    27   val string_of_thm: Proof.context -> thm -> string
    28   val string_of_thm_global: theory -> thm -> string
    29   val string_of_thm_without_context: thm -> string
    30   val pretty_thms_aux: Proof.context -> bool -> thm list -> Pretty.T
    31   val pretty_thms: Proof.context -> thm list -> Pretty.T
    32   val print_syntax: theory -> unit
    33   val pretty_full_theory: bool -> theory -> Pretty.T list
    34 end;
    35 
    36 structure Display: DISPLAY =
    37 struct
    38 
    39 (** options **)
    40 
    41 val show_consts_default = Goal_Display.show_consts_default;
    42 val show_consts = Goal_Display.show_consts;
    43 
    44 val show_hyps_raw = Config.declare "show_hyps" (fn _ => Config.Bool false);
    45 val show_hyps = Config.bool show_hyps_raw;
    46 
    47 val show_tags_raw = Config.declare "show_tags" (fn _ => Config.Bool false);
    48 val show_tags = Config.bool show_tags_raw;
    49 
    50 
    51 
    52 (** print thm **)
    53 
    54 fun pretty_tag (name, arg) = Pretty.strs [name, quote arg];
    55 val pretty_tags = Pretty.list "[" "]" o map pretty_tag;
    56 
    57 fun display_status _ false _ = ""
    58   | display_status show_hyps true th =
    59       let
    60         val {oracle = oracle0, unfinished, failed} = Thm.status_of th;
    61         val oracle = oracle0 andalso (not (! quick_and_dirty) orelse show_hyps);
    62       in
    63         if failed then "!!"
    64         else if oracle andalso unfinished then "!?"
    65         else if oracle then "!"
    66         else if unfinished then "?"
    67         else ""
    68       end;
    69 
    70 fun pretty_thm_raw ctxt {quote, show_hyps = show_hyps', show_status} raw_th =
    71   let
    72     val show_tags = Config.get ctxt show_tags;
    73     val show_hyps = Config.get ctxt show_hyps;
    74 
    75     val th = Thm.strip_shyps raw_th;
    76     val {hyps, tpairs, prop, ...} = Thm.rep_thm th;
    77     val xshyps = Thm.extra_shyps th;
    78     val tags = Thm.get_tags th;
    79 
    80     val q = if quote then Pretty.quote else I;
    81     val prt_term = q o Syntax.pretty_term ctxt;
    82 
    83     val asms = map Thm.term_of (Assumption.all_assms_of ctxt);
    84     val hyps' = if show_hyps then hyps else subtract (op aconv) asms hyps;
    85     val status = display_status show_hyps show_status th;
    86 
    87     val hlen = length xshyps + length hyps' + length tpairs;
    88     val hsymbs =
    89       if hlen = 0 andalso status = "" then []
    90       else if show_hyps orelse show_hyps' then
    91         [Pretty.brk 2, Pretty.list "[" "]"
    92           (map (q o Goal_Display.pretty_flexpair ctxt) tpairs @ map prt_term hyps' @
    93            map (Syntax.pretty_sort ctxt) xshyps @
    94             (if status = "" then [] else [Pretty.str status]))]
    95       else [Pretty.brk 2, Pretty.str ("[" ^ implode (replicate hlen ".") ^ status ^ "]")];
    96     val tsymbs =
    97       if null tags orelse not show_tags then []
    98       else [Pretty.brk 1, pretty_tags tags];
    99   in Pretty.block (prt_term prop :: (hsymbs @ tsymbs)) end;
   100 
   101 fun pretty_thm_aux ctxt show_status =
   102   pretty_thm_raw ctxt {quote = false, show_hyps = true, show_status = show_status};
   103 
   104 fun pretty_thm ctxt = pretty_thm_aux ctxt true;
   105 
   106 fun pretty_thm_global thy =
   107   pretty_thm_raw (Syntax.init_pretty_global thy)
   108     {quote = false, show_hyps = false, show_status = true};
   109 
   110 fun pretty_thm_without_context th = pretty_thm_global (Thm.theory_of_thm th) th;
   111 
   112 val string_of_thm = Pretty.string_of oo pretty_thm;
   113 val string_of_thm_global = Pretty.string_of oo pretty_thm_global;
   114 val string_of_thm_without_context = Pretty.string_of o pretty_thm_without_context;
   115 
   116 
   117 (* multiple theorems *)
   118 
   119 fun pretty_thms_aux ctxt flag [th] = pretty_thm_aux ctxt flag th
   120   | pretty_thms_aux ctxt flag ths =
   121       Pretty.blk (0, Pretty.fbreaks (map (pretty_thm_aux ctxt flag) ths));
   122 
   123 fun pretty_thms ctxt = pretty_thms_aux ctxt true;
   124 
   125 
   126 
   127 (** print theory **)
   128 
   129 val print_syntax = Syntax.print_syntax o Sign.syn_of;
   130 
   131 
   132 (* pretty_full_theory *)
   133 
   134 fun pretty_full_theory verbose thy =
   135   let
   136     val ctxt = Syntax.init_pretty_global thy;
   137 
   138     fun prt_cls c = Syntax.pretty_sort ctxt [c];
   139     fun prt_sort S = Syntax.pretty_sort ctxt S;
   140     fun prt_arity t (c, Ss) = Syntax.pretty_arity ctxt (t, Ss, [c]);
   141     fun prt_typ ty = Pretty.quote (Syntax.pretty_typ ctxt ty);
   142     val prt_typ_no_tvars = prt_typ o Logic.unvarifyT_global;
   143     fun prt_term t = Pretty.quote (Syntax.pretty_term ctxt t);
   144     val prt_term_no_vars = prt_term o Logic.unvarify_global;
   145     fun prt_const (c, ty) = [Pretty.str c, Pretty.str " ::", Pretty.brk 1, prt_typ_no_tvars ty];
   146     val prt_const' = Defs.pretty_const ctxt;
   147 
   148     fun pretty_classrel (c, []) = prt_cls c
   149       | pretty_classrel (c, cs) = Pretty.block
   150           (prt_cls c :: Pretty.str " <" :: Pretty.brk 1 ::
   151             Pretty.commas (map prt_cls cs));
   152 
   153     fun pretty_default S = Pretty.block
   154       [Pretty.str "default sort:", Pretty.brk 1, prt_sort S];
   155 
   156     val tfrees = map (fn v => TFree (v, []));
   157     fun pretty_type syn (t, (Type.LogicalType n)) =
   158           if syn then NONE
   159           else SOME (prt_typ (Type (t, tfrees (Name.invent Name.context Name.aT n))))
   160       | pretty_type syn (t, (Type.Abbreviation (vs, U, syn'))) =
   161           if syn <> syn' then NONE
   162           else SOME (Pretty.block
   163             [prt_typ (Type (t, tfrees vs)), Pretty.str " =", Pretty.brk 1, prt_typ U])
   164       | pretty_type syn (t, Type.Nonterminal) =
   165           if not syn then NONE
   166           else SOME (prt_typ (Type (t, [])));
   167 
   168     val pretty_arities = maps (fn (t, ars) => map (prt_arity t) ars);
   169 
   170     fun pretty_abbrev (c, (ty, t)) = Pretty.block
   171       (prt_const (c, ty) @ [Pretty.str " ==", Pretty.brk 1, prt_term_no_vars t]);
   172 
   173     fun pretty_axm (a, t) = Pretty.block [Pretty.str (a ^ ":"), Pretty.brk 1, prt_term_no_vars t];
   174 
   175     fun pretty_finals reds = Pretty.block
   176       (Pretty.str "final:" :: Pretty.brk 1 :: Pretty.commas (map (prt_const' o fst) reds));
   177 
   178     fun pretty_reduct (lhs, rhs) = Pretty.block
   179       ([prt_const' lhs, Pretty.str "  ->", Pretty.brk 2] @
   180         Pretty.commas (map prt_const' (sort_wrt #1 rhs)));
   181 
   182     fun pretty_restrict (const, name) =
   183       Pretty.block ([prt_const' const, Pretty.brk 2, Pretty.str ("(from " ^ quote name ^ ")")]);
   184 
   185     val axioms = (Theory.axiom_space thy, Theory.axiom_table thy);
   186     val defs = Theory.defs_of thy;
   187     val {restricts, reducts} = Defs.dest defs;
   188     val tsig = Sign.tsig_of thy;
   189     val consts = Sign.consts_of thy;
   190     val {constants, constraints} = Consts.dest consts;
   191     val extern_const = Name_Space.extern ctxt (#1 constants);
   192     val {classes, default, types, ...} = Type.rep_tsig tsig;
   193     val (class_space, class_algebra) = classes;
   194     val classes = Sorts.classes_of class_algebra;
   195     val arities = Sorts.arities_of class_algebra;
   196 
   197     val clsses = Name_Space.dest_table ctxt (class_space, Symtab.make (Graph.dest classes));
   198     val tdecls = Name_Space.dest_table ctxt types;
   199     val arties = Name_Space.dest_table ctxt (Type.type_space tsig, arities);
   200 
   201     fun prune_const c = not verbose andalso Consts.is_concealed consts c;
   202     val cnsts = Name_Space.extern_table ctxt (#1 constants,
   203       Symtab.make (filter_out (prune_const o fst) (Symtab.dest (#2 constants))));
   204 
   205     val log_cnsts = map_filter (fn (c, (ty, NONE)) => SOME (c, ty) | _ => NONE) cnsts;
   206     val abbrevs = map_filter (fn (c, (ty, SOME t)) => SOME (c, (ty, t)) | _ => NONE) cnsts;
   207     val cnstrs = Name_Space.extern_table ctxt constraints;
   208 
   209     val axms = Name_Space.extern_table ctxt axioms;
   210 
   211     val (reds0, (reds1, reds2)) = filter_out (prune_const o fst o fst) reducts
   212       |> map (fn (lhs, rhs) =>
   213         (apfst extern_const lhs, map (apfst extern_const) (filter_out (prune_const o fst) rhs)))
   214       |> sort_wrt (#1 o #1)
   215       |> List.partition (null o #2)
   216       ||> List.partition (Defs.plain_args o #2 o #1);
   217     val rests = restricts |> map (apfst (apfst extern_const)) |> sort_wrt (#1 o #1);
   218   in
   219     [Pretty.strs ("names:" :: Context.display_names thy)] @
   220     [Pretty.big_list "classes:" (map pretty_classrel clsses),
   221       pretty_default default,
   222       Pretty.big_list "syntactic types:" (map_filter (pretty_type true) tdecls),
   223       Pretty.big_list "logical types:" (map_filter (pretty_type false) tdecls),
   224       Pretty.big_list "type arities:" (pretty_arities arties),
   225       Pretty.big_list "logical consts:" (map (Pretty.block o prt_const) log_cnsts),
   226       Pretty.big_list "abbreviations:" (map pretty_abbrev abbrevs),
   227       Pretty.big_list "constraints:" (map (Pretty.block o prt_const) cnstrs),
   228       Pretty.big_list "axioms:" (map pretty_axm axms),
   229       Pretty.strs ("oracles:" :: Thm.extern_oracles ctxt),
   230       Pretty.big_list "definitions:"
   231         [pretty_finals reds0,
   232          Pretty.big_list "non-overloaded:" (map pretty_reduct reds1),
   233          Pretty.big_list "overloaded:" (map pretty_reduct reds2),
   234          Pretty.big_list "pattern restrictions:" (map pretty_restrict rests)]]
   235   end;
   236 
   237 end;
   238 
   239 structure Basic_Display: BASIC_DISPLAY = Display;
   240 open Basic_Display;