src/HOL/Tools/ATP/atp_util.ML
author nipkow
Thu, 26 Jun 2025 17:30:33 +0200
changeset 82772 59b937edcff8
parent 81519 cdc43c0fdbfc
permissions -rw-r--r--
merged
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
79799
2746dfc9ceae optional cartouche syntax and proper name printing in atp Isar output
Simon Wimmer <wimmers@in.tum.de>
parents: 77918
diff changeset
     9
  val proof_cartouches: bool Config.T
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    10
  val timestamp : unit -> string
53514
fa5b34ffe4a4 sorted out dependencies
blanchet
parents: 53505
diff changeset
    11
  val hashw : word * word -> word
fa5b34ffe4a4 sorted out dependencies
blanchet
parents: 53505
diff changeset
    12
  val hashw_string : string * word -> word
43827
62d64709af3b added option to control which lambda translation to use (for experiments)
blanchet
parents: 43572
diff changeset
    13
  val hash_string : string -> int
48323
7b5f7ca25d17 optimized MaSh output by chunking it
blanchet
parents: 48316
diff changeset
    14
  val chunk_list : int -> 'a list -> 'a list list
48251
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48247
diff changeset
    15
  val stringN_of_int : int -> int -> string
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    16
  val strip_spaces : bool -> (char -> bool) -> string -> string
44784
blanchet
parents: 44500
diff changeset
    17
  val strip_spaces_except_between_idents : string -> string
48316
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48251
diff changeset
    18
  val elide_string : int -> string -> string
52077
788b27dfaefa parse agsyHOL proofs (as unsat cores)
blanchet
parents: 52076
diff changeset
    19
  val find_enclosed : string -> string -> string -> string list
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    20
  val nat_subscript : int -> string
52076
bfa28e1cba77 freeze types in Sledgehammer goal, not just terms
blanchet
parents: 52031
diff changeset
    21
  val unquote_tvar : string -> string
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
    22
  val content_of_pretty : Pretty.T -> string
79799
2746dfc9ceae optional cartouche syntax and proper name printing in atp Isar output
Simon Wimmer <wimmers@in.tum.de>
parents: 77918
diff changeset
    23
  val maybe_quote : Proof.context -> string -> string
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
    24
  val pretty_maybe_quote : Proof.context -> Pretty.T -> Pretty.T
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51209
diff changeset
    25
  val string_of_ext_time : bool * Time.time -> string
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51209
diff changeset
    26
  val string_of_time : Time.time -> string
47150
6df6e56fd7cd tuning (in particular, Symtab instead of AList)
blanchet
parents: 46711
diff changeset
    27
  val type_instance : theory -> typ -> typ -> bool
6df6e56fd7cd tuning (in particular, Symtab instead of AList)
blanchet
parents: 46711
diff changeset
    28
  val type_generalization : theory -> typ -> typ -> bool
6df6e56fd7cd tuning (in particular, Symtab instead of AList)
blanchet
parents: 46711
diff changeset
    29
  val type_intersect : theory -> typ -> typ -> bool
6df6e56fd7cd tuning (in particular, Symtab instead of AList)
blanchet
parents: 46711
diff changeset
    30
  val type_equiv : theory -> typ * typ -> bool
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    31
  val varify_type : Proof.context -> typ -> typ
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    32
  val instantiate_type : theory -> typ -> typ -> typ -> typ
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    33
  val varify_and_instantiate_type : Proof.context -> typ -> typ -> typ -> typ
44393
23adec5984f1 make sound mode more sound (and clean up code)
blanchet
parents: 44392
diff changeset
    34
  val is_type_surely_finite : Proof.context -> typ -> bool
44399
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
    35
  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
    36
  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
    37
  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
    38
  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
    39
  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
    40
  val s_iff : term * term -> term
49983
33e18e9916a8 use metaquantification when possible in Isar proofs
blanchet
parents: 49982
diff changeset
    41
  val close_form : term -> term
49982
724cfe013182 tuned code
blanchet
parents: 48902
diff changeset
    42
  val hol_close_form_prefix : string
724cfe013182 tuned code
blanchet
parents: 48902
diff changeset
    43
  val hol_close_form : term -> term
724cfe013182 tuned code
blanchet
parents: 48902
diff changeset
    44
  val hol_open_form : (string -> string) -> term -> term
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    45
  val eta_expand : typ list -> term -> int -> term
59632
5980e75a204e clarified context;
wenzelm
parents: 59621
diff changeset
    46
  val cong_extensionalize_term : Proof.context -> term -> term
47953
a2c3706c4cb1 added "ext_cong_neq" lemma (not used yet); tuning
blanchet
parents: 47718
diff changeset
    47
  val abs_extensionalize_term : Proof.context -> term -> term
47991
3eb598b044ad make Nitpick's handling of definitions more robust in the face of formulas that don't have the expected format (needed for soundness, cf. RNG100+1)
blanchet
parents: 47954
diff changeset
    48
  val unextensionalize_def : term -> term
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    49
  val transform_elim_prop : term -> term
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    50
  val specialize_type : theory -> (string * typ) -> term -> term
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 54768
diff changeset
    51
  val strip_subgoal : thm -> int -> Proof.context -> (string * typ) list * term list * term
57541
147e3f1e0459 lambda-lifting for Z3 Isar proofs
blanchet
parents: 56126
diff changeset
    52
  val extract_lambda_def : (term -> string * typ) -> term -> string * term
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 54768
diff changeset
    53
  val short_thm_name : Proof.context -> thm -> string
74328
404ce20efc4c proper constants in TPTP $let binding
desharna
parents: 69597
diff changeset
    54
  val map_prod : ('a -> 'b) -> ('c -> 'd) -> 'a * 'c -> 'b * 'd
75005
4106bc2a9cc8 optimized app_op_level selection in TPTP generation
desharna
parents: 74379
diff changeset
    55
  val compare_length_with : 'a list -> int -> order
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
    56
  val scan_and_trace_multi_thm : Context.generic * Token.T list ->
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
    57
    (thm list * Token.T list) * (Context.generic * Token.T list)
77918
55b81d14a1b8 tuned; avoided intermediate list and list traversal
desharna
parents: 77430
diff changeset
    58
  val forall2 : ('a -> 'b -> bool) -> 'a list -> 'b list -> bool
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    59
end;
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    60
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    61
structure ATP_Util : ATP_UTIL =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    62
struct
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    63
79799
2746dfc9ceae optional cartouche syntax and proper name printing in atp Isar output
Simon Wimmer <wimmers@in.tum.de>
parents: 77918
diff changeset
    64
val proof_cartouches = Attrib.setup_config_bool \<^binding>\<open>atp_proof_cartouches\<close> (K false)
2746dfc9ceae optional cartouche syntax and proper name printing in atp Isar output
Simon Wimmer <wimmers@in.tum.de>
parents: 77918
diff changeset
    65
77918
55b81d14a1b8 tuned; avoided intermediate list and list traversal
desharna
parents: 77430
diff changeset
    66
fun forall2 _ [] [] = true
55b81d14a1b8 tuned; avoided intermediate list and list traversal
desharna
parents: 77430
diff changeset
    67
  | forall2 P (x :: xs) (y :: ys) = P x y andalso forall2 P xs ys
55b81d14a1b8 tuned; avoided intermediate list and list traversal
desharna
parents: 77430
diff changeset
    68
  | forall2 _ _ _ = false
55b81d14a1b8 tuned; avoided intermediate list and list traversal
desharna
parents: 77430
diff changeset
    69
53800
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53514
diff changeset
    70
fun timestamp_format time =
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53514
diff changeset
    71
  Date.fmt "%Y-%m-%d %H:%M:%S." (Date.fromTimeLocal time) ^
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53514
diff changeset
    72
  (StringCvt.padLeft #"0" 3 (string_of_int (Time.toMilliseconds time - 1000 * Time.toSeconds time)))
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53514
diff changeset
    73
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53514
diff changeset
    74
val timestamp = timestamp_format o Time.now
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    75
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    76
(* This hash function is recommended in "Compilers: Principles, Techniques, and
77430
51dac6fcdd0e reverted 0506c3273814 -- the message is still useful
blanchet
parents: 75125
diff changeset
    77
   Tools" by Aho, Sethi, and Ullman. *)
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    78
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
    79
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
    80
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
    81
fun hash_string s = Word.toInt (hashw_string (s, 0w0))
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    82
48323
7b5f7ca25d17 optimized MaSh output by chunking it
blanchet
parents: 48316
diff changeset
    83
fun chunk_list _ [] = []
7b5f7ca25d17 optimized MaSh output by chunking it
blanchet
parents: 48316
diff changeset
    84
  | chunk_list k xs =
7b5f7ca25d17 optimized MaSh output by chunking it
blanchet
parents: 48316
diff changeset
    85
    let val (xs1, xs2) = chop k xs in xs1 :: chunk_list k xs2 end
7b5f7ca25d17 optimized MaSh output by chunking it
blanchet
parents: 48316
diff changeset
    86
48251
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48247
diff changeset
    87
fun stringN_of_int 0 _ = ""
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48247
diff changeset
    88
  | stringN_of_int k n =
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48247
diff changeset
    89
    stringN_of_int (k - 1) (n div 10) ^ string_of_int (n mod 10)
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48247
diff changeset
    90
57728
c21e7bdb40ad remove YXML formatting when parsing backquoted facts supplied manually to Sledgehammer
blanchet
parents: 57547
diff changeset
    91
fun is_spaceish c = Char.isSpace c orelse c = #"\127" (* DEL -- no idea where these come from *)
c21e7bdb40ad remove YXML formatting when parsing backquoted facts supplied manually to Sledgehammer
blanchet
parents: 57547
diff changeset
    92
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
    93
fun strip_spaces skip_comments is_evil =
44935
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    94
  let
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    95
    fun strip_c_style_comment [] accum = accum
57728
c21e7bdb40ad remove YXML formatting when parsing backquoted facts supplied manually to Sledgehammer
blanchet
parents: 57547
diff changeset
    96
      | strip_c_style_comment (#"*" :: #"/" :: cs) accum = strip_spaces_in_list true cs accum
44935
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    97
      | strip_c_style_comment (_ :: cs) accum = strip_c_style_comment cs accum
48766
553ad5f99968 fixed "double rev" bug that arose in situations where a % comment arose on the last line of a file without \n at the end
blanchet
parents: 48323
diff changeset
    98
    and strip_spaces_in_list _ [] accum = accum
44935
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
    99
      | strip_spaces_in_list true (#"%" :: cs) accum =
67522
9e712280cc37 clarified take/drop/chop prefix/suffix;
wenzelm
parents: 66020
diff changeset
   100
        strip_spaces_in_list true (cs |> drop_prefix (not_equal #"\n")) accum
57728
c21e7bdb40ad remove YXML formatting when parsing backquoted facts supplied manually to Sledgehammer
blanchet
parents: 57547
diff changeset
   101
      | strip_spaces_in_list true (#"/" :: #"*" :: cs) accum = strip_c_style_comment cs accum
c21e7bdb40ad remove YXML formatting when parsing backquoted facts supplied manually to Sledgehammer
blanchet
parents: 57547
diff changeset
   102
      | strip_spaces_in_list _ [c1] accum = accum |> not (is_spaceish c1) ? cons c1
44935
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
   103
      | strip_spaces_in_list skip_comments (cs as [_, _]) accum =
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
   104
        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
   105
      | strip_spaces_in_list skip_comments (c1 :: c2 :: c3 :: cs) accum =
57728
c21e7bdb40ad remove YXML formatting when parsing backquoted facts supplied manually to Sledgehammer
blanchet
parents: 57547
diff changeset
   106
        if is_spaceish c1 then
44935
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
   107
          strip_spaces_in_list skip_comments (c2 :: c3 :: cs) accum
57728
c21e7bdb40ad remove YXML formatting when parsing backquoted facts supplied manually to Sledgehammer
blanchet
parents: 57547
diff changeset
   108
        else if is_spaceish c2 then
c21e7bdb40ad remove YXML formatting when parsing backquoted facts supplied manually to Sledgehammer
blanchet
parents: 57547
diff changeset
   109
          if is_spaceish c3 then
44935
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
   110
            strip_spaces_in_list skip_comments (c1 :: c3 :: cs) accum
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
   111
          else
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
   112
            strip_spaces_in_list skip_comments (c3 :: cs)
57728
c21e7bdb40ad remove YXML formatting when parsing backquoted facts supplied manually to Sledgehammer
blanchet
parents: 57547
diff changeset
   113
              (c1 :: accum |> forall is_evil [c1, c3] ? cons #" ")
44935
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
   114
        else
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
   115
          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
   116
  in
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
   117
    String.explode
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
   118
    #> rpair [] #-> strip_spaces_in_list skip_comments
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
   119
    #> rev #> String.implode
2e812384afa8 tail recursive proof preprocessing (needed for huge proofs)
blanchet
parents: 44893
diff changeset
   120
  end
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   121
44784
blanchet
parents: 44500
diff changeset
   122
fun is_ident_char c = Char.isAlphaNum c orelse c = #"_"
blanchet
parents: 44500
diff changeset
   123
val strip_spaces_except_between_idents = strip_spaces true is_ident_char
blanchet
parents: 44500
diff changeset
   124
48316
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48251
diff changeset
   125
fun elide_string threshold s =
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48251
diff changeset
   126
  if size s > threshold then
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48251
diff changeset
   127
    String.extract (s, 0, SOME (threshold div 2 - 5)) ^ " ...... " ^
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48251
diff changeset
   128
    String.extract (s, size s - (threshold + 1) div 2 + 6, NONE)
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48251
diff changeset
   129
  else
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48251
diff changeset
   130
    s
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48251
diff changeset
   131
52077
788b27dfaefa parse agsyHOL proofs (as unsat cores)
blanchet
parents: 52076
diff changeset
   132
fun find_enclosed left right s =
788b27dfaefa parse agsyHOL proofs (as unsat cores)
blanchet
parents: 52076
diff changeset
   133
  case first_field left s of
788b27dfaefa parse agsyHOL proofs (as unsat cores)
blanchet
parents: 52076
diff changeset
   134
    SOME (_, s) =>
788b27dfaefa parse agsyHOL proofs (as unsat cores)
blanchet
parents: 52076
diff changeset
   135
    (case first_field right s of
788b27dfaefa parse agsyHOL proofs (as unsat cores)
blanchet
parents: 52076
diff changeset
   136
       SOME (enclosed, s) => enclosed :: find_enclosed left right s
788b27dfaefa parse agsyHOL proofs (as unsat cores)
blanchet
parents: 52076
diff changeset
   137
     | NONE => [])
788b27dfaefa parse agsyHOL proofs (as unsat cores)
blanchet
parents: 52076
diff changeset
   138
  | NONE => []
788b27dfaefa parse agsyHOL proofs (as unsat cores)
blanchet
parents: 52076
diff changeset
   139
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52196
diff changeset
   140
val subscript = implode o map (prefix "\<^sub>") o raw_explode  (* FIXME Symbol.explode (?) *)
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   141
fun nat_subscript n =
66020
a31760eee09d discontinued obsolete print mode;
wenzelm
parents: 61770
diff changeset
   142
  n |> string_of_int |> not (print_mode_active Print_Mode.ASCII) ? subscript
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   143
52076
bfa28e1cba77 freeze types in Sledgehammer goal, not just terms
blanchet
parents: 52031
diff changeset
   144
val unquote_tvar = perhaps (try (unprefix "'"))
bfa28e1cba77 freeze types in Sledgehammer goal, not just terms
blanchet
parents: 52031
diff changeset
   145
val unquery_var = perhaps (try (unprefix "?"))
bfa28e1cba77 freeze types in Sledgehammer goal, not just terms
blanchet
parents: 52031
diff changeset
   146
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
   147
(* unformatted string without markup *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
   148
val content_of_pretty = Protocol_Message.clean_output o Pretty.unformatted_string_of
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
   149
50239
fb579401dc26 tuned signature;
wenzelm
parents: 49983
diff changeset
   150
val is_long_identifier = forall Symbol_Pos.is_identifier o Long_Name.explode
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
   151
fun gen_maybe_quote content_of cartouche quote ctxt y =
79799
2746dfc9ceae optional cartouche syntax and proper name printing in atp Isar output
Simon Wimmer <wimmers@in.tum.de>
parents: 77918
diff changeset
   152
  let
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
   153
    val s = content_of y
79825
wenzelm
parents: 79799
diff changeset
   154
    val is_literal = Keyword.is_literal (Thy_Header.get_keywords' ctxt)
wenzelm
parents: 79799
diff changeset
   155
    val q = if Config.get ctxt proof_cartouches then cartouche else quote
79799
2746dfc9ceae optional cartouche syntax and proper name printing in atp Isar output
Simon Wimmer <wimmers@in.tum.de>
parents: 77918
diff changeset
   156
  in
52076
bfa28e1cba77 freeze types in Sledgehammer goal, not just terms
blanchet
parents: 52031
diff changeset
   157
    y |> ((not (is_long_identifier (unquote_tvar s)) andalso
bfa28e1cba77 freeze types in Sledgehammer goal, not just terms
blanchet
parents: 52031
diff changeset
   158
           not (is_long_identifier (unquery_var s))) orelse
79825
wenzelm
parents: 79799
diff changeset
   159
           is_literal s) ? q
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   160
  end
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
   161
val maybe_quote = gen_maybe_quote Protocol_Message.clean_output cartouche quote
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
   162
val pretty_maybe_quote = gen_maybe_quote content_of_pretty Pretty.cartouche Pretty.quote
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   163
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51209
diff changeset
   164
fun string_of_ext_time (plus, time) =
58081
aa239fee063a show microseconds as well (useful when playing with Isar proofs)
blanchet
parents: 57728
diff changeset
   165
  let val us = Time.toMicroseconds time in
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   166
    (if plus then "> " else "") ^
58088
f9e4a9621c75 prefer '0.2 ms' to '249 \<mu>s'
blanchet
parents: 58086
diff changeset
   167
    (if us < 1000 then string_of_real (Real.fromInt (us div 100) / 10.0) ^ " ms"
58081
aa239fee063a show microseconds as well (useful when playing with Isar proofs)
blanchet
parents: 57728
diff changeset
   168
     else if us < 1000000 then signed_string_of_int (us div 1000) ^ " ms"
58088
f9e4a9621c75 prefer '0.2 ms' to '249 \<mu>s'
blanchet
parents: 58086
diff changeset
   169
     else string_of_real (Real.fromInt (us div 100000) / 10.0) ^ " s")
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   170
  end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   171
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51209
diff changeset
   172
val string_of_time = string_of_ext_time o pair false
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   173
47150
6df6e56fd7cd tuning (in particular, Symtab instead of AList)
blanchet
parents: 46711
diff changeset
   174
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
   175
fun type_generalization thy T T' = Sign.typ_instance thy (T', T)
48247
8f37d2ddabc8 optimized type intersection, hoping this will reduce the number of sudden Interrupts in the "incr_tvar" code
blanchet
parents: 48238
diff changeset
   176
8f37d2ddabc8 optimized type intersection, hoping this will reduce the number of sudden Interrupts in the "incr_tvar" code
blanchet
parents: 48238
diff changeset
   177
fun type_intersect _ (TVar _) _ = true
8f37d2ddabc8 optimized type intersection, hoping this will reduce the number of sudden Interrupts in the "incr_tvar" code
blanchet
parents: 48238
diff changeset
   178
  | type_intersect _ _ (TVar _) = true
8f37d2ddabc8 optimized type intersection, hoping this will reduce the number of sudden Interrupts in the "incr_tvar" code
blanchet
parents: 48238
diff changeset
   179
  | type_intersect thy T T' =
8f37d2ddabc8 optimized type intersection, hoping this will reduce the number of sudden Interrupts in the "incr_tvar" code
blanchet
parents: 48238
diff changeset
   180
    let
8f37d2ddabc8 optimized type intersection, hoping this will reduce the number of sudden Interrupts in the "incr_tvar" code
blanchet
parents: 48238
diff changeset
   181
      val tvars = Term.add_tvar_namesT T []
8f37d2ddabc8 optimized type intersection, hoping this will reduce the number of sudden Interrupts in the "incr_tvar" code
blanchet
parents: 48238
diff changeset
   182
      val tvars' = Term.add_tvar_namesT T' []
50968
3686bc0d4df2 pass correct index to "Sign.typ_unify" -- this is important to avoid what appears to be an infinite loop in the unifier
blanchet
parents: 50239
diff changeset
   183
      val maxidx' = maxidx_of_typ T'
48247
8f37d2ddabc8 optimized type intersection, hoping this will reduce the number of sudden Interrupts in the "incr_tvar" code
blanchet
parents: 48238
diff changeset
   184
      val T =
50968
3686bc0d4df2 pass correct index to "Sign.typ_unify" -- this is important to avoid what appears to be an infinite loop in the unifier
blanchet
parents: 50239
diff changeset
   185
        T |> exists (member (op =) tvars') tvars ? Logic.incr_tvar (maxidx' + 1)
3686bc0d4df2 pass correct index to "Sign.typ_unify" -- this is important to avoid what appears to be an infinite loop in the unifier
blanchet
parents: 50239
diff changeset
   186
      val maxidx = Integer.max (maxidx_of_typ T) maxidx'
3686bc0d4df2 pass correct index to "Sign.typ_unify" -- this is important to avoid what appears to be an infinite loop in the unifier
blanchet
parents: 50239
diff changeset
   187
    in can (Sign.typ_unify thy (T, T')) (Vartab.empty, maxidx) end
48247
8f37d2ddabc8 optimized type intersection, hoping this will reduce the number of sudden Interrupts in the "incr_tvar" code
blanchet
parents: 48238
diff changeset
   188
47150
6df6e56fd7cd tuning (in particular, Symtab instead of AList)
blanchet
parents: 46711
diff changeset
   189
val type_equiv = Sign.typ_equiv
44399
cd1e32b8d4c4 added caching for (in)finiteness checks
blanchet
parents: 44393
diff changeset
   190
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   191
fun varify_type ctxt T =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   192
  Variable.polymorphic_types ctxt [Const (\<^const_name>\<open>undefined\<close>, T)]
80636
4041e7c8059d tuned: more explicit dest_Const_name and dest_Const_type;
wenzelm
parents: 80306
diff changeset
   193
  |> snd |> the_single |> dest_Const_type
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   194
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   195
(* TODO: use "Term_Subst.instantiateT" instead? *)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   196
fun instantiate_type thy T1 T1' T2 =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   197
  Same.commit (Envir.subst_type_same
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   198
                   (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
   199
  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
   200
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   201
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
   202
  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
   203
    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
   204
  end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   205
54554
b8d0d8407c3b eliminated Sledgehammer's dependency on old-style datatypes
blanchet
parents: 53800
diff changeset
   206
fun free_constructors_of ctxt (Type (s, Ts)) =
b8d0d8407c3b eliminated Sledgehammer's dependency on old-style datatypes
blanchet
parents: 53800
diff changeset
   207
    (case Ctr_Sugar.ctr_sugar_of ctxt s of
b8d0d8407c3b eliminated Sledgehammer's dependency on old-style datatypes
blanchet
parents: 53800
diff changeset
   208
      SOME {ctrs, ...} => map_filter (try dest_Const o Ctr_Sugar.mk_ctr Ts) ctrs
b8d0d8407c3b eliminated Sledgehammer's dependency on old-style datatypes
blanchet
parents: 53800
diff changeset
   209
    | NONE => [])
b8d0d8407c3b eliminated Sledgehammer's dependency on old-style datatypes
blanchet
parents: 53800
diff changeset
   210
  | free_constructors_of _ _ = []
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   211
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   212
(* 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
   213
   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
   214
   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
   215
   cardinality of the type can't be determined. *)
44500
dbd98b549597 make default unsound mode less unsound
blanchet
parents: 44491
diff changeset
   216
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
   217
  let
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   218
    val thy = Proof_Context.theory_of ctxt
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   219
    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
   220
    fun aux slack avoid T =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   221
      if member (op =) avoid T then
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   222
        0
47150
6df6e56fd7cd tuning (in particular, Symtab instead of AList)
blanchet
parents: 46711
diff changeset
   223
      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
   224
        SOME k => k
23adec5984f1 make sound mode more sound (and clean up code)
blanchet
parents: 44392
diff changeset
   225
      | NONE =>
44392
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   226
        case T of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   227
          Type (\<^type_name>\<open>fun\<close>, [T1, T2]) =>
44392
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   228
          (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
   229
             (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
   230
           | (0, _) => 0
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   231
           | (_, 0) => 0
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   232
           | (k1, k2) =>
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   233
             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
   234
             else Int.min (max, Integer.pow k2 k1))
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   235
        | Type (\<^type_name>\<open>set\<close>, [T']) => aux slack avoid (T' --> \<^typ>\<open>bool\<close>)
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   236
        | \<^typ>\<open>prop\<close> => 2
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   237
        | \<^typ>\<open>bool\<close> => 2 (* optimization *)
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   238
        | \<^typ>\<open>nat\<close> => 0 (* optimization *)
44392
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   239
        | 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
   240
        | Type (s, _) =>
54554
b8d0d8407c3b eliminated Sledgehammer's dependency on old-style datatypes
blanchet
parents: 53800
diff changeset
   241
          (case free_constructors_of ctxt T of
44392
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   242
             constrs as _ :: _ =>
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   243
             let
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   244
               val constr_cards =
54554
b8d0d8407c3b eliminated Sledgehammer's dependency on old-style datatypes
blanchet
parents: 53800
diff changeset
   245
                 map (Integer.prod o map (aux slack (T :: avoid)) o binder_types o snd) constrs
44392
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   246
             in
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   247
               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
   248
               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
   249
             end
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   250
           | [] =>
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   251
             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
   252
               ({abs_type, rep_type, ...}, _) :: _ =>
45299
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   253
               if not sound then
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   254
                 (* We cheat here by assuming that typedef types are infinite if
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   255
                    their underlying type is infinite. This is unsound in
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   256
                    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
   257
                    this would not be the case. We are also slack with
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   258
                    representation types: If a representation type has the form
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   259
                    "sigma => tau", we consider it enough to check "sigma" for
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   260
                    infiniteness. *)
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   261
                 (case varify_and_instantiate_type ctxt
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   262
                           (Logic.varifyT_global abs_type) T
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   263
                           (Logic.varifyT_global rep_type)
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   264
                       |> aux true avoid of
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   265
                    0 => 0
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   266
                  | 1 => 1
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   267
                  | _ => default_card)
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   268
               else
ee584ff987c3 check "sound" flag before doing something unsound...
blanchet
parents: 44935
diff changeset
   269
                 default_card
44392
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   270
             | [] => default_card)
54554
b8d0d8407c3b eliminated Sledgehammer's dependency on old-style datatypes
blanchet
parents: 53800
diff changeset
   271
        | TFree _ =>
44392
6750b4297691 reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
blanchet
parents: 43864
diff changeset
   272
          (* 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
   273
             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
   274
             likely have used "unit" directly anyway.) *)
44500
dbd98b549597 make default unsound mode less unsound
blanchet
parents: 44491
diff changeset
   275
          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
   276
        | TVar _ => default_card
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   277
  in Int.min (max, aux false [] T) end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   278
44500
dbd98b549597 make default unsound mode less unsound
blanchet
parents: 44491
diff changeset
   279
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
   280
fun is_type_surely_infinite ctxt sound infinite_Ts T =
dbd98b549597 make default unsound mode less unsound
blanchet
parents: 44491
diff changeset
   281
  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
   282
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
   283
(* 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
   284
   spurious "True"s. *)
74379
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   285
fun s_not \<^Const_>\<open>All T for \<open>Abs (s, T', t')\<close>\<close> = \<^Const>\<open>Ex T for \<open>Abs (s, T', s_not t')\<close>\<close>
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   286
  | s_not \<^Const_>\<open>Ex T for \<open>Abs (s, T', t')\<close>\<close> = \<^Const>\<open>All T for \<open>Abs (s, T', s_not t')\<close>\<close>
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   287
  | s_not \<^Const_>\<open>implies for t1 t2\<close> = \<^Const>\<open>conj for t1 \<open>s_not t2\<close>\<close>
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   288
  | s_not \<^Const_>\<open>conj for t1 t2\<close> = \<^Const>\<open>disj for \<open>s_not t1\<close> \<open>s_not t2\<close>\<close> 
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   289
  | s_not \<^Const_>\<open>disj for t1 t2\<close> = \<^Const>\<open>conj for \<open>s_not t1\<close> \<open>s_not t2\<close>\<close>
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   290
  | s_not \<^Const_>\<open>False\<close> = \<^Const>\<open>True\<close>
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   291
  | s_not \<^Const_>\<open>True\<close> = \<^Const>\<open>False\<close>
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   292
  | s_not \<^Const_>\<open>Not for t\<close> = t
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   293
  | s_not t = \<^Const>\<open>Not for t\<close>
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   294
74379
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   295
fun s_conj (\<^Const_>\<open>True\<close>, t2) = t2
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   296
  | s_conj (t1, \<^Const_>\<open>True\<close>) = t1
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   297
  | s_conj (\<^Const_>\<open>False\<close>, _) = \<^Const>\<open>False\<close>
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   298
  | s_conj (_, \<^Const_>\<open>False\<close>) = \<^Const>\<open>False\<close>
51197
1c6031e5d284 optimize Isar output some more
blanchet
parents: 50968
diff changeset
   299
  | s_conj (t1, t2) = if t1 aconv t2 then t1 else HOLogic.mk_conj (t1, t2)
54758
ba488d89614a simplify generated propositions
blanchet
parents: 54757
diff changeset
   300
74379
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   301
fun s_disj (\<^Const_>\<open>False\<close>, t2) = t2
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   302
  | s_disj (t1, \<^Const_>\<open>False\<close>) = t1
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   303
  | s_disj (\<^Const_>\<open>True\<close>, _) = \<^Const>\<open>True\<close>
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   304
  | s_disj (_, \<^Const_>\<open>True\<close>) = \<^Const>\<open>True\<close>
51197
1c6031e5d284 optimize Isar output some more
blanchet
parents: 50968
diff changeset
   305
  | s_disj (t1, t2) = if t1 aconv t2 then t1 else HOLogic.mk_disj (t1, t2)
54757
4960647932ec use 'prop' rather than 'bool' systematically in Isar reconstruction code
blanchet
parents: 54554
diff changeset
   306
74379
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   307
fun s_imp (\<^Const_>\<open>True\<close>, t2) = t2
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   308
  | s_imp (t1, \<^Const_>\<open>False\<close>) = s_not t1
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   309
  | s_imp (\<^Const_>\<open>False\<close>, _) = \<^Const>\<open>True\<close>
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   310
  | s_imp (_, \<^Const_>\<open>True\<close>) = \<^Const>\<open>True\<close>
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
   311
  | s_imp p = HOLogic.mk_imp p
54757
4960647932ec use 'prop' rather than 'bool' systematically in Isar reconstruction code
blanchet
parents: 54554
diff changeset
   312
74379
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   313
fun s_iff (\<^Const_>\<open>True\<close>, t2) = t2
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   314
  | s_iff (t1, \<^Const_>\<open>True\<close>) = t1
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   315
  | s_iff (\<^Const_>\<open>False\<close>, t2) = s_not t2
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   316
  | s_iff (t1, \<^Const_>\<open>False\<close>) = s_not t1
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   317
  | s_iff (t1, t2) = if t1 aconv t2 then \<^Const>\<open>True\<close> else HOLogic.eq_const HOLogic.boolT $ t1 $ t2
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
   318
49983
33e18e9916a8 use metaquantification when possible in Isar proofs
blanchet
parents: 49982
diff changeset
   319
fun close_form t =
33e18e9916a8 use metaquantification when possible in Isar proofs
blanchet
parents: 49982
diff changeset
   320
  fold (fn ((s, i), T) => fn t' =>
54757
4960647932ec use 'prop' rather than 'bool' systematically in Isar reconstruction code
blanchet
parents: 54554
diff changeset
   321
      Logic.all_const T $ Abs (s, T, abstract_over (Var ((s, i), T), t')))
4960647932ec use 'prop' rather than 'bool' systematically in Isar reconstruction code
blanchet
parents: 54554
diff changeset
   322
    (Term.add_vars t []) t
49983
33e18e9916a8 use metaquantification when possible in Isar proofs
blanchet
parents: 49982
diff changeset
   323
58091
ecf5826ba234 reworked unskolemization for SPASS
blanchet
parents: 58088
diff changeset
   324
val hol_close_form_prefix = "ATP."
46385
0ccf458a3633 third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
blanchet
parents: 45896
diff changeset
   325
49982
724cfe013182 tuned code
blanchet
parents: 48902
diff changeset
   326
fun hol_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
   327
  fold (fn ((s, i), T) => fn t' =>
45511
9b0f8ca4388e continued implementation of lambda-lifting in Metis
blanchet
parents: 45299
diff changeset
   328
           HOLogic.all_const T
49982
724cfe013182 tuned code
blanchet
parents: 48902
diff changeset
   329
           $ Abs (hol_close_form_prefix ^ s, T,
46385
0ccf458a3633 third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
blanchet
parents: 45896
diff changeset
   330
                  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
   331
       (Term.add_vars t []) t
58a7b3fdc193 fixed lambda-liftg: must ensure the formulas are in close form
blanchet
parents: 43863
diff changeset
   332
49982
724cfe013182 tuned code
blanchet
parents: 48902
diff changeset
   333
fun hol_open_form unprefix
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   334
      (t as Const (\<^const_name>\<open>All\<close>, _) $ Abs (s, T, t')) =
47718
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
   335
    (case try unprefix s of
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
   336
       SOME s =>
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
   337
       let
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
   338
         val names = Name.make_context (map fst (Term.add_var_names t' []))
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
   339
         val (s, _) = Name.variant s names
49982
724cfe013182 tuned code
blanchet
parents: 48902
diff changeset
   340
       in hol_open_form unprefix (subst_bound (Var ((s, 0), T), t')) end
47718
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
   341
     | NONE => t)
49982
724cfe013182 tuned code
blanchet
parents: 48902
diff changeset
   342
  | hol_open_form _ t = t
47718
39229c760636 smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
blanchet
parents: 47715
diff changeset
   343
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   344
fun eta_expand _ t 0 = t
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   345
  | eta_expand Ts (Abs (s, T, t')) n =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   346
    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
   347
  | eta_expand Ts t n =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   348
    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
   349
             (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
   350
             (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
   351
59632
5980e75a204e clarified context;
wenzelm
parents: 59621
diff changeset
   352
fun cong_extensionalize_term ctxt t =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   353
  if exists_Const (fn (s, _) => s = \<^const_name>\<open>Not\<close>) t then
59632
5980e75a204e clarified context;
wenzelm
parents: 59621
diff changeset
   354
    t |> Skip_Proof.make_thm (Proof_Context.theory_of ctxt)
5980e75a204e clarified context;
wenzelm
parents: 59621
diff changeset
   355
      |> Meson.cong_extensionalize_thm ctxt
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59433
diff changeset
   356
      |> Thm.prop_of
47954
aada9fd08b58 make higher-order goals more first-order via extensionality
blanchet
parents: 47953
diff changeset
   357
  else
aada9fd08b58 make higher-order goals more first-order via extensionality
blanchet
parents: 47953
diff changeset
   358
    t
aada9fd08b58 make higher-order goals more first-order via extensionality
blanchet
parents: 47953
diff changeset
   359
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   360
fun is_fun_equality (\<^const_name>\<open>HOL.eq\<close>,
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   361
                     Type (_, [Type (\<^type_name>\<open>fun\<close>, _), _])) = true
47715
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   362
  | is_fun_equality _ = false
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   363
47953
a2c3706c4cb1 added "ext_cong_neq" lemma (not used yet); tuning
blanchet
parents: 47718
diff changeset
   364
fun abs_extensionalize_term ctxt t =
47715
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   365
  if exists_Const is_fun_equality t then
59632
5980e75a204e clarified context;
wenzelm
parents: 59621
diff changeset
   366
    t |> Thm.cterm_of ctxt |> Meson.abs_extensionalize_conv ctxt
5980e75a204e clarified context;
wenzelm
parents: 59621
diff changeset
   367
      |> Thm.prop_of |> Logic.dest_equals |> snd
47715
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   368
  else
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   369
    t
04400144c6fc handle TPTP definitions as definitions in Nitpick rather than as axioms
blanchet
parents: 47150
diff changeset
   370
47991
3eb598b044ad make Nitpick's handling of definitions more robust in the face of formulas that don't have the expected format (needed for soundness, cf. RNG100+1)
blanchet
parents: 47954
diff changeset
   371
fun unextensionalize_def t =
74379
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   372
  (case t of
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   373
    \<^Const_>\<open>Trueprop for \<^Const_>\<open>HOL.eq _ for lhs rhs\<close>\<close> =>
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   374
      (case strip_comb lhs of
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   375
        (c as Const (_, T), args) =>
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   376
          if forall is_Var args andalso not (has_duplicates (op =) args) then
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   377
            \<^Const>\<open>Trueprop for \<^Const>\<open>HOL.eq T for c \<open>fold_rev lambda args rhs\<close>\<close>\<close>
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   378
          else t
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   379
      | _ => t)
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   380
  | _ => t)
47991
3eb598b044ad make Nitpick's handling of definitions more robust in the face of formulas that don't have the expected format (needed for soundness, cf. RNG100+1)
blanchet
parents: 47954
diff changeset
   381
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   382
(* 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
   383
   predicate variable. Leaves other theorems unchanged. We simply instantiate
44460
blanchet
parents: 44399
diff changeset
   384
   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
   385
   "Meson_Clausify".) *)
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   386
fun transform_elim_prop t =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   387
  case Logic.strip_imp_concl t of
74379
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   388
    \<^Const_>\<open>Trueprop for \<open>Var (z, \<^typ>\<open>bool\<close>)\<close>\<close> => subst_Vars [(z, \<^Const>\<open>False\<close>)] t
9ea507f63bcb clarified antiquotations;
wenzelm
parents: 74328
diff changeset
   389
  | Var (z, \<^Type>\<open>prop\<close>) => subst_Vars [(z, \<^prop>\<open>False\<close>)] t
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   390
  | _ => t
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   391
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   392
fun specialize_type thy (s, T) t =
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   393
  let
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   394
    fun subst_for (Const (s', T')) =
61769
2cd36f4c5d65 tuned whitespace
blanchet
parents: 59632
diff changeset
   395
        if s = s' then
2cd36f4c5d65 tuned whitespace
blanchet
parents: 59632
diff changeset
   396
          SOME (Sign.typ_match thy (T', T) Vartab.empty)
2cd36f4c5d65 tuned whitespace
blanchet
parents: 59632
diff changeset
   397
          handle Type.TYPE_MATCH => NONE
2cd36f4c5d65 tuned whitespace
blanchet
parents: 59632
diff changeset
   398
        else
2cd36f4c5d65 tuned whitespace
blanchet
parents: 59632
diff changeset
   399
          NONE
2cd36f4c5d65 tuned whitespace
blanchet
parents: 59632
diff changeset
   400
      | subst_for (t1 $ t2) = (case subst_for t1 of SOME x => SOME x | NONE => subst_for t2)
2cd36f4c5d65 tuned whitespace
blanchet
parents: 59632
diff changeset
   401
      | subst_for (Abs (_, _, t')) = subst_for t'
2cd36f4c5d65 tuned whitespace
blanchet
parents: 59632
diff changeset
   402
      | subst_for _ = NONE
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   403
  in
61769
2cd36f4c5d65 tuned whitespace
blanchet
parents: 59632
diff changeset
   404
    (case subst_for t of
61770
a20048c78891 removed needless ML function
blanchet
parents: 61769
diff changeset
   405
      SOME subst => Envir.subst_term_types subst t
61769
2cd36f4c5d65 tuned whitespace
blanchet
parents: 59632
diff changeset
   406
    | NONE => raise Type.TYPE_MATCH)
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   407
  end
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   408
52125
ac7830871177 improved handling of free variables' types in Isar proofs
blanchet
parents: 52077
diff changeset
   409
fun strip_subgoal goal i ctxt =
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   410
  let
52196
2281f33e8da6 redid rac7830871177 to avoid duplicate fixed variable (e.g. lemma "P (a::nat)" proof - have "!!a::int. Q a" sledgehammer [e])
blanchet
parents: 52125
diff changeset
   411
    val (t, (frees, params)) =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59433
diff changeset
   412
      Logic.goal_params (Thm.prop_of goal) i
81519
cdc43c0fdbfc clarified signature;
wenzelm
parents: 81254
diff changeset
   413
      ||> (map dest_Free #> Variable.variant_names ctxt #> `(map Free))
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   414
    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
   415
    val concl_t = t |> Logic.strip_assums_concl |> curry subst_bounds frees
52196
2281f33e8da6 redid rac7830871177 to avoid duplicate fixed variable (e.g. lemma "P (a::nat)" proof - have "!!a::int. Q a" sledgehammer [e])
blanchet
parents: 52125
diff changeset
   416
  in (rev params, hyp_ts, concl_t) end
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   417
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67522
diff changeset
   418
fun extract_lambda_def dest_head (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ t $ u) =
57541
147e3f1e0459 lambda-lifting for Z3 Isar proofs
blanchet
parents: 56126
diff changeset
   419
    let val (head, args) = strip_comb t in
147e3f1e0459 lambda-lifting for Z3 Isar proofs
blanchet
parents: 56126
diff changeset
   420
      (head |> dest_head |> fst,
147e3f1e0459 lambda-lifting for Z3 Isar proofs
blanchet
parents: 56126
diff changeset
   421
       fold_rev (fn t as Var ((s, _), T) =>
147e3f1e0459 lambda-lifting for Z3 Isar proofs
blanchet
parents: 56126
diff changeset
   422
                    (fn u => Abs (s, T, abstract_over (t, u)))
147e3f1e0459 lambda-lifting for Z3 Isar proofs
blanchet
parents: 56126
diff changeset
   423
                  | _ => raise Fail "expected \"Var\"") args u)
147e3f1e0459 lambda-lifting for Z3 Isar proofs
blanchet
parents: 56126
diff changeset
   424
    end
147e3f1e0459 lambda-lifting for Z3 Isar proofs
blanchet
parents: 56126
diff changeset
   425
  | extract_lambda_def _ _ = raise Fail "malformed lifted lambda"
147e3f1e0459 lambda-lifting for Z3 Isar proofs
blanchet
parents: 56126
diff changeset
   426
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 54768
diff changeset
   427
fun short_thm_name ctxt th =
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 54768
diff changeset
   428
  let
80306
c2537860ccf8 more accurate thm "name_hint", using Thm_Name.T;
wenzelm
parents: 79825
diff changeset
   429
    val long = Thm_Name.short (Thm.get_name_hint th)
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 54768
diff changeset
   430
    val short = Long_Name.base_name long
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 54768
diff changeset
   431
  in
75125
18cd39e55eca have Sledgehammer honor 'smt_nat_as_int' option
blanchet
parents: 75005
diff changeset
   432
    (case try (singleton (Attrib.eval_thms ctxt)) (Facts.named short, []) of
18cd39e55eca have Sledgehammer honor 'smt_nat_as_int' option
blanchet
parents: 75005
diff changeset
   433
      SOME th' => if Thm.eq_thm_prop (th, th') then short else long
18cd39e55eca have Sledgehammer honor 'smt_nat_as_int' option
blanchet
parents: 75005
diff changeset
   434
    | _ => long)
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 54768
diff changeset
   435
  end
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 54768
diff changeset
   436
75005
4106bc2a9cc8 optimized app_op_level selection in TPTP generation
desharna
parents: 74379
diff changeset
   437
val map_prod = Ctr_Sugar_Util.map_prod
4106bc2a9cc8 optimized app_op_level selection in TPTP generation
desharna
parents: 74379
diff changeset
   438
77430
51dac6fcdd0e reverted 0506c3273814 -- the message is still useful
blanchet
parents: 75125
diff changeset
   439
(* Compare the length of a list with an integer n while traversing at most n
51dac6fcdd0e reverted 0506c3273814 -- the message is still useful
blanchet
parents: 75125
diff changeset
   440
elements of the list. *)
75005
4106bc2a9cc8 optimized app_op_level selection in TPTP generation
desharna
parents: 74379
diff changeset
   441
fun compare_length_with [] n = if n < 0 then GREATER else if n = 0 then EQUAL else LESS
4106bc2a9cc8 optimized app_op_level selection in TPTP generation
desharna
parents: 74379
diff changeset
   442
  | compare_length_with (_ :: xs) n = if n <= 0 then GREATER else compare_length_with xs (n - 1)
74328
404ce20efc4c proper constants in TPTP $let binding
desharna
parents: 69597
diff changeset
   443
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
   444
(* Scan Attrib.multi_thm and store the input tokens *)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
   445
fun scan_and_trace_multi_thm (context, toks) =
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
   446
  let
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
   447
    val (thms, (context', toks')) = Attrib.multi_thm (context, toks)
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
   448
  in
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
   449
    ((thms, take (length toks - length toks') toks), (context', toks'))
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
   450
  end
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 80820
diff changeset
   451
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents:
diff changeset
   452
end;