| 8745 |      1 | (*<*)
 | 
|  |      2 | theory def_rewr = Main:;
 | 
|  |      3 | 
 | 
|  |      4 | (*>*)text{*\noindent Constant definitions (\S\ref{sec:ConstDefinitions}) can
 | 
|  |      5 | be used as simplification rules, but by default they are not.  Hence the
 | 
|  |      6 | simplifier does not expand them automatically, just as it should be:
 | 
|  |      7 | definitions are introduced for the purpose of abbreviating complex
 | 
|  |      8 | concepts. Of course we need to expand the definitions initially to derive
 | 
|  |      9 | enough lemmas that characterize the concept sufficiently for us to forget the
 | 
|  |     10 | original definition. For example, given
 | 
|  |     11 | *}
 | 
|  |     12 | 
 | 
|  |     13 | constdefs exor :: "bool \\<Rightarrow> bool \\<Rightarrow> bool"
 | 
|  |     14 |          "exor A B \\<equiv> (A \\<and> \\<not>B) \\<or> (\\<not>A \\<and> B)";
 | 
|  |     15 | 
 | 
|  |     16 | text{*\noindent
 | 
|  |     17 | we may want to prove
 | 
|  |     18 | *}
 | 
|  |     19 | 
 | 
|  |     20 | lemma "exor A (\\<not>A)";
 | 
|  |     21 | 
 | 
|  |     22 | txt{*\noindent
 | 
|  |     23 | There is a special method for \emph{unfolding} definitions, which we need to
 | 
|  |     24 | get started:\indexbold{*unfold}\indexbold{definition!unfolding}
 | 
|  |     25 | *}
 | 
|  |     26 | 
 | 
|  |     27 | apply(unfold exor_def);
 | 
|  |     28 | 
 | 
|  |     29 | txt{*\noindent
 | 
|  |     30 | It unfolds the given list of definitions (here merely one) in all subgoals:
 | 
|  |     31 | \begin{isabellepar}%
 | 
|  |     32 | ~1.~A~{\isasymand}~{\isasymnot}~{\isasymnot}~A~{\isasymor}~{\isasymnot}~A~{\isasymand}~{\isasymnot}~A%
 | 
|  |     33 | \end{isabellepar}%
 | 
|  |     34 | The resulting goal can be proved by simplification.
 | 
|  |     35 | 
 | 
|  |     36 | In case we want to expand a definition in the middle of a proof, we can
 | 
|  |     37 | simply include it locally:*}(*<*)oops;lemma "exor A (\\<not>A)";(*>*)
 | 
|  |     38 | apply(simp add: exor_def);(*<*).(*>*)
 | 
|  |     39 | 
 | 
|  |     40 | text{*\noindent
 | 
|  |     41 | In fact, this one command proves the above lemma directly.
 | 
|  |     42 | 
 | 
|  |     43 | You should normally not turn a definition permanently into a simplification
 | 
|  |     44 | rule because this defeats the whole purpose of an abbreviation.
 | 
|  |     45 | *}
 | 
|  |     46 | (*<*)end(*>*)
 |