| 17914 |      1 | (*<*)theory Mutual imports Main begin(*>*)
 | 
| 10762 |      2 | 
 | 
| 10884 |      3 | subsection{*Mutually Inductive Definitions*}
 | 
| 10762 |      4 | 
 | 
|  |      5 | text{*
 | 
|  |      6 | Just as there are datatypes defined by mutual recursion, there are sets defined
 | 
| 10790 |      7 | by mutual induction. As a trivial example we consider the even and odd
 | 
|  |      8 | natural numbers:
 | 
| 10762 |      9 | *}
 | 
|  |     10 | 
 | 
| 23733 |     11 | inductive_set
 | 
|  |     12 |   Even :: "nat set"
 | 
|  |     13 |   and Odd  :: "nat set"
 | 
|  |     14 | where
 | 
|  |     15 |   zero:  "0 \<in> Even"
 | 
|  |     16 | | EvenI: "n \<in> Odd \<Longrightarrow> Suc n \<in> Even"
 | 
|  |     17 | | OddI:  "n \<in> Even \<Longrightarrow> Suc n \<in> Odd"
 | 
| 10762 |     18 | 
 | 
|  |     19 | text{*\noindent
 | 
| 10790 |     20 | The mutually inductive definition of multiple sets is no different from
 | 
|  |     21 | that of a single set, except for induction: just as for mutually recursive
 | 
|  |     22 | datatypes, induction needs to involve all the simultaneously defined sets. In
 | 
| 19389 |     23 | the above case, the induction rule is called @{thm[source]Even_Odd.induct}
 | 
| 10790 |     24 | (simply concatenate the names of the sets involved) and has the conclusion
 | 
| 19389 |     25 | @{text[display]"(?x \<in> Even \<longrightarrow> ?P ?x) \<and> (?y \<in> Odd \<longrightarrow> ?Q ?y)"}
 | 
| 10762 |     26 | 
 | 
| 11494 |     27 | If we want to prove that all even numbers are divisible by two, we have to
 | 
| 10790 |     28 | generalize the statement as follows:
 | 
| 10762 |     29 | *}
 | 
|  |     30 | 
 | 
| 19389 |     31 | lemma "(m \<in> Even \<longrightarrow> 2 dvd m) \<and> (n \<in> Odd \<longrightarrow> 2 dvd (Suc n))"
 | 
| 10762 |     32 | 
 | 
|  |     33 | txt{*\noindent
 | 
| 10790 |     34 | The proof is by rule induction. Because of the form of the induction theorem,
 | 
|  |     35 | it is applied by @{text rule} rather than @{text erule} as for ordinary
 | 
|  |     36 | inductive definitions:
 | 
| 10762 |     37 | *}
 | 
|  |     38 | 
 | 
| 19389 |     39 | apply(rule Even_Odd.induct)
 | 
| 10762 |     40 | 
 | 
|  |     41 | txt{*
 | 
|  |     42 | @{subgoals[display,indent=0]}
 | 
|  |     43 | The first two subgoals are proved by simplification and the final one can be
 | 
|  |     44 | proved in the same manner as in \S\ref{sec:rule-induction}
 | 
|  |     45 | where the same subgoal was encountered before.
 | 
|  |     46 | We do not show the proof script.
 | 
|  |     47 | *}
 | 
|  |     48 | (*<*)
 | 
|  |     49 |   apply simp
 | 
|  |     50 |  apply simp
 | 
| 12815 |     51 | apply(simp add: dvd_def)
 | 
| 10762 |     52 | apply(clarify)
 | 
|  |     53 | apply(rule_tac x = "Suc k" in exI)
 | 
|  |     54 | apply simp
 | 
|  |     55 | done
 | 
|  |     56 | (*>*)
 | 
|  |     57 | (*
 | 
|  |     58 | Exercise: 1 : odd
 | 
|  |     59 | *)
 | 
| 10790 |     60 | (*<*)end(*>*)
 |