| author | wenzelm | 
| Tue, 29 Sep 2020 19:54:59 +0200 | |
| changeset 72339 | 626920749f5d | 
| parent 69593 | 3dda49e08b9d | 
| child 79439 | 739b1703866e | 
| permissions | -rw-r--r-- | 
| 66646 | 1  | 
(* Title: HOL/Tools/Nunchaku/nunchaku_util.ML  | 
| 
66614
 
1f1c5d85d232
moved Nunchaku to Main; the goal is to move Nitpick out in the next 1-2 years
 
blanchet 
parents: 
64389 
diff
changeset
 | 
2  | 
Author: Jasmin Blanchette, VU Amsterdam  | 
| 
 
1f1c5d85d232
moved Nunchaku to Main; the goal is to move Nitpick out in the next 1-2 years
 
blanchet 
parents: 
64389 
diff
changeset
 | 
3  | 
Copyright 2015, 2016, 2017  | 
| 64389 | 4  | 
|
5  | 
General-purpose functions used by Nunchaku.  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
signature NUNCHAKU_UTIL =  | 
|
9  | 
sig  | 
|
10  | 
val elide_string: int -> string -> string  | 
|
11  | 
val nat_subscript: int -> string  | 
|
12  | 
val timestamp: unit -> string  | 
|
13  | 
val parse_bool_option: bool -> string -> string -> bool option  | 
|
14  | 
val parse_time: string -> string -> Time.time  | 
|
15  | 
val string_of_time: Time.time -> string  | 
|
16  | 
val simplify_spaces: string -> string  | 
|
17  | 
val ascii_of: string -> string  | 
|
18  | 
val unascii_of: string -> string  | 
|
19  | 
  val double_lookup: ('a * 'a -> bool) -> ('a option * 'b) list -> 'a -> 'b option
 | 
|
20  | 
  val triple_lookup: (''a * ''a -> bool) -> (''a option * 'b) list -> ''a -> 'b option
 | 
|
21  | 
val plural_s_for_list: 'a list -> string  | 
|
22  | 
val with_overlord_file: string -> string -> (Path.T -> 'a) -> 'a  | 
|
23  | 
val with_tmp_or_overlord_file: bool -> string -> string -> (Path.T -> 'a) -> 'a  | 
|
24  | 
val num_binder_types: typ -> int  | 
|
25  | 
val strip_fun_type: typ -> typ list * typ  | 
|
26  | 
val attach_typeS: term -> term  | 
|
27  | 
val specialize_type: theory -> string * typ -> term -> term  | 
|
28  | 
val typ_match: theory -> typ * typ -> bool  | 
|
29  | 
val term_match: theory -> term * term -> bool  | 
|
30  | 
val const_match: theory -> (string * typ) * (string * typ) -> bool  | 
|
31  | 
val DETERM_TIMEOUT: Time.time -> tactic -> tactic  | 
|
32  | 
val spying: bool -> (unit -> Proof.state * int * string) -> unit  | 
|
33  | 
end;  | 
|
34  | 
||
35  | 
structure Nunchaku_Util : NUNCHAKU_UTIL =  | 
|
36  | 
struct  | 
|
37  | 
||
38  | 
val elide_string = ATP_Util.elide_string;  | 
|
39  | 
val nat_subscript = Nitpick_Util.nat_subscript;  | 
|
40  | 
val timestamp = ATP_Util.timestamp;  | 
|
41  | 
||
42  | 
val parse_bool_option = Sledgehammer_Util.parse_bool_option;  | 
|
43  | 
val parse_time = Sledgehammer_Util.parse_time;  | 
|
44  | 
val string_of_time = ATP_Util.string_of_time;  | 
|
45  | 
val simplify_spaces = Sledgehammer_Util.simplify_spaces;  | 
|
46  | 
val ascii_of = ATP_Problem_Generate.ascii_of;  | 
|
47  | 
val unascii_of = ATP_Problem_Generate.unascii_of;  | 
|
48  | 
val double_lookup = Nitpick_Util.double_lookup;  | 
|
49  | 
val triple_lookup = Nitpick_Util.triple_lookup;  | 
|
50  | 
val plural_s_for_list = Nitpick_Util.plural_s_for_list;  | 
|
51  | 
||
52  | 
fun with_overlord_file name ext f =  | 
|
53  | 
  f (Path.explode ("$ISABELLE_HOME_USER/" ^ name ^ "." ^ ext));
 | 
|
54  | 
||
55  | 
fun with_tmp_or_overlord_file overlord =  | 
|
56  | 
if overlord then with_overlord_file else Isabelle_System.with_tmp_file;  | 
|
57  | 
||
58  | 
val num_binder_types = BNF_Util.num_binder_types  | 
|
59  | 
val strip_fun_type = BNF_Util.strip_fun_type;  | 
|
60  | 
||
61  | 
(* Clone from "HOL/Tools/inductive_realizer.ML". *)  | 
|
62  | 
val attach_typeS =  | 
|
63  | 
map_types (map_atyps  | 
|
| 69593 | 64  | 
(fn TFree (s, []) => TFree (s, \<^sort>\<open>type\<close>)  | 
65  | 
| TVar (ixn, []) => TVar (ixn, \<^sort>\<open>type\<close>)  | 
|
| 64389 | 66  | 
| T => T));  | 
67  | 
||
68  | 
val specialize_type = ATP_Util.specialize_type;  | 
|
69  | 
||
70  | 
fun typ_match thy TU = can (Sign.typ_match thy TU) Vartab.empty;  | 
|
71  | 
fun term_match thy tu = can (Pattern.match thy tu) (Vartab.empty, Vartab.empty);  | 
|
72  | 
fun const_match thy = term_match thy o apply2 Const;  | 
|
73  | 
||
74  | 
val DETERM_TIMEOUT = Nitpick_Util.DETERM_TIMEOUT;  | 
|
75  | 
||
76  | 
val spying_version = "a"  | 
|
77  | 
||
78  | 
val hackish_string_of_term = Sledgehammer_Util.hackish_string_of_term;  | 
|
79  | 
||
80  | 
fun spying spy f =  | 
|
81  | 
if spy then  | 
|
82  | 
let  | 
|
83  | 
val (state, i, message) = f ();  | 
|
84  | 
val ctxt = Proof.context_of state;  | 
|
85  | 
val goal = Logic.get_goal (Thm.prop_of (#goal (Proof.goal state))) i;  | 
|
86  | 
val hash =  | 
|
87  | 
String.substring (SHA1.rep (SHA1.digest (hackish_string_of_term ctxt goal)), 0, 12);  | 
|
88  | 
in  | 
|
89  | 
File.append (Path.explode "$ISABELLE_HOME_USER/spy_nunchaku")  | 
|
90  | 
(spying_version ^ " " ^ timestamp () ^ ": " ^ hash ^ ": " ^ message ^ "\n")  | 
|
91  | 
end  | 
|
92  | 
else  | 
|
93  | 
();  | 
|
94  | 
||
95  | 
end;  |