src/HOL/Tools/ATP/atp_util.ML
author wenzelm
Fri, 26 Aug 2011 16:06:58 +0200
changeset 44481 bb42bc831570
parent 44460 5d0754cf994a
child 44491 ba22ed224b20
permissions -rw-r--r--
tuned signature;
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
44399
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
    18
  val type_instance : Proof.context -> typ -> typ -> bool
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
    19
  val type_generalization : Proof.context -> typ -> typ -> bool
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
    20
  val type_aconv : Proof.context -> typ * typ -> bool
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    21
  val varify_type : Proof.context -> typ -> typ
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    22
  val instantiate_type : theory -> typ -> typ -> typ -> typ
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    23
  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
    24
  val typ_of_dtyp :
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    25
    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
    26
    -> typ
44393
23adec5984f1 make sound mode more sound (and clean up code)
blanchet
parents: 44392
diff changeset
    27
  val is_type_surely_finite : Proof.context -> typ -> bool
44399
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
    28
  val is_type_surely_infinite : Proof.context -> bool -> typ list -> typ -> bool
43863
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
    29
  val s_not : term -> term
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
    30
  val s_conj : term * term -> term
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
    31
  val s_disj : term * term -> term
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
    32
  val s_imp : term * term -> term
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
    33
  val s_iff : term * term -> term
43864
58a7b3fdc193 fixed lambda-liftg: must ensure the formulas are in close form
blanchet
parents: 43863
diff changeset
    34
  val close_form : term -> term
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    35
  val monomorphic_term : Type.tyenv -> term -> term
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    36
  val eta_expand : typ list -> term -> int -> term
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    37
  val transform_elim_prop : term -> term
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    38
  val specialize_type : theory -> (string * typ) -> term -> term
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    39
  val strip_subgoal :
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    40
    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
    41
end;
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    42
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    43
structure ATP_Util : ATP_UTIL =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    44
struct
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    45
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    46
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
    47
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    48
(* 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
    49
   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
    50
   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
    51
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
    52
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
    53
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
    54
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
    55
  | 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
    56
  | 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
    57
  | hashw_term _ = 0w0
62d64709af3b added option to control which lambda translation to use (for experiments)
blanchet
parents: 43572
diff changeset
    58
62d64709af3b added option to control which lambda translation to use (for experiments)
blanchet
parents: 43572
diff changeset
    59
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
    60
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
    61
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    62
fun strip_c_style_comment _ [] = []
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    63
  | strip_c_style_comment is_evil (#"*" :: #"/" :: cs) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    64
    strip_spaces_in_list true is_evil cs
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    65
  | 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
    66
and strip_spaces_in_list _ _ [] = []
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    67
  | strip_spaces_in_list true is_evil (#"%" :: cs) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    68
    strip_spaces_in_list true is_evil
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    69
                         (cs |> chop_while (not_equal #"\n") |> snd)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    70
  | strip_spaces_in_list true is_evil (#"/" :: #"*" :: cs) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    71
    strip_c_style_comment is_evil cs
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    72
  | 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
    73
  | 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
    74
    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
    75
    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
    76
  | 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
    77
    if Char.isSpace c1 then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    78
      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
    79
    else if Char.isSpace c2 then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    80
      if Char.isSpace c3 then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    81
        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
    82
      else
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    83
        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
    84
        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
    85
    else
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    86
      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
    87
fun strip_spaces skip_comments is_evil =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    88
  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
    89
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    90
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
    91
fun nat_subscript n =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    92
  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
    93
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    94
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
    95
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    96
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
    97
fun maybe_quote y =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    98
  let val s = unyxml y in
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    99
    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
   100
           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
   101
           Keyword.is_keyword s) ? quote
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   102
  end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   103
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   104
fun string_from_ext_time (plus, time) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   105
  let val ms = Time.toMilliseconds time in
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   106
    (if plus then "> " else "") ^
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   107
    (if plus andalso ms mod 1000 = 0 then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   108
       signed_string_of_int (ms div 1000) ^ " s"
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   109
     else if ms < 1000 then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   110
       signed_string_of_int ms ^ " ms"
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   111
     else
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   112
       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
   113
  end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   114
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   115
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
   116
44399
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
   117
fun type_instance ctxt T T' =
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
   118
  Sign.typ_instance (Proof_Context.theory_of ctxt) (T, T')
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
   119
fun type_generalization ctxt T T' = type_instance ctxt T' T
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
   120
fun type_aconv ctxt (T, T') =
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
   121
  type_instance ctxt T T' andalso type_instance ctxt T' T
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
   122
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   123
fun varify_type ctxt T =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   124
  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
   125
  |> snd |> the_single |> dest_Const |> snd
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   126
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   127
(* TODO: use "Term_Subst.instantiateT" instead? *)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   128
fun instantiate_type thy T1 T1' T2 =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   129
  Same.commit (Envir.subst_type_same
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   130
                   (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
   131
  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
   132
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   133
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
   134
  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
   135
    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
   136
  end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   137
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   138
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
   139
    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
   140
  | 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
   141
    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
   142
  | 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
   143
    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
   144
      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
   145
    end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   146
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   147
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
   148
    (case Datatype.get_info thy s of
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   149
       SOME {index, descr, ...} =>
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   150
       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
   151
         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
   152
             constrs
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   153
       end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   154
     | NONE => [])
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   155
  | datatype_constrs _ _ = []
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   156
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   157
(* 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
   158
   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
   159
   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
   160
   cardinality of the type can't be determined. *)
44399
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
   161
fun tiny_card_of_type ctxt inst_tvars assigns default_card T =
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   162
  let
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   163
    val thy = Proof_Context.theory_of ctxt
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   164
    val max = 2 (* 1 would be too small for the "fun" case *)
44399
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
   165
    val type_cmp =
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
   166
      if inst_tvars then uncurry (type_generalization ctxt) else type_aconv ctxt
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   167
    fun aux slack avoid T =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   168
      if member (op =) avoid T then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   169
        0
44399
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
   170
      else case AList.lookup type_cmp assigns T of
44393
23adec5984f1 make sound mode more sound (and clean up code)
blanchet
parents: 44392
diff changeset
   171
        SOME k => k
23adec5984f1 make sound mode more sound (and clean up code)
blanchet
parents: 44392
diff changeset
   172
      | NONE =>
44392
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   173
        case T of
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   174
          Type (@{type_name fun}, [T1, T2]) =>
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   175
          (case (aux slack avoid T1, aux slack avoid T2) of
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   176
             (k, 1) => if slack andalso k = 0 then 0 else 1
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   177
           | (0, _) => 0
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   178
           | (_, 0) => 0
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   179
           | (k1, k2) =>
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   180
             if k1 >= max orelse k2 >= max then max
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   181
             else Int.min (max, Integer.pow k2 k1))
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   182
        | @{typ prop} => 2
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   183
        | @{typ bool} => 2 (* optimization *)
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   184
        | @{typ nat} => 0 (* optimization *)
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   185
        | Type ("Int.int", []) => 0 (* optimization *)
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   186
        | Type (s, _) =>
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   187
          (case datatype_constrs thy T of
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   188
             constrs as _ :: _ =>
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   189
             let
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   190
               val constr_cards =
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   191
                 map (Integer.prod o map (aux slack (T :: avoid)) o binder_types
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   192
                      o snd) constrs
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   193
             in
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   194
               if exists (curry (op =) 0) constr_cards then 0
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   195
               else Int.min (max, Integer.sum constr_cards)
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   196
             end
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   197
           | [] =>
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   198
             case Typedef.get_info ctxt s of
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   199
               ({abs_type, rep_type, ...}, _) :: _ =>
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   200
               (* We cheat here by assuming that typedef types are infinite if
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   201
                  their underlying type is infinite. This is unsound in general
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   202
                  but it's hard to think of a realistic example where this would
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   203
                  not be the case. We are also slack with representation types:
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   204
                  If a representation type has the form "sigma => tau", we
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   205
                  consider it enough to check "sigma" for infiniteness. (Look
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   206
                  for "slack" in this function.) *)
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   207
               (case varify_and_instantiate_type ctxt
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   208
                         (Logic.varifyT_global abs_type) T
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   209
                         (Logic.varifyT_global rep_type)
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   210
                     |> aux true avoid of
44393
23adec5984f1 make sound mode more sound (and clean up code)
blanchet
parents: 44392
diff changeset
   211
                  0 => 0
44392
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   212
                | 1 => 1
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   213
                | _ => default_card)
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   214
             | [] => default_card)
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   215
          (* Very slightly unsound: Type variables are assumed not to be
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   216
             constrained to cardinality 1. (In practice, the user would most
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   217
             likely have used "unit" directly anyway.) *)
44393
23adec5984f1 make sound mode more sound (and clean up code)
blanchet
parents: 44392
diff changeset
   218
        | TFree _ => if default_card = 1 then 2 else default_card
44392
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   219
        | TVar _ => default_card
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   220
  in Int.min (max, aux false [] T) end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   221
44399
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
   222
fun is_type_surely_finite ctxt T = tiny_card_of_type ctxt false [] 0 T <> 0
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
   223
fun is_type_surely_infinite ctxt inst_tvars infinite_Ts T =
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
   224
  tiny_card_of_type ctxt inst_tvars (map (rpair 0) infinite_Ts) 1 T = 0
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   225
43863
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   226
(* Simple simplifications to ensure that sort annotations don't leave a trail of
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   227
   spurious "True"s. *)
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   228
fun s_not (Const (@{const_name All}, T) $ Abs (s, T', t')) =
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   229
    Const (@{const_name Ex}, T) $ Abs (s, T', s_not t')
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   230
  | s_not (Const (@{const_name Ex}, T) $ Abs (s, T', t')) =
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   231
    Const (@{const_name All}, T) $ Abs (s, T', s_not t')
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   232
  | s_not (@{const HOL.implies} $ t1 $ t2) = @{const HOL.conj} $ t1 $ s_not t2
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   233
  | s_not (@{const HOL.conj} $ t1 $ t2) =
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   234
    @{const HOL.disj} $ s_not t1 $ s_not t2
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   235
  | s_not (@{const HOL.disj} $ t1 $ t2) =
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   236
    @{const HOL.conj} $ s_not t1 $ s_not t2
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   237
  | s_not (@{const False}) = @{const True}
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   238
  | s_not (@{const True}) = @{const False}
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   239
  | s_not (@{const Not} $ t) = t
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   240
  | s_not t = @{const Not} $ t
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   241
fun s_conj (@{const True}, t2) = t2
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   242
  | s_conj (t1, @{const True}) = t1
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   243
  | s_conj p = HOLogic.mk_conj p
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   244
fun s_disj (@{const False}, t2) = t2
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   245
  | s_disj (t1, @{const False}) = t1
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   246
  | s_disj p = HOLogic.mk_disj p
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   247
fun s_imp (@{const True}, t2) = t2
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   248
  | s_imp (t1, @{const False}) = s_not t1
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   249
  | s_imp p = HOLogic.mk_imp p
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   250
fun s_iff (@{const True}, t2) = t2
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   251
  | s_iff (t1, @{const True}) = t1
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   252
  | s_iff (t1, t2) = HOLogic.eq_const HOLogic.boolT $ t1 $ t2
a43d61270142 ensure that the lambda translation procedure is called only once with all the facts, which is necessary for soundness of lambda-lifting (freshness of new names)
blanchet
parents: 43827
diff changeset
   253
43864
58a7b3fdc193 fixed lambda-liftg: must ensure the formulas are in close form
blanchet
parents: 43863
diff changeset
   254
fun close_form t =
58a7b3fdc193 fixed lambda-liftg: must ensure the formulas are in close form
blanchet
parents: 43863
diff changeset
   255
  fold (fn ((x, i), T) => fn t' =>
58a7b3fdc193 fixed lambda-liftg: must ensure the formulas are in close form
blanchet
parents: 43863
diff changeset
   256
           HOLogic.all_const T $ Abs (x, T, abstract_over (Var ((x, i), T), t')))
58a7b3fdc193 fixed lambda-liftg: must ensure the formulas are in close form
blanchet
parents: 43863
diff changeset
   257
       (Term.add_vars t []) t
58a7b3fdc193 fixed lambda-liftg: must ensure the formulas are in close form
blanchet
parents: 43863
diff changeset
   258
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
   259
fun monomorphic_term subst =
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   260
  map_types (map_type_tvar (fn v =>
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   261
      case Type.lookup subst v of
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   262
        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
   263
      | NONE => TVar v))
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   264
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   265
fun eta_expand _ t 0 = t
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   266
  | eta_expand Ts (Abs (s, T, t')) n =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   267
    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
   268
  | eta_expand Ts t n =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   269
    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
   270
             (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
   271
             (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
   272
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   273
(* 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
   274
   predicate variable. Leaves other theorems unchanged. We simply instantiate
44460
blanchet
parents: 44399
diff changeset
   275
   the conclusion variable to "False". (Cf. "transform_elim_theorem" in
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   276
   "Meson_Clausify".) *)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   277
fun transform_elim_prop t =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   278
  case Logic.strip_imp_concl t of
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   279
    @{const Trueprop} $ Var (z, @{typ bool}) =>
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   280
    subst_Vars [(z, @{const False})] t
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   281
  | 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
   282
  | _ => t
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   283
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   284
fun specialize_type thy (s, T) t =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   285
  let
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   286
    fun subst_for (Const (s', T')) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   287
      if s = s' then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   288
        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
   289
        handle Type.TYPE_MATCH => NONE
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   290
      else
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   291
        NONE
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   292
    | subst_for (t1 $ t2) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   293
      (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
   294
    | subst_for (Abs (_, _, t')) = subst_for t'
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   295
    | subst_for _ = NONE
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   296
  in
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   297
    case subst_for t of
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   298
      SOME subst => monomorphic_term subst t
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   299
    | NONE => raise Type.TYPE_MATCH
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   300
  end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   301
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   302
fun strip_subgoal ctxt goal i =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   303
  let
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   304
    val (t, (frees, params)) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   305
      Logic.goal_params (prop_of goal) i
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   306
      ||> (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
   307
    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
   308
    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
   309
  in (rev params, hyp_ts, concl_t) end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   310
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   311
end;