src/HOL/Tools/Nitpick/nitpick_scope.ML
author blanchet
Wed, 17 Feb 2010 20:46:50 +0100
changeset 35190 ce653cc27a94
parent 35179 4b198af5beb5
child 35220 2bcdae5f4fdb
permissions -rw-r--r--
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet@33982
     1
(*  Title:      HOL/Tools/Nitpick/nitpick_scope.ML
blanchet@33192
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@34969
     3
    Copyright   2008, 2009, 2010
blanchet@33192
     4
blanchet@33192
     5
Scope enumerator for Nitpick.
blanchet@33192
     6
*)
blanchet@33192
     7
blanchet@33192
     8
signature NITPICK_SCOPE =
blanchet@33192
     9
sig
blanchet@33705
    10
  type styp = Nitpick_Util.styp
blanchet@35067
    11
  type hol_context = Nitpick_HOL.hol_context
blanchet@33192
    12
blanchet@33192
    13
  type constr_spec = {
blanchet@33192
    14
    const: styp,
blanchet@33192
    15
    delta: int,
blanchet@33192
    16
    epsilon: int,
blanchet@33192
    17
    exclusive: bool,
blanchet@33192
    18
    explicit_max: int,
blanchet@33192
    19
    total: bool}
blanchet@33192
    20
blanchet@33192
    21
  type dtype_spec = {
blanchet@33192
    22
    typ: typ,
blanchet@33192
    23
    card: int,
blanchet@33192
    24
    co: bool,
blanchet@35179
    25
    standard: bool,
blanchet@34120
    26
    complete: bool,
blanchet@34120
    27
    concrete: bool,
blanchet@34969
    28
    deep: bool,
blanchet@33192
    29
    constrs: constr_spec list}
blanchet@33192
    30
blanchet@33192
    31
  type scope = {
blanchet@35067
    32
    hol_ctxt: hol_context,
blanchet@35190
    33
    binarize: bool,
blanchet@33192
    34
    card_assigns: (typ * int) list,
blanchet@34121
    35
    bits: int,
blanchet@33192
    36
    bisim_depth: int,
blanchet@33192
    37
    datatypes: dtype_spec list,
blanchet@33192
    38
    ofs: int Typtab.table}
blanchet@33192
    39
blanchet@33192
    40
  val datatype_spec : dtype_spec list -> typ -> dtype_spec option
blanchet@33192
    41
  val constr_spec : dtype_spec list -> styp -> constr_spec
blanchet@34120
    42
  val is_complete_type : dtype_spec list -> typ -> bool
blanchet@34120
    43
  val is_concrete_type : dtype_spec list -> typ -> bool
blanchet@34120
    44
  val is_exact_type : dtype_spec list -> typ -> bool
blanchet@33192
    45
  val offset_of_type : int Typtab.table -> typ -> int
blanchet@33192
    46
  val spec_of_type : scope -> typ -> int * int
blanchet@33192
    47
  val pretties_for_scope : scope -> bool -> Pretty.T list
blanchet@33192
    48
  val multiline_string_for_scope : scope -> string
blanchet@33192
    49
  val scopes_equivalent : scope -> scope -> bool
blanchet@33192
    50
  val scope_less_eq : scope -> scope -> bool
blanchet@33192
    51
  val all_scopes :
blanchet@35190
    52
    hol_context -> bool -> int -> (typ option * int list) list
blanchet@33192
    53
    -> (styp option * int list) list -> (styp option * int list) list
blanchet@34121
    54
    -> int list -> int list -> typ list -> typ list -> typ list
blanchet@34121
    55
    -> int * scope list
blanchet@33192
    56
end;
blanchet@33192
    57
blanchet@33224
    58
structure Nitpick_Scope : NITPICK_SCOPE =
blanchet@33192
    59
struct
blanchet@33192
    60
blanchet@33224
    61
open Nitpick_Util
blanchet@33224
    62
open Nitpick_HOL
blanchet@33192
    63
blanchet@33192
    64
type constr_spec = {
blanchet@33192
    65
  const: styp,
blanchet@33192
    66
  delta: int,
blanchet@33192
    67
  epsilon: int,
blanchet@33192
    68
  exclusive: bool,
blanchet@33192
    69
  explicit_max: int,
blanchet@33192
    70
  total: bool}
blanchet@33192
    71
blanchet@33192
    72
type dtype_spec = {
blanchet@33192
    73
  typ: typ,
blanchet@33192
    74
  card: int,
blanchet@33192
    75
  co: bool,
blanchet@35179
    76
  standard: bool,
blanchet@34120
    77
  complete: bool,
blanchet@34120
    78
  concrete: bool,
blanchet@34969
    79
  deep: bool,
blanchet@33192
    80
  constrs: constr_spec list}
blanchet@33192
    81
blanchet@33192
    82
type scope = {
blanchet@35067
    83
  hol_ctxt: hol_context,
blanchet@35190
    84
  binarize: bool,
blanchet@33192
    85
  card_assigns: (typ * int) list,
blanchet@34121
    86
  bits: int,
blanchet@33192
    87
  bisim_depth: int,
blanchet@33192
    88
  datatypes: dtype_spec list,
blanchet@33192
    89
  ofs: int Typtab.table}
blanchet@33192
    90
blanchet@33192
    91
datatype row_kind = Card of typ | Max of styp
blanchet@33192
    92
blanchet@33192
    93
type row = row_kind * int list
blanchet@33192
    94
type block = row list
blanchet@33192
    95
blanchet@33192
    96
(* dtype_spec list -> typ -> dtype_spec option *)
blanchet@33192
    97
fun datatype_spec (dtypes : dtype_spec list) T =
blanchet@34118
    98
  List.find (curry (op =) T o #typ) dtypes
blanchet@33192
    99
blanchet@33192
   100
(* dtype_spec list -> styp -> constr_spec *)
blanchet@33224
   101
fun constr_spec [] x = raise TERM ("Nitpick_Scope.constr_spec", [Const x])
blanchet@33192
   102
  | constr_spec ({constrs, ...} :: dtypes : dtype_spec list) (x as (s, T)) =
blanchet@34118
   103
    case List.find (curry (op =) (s, body_type T) o (apsnd body_type o #const))
blanchet@33192
   104
                   constrs of
blanchet@33192
   105
      SOME c => c
blanchet@33192
   106
    | NONE => constr_spec dtypes x
blanchet@33192
   107
blanchet@33192
   108
(* dtype_spec list -> typ -> bool *)
blanchet@34120
   109
fun is_complete_type dtypes (Type ("fun", [T1, T2])) =
blanchet@34120
   110
    is_concrete_type dtypes T1 andalso is_complete_type dtypes T2
blanchet@34120
   111
  | is_complete_type dtypes (Type ("*", Ts)) =
blanchet@34120
   112
    forall (is_complete_type dtypes) Ts
blanchet@34120
   113
  | is_complete_type dtypes T =
blanchet@34923
   114
    not (is_integer_type T) andalso not (is_bit_type T) andalso
blanchet@34923
   115
    #complete (the (datatype_spec dtypes T))
blanchet@33192
   116
    handle Option.Option => true
blanchet@34120
   117
and is_concrete_type dtypes (Type ("fun", [T1, T2])) =
blanchet@34120
   118
    is_complete_type dtypes T1 andalso is_concrete_type dtypes T2
blanchet@34120
   119
  | is_concrete_type dtypes (Type ("*", Ts)) =
blanchet@34120
   120
    forall (is_concrete_type dtypes) Ts
blanchet@34120
   121
  | is_concrete_type dtypes T =
blanchet@34120
   122
    #concrete (the (datatype_spec dtypes T)) handle Option.Option => true
blanchet@34120
   123
fun is_exact_type dtypes = is_complete_type dtypes andf is_concrete_type dtypes
blanchet@33192
   124
blanchet@33192
   125
(* int Typtab.table -> typ -> int *)
blanchet@33192
   126
fun offset_of_type ofs T =
blanchet@33192
   127
  case Typtab.lookup ofs T of
blanchet@33192
   128
    SOME j0 => j0
blanchet@33192
   129
  | NONE => Typtab.lookup ofs dummyT |> the_default 0
blanchet@33192
   130
blanchet@33192
   131
(* scope -> typ -> int * int *)
blanchet@33192
   132
fun spec_of_type ({card_assigns, ofs, ...} : scope) T =
blanchet@33192
   133
  (card_of_type card_assigns T
blanchet@33224
   134
   handle TYPE ("Nitpick_HOL.card_of_type", _, _) => ~1, offset_of_type ofs T)
blanchet@33192
   135
blanchet@33192
   136
(* (string -> string) -> scope
blanchet@33192
   137
   -> string list * string list * string list * string list * string list *)
blanchet@35067
   138
fun quintuple_for_scope quote ({hol_ctxt as {thy, ctxt, ...}, card_assigns,
blanchet@34121
   139
                                bits, bisim_depth, datatypes, ...} : scope) =
blanchet@33192
   140
  let
blanchet@35072
   141
    val boring_Ts = [@{typ unsigned_bit}, @{typ signed_bit},
blanchet@34121
   142
                     @{typ bisim_iterator}]
blanchet@34120
   143
    val (iter_assigns, card_assigns) =
blanchet@34121
   144
      card_assigns |> filter_out (member (op =) boring_Ts o fst)
blanchet@33192
   145
                   |> List.partition (is_fp_iterator_type o fst)
blanchet@34121
   146
    val (secondary_card_assigns, primary_card_assigns) =
blanchet@34121
   147
      card_assigns |> List.partition ((is_integer_type orf is_datatype thy)
blanchet@34121
   148
                                      o fst)
blanchet@33192
   149
    val cards =
blanchet@33192
   150
      map (fn (T, k) => quote (string_for_type ctxt T) ^ " = " ^
blanchet@33192
   151
                        string_of_int k)
blanchet@33192
   152
    fun maxes () =
blanchet@33192
   153
      maps (map_filter
blanchet@33192
   154
                (fn {const, explicit_max, ...} =>
blanchet@33192
   155
                    if explicit_max < 0 then
blanchet@33192
   156
                      NONE
blanchet@33192
   157
                    else
blanchet@33192
   158
                      SOME (Syntax.string_of_term ctxt (Const const) ^ " = " ^
blanchet@33192
   159
                            string_of_int explicit_max))
blanchet@33192
   160
                 o #constrs) datatypes
blanchet@33192
   161
    fun iters () =
blanchet@33192
   162
      map (fn (T, k) =>
blanchet@33192
   163
              quote (Syntax.string_of_term ctxt
blanchet@33192
   164
                         (Const (const_for_iterator_type T))) ^ " = " ^
blanchet@34120
   165
              string_of_int (k - 1)) iter_assigns
blanchet@34121
   166
    fun miscs () =
blanchet@34121
   167
      (if bits = 0 then [] else ["bits = " ^ string_of_int bits]) @
blanchet@34121
   168
      (if bisim_depth < 0 andalso forall (not o #co) datatypes then []
blanchet@35178
   169
       else ["bisim_depth = " ^ signed_string_of_int bisim_depth])
blanchet@33192
   170
  in
blanchet@33192
   171
    setmp_show_all_types
blanchet@34121
   172
        (fn () => (cards primary_card_assigns, cards secondary_card_assigns,
blanchet@34121
   173
                   maxes (), iters (), miscs ())) ()
blanchet@33192
   174
  end
blanchet@33192
   175
blanchet@33192
   176
(* scope -> bool -> Pretty.T list *)
blanchet@33192
   177
fun pretties_for_scope scope verbose =
blanchet@33192
   178
  let
blanchet@34121
   179
    val (primary_cards, secondary_cards, maxes, iters, bisim_depths) =
blanchet@33192
   180
      quintuple_for_scope maybe_quote scope
blanchet@34121
   181
    val ss = map (prefix "card ") primary_cards @
blanchet@33192
   182
             (if verbose then
blanchet@34121
   183
                map (prefix "card ") secondary_cards @
blanchet@33192
   184
                map (prefix "max ") maxes @
blanchet@33192
   185
                map (prefix "iter ") iters @
blanchet@33192
   186
                bisim_depths
blanchet@33192
   187
              else
blanchet@33192
   188
                [])
blanchet@33192
   189
  in
blanchet@33192
   190
    if null ss then []
blanchet@33192
   191
    else serial_commas "and" ss |> map Pretty.str |> Pretty.breaks
blanchet@33192
   192
  end
blanchet@33192
   193
blanchet@33192
   194
(* scope -> string *)
blanchet@33192
   195
fun multiline_string_for_scope scope =
blanchet@33192
   196
  let
blanchet@34121
   197
    val (primary_cards, secondary_cards, maxes, iters, bisim_depths) =
blanchet@33192
   198
      quintuple_for_scope I scope
blanchet@34121
   199
    val cards = primary_cards @ secondary_cards
blanchet@33192
   200
  in
blanchet@33192
   201
    case (if null cards then [] else ["card: " ^ commas cards]) @
blanchet@33192
   202
         (if null maxes then [] else ["max: " ^ commas maxes]) @
blanchet@33192
   203
         (if null iters then [] else ["iter: " ^ commas iters]) @
blanchet@33192
   204
         bisim_depths of
blanchet@33192
   205
      [] => "empty"
blanchet@33192
   206
    | lines => space_implode "\n" lines
blanchet@33192
   207
  end
blanchet@33192
   208
blanchet@33192
   209
(* scope -> scope -> bool *)
blanchet@33192
   210
fun scopes_equivalent (s1 : scope) (s2 : scope) =
blanchet@33192
   211
  #datatypes s1 = #datatypes s2 andalso #card_assigns s1 = #card_assigns s2
blanchet@33192
   212
fun scope_less_eq (s1 : scope) (s2 : scope) =
blanchet@33192
   213
  (s1, s2) |> pairself (map snd o #card_assigns) |> op ~~ |> forall (op <=)
blanchet@33192
   214
blanchet@33192
   215
(* row -> int *)
blanchet@33192
   216
fun rank_of_row (_, ks) = length ks
blanchet@33192
   217
(* block -> int *)
blanchet@33192
   218
fun rank_of_block block = fold Integer.max (map rank_of_row block) 1
blanchet@33192
   219
(* int -> typ * int list -> typ * int list *)
blanchet@33192
   220
fun project_row column (y, ks) = (y, [nth ks (Int.min (column, length ks - 1))])
blanchet@33192
   221
(* int -> block -> block *)
blanchet@33192
   222
fun project_block (column, block) = map (project_row column) block
blanchet@33192
   223
blanchet@33192
   224
(* (''a * ''a -> bool) -> (''a option * int list) list -> ''a -> int list *)
blanchet@34120
   225
fun lookup_ints_assign eq assigns key =
blanchet@34120
   226
  case triple_lookup eq assigns key of
blanchet@33192
   227
    SOME ks => ks
blanchet@33224
   228
  | NONE => raise ARG ("Nitpick_Scope.lookup_ints_assign", "")
blanchet@33192
   229
(* theory -> (typ option * int list) list -> typ -> int list *)
blanchet@34120
   230
fun lookup_type_ints_assign thy assigns T =
blanchet@34120
   231
  map (curry Int.max 1) (lookup_ints_assign (type_match thy) assigns T)
blanchet@33224
   232
  handle ARG ("Nitpick_Scope.lookup_ints_assign", _) =>
blanchet@33224
   233
         raise TYPE ("Nitpick_Scope.lookup_type_ints_assign", [T], [])
blanchet@33192
   234
(* theory -> (styp option * int list) list -> styp -> int list *)
blanchet@34120
   235
fun lookup_const_ints_assign thy assigns x =
blanchet@34120
   236
  lookup_ints_assign (const_match thy) assigns x
blanchet@33224
   237
  handle ARG ("Nitpick_Scope.lookup_ints_assign", _) =>
blanchet@33224
   238
         raise TERM ("Nitpick_Scope.lookup_const_ints_assign", [Const x])
blanchet@33192
   239
blanchet@33192
   240
(* theory -> (styp option * int list) list -> styp -> row option *)
blanchet@34120
   241
fun row_for_constr thy maxes_assigns constr =
blanchet@34120
   242
  SOME (Max constr, lookup_const_ints_assign thy maxes_assigns constr)
blanchet@33192
   243
  handle TERM ("lookup_const_ints_assign", _) => NONE
blanchet@33192
   244
blanchet@34121
   245
val max_bits = 31 (* Kodkod limit *)
blanchet@34121
   246
blanchet@35190
   247
(* hol_context -> bool -> (typ option * int list) list
blanchet@35190
   248
   -> (styp option * int list) list -> (styp option * int list) list -> int list
blanchet@35190
   249
   -> int list -> typ -> block *)
blanchet@35190
   250
fun block_for_type (hol_ctxt as {thy, ...}) binarize cards_assigns maxes_assigns
blanchet@34121
   251
                   iters_assigns bitss bisim_depths T =
blanchet@34121
   252
  if T = @{typ unsigned_bit} then
blanchet@34121
   253
    [(Card T, map (Integer.min max_bits o Integer.max 1) bitss)]
blanchet@34121
   254
  else if T = @{typ signed_bit} then
blanchet@34121
   255
    [(Card T, map (Integer.add 1 o Integer.min max_bits o Integer.max 1) bitss)]
blanchet@34123
   256
  else if T = @{typ "unsigned_bit word"} then
blanchet@34123
   257
    [(Card T, lookup_type_ints_assign thy cards_assigns nat_T)]
blanchet@34123
   258
  else if T = @{typ "signed_bit word"} then
blanchet@34123
   259
    [(Card T, lookup_type_ints_assign thy cards_assigns int_T)]
blanchet@34121
   260
  else if T = @{typ bisim_iterator} then
blanchet@34121
   261
    [(Card T, map (Integer.add 1 o Integer.max 0) bisim_depths)]
blanchet@33571
   262
  else if is_fp_iterator_type T then
blanchet@34121
   263
    [(Card T, map (Integer.add 1 o Integer.max 0)
blanchet@34120
   264
                  (lookup_const_ints_assign thy iters_assigns
blanchet@33571
   265
                                            (const_for_iterator_type T)))]
blanchet@33571
   266
  else
blanchet@34120
   267
    (Card T, lookup_type_ints_assign thy cards_assigns T) ::
blanchet@35190
   268
    (case binarized_and_boxed_datatype_constrs hol_ctxt binarize T of
blanchet@33571
   269
       [_] => []
blanchet@34120
   270
     | constrs => map_filter (row_for_constr thy maxes_assigns) constrs)
blanchet@33192
   271
blanchet@35190
   272
(* hol_context -> bool -> (typ option * int list) list
blanchet@35190
   273
   -> (styp option * int list) list -> (styp option * int list) list -> int list
blanchet@35190
   274
   -> int list -> typ list -> typ list -> block list *)
blanchet@35190
   275
fun blocks_for_types hol_ctxt binarize cards_assigns maxes_assigns iters_assigns
blanchet@35190
   276
                     bitss bisim_depths mono_Ts nonmono_Ts =
blanchet@33192
   277
  let
blanchet@33192
   278
    (* typ -> block *)
blanchet@35190
   279
    val block_for = block_for_type hol_ctxt binarize cards_assigns maxes_assigns
blanchet@34121
   280
                                   iters_assigns bitss bisim_depths
blanchet@33192
   281
    val mono_block = maps block_for mono_Ts
blanchet@33192
   282
    val nonmono_blocks = map block_for nonmono_Ts
blanchet@33192
   283
  in mono_block :: nonmono_blocks end
blanchet@33192
   284
blanchet@33192
   285
val sync_threshold = 5
blanchet@33192
   286
blanchet@33192
   287
(* int list -> int list list *)
blanchet@33192
   288
fun all_combinations_ordered_smartly ks =
blanchet@33192
   289
  let
blanchet@33192
   290
    (* int list -> int *)
blanchet@33192
   291
    fun cost_with_monos [] = 0
blanchet@33192
   292
      | cost_with_monos (k :: ks) =
blanchet@34118
   293
        if k < sync_threshold andalso forall (curry (op =) k) ks then
blanchet@33192
   294
          k - sync_threshold
blanchet@33192
   295
        else
blanchet@33192
   296
          k * (k + 1) div 2 + Integer.sum ks
blanchet@33192
   297
    fun cost_without_monos [] = 0
blanchet@33192
   298
      | cost_without_monos [k] = k
blanchet@33192
   299
      | cost_without_monos (_ :: k :: ks) =
blanchet@34118
   300
        if k < sync_threshold andalso forall (curry (op =) k) ks then
blanchet@33192
   301
          k - sync_threshold
blanchet@33192
   302
        else
blanchet@33192
   303
          Integer.sum (k :: ks)
blanchet@33192
   304
  in
blanchet@33192
   305
    ks |> all_combinations
blanchet@33192
   306
       |> map (`(if fst (hd ks) > 1 then cost_with_monos
blanchet@33192
   307
                 else cost_without_monos))
blanchet@33192
   308
       |> sort (int_ord o pairself fst) |> map snd
blanchet@33192
   309
  end
blanchet@33192
   310
blanchet@33192
   311
(* typ -> bool *)
blanchet@33192
   312
fun is_self_recursive_constr_type T =
blanchet@34118
   313
  exists (exists_subtype (curry (op =) (body_type T))) (binder_types T)
blanchet@33192
   314
blanchet@33192
   315
(* (styp * int) list -> styp -> int *)
blanchet@33192
   316
fun constr_max maxes x = the_default ~1 (AList.lookup (op =) maxes x)
blanchet@33192
   317
blanchet@33192
   318
type scope_desc = (typ * int) list * (styp * int) list
blanchet@33192
   319
blanchet@35190
   320
(* hol_context -> bool -> scope_desc -> typ * int -> bool *)
blanchet@35190
   321
fun is_surely_inconsistent_card_assign hol_ctxt binarize
blanchet@35190
   322
                                       (card_assigns, max_assigns) (T, k) =
blanchet@35190
   323
  case binarized_and_boxed_datatype_constrs hol_ctxt binarize T of
blanchet@33192
   324
    [] => false
blanchet@33192
   325
  | xs =>
blanchet@33192
   326
    let
blanchet@34123
   327
      val dom_cards =
blanchet@34123
   328
        map (Integer.prod o map (bounded_card_of_type k ~1 card_assigns)
blanchet@33192
   329
             o binder_types o snd) xs
blanchet@34120
   330
      val maxes = map (constr_max max_assigns) xs
blanchet@33192
   331
      (* int -> int -> int *)
blanchet@34123
   332
      fun effective_max card ~1 = card
blanchet@33192
   333
        | effective_max card max = Int.min (card, max)
blanchet@34123
   334
      val max = map2 effective_max dom_cards maxes |> Integer.sum
blanchet@34123
   335
    in max < k end
blanchet@35190
   336
(* hol_context -> bool -> (typ * int) list -> (typ * int) list
blanchet@35190
   337
   -> (styp * int) list -> bool *)
blanchet@35190
   338
fun is_surely_inconsistent_scope_description hol_ctxt binarize seen rest
blanchet@35190
   339
                                             max_assigns =
blanchet@35190
   340
  exists (is_surely_inconsistent_card_assign hol_ctxt binarize
blanchet@34123
   341
                                             (seen @ rest, max_assigns)) seen
blanchet@33192
   342
blanchet@35190
   343
(* hol_context -> bool -> scope_desc -> (typ * int) list option *)
blanchet@35190
   344
fun repair_card_assigns hol_ctxt binarize (card_assigns, max_assigns) =
blanchet@33192
   345
  let
blanchet@33192
   346
    (* (typ * int) list -> (typ * int) list -> (typ * int) list option *)
blanchet@33192
   347
    fun aux seen [] = SOME seen
blanchet@33192
   348
      | aux seen ((T, 0) :: _) = NONE
blanchet@34123
   349
      | aux seen ((T, k) :: rest) =
blanchet@35190
   350
        (if is_surely_inconsistent_scope_description hol_ctxt binarize
blanchet@35190
   351
                ((T, k) :: seen) rest max_assigns then
blanchet@33192
   352
           raise SAME ()
blanchet@33192
   353
         else
blanchet@34123
   354
           case aux ((T, k) :: seen) rest of
blanchet@34120
   355
             SOME assigns => SOME assigns
blanchet@33192
   356
           | NONE => raise SAME ())
blanchet@34123
   357
        handle SAME () => aux seen ((T, k - 1) :: rest)
blanchet@34120
   358
  in aux [] (rev card_assigns) end
blanchet@33192
   359
blanchet@33192
   360
(* theory -> (typ * int) list -> typ * int -> typ * int *)
blanchet@34120
   361
fun repair_iterator_assign thy assigns (T as Type (s, Ts), k) =
blanchet@33192
   362
    (T, if T = @{typ bisim_iterator} then
blanchet@34120
   363
          let
blanchet@34120
   364
            val co_cards = map snd (filter (is_codatatype thy o fst) assigns)
blanchet@34120
   365
          in Int.min (k, Integer.sum co_cards) end
blanchet@33192
   366
        else if is_fp_iterator_type T then
blanchet@33192
   367
          case Ts of
blanchet@33192
   368
            [] => 1
blanchet@34120
   369
          | _ => bounded_card_of_type k ~1 assigns (foldr1 HOLogic.mk_prodT Ts)
blanchet@33192
   370
        else
blanchet@33192
   371
          k)
blanchet@34120
   372
  | repair_iterator_assign _ _ assign = assign
blanchet@33192
   373
blanchet@33192
   374
(* row -> scope_desc -> scope_desc *)
blanchet@34120
   375
fun add_row_to_scope_descriptor (kind, ks) (card_assigns, max_assigns) =
blanchet@33192
   376
  case kind of
blanchet@34120
   377
    Card T => ((T, the_single ks) :: card_assigns, max_assigns)
blanchet@34120
   378
  | Max x => (card_assigns, (x, the_single ks) :: max_assigns)
blanchet@33192
   379
(* block -> scope_desc *)
blanchet@33192
   380
fun scope_descriptor_from_block block =
blanchet@33192
   381
  fold_rev add_row_to_scope_descriptor block ([], [])
blanchet@35190
   382
(* hol_context -> bool -> block list -> int list -> scope_desc option *)
blanchet@35190
   383
fun scope_descriptor_from_combination (hol_ctxt as {thy, ...}) binarize blocks
blanchet@35190
   384
                                      columns =
blanchet@33192
   385
  let
blanchet@34120
   386
    val (card_assigns, max_assigns) =
blanchet@33192
   387
      maps project_block (columns ~~ blocks) |> scope_descriptor_from_block
blanchet@35190
   388
    val card_assigns =
blanchet@35190
   389
      repair_card_assigns hol_ctxt binarize (card_assigns, max_assigns) |> the
blanchet@33192
   390
  in
blanchet@34120
   391
    SOME (map (repair_iterator_assign thy card_assigns) card_assigns,
blanchet@34120
   392
          max_assigns)
blanchet@33192
   393
  end
blanchet@33192
   394
  handle Option.Option => NONE
blanchet@33192
   395
blanchet@33192
   396
(* theory -> (typ * int) list -> dtype_spec list -> int Typtab.table *)
blanchet@34120
   397
fun offset_table_for_card_assigns thy assigns dtypes =
blanchet@33192
   398
  let
blanchet@33192
   399
    (* int -> (int * int) list -> (typ * int) list -> int Typtab.table
blanchet@33192
   400
       -> int Typtab.table *)
blanchet@33192
   401
    fun aux next _ [] = Typtab.update_new (dummyT, next)
blanchet@34120
   402
      | aux next reusable ((T, k) :: assigns) =
blanchet@34121
   403
        if k = 1 orelse is_integer_type T orelse is_bit_type T then
blanchet@34120
   404
          aux next reusable assigns
blanchet@33192
   405
        else if length (these (Option.map #constrs (datatype_spec dtypes T)))
blanchet@33192
   406
                > 1 then
blanchet@34120
   407
          Typtab.update_new (T, next) #> aux (next + k) reusable assigns
blanchet@33192
   408
        else
blanchet@33192
   409
          case AList.lookup (op =) reusable k of
blanchet@34120
   410
            SOME j0 => Typtab.update_new (T, j0) #> aux next reusable assigns
blanchet@33192
   411
          | NONE => Typtab.update_new (T, next)
blanchet@34120
   412
                    #> aux (next + k) ((k, next) :: reusable) assigns
blanchet@34120
   413
  in aux 0 [] assigns Typtab.empty end
blanchet@33192
   414
blanchet@33192
   415
(* int -> (typ * int) list -> typ -> int *)
blanchet@34120
   416
fun domain_card max card_assigns =
blanchet@34120
   417
  Integer.prod o map (bounded_card_of_type max max card_assigns) o binder_types
blanchet@33192
   418
blanchet@33192
   419
(* scope_desc -> bool -> int -> (int -> int) -> int -> int -> bool * styp
blanchet@33192
   420
   -> constr_spec list -> constr_spec list *)
blanchet@34120
   421
fun add_constr_spec (card_assigns, max_assigns) co card sum_dom_cards
blanchet@34120
   422
                    num_self_recs num_non_self_recs (self_rec, x as (s, T))
blanchet@34120
   423
                    constrs =
blanchet@33192
   424
  let
blanchet@34120
   425
    val max = constr_max max_assigns x
blanchet@33192
   426
    (* int -> int *)
blanchet@33192
   427
    fun bound k = Int.min (card, (max >= 0 ? curry Int.min max) k)
blanchet@33192
   428
    (* unit -> int *)
blanchet@33192
   429
    fun next_delta () = if null constrs then 0 else #epsilon (hd constrs)
blanchet@33192
   430
    val {delta, epsilon, exclusive, total} =
blanchet@33192
   431
      if max = 0 then
blanchet@33192
   432
        let val delta = next_delta () in
blanchet@33192
   433
          {delta = delta, epsilon = delta, exclusive = true, total = false}
blanchet@33192
   434
        end
blanchet@33192
   435
      else if not co andalso num_self_recs > 0 then
blanchet@35069
   436
        (if num_self_recs = 1 andalso num_non_self_recs = 1 then
blanchet@35069
   437
           if self_rec then
blanchet@35069
   438
             case constrs of
blanchet@35069
   439
               [{delta = 0, epsilon = 1, exclusive = true, ...}] =>
blanchet@35069
   440
               {delta = 1, epsilon = card, exclusive = true, total = false}
blanchet@35069
   441
             | _ => raise SAME ()
blanchet@35069
   442
           else
blanchet@35069
   443
             if domain_card 2 card_assigns T = 1 then
blanchet@35069
   444
               {delta = 0, epsilon = 1, exclusive = true, total = true}
blanchet@35069
   445
             else
blanchet@35069
   446
               raise SAME ()
blanchet@35069
   447
         else
blanchet@35069
   448
           raise SAME ())
blanchet@35069
   449
        handle SAME () =>
blanchet@35069
   450
               {delta = 0, epsilon = card, exclusive = false, total = false}
blanchet@33192
   451
      else if card = sum_dom_cards (card + 1) then
blanchet@33192
   452
        let val delta = next_delta () in
blanchet@34120
   453
          {delta = delta, epsilon = delta + domain_card card card_assigns T,
blanchet@33192
   454
           exclusive = true, total = true}
blanchet@33192
   455
        end
blanchet@33192
   456
      else
blanchet@33192
   457
        {delta = 0, epsilon = card,
blanchet@33192
   458
         exclusive = (num_self_recs + num_non_self_recs = 1), total = false}
blanchet@33192
   459
  in
blanchet@33192
   460
    {const = x, delta = delta, epsilon = epsilon, exclusive = exclusive,
blanchet@33192
   461
     explicit_max = max, total = total} :: constrs
blanchet@33192
   462
  end
blanchet@33192
   463
blanchet@35067
   464
(* hol_context -> (typ * int) list -> typ -> bool *)
blanchet@35067
   465
fun has_exact_card hol_ctxt card_assigns T =
blanchet@34120
   466
  let val card = card_of_type card_assigns T in
blanchet@35067
   467
    card = bounded_exact_card_of_type hol_ctxt (card + 1) 0 card_assigns T
blanchet@34120
   468
  end
blanchet@34120
   469
blanchet@35190
   470
(* hol_context -> bool -> typ list -> scope_desc -> typ * int -> dtype_spec *)
blanchet@35190
   471
fun datatype_spec_from_scope_descriptor (hol_ctxt as {thy, ...}) binarize
blanchet@35190
   472
        deep_dataTs (desc as (card_assigns, _)) (T, card) =
blanchet@33192
   473
  let
blanchet@34969
   474
    val deep = member (op =) deep_dataTs T
blanchet@33192
   475
    val co = is_codatatype thy T
blanchet@35179
   476
    val standard = is_standard_datatype hol_ctxt T
blanchet@35190
   477
    val xs = binarized_and_boxed_datatype_constrs hol_ctxt binarize T
blanchet@33192
   478
    val self_recs = map (is_self_recursive_constr_type o snd) xs
blanchet@33192
   479
    val (num_self_recs, num_non_self_recs) =
blanchet@34120
   480
      List.partition I self_recs |> pairself length
blanchet@35067
   481
    val complete = has_exact_card hol_ctxt card_assigns T
blanchet@34120
   482
    val concrete = xs |> maps (binder_types o snd) |> maps binder_types
blanchet@35067
   483
                      |> forall (has_exact_card hol_ctxt card_assigns)
blanchet@33192
   484
    (* int -> int *)
blanchet@33192
   485
    fun sum_dom_cards max =
blanchet@34120
   486
      map (domain_card max card_assigns o snd) xs |> Integer.sum
blanchet@33192
   487
    val constrs =
blanchet@33192
   488
      fold_rev (add_constr_spec desc co card sum_dom_cards num_self_recs
blanchet@35069
   489
                                num_non_self_recs)
blanchet@35069
   490
               (sort (bool_ord o swap o pairself fst) (self_recs ~~ xs)) []
blanchet@33549
   491
  in
blanchet@35179
   492
    {typ = T, card = card, co = co, standard = standard, complete = complete,
blanchet@35179
   493
     concrete = concrete, deep = deep, constrs = constrs}
blanchet@33549
   494
  end
blanchet@33192
   495
blanchet@34121
   496
(* int -> int *)
blanchet@34121
   497
fun min_bits_for_nat_value n = if n <= 0 then 0 else IntInf.log2 n + 1
blanchet@34121
   498
(* scope_desc -> int *)
blanchet@34121
   499
fun min_bits_for_max_assigns (_, []) = 0
blanchet@34121
   500
  | min_bits_for_max_assigns (card_assigns, max_assigns) =
blanchet@34121
   501
    min_bits_for_nat_value (fold Integer.max
blanchet@34121
   502
        (map snd card_assigns @ map snd max_assigns) 0)
blanchet@34121
   503
blanchet@35190
   504
(* hol_context -> bool -> int -> typ list -> scope_desc -> scope *)
blanchet@35190
   505
fun scope_from_descriptor (hol_ctxt as {thy, ...}) binarize sym_break
blanchet@35190
   506
                          deep_dataTs (desc as (card_assigns, _)) =
blanchet@33192
   507
  let
blanchet@33549
   508
    val datatypes =
blanchet@35190
   509
      map (datatype_spec_from_scope_descriptor hol_ctxt binarize deep_dataTs
blanchet@35190
   510
               desc) (filter (is_datatype thy o fst) card_assigns)
blanchet@34121
   511
    val bits = card_of_type card_assigns @{typ signed_bit} - 1
blanchet@34121
   512
               handle TYPE ("Nitpick_HOL.card_of_type", _, _) =>
blanchet@34121
   513
                      card_of_type card_assigns @{typ unsigned_bit}
blanchet@34121
   514
                      handle TYPE ("Nitpick_HOL.card_of_type", _, _) => 0
blanchet@34120
   515
    val bisim_depth = card_of_type card_assigns @{typ bisim_iterator} - 1
blanchet@33192
   516
  in
blanchet@35190
   517
    {hol_ctxt = hol_ctxt, binarize = binarize, card_assigns = card_assigns,
blanchet@35190
   518
     datatypes = datatypes, bits = bits, bisim_depth = bisim_depth,
blanchet@33192
   519
     ofs = if sym_break <= 0 then Typtab.empty
blanchet@34120
   520
           else offset_table_for_card_assigns thy card_assigns datatypes}
blanchet@33192
   521
  end
blanchet@33192
   522
blanchet@33192
   523
(* theory -> typ list -> (typ option * int list) list
blanchet@33192
   524
   -> (typ option * int list) list *)
blanchet@34121
   525
fun repair_cards_assigns_wrt_boxing _ _ [] = []
blanchet@34121
   526
  | repair_cards_assigns_wrt_boxing thy Ts ((SOME T, ks) :: cards_assigns) =
blanchet@33192
   527
    (if is_fun_type T orelse is_pair_type T then
blanchet@35190
   528
       Ts |> filter (curry (type_match thy o swap) T o unarize_and_unbox_type)
blanchet@33192
   529
          |> map (rpair ks o SOME)
blanchet@33192
   530
     else
blanchet@34121
   531
       [(SOME T, ks)]) @ repair_cards_assigns_wrt_boxing thy Ts cards_assigns
blanchet@34121
   532
  | repair_cards_assigns_wrt_boxing thy Ts ((NONE, ks) :: cards_assigns) =
blanchet@34121
   533
    (NONE, ks) :: repair_cards_assigns_wrt_boxing thy Ts cards_assigns
blanchet@33192
   534
blanchet@33571
   535
val max_scopes = 4096
blanchet@33192
   536
val distinct_threshold = 512
blanchet@33192
   537
blanchet@35190
   538
(* hol_context -> bool -> int -> (typ option * int list) list
blanchet@33192
   539
   -> (styp option * int list) list -> (styp option * int list) list -> int list
blanchet@33571
   540
   -> typ list -> typ list -> typ list -> int * scope list *)
blanchet@35190
   541
fun all_scopes (hol_ctxt as {thy, ...}) binarize sym_break cards_assigns
blanchet@35190
   542
               maxes_assigns iters_assigns bitss bisim_depths mono_Ts nonmono_Ts
blanchet@35190
   543
               deep_dataTs =
blanchet@33192
   544
  let
blanchet@34121
   545
    val cards_assigns = repair_cards_assigns_wrt_boxing thy mono_Ts
blanchet@34121
   546
                                                        cards_assigns
blanchet@35190
   547
    val blocks = blocks_for_types hol_ctxt binarize cards_assigns maxes_assigns
blanchet@34121
   548
                                  iters_assigns bitss bisim_depths mono_Ts
blanchet@34121
   549
                                  nonmono_Ts
blanchet@33192
   550
    val ranks = map rank_of_block blocks
blanchet@33571
   551
    val all = all_combinations_ordered_smartly (map (rpair 0) ranks)
haftmann@33956
   552
    val head = take max_scopes all
blanchet@35190
   553
    val descs =
blanchet@35190
   554
      map_filter (scope_descriptor_from_combination hol_ctxt binarize blocks)
blanchet@35190
   555
                 head
blanchet@33192
   556
  in
blanchet@33571
   557
    (length all - length head,
blanchet@33571
   558
     descs |> length descs <= distinct_threshold ? distinct (op =)
blanchet@35190
   559
           |> map (scope_from_descriptor hol_ctxt binarize sym_break
blanchet@35190
   560
                                         deep_dataTs))
blanchet@33192
   561
  end
blanchet@33192
   562
blanchet@33192
   563
end;