src/HOL/Tools/ATP/atp_util.ML
author blanchet
Tue, 24 Apr 2012 09:47:40 +0200
changeset 47715 04400144c6fc
parent 47150 6df6e56fd7cd
child 47718 39229c760636
permissions -rw-r--r--
handle TPTP definitions as definitions in Nitpick rather than as axioms
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
44784
blanchet
parents: 44500
diff changeset
    13
  val strip_spaces_except_between_idents : string -> string
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    14
  val nat_subscript : int -> string
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    15
  val unyxml : string -> string
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    16
  val maybe_quote : string -> string
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    17
  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
    18
  val string_from_time : Time.time -> string
47150
6df6e56fd7cd tuning (in particular, Symtab instead of AList)
blanchet
parents: 46711
diff changeset
    19
  val type_instance : theory -> typ -> typ -> bool
6df6e56fd7cd tuning (in particular, Symtab instead of AList)
blanchet
parents: 46711
diff changeset
    20
  val type_generalization : theory -> typ -> typ -> bool
6df6e56fd7cd tuning (in particular, Symtab instead of AList)
blanchet
parents: 46711
diff changeset
    21
  val type_intersect : theory -> typ -> typ -> bool
6df6e56fd7cd tuning (in particular, Symtab instead of AList)
blanchet
parents: 46711
diff changeset
    22
  val type_equiv : theory -> typ * typ -> bool
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    23
  val varify_type : Proof.context -> typ -> typ
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    24
  val instantiate_type : theory -> typ -> typ -> typ -> typ
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    25
  val varify_and_instantiate_type : Proof.context -> typ -> typ -> typ -> typ
45896
100fb1f33e3e tuned signature;
wenzelm
parents: 45570
diff changeset
    26
  val typ_of_dtyp : Datatype.descr -> (Datatype.dtyp * typ) list -> Datatype.dtyp -> 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
46385
0ccf458a3633 third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
blanchet
parents: 45896
diff changeset
    34
  val close_form_prefix : string
43864
58a7b3fdc193 fixed lambda-liftg: must ensure the formulas are in close form
blanchet
parents: 43863
diff changeset
    35
  val close_form : term -> term
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    36
  val monomorphic_term : Type.tyenv -> term -> term
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    37
  val eta_expand : typ list -> term -> int -> term
47715
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
    38
  val extensionalize_term : Proof.context -> term -> term
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    39
  val transform_elim_prop : term -> term
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    40
  val specialize_type : theory -> (string * typ) -> term -> term
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    41
  val strip_subgoal :
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    42
    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
    43
end;
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    44
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    45
structure ATP_Util : ATP_UTIL =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    46
struct
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
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
    49
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    50
(* 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
    51
   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
    52
   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
    53
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
    54
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
    55
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
    56
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
    57
  | 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
    58
  | 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
    59
  | hashw_term _ = 0w0
62d64709af3b added option to control which lambda translation to use (for experiments)
blanchet
parents: 43572
diff changeset
    60
62d64709af3b added option to control which lambda translation to use (for experiments)
blanchet
parents: 43572
diff changeset
    61
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
    62
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
    63
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    64
fun strip_spaces skip_comments is_evil =
44935
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    65
  let
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    66
    fun strip_c_style_comment [] accum = accum
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    67
      | strip_c_style_comment (#"*" :: #"/" :: cs) accum =
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    68
        strip_spaces_in_list true cs accum
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    69
      | strip_c_style_comment (_ :: cs) accum = strip_c_style_comment cs accum
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    70
    and strip_spaces_in_list _ [] accum = rev accum
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    71
      | strip_spaces_in_list true (#"%" :: cs) accum =
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    72
        strip_spaces_in_list true (cs |> chop_while (not_equal #"\n") |> snd)
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    73
                             accum
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    74
      | strip_spaces_in_list true (#"/" :: #"*" :: cs) accum =
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    75
        strip_c_style_comment cs accum
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    76
      | strip_spaces_in_list _ [c1] accum =
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    77
        accum |> not (Char.isSpace c1) ? cons c1
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    78
      | strip_spaces_in_list skip_comments (cs as [_, _]) accum =
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    79
        accum |> fold (strip_spaces_in_list skip_comments o single) cs
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    80
      | strip_spaces_in_list skip_comments (c1 :: c2 :: c3 :: cs) accum =
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    81
        if Char.isSpace c1 then
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    82
          strip_spaces_in_list skip_comments (c2 :: c3 :: cs) accum
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    83
        else if Char.isSpace c2 then
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    84
          if Char.isSpace c3 then
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    85
            strip_spaces_in_list skip_comments (c1 :: c3 :: cs) accum
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    86
          else
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    87
            strip_spaces_in_list skip_comments (c3 :: cs)
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    88
                (c1 :: accum |> forall is_evil [c1, c3] ? cons #" ")
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    89
        else
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    90
          strip_spaces_in_list skip_comments (c2 :: c3 :: cs) (cons c1 accum)
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    91
  in
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    92
    String.explode
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    93
    #> rpair [] #-> strip_spaces_in_list skip_comments
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    94
    #> rev #> String.implode
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    95
  end
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    96
44784
blanchet
parents: 44500
diff changeset
    97
fun is_ident_char c = Char.isAlphaNum c orelse c = #"_"
blanchet
parents: 44500
diff changeset
    98
val strip_spaces_except_between_idents = strip_spaces true is_ident_char
blanchet
parents: 44500
diff changeset
    99
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   100
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
   101
fun nat_subscript n =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   102
  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
   103
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   104
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
   105
46711
f745bcc4a1e5 more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
wenzelm
parents: 46385
diff changeset
   106
val is_long_identifier = forall Lexicon.is_identifier o Long_Name.explode
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   107
fun maybe_quote y =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   108
  let val s = unyxml y in
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   109
    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
   110
           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
   111
           Keyword.is_keyword s) ? quote
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   112
  end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   113
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   114
fun string_from_ext_time (plus, time) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   115
  let val ms = Time.toMilliseconds time in
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   116
    (if plus then "> " else "") ^
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   117
    (if plus andalso ms mod 1000 = 0 then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   118
       signed_string_of_int (ms div 1000) ^ " s"
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   119
     else if ms < 1000 then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   120
       signed_string_of_int ms ^ " ms"
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   121
     else
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   122
       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
   123
  end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   124
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   125
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
   126
47150
6df6e56fd7cd tuning (in particular, Symtab instead of AList)
blanchet
parents: 46711
diff changeset
   127
fun type_instance thy T T' = Sign.typ_instance thy (T, T')
6df6e56fd7cd tuning (in particular, Symtab instead of AList)
blanchet
parents: 46711
diff changeset
   128
fun type_generalization thy T T' = Sign.typ_instance thy (T', T)
6df6e56fd7cd tuning (in particular, Symtab instead of AList)
blanchet
parents: 46711
diff changeset
   129
fun type_intersect thy T T' =
6df6e56fd7cd tuning (in particular, Symtab instead of AList)
blanchet
parents: 46711
diff changeset
   130
  can (Sign.typ_unify thy (T, Logic.incr_tvar (maxidx_of_typ T + 1) T'))
44893
bdc64c34ccae fixed type intersection (again)
blanchet
parents: 44859
diff changeset
   131
      (Vartab.empty, 0)
47150
6df6e56fd7cd tuning (in particular, Symtab instead of AList)
blanchet
parents: 46711
diff changeset
   132
val type_equiv = Sign.typ_equiv
44399
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
   133
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   134
fun varify_type ctxt T =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   135
  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
   136
  |> snd |> the_single |> dest_Const |> snd
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
(* TODO: use "Term_Subst.instantiateT" instead? *)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   139
fun instantiate_type thy T1 T1' T2 =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   140
  Same.commit (Envir.subst_type_same
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   141
                   (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
   142
  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
   143
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   144
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
   145
  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
   146
    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
   147
  end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   148
45896
100fb1f33e3e tuned signature;
wenzelm
parents: 45570
diff changeset
   149
fun typ_of_dtyp _ typ_assoc (Datatype.DtTFree a) =
100fb1f33e3e tuned signature;
wenzelm
parents: 45570
diff changeset
   150
    the (AList.lookup (op =) typ_assoc (Datatype.DtTFree a))
100fb1f33e3e tuned signature;
wenzelm
parents: 45570
diff changeset
   151
  | typ_of_dtyp descr typ_assoc (Datatype.DtType (s, Us)) =
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   152
    Type (s, map (typ_of_dtyp descr typ_assoc) Us)
45896
100fb1f33e3e tuned signature;
wenzelm
parents: 45570
diff changeset
   153
  | typ_of_dtyp descr typ_assoc (Datatype.DtRec i) =
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   154
    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
   155
      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
   156
    end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   157
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   158
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
   159
    (case Datatype.get_info thy s of
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   160
       SOME {index, descr, ...} =>
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   161
       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
   162
         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
   163
             constrs
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   164
       end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   165
     | NONE => [])
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   166
  | datatype_constrs _ _ = []
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   167
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   168
(* 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
   169
   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
   170
   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
   171
   cardinality of the type can't be determined. *)
44500
dbd98b549597 make default unsound mode less unsound
blanchet
parents: 44491
diff changeset
   172
fun tiny_card_of_type ctxt sound assigns default_card T =
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   173
  let
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   174
    val thy = Proof_Context.theory_of ctxt
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   175
    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
   176
    fun aux slack avoid T =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   177
      if member (op =) avoid T then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   178
        0
47150
6df6e56fd7cd tuning (in particular, Symtab instead of AList)
blanchet
parents: 46711
diff changeset
   179
      else case AList.lookup (type_equiv thy) assigns T of
44393
23adec5984f1 make sound mode more sound (and clean up code)
blanchet
parents: 44392
diff changeset
   180
        SOME k => k
23adec5984f1 make sound mode more sound (and clean up code)
blanchet
parents: 44392
diff changeset
   181
      | NONE =>
44392
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   182
        case T of
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   183
          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
   184
          (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
   185
             (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
   186
           | (0, _) => 0
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   187
           | (_, 0) => 0
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   188
           | (k1, k2) =>
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   189
             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
   190
             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
   191
        | @{typ prop} => 2
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   192
        | @{typ bool} => 2 (* optimization *)
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   193
        | @{typ nat} => 0 (* optimization *)
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   194
        | 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
   195
        | Type (s, _) =>
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   196
          (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
   197
             constrs as _ :: _ =>
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   198
             let
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   199
               val constr_cards =
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   200
                 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
   201
                      o snd) constrs
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   202
             in
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   203
               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
   204
               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
   205
             end
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   206
           | [] =>
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   207
             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
   208
               ({abs_type, rep_type, ...}, _) :: _ =>
45299
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   209
               if not sound then
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   210
                 (* We cheat here by assuming that typedef types are infinite if
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   211
                    their underlying type is infinite. This is unsound in
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   212
                    general but it's hard to think of a realistic example where
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   213
                    this would not be the case. We are also slack with
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   214
                    representation types: If a representation type has the form
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   215
                    "sigma => tau", we consider it enough to check "sigma" for
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   216
                    infiniteness. *)
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   217
                 (case varify_and_instantiate_type ctxt
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   218
                           (Logic.varifyT_global abs_type) T
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   219
                           (Logic.varifyT_global rep_type)
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   220
                       |> aux true avoid of
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   221
                    0 => 0
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   222
                  | 1 => 1
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   223
                  | _ => default_card)
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   224
               else
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   225
                 default_card
44392
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   226
             | [] => default_card)
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   227
          (* 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
   228
             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
   229
             likely have used "unit" directly anyway.) *)
44500
dbd98b549597 make default unsound mode less unsound
blanchet
parents: 44491
diff changeset
   230
        | TFree _ =>
dbd98b549597 make default unsound mode less unsound
blanchet
parents: 44491
diff changeset
   231
          if not sound andalso 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
   232
        | TVar _ => default_card
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   233
  in Int.min (max, aux false [] T) end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   234
44500
dbd98b549597 make default unsound mode less unsound
blanchet
parents: 44491
diff changeset
   235
fun is_type_surely_finite ctxt T = tiny_card_of_type ctxt true [] 0 T <> 0
dbd98b549597 make default unsound mode less unsound
blanchet
parents: 44491
diff changeset
   236
fun is_type_surely_infinite ctxt sound infinite_Ts T =
dbd98b549597 make default unsound mode less unsound
blanchet
parents: 44491
diff changeset
   237
  tiny_card_of_type ctxt sound (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
   238
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
   239
(* 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
   240
   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
   241
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
   242
    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
   243
  | 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
   244
    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
   245
  | 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
   246
  | 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
   247
    @{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
   248
  | 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
   249
    @{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
   250
  | 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
   251
  | 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
   252
  | 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
   253
  | 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
   254
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
   255
  | 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
   256
  | 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
   257
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
   258
  | 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
   259
  | 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
   260
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
   261
  | 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
   262
  | 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
   263
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
   264
  | 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
   265
  | 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
   266
46385
0ccf458a3633 third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
blanchet
parents: 45896
diff changeset
   267
val close_form_prefix = "ATP.close_form."
0ccf458a3633 third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
blanchet
parents: 45896
diff changeset
   268
43864
58a7b3fdc193 fixed lambda-liftg: must ensure the formulas are in close form
blanchet
parents: 43863
diff changeset
   269
fun close_form t =
45570
6d95a66cce00 pull variables (Var) out of lambdas, so that the Isabelle theorems closely mirror the Metis lambda-lifted ones
blanchet
parents: 45511
diff changeset
   270
  fold (fn ((s, i), T) => fn t' =>
45511
9b0f8ca4388e continued implementation of lambda-lifting in Metis
blanchet
parents: 45299
diff changeset
   271
           HOLogic.all_const T
46385
0ccf458a3633 third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
blanchet
parents: 45896
diff changeset
   272
           $ Abs (close_form_prefix ^ s, T,
0ccf458a3633 third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
blanchet
parents: 45896
diff changeset
   273
                  abstract_over (Var ((s, i), T), t')))
43864
58a7b3fdc193 fixed lambda-liftg: must ensure the formulas are in close form
blanchet
parents: 43863
diff changeset
   274
       (Term.add_vars t []) t
58a7b3fdc193 fixed lambda-liftg: must ensure the formulas are in close form
blanchet
parents: 43863
diff changeset
   275
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
   276
fun monomorphic_term subst =
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   277
  map_types (map_type_tvar (fn v =>
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   278
      case Type.lookup subst v of
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   279
        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
   280
      | NONE => TVar v))
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   281
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   282
fun eta_expand _ t 0 = t
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   283
  | eta_expand Ts (Abs (s, T, t')) n =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   284
    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
   285
  | eta_expand Ts t n =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   286
    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
   287
             (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
   288
             (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
   289
47715
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   290
fun is_fun_equality (@{const_name HOL.eq},
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   291
                     Type (_, [Type (@{type_name fun}, _), _])) = true
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   292
  | is_fun_equality _ = false
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   293
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   294
fun extensionalize_term ctxt t =
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   295
  if exists_Const is_fun_equality t then
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   296
    let val thy = Proof_Context.theory_of ctxt in
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   297
      t |> cterm_of thy |> Meson.extensionalize_conv ctxt
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   298
        |> prop_of |> Logic.dest_equals |> snd
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   299
    end
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   300
  else
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   301
    t
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   302
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   303
(* 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
   304
   predicate variable. Leaves other theorems unchanged. We simply instantiate
44460
blanchet
parents: 44399
diff changeset
   305
   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
   306
   "Meson_Clausify".) *)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   307
fun transform_elim_prop t =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   308
  case Logic.strip_imp_concl t of
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   309
    @{const Trueprop} $ Var (z, @{typ bool}) =>
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   310
    subst_Vars [(z, @{const False})] t
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   311
  | 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
   312
  | _ => t
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   313
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   314
fun specialize_type thy (s, T) t =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   315
  let
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   316
    fun subst_for (Const (s', T')) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   317
      if s = s' then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   318
        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
   319
        handle Type.TYPE_MATCH => NONE
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   320
      else
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   321
        NONE
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   322
    | subst_for (t1 $ t2) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   323
      (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
   324
    | subst_for (Abs (_, _, t')) = subst_for t'
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   325
    | subst_for _ = NONE
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   326
  in
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   327
    case subst_for t of
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   328
      SOME subst => monomorphic_term subst t
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   329
    | NONE => raise Type.TYPE_MATCH
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   330
  end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   331
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   332
fun strip_subgoal ctxt goal i =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   333
  let
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   334
    val (t, (frees, params)) =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   335
      Logic.goal_params (prop_of goal) i
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   336
      ||> (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
   337
    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
   338
    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
   339
  in (rev params, hyp_ts, concl_t) end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   340
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   341
end;