src/HOL/Tools/ATP/atp_util.ML
author blanchet
Thu, 14 Jul 2011 16:50:05 +0200
changeset 43827 62d64709af3b
parent 43572 ae612a423dad
child 43863 a43d61270142
permissions -rw-r--r--
added option to control which lambda translation to use (for experiments)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/Tools/ATP/atp_util.ML
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
     3
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
     4
General-purpose functions used by the ATP module.
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
     5
*)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
     6
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
     7
signature ATP_UTIL =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
     8
sig
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
     9
  val timestamp : unit -> string
43827
62d64709af3b added option to control which lambda translation to use (for experiments)
blanchet
parents: 43572
diff changeset
    10
  val hash_string : string -> int
62d64709af3b added option to control which lambda translation to use (for experiments)
blanchet
parents: 43572
diff changeset
    11
  val hash_term : term -> int
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    12
  val strip_spaces : bool -> (char -> bool) -> string -> string
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    13
  val nat_subscript : int -> string
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    14
  val unyxml : string -> string
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    15
  val maybe_quote : string -> string
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    16
  val string_from_ext_time : bool * Time.time -> string
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    17
  val string_from_time : Time.time -> string
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    18
  val varify_type : Proof.context -> typ -> typ
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    19
  val instantiate_type : theory -> typ -> typ -> typ -> typ
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    20
  val varify_and_instantiate_type : Proof.context -> typ -> typ -> typ -> typ
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    21
  val typ_of_dtyp :
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    22
    Datatype_Aux.descr -> (Datatype_Aux.dtyp * typ) list -> Datatype_Aux.dtyp
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    23
    -> typ
43572
ae612a423dad added "sound" option to force Sledgehammer to be pedantically sound
blanchet
parents: 43423
diff changeset
    24
  val is_type_surely_finite : Proof.context -> bool -> typ -> bool
ae612a423dad added "sound" option to force Sledgehammer to be pedantically sound
blanchet
parents: 43423
diff changeset
    25
  val is_type_surely_infinite : Proof.context -> bool -> typ -> bool
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    26
  val monomorphic_term : Type.tyenv -> term -> term
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    27
  val eta_expand : typ list -> term -> int -> term
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    28
  val transform_elim_prop : term -> term
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    29
  val specialize_type : theory -> (string * typ) -> term -> term
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    30
  val strip_subgoal :
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    31
    Proof.context -> thm -> int -> (string * typ) list * term list * term
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    32
end;
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    33
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    34
structure ATP_Util : ATP_UTIL =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    35
struct
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    36
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    37
val timestamp = Date.fmt "%Y-%m-%d %H:%M:%S" o Date.fromTimeLocal o Time.now
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    38
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    39
(* This hash function is recommended in "Compilers: Principles, Techniques, and
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    40
   Tools" by Aho, Sethi, and Ullman. The "hashpjw" function, which they
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    41
   particularly recommend, triggers a bug in versions of Poly/ML up to 4.2.0. *)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    42
fun hashw (u, w) = Word.+ (u, Word.* (0w65599, w))
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    43
fun hashw_char (c, w) = hashw (Word.fromInt (Char.ord c), w)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    44
fun hashw_string (s : string, w) = CharVector.foldl hashw_char w s
43827
62d64709af3b added option to control which lambda translation to use (for experiments)
blanchet
parents: 43572
diff changeset
    45
fun hashw_term (t1 $ t2) = hashw (hashw_term t1, hashw_term t2)
62d64709af3b added option to control which lambda translation to use (for experiments)
blanchet
parents: 43572
diff changeset
    46
  | hashw_term (Const (s, _)) = hashw_string (s, 0w0)
62d64709af3b added option to control which lambda translation to use (for experiments)
blanchet
parents: 43572
diff changeset
    47
  | hashw_term (Free (s, _)) = hashw_string (s, 0w0)
62d64709af3b added option to control which lambda translation to use (for experiments)
blanchet
parents: 43572
diff changeset
    48
  | hashw_term _ = 0w0
62d64709af3b added option to control which lambda translation to use (for experiments)
blanchet
parents: 43572
diff changeset
    49
62d64709af3b added option to control which lambda translation to use (for experiments)
blanchet
parents: 43572
diff changeset
    50
fun hash_string s = Word.toInt (hashw_string (s, 0w0))
62d64709af3b added option to control which lambda translation to use (for experiments)
blanchet
parents: 43572
diff changeset
    51
val hash_term = Word.toInt o hashw_term
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    52
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    53
fun strip_c_style_comment _ [] = []
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    54
  | strip_c_style_comment is_evil (#"*" :: #"/" :: cs) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    55
    strip_spaces_in_list true is_evil cs
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    56
  | strip_c_style_comment is_evil (_ :: cs) = strip_c_style_comment is_evil cs
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    57
and strip_spaces_in_list _ _ [] = []
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    58
  | strip_spaces_in_list true is_evil (#"%" :: cs) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    59
    strip_spaces_in_list true is_evil
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    60
                         (cs |> chop_while (not_equal #"\n") |> snd)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    61
  | strip_spaces_in_list true is_evil (#"/" :: #"*" :: cs) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    62
    strip_c_style_comment is_evil cs
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    63
  | strip_spaces_in_list _ _ [c1] = if Char.isSpace c1 then [] else [str c1]
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    64
  | strip_spaces_in_list skip_comments is_evil [c1, c2] =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    65
    strip_spaces_in_list skip_comments is_evil [c1] @
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    66
    strip_spaces_in_list skip_comments is_evil [c2]
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    67
  | strip_spaces_in_list skip_comments is_evil (c1 :: c2 :: c3 :: cs) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    68
    if Char.isSpace c1 then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    69
      strip_spaces_in_list skip_comments is_evil (c2 :: c3 :: cs)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    70
    else if Char.isSpace c2 then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    71
      if Char.isSpace c3 then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    72
        strip_spaces_in_list skip_comments is_evil (c1 :: c3 :: cs)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    73
      else
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    74
        str c1 :: (if forall is_evil [c1, c3] then [" "] else []) @
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    75
        strip_spaces_in_list skip_comments is_evil (c3 :: cs)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    76
    else
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    77
      str c1 :: strip_spaces_in_list skip_comments is_evil (c2 :: c3 :: cs)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    78
fun strip_spaces skip_comments is_evil =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    79
  implode o strip_spaces_in_list skip_comments is_evil o String.explode
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    80
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    81
val subscript = implode o map (prefix "\<^isub>") o raw_explode  (* FIXME Symbol.explode (?) *)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    82
fun nat_subscript n =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    83
  n |> string_of_int |> print_mode_active Symbol.xsymbolsN ? subscript
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    84
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    85
val unyxml = XML.content_of o YXML.parse_body
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    86
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    87
val is_long_identifier = forall Lexicon.is_identifier o space_explode "."
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    88
fun maybe_quote y =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    89
  let val s = unyxml y in
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    90
    y |> ((not (is_long_identifier (perhaps (try (unprefix "'")) s)) andalso
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    91
           not (is_long_identifier (perhaps (try (unprefix "?")) s))) orelse
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    92
           Keyword.is_keyword s) ? quote
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    93
  end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    94
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    95
fun string_from_ext_time (plus, time) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    96
  let val ms = Time.toMilliseconds time in
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    97
    (if plus then "> " else "") ^
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    98
    (if plus andalso ms mod 1000 = 0 then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    99
       signed_string_of_int (ms div 1000) ^ " s"
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   100
     else if ms < 1000 then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   101
       signed_string_of_int ms ^ " ms"
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   102
     else
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   103
       string_of_real (0.01 * Real.fromInt (ms div 10)) ^ " s")
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   104
  end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   105
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   106
val string_from_time = string_from_ext_time o pair false
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   107
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   108
fun varify_type ctxt T =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   109
  Variable.polymorphic_types ctxt [Const (@{const_name undefined}, T)]
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   110
  |> snd |> the_single |> dest_Const |> snd
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   111
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   112
(* TODO: use "Term_Subst.instantiateT" instead? *)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   113
fun instantiate_type thy T1 T1' T2 =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   114
  Same.commit (Envir.subst_type_same
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   115
                   (Sign.typ_match thy (T1, T1') Vartab.empty)) T2
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   116
  handle Type.TYPE_MATCH => raise TYPE ("instantiate_type", [T1, T1'], [])
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   117
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   118
fun varify_and_instantiate_type ctxt T1 T1' T2 =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   119
  let val thy = Proof_Context.theory_of ctxt in
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   120
    instantiate_type thy (varify_type ctxt T1) T1' (varify_type ctxt T2)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   121
  end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   122
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   123
fun typ_of_dtyp _ typ_assoc (Datatype_Aux.DtTFree a) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   124
    the (AList.lookup (op =) typ_assoc (Datatype_Aux.DtTFree a))
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   125
  | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtType (s, Us)) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   126
    Type (s, map (typ_of_dtyp descr typ_assoc) Us)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   127
  | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtRec i) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   128
    let val (s, ds, _) = the (AList.lookup (op =) descr i) in
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   129
      Type (s, map (typ_of_dtyp descr typ_assoc) ds)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   130
    end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   131
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   132
fun datatype_constrs thy (T as Type (s, Ts)) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   133
    (case Datatype.get_info thy s of
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   134
       SOME {index, descr, ...} =>
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   135
       let val (_, dtyps, constrs) = AList.lookup (op =) descr index |> the in
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   136
         map (apsnd (fn Us => map (typ_of_dtyp descr (dtyps ~~ Ts)) Us ---> T))
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   137
             constrs
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   138
       end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   139
     | NONE => [])
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   140
  | datatype_constrs _ _ = []
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   141
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   142
(* Similar to "Nitpick_HOL.bounded_exact_card_of_type".
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   143
   0 means infinite type, 1 means singleton type (e.g., "unit"), and 2 means
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   144
   cardinality 2 or more. The specified default cardinality is returned if the
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   145
   cardinality of the type can't be determined. *)
43572
ae612a423dad added "sound" option to force Sledgehammer to be pedantically sound
blanchet
parents: 43423
diff changeset
   146
fun tiny_card_of_type ctxt sound default_card T =
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   147
  let
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   148
    val thy = Proof_Context.theory_of ctxt
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   149
    val max = 2 (* 1 would be too small for the "fun" case *)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   150
    fun aux slack avoid T =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   151
      if member (op =) avoid T then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   152
        0
43423
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   153
      else case T of
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   154
        Type (@{type_name fun}, [T1, T2]) =>
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   155
        (case (aux slack avoid T1, aux slack avoid T2) of
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   156
           (k, 1) => if slack andalso k = 0 then 0 else 1
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   157
         | (0, _) => 0
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   158
         | (_, 0) => 0
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   159
         | (k1, k2) =>
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   160
           if k1 >= max orelse k2 >= max then max
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   161
           else Int.min (max, Integer.pow k2 k1))
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   162
      | @{typ prop} => 2
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   163
      | @{typ bool} => 2 (* optimization *)
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   164
      | @{typ nat} => 0 (* optimization *)
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   165
      | Type ("Int.int", []) => 0 (* optimization *)
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   166
      | Type (s, _) =>
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   167
        (case datatype_constrs thy T of
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   168
           constrs as _ :: _ =>
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   169
           let
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   170
             val constr_cards =
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   171
               map (Integer.prod o map (aux slack (T :: avoid)) o binder_types
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   172
                    o snd) constrs
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   173
           in
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   174
             if exists (curry (op =) 0) constr_cards then 0
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   175
             else Int.min (max, Integer.sum constr_cards)
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   176
           end
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   177
         | [] =>
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   178
           case Typedef.get_info ctxt s of
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   179
             ({abs_type, rep_type, ...}, _) :: _ =>
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   180
             (* We cheat here by assuming that typedef types are infinite if
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   181
                their underlying type is infinite. This is unsound in general
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   182
                but it's hard to think of a realistic example where this would
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   183
                not be the case. We are also slack with representation types:
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   184
                If a representation type has the form "sigma => tau", we
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   185
                consider it enough to check "sigma" for infiniteness. (Look
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   186
                for "slack" in this function.) *)
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   187
             (case varify_and_instantiate_type ctxt
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   188
                       (Logic.varifyT_global abs_type) T
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   189
                       (Logic.varifyT_global rep_type)
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   190
                   |> aux true avoid of
43572
ae612a423dad added "sound" option to force Sledgehammer to be pedantically sound
blanchet
parents: 43423
diff changeset
   191
                0 => if sound then default_card else 0
43423
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   192
              | 1 => 1
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   193
              | _ => default_card)
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   194
           | [] => default_card)
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   195
        (* Very slightly unsound: Type variables are assumed not to be
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   196
           constrained to cardinality 1. (In practice, the user would most
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   197
           likely have used "unit" directly anyway.) *)
43572
ae612a423dad added "sound" option to force Sledgehammer to be pedantically sound
blanchet
parents: 43423
diff changeset
   198
      | TFree _ =>
ae612a423dad added "sound" option to force Sledgehammer to be pedantically sound
blanchet
parents: 43423
diff changeset
   199
        if default_card = 1 andalso not sound then 2 else default_card
43423
717880e98e6b gave up an optimization that sometimes lead to unsound proofs -- in short, facts talking about a schematic type variable can encode a cardinality constraint and be consistent with HOL, e.g. "card (UNIV::?'a set) = 1 ==> ALL x y. x = y"
blanchet
parents: 43171
diff changeset
   200
      | TVar _ => default_card
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   201
  in Int.min (max, aux false [] T) end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   202
43572
ae612a423dad added "sound" option to force Sledgehammer to be pedantically sound
blanchet
parents: 43423
diff changeset
   203
fun is_type_surely_finite ctxt sound T = tiny_card_of_type ctxt sound 0 T <> 0
ae612a423dad added "sound" option to force Sledgehammer to be pedantically sound
blanchet
parents: 43423
diff changeset
   204
fun is_type_surely_infinite ctxt sound T = tiny_card_of_type ctxt sound 1 T = 0
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   205
43171
37e1431cc213 gracefully handle the case where a constant is partially or not instantiated at all, as may happen when reconstructing Metis proofs for polymorphic type encodings
blanchet
parents: 43085
diff changeset
   206
fun monomorphic_term subst =
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   207
  map_types (map_type_tvar (fn v =>
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   208
      case Type.lookup subst v of
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   209
        SOME typ => typ
43171
37e1431cc213 gracefully handle the case where a constant is partially or not instantiated at all, as may happen when reconstructing Metis proofs for polymorphic type encodings
blanchet
parents: 43085
diff changeset
   210
      | NONE => TVar v))
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   211
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   212
fun eta_expand _ t 0 = t
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   213
  | eta_expand Ts (Abs (s, T, t')) n =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   214
    Abs (s, T, eta_expand (T :: Ts) t' (n - 1))
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   215
  | eta_expand Ts t n =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   216
    fold_rev (fn T => fn t' => Abs ("x" ^ nat_subscript n, T, t'))
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   217
             (List.take (binder_types (fastype_of1 (Ts, t)), n))
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   218
             (list_comb (incr_boundvars n t, map Bound (n - 1 downto 0)))
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   219
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   220
(* Converts an elim-rule into an equivalent theorem that does not have the
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   221
   predicate variable. Leaves other theorems unchanged. We simply instantiate
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   222
   the conclusion variable to False. (Cf. "transform_elim_theorem" in
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   223
   "Meson_Clausify".) *)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   224
fun transform_elim_prop t =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   225
  case Logic.strip_imp_concl t of
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   226
    @{const Trueprop} $ Var (z, @{typ bool}) =>
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   227
    subst_Vars [(z, @{const False})] t
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   228
  | Var (z, @{typ prop}) => subst_Vars [(z, @{prop False})] t
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   229
  | _ => t
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   230
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   231
fun specialize_type thy (s, T) t =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   232
  let
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   233
    fun subst_for (Const (s', T')) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   234
      if s = s' then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   235
        SOME (Sign.typ_match thy (T', T) Vartab.empty)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   236
        handle Type.TYPE_MATCH => NONE
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   237
      else
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   238
        NONE
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   239
    | subst_for (t1 $ t2) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   240
      (case subst_for t1 of SOME x => SOME x | NONE => subst_for t2)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   241
    | subst_for (Abs (_, _, t')) = subst_for t'
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   242
    | subst_for _ = NONE
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   243
  in
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   244
    case subst_for t of
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   245
      SOME subst => monomorphic_term subst t
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   246
    | NONE => raise Type.TYPE_MATCH
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   247
  end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   248
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   249
fun strip_subgoal ctxt goal i =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   250
  let
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   251
    val (t, (frees, params)) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   252
      Logic.goal_params (prop_of goal) i
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   253
      ||> (map dest_Free #> Variable.variant_frees ctxt [] #> `(map Free))
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   254
    val hyp_ts = t |> Logic.strip_assums_hyp |> map (curry subst_bounds frees)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   255
    val concl_t = t |> Logic.strip_assums_concl |> curry subst_bounds frees
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   256
  in (rev params, hyp_ts, concl_t) end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   257
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   258
end;