src/HOL/Tools/Nitpick/nitpick_model.ML
author blanchet
Fri, 18 Dec 2009 12:00:29 +0100
changeset 34123 8a2c5d7aff51
parent 34121 c4628a1dcf75
child 34923 c4f04bee79f3
permissions -rw-r--r--
polished Nitpick's binary integer support etc.;
etc. = improve inconsistent scope correction + sort values nicely in output
+ handle "mod" using the characterization "x mod y = x - x div y * y"
(instead of explicit code in Nitpick)
+ introduce KK = Kodkod as abbreviation
     1 (*  Title:      HOL/Tools/Nitpick/nitpick_model.ML
     2     Author:     Jasmin Blanchette, TU Muenchen
     3     Copyright   2009
     4 
     5 Model reconstruction for Nitpick.
     6 *)
     7 
     8 signature NITPICK_MODEL =
     9 sig
    10   type styp = Nitpick_Util.styp
    11   type scope = Nitpick_Scope.scope
    12   type rep = Nitpick_Rep.rep
    13   type nut = Nitpick_Nut.nut
    14 
    15   type params = {
    16     show_skolems: bool,
    17     show_datatypes: bool,
    18     show_consts: bool}
    19 
    20   structure NameTable : TABLE
    21 
    22   val tuple_list_for_name :
    23     nut NameTable.table -> Kodkod.raw_bound list -> nut -> int list list
    24   val reconstruct_hol_model :
    25     params -> scope -> (term option * int list) list -> styp list -> nut list
    26     -> nut list -> nut list -> nut NameTable.table -> Kodkod.raw_bound list
    27     -> Pretty.T * bool
    28   val prove_hol_model :
    29     scope -> Time.time option -> nut list -> nut list -> nut NameTable.table
    30     -> Kodkod.raw_bound list -> term -> bool option
    31 end;
    32 
    33 structure Nitpick_Model : NITPICK_MODEL =
    34 struct
    35 
    36 open Nitpick_Util
    37 open Nitpick_HOL
    38 open Nitpick_Scope
    39 open Nitpick_Peephole
    40 open Nitpick_Rep
    41 open Nitpick_Nut
    42 
    43 structure KK = Kodkod
    44 
    45 type params = {
    46   show_skolems: bool,
    47   show_datatypes: bool,
    48   show_consts: bool}
    49 
    50 val unknown = "?"
    51 val unrep = "\<dots>"
    52 val maybe_mixfix = "_\<^sup>?"
    53 val base_mixfix = "_\<^bsub>base\<^esub>"
    54 val step_mixfix = "_\<^bsub>step\<^esub>"
    55 val abs_mixfix = "\<guillemotleft>_\<guillemotright>"
    56 val non_opt_name = nitpick_prefix ^ "non_opt"
    57 
    58 (* string -> int -> string *)
    59 fun atom_suffix s j =
    60   nat_subscript (j + 1)
    61   |> (s <> "" andalso Symbol.is_ascii_digit (List.last (explode s)))
    62      ? prefix "\<^isub>,"
    63 (* string -> typ -> int -> string *)
    64 fun atom_name prefix (T as Type (s, _)) j =
    65     prefix ^ substring (shortest_name s, 0, 1) ^ atom_suffix s j
    66   | atom_name prefix (T as TFree (s, _)) j =
    67     prefix ^ perhaps (try (unprefix "'")) s ^ atom_suffix s j
    68   | atom_name _ T _ = raise TYPE ("Nitpick_Model.atom_name", [T], [])
    69 (* bool -> typ -> int -> term *)
    70 fun atom for_auto T j =
    71   if for_auto then
    72     Free (atom_name (hd (space_explode "." nitpick_prefix)) T j, T)
    73   else
    74     Const (atom_name "" T j, T)
    75 
    76 (* term -> real *)
    77 fun extract_real_number (Const (@{const_name HOL.divide}, _) $ t1 $ t2) =
    78     real (snd (HOLogic.dest_number t1)) / real (snd (HOLogic.dest_number t2))
    79   | extract_real_number t = real (snd (HOLogic.dest_number t))
    80 (* term * term -> order *)
    81 fun nice_term_ord (Abs (_, _, t1), Abs (_, _, t2)) = nice_term_ord (t1, t2)
    82   | nice_term_ord tp = Real.compare (pairself extract_real_number tp)
    83     handle TERM ("dest_number", _) =>
    84            case tp of
    85              (t11 $ t12, t21 $ t22) =>
    86              (case nice_term_ord (t11, t21) of
    87                 EQUAL => nice_term_ord (t12, t22)
    88               | ord => ord)
    89            | _ => TermOrd.fast_term_ord tp
    90 
    91 (* nut NameTable.table -> KK.raw_bound list -> nut -> int list list *)
    92 fun tuple_list_for_name rel_table bounds name =
    93   the (AList.lookup (op =) bounds (the_rel rel_table name)) handle NUT _ => [[]]
    94 
    95 (* term -> term *)
    96 fun unbit_and_unbox_term (Const (@{const_name FunBox}, _) $ t1) =
    97     unbit_and_unbox_term t1
    98   | unbit_and_unbox_term (Const (@{const_name PairBox},
    99                           Type ("fun", [T1, Type ("fun", [T2, T3])]))
   100                           $ t1 $ t2) =
   101     let val Ts = map unbit_and_unbox_type [T1, T2] in
   102       Const (@{const_name Pair}, Ts ---> Type ("*", Ts))
   103       $ unbit_and_unbox_term t1 $ unbit_and_unbox_term t2
   104     end
   105   | unbit_and_unbox_term (Const (s, T)) = Const (s, unbit_and_unbox_type T)
   106   | unbit_and_unbox_term (t1 $ t2) =
   107     unbit_and_unbox_term t1 $ unbit_and_unbox_term t2
   108   | unbit_and_unbox_term (Free (s, T)) = Free (s, unbit_and_unbox_type T)
   109   | unbit_and_unbox_term (Var (x, T)) = Var (x, unbit_and_unbox_type T)
   110   | unbit_and_unbox_term (Bound j) = Bound j
   111   | unbit_and_unbox_term (Abs (s, T, t')) =
   112     Abs (s, unbit_and_unbox_type T, unbit_and_unbox_term t')
   113 
   114 (* typ -> typ -> (typ * typ) * (typ * typ) *)
   115 fun factor_out_types (T1 as Type ("*", [T11, T12]))
   116                      (T2 as Type ("*", [T21, T22])) =
   117     let val (n1, n2) = pairself num_factors_in_type (T11, T21) in
   118       if n1 = n2 then
   119         let
   120           val ((T11', opt_T12'), (T21', opt_T22')) = factor_out_types T12 T22
   121         in
   122           ((Type ("*", [T11, T11']), opt_T12'),
   123            (Type ("*", [T21, T21']), opt_T22'))
   124         end
   125       else if n1 < n2 then
   126         case factor_out_types T1 T21 of
   127           (p1, (T21', NONE)) => (p1, (T21', SOME T22))
   128         | (p1, (T21', SOME T22')) =>
   129           (p1, (T21', SOME (Type ("*", [T22', T22]))))
   130       else
   131         swap (factor_out_types T2 T1)
   132     end
   133   | factor_out_types (Type ("*", [T11, T12])) T2 = ((T11, SOME T12), (T2, NONE))
   134   | factor_out_types T1 (Type ("*", [T21, T22])) = ((T1, NONE), (T21, SOME T22))
   135   | factor_out_types T1 T2 = ((T1, NONE), (T2, NONE))
   136 
   137 (* bool -> typ -> typ -> (term * term) list -> term *)
   138 fun make_plain_fun maybe_opt T1 T2 =
   139   let
   140     (* typ -> typ -> (term * term) list -> term *)
   141     fun aux T1 T2 [] =
   142         Const (if maybe_opt orelse T2 <> bool_T then @{const_name undefined}
   143                else non_opt_name, T1 --> T2)
   144       | aux T1 T2 ((t1, t2) :: ps) =
   145         Const (@{const_name fun_upd}, (T1 --> T2) --> T1 --> T2 --> T1 --> T2)
   146         $ aux T1 T2 ps $ t1 $ t2
   147   in aux T1 T2 o rev end
   148 (* term -> bool *)
   149 fun is_plain_fun (Const (s, _)) =
   150     (s = @{const_name undefined} orelse s = non_opt_name)
   151   | is_plain_fun (Const (@{const_name fun_upd}, _) $ t0 $ _ $ _) =
   152     is_plain_fun t0
   153   | is_plain_fun _ = false
   154 (* term -> bool * (term list * term list) *)
   155 val dest_plain_fun =
   156   let
   157     (* term -> term list * term list *)
   158     fun aux (Const (s, _)) = (s <> non_opt_name, ([], []))
   159       | aux (Const (@{const_name fun_upd}, _) $ t0 $ t1 $ t2) =
   160         let val (s, (ts1, ts2)) = aux t0 in (s, (t1 :: ts1, t2 :: ts2)) end
   161       | aux t = raise TERM ("Nitpick_Model.dest_plain_fun", [t])
   162   in apsnd (pairself rev) o aux end
   163 
   164 (* typ -> typ -> typ -> term -> term * term *)
   165 fun break_in_two T T1 T2 t =
   166   let
   167     val ps = HOLogic.flat_tupleT_paths T
   168     val cut = length (HOLogic.strip_tupleT T1)
   169     val (ps1, ps2) = pairself HOLogic.flat_tupleT_paths (T1, T2)
   170     val (ts1, ts2) = t |> HOLogic.strip_ptuple ps |> chop cut
   171   in (HOLogic.mk_ptuple ps1 T1 ts1, HOLogic.mk_ptuple ps2 T2 ts2) end
   172 (* typ -> term -> term -> term *)
   173 fun pair_up (Type ("*", [T1', T2']))
   174             (t1 as Const (@{const_name Pair},
   175                           Type ("fun", [_, Type ("fun", [_, T1])])) $ t11 $ t12)
   176             t2 =
   177     if T1 = T1' then HOLogic.mk_prod (t1, t2)
   178     else HOLogic.mk_prod (t11, pair_up T2' t12 t2)
   179   | pair_up _ t1 t2 = HOLogic.mk_prod (t1, t2)
   180 (* typ -> term -> term list * term list -> (term * term) list*)
   181 fun multi_pair_up T1 t1 (ts2, ts3) = map2 (pair o pair_up T1 t1) ts2 ts3
   182 
   183 (* typ -> typ -> typ -> term -> term *)
   184 fun typecast_fun (Type ("fun", [T1', T2'])) T1 T2 t =
   185     let
   186       (* typ -> typ -> typ -> typ -> term -> term *)
   187       fun do_curry T1 T1a T1b T2 t =
   188         let
   189           val (maybe_opt, ps) = dest_plain_fun t
   190           val ps =
   191             ps |>> map (break_in_two T1 T1a T1b)
   192                |> uncurry (map2 (fn (t1a, t1b) => fn t2 => (t1a, (t1b, t2))))
   193                |> AList.coalesce (op =)
   194                |> map (apsnd (make_plain_fun maybe_opt T1b T2))
   195         in make_plain_fun maybe_opt T1a (T1b --> T2) ps end
   196       (* typ -> typ -> term -> term *)
   197       and do_uncurry T1 T2 t =
   198         let
   199           val (maybe_opt, tsp) = dest_plain_fun t
   200           val ps =
   201             tsp |> op ~~
   202                 |> maps (fn (t1, t2) =>
   203                             multi_pair_up T1 t1 (snd (dest_plain_fun t2)))
   204         in make_plain_fun maybe_opt T1 T2 ps end
   205       (* typ -> typ -> typ -> typ -> term -> term *)
   206       and do_arrow T1' T2' _ _ (Const (s, _)) = Const (s, T1' --> T2')
   207         | do_arrow T1' T2' T1 T2
   208                    (Const (@{const_name fun_upd}, _) $ t0 $ t1 $ t2) =
   209           Const (@{const_name fun_upd},
   210                  (T1' --> T2') --> T1' --> T2' --> T1' --> T2')
   211           $ do_arrow T1' T2' T1 T2 t0 $ do_term T1' T1 t1 $ do_term T2' T2 t2
   212         | do_arrow _ _ _ _ t =
   213           raise TERM ("Nitpick_Model.typecast_fun.do_arrow", [t])
   214       and do_fun T1' T2' T1 T2 t =
   215         case factor_out_types T1' T1 of
   216           ((_, NONE), (_, NONE)) => t |> do_arrow T1' T2' T1 T2
   217         | ((_, NONE), (T1a, SOME T1b)) =>
   218           t |> do_curry T1 T1a T1b T2 |> do_arrow T1' T2' T1a (T1b --> T2)
   219         | ((T1a', SOME T1b'), (_, NONE)) =>
   220           t |> do_arrow T1a' (T1b' --> T2') T1 T2 |> do_uncurry T1' T2'
   221         | _ => raise TYPE ("Nitpick_Model.typecast_fun.do_fun", [T1, T1'], [])
   222       (* typ -> typ -> term -> term *)
   223       and do_term (Type ("fun", [T1', T2'])) (Type ("fun", [T1, T2])) t =
   224           do_fun T1' T2' T1 T2 t
   225         | do_term (T' as Type ("*", Ts' as [T1', T2'])) (Type ("*", [T1, T2]))
   226                   (Const (@{const_name Pair}, _) $ t1 $ t2) =
   227           Const (@{const_name Pair}, Ts' ---> T')
   228           $ do_term T1' T1 t1 $ do_term T2' T2 t2
   229         | do_term T' T t =
   230           if T = T' then t
   231           else raise TYPE ("Nitpick_Model.typecast_fun.do_term", [T, T'], [])
   232     in if T1' = T1 andalso T2' = T2 then t else do_fun T1' T2' T1 T2 t end
   233   | typecast_fun T' _ _ _ = raise TYPE ("Nitpick_Model.typecast_fun", [T'], [])
   234 
   235 (* term -> string *)
   236 fun truth_const_sort_key @{const True} = "0"
   237   | truth_const_sort_key @{const False} = "2"
   238   | truth_const_sort_key _ = "1"
   239 
   240 (* typ -> term list -> term *)
   241 fun mk_tuple (Type ("*", [T1, T2])) ts =
   242     HOLogic.mk_prod (mk_tuple T1 ts,
   243         mk_tuple T2 (List.drop (ts, length (HOLogic.flatten_tupleT T1))))
   244   | mk_tuple _ (t :: _) = t
   245   | mk_tuple T [] = raise TYPE ("Nitpick_Model.mk_tuple", [T], [])
   246 
   247 (* string * string * string * string -> scope -> nut list -> nut list
   248    -> nut list -> nut NameTable.table -> KK.raw_bound list -> typ -> typ -> rep
   249    -> int list list -> term *)
   250 fun reconstruct_term (maybe_name, base_name, step_name, abs_name)
   251         ({ext_ctxt as {thy, ctxt, ...}, card_assigns, bits, datatypes, ofs, ...}
   252          : scope) sel_names rel_table bounds =
   253   let
   254     val for_auto = (maybe_name = "")
   255     (* int list list -> int *)
   256     fun value_of_bits jss =
   257       let
   258         val j0 = offset_of_type ofs @{typ unsigned_bit}
   259         val js = map (Integer.add (~ j0) o the_single) jss
   260       in
   261         fold (fn j => Integer.add (reasonable_power 2 j |> j = bits ? op ~))
   262              js 0
   263       end
   264     (* bool -> typ -> typ -> (term * term) list -> term *)
   265     fun make_set maybe_opt T1 T2 =
   266       let
   267         val empty_const = Const (@{const_name Set.empty}, T1 --> T2)
   268         val insert_const = Const (@{const_name insert},
   269                                   T1 --> (T1 --> T2) --> T1 --> T2)
   270         (* (term * term) list -> term *)
   271         fun aux [] =
   272             if maybe_opt andalso not (is_complete_type datatypes T1) then
   273               insert_const $ Const (unrep, T1) $ empty_const
   274             else
   275               empty_const
   276           | aux ((t1, t2) :: zs) =
   277             aux zs |> t2 <> @{const False}
   278                       ? curry (op $) (insert_const
   279                                       $ (t1 |> t2 <> @{const True}
   280                                                ? curry (op $)
   281                                                        (Const (maybe_name,
   282                                                                T1 --> T1))))
   283       in aux end
   284     (* typ -> typ -> typ -> (term * term) list -> term *)
   285     fun make_map T1 T2 T2' =
   286       let
   287         val update_const = Const (@{const_name fun_upd},
   288                                   (T1 --> T2) --> T1 --> T2 --> T1 --> T2)
   289         (* (term * term) list -> term *)
   290         fun aux' [] = Const (@{const_name Map.empty}, T1 --> T2)
   291           | aux' ((t1, t2) :: ps) =
   292             (case t2 of
   293                Const (@{const_name None}, _) => aux' ps
   294              | _ => update_const $ aux' ps $ t1 $ t2)
   295         fun aux ps =
   296           if not (is_complete_type datatypes T1) then
   297             update_const $ aux' ps $ Const (unrep, T1)
   298             $ (Const (@{const_name Some}, T2' --> T2) $ Const (unknown, T2'))
   299           else
   300             aux' ps
   301       in aux end
   302     (* typ list -> term -> term *)
   303     fun setify_mapify_funs Ts t =
   304       (case fastype_of1 (Ts, t) of
   305          Type ("fun", [T1, T2]) =>
   306          if is_plain_fun t then
   307            case T2 of
   308              @{typ bool} =>
   309              let
   310                val (maybe_opt, ts_pair) =
   311                  dest_plain_fun t ||> pairself (map (setify_mapify_funs Ts))
   312              in
   313                make_set maybe_opt T1 T2
   314                         (sort_wrt (truth_const_sort_key o snd) (op ~~ ts_pair))
   315              end
   316            | Type (@{type_name option}, [T2']) =>
   317              let
   318                val ts_pair = snd (dest_plain_fun t)
   319                              |> pairself (map (setify_mapify_funs Ts))
   320              in make_map T1 T2 T2' (rev (op ~~ ts_pair)) end
   321            | _ => raise SAME ()
   322          else
   323            raise SAME ()
   324        | _ => raise SAME ())
   325       handle SAME () =>
   326              case t of
   327                t1 $ t2 => setify_mapify_funs Ts t1 $ setify_mapify_funs Ts t2
   328              | Abs (s, T, t') => Abs (s, T, setify_mapify_funs (T :: Ts) t')
   329              | _ => t
   330     (* bool -> typ -> typ -> typ -> term list -> term list -> term *)
   331     fun make_fun maybe_opt T1 T2 T' ts1 ts2 =
   332       ts1 ~~ ts2 |> sort (nice_term_ord o pairself fst)
   333                  |> make_plain_fun (maybe_opt andalso not for_auto) T1 T2
   334                  |> unbit_and_unbox_term
   335                  |> typecast_fun (unbit_and_unbox_type T')
   336                                  (unbit_and_unbox_type T1)
   337                                  (unbit_and_unbox_type T2)
   338     (* (typ * int) list -> typ -> typ -> int -> term *)
   339     fun term_for_atom seen (T as Type ("fun", [T1, T2])) T' j =
   340         let
   341           val k1 = card_of_type card_assigns T1
   342           val k2 = card_of_type card_assigns T2
   343         in
   344           term_for_rep seen T T' (Vect (k1, Atom (k2, 0)))
   345                        [nth_combination (replicate k1 (k2, 0)) j]
   346           handle General.Subscript =>
   347                  raise ARG ("Nitpick_Model.reconstruct_term.term_for_atom",
   348                             signed_string_of_int j ^ " for " ^
   349                             string_for_rep (Vect (k1, Atom (k2, 0))))
   350         end
   351       | term_for_atom seen (Type ("*", [T1, T2])) _ j =
   352         let val k1 = card_of_type card_assigns T1 in
   353           list_comb (HOLogic.pair_const T1 T2,
   354                      map2 (fn T => term_for_atom seen T T) [T1, T2]
   355                           [j div k1, j mod k1])
   356         end
   357       | term_for_atom seen @{typ prop} _ j =
   358         HOLogic.mk_Trueprop (term_for_atom seen bool_T bool_T j)
   359       | term_for_atom _ @{typ bool} _ j =
   360         if j = 0 then @{const False} else @{const True}
   361       | term_for_atom _ @{typ unit} _ _ = @{const Unity}
   362       | term_for_atom seen T _ j =
   363         if T = nat_T then
   364           HOLogic.mk_number nat_T j
   365         else if T = int_T then
   366           HOLogic.mk_number int_T
   367               (int_for_atom (card_of_type card_assigns int_T, 0) j)
   368         else if is_fp_iterator_type T then
   369           HOLogic.mk_number nat_T (card_of_type card_assigns T - j - 1)
   370         else if T = @{typ bisim_iterator} then
   371           HOLogic.mk_number nat_T j
   372         else case datatype_spec datatypes T of
   373           NONE => atom for_auto T j
   374         | SOME {shallow = true, ...} => atom for_auto T j
   375         | SOME {co, constrs, ...} =>
   376           let
   377             (* styp -> int list *)
   378             fun tuples_for_const (s, T) =
   379               tuple_list_for_name rel_table bounds (ConstName (s, T, Any))
   380             (* unit -> indexname * typ *)
   381             fun var () = ((atom_name "" T j, 0), T)
   382             val discr_jsss = map (tuples_for_const o discr_for_constr o #const)
   383                                  constrs
   384             val real_j = j + offset_of_type ofs T
   385             val constr_x as (constr_s, constr_T) =
   386               get_first (fn (jss, {const, ...}) =>
   387                             if member (op =) jss [real_j] then SOME const
   388                             else NONE)
   389                         (discr_jsss ~~ constrs) |> the
   390             val arg_Ts = curried_binder_types constr_T
   391             val sel_xs = map (boxed_nth_sel_for_constr ext_ctxt constr_x)
   392                              (index_seq 0 (length arg_Ts))
   393             val sel_Rs =
   394               map (fn x => get_first
   395                                (fn ConstName (s', T', R) =>
   396                                    if (s', T') = x then SOME R else NONE
   397                                  | u => raise NUT ("Nitpick_Model.reconstruct_\
   398                                                    \term.term_for_atom", [u]))
   399                                sel_names |> the) sel_xs
   400             val arg_Rs = map (snd o dest_Func) sel_Rs
   401             val sel_jsss = map tuples_for_const sel_xs
   402             val arg_jsss =
   403               map (map_filter (fn js => if hd js = real_j then SOME (tl js)
   404                                         else NONE)) sel_jsss
   405             val uncur_arg_Ts = binder_types constr_T
   406           in
   407             if co andalso member (op =) seen (T, j) then
   408               Var (var ())
   409             else if constr_s = @{const_name Word} then
   410               HOLogic.mk_number
   411                   (if T = @{typ "unsigned_bit word"} then nat_T else int_T)
   412                   (value_of_bits (the_single arg_jsss))
   413             else
   414               let
   415                 val seen = seen |> co ? cons (T, j)
   416                 val ts =
   417                   if length arg_Ts = 0 then
   418                     []
   419                   else
   420                     map3 (fn Ts => term_for_rep seen Ts Ts) arg_Ts arg_Rs
   421                          arg_jsss
   422                     |> mk_tuple (HOLogic.mk_tupleT uncur_arg_Ts)
   423                     |> dest_n_tuple (length uncur_arg_Ts)
   424                 val t =
   425                   if constr_s = @{const_name Abs_Frac} then
   426                     let
   427                       val num_T = body_type T
   428                       (* int -> term *)
   429                       val mk_num = HOLogic.mk_number num_T
   430                     in
   431                       case ts of
   432                         [Const (@{const_name Pair}, _) $ t1 $ t2] =>
   433                         (case snd (HOLogic.dest_number t1) of
   434                            0 => mk_num 0
   435                          | n1 => case HOLogic.dest_number t2 |> snd of
   436                                    1 => mk_num n1
   437                                  | n2 => Const (@{const_name HOL.divide},
   438                                                 num_T --> num_T --> num_T)
   439                                          $ mk_num n1 $ mk_num n2)
   440                       | _ => raise TERM ("Nitpick_Model.reconstruct_term.term_\
   441                                          \for_atom (Abs_Frac)", ts)
   442                     end
   443                   else if not for_auto andalso is_abs_fun thy constr_x then
   444                     Const (abs_name, constr_T) $ the_single ts
   445                   else
   446                     list_comb (Const constr_x, ts)
   447               in
   448                 if co then
   449                   let val var = var () in
   450                     if exists_subterm (curry (op =) (Var var)) t then
   451                       Const (@{const_name The}, (T --> bool_T) --> T)
   452                       $ Abs ("\<omega>", T,
   453                              Const (@{const_name "op ="}, T --> T --> bool_T)
   454                              $ Bound 0 $ abstract_over (Var var, t))
   455                     else
   456                       t
   457                   end
   458                 else
   459                   t
   460               end
   461           end
   462     (* (typ * int) list -> int -> rep -> typ -> typ -> typ -> int list
   463        -> term *)
   464     and term_for_vect seen k R T1 T2 T' js =
   465       make_fun true T1 T2 T' (map (term_for_atom seen T1 T1) (index_seq 0 k))
   466                (map (term_for_rep seen T2 T2 R o single)
   467                     (batch_list (arity_of_rep R) js))
   468     (* (typ * int) list -> typ -> typ -> rep -> int list list -> term *)
   469     and term_for_rep seen T T' Unit [[]] = term_for_atom seen T T' 0
   470       | term_for_rep seen T T' (R as Atom (k, j0)) [[j]] =
   471         if j >= j0 andalso j < j0 + k then term_for_atom seen T T' (j - j0)
   472         else raise REP ("Nitpick_Model.reconstruct_term.term_for_rep", [R])
   473       | term_for_rep seen (Type ("*", [T1, T2])) _ (Struct [R1, R2]) [js] =
   474         let
   475           val arity1 = arity_of_rep R1
   476           val (js1, js2) = chop arity1 js
   477         in
   478           list_comb (HOLogic.pair_const T1 T2,
   479                      map3 (fn T => term_for_rep seen T T) [T1, T2] [R1, R2]
   480                           [[js1], [js2]])
   481         end
   482       | term_for_rep seen (Type ("fun", [T1, T2])) T' (R as Vect (k, R')) [js] =
   483         term_for_vect seen k R' T1 T2 T' js
   484       | term_for_rep seen (Type ("fun", [T1, T2])) T' (Func (R1, Formula Neut))
   485                      jss =
   486         let
   487           val jss1 = all_combinations_for_rep R1
   488           val ts1 = map (term_for_rep seen T1 T1 R1 o single) jss1
   489           val ts2 =
   490             map (fn js => term_for_rep seen T2 T2 (Atom (2, 0))
   491                                        [[int_for_bool (member (op =) jss js)]])
   492                 jss1
   493         in make_fun false T1 T2 T' ts1 ts2 end
   494       | term_for_rep seen (Type ("fun", [T1, T2])) T' (Func (R1, R2)) jss =
   495         let
   496           val arity1 = arity_of_rep R1
   497           val jss1 = all_combinations_for_rep R1
   498           val ts1 = map (term_for_rep seen T1 T1 R1 o single) jss1
   499           val grouped_jss2 = AList.group (op =) (map (chop arity1) jss)
   500           val ts2 = map (term_for_rep seen T2 T2 R2 o the_default []
   501                          o AList.lookup (op =) grouped_jss2) jss1
   502         in make_fun true T1 T2 T' ts1 ts2 end
   503       | term_for_rep seen T T' (Opt R) jss =
   504         if null jss then Const (unknown, T) else term_for_rep seen T T' R jss
   505       | term_for_rep seen T _ R jss =
   506         raise ARG ("Nitpick_Model.reconstruct_term.term_for_rep",
   507                    Refute.string_of_typ T ^ " " ^ string_for_rep R ^ " " ^
   508                    string_of_int (length jss))
   509   in
   510     (not for_auto ? setify_mapify_funs []) o unbit_and_unbox_term
   511     oooo term_for_rep []
   512   end
   513 
   514 (* scope -> nut list -> nut NameTable.table -> KK.raw_bound list -> nut
   515    -> term *)
   516 fun term_for_name scope sel_names rel_table bounds name =
   517   let val T = type_of name in
   518     tuple_list_for_name rel_table bounds name
   519     |> reconstruct_term ("", "", "", "") scope sel_names rel_table bounds T T
   520                         (rep_of name)
   521   end
   522 
   523 (* Proof.context
   524    -> (string * string * string * string * string) * Proof.context *)
   525 fun add_wacky_syntax ctxt =
   526   let
   527     (* term -> string *)
   528     val name_of = fst o dest_Const
   529     val thy = ProofContext.theory_of ctxt |> Context.reject_draft
   530     val (maybe_t, thy) =
   531       Sign.declare_const ((@{binding nitpick_maybe}, @{typ "'a => 'a"}),
   532                           Mixfix (maybe_mixfix, [1000], 1000)) thy
   533     val (base_t, thy) =
   534       Sign.declare_const ((@{binding nitpick_base}, @{typ "'a => 'a"}),
   535                           Mixfix (base_mixfix, [1000], 1000)) thy
   536     val (step_t, thy) =
   537       Sign.declare_const ((@{binding nitpick_step}, @{typ "'a => 'a"}),
   538                           Mixfix (step_mixfix, [1000], 1000)) thy
   539     val (abs_t, thy) =
   540       Sign.declare_const ((@{binding nitpick_abs}, @{typ "'a => 'b"}),
   541                           Mixfix (abs_mixfix, [40], 40)) thy
   542   in
   543     ((name_of maybe_t, name_of base_t, name_of step_t, name_of abs_t),
   544      ProofContext.transfer_syntax thy ctxt)
   545   end
   546 
   547 (* term -> term *)
   548 fun unfold_outer_the_binders (t as Const (@{const_name The}, _)
   549                                    $ Abs (s, T, Const (@{const_name "op ="}, _)
   550                                                 $ Bound 0 $ t')) =
   551     betapply (Abs (s, T, t'), t) |> unfold_outer_the_binders
   552   | unfold_outer_the_binders t = t
   553 (* typ list -> int -> term * term -> bool *)
   554 fun bisimilar_values _ 0 _ = true
   555   | bisimilar_values coTs max_depth (t1, t2) =
   556     let val T = fastype_of t1 in
   557       if exists_subtype (member (op =) coTs) T then
   558         let
   559           val ((head1, args1), (head2, args2)) =
   560             pairself (strip_comb o unfold_outer_the_binders) (t1, t2)
   561           val max_depth = max_depth - (if member (op =) coTs T then 1 else 0)
   562         in
   563           head1 = head2
   564           andalso forall (bisimilar_values coTs max_depth) (args1 ~~ args2)
   565         end
   566       else
   567         t1 = t2
   568     end
   569 
   570 (* params -> scope -> (term option * int list) list -> styp list -> nut list
   571   -> nut list -> nut list -> nut NameTable.table -> KK.raw_bound list
   572   -> Pretty.T * bool *)
   573 fun reconstruct_hol_model {show_skolems, show_datatypes, show_consts}
   574         ({ext_ctxt as {thy, ctxt, max_bisim_depth, boxes, wfs, user_axioms,
   575                        debug, binary_ints, destroy_constrs, specialize,
   576                        skolemize, star_linear_preds, uncurry, fast_descrs,
   577                        tac_timeout, evals, case_names, def_table, nondef_table,
   578                        user_nondefs, simp_table, psimp_table, intro_table,
   579                        ground_thm_table, ersatz_table, skolems, special_funs,
   580                        unrolled_preds, wf_cache, constr_cache},
   581          card_assigns, bits, bisim_depth, datatypes, ofs} : scope)
   582         formats all_frees free_names sel_names nonsel_names rel_table bounds =
   583   let
   584     val (wacky_names as (_, base_name, step_name, _), ctxt) =
   585       add_wacky_syntax ctxt
   586     val ext_ctxt =
   587       {thy = thy, ctxt = ctxt, max_bisim_depth = max_bisim_depth, boxes = boxes,
   588        wfs = wfs, user_axioms = user_axioms, debug = debug,
   589        binary_ints = binary_ints, destroy_constrs = destroy_constrs,
   590        specialize = specialize, skolemize = skolemize,
   591        star_linear_preds = star_linear_preds, uncurry = uncurry,
   592        fast_descrs = fast_descrs, tac_timeout = tac_timeout, evals = evals,
   593        case_names = case_names, def_table = def_table,
   594        nondef_table = nondef_table, user_nondefs = user_nondefs,
   595        simp_table = simp_table, psimp_table = psimp_table,
   596        intro_table = intro_table, ground_thm_table = ground_thm_table,
   597        ersatz_table = ersatz_table, skolems = skolems,
   598        special_funs = special_funs, unrolled_preds = unrolled_preds,
   599        wf_cache = wf_cache, constr_cache = constr_cache}
   600     val scope = {ext_ctxt = ext_ctxt, card_assigns = card_assigns,
   601                  bits = bits, bisim_depth = bisim_depth, datatypes = datatypes,
   602                  ofs = ofs}
   603     (* typ -> typ -> rep -> int list list -> term *)
   604     val term_for_rep = reconstruct_term wacky_names scope sel_names rel_table
   605                                         bounds
   606     (* nat -> typ -> nat -> typ *)
   607     fun nth_value_of_type card T n = term_for_rep T T (Atom (card, 0)) [[n]]
   608     (* nat -> typ -> typ list *)
   609     fun all_values_of_type card T =
   610       index_seq 0 card |> map (nth_value_of_type card T) |> sort nice_term_ord
   611     (* dtype_spec list -> dtype_spec -> bool *)
   612     fun is_codatatype_wellformed (cos : dtype_spec list)
   613                                  ({typ, card, ...} : dtype_spec) =
   614       let
   615         val ts = all_values_of_type card typ
   616         val max_depth = Integer.sum (map #card cos)
   617       in
   618         forall (not o bisimilar_values (map #typ cos) max_depth)
   619                (all_distinct_unordered_pairs_of ts)
   620       end
   621     (* string -> Pretty.T *)
   622     fun pretty_for_assign name =
   623       let
   624         val (oper, (t1, T'), T) =
   625           case name of
   626             FreeName (s, T, _) =>
   627             let val t = Free (s, unbit_and_unbox_type T) in
   628               ("=", (t, format_term_type thy def_table formats t), T)
   629             end
   630           | ConstName (s, T, _) =>
   631             (assign_operator_for_const (s, T),
   632              user_friendly_const ext_ctxt (base_name, step_name) formats (s, T),
   633              T)
   634           | _ => raise NUT ("Nitpick_Model.reconstruct_hol_model.\
   635                             \pretty_for_assign", [name])
   636         val t2 = if rep_of name = Any then
   637                    Const (@{const_name undefined}, T')
   638                  else
   639                    tuple_list_for_name rel_table bounds name
   640                    |> term_for_rep T T' (rep_of name)
   641       in
   642         Pretty.block (Pretty.breaks
   643             [setmp_show_all_types (Syntax.pretty_term ctxt) t1,
   644              Pretty.str oper, Syntax.pretty_term ctxt t2])
   645       end
   646     (* dtype_spec -> Pretty.T *)
   647     fun pretty_for_datatype ({typ, card, complete, ...} : dtype_spec) =
   648       Pretty.block (Pretty.breaks
   649           [Syntax.pretty_typ ctxt (unbit_and_unbox_type typ), Pretty.str "=",
   650            Pretty.enum "," "{" "}"
   651                (map (Syntax.pretty_term ctxt) (all_values_of_type card typ)
   652                 @ (if complete then [] else [Pretty.str unrep]))])
   653     (* typ -> dtype_spec list *)
   654     fun integer_datatype T =
   655       [{typ = T, card = card_of_type card_assigns T, co = false,
   656         complete = false, concrete = true, shallow = false, constrs = []}]
   657       handle TYPE ("Nitpick_HOL.card_of_type", _, _) => []
   658     val (codatatypes, datatypes) =
   659       datatypes |> filter_out #shallow
   660                 |> List.partition #co
   661                 ||> append (integer_datatype nat_T @ integer_datatype int_T)
   662     val block_of_datatypes =
   663       if show_datatypes andalso not (null datatypes) then
   664         [Pretty.big_list ("Datatype" ^ plural_s_for_list datatypes ^ ":")
   665                          (map pretty_for_datatype datatypes)]
   666       else
   667         []
   668     val block_of_codatatypes =
   669       if show_datatypes andalso not (null codatatypes) then
   670         [Pretty.big_list ("Codatatype" ^ plural_s_for_list codatatypes ^ ":")
   671                          (map pretty_for_datatype codatatypes)]
   672       else
   673         []
   674     (* bool -> string -> nut list -> Pretty.T list *)
   675     fun block_of_names show title names =
   676       if show andalso not (null names) then
   677         Pretty.str (title ^ plural_s_for_list names ^ ":")
   678         :: map (Pretty.indent indent_size o pretty_for_assign)
   679                (sort_wrt (original_name o nickname_of) names)
   680       else
   681         []
   682     val (skolem_names, nonskolem_nonsel_names) =
   683       List.partition is_skolem_name nonsel_names
   684     val (eval_names, noneval_nonskolem_nonsel_names) =
   685       List.partition (String.isPrefix eval_prefix o nickname_of)
   686                      nonskolem_nonsel_names
   687       ||> filter_out (curry (op =) @{const_name bisim_iterator_max}
   688                       o nickname_of)
   689     val free_names =
   690       map (fn x as (s, T) =>
   691               case filter (curry (op =) x
   692                            o pairf nickname_of (unbit_and_unbox_type o type_of))
   693                           free_names of
   694                 [name] => name
   695               | [] => FreeName (s, T, Any)
   696               | _ => raise TERM ("Nitpick_Model.reconstruct_hol_model",
   697                                  [Const x])) all_frees
   698     val chunks = block_of_names true "Free variable" free_names @
   699                  block_of_names show_skolems "Skolem constant" skolem_names @
   700                  block_of_names true "Evaluated term" eval_names @
   701                  block_of_datatypes @ block_of_codatatypes @
   702                  block_of_names show_consts "Constant"
   703                                 noneval_nonskolem_nonsel_names
   704   in
   705     (Pretty.chunks (if null chunks then [Pretty.str "Empty assignment"]
   706                     else chunks),
   707      bisim_depth >= 0
   708      orelse forall (is_codatatype_wellformed codatatypes) codatatypes)
   709   end
   710 
   711 (* scope -> Time.time option -> nut list -> nut list -> nut NameTable.table
   712    -> KK.raw_bound list -> term -> bool option *)
   713 fun prove_hol_model (scope as {ext_ctxt as {thy, ctxt, ...}, card_assigns, ...})
   714                     auto_timeout free_names sel_names rel_table bounds prop =
   715   let
   716     (* typ * int -> term *)
   717     fun free_type_assm (T, k) =
   718       let
   719         (* int -> term *)
   720         val atom = atom true T
   721         fun equation_for_atom j = HOLogic.eq_const T $ Bound 0 $ atom j
   722         val eqs = map equation_for_atom (index_seq 0 k)
   723         val compreh_assm =
   724           Const (@{const_name All}, (T --> bool_T) --> bool_T)
   725               $ Abs ("x", T, foldl1 HOLogic.mk_disj eqs)
   726         val distinct_assm = distinctness_formula T (map atom (index_seq 0 k))
   727       in HOLogic.mk_conj (compreh_assm, distinct_assm) end
   728     (* nut -> term *)
   729     fun free_name_assm name =
   730       HOLogic.mk_eq (Free (nickname_of name, type_of name),
   731                      term_for_name scope sel_names rel_table bounds name)
   732     val freeT_assms = map free_type_assm (filter (is_TFree o fst) card_assigns)
   733     val model_assms = map free_name_assm free_names
   734     val assm = List.foldr HOLogic.mk_conj @{const True}
   735                           (freeT_assms @ model_assms)
   736     (* bool -> bool *)
   737     fun try_out negate =
   738       let
   739         val concl = (negate ? curry (op $) @{const Not})
   740                     (ObjectLogic.atomize_term thy prop)
   741         val goal = HOLogic.mk_Trueprop (HOLogic.mk_imp (assm, concl))
   742                    |> map_types (map_type_tfree
   743                           (fn (s, []) => TFree (s, HOLogic.typeS)
   744                             | x => TFree x))
   745                    |> cterm_of thy |> Goal.init
   746       in
   747         (goal |> SINGLE (DETERM_TIMEOUT auto_timeout
   748                                         (auto_tac (clasimpset_of ctxt)))
   749               |> the |> Goal.finish ctxt; true)
   750         handle THM _ => false
   751              | TimeLimit.TimeOut => false
   752       end
   753   in
   754     if try_out false then SOME true
   755     else if try_out true then SOME false
   756     else NONE
   757   end
   758 
   759 end;