| author | blanchet | 
| Fri, 01 Jul 2011 15:53:37 +0200 | |
| changeset 43625 | c3e28c869499 | 
| parent 41899 | 83dd157ec9ab | 
| child 45263 | 93ac73160d78 | 
| permissions | -rw-r--r-- | 
| 40662 | 1 | (* Title: HOL/Tools/SMT/z3_proof_methods.ML | 
| 2 | Author: Sascha Boehme, TU Muenchen | |
| 3 | ||
| 4 | Proof methods for Z3 proof reconstruction. | |
| 5 | *) | |
| 6 | ||
| 7 | signature Z3_PROOF_METHODS = | |
| 8 | sig | |
| 9 | val prove_injectivity: Proof.context -> cterm -> thm | |
| 10 | end | |
| 11 | ||
| 12 | structure Z3_Proof_Methods: Z3_PROOF_METHODS = | |
| 13 | struct | |
| 14 | ||
| 15 | ||
| 16 | fun apply tac st = | |
| 17 | (case Seq.pull (tac 1 st) of | |
| 18 |     NONE => raise THM ("tactic failed", 1, [st])
 | |
| 19 | | SOME (st', _) => st') | |
| 20 | ||
| 21 | ||
| 22 | ||
| 23 | (* injectivity *) | |
| 24 | ||
| 25 | local | |
| 26 | ||
| 27 | val B = @{typ bool}
 | |
| 28 | fun mk_univ T = Const (@{const_name top}, T --> B)
 | |
| 29 | fun mk_inj_on T U = | |
| 30 |   Const (@{const_name inj_on}, (T --> U) --> (T --> B) --> B)
 | |
| 31 | fun mk_inv_into T U = | |
| 32 |   Const (@{const_name inv_into}, [T --> B, T --> U, U] ---> T)
 | |
| 33 | ||
| 34 | fun mk_inv_of ctxt ct = | |
| 35 | let | |
| 41328 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 36 | val (dT, rT) = Term.dest_funT (SMT_Utils.typ_of ct) | 
| 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 37 | val inv = SMT_Utils.certify ctxt (mk_inv_into dT rT) | 
| 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 38 | val univ = SMT_Utils.certify ctxt (mk_univ dT) | 
| 40662 | 39 | in Thm.mk_binop inv univ ct end | 
| 40 | ||
| 41 | fun mk_inj_prop ctxt ct = | |
| 42 | let | |
| 41328 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 43 | val (dT, rT) = Term.dest_funT (SMT_Utils.typ_of ct) | 
| 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 44 | val inj = SMT_Utils.certify ctxt (mk_inj_on dT rT) | 
| 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 45 | val univ = SMT_Utils.certify ctxt (mk_univ dT) | 
| 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 46 | in SMT_Utils.mk_cprop (Thm.mk_binop inj ct univ) end | 
| 40662 | 47 | |
| 48 | ||
| 49 | val disjE = @{lemma "~P | Q ==> P ==> Q" by fast}
 | |
| 50 | ||
| 40663 | 51 | fun prove_inj_prop ctxt def lhs = | 
| 40662 | 52 | let | 
| 41328 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 53 | val (ct, ctxt') = SMT_Utils.dest_all_cabs (Thm.rhs_of def) ctxt | 
| 40662 | 54 | val rule = disjE OF [Object_Logic.rulify (Thm.assume lhs)] | 
| 55 | in | |
| 56 | Goal.init (mk_inj_prop ctxt' (Thm.dest_arg ct)) | |
| 57 |     |> apply (Tactic.rtac @{thm injI})
 | |
| 58 |     |> apply (Tactic.solve_tac [rule, rule RS @{thm sym}])
 | |
| 59 | |> Goal.norm_result o Goal.finish ctxt' | |
| 60 | |> singleton (Variable.export ctxt' ctxt) | |
| 61 | end | |
| 62 | ||
| 40663 | 63 | fun prove_rhs ctxt def lhs = | 
| 41328 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 64 | Z3_Proof_Tools.by_tac ( | 
| 40663 | 65 | CONVERSION (Conv.top_sweep_conv (K (Conv.rewr_conv def)) ctxt) | 
| 66 |     THEN' REPEAT_ALL_NEW (Tactic.match_tac @{thms allI})
 | |
| 41899 
83dd157ec9ab
finished and tested Z3 proof reconstruction for injective functions;
 boehmes parents: 
41328diff
changeset | 67 |     THEN' Tactic.rtac (@{thm inv_f_f} OF [prove_inj_prop ctxt def lhs]))
 | 
| 40662 | 68 | |
| 69 | ||
| 70 | fun expand thm ct = | |
| 71 | let | |
| 72 | val cpat = Thm.dest_arg (Thm.rhs_of thm) | |
| 73 | val (cl, cr) = Thm.dest_binop (Thm.dest_arg (Thm.dest_arg1 ct)) | |
| 74 | val thm1 = Thm.instantiate (Thm.match (cpat, cl)) thm | |
| 75 | val thm2 = Thm.instantiate (Thm.match (cpat, cr)) thm | |
| 76 | in Conv.arg_conv (Conv.binop_conv (Conv.rewrs_conv [thm1, thm2])) ct end | |
| 77 | ||
| 40663 | 78 | fun prove_lhs ctxt rhs = | 
| 40662 | 79 | let | 
| 80 | val eq = Thm.symmetric (mk_meta_eq (Object_Logic.rulify (Thm.assume rhs))) | |
| 41328 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 81 | val conv = SMT_Utils.binders_conv (K (expand eq)) ctxt | 
| 40662 | 82 | in | 
| 41328 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 83 | Z3_Proof_Tools.by_tac ( | 
| 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 84 | CONVERSION (SMT_Utils.prop_conv conv) | 
| 40663 | 85 | THEN' Simplifier.simp_tac HOL_ss) | 
| 40662 | 86 | end | 
| 87 | ||
| 88 | ||
| 40663 | 89 | fun mk_inv_def ctxt rhs = | 
| 40662 | 90 | let | 
| 41328 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 91 | val (ct, ctxt') = | 
| 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 92 | SMT_Utils.dest_all_cbinders (SMT_Utils.dest_cprop rhs) ctxt | 
| 40662 | 93 | val (cl, cv) = Thm.dest_binop ct | 
| 94 | val (cg, (cargs, cf)) = Drule.strip_comb cl ||> split_last | |
| 95 | val cu = fold_rev Thm.cabs cargs (mk_inv_of ctxt' (Thm.cabs cv cf)) | |
| 41328 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 96 | in Thm.assume (SMT_Utils.mk_cequals cg cu) end | 
| 40662 | 97 | |
| 98 | fun prove_inj_eq ctxt ct = | |
| 99 | let | |
| 41328 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 100 | val (lhs, rhs) = | 
| 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 101 | pairself SMT_Utils.mk_cprop (Thm.dest_binop (SMT_Utils.dest_cprop ct)) | 
| 41899 
83dd157ec9ab
finished and tested Z3 proof reconstruction for injective functions;
 boehmes parents: 
41328diff
changeset | 102 | val lhs_thm = Thm.implies_intr rhs (prove_lhs ctxt rhs lhs) | 
| 
83dd157ec9ab
finished and tested Z3 proof reconstruction for injective functions;
 boehmes parents: 
41328diff
changeset | 103 | val rhs_thm = | 
| 
83dd157ec9ab
finished and tested Z3 proof reconstruction for injective functions;
 boehmes parents: 
41328diff
changeset | 104 | Thm.implies_intr lhs (prove_rhs ctxt (mk_inv_def ctxt rhs) lhs rhs) | 
| 40662 | 105 |   in lhs_thm COMP (rhs_thm COMP @{thm iffI}) end
 | 
| 106 | ||
| 107 | ||
| 108 | val swap_eq_thm = mk_meta_eq @{thm eq_commute}
 | |
| 109 | val swap_disj_thm = mk_meta_eq @{thm disj_commute}
 | |
| 110 | ||
| 111 | fun swap_conv dest eq = | |
| 41328 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 112 | SMT_Utils.if_true_conv ((op <) o pairself Term.size_of_term o dest) | 
| 40662 | 113 | (Conv.rewr_conv eq) | 
| 114 | ||
| 115 | val swap_eq_conv = swap_conv HOLogic.dest_eq swap_eq_thm | |
| 41328 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 116 | val swap_disj_conv = swap_conv SMT_Utils.dest_disj swap_disj_thm | 
| 40662 | 117 | |
| 118 | fun norm_conv ctxt = | |
| 119 | swap_eq_conv then_conv | |
| 41328 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 120 | Conv.arg1_conv (SMT_Utils.binders_conv (K swap_disj_conv) ctxt) then_conv | 
| 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 121 | Conv.arg_conv (SMT_Utils.binders_conv (K swap_eq_conv) ctxt) | 
| 40662 | 122 | |
| 123 | in | |
| 124 | ||
| 40663 | 125 | fun prove_injectivity ctxt = | 
| 41328 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 126 | Z3_Proof_Tools.by_tac ( | 
| 
6792a5c92a58
avoid ML structure aliases (especially single-letter abbreviations)
 boehmes parents: 
40840diff
changeset | 127 | CONVERSION (SMT_Utils.prop_conv (norm_conv ctxt)) | 
| 40663 | 128 | THEN' CSUBGOAL (uncurry (Tactic.rtac o prove_inj_eq ctxt))) | 
| 40662 | 129 | |
| 130 | end | |
| 131 | ||
| 132 | end |