author | bulwahn |
Mon, 30 Jan 2012 13:55:20 +0100 | |
changeset 46357 | 2795607a1f40 |
parent 46356 | 48fcca8965f4 |
child 46358 | b2a936486685 |
src/HOL/Enum.thy | file | annotate | diff | comparison | revisions |
--- a/src/HOL/Enum.thy Mon Jan 30 13:55:19 2012 +0100 +++ b/src/HOL/Enum.thy Mon Jan 30 13:55:20 2012 +0100 @@ -787,6 +787,9 @@ "Collect P = set (filter P enum)" by (auto simp add: enum_UNIV) +lemma tranclp_unfold [code, no_atp]: + "tranclp r a b \<equiv> (a, b) \<in> trancl {(x, y). r x y}" +by (simp add: trancl_def) subsection {* Executable accessible part *} (* FIXME: should be moved somewhere else !? *)