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