diff -r 7263c856787e -r b9fd52525b69 doc-src/TutorialI/Inductive/Star.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/TutorialI/Inductive/Star.thy Mon Oct 16 13:21:01 2000 +0200 @@ -0,0 +1,56 @@ +(*<*)theory Star = Main:(*>*) + +section{*The transitive and reflexive closure*} + +text{* +A perfect example of an inductive definition is the transitive, reflexive +closure of a relation. This concept was already introduced in +\S\ref{sec:rtrancl}, but it was not shown how it is defined. In fact, +the operator @{text"^*"} is not defined inductively but via a least +fixpoint because at that point in the theory hierarchy +inductive definitions are not yet available. But now they are: +*} + +consts trc :: "('a \ 'a)set \ ('a \ 'a)set" ("_*" [1000] 999) +inductive "r*" +intros +trc_refl[intro!]: "(x,x) \ r*" +trc_step: "\ (x,y) \ r*; (y,z) \ r \ \ (x,z) \ r*" + +lemma [intro]: "(x,y) : r \ (x,y) \ r*" +by(blast intro: trc_step); + +lemma step2[rule_format]: + "(y,z) \ r* \ (x,y) \ r \ (x,z) \ r*" +apply(erule trc.induct) + apply(blast); +apply(blast intro: trc_step); +done + +lemma trc_trans[rule_format]: + "(x,y) \ r* \ \z. (y,z) \ r* \ (x,z) \ r*" +apply(erule trc.induct) + apply blast; +apply(blast intro: step2); +done + +consts trc2 :: "('a \ 'a)set \ ('a \ 'a)set" +inductive "trc2 r" +intros +"(x,y) \ r \ (x,y) \ trc2 r" +"(x,x) \ trc2 r" +"\ (x,y) \ trc2 r; (y,z) \ trc2 r \ \ (x,z) \ trc2 r" + + +lemma "((x,y) \ trc2 r) = ((x,y) \ r*)" +apply(rule iffI); + apply(erule trc2.induct); + apply(blast); + apply(blast); + apply(blast intro: trc_trans); +apply(erule trc.induct); + apply(blast intro: trc2.intros); +apply(blast intro: trc2.intros); +done + +(*<*)end(*>*)