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