1 (* Title: HOL/Tools/ATP/atp_util.ML
2 Author: Jasmin Blanchette, TU Muenchen
4 General-purpose functions used by the ATP module.
9 val timestamp : unit -> string
10 val hash_string : string -> int
11 val hash_term : term -> int
12 val strip_spaces : bool -> (char -> bool) -> string -> string
13 val nat_subscript : int -> string
14 val unyxml : string -> string
15 val maybe_quote : string -> string
16 val string_from_ext_time : bool * Time.time -> string
17 val string_from_time : Time.time -> string
18 val varify_type : Proof.context -> typ -> typ
19 val instantiate_type : theory -> typ -> typ -> typ -> typ
20 val varify_and_instantiate_type : Proof.context -> typ -> typ -> typ -> typ
22 Datatype_Aux.descr -> (Datatype_Aux.dtyp * typ) list -> Datatype_Aux.dtyp
24 val is_type_surely_finite : Proof.context -> bool -> typ -> bool
25 val is_type_surely_infinite : Proof.context -> bool -> typ -> bool
26 val s_not : term -> term
27 val s_conj : term * term -> term
28 val s_disj : term * term -> term
29 val s_imp : term * term -> term
30 val s_iff : term * term -> term
31 val close_form : term -> term
32 val monomorphic_term : Type.tyenv -> term -> term
33 val eta_expand : typ list -> term -> int -> term
34 val transform_elim_prop : term -> term
35 val specialize_type : theory -> (string * typ) -> term -> term
37 Proof.context -> thm -> int -> (string * typ) list * term list * term
40 structure ATP_Util : ATP_UTIL =
43 val timestamp = Date.fmt "%Y-%m-%d %H:%M:%S" o Date.fromTimeLocal o Time.now
45 (* This hash function is recommended in "Compilers: Principles, Techniques, and
46 Tools" by Aho, Sethi, and Ullman. The "hashpjw" function, which they
47 particularly recommend, triggers a bug in versions of Poly/ML up to 4.2.0. *)
48 fun hashw (u, w) = Word.+ (u, Word.* (0w65599, w))
49 fun hashw_char (c, w) = hashw (Word.fromInt (Char.ord c), w)
50 fun hashw_string (s : string, w) = CharVector.foldl hashw_char w s
51 fun hashw_term (t1 $ t2) = hashw (hashw_term t1, hashw_term t2)
52 | hashw_term (Const (s, _)) = hashw_string (s, 0w0)
53 | hashw_term (Free (s, _)) = hashw_string (s, 0w0)
56 fun hash_string s = Word.toInt (hashw_string (s, 0w0))
57 val hash_term = Word.toInt o hashw_term
59 fun strip_c_style_comment _ [] = []
60 | strip_c_style_comment is_evil (#"*" :: #"/" :: cs) =
61 strip_spaces_in_list true is_evil cs
62 | strip_c_style_comment is_evil (_ :: cs) = strip_c_style_comment is_evil cs
63 and strip_spaces_in_list _ _ [] = []
64 | strip_spaces_in_list true is_evil (#"%" :: cs) =
65 strip_spaces_in_list true is_evil
66 (cs |> chop_while (not_equal #"\n") |> snd)
67 | strip_spaces_in_list true is_evil (#"/" :: #"*" :: cs) =
68 strip_c_style_comment is_evil cs
69 | strip_spaces_in_list _ _ [c1] = if Char.isSpace c1 then [] else [str c1]
70 | strip_spaces_in_list skip_comments is_evil [c1, c2] =
71 strip_spaces_in_list skip_comments is_evil [c1] @
72 strip_spaces_in_list skip_comments is_evil [c2]
73 | strip_spaces_in_list skip_comments is_evil (c1 :: c2 :: c3 :: cs) =
74 if Char.isSpace c1 then
75 strip_spaces_in_list skip_comments is_evil (c2 :: c3 :: cs)
76 else if Char.isSpace c2 then
77 if Char.isSpace c3 then
78 strip_spaces_in_list skip_comments is_evil (c1 :: c3 :: cs)
80 str c1 :: (if forall is_evil [c1, c3] then [" "] else []) @
81 strip_spaces_in_list skip_comments is_evil (c3 :: cs)
83 str c1 :: strip_spaces_in_list skip_comments is_evil (c2 :: c3 :: cs)
84 fun strip_spaces skip_comments is_evil =
85 implode o strip_spaces_in_list skip_comments is_evil o String.explode
87 val subscript = implode o map (prefix "\<^isub>") o raw_explode (* FIXME Symbol.explode (?) *)
89 n |> string_of_int |> print_mode_active Symbol.xsymbolsN ? subscript
91 val unyxml = XML.content_of o YXML.parse_body
93 val is_long_identifier = forall Lexicon.is_identifier o space_explode "."
95 let val s = unyxml y in
96 y |> ((not (is_long_identifier (perhaps (try (unprefix "'")) s)) andalso
97 not (is_long_identifier (perhaps (try (unprefix "?")) s))) orelse
98 Keyword.is_keyword s) ? quote
101 fun string_from_ext_time (plus, time) =
102 let val ms = Time.toMilliseconds time in
103 (if plus then "> " else "") ^
104 (if plus andalso ms mod 1000 = 0 then
105 signed_string_of_int (ms div 1000) ^ " s"
106 else if ms < 1000 then
107 signed_string_of_int ms ^ " ms"
109 string_of_real (0.01 * Real.fromInt (ms div 10)) ^ " s")
112 val string_from_time = string_from_ext_time o pair false
114 fun varify_type ctxt T =
115 Variable.polymorphic_types ctxt [Const (@{const_name undefined}, T)]
116 |> snd |> the_single |> dest_Const |> snd
118 (* TODO: use "Term_Subst.instantiateT" instead? *)
119 fun instantiate_type thy T1 T1' T2 =
120 Same.commit (Envir.subst_type_same
121 (Sign.typ_match thy (T1, T1') Vartab.empty)) T2
122 handle Type.TYPE_MATCH => raise TYPE ("instantiate_type", [T1, T1'], [])
124 fun varify_and_instantiate_type ctxt T1 T1' T2 =
125 let val thy = Proof_Context.theory_of ctxt in
126 instantiate_type thy (varify_type ctxt T1) T1' (varify_type ctxt T2)
129 fun typ_of_dtyp _ typ_assoc (Datatype_Aux.DtTFree a) =
130 the (AList.lookup (op =) typ_assoc (Datatype_Aux.DtTFree a))
131 | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtType (s, Us)) =
132 Type (s, map (typ_of_dtyp descr typ_assoc) Us)
133 | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtRec i) =
134 let val (s, ds, _) = the (AList.lookup (op =) descr i) in
135 Type (s, map (typ_of_dtyp descr typ_assoc) ds)
138 fun datatype_constrs thy (T as Type (s, Ts)) =
139 (case Datatype.get_info thy s of
140 SOME {index, descr, ...} =>
141 let val (_, dtyps, constrs) = AList.lookup (op =) descr index |> the in
142 map (apsnd (fn Us => map (typ_of_dtyp descr (dtyps ~~ Ts)) Us ---> T))
146 | datatype_constrs _ _ = []
148 (* Similar to "Nitpick_HOL.bounded_exact_card_of_type".
149 0 means infinite type, 1 means singleton type (e.g., "unit"), and 2 means
150 cardinality 2 or more. The specified default cardinality is returned if the
151 cardinality of the type can't be determined. *)
152 fun tiny_card_of_type ctxt sound default_card T =
154 val thy = Proof_Context.theory_of ctxt
155 val max = 2 (* 1 would be too small for the "fun" case *)
156 fun aux slack avoid T =
157 if member (op =) avoid T then
160 Type (@{type_name fun}, [T1, T2]) =>
161 (case (aux slack avoid T1, aux slack avoid T2) of
162 (k, 1) => if slack andalso k = 0 then 0 else 1
166 if k1 >= max orelse k2 >= max then max
167 else Int.min (max, Integer.pow k2 k1))
169 | @{typ bool} => 2 (* optimization *)
170 | @{typ nat} => 0 (* optimization *)
171 | Type ("Int.int", []) => 0 (* optimization *)
173 (case datatype_constrs thy T of
177 map (Integer.prod o map (aux slack (T :: avoid)) o binder_types
180 if exists (curry (op =) 0) constr_cards then 0
181 else Int.min (max, Integer.sum constr_cards)
184 case Typedef.get_info ctxt s of
185 ({abs_type, rep_type, ...}, _) :: _ =>
186 (* We cheat here by assuming that typedef types are infinite if
187 their underlying type is infinite. This is unsound in general
188 but it's hard to think of a realistic example where this would
189 not be the case. We are also slack with representation types:
190 If a representation type has the form "sigma => tau", we
191 consider it enough to check "sigma" for infiniteness. (Look
192 for "slack" in this function.) *)
193 (case varify_and_instantiate_type ctxt
194 (Logic.varifyT_global abs_type) T
195 (Logic.varifyT_global rep_type)
197 0 => if sound then default_card else 0
200 | [] => default_card)
201 (* Very slightly unsound: Type variables are assumed not to be
202 constrained to cardinality 1. (In practice, the user would most
203 likely have used "unit" directly anyway.) *)
205 if default_card = 1 andalso not sound then 2 else default_card
206 | TVar _ => default_card
207 in Int.min (max, aux false [] T) end
209 fun is_type_surely_finite ctxt sound T = tiny_card_of_type ctxt sound 0 T <> 0
210 fun is_type_surely_infinite ctxt sound T = tiny_card_of_type ctxt sound 1 T = 0
212 (* Simple simplifications to ensure that sort annotations don't leave a trail of
214 fun s_not (Const (@{const_name All}, T) $ Abs (s, T', t')) =
215 Const (@{const_name Ex}, T) $ Abs (s, T', s_not t')
216 | s_not (Const (@{const_name Ex}, T) $ Abs (s, T', t')) =
217 Const (@{const_name All}, T) $ Abs (s, T', s_not t')
218 | s_not (@{const HOL.implies} $ t1 $ t2) = @{const HOL.conj} $ t1 $ s_not t2
219 | s_not (@{const HOL.conj} $ t1 $ t2) =
220 @{const HOL.disj} $ s_not t1 $ s_not t2
221 | s_not (@{const HOL.disj} $ t1 $ t2) =
222 @{const HOL.conj} $ s_not t1 $ s_not t2
223 | s_not (@{const False}) = @{const True}
224 | s_not (@{const True}) = @{const False}
225 | s_not (@{const Not} $ t) = t
226 | s_not t = @{const Not} $ t
227 fun s_conj (@{const True}, t2) = t2
228 | s_conj (t1, @{const True}) = t1
229 | s_conj p = HOLogic.mk_conj p
230 fun s_disj (@{const False}, t2) = t2
231 | s_disj (t1, @{const False}) = t1
232 | s_disj p = HOLogic.mk_disj p
233 fun s_imp (@{const True}, t2) = t2
234 | s_imp (t1, @{const False}) = s_not t1
235 | s_imp p = HOLogic.mk_imp p
236 fun s_iff (@{const True}, t2) = t2
237 | s_iff (t1, @{const True}) = t1
238 | s_iff (t1, t2) = HOLogic.eq_const HOLogic.boolT $ t1 $ t2
241 fold (fn ((x, i), T) => fn t' =>
242 HOLogic.all_const T $ Abs (x, T, abstract_over (Var ((x, i), T), t')))
243 (Term.add_vars t []) t
245 fun monomorphic_term subst =
246 map_types (map_type_tvar (fn v =>
247 case Type.lookup subst v of
251 fun eta_expand _ t 0 = t
252 | eta_expand Ts (Abs (s, T, t')) n =
253 Abs (s, T, eta_expand (T :: Ts) t' (n - 1))
254 | eta_expand Ts t n =
255 fold_rev (fn T => fn t' => Abs ("x" ^ nat_subscript n, T, t'))
256 (List.take (binder_types (fastype_of1 (Ts, t)), n))
257 (list_comb (incr_boundvars n t, map Bound (n - 1 downto 0)))
259 (* Converts an elim-rule into an equivalent theorem that does not have the
260 predicate variable. Leaves other theorems unchanged. We simply instantiate
261 the conclusion variable to False. (Cf. "transform_elim_theorem" in
262 "Meson_Clausify".) *)
263 fun transform_elim_prop t =
264 case Logic.strip_imp_concl t of
265 @{const Trueprop} $ Var (z, @{typ bool}) =>
266 subst_Vars [(z, @{const False})] t
267 | Var (z, @{typ prop}) => subst_Vars [(z, @{prop False})] t
270 fun specialize_type thy (s, T) t =
272 fun subst_for (Const (s', T')) =
274 SOME (Sign.typ_match thy (T', T) Vartab.empty)
275 handle Type.TYPE_MATCH => NONE
278 | subst_for (t1 $ t2) =
279 (case subst_for t1 of SOME x => SOME x | NONE => subst_for t2)
280 | subst_for (Abs (_, _, t')) = subst_for t'
284 SOME subst => monomorphic_term subst t
285 | NONE => raise Type.TYPE_MATCH
288 fun strip_subgoal ctxt goal i =
290 val (t, (frees, params)) =
291 Logic.goal_params (prop_of goal) i
292 ||> (map dest_Free #> Variable.variant_frees ctxt [] #> `(map Free))
293 val hyp_ts = t |> Logic.strip_assums_hyp |> map (curry subst_bounds frees)
294 val concl_t = t |> Logic.strip_assums_concl |> curry subst_bounds frees
295 in (rev params, hyp_ts, concl_t) end