diff -r 6902638af59e -r 860c65c7388a doc-src/TutorialI/Overview/Ind.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/TutorialI/Overview/Ind.thy Fri Mar 30 16:12:57 2001 +0200 @@ -0,0 +1,101 @@ +theory Ind = Main: + +section{*Inductive Definitions*} + + +subsection{*Even Numbers*} + +subsubsection{*The Definition*} + +consts even :: "nat set" +inductive even +intros +zero[intro!]: "0 \ even" +step[intro!]: "n \ even \ Suc(Suc n) \ even" + +lemma [simp,intro!]: "2 dvd n \ 2 dvd Suc(Suc n)" +apply (unfold dvd_def) +apply clarify +apply (rule_tac x = "Suc k" in exI, simp) +done + +subsubsection{*Rule Induction*} + +thm even.induct + +lemma even_imp_dvd: "n \ even \ 2 dvd n" +apply (erule even.induct) +apply simp_all +done + +subsubsection{*Rule Inversion*} + +inductive_cases Suc_Suc_case [elim!]: "Suc(Suc n) \ even" +thm Suc_Suc_case + +lemma "Suc(Suc n) \ even \ n \ even" +by blast + + +subsection{*Mutually Inductive Definitions*} + +consts evn :: "nat set" + odd :: "nat set" + +inductive evn odd +intros +zero: "0 \ evn" +evnI: "n \ odd \ Suc n \ evn" +oddI: "n \ evn \ Suc n \ odd" + +lemma "(m \ evn \ 2 dvd m) \ (n \ odd \ 2 dvd (Suc n))" +apply(rule evn_odd.induct) +by simp_all + + +subsection{*The Reflexive Transitive Closure*} + +consts rtc :: "('a \ 'a)set \ ('a \ 'a)set" ("_*" [1000] 999) +inductive "r*" +intros +rtc_refl[iff]: "(x,x) \ r*" +rtc_step: "\ (x,y) \ r; (y,z) \ r* \ \ (x,z) \ r*" + +lemma [intro]: "(x,y) : r \ (x,y) \ r*" +by(blast intro: rtc_step); + +lemma rtc_trans: "\ (x,y) \ r*; (y,z) \ r* \ \ (x,z) \ r*" +apply(erule rtc.induct) +oops + +lemma rtc_trans[rule_format]: + "(x,y) \ r* \ (y,z) \ r* \ (x,z) \ r*" +apply(erule rtc.induct) + apply(blast); +apply(blast intro: rtc_step); +done + +text{* +\begin{exercise} +Show that the converse of @{thm[source]rtc_step} also holds: +@{prop[display]"[| (x,y) : r*; (y,z) : r |] ==> (x,z) : r*"} +\end{exercise} +*} + +subsection{*The accessible part of a relation*} + +consts + acc :: "('a \ 'a) set \ 'a set" +inductive "acc r" +intros + "(\y. (x,y) \ r \ y \ acc r) \ x \ acc r" + + +consts + accs :: "('a \ 'a) set \ 'a set" +inductive "accs r" +intros + "r``{x} \ Pow(accs r) \ x \ accs r" +monos Pow_mono + +end