src/HOL/Tools/Nitpick/nitpick_scope.ML
author blanchet
Sat, 24 Apr 2010 16:43:03 +0200
changeset 36386 2132f15b366f
parent 36385 ff5f88702590
child 36390 eee4ee6a5cbe
permissions -rw-r--r--
Fruhjahrsputz: remove three mostly useless Nitpick options
     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 -> (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 fun datatype_spec (dtypes : dtype_spec list) T =
    97   List.find (curry (op =) T o #typ) dtypes
    98 
    99 fun constr_spec [] x = raise TERM ("Nitpick_Scope.constr_spec", [Const x])
   100   | constr_spec ({constrs, ...} :: dtypes : dtype_spec list) (x as (s, T)) =
   101     case List.find (curry (op =) (s, body_type T) o (apsnd body_type o #const))
   102                    constrs of
   103       SOME c => c
   104     | NONE => constr_spec dtypes x
   105 
   106 fun is_complete_type dtypes facto (Type (@{type_name fun}, [T1, T2])) =
   107     is_concrete_type dtypes facto T1 andalso is_complete_type dtypes facto T2
   108   | is_complete_type dtypes facto (Type (@{type_name fin_fun}, [T1, T2])) =
   109     is_exact_type dtypes facto T1 andalso is_complete_type dtypes facto T2
   110   | is_complete_type dtypes facto (Type (@{type_name "*"}, Ts)) =
   111     forall (is_complete_type dtypes facto) Ts
   112   | is_complete_type dtypes facto T =
   113     not (is_integer_like_type T) andalso not (is_bit_type T) andalso
   114     fun_from_pair (#complete (the (datatype_spec dtypes T))) facto
   115     handle Option.Option => true
   116 and is_concrete_type dtypes facto (Type (@{type_name fun}, [T1, T2])) =
   117     is_complete_type dtypes facto T1 andalso is_concrete_type dtypes facto T2
   118   | is_concrete_type dtypes facto (Type (@{type_name fin_fun}, [_, T2])) =
   119     is_concrete_type dtypes facto T2
   120   | is_concrete_type dtypes facto (Type (@{type_name "*"}, Ts)) =
   121     forall (is_concrete_type dtypes facto) Ts
   122   | is_concrete_type dtypes facto T =
   123     fun_from_pair (#concrete (the (datatype_spec dtypes T))) facto
   124     handle Option.Option => true
   125 and is_exact_type dtypes facto =
   126   is_complete_type dtypes facto andf is_concrete_type dtypes facto
   127 
   128 fun offset_of_type ofs T =
   129   case Typtab.lookup ofs T of
   130     SOME j0 => j0
   131   | NONE => Typtab.lookup ofs dummyT |> the_default 0
   132 
   133 fun spec_of_type ({card_assigns, ofs, ...} : scope) T =
   134   (card_of_type card_assigns T
   135    handle TYPE ("Nitpick_HOL.card_of_type", _, _) => ~1, offset_of_type ofs T)
   136 
   137 fun quintuple_for_scope quote
   138         ({hol_ctxt = {thy, ctxt, stds, ...}, card_assigns, bits, bisim_depth,
   139          datatypes, ...} : scope) =
   140   let
   141     val boring_Ts = [@{typ unsigned_bit}, @{typ signed_bit},
   142                      @{typ bisim_iterator}]
   143     val (iter_assigns, card_assigns) =
   144       card_assigns |> filter_out (member (op =) boring_Ts o fst)
   145                    |> List.partition (is_fp_iterator_type o fst)
   146     val (secondary_card_assigns, primary_card_assigns) =
   147       card_assigns |> List.partition ((is_integer_type orf is_datatype thy stds)
   148                                       o fst)
   149     val cards =
   150       map (fn (T, k) => quote (string_for_type ctxt T) ^ " = " ^
   151                         string_of_int k)
   152     fun maxes () =
   153       maps (map_filter
   154                 (fn {const, explicit_max, ...} =>
   155                     if explicit_max < 0 then
   156                       NONE
   157                     else
   158                       SOME (Syntax.string_of_term ctxt (Const const) ^ " = " ^
   159                             string_of_int explicit_max))
   160                  o #constrs) datatypes
   161     fun iters () =
   162       map (fn (T, k) =>
   163               quote (Syntax.string_of_term ctxt
   164                          (Const (const_for_iterator_type T))) ^ " = " ^
   165               string_of_int (k - 1)) iter_assigns
   166     fun miscs () =
   167       (if bits = 0 then [] else ["bits = " ^ string_of_int bits]) @
   168       (if bisim_depth < 0 andalso forall (not o #co) datatypes then []
   169        else ["bisim_depth = " ^ signed_string_of_int bisim_depth])
   170   in
   171     setmp_show_all_types
   172         (fn () => (cards primary_card_assigns, cards secondary_card_assigns,
   173                    maxes (), iters (), miscs ())) ()
   174   end
   175 
   176 fun pretties_for_scope scope verbose =
   177   let
   178     val (primary_cards, secondary_cards, maxes, iters, bisim_depths) =
   179       quintuple_for_scope maybe_quote scope
   180     val ss = map (prefix "card ") primary_cards @
   181              (if verbose then
   182                 map (prefix "card ") secondary_cards @
   183                 map (prefix "max ") maxes @
   184                 map (prefix "iter ") iters @
   185                 bisim_depths
   186               else
   187                 [])
   188   in
   189     if null ss then []
   190     else serial_commas "and" ss |> map Pretty.str |> Pretty.breaks
   191   end
   192 
   193 fun multiline_string_for_scope scope =
   194   let
   195     val (primary_cards, secondary_cards, maxes, iters, bisim_depths) =
   196       quintuple_for_scope I scope
   197     val cards = primary_cards @ secondary_cards
   198   in
   199     case (if null cards then [] else ["card: " ^ commas cards]) @
   200          (if null maxes then [] else ["max: " ^ commas maxes]) @
   201          (if null iters then [] else ["iter: " ^ commas iters]) @
   202          bisim_depths of
   203       [] => "empty"
   204     | lines => space_implode "\n" lines
   205   end
   206 
   207 fun scopes_equivalent (s1 : scope, s2 : scope) =
   208   #datatypes s1 = #datatypes s2 andalso #card_assigns s1 = #card_assigns s2
   209 fun scope_less_eq (s1 : scope) (s2 : scope) =
   210   (s1, s2) |> pairself (map snd o #card_assigns) |> op ~~ |> forall (op <=)
   211 
   212 fun rank_of_row (_, ks) = length ks
   213 fun rank_of_block block = fold Integer.max (map rank_of_row block) 1
   214 fun project_row column (y, ks) = (y, [nth ks (Int.min (column, length ks - 1))])
   215 fun project_block (column, block) = map (project_row column) block
   216 
   217 fun lookup_ints_assign eq assigns key =
   218   case triple_lookup eq assigns key of
   219     SOME ks => ks
   220   | NONE => raise ARG ("Nitpick_Scope.lookup_ints_assign", "")
   221 fun lookup_type_ints_assign thy assigns T =
   222   map (Integer.max 1) (lookup_ints_assign (type_match thy) assigns T)
   223   handle ARG ("Nitpick_Scope.lookup_ints_assign", _) =>
   224          raise TYPE ("Nitpick_Scope.lookup_type_ints_assign", [T], [])
   225 fun lookup_const_ints_assign thy assigns x =
   226   lookup_ints_assign (const_match thy) assigns x
   227   handle ARG ("Nitpick_Scope.lookup_ints_assign", _) =>
   228          raise TERM ("Nitpick_Scope.lookup_const_ints_assign", [Const x])
   229 
   230 fun row_for_constr thy maxes_assigns constr =
   231   SOME (Max constr, lookup_const_ints_assign thy maxes_assigns constr)
   232   handle TERM ("lookup_const_ints_assign", _) => NONE
   233 
   234 val max_bits = 31 (* Kodkod limit *)
   235 
   236 fun block_for_type (hol_ctxt as {thy, ...}) binarize cards_assigns maxes_assigns
   237                    iters_assigns bitss bisim_depths T =
   238   if T = @{typ unsigned_bit} then
   239     [(Card T, map (Integer.min max_bits o Integer.max 1) bitss)]
   240   else if T = @{typ signed_bit} then
   241     [(Card T, map (Integer.add 1 o Integer.min max_bits o Integer.max 1) bitss)]
   242   else if T = @{typ "unsigned_bit word"} then
   243     [(Card T, lookup_type_ints_assign thy cards_assigns nat_T)]
   244   else if T = @{typ "signed_bit word"} then
   245     [(Card T, lookup_type_ints_assign thy cards_assigns int_T)]
   246   else if T = @{typ bisim_iterator} then
   247     [(Card T, map (Integer.add 1 o Integer.max 0) bisim_depths)]
   248   else if is_fp_iterator_type T then
   249     [(Card T, map (Integer.add 1 o Integer.max 0)
   250                   (lookup_const_ints_assign thy iters_assigns
   251                                             (const_for_iterator_type T)))]
   252   else
   253     (Card T, lookup_type_ints_assign thy cards_assigns T) ::
   254     (case binarized_and_boxed_datatype_constrs hol_ctxt binarize T of
   255        [_] => []
   256      | constrs => map_filter (row_for_constr thy maxes_assigns) constrs)
   257 
   258 fun blocks_for_types hol_ctxt binarize cards_assigns maxes_assigns iters_assigns
   259                      bitss bisim_depths mono_Ts nonmono_Ts =
   260   let
   261     val block_for = block_for_type hol_ctxt binarize cards_assigns maxes_assigns
   262                                    iters_assigns bitss bisim_depths
   263     val mono_block = maps block_for mono_Ts
   264     val nonmono_blocks = map block_for nonmono_Ts
   265   in mono_block :: nonmono_blocks end
   266 
   267 val sync_threshold = 5
   268 
   269 fun all_combinations_ordered_smartly ks =
   270   let
   271     fun cost_with_monos [] = 0
   272       | cost_with_monos (k :: ks) =
   273         if k < sync_threshold andalso forall (curry (op =) k) ks then
   274           k - sync_threshold
   275         else
   276           k * (k + 1) div 2 + Integer.sum ks
   277     fun cost_without_monos [] = 0
   278       | cost_without_monos [k] = k
   279       | cost_without_monos (_ :: k :: ks) =
   280         if k < sync_threshold andalso forall (curry (op =) k) ks then
   281           k - sync_threshold
   282         else
   283           Integer.sum (k :: ks)
   284   in
   285     ks |> all_combinations
   286        |> map (`(if fst (hd ks) > 1 then cost_with_monos
   287                  else cost_without_monos))
   288        |> sort (int_ord o pairself fst) |> map snd
   289   end
   290 
   291 fun is_self_recursive_constr_type T =
   292   exists (exists_subtype (curry (op =) (body_type T))) (binder_types T)
   293 
   294 fun constr_max maxes x = the_default ~1 (AList.lookup (op =) maxes x)
   295 
   296 type scope_desc = (typ * int) list * (styp * int) list
   297 
   298 fun is_surely_inconsistent_card_assign hol_ctxt binarize
   299                                        (card_assigns, max_assigns) (T, k) =
   300   case binarized_and_boxed_datatype_constrs hol_ctxt binarize T of
   301     [] => false
   302   | xs =>
   303     let
   304       val dom_cards =
   305         map (Integer.prod o map (bounded_card_of_type k ~1 card_assigns)
   306              o binder_types o snd) xs
   307       val maxes = map (constr_max max_assigns) xs
   308       fun effective_max card ~1 = card
   309         | effective_max card max = Int.min (card, max)
   310       val max = map2 effective_max dom_cards maxes |> Integer.sum
   311     in max < k end
   312 fun is_surely_inconsistent_scope_description hol_ctxt binarize seen rest
   313                                              max_assigns =
   314   exists (is_surely_inconsistent_card_assign hol_ctxt binarize
   315                                              (seen @ rest, max_assigns)) seen
   316 
   317 fun repair_card_assigns hol_ctxt binarize (card_assigns, max_assigns) =
   318   let
   319     fun aux seen [] = SOME seen
   320       | aux _ ((_, 0) :: _) = NONE
   321       | aux seen ((T, k) :: rest) =
   322         (if is_surely_inconsistent_scope_description hol_ctxt binarize
   323                 ((T, k) :: seen) rest max_assigns then
   324            raise SAME ()
   325          else
   326            case aux ((T, k) :: seen) rest of
   327              SOME assigns => SOME assigns
   328            | NONE => raise SAME ())
   329         handle SAME () => aux seen ((T, k - 1) :: rest)
   330   in aux [] (rev card_assigns) end
   331 
   332 fun repair_iterator_assign thy assigns (T as Type (_, Ts), k) =
   333     (T, if T = @{typ bisim_iterator} then
   334           let
   335             val co_cards = map snd (filter (is_codatatype thy o fst) assigns)
   336           in Int.min (k, Integer.sum co_cards) end
   337         else if is_fp_iterator_type T then
   338           case Ts of
   339             [] => 1
   340           | _ => bounded_card_of_type k ~1 assigns (foldr1 HOLogic.mk_prodT Ts)
   341         else
   342           k)
   343   | repair_iterator_assign _ _ assign = assign
   344 
   345 fun add_row_to_scope_descriptor (kind, ks) (card_assigns, max_assigns) =
   346   case kind of
   347     Card T => ((T, the_single ks) :: card_assigns, max_assigns)
   348   | Max x => (card_assigns, (x, the_single ks) :: max_assigns)
   349 fun scope_descriptor_from_block block =
   350   fold_rev add_row_to_scope_descriptor block ([], [])
   351 fun scope_descriptor_from_combination (hol_ctxt as {thy, ...}) binarize blocks
   352                                       columns =
   353   let
   354     val (card_assigns, max_assigns) =
   355       maps project_block (columns ~~ blocks) |> scope_descriptor_from_block
   356     val card_assigns =
   357       repair_card_assigns hol_ctxt binarize (card_assigns, max_assigns) |> the
   358   in
   359     SOME (map (repair_iterator_assign thy card_assigns) card_assigns,
   360           max_assigns)
   361   end
   362   handle Option.Option => NONE
   363 
   364 fun offset_table_for_card_assigns assigns dtypes =
   365   let
   366     fun aux next _ [] = Typtab.update_new (dummyT, next)
   367       | aux next reusable ((T, k) :: assigns) =
   368         if k = 1 orelse is_iterator_type T orelse is_integer_type T
   369            orelse is_bit_type T then
   370           aux next reusable assigns
   371         else if length (these (Option.map #constrs (datatype_spec dtypes T)))
   372                 > 1 then
   373           Typtab.update_new (T, next) #> aux (next + k) reusable assigns
   374         else
   375           case AList.lookup (op =) reusable k of
   376             SOME j0 => Typtab.update_new (T, j0) #> aux next reusable assigns
   377           | NONE => Typtab.update_new (T, next)
   378                     #> aux (next + k) ((k, next) :: reusable) assigns
   379   in aux 0 [] assigns Typtab.empty end
   380 
   381 fun domain_card max card_assigns =
   382   Integer.prod o map (bounded_card_of_type max max card_assigns) o binder_types
   383 
   384 fun add_constr_spec (card_assigns, max_assigns) co card sum_dom_cards
   385                     num_self_recs num_non_self_recs (self_rec, x as (_, T))
   386                     constrs =
   387   let
   388     val max = constr_max max_assigns x
   389     fun next_delta () = if null constrs then 0 else #epsilon (hd constrs)
   390     val {delta, epsilon, exclusive, total} =
   391       if max = 0 then
   392         let val delta = next_delta () in
   393           {delta = delta, epsilon = delta, exclusive = true, total = false}
   394         end
   395       else if not co andalso num_self_recs > 0 then
   396         (if num_self_recs = 1 andalso num_non_self_recs = 1 then
   397            if self_rec then
   398              case constrs of
   399                [{delta = 0, epsilon = 1, exclusive = true, ...}] =>
   400                {delta = 1, epsilon = card, exclusive = true, total = false}
   401              | _ => raise SAME ()
   402            else
   403              if domain_card 2 card_assigns T = 1 then
   404                {delta = 0, epsilon = 1, exclusive = true, total = true}
   405              else
   406                raise SAME ()
   407          else
   408            raise SAME ())
   409         handle SAME () =>
   410                {delta = 0, epsilon = card, exclusive = false, total = false}
   411       else if card = sum_dom_cards (card + 1) then
   412         let val delta = next_delta () in
   413           {delta = delta, epsilon = delta + domain_card card card_assigns T,
   414            exclusive = true, total = true}
   415         end
   416       else
   417         {delta = 0, epsilon = card,
   418          exclusive = (num_self_recs + num_non_self_recs = 1), total = false}
   419   in
   420     {const = x, delta = delta, epsilon = epsilon, exclusive = exclusive,
   421      explicit_max = max, total = total} :: constrs
   422   end
   423 
   424 fun has_exact_card hol_ctxt facto finitizable_dataTs card_assigns T =
   425   let val card = card_of_type card_assigns T in
   426     card = bounded_exact_card_of_type hol_ctxt
   427                (if facto then finitizable_dataTs else []) (card + 1) 0
   428                card_assigns T
   429   end
   430 
   431 fun datatype_spec_from_scope_descriptor (hol_ctxt as {thy, stds, ...}) binarize
   432         deep_dataTs finitizable_dataTs (desc as (card_assigns, _)) (T, card) =
   433   let
   434     val deep = member (op =) deep_dataTs T
   435     val co = is_codatatype thy T
   436     val standard = is_standard_datatype thy stds T
   437     val xs = binarized_and_boxed_datatype_constrs hol_ctxt binarize T
   438     val self_recs = map (is_self_recursive_constr_type o snd) xs
   439     val (num_self_recs, num_non_self_recs) =
   440       List.partition I self_recs |> pairself length
   441     fun is_complete facto =
   442       has_exact_card hol_ctxt facto finitizable_dataTs card_assigns T
   443     fun is_concrete facto =
   444       is_word_type T orelse
   445       xs |> maps (binder_types o snd) |> maps binder_types
   446          |> forall (has_exact_card hol_ctxt facto finitizable_dataTs
   447                                    card_assigns)
   448     val complete = pair_from_fun is_complete
   449     val concrete = pair_from_fun is_concrete
   450     fun sum_dom_cards max =
   451       map (domain_card max card_assigns o snd) xs |> Integer.sum
   452     val constrs =
   453       fold_rev (add_constr_spec desc co card sum_dom_cards num_self_recs
   454                                 num_non_self_recs)
   455                (sort (bool_ord o swap o pairself fst) (self_recs ~~ xs)) []
   456   in
   457     {typ = T, card = card, co = co, standard = standard, complete = complete,
   458      concrete = concrete, deep = deep, constrs = constrs}
   459   end
   460 
   461 fun scope_from_descriptor (hol_ctxt as {thy, stds, ...}) binarize deep_dataTs
   462                           finitizable_dataTs (desc as (card_assigns, _)) =
   463   let
   464     val datatypes =
   465       map (datatype_spec_from_scope_descriptor hol_ctxt binarize deep_dataTs
   466                                                finitizable_dataTs desc)
   467           (filter (is_datatype thy stds o fst) card_assigns)
   468     val bits = card_of_type card_assigns @{typ signed_bit} - 1
   469                handle TYPE ("Nitpick_HOL.card_of_type", _, _) =>
   470                       card_of_type card_assigns @{typ unsigned_bit}
   471                       handle TYPE ("Nitpick_HOL.card_of_type", _, _) => 0
   472     val bisim_depth = card_of_type card_assigns @{typ bisim_iterator} - 1
   473   in
   474     {hol_ctxt = hol_ctxt, binarize = binarize, card_assigns = card_assigns,
   475      datatypes = datatypes, bits = bits, bisim_depth = bisim_depth,
   476      ofs = offset_table_for_card_assigns card_assigns datatypes}
   477   end
   478 
   479 fun repair_cards_assigns_wrt_boxing_etc _ _ [] = []
   480   | repair_cards_assigns_wrt_boxing_etc thy Ts ((SOME T, ks) :: cards_assigns) =
   481     (if is_fun_type T orelse is_pair_type T then
   482        Ts |> filter (curry (type_match thy o swap) T) |> map (rpair ks o SOME)
   483      else
   484        [(SOME T, ks)]) @
   485        repair_cards_assigns_wrt_boxing_etc thy Ts cards_assigns
   486   | repair_cards_assigns_wrt_boxing_etc thy Ts ((NONE, ks) :: cards_assigns) =
   487     (NONE, ks) :: repair_cards_assigns_wrt_boxing_etc thy Ts cards_assigns
   488 
   489 val max_scopes = 4096
   490 val distinct_threshold = 512
   491 
   492 fun all_scopes (hol_ctxt as {thy, ...}) binarize cards_assigns maxes_assigns
   493                iters_assigns bitss bisim_depths mono_Ts nonmono_Ts deep_dataTs
   494                finitizable_dataTs =
   495   let
   496     val cards_assigns = repair_cards_assigns_wrt_boxing_etc thy mono_Ts
   497                                                             cards_assigns
   498     val blocks = blocks_for_types hol_ctxt binarize cards_assigns maxes_assigns
   499                                   iters_assigns bitss bisim_depths mono_Ts
   500                                   nonmono_Ts
   501     val ranks = map rank_of_block blocks
   502     val all = all_combinations_ordered_smartly (map (rpair 0) ranks)
   503     val head = take max_scopes all
   504     val descs =
   505       map_filter (scope_descriptor_from_combination hol_ctxt binarize blocks)
   506                  head
   507   in
   508     (length all - length head,
   509      descs |> length descs <= distinct_threshold ? distinct (op =)
   510            |> map (scope_from_descriptor hol_ctxt binarize deep_dataTs
   511                                          finitizable_dataTs))
   512   end
   513 
   514 end;