src/HOL/Tools/Nitpick/nitpick_scope.ML
author haftmann
Wed, 25 Nov 2009 09:13:46 +0100
changeset 33956 e9afca2118d4
parent 33954 fff6f11b1f09
child 33982 1ae222745c4a
permissions -rw-r--r--
normalized uncurry take/drop
blanchet@33192
     1
(*  Title:      HOL/Nitpick/Tools/nitpick_scope.ML
blanchet@33192
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@33192
     3
    Copyright   2008, 2009
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@33224
    11
  type extended_context = Nitpick_HOL.extended_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@33192
    25
    precise: bool,
blanchet@33549
    26
    shallow: bool,
blanchet@33192
    27
    constrs: constr_spec list}
blanchet@33192
    28
blanchet@33192
    29
  type scope = {
blanchet@33192
    30
    ext_ctxt: extended_context,
blanchet@33192
    31
    card_assigns: (typ * int) list,
blanchet@33192
    32
    bisim_depth: int,
blanchet@33192
    33
    datatypes: dtype_spec list,
blanchet@33192
    34
    ofs: int Typtab.table}
blanchet@33192
    35
blanchet@33192
    36
  val datatype_spec : dtype_spec list -> typ -> dtype_spec option
blanchet@33192
    37
  val constr_spec : dtype_spec list -> styp -> constr_spec
blanchet@33192
    38
  val is_precise_type : dtype_spec list -> typ -> bool
blanchet@33192
    39
  val is_fully_comparable_type : dtype_spec list -> typ -> bool
blanchet@33192
    40
  val offset_of_type : int Typtab.table -> typ -> int
blanchet@33192
    41
  val spec_of_type : scope -> typ -> int * int
blanchet@33192
    42
  val pretties_for_scope : scope -> bool -> Pretty.T list
blanchet@33192
    43
  val multiline_string_for_scope : scope -> string
blanchet@33192
    44
  val scopes_equivalent : scope -> scope -> bool
blanchet@33192
    45
  val scope_less_eq : scope -> scope -> bool
blanchet@33192
    46
  val all_scopes :
blanchet@33192
    47
    extended_context -> int -> (typ option * int list) list
blanchet@33192
    48
    -> (styp option * int list) list -> (styp option * int list) list
blanchet@33571
    49
    -> int list -> typ list -> typ list -> typ list -> int * scope list
blanchet@33192
    50
end;
blanchet@33192
    51
blanchet@33224
    52
structure Nitpick_Scope : NITPICK_SCOPE =
blanchet@33192
    53
struct
blanchet@33192
    54
blanchet@33224
    55
open Nitpick_Util
blanchet@33224
    56
open Nitpick_HOL
blanchet@33192
    57
blanchet@33192
    58
type constr_spec = {
blanchet@33192
    59
  const: styp,
blanchet@33192
    60
  delta: int,
blanchet@33192
    61
  epsilon: int,
blanchet@33192
    62
  exclusive: bool,
blanchet@33192
    63
  explicit_max: int,
blanchet@33192
    64
  total: bool}
blanchet@33192
    65
blanchet@33192
    66
type dtype_spec = {
blanchet@33192
    67
  typ: typ,
blanchet@33192
    68
  card: int,
blanchet@33192
    69
  co: bool,
blanchet@33192
    70
  precise: bool,
blanchet@33549
    71
  shallow: bool,
blanchet@33192
    72
  constrs: constr_spec list}
blanchet@33192
    73
blanchet@33192
    74
type scope = {
blanchet@33192
    75
  ext_ctxt: extended_context,
blanchet@33192
    76
  card_assigns: (typ * int) list,
blanchet@33192
    77
  bisim_depth: int,
blanchet@33192
    78
  datatypes: dtype_spec list,
blanchet@33192
    79
  ofs: int Typtab.table}
blanchet@33192
    80
blanchet@33192
    81
datatype row_kind = Card of typ | Max of styp
blanchet@33192
    82
blanchet@33192
    83
type row = row_kind * int list
blanchet@33192
    84
type block = row list
blanchet@33192
    85
blanchet@33192
    86
(* dtype_spec list -> typ -> dtype_spec option *)
blanchet@33192
    87
fun datatype_spec (dtypes : dtype_spec list) T =
blanchet@33192
    88
  List.find (equal T o #typ) dtypes
blanchet@33192
    89
blanchet@33192
    90
(* dtype_spec list -> styp -> constr_spec *)
blanchet@33224
    91
fun constr_spec [] x = raise TERM ("Nitpick_Scope.constr_spec", [Const x])
blanchet@33192
    92
  | constr_spec ({constrs, ...} :: dtypes : dtype_spec list) (x as (s, T)) =
blanchet@33192
    93
    case List.find (equal (s, body_type T) o (apsnd body_type o #const))
blanchet@33192
    94
                   constrs of
blanchet@33192
    95
      SOME c => c
blanchet@33192
    96
    | NONE => constr_spec dtypes x
blanchet@33192
    97
blanchet@33192
    98
(* dtype_spec list -> typ -> bool *)
blanchet@33192
    99
fun is_precise_type dtypes (Type ("fun", Ts)) =
blanchet@33192
   100
    forall (is_precise_type dtypes) Ts
blanchet@33192
   101
  | is_precise_type dtypes (Type ("*", Ts)) = forall (is_precise_type dtypes) Ts
blanchet@33192
   102
  | is_precise_type dtypes T =
blanchet@33853
   103
    not (is_integer_type T) andalso #precise (the (datatype_spec dtypes T))
blanchet@33192
   104
    handle Option.Option => true
blanchet@33192
   105
fun is_fully_comparable_type dtypes (Type ("fun", [T1, T2])) =
blanchet@33192
   106
    is_precise_type dtypes T1 andalso is_fully_comparable_type dtypes T2
blanchet@33192
   107
  | is_fully_comparable_type dtypes (Type ("*", Ts)) =
blanchet@33192
   108
    forall (is_fully_comparable_type dtypes) Ts
blanchet@33192
   109
  | is_fully_comparable_type _ _ = true
blanchet@33192
   110
blanchet@33192
   111
(* int Typtab.table -> typ -> int *)
blanchet@33192
   112
fun offset_of_type ofs T =
blanchet@33192
   113
  case Typtab.lookup ofs T of
blanchet@33192
   114
    SOME j0 => j0
blanchet@33192
   115
  | NONE => Typtab.lookup ofs dummyT |> the_default 0
blanchet@33192
   116
blanchet@33192
   117
(* scope -> typ -> int * int *)
blanchet@33192
   118
fun spec_of_type ({card_assigns, ofs, ...} : scope) T =
blanchet@33192
   119
  (card_of_type card_assigns T
blanchet@33224
   120
   handle TYPE ("Nitpick_HOL.card_of_type", _, _) => ~1, offset_of_type ofs T)
blanchet@33192
   121
blanchet@33192
   122
(* (string -> string) -> scope
blanchet@33192
   123
   -> string list * string list * string list * string list * string list *)
blanchet@33192
   124
fun quintuple_for_scope quote ({ext_ctxt as {thy, ctxt, ...}, card_assigns,
blanchet@33192
   125
                                bisim_depth, datatypes, ...} : scope) =
blanchet@33192
   126
  let
blanchet@33192
   127
    val (iter_asgns, card_asgns) =
blanchet@33192
   128
      card_assigns |> filter_out (equal @{typ bisim_iterator} o fst)
blanchet@33192
   129
                   |> List.partition (is_fp_iterator_type o fst)
blanchet@33192
   130
    val (unimportant_card_asgns, important_card_asgns) =
blanchet@33549
   131
      card_asgns |> List.partition ((is_integer_type orf is_datatype thy) o fst)
blanchet@33192
   132
    val cards =
blanchet@33192
   133
      map (fn (T, k) => quote (string_for_type ctxt T) ^ " = " ^
blanchet@33192
   134
                        string_of_int k)
blanchet@33192
   135
    fun maxes () =
blanchet@33192
   136
      maps (map_filter
blanchet@33192
   137
                (fn {const, explicit_max, ...} =>
blanchet@33192
   138
                    if explicit_max < 0 then
blanchet@33192
   139
                      NONE
blanchet@33192
   140
                    else
blanchet@33192
   141
                      SOME (Syntax.string_of_term ctxt (Const const) ^ " = " ^
blanchet@33192
   142
                            string_of_int explicit_max))
blanchet@33192
   143
                 o #constrs) datatypes
blanchet@33192
   144
    fun iters () =
blanchet@33192
   145
      map (fn (T, k) =>
blanchet@33192
   146
              quote (Syntax.string_of_term ctxt
blanchet@33192
   147
                         (Const (const_for_iterator_type T))) ^ " = " ^
blanchet@33192
   148
              string_of_int (k - 1)) iter_asgns
blanchet@33192
   149
    fun bisims () =
blanchet@33192
   150
      if bisim_depth < 0 andalso forall (not o #co) datatypes then []
blanchet@33192
   151
      else ["bisim_depth = " ^ string_of_int bisim_depth]
blanchet@33192
   152
  in
blanchet@33192
   153
    setmp_show_all_types
blanchet@33192
   154
        (fn () => (cards important_card_asgns, cards unimportant_card_asgns,
blanchet@33192
   155
                   maxes (), iters (), bisims ())) ()
blanchet@33192
   156
  end
blanchet@33192
   157
blanchet@33192
   158
(* scope -> bool -> Pretty.T list *)
blanchet@33192
   159
fun pretties_for_scope scope verbose =
blanchet@33192
   160
  let
blanchet@33192
   161
    val (important_cards, unimportant_cards, maxes, iters, bisim_depths) =
blanchet@33192
   162
      quintuple_for_scope maybe_quote scope
blanchet@33192
   163
    val ss = map (prefix "card ") important_cards @
blanchet@33192
   164
             (if verbose then
blanchet@33192
   165
                map (prefix "card ") unimportant_cards @
blanchet@33192
   166
                map (prefix "max ") maxes @
blanchet@33192
   167
                map (prefix "iter ") iters @
blanchet@33192
   168
                bisim_depths
blanchet@33192
   169
              else
blanchet@33192
   170
                [])
blanchet@33192
   171
  in
blanchet@33192
   172
    if null ss then []
blanchet@33192
   173
    else serial_commas "and" ss |> map Pretty.str |> Pretty.breaks
blanchet@33192
   174
  end
blanchet@33192
   175
blanchet@33192
   176
(* scope -> string *)
blanchet@33192
   177
fun multiline_string_for_scope scope =
blanchet@33192
   178
  let
blanchet@33192
   179
    val (important_cards, unimportant_cards, maxes, iters, bisim_depths) =
blanchet@33192
   180
      quintuple_for_scope I scope
blanchet@33192
   181
    val cards = important_cards @ unimportant_cards
blanchet@33192
   182
  in
blanchet@33192
   183
    case (if null cards then [] else ["card: " ^ commas cards]) @
blanchet@33192
   184
         (if null maxes then [] else ["max: " ^ commas maxes]) @
blanchet@33192
   185
         (if null iters then [] else ["iter: " ^ commas iters]) @
blanchet@33192
   186
         bisim_depths of
blanchet@33192
   187
      [] => "empty"
blanchet@33192
   188
    | lines => space_implode "\n" lines
blanchet@33192
   189
  end
blanchet@33192
   190
blanchet@33192
   191
(* scope -> scope -> bool *)
blanchet@33192
   192
fun scopes_equivalent (s1 : scope) (s2 : scope) =
blanchet@33192
   193
  #datatypes s1 = #datatypes s2 andalso #card_assigns s1 = #card_assigns s2
blanchet@33192
   194
fun scope_less_eq (s1 : scope) (s2 : scope) =
blanchet@33192
   195
  (s1, s2) |> pairself (map snd o #card_assigns) |> op ~~ |> forall (op <=)
blanchet@33192
   196
blanchet@33192
   197
(* row -> int *)
blanchet@33192
   198
fun rank_of_row (_, ks) = length ks
blanchet@33192
   199
(* block -> int *)
blanchet@33192
   200
fun rank_of_block block = fold Integer.max (map rank_of_row block) 1
blanchet@33192
   201
(* int -> typ * int list -> typ * int list *)
blanchet@33192
   202
fun project_row column (y, ks) = (y, [nth ks (Int.min (column, length ks - 1))])
blanchet@33192
   203
(* int -> block -> block *)
blanchet@33192
   204
fun project_block (column, block) = map (project_row column) block
blanchet@33192
   205
blanchet@33192
   206
(* (''a * ''a -> bool) -> (''a option * int list) list -> ''a -> int list *)
blanchet@33192
   207
fun lookup_ints_assign eq asgns key =
blanchet@33192
   208
  case triple_lookup eq asgns key of
blanchet@33192
   209
    SOME ks => ks
blanchet@33224
   210
  | NONE => raise ARG ("Nitpick_Scope.lookup_ints_assign", "")
blanchet@33192
   211
(* theory -> (typ option * int list) list -> typ -> int list *)
blanchet@33192
   212
fun lookup_type_ints_assign thy asgns T =
blanchet@33192
   213
  map (curry Int.max 1) (lookup_ints_assign (type_match thy) asgns T)
blanchet@33224
   214
  handle ARG ("Nitpick_Scope.lookup_ints_assign", _) =>
blanchet@33224
   215
         raise TYPE ("Nitpick_Scope.lookup_type_ints_assign", [T], [])
blanchet@33192
   216
(* theory -> (styp option * int list) list -> styp -> int list *)
blanchet@33192
   217
fun lookup_const_ints_assign thy asgns x =
blanchet@33192
   218
  lookup_ints_assign (const_match thy) asgns x
blanchet@33224
   219
  handle ARG ("Nitpick_Scope.lookup_ints_assign", _) =>
blanchet@33224
   220
         raise TERM ("Nitpick_Scope.lookup_const_ints_assign", [Const x])
blanchet@33192
   221
blanchet@33192
   222
(* theory -> (styp option * int list) list -> styp -> row option *)
blanchet@33192
   223
fun row_for_constr thy maxes_asgns constr =
blanchet@33192
   224
  SOME (Max constr, lookup_const_ints_assign thy maxes_asgns constr)
blanchet@33192
   225
  handle TERM ("lookup_const_ints_assign", _) => NONE
blanchet@33192
   226
blanchet@33571
   227
(* extended_context -> (typ option * int list) list
blanchet@33192
   228
   -> (styp option * int list) list -> (styp option * int list) list -> int list
blanchet@33192
   229
   -> typ -> block *)
blanchet@33571
   230
fun block_for_type (ext_ctxt as {thy, ...}) cards_asgns maxes_asgns iters_asgns
blanchet@33571
   231
                   bisim_depths T =
blanchet@33571
   232
  if T = @{typ bisim_iterator} then
blanchet@33571
   233
    [(Card T, map (fn k => Int.max (0, k) + 1) bisim_depths)]
blanchet@33571
   234
  else if is_fp_iterator_type T then
blanchet@33571
   235
    [(Card T, map (fn k => Int.max (0, k) + 1)
blanchet@33571
   236
                  (lookup_const_ints_assign thy iters_asgns
blanchet@33571
   237
                                            (const_for_iterator_type T)))]
blanchet@33571
   238
  else
blanchet@33571
   239
    (Card T, lookup_type_ints_assign thy cards_asgns T) ::
blanchet@33571
   240
    (case datatype_constrs ext_ctxt T of
blanchet@33571
   241
       [_] => []
blanchet@33571
   242
     | constrs => map_filter (row_for_constr thy maxes_asgns) constrs)
blanchet@33192
   243
blanchet@33571
   244
(* extended_context -> (typ option * int list) list
blanchet@33192
   245
   -> (styp option * int list) list -> (styp option * int list) list -> int list
blanchet@33192
   246
   -> typ list -> typ list -> block list *)
blanchet@33571
   247
fun blocks_for_types ext_ctxt cards_asgns maxes_asgns iters_asgns bisim_depths
blanchet@33192
   248
                     mono_Ts nonmono_Ts =
blanchet@33192
   249
  let
blanchet@33192
   250
    (* typ -> block *)
blanchet@33571
   251
    val block_for = block_for_type ext_ctxt cards_asgns maxes_asgns iters_asgns
blanchet@33192
   252
                                   bisim_depths
blanchet@33192
   253
    val mono_block = maps block_for mono_Ts
blanchet@33192
   254
    val nonmono_blocks = map block_for nonmono_Ts
blanchet@33192
   255
  in mono_block :: nonmono_blocks end
blanchet@33192
   256
blanchet@33192
   257
val sync_threshold = 5
blanchet@33192
   258
blanchet@33192
   259
(* int list -> int list list *)
blanchet@33192
   260
fun all_combinations_ordered_smartly ks =
blanchet@33192
   261
  let
blanchet@33192
   262
    (* int list -> int *)
blanchet@33192
   263
    fun cost_with_monos [] = 0
blanchet@33192
   264
      | cost_with_monos (k :: ks) =
blanchet@33192
   265
        if k < sync_threshold andalso forall (equal k) ks then
blanchet@33192
   266
          k - sync_threshold
blanchet@33192
   267
        else
blanchet@33192
   268
          k * (k + 1) div 2 + Integer.sum ks
blanchet@33192
   269
    fun cost_without_monos [] = 0
blanchet@33192
   270
      | cost_without_monos [k] = k
blanchet@33192
   271
      | cost_without_monos (_ :: k :: ks) =
blanchet@33192
   272
        if k < sync_threshold andalso forall (equal k) ks then
blanchet@33192
   273
          k - sync_threshold
blanchet@33192
   274
        else
blanchet@33192
   275
          Integer.sum (k :: ks)
blanchet@33192
   276
  in
blanchet@33192
   277
    ks |> all_combinations
blanchet@33192
   278
       |> map (`(if fst (hd ks) > 1 then cost_with_monos
blanchet@33192
   279
                 else cost_without_monos))
blanchet@33192
   280
       |> sort (int_ord o pairself fst) |> map snd
blanchet@33192
   281
  end
blanchet@33192
   282
blanchet@33192
   283
(* typ -> bool *)
blanchet@33192
   284
fun is_self_recursive_constr_type T =
blanchet@33192
   285
  exists (exists_subtype (equal (body_type T))) (binder_types T)
blanchet@33192
   286
blanchet@33192
   287
(* (styp * int) list -> styp -> int *)
blanchet@33192
   288
fun constr_max maxes x = the_default ~1 (AList.lookup (op =) maxes x)
blanchet@33192
   289
blanchet@33192
   290
type scope_desc = (typ * int) list * (styp * int) list
blanchet@33192
   291
blanchet@33571
   292
(* extended_context -> scope_desc -> typ * int -> bool *)
blanchet@33571
   293
fun is_surely_inconsistent_card_assign ext_ctxt (card_asgns, max_asgns) (T, k) =
blanchet@33571
   294
  case datatype_constrs ext_ctxt T of
blanchet@33192
   295
    [] => false
blanchet@33192
   296
  | xs =>
blanchet@33192
   297
    let
blanchet@33192
   298
      val precise_cards =
blanchet@33192
   299
        map (Integer.prod
blanchet@33571
   300
             o map (bounded_precise_card_of_type ext_ctxt k 0 card_asgns)
blanchet@33192
   301
             o binder_types o snd) xs
blanchet@33192
   302
      val maxes = map (constr_max max_asgns) xs
blanchet@33192
   303
      (* int -> int -> int *)
blanchet@33192
   304
      fun effective_max 0 ~1 = k
blanchet@33192
   305
        | effective_max 0 max = max
blanchet@33192
   306
        | effective_max card ~1 = card
blanchet@33192
   307
        | effective_max card max = Int.min (card, max)
blanchet@33192
   308
      val max = map2 effective_max precise_cards maxes |> Integer.sum
blanchet@33192
   309
      (* unit -> int *)
blanchet@33192
   310
      fun doms_card () =
blanchet@33192
   311
        xs |> map (Integer.prod o map (bounded_card_of_type k ~1 card_asgns)
blanchet@33192
   312
                   o binder_types o snd)
blanchet@33192
   313
           |> Integer.sum
blanchet@33192
   314
    in
blanchet@33192
   315
      max < k
blanchet@33192
   316
      orelse (forall (not_equal 0) precise_cards andalso doms_card () < k)
blanchet@33192
   317
    end
blanchet@33224
   318
    handle TYPE ("Nitpick_HOL.card_of_type", _, _) => false
blanchet@33192
   319
blanchet@33571
   320
(* extended_context -> scope_desc -> bool *)
blanchet@33571
   321
fun is_surely_inconsistent_scope_description ext_ctxt
blanchet@33571
   322
                                             (desc as (card_asgns, _)) =
blanchet@33571
   323
  exists (is_surely_inconsistent_card_assign ext_ctxt desc) card_asgns
blanchet@33192
   324
blanchet@33571
   325
(* extended_context -> scope_desc -> (typ * int) list option *)
blanchet@33571
   326
fun repair_card_assigns ext_ctxt (card_asgns, max_asgns) =
blanchet@33192
   327
  let
blanchet@33192
   328
    (* (typ * int) list -> (typ * int) list -> (typ * int) list option *)
blanchet@33192
   329
    fun aux seen [] = SOME seen
blanchet@33192
   330
      | aux seen ((T, 0) :: _) = NONE
blanchet@33192
   331
      | aux seen ((T, k) :: asgns) =
blanchet@33571
   332
        (if is_surely_inconsistent_scope_description ext_ctxt
blanchet@33192
   333
                ((T, k) :: seen, max_asgns) then
blanchet@33192
   334
           raise SAME ()
blanchet@33192
   335
         else
blanchet@33192
   336
           case aux ((T, k) :: seen) asgns of
blanchet@33192
   337
             SOME asgns => SOME asgns
blanchet@33192
   338
           | NONE => raise SAME ())
blanchet@33192
   339
        handle SAME () => aux seen ((T, k - 1) :: asgns)
blanchet@33192
   340
  in aux [] (rev card_asgns) end
blanchet@33192
   341
blanchet@33192
   342
(* theory -> (typ * int) list -> typ * int -> typ * int *)
blanchet@33192
   343
fun repair_iterator_assign thy asgns (T as Type (s, Ts), k) =
blanchet@33192
   344
    (T, if T = @{typ bisim_iterator} then
blanchet@33192
   345
          let val co_cards = map snd (filter (is_codatatype thy o fst) asgns) in
blanchet@33192
   346
            Int.min (k, Integer.sum co_cards)
blanchet@33192
   347
          end
blanchet@33192
   348
        else if is_fp_iterator_type T then
blanchet@33192
   349
          case Ts of
blanchet@33192
   350
            [] => 1
blanchet@33192
   351
          | _ => bounded_card_of_type k ~1 asgns (foldr1 HOLogic.mk_prodT Ts)
blanchet@33192
   352
        else
blanchet@33192
   353
          k)
blanchet@33192
   354
  | repair_iterator_assign _ _ asgn = asgn
blanchet@33192
   355
blanchet@33192
   356
(* row -> scope_desc -> scope_desc *)
blanchet@33192
   357
fun add_row_to_scope_descriptor (kind, ks) (card_asgns, max_asgns) =
blanchet@33192
   358
  case kind of
blanchet@33192
   359
    Card T => ((T, the_single ks) :: card_asgns, max_asgns)
blanchet@33192
   360
  | Max x => (card_asgns, (x, the_single ks) :: max_asgns)
blanchet@33192
   361
(* block -> scope_desc *)
blanchet@33192
   362
fun scope_descriptor_from_block block =
blanchet@33192
   363
  fold_rev add_row_to_scope_descriptor block ([], [])
blanchet@33571
   364
(* extended_context -> block list -> int list -> scope_desc option *)
blanchet@33571
   365
fun scope_descriptor_from_combination (ext_ctxt as {thy, ...}) blocks columns =
blanchet@33192
   366
  let
blanchet@33192
   367
    val (card_asgns, max_asgns) =
blanchet@33192
   368
      maps project_block (columns ~~ blocks) |> scope_descriptor_from_block
blanchet@33571
   369
    val card_asgns = repair_card_assigns ext_ctxt (card_asgns, max_asgns) |> the
blanchet@33192
   370
  in
blanchet@33192
   371
    SOME (map (repair_iterator_assign thy card_asgns) card_asgns, max_asgns)
blanchet@33192
   372
  end
blanchet@33192
   373
  handle Option.Option => NONE
blanchet@33192
   374
blanchet@33192
   375
(* theory -> (typ * int) list -> dtype_spec list -> int Typtab.table *)
blanchet@33192
   376
fun offset_table_for_card_assigns thy asgns dtypes =
blanchet@33192
   377
  let
blanchet@33192
   378
    (* int -> (int * int) list -> (typ * int) list -> int Typtab.table
blanchet@33192
   379
       -> int Typtab.table *)
blanchet@33192
   380
    fun aux next _ [] = Typtab.update_new (dummyT, next)
blanchet@33192
   381
      | aux next reusable ((T, k) :: asgns) =
blanchet@33192
   382
        if k = 1 orelse is_integer_type T then
blanchet@33192
   383
          aux next reusable asgns
blanchet@33192
   384
        else if length (these (Option.map #constrs (datatype_spec dtypes T)))
blanchet@33192
   385
                > 1 then
blanchet@33192
   386
          Typtab.update_new (T, next) #> aux (next + k) reusable asgns
blanchet@33192
   387
        else
blanchet@33192
   388
          case AList.lookup (op =) reusable k of
blanchet@33192
   389
            SOME j0 => Typtab.update_new (T, j0) #> aux next reusable asgns
blanchet@33192
   390
          | NONE => Typtab.update_new (T, next)
blanchet@33192
   391
                    #> aux (next + k) ((k, next) :: reusable) asgns
blanchet@33192
   392
  in aux 0 [] asgns Typtab.empty end
blanchet@33192
   393
blanchet@33192
   394
(* int -> (typ * int) list -> typ -> int *)
blanchet@33192
   395
fun domain_card max card_asgns =
blanchet@33192
   396
  Integer.prod o map (bounded_card_of_type max max card_asgns) o binder_types
blanchet@33192
   397
blanchet@33192
   398
(* scope_desc -> bool -> int -> (int -> int) -> int -> int -> bool * styp
blanchet@33192
   399
   -> constr_spec list -> constr_spec list *)
blanchet@33192
   400
fun add_constr_spec (card_asgns, max_asgns) co card sum_dom_cards num_self_recs
blanchet@33192
   401
                    num_non_self_recs (self_rec, x as (s, T)) constrs =
blanchet@33192
   402
  let
blanchet@33192
   403
    val max = constr_max max_asgns x
blanchet@33192
   404
    (* int -> int *)
blanchet@33192
   405
    fun bound k = Int.min (card, (max >= 0 ? curry Int.min max) k)
blanchet@33192
   406
    (* unit -> int *)
blanchet@33192
   407
    fun next_delta () = if null constrs then 0 else #epsilon (hd constrs)
blanchet@33192
   408
    val {delta, epsilon, exclusive, total} =
blanchet@33192
   409
      if max = 0 then
blanchet@33192
   410
        let val delta = next_delta () in
blanchet@33192
   411
          {delta = delta, epsilon = delta, exclusive = true, total = false}
blanchet@33192
   412
        end
blanchet@33192
   413
      else if not co andalso num_self_recs > 0 then
blanchet@33192
   414
        if not self_rec andalso num_non_self_recs = 1
blanchet@33192
   415
           andalso domain_card 2 card_asgns T = 1 then
blanchet@33192
   416
          {delta = 0, epsilon = 1, exclusive = (s = @{const_name Nil}),
blanchet@33192
   417
           total = true}
blanchet@33192
   418
        else if s = @{const_name Cons} then
blanchet@33192
   419
          {delta = 1, epsilon = card, exclusive = true, total = false}
blanchet@33192
   420
        else
blanchet@33192
   421
          {delta = 0, epsilon = card, exclusive = false, total = false}
blanchet@33192
   422
      else if card = sum_dom_cards (card + 1) then
blanchet@33192
   423
        let val delta = next_delta () in
blanchet@33192
   424
          {delta = delta, epsilon = delta + domain_card card card_asgns T,
blanchet@33192
   425
           exclusive = true, total = true}
blanchet@33192
   426
        end
blanchet@33192
   427
      else
blanchet@33192
   428
        {delta = 0, epsilon = card,
blanchet@33192
   429
         exclusive = (num_self_recs + num_non_self_recs = 1), total = false}
blanchet@33192
   430
  in
blanchet@33192
   431
    {const = x, delta = delta, epsilon = epsilon, exclusive = exclusive,
blanchet@33192
   432
     explicit_max = max, total = total} :: constrs
blanchet@33192
   433
  end
blanchet@33192
   434
blanchet@33549
   435
(* extended_context -> typ list -> scope_desc -> typ * int -> dtype_spec *)
blanchet@33549
   436
fun datatype_spec_from_scope_descriptor (ext_ctxt as {thy, ...}) shallow_dataTs
blanchet@33192
   437
                                        (desc as (card_asgns, _)) (T, card) =
blanchet@33192
   438
  let
blanchet@33549
   439
    val shallow = T mem shallow_dataTs
blanchet@33192
   440
    val co = is_codatatype thy T
blanchet@33192
   441
    val xs = boxed_datatype_constrs ext_ctxt T
blanchet@33192
   442
    val self_recs = map (is_self_recursive_constr_type o snd) xs
blanchet@33192
   443
    val (num_self_recs, num_non_self_recs) =
blanchet@33192
   444
      List.partition (equal true) self_recs |> pairself length
blanchet@33571
   445
    val precise = (card = bounded_precise_card_of_type ext_ctxt (card + 1) 0
blanchet@33192
   446
                                                       card_asgns T)
blanchet@33192
   447
    (* int -> int *)
blanchet@33192
   448
    fun sum_dom_cards max =
blanchet@33192
   449
      map (domain_card max card_asgns o snd) xs |> Integer.sum
blanchet@33192
   450
    val constrs =
blanchet@33192
   451
      fold_rev (add_constr_spec desc co card sum_dom_cards num_self_recs
blanchet@33192
   452
                                num_non_self_recs) (self_recs ~~ xs) []
blanchet@33549
   453
  in
blanchet@33549
   454
    {typ = T, card = card, co = co, precise = precise, shallow = shallow,
blanchet@33549
   455
     constrs = constrs}
blanchet@33549
   456
  end
blanchet@33192
   457
blanchet@33549
   458
(* extended_context -> int -> typ list -> scope_desc -> scope *)
blanchet@33549
   459
fun scope_from_descriptor (ext_ctxt as {thy, ...}) sym_break shallow_dataTs
blanchet@33192
   460
                          (desc as (card_asgns, _)) =
blanchet@33192
   461
  let
blanchet@33549
   462
    val datatypes =
blanchet@33549
   463
      map (datatype_spec_from_scope_descriptor ext_ctxt shallow_dataTs desc)
blanchet@33549
   464
          (filter (is_datatype thy o fst) card_asgns)
blanchet@33192
   465
    val bisim_depth = card_of_type card_asgns @{typ bisim_iterator} - 1
blanchet@33192
   466
  in
blanchet@33192
   467
    {ext_ctxt = ext_ctxt, card_assigns = card_asgns, datatypes = datatypes,
blanchet@33192
   468
     bisim_depth = bisim_depth,
blanchet@33192
   469
     ofs = if sym_break <= 0 then Typtab.empty
blanchet@33192
   470
           else offset_table_for_card_assigns thy card_asgns datatypes}
blanchet@33192
   471
  end
blanchet@33192
   472
blanchet@33192
   473
(* theory -> typ list -> (typ option * int list) list
blanchet@33192
   474
   -> (typ option * int list) list *)
blanchet@33192
   475
fun fix_cards_assigns_wrt_boxing _ _ [] = []
blanchet@33192
   476
  | fix_cards_assigns_wrt_boxing thy Ts ((SOME T, ks) :: cards_asgns) =
blanchet@33192
   477
    (if is_fun_type T orelse is_pair_type T then
blanchet@33192
   478
       Ts |> filter (curry (type_match thy o swap) T o unbox_type)
blanchet@33192
   479
          |> map (rpair ks o SOME)
blanchet@33192
   480
     else
blanchet@33192
   481
       [(SOME T, ks)]) @ fix_cards_assigns_wrt_boxing thy Ts cards_asgns
blanchet@33192
   482
  | fix_cards_assigns_wrt_boxing thy Ts ((NONE, ks) :: cards_asgns) =
blanchet@33192
   483
    (NONE, ks) :: fix_cards_assigns_wrt_boxing thy Ts cards_asgns
blanchet@33192
   484
blanchet@33571
   485
val max_scopes = 4096
blanchet@33192
   486
val distinct_threshold = 512
blanchet@33192
   487
blanchet@33192
   488
(* extended_context -> int -> (typ option * int list) list
blanchet@33192
   489
   -> (styp option * int list) list -> (styp option * int list) list -> int list
blanchet@33571
   490
   -> typ list -> typ list -> typ list -> int * scope list *)
blanchet@33571
   491
fun all_scopes (ext_ctxt as {thy, ...}) sym_break cards_asgns maxes_asgns
blanchet@33549
   492
               iters_asgns bisim_depths mono_Ts nonmono_Ts shallow_dataTs =
blanchet@33192
   493
  let
blanchet@33192
   494
    val cards_asgns = fix_cards_assigns_wrt_boxing thy mono_Ts cards_asgns
blanchet@33571
   495
    val blocks = blocks_for_types ext_ctxt cards_asgns maxes_asgns iters_asgns
blanchet@33192
   496
                                  bisim_depths mono_Ts nonmono_Ts
blanchet@33192
   497
    val ranks = map rank_of_block blocks
blanchet@33571
   498
    val all = all_combinations_ordered_smartly (map (rpair 0) ranks)
haftmann@33956
   499
    val head = take max_scopes all
blanchet@33571
   500
    val descs = map_filter (scope_descriptor_from_combination ext_ctxt blocks)
blanchet@33571
   501
                           head
blanchet@33192
   502
  in
blanchet@33571
   503
    (length all - length head,
blanchet@33571
   504
     descs |> length descs <= distinct_threshold ? distinct (op =)
blanchet@33571
   505
           |> map (scope_from_descriptor ext_ctxt sym_break shallow_dataTs))
blanchet@33192
   506
  end
blanchet@33192
   507
blanchet@33192
   508
end;