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