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