src/HOL/Tools/Nunchaku/nunchaku_display.ML
changeset 66614 1f1c5d85d232
parent 64389 6273d4c8325b
child 66623 8fc868e9e1bf
equal deleted inserted replaced
66613:db3969568560 66614:1f1c5d85d232
       
     1 (*  Title:      HOL/Nunchaku/Tools/nunchaku_display.ML
       
     2     Author:     Jasmin Blanchette, VU Amsterdam
       
     3     Copyright   2015, 2016, 2017
       
     4 
       
     5 Pretty printing of Isabelle/HOL models for Nunchaku.
       
     6 *)
       
     7 
       
     8 signature NUNCHAKU_DISPLAY =
       
     9 sig
       
    10   type isa_model = Nunchaku_Reconstruct.isa_model
       
    11 
       
    12   val pretty_of_isa_model_opt: Proof.context -> isa_model option -> Pretty.T
       
    13 end;
       
    14 
       
    15 structure Nunchaku_Display : NUNCHAKU_DISPLAY =
       
    16 struct
       
    17 
       
    18 open Nunchaku_Util;
       
    19 open Nunchaku_Reconstruct;
       
    20 
       
    21 val indent_size = 2;
       
    22 
       
    23 val pretty_indent = Pretty.indent indent_size;
       
    24 
       
    25 fun sorting_str_of_typ (TFree (s, _)) = "a" ^ s
       
    26   | sorting_str_of_typ (Type (s, Ts)) = "b" ^ s ^ space_implode " " (map sorting_str_of_typ Ts)
       
    27   | sorting_str_of_typ (TVar _) = "X";
       
    28 
       
    29 fun sorting_str_of_term (Const (s, T)) = "b" ^ s ^ sorting_str_of_typ T
       
    30   | sorting_str_of_term (Free (s, _)) = "a" ^ s
       
    31   | sorting_str_of_term (t $ u) = sorting_str_of_term t ^ " " ^ sorting_str_of_term u
       
    32   | sorting_str_of_term (Abs (_, T, t)) = "c" ^ sorting_str_of_typ T ^ " " ^ sorting_str_of_term t
       
    33   | sorting_str_of_term _ = "X";
       
    34 
       
    35 fun pretty_of_isa_model_opt _ NONE =
       
    36     pretty_indent (Pretty.str "Model unavailable (internal error)")
       
    37   | pretty_of_isa_model_opt ctxt0
       
    38       (SOME {type_model, free_model, pat_complete_model, pat_incomplete_model, skolem_model}) =
       
    39     let
       
    40       val ctxt = ctxt0 |> Config.put show_question_marks false;
       
    41 
       
    42       val pat_incomplete_model' = pat_incomplete_model
       
    43         |> filter_out (can (fn Const (@{const_name unreachable}, _) => ()) o fst);
       
    44 
       
    45       fun pretty_of_typ_entry (T, atoms) =
       
    46         Pretty.block (Pretty.breaks [Syntax.pretty_typ ctxt T, Pretty.str "=",
       
    47            Pretty.enum "," "{" "}" (map (Syntax.pretty_term ctxt) atoms)]);
       
    48 
       
    49       fun pretty_of_term_entry (t, value) =
       
    50         let
       
    51           val no_types_ctxt = ctxt |> Config.put show_types false;
       
    52           val schematic_ctxt = ctxt |> Proof_Context.set_mode Proof_Context.mode_schematic;
       
    53 
       
    54           val show_types = Config.get ctxt show_types;
       
    55           val value' = value |> perhaps (try (Syntax.check_term schematic_ctxt));
       
    56           val T = fastype_of t;
       
    57           val T' = if T = dummyT then try fastype_of value' |> the_default T else T;
       
    58           val t' = t |> show_types ? Type.constraint T';
       
    59         in
       
    60           Pretty.block (Pretty.breaks
       
    61             [Syntax.pretty_term ctxt t'
       
    62              |> (show_types andalso T' <> dummyT) ? (single #> Pretty.enclose "(" ")"),
       
    63              Pretty.str "=", Syntax.pretty_term no_types_ctxt value'])
       
    64         end;
       
    65 
       
    66       fun chunks_of_entries sorting_str_of pretty_of title entries =
       
    67         if not (null entries) then
       
    68           (if title = "" then [] else [Pretty.str (title ^ plural_s_for_list entries ^ ":")]) @
       
    69           map (pretty_indent o pretty_of) (sort_by (sorting_str_of o fst) entries)
       
    70         else
       
    71           [];
       
    72 
       
    73       val chunks =
       
    74         (if null free_model then
       
    75            [pretty_indent (Pretty.str "No free variables")]
       
    76          else
       
    77            chunks_of_entries sorting_str_of_term pretty_of_term_entry "" free_model) @
       
    78         chunks_of_entries sorting_str_of_term pretty_of_term_entry "Skolem constant" skolem_model @
       
    79         chunks_of_entries sorting_str_of_term pretty_of_term_entry "Underspecified constant"
       
    80           pat_incomplete_model' @
       
    81         (if Config.get ctxt show_consts then
       
    82            chunks_of_entries sorting_str_of_term pretty_of_term_entry "Fully specified constant"
       
    83              pat_complete_model
       
    84          else
       
    85            []) @
       
    86         chunks_of_entries sorting_str_of_typ pretty_of_typ_entry "Type" type_model;
       
    87     in
       
    88       Pretty.chunks chunks
       
    89     end;
       
    90 
       
    91 end;