author | wenzelm |
Fri, 13 May 2011 22:55:00 +0200 | |
changeset 42793 | 88bee9f6eec7 |
parent 42415 | 10accf397ab6 |
child 45479 | 3387d482e0a9 |
permissions | -rw-r--r-- |
33982 | 1 |
(* Title: HOL/Tools/Nitpick/nitpick_model.ML |
33192 | 2 |
Author: Jasmin Blanchette, TU Muenchen |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
3 |
Copyright 2009, 2010 |
33192 | 4 |
|
5 |
Model reconstruction for Nitpick. |
|
6 |
*) |
|
7 |
||
8 |
signature NITPICK_MODEL = |
|
9 |
sig |
|
33705
947184dc75c9
removed a few global names in Nitpick (styp, nat_less, pairf)
blanchet
parents:
33580
diff
changeset
|
10 |
type styp = Nitpick_Util.styp |
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
11 |
type scope = Nitpick_Scope.scope |
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
12 |
type rep = Nitpick_Rep.rep |
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
13 |
type nut = Nitpick_Nut.nut |
33192 | 14 |
|
36390
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36389
diff
changeset
|
15 |
type params = |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36389
diff
changeset
|
16 |
{show_datatypes: bool, |
41993
bd6296de1432
reintroduced "show_skolems" option -- useful when too many Skolems are displayed
blanchet
parents:
41875
diff
changeset
|
17 |
show_skolems: bool, |
36390
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36389
diff
changeset
|
18 |
show_consts: bool} |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36389
diff
changeset
|
19 |
|
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
20 |
type term_postprocessor = |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
21 |
Proof.context -> string -> (typ -> term list) -> typ -> term -> term |
33192 | 22 |
|
23 |
structure NameTable : TABLE |
|
24 |
||
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
25 |
val irrelevant : string |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
26 |
val unknown : string |
37261 | 27 |
val unrep : unit -> string |
35711
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35665
diff
changeset
|
28 |
val register_term_postprocessor : |
38284
9f98107ad8b4
use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents:
38240
diff
changeset
|
29 |
typ -> term_postprocessor -> morphism -> Context.generic -> Context.generic |
38240
a44d108a8d39
local versions of Nitpick.register_xxx functions
blanchet
parents:
38209
diff
changeset
|
30 |
val register_term_postprocessor_global : |
35711
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35665
diff
changeset
|
31 |
typ -> term_postprocessor -> theory -> theory |
38284
9f98107ad8b4
use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents:
38240
diff
changeset
|
32 |
val unregister_term_postprocessor : |
9f98107ad8b4
use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents:
38240
diff
changeset
|
33 |
typ -> morphism -> Context.generic -> Context.generic |
38240
a44d108a8d39
local versions of Nitpick.register_xxx functions
blanchet
parents:
38209
diff
changeset
|
34 |
val unregister_term_postprocessor_global : typ -> theory -> theory |
33192 | 35 |
val tuple_list_for_name : |
36 |
nut NameTable.table -> Kodkod.raw_bound list -> nut -> int list list |
|
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
37 |
val dest_plain_fun : term -> bool * (term list * term list) |
33192 | 38 |
val reconstruct_hol_model : |
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
39 |
params -> scope -> (term option * int list) list |
38170 | 40 |
-> (typ option * string list) list -> styp list -> styp list -> nut list |
41 |
-> nut list -> nut list -> nut NameTable.table -> Kodkod.raw_bound list |
|
33192 | 42 |
-> Pretty.T * bool |
43 |
val prove_hol_model : |
|
44 |
scope -> Time.time option -> nut list -> nut list -> nut NameTable.table |
|
45 |
-> Kodkod.raw_bound list -> term -> bool option |
|
46 |
end; |
|
47 |
||
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
48 |
structure Nitpick_Model : NITPICK_MODEL = |
33192 | 49 |
struct |
50 |
||
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
51 |
open Nitpick_Util |
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
52 |
open Nitpick_HOL |
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
53 |
open Nitpick_Scope |
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
54 |
open Nitpick_Peephole |
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
55 |
open Nitpick_Rep |
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
56 |
open Nitpick_Nut |
33192 | 57 |
|
34126 | 58 |
structure KK = Kodkod |
59 |
||
36390
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36389
diff
changeset
|
60 |
type params = |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36389
diff
changeset
|
61 |
{show_datatypes: bool, |
41993
bd6296de1432
reintroduced "show_skolems" option -- useful when too many Skolems are displayed
blanchet
parents:
41875
diff
changeset
|
62 |
show_skolems: bool, |
36390
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36389
diff
changeset
|
63 |
show_consts: bool} |
33192 | 64 |
|
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
65 |
type term_postprocessor = |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
66 |
Proof.context -> string -> (typ -> term list) -> typ -> term -> term |
35711
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35665
diff
changeset
|
67 |
|
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
41052
diff
changeset
|
68 |
structure Data = Generic_Data |
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
41052
diff
changeset
|
69 |
( |
35711
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35665
diff
changeset
|
70 |
type T = (typ * term_postprocessor) list |
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35665
diff
changeset
|
71 |
val empty = [] |
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35665
diff
changeset
|
72 |
val extend = I |
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
41052
diff
changeset
|
73 |
fun merge data = AList.merge (op =) (K true) data |
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
41052
diff
changeset
|
74 |
) |
35711
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35665
diff
changeset
|
75 |
|
37261 | 76 |
fun xsym s s' () = if print_mode_active Symbol.xsymbolsN then s else s' |
77 |
||
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
78 |
val irrelevant = "_" |
33192 | 79 |
val unknown = "?" |
37261 | 80 |
val unrep = xsym "\<dots>" "..." |
81 |
val maybe_mixfix = xsym "_\<^sup>?" "_?" |
|
82 |
val base_mixfix = xsym "_\<^bsub>base\<^esub>" "_.base" |
|
83 |
val step_mixfix = xsym "_\<^bsub>step\<^esub>" "_.step" |
|
84 |
val abs_mixfix = xsym "\<guillemotleft>_\<guillemotright>" "\"_\"" |
|
35718 | 85 |
val arg_var_prefix = "x" |
37261 | 86 |
val cyclic_co_val_name = xsym "\<omega>" "w" |
87 |
val cyclic_const_prefix = xsym "\<xi>" "X" |
|
88 |
fun cyclic_type_name () = nitpick_prefix ^ cyclic_const_prefix () |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
89 |
val opt_flag = nitpick_prefix ^ "opt" |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
90 |
val non_opt_flag = nitpick_prefix ^ "non_opt" |
33192 | 91 |
|
35076
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
35075
diff
changeset
|
92 |
type atom_pool = ((string * int) * int list) list |
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
35075
diff
changeset
|
93 |
|
35718 | 94 |
fun add_wacky_syntax ctxt = |
95 |
let |
|
96 |
val name_of = fst o dest_Const |
|
42361 | 97 |
val thy = Proof_Context.theory_of ctxt |> Context.reject_draft |
35718 | 98 |
val (maybe_t, thy) = |
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
99 |
Sign.declare_const_global ((@{binding nitpick_maybe}, @{typ "'a => 'a"}), |
37261 | 100 |
Mixfix (maybe_mixfix (), [1000], 1000)) thy |
35718 | 101 |
val (abs_t, thy) = |
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
102 |
Sign.declare_const_global ((@{binding nitpick_abs}, @{typ "'a => 'b"}), |
37261 | 103 |
Mixfix (abs_mixfix (), [40], 40)) thy |
35718 | 104 |
val (base_t, thy) = |
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
105 |
Sign.declare_const_global ((@{binding nitpick_base}, @{typ "'a => 'a"}), |
37261 | 106 |
Mixfix (base_mixfix (), [1000], 1000)) thy |
35718 | 107 |
val (step_t, thy) = |
42375
774df7c59508
report Name_Space.declare/define, relatively to context;
wenzelm
parents:
42361
diff
changeset
|
108 |
Sign.declare_const_global ((@{binding nitpick_step}, @{typ "'a => 'a"}), |
37261 | 109 |
Mixfix (step_mixfix (), [1000], 1000)) thy |
35718 | 110 |
in |
111 |
(pairself (pairself name_of) ((maybe_t, abs_t), (base_t, step_t)), |
|
42361 | 112 |
Proof_Context.transfer_syntax thy ctxt) |
35718 | 113 |
end |
114 |
||
115 |
(** Term reconstruction **) |
|
116 |
||
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
117 |
fun nth_atom_number pool T j = |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
118 |
case AList.lookup (op =) (!pool) T of |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
119 |
SOME js => |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
120 |
(case find_index (curry (op =) j) js of |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
121 |
~1 => (Unsynchronized.change pool (cons (T, j :: js)); |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
122 |
length js + 1) |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
123 |
| n => length js - n) |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
124 |
| NONE => (Unsynchronized.change pool (cons (T, [j])); 1) |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
125 |
fun atom_suffix s = |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
126 |
nat_subscript |
40627
becf5d5187cc
renamed raw "explode" function to "raw_explode" to emphasize its meaning;
wenzelm
parents:
40132
diff
changeset
|
127 |
#> (s <> "" andalso Symbol.is_ascii_digit (List.last (raw_explode s))) (* FIXME Symbol.explode (?) *) |
33884
a0c43f185fef
generate clearer atom names in Nitpick for types that end with a digit;
blanchet
parents:
33705
diff
changeset
|
128 |
? prefix "\<^isub>," |
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
129 |
fun nth_atom_name thy atomss pool prefix T j = |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
130 |
let |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
131 |
val ss = these (triple_lookup (type_match thy) atomss T) |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
132 |
val m = nth_atom_number pool T j |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
133 |
in |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
134 |
if m <= length ss then |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
135 |
nth ss (m - 1) |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
136 |
else case T of |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
137 |
Type (s, _) => |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
138 |
let val s' = shortest_name s in |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
139 |
prefix ^ |
41039
405a9f41ad6b
show strings as "s_1" etc. rather than "l_1" etc.
blanchet
parents:
40627
diff
changeset
|
140 |
(if T = @{typ string} then "s" |
405a9f41ad6b
show strings as "s_1" etc. rather than "l_1" etc.
blanchet
parents:
40627
diff
changeset
|
141 |
else if String.isPrefix "\\" s' then s' |
405a9f41ad6b
show strings as "s_1" etc. rather than "l_1" etc.
blanchet
parents:
40627
diff
changeset
|
142 |
else substring (s', 0, 1)) ^ atom_suffix s m |
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
143 |
end |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
144 |
| TFree (s, _) => prefix ^ perhaps (try (unprefix "'")) s ^ atom_suffix s m |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
145 |
| _ => raise TYPE ("Nitpick_Model.nth_atom_name", [T], []) |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
146 |
end |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
147 |
fun nth_atom thy atomss pool for_auto T j = |
33192 | 148 |
if for_auto then |
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
149 |
Free (nth_atom_name thy atomss pool (hd (space_explode "." nitpick_prefix)) |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
150 |
T j, T) |
33192 | 151 |
else |
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
152 |
Const (nth_atom_name thy atomss pool "" T j, T) |
33192 | 153 |
|
35177 | 154 |
fun extract_real_number (Const (@{const_name divide}, _) $ t1 $ t2) = |
34126 | 155 |
real (snd (HOLogic.dest_number t1)) / real (snd (HOLogic.dest_number t2)) |
156 |
| extract_real_number t = real (snd (HOLogic.dest_number t)) |
|
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
157 |
fun nice_term_ord (Abs (_, _, t1), Abs (_, _, t2)) = nice_term_ord (t1, t2) |
34126 | 158 |
| nice_term_ord tp = Real.compare (pairself extract_real_number tp) |
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
159 |
handle TERM ("dest_number", _) => |
34126 | 160 |
case tp of |
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
161 |
(t11 $ t12, t21 $ t22) => |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
162 |
(case nice_term_ord (t11, t21) of |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
163 |
EQUAL => nice_term_ord (t12, t22) |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
164 |
| ord => ord) |
35408 | 165 |
| _ => Term_Ord.fast_term_ord tp |
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
166 |
|
38284
9f98107ad8b4
use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents:
38240
diff
changeset
|
167 |
fun register_term_postprocessor_generic T postproc = |
9f98107ad8b4
use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents:
38240
diff
changeset
|
168 |
Data.map (cons (T, postproc)) |
9f98107ad8b4
use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents:
38240
diff
changeset
|
169 |
(* TODO: Consider morphism. *) |
9f98107ad8b4
use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents:
38240
diff
changeset
|
170 |
fun register_term_postprocessor T postproc (_ : morphism) = |
9f98107ad8b4
use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents:
38240
diff
changeset
|
171 |
register_term_postprocessor_generic T postproc |
38240
a44d108a8d39
local versions of Nitpick.register_xxx functions
blanchet
parents:
38209
diff
changeset
|
172 |
val register_term_postprocessor_global = |
a44d108a8d39
local versions of Nitpick.register_xxx functions
blanchet
parents:
38209
diff
changeset
|
173 |
Context.theory_map oo register_term_postprocessor_generic |
a44d108a8d39
local versions of Nitpick.register_xxx functions
blanchet
parents:
38209
diff
changeset
|
174 |
|
a44d108a8d39
local versions of Nitpick.register_xxx functions
blanchet
parents:
38209
diff
changeset
|
175 |
fun unregister_term_postprocessor_generic T = Data.map (AList.delete (op =) T) |
38284
9f98107ad8b4
use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents:
38240
diff
changeset
|
176 |
(* TODO: Consider morphism. *) |
9f98107ad8b4
use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents:
38240
diff
changeset
|
177 |
fun unregister_term_postprocessor T (_ : morphism) = |
9f98107ad8b4
use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents:
38240
diff
changeset
|
178 |
unregister_term_postprocessor_generic T |
38240
a44d108a8d39
local versions of Nitpick.register_xxx functions
blanchet
parents:
38209
diff
changeset
|
179 |
val unregister_term_postprocessor_global = |
a44d108a8d39
local versions of Nitpick.register_xxx functions
blanchet
parents:
38209
diff
changeset
|
180 |
Context.theory_map o unregister_term_postprocessor_generic |
35711
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35665
diff
changeset
|
181 |
|
33192 | 182 |
fun tuple_list_for_name rel_table bounds name = |
183 |
the (AList.lookup (op =) bounds (the_rel rel_table name)) handle NUT _ => [[]] |
|
184 |
||
41052
3db267a01c1d
remove the "fin_fun" optimization in Nitpick -- it was always a hack and didn't help much
blanchet
parents:
41039
diff
changeset
|
185 |
fun unarize_unbox_etc_term (Const (@{const_name FunBox}, _) $ t1) = |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
186 |
unarize_unbox_etc_term t1 |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
187 |
| unarize_unbox_etc_term |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
188 |
(Const (@{const_name PairBox}, |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
189 |
Type (@{type_name fun}, [T1, Type (@{type_name fun}, [T2, _])])) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
190 |
$ t1 $ t2) = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
191 |
let val Ts = map uniterize_unarize_unbox_etc_type [T1, T2] in |
38190
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents:
38189
diff
changeset
|
192 |
Const (@{const_name Pair}, Ts ---> Type (@{type_name prod}, Ts)) |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
193 |
$ unarize_unbox_etc_term t1 $ unarize_unbox_etc_term t2 |
33192 | 194 |
end |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
195 |
| unarize_unbox_etc_term (Const (s, T)) = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
196 |
Const (s, uniterize_unarize_unbox_etc_type T) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
197 |
| unarize_unbox_etc_term (t1 $ t2) = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
198 |
unarize_unbox_etc_term t1 $ unarize_unbox_etc_term t2 |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
199 |
| unarize_unbox_etc_term (Free (s, T)) = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
200 |
Free (s, uniterize_unarize_unbox_etc_type T) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
201 |
| unarize_unbox_etc_term (Var (x, T)) = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
202 |
Var (x, uniterize_unarize_unbox_etc_type T) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
203 |
| unarize_unbox_etc_term (Bound j) = Bound j |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
204 |
| unarize_unbox_etc_term (Abs (s, T, t')) = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
205 |
Abs (s, uniterize_unarize_unbox_etc_type T, unarize_unbox_etc_term t') |
33192 | 206 |
|
38190
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents:
38189
diff
changeset
|
207 |
fun factor_out_types (T1 as Type (@{type_name prod}, [T11, T12])) |
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents:
38189
diff
changeset
|
208 |
(T2 as Type (@{type_name prod}, [T21, T22])) = |
33192 | 209 |
let val (n1, n2) = pairself num_factors_in_type (T11, T21) in |
210 |
if n1 = n2 then |
|
211 |
let |
|
212 |
val ((T11', opt_T12'), (T21', opt_T22')) = factor_out_types T12 T22 |
|
213 |
in |
|
38190
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents:
38189
diff
changeset
|
214 |
((Type (@{type_name prod}, [T11, T11']), opt_T12'), |
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents:
38189
diff
changeset
|
215 |
(Type (@{type_name prod}, [T21, T21']), opt_T22')) |
33192 | 216 |
end |
217 |
else if n1 < n2 then |
|
218 |
case factor_out_types T1 T21 of |
|
219 |
(p1, (T21', NONE)) => (p1, (T21', SOME T22)) |
|
220 |
| (p1, (T21', SOME T22')) => |
|
38190
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents:
38189
diff
changeset
|
221 |
(p1, (T21', SOME (Type (@{type_name prod}, [T22', T22])))) |
33192 | 222 |
else |
223 |
swap (factor_out_types T2 T1) |
|
224 |
end |
|
38190
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents:
38189
diff
changeset
|
225 |
| factor_out_types (Type (@{type_name prod}, [T11, T12])) T2 = |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
226 |
((T11, SOME T12), (T2, NONE)) |
38190
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents:
38189
diff
changeset
|
227 |
| factor_out_types T1 (Type (@{type_name prod}, [T21, T22])) = |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
228 |
((T1, NONE), (T21, SOME T22)) |
33192 | 229 |
| factor_out_types T1 T2 = ((T1, NONE), (T2, NONE)) |
230 |
||
231 |
fun make_plain_fun maybe_opt T1 T2 = |
|
232 |
let |
|
233 |
fun aux T1 T2 [] = |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
234 |
Const (if maybe_opt then opt_flag else non_opt_flag, T1 --> T2) |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
235 |
| aux T1 T2 ((t1, t2) :: tps) = |
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
236 |
Const (@{const_name fun_upd}, (T1 --> T2) --> T1 --> T2 --> T1 --> T2) |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
237 |
$ aux T1 T2 tps $ t1 $ t2 |
33192 | 238 |
in aux T1 T2 o rev end |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
239 |
fun is_plain_fun (Const (s, _)) = (s = opt_flag orelse s = non_opt_flag) |
33192 | 240 |
| is_plain_fun (Const (@{const_name fun_upd}, _) $ t0 $ _ $ _) = |
241 |
is_plain_fun t0 |
|
242 |
| is_plain_fun _ = false |
|
243 |
val dest_plain_fun = |
|
244 |
let |
|
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
245 |
fun aux (Abs (_, _, Const (s, _))) = (s <> irrelevant, ([], [])) |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
246 |
| aux (Const (s, _)) = (s <> non_opt_flag, ([], [])) |
33192 | 247 |
| aux (Const (@{const_name fun_upd}, _) $ t0 $ t1 $ t2) = |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
248 |
let val (maybe_opt, (ts1, ts2)) = aux t0 in |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
249 |
(maybe_opt, (t1 :: ts1, t2 :: ts2)) |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
250 |
end |
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
251 |
| aux t = raise TERM ("Nitpick_Model.dest_plain_fun", [t]) |
33192 | 252 |
in apsnd (pairself rev) o aux end |
253 |
||
33565
5fad8e36dfb1
fixed error in Nitpick's display of uncurried constants, which resulted in an exception
blanchet
parents:
33558
diff
changeset
|
254 |
fun break_in_two T T1 T2 t = |
5fad8e36dfb1
fixed error in Nitpick's display of uncurried constants, which resulted in an exception
blanchet
parents:
33558
diff
changeset
|
255 |
let |
5fad8e36dfb1
fixed error in Nitpick's display of uncurried constants, which resulted in an exception
blanchet
parents:
33558
diff
changeset
|
256 |
val ps = HOLogic.flat_tupleT_paths T |
5fad8e36dfb1
fixed error in Nitpick's display of uncurried constants, which resulted in an exception
blanchet
parents:
33558
diff
changeset
|
257 |
val cut = length (HOLogic.strip_tupleT T1) |
5fad8e36dfb1
fixed error in Nitpick's display of uncurried constants, which resulted in an exception
blanchet
parents:
33558
diff
changeset
|
258 |
val (ps1, ps2) = pairself HOLogic.flat_tupleT_paths (T1, T2) |
5fad8e36dfb1
fixed error in Nitpick's display of uncurried constants, which resulted in an exception
blanchet
parents:
33558
diff
changeset
|
259 |
val (ts1, ts2) = t |> HOLogic.strip_ptuple ps |> chop cut |
5fad8e36dfb1
fixed error in Nitpick's display of uncurried constants, which resulted in an exception
blanchet
parents:
33558
diff
changeset
|
260 |
in (HOLogic.mk_ptuple ps1 T1 ts1, HOLogic.mk_ptuple ps2 T2 ts2) end |
38190
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents:
38189
diff
changeset
|
261 |
fun pair_up (Type (@{type_name prod}, [T1', T2'])) |
33192 | 262 |
(t1 as Const (@{const_name Pair}, |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
263 |
Type (@{type_name fun}, |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
264 |
[_, Type (@{type_name fun}, [_, T1])])) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
265 |
$ t11 $ t12) t2 = |
33192 | 266 |
if T1 = T1' then HOLogic.mk_prod (t1, t2) |
267 |
else HOLogic.mk_prod (t11, pair_up T2' t12 t2) |
|
268 |
| pair_up _ t1 t2 = HOLogic.mk_prod (t1, t2) |
|
269 |
fun multi_pair_up T1 t1 (ts2, ts3) = map2 (pair o pair_up T1 t1) ts2 ts3 |
|
270 |
||
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
271 |
fun typecast_fun (Type (@{type_name fun}, [T1', T2'])) T1 T2 t = |
33192 | 272 |
let |
33565
5fad8e36dfb1
fixed error in Nitpick's display of uncurried constants, which resulted in an exception
blanchet
parents:
33558
diff
changeset
|
273 |
fun do_curry T1 T1a T1b T2 t = |
33192 | 274 |
let |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
275 |
val (maybe_opt, tsp) = dest_plain_fun t |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
276 |
val tps = |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
277 |
tsp |>> map (break_in_two T1 T1a T1b) |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
278 |
|> uncurry (map2 (fn (t1a, t1b) => fn t2 => (t1a, (t1b, t2)))) |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
279 |
|> AList.coalesce (op =) |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
280 |
|> map (apsnd (make_plain_fun maybe_opt T1b T2)) |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
281 |
in make_plain_fun maybe_opt T1a (T1b --> T2) tps end |
33192 | 282 |
and do_uncurry T1 T2 t = |
283 |
let |
|
284 |
val (maybe_opt, tsp) = dest_plain_fun t |
|
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
285 |
val tps = |
33192 | 286 |
tsp |> op ~~ |
287 |
|> maps (fn (t1, t2) => |
|
288 |
multi_pair_up T1 t1 (snd (dest_plain_fun t2))) |
|
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
289 |
in make_plain_fun maybe_opt T1 T2 tps end |
33192 | 290 |
and do_arrow T1' T2' _ _ (Const (s, _)) = Const (s, T1' --> T2') |
291 |
| do_arrow T1' T2' T1 T2 |
|
292 |
(Const (@{const_name fun_upd}, _) $ t0 $ t1 $ t2) = |
|
293 |
Const (@{const_name fun_upd}, |
|
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
294 |
(T1' --> T2') --> T1' --> T2' --> T1' --> T2') |
33192 | 295 |
$ do_arrow T1' T2' T1 T2 t0 $ do_term T1' T1 t1 $ do_term T2' T2 t2 |
296 |
| do_arrow _ _ _ _ t = |
|
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
297 |
raise TERM ("Nitpick_Model.typecast_fun.do_arrow", [t]) |
33192 | 298 |
and do_fun T1' T2' T1 T2 t = |
299 |
case factor_out_types T1' T1 of |
|
300 |
((_, NONE), (_, NONE)) => t |> do_arrow T1' T2' T1 T2 |
|
301 |
| ((_, NONE), (T1a, SOME T1b)) => |
|
33565
5fad8e36dfb1
fixed error in Nitpick's display of uncurried constants, which resulted in an exception
blanchet
parents:
33558
diff
changeset
|
302 |
t |> do_curry T1 T1a T1b T2 |> do_arrow T1' T2' T1a (T1b --> T2) |
33192 | 303 |
| ((T1a', SOME T1b'), (_, NONE)) => |
304 |
t |> do_arrow T1a' (T1b' --> T2') T1 T2 |> do_uncurry T1' T2' |
|
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
305 |
| _ => raise TYPE ("Nitpick_Model.typecast_fun.do_fun", [T1, T1'], []) |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
306 |
and do_term (Type (@{type_name fun}, [T1', T2'])) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
307 |
(Type (@{type_name fun}, [T1, T2])) t = |
33192 | 308 |
do_fun T1' T2' T1 T2 t |
38190
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents:
38189
diff
changeset
|
309 |
| do_term (T' as Type (@{type_name prod}, Ts' as [T1', T2'])) |
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents:
38189
diff
changeset
|
310 |
(Type (@{type_name prod}, [T1, T2])) |
33192 | 311 |
(Const (@{const_name Pair}, _) $ t1 $ t2) = |
312 |
Const (@{const_name Pair}, Ts' ---> T') |
|
313 |
$ do_term T1' T1 t1 $ do_term T2' T2 t2 |
|
314 |
| do_term T' T t = |
|
315 |
if T = T' then t |
|
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
316 |
else raise TYPE ("Nitpick_Model.typecast_fun.do_term", [T, T'], []) |
33192 | 317 |
in if T1' = T1 andalso T2' = T2 then t else do_fun T1' T2' T1 T2 t end |
34998 | 318 |
| typecast_fun T' _ _ _ = |
319 |
raise TYPE ("Nitpick_Model.typecast_fun", [T'], []) |
|
33192 | 320 |
|
321 |
fun truth_const_sort_key @{const True} = "0" |
|
322 |
| truth_const_sort_key @{const False} = "2" |
|
323 |
| truth_const_sort_key _ = "1" |
|
324 |
||
38190
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents:
38189
diff
changeset
|
325 |
fun mk_tuple (Type (@{type_name prod}, [T1, T2])) ts = |
33192 | 326 |
HOLogic.mk_prod (mk_tuple T1 ts, |
327 |
mk_tuple T2 (List.drop (ts, length (HOLogic.flatten_tupleT T1)))) |
|
328 |
| mk_tuple _ (t :: _) = t |
|
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
329 |
| mk_tuple T [] = raise TYPE ("Nitpick_Model.mk_tuple", [T], []) |
33192 | 330 |
|
38240
a44d108a8d39
local versions of Nitpick.register_xxx functions
blanchet
parents:
38209
diff
changeset
|
331 |
fun varified_type_match ctxt (candid_T, pat_T) = |
42361 | 332 |
let val thy = Proof_Context.theory_of ctxt in |
38240
a44d108a8d39
local versions of Nitpick.register_xxx functions
blanchet
parents:
38209
diff
changeset
|
333 |
strict_type_match thy (candid_T, varify_type ctxt pat_T) |
a44d108a8d39
local versions of Nitpick.register_xxx functions
blanchet
parents:
38209
diff
changeset
|
334 |
end |
35711
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35665
diff
changeset
|
335 |
|
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
336 |
fun all_values_of_type pool wacky_names (scope as {card_assigns, ...} : scope) |
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
337 |
atomss sel_names rel_table bounds card T = |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
338 |
let |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
339 |
val card = if card = 0 then card_of_type card_assigns T else card |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
340 |
fun nth_value_of_type n = |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
341 |
let |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
342 |
fun term unfold = |
37262
c0fe8fa35771
don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents:
37261
diff
changeset
|
343 |
reconstruct_term true unfold pool wacky_names scope atomss sel_names |
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
344 |
rel_table bounds T T (Atom (card, 0)) [[n]] |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
345 |
in |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
346 |
case term false of |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
347 |
t as Const (s, _) => |
37261 | 348 |
if String.isPrefix (cyclic_const_prefix ()) s then |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
349 |
HOLogic.mk_eq (t, term true) |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
350 |
else |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
351 |
t |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
352 |
| t => t |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
353 |
end |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
354 |
in index_seq 0 card |> map nth_value_of_type |> sort nice_term_ord end |
37262
c0fe8fa35771
don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents:
37261
diff
changeset
|
355 |
and reconstruct_term maybe_opt unfold pool |
c0fe8fa35771
don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents:
37261
diff
changeset
|
356 |
(wacky_names as ((maybe_name, abs_name), _)) |
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
37170
diff
changeset
|
357 |
(scope as {hol_ctxt as {thy, ctxt, stds, ...}, binarize, card_assigns, |
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
358 |
bits, datatypes, ofs, ...}) |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
359 |
atomss sel_names rel_table bounds = |
33192 | 360 |
let |
361 |
val for_auto = (maybe_name = "") |
|
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
362 |
fun value_of_bits jss = |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
363 |
let |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
364 |
val j0 = offset_of_type ofs @{typ unsigned_bit} |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
365 |
val js = map (Integer.add (~ j0) o the_single) jss |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
366 |
in |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
367 |
fold (fn j => Integer.add (reasonable_power 2 j |> j = bits ? op ~)) |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
368 |
js 0 |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
369 |
end |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
370 |
val all_values = |
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
371 |
all_values_of_type pool wacky_names scope atomss sel_names rel_table |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
372 |
bounds 0 |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
373 |
fun postprocess_term (Type (@{type_name fun}, _)) = I |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
374 |
| postprocess_term T = |
38240
a44d108a8d39
local versions of Nitpick.register_xxx functions
blanchet
parents:
38209
diff
changeset
|
375 |
case Data.get (Context.Proof ctxt) of |
a44d108a8d39
local versions of Nitpick.register_xxx functions
blanchet
parents:
38209
diff
changeset
|
376 |
[] => I |
a44d108a8d39
local versions of Nitpick.register_xxx functions
blanchet
parents:
38209
diff
changeset
|
377 |
| postprocs => |
a44d108a8d39
local versions of Nitpick.register_xxx functions
blanchet
parents:
38209
diff
changeset
|
378 |
case AList.lookup (varified_type_match ctxt) postprocs T of |
a44d108a8d39
local versions of Nitpick.register_xxx functions
blanchet
parents:
38209
diff
changeset
|
379 |
SOME postproc => postproc ctxt maybe_name all_values T |
a44d108a8d39
local versions of Nitpick.register_xxx functions
blanchet
parents:
38209
diff
changeset
|
380 |
| NONE => I |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
381 |
fun postprocess_subterms Ts (t1 $ t2) = |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
382 |
let val t = postprocess_subterms Ts t1 $ postprocess_subterms Ts t2 in |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
383 |
postprocess_term (fastype_of1 (Ts, t)) t |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
384 |
end |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
385 |
| postprocess_subterms Ts (Abs (s, T, t')) = |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
386 |
Abs (s, T, postprocess_subterms (T :: Ts) t') |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
387 |
| postprocess_subterms Ts t = postprocess_term (fastype_of1 (Ts, t)) t |
35388
42d39948cace
use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents:
35385
diff
changeset
|
388 |
fun make_set maybe_opt T1 T2 tps = |
33192 | 389 |
let |
35402 | 390 |
val empty_const = Const (@{const_abbrev Set.empty}, T1 --> T2) |
33192 | 391 |
val insert_const = Const (@{const_name insert}, |
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
392 |
T1 --> (T1 --> T2) --> T1 --> T2) |
33192 | 393 |
fun aux [] = |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
394 |
if maybe_opt andalso not (is_complete_type datatypes false T1) then |
37261 | 395 |
insert_const $ Const (unrep (), T1) $ empty_const |
33192 | 396 |
else |
397 |
empty_const |
|
398 |
| aux ((t1, t2) :: zs) = |
|
35388
42d39948cace
use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents:
35385
diff
changeset
|
399 |
aux zs |
42d39948cace
use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents:
35385
diff
changeset
|
400 |
|> t2 <> @{const False} |
42d39948cace
use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents:
35385
diff
changeset
|
401 |
? curry (op $) |
42d39948cace
use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents:
35385
diff
changeset
|
402 |
(insert_const |
42d39948cace
use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents:
35385
diff
changeset
|
403 |
$ (t1 |> t2 <> @{const True} |
42d39948cace
use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents:
35385
diff
changeset
|
404 |
? curry (op $) |
42d39948cace
use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents:
35385
diff
changeset
|
405 |
(Const (maybe_name, T1 --> T1)))) |
42d39948cace
use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents:
35385
diff
changeset
|
406 |
in |
42d39948cace
use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents:
35385
diff
changeset
|
407 |
if forall (fn (_, t) => t <> @{const True} andalso t <> @{const False}) |
42d39948cace
use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents:
35385
diff
changeset
|
408 |
tps then |
42d39948cace
use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents:
35385
diff
changeset
|
409 |
Const (unknown, T1 --> T2) |
42d39948cace
use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents:
35385
diff
changeset
|
410 |
else |
42d39948cace
use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents:
35385
diff
changeset
|
411 |
aux tps |
42d39948cace
use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents:
35385
diff
changeset
|
412 |
end |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
413 |
fun make_map maybe_opt T1 T2 T2' = |
33192 | 414 |
let |
415 |
val update_const = Const (@{const_name fun_upd}, |
|
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
416 |
(T1 --> T2) --> T1 --> T2 --> T1 --> T2) |
35402 | 417 |
fun aux' [] = Const (@{const_abbrev Map.empty}, T1 --> T2) |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
418 |
| aux' ((t1, t2) :: tps) = |
33192 | 419 |
(case t2 of |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
420 |
Const (@{const_name None}, _) => aux' tps |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
421 |
| _ => update_const $ aux' tps $ t1 $ t2) |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
422 |
fun aux tps = |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
423 |
if maybe_opt andalso not (is_complete_type datatypes false T1) then |
37261 | 424 |
update_const $ aux' tps $ Const (unrep (), T1) |
33192 | 425 |
$ (Const (@{const_name Some}, T2' --> T2) $ Const (unknown, T2')) |
426 |
else |
|
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
427 |
aux' tps |
33192 | 428 |
in aux end |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
429 |
fun polish_funs Ts t = |
33192 | 430 |
(case fastype_of1 (Ts, t) of |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
431 |
Type (@{type_name fun}, [T1, T2]) => |
33192 | 432 |
if is_plain_fun t then |
433 |
case T2 of |
|
434 |
@{typ bool} => |
|
435 |
let |
|
436 |
val (maybe_opt, ts_pair) = |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
437 |
dest_plain_fun t ||> pairself (map (polish_funs Ts)) |
33192 | 438 |
in |
439 |
make_set maybe_opt T1 T2 |
|
440 |
(sort_wrt (truth_const_sort_key o snd) (op ~~ ts_pair)) |
|
441 |
end |
|
442 |
| Type (@{type_name option}, [T2']) => |
|
443 |
let |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
444 |
val (maybe_opt, ts_pair) = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
445 |
dest_plain_fun t ||> pairself (map (polish_funs Ts)) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
446 |
in make_map maybe_opt T1 T2 T2' (rev (op ~~ ts_pair)) end |
33192 | 447 |
| _ => raise SAME () |
448 |
else |
|
449 |
raise SAME () |
|
450 |
| _ => raise SAME ()) |
|
451 |
handle SAME () => |
|
452 |
case t of |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
453 |
(t1 as Const (@{const_name fun_upd}, _) $ t11 $ _) |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
454 |
$ (t2 as Const (s, _)) => |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
455 |
if s = unknown then polish_funs Ts t11 |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
456 |
else polish_funs Ts t1 $ polish_funs Ts t2 |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
457 |
| t1 $ t2 => polish_funs Ts t1 $ polish_funs Ts t2 |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
458 |
| Abs (s, T, t') => Abs (s, T, polish_funs (T :: Ts) t') |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
459 |
| Const (s, Type (@{type_name fun}, [T1, T2])) => |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
460 |
if s = opt_flag orelse s = non_opt_flag then |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
461 |
Abs ("x", T1, |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
462 |
Const (if is_complete_type datatypes false T1 then |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
463 |
irrelevant |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
464 |
else |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
465 |
unknown, T2)) |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
466 |
else |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
467 |
t |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
468 |
| t => t |
33192 | 469 |
fun make_fun maybe_opt T1 T2 T' ts1 ts2 = |
34126 | 470 |
ts1 ~~ ts2 |> sort (nice_term_ord o pairself fst) |
34998 | 471 |
|> make_plain_fun maybe_opt T1 T2 |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
472 |
|> unarize_unbox_etc_term |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
473 |
|> typecast_fun (uniterize_unarize_unbox_etc_type T') |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
474 |
(uniterize_unarize_unbox_etc_type T1) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
475 |
(uniterize_unarize_unbox_etc_type T2) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
476 |
fun term_for_atom seen (T as Type (@{type_name fun}, [T1, T2])) T' j _ = |
33192 | 477 |
let |
478 |
val k1 = card_of_type card_assigns T1 |
|
479 |
val k2 = card_of_type card_assigns T2 |
|
480 |
in |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
481 |
term_for_rep true seen T T' (Vect (k1, Atom (k2, 0))) |
33192 | 482 |
[nth_combination (replicate k1 (k2, 0)) j] |
483 |
handle General.Subscript => |
|
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
484 |
raise ARG ("Nitpick_Model.reconstruct_term.term_for_atom", |
33192 | 485 |
signed_string_of_int j ^ " for " ^ |
486 |
string_for_rep (Vect (k1, Atom (k2, 0)))) |
|
487 |
end |
|
38190
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents:
38189
diff
changeset
|
488 |
| term_for_atom seen (Type (@{type_name prod}, [T1, T2])) _ j k = |
35075
888802be2019
handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents:
35070
diff
changeset
|
489 |
let |
888802be2019
handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents:
35070
diff
changeset
|
490 |
val k1 = card_of_type card_assigns T1 |
888802be2019
handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents:
35070
diff
changeset
|
491 |
val k2 = k div k1 |
888802be2019
handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents:
35070
diff
changeset
|
492 |
in |
33192 | 493 |
list_comb (HOLogic.pair_const T1 T2, |
35075
888802be2019
handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents:
35070
diff
changeset
|
494 |
map3 (fn T => term_for_atom seen T T) [T1, T2] |
888802be2019
handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents:
35070
diff
changeset
|
495 |
[j div k2, j mod k2] [k1, k2]) (* ### k2 or k1? FIXME *) |
33192 | 496 |
end |
35075
888802be2019
handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents:
35070
diff
changeset
|
497 |
| term_for_atom seen @{typ prop} _ j k = |
888802be2019
handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents:
35070
diff
changeset
|
498 |
HOLogic.mk_Trueprop (term_for_atom seen bool_T bool_T j k) |
888802be2019
handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents:
35070
diff
changeset
|
499 |
| term_for_atom _ @{typ bool} _ j _ = |
33192 | 500 |
if j = 0 then @{const False} else @{const True} |
35075
888802be2019
handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents:
35070
diff
changeset
|
501 |
| term_for_atom seen T _ j k = |
35220
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35190
diff
changeset
|
502 |
if T = nat_T andalso is_standard_datatype thy stds nat_T then |
33192 | 503 |
HOLogic.mk_number nat_T j |
504 |
else if T = int_T then |
|
35075
888802be2019
handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents:
35070
diff
changeset
|
505 |
HOLogic.mk_number int_T (int_for_atom (k, 0) j) |
33192 | 506 |
else if is_fp_iterator_type T then |
35075
888802be2019
handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents:
35070
diff
changeset
|
507 |
HOLogic.mk_number nat_T (k - j - 1) |
33192 | 508 |
else if T = @{typ bisim_iterator} then |
509 |
HOLogic.mk_number nat_T j |
|
510 |
else case datatype_spec datatypes T of |
|
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
511 |
NONE => nth_atom thy atomss pool for_auto T j |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
512 |
| SOME {deep = false, ...} => nth_atom thy atomss pool for_auto T j |
35179
4b198af5beb5
redo Nitpick's nonstandard values as cyclic values (instead of additional constructors)
blanchet
parents:
35177
diff
changeset
|
513 |
| SOME {co, standard, constrs, ...} => |
33192 | 514 |
let |
515 |
fun tuples_for_const (s, T) = |
|
516 |
tuple_list_for_name rel_table bounds (ConstName (s, T, Any)) |
|
35180
c57dba973391
more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents:
35179
diff
changeset
|
517 |
fun cyclic_atom () = |
37261 | 518 |
nth_atom thy atomss pool for_auto (Type (cyclic_type_name (), [])) |
519 |
j |
|
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
520 |
fun cyclic_var () = |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
521 |
Var ((nth_atom_name thy atomss pool "" T j, 0), T) |
33192 | 522 |
val discr_jsss = map (tuples_for_const o discr_for_constr o #const) |
523 |
constrs |
|
524 |
val real_j = j + offset_of_type ofs T |
|
525 |
val constr_x as (constr_s, constr_T) = |
|
526 |
get_first (fn (jss, {const, ...}) => |
|
34121
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents:
33982
diff
changeset
|
527 |
if member (op =) jss [real_j] then SOME const |
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents:
33982
diff
changeset
|
528 |
else NONE) |
33192 | 529 |
(discr_jsss ~~ constrs) |> the |
530 |
val arg_Ts = curried_binder_types constr_T |
|
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35188
diff
changeset
|
531 |
val sel_xs = |
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35188
diff
changeset
|
532 |
map (binarized_and_boxed_nth_sel_for_constr hol_ctxt binarize |
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35188
diff
changeset
|
533 |
constr_x) |
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35188
diff
changeset
|
534 |
(index_seq 0 (length arg_Ts)) |
33192 | 535 |
val sel_Rs = |
536 |
map (fn x => get_first |
|
537 |
(fn ConstName (s', T', R) => |
|
538 |
if (s', T') = x then SOME R else NONE |
|
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
539 |
| u => raise NUT ("Nitpick_Model.reconstruct_\ |
33192 | 540 |
\term.term_for_atom", [u])) |
541 |
sel_names |> the) sel_xs |
|
542 |
val arg_Rs = map (snd o dest_Func) sel_Rs |
|
543 |
val sel_jsss = map tuples_for_const sel_xs |
|
544 |
val arg_jsss = |
|
545 |
map (map_filter (fn js => if hd js = real_j then SOME (tl js) |
|
546 |
else NONE)) sel_jsss |
|
547 |
val uncur_arg_Ts = binder_types constr_T |
|
35179
4b198af5beb5
redo Nitpick's nonstandard values as cyclic values (instead of additional constructors)
blanchet
parents:
35177
diff
changeset
|
548 |
val maybe_cyclic = co orelse not standard |
33192 | 549 |
in |
35180
c57dba973391
more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents:
35179
diff
changeset
|
550 |
if maybe_cyclic andalso not (null seen) andalso |
35188
8c70a34931b1
improve Nitpick's "Datatypes" rendering for elements containing cycles
blanchet
parents:
35180
diff
changeset
|
551 |
member (op =) (seen |> unfold ? (fst o split_last)) (T, j) then |
35180
c57dba973391
more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents:
35179
diff
changeset
|
552 |
cyclic_var () |
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
553 |
else if constr_s = @{const_name Word} then |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
554 |
HOLogic.mk_number |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
555 |
(if T = @{typ "unsigned_bit word"} then nat_T else int_T) |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
556 |
(value_of_bits (the_single arg_jsss)) |
33192 | 557 |
else |
558 |
let |
|
35179
4b198af5beb5
redo Nitpick's nonstandard values as cyclic values (instead of additional constructors)
blanchet
parents:
35177
diff
changeset
|
559 |
val seen = seen |> maybe_cyclic ? cons (T, j) |
33192 | 560 |
val ts = |
561 |
if length arg_Ts = 0 then |
|
562 |
[] |
|
563 |
else |
|
41052
3db267a01c1d
remove the "fin_fun" optimization in Nitpick -- it was always a hack and didn't help much
blanchet
parents:
41039
diff
changeset
|
564 |
map3 (fn Ts => term_for_rep true seen Ts Ts) arg_Ts arg_Rs |
3db267a01c1d
remove the "fin_fun" optimization in Nitpick -- it was always a hack and didn't help much
blanchet
parents:
41039
diff
changeset
|
565 |
arg_jsss |
33192 | 566 |
|> mk_tuple (HOLogic.mk_tupleT uncur_arg_Ts) |
567 |
|> dest_n_tuple (length uncur_arg_Ts) |
|
568 |
val t = |
|
569 |
if constr_s = @{const_name Abs_Frac} then |
|
35711
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35665
diff
changeset
|
570 |
case ts of |
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35665
diff
changeset
|
571 |
[Const (@{const_name Pair}, _) $ t1 $ t2] => |
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35665
diff
changeset
|
572 |
frac_from_term_pair (body_type T) t1 t2 |
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35665
diff
changeset
|
573 |
| _ => raise TERM ("Nitpick_Model.reconstruct_term.\ |
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents:
35665
diff
changeset
|
574 |
\term_for_atom (Abs_Frac)", ts) |
34936
c4f04bee79f3
some work on Nitpick's support for quotient types;
blanchet
parents:
34126
diff
changeset
|
575 |
else if not for_auto andalso |
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
37170
diff
changeset
|
576 |
(is_abs_fun ctxt constr_x orelse |
34936
c4f04bee79f3
some work on Nitpick's support for quotient types;
blanchet
parents:
34126
diff
changeset
|
577 |
constr_s = @{const_name Quot}) then |
33192 | 578 |
Const (abs_name, constr_T) $ the_single ts |
579 |
else |
|
580 |
list_comb (Const constr_x, ts) |
|
581 |
in |
|
35179
4b198af5beb5
redo Nitpick's nonstandard values as cyclic values (instead of additional constructors)
blanchet
parents:
35177
diff
changeset
|
582 |
if maybe_cyclic then |
35180
c57dba973391
more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents:
35179
diff
changeset
|
583 |
let val var = cyclic_var () in |
35188
8c70a34931b1
improve Nitpick's "Datatypes" rendering for elements containing cycles
blanchet
parents:
35180
diff
changeset
|
584 |
if unfold andalso not standard andalso |
35180
c57dba973391
more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents:
35179
diff
changeset
|
585 |
length seen = 1 andalso |
37261 | 586 |
exists_subterm |
587 |
(fn Const (s, _) => |
|
588 |
String.isPrefix (cyclic_const_prefix ()) s |
|
589 |
| t' => t' = var) t then |
|
35188
8c70a34931b1
improve Nitpick's "Datatypes" rendering for elements containing cycles
blanchet
parents:
35180
diff
changeset
|
590 |
subst_atomic [(var, cyclic_atom ())] t |
35180
c57dba973391
more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents:
35179
diff
changeset
|
591 |
else if exists_subterm (curry (op =) var) t then |
35179
4b198af5beb5
redo Nitpick's nonstandard values as cyclic values (instead of additional constructors)
blanchet
parents:
35177
diff
changeset
|
592 |
if co then |
4b198af5beb5
redo Nitpick's nonstandard values as cyclic values (instead of additional constructors)
blanchet
parents:
35177
diff
changeset
|
593 |
Const (@{const_name The}, (T --> bool_T) --> T) |
37261 | 594 |
$ Abs (cyclic_co_val_name (), T, |
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38284
diff
changeset
|
595 |
Const (@{const_name HOL.eq}, T --> T --> bool_T) |
35180
c57dba973391
more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents:
35179
diff
changeset
|
596 |
$ Bound 0 $ abstract_over (var, t)) |
35179
4b198af5beb5
redo Nitpick's nonstandard values as cyclic values (instead of additional constructors)
blanchet
parents:
35177
diff
changeset
|
597 |
else |
35180
c57dba973391
more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents:
35179
diff
changeset
|
598 |
cyclic_atom () |
33192 | 599 |
else |
600 |
t |
|
601 |
end |
|
602 |
else |
|
603 |
t |
|
604 |
end |
|
605 |
end |
|
606 |
and term_for_vect seen k R T1 T2 T' js = |
|
35075
888802be2019
handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents:
35070
diff
changeset
|
607 |
make_fun true T1 T2 T' |
888802be2019
handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents:
35070
diff
changeset
|
608 |
(map (fn j => term_for_atom seen T1 T1 j k) (index_seq 0 k)) |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
609 |
(map (term_for_rep true seen T2 T2 R o single) |
33192 | 610 |
(batch_list (arity_of_rep R) js)) |
38190
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents:
38189
diff
changeset
|
611 |
and term_for_rep _ seen T T' (R as Atom (k, j0)) [[j]] = |
35075
888802be2019
handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents:
35070
diff
changeset
|
612 |
if j >= j0 andalso j < j0 + k then term_for_atom seen T T' (j - j0) k |
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
613 |
else raise REP ("Nitpick_Model.reconstruct_term.term_for_rep", [R]) |
38190
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents:
38189
diff
changeset
|
614 |
| term_for_rep _ seen (Type (@{type_name prod}, [T1, T2])) _ |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
615 |
(Struct [R1, R2]) [js] = |
33192 | 616 |
let |
617 |
val arity1 = arity_of_rep R1 |
|
618 |
val (js1, js2) = chop arity1 js |
|
619 |
in |
|
620 |
list_comb (HOLogic.pair_const T1 T2, |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
621 |
map3 (fn T => term_for_rep true seen T T) [T1, T2] [R1, R2] |
33192 | 622 |
[[js1], [js2]]) |
623 |
end |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
624 |
| term_for_rep _ seen (Type (@{type_name fun}, [T1, T2])) T' |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
625 |
(Vect (k, R')) [js] = |
33192 | 626 |
term_for_vect seen k R' T1 T2 T' js |
37170
38ba15040455
Nitpick: show "..." in datatype values (e.g., [{0::nat, ...}]), since these are really equivalence classes
blanchet
parents:
36607
diff
changeset
|
627 |
| term_for_rep maybe_opt seen (Type (@{type_name fun}, [T1, T2])) T' |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
628 |
(Func (R1, Formula Neut)) jss = |
33192 | 629 |
let |
630 |
val jss1 = all_combinations_for_rep R1 |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
631 |
val ts1 = map (term_for_rep true seen T1 T1 R1 o single) jss1 |
33192 | 632 |
val ts2 = |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
633 |
map (fn js => term_for_rep true seen T2 T2 (Atom (2, 0)) |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
634 |
[[int_from_bool (member (op =) jss js)]]) |
34121
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents:
33982
diff
changeset
|
635 |
jss1 |
37170
38ba15040455
Nitpick: show "..." in datatype values (e.g., [{0::nat, ...}]), since these are really equivalence classes
blanchet
parents:
36607
diff
changeset
|
636 |
in make_fun maybe_opt T1 T2 T' ts1 ts2 end |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
637 |
| term_for_rep maybe_opt seen (Type (@{type_name fun}, [T1, T2])) T' |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
638 |
(Func (R1, R2)) jss = |
33192 | 639 |
let |
640 |
val arity1 = arity_of_rep R1 |
|
641 |
val jss1 = all_combinations_for_rep R1 |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
642 |
val ts1 = map (term_for_rep false seen T1 T1 R1 o single) jss1 |
33192 | 643 |
val grouped_jss2 = AList.group (op =) (map (chop arity1) jss) |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
644 |
val ts2 = map (term_for_rep false seen T2 T2 R2 o the_default [] |
33192 | 645 |
o AList.lookup (op =) grouped_jss2) jss1 |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
646 |
in make_fun maybe_opt T1 T2 T' ts1 ts2 end |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
647 |
| term_for_rep _ seen T T' (Opt R) jss = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
648 |
if null jss then Const (unknown, T) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
649 |
else term_for_rep true seen T T' R jss |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
650 |
| term_for_rep _ _ T _ R jss = |
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
651 |
raise ARG ("Nitpick_Model.reconstruct_term.term_for_rep", |
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
652 |
Syntax.string_of_typ ctxt T ^ " " ^ string_for_rep R ^ " " ^ |
33192 | 653 |
string_of_int (length jss)) |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
654 |
in |
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
655 |
postprocess_subterms [] o polish_funs [] o unarize_unbox_etc_term |
37262
c0fe8fa35771
don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents:
37261
diff
changeset
|
656 |
oooo term_for_rep maybe_opt [] |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
657 |
end |
33192 | 658 |
|
35718 | 659 |
(** Constant postprocessing **) |
660 |
||
661 |
fun dest_n_tuple_type 1 T = [T] |
|
662 |
| dest_n_tuple_type n (Type (_, [T1, T2])) = |
|
663 |
T1 :: dest_n_tuple_type (n - 1) T2 |
|
664 |
| dest_n_tuple_type _ T = |
|
665 |
raise TYPE ("Nitpick_Model.dest_n_tuple_type", [T], []) |
|
666 |
||
41791
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
blanchet
parents:
41472
diff
changeset
|
667 |
fun const_format thy def_tables (x as (s, T)) = |
35718 | 668 |
if String.isPrefix unrolled_prefix s then |
41791
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
blanchet
parents:
41472
diff
changeset
|
669 |
const_format thy def_tables (original_name s, range_type T) |
35718 | 670 |
else if String.isPrefix skolem_prefix s then |
671 |
let |
|
672 |
val k = unprefix skolem_prefix s |
|
673 |
|> strip_first_name_sep |> fst |> space_explode "@" |
|
674 |
|> hd |> Int.fromString |> the |
|
675 |
in [k, num_binder_types T - k] end |
|
676 |
else if original_name s <> s then |
|
677 |
[num_binder_types T] |
|
41791
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
blanchet
parents:
41472
diff
changeset
|
678 |
else case def_of_const thy def_tables x of |
35718 | 679 |
SOME t' => if fixpoint_kind_of_rhs t' <> NoFp then |
680 |
let val k = length (strip_abs_vars t') in |
|
681 |
[k, num_binder_types T - k] |
|
682 |
end |
|
683 |
else |
|
684 |
[num_binder_types T] |
|
685 |
| NONE => [num_binder_types T] |
|
686 |
fun intersect_formats _ [] = [] |
|
687 |
| intersect_formats [] _ = [] |
|
688 |
| intersect_formats ks1 ks2 = |
|
689 |
let val ((ks1', k1), (ks2', k2)) = pairself split_last (ks1, ks2) in |
|
690 |
intersect_formats (ks1' @ (if k1 > k2 then [k1 - k2] else [])) |
|
691 |
(ks2' @ (if k2 > k1 then [k2 - k1] else [])) @ |
|
692 |
[Int.min (k1, k2)] |
|
693 |
end |
|
694 |
||
41791
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
blanchet
parents:
41472
diff
changeset
|
695 |
fun lookup_format thy def_tables formats t = |
35718 | 696 |
case AList.lookup (fn (SOME x, SOME y) => |
697 |
(term_match thy) (x, y) | _ => false) |
|
698 |
formats (SOME t) of |
|
699 |
SOME format => format |
|
700 |
| NONE => let val format = the (AList.lookup (op =) formats NONE) in |
|
701 |
case t of |
|
702 |
Const x => intersect_formats format |
|
41791
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
blanchet
parents:
41472
diff
changeset
|
703 |
(const_format thy def_tables x) |
35718 | 704 |
| _ => format |
705 |
end |
|
706 |
||
707 |
fun format_type default_format format T = |
|
708 |
let |
|
709 |
val T = uniterize_unarize_unbox_etc_type T |
|
710 |
val format = format |> filter (curry (op <) 0) |
|
711 |
in |
|
712 |
if forall (curry (op =) 1) format then |
|
713 |
T |
|
714 |
else |
|
715 |
let |
|
716 |
val (binder_Ts, body_T) = strip_type T |
|
717 |
val batched = |
|
718 |
binder_Ts |
|
719 |
|> map (format_type default_format default_format) |
|
720 |
|> rev |> chunk_list_unevenly (rev format) |
|
721 |
|> map (HOLogic.mk_tupleT o rev) |
|
722 |
in List.foldl (op -->) body_T batched end |
|
723 |
end |
|
41791
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
blanchet
parents:
41472
diff
changeset
|
724 |
fun format_term_type thy def_tables formats t = |
35718 | 725 |
format_type (the (AList.lookup (op =) formats NONE)) |
41791
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
blanchet
parents:
41472
diff
changeset
|
726 |
(lookup_format thy def_tables formats t) (fastype_of t) |
35718 | 727 |
|
728 |
fun repair_special_format js m format = |
|
729 |
m - 1 downto 0 |> chunk_list_unevenly (rev format) |
|
730 |
|> map (rev o filter_out (member (op =) js)) |
|
731 |
|> filter_out null |> map length |> rev |
|
732 |
||
41791
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
blanchet
parents:
41472
diff
changeset
|
733 |
fun user_friendly_const ({thy, evals, def_tables, skolems, special_funs, ...} |
35718 | 734 |
: hol_context) (base_name, step_name) formats = |
735 |
let |
|
736 |
val default_format = the (AList.lookup (op =) formats NONE) |
|
737 |
fun do_const (x as (s, T)) = |
|
738 |
(if String.isPrefix special_prefix s then |
|
739 |
let |
|
740 |
val do_term = map_aterms (fn Const x => fst (do_const x) | t' => t') |
|
741 |
val (x' as (_, T'), js, ts) = |
|
742 |
AList.find (op =) (!special_funs) (s, unarize_unbox_etc_type T) |
|
743 |
|> the_single |
|
744 |
val max_j = List.last js |
|
745 |
val Ts = List.take (binder_types T', max_j + 1) |
|
746 |
val missing_js = filter_out (member (op =) js) (0 upto max_j) |
|
747 |
val missing_Ts = filter_indices missing_js Ts |
|
748 |
fun nth_missing_var n = |
|
749 |
((arg_var_prefix ^ nat_subscript (n + 1), 0), nth missing_Ts n) |
|
750 |
val missing_vars = map nth_missing_var (0 upto length missing_js - 1) |
|
751 |
val vars = special_bounds ts @ missing_vars |
|
752 |
val ts' = |
|
753 |
map (fn j => |
|
754 |
case AList.lookup (op =) (js ~~ ts) j of |
|
755 |
SOME t => do_term t |
|
756 |
| NONE => |
|
757 |
Var (nth missing_vars |
|
758 |
(find_index (curry (op =) j) missing_js))) |
|
759 |
(0 upto max_j) |
|
760 |
val t = do_const x' |> fst |
|
761 |
val format = |
|
762 |
case AList.lookup (fn (SOME t1, SOME t2) => term_match thy (t1, t2) |
|
763 |
| _ => false) formats (SOME t) of |
|
764 |
SOME format => |
|
765 |
repair_special_format js (num_binder_types T') format |
|
766 |
| NONE => |
|
41791
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
blanchet
parents:
41472
diff
changeset
|
767 |
const_format thy def_tables x' |
35718 | 768 |
|> repair_special_format js (num_binder_types T') |
769 |
|> intersect_formats default_format |
|
770 |
in |
|
771 |
(list_comb (t, ts') |> fold_rev abs_var vars, |
|
772 |
format_type default_format format T) |
|
773 |
end |
|
774 |
else if String.isPrefix uncurry_prefix s then |
|
775 |
let |
|
776 |
val (ss, s') = unprefix uncurry_prefix s |
|
777 |
|> strip_first_name_sep |>> space_explode "@" |
|
778 |
in |
|
779 |
if String.isPrefix step_prefix s' then |
|
780 |
do_const (s', T) |
|
781 |
else |
|
782 |
let |
|
783 |
val k = the (Int.fromString (hd ss)) |
|
784 |
val j = the (Int.fromString (List.last ss)) |
|
785 |
val (before_Ts, (tuple_T, rest_T)) = |
|
786 |
strip_n_binders j T ||> (strip_n_binders 1 #>> hd) |
|
787 |
val T' = before_Ts ---> dest_n_tuple_type k tuple_T ---> rest_T |
|
788 |
in do_const (s', T') end |
|
789 |
end |
|
790 |
else if String.isPrefix unrolled_prefix s then |
|
791 |
let val t = Const (original_name s, range_type T) in |
|
792 |
(lambda (Free (iter_var_prefix, nat_T)) t, |
|
793 |
format_type default_format |
|
41791
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
blanchet
parents:
41472
diff
changeset
|
794 |
(lookup_format thy def_tables formats t) T) |
35718 | 795 |
end |
796 |
else if String.isPrefix base_prefix s then |
|
797 |
(Const (base_name, T --> T) $ Const (unprefix base_prefix s, T), |
|
798 |
format_type default_format default_format T) |
|
799 |
else if String.isPrefix step_prefix s then |
|
800 |
(Const (step_name, T --> T) $ Const (unprefix step_prefix s, T), |
|
801 |
format_type default_format default_format T) |
|
802 |
else if String.isPrefix quot_normal_prefix s then |
|
38207
792b78e355e7
added support for "Abs_" and "Rep_" functions on quotient types
blanchet
parents:
38190
diff
changeset
|
803 |
let val t = Const (nitpick_prefix ^ "quotient normal form", T) in |
41791
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
blanchet
parents:
41472
diff
changeset
|
804 |
(t, format_term_type thy def_tables formats t) |
35718 | 805 |
end |
806 |
else if String.isPrefix skolem_prefix s then |
|
807 |
let |
|
808 |
val ss = the (AList.lookup (op =) (!skolems) s) |
|
809 |
val (Ts, Ts') = chop (length ss) (binder_types T) |
|
810 |
val frees = map Free (ss ~~ Ts) |
|
811 |
val s' = original_name s |
|
812 |
in |
|
813 |
(fold lambda frees (Const (s', Ts' ---> T)), |
|
814 |
format_type default_format |
|
41791
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
blanchet
parents:
41472
diff
changeset
|
815 |
(lookup_format thy def_tables formats (Const x)) T) |
35718 | 816 |
end |
817 |
else if String.isPrefix eval_prefix s then |
|
818 |
let |
|
819 |
val t = nth evals (the (Int.fromString (unprefix eval_prefix s))) |
|
41791
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
blanchet
parents:
41472
diff
changeset
|
820 |
in (t, format_term_type thy def_tables formats t) end |
35718 | 821 |
else |
822 |
let val t = Const (original_name s, T) in |
|
41791
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
blanchet
parents:
41472
diff
changeset
|
823 |
(t, format_term_type thy def_tables formats t) |
35718 | 824 |
end) |
825 |
|>> map_types uniterize_unarize_unbox_etc_type |
|
826 |
|>> shorten_names_in_term |>> Term.map_abs_vars shortest_name |
|
827 |
in do_const end |
|
828 |
||
829 |
fun assign_operator_for_const (s, T) = |
|
830 |
if String.isPrefix ubfp_prefix s then |
|
37261 | 831 |
if is_fun_type T then xsym "\<subseteq>" "<=" () |
832 |
else xsym "\<le>" "<=" () |
|
35718 | 833 |
else if String.isPrefix lbfp_prefix s then |
37261 | 834 |
if is_fun_type T then xsym "\<supseteq>" ">=" () |
835 |
else xsym "\<ge>" ">=" () |
|
35718 | 836 |
else if original_name s <> s then |
837 |
assign_operator_for_const (strip_first_name_sep s |> snd, T) |
|
838 |
else |
|
839 |
"=" |
|
840 |
||
841 |
(** Model reconstruction **) |
|
842 |
||
33192 | 843 |
fun unfold_outer_the_binders (t as Const (@{const_name The}, _) |
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38284
diff
changeset
|
844 |
$ Abs (s, T, Const (@{const_name HOL.eq}, _) |
33192 | 845 |
$ Bound 0 $ t')) = |
846 |
betapply (Abs (s, T, t'), t) |> unfold_outer_the_binders |
|
847 |
| unfold_outer_the_binders t = t |
|
848 |
fun bisimilar_values _ 0 _ = true |
|
849 |
| bisimilar_values coTs max_depth (t1, t2) = |
|
850 |
let val T = fastype_of t1 in |
|
851 |
if exists_subtype (member (op =) coTs) T then |
|
852 |
let |
|
853 |
val ((head1, args1), (head2, args2)) = |
|
854 |
pairself (strip_comb o unfold_outer_the_binders) (t1, t2) |
|
34121
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents:
33982
diff
changeset
|
855 |
val max_depth = max_depth - (if member (op =) coTs T then 1 else 0) |
33192 | 856 |
in |
34936
c4f04bee79f3
some work on Nitpick's support for quotient types;
blanchet
parents:
34126
diff
changeset
|
857 |
head1 = head2 andalso |
c4f04bee79f3
some work on Nitpick's support for quotient types;
blanchet
parents:
34126
diff
changeset
|
858 |
forall (bisimilar_values coTs max_depth) (args1 ~~ args2) |
33192 | 859 |
end |
860 |
else |
|
861 |
t1 = t2 |
|
862 |
end |
|
863 |
||
41993
bd6296de1432
reintroduced "show_skolems" option -- useful when too many Skolems are displayed
blanchet
parents:
41875
diff
changeset
|
864 |
fun reconstruct_hol_model {show_datatypes, show_skolems, show_consts} |
35280
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
865 |
({hol_ctxt = {thy, ctxt, max_bisim_depth, boxes, stds, wfs, user_axioms, |
38209 | 866 |
debug, whacks, binary_ints, destroy_constrs, specialize, |
41875 | 867 |
star_linear_preds, total_consts, needs, tac_timeout, |
42415
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
blanchet
parents:
42375
diff
changeset
|
868 |
evals, case_names, def_tables, nondef_table, nondefs, |
41803
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
blanchet
parents:
41791
diff
changeset
|
869 |
simp_table, psimp_table, choice_spec_table, intro_table, |
36388 | 870 |
ground_thm_table, ersatz_table, skolems, special_funs, |
41803
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
blanchet
parents:
41791
diff
changeset
|
871 |
unrolled_preds, wf_cache, constr_cache}, binarize, |
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
blanchet
parents:
41791
diff
changeset
|
872 |
card_assigns, bits, bisim_depth, datatypes, ofs} : scope) |
38170 | 873 |
formats atomss real_frees pseudo_frees free_names sel_names nonsel_names |
874 |
rel_table bounds = |
|
33192 | 875 |
let |
35180
c57dba973391
more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents:
35179
diff
changeset
|
876 |
val pool = Unsynchronized.ref [] |
35280
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
877 |
val (wacky_names as (_, base_step_names), ctxt) = |
33192 | 878 |
add_wacky_syntax ctxt |
35070
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
blanchet
parents:
34998
diff
changeset
|
879 |
val hol_ctxt = |
33192 | 880 |
{thy = thy, ctxt = ctxt, max_bisim_depth = max_bisim_depth, boxes = boxes, |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
881 |
stds = stds, wfs = wfs, user_axioms = user_axioms, debug = debug, |
38209 | 882 |
whacks = whacks, binary_ints = binary_ints, |
883 |
destroy_constrs = destroy_constrs, specialize = specialize, |
|
41871
394eef237bd1
lower threshold for implicitly using "nitpick_simp" for predicate definitions when "total_consts" is on
blanchet
parents:
41803
diff
changeset
|
884 |
star_linear_preds = star_linear_preds, total_consts = total_consts, |
41875 | 885 |
needs = needs, tac_timeout = tac_timeout, evals = evals, |
41871
394eef237bd1
lower threshold for implicitly using "nitpick_simp" for predicate definitions when "total_consts" is on
blanchet
parents:
41803
diff
changeset
|
886 |
case_names = case_names, def_tables = def_tables, |
42415
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
blanchet
parents:
42375
diff
changeset
|
887 |
nondef_table = nondef_table, nondefs = nondefs, simp_table = simp_table, |
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
blanchet
parents:
42375
diff
changeset
|
888 |
psimp_table = psimp_table, choice_spec_table = choice_spec_table, |
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
blanchet
parents:
42375
diff
changeset
|
889 |
intro_table = intro_table, ground_thm_table = ground_thm_table, |
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
blanchet
parents:
42375
diff
changeset
|
890 |
ersatz_table = ersatz_table, skolems = skolems, |
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
blanchet
parents:
42375
diff
changeset
|
891 |
special_funs = special_funs, unrolled_preds = unrolled_preds, |
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
blanchet
parents:
42375
diff
changeset
|
892 |
wf_cache = wf_cache, constr_cache = constr_cache} |
36388 | 893 |
val scope = |
894 |
{hol_ctxt = hol_ctxt, binarize = binarize, card_assigns = card_assigns, |
|
895 |
bits = bits, bisim_depth = bisim_depth, datatypes = datatypes, ofs = ofs} |
|
37262
c0fe8fa35771
don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents:
37261
diff
changeset
|
896 |
fun term_for_rep maybe_opt unfold = |
c0fe8fa35771
don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents:
37261
diff
changeset
|
897 |
reconstruct_term maybe_opt unfold pool wacky_names scope atomss |
c0fe8fa35771
don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents:
37261
diff
changeset
|
898 |
sel_names rel_table bounds |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
899 |
val all_values = |
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
900 |
all_values_of_type pool wacky_names scope atomss sel_names rel_table |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
901 |
bounds |
38126 | 902 |
fun is_codatatype_wellformed (cos : datatype_spec list) |
903 |
({typ, card, ...} : datatype_spec) = |
|
33192 | 904 |
let |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
905 |
val ts = all_values card typ |
33192 | 906 |
val max_depth = Integer.sum (map #card cos) |
907 |
in |
|
908 |
forall (not o bisimilar_values (map #typ cos) max_depth) |
|
909 |
(all_distinct_unordered_pairs_of ts) |
|
910 |
end |
|
911 |
fun pretty_for_assign name = |
|
912 |
let |
|
913 |
val (oper, (t1, T'), T) = |
|
914 |
case name of |
|
915 |
FreeName (s, T, _) => |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
916 |
let val t = Free (s, uniterize_unarize_unbox_etc_type T) in |
41791
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
blanchet
parents:
41472
diff
changeset
|
917 |
("=", (t, format_term_type thy def_tables formats t), T) |
33192 | 918 |
end |
919 |
| ConstName (s, T, _) => |
|
920 |
(assign_operator_for_const (s, T), |
|
35280
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
921 |
user_friendly_const hol_ctxt base_step_names formats (s, T), T) |
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
922 |
| _ => raise NUT ("Nitpick_Model.reconstruct_hol_model.\ |
33192 | 923 |
\pretty_for_assign", [name]) |
924 |
val t2 = if rep_of name = Any then |
|
925 |
Const (@{const_name undefined}, T') |
|
926 |
else |
|
927 |
tuple_list_for_name rel_table bounds name |
|
37262
c0fe8fa35771
don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents:
37261
diff
changeset
|
928 |
|> term_for_rep (not (is_fully_representable_set name)) false |
c0fe8fa35771
don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents:
37261
diff
changeset
|
929 |
T T' (rep_of name) |
33192 | 930 |
in |
931 |
Pretty.block (Pretty.breaks |
|
39118
12f3788be67b
turned show_all_types into proper configuration option;
wenzelm
parents:
38864
diff
changeset
|
932 |
[Syntax.pretty_term (set_show_all_types ctxt) t1, |
33192 | 933 |
Pretty.str oper, Syntax.pretty_term ctxt t2]) |
934 |
end |
|
38126 | 935 |
fun pretty_for_datatype ({typ, card, complete, ...} : datatype_spec) = |
33192 | 936 |
Pretty.block (Pretty.breaks |
38189 | 937 |
(pretty_for_type ctxt typ :: |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
938 |
(case typ of |
41052
3db267a01c1d
remove the "fin_fun" optimization in Nitpick -- it was always a hack and didn't help much
blanchet
parents:
41039
diff
changeset
|
939 |
Type (@{type_name fun_box}, _) => [Pretty.str "[boxed]"] |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
940 |
| Type (@{type_name pair_box}, _) => [Pretty.str "[boxed]"] |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
941 |
| _ => []) @ |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
942 |
[Pretty.str "=", |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
943 |
Pretty.enum "," "{" "}" |
35712
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents:
35711
diff
changeset
|
944 |
(map (Syntax.pretty_term ctxt) (all_values card typ) @ |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
945 |
(if fun_from_pair complete false then [] |
37261 | 946 |
else [Pretty.str (unrep ())]))])) |
33192 | 947 |
fun integer_datatype T = |
948 |
[{typ = T, card = card_of_type card_assigns T, co = false, |
|
38126 | 949 |
standard = true, self_rec = true, complete = (false, false), |
950 |
concrete = (true, true), deep = true, constrs = []}] |
|
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33202
diff
changeset
|
951 |
handle TYPE ("Nitpick_HOL.card_of_type", _, _) => [] |
33192 | 952 |
val (codatatypes, datatypes) = |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34974
diff
changeset
|
953 |
datatypes |> filter #deep |> List.partition #co |
35220
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35190
diff
changeset
|
954 |
||> append (integer_datatype int_T |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35190
diff
changeset
|
955 |
|> is_standard_datatype thy stds nat_T |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35190
diff
changeset
|
956 |
? append (integer_datatype nat_T)) |
33192 | 957 |
val block_of_datatypes = |
958 |
if show_datatypes andalso not (null datatypes) then |
|
959 |
[Pretty.big_list ("Datatype" ^ plural_s_for_list datatypes ^ ":") |
|
960 |
(map pretty_for_datatype datatypes)] |
|
961 |
else |
|
962 |
[] |
|
963 |
val block_of_codatatypes = |
|
964 |
if show_datatypes andalso not (null codatatypes) then |
|
965 |
[Pretty.big_list ("Codatatype" ^ plural_s_for_list codatatypes ^ ":") |
|
966 |
(map pretty_for_datatype codatatypes)] |
|
967 |
else |
|
968 |
[] |
|
969 |
fun block_of_names show title names = |
|
970 |
if show andalso not (null names) then |
|
971 |
Pretty.str (title ^ plural_s_for_list names ^ ":") |
|
972 |
:: map (Pretty.indent indent_size o pretty_for_assign) |
|
973 |
(sort_wrt (original_name o nickname_of) names) |
|
974 |
else |
|
975 |
[] |
|
38170 | 976 |
fun free_name_for_term keep_all (x as (s, T)) = |
977 |
case filter (curry (op =) x |
|
978 |
o pairf nickname_of (uniterize_unarize_unbox_etc_type |
|
979 |
o type_of)) free_names of |
|
980 |
[name] => SOME name |
|
981 |
| [] => if keep_all then SOME (FreeName (s, T, Any)) else NONE |
|
982 |
| _ => raise TERM ("Nitpick_Model.reconstruct_hol_model.\ |
|
983 |
\free_name_for_term", [Const x]) |
|
33192 | 984 |
val (skolem_names, nonskolem_nonsel_names) = |
985 |
List.partition is_skolem_name nonsel_names |
|
986 |
val (eval_names, noneval_nonskolem_nonsel_names) = |
|
987 |
List.partition (String.isPrefix eval_prefix o nickname_of) |
|
988 |
nonskolem_nonsel_names |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
989 |
||> filter_out (member (op =) [@{const_name bisim}, |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35625
diff
changeset
|
990 |
@{const_name bisim_iterator_max}] |
34121
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents:
33982
diff
changeset
|
991 |
o nickname_of) |
38170 | 992 |
||> append (map_filter (free_name_for_term false) pseudo_frees) |
993 |
val real_free_names = map_filter (free_name_for_term true) real_frees |
|
994 |
val chunks = block_of_names true "Free variable" real_free_names @ |
|
41993
bd6296de1432
reintroduced "show_skolems" option -- useful when too many Skolems are displayed
blanchet
parents:
41875
diff
changeset
|
995 |
block_of_names show_skolems "Skolem constant" skolem_names @ |
33192 | 996 |
block_of_names true "Evaluated term" eval_names @ |
997 |
block_of_datatypes @ block_of_codatatypes @ |
|
998 |
block_of_names show_consts "Constant" |
|
999 |
noneval_nonskolem_nonsel_names |
|
1000 |
in |
|
1001 |
(Pretty.chunks (if null chunks then [Pretty.str "Empty assignment"] |
|
1002 |
else chunks), |
|
34936
c4f04bee79f3
some work on Nitpick's support for quotient types;
blanchet
parents:
34126
diff
changeset
|
1003 |
bisim_depth >= 0 orelse |
c4f04bee79f3
some work on Nitpick's support for quotient types;
blanchet
parents:
34126
diff
changeset
|
1004 |
forall (is_codatatype_wellformed codatatypes) codatatypes) |
33192 | 1005 |
end |
1006 |
||
37262
c0fe8fa35771
don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents:
37261
diff
changeset
|
1007 |
fun term_for_name pool scope atomss sel_names rel_table bounds name = |
c0fe8fa35771
don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents:
37261
diff
changeset
|
1008 |
let val T = type_of name in |
c0fe8fa35771
don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents:
37261
diff
changeset
|
1009 |
tuple_list_for_name rel_table bounds name |
c0fe8fa35771
don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents:
37261
diff
changeset
|
1010 |
|> reconstruct_term (not (is_fully_representable_set name)) false pool |
c0fe8fa35771
don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents:
37261
diff
changeset
|
1011 |
(("", ""), ("", "")) scope atomss sel_names rel_table |
c0fe8fa35771
don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents:
37261
diff
changeset
|
1012 |
bounds T T (rep_of name) |
c0fe8fa35771
don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents:
37261
diff
changeset
|
1013 |
end |
c0fe8fa35771
don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents:
37261
diff
changeset
|
1014 |
|
35280
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
1015 |
fun prove_hol_model (scope as {hol_ctxt = {thy, ctxt, debug, ...}, |
34998 | 1016 |
card_assigns, ...}) |
33192 | 1017 |
auto_timeout free_names sel_names rel_table bounds prop = |
1018 |
let |
|
35076
cc19e2aef17e
added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents:
35075
diff
changeset
|
1019 |
val pool = Unsynchronized.ref [] |
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
1020 |
val atomss = [(NONE, [])] |
33192 | 1021 |
fun free_type_assm (T, k) = |
1022 |
let |
|
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
1023 |
fun atom j = nth_atom thy atomss pool true T j |
33192 | 1024 |
fun equation_for_atom j = HOLogic.eq_const T $ Bound 0 $ atom j |
1025 |
val eqs = map equation_for_atom (index_seq 0 k) |
|
1026 |
val compreh_assm = |
|
1027 |
Const (@{const_name All}, (T --> bool_T) --> bool_T) |
|
1028 |
$ Abs ("x", T, foldl1 HOLogic.mk_disj eqs) |
|
1029 |
val distinct_assm = distinctness_formula T (map atom (index_seq 0 k)) |
|
34998 | 1030 |
in s_conj (compreh_assm, distinct_assm) end |
33192 | 1031 |
fun free_name_assm name = |
1032 |
HOLogic.mk_eq (Free (nickname_of name, type_of name), |
|
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
1033 |
term_for_name pool scope atomss sel_names rel_table bounds |
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
1034 |
name) |
33192 | 1035 |
val freeT_assms = map free_type_assm (filter (is_TFree o fst) card_assigns) |
1036 |
val model_assms = map free_name_assm free_names |
|
34998 | 1037 |
val assm = foldr1 s_conj (freeT_assms @ model_assms) |
33192 | 1038 |
fun try_out negate = |
1039 |
let |
|
1040 |
val concl = (negate ? curry (op $) @{const Not}) |
|
35625 | 1041 |
(Object_Logic.atomize_term thy prop) |
34998 | 1042 |
val prop = HOLogic.mk_Trueprop (HOLogic.mk_imp (assm, concl)) |
33192 | 1043 |
|> map_types (map_type_tfree |
34998 | 1044 |
(fn (s, []) => TFree (s, HOLogic.typeS) |
1045 |
| x => TFree x)) |
|
1046 |
val _ = if debug then |
|
40132
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents:
39360
diff
changeset
|
1047 |
Output.urgent_message ((if negate then "Genuineness" else "Spuriousness") ^ |
34998 | 1048 |
" goal: " ^ Syntax.string_of_term ctxt prop ^ ".") |
1049 |
else |
|
1050 |
() |
|
1051 |
val goal = prop |> cterm_of thy |> Goal.init |
|
33192 | 1052 |
in |
42793 | 1053 |
(goal |> SINGLE (DETERM_TIMEOUT auto_timeout (auto_tac ctxt)) |
33192 | 1054 |
|> the |> Goal.finish ctxt; true) |
1055 |
handle THM _ => false |
|
1056 |
| TimeLimit.TimeOut => false |
|
1057 |
end |
|
1058 |
in |
|
33705
947184dc75c9
removed a few global names in Nitpick (styp, nat_less, pairf)
blanchet
parents:
33580
diff
changeset
|
1059 |
if try_out false then SOME true |
947184dc75c9
removed a few global names in Nitpick (styp, nat_less, pairf)
blanchet
parents:
33580
diff
changeset
|
1060 |
else if try_out true then SOME false |
33192 | 1061 |
else NONE |
1062 |
end |
|
1063 |
||
1064 |
end; |