diff -r 504d72a39638 -r 11fce8564415 src/HOL/IMP/Star.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/IMP/Star.thy Wed Jun 01 21:35:34 2011 +0200 @@ -0,0 +1,27 @@ +theory Star imports Main +begin + +inductive + star :: "('a \ 'a \ bool) \ 'a \ 'a \ bool" +for r where +refl: "star r x x" | +step: "r x y \ star r y z \ star r x z" + +lemma star_trans: + "star r x y \ star r y z \ star r x z" +proof(induct rule: star.induct) + case refl thus ?case . +next + case step thus ?case by (metis star.step) +qed + +lemmas star_induct = star.induct[of "r:: 'a*'b \ 'a*'b \ bool", split_format(complete)] + +declare star.refl[simp,intro] + +lemma step1[simp, intro]: "r x y \ star r x y" +by(metis refl step) + +code_pred star . + +end