author | haftmann |
Sat, 19 Jul 2025 18:41:55 +0200 | |
changeset 82886 | 8d1e295aab70 |
parent 80910 | 406a85a25189 |
permissions | -rw-r--r-- |
66646 | 1 |
(* Title: HOL/Tools/Nunchaku/nunchaku_display.ML |
66614
1f1c5d85d232
moved Nunchaku to Main; the goal is to move Nitpick out in the next 1-2 years
blanchet
parents:
64389
diff
changeset
|
2 |
Author: Jasmin Blanchette, VU Amsterdam |
1f1c5d85d232
moved Nunchaku to Main; the goal is to move Nitpick out in the next 1-2 years
blanchet
parents:
64389
diff
changeset
|
3 |
Copyright 2015, 2016, 2017 |
64389 | 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 |
|
80910 | 26 |
| sorting_str_of_typ (Type (s, Ts)) = "b" ^ s ^ implode_space (map sorting_str_of_typ Ts) |
64389 | 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 |
||
66634
56456f388867
eliminate artifact of translation in printed Nunchaku model
blanchet
parents:
66623
diff
changeset
|
35 |
val basic_defs_sym = @{thms rmember_def[abs_def]}; |
56456f388867
eliminate artifact of translation in printed Nunchaku model
blanchet
parents:
66623
diff
changeset
|
36 |
|
56456f388867
eliminate artifact of translation in printed Nunchaku model
blanchet
parents:
66623
diff
changeset
|
37 |
fun fold_basic_def ctxt = |
56456f388867
eliminate artifact of translation in printed Nunchaku model
blanchet
parents:
66623
diff
changeset
|
38 |
let val thy = Proof_Context.theory_of ctxt in |
56456f388867
eliminate artifact of translation in printed Nunchaku model
blanchet
parents:
66623
diff
changeset
|
39 |
Pattern.rewrite_term thy (map (Logic.dest_equals o Thm.prop_of) basic_defs_sym) [] |
56456f388867
eliminate artifact of translation in printed Nunchaku model
blanchet
parents:
66623
diff
changeset
|
40 |
end; |
56456f388867
eliminate artifact of translation in printed Nunchaku model
blanchet
parents:
66623
diff
changeset
|
41 |
|
64389 | 42 |
fun pretty_of_isa_model_opt _ NONE = |
66623 | 43 |
pretty_indent (Pretty.str "Model unavailable due to internal error") |
44 |
| pretty_of_isa_model_opt ctxt0 (SOME {type_model, free_model, pat_complete_model, |
|
45 |
pat_incomplete_model, skolem_model, auxiliary_model}) = |
|
64389 | 46 |
let |
47 |
val ctxt = ctxt0 |> Config.put show_question_marks false; |
|
48 |
fun pretty_of_typ_entry (T, atoms) = |
|
49 |
Pretty.block (Pretty.breaks [Syntax.pretty_typ ctxt T, Pretty.str "=", |
|
50 |
Pretty.enum "," "{" "}" (map (Syntax.pretty_term ctxt) atoms)]); |
|
51 |
||
52 |
fun pretty_of_term_entry (t, value) = |
|
53 |
let |
|
54 |
val no_types_ctxt = ctxt |> Config.put show_types false; |
|
55 |
val schematic_ctxt = ctxt |> Proof_Context.set_mode Proof_Context.mode_schematic; |
|
56 |
||
57 |
val show_types = Config.get ctxt show_types; |
|
66634
56456f388867
eliminate artifact of translation in printed Nunchaku model
blanchet
parents:
66623
diff
changeset
|
58 |
val value' = |
56456f388867
eliminate artifact of translation in printed Nunchaku model
blanchet
parents:
66623
diff
changeset
|
59 |
value |
56456f388867
eliminate artifact of translation in printed Nunchaku model
blanchet
parents:
66623
diff
changeset
|
60 |
|> perhaps (try (Syntax.check_term schematic_ctxt)) |
56456f388867
eliminate artifact of translation in printed Nunchaku model
blanchet
parents:
66623
diff
changeset
|
61 |
|> perhaps (try (fold_basic_def ctxt)); |
64389 | 62 |
val T = fastype_of t; |
63 |
val T' = if T = dummyT then try fastype_of value' |> the_default T else T; |
|
64 |
val t' = t |> show_types ? Type.constraint T'; |
|
65 |
in |
|
66 |
Pretty.block (Pretty.breaks |
|
67 |
[Syntax.pretty_term ctxt t' |
|
68 |
|> (show_types andalso T' <> dummyT) ? (single #> Pretty.enclose "(" ")"), |
|
69 |
Pretty.str "=", Syntax.pretty_term no_types_ctxt value']) |
|
70 |
end; |
|
71 |
||
72 |
fun chunks_of_entries sorting_str_of pretty_of title entries = |
|
73 |
if not (null entries) then |
|
74 |
(if title = "" then [] else [Pretty.str (title ^ plural_s_for_list entries ^ ":")]) @ |
|
75 |
map (pretty_indent o pretty_of) (sort_by (sorting_str_of o fst) entries) |
|
76 |
else |
|
77 |
[]; |
|
78 |
||
79 |
val chunks = |
|
80 |
(if null free_model then |
|
81 |
[pretty_indent (Pretty.str "No free variables")] |
|
82 |
else |
|
83 |
chunks_of_entries sorting_str_of_term pretty_of_term_entry "" free_model) @ |
|
84 |
chunks_of_entries sorting_str_of_term pretty_of_term_entry "Skolem constant" skolem_model @ |
|
66623 | 85 |
chunks_of_entries sorting_str_of_term pretty_of_term_entry |
86 |
"Potentially underspecified constant" pat_incomplete_model @ |
|
87 |
chunks_of_entries sorting_str_of_term pretty_of_term_entry "Fully specified constant" |
|
88 |
pat_complete_model @ |
|
89 |
chunks_of_entries sorting_str_of_term pretty_of_term_entry "Auxiliary variable" |
|
90 |
auxiliary_model @ |
|
64389 | 91 |
chunks_of_entries sorting_str_of_typ pretty_of_typ_entry "Type" type_model; |
92 |
in |
|
93 |
Pretty.chunks chunks |
|
94 |
end; |
|
95 |
||
96 |
end; |