| author | blanchet | 
| Fri, 24 May 2013 16:43:37 +0200 | |
| changeset 52125 | ac7830871177 | 
| parent 52077 | 788b27dfaefa | 
| child 52196 | 2281f33e8da6 | 
| permissions | -rw-r--r-- | 
| 
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  | 
| 48323 | 12  | 
val chunk_list : int -> 'a list -> 'a list list  | 
| 
48251
 
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
 
blanchet 
parents: 
48247 
diff
changeset
 | 
13  | 
val stringN_of_int : int -> int -> string  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
14  | 
val strip_spaces : bool -> (char -> bool) -> string -> string  | 
| 44784 | 15  | 
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
 | 
16  | 
val elide_string : int -> string -> string  | 
| 52077 | 17  | 
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
 | 
18  | 
val nat_subscript : int -> string  | 
| 
52076
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52031 
diff
changeset
 | 
19  | 
val unquote_tvar : string -> string  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
20  | 
val unyxml : string -> string  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
21  | 
val maybe_quote : string -> string  | 
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
51209 
diff
changeset
 | 
22  | 
val string_of_ext_time : bool * Time.time -> string  | 
| 
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
51209 
diff
changeset
 | 
23  | 
val string_of_time : Time.time -> string  | 
| 47150 | 24  | 
val type_instance : theory -> typ -> typ -> bool  | 
25  | 
val type_generalization : theory -> typ -> typ -> bool  | 
|
26  | 
val type_intersect : theory -> typ -> typ -> bool  | 
|
27  | 
val type_equiv : theory -> typ * typ -> bool  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
28  | 
val varify_type : Proof.context -> typ -> typ  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
29  | 
val instantiate_type : theory -> typ -> typ -> typ -> typ  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
30  | 
val varify_and_instantiate_type : Proof.context -> typ -> typ -> typ -> typ  | 
| 45896 | 31  | 
val typ_of_dtyp : Datatype.descr -> (Datatype.dtyp * typ) list -> Datatype.dtyp -> typ  | 
| 44393 | 32  | 
val is_type_surely_finite : Proof.context -> typ -> bool  | 
| 44399 | 33  | 
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
 | 
34  | 
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
 | 
35  | 
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
 | 
36  | 
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
 | 
37  | 
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
 | 
38  | 
val s_iff : term * term -> term  | 
| 
49983
 
33e18e9916a8
use metaquantification when possible in Isar proofs
 
blanchet 
parents: 
49982 
diff
changeset
 | 
39  | 
val close_form : term -> term  | 
| 49982 | 40  | 
val hol_close_form_prefix : string  | 
41  | 
val hol_close_form : term -> term  | 
|
42  | 
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
 | 
43  | 
val monomorphic_term : Type.tyenv -> term -> term  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
44  | 
val eta_expand : typ list -> term -> int -> term  | 
| 
47954
 
aada9fd08b58
make higher-order goals more first-order via extensionality
 
blanchet 
parents: 
47953 
diff
changeset
 | 
45  | 
val cong_extensionalize_term : theory -> term -> term  | 
| 
47953
 
a2c3706c4cb1
added "ext_cong_neq" lemma (not used yet); tuning
 
blanchet 
parents: 
47718 
diff
changeset
 | 
46  | 
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
 | 
47  | 
val unextensionalize_def : term -> term  | 
| 
 
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 is_legitimate_tptp_def : term -> bool  | 
| 
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  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
51  | 
val strip_subgoal :  | 
| 
52125
 
ac7830871177
improved handling of free variables' types in Isar proofs
 
blanchet 
parents: 
52077 
diff
changeset
 | 
52  | 
thm -> int -> Proof.context  | 
| 
 
ac7830871177
improved handling of free variables' types in Isar proofs
 
blanchet 
parents: 
52077 
diff
changeset
 | 
53  | 
-> ((string * typ) list * term list * term) * Proof.context  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
54  | 
end;  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
55  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
56  | 
structure ATP_Util : ATP_UTIL =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
57  | 
struct  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
58  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
59  | 
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
 | 
60  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
61  | 
(* 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
 | 
62  | 
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
 | 
63  | 
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
 | 
64  | 
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
 | 
65  | 
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
 | 
66  | 
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
 | 
67  | 
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
 | 
68  | 
| 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
 | 
69  | 
| 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
 | 
70  | 
| hashw_term _ = 0w0  | 
| 
 
62d64709af3b
added option to control which lambda translation to use (for experiments)
 
blanchet 
parents: 
43572 
diff
changeset
 | 
71  | 
|
| 
 
62d64709af3b
added option to control which lambda translation to use (for experiments)
 
blanchet 
parents: 
43572 
diff
changeset
 | 
72  | 
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
 | 
73  | 
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
 | 
74  | 
|
| 48323 | 75  | 
fun chunk_list _ [] = []  | 
76  | 
| chunk_list k xs =  | 
|
77  | 
let val (xs1, xs2) = chop k xs in xs1 :: chunk_list k xs2 end  | 
|
78  | 
||
| 
48251
 
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
 
blanchet 
parents: 
48247 
diff
changeset
 | 
79  | 
fun stringN_of_int 0 _ = ""  | 
| 
 
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
 
blanchet 
parents: 
48247 
diff
changeset
 | 
80  | 
| stringN_of_int k n =  | 
| 
 
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
 
blanchet 
parents: 
48247 
diff
changeset
 | 
81  | 
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
 | 
82  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
83  | 
fun strip_spaces skip_comments is_evil =  | 
| 
44935
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
84  | 
let  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
85  | 
fun strip_c_style_comment [] accum = accum  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
86  | 
| strip_c_style_comment (#"*" :: #"/" :: cs) accum =  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
87  | 
strip_spaces_in_list true cs accum  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
88  | 
| 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
 | 
89  | 
and strip_spaces_in_list _ [] accum = accum  | 
| 
44935
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
90  | 
| strip_spaces_in_list true (#"%" :: cs) accum =  | 
| 
48902
 
44a6967240b7
prefer classic take_prefix/take_suffix over chop_while (cf. 0659e84bdc5f);
 
wenzelm 
parents: 
48766 
diff
changeset
 | 
91  | 
strip_spaces_in_list true (cs |> take_prefix (not_equal #"\n") |> snd)  | 
| 
44935
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
92  | 
accum  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
93  | 
| strip_spaces_in_list true (#"/" :: #"*" :: cs) accum =  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
94  | 
strip_c_style_comment cs accum  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
95  | 
| strip_spaces_in_list _ [c1] accum =  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
96  | 
accum |> not (Char.isSpace c1) ? cons c1  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
97  | 
| strip_spaces_in_list skip_comments (cs as [_, _]) accum =  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
98  | 
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
 | 
99  | 
| 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
 | 
100  | 
if Char.isSpace c1 then  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
101  | 
strip_spaces_in_list skip_comments (c2 :: c3 :: cs) accum  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
102  | 
else if Char.isSpace c2 then  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
103  | 
if Char.isSpace c3 then  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
104  | 
strip_spaces_in_list skip_comments (c1 :: c3 :: cs) accum  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
105  | 
else  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
106  | 
strip_spaces_in_list skip_comments (c3 :: cs)  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
107  | 
(c1 :: accum |> forall is_evil [c1, c3] ? cons #" ")  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
108  | 
else  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
109  | 
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
 | 
110  | 
in  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
111  | 
String.explode  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
112  | 
#> rpair [] #-> strip_spaces_in_list skip_comments  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
113  | 
#> rev #> String.implode  | 
| 
 
2e812384afa8
tail recursive proof preprocessing (needed for huge proofs)
 
blanchet 
parents: 
44893 
diff
changeset
 | 
114  | 
end  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
115  | 
|
| 44784 | 116  | 
fun is_ident_char c = Char.isAlphaNum c orelse c = #"_"  | 
117  | 
val strip_spaces_except_between_idents = strip_spaces true is_ident_char  | 
|
118  | 
||
| 
48316
 
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
 
blanchet 
parents: 
48251 
diff
changeset
 | 
119  | 
fun elide_string threshold s =  | 
| 
 
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
 
blanchet 
parents: 
48251 
diff
changeset
 | 
120  | 
if size s > threshold then  | 
| 
 
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
 
blanchet 
parents: 
48251 
diff
changeset
 | 
121  | 
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
 | 
122  | 
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
 | 
123  | 
else  | 
| 
 
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
 
blanchet 
parents: 
48251 
diff
changeset
 | 
124  | 
s  | 
| 
 
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
 
blanchet 
parents: 
48251 
diff
changeset
 | 
125  | 
|
| 52077 | 126  | 
fun find_enclosed left right s =  | 
127  | 
case first_field left s of  | 
|
128  | 
SOME (_, s) =>  | 
|
129  | 
(case first_field right s of  | 
|
130  | 
SOME (enclosed, s) => enclosed :: find_enclosed left right s  | 
|
131  | 
| NONE => [])  | 
|
132  | 
| NONE => []  | 
|
133  | 
||
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
134  | 
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
 | 
135  | 
fun nat_subscript n =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
136  | 
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
 | 
137  | 
|
| 
52076
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52031 
diff
changeset
 | 
138  | 
val unquote_tvar = perhaps (try (unprefix "'"))  | 
| 
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52031 
diff
changeset
 | 
139  | 
val unquery_var = perhaps (try (unprefix "?"))  | 
| 
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52031 
diff
changeset
 | 
140  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
141  | 
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
 | 
142  | 
|
| 50239 | 143  | 
val is_long_identifier = forall Symbol_Pos.is_identifier o Long_Name.explode  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
144  | 
fun maybe_quote y =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
145  | 
let val s = unyxml y in  | 
| 
52076
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52031 
diff
changeset
 | 
146  | 
y |> ((not (is_long_identifier (unquote_tvar s)) andalso  | 
| 
 
bfa28e1cba77
freeze types in Sledgehammer goal, not just terms
 
blanchet 
parents: 
52031 
diff
changeset
 | 
147  | 
not (is_long_identifier (unquery_var s))) orelse  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
148  | 
Keyword.is_keyword s) ? quote  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
149  | 
end  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
150  | 
|
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
51209 
diff
changeset
 | 
151  | 
fun string_of_ext_time (plus, time) =  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
152  | 
let val ms = Time.toMilliseconds time in  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
153  | 
(if plus then "> " else "") ^  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
154  | 
(if plus andalso ms mod 1000 = 0 then  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
155  | 
signed_string_of_int (ms div 1000) ^ " s"  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
156  | 
else if ms < 1000 then  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
157  | 
signed_string_of_int ms ^ " ms"  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
158  | 
else  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
159  | 
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
 | 
160  | 
end  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
161  | 
|
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
51209 
diff
changeset
 | 
162  | 
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
 | 
163  | 
|
| 47150 | 164  | 
fun type_instance thy T T' = Sign.typ_instance thy (T, T')  | 
165  | 
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
 | 
166  | 
|
| 
 
8f37d2ddabc8
optimized type intersection, hoping this will reduce the number of sudden Interrupts in the "incr_tvar" code
 
blanchet 
parents: 
48238 
diff
changeset
 | 
167  | 
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
 | 
168  | 
| 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
 | 
169  | 
| 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
 | 
170  | 
let  | 
| 
 
8f37d2ddabc8
optimized type intersection, hoping this will reduce the number of sudden Interrupts in the "incr_tvar" code
 
blanchet 
parents: 
48238 
diff
changeset
 | 
171  | 
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
 | 
172  | 
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
 | 
173  | 
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
 | 
174  | 
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
 | 
175  | 
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
 | 
176  | 
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
 | 
177  | 
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
 | 
178  | 
|
| 47150 | 179  | 
val type_equiv = Sign.typ_equiv  | 
| 44399 | 180  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
181  | 
fun varify_type ctxt T =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
182  | 
  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
 | 
183  | 
|> snd |> the_single |> dest_Const |> snd  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
184  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
185  | 
(* TODO: use "Term_Subst.instantiateT" instead? *)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
186  | 
fun instantiate_type thy T1 T1' T2 =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
187  | 
Same.commit (Envir.subst_type_same  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
188  | 
(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
 | 
189  | 
  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
 | 
190  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
191  | 
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
 | 
192  | 
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
 | 
193  | 
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
 | 
194  | 
end  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
195  | 
|
| 45896 | 196  | 
fun typ_of_dtyp _ typ_assoc (Datatype.DtTFree a) =  | 
197  | 
the (AList.lookup (op =) typ_assoc (Datatype.DtTFree a))  | 
|
198  | 
| 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
 | 
199  | 
Type (s, map (typ_of_dtyp descr typ_assoc) Us)  | 
| 45896 | 200  | 
| 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
 | 
201  | 
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
 | 
202  | 
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
 | 
203  | 
end  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
204  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
205  | 
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
 | 
206  | 
(case Datatype.get_info thy s of  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
207  | 
       SOME {index, descr, ...} =>
 | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
208  | 
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
 | 
209  | 
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
 | 
210  | 
constrs  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
211  | 
end  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
212  | 
| NONE => [])  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
213  | 
| datatype_constrs _ _ = []  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
214  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
215  | 
(* 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
 | 
216  | 
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
 | 
217  | 
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
 | 
218  | 
cardinality of the type can't be determined. *)  | 
| 44500 | 219  | 
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
 | 
220  | 
let  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
221  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
222  | 
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
 | 
223  | 
fun aux slack avoid T =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
224  | 
if member (op =) avoid T then  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
225  | 
0  | 
| 47150 | 226  | 
else case AList.lookup (type_equiv thy) assigns T of  | 
| 44393 | 227  | 
SOME k => k  | 
228  | 
| NONE =>  | 
|
| 
44392
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
229  | 
case T of  | 
| 
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
230  | 
          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
 | 
231  | 
(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
 | 
232  | 
(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
 | 
233  | 
| (0, _) => 0  | 
| 
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
234  | 
| (_, 0) => 0  | 
| 
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
235  | 
| (k1, k2) =>  | 
| 
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
236  | 
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
 | 
237  | 
else Int.min (max, Integer.pow k2 k1))  | 
| 
48230
 
0feb93dfb268
gracefully compute cardinality of sets (to avoid type protectors)
 
blanchet 
parents: 
47991 
diff
changeset
 | 
238  | 
        | Type (@{type_name set}, [T']) => aux slack avoid (T' --> @{typ bool})
 | 
| 
44392
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
239  | 
        | @{typ prop} => 2
 | 
| 
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
240  | 
        | @{typ bool} => 2 (* optimization *)
 | 
| 
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
241  | 
        | @{typ nat} => 0 (* optimization *)
 | 
| 
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
242  | 
        | 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
 | 
243  | 
| Type (s, _) =>  | 
| 
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
244  | 
(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
 | 
245  | 
constrs as _ :: _ =>  | 
| 
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
246  | 
let  | 
| 
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
247  | 
val constr_cards =  | 
| 
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
248  | 
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
 | 
249  | 
o snd) constrs  | 
| 
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
250  | 
in  | 
| 
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
251  | 
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
 | 
252  | 
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
 | 
253  | 
end  | 
| 
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
254  | 
| [] =>  | 
| 
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
255  | 
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
 | 
256  | 
               ({abs_type, rep_type, ...}, _) :: _ =>
 | 
| 
45299
 
ee584ff987c3
check "sound" flag before doing something unsound...
 
blanchet 
parents: 
44935 
diff
changeset
 | 
257  | 
if not sound then  | 
| 
 
ee584ff987c3
check "sound" flag before doing something unsound...
 
blanchet 
parents: 
44935 
diff
changeset
 | 
258  | 
(* We cheat here by assuming that typedef types are infinite if  | 
| 
 
ee584ff987c3
check "sound" flag before doing something unsound...
 
blanchet 
parents: 
44935 
diff
changeset
 | 
259  | 
their underlying type is infinite. This is unsound in  | 
| 
 
ee584ff987c3
check "sound" flag before doing something unsound...
 
blanchet 
parents: 
44935 
diff
changeset
 | 
260  | 
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
 | 
261  | 
this would not be the case. We are also slack with  | 
| 
 
ee584ff987c3
check "sound" flag before doing something unsound...
 
blanchet 
parents: 
44935 
diff
changeset
 | 
262  | 
representation types: If a representation type has the form  | 
| 
 
ee584ff987c3
check "sound" flag before doing something unsound...
 
blanchet 
parents: 
44935 
diff
changeset
 | 
263  | 
"sigma => tau", we consider it enough to check "sigma" for  | 
| 
 
ee584ff987c3
check "sound" flag before doing something unsound...
 
blanchet 
parents: 
44935 
diff
changeset
 | 
264  | 
infiniteness. *)  | 
| 
 
ee584ff987c3
check "sound" flag before doing something unsound...
 
blanchet 
parents: 
44935 
diff
changeset
 | 
265  | 
(case varify_and_instantiate_type ctxt  | 
| 
 
ee584ff987c3
check "sound" flag before doing something unsound...
 
blanchet 
parents: 
44935 
diff
changeset
 | 
266  | 
(Logic.varifyT_global abs_type) T  | 
| 
 
ee584ff987c3
check "sound" flag before doing something unsound...
 
blanchet 
parents: 
44935 
diff
changeset
 | 
267  | 
(Logic.varifyT_global rep_type)  | 
| 
 
ee584ff987c3
check "sound" flag before doing something unsound...
 
blanchet 
parents: 
44935 
diff
changeset
 | 
268  | 
|> aux true avoid of  | 
| 
 
ee584ff987c3
check "sound" flag before doing something unsound...
 
blanchet 
parents: 
44935 
diff
changeset
 | 
269  | 
0 => 0  | 
| 
 
ee584ff987c3
check "sound" flag before doing something unsound...
 
blanchet 
parents: 
44935 
diff
changeset
 | 
270  | 
| 1 => 1  | 
| 
 
ee584ff987c3
check "sound" flag before doing something unsound...
 
blanchet 
parents: 
44935 
diff
changeset
 | 
271  | 
| _ => default_card)  | 
| 
 
ee584ff987c3
check "sound" flag before doing something unsound...
 
blanchet 
parents: 
44935 
diff
changeset
 | 
272  | 
else  | 
| 
 
ee584ff987c3
check "sound" flag before doing something unsound...
 
blanchet 
parents: 
44935 
diff
changeset
 | 
273  | 
default_card  | 
| 
44392
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
274  | 
| [] => default_card)  | 
| 
 
6750b4297691
reintroduced slightly unsound optimization taken out in 717880e98e6b, but only if "sound" is false
 
blanchet 
parents: 
43864 
diff
changeset
 | 
275  | 
(* 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
 | 
276  | 
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
 | 
277  | 
likely have used "unit" directly anyway.) *)  | 
| 44500 | 278  | 
| TFree _ =>  | 
279  | 
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
 | 
280  | 
| TVar _ => default_card  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
281  | 
in Int.min (max, aux false [] T) end  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
282  | 
|
| 44500 | 283  | 
fun is_type_surely_finite ctxt T = tiny_card_of_type ctxt true [] 0 T <> 0  | 
284  | 
fun is_type_surely_infinite ctxt sound infinite_Ts T =  | 
|
285  | 
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
 | 
286  | 
|
| 
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
 | 
287  | 
(* 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
 | 
288  | 
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
 | 
289  | 
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
 | 
290  | 
    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
 | 
291  | 
  | 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
 | 
292  | 
    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
 | 
293  | 
  | 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
 | 
294  | 
  | 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
 | 
295  | 
    @{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
 | 
296  | 
  | 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
 | 
297  | 
    @{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
 | 
298  | 
  | 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
 | 
299  | 
  | 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
 | 
300  | 
  | 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
 | 
301  | 
  | 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
 | 
302  | 
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
 | 
303  | 
  | s_conj (t1, @{const True}) = t1
 | 
| 51209 | 304  | 
  | s_conj (@{const False}, _) = @{const False}
 | 
305  | 
  | s_conj (_, @{const False}) = @{const False}
 | 
|
| 51197 | 306  | 
| s_conj (t1, t2) = if t1 aconv t2 then t1 else HOLogic.mk_conj (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
 | 
307  | 
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
 | 
308  | 
  | s_disj (t1, @{const False}) = t1
 | 
| 51209 | 309  | 
  | s_disj (@{const True}, _) = @{const True}
 | 
310  | 
  | s_disj (_, @{const True}) = @{const True}
 | 
|
| 51197 | 311  | 
| s_disj (t1, t2) = if t1 aconv t2 then t1 else HOLogic.mk_disj (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
 | 
312  | 
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
 | 
313  | 
  | s_imp (t1, @{const False}) = s_not t1
 | 
| 51209 | 314  | 
  | s_imp (@{const False}, _) = @{const True}
 | 
315  | 
  | s_imp (_, @{const True}) = @{const True}
 | 
|
| 
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
 | 
316  | 
| 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
 | 
317  | 
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
 | 
318  | 
  | s_iff (t1, @{const True}) = t1
 | 
| 51209 | 319  | 
  | s_iff (@{const False}, t2) = s_not t2
 | 
320  | 
  | s_iff (t1, @{const False}) = s_not t1
 | 
|
| 
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
 | 
321  | 
| 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
 | 
322  | 
|
| 
49983
 
33e18e9916a8
use metaquantification when possible in Isar proofs
 
blanchet 
parents: 
49982 
diff
changeset
 | 
323  | 
(* cf. "close_form" in "refute.ML" *)  | 
| 
 
33e18e9916a8
use metaquantification when possible in Isar proofs
 
blanchet 
parents: 
49982 
diff
changeset
 | 
324  | 
fun close_form t =  | 
| 
 
33e18e9916a8
use metaquantification when possible in Isar proofs
 
blanchet 
parents: 
49982 
diff
changeset
 | 
325  | 
fold (fn ((s, i), T) => fn t' =>  | 
| 
 
33e18e9916a8
use metaquantification when possible in Isar proofs
 
blanchet 
parents: 
49982 
diff
changeset
 | 
326  | 
Logic.all_const T $ Abs (s, T, abstract_over (Var ((s, i), T), t')))  | 
| 
 
33e18e9916a8
use metaquantification when possible in Isar proofs
 
blanchet 
parents: 
49982 
diff
changeset
 | 
327  | 
(Term.add_vars t []) t  | 
| 
 
33e18e9916a8
use metaquantification when possible in Isar proofs
 
blanchet 
parents: 
49982 
diff
changeset
 | 
328  | 
|
| 49982 | 329  | 
val hol_close_form_prefix = "ATP.close_form."  | 
| 
46385
 
0ccf458a3633
third attempt at lambda lifting that works for both Sledgehammer and Metis (cf. dce6c3a460a9)
 
blanchet 
parents: 
45896 
diff
changeset
 | 
330  | 
|
| 49982 | 331  | 
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
 | 
332  | 
fold (fn ((s, i), T) => fn t' =>  | 
| 
45511
 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 
blanchet 
parents: 
45299 
diff
changeset
 | 
333  | 
HOLogic.all_const T  | 
| 49982 | 334  | 
$ 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
 | 
335  | 
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
 | 
336  | 
(Term.add_vars t []) t  | 
| 
 
58a7b3fdc193
fixed lambda-liftg: must ensure the formulas are in close form
 
blanchet 
parents: 
43863 
diff
changeset
 | 
337  | 
|
| 49982 | 338  | 
fun hol_open_form unprefix  | 
339  | 
      (t as Const (@{const_name All}, _) $ Abs (s, T, t')) =
 | 
|
| 
47718
 
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
 
blanchet 
parents: 
47715 
diff
changeset
 | 
340  | 
(case try unprefix s of  | 
| 
 
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
 
blanchet 
parents: 
47715 
diff
changeset
 | 
341  | 
SOME s =>  | 
| 
 
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
 
blanchet 
parents: 
47715 
diff
changeset
 | 
342  | 
let  | 
| 
 
39229c760636
smoother handling of conjecture, so that its Skolem constants get displayed in countermodels
 
blanchet 
parents: 
47715 
diff
changeset
 | 
343  | 
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
 | 
344  | 
val (s, _) = Name.variant s names  | 
| 49982 | 345  | 
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
 | 
346  | 
| NONE => t)  | 
| 49982 | 347  | 
| 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
 | 
348  | 
|
| 
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
 | 
349  | 
fun monomorphic_term subst =  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
350  | 
map_types (map_type_tvar (fn v =>  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
351  | 
case Type.lookup subst v of  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
352  | 
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
 | 
353  | 
| NONE => TVar v))  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
354  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
355  | 
fun eta_expand _ t 0 = t  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
356  | 
| eta_expand Ts (Abs (s, T, t')) n =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
357  | 
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
 | 
358  | 
| eta_expand Ts t n =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
359  | 
    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
 | 
360  | 
(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
 | 
361  | 
(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
 | 
362  | 
|
| 
47954
 
aada9fd08b58
make higher-order goals more first-order via extensionality
 
blanchet 
parents: 
47953 
diff
changeset
 | 
363  | 
fun cong_extensionalize_term thy t =  | 
| 
 
aada9fd08b58
make higher-order goals more first-order via extensionality
 
blanchet 
parents: 
47953 
diff
changeset
 | 
364  | 
  if exists_Const (fn (s, _) => s = @{const_name Not}) t then
 | 
| 
 
aada9fd08b58
make higher-order goals more first-order via extensionality
 
blanchet 
parents: 
47953 
diff
changeset
 | 
365  | 
t |> Skip_Proof.make_thm thy  | 
| 
 
aada9fd08b58
make higher-order goals more first-order via extensionality
 
blanchet 
parents: 
47953 
diff
changeset
 | 
366  | 
|> Meson.cong_extensionalize_thm thy  | 
| 
 
aada9fd08b58
make higher-order goals more first-order via extensionality
 
blanchet 
parents: 
47953 
diff
changeset
 | 
367  | 
|> prop_of  | 
| 
 
aada9fd08b58
make higher-order goals more first-order via extensionality
 
blanchet 
parents: 
47953 
diff
changeset
 | 
368  | 
else  | 
| 
 
aada9fd08b58
make higher-order goals more first-order via extensionality
 
blanchet 
parents: 
47953 
diff
changeset
 | 
369  | 
t  | 
| 
 
aada9fd08b58
make higher-order goals more first-order via extensionality
 
blanchet 
parents: 
47953 
diff
changeset
 | 
370  | 
|
| 
47715
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47150 
diff
changeset
 | 
371  | 
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
 | 
372  | 
                     Type (_, [Type (@{type_name fun}, _), _])) = true
 | 
| 
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47150 
diff
changeset
 | 
373  | 
| is_fun_equality _ = false  | 
| 
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47150 
diff
changeset
 | 
374  | 
|
| 
47953
 
a2c3706c4cb1
added "ext_cong_neq" lemma (not used yet); tuning
 
blanchet 
parents: 
47718 
diff
changeset
 | 
375  | 
fun abs_extensionalize_term ctxt t =  | 
| 
47715
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47150 
diff
changeset
 | 
376  | 
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
 | 
377  | 
let val thy = Proof_Context.theory_of ctxt in  | 
| 
47953
 
a2c3706c4cb1
added "ext_cong_neq" lemma (not used yet); tuning
 
blanchet 
parents: 
47718 
diff
changeset
 | 
378  | 
t |> cterm_of thy |> Meson.abs_extensionalize_conv ctxt  | 
| 
47715
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47150 
diff
changeset
 | 
379  | 
|> prop_of |> Logic.dest_equals |> snd  | 
| 
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47150 
diff
changeset
 | 
380  | 
end  | 
| 
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47150 
diff
changeset
 | 
381  | 
else  | 
| 
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47150 
diff
changeset
 | 
382  | 
t  | 
| 
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47150 
diff
changeset
 | 
383  | 
|
| 
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
 | 
384  | 
fun unextensionalize_def t =  | 
| 
 
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
 | 
385  | 
case t of  | 
| 
 
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
 | 
386  | 
    @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ lhs $ rhs) =>
 | 
| 
 
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
 | 
387  | 
(case strip_comb lhs of  | 
| 
 
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
 | 
388  | 
(c as Const (_, T), args) =>  | 
| 
 
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
 | 
389  | 
if forall is_Var args andalso not (has_duplicates (op =) args) then  | 
| 
 
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
 | 
390  | 
         @{const Trueprop}
 | 
| 
 
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
 | 
391  | 
         $ (Const (@{const_name HOL.eq}, T --> T --> @{typ bool})
 | 
| 
 
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
 | 
392  | 
$ c $ fold_rev lambda args rhs)  | 
| 
 
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
 | 
393  | 
else  | 
| 
 
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
 | 
394  | 
t  | 
| 
 
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
 | 
395  | 
| _ => t)  | 
| 
 
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
 | 
396  | 
| _ => t  | 
| 
 
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
 | 
397  | 
|
| 
 
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
 | 
398  | 
fun is_legitimate_tptp_def (@{const Trueprop} $ t) = is_legitimate_tptp_def t
 | 
| 
 
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
 | 
399  | 
  | is_legitimate_tptp_def (Const (@{const_name HOL.eq}, _) $ t $ u) =
 | 
| 
 
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
 | 
400  | 
(is_Const t orelse is_Free t) andalso  | 
| 
 
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
 | 
401  | 
not (exists_subterm (curry (op =) t) u)  | 
| 
 
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
 | 
402  | 
| is_legitimate_tptp_def _ = false  | 
| 
 
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
 | 
403  | 
|
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
404  | 
(* 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
 | 
405  | 
predicate variable. Leaves other theorems unchanged. We simply instantiate  | 
| 44460 | 406  | 
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
 | 
407  | 
"Meson_Clausify".) *)  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
408  | 
fun transform_elim_prop t =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
409  | 
case Logic.strip_imp_concl t of  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
410  | 
    @{const Trueprop} $ Var (z, @{typ bool}) =>
 | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
411  | 
    subst_Vars [(z, @{const False})] t
 | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
412  | 
  | 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
 | 
413  | 
| _ => t  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
414  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
415  | 
fun specialize_type thy (s, T) t =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
416  | 
let  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
417  | 
fun subst_for (Const (s', T')) =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
418  | 
if s = s' then  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
419  | 
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
 | 
420  | 
handle Type.TYPE_MATCH => NONE  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
421  | 
else  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
422  | 
NONE  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
423  | 
| subst_for (t1 $ t2) =  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
424  | 
(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
 | 
425  | 
| subst_for (Abs (_, _, t')) = subst_for t'  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
426  | 
| subst_for _ = NONE  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
427  | 
in  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
428  | 
case subst_for t of  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
429  | 
SOME subst => monomorphic_term subst t  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
430  | 
| NONE => raise Type.TYPE_MATCH  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
431  | 
end  | 
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
432  | 
|
| 
52125
 
ac7830871177
improved handling of free variables' types in Isar proofs
 
blanchet 
parents: 
52077 
diff
changeset
 | 
433  | 
fun strip_subgoal goal i ctxt =  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
434  | 
let  | 
| 
52125
 
ac7830871177
improved handling of free variables' types in Isar proofs
 
blanchet 
parents: 
52077 
diff
changeset
 | 
435  | 
val (t, ((frees, params), ctxt)) =  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
436  | 
Logic.goal_params (prop_of goal) i  | 
| 
52125
 
ac7830871177
improved handling of free variables' types in Isar proofs
 
blanchet 
parents: 
52077 
diff
changeset
 | 
437  | 
||> map dest_Free  | 
| 
 
ac7830871177
improved handling of free variables' types in Isar proofs
 
blanchet 
parents: 
52077 
diff
changeset
 | 
438  | 
||> map (fn (s, T) => (Binding.name s, SOME T, NoSyn))  | 
| 
 
ac7830871177
improved handling of free variables' types in Isar proofs
 
blanchet 
parents: 
52077 
diff
changeset
 | 
439  | 
||> (fn fixes =>  | 
| 
 
ac7830871177
improved handling of free variables' types in Isar proofs
 
blanchet 
parents: 
52077 
diff
changeset
 | 
440  | 
ctxt |> Variable.set_body false  | 
| 
 
ac7830871177
improved handling of free variables' types in Isar proofs
 
blanchet 
parents: 
52077 
diff
changeset
 | 
441  | 
|> Proof_Context.add_fixes fixes  | 
| 
 
ac7830871177
improved handling of free variables' types in Isar proofs
 
blanchet 
parents: 
52077 
diff
changeset
 | 
442  | 
|>> map2 (fn (_, SOME T, _) => fn s => (s, T)) fixes)  | 
| 
 
ac7830871177
improved handling of free variables' types in Isar proofs
 
blanchet 
parents: 
52077 
diff
changeset
 | 
443  | 
||> apfst (`(map Free))  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
444  | 
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
 | 
445  | 
val concl_t = t |> Logic.strip_assums_concl |> curry subst_bounds frees  | 
| 
52125
 
ac7830871177
improved handling of free variables' types in Isar proofs
 
blanchet 
parents: 
52077 
diff
changeset
 | 
446  | 
in ((rev params, hyp_ts, concl_t), ctxt) end  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
447  | 
|
| 
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents:  
diff
changeset
 | 
448  | 
end;  |