src/Pure/display.ML
changeset 61266 eb9522a9d997
parent 61252 c165f0472d57
equal deleted inserted replaced
61265:996d73a96b4f 61266:eb9522a9d997
    21   val pretty_thm: Proof.context -> thm -> Pretty.T
    21   val pretty_thm: Proof.context -> thm -> Pretty.T
    22   val pretty_thm_item: Proof.context -> thm -> Pretty.T
    22   val pretty_thm_item: Proof.context -> thm -> Pretty.T
    23   val pretty_thm_global: theory -> thm -> Pretty.T
    23   val pretty_thm_global: theory -> thm -> Pretty.T
    24   val string_of_thm: Proof.context -> thm -> string
    24   val string_of_thm: Proof.context -> thm -> string
    25   val string_of_thm_global: theory -> thm -> string
    25   val string_of_thm_global: theory -> thm -> string
    26   val pretty_full_theory: bool -> theory -> Pretty.T list
       
    27 end;
    26 end;
    28 
    27 
    29 structure Display: DISPLAY =
    28 structure Display: DISPLAY =
    30 struct
    29 struct
    31 
    30 
    84   pretty_thm_raw (Syntax.init_pretty_global thy) {quote = false, show_hyps = false};
    83   pretty_thm_raw (Syntax.init_pretty_global thy) {quote = false, show_hyps = false};
    85 
    84 
    86 val string_of_thm = Pretty.string_of oo pretty_thm;
    85 val string_of_thm = Pretty.string_of oo pretty_thm;
    87 val string_of_thm_global = Pretty.string_of oo pretty_thm_global;
    86 val string_of_thm_global = Pretty.string_of oo pretty_thm_global;
    88 
    87 
    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 
       
   108     fun pretty_classrel (c, []) = prt_cls c
       
   109       | pretty_classrel (c, cs) = Pretty.block
       
   110           (prt_cls c :: Pretty.str " <" :: Pretty.brk 1 :: Pretty.commas (map prt_cls cs));
       
   111 
       
   112     fun pretty_default S = Pretty.block
       
   113       [Pretty.str "default sort:", Pretty.brk 1, prt_sort S];
       
   114 
       
   115     val tfrees = map (fn v => TFree (v, []));
       
   116     fun pretty_type syn (t, Type.LogicalType n) =
       
   117           if syn then NONE
       
   118           else SOME (prt_typ (Type (t, tfrees (Name.invent Name.context Name.aT n))))
       
   119       | pretty_type syn (t, Type.Abbreviation (vs, U, syn')) =
       
   120           if syn <> syn' then NONE
       
   121           else SOME (Pretty.block
       
   122             [prt_typ (Type (t, tfrees vs)), Pretty.str " =", Pretty.brk 1, prt_typ U])
       
   123       | pretty_type syn (t, Type.Nonterminal) =
       
   124           if not syn then NONE
       
   125           else SOME (prt_typ (Type (t, [])));
       
   126 
       
   127     val pretty_arities = maps (fn (t, ars) => map (prt_arity t) ars);
       
   128 
       
   129     fun pretty_abbrev (c, (ty, t)) = Pretty.block
       
   130       (prt_const (c, ty) @ [Pretty.str " ==", Pretty.brk 1, prt_term_no_vars t]);
       
   131 
       
   132     fun pretty_axm (a, t) =
       
   133       Pretty.block [Pretty.mark_str a, Pretty.str ":", Pretty.brk 1, prt_term_no_vars t];
       
   134 
       
   135     val tsig = Sign.tsig_of thy;
       
   136     val consts = Sign.consts_of thy;
       
   137     val {const_space, constants, constraints} = Consts.dest consts;
       
   138     val {classes, default, types, ...} = Type.rep_tsig tsig;
       
   139     val type_space = Type.type_space tsig;
       
   140     val (class_space, class_algebra) = classes;
       
   141     val classes = Sorts.classes_of class_algebra;
       
   142     val arities = Sorts.arities_of class_algebra;
       
   143 
       
   144     val clsses =
       
   145       Name_Space.extern_entries verbose ctxt class_space
       
   146         (map (fn ((c, _), cs) => (c, Sign.minimize_sort thy cs)) (Graph.dest classes))
       
   147       |> map (apfst #1);
       
   148     val tdecls = Name_Space.extern_table verbose ctxt types |> map (apfst #1);
       
   149     val arties =
       
   150       Name_Space.extern_entries verbose ctxt type_space (Symtab.dest arities)
       
   151       |> map (apfst #1);
       
   152 
       
   153     val cnsts = Name_Space.markup_entries verbose ctxt const_space constants;
       
   154     val log_cnsts = map_filter (fn (c, (ty, NONE)) => SOME (c, ty) | _ => NONE) cnsts;
       
   155     val abbrevs = map_filter (fn (c, (ty, SOME t)) => SOME (c, (ty, t)) | _ => NONE) cnsts;
       
   156     val cnstrs = Name_Space.markup_entries verbose ctxt const_space constraints;
       
   157 
       
   158     val axms = Name_Space.markup_table verbose ctxt (Theory.axiom_table thy);
       
   159   in
       
   160     [Pretty.strs ("names:" :: Context.display_names thy)] @
       
   161     [Pretty.big_list "classes:" (map pretty_classrel clsses),
       
   162       pretty_default default,
       
   163       Pretty.big_list "syntactic types:" (map_filter (pretty_type true) tdecls),
       
   164       Pretty.big_list "logical types:" (map_filter (pretty_type false) tdecls),
       
   165       Pretty.big_list "type arities:" (pretty_arities arties),
       
   166       Pretty.big_list "logical consts:" (map (Pretty.block o prt_const) log_cnsts),
       
   167       Pretty.big_list "abbreviations:" (map pretty_abbrev abbrevs),
       
   168       Pretty.big_list "constraints:" (map (Pretty.block o prt_const) cnstrs),
       
   169       Pretty.big_list "axioms:" (map pretty_axm axms),
       
   170       Pretty.block
       
   171         (Pretty.breaks (Pretty.str "oracles:" ::
       
   172           map Pretty.mark_str (Thm.extern_oracles verbose ctxt)))]
       
   173   end;
       
   174 
       
   175 end;
    88 end;
   176 
    89 
   177 structure Basic_Display: BASIC_DISPLAY = Display;
    90 structure Basic_Display: BASIC_DISPLAY = Display;
   178 open Basic_Display;
    91 open Basic_Display;