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