| author | bulwahn | 
| Mon, 22 Nov 2010 11:35:09 +0100 | |
| changeset 40658 | 5ccfc3ee7fe6 | 
| parent 40579 | 98ebd2300823 | 
| child 40662 | 798aad2229c0 | 
| permissions | -rw-r--r-- | 
| 36898 | 1 | (* Title: HOL/Tools/SMT/z3_proof_tools.ML | 
| 2 | Author: Sascha Boehme, TU Muenchen | |
| 3 | ||
| 4 | Helper functions required for Z3 proof reconstruction. | |
| 5 | *) | |
| 6 | ||
| 7 | signature Z3_PROOF_TOOLS = | |
| 8 | sig | |
| 9 | (* accessing and modifying terms *) | |
| 10 | val term_of: cterm -> term | |
| 11 | val prop_of: thm -> term | |
| 12 | val mk_prop: cterm -> cterm | |
| 13 | val as_meta_eq: cterm -> cterm | |
| 14 | ||
| 15 | (* theorem nets *) | |
| 40164 
57f5db2a48a3
added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
 boehmes parents: 
38864diff
changeset | 16 |   val thm_net_of: ('a -> thm) -> 'a list -> 'a Net.net
 | 
| 
57f5db2a48a3
added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
 boehmes parents: 
38864diff
changeset | 17 | val net_instance': ((thm -> thm option) -> 'a -> 'a option) -> 'a Net.net -> | 
| 
57f5db2a48a3
added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
 boehmes parents: 
38864diff
changeset | 18 | cterm -> 'a option | 
| 36898 | 19 | val net_instance: thm Net.net -> cterm -> thm option | 
| 20 | ||
| 21 | (* proof combinators *) | |
| 22 | val under_assumption: (thm -> thm) -> cterm -> thm | |
| 23 | val with_conv: conv -> (cterm -> thm) -> cterm -> thm | |
| 24 | val discharge: thm -> thm -> thm | |
| 25 | val varify: string list -> thm -> thm | |
| 26 | val unfold_eqs: Proof.context -> thm list -> conv | |
| 27 | val match_instantiate: (cterm -> cterm) -> cterm -> thm -> thm | |
| 28 | val by_tac: (int -> tactic) -> cterm -> thm | |
| 29 | val make_hyp_def: thm -> Proof.context -> thm * Proof.context | |
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 30 | val by_abstraction: bool * bool -> Proof.context -> thm list -> | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 31 | (Proof.context -> cterm -> thm) -> cterm -> thm | 
| 36898 | 32 | |
| 33 | (* a faster COMP *) | |
| 34 | type compose_data | |
| 35 | val precompose: (cterm -> cterm list) -> thm -> compose_data | |
| 36 | val precompose2: (cterm -> cterm * cterm) -> thm -> compose_data | |
| 37 | val compose: compose_data -> thm -> thm | |
| 38 | ||
| 39 | (* unfolding of 'distinct' *) | |
| 40 | val unfold_distinct_conv: conv | |
| 41 | ||
| 42 | (* simpset *) | |
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 43 | val add_simproc: Simplifier.simproc -> Context.generic -> Context.generic | 
| 36898 | 44 | val make_simpset: Proof.context -> thm list -> simpset | 
| 45 | end | |
| 46 | ||
| 47 | structure Z3_Proof_Tools: Z3_PROOF_TOOLS = | |
| 48 | struct | |
| 49 | ||
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 50 | structure I = Z3_Interface | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 51 | |
| 36898 | 52 | |
| 53 | ||
| 54 | (* accessing terms *) | |
| 55 | ||
| 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: 
40274diff
changeset | 56 | val dest_prop = (fn @{const Trueprop} $ t => t | t => t)
 | 
| 36898 | 57 | |
| 58 | fun term_of ct = dest_prop (Thm.term_of ct) | |
| 59 | fun prop_of thm = dest_prop (Thm.prop_of thm) | |
| 60 | ||
| 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: 
40274diff
changeset | 61 | val mk_prop = Thm.capply (Thm.cterm_of @{theory} @{const Trueprop})
 | 
| 36898 | 62 | |
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 63 | val eq = I.mk_inst_pair I.destT1 @{cpat "op =="}
 | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 64 | fun mk_meta_eq_cterm ct cu = Thm.mk_binop (I.instT' ct eq) ct cu | 
| 36898 | 65 | |
| 66 | fun as_meta_eq ct = uncurry mk_meta_eq_cterm (Thm.dest_binop (Thm.dest_arg ct)) | |
| 67 | ||
| 68 | ||
| 69 | ||
| 70 | (* theorem nets *) | |
| 71 | ||
| 40164 
57f5db2a48a3
added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
 boehmes parents: 
38864diff
changeset | 72 | fun thm_net_of f xthms = | 
| 
57f5db2a48a3
added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
 boehmes parents: 
38864diff
changeset | 73 | let fun insert xthm = Net.insert_term (K false) (Thm.prop_of (f xthm), xthm) | 
| 
57f5db2a48a3
added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
 boehmes parents: 
38864diff
changeset | 74 | in fold insert xthms Net.empty end | 
| 36898 | 75 | |
| 76 | fun maybe_instantiate ct thm = | |
| 77 | try Thm.first_order_match (Thm.cprop_of thm, ct) | |
| 78 | |> Option.map (fn inst => Thm.instantiate inst thm) | |
| 79 | ||
| 40164 
57f5db2a48a3
added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
 boehmes parents: 
38864diff
changeset | 80 | fun net_instance' f net ct = | 
| 
57f5db2a48a3
added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
 boehmes parents: 
38864diff
changeset | 81 | let fun first_of f xthms ct = get_first (f (maybe_instantiate ct)) xthms | 
| 
57f5db2a48a3
added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
 boehmes parents: 
38864diff
changeset | 82 | in first_of f (Net.match_term net (Thm.term_of ct)) ct end | 
| 
57f5db2a48a3
added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
 boehmes parents: 
38864diff
changeset | 83 | val net_instance = net_instance' I | 
| 36898 | 84 | |
| 85 | ||
| 86 | ||
| 87 | (* proof combinators *) | |
| 88 | ||
| 89 | fun under_assumption f ct = | |
| 90 | let val ct' = mk_prop ct | |
| 91 | in Thm.implies_intr ct' (f (Thm.assume ct')) end | |
| 92 | ||
| 93 | fun with_conv conv prove ct = | |
| 94 | let val eq = Thm.symmetric (conv ct) | |
| 95 | in Thm.equal_elim eq (prove (Thm.lhs_of eq)) end | |
| 96 | ||
| 97 | fun discharge p pq = Thm.implies_elim pq p | |
| 98 | ||
| 99 | fun varify vars = Drule.generalize ([], vars) | |
| 100 | ||
| 101 | fun unfold_eqs _ [] = Conv.all_conv | |
| 102 | | unfold_eqs ctxt eqs = | |
| 36936 
c52d1c130898
incorporated further conversions and conversionals, after some minor tuning;
 wenzelm parents: 
36899diff
changeset | 103 | Conv.top_sweep_conv (K (Conv.rewrs_conv eqs)) ctxt | 
| 36898 | 104 | |
| 105 | fun match_instantiate f ct thm = | |
| 106 | Thm.instantiate (Thm.match (f (Thm.cprop_of thm), ct)) thm | |
| 107 | ||
| 108 | fun by_tac tac ct = Goal.norm_result (Goal.prove_internal [] ct (K (tac 1))) | |
| 109 | ||
| 110 | (* |- c x == t x ==> P (c x) ~~> c == t |- P (c x) *) | |
| 111 | fun make_hyp_def thm ctxt = | |
| 112 | let | |
| 113 | val (lhs, rhs) = Thm.dest_binop (Thm.cprem_of thm 1) | |
| 114 | val (cf, cvs) = Drule.strip_comb lhs | |
| 115 | val eq = mk_meta_eq_cterm cf (fold_rev Thm.cabs cvs rhs) | |
| 116 | fun apply cv th = | |
| 117 | Thm.combination th (Thm.reflexive cv) | |
| 118 | |> Conv.fconv_rule (Conv.arg_conv (Thm.beta_conversion false)) | |
| 119 | in | |
| 120 | yield_singleton Assumption.add_assumes eq ctxt | |
| 121 | |>> Thm.implies_elim thm o fold apply cvs | |
| 122 | end | |
| 123 | ||
| 124 | ||
| 125 | ||
| 126 | (* abstraction *) | |
| 127 | ||
| 128 | local | |
| 129 | ||
| 130 | fun typ_of ct = #T (Thm.rep_cterm ct) | |
| 131 | fun certify ctxt = Thm.cterm_of (ProofContext.theory_of ctxt) | |
| 132 | ||
| 133 | fun abs_context ctxt = (ctxt, Termtab.empty, 1, false) | |
| 134 | ||
| 135 | fun context_of (ctxt, _, _, _) = ctxt | |
| 136 | ||
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 137 | fun replace (_, (cv, ct)) = Thm.forall_elim ct o Thm.forall_intr cv | 
| 36898 | 138 | |
| 139 | fun abs_instantiate (_, tab, _, beta_norm) = | |
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 140 | fold replace (Termtab.dest tab) #> | 
| 36898 | 141 | beta_norm ? Conv.fconv_rule (Thm.beta_conversion true) | 
| 142 | ||
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 143 | fun lambda_abstract cvs t = | 
| 36898 | 144 | let | 
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 145 | val frees = map Free (Term.add_frees t []) | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 146 | val cvs' = filter (fn cv => member (op aconv) frees (Thm.term_of cv)) cvs | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 147 | val vs = map (Term.dest_Free o Thm.term_of) cvs' | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 148 | in (Term.list_abs_free (vs, t), cvs') end | 
| 36898 | 149 | |
| 150 | fun fresh_abstraction cvs ct (cx as (ctxt, tab, idx, beta_norm)) = | |
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 151 | let val (t, cvs') = lambda_abstract cvs (Thm.term_of ct) | 
| 36898 | 152 | in | 
| 153 | (case Termtab.lookup tab t of | |
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 154 | SOME (cv, _) => (Drule.list_comb (cv, cvs'), cx) | 
| 36898 | 155 | | NONE => | 
| 156 | let | |
| 157 | val (n, ctxt') = yield_singleton Variable.variant_fixes "x" ctxt | |
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 158 | val cv = certify ctxt' (Free (n, map typ_of cvs' ---> typ_of ct)) | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 159 | val cu = Drule.list_comb (cv, cvs') | 
| 36898 | 160 | val e = (t, (cv, fold_rev Thm.cabs cvs' ct)) | 
| 161 | val beta_norm' = beta_norm orelse not (null cvs') | |
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 162 | in (cu, (ctxt', Termtab.update e tab, idx + 1, beta_norm')) end) | 
| 36898 | 163 | end | 
| 164 | ||
| 165 | fun abs_comb f g cvs ct = | |
| 166 | let val (cf, cu) = Thm.dest_comb ct | |
| 167 | in f cvs cf ##>> g cvs cu #>> uncurry Thm.capply end | |
| 168 | ||
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 169 | fun abs_arg f = abs_comb (K pair) f | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 170 | |
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 171 | fun abs_args f cvs ct = | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 172 | (case Thm.term_of ct of | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 173 | _ $ _ => abs_comb (abs_args f) f cvs ct | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 174 | | _ => pair ct) | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 175 | |
| 36898 | 176 | fun abs_list f g cvs ct = | 
| 177 | (case Thm.term_of ct of | |
| 178 |     Const (@{const_name Nil}, _) => pair ct
 | |
| 179 |   | Const (@{const_name Cons}, _) $ _ $ _ =>
 | |
| 180 | abs_comb (abs_arg f) (abs_list f g) cvs ct | |
| 181 | | _ => g cvs ct) | |
| 182 | ||
| 183 | fun abs_abs f cvs ct = | |
| 184 | let val (cv, cu) = Thm.dest_abs NONE ct | |
| 185 | in f (cv :: cvs) cu #>> Thm.cabs cv end | |
| 186 | ||
| 187 | val is_atomic = (fn _ $ _ => false | Abs _ => false | _ => true) | |
| 188 | ||
| 189 | fun abstract (ext_logic, with_theories) = | |
| 190 | let | |
| 191 | fun abstr1 cvs ct = abs_arg abstr cvs ct | |
| 192 | and abstr2 cvs ct = abs_comb abstr1 abstr cvs ct | |
| 193 | and abstr3 cvs ct = abs_comb abstr2 abstr cvs ct | |
| 194 | and abstr_abs cvs ct = abs_arg (abs_abs abstr) cvs ct | |
| 195 | ||
| 196 | and abstr cvs ct = | |
| 197 | (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: 
40274diff
changeset | 198 |         @{const Trueprop} $ _ => abstr1 cvs ct
 | 
| 
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: 
40274diff
changeset | 199 |       | @{const "==>"} $ _ $ _ => abstr2 cvs ct
 | 
| 
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: 
40274diff
changeset | 200 |       | @{const True} => pair ct
 | 
| 
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: 
40274diff
changeset | 201 |       | @{const False} => pair ct
 | 
| 
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: 
40274diff
changeset | 202 |       | @{const Not} $ _ => abstr1 cvs ct
 | 
| 
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: 
40274diff
changeset | 203 |       | @{const HOL.conj} $ _ $ _ => abstr2 cvs ct
 | 
| 
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: 
40274diff
changeset | 204 |       | @{const HOL.disj} $ _ $ _ => abstr2 cvs ct
 | 
| 
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: 
40274diff
changeset | 205 |       | @{const HOL.implies} $ _ $ _ => abstr2 cvs ct
 | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 206 |       | Const (@{const_name HOL.eq}, _) $ _ $ _ => abstr2 cvs ct
 | 
| 40274 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40164diff
changeset | 207 |       | Const (@{const_name SMT.distinct}, _) $ _ =>
 | 
| 36898 | 208 | if ext_logic then abs_arg (abs_list abstr fresh_abstraction) cvs ct | 
| 209 | else fresh_abstraction cvs ct | |
| 210 |       | Const (@{const_name If}, _) $ _ $ _ $ _ =>
 | |
| 211 | if ext_logic then abstr3 cvs ct else fresh_abstraction cvs ct | |
| 212 |       | Const (@{const_name All}, _) $ _ =>
 | |
| 213 | if ext_logic then abstr_abs cvs ct else fresh_abstraction cvs ct | |
| 214 |       | Const (@{const_name Ex}, _) $ _ =>
 | |
| 215 | if ext_logic then abstr_abs cvs ct else fresh_abstraction cvs ct | |
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 216 | | t => (fn cx => | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 217 | if is_atomic t orelse can HOLogic.dest_number t then (ct, cx) | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 218 | else if with_theories andalso | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 219 | I.is_builtin_theory_term (context_of cx) t | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 220 | then abs_args abstr cvs ct cx | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 221 | else fresh_abstraction cvs ct cx)) | 
| 36898 | 222 | in abstr [] end | 
| 223 | ||
| 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: 
40274diff
changeset | 224 | val cimp = Thm.cterm_of @{theory} @{const "==>"}
 | 
| 
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: 
40274diff
changeset | 225 | |
| 36898 | 226 | fun with_prems thms f ct = | 
| 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: 
40274diff
changeset | 227 | fold_rev (Thm.mk_binop cimp o Thm.cprop_of) thms ct | 
| 36898 | 228 | |> f | 
| 229 | |> fold (fn prem => fn th => Thm.implies_elim th prem) thms | |
| 230 | ||
| 231 | in | |
| 232 | ||
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 233 | fun by_abstraction mode ctxt thms prove = with_prems thms (fn ct => | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 234 | let val (cu, cx) = abstract mode ct (abs_context ctxt) | 
| 36898 | 235 | in abs_instantiate cx (prove (context_of cx) cu) end) | 
| 236 | ||
| 237 | end | |
| 238 | ||
| 239 | ||
| 240 | ||
| 241 | (* a faster COMP *) | |
| 242 | ||
| 243 | type compose_data = cterm list * (cterm -> cterm list) * thm | |
| 244 | ||
| 245 | fun list2 (x, y) = [x, y] | |
| 246 | ||
| 247 | fun precompose f rule = (f (Thm.cprem_of rule 1), f, rule) | |
| 248 | fun precompose2 f rule = precompose (list2 o f) rule | |
| 249 | ||
| 250 | fun compose (cvs, f, rule) thm = | |
| 251 | discharge thm (Thm.instantiate ([], cvs ~~ f (Thm.cprop_of thm)) rule) | |
| 252 | ||
| 253 | ||
| 254 | ||
| 255 | (* unfolding of 'distinct' *) | |
| 256 | ||
| 257 | local | |
| 258 |   val set1 = @{lemma "x ~: set [] == ~False" by simp}
 | |
| 259 |   val set2 = @{lemma "x ~: set [x] == False" by simp}
 | |
| 260 |   val set3 = @{lemma "x ~: set [y] == x ~= y" by simp}
 | |
| 261 |   val set4 = @{lemma "x ~: set (x # ys) == False" by simp}
 | |
| 262 |   val set5 = @{lemma "x ~: set (y # ys) == x ~= y & x ~: set ys" by simp}
 | |
| 263 | ||
| 264 | fun set_conv ct = | |
| 36936 
c52d1c130898
incorporated further conversions and conversionals, after some minor tuning;
 wenzelm parents: 
36899diff
changeset | 265 | (Conv.rewrs_conv [set1, set2, set3, set4] else_conv | 
| 36898 | 266 | (Conv.rewr_conv set5 then_conv Conv.arg_conv set_conv)) ct | 
| 267 | ||
| 40274 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40164diff
changeset | 268 |   val dist1 = @{lemma "SMT.distinct [] == ~False" by (simp add: distinct_def)}
 | 
| 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40164diff
changeset | 269 |   val dist2 = @{lemma "SMT.distinct [x] == ~False" by (simp add: distinct_def)}
 | 
| 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40164diff
changeset | 270 |   val dist3 = @{lemma "SMT.distinct (x # xs) == x ~: set xs & distinct xs"
 | 
| 
6486c610a549
introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
 boehmes parents: 
40164diff
changeset | 271 | by (simp add: distinct_def)} | 
| 36898 | 272 | |
| 273 | fun binop_conv cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2 | |
| 274 | in | |
| 275 | fun unfold_distinct_conv ct = | |
| 36936 
c52d1c130898
incorporated further conversions and conversionals, after some minor tuning;
 wenzelm parents: 
36899diff
changeset | 276 | (Conv.rewrs_conv [dist1, dist2] else_conv | 
| 36898 | 277 | (Conv.rewr_conv dist3 then_conv binop_conv set_conv unfold_distinct_conv)) ct | 
| 278 | end | |
| 279 | ||
| 280 | ||
| 281 | ||
| 282 | (* simpset *) | |
| 283 | ||
| 284 | local | |
| 285 |   val antisym_le1 = mk_meta_eq @{thm order_class.antisym_conv}
 | |
| 286 |   val antisym_le2 = mk_meta_eq @{thm linorder_class.antisym_conv2}
 | |
| 287 |   val antisym_less1 = mk_meta_eq @{thm linorder_class.antisym_conv1}
 | |
| 288 |   val antisym_less2 = mk_meta_eq @{thm linorder_class.antisym_conv3}
 | |
| 289 | ||
| 290 | fun eq_prop t thm = HOLogic.mk_Trueprop t aconv Thm.prop_of thm | |
| 291 | fun dest_binop ((c as Const _) $ t $ u) = (c, t, u) | |
| 292 |     | dest_binop t = raise TERM ("dest_binop", [t])
 | |
| 293 | ||
| 294 | fun prove_antisym_le ss t = | |
| 295 | let | |
| 296 | val (le, r, s) = dest_binop t | |
| 297 |       val less = Const (@{const_name less}, Term.fastype_of le)
 | |
| 298 | val prems = Simplifier.prems_of_ss ss | |
| 299 | in | |
| 300 | (case find_first (eq_prop (le $ s $ r)) prems of | |
| 301 | NONE => | |
| 302 | find_first (eq_prop (HOLogic.mk_not (less $ r $ s))) prems | |
| 303 | |> Option.map (fn thm => thm RS antisym_less1) | |
| 304 | | SOME thm => SOME (thm RS antisym_le1)) | |
| 305 | end | |
| 306 | handle THM _ => NONE | |
| 307 | ||
| 308 | fun prove_antisym_less ss t = | |
| 309 | let | |
| 310 | val (less, r, s) = dest_binop (HOLogic.dest_not t) | |
| 311 |       val le = Const (@{const_name less_eq}, Term.fastype_of less)
 | |
| 312 | val prems = prems_of_ss ss | |
| 313 | in | |
| 314 | (case find_first (eq_prop (le $ r $ s)) prems of | |
| 315 | NONE => | |
| 316 | find_first (eq_prop (HOLogic.mk_not (less $ s $ r))) prems | |
| 317 | |> Option.map (fn thm => thm RS antisym_less2) | |
| 318 | | SOME thm => SOME (thm RS antisym_le2)) | |
| 319 | end | |
| 320 | handle THM _ => NONE | |
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 321 | |
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 322 |   val basic_simpset = HOL_ss addsimps @{thms field_simps}
 | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 323 |     addsimps [@{thm times_divide_eq_right}, @{thm times_divide_eq_left}]
 | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 324 |     addsimps @{thms arith_special} addsimps @{thms less_bin_simps}
 | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 325 |     addsimps @{thms le_bin_simps} addsimps @{thms eq_bin_simps}
 | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 326 |     addsimps @{thms add_bin_simps} addsimps @{thms succ_bin_simps}
 | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 327 |     addsimps @{thms minus_bin_simps} addsimps @{thms pred_bin_simps}
 | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 328 |     addsimps @{thms mult_bin_simps} addsimps @{thms iszero_simps}
 | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 329 |     addsimps @{thms array_rules}
 | 
| 37151 | 330 |     addsimps @{thms z3div_def} addsimps @{thms z3mod_def}
 | 
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 331 | addsimprocs [ | 
| 38715 
6513ea67d95d
renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
 wenzelm parents: 
37151diff
changeset | 332 |       Simplifier.simproc_global @{theory} "fast_int_arith" [
 | 
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 333 | "(m::int) < n", "(m::int) <= n", "(m::int) = n"] (K Lin_Arith.simproc), | 
| 38715 
6513ea67d95d
renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
 wenzelm parents: 
37151diff
changeset | 334 |       Simplifier.simproc_global @{theory} "antisym_le" ["(x::'a::order) <= y"]
 | 
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 335 | (K prove_antisym_le), | 
| 38715 
6513ea67d95d
renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
 wenzelm parents: 
37151diff
changeset | 336 |       Simplifier.simproc_global @{theory} "antisym_less" ["~ (x::'a::linorder) < y"]
 | 
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 337 | (K prove_antisym_less)] | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 338 | |
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 339 | structure Simpset = Generic_Data | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 340 | ( | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 341 | type T = simpset | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 342 | val empty = basic_simpset | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 343 | val extend = I | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 344 | val merge = Simplifier.merge_ss | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 345 | ) | 
| 36898 | 346 | in | 
| 347 | ||
| 36899 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 348 | fun add_simproc simproc = Simpset.map (fn ss => ss addsimprocs [simproc]) | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 349 | |
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 350 | fun make_simpset ctxt rules = | 
| 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 boehmes parents: 
36898diff
changeset | 351 | Simplifier.context ctxt (Simpset.get (Context.Proof ctxt)) addsimps rules | 
| 36898 | 352 | |
| 353 | end | |
| 354 | ||
| 355 | end |