| author | wenzelm | 
| Wed, 19 Jul 2023 11:31:19 +0200 | |
| changeset 78403 | a09929ad05b6 | 
| parent 77879 | dd222e2af01a | 
| child 81942 | da3c3948a39c | 
| permissions | -rw-r--r-- | 
| 58061 | 1  | 
(* Title: HOL/Tools/SMT/smt_normalize.ML  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
2  | 
Author: Sascha Boehme, TU Muenchen  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
3  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
4  | 
Normalization steps on theorems required by SMT solvers.  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
6  | 
|
| 58061 | 7  | 
signature SMT_NORMALIZE =  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
8  | 
sig  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
9  | 
val drop_fact_warning: Proof.context -> thm -> unit  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
10  | 
val atomize_conv: Proof.context -> conv  | 
| 
56981
 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 
blanchet 
parents: 
56245 
diff
changeset
 | 
11  | 
|
| 
 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 
blanchet 
parents: 
56245 
diff
changeset
 | 
12  | 
val special_quant_table: (string * thm) list  | 
| 
 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 
blanchet 
parents: 
56245 
diff
changeset
 | 
13  | 
val case_bool_entry: string * thm  | 
| 
 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 
blanchet 
parents: 
56245 
diff
changeset
 | 
14  | 
val abs_min_max_table: (string * thm) list  | 
| 
 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 
blanchet 
parents: 
56245 
diff
changeset
 | 
15  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
16  | 
type extra_norm = Proof.context -> thm list * thm list -> thm list * thm list  | 
| 58061 | 17  | 
val add_extra_norm: SMT_Util.class * extra_norm -> Context.generic -> Context.generic  | 
| 
75274
 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 
desharna 
parents: 
74561 
diff
changeset
 | 
18  | 
val normalize: Proof.context -> (SMT_Util.role * thm) list -> ((int * SMT_Util.role) * thm) list  | 
| 57229 | 19  | 
end;  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
20  | 
|
| 58061 | 21  | 
structure SMT_Normalize: SMT_NORMALIZE =  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
22  | 
struct  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
23  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
24  | 
fun drop_fact_warning ctxt =  | 
| 58061 | 25  | 
SMT_Config.verbose_msg ctxt (prefix "Warning: dropping assumption: " o  | 
| 61268 | 26  | 
Thm.string_of_thm ctxt)  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
27  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
28  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
29  | 
(* general theorem normalizations *)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
30  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
31  | 
(** instantiate elimination rules **)  | 
| 57230 | 32  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
33  | 
local  | 
| 74382 | 34  | 
val (cpfalse, cfalse) = `SMT_Util.mk_cprop \<^cterm>\<open>False\<close>  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
35  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
36  | 
fun inst f ct thm =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
37  | 
let val cv = f (Drule.strip_imp_concl (Thm.cprop_of thm))  | 
| 77879 | 38  | 
in Thm.instantiate (TVars.empty, Vars.make1 (dest_Var (Thm.term_of cv), ct)) thm end  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
39  | 
in  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
40  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
41  | 
fun instantiate_elim thm =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
42  | 
(case Thm.concl_of thm of  | 
| 74382 | 43  | 
\<^Const_>\<open>Trueprop for \<open>Var (_, \<^Type>\<open>bool\<close>)\<close>\<close> => inst Thm.dest_arg cfalse thm  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
44  | 
| Var _ => inst I cpfalse thm  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
45  | 
| _ => thm)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
46  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
47  | 
end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
48  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
49  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
50  | 
(** normalize definitions **)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
52  | 
fun norm_def thm =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
53  | 
(case Thm.prop_of thm of  | 
| 74382 | 54  | 
    \<^Const_>\<open>Trueprop for \<^Const>\<open>HOL.eq _ for _ \<open>Abs _\<close>\<close>\<close> => norm_def (thm RS @{thm fun_cong})
 | 
55  | 
| \<^Const_>\<open>Pure.eq _ for _ \<open>Abs _\<close>\<close> => norm_def (HOLogic.mk_obj_eq thm)  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
56  | 
| _ => thm)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
57  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
58  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
59  | 
(** atomization **)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
60  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
61  | 
fun atomize_conv ctxt ct =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
62  | 
(case Thm.term_of ct of  | 
| 74382 | 63  | 
\<^Const_>\<open>Pure.imp for _ _\<close> =>  | 
| 
56103
 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 
blanchet 
parents: 
56100 
diff
changeset
 | 
64  | 
      Conv.binop_conv (atomize_conv ctxt) then_conv Conv.rewr_conv @{thm atomize_imp}
 | 
| 74382 | 65  | 
| \<^Const_>\<open>Pure.eq _ for _ _\<close> =>  | 
| 
56103
 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 
blanchet 
parents: 
56100 
diff
changeset
 | 
66  | 
      Conv.binop_conv (atomize_conv ctxt) then_conv Conv.rewr_conv @{thm atomize_eq}
 | 
| 74382 | 67  | 
| \<^Const_>\<open>Pure.all _ for \<open>Abs _\<close>\<close> =>  | 
| 
56103
 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 
blanchet 
parents: 
56100 
diff
changeset
 | 
68  | 
      Conv.binder_conv (atomize_conv o snd) ctxt then_conv Conv.rewr_conv @{thm atomize_all}
 | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
69  | 
| _ => Conv.all_conv) ct  | 
| 
61033
 
fd7fe96ca7b9
generate proper error instead of exception if goal cannot be atomized
 
blanchet 
parents: 
60642 
diff
changeset
 | 
70  | 
handle CTERM _ => Conv.all_conv ct  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
71  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
72  | 
val setup_atomize =  | 
| 67149 | 73  | 
fold SMT_Builtin.add_builtin_fun_ext'' [\<^const_name>\<open>Pure.imp\<close>, \<^const_name>\<open>Pure.eq\<close>,  | 
74  | 
\<^const_name>\<open>Pure.all\<close>, \<^const_name>\<open>Trueprop\<close>]  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
75  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
76  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
77  | 
(** unfold special quantifiers **)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
78  | 
|
| 
56981
 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 
blanchet 
parents: 
56245 
diff
changeset
 | 
79  | 
val special_quant_table = [  | 
| 67149 | 80  | 
  (\<^const_name>\<open>Ex1\<close>, @{thm Ex1_def_raw}),
 | 
81  | 
  (\<^const_name>\<open>Ball\<close>, @{thm Ball_def_raw}),
 | 
|
82  | 
  (\<^const_name>\<open>Bex\<close>, @{thm Bex_def_raw})]
 | 
|
| 
56981
 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 
blanchet 
parents: 
56245 
diff
changeset
 | 
83  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
84  | 
local  | 
| 
56981
 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 
blanchet 
parents: 
56245 
diff
changeset
 | 
85  | 
fun special_quant (Const (n, _)) = AList.lookup (op =) special_quant_table n  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
86  | 
| special_quant _ = NONE  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
87  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
88  | 
fun special_quant_conv _ ct =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
89  | 
(case special_quant (Thm.term_of ct) of  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
90  | 
SOME thm => Conv.rewr_conv thm  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
91  | 
| NONE => Conv.all_conv) ct  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
92  | 
in  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
93  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
94  | 
fun unfold_special_quants_conv ctxt =  | 
| 58061 | 95  | 
SMT_Util.if_exists_conv (is_some o special_quant) (Conv.top_conv special_quant_conv ctxt)  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
96  | 
|
| 58061 | 97  | 
val setup_unfolded_quants = fold (SMT_Builtin.add_builtin_fun_ext'' o fst) special_quant_table  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
98  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
99  | 
end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
100  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
101  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
102  | 
(** trigger inference **)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
103  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
104  | 
local  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
105  | 
(*** check trigger syntax ***)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
106  | 
|
| 67149 | 107  | 
fun dest_trigger (Const (\<^const_name>\<open>pat\<close>, _) $ _) = SOME true  | 
108  | 
| dest_trigger (Const (\<^const_name>\<open>nopat\<close>, _) $ _) = SOME false  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
109  | 
| dest_trigger _ = NONE  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
110  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
111  | 
fun eq_list [] = false  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
112  | 
| eq_list (b :: bs) = forall (equal b) bs  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
113  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
114  | 
fun proper_trigger t =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
115  | 
t  | 
| 58061 | 116  | 
|> these o try SMT_Util.dest_symb_list  | 
117  | 
|> map (map_filter dest_trigger o these o try SMT_Util.dest_symb_list)  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
118  | 
|> (fn [] => false | bss => forall eq_list bss)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
119  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
120  | 
fun proper_quant inside f t =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
121  | 
(case t of  | 
| 74382 | 122  | 
\<^Const_>\<open>All _ for \<open>Abs (_, _, u)\<close>\<close> => proper_quant true f u  | 
123  | 
| \<^Const_>\<open>Ex _ for \<open>Abs (_, _, u)\<close>\<close> => proper_quant true f u  | 
|
124  | 
| \<^Const_>\<open>trigger for p u\<close> =>  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
125  | 
(if inside then f p else false) andalso proper_quant false f u  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
126  | 
| Abs (_, _, u) => proper_quant false f u  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
127  | 
| u1 $ u2 => proper_quant false f u1 andalso proper_quant false f u2  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
128  | 
| _ => true)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
129  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
130  | 
fun check_trigger_error ctxt t =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
131  | 
    error ("SMT triggers must only occur under quantifier and multipatterns " ^
 | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
132  | 
"must have the same kind: " ^ Syntax.string_of_term ctxt t)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
133  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
134  | 
fun check_trigger_conv ctxt ct =  | 
| 58061 | 135  | 
if proper_quant false proper_trigger (SMT_Util.term_of ct) then Conv.all_conv ct  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
136  | 
else check_trigger_error ctxt (Thm.term_of ct)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
137  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
138  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
139  | 
(*** infer simple triggers ***)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
140  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
141  | 
fun dest_cond_eq ct =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
142  | 
(case Thm.term_of ct of  | 
| 67149 | 143  | 
Const (\<^const_name>\<open>HOL.eq\<close>, _) $ _ $ _ => Thm.dest_binop ct  | 
| 74382 | 144  | 
| \<^Const_>\<open>implies for _ _\<close> => dest_cond_eq (Thm.dest_arg ct)  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
145  | 
    | _ => raise CTERM ("no equation", [ct]))
 | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
146  | 
|
| 58132 | 147  | 
fun get_constrs thy (Type (n, _)) = these (BNF_LFP_Compat.get_constrs thy n)  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
148  | 
| get_constrs _ _ = []  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
149  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
150  | 
fun is_constr thy (n, T) =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
151  | 
let fun match (m, U) = m = n andalso Sign.typ_instance thy (T, U)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
152  | 
in can (the o find_first match o get_constrs thy o Term.body_type) T end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
153  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
154  | 
fun is_constr_pat thy t =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
155  | 
(case Term.strip_comb t of  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
156  | 
(Free _, []) => true  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
157  | 
| (Const c, ts) => is_constr thy c andalso forall (is_constr_pat thy) ts  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
158  | 
| _ => false)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
159  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
160  | 
fun is_simp_lhs ctxt t =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
161  | 
(case Term.strip_comb t of  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
162  | 
(Const c, ts as _ :: _) =>  | 
| 58061 | 163  | 
not (SMT_Builtin.is_builtin_fun_ext ctxt c ts) andalso  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
164  | 
forall (is_constr_pat (Proof_Context.theory_of ctxt)) ts  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
165  | 
| _ => false)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
166  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
167  | 
fun has_all_vars vs t =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
168  | 
subset (op aconv) (vs, map Free (Term.add_frees t []))  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
169  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
170  | 
fun minimal_pats vs ct =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
171  | 
if has_all_vars vs (Thm.term_of ct) then  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
172  | 
(case Thm.term_of ct of  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
173  | 
_ $ _ =>  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58132 
diff
changeset
 | 
174  | 
(case apply2 (minimal_pats vs) (Thm.dest_comb ct) of  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
175  | 
([], []) => [[ct]]  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
176  | 
| (ctss, ctss') => union (eq_set (op aconvc)) ctss ctss')  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
177  | 
| _ => [])  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
178  | 
else []  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
179  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
180  | 
fun proper_mpat _ _ _ [] = false  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
181  | 
| proper_mpat thy gen u cts =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
182  | 
let  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
183  | 
val tps = (op ~~) (`gen (map Thm.term_of cts))  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
184  | 
fun some_match u = tps |> exists (fn (t', t) =>  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
185  | 
Pattern.matches thy (t', u) andalso not (t aconv u))  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
186  | 
in not (Term.exists_subterm some_match u) end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
187  | 
|
| 70159 | 188  | 
val pat = SMT_Util.mk_const_pat \<^theory> \<^const_name>\<open>pat\<close> Thm.dest_ctyp0  | 
| 58061 | 189  | 
fun mk_pat ct = Thm.apply (SMT_Util.instT' ct pat) ct  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
190  | 
|
| 57230 | 191  | 
fun mk_clist T =  | 
| 69593 | 192  | 
apply2 (Thm.cterm_of \<^context>) (SMT_Util.symb_cons_const T, SMT_Util.symb_nil_const T)  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
193  | 
fun mk_list (ccons, cnil) f cts = fold_rev (Thm.mk_binop ccons o f) cts cnil  | 
| 67149 | 194  | 
val mk_pat_list = mk_list (mk_clist \<^typ>\<open>pattern\<close>)  | 
195  | 
val mk_mpat_list = mk_list (mk_clist \<^typ>\<open>pattern symb_list\<close>)  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
196  | 
fun mk_trigger ctss = mk_mpat_list (mk_pat_list mk_pat) ctss  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
197  | 
|
| 57230 | 198  | 
  val trigger_eq = mk_meta_eq @{lemma "p = trigger t p" by (simp add: trigger_def)}
 | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
199  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
200  | 
fun insert_trigger_conv [] ct = Conv.all_conv ct  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
201  | 
| insert_trigger_conv ctss ct =  | 
| 
60642
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
59632 
diff
changeset
 | 
202  | 
let  | 
| 
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
59632 
diff
changeset
 | 
203  | 
val (ctr, cp) = Thm.dest_binop (Thm.rhs_of trigger_eq) ||> rpair ct  | 
| 
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
59632 
diff
changeset
 | 
204  | 
val inst = map (apfst (dest_Var o Thm.term_of)) [cp, (ctr, mk_trigger ctss)]  | 
| 74282 | 205  | 
in Thm.instantiate (TVars.empty, Vars.make inst) trigger_eq end  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
206  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
207  | 
fun infer_trigger_eq_conv outer_ctxt (ctxt, cvs) ct =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
208  | 
let  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
209  | 
val (lhs, rhs) = dest_cond_eq ct  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
210  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
211  | 
val vs = map Thm.term_of cvs  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
212  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
213  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
214  | 
fun get_mpats ct =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
215  | 
if is_simp_lhs ctxt (Thm.term_of ct) then minimal_pats vs ct  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
216  | 
else []  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
217  | 
val gen = Variable.export_terms ctxt outer_ctxt  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
218  | 
val filter_mpats = filter (proper_mpat thy gen (Thm.term_of rhs))  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
219  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
220  | 
in insert_trigger_conv (filter_mpats (get_mpats lhs)) ct end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
221  | 
|
| 74382 | 222  | 
fun has_trigger \<^Const_>\<open>trigger for _ _\<close> = true  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
223  | 
| has_trigger _ = false  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
224  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
225  | 
fun try_trigger_conv cv ct =  | 
| 58061 | 226  | 
if SMT_Util.under_quant has_trigger (SMT_Util.term_of ct) then Conv.all_conv ct  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
227  | 
else Conv.try_conv cv ct  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
228  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
229  | 
fun infer_trigger_conv ctxt =  | 
| 58061 | 230  | 
if Config.get ctxt SMT_Config.infer_triggers then  | 
231  | 
try_trigger_conv (SMT_Util.under_quant_conv (infer_trigger_eq_conv ctxt) ctxt)  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
232  | 
else Conv.all_conv  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
233  | 
in  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
234  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
235  | 
fun trigger_conv ctxt =  | 
| 58061 | 236  | 
SMT_Util.prop_conv (check_trigger_conv ctxt then_conv infer_trigger_conv ctxt)  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
237  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
238  | 
val setup_trigger =  | 
| 58061 | 239  | 
fold SMT_Builtin.add_builtin_fun_ext''  | 
| 67149 | 240  | 
[\<^const_name>\<open>pat\<close>, \<^const_name>\<open>nopat\<close>, \<^const_name>\<open>trigger\<close>]  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
241  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
242  | 
end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
243  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
244  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
245  | 
(** combined general normalizations **)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
246  | 
|
| 
57165
 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 
blanchet 
parents: 
56981 
diff
changeset
 | 
247  | 
fun gen_normalize1_conv ctxt =  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
248  | 
atomize_conv ctxt then_conv  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
249  | 
unfold_special_quants_conv ctxt then_conv  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
250  | 
Thm.beta_conversion true then_conv  | 
| 
57165
 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 
blanchet 
parents: 
56981 
diff
changeset
 | 
251  | 
trigger_conv ctxt  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
252  | 
|
| 
57165
 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 
blanchet 
parents: 
56981 
diff
changeset
 | 
253  | 
fun gen_normalize1 ctxt =  | 
| 56107 | 254  | 
instantiate_elim #>  | 
255  | 
norm_def #>  | 
|
256  | 
Conv.fconv_rule (Thm.beta_conversion true then_conv Thm.eta_conversion) #>  | 
|
| 74245 | 257  | 
Thm.forall_intr_vars #>  | 
| 
57165
 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 
blanchet 
parents: 
56981 
diff
changeset
 | 
258  | 
Conv.fconv_rule (gen_normalize1_conv ctxt) #>  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
259  | 
(* Z3 4.3.1 silently normalizes "P --> Q --> R" to "P & Q --> R" *)  | 
| 56107 | 260  | 
  Raw_Simplifier.rewrite_rule ctxt @{thms HOL.imp_conjL[symmetric, THEN eq_reflection]}
 | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
261  | 
|
| 
57165
 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 
blanchet 
parents: 
56981 
diff
changeset
 | 
262  | 
fun gen_norm1_safe ctxt (i, thm) =  | 
| 
 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 
blanchet 
parents: 
56981 
diff
changeset
 | 
263  | 
(case try (gen_normalize1 ctxt) thm of  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
264  | 
SOME thm' => SOME (i, thm')  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
265  | 
| NONE => (drop_fact_warning ctxt thm; NONE))  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
266  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
267  | 
fun gen_normalize ctxt iwthms = map_filter (gen_norm1_safe ctxt) iwthms  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
268  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
269  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
270  | 
(* unfolding of definitions and theory-specific rewritings *)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
271  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
272  | 
fun expand_head_conv cv ct =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
273  | 
(case Thm.term_of ct of  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
274  | 
_ $ _ =>  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
275  | 
Conv.fun_conv (expand_head_conv cv) then_conv  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
276  | 
Conv.try_conv (Thm.beta_conversion false)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
277  | 
| _ => cv) ct  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
278  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
279  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
280  | 
(** rewrite bool case expressions as if expressions **)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
281  | 
|
| 67149 | 282  | 
val case_bool_entry = (\<^const_name>\<open>bool.case_bool\<close>, @{thm case_bool_if})
 | 
| 
56981
 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 
blanchet 
parents: 
56245 
diff
changeset
 | 
283  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
284  | 
local  | 
| 67149 | 285  | 
fun is_case_bool (Const (\<^const_name>\<open>bool.case_bool\<close>, _)) = true  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
286  | 
| is_case_bool _ = false  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
287  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
288  | 
fun unfold_conv _ =  | 
| 58061 | 289  | 
SMT_Util.if_true_conv (is_case_bool o Term.head_of)  | 
| 
56103
 
6689512f3710
move lemmas to theory file, towards textual proof reconstruction
 
blanchet 
parents: 
56100 
diff
changeset
 | 
290  | 
      (expand_head_conv (Conv.rewr_conv @{thm case_bool_if}))
 | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
291  | 
in  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
292  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
293  | 
fun rewrite_case_bool_conv ctxt =  | 
| 58061 | 294  | 
SMT_Util.if_exists_conv is_case_bool (Conv.top_conv unfold_conv ctxt)  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
295  | 
|
| 67149 | 296  | 
val setup_case_bool = SMT_Builtin.add_builtin_fun_ext'' \<^const_name>\<open>bool.case_bool\<close>  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
297  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
298  | 
end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
299  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
300  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
301  | 
(** unfold abs, min and max **)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
302  | 
|
| 
56981
 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 
blanchet 
parents: 
56245 
diff
changeset
 | 
303  | 
val abs_min_max_table = [  | 
| 67149 | 304  | 
  (\<^const_name>\<open>min\<close>, @{thm min_def_raw}),
 | 
305  | 
  (\<^const_name>\<open>max\<close>, @{thm max_def_raw}),
 | 
|
306  | 
  (\<^const_name>\<open>abs\<close>, @{thm abs_if_raw})]
 | 
|
| 
56981
 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 
blanchet 
parents: 
56245 
diff
changeset
 | 
307  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
308  | 
local  | 
| 67149 | 309  | 
fun abs_min_max ctxt (Const (n, Type (\<^type_name>\<open>fun\<close>, [T, _]))) =  | 
| 
56981
 
3ef45ce002b5
honor original format of conjecture or hypotheses in Z3-to-Isar proofs
 
blanchet 
parents: 
56245 
diff
changeset
 | 
310  | 
(case AList.lookup (op =) abs_min_max_table n of  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
311  | 
NONE => NONE  | 
| 58061 | 312  | 
| SOME thm => if SMT_Builtin.is_builtin_typ_ext ctxt T then SOME thm else NONE)  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
313  | 
| abs_min_max _ _ = NONE  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
314  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
315  | 
fun unfold_amm_conv ctxt ct =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
316  | 
(case abs_min_max ctxt (Term.head_of (Thm.term_of ct)) of  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
317  | 
SOME thm => expand_head_conv (Conv.rewr_conv thm)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
318  | 
| NONE => Conv.all_conv) ct  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
319  | 
in  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
320  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
321  | 
fun unfold_abs_min_max_conv ctxt =  | 
| 58061 | 322  | 
SMT_Util.if_exists_conv (is_some o abs_min_max ctxt) (Conv.top_conv unfold_amm_conv ctxt)  | 
| 57230 | 323  | 
|
| 58061 | 324  | 
val setup_abs_min_max = fold (SMT_Builtin.add_builtin_fun_ext'' o fst) abs_min_max_table  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
325  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
326  | 
end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
327  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
328  | 
|
| 66298 | 329  | 
(** embedding of standard natural number operations into integer operations **)  | 
330  | 
||
331  | 
local  | 
|
332  | 
val simple_nat_ops = [  | 
|
| 74382 | 333  | 
\<^Const>\<open>HOL.eq \<^Type>\<open>nat\<close>\<close>, \<^Const>\<open>less \<^Type>\<open>nat\<close>\<close>, \<^Const>\<open>less_eq \<^Type>\<open>nat\<close>\<close>,  | 
334  | 
\<^Const>\<open>Suc\<close>, \<^Const>\<open>plus \<^Type>\<open>nat\<close>\<close>, \<^Const>\<open>minus \<^Type>\<open>nat\<close>\<close>]  | 
|
| 66298 | 335  | 
|
| 
67972
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
336  | 
val nat_consts = simple_nat_ops @  | 
| 74382 | 337  | 
[\<^Const>\<open>numeral \<^Type>\<open>nat\<close>\<close>, \<^Const>\<open>zero_class.zero \<^Type>\<open>nat\<close>\<close>, \<^Const>\<open>one_class.one \<^Type>\<open>nat\<close>\<close>] @  | 
338  | 
[\<^Const>\<open>times \<^Type>\<open>nat\<close>\<close>, \<^Const>\<open>divide \<^Type>\<open>nat\<close>\<close>, \<^Const>\<open>modulo \<^Type>\<open>nat\<close>\<close>]  | 
|
| 66298 | 339  | 
|
340  | 
val is_nat_const = member (op aconv) nat_consts  | 
|
341  | 
||
| 
67972
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
342  | 
  val nat_int_thm = Thm.symmetric (mk_meta_eq @{thm nat_int})
 | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
343  | 
  val nat_int_comp_thms = map mk_meta_eq @{thms nat_int_comparison}
 | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
344  | 
  val int_ops_thms = map mk_meta_eq @{thms int_ops}
 | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
345  | 
  val int_if_thm = mk_meta_eq @{thm int_if}
 | 
| 66298 | 346  | 
|
| 
67972
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
347  | 
fun if_conv cv1 cv2 = Conv.combination_conv (Conv.combination_conv (Conv.arg_conv cv1) cv2) cv2  | 
| 66298 | 348  | 
|
| 
67972
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
349  | 
fun int_ops_conv cv ctxt ct =  | 
| 66298 | 350  | 
(case Thm.term_of ct of  | 
| 74382 | 351  | 
\<^Const_>\<open>of_nat \<^Type>\<open>int\<close> for \<open>\<^Const_>\<open>If _ for _ _ _\<close>\<close>\<close> =>  | 
| 
67972
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
352  | 
Conv.rewr_conv int_if_thm then_conv  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
353  | 
if_conv (cv ctxt) (int_ops_conv cv ctxt)  | 
| 74382 | 354  | 
| \<^Const>\<open>of_nat \<^Type>\<open>int\<close> for _\<close> =>  | 
| 
67972
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
355  | 
(Conv.rewrs_conv int_ops_thms then_conv  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
356  | 
Conv.top_sweep_conv (int_ops_conv cv) ctxt) else_conv  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
357  | 
Conv.arg_conv (Conv.sub_conv cv ctxt)  | 
| 66298 | 358  | 
| _ => Conv.no_conv) ct  | 
359  | 
||
| 
67972
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
360  | 
  val unfold_nat_let_conv = Conv.rewr_conv @{lemma "Let (n::nat) f \<equiv> f n" by (rule Let_def)}
 | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
361  | 
val drop_nat_int_conv = Conv.rewr_conv (Thm.symmetric nat_int_thm)  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
362  | 
|
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
363  | 
fun nat_to_int_conv ctxt ct = (  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
364  | 
Conv.top_conv (K (Conv.try_conv unfold_nat_let_conv)) ctxt then_conv  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
365  | 
Conv.top_sweep_conv nat_conv ctxt then_conv  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
366  | 
Conv.top_conv (K (Conv.try_conv drop_nat_int_conv)) ctxt) ct  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
367  | 
|
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
368  | 
and nat_conv ctxt ct = (  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
369  | 
Conv.rewrs_conv (nat_int_thm :: nat_int_comp_thms) then_conv  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
370  | 
Conv.top_sweep_conv (int_ops_conv nat_to_int_conv) ctxt) ct  | 
| 66298 | 371  | 
|
| 
67972
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
372  | 
fun add_int_of_nat vs ct cu (q, cts) =  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
373  | 
(case Thm.term_of ct of  | 
| 74382 | 374  | 
\<^Const>\<open>of_nat \<^Type>\<open>int\<close>\<close> =>  | 
| 
67972
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
375  | 
if Term.exists_subterm (member (op aconv) vs) (Thm.term_of cu) then (true, cts)  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
376  | 
else (q, insert (op aconvc) cu cts)  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
377  | 
| _ => (q, cts))  | 
| 66298 | 378  | 
|
| 
67972
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
379  | 
fun add_apps f vs ct =  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
380  | 
(case Thm.term_of ct of  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
381  | 
_ $ _ =>  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
382  | 
let val (cu1, cu2) = Thm.dest_comb ct  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
383  | 
in f vs cu1 cu2 #> add_apps f vs cu1 #> add_apps f vs cu2 end  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
384  | 
| Abs _ =>  | 
| 
74525
 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 
wenzelm 
parents: 
74518 
diff
changeset
 | 
385  | 
let val (cv, cu) = Thm.dest_abs_global ct  | 
| 
67972
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
386  | 
in add_apps f (Thm.term_of cv :: vs) cu end  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
387  | 
| _ => I)  | 
| 66298 | 388  | 
|
| 
67972
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
389  | 
  val int_thm = @{lemma "(0::int) <= int (n::nat)" by simp}
 | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
390  | 
  val nat_int_thms = @{lemma
 | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
391  | 
"\<forall>n::nat. (0::int) <= int n"  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
392  | 
"\<forall>n::nat. nat (int n) = n"  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
393  | 
"\<forall>i::int. int (nat i) = (if 0 <= i then i else 0)"  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
394  | 
by simp_all}  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
395  | 
val var = Term.dest_Var (Thm.term_of (funpow 3 Thm.dest_arg (Thm.cprop_of int_thm)))  | 
| 66298 | 396  | 
in  | 
397  | 
||
| 
67972
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
398  | 
fun nat_as_int_conv ctxt = SMT_Util.if_exists_conv is_nat_const (nat_to_int_conv ctxt)  | 
| 66298 | 399  | 
|
| 
67972
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
400  | 
fun add_int_of_nat_constraints thms =  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
401  | 
let val (q, cts) = fold (add_apps add_int_of_nat [] o Thm.cprop_of) thms (false, [])  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
402  | 
in  | 
| 
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
403  | 
if q then (thms, nat_int_thms)  | 
| 77879 | 404  | 
else (thms, map (fn ct => Thm.instantiate (TVars.empty, Vars.make1 (var, ct)) int_thm) cts)  | 
| 
67972
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
405  | 
end  | 
| 66298 | 406  | 
|
407  | 
val setup_nat_as_int =  | 
|
| 67149 | 408  | 
SMT_Builtin.add_builtin_typ_ext (\<^typ>\<open>nat\<close>,  | 
| 66298 | 409  | 
fn ctxt => K (Config.get ctxt SMT_Config.nat_as_int)) #>  | 
| 
67972
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
410  | 
fold (SMT_Builtin.add_builtin_fun_ext' o Term.dest_Const) simple_nat_ops  | 
| 66298 | 411  | 
|
412  | 
end  | 
|
413  | 
||
414  | 
||
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
415  | 
(** normalize numerals **)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
416  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
417  | 
local  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
418  | 
(*  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
419  | 
rewrite Numeral1 into 1  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
420  | 
rewrite - 0 into 0  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
421  | 
*)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
422  | 
|
| 67149 | 423  | 
fun is_irregular_number (Const (\<^const_name>\<open>numeral\<close>, _) $ Const (\<^const_name>\<open>num.One\<close>, _)) =  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
424  | 
true  | 
| 67149 | 425  | 
| is_irregular_number (Const (\<^const_name>\<open>uminus\<close>, _) $ Const (\<^const_name>\<open>Groups.zero\<close>, _)) =  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
426  | 
true  | 
| 
57696
 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 
blanchet 
parents: 
57230 
diff
changeset
 | 
427  | 
| is_irregular_number _ = false  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
428  | 
|
| 58061 | 429  | 
fun is_strange_number ctxt t = is_irregular_number t andalso SMT_Builtin.is_builtin_num ctxt t  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
430  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
431  | 
val proper_num_ss =  | 
| 69593 | 432  | 
    simpset_of (put_simpset HOL_ss \<^context> addsimps @{thms Num.numeral_One minus_zero})
 | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
433  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
434  | 
fun norm_num_conv ctxt =  | 
| 58061 | 435  | 
SMT_Util.if_conv (is_strange_number ctxt) (Simplifier.rewrite (put_simpset proper_num_ss ctxt))  | 
| 56090 | 436  | 
Conv.no_conv  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
437  | 
in  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
438  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
439  | 
fun normalize_numerals_conv ctxt =  | 
| 58061 | 440  | 
SMT_Util.if_exists_conv (is_strange_number ctxt) (Conv.top_sweep_conv norm_num_conv ctxt)  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
441  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
442  | 
end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
443  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
444  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
445  | 
(** combined unfoldings and rewritings **)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
446  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
447  | 
fun burrow_ids f ithms =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
448  | 
let  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
449  | 
val (is, thms) = split_list ithms  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
450  | 
val (thms', extra_thms) = f thms  | 
| 
75274
 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 
desharna 
parents: 
74561 
diff
changeset
 | 
451  | 
in (is ~~ thms') @ map (pair (~1, SMT_Util.Axiom)) extra_thms end  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
452  | 
|
| 66298 | 453  | 
fun unfold_conv ctxt =  | 
454  | 
rewrite_case_bool_conv ctxt then_conv  | 
|
455  | 
unfold_abs_min_max_conv ctxt then_conv  | 
|
456  | 
(if Config.get ctxt SMT_Config.nat_as_int then nat_as_int_conv ctxt  | 
|
457  | 
else Conv.all_conv) then_conv  | 
|
458  | 
Thm.beta_conversion true  | 
|
459  | 
||
460  | 
fun unfold_polymorph ctxt = map (apsnd (Conv.fconv_rule (unfold_conv ctxt)))  | 
|
461  | 
fun unfold_monomorph ctxt =  | 
|
462  | 
map (apsnd (Conv.fconv_rule (normalize_numerals_conv ctxt)))  | 
|
| 
67972
 
959b0aed2ce5
avoid adding unnecessary quantified lemmas when embedding natural number terms into integer terms: quantified lemmas can cause Z3 to produce complex proofs that are hard to replay in Isabelle
 
boehmes 
parents: 
67710 
diff
changeset
 | 
463  | 
#> Config.get ctxt SMT_Config.nat_as_int ? burrow_ids add_int_of_nat_constraints  | 
| 66298 | 464  | 
|
465  | 
||
466  | 
(* overall normalization *)  | 
|
467  | 
||
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
468  | 
type extra_norm = Proof.context -> thm list * thm list -> thm list * thm list  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
469  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
470  | 
structure Extra_Norms = Generic_Data  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
471  | 
(  | 
| 58061 | 472  | 
type T = extra_norm SMT_Util.dict  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
473  | 
val empty = []  | 
| 58061 | 474  | 
fun merge data = SMT_Util.dict_merge fst data  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
475  | 
)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
476  | 
|
| 58061 | 477  | 
fun add_extra_norm (cs, norm) = Extra_Norms.map (SMT_Util.dict_update (cs, norm))  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
478  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
479  | 
fun apply_extra_norms ctxt ithms =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
480  | 
let  | 
| 58061 | 481  | 
val cs = SMT_Config.solver_class_of ctxt  | 
482  | 
val es = SMT_Util.dict_lookup (Extra_Norms.get (Context.Proof ctxt)) cs  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
483  | 
in burrow_ids (fold (fn e => e ctxt) es o rpair []) ithms end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
484  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
485  | 
local  | 
| 67149 | 486  | 
val ignored = member (op =) [\<^const_name>\<open>All\<close>, \<^const_name>\<open>Ex\<close>,  | 
487  | 
\<^const_name>\<open>Let\<close>, \<^const_name>\<open>If\<close>, \<^const_name>\<open>HOL.eq\<close>]  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
488  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
489  | 
val schematic_consts_of =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
490  | 
let  | 
| 74382 | 491  | 
fun collect \<^Const_>\<open>trigger for p t\<close> = collect_trigger p #> collect t  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
492  | 
| collect (t $ u) = collect t #> collect u  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
493  | 
| collect (Abs (_, _, t)) = collect t  | 
| 57230 | 494  | 
| collect (t as Const (n, _)) =  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
495  | 
if not (ignored n) then Monomorph.add_schematic_consts_of t else I  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
496  | 
| collect _ = I  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
497  | 
and collect_trigger t =  | 
| 58061 | 498  | 
let val dest = these o try SMT_Util.dest_symb_list  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
499  | 
in fold (fold collect_pat o dest) (dest t) end  | 
| 67149 | 500  | 
and collect_pat (Const (\<^const_name>\<open>pat\<close>, _) $ t) = collect t  | 
501  | 
| collect_pat (Const (\<^const_name>\<open>nopat\<close>, _) $ t) = collect t  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
502  | 
| collect_pat _ = I  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
503  | 
in (fn t => collect t Symtab.empty) end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
504  | 
in  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
505  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
506  | 
fun monomorph ctxt xthms =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
507  | 
let val (xs, thms) = split_list xthms  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
508  | 
in  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
509  | 
map (pair 1) thms  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
510  | 
|> Monomorph.monomorph schematic_consts_of ctxt  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
511  | 
|> maps (uncurry (map o pair)) o map2 pair xs o map (map snd)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
512  | 
end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
513  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
514  | 
end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
515  | 
|
| 56106 | 516  | 
fun normalize ctxt wthms =  | 
517  | 
wthms  | 
|
| 
75274
 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 
desharna 
parents: 
74561 
diff
changeset
 | 
518  | 
|> map_index (fn (n, (role, thm)) => ((n, role), thm))  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
519  | 
|> gen_normalize ctxt  | 
| 
57696
 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 
blanchet 
parents: 
57230 
diff
changeset
 | 
520  | 
|> unfold_polymorph ctxt  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
521  | 
|> monomorph ctxt  | 
| 
57696
 
fb71c6f100f8
do not embed 'nat' into 'int's in 'smt2' method -- this is highly inefficient and decreases the Sledgehammer success rate significantly
 
blanchet 
parents: 
57230 
diff
changeset
 | 
522  | 
|> unfold_monomorph ctxt  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
523  | 
|> apply_extra_norms ctxt  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
524  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
525  | 
val _ = Theory.setup (Context.theory_map (  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
526  | 
setup_atomize #>  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
527  | 
setup_unfolded_quants #>  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
528  | 
setup_trigger #>  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
529  | 
setup_case_bool #>  | 
| 66298 | 530  | 
setup_abs_min_max #>  | 
531  | 
setup_nat_as_int))  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
532  | 
|
| 57229 | 533  | 
end;  |