| author | boehmes | 
| Mon, 20 Dec 2010 22:02:57 +0100 | |
| changeset 41328 | 6792a5c92a58 | 
| parent 41327 | 49feace87649 | 
| child 42183 | 173b0f488428 | 
| permissions | -rw-r--r-- | 
| 36898 | 1  | 
(* Title: HOL/Tools/SMT/smt_normalize.ML  | 
2  | 
Author: Sascha Boehme, TU Muenchen  | 
|
3  | 
||
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
4  | 
Normalization steps on theorems required by SMT solvers.  | 
| 36898 | 5  | 
*)  | 
6  | 
||
7  | 
signature SMT_NORMALIZE =  | 
|
8  | 
sig  | 
|
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
36898 
diff
changeset
 | 
9  | 
val atomize_conv: Proof.context -> conv  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
10  | 
type extra_norm = Proof.context -> thm list * thm list -> thm list * thm list  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
11  | 
val add_extra_norm: SMT_Utils.class * extra_norm -> Context.generic ->  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
12  | 
Context.generic  | 
| 
41300
 
528f5d00b542
perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
 
boehmes 
parents: 
41280 
diff
changeset
 | 
13  | 
val normalize: (int * (int option * thm)) list -> Proof.context ->  | 
| 
 
528f5d00b542
perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
 
boehmes 
parents: 
41280 
diff
changeset
 | 
14  | 
(int * thm) list * Proof.context  | 
| 
41059
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40686 
diff
changeset
 | 
15  | 
val setup: theory -> theory  | 
| 36898 | 16  | 
end  | 
17  | 
||
18  | 
structure SMT_Normalize: SMT_NORMALIZE =  | 
|
19  | 
struct  | 
|
20  | 
||
21  | 
||
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
22  | 
(* general theorem normalizations *)  | 
| 36898 | 23  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
24  | 
(** instantiate elimination rules **)  | 
| 
40685
 
dcb27631cb45
instantiate elimination rules (reduces number of quantified variables, and makes such theorems better amenable for SMT solvers)
 
boehmes 
parents: 
40681 
diff
changeset
 | 
25  | 
|
| 
 
dcb27631cb45
instantiate elimination rules (reduces number of quantified variables, and makes such theorems better amenable for SMT solvers)
 
boehmes 
parents: 
40681 
diff
changeset
 | 
26  | 
local  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
27  | 
val (cpfalse, cfalse) =  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
28  | 
    `SMT_Utils.mk_cprop (Thm.cterm_of @{theory} @{const False})
 | 
| 
40685
 
dcb27631cb45
instantiate elimination rules (reduces number of quantified variables, and makes such theorems better amenable for SMT solvers)
 
boehmes 
parents: 
40681 
diff
changeset
 | 
29  | 
|
| 
 
dcb27631cb45
instantiate elimination rules (reduces number of quantified variables, and makes such theorems better amenable for SMT solvers)
 
boehmes 
parents: 
40681 
diff
changeset
 | 
30  | 
fun inst f ct thm =  | 
| 
 
dcb27631cb45
instantiate elimination rules (reduces number of quantified variables, and makes such theorems better amenable for SMT solvers)
 
boehmes 
parents: 
40681 
diff
changeset
 | 
31  | 
let val cv = f (Drule.strip_imp_concl (Thm.cprop_of thm))  | 
| 
 
dcb27631cb45
instantiate elimination rules (reduces number of quantified variables, and makes such theorems better amenable for SMT solvers)
 
boehmes 
parents: 
40681 
diff
changeset
 | 
32  | 
in Thm.instantiate ([], [(cv, ct)]) thm end  | 
| 
 
dcb27631cb45
instantiate elimination rules (reduces number of quantified variables, and makes such theorems better amenable for SMT solvers)
 
boehmes 
parents: 
40681 
diff
changeset
 | 
33  | 
in  | 
| 
 
dcb27631cb45
instantiate elimination rules (reduces number of quantified variables, and makes such theorems better amenable for SMT solvers)
 
boehmes 
parents: 
40681 
diff
changeset
 | 
34  | 
|
| 
 
dcb27631cb45
instantiate elimination rules (reduces number of quantified variables, and makes such theorems better amenable for SMT solvers)
 
boehmes 
parents: 
40681 
diff
changeset
 | 
35  | 
fun instantiate_elim thm =  | 
| 
 
dcb27631cb45
instantiate elimination rules (reduces number of quantified variables, and makes such theorems better amenable for SMT solvers)
 
boehmes 
parents: 
40681 
diff
changeset
 | 
36  | 
(case Thm.concl_of thm of  | 
| 
 
dcb27631cb45
instantiate elimination rules (reduces number of quantified variables, and makes such theorems better amenable for SMT solvers)
 
boehmes 
parents: 
40681 
diff
changeset
 | 
37  | 
    @{const Trueprop} $ Var (_, @{typ bool}) => inst Thm.dest_arg cfalse thm
 | 
| 
 
dcb27631cb45
instantiate elimination rules (reduces number of quantified variables, and makes such theorems better amenable for SMT solvers)
 
boehmes 
parents: 
40681 
diff
changeset
 | 
38  | 
| Var _ => inst I cpfalse thm  | 
| 
 
dcb27631cb45
instantiate elimination rules (reduces number of quantified variables, and makes such theorems better amenable for SMT solvers)
 
boehmes 
parents: 
40681 
diff
changeset
 | 
39  | 
| _ => thm)  | 
| 
 
dcb27631cb45
instantiate elimination rules (reduces number of quantified variables, and makes such theorems better amenable for SMT solvers)
 
boehmes 
parents: 
40681 
diff
changeset
 | 
40  | 
|
| 
 
dcb27631cb45
instantiate elimination rules (reduces number of quantified variables, and makes such theorems better amenable for SMT solvers)
 
boehmes 
parents: 
40681 
diff
changeset
 | 
41  | 
end  | 
| 
 
dcb27631cb45
instantiate elimination rules (reduces number of quantified variables, and makes such theorems better amenable for SMT solvers)
 
boehmes 
parents: 
40681 
diff
changeset
 | 
42  | 
|
| 
 
dcb27631cb45
instantiate elimination rules (reduces number of quantified variables, and makes such theorems better amenable for SMT solvers)
 
boehmes 
parents: 
40681 
diff
changeset
 | 
43  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
44  | 
(** normalize definitions **)  | 
| 36898 | 45  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
46  | 
fun norm_def thm =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
47  | 
(case Thm.prop_of thm of  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
48  | 
    @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ _ $ Abs _) =>
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
49  | 
      norm_def (thm RS @{thm fun_cong})
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
50  | 
  | Const (@{const_name "=="}, _) $ _ $ Abs _ =>
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
51  | 
      norm_def (thm RS @{thm meta_eq_to_obj_eq})
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
52  | 
| _ => thm)  | 
| 36898 | 53  | 
|
54  | 
||
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
55  | 
(** atomization **)  | 
| 36898 | 56  | 
|
57  | 
fun atomize_conv ctxt ct =  | 
|
58  | 
(case Thm.term_of ct of  | 
|
| 
40579
 
98ebd2300823
use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
 
boehmes 
parents: 
40424 
diff
changeset
 | 
59  | 
    @{const "==>"} $ _ $ _ =>
 | 
| 36898 | 60  | 
Conv.binop_conv (atomize_conv ctxt) then_conv  | 
61  | 
      Conv.rewr_conv @{thm atomize_imp}
 | 
|
62  | 
  | Const (@{const_name "=="}, _) $ _ $ _ =>
 | 
|
63  | 
Conv.binop_conv (atomize_conv ctxt) then_conv  | 
|
64  | 
      Conv.rewr_conv @{thm atomize_eq}
 | 
|
65  | 
  | Const (@{const_name all}, _) $ Abs _ =>
 | 
|
| 
36936
 
c52d1c130898
incorporated further conversions and conversionals, after some minor tuning;
 
wenzelm 
parents: 
36899 
diff
changeset
 | 
66  | 
Conv.binder_conv (atomize_conv o snd) ctxt then_conv  | 
| 36898 | 67  | 
      Conv.rewr_conv @{thm atomize_all}
 | 
68  | 
| _ => Conv.all_conv) ct  | 
|
69  | 
||
| 
41059
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40686 
diff
changeset
 | 
70  | 
val setup_atomize =  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
71  | 
  fold SMT_Builtin.add_builtin_fun_ext'' [@{const_name "==>"},
 | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
72  | 
    @{const_name "=="}, @{const_name all}, @{const_name Trueprop}]
 | 
| 
41059
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40686 
diff
changeset
 | 
73  | 
|
| 36898 | 74  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
75  | 
(** unfold special quantifiers **)  | 
| 36898 | 76  | 
|
77  | 
local  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
78  | 
  val ex1_def = mk_meta_eq @{lemma
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
79  | 
"Ex1 = (%P. EX x. P x & (ALL y. P y --> y = x))"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
80  | 
by (rule ext) (simp only: Ex1_def)}  | 
| 36898 | 81  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
82  | 
  val ball_def = mk_meta_eq @{lemma "Ball = (%A P. ALL x. x : A --> P x)"
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
83  | 
by (rule ext)+ (rule Ball_def)}  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
84  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
85  | 
  val bex_def = mk_meta_eq @{lemma "Bex = (%A P. EX x. x : A & P x)"
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
86  | 
by (rule ext)+ (rule Bex_def)}  | 
| 36898 | 87  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
88  | 
  val special_quants = [(@{const_name Ex1}, ex1_def),
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
89  | 
    (@{const_name Ball}, ball_def), (@{const_name Bex}, bex_def)]
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
90  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
91  | 
fun special_quant (Const (n, _)) = AList.lookup (op =) special_quants n  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
92  | 
| special_quant _ = NONE  | 
| 36898 | 93  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
94  | 
fun special_quant_conv _ ct =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
95  | 
(case special_quant (Thm.term_of ct) of  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
96  | 
SOME thm => Conv.rewr_conv thm  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
97  | 
| NONE => Conv.all_conv) ct  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
98  | 
in  | 
| 36898 | 99  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
100  | 
fun unfold_special_quants_conv ctxt =  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
101  | 
SMT_Utils.if_exists_conv (is_some o special_quant)  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
102  | 
(Conv.top_conv special_quant_conv ctxt)  | 
| 36898 | 103  | 
|
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
104  | 
val setup_unfolded_quants =  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
105  | 
fold (SMT_Builtin.add_builtin_fun_ext'' o fst) special_quants  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
106  | 
|
| 36898 | 107  | 
end  | 
108  | 
||
109  | 
||
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
110  | 
(** trigger inference **)  | 
| 36898 | 111  | 
|
112  | 
local  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
113  | 
(*** check trigger syntax ***)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
114  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
115  | 
  fun dest_trigger (Const (@{const_name pat}, _) $ _) = SOME true
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
116  | 
    | dest_trigger (Const (@{const_name nopat}, _) $ _) = SOME false
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
117  | 
| dest_trigger _ = NONE  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
118  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
119  | 
fun eq_list [] = false  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
120  | 
| eq_list (b :: bs) = forall (equal b) bs  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
121  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
122  | 
fun proper_trigger t =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
123  | 
t  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
124  | 
|> these o try HOLogic.dest_list  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
125  | 
|> map (map_filter dest_trigger o these o try HOLogic.dest_list)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
126  | 
|> (fn [] => false | bss => forall eq_list bss)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
127  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
128  | 
fun proper_quant inside f t =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
129  | 
(case t of  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
130  | 
      Const (@{const_name All}, _) $ Abs (_, _, u) => proper_quant true f u
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
131  | 
    | Const (@{const_name Ex}, _) $ Abs (_, _, u) => proper_quant true f u
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
132  | 
    | @{const trigger} $ p $ u =>
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
133  | 
(if inside then f p else false) andalso proper_quant false f u  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
134  | 
| Abs (_, _, u) => proper_quant false f u  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
135  | 
| u1 $ u2 => proper_quant false f u1 andalso proper_quant false f u2  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
136  | 
| _ => true)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
137  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
138  | 
fun check_trigger_error ctxt t =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
139  | 
    error ("SMT triggers must only occur under quantifier and multipatterns " ^
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
140  | 
"must have the same kind: " ^ Syntax.string_of_term ctxt t)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
141  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
142  | 
fun check_trigger_conv ctxt ct =  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
143  | 
if proper_quant false proper_trigger (SMT_Utils.term_of ct) then  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
144  | 
Conv.all_conv ct  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
145  | 
else check_trigger_error ctxt (Thm.term_of ct)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
146  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
147  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
148  | 
(*** infer simple triggers ***)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
149  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
150  | 
fun dest_cond_eq ct =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
151  | 
(case Thm.term_of ct of  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
152  | 
      Const (@{const_name HOL.eq}, _) $ _ $ _ => Thm.dest_binop ct
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
153  | 
    | @{const HOL.implies} $ _ $ _ => dest_cond_eq (Thm.dest_arg ct)
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
154  | 
    | _ => raise CTERM ("no equation", [ct]))
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
155  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
156  | 
fun get_constrs thy (Type (n, _)) = these (Datatype.get_constrs thy n)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
157  | 
| get_constrs _ _ = []  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
158  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
159  | 
fun is_constr thy (n, T) =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
160  | 
let fun match (m, U) = m = n andalso Sign.typ_instance thy (T, U)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
161  | 
in can (the o find_first match o get_constrs thy o Term.body_type) T end  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
162  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
163  | 
fun is_constr_pat thy t =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
164  | 
(case Term.strip_comb t of  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
165  | 
(Free _, []) => true  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
166  | 
| (Const c, ts) => is_constr thy c andalso forall (is_constr_pat thy) ts  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
167  | 
| _ => false)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
168  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
169  | 
fun is_simp_lhs ctxt t =  | 
| 36898 | 170  | 
(case Term.strip_comb t of  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
171  | 
(Const c, ts as _ :: _) =>  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
172  | 
not (SMT_Builtin.is_builtin_fun_ext ctxt c ts) andalso  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
173  | 
forall (is_constr_pat (ProofContext.theory_of ctxt)) ts  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
174  | 
| _ => false)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
175  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
176  | 
fun has_all_vars vs t =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
177  | 
subset (op aconv) (vs, map Free (Term.add_frees t []))  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
178  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
179  | 
fun minimal_pats vs ct =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
180  | 
if has_all_vars vs (Thm.term_of ct) then  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
181  | 
(case Thm.term_of ct of  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
182  | 
_ $ _ =>  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
183  | 
(case pairself (minimal_pats vs) (Thm.dest_comb ct) of  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
184  | 
([], []) => [[ct]]  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
185  | 
| (ctss, ctss') => union (eq_set (op aconvc)) ctss ctss')  | 
| 
41174
 
10eb369f8c01
fixed trigger inference: testing if a theorem already has a trigger was too strict;
 
boehmes 
parents: 
41173 
diff
changeset
 | 
186  | 
| _ => [])  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
187  | 
else []  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
188  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
189  | 
fun proper_mpat _ _ _ [] = false  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
190  | 
| proper_mpat thy gen u cts =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
191  | 
let  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
192  | 
val tps = (op ~~) (`gen (map Thm.term_of cts))  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
193  | 
fun some_match u = tps |> exists (fn (t', t) =>  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
194  | 
Pattern.matches thy (t', u) andalso not (t aconv u))  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
195  | 
in not (Term.exists_subterm some_match u) end  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
196  | 
|
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
197  | 
val pat =  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
198  | 
    SMT_Utils.mk_const_pat @{theory} @{const_name SMT.pat} SMT_Utils.destT1
 | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
199  | 
fun mk_pat ct = Thm.capply (SMT_Utils.instT' ct pat) ct  | 
| 36898 | 200  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
201  | 
  fun mk_clist T = pairself (Thm.cterm_of @{theory})
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
202  | 
(HOLogic.cons_const T, HOLogic.nil_const T)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
203  | 
fun mk_list (ccons, cnil) f cts = fold_rev (Thm.mk_binop ccons o f) cts cnil  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
204  | 
  val mk_pat_list = mk_list (mk_clist @{typ SMT.pattern})
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
205  | 
  val mk_mpat_list = mk_list (mk_clist @{typ "SMT.pattern list"})  
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
206  | 
fun mk_trigger ctss = mk_mpat_list (mk_pat_list mk_pat) ctss  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
207  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
208  | 
val trigger_eq =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
209  | 
    mk_meta_eq @{lemma "p = SMT.trigger t p" by (simp add: trigger_def)}
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
210  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
211  | 
fun insert_trigger_conv [] ct = Conv.all_conv ct  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
212  | 
| insert_trigger_conv ctss ct =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
213  | 
let val (ctr, cp) = Thm.dest_binop (Thm.rhs_of trigger_eq) ||> rpair ct  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
214  | 
in Thm.instantiate ([], [cp, (ctr, mk_trigger ctss)]) trigger_eq end  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
215  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
216  | 
fun infer_trigger_eq_conv outer_ctxt (ctxt, cvs) ct =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
217  | 
let  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
218  | 
val (lhs, rhs) = dest_cond_eq ct  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
219  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
220  | 
val vs = map Thm.term_of cvs  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
221  | 
val thy = ProofContext.theory_of ctxt  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
222  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
223  | 
fun get_mpats ct =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
224  | 
if is_simp_lhs ctxt (Thm.term_of ct) then minimal_pats vs ct  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
225  | 
else []  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
226  | 
val gen = Variable.export_terms ctxt outer_ctxt  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
227  | 
val filter_mpats = filter (proper_mpat thy gen (Thm.term_of rhs))  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
228  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
229  | 
in insert_trigger_conv (filter_mpats (get_mpats lhs)) ct end  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
230  | 
|
| 
41174
 
10eb369f8c01
fixed trigger inference: testing if a theorem already has a trigger was too strict;
 
boehmes 
parents: 
41173 
diff
changeset
 | 
231  | 
  fun has_trigger (@{const SMT.trigger} $ _ $ _) = true
 | 
| 
 
10eb369f8c01
fixed trigger inference: testing if a theorem already has a trigger was too strict;
 
boehmes 
parents: 
41173 
diff
changeset
 | 
232  | 
| has_trigger _ = false  | 
| 
 
10eb369f8c01
fixed trigger inference: testing if a theorem already has a trigger was too strict;
 
boehmes 
parents: 
41173 
diff
changeset
 | 
233  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
234  | 
fun try_trigger_conv cv ct =  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
235  | 
if SMT_Utils.under_quant has_trigger (SMT_Utils.term_of ct) then  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
236  | 
Conv.all_conv ct  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
237  | 
else Conv.try_conv cv ct  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
238  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
239  | 
fun infer_trigger_conv ctxt =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
240  | 
if Config.get ctxt SMT_Config.infer_triggers then  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
241  | 
try_trigger_conv  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
242  | 
(SMT_Utils.under_quant_conv (infer_trigger_eq_conv ctxt) ctxt)  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
243  | 
else Conv.all_conv  | 
| 36898 | 244  | 
in  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
245  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
246  | 
fun trigger_conv ctxt =  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
247  | 
SMT_Utils.prop_conv  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
248  | 
(check_trigger_conv ctxt then_conv infer_trigger_conv ctxt)  | 
| 36898 | 249  | 
|
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
250  | 
val setup_trigger =  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
251  | 
fold SMT_Builtin.add_builtin_fun_ext''  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
252  | 
    [@{const_name SMT.pat}, @{const_name SMT.nopat}, @{const_name SMT.trigger}]
 | 
| 36898 | 253  | 
|
254  | 
end  | 
|
255  | 
||
256  | 
||
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
257  | 
(** adding quantifier weights **)  | 
| 
39483
 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 
boehmes 
parents: 
38864 
diff
changeset
 | 
258  | 
|
| 
 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 
boehmes 
parents: 
38864 
diff
changeset
 | 
259  | 
local  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
260  | 
(*** check weight syntax ***)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
261  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
262  | 
val has_no_weight =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
263  | 
    not o Term.exists_subterm (fn @{const SMT.weight} => true | _ => false)
 | 
| 
39483
 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 
boehmes 
parents: 
38864 
diff
changeset
 | 
264  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
265  | 
  fun is_weight (@{const SMT.weight} $ w $ t) =
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
266  | 
(case try HOLogic.dest_number w of  | 
| 
41173
 
7c6178d45cc8
fixed checking and translation of weights (previously, weights occurring in terms were rejected, and weight numbers were unintended translated into Vars)
 
boehmes 
parents: 
41126 
diff
changeset
 | 
267  | 
SOME (_, i) => i >= 0 andalso has_no_weight t  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
268  | 
| _ => false)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
269  | 
| is_weight t = has_no_weight t  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
270  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
271  | 
  fun proper_trigger (@{const SMT.trigger} $ _ $ t) = is_weight t
 | 
| 
41173
 
7c6178d45cc8
fixed checking and translation of weights (previously, weights occurring in terms were rejected, and weight numbers were unintended translated into Vars)
 
boehmes 
parents: 
41126 
diff
changeset
 | 
272  | 
| proper_trigger t = is_weight t  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
273  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
274  | 
fun check_weight_error ctxt t =  | 
| 
41173
 
7c6178d45cc8
fixed checking and translation of weights (previously, weights occurring in terms were rejected, and weight numbers were unintended translated into Vars)
 
boehmes 
parents: 
41126 
diff
changeset
 | 
275  | 
    error ("SMT weight must be a non-negative number and must only occur " ^
 | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
276  | 
"under the top-most quantifier and an optional trigger: " ^  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
277  | 
Syntax.string_of_term ctxt t)  | 
| 
39483
 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 
boehmes 
parents: 
38864 
diff
changeset
 | 
278  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
279  | 
fun check_weight_conv ctxt ct =  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
280  | 
if SMT_Utils.under_quant proper_trigger (SMT_Utils.term_of ct) then  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
281  | 
Conv.all_conv ct  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
282  | 
else check_weight_error ctxt (Thm.term_of ct)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
283  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
284  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
285  | 
(*** insertion of weights ***)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
286  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
287  | 
fun under_trigger_conv cv ct =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
288  | 
(case Thm.term_of ct of  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
289  | 
      @{const SMT.trigger} $ _ $ _ => Conv.arg_conv cv
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
290  | 
| _ => cv) ct  | 
| 
39483
 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 
boehmes 
parents: 
38864 
diff
changeset
 | 
291  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
292  | 
val weight_eq =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
293  | 
    mk_meta_eq @{lemma "p = SMT.weight i p" by (simp add: weight_def)}
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
294  | 
fun mk_weight_eq w =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
295  | 
let val cv = Thm.dest_arg1 (Thm.rhs_of weight_eq)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
296  | 
in  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
297  | 
      Thm.instantiate ([], [(cv, Numeral.mk_cnumber @{ctyp int} w)]) weight_eq
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
298  | 
end  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
299  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
300  | 
fun add_weight_conv NONE _ = Conv.all_conv  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
301  | 
| add_weight_conv (SOME weight) ctxt =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
302  | 
let val cv = Conv.rewr_conv (mk_weight_eq weight)  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
303  | 
in SMT_Utils.under_quant_conv (K (under_trigger_conv cv)) ctxt end  | 
| 
39483
 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 
boehmes 
parents: 
38864 
diff
changeset
 | 
304  | 
in  | 
| 
 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 
boehmes 
parents: 
38864 
diff
changeset
 | 
305  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
306  | 
fun weight_conv weight ctxt =  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
307  | 
SMT_Utils.prop_conv  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
308  | 
(check_weight_conv ctxt then_conv add_weight_conv weight ctxt)  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
309  | 
|
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
310  | 
val setup_weight = SMT_Builtin.add_builtin_fun_ext'' @{const_name SMT.weight}
 | 
| 
39483
 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 
boehmes 
parents: 
38864 
diff
changeset
 | 
311  | 
|
| 
 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 
boehmes 
parents: 
38864 
diff
changeset
 | 
312  | 
end  | 
| 
 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 
boehmes 
parents: 
38864 
diff
changeset
 | 
313  | 
|
| 
 
9f0e5684f04b
add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
 
boehmes 
parents: 
38864 
diff
changeset
 | 
314  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
315  | 
(** combined general normalizations **)  | 
| 36898 | 316  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
317  | 
fun gen_normalize1_conv ctxt weight =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
318  | 
atomize_conv ctxt then_conv  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
319  | 
unfold_special_quants_conv ctxt then_conv  | 
| 
41327
 
49feace87649
added an additional beta reduction: unfolding of special quantifiers might leave terms unnormalized wrt to beta reductions
 
boehmes 
parents: 
41300 
diff
changeset
 | 
320  | 
Thm.beta_conversion true then_conv  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
321  | 
trigger_conv ctxt then_conv  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
322  | 
weight_conv weight ctxt  | 
| 36898 | 323  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
324  | 
fun gen_normalize1 ctxt weight thm =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
325  | 
thm  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
326  | 
|> instantiate_elim  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
327  | 
|> norm_def  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
328  | 
|> Conv.fconv_rule (Thm.beta_conversion true then_conv Thm.eta_conversion)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
329  | 
|> Drule.forall_intr_vars  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
330  | 
|> Conv.fconv_rule (gen_normalize1_conv ctxt weight)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
331  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
332  | 
fun drop_fact_warning ctxt =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
333  | 
let val pre = prefix "Warning: dropping assumption: "  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
334  | 
in SMT_Config.verbose_msg ctxt (pre o Display.string_of_thm ctxt) end  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
335  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
336  | 
fun gen_norm1_safe ctxt (i, (weight, thm)) =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
337  | 
if Config.get ctxt SMT_Config.drop_bad_facts then  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
338  | 
(case try (gen_normalize1 ctxt weight) thm of  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
339  | 
SOME thm' => SOME (i, thm')  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
340  | 
| NONE => (drop_fact_warning ctxt thm; NONE))  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
341  | 
else SOME (i, gen_normalize1 ctxt weight thm)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
342  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
343  | 
fun gen_normalize ctxt iwthms = map_filter (gen_norm1_safe ctxt) iwthms  | 
| 36898 | 344  | 
|
| 
41059
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40686 
diff
changeset
 | 
345  | 
|
| 
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40686 
diff
changeset
 | 
346  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
347  | 
(* unfolding of definitions and theory-specific rewritings *)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
348  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
349  | 
(** unfold trivial distincts **)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
350  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
351  | 
local  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
352  | 
  fun is_trivial_distinct (Const (@{const_name distinct}, _) $ t) =
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
353  | 
(case try HOLogic.dest_list t of  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
354  | 
SOME [] => true  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
355  | 
| SOME [_] => true  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
356  | 
| _ => false)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
357  | 
| is_trivial_distinct _ = false  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
358  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
359  | 
  val thms = map mk_meta_eq @{lemma
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
360  | 
"distinct [] = True"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
361  | 
"distinct [x] = True"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
362  | 
"distinct [x, y] = (x ~= y)"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
363  | 
by simp_all}  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
364  | 
fun distinct_conv _ =  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
365  | 
SMT_Utils.if_true_conv is_trivial_distinct (Conv.rewrs_conv thms)  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
366  | 
in  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
367  | 
|
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
368  | 
fun trivial_distinct_conv ctxt =  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
369  | 
SMT_Utils.if_exists_conv is_trivial_distinct  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
370  | 
(Conv.top_conv distinct_conv ctxt)  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
371  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
372  | 
end  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
373  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
374  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
375  | 
(** rewrite bool case expressions as if expressions **)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
376  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
377  | 
local  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
378  | 
  fun is_bool_case (Const (@{const_name "bool.bool_case"}, _)) = true
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
379  | 
| is_bool_case _ = false  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
380  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
381  | 
  val thm = mk_meta_eq @{lemma
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
382  | 
"bool_case = (%x y P. if P then x else y)" by (rule ext)+ simp}  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
383  | 
|
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
384  | 
fun unfold_conv _ = SMT_Utils.if_true_conv is_bool_case (Conv.rewr_conv thm)  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
385  | 
in  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
386  | 
|
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
387  | 
fun rewrite_bool_case_conv ctxt =  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
388  | 
SMT_Utils.if_exists_conv is_bool_case (Conv.top_conv unfold_conv ctxt)  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
389  | 
|
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
390  | 
val setup_bool_case =  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
391  | 
  SMT_Builtin.add_builtin_fun_ext'' @{const_name "bool.bool_case"}
 | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
392  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
393  | 
end  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
394  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
395  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
396  | 
(** unfold abs, min and max **)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
397  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
398  | 
local  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
399  | 
  val abs_def = mk_meta_eq @{lemma
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
400  | 
"abs = (%a::'a::abs_if. if a < 0 then - a else a)"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
401  | 
by (rule ext) (rule abs_if)}  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
402  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
403  | 
  val min_def = mk_meta_eq @{lemma "min = (%a b. if a <= b then a else b)"
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
404  | 
by (rule ext)+ (rule min_def)}  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
405  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
406  | 
  val max_def = mk_meta_eq  @{lemma "max = (%a b. if a <= b then b else a)"
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
407  | 
by (rule ext)+ (rule max_def)}  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
408  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
409  | 
  val defs = [(@{const_name min}, min_def), (@{const_name max}, max_def),
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
410  | 
    (@{const_name abs}, abs_def)]
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
411  | 
|
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
412  | 
fun is_builtinT ctxt T =  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
413  | 
SMT_Builtin.is_builtin_typ_ext ctxt (Term.domain_type T)  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
414  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
415  | 
fun abs_min_max ctxt (Const (n, T)) =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
416  | 
(case AList.lookup (op =) defs n of  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
417  | 
NONE => NONE  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
418  | 
| SOME thm => if is_builtinT ctxt T then SOME thm else NONE)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
419  | 
| abs_min_max _ _ = NONE  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
420  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
421  | 
fun unfold_amm_conv ctxt ct =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
422  | 
(case abs_min_max ctxt (Thm.term_of ct) of  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
423  | 
SOME thm => Conv.rewr_conv thm  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
424  | 
| NONE => Conv.all_conv) ct  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
425  | 
in  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
426  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
427  | 
fun unfold_abs_min_max_conv ctxt =  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
428  | 
SMT_Utils.if_exists_conv (is_some o abs_min_max ctxt)  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
429  | 
(Conv.top_conv unfold_amm_conv ctxt)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
430  | 
|
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
431  | 
val setup_abs_min_max = fold (SMT_Builtin.add_builtin_fun_ext'' o fst) defs  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
432  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
433  | 
end  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
434  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
435  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
436  | 
(** embedding of standard natural number operations into integer operations **)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
437  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
438  | 
local  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
439  | 
  val nat_embedding = @{lemma
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
440  | 
"ALL n. nat (int n) = n"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
441  | 
"ALL i. i >= 0 --> int (nat i) = i"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
442  | 
"ALL i. i < 0 --> int (nat i) = 0"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
443  | 
by simp_all}  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
444  | 
|
| 
41280
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41224 
diff
changeset
 | 
445  | 
val simple_nat_ops = [  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
446  | 
    @{const less (nat)}, @{const less_eq (nat)},
 | 
| 
41280
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41224 
diff
changeset
 | 
447  | 
    @{const Suc}, @{const plus (nat)}, @{const minus (nat)}]
 | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41224 
diff
changeset
 | 
448  | 
|
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41224 
diff
changeset
 | 
449  | 
val mult_nat_ops =  | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41224 
diff
changeset
 | 
450  | 
    [@{const times (nat)}, @{const div (nat)}, @{const mod (nat)}]
 | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41224 
diff
changeset
 | 
451  | 
|
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41224 
diff
changeset
 | 
452  | 
val nat_ops = simple_nat_ops @ mult_nat_ops  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
453  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
454  | 
  val nat_consts = nat_ops @ [@{const number_of (nat)},
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
455  | 
    @{const zero_class.zero (nat)}, @{const one_class.one (nat)}]
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
456  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
457  | 
  val nat_int_coercions = [@{const of_nat (int)}, @{const nat}]
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
458  | 
|
| 
41280
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41224 
diff
changeset
 | 
459  | 
val builtin_nat_ops = nat_int_coercions @ simple_nat_ops  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
460  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
461  | 
val is_nat_const = member (op aconv) nat_consts  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
462  | 
|
| 
41280
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41224 
diff
changeset
 | 
463  | 
  fun is_nat_const' @{const of_nat (int)} = true
 | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41224 
diff
changeset
 | 
464  | 
| is_nat_const' t = is_nat_const t  | 
| 
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41224 
diff
changeset
 | 
465  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
466  | 
  val expands = map mk_meta_eq @{lemma
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
467  | 
"0 = nat 0"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
468  | 
"1 = nat 1"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
469  | 
"(number_of :: int => nat) = (%i. nat (number_of i))"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
470  | 
"op < = (%a b. int a < int b)"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
471  | 
"op <= = (%a b. int a <= int b)"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
472  | 
"Suc = (%a. nat (int a + 1))"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
473  | 
"op + = (%a b. nat (int a + int b))"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
474  | 
"op - = (%a b. nat (int a - int b))"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
475  | 
"op * = (%a b. nat (int a * int b))"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
476  | 
"op div = (%a b. nat (int a div int b))"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
477  | 
"op mod = (%a b. nat (int a mod int b))"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
478  | 
by (auto intro!: ext simp add: nat_mult_distrib nat_div_distrib  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
479  | 
nat_mod_distrib)}  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
480  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
481  | 
  val ints = map mk_meta_eq @{lemma
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
482  | 
"int 0 = 0"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
483  | 
"int 1 = 1"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
484  | 
"int (Suc n) = int n + 1"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
485  | 
"int (n + m) = int n + int m"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
486  | 
"int (n - m) = int (nat (int n - int m))"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
487  | 
"int (n * m) = int n * int m"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
488  | 
"int (n div m) = int n div int m"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
489  | 
"int (n mod m) = int n mod int m"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
490  | 
"int (if P then n else m) = (if P then int n else int m)"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
491  | 
by (auto simp add: int_mult zdiv_int zmod_int)}  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
492  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
493  | 
fun mk_number_eq ctxt i lhs =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
494  | 
let  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
495  | 
      val eq = SMT_Utils.mk_cequals lhs (Numeral.mk_cnumber @{ctyp int} i)
 | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
496  | 
val ss = HOL_ss  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
497  | 
        addsimps [@{thm Nat_Numeral.int_nat_number_of}]
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
498  | 
        addsimps @{thms neg_simps}
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
499  | 
fun tac _ = Simplifier.simp_tac (Simplifier.context ctxt ss) 1  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
500  | 
in Goal.norm_result (Goal.prove_internal [] eq tac) end  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
501  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
502  | 
fun expand_head_conv cv ct =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
503  | 
(case Thm.term_of ct of  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
504  | 
_ $ _ =>  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
505  | 
Conv.fun_conv (expand_head_conv cv) then_conv  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
506  | 
Thm.beta_conversion false  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
507  | 
| _ => cv) ct  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
508  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
509  | 
fun int_conv ctxt ct =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
510  | 
(case Thm.term_of ct of  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
511  | 
      @{const of_nat (int)} $ (n as (@{const number_of (nat)} $ _)) =>
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
512  | 
Conv.rewr_conv (mk_number_eq ctxt (snd (HOLogic.dest_number n)) ct)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
513  | 
    | @{const of_nat (int)} $ _ =>
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
514  | 
(Conv.rewrs_conv ints then_conv Conv.sub_conv ints_conv ctxt) else_conv  | 
| 
41280
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41224 
diff
changeset
 | 
515  | 
Conv.sub_conv (Conv.top_sweep_conv nat_conv) ctxt  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
516  | 
| _ => Conv.no_conv) ct  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
517  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
518  | 
and ints_conv ctxt = Conv.top_sweep_conv int_conv ctxt  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
519  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
520  | 
and expand_conv ctxt =  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
521  | 
SMT_Utils.if_conv (is_nat_const o Term.head_of)  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
522  | 
(expand_head_conv (Conv.rewrs_conv expands) then_conv ints_conv ctxt)  | 
| 
41280
 
a7de9d36f4f2
only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
 
boehmes 
parents: 
41224 
diff
changeset
 | 
523  | 
(int_conv ctxt)  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
524  | 
|
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
525  | 
and nat_conv ctxt = SMT_Utils.if_exists_conv is_nat_const'  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
526  | 
(Conv.top_sweep_conv expand_conv ctxt)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
527  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
528  | 
val uses_nat_int = Term.exists_subterm (member (op aconv) nat_int_coercions)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
529  | 
in  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
530  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
531  | 
val nat_as_int_conv = nat_conv  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
532  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
533  | 
fun add_nat_embedding thms =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
534  | 
if exists (uses_nat_int o Thm.prop_of) thms then (thms, nat_embedding)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
535  | 
else (thms, [])  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
536  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
537  | 
val setup_nat_as_int =  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
538  | 
  SMT_Builtin.add_builtin_typ_ext (@{typ nat}, K true) #>
 | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
539  | 
fold (SMT_Builtin.add_builtin_fun_ext' o Term.dest_Const) builtin_nat_ops  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
540  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
541  | 
end  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
542  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
543  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
544  | 
(** normalize numerals **)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
545  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
546  | 
local  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
547  | 
(*  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
548  | 
rewrite negative numerals into positive numerals,  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
549  | 
rewrite Numeral0 into 0  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
550  | 
rewrite Numeral1 into 1  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
551  | 
*)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
552  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
553  | 
  fun is_strange_number ctxt (t as Const (@{const_name number_of}, _) $ _) =
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
554  | 
(case try HOLogic.dest_number t of  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
555  | 
SOME (_, i) => SMT_Builtin.is_builtin_num ctxt t andalso i < 2  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
556  | 
| NONE => false)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
557  | 
| is_strange_number _ _ = false  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
558  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
559  | 
val pos_num_ss = HOL_ss  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
560  | 
    addsimps [@{thm Int.number_of_minus}, @{thm Int.number_of_Min}]
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
561  | 
    addsimps [@{thm Int.number_of_Pls}, @{thm Int.numeral_1_eq_1}]
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
562  | 
    addsimps @{thms Int.pred_bin_simps}
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
563  | 
    addsimps @{thms Int.normalize_bin_simps}
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
564  | 
    addsimps @{lemma
 | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
565  | 
"Int.Min = - Int.Bit1 Int.Pls"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
566  | 
"Int.Bit0 (- Int.Pls) = - Int.Pls"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
567  | 
"Int.Bit0 (- k) = - Int.Bit0 k"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
568  | 
"Int.Bit1 (- k) = - Int.Bit1 (Int.pred k)"  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
569  | 
by simp_all (simp add: pred_def)}  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
570  | 
|
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
571  | 
fun norm_num_conv ctxt =  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
572  | 
SMT_Utils.if_conv (is_strange_number ctxt)  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
573  | 
(Simplifier.rewrite (Simplifier.context ctxt pos_num_ss)) Conv.no_conv  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
574  | 
in  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
575  | 
|
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
576  | 
fun normalize_numerals_conv ctxt =  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
577  | 
SMT_Utils.if_exists_conv (is_strange_number ctxt)  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
578  | 
(Conv.top_sweep_conv norm_num_conv ctxt)  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
579  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
580  | 
end  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
581  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
582  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
583  | 
(** combined unfoldings and rewritings **)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
584  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
585  | 
fun unfold_conv ctxt =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
586  | 
trivial_distinct_conv ctxt then_conv  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
587  | 
rewrite_bool_case_conv ctxt then_conv  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
588  | 
unfold_abs_min_max_conv ctxt then_conv  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
589  | 
nat_as_int_conv ctxt then_conv  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
590  | 
Thm.beta_conversion true  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
591  | 
|
| 
41300
 
528f5d00b542
perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
 
boehmes 
parents: 
41280 
diff
changeset
 | 
592  | 
fun unfold1 ctxt = map (apsnd (Conv.fconv_rule (unfold_conv ctxt)))  | 
| 
 
528f5d00b542
perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
 
boehmes 
parents: 
41280 
diff
changeset
 | 
593  | 
|
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
594  | 
fun burrow_ids f ithms =  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
595  | 
let  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
596  | 
val (is, thms) = split_list ithms  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
597  | 
val (thms', extra_thms) = f thms  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
598  | 
in (is ~~ thms') @ map (pair ~1) extra_thms end  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
599  | 
|
| 
41300
 
528f5d00b542
perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
 
boehmes 
parents: 
41280 
diff
changeset
 | 
600  | 
fun unfold2 ithms ctxt =  | 
| 
 
528f5d00b542
perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
 
boehmes 
parents: 
41280 
diff
changeset
 | 
601  | 
ithms  | 
| 
 
528f5d00b542
perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
 
boehmes 
parents: 
41280 
diff
changeset
 | 
602  | 
|> map (apsnd (Conv.fconv_rule (normalize_numerals_conv ctxt)))  | 
| 
 
528f5d00b542
perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
 
boehmes 
parents: 
41280 
diff
changeset
 | 
603  | 
|> burrow_ids add_nat_embedding  | 
| 
 
528f5d00b542
perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
 
boehmes 
parents: 
41280 
diff
changeset
 | 
604  | 
|> rpair ctxt  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
605  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
606  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
607  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
608  | 
(* overall normalization *)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
609  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
610  | 
type extra_norm = Proof.context -> thm list * thm list -> thm list * thm list  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
611  | 
|
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
612  | 
structure Extra_Norms = Generic_Data  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
613  | 
(  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
614  | 
type T = extra_norm SMT_Utils.dict  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
615  | 
val empty = []  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
616  | 
val extend = I  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
617  | 
fun merge data = SMT_Utils.dict_merge fst data  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
618  | 
)  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
619  | 
|
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
620  | 
fun add_extra_norm (cs, norm) =  | 
| 
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
621  | 
Extra_Norms.map (SMT_Utils.dict_update (cs, norm))  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
622  | 
|
| 
41300
 
528f5d00b542
perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
 
boehmes 
parents: 
41280 
diff
changeset
 | 
623  | 
fun apply_extra_norms ithms ctxt =  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
624  | 
let  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
625  | 
val cs = SMT_Config.solver_class_of ctxt  | 
| 
41328
 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 
boehmes 
parents: 
41327 
diff
changeset
 | 
626  | 
val es = SMT_Utils.dict_lookup (Extra_Norms.get (Context.Proof ctxt)) cs  | 
| 
41300
 
528f5d00b542
perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
 
boehmes 
parents: 
41280 
diff
changeset
 | 
627  | 
in (burrow_ids (fold (fn e => e ctxt) es o rpair []) ithms, ctxt) end  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
628  | 
|
| 
41300
 
528f5d00b542
perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
 
boehmes 
parents: 
41280 
diff
changeset
 | 
629  | 
fun normalize iwthms ctxt =  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
630  | 
iwthms  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
631  | 
|> gen_normalize ctxt  | 
| 
41300
 
528f5d00b542
perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
 
boehmes 
parents: 
41280 
diff
changeset
 | 
632  | 
|> unfold1 ctxt  | 
| 
 
528f5d00b542
perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
 
boehmes 
parents: 
41280 
diff
changeset
 | 
633  | 
|> rpair ctxt  | 
| 
 
528f5d00b542
perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
 
boehmes 
parents: 
41280 
diff
changeset
 | 
634  | 
|-> SMT_Monomorph.monomorph  | 
| 
 
528f5d00b542
perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
 
boehmes 
parents: 
41280 
diff
changeset
 | 
635  | 
|-> unfold2  | 
| 
 
528f5d00b542
perform monomorphization during normalization: schematic numerals might be monomorphized into built-in numerals and then numeral normalization is required
 
boehmes 
parents: 
41280 
diff
changeset
 | 
636  | 
|-> apply_extra_norms  | 
| 
41059
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40686 
diff
changeset
 | 
637  | 
|
| 
41072
 
9f9bc1bdacef
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
 
boehmes 
parents: 
41059 
diff
changeset
 | 
638  | 
val setup = Context.theory_map (  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
639  | 
setup_atomize #>  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
640  | 
setup_unfolded_quants #>  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
641  | 
setup_trigger #>  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
642  | 
setup_weight #>  | 
| 
41059
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40686 
diff
changeset
 | 
643  | 
setup_bool_case #>  | 
| 
41126
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
644  | 
setup_abs_min_max #>  | 
| 
 
e0bd443c0fdd
re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
 
boehmes 
parents: 
41072 
diff
changeset
 | 
645  | 
setup_nat_as_int)  | 
| 
41059
 
d2b1fc1b8e19
centralized handling of built-in types and constants;
 
boehmes 
parents: 
40686 
diff
changeset
 | 
646  | 
|
| 36898 | 647  | 
end  |