src/HOL/Tools/Nitpick/nitpick_scope.ML
author blanchet
Mon, 03 Mar 2014 22:33:22 +0100
changeset 55888 cac1add157e8
parent 55628 8103021024c1
child 55889 6bfbec3dff62
permissions -rw-r--r--
removed nonstandard models from Nitpick
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33982
1ae222745c4a fixed paths in Nitpick's ML file headers
blanchet
parents: 33957
diff changeset
     1
(*  Title:      HOL/Tools/Nitpick/nitpick_scope.ML
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
     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: 34936
diff changeset
     3
    Copyright   2008, 2009, 2010
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
     4
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
     5
Scope enumerator for Nitpick.
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
     6
*)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
     7
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
     8
signature NITPICK_SCOPE =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
     9
sig
33705
947184dc75c9 removed a few global names in Nitpick (styp, nat_less, pairf)
blanchet
parents: 33580
diff changeset
    10
  type styp = Nitpick_Util.styp
35070
96136eb6218f split "nitpick_hol.ML" into two files to make it more manageable;
blanchet
parents: 34982
diff changeset
    11
  type hol_context = Nitpick_HOL.hol_context
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    12
36390
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    13
  type constr_spec =
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    14
    {const: styp,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    15
     delta: int,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    16
     epsilon: int,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    17
     exclusive: bool,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    18
     explicit_max: int,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    19
     total: bool}
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    20
38126
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 38124
diff changeset
    21
  type datatype_spec =
36390
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    22
    {typ: typ,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    23
     card: int,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    24
     co: bool,
38126
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 38124
diff changeset
    25
     self_rec: bool,
36390
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    26
     complete: bool * bool,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    27
     concrete: bool * bool,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    28
     deep: bool,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    29
     constrs: constr_spec list}
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    30
36390
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    31
  type scope =
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    32
    {hol_ctxt: hol_context,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    33
     binarize: bool,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    34
     card_assigns: (typ * int) list,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    35
     bits: int,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    36
     bisim_depth: int,
38126
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 38124
diff changeset
    37
     datatypes: datatype_spec list,
36390
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    38
     ofs: int Typtab.table}
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    39
38127
9f9f696fc4e8 tweak datatype sym break code
blanchet
parents: 38126
diff changeset
    40
  val is_asymmetric_nondatatype : typ -> bool
38126
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 38124
diff changeset
    41
  val datatype_spec : datatype_spec list -> typ -> datatype_spec option
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 38124
diff changeset
    42
  val constr_spec : datatype_spec list -> styp -> constr_spec
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 38124
diff changeset
    43
  val is_complete_type : datatype_spec list -> bool -> typ -> bool
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 38124
diff changeset
    44
  val is_concrete_type : datatype_spec list -> bool -> typ -> bool
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 38124
diff changeset
    45
  val is_exact_type : datatype_spec list -> bool -> typ -> bool
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    46
  val offset_of_type : int Typtab.table -> typ -> int
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    47
  val spec_of_type : scope -> typ -> int * int
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    48
  val pretties_for_scope : scope -> bool -> Pretty.T list
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    49
  val multiline_string_for_scope : scope -> string
35814
234eaa508359 added one-entry cache around Kodkod invocation
blanchet
parents: 35665
diff changeset
    50
  val scopes_equivalent : scope * scope -> bool
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    51
  val scope_less_eq : scope -> scope -> bool
38126
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 38124
diff changeset
    52
  val is_self_recursive_constr_type : typ -> bool
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    53
  val all_scopes :
36386
2132f15b366f Fruhjahrsputz: remove three mostly useless Nitpick options
blanchet
parents: 36385
diff changeset
    54
    hol_context -> bool -> (typ option * int list) list
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    55
    -> (styp option * int list) list -> (styp option * int list) list
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
    56
    -> int list -> int list -> typ list -> typ list -> typ list -> typ list
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
    57
    -> int * scope list
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    58
end;
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    59
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
    60
structure Nitpick_Scope : NITPICK_SCOPE =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    61
struct
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    62
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
    63
open Nitpick_Util
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
    64
open Nitpick_HOL
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    65
36390
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    66
type constr_spec =
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    67
  {const: styp,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    68
   delta: int,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    69
   epsilon: int,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    70
   exclusive: bool,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    71
   explicit_max: int,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    72
   total: bool}
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    73
38126
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 38124
diff changeset
    74
type datatype_spec =
36390
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    75
  {typ: typ,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    76
   card: int,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    77
   co: bool,
38126
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 38124
diff changeset
    78
   self_rec: bool,
36390
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    79
   complete: bool * bool,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    80
   concrete: bool * bool,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    81
   deep: bool,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    82
   constrs: constr_spec list}
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    83
36390
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    84
type scope =
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    85
  {hol_ctxt: hol_context,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    86
   binarize: bool,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    87
   card_assigns: (typ * int) list,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    88
   bits: int,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    89
   bisim_depth: int,
38126
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 38124
diff changeset
    90
   datatypes: datatype_spec list,
36390
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36386
diff changeset
    91
   ofs: int Typtab.table}
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    92
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    93
datatype row_kind = Card of typ | Max of styp
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    94
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    95
type row = row_kind * int list
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    96
type block = row list
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    97
38127
9f9f696fc4e8 tweak datatype sym break code
blanchet
parents: 38126
diff changeset
    98
val is_asymmetric_nondatatype =
9f9f696fc4e8 tweak datatype sym break code
blanchet
parents: 38126
diff changeset
    99
  is_iterator_type orf is_integer_type orf is_bit_type
9f9f696fc4e8 tweak datatype sym break code
blanchet
parents: 38126
diff changeset
   100
38126
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 38124
diff changeset
   101
fun datatype_spec (dtypes : datatype_spec list) T =
34121
5e831d805118 get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents: 33982
diff changeset
   102
  List.find (curry (op =) T o #typ) dtypes
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   103
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
   104
fun constr_spec [] x = raise TERM ("Nitpick_Scope.constr_spec", [Const x])
38126
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 38124
diff changeset
   105
  | constr_spec ({constrs, ...} :: dtypes : datatype_spec list) (x as (s, T)) =
34121
5e831d805118 get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents: 33982
diff changeset
   106
    case List.find (curry (op =) (s, body_type T) o (apsnd body_type o #const))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   107
                   constrs of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   108
      SOME c => c
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   109
    | NONE => constr_spec dtypes x
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   110
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35385
diff changeset
   111
fun is_complete_type dtypes facto (Type (@{type_name fun}, [T1, T2])) =
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   112
    is_concrete_type dtypes facto T1 andalso is_complete_type dtypes facto T2
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38188
diff changeset
   113
  | is_complete_type dtypes facto (Type (@{type_name prod}, Ts)) =
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   114
    forall (is_complete_type dtypes facto) Ts
46083
efeaa79f021b port part of Nitpick to "set" type constructor
blanchet
parents: 45402
diff changeset
   115
  | is_complete_type dtypes facto (Type (@{type_name set}, [T'])) =
efeaa79f021b port part of Nitpick to "set" type constructor
blanchet
parents: 45402
diff changeset
   116
    is_concrete_type dtypes facto T'
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   117
  | is_complete_type dtypes facto T =
35220
2bcdae5f4fdb added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents: 35190
diff changeset
   118
    not (is_integer_like_type T) andalso not (is_bit_type T) andalso
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   119
    fun_from_pair (#complete (the (datatype_spec dtypes T))) facto
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   120
    handle Option.Option => true
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35385
diff changeset
   121
and is_concrete_type dtypes facto (Type (@{type_name fun}, [T1, T2])) =
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   122
    is_complete_type dtypes facto T1 andalso is_concrete_type dtypes facto T2
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38188
diff changeset
   123
  | is_concrete_type dtypes facto (Type (@{type_name prod}, Ts)) =
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   124
    forall (is_concrete_type dtypes facto) Ts
46083
efeaa79f021b port part of Nitpick to "set" type constructor
blanchet
parents: 45402
diff changeset
   125
  | is_concrete_type dtypes facto (Type (@{type_name set}, [T'])) =
efeaa79f021b port part of Nitpick to "set" type constructor
blanchet
parents: 45402
diff changeset
   126
    is_complete_type dtypes facto T'
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   127
  | is_concrete_type dtypes facto T =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   128
    fun_from_pair (#concrete (the (datatype_spec dtypes T))) facto
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   129
    handle Option.Option => true
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35385
diff changeset
   130
and is_exact_type dtypes facto =
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   131
  is_complete_type dtypes facto andf is_concrete_type dtypes facto
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   132
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   133
fun offset_of_type ofs T =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   134
  case Typtab.lookup ofs T of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   135
    SOME j0 => j0
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   136
  | NONE => Typtab.lookup ofs dummyT |> the_default 0
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   137
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   138
fun spec_of_type ({card_assigns, ofs, ...} : scope) T =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   139
  (card_of_type card_assigns T
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
   140
   handle TYPE ("Nitpick_HOL.card_of_type", _, _) => ~1, offset_of_type ofs T)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   141
38188
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   142
fun quintuple_for_scope code_type code_term code_string
55888
cac1add157e8 removed nonstandard models from Nitpick
blanchet
parents: 55628
diff changeset
   143
        ({hol_ctxt = {ctxt, ...}, card_assigns, bits, bisim_depth,
35220
2bcdae5f4fdb added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents: 35190
diff changeset
   144
         datatypes, ...} : scope) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   145
  let
35075
888802be2019 handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents: 35072
diff changeset
   146
    val boring_Ts = [@{typ unsigned_bit}, @{typ signed_bit},
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   147
                     @{typ bisim_iterator}]
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   148
    val (iter_assigns, card_assigns) =
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   149
      card_assigns |> filter_out (member (op =) boring_Ts o fst)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   150
                   |> List.partition (is_fp_iterator_type o fst)
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   151
    val (secondary_card_assigns, primary_card_assigns) =
37256
0dca1ec52999 thread along context instead of theory for typedef lookup
blanchet
parents: 36390
diff changeset
   152
      card_assigns
55888
cac1add157e8 removed nonstandard models from Nitpick
blanchet
parents: 55628
diff changeset
   153
      |> List.partition ((is_integer_type orf is_datatype ctxt) o fst)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   154
    val cards =
38188
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   155
      map (fn (T, k) =>
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   156
              [code_type ctxt T, code_string (" = " ^ string_of_int k)])
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   157
    fun maxes () =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   158
      maps (map_filter
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   159
                (fn {const, explicit_max, ...} =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   160
                    if explicit_max < 0 then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   161
                      NONE
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   162
                    else
38188
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   163
                      SOME [code_term ctxt (Const const),
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   164
                            code_string (" = " ^ string_of_int explicit_max)])
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   165
                 o #constrs) datatypes
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   166
    fun iters () =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   167
      map (fn (T, k) =>
38188
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   168
              [code_term ctxt (Const (const_for_iterator_type T)),
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   169
               code_string (" = " ^ string_of_int (k - 1))]) iter_assigns
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   170
    fun miscs () =
38188
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   171
      (if bits = 0 then []
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   172
       else [code_string ("bits = " ^ string_of_int bits)]) @
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   173
      (if bisim_depth < 0 andalso forall (not o #co) datatypes then []
38188
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   174
       else [code_string ("bisim_depth = " ^ signed_string_of_int bisim_depth)])
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   175
  in
39118
12f3788be67b turned show_all_types into proper configuration option;
wenzelm
parents: 38240
diff changeset
   176
    (cards primary_card_assigns, cards secondary_card_assigns,
12f3788be67b turned show_all_types into proper configuration option;
wenzelm
parents: 38240
diff changeset
   177
     maxes (), iters (), miscs ())
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   178
  end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   179
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   180
fun pretties_for_scope scope verbose =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   181
  let
38188
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   182
    fun standard_blocks s = map (Pretty.block o cons (Pretty.str (s ^ " ")))
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   183
    val (primary_cards, secondary_cards, maxes, iters, miscs) =
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   184
      quintuple_for_scope (pretty_maybe_quote oo pretty_for_type)
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   185
                          (pretty_maybe_quote oo Syntax.pretty_term)
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   186
                          Pretty.str scope
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   187
  in
38188
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   188
    standard_blocks "card" primary_cards @
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   189
    (if verbose then
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   190
       standard_blocks "card" secondary_cards @
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   191
       standard_blocks "max" maxes @
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   192
       standard_blocks "iter" iters @
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   193
       miscs
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   194
     else
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   195
       [])
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   196
    |> pretty_serial_commas "and"
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   197
  end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   198
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   199
fun multiline_string_for_scope scope =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   200
  let
38188
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   201
    val (primary_cards, secondary_cards, maxes, iters, miscs) =
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   202
      quintuple_for_scope Syntax.string_of_typ Syntax.string_of_term I scope
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   203
    val cards = primary_cards @ secondary_cards
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   204
  in
38188
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   205
    case (if null cards then [] else ["card: " ^ commas (map implode cards)]) @
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   206
         (if null maxes then [] else ["max: " ^ commas (map implode maxes)]) @
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   207
         (if null iters then [] else ["iter: " ^ commas (map implode iters)]) @
7f12a03c513c better "Pretty" handling
blanchet
parents: 38187
diff changeset
   208
         miscs of
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   209
      [] => "empty"
55628
8103021024c1 prefer cat_lines;
wenzelm
parents: 52174
diff changeset
   210
    | lines => cat_lines lines
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   211
  end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   212
35814
234eaa508359 added one-entry cache around Kodkod invocation
blanchet
parents: 35665
diff changeset
   213
fun scopes_equivalent (s1 : scope, s2 : scope) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   214
  #datatypes s1 = #datatypes s2 andalso #card_assigns s1 = #card_assigns s2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   215
fun scope_less_eq (s1 : scope) (s2 : scope) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   216
  (s1, s2) |> pairself (map snd o #card_assigns) |> op ~~ |> forall (op <=)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   217
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   218
fun rank_of_row (_, ks) = length ks
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   219
fun rank_of_block block = fold Integer.max (map rank_of_row block) 1
41991
ea02b9ee3085 prevent an exception if "card" is empty (e.g., "nitpick [card]")
blanchet
parents: 41052
diff changeset
   220
fun project_row _ (y, []) = (y, [1]) (* desperate measure *)
ea02b9ee3085 prevent an exception if "card" is empty (e.g., "nitpick [card]")
blanchet
parents: 41052
diff changeset
   221
  | project_row column (y, ks) = (y, [nth ks (Int.min (column, length ks - 1))])
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   222
fun project_block (column, block) = map (project_row column) block
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   223
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   224
fun lookup_ints_assign eq assigns key =
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   225
  case triple_lookup eq assigns key of
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   226
    SOME ks => ks
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
   227
  | NONE => raise ARG ("Nitpick_Scope.lookup_ints_assign", "")
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   228
fun lookup_type_ints_assign thy assigns T =
36384
76d5fd5a45fb cosmetics
blanchet
parents: 36380
diff changeset
   229
  map (Integer.max 1) (lookup_ints_assign (type_match thy) assigns T)
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
   230
  handle ARG ("Nitpick_Scope.lookup_ints_assign", _) =>
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
   231
         raise TYPE ("Nitpick_Scope.lookup_type_ints_assign", [T], [])
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   232
fun lookup_const_ints_assign thy assigns x =
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   233
  lookup_ints_assign (const_match thy) assigns x
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
   234
  handle ARG ("Nitpick_Scope.lookup_ints_assign", _) =>
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
   235
         raise TERM ("Nitpick_Scope.lookup_const_ints_assign", [Const x])
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   236
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   237
fun row_for_constr thy maxes_assigns constr =
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   238
  SOME (Max constr, lookup_const_ints_assign thy maxes_assigns constr)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   239
  handle TERM ("lookup_const_ints_assign", _) => NONE
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   240
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   241
val max_bits = 31 (* Kodkod limit *)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   242
35190
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   243
fun block_for_type (hol_ctxt as {thy, ...}) binarize cards_assigns maxes_assigns
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   244
                   iters_assigns bitss bisim_depths T =
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   245
  case T of
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   246
    @{typ unsigned_bit} =>
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   247
    [(Card T, map (Integer.min max_bits o Integer.max 1) bitss)]
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   248
  | @{typ signed_bit} =>
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   249
    [(Card T, map (Integer.add 1 o Integer.min max_bits o Integer.max 1) bitss)]
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   250
  | @{typ "unsigned_bit word"} =>
34126
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   251
    [(Card T, lookup_type_ints_assign thy cards_assigns nat_T)]
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   252
  | @{typ "signed_bit word"} =>
34126
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   253
    [(Card T, lookup_type_ints_assign thy cards_assigns int_T)]
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   254
  | @{typ bisim_iterator} =>
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   255
    [(Card T, map (Integer.add 1 o Integer.max 0) bisim_depths)]
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   256
  | _ =>
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   257
    if is_fp_iterator_type T then
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   258
      [(Card T, map (Integer.add 1 o Integer.max 0)
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   259
                    (lookup_const_ints_assign thy iters_assigns
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   260
                                              (const_for_iterator_type T)))]
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   261
    else
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   262
      (Card T, lookup_type_ints_assign thy cards_assigns T) ::
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   263
      (case binarized_and_boxed_datatype_constrs hol_ctxt binarize T of
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   264
         [_] => []
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   265
       | constrs => map_filter (row_for_constr thy maxes_assigns) constrs)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   266
35190
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   267
fun blocks_for_types hol_ctxt binarize cards_assigns maxes_assigns iters_assigns
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   268
                     bitss bisim_depths mono_Ts nonmono_Ts =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   269
  let
35190
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   270
    val block_for = block_for_type hol_ctxt binarize cards_assigns maxes_assigns
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   271
                                   iters_assigns bitss bisim_depths
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   272
    val mono_block = maps block_for mono_Ts
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   273
    val nonmono_blocks = map block_for nonmono_Ts
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   274
  in mono_block :: nonmono_blocks end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   275
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   276
val sync_threshold = 5
38186
c28018f5a1d6 example tweaking -- also prevents Nitpick_Tests from using more than 1 thread
blanchet
parents: 38180
diff changeset
   277
val linearity = 5
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   278
38186
c28018f5a1d6 example tweaking -- also prevents Nitpick_Tests from using more than 1 thread
blanchet
parents: 38180
diff changeset
   279
val all_combinations_ordered_smartly =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   280
  let
38186
c28018f5a1d6 example tweaking -- also prevents Nitpick_Tests from using more than 1 thread
blanchet
parents: 38180
diff changeset
   281
    fun cost [] = 0
c28018f5a1d6 example tweaking -- also prevents Nitpick_Tests from using more than 1 thread
blanchet
parents: 38180
diff changeset
   282
      | cost [k] = k
c28018f5a1d6 example tweaking -- also prevents Nitpick_Tests from using more than 1 thread
blanchet
parents: 38180
diff changeset
   283
      | cost (k :: ks) =
34121
5e831d805118 get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents: 33982
diff changeset
   284
        if k < sync_threshold andalso forall (curry (op =) k) ks then
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   285
          k - sync_threshold
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   286
        else
38187
fd28328daf73 change formula for enumerating scopes
blanchet
parents: 38186
diff changeset
   287
          k :: ks |> map (fn k => (k + linearity) * (k + linearity))
fd28328daf73 change formula for enumerating scopes
blanchet
parents: 38186
diff changeset
   288
                  |> Integer.sum
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   289
  in
38187
fd28328daf73 change formula for enumerating scopes
blanchet
parents: 38186
diff changeset
   290
    all_combinations #> map (`cost) #> sort (int_ord o pairself fst) #> map snd
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   291
  end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   292
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   293
fun is_self_recursive_constr_type T =
34121
5e831d805118 get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents: 33982
diff changeset
   294
  exists (exists_subtype (curry (op =) (body_type T))) (binder_types T)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   295
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   296
fun constr_max maxes x = the_default ~1 (AList.lookup (op =) maxes x)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   297
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   298
type scope_desc = (typ * int) list * (styp * int) list
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   299
35190
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   300
fun is_surely_inconsistent_card_assign hol_ctxt binarize
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   301
                                       (card_assigns, max_assigns) (T, k) =
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   302
  case binarized_and_boxed_datatype_constrs hol_ctxt binarize T of
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   303
    [] => false
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   304
  | xs =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   305
    let
34126
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   306
      val dom_cards =
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   307
        map (Integer.prod o map (bounded_card_of_type k ~1 card_assigns)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   308
             o binder_types o snd) xs
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   309
      val maxes = map (constr_max max_assigns) xs
34126
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   310
      fun effective_max card ~1 = card
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   311
        | effective_max card max = Int.min (card, max)
34126
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   312
      val max = map2 effective_max dom_cards maxes |> Integer.sum
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   313
    in max < k end
35190
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   314
fun is_surely_inconsistent_scope_description hol_ctxt binarize seen rest
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   315
                                             max_assigns =
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   316
  exists (is_surely_inconsistent_card_assign hol_ctxt binarize
34126
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   317
                                             (seen @ rest, max_assigns)) seen
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   318
35190
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   319
fun repair_card_assigns hol_ctxt binarize (card_assigns, max_assigns) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   320
  let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   321
    fun aux seen [] = SOME seen
35280
54ab4921f826 fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents: 35220
diff changeset
   322
      | aux _ ((_, 0) :: _) = NONE
34126
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   323
      | aux seen ((T, k) :: rest) =
35190
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   324
        (if is_surely_inconsistent_scope_description hol_ctxt binarize
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   325
                ((T, k) :: seen) rest max_assigns then
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   326
           raise SAME ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   327
         else
34126
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   328
           case aux ((T, k) :: seen) rest of
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   329
             SOME assigns => SOME assigns
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   330
           | NONE => raise SAME ())
34126
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   331
        handle SAME () => aux seen ((T, k - 1) :: rest)
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   332
  in aux [] (rev card_assigns) end
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   333
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   334
fun repair_iterator_assign ctxt assigns (T as Type (_, Ts), k) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   335
    (T, if T = @{typ bisim_iterator} then
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   336
          let
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   337
            val co_cards = map snd (filter (is_codatatype ctxt o fst) assigns)
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   338
          in Int.min (k, Integer.sum co_cards) end
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   339
        else if is_fp_iterator_type T then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   340
          case Ts of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   341
            [] => 1
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   342
          | _ => bounded_card_of_type k ~1 assigns (foldr1 HOLogic.mk_prodT Ts)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   343
        else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   344
          k)
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   345
  | repair_iterator_assign _ _ assign = assign
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   346
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   347
fun add_row_to_scope_descriptor (kind, ks) (card_assigns, max_assigns) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   348
  case kind of
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   349
    Card T => ((T, the_single ks) :: card_assigns, max_assigns)
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   350
  | Max x => (card_assigns, (x, the_single ks) :: max_assigns)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   351
fun scope_descriptor_from_block block =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   352
  fold_rev add_row_to_scope_descriptor block ([], [])
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   353
fun scope_descriptor_from_combination (hol_ctxt as {ctxt, ...}) binarize blocks
35190
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   354
                                      columns =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   355
  let
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   356
    val (card_assigns, max_assigns) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   357
      maps project_block (columns ~~ blocks) |> scope_descriptor_from_block
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   358
  in
41991
ea02b9ee3085 prevent an exception if "card" is empty (e.g., "nitpick [card]")
blanchet
parents: 41052
diff changeset
   359
    (card_assigns, max_assigns)
ea02b9ee3085 prevent an exception if "card" is empty (e.g., "nitpick [card]")
blanchet
parents: 41052
diff changeset
   360
    |> repair_card_assigns hol_ctxt binarize
ea02b9ee3085 prevent an exception if "card" is empty (e.g., "nitpick [card]")
blanchet
parents: 41052
diff changeset
   361
    |> Option.map
ea02b9ee3085 prevent an exception if "card" is empty (e.g., "nitpick [card]")
blanchet
parents: 41052
diff changeset
   362
           (fn card_assigns =>
ea02b9ee3085 prevent an exception if "card" is empty (e.g., "nitpick [card]")
blanchet
parents: 41052
diff changeset
   363
               (map (repair_iterator_assign ctxt card_assigns) card_assigns,
ea02b9ee3085 prevent an exception if "card" is empty (e.g., "nitpick [card]")
blanchet
parents: 41052
diff changeset
   364
                max_assigns))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   365
  end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   366
38124
6538e25cf5dd started implementation of custom sym break
blanchet
parents: 37678
diff changeset
   367
fun offset_table_for_card_assigns dtypes assigns =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   368
  let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   369
    fun aux next _ [] = Typtab.update_new (dummyT, next)
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   370
      | aux next reusable ((T, k) :: assigns) =
38127
9f9f696fc4e8 tweak datatype sym break code
blanchet
parents: 38126
diff changeset
   371
        if k = 1 orelse is_asymmetric_nondatatype T then
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   372
          aux next reusable assigns
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   373
        else if length (these (Option.map #constrs (datatype_spec dtypes T)))
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   374
                > 1 then
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   375
          Typtab.update_new (T, next) #> aux (next + k) reusable assigns
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   376
        else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   377
          case AList.lookup (op =) reusable k of
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   378
            SOME j0 => Typtab.update_new (T, j0) #> aux next reusable assigns
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   379
          | NONE => Typtab.update_new (T, next)
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   380
                    #> aux (next + k) ((k, next) :: reusable) assigns
38124
6538e25cf5dd started implementation of custom sym break
blanchet
parents: 37678
diff changeset
   381
  in Typtab.empty |> aux 0 [] assigns end
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   382
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   383
fun domain_card max card_assigns =
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   384
  Integer.prod o map (bounded_card_of_type max max card_assigns) o binder_types
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   385
38162
824e940a3dd0 minor symmetry breaking for codatatypes like llist
blanchet
parents: 38127
diff changeset
   386
fun add_constr_spec (card_assigns, max_assigns) acyclic card sum_dom_cards
35280
54ab4921f826 fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents: 35220
diff changeset
   387
                    num_self_recs num_non_self_recs (self_rec, x as (_, T))
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   388
                    constrs =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   389
  let
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   390
    val max = constr_max max_assigns x
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   391
    fun next_delta () = if null constrs then 0 else #epsilon (hd constrs)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   392
    val {delta, epsilon, exclusive, total} =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   393
      if max = 0 then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   394
        let val delta = next_delta () in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   395
          {delta = delta, epsilon = delta, exclusive = true, total = false}
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   396
        end
38162
824e940a3dd0 minor symmetry breaking for codatatypes like llist
blanchet
parents: 38127
diff changeset
   397
      else if num_self_recs > 0 then
38193
44d635ce6da4 improve datatype symmetry breaking;
blanchet
parents: 38190
diff changeset
   398
        (if num_non_self_recs = 1 then
35072
d79308423aea optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents: 35071
diff changeset
   399
           if self_rec then
38193
44d635ce6da4 improve datatype symmetry breaking;
blanchet
parents: 38190
diff changeset
   400
             case List.last constrs of
44d635ce6da4 improve datatype symmetry breaking;
blanchet
parents: 38190
diff changeset
   401
               {delta = 0, epsilon = 1, exclusive = true, ...} =>
44d635ce6da4 improve datatype symmetry breaking;
blanchet
parents: 38190
diff changeset
   402
               {delta = 1, epsilon = card, exclusive = (num_self_recs = 1),
44d635ce6da4 improve datatype symmetry breaking;
blanchet
parents: 38190
diff changeset
   403
                total = false}
35072
d79308423aea optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents: 35071
diff changeset
   404
             | _ => raise SAME ()
d79308423aea optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents: 35071
diff changeset
   405
           else
d79308423aea optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents: 35071
diff changeset
   406
             if domain_card 2 card_assigns T = 1 then
38162
824e940a3dd0 minor symmetry breaking for codatatypes like llist
blanchet
parents: 38127
diff changeset
   407
               {delta = 0, epsilon = 1, exclusive = acyclic, total = acyclic}
35072
d79308423aea optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents: 35071
diff changeset
   408
             else
d79308423aea optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents: 35071
diff changeset
   409
               raise SAME ()
d79308423aea optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents: 35071
diff changeset
   410
         else
d79308423aea optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents: 35071
diff changeset
   411
           raise SAME ())
d79308423aea optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents: 35071
diff changeset
   412
        handle SAME () =>
d79308423aea optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents: 35071
diff changeset
   413
               {delta = 0, epsilon = card, exclusive = false, total = false}
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   414
      else if card = sum_dom_cards (card + 1) then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   415
        let val delta = next_delta () in
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   416
          {delta = delta, epsilon = delta + domain_card card card_assigns T,
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   417
           exclusive = true, total = true}
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   418
        end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   419
      else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   420
        {delta = 0, epsilon = card,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   421
         exclusive = (num_self_recs + num_non_self_recs = 1), total = false}
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   422
  in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   423
    {const = x, delta = delta, epsilon = epsilon, exclusive = exclusive,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   424
     explicit_max = max, total = total} :: constrs
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   425
  end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   426
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   427
fun has_exact_card hol_ctxt facto finitizable_dataTs card_assigns T =
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   428
  let val card = card_of_type card_assigns T in
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   429
    card = bounded_exact_card_of_type hol_ctxt
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   430
               (if facto then finitizable_dataTs else []) (card + 1) 0
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   431
               card_assigns T
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   432
  end
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   433
55888
cac1add157e8 removed nonstandard models from Nitpick
blanchet
parents: 55628
diff changeset
   434
fun datatype_spec_from_scope_descriptor (hol_ctxt as {ctxt, ...})
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   435
        binarize deep_dataTs finitizable_dataTs (desc as (card_assigns, _))
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   436
        (T, card) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   437
  let
34982
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34936
diff changeset
   438
    val deep = member (op =) deep_dataTs T
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   439
    val co = is_codatatype ctxt T
35190
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   440
    val xs = binarized_and_boxed_datatype_constrs hol_ctxt binarize T
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   441
    val self_recs = map (is_self_recursive_constr_type o snd) xs
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   442
    val (num_self_recs, num_non_self_recs) =
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   443
      List.partition I self_recs |> pairself length
38126
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 38124
diff changeset
   444
    val self_rec = num_self_recs > 0
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   445
    fun is_complete facto =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   446
      has_exact_card hol_ctxt facto finitizable_dataTs card_assigns T
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   447
    fun is_concrete facto =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   448
      is_word_type T orelse
47909
5f1afeebafbc fixed "real" after they were redefined as a 'quotient_type'
blanchet
parents: 46083
diff changeset
   449
      (* FIXME: looks wrong; other types than just functions might be
45402
1fac64bbdb4f tweaked comment
blanchet
parents: 45399
diff changeset
   450
         abstract. "is_complete" is also suspicious. *)
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   451
      xs |> maps (binder_types o snd) |> maps binder_types
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   452
         |> forall (has_exact_card hol_ctxt facto finitizable_dataTs
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   453
                                   card_assigns)
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   454
    val complete = pair_from_fun is_complete
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   455
    val concrete = pair_from_fun is_concrete
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   456
    fun sum_dom_cards max =
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   457
      map (domain_card max card_assigns o snd) xs |> Integer.sum
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   458
    val constrs =
55888
cac1add157e8 removed nonstandard models from Nitpick
blanchet
parents: 55628
diff changeset
   459
      fold_rev (add_constr_spec desc (not co) card sum_dom_cards num_self_recs
cac1add157e8 removed nonstandard models from Nitpick
blanchet
parents: 55628
diff changeset
   460
                                num_non_self_recs)
35072
d79308423aea optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents: 35071
diff changeset
   461
               (sort (bool_ord o swap o pairself fst) (self_recs ~~ xs)) []
33558
a2db56854b83 optimized Nitpick's encoding and rendering of datatypes whose constructors don't appear in the problem
blanchet
parents: 33232
diff changeset
   462
  in
55888
cac1add157e8 removed nonstandard models from Nitpick
blanchet
parents: 55628
diff changeset
   463
    {typ = T, card = card, co = co, self_rec = self_rec, complete = complete,
cac1add157e8 removed nonstandard models from Nitpick
blanchet
parents: 55628
diff changeset
   464
     concrete = concrete, deep = deep, constrs = constrs}
33558
a2db56854b83 optimized Nitpick's encoding and rendering of datatypes whose constructors don't appear in the problem
blanchet
parents: 33232
diff changeset
   465
  end
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   466
55888
cac1add157e8 removed nonstandard models from Nitpick
blanchet
parents: 55628
diff changeset
   467
fun scope_from_descriptor (hol_ctxt as {ctxt, ...}) binarize deep_dataTs
36386
2132f15b366f Fruhjahrsputz: remove three mostly useless Nitpick options
blanchet
parents: 36385
diff changeset
   468
                          finitizable_dataTs (desc as (card_assigns, _)) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   469
  let
33558
a2db56854b83 optimized Nitpick's encoding and rendering of datatypes whose constructors don't appear in the problem
blanchet
parents: 33232
diff changeset
   470
    val datatypes =
35190
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   471
      map (datatype_spec_from_scope_descriptor hol_ctxt binarize deep_dataTs
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   472
                                               finitizable_dataTs desc)
55888
cac1add157e8 removed nonstandard models from Nitpick
blanchet
parents: 55628
diff changeset
   473
          (filter (is_datatype ctxt o fst) card_assigns)
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   474
    val bits = card_of_type card_assigns @{typ signed_bit} - 1
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   475
               handle TYPE ("Nitpick_HOL.card_of_type", _, _) =>
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   476
                      card_of_type card_assigns @{typ unsigned_bit}
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   477
                      handle TYPE ("Nitpick_HOL.card_of_type", _, _) => 0
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   478
    val bisim_depth = card_of_type card_assigns @{typ bisim_iterator} - 1
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   479
  in
35190
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   480
    {hol_ctxt = hol_ctxt, binarize = binarize, card_assigns = card_assigns,
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   481
     datatypes = datatypes, bits = bits, bisim_depth = bisim_depth,
38124
6538e25cf5dd started implementation of custom sym break
blanchet
parents: 37678
diff changeset
   482
     ofs = offset_table_for_card_assigns datatypes card_assigns}
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   483
  end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   484
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35385
diff changeset
   485
fun repair_cards_assigns_wrt_boxing_etc _ _ [] = []
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35385
diff changeset
   486
  | repair_cards_assigns_wrt_boxing_etc thy Ts ((SOME T, ks) :: cards_assigns) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   487
    (if is_fun_type T orelse is_pair_type T then
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35385
diff changeset
   488
       Ts |> filter (curry (type_match thy o swap) T) |> map (rpair ks o SOME)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   489
     else
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35385
diff changeset
   490
       [(SOME T, ks)]) @
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35385
diff changeset
   491
       repair_cards_assigns_wrt_boxing_etc thy Ts cards_assigns
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35385
diff changeset
   492
  | repair_cards_assigns_wrt_boxing_etc thy Ts ((NONE, ks) :: cards_assigns) =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35385
diff changeset
   493
    (NONE, ks) :: repair_cards_assigns_wrt_boxing_etc thy Ts cards_assigns
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   494
38180
7a88032f9265 bump up the max cardinalities, to use up more of the time given to us by the user
blanchet
parents: 38162
diff changeset
   495
val max_scopes = 5000
7a88032f9265 bump up the max cardinalities, to use up more of the time given to us by the user
blanchet
parents: 38162
diff changeset
   496
val distinct_threshold = 1000
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   497
36386
2132f15b366f Fruhjahrsputz: remove three mostly useless Nitpick options
blanchet
parents: 36385
diff changeset
   498
fun all_scopes (hol_ctxt as {thy, ...}) binarize cards_assigns maxes_assigns
2132f15b366f Fruhjahrsputz: remove three mostly useless Nitpick options
blanchet
parents: 36385
diff changeset
   499
               iters_assigns bitss bisim_depths mono_Ts nonmono_Ts deep_dataTs
2132f15b366f Fruhjahrsputz: remove three mostly useless Nitpick options
blanchet
parents: 36385
diff changeset
   500
               finitizable_dataTs =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   501
  let
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   502
    val cards_assigns =
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   503
      repair_cards_assigns_wrt_boxing_etc thy mono_Ts cards_assigns
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   504
    val blocks =
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   505
      blocks_for_types hol_ctxt binarize cards_assigns maxes_assigns
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38193
diff changeset
   506
                       iters_assigns bitss bisim_depths mono_Ts nonmono_Ts
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   507
    val ranks = map rank_of_block blocks
33580
45c33e97cb86 added datatype constructor cache in Nitpick (to speed up the scope enumeration) and never test more than 4096 scopes
blanchet
parents: 33558
diff changeset
   508
    val all = all_combinations_ordered_smartly (map (rpair 0) ranks)
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
   509
    val head = take max_scopes all
35190
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   510
    val descs =
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   511
      map_filter (scope_descriptor_from_combination hol_ctxt binarize blocks)
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35179
diff changeset
   512
                 head
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   513
  in
33580
45c33e97cb86 added datatype constructor cache in Nitpick (to speed up the scope enumeration) and never test more than 4096 scopes
blanchet
parents: 33558
diff changeset
   514
    (length all - length head,
45c33e97cb86 added datatype constructor cache in Nitpick (to speed up the scope enumeration) and never test more than 4096 scopes
blanchet
parents: 33558
diff changeset
   515
     descs |> length descs <= distinct_threshold ? distinct (op =)
36386
2132f15b366f Fruhjahrsputz: remove three mostly useless Nitpick options
blanchet
parents: 36385
diff changeset
   516
           |> map (scope_from_descriptor hol_ctxt binarize deep_dataTs
2132f15b366f Fruhjahrsputz: remove three mostly useless Nitpick options
blanchet
parents: 36385
diff changeset
   517
                                         finitizable_dataTs))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   518
  end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   519
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   520
end;