src/HOL/Tools/Nitpick/nitpick_mono.ML
author blanchet
Sat, 03 Jul 2010 00:50:35 +0200
changeset 37695 c6161bee8486
parent 37678 0040bafffdef
child 38406 7207321df8af
permissions -rw-r--r--
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
the code in "Nitpick_Preproc", which sorted the types using "typ_ord", was wrong and evil; it seems to have worked only because "*" was called "*"
     1 (*  Title:      HOL/Tools/Nitpick/nitpick_mono.ML
     2     Author:     Jasmin Blanchette, TU Muenchen
     3     Copyright   2009, 2010
     4 
     5 Monotonicity inference for higher-order logic.
     6 *)
     7 
     8 signature NITPICK_MONO =
     9 sig
    10   type hol_context = Nitpick_HOL.hol_context
    11 
    12   val formulas_monotonic :
    13     hol_context -> bool -> typ -> term list * term list -> bool
    14   val finitize_funs :
    15     hol_context -> bool -> (typ option * bool option) list -> typ
    16     -> term list * term list -> term list * term list
    17 end;
    18 
    19 structure Nitpick_Mono : NITPICK_MONO =
    20 struct
    21 
    22 open Nitpick_Util
    23 open Nitpick_HOL
    24 
    25 type var = int
    26 
    27 datatype sign = Plus | Minus
    28 datatype sign_atom = S of sign | V of var
    29 
    30 type literal = var * sign
    31 
    32 datatype mtyp =
    33   MAlpha |
    34   MFun of mtyp * sign_atom * mtyp |
    35   MPair of mtyp * mtyp |
    36   MType of string * mtyp list |
    37   MRec of string * typ list
    38 
    39 datatype mterm =
    40   MRaw of term * mtyp |
    41   MAbs of string * typ * mtyp * sign_atom * mterm |
    42   MApp of mterm * mterm
    43 
    44 type mdata =
    45   {hol_ctxt: hol_context,
    46    binarize: bool,
    47    alpha_T: typ,
    48    no_harmless: bool,
    49    max_fresh: int Unsynchronized.ref,
    50    datatype_mcache: ((string * typ list) * mtyp) list Unsynchronized.ref,
    51    constr_mcache: (styp * mtyp) list Unsynchronized.ref}
    52 
    53 exception UNSOLVABLE of unit
    54 exception MTYPE of string * mtyp list * typ list
    55 exception MTERM of string * mterm list
    56 
    57 val debug_mono = false
    58 
    59 fun print_g f = () |> debug_mono ? tracing o f
    60 
    61 val string_for_var = signed_string_of_int
    62 fun string_for_vars sep [] = "0\<^bsub>" ^ sep ^ "\<^esub>"
    63   | string_for_vars sep xs = space_implode sep (map string_for_var xs)
    64 fun subscript_string_for_vars sep xs =
    65   if null xs then "" else "\<^bsub>" ^ string_for_vars sep xs ^ "\<^esub>"
    66 
    67 fun string_for_sign Plus = "+"
    68   | string_for_sign Minus = "-"
    69 
    70 fun xor sn1 sn2 = if sn1 = sn2 then Plus else Minus
    71 val negate = xor Minus
    72 
    73 fun string_for_sign_atom (S sn) = string_for_sign sn
    74   | string_for_sign_atom (V x) = string_for_var x
    75 
    76 fun string_for_literal (x, sn) = string_for_var x ^ " = " ^ string_for_sign sn
    77 
    78 val bool_M = MType (@{type_name bool}, [])
    79 val dummy_M = MType (nitpick_prefix ^ "dummy", [])
    80 
    81 fun is_MRec (MRec _) = true
    82   | is_MRec _ = false
    83 fun dest_MFun (MFun z) = z
    84   | dest_MFun M = raise MTYPE ("Nitpick_Mono.dest_MFun", [M], [])
    85 
    86 val no_prec = 100
    87 
    88 fun precedence_of_mtype (MFun _) = 1
    89   | precedence_of_mtype (MPair _) = 2
    90   | precedence_of_mtype _ = no_prec
    91 
    92 val string_for_mtype =
    93   let
    94     fun aux outer_prec M =
    95       let
    96         val prec = precedence_of_mtype M
    97         val need_parens = (prec < outer_prec)
    98       in
    99         (if need_parens then "(" else "") ^
   100         (if M = dummy_M then
   101            "_"
   102          else case M of
   103              MAlpha => "\<alpha>"
   104            | MFun (M1, a, M2) =>
   105              aux (prec + 1) M1 ^ " \<Rightarrow>\<^bsup>" ^
   106              string_for_sign_atom a ^ "\<^esup> " ^ aux prec M2
   107            | MPair (M1, M2) => aux (prec + 1) M1 ^ " \<times> " ^ aux prec M2
   108            | MType (s, []) =>
   109              if s = @{type_name prop} orelse s = @{type_name bool} then "o"
   110              else s
   111            | MType (s, Ms) => "(" ^ commas (map (aux 0) Ms) ^ ") " ^ s
   112            | MRec (s, _) => "[" ^ s ^ "]") ^
   113         (if need_parens then ")" else "")
   114       end
   115   in aux 0 end
   116 
   117 fun flatten_mtype (MPair (M1, M2)) = maps flatten_mtype [M1, M2]
   118   | flatten_mtype (MType (_, Ms)) = maps flatten_mtype Ms
   119   | flatten_mtype M = [M]
   120 
   121 fun precedence_of_mterm (MRaw _) = no_prec
   122   | precedence_of_mterm (MAbs _) = 1
   123   | precedence_of_mterm (MApp _) = 2
   124 
   125 fun string_for_mterm ctxt =
   126   let
   127     fun mtype_annotation M = "\<^bsup>" ^ string_for_mtype M ^ "\<^esup>"
   128     fun aux outer_prec m =
   129       let
   130         val prec = precedence_of_mterm m
   131         val need_parens = (prec < outer_prec)
   132       in
   133         (if need_parens then "(" else "") ^
   134         (case m of
   135            MRaw (t, M) => Syntax.string_of_term ctxt t ^ mtype_annotation M
   136          | MAbs (s, _, M, a, m) =>
   137            "\<lambda>" ^ s ^ mtype_annotation M ^ ".\<^bsup>" ^
   138            string_for_sign_atom a ^ "\<^esup> " ^ aux prec m
   139          | MApp (m1, m2) => aux prec m1 ^ " " ^ aux (prec + 1) m2) ^
   140         (if need_parens then ")" else "")
   141       end
   142   in aux 0 end
   143 
   144 fun mtype_of_mterm (MRaw (_, M)) = M
   145   | mtype_of_mterm (MAbs (_, _, M, a, m)) = MFun (M, a, mtype_of_mterm m)
   146   | mtype_of_mterm (MApp (m1, _)) =
   147     case mtype_of_mterm m1 of
   148       MFun (_, _, M12) => M12
   149     | M1 => raise MTYPE ("Nitpick_Mono.mtype_of_mterm", [M1], [])
   150 
   151 fun strip_mcomb (MApp (m1, m2)) = strip_mcomb m1 ||> (fn ms => append ms [m2])
   152   | strip_mcomb m = (m, [])
   153 
   154 fun initial_mdata hol_ctxt binarize no_harmless alpha_T =
   155   ({hol_ctxt = hol_ctxt, binarize = binarize, alpha_T = alpha_T,
   156     no_harmless = no_harmless, max_fresh = Unsynchronized.ref 0,
   157     datatype_mcache = Unsynchronized.ref [],
   158     constr_mcache = Unsynchronized.ref []} : mdata)
   159 
   160 fun could_exist_alpha_subtype alpha_T (T as Type (_, Ts)) =
   161     T = alpha_T orelse (not (is_fp_iterator_type T) andalso
   162                         exists (could_exist_alpha_subtype alpha_T) Ts)
   163   | could_exist_alpha_subtype alpha_T T = (T = alpha_T)
   164 fun could_exist_alpha_sub_mtype _ (alpha_T as TFree _) T =
   165     could_exist_alpha_subtype alpha_T T
   166   | could_exist_alpha_sub_mtype ctxt alpha_T T =
   167     (T = alpha_T orelse is_datatype ctxt [(NONE, true)] T)
   168 
   169 fun exists_alpha_sub_mtype MAlpha = true
   170   | exists_alpha_sub_mtype (MFun (M1, _, M2)) =
   171     exists exists_alpha_sub_mtype [M1, M2]
   172   | exists_alpha_sub_mtype (MPair (M1, M2)) =
   173     exists exists_alpha_sub_mtype [M1, M2]
   174   | exists_alpha_sub_mtype (MType (_, Ms)) = exists exists_alpha_sub_mtype Ms
   175   | exists_alpha_sub_mtype (MRec _) = true
   176 
   177 fun exists_alpha_sub_mtype_fresh MAlpha = true
   178   | exists_alpha_sub_mtype_fresh (MFun (_, V _, _)) = true
   179   | exists_alpha_sub_mtype_fresh (MFun (_, _, M2)) =
   180     exists_alpha_sub_mtype_fresh M2
   181   | exists_alpha_sub_mtype_fresh (MPair (M1, M2)) =
   182     exists exists_alpha_sub_mtype_fresh [M1, M2]
   183   | exists_alpha_sub_mtype_fresh (MType (_, Ms)) =
   184     exists exists_alpha_sub_mtype_fresh Ms
   185   | exists_alpha_sub_mtype_fresh (MRec _) = true
   186 
   187 fun constr_mtype_for_binders z Ms =
   188   fold_rev (fn M => curry3 MFun M (S Minus)) Ms (MRec z)
   189 
   190 fun repair_mtype _ _ MAlpha = MAlpha
   191   | repair_mtype cache seen (MFun (M1, a, M2)) =
   192     MFun (repair_mtype cache seen M1, a, repair_mtype cache seen M2)
   193   | repair_mtype cache seen (MPair Mp) =
   194     MPair (pairself (repair_mtype cache seen) Mp)
   195   | repair_mtype cache seen (MType (s, Ms)) =
   196     MType (s, maps (flatten_mtype o repair_mtype cache seen) Ms)
   197   | repair_mtype cache seen (MRec (z as (s, _))) =
   198     case AList.lookup (op =) cache z |> the of
   199       MRec _ => MType (s, [])
   200     | M => if member (op =) seen M then MType (s, [])
   201            else repair_mtype cache (M :: seen) M
   202 
   203 fun repair_datatype_mcache cache =
   204   let
   205     fun repair_one (z, M) =
   206       Unsynchronized.change cache
   207           (AList.update (op =) (z, repair_mtype (!cache) [] M))
   208   in List.app repair_one (rev (!cache)) end
   209 
   210 fun repair_constr_mcache dtype_cache constr_mcache =
   211   let
   212     fun repair_one (x, M) =
   213       Unsynchronized.change constr_mcache
   214           (AList.update (op =) (x, repair_mtype dtype_cache [] M))
   215   in List.app repair_one (!constr_mcache) end
   216 
   217 fun is_fin_fun_supported_type @{typ prop} = true
   218   | is_fin_fun_supported_type @{typ bool} = true
   219   | is_fin_fun_supported_type (Type (@{type_name option}, _)) = true
   220   | is_fin_fun_supported_type _ = false
   221 fun fin_fun_body _ _ (t as @{term False}) = SOME t
   222   | fin_fun_body _ _ (t as Const (@{const_name None}, _)) = SOME t
   223   | fin_fun_body dom_T ran_T
   224                  ((t0 as Const (@{const_name If}, _))
   225                   $ (t1 as Const (@{const_name "op ="}, _) $ Bound 0 $ t1')
   226                   $ t2 $ t3) =
   227     (if loose_bvar1 (t1', 0) then
   228        NONE
   229      else case fin_fun_body dom_T ran_T t3 of
   230        NONE => NONE
   231      | SOME t3 =>
   232        SOME (t0 $ (Const (@{const_name is_unknown}, dom_T --> bool_T) $ t1')
   233                 $ (Const (@{const_name unknown}, ran_T)) $ (t0 $ t1 $ t2 $ t3)))
   234   | fin_fun_body _ _ _ = NONE
   235 
   236 fun fresh_mfun_for_fun_type (mdata as {max_fresh, ...} : mdata) all_minus
   237                             T1 T2 =
   238   let
   239     val M1 = fresh_mtype_for_type mdata all_minus T1
   240     val M2 = fresh_mtype_for_type mdata all_minus T2
   241     val a = if not all_minus andalso exists_alpha_sub_mtype_fresh M1 andalso
   242                is_fin_fun_supported_type (body_type T2) then
   243               V (Unsynchronized.inc max_fresh)
   244             else
   245               S Minus
   246   in (M1, a, M2) end
   247 and fresh_mtype_for_type (mdata as {hol_ctxt as {ctxt, ...}, binarize, alpha_T,
   248                                     datatype_mcache, constr_mcache, ...})
   249                          all_minus =
   250   let
   251     fun do_type T =
   252       if T = alpha_T then
   253         MAlpha
   254       else case T of
   255         Type (@{type_name fun}, [T1, T2]) =>
   256         MFun (fresh_mfun_for_fun_type mdata false T1 T2)
   257       | Type (@{type_name Product_Type.prod}, [T1, T2]) => MPair (pairself do_type (T1, T2))
   258       | Type (z as (s, _)) =>
   259         if could_exist_alpha_sub_mtype ctxt alpha_T T then
   260           case AList.lookup (op =) (!datatype_mcache) z of
   261             SOME M => M
   262           | NONE =>
   263             let
   264               val _ = Unsynchronized.change datatype_mcache (cons (z, MRec z))
   265               val xs = binarized_and_boxed_datatype_constrs hol_ctxt binarize T
   266               val (all_Ms, constr_Ms) =
   267                 fold_rev (fn (_, T') => fn (all_Ms, constr_Ms) =>
   268                              let
   269                                val binder_Ms = map do_type (binder_types T')
   270                                val new_Ms = filter exists_alpha_sub_mtype_fresh
   271                                                    binder_Ms
   272                                val constr_M = constr_mtype_for_binders z
   273                                                                        binder_Ms
   274                              in
   275                                (union (op =) new_Ms all_Ms,
   276                                 constr_M :: constr_Ms)
   277                              end)
   278                          xs ([], [])
   279               val M = MType (s, all_Ms)
   280               val _ = Unsynchronized.change datatype_mcache
   281                           (AList.update (op =) (z, M))
   282               val _ = Unsynchronized.change constr_mcache
   283                           (append (xs ~~ constr_Ms))
   284             in
   285               if forall (not o is_MRec o snd) (!datatype_mcache) then
   286                 (repair_datatype_mcache datatype_mcache;
   287                  repair_constr_mcache (!datatype_mcache) constr_mcache;
   288                  AList.lookup (op =) (!datatype_mcache) z |> the)
   289               else
   290                 M
   291             end
   292         else
   293           MType (s, [])
   294       | _ => MType (simple_string_of_typ T, [])
   295   in do_type end
   296 
   297 fun prodM_factors (MPair (M1, M2)) = maps prodM_factors [M1, M2]
   298   | prodM_factors M = [M]
   299 fun curried_strip_mtype (MFun (M1, _, M2)) =
   300     curried_strip_mtype M2 |>> append (prodM_factors M1)
   301   | curried_strip_mtype M = ([], M)
   302 fun sel_mtype_from_constr_mtype s M =
   303   let val (arg_Ms, dataM) = curried_strip_mtype M in
   304     MFun (dataM, S Minus,
   305           case sel_no_from_name s of ~1 => bool_M | n => nth arg_Ms n)
   306   end
   307 
   308 fun mtype_for_constr (mdata as {hol_ctxt = {ctxt, ...}, alpha_T, constr_mcache,
   309                                 ...}) (x as (_, T)) =
   310   if could_exist_alpha_sub_mtype ctxt alpha_T T then
   311     case AList.lookup (op =) (!constr_mcache) x of
   312       SOME M => M
   313     | NONE => if T = alpha_T then
   314                 let val M = fresh_mtype_for_type mdata false T in
   315                   (Unsynchronized.change constr_mcache (cons (x, M)); M)
   316                 end
   317               else
   318                 (fresh_mtype_for_type mdata false (body_type T);
   319                  AList.lookup (op =) (!constr_mcache) x |> the)
   320   else
   321     fresh_mtype_for_type mdata false T
   322 fun mtype_for_sel (mdata as {hol_ctxt, binarize, ...}) (x as (s, _)) =
   323   x |> binarized_and_boxed_constr_for_sel hol_ctxt binarize
   324     |> mtype_for_constr mdata |> sel_mtype_from_constr_mtype s
   325 
   326 fun resolve_sign_atom lits (V x) =
   327     x |> AList.lookup (op =) lits |> Option.map S |> the_default (V x)
   328   | resolve_sign_atom _ a = a
   329 fun resolve_mtype lits =
   330   let
   331     fun aux MAlpha = MAlpha
   332       | aux (MFun (M1, a, M2)) = MFun (aux M1, resolve_sign_atom lits a, aux M2)
   333       | aux (MPair Mp) = MPair (pairself aux Mp)
   334       | aux (MType (s, Ms)) = MType (s, map aux Ms)
   335       | aux (MRec z) = MRec z
   336   in aux end
   337 
   338 datatype comp_op = Eq | Leq
   339 
   340 type comp = sign_atom * sign_atom * comp_op * var list
   341 type sign_expr = literal list
   342 
   343 type constraint_set = literal list * comp list * sign_expr list
   344 
   345 fun string_for_comp_op Eq = "="
   346   | string_for_comp_op Leq = "\<le>"
   347 
   348 fun string_for_sign_expr [] = "\<bot>"
   349   | string_for_sign_expr lits =
   350     space_implode " \<or> " (map string_for_literal lits)
   351 
   352 fun do_literal _ NONE = NONE
   353   | do_literal (x, sn) (SOME lits) =
   354     case AList.lookup (op =) lits x of
   355       SOME sn' => if sn = sn' then SOME lits else NONE
   356     | NONE => SOME ((x, sn) :: lits)
   357 
   358 fun do_sign_atom_comp Eq [] a1 a2 (accum as (lits, comps)) =
   359     (case (a1, a2) of
   360        (S sn1, S sn2) => if sn1 = sn2 then SOME accum else NONE
   361      | (V x1, S sn2) =>
   362        Option.map (rpair comps) (do_literal (x1, sn2) (SOME lits))
   363      | (V _, V _) => SOME (lits, insert (op =) (a1, a2, Eq, []) comps)
   364      | _ => do_sign_atom_comp Eq [] a2 a1 accum)
   365   | do_sign_atom_comp Leq [] a1 a2 (accum as (lits, comps)) =
   366     (case (a1, a2) of
   367        (_, S Minus) => SOME accum
   368      | (S Plus, _) => SOME accum
   369      | (S Minus, S Plus) => NONE
   370      | (V _, V _) => SOME (lits, insert (op =) (a1, a2, Leq, []) comps)
   371      | _ => do_sign_atom_comp Eq [] a1 a2 accum)
   372   | do_sign_atom_comp cmp xs a1 a2 (lits, comps) =
   373     SOME (lits, insert (op =) (a1, a2, cmp, xs) comps)
   374 
   375 fun do_mtype_comp _ _ _ _ NONE = NONE
   376   | do_mtype_comp _ _ MAlpha MAlpha accum = accum
   377   | do_mtype_comp Eq xs (MFun (M11, a1, M12)) (MFun (M21, a2, M22))
   378                   (SOME accum) =
   379      accum |> do_sign_atom_comp Eq xs a1 a2 |> do_mtype_comp Eq xs M11 M21
   380            |> do_mtype_comp Eq xs M12 M22
   381   | do_mtype_comp Leq xs (MFun (M11, a1, M12)) (MFun (M21, a2, M22))
   382                   (SOME accum) =
   383     (if exists_alpha_sub_mtype M11 then
   384        accum |> do_sign_atom_comp Leq xs a1 a2
   385              |> do_mtype_comp Leq xs M21 M11
   386              |> (case a2 of
   387                    S Minus => I
   388                  | S Plus => do_mtype_comp Leq xs M11 M21
   389                  | V x => do_mtype_comp Leq (x :: xs) M11 M21)
   390      else
   391        SOME accum)
   392     |> do_mtype_comp Leq xs M12 M22
   393   | do_mtype_comp cmp xs (M1 as MPair (M11, M12)) (M2 as MPair (M21, M22))
   394                   accum =
   395     (accum |> fold (uncurry (do_mtype_comp cmp xs)) [(M11, M21), (M12, M22)]
   396      handle Library.UnequalLengths =>
   397             raise MTYPE ("Nitpick_Mono.do_mtype_comp", [M1, M2], []))
   398   | do_mtype_comp _ _ (MType _) (MType _) accum =
   399     accum (* no need to compare them thanks to the cache *)
   400   | do_mtype_comp cmp _ M1 M2 _ =
   401     raise MTYPE ("Nitpick_Mono.do_mtype_comp (" ^ string_for_comp_op cmp ^ ")",
   402                  [M1, M2], [])
   403 
   404 fun add_mtype_comp cmp M1 M2 ((lits, comps, sexps) : constraint_set) =
   405     (print_g (fn () => "*** Add " ^ string_for_mtype M1 ^ " " ^
   406                        string_for_comp_op cmp ^ " " ^ string_for_mtype M2);
   407      case do_mtype_comp cmp [] M1 M2 (SOME (lits, comps)) of
   408        NONE => (print_g (K "**** Unsolvable"); raise UNSOLVABLE ())
   409      | SOME (lits, comps) => (lits, comps, sexps))
   410 
   411 val add_mtypes_equal = add_mtype_comp Eq
   412 val add_is_sub_mtype = add_mtype_comp Leq
   413 
   414 fun do_notin_mtype_fv _ _ _ NONE = NONE
   415   | do_notin_mtype_fv Minus _ MAlpha accum = accum
   416   | do_notin_mtype_fv Plus [] MAlpha _ = NONE
   417   | do_notin_mtype_fv Plus [(x, sn)] MAlpha (SOME (lits, sexps)) =
   418     SOME lits |> do_literal (x, sn) |> Option.map (rpair sexps)
   419   | do_notin_mtype_fv Plus sexp MAlpha (SOME (lits, sexps)) =
   420     SOME (lits, insert (op =) sexp sexps)
   421   | do_notin_mtype_fv sn sexp (MFun (M1, S sn', M2)) accum =
   422     accum |> (if sn' = Plus andalso sn = Plus then
   423                 do_notin_mtype_fv Plus sexp M1
   424               else
   425                 I)
   426           |> (if sn' = Minus orelse sn = Plus then
   427                 do_notin_mtype_fv Minus sexp M1
   428               else
   429                 I)
   430           |> do_notin_mtype_fv sn sexp M2
   431   | do_notin_mtype_fv Plus sexp (MFun (M1, V x, M2)) accum =
   432     accum |> (case do_literal (x, Minus) (SOME sexp) of
   433                 NONE => I
   434               | SOME sexp' => do_notin_mtype_fv Plus sexp' M1)
   435           |> do_notin_mtype_fv Minus sexp M1
   436           |> do_notin_mtype_fv Plus sexp M2
   437   | do_notin_mtype_fv Minus sexp (MFun (M1, V x, M2)) accum =
   438     accum |> (case do_literal (x, Plus) (SOME sexp) of
   439                 NONE => I
   440               | SOME sexp' => do_notin_mtype_fv Plus sexp' M1)
   441           |> do_notin_mtype_fv Minus sexp M2
   442   | do_notin_mtype_fv sn sexp (MPair (M1, M2)) accum =
   443     accum |> fold (do_notin_mtype_fv sn sexp) [M1, M2]
   444   | do_notin_mtype_fv sn sexp (MType (_, Ms)) accum =
   445     accum |> fold (do_notin_mtype_fv sn sexp) Ms
   446   | do_notin_mtype_fv _ _ M _ =
   447     raise MTYPE ("Nitpick_Mono.do_notin_mtype_fv", [M], [])
   448 
   449 fun add_notin_mtype_fv sn M ((lits, comps, sexps) : constraint_set) =
   450     (print_g (fn () => "*** Add " ^ string_for_mtype M ^ " is " ^
   451                        (case sn of Minus => "concrete" | Plus => "complete") ^
   452                        ".");
   453      case do_notin_mtype_fv sn [] M (SOME (lits, sexps)) of
   454        NONE => (print_g (K "**** Unsolvable"); raise UNSOLVABLE ())
   455      | SOME (lits, sexps) => (lits, comps, sexps))
   456 
   457 val add_mtype_is_concrete = add_notin_mtype_fv Minus
   458 val add_mtype_is_complete = add_notin_mtype_fv Plus
   459 
   460 val bool_from_minus = true
   461 
   462 fun bool_from_sign Plus = not bool_from_minus
   463   | bool_from_sign Minus = bool_from_minus
   464 fun sign_from_bool b = if b = bool_from_minus then Minus else Plus
   465 
   466 fun prop_for_literal (x, sn) =
   467   (not (bool_from_sign sn) ? PropLogic.Not) (PropLogic.BoolVar x)
   468 fun prop_for_sign_atom_eq (S sn', sn) =
   469     if sn = sn' then PropLogic.True else PropLogic.False
   470   | prop_for_sign_atom_eq (V x, sn) = prop_for_literal (x, sn)
   471 fun prop_for_sign_expr xs = PropLogic.exists (map prop_for_literal xs)
   472 fun prop_for_exists_eq xs sn =
   473   PropLogic.exists (map (fn x => prop_for_literal (x, sn)) xs)
   474 fun prop_for_comp (a1, a2, Eq, []) =
   475     PropLogic.SAnd (prop_for_comp (a1, a2, Leq, []),
   476                     prop_for_comp (a2, a1, Leq, []))
   477   | prop_for_comp (a1, a2, Leq, []) =
   478     PropLogic.SOr (prop_for_sign_atom_eq (a1, Plus),
   479                    prop_for_sign_atom_eq (a2, Minus))
   480   | prop_for_comp (a1, a2, cmp, xs) =
   481     PropLogic.SOr (prop_for_exists_eq xs Minus, prop_for_comp (a1, a2, cmp, []))
   482 
   483 fun literals_from_assignments max_var assigns lits =
   484   fold (fn x => fn accum =>
   485            if AList.defined (op =) lits x then
   486              accum
   487            else case assigns x of
   488              SOME b => (x, sign_from_bool b) :: accum
   489            | NONE => accum) (max_var downto 1) lits
   490 
   491 fun string_for_comp (a1, a2, cmp, xs) =
   492   string_for_sign_atom a1 ^ " " ^ string_for_comp_op cmp ^
   493   subscript_string_for_vars " \<and> " xs ^ " " ^ string_for_sign_atom a2
   494 
   495 fun print_problem lits comps sexps =
   496   print_g (fn () => "*** Problem:\n" ^
   497                     cat_lines (map string_for_literal lits @
   498                                map string_for_comp comps @
   499                                map string_for_sign_expr sexps))
   500 
   501 fun print_solution lits =
   502   let val (pos, neg) = List.partition (curry (op =) Plus o snd) lits in
   503     print_g (fn () => "*** Solution:\n" ^
   504                       "+: " ^ commas (map (string_for_var o fst) pos) ^ "\n" ^
   505                       "-: " ^ commas (map (string_for_var o fst) neg))
   506   end
   507 
   508 fun solve max_var (lits, comps, sexps) =
   509     let
   510       fun do_assigns assigns =
   511         SOME (literals_from_assignments max_var assigns lits
   512               |> tap print_solution)
   513       val _ = print_problem lits comps sexps
   514       val prop = PropLogic.all (map prop_for_literal lits @
   515                                 map prop_for_comp comps @
   516                                 map prop_for_sign_expr sexps)
   517       val default_val = bool_from_sign Minus
   518     in
   519       if PropLogic.eval (K default_val) prop then
   520         do_assigns (K (SOME default_val))
   521       else
   522         let
   523           (* use the first ML solver (to avoid startup overhead) *)
   524           val solvers = !SatSolver.solvers
   525                         |> filter (member (op =) ["dptsat", "dpll"] o fst)
   526         in
   527           case snd (hd solvers) prop of
   528             SatSolver.SATISFIABLE assigns => do_assigns assigns
   529           | _ => NONE
   530         end
   531     end
   532 
   533 type mtype_schema = mtyp * constraint_set
   534 type mtype_context =
   535   {bound_Ts: typ list,
   536    bound_Ms: mtyp list,
   537    frees: (styp * mtyp) list,
   538    consts: (styp * mtyp) list}
   539 
   540 type accumulator = mtype_context * constraint_set
   541 
   542 val initial_gamma = {bound_Ts = [], bound_Ms = [], frees = [], consts = []}
   543 
   544 fun push_bound T M {bound_Ts, bound_Ms, frees, consts} =
   545   {bound_Ts = T :: bound_Ts, bound_Ms = M :: bound_Ms, frees = frees,
   546    consts = consts}
   547 fun pop_bound {bound_Ts, bound_Ms, frees, consts} =
   548   {bound_Ts = tl bound_Ts, bound_Ms = tl bound_Ms, frees = frees,
   549    consts = consts}
   550   handle List.Empty => initial_gamma (* FIXME: needed? *)
   551 
   552 fun consider_term (mdata as {hol_ctxt as {thy, ctxt, stds, fast_descrs,
   553                                           def_table, ...},
   554                              alpha_T, max_fresh, ...}) =
   555   let
   556     fun is_enough_eta_expanded t =
   557       case strip_comb t of
   558         (Const x, ts) =>
   559         the_default 0 (arity_of_built_in_const thy stds fast_descrs x)
   560         <= length ts
   561       | _ => true
   562     val mtype_for = fresh_mtype_for_type mdata false
   563     fun plus_set_mtype_for_dom M =
   564       MFun (M, S (if exists_alpha_sub_mtype M then Plus else Minus), bool_M)
   565     fun do_all T (gamma, cset) =
   566       let
   567         val abs_M = mtype_for (domain_type (domain_type T))
   568         val body_M = mtype_for (body_type T)
   569       in
   570         (MFun (MFun (abs_M, S Minus, body_M), S Minus, body_M),
   571          (gamma, cset |> add_mtype_is_complete abs_M))
   572       end
   573     fun do_equals T (gamma, cset) =
   574       let val M = mtype_for (domain_type T) in
   575         (MFun (M, S Minus, MFun (M, V (Unsynchronized.inc max_fresh),
   576                                  mtype_for (nth_range_type 2 T))),
   577          (gamma, cset |> add_mtype_is_concrete M))
   578       end
   579     fun do_robust_set_operation T (gamma, cset) =
   580       let
   581         val set_T = domain_type T
   582         val M1 = mtype_for set_T
   583         val M2 = mtype_for set_T
   584         val M3 = mtype_for set_T
   585       in
   586         (MFun (M1, S Minus, MFun (M2, S Minus, M3)),
   587          (gamma, cset |> add_is_sub_mtype M1 M3 |> add_is_sub_mtype M2 M3))
   588       end
   589     fun do_fragile_set_operation T (gamma, cset) =
   590       let
   591         val set_T = domain_type T
   592         val set_M = mtype_for set_T
   593         fun custom_mtype_for (T as Type (@{type_name fun}, [T1, T2])) =
   594             if T = set_T then set_M
   595             else MFun (custom_mtype_for T1, S Minus, custom_mtype_for T2)
   596           | custom_mtype_for T = mtype_for T
   597       in
   598         (custom_mtype_for T, (gamma, cset |> add_mtype_is_concrete set_M))
   599       end
   600     fun do_pair_constr T accum =
   601       case mtype_for (nth_range_type 2 T) of
   602         M as MPair (a_M, b_M) =>
   603         (MFun (a_M, S Minus, MFun (b_M, S Minus, M)), accum)
   604       | M => raise MTYPE ("Nitpick_Mono.consider_term.do_pair_constr", [M], [])
   605     fun do_nth_pair_sel n T =
   606       case mtype_for (domain_type T) of
   607         M as MPair (a_M, b_M) =>
   608         pair (MFun (M, S Minus, if n = 0 then a_M else b_M))
   609       | M => raise MTYPE ("Nitpick_Mono.consider_term.do_nth_pair_sel", [M], [])
   610     fun do_bounded_quantifier t0 abs_s abs_T connective_t bound_t body_t accum =
   611       let
   612         val abs_M = mtype_for abs_T
   613         val (bound_m, accum) =
   614           accum |>> push_bound abs_T abs_M |> do_term bound_t
   615         val expected_bound_M = plus_set_mtype_for_dom abs_M
   616         val (body_m, accum) =
   617           accum ||> add_mtypes_equal expected_bound_M (mtype_of_mterm bound_m)
   618                 |> do_term body_t ||> apfst pop_bound
   619         val bound_M = mtype_of_mterm bound_m
   620         val (M1, a, M2) = dest_MFun bound_M
   621       in
   622         (MApp (MRaw (t0, MFun (bound_M, S Minus, bool_M)),
   623                MAbs (abs_s, abs_T, M1, a,
   624                      MApp (MApp (MRaw (connective_t,
   625                                        mtype_for (fastype_of connective_t)),
   626                                  MApp (bound_m, MRaw (Bound 0, M1))),
   627                            body_m))), accum)
   628       end
   629     and do_term t (accum as (gamma as {bound_Ts, bound_Ms, frees, consts},
   630                              cset)) =
   631         (print_g (fn () => "  \<Gamma> \<turnstile> " ^
   632                            Syntax.string_of_term ctxt t ^ " : _?");
   633          case t of
   634            Const (x as (s, T)) =>
   635            (case AList.lookup (op =) consts x of
   636               SOME M => (M, accum)
   637             | NONE =>
   638               if not (could_exist_alpha_subtype alpha_T T) then
   639                 (mtype_for T, accum)
   640               else case s of
   641                 @{const_name all} => do_all T accum
   642               | @{const_name "=="} => do_equals T accum
   643               | @{const_name All} => do_all T accum
   644               | @{const_name Ex} =>
   645                 let val set_T = domain_type T in
   646                   do_term (Abs (Name.uu, set_T,
   647                                 @{const Not} $ (HOLogic.mk_eq
   648                                     (Abs (Name.uu, domain_type set_T,
   649                                           @{const False}),
   650                                      Bound 0)))) accum
   651                   |>> mtype_of_mterm
   652                 end
   653               | @{const_name "op ="} => do_equals T accum
   654               | @{const_name The} =>
   655                 (print_g (K "*** The"); raise UNSOLVABLE ())
   656               | @{const_name Eps} =>
   657                 (print_g (K "*** Eps"); raise UNSOLVABLE ())
   658               | @{const_name If} =>
   659                 do_robust_set_operation (range_type T) accum
   660                 |>> curry3 MFun bool_M (S Minus)
   661               | @{const_name Pair} => do_pair_constr T accum
   662               | @{const_name fst} => do_nth_pair_sel 0 T accum
   663               | @{const_name snd} => do_nth_pair_sel 1 T accum 
   664               | @{const_name Id} =>
   665                 (MFun (mtype_for (domain_type T), S Minus, bool_M), accum)
   666               | @{const_name converse} =>
   667                 let
   668                   val x = Unsynchronized.inc max_fresh
   669                   fun mtype_for_set T =
   670                     MFun (mtype_for (domain_type T), V x, bool_M)
   671                   val ab_set_M = domain_type T |> mtype_for_set
   672                   val ba_set_M = range_type T |> mtype_for_set
   673                 in (MFun (ab_set_M, S Minus, ba_set_M), accum) end
   674               | @{const_name trancl} => do_fragile_set_operation T accum
   675               | @{const_name rel_comp} =>
   676                 let
   677                   val x = Unsynchronized.inc max_fresh
   678                   fun mtype_for_set T =
   679                     MFun (mtype_for (domain_type T), V x, bool_M)
   680                   val bc_set_M = domain_type T |> mtype_for_set
   681                   val ab_set_M = domain_type (range_type T) |> mtype_for_set
   682                   val ac_set_M = nth_range_type 2 T |> mtype_for_set
   683                 in
   684                   (MFun (bc_set_M, S Minus, MFun (ab_set_M, S Minus, ac_set_M)),
   685                    accum)
   686                 end
   687               | @{const_name image} =>
   688                 let
   689                   val a_M = mtype_for (domain_type (domain_type T))
   690                   val b_M = mtype_for (range_type (domain_type T))
   691                 in
   692                   (MFun (MFun (a_M, S Minus, b_M), S Minus,
   693                          MFun (plus_set_mtype_for_dom a_M, S Minus,
   694                                plus_set_mtype_for_dom b_M)), accum)
   695                 end
   696               | @{const_name finite} =>
   697                 let val M1 = mtype_for (domain_type (domain_type T)) in
   698                   (MFun (plus_set_mtype_for_dom M1, S Minus, bool_M), accum)
   699                 end
   700               | @{const_name Sigma} =>
   701                 let
   702                   val x = Unsynchronized.inc max_fresh
   703                   fun mtype_for_set T =
   704                     MFun (mtype_for (domain_type T), V x, bool_M)
   705                   val a_set_T = domain_type T
   706                   val a_M = mtype_for (domain_type a_set_T)
   707                   val b_set_M = mtype_for_set (range_type (domain_type
   708                                                                (range_type T)))
   709                   val a_set_M = mtype_for_set a_set_T
   710                   val a_to_b_set_M = MFun (a_M, S Minus, b_set_M)
   711                   val ab_set_M = mtype_for_set (nth_range_type 2 T)
   712                 in
   713                   (MFun (a_set_M, S Minus,
   714                          MFun (a_to_b_set_M, S Minus, ab_set_M)), accum)
   715                 end
   716               | _ =>
   717                 if s = @{const_name safe_The} orelse
   718                    s = @{const_name safe_Eps} then
   719                   let
   720                     val a_set_M = mtype_for (domain_type T)
   721                     val a_M = dest_MFun a_set_M |> #1
   722                   in (MFun (a_set_M, S Minus, a_M), accum) end
   723                 else if s = @{const_name ord_class.less_eq} andalso
   724                         is_set_type (domain_type T) then
   725                   do_fragile_set_operation T accum
   726                 else if is_sel s then
   727                   (mtype_for_sel mdata x, accum)
   728                 else if is_constr ctxt stds x then
   729                   (mtype_for_constr mdata x, accum)
   730                 else if is_built_in_const thy stds fast_descrs x then
   731                   (fresh_mtype_for_type mdata true T, accum)
   732                 else
   733                   let val M = mtype_for T in
   734                     (M, ({bound_Ts = bound_Ts, bound_Ms = bound_Ms,
   735                           frees = frees, consts = (x, M) :: consts}, cset))
   736                   end) |>> curry MRaw t
   737          | Free (x as (_, T)) =>
   738            (case AList.lookup (op =) frees x of
   739               SOME M => (M, accum)
   740             | NONE =>
   741               let val M = mtype_for T in
   742                 (M, ({bound_Ts = bound_Ts, bound_Ms = bound_Ms,
   743                       frees = (x, M) :: frees, consts = consts}, cset))
   744               end) |>> curry MRaw t
   745          | Var _ => (print_g (K "*** Var"); raise UNSOLVABLE ())
   746          | Bound j => (MRaw (t, nth bound_Ms j), accum)
   747          | Abs (s, T, t') =>
   748            (case fin_fun_body T (fastype_of1 (T :: bound_Ts, t')) t' of
   749               SOME t' =>
   750               let
   751                 val M = mtype_for T
   752                 val a = V (Unsynchronized.inc max_fresh)
   753                 val (m', accum) = do_term t' (accum |>> push_bound T M)
   754               in (MAbs (s, T, M, a, m'), accum |>> pop_bound) end
   755             | NONE =>
   756               ((case t' of
   757                   t1' $ Bound 0 =>
   758                   if not (loose_bvar1 (t1', 0)) andalso
   759                      is_enough_eta_expanded t1' then
   760                     do_term (incr_boundvars ~1 t1') accum
   761                   else
   762                     raise SAME ()
   763                 | (t11 as Const (@{const_name "op ="}, _)) $ Bound 0 $ t13 =>
   764                   if not (loose_bvar1 (t13, 0)) then
   765                     do_term (incr_boundvars ~1 (t11 $ t13)) accum
   766                   else
   767                     raise SAME ()
   768                 | _ => raise SAME ())
   769                handle SAME () =>
   770                       let
   771                         val M = mtype_for T
   772                         val (m', accum) = do_term t' (accum |>> push_bound T M)
   773                       in
   774                         (MAbs (s, T, M, S Minus, m'), accum |>> pop_bound)
   775                       end))
   776          | (t0 as Const (@{const_name All}, _))
   777            $ Abs (s', T', (t10 as @{const "op -->"}) $ (t11 $ Bound 0) $ t12) =>
   778            do_bounded_quantifier t0 s' T' t10 t11 t12 accum
   779          | (t0 as Const (@{const_name Ex}, _))
   780            $ Abs (s', T', (t10 as @{const "op &"}) $ (t11 $ Bound 0) $ t12) =>
   781            do_bounded_quantifier t0 s' T' t10 t11 t12 accum
   782          | Const (@{const_name Let}, _) $ t1 $ t2 =>
   783            do_term (betapply (t2, t1)) accum
   784          | t1 $ t2 =>
   785            let
   786              val (m1, accum) = do_term t1 accum
   787              val (m2, accum) = do_term t2 accum
   788            in
   789              let
   790                val T11 = domain_type (fastype_of1 (bound_Ts, t1))
   791                val T2 = fastype_of1 (bound_Ts, t2)
   792                val M11 = mtype_of_mterm m1 |> dest_MFun |> #1
   793                val M2 = mtype_of_mterm m2
   794              in (MApp (m1, m2), accum ||> add_is_sub_mtype M2 M11) end
   795            end)
   796         |> tap (fn (m, _) => print_g (fn () => "  \<Gamma> \<turnstile> " ^
   797                                                string_for_mterm ctxt m))
   798   in do_term end
   799 
   800 fun force_minus_funs 0 _ = I
   801   | force_minus_funs n (M as MFun (M1, _, M2)) =
   802     add_mtypes_equal M (MFun (M1, S Minus, M2))
   803     #> force_minus_funs (n - 1) M2
   804   | force_minus_funs _ M =
   805     raise MTYPE ("Nitpick_Mono.force_minus_funs", [M], [])
   806 fun consider_general_equals mdata def (x as (_, T)) t1 t2 accum =
   807   let
   808     val (m1, accum) = consider_term mdata t1 accum
   809     val (m2, accum) = consider_term mdata t2 accum
   810     val M1 = mtype_of_mterm m1
   811     val M2 = mtype_of_mterm m2
   812     val accum = accum ||> add_mtypes_equal M1 M2
   813     val body_M = fresh_mtype_for_type mdata false (nth_range_type 2 T)
   814     val m = MApp (MApp (MRaw (Const x,
   815                 MFun (M1, S Minus, MFun (M2, S Minus, body_M))), m1), m2)
   816   in
   817     (m, if def then
   818           let val (head_m, arg_ms) = strip_mcomb m1 in
   819             accum ||> force_minus_funs (length arg_ms) (mtype_of_mterm head_m)
   820           end
   821         else
   822           accum)
   823   end
   824 
   825 fun consider_general_formula (mdata as {hol_ctxt = {ctxt, ...}, ...}) =
   826   let
   827     val mtype_for = fresh_mtype_for_type mdata false
   828     val do_term = consider_term mdata
   829     fun do_formula sn t accum =
   830         let
   831           fun do_quantifier (quant_x as (quant_s, _)) abs_s abs_T body_t =
   832             let
   833               val abs_M = mtype_for abs_T
   834               val side_cond = ((sn = Minus) = (quant_s = @{const_name Ex}))
   835               val (body_m, accum) =
   836                 accum ||> side_cond ? add_mtype_is_complete abs_M
   837                       |>> push_bound abs_T abs_M |> do_formula sn body_t
   838               val body_M = mtype_of_mterm body_m
   839             in
   840               (MApp (MRaw (Const quant_x,
   841                            MFun (MFun (abs_M, S Minus, body_M), S Minus,
   842                                  body_M)),
   843                      MAbs (abs_s, abs_T, abs_M, S Minus, body_m)),
   844                accum |>> pop_bound)
   845             end
   846           fun do_equals x t1 t2 =
   847             case sn of
   848               Plus => do_term t accum
   849             | Minus => consider_general_equals mdata false x t1 t2 accum
   850         in
   851           (print_g (fn () => "  \<Gamma> \<turnstile> " ^
   852                            Syntax.string_of_term ctxt t ^ " : o\<^sup>" ^
   853                            string_for_sign sn ^ "?");
   854            case t of
   855              Const (x as (@{const_name all}, _)) $ Abs (s1, T1, t1) =>
   856              do_quantifier x s1 T1 t1
   857            | Const (x as (@{const_name "=="}, _)) $ t1 $ t2 => do_equals x t1 t2
   858            | @{const Trueprop} $ t1 =>
   859              let val (m1, accum) = do_formula sn t1 accum in
   860                (MApp (MRaw (@{const Trueprop}, mtype_for (bool_T --> prop_T)),
   861                       m1), accum)
   862              end
   863            | @{const Not} $ t1 =>
   864              let val (m1, accum) = do_formula (negate sn) t1 accum in
   865                (MApp (MRaw (@{const Not}, mtype_for (bool_T --> bool_T)), m1),
   866                 accum)
   867              end
   868            | Const (x as (@{const_name All}, _)) $ Abs (s1, T1, t1) =>
   869              do_quantifier x s1 T1 t1
   870            | Const (x0 as (s0 as @{const_name Ex}, T0))
   871              $ (t1 as Abs (s1, T1, t1')) =>
   872              (case sn of
   873                 Plus => do_quantifier x0 s1 T1 t1'
   874               | Minus =>
   875                 (* FIXME: Move elsewhere *)
   876                 do_term (@{const Not}
   877                          $ (HOLogic.eq_const (domain_type T0) $ t1
   878                             $ Abs (Name.uu, T1, @{const False}))) accum)
   879            | Const (x as (@{const_name "op ="}, _)) $ t1 $ t2 =>
   880              do_equals x t1 t2
   881            | Const (@{const_name Let}, _) $ t1 $ t2 =>
   882              do_formula sn (betapply (t2, t1)) accum
   883            | (t0 as Const (s0, _)) $ t1 $ t2 =>
   884              if s0 = @{const_name "==>"} orelse s0 = @{const_name "op &"} orelse
   885                 s0 = @{const_name "op |"} orelse
   886                 s0 = @{const_name "op -->"} then
   887                let
   888                  val impl = (s0 = @{const_name "==>"} orelse
   889                              s0 = @{const_name "op -->"})
   890                  val (m1, accum) = do_formula (sn |> impl ? negate) t1 accum
   891                  val (m2, accum) = do_formula sn t2 accum
   892                in
   893                  (MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2),
   894                   accum)
   895                end 
   896              else
   897                do_term t accum
   898            | _ => do_term t accum)
   899         end
   900         |> tap (fn (m, _) =>
   901                    print_g (fn () => "\<Gamma> \<turnstile> " ^
   902                                      string_for_mterm ctxt m ^ " : o\<^sup>" ^
   903                                      string_for_sign sn))
   904   in do_formula end
   905 
   906 (* The harmless axiom optimization below is somewhat too aggressive in the face
   907    of (rather peculiar) user-defined axioms. *)
   908 val harmless_consts =
   909   [@{const_name ord_class.less}, @{const_name ord_class.less_eq}]
   910 val bounteous_consts = [@{const_name bisim}]
   911 
   912 fun is_harmless_axiom ({no_harmless = true, ...} : mdata) _ = false
   913   | is_harmless_axiom {hol_ctxt = {thy, stds, fast_descrs, ...}, ...} t =
   914     Term.add_consts t []
   915     |> filter_out (is_built_in_const thy stds fast_descrs)
   916     |> (forall (member (op =) harmless_consts o original_name o fst) orf
   917         exists (member (op =) bounteous_consts o fst))
   918 
   919 fun consider_nondefinitional_axiom mdata t =
   920   if is_harmless_axiom mdata t then pair (MRaw (t, dummy_M))
   921   else consider_general_formula mdata Plus t
   922 
   923 fun consider_definitional_axiom (mdata as {hol_ctxt = {ctxt, ...}, ...}) t =
   924   if not (is_constr_pattern_formula ctxt t) then
   925     consider_nondefinitional_axiom mdata t
   926   else if is_harmless_axiom mdata t then
   927     pair (MRaw (t, dummy_M))
   928   else
   929     let
   930       val mtype_for = fresh_mtype_for_type mdata false
   931       val do_term = consider_term mdata
   932       fun do_all quant_t abs_s abs_T body_t accum =
   933         let
   934           val abs_M = mtype_for abs_T
   935           val (body_m, accum) =
   936             accum |>> push_bound abs_T abs_M |> do_formula body_t
   937           val body_M = mtype_of_mterm body_m
   938         in
   939           (MApp (MRaw (quant_t,
   940                        MFun (MFun (abs_M, S Minus, body_M), S Minus, body_M)),
   941                  MAbs (abs_s, abs_T, abs_M, S Minus, body_m)),
   942            accum |>> pop_bound)
   943         end
   944       and do_conjunction t0 t1 t2 accum =
   945         let
   946           val (m1, accum) = do_formula t1 accum
   947           val (m2, accum) = do_formula t2 accum
   948         in
   949           (MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2), accum)
   950         end
   951       and do_implies t0 t1 t2 accum =
   952         let
   953           val (m1, accum) = do_term t1 accum
   954           val (m2, accum) = do_formula t2 accum
   955         in
   956           (MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2), accum)
   957         end
   958       and do_formula t accum =
   959           case t of
   960             (t0 as Const (@{const_name all}, _)) $ Abs (s1, T1, t1) =>
   961             do_all t0 s1 T1 t1 accum
   962           | @{const Trueprop} $ t1 =>
   963             let val (m1, accum) = do_formula t1 accum in
   964               (MApp (MRaw (@{const Trueprop}, mtype_for (bool_T --> prop_T)),
   965                      m1), accum)
   966             end
   967           | Const (x as (@{const_name "=="}, _)) $ t1 $ t2 =>
   968             consider_general_equals mdata true x t1 t2 accum
   969           | (t0 as @{const "==>"}) $ t1 $ t2 => do_implies t0 t1 t2 accum
   970           | (t0 as @{const Pure.conjunction}) $ t1 $ t2 =>
   971             do_conjunction t0 t1 t2 accum
   972           | (t0 as Const (@{const_name All}, _)) $ Abs (s0, T1, t1) =>
   973             do_all t0 s0 T1 t1 accum
   974           | Const (x as (@{const_name "op ="}, _)) $ t1 $ t2 =>
   975             consider_general_equals mdata true x t1 t2 accum
   976           | (t0 as @{const "op &"}) $ t1 $ t2 => do_conjunction t0 t1 t2 accum
   977           | (t0 as @{const "op -->"}) $ t1 $ t2 => do_implies t0 t1 t2 accum
   978           | _ => raise TERM ("Nitpick_Mono.consider_definitional_axiom.\
   979                              \do_formula", [t])
   980     in do_formula t end
   981 
   982 fun string_for_mtype_of_term ctxt lits t M =
   983   Syntax.string_of_term ctxt t ^ " : " ^ string_for_mtype (resolve_mtype lits M)
   984 
   985 fun print_mtype_context ctxt lits ({frees, consts, ...} : mtype_context) =
   986   print_g (fn () =>
   987       map (fn (x, M) => string_for_mtype_of_term ctxt lits (Free x) M) frees @
   988       map (fn (x, M) => string_for_mtype_of_term ctxt lits (Const x) M) consts
   989       |> cat_lines)
   990 
   991 fun amass f t (ms, accum) =
   992   let val (m, accum) = f t accum in (m :: ms, accum) end
   993 
   994 fun infer which no_harmless (hol_ctxt as {ctxt, ...}) binarize alpha_T
   995           (nondef_ts, def_ts) =
   996   let
   997     val _ = print_g (fn () => "****** " ^ which ^ " analysis: " ^
   998                               string_for_mtype MAlpha ^ " is " ^
   999                               Syntax.string_of_typ ctxt alpha_T)
  1000     val mdata as {max_fresh, constr_mcache, ...} =
  1001       initial_mdata hol_ctxt binarize no_harmless alpha_T
  1002     val accum = (initial_gamma, ([], [], []))
  1003     val (nondef_ms, accum) =
  1004       ([], accum) |> amass (consider_general_formula mdata Plus) (hd nondef_ts)
  1005                   |> fold (amass (consider_nondefinitional_axiom mdata))
  1006                           (tl nondef_ts)
  1007     val (def_ms, (gamma, cset)) =
  1008       ([], accum) |> fold (amass (consider_definitional_axiom mdata)) def_ts
  1009   in
  1010     case solve (!max_fresh) cset of
  1011       SOME lits => (print_mtype_context ctxt lits gamma;
  1012                     SOME (lits, (nondef_ms, def_ms), !constr_mcache))
  1013     | _ => NONE
  1014   end
  1015   handle UNSOLVABLE () => NONE
  1016        | MTYPE (loc, Ms, Ts) =>
  1017          raise BAD (loc, commas (map string_for_mtype Ms @
  1018                                  map (Syntax.string_of_typ ctxt) Ts))
  1019        | MTERM (loc, ms) =>
  1020          raise BAD (loc, commas (map (string_for_mterm ctxt) ms))
  1021 
  1022 val formulas_monotonic = is_some oooo infer "Monotonicity" false
  1023 
  1024 fun fin_fun_constr T1 T2 =
  1025   (@{const_name FinFun}, (T1 --> T2) --> Type (@{type_name fin_fun}, [T1, T2]))
  1026 
  1027 fun finitize_funs (hol_ctxt as {thy, ctxt, stds, fast_descrs, constr_cache,
  1028                                 ...})
  1029                   binarize finitizes alpha_T tsp =
  1030   case infer "Finiteness" true hol_ctxt binarize alpha_T tsp of
  1031     SOME (lits, msp, constr_mtypes) =>
  1032     if forall (curry (op =) Minus o snd) lits then
  1033       tsp
  1034     else
  1035       let
  1036         fun should_finitize T a =
  1037           case triple_lookup (type_match thy) finitizes T of
  1038             SOME (SOME false) => false
  1039           | _ => resolve_sign_atom lits a = S Plus
  1040         fun type_from_mtype T M =
  1041           case (M, T) of
  1042             (MAlpha, _) => T
  1043           | (MFun (M1, a, M2), Type (@{type_name fun}, Ts)) =>
  1044             Type (if should_finitize T a then @{type_name fin_fun}
  1045                   else @{type_name fun}, map2 type_from_mtype Ts [M1, M2])
  1046           | (MPair (M1, M2), Type (@{type_name Product_Type.prod}, Ts)) =>
  1047             Type (@{type_name Product_Type.prod}, map2 type_from_mtype Ts [M1, M2])
  1048           | (MType _, _) => T
  1049           | _ => raise MTYPE ("Nitpick_Mono.finitize_funs.type_from_mtype",
  1050                               [M], [T])
  1051         fun finitize_constr (x as (s, T)) =
  1052           (s, case AList.lookup (op =) constr_mtypes x of
  1053                 SOME M => type_from_mtype T M
  1054               | NONE => T)
  1055         fun term_from_mterm new_Ts old_Ts m =
  1056           case m of
  1057             MRaw (t, M) =>
  1058             let
  1059               val T = fastype_of1 (old_Ts, t)
  1060               val T' = type_from_mtype T M
  1061             in
  1062               case t of
  1063                 Const (x as (s, _)) =>
  1064                 if s = @{const_name finite} then
  1065                   case domain_type T' of
  1066                     set_T' as Type (@{type_name fin_fun}, _) =>
  1067                     Abs (Name.uu, set_T', @{const True})
  1068                   | _ => Const (s, T')
  1069                 else if s = @{const_name "=="} orelse
  1070                         s = @{const_name "op ="} then
  1071                   let
  1072                     val T =
  1073                       case T' of
  1074                         Type (_, [T1, Type (_, [T2, T3])]) =>
  1075                         T1 --> T2 --> T3
  1076                       | _ => raise TYPE ("Nitpick_Mono.finitize_funs.\
  1077                                          \term_from_mterm", [T, T'], [])
  1078                   in coerce_term hol_ctxt new_Ts T' T (Const (s, T)) end
  1079                 else if is_built_in_const thy stds fast_descrs x then
  1080                   coerce_term hol_ctxt new_Ts T' T t
  1081                 else if is_constr ctxt stds x then
  1082                   Const (finitize_constr x)
  1083                 else if is_sel s then
  1084                   let
  1085                     val n = sel_no_from_name s
  1086                     val x' =
  1087                       x |> binarized_and_boxed_constr_for_sel hol_ctxt binarize
  1088                         |> finitize_constr
  1089                     val x'' =
  1090                       binarized_and_boxed_nth_sel_for_constr hol_ctxt binarize
  1091                                                              x' n
  1092                   in Const x'' end
  1093                 else
  1094                   Const (s, T')
  1095               | Free (s, T) => Free (s, type_from_mtype T M)
  1096               | Bound _ => t
  1097               | _ => raise MTERM ("Nitpick_Mono.finitize_funs.term_from_mterm",
  1098                                   [m])
  1099             end
  1100           | MApp (m1, m2) =>
  1101             let
  1102               val (t1, t2) = pairself (term_from_mterm new_Ts old_Ts) (m1, m2)
  1103               val (T1, T2) = pairself (curry fastype_of1 new_Ts) (t1, t2)
  1104               val (t1', T2') =
  1105                 case T1 of
  1106                   Type (s, [T11, T12]) => 
  1107                   (if s = @{type_name fin_fun} then
  1108                      select_nth_constr_arg ctxt stds (fin_fun_constr T11 T12) t1
  1109                                            0 (T11 --> T12)
  1110                    else
  1111                      t1, T11)
  1112                 | _ => raise TYPE ("Nitpick_Mono.finitize_funs.term_from_mterm",
  1113                                    [T1], [])
  1114             in betapply (t1', coerce_term hol_ctxt new_Ts T2' T2 t2) end
  1115           | MAbs (s, old_T, M, a, m') =>
  1116             let
  1117               val new_T = type_from_mtype old_T M
  1118               val t' = term_from_mterm (new_T :: new_Ts) (old_T :: old_Ts) m'
  1119               val T' = fastype_of1 (new_T :: new_Ts, t')
  1120             in
  1121               Abs (s, new_T, t')
  1122               |> should_finitize (new_T --> T') a
  1123                  ? construct_value ctxt stds (fin_fun_constr new_T T') o single
  1124             end
  1125       in
  1126         Unsynchronized.change constr_cache (map (apsnd (map finitize_constr)));
  1127         pairself (map (term_from_mterm [] [])) msp
  1128       end
  1129   | NONE => tsp
  1130 
  1131 end;