| author | Christian Urban <urbanc@in.tum.de> | 
| Mon, 17 Sep 2012 00:07:40 +0100 | |
| changeset 537 | 308ba2488d40 | 
| parent 535 | 5734ab5dd86d | 
| child 538 | e9fd5eff62c1 | 
| permissions | -rw-r--r-- | 
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1 | theory Oracle | 
| 346 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
211diff
changeset | 2 | imports "../Appendix" | 
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 3 | uses ("external_solver.ML")
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 4 | begin | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 5 | |
| 211 
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 6 | section {* Writing an Oracle (TBD)\label{rec:oracle} *} 
 | 
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 7 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 8 | text {*
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 9 |   {\bf Problem:}
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 10 | You want to use a fast, new decision procedure not based one Isabelle's | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 11 | tactics, and you do not care whether it is sound. | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 12 | \smallskip | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 13 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 14 |   {\bf Solution:} Isabelle provides the oracle mechanisms to bypass the
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 15 | inference kernel. Note that theorems proven by an oracle carry a special | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 16 | mark to inform the user of their potential incorrectness. | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 17 | \smallskip | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 18 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 19 |   \begin{readmore}
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 20 | A short introduction to oracles can be found in [isar-ref: no suitable label | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 21 | for section 3.11]. A simple example, which we will slightly extend here, | 
| 457 | 22 |   is given in @{ML_file "HOL/ex/Iff_Oracle.thy"}. The raw interface for adding
 | 
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 23 |   oracles is @{ML add_oracle in Thm} in @{ML_file "Pure/thm.ML"}.
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 24 |   \end{readmore}
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 25 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 26 | For our explanation here, we restrict ourselves to decide propositional | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 27 | formulae which consist only of equivalences between propositional variables, | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 28 |   i.e. we want to decide whether @{term "P = (Q = P) = Q"} is a tautology.
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 29 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 30 | Assume, that we have a decision procedure for such formulae, implemented | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 31 | in ML. Since we do not care how it works, we will use it here as an | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 32 | ``external solver'': | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 33 | *} | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 34 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 35 | use "external_solver.ML" | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 36 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 37 | text {*
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 38 | We do, however, know that the solver provides a function | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 39 |   @{ML IffSolver.decide}.
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 40 | It takes a string representation of a formula and returns either | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 41 |   @{ML true} if the formula is a tautology or
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 42 |   @{ML false} otherwise. The input syntax is specified as follows:
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 43 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 44 | formula $::=$ atom $\mid$ \verb|(| formula \verb|<=>| formula \verb|)| | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 45 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 46 | and all token are separated by at least one space. | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 47 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 48 | (FIXME: is there a better way for describing the syntax?) | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 49 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 50 | We will proceed in the following way. We start by translating a HOL formula | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 51 | into the string representation expected by the solver. The solver's result | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 52 | is then used to build an oracle, which we will subsequently use as a core | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 53 | for an Isar method to be able to apply the oracle in proving theorems. | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 54 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 55 | Let us start with the translation function from Isabelle propositions into | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 56 | the solver's string representation. To increase efficiency while building | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 57 |   the string, we use functions from the @{text Buffer} module.
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 58 | *} | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 59 | |
| 535 
5734ab5dd86d
adapted to new build framework
 Christian Urban <urbanc@in.tum.de> parents: 
517diff
changeset | 60 | ML %grayML{*fun translate t =
 | 
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 61 | let | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 62 | fun trans t = | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 63 | (case t of | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 64 |         @{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ t $ u =>
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 65 |           Buffer.add " (" #>
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 66 | trans t #> | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 67 | Buffer.add "<=>" #> | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 68 | trans u #> | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 69 | Buffer.add ") " | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 70 |       | Free (n, @{typ bool}) =>
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 71 | Buffer.add " " #> | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 72 | Buffer.add n #> | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 73 | Buffer.add " " | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 74 | | _ => error "inacceptable term") | 
| 535 
5734ab5dd86d
adapted to new build framework
 Christian Urban <urbanc@in.tum.de> parents: 
517diff
changeset | 75 | in Buffer.content (trans t Buffer.empty) end*} | 
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 76 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 77 | text {*
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 78 |   Here is the string representation of the term @{term "p = (q = p)"}:
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 79 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 80 |   @{ML_response 
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 81 |     "translate @{term \"p = (q = p)\"}" 
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 82 | "\" ( p <=> ( q <=> p ) ) \""} | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 83 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 84 | Let us check, what the solver returns when given a tautology: | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 85 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 86 |   @{ML_response 
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 87 |     "IffSolver.decide (translate @{term \"p = (q = p) = q\"})"
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 88 | "true"} | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 89 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 90 | And here is what it returns for a formula which is not valid: | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 91 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 92 |   @{ML_response 
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 93 |     "IffSolver.decide (translate @{term \"p = (q = p)\"})" 
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 94 | "false"} | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 95 | *} | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 96 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 97 | text {* 
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 98 | Now, we combine these functions into an oracle. In general, an oracle may | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 99 | be given any input, but it has to return a certified proposition (a | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 100 | special term which is type-checked), out of which Isabelle's inference | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 101 | kernel ``magically'' makes a theorem. | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 102 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 103 | Here, we take the proposition to be show as input. Note that we have | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 104 | to first extract the term which is then passed to the translation and | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 105 | decision procedure. If the solver finds this term to be valid, we return | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 106 | the given proposition unchanged to be turned then into a theorem: | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 107 | *} | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 108 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 109 | oracle iff_oracle = {* fn ct =>
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 110 | if IffSolver.decide (translate (HOLogic.dest_Trueprop (Thm.term_of ct))) | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 111 | then ct | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 112 | else error "Proof failed."*} | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 113 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 114 | text {*
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 115 | Here is what we get when applying the oracle: | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 116 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 117 |   @{ML_response_fake "iff_oracle @{cprop \"p = (p::bool)\"}" "p = p"}
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 118 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 119 | (FIXME: is there a better way to present the theorem?) | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 120 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 121 | To be able to use our oracle for Isar proofs, we wrap it into a tactic: | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 122 | *} | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 123 | |
| 517 
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
 Christian Urban <urbanc@in.tum.de> parents: 
457diff
changeset | 124 | ML %grayML{*val iff_oracle_tac =
 | 
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 125 | CSUBGOAL (fn (goal, i) => | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 126 | (case try iff_oracle goal of | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 127 | NONE => no_tac | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 128 | | SOME thm => rtac thm i))*} | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 129 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 130 | text {*
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 131 | and create a new method solely based on this tactic: | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 132 | *} | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 133 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 134 | method_setup iff_oracle = {*
 | 
| 181 
5baaabe1ab92
updated to new method_setup
 Christian Urban <urbanc@in.tum.de> parents: 
154diff
changeset | 135 | Scan.succeed (K (Method.SIMPLE_METHOD' iff_oracle_tac)) | 
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 136 | *} "Oracle-based decision procedure for chains of equivalences" | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 137 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 138 | text {*
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 139 | Finally, we can test our oracle to prove some theorems: | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 140 | *} | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 141 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 142 | lemma "p = (p::bool)" | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 143 | by iff_oracle | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 144 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 145 | lemma "p = (q = p) = q" | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 146 | by iff_oracle | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 147 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 148 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 149 | text {*
 | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 150 | (FIXME: say something about what the proof of the oracle is ... what do you mean?) | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 151 | *} | 
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 152 | |
| 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 153 | |
| 457 | 154 | end |